X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakToCore.v;h=b6add9448794403c5c9de6c77baff31dfdd1447c;hp=f0a8300886a15eb82d1d717982d37e7dcb693d92;hb=635ee434c9edbad1bc6c9bf5ba2b91cb8c51be8e;hpb=bcb16a7fa1ff772f12807c4587609fd756b7762e diff --git a/src/HaskWeakToCore.v b/src/HaskWeakToCore.v index f0a8300..b6add94 100644 --- a/src/HaskWeakToCore.v +++ b/src/HaskWeakToCore.v @@ -21,32 +21,20 @@ Variable mkCoreLet : @CoreBind CoreVar -> @CoreExpr CoreVar -> @CoreExpr CoreVar Extract Inlined Constant mkCoreLet => "MkCore.mkCoreLet". Variable sortAlts : forall {a}{b}, list (@triple AltCon a b) -> list (@triple AltCon a b). - Extract Inlined Constant mkCoreLet => "sortAlts". + Extract Inlined Constant sortAlts => "sortAlts". Implicit Arguments sortAlts [[a][b]]. -Variable trustMeCoercion : CoreCoercion. - Extract Inlined Constant trustMeCoercion => "Coercion.unsafeCoerce". +Variable mkUnsafeCoercion : CoreType -> CoreType -> CoreCoercion. + Extract Inlined Constant mkUnsafeCoercion => "Coercion.mkUnsafeCoercion". (* Coercion and Type are actually the same thing in GHC, but we don't tell Coq about that. This lets us get around it. *) Variable coreCoercionsAreReallyTypes : CoreCoercion -> CoreType. Extract Inlined Constant coreCoercionsAreReallyTypes => "(\x -> x)". -(* a dummy variable which is never referenced but needed for a binder *) -Variable dummyVariable : CoreVar. - (* FIXME this doesn't actually work *) - Extract Inlined Constant dummyVariable => "(Prelude.error ""dummyVariable"")". +Definition weakCoercionToCoreCoercion (wc:WeakCoercion) : CoreCoercion := + mkUnsafeCoercion (weakTypeToCoreType (fst (weakCoercionTypes wc))) (weakTypeToCoreType (snd (weakCoercionTypes wc))). -Section HaskWeakToCore. - - (* the CoreVar for the "magic" bracket/escape identifiers must be created from within one of the monads *) - Context (hetmet_brak_var : CoreVar). - Context (hetmet_esc_var : CoreVar). - - (* FIXME: do something more intelligent here *) - Definition weakCoercionToCoreCoercion : WeakCoercion -> CoreCoercion := - fun _ => trustMeCoercion. - - Fixpoint weakExprToCoreExpr (me:WeakExpr) : @CoreExpr CoreVar := +Fixpoint weakExprToCoreExpr (me:WeakExpr) : @CoreExpr CoreVar := match me with | WEVar (weakExprVar v _) => CoreEVar v | WELit lit => CoreELit lit @@ -57,14 +45,23 @@ Section HaskWeakToCore. | WENote n e => CoreENote n (weakExprToCoreExpr e ) | WELam (weakExprVar ev _ ) e => CoreELam ev (weakExprToCoreExpr e ) | WETyLam (weakTypeVar tv _ ) e => CoreELam tv (weakExprToCoreExpr e ) - | WECoLam (weakCoerVar cv _ _) e => CoreELam cv (weakExprToCoreExpr e ) + | WECoLam (weakCoerVar cv _ _ _) e => CoreELam cv (weakExprToCoreExpr e ) | WECast e co => CoreECast (weakExprToCoreExpr e ) (weakCoercionToCoreCoercion co) - | WEBrak (weakTypeVar ec _) e t => CoreEApp (CoreEApp (CoreEVar hetmet_brak_var) - (CoreEType (TyVarTy ec))) (CoreEType (weakTypeToCoreType t)) - | WEEsc (weakTypeVar ec _) e t => CoreEApp (CoreEApp (CoreEVar hetmet_esc_var) - (CoreEType (TyVarTy ec))) (CoreEType (weakTypeToCoreType t)) + | WEBrak v (weakTypeVar ec _) e t => fold_left CoreEApp + ((CoreEType (TyVarTy ec)):: + (CoreEType (weakTypeToCoreType t)):: + (weakExprToCoreExpr e):: + nil) + (CoreEVar v) + | WEEsc v (weakTypeVar ec _) e t => fold_left CoreEApp + ((CoreEType (TyVarTy ec)):: + (CoreEType (weakTypeToCoreType t)):: + (weakExprToCoreExpr e):: + nil) + (CoreEVar v) | WELet (weakExprVar v _) ve e => mkCoreLet (CoreNonRec v (weakExprToCoreExpr ve)) (weakExprToCoreExpr e) - | WECase e tbranches tc types alts => CoreECase (weakExprToCoreExpr e) dummyVariable (weakTypeToCoreType tbranches) + | WECase vscrut e tbranches tc types alts => + CoreECase (weakExprToCoreExpr e) vscrut (weakTypeToCoreType tbranches) (sortAlts (( fix mkCaseBranches (alts:Tree ??(AltCon*list WeakTypeVar*list WeakCoerVar*list WeakExprVar*WeakExpr)) := @@ -90,6 +87,9 @@ Section HaskWeakToCore. (weakExprToCoreExpr e) end. -End HaskWeakToCore. +Definition weakTypeOfWeakExpr (we:WeakExpr) : ???WeakType := + coreTypeToWeakType (coreTypeOfCoreExpr (weakExprToCoreExpr we)). +Instance weakExprToString : ToString WeakExpr := + { toString := fun we => toString (weakExprToCoreExpr we) }.