X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakToCore.v;h=c3e90a43763c820af3ea4c561b465986d05a73ad;hp=c97a63cef20a810160a3b07d51e277358f434939;hb=5cb97fa6ed28f55ca888bdadc4f145396cc02236;hpb=14a87dd821c4194382f29eef2d59fe932d4124c1 diff --git a/src/HaskWeakToCore.v b/src/HaskWeakToCore.v index c97a63c..c3e90a4 100644 --- a/src/HaskWeakToCore.v +++ b/src/HaskWeakToCore.v @@ -8,7 +8,8 @@ Require Import General. Require Import Coq.Strings.String. Require Import Coq.Lists.List. Require Import HaskKinds. -Require Import HaskLiteralsAndTyCons. +Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskCoreVars. Require Import HaskCoreTypes. Require Import HaskCore. @@ -24,16 +25,6 @@ Variable sortAlts : forall {a}{b}, list (@triple CoreAltCon a b) -> list (@trip Extract Inlined Constant sortAlts => "sortAlts". Implicit Arguments sortAlts [[a][b]]. -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)". - -Variable ModalBoxTyCon : TyCon. Extract Inlined Constant ModalBoxTyCon => "TysWiredIn.hetMetCodeTypeTyCon". -Variable ArrowTyCon : TyCon. Extract Constant ArrowTyCon => "Type.funTyCon". - Definition weakAltConToCoreAltCon (wa:WeakAltCon) : CoreAltCon := match wa with | WeakDataAlt cdc => DataAlt cdc @@ -61,7 +52,7 @@ Fixpoint weakTypeToCoreType (wt:WeakType) : CoreType := end. Definition weakCoercionToCoreCoercion (wc:WeakCoercion) : CoreCoercion := - mkUnsafeCoercion (weakTypeToCoreType (fst (weakCoercionTypes wc))) (weakTypeToCoreType (snd (weakCoercionTypes wc))). + CoreCoercionUnsafeCo (weakTypeToCoreType (fst (weakCoercionTypes wc))) (weakTypeToCoreType (snd (weakCoercionTypes wc))). Fixpoint weakExprToCoreExpr (me:WeakExpr) : @CoreExpr CoreVar := match me with @@ -69,12 +60,11 @@ Fixpoint weakExprToCoreExpr (me:WeakExpr) : @CoreExpr CoreVar := | WELit lit => CoreELit lit | WEApp e1 e2 => CoreEApp (weakExprToCoreExpr e1) (weakExprToCoreExpr e2) | WETyApp e t => CoreEApp (weakExprToCoreExpr e ) (CoreEType (weakTypeToCoreType t)) - | WECoApp e co => CoreEApp (weakExprToCoreExpr e ) - (CoreEType (coreCoercionsAreReallyTypes (weakCoercionToCoreCoercion co))) + | WECoApp e co => CoreEApp (weakExprToCoreExpr e ) (CoreECoercion (weakCoercionToCoreCoercion co)) | 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 v (weakTypeVar ec _) e t => fold_left CoreEApp ((CoreEType (TyVarTy ec)):: @@ -95,8 +85,8 @@ Fixpoint weakExprToCoreExpr (me:WeakExpr) : @CoreExpr CoreVar := nil) (CoreEVar v) | WELet (weakExprVar v _) ve e => mkCoreLet (CoreNonRec v (weakExprToCoreExpr ve)) (weakExprToCoreExpr e) - | WECase vscrut e tbranches tc types alts => - CoreECase (weakExprToCoreExpr e) vscrut (weakTypeToCoreType tbranches) + | WECase vscrut escrut tbranches tc types alts => + CoreECase (weakExprToCoreExpr escrut) vscrut (weakTypeToCoreType tbranches) (sortAlts (( fix mkCaseBranches (alts:Tree ??(WeakAltCon*list WeakTypeVar*list WeakCoerVar*list WeakExprVar*WeakExpr)) :=