X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;f=src%2FHaskCoreToWeak.v;h=1cbaf228293a9bb6c296ce7efc69523ecf3a2292;hb=48bc98e014fd0c21ca75017bf689e8e6e80f54e3;hp=a812483bc7ded9f0e91dddeb8c1188e55628ac23;hpb=ee5aaad57d76400e9b8736d4a12d2804f99f329c;p=coq-hetmet.git diff --git a/src/HaskCoreToWeak.v b/src/HaskCoreToWeak.v index a812483..1cbaf22 100644 --- a/src/HaskCoreToWeak.v +++ b/src/HaskCoreToWeak.v @@ -43,7 +43,7 @@ Fixpoint coreTypeToWeakType' (ct:CoreType) : ???WeakType := | WExprVar _ => Error "encountered expression variable in a modal box type" | WCoerVar _ => Error "encountered coercion variable in a modal box type" end - | _ => Error "mis-applied modal box tycon" + | _ => Error ("mis-applied modal box tycon: " +++ ct) end else let tc' := if eqd_dec tc ArrowTyCon then WFunTyCon @@ -98,11 +98,14 @@ match ce with | _ => None end. -Definition coreCoercionToWeakCoercion (cc:CoreCoercion) : ???WeakCoercion := +Variable coreCoercionToWeakCoercion : CoreCoercion -> WeakCoercion. + Extract Inlined Constant coreCoercionToWeakCoercion => "coreCoercionToWeakCoercion". +(* let (t1,t2) := coreCoercionKind cc in coreTypeToWeakType t1 >>= fun t1' => - coreTypeToWeakType t2 >>= fun t2' => - OK (weakCoercion (kindOfCoreType t1) t1' t2' cc). + coreTypeToWeakType t2 >>= fun t2' => + OK (WCoUnsafe t1' t2'). +*) Fixpoint coreExprToWeakExpr (ce:@CoreExpr CoreVar) : ???WeakExpr := match ce with @@ -110,8 +113,7 @@ Fixpoint coreExprToWeakExpr (ce:@CoreExpr CoreVar) : ???WeakExpr := | CoreENote n e => coreExprToWeakExpr e >>= fun e' => OK (WENote n e') | CoreEType t => Error "encountered CoreEType in a position where an Expr should have been" | CoreECast e co => coreExprToWeakExpr e >>= fun e' => - coreCoercionToWeakCoercion co >>= fun co' => - OK (WECast e' co') + OK (WECast e' (coreCoercionToWeakCoercion co)) | CoreEVar v => match coreVarToWeakVar v with | WExprVar ev => OK (WEVar ev) @@ -174,8 +176,8 @@ Fixpoint coreExprToWeakExpr (ce:@CoreExpr CoreVar) : ???WeakExpr := | WTypeVar _ => Error "found a type variable in a case" | WCoerVar _ => Error "found a coercion variable in a case" | WExprVar ev => - coreExprToWeakExpr e >>= fun e' => - weakTypeOfWeakExpr e' >>= fun te' => + coreTypeToWeakType (coreTypeOfCoreExpr e) >>= fun te' => + coreExprToWeakExpr e >>= fun e' => (match isTyConApp te' nil with None => Error "expectTyConApp encountered strange type" | Some x => OK x end) >>= fun tca => let (tc,lt) := tca in