X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeak.v;h=077f7b509a22032859a83729226a7fc82b2620fe;hp=b4503303eec5344655a7aacf1197705a1e6a33a1;hb=1a2754d2e135ef3c5fd7ef817e1129af93b533a5;hpb=112daf37524662d6d2267d3f7e50ff3522683b8f diff --git a/src/HaskWeak.v b/src/HaskWeak.v index b450330..077f7b5 100644 --- a/src/HaskWeak.v +++ b/src/HaskWeak.v @@ -5,165 +5,107 @@ Generalizable All Variables. Require Import Preamble. Require Import General. -Require Import HaskGeneral. +Require Import Coq.Lists.List. +Require Import HaskKinds. Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskCoreVars. -Require Import HaskCoreTypes. +Require Import HaskWeakVars. +Require Import HaskWeakTypes. + +Inductive WeakAltCon := +| WeakDataAlt : CoreDataCon -> WeakAltCon +| WeakLitAlt : HaskLiteral -> WeakAltCon +| WeakDEFAULT : WeakAltCon. Inductive WeakExpr := -| WEVar : CoreVar -> WeakExpr -| WELit : HaskLiteral -> WeakExpr -| WELet : CoreVar -> WeakExpr -> WeakExpr -> WeakExpr -| WELetRec : Tree ??(CoreVar * WeakExpr) -> WeakExpr -> WeakExpr -| WECast : WeakExpr -> CoreCoercion -> WeakExpr -| WENote : Note -> WeakExpr -> WeakExpr -| WEApp : WeakExpr -> WeakExpr -> WeakExpr -| WETyApp : WeakExpr -> CoreType -> WeakExpr -| WECoApp : WeakExpr -> CoreCoercion -> WeakExpr -| WELam : CoreVar -> WeakExpr -> WeakExpr -| WETyLam : CoreVar -> WeakExpr -> WeakExpr -| WECoLam : CoreVar -> WeakExpr -> WeakExpr -| WEBrak : CoreVar -> WeakExpr -> WeakExpr -| WEEsc : CoreVar -> WeakExpr -> WeakExpr -| WECase : forall (scrutinee:WeakExpr) - (tbranches:CoreType) - {n:nat}(tc:TyCon n) - (type_params:vec CoreType n) - (alts : Tree ??(AltCon*list CoreVar*WeakExpr)), - WeakExpr. +| WEVar : WeakExprVar -> WeakExpr +| WELit : HaskLiteral -> WeakExpr -(* calculate the free CoreVar's in a WeakExpr *) -Fixpoint getWeakExprFreeVars (me:WeakExpr) : list CoreVar := - match me with - | WELit _ => nil - | WEVar cv => cv::nil - | WECast e co => getWeakExprFreeVars e - | WENote n e => getWeakExprFreeVars e - | WETyApp e t => getWeakExprFreeVars e - | WECoApp e co => getWeakExprFreeVars e - | WEBrak ec e => getWeakExprFreeVars e - | WEEsc ec e => getWeakExprFreeVars e - | WELet cv e1 e2 => mergeDistinctLists (getWeakExprFreeVars e1) (removeFromDistinctList cv (getWeakExprFreeVars e2)) - | WEApp e1 e2 => mergeDistinctLists (getWeakExprFreeVars e1) (getWeakExprFreeVars e2) - | WELam cv e => removeFromDistinctList cv (getWeakExprFreeVars e) - | WETyLam cv e => removeFromDistinctList cv (getWeakExprFreeVars e) - | WECoLam cv e => removeFromDistinctList cv (getWeakExprFreeVars e) +(* TO DO: add a WEWhere and use the source location to detect which one the user used *) +| WELet : WeakExprVar -> WeakExpr -> WeakExpr -> WeakExpr +| WELetRec : Tree ??(WeakExprVar * WeakExpr) -> WeakExpr -> WeakExpr +| WECast : WeakExpr -> WeakCoercion -> WeakExpr +| WENote : Note -> WeakExpr -> WeakExpr +| WEApp : WeakExpr -> WeakExpr -> WeakExpr +| WETyApp : WeakExpr -> WeakType -> WeakExpr +| WECoApp : WeakExpr -> WeakCoercion -> WeakExpr +| WELam : WeakExprVar -> WeakExpr -> WeakExpr +| WETyLam : WeakTypeVar -> WeakExpr -> WeakExpr +| WECoLam : WeakCoerVar -> WeakExpr -> WeakExpr - (* the messy fixpoints below are required by Coq's termination conditions *) - | WECase scrutinee tbranches n tc type_params alts => - mergeDistinctLists (getWeakExprFreeVars scrutinee) ( - ((fix getWeakExprFreeVarsAlts (alts:Tree ??(AltCon*list CoreVar*WeakExpr)) {struct alts} : list CoreVar := - match alts with - | T_Leaf None => nil - | T_Leaf (Some (DEFAULT,_,e)) => getWeakExprFreeVars e - | T_Leaf (Some (LitAlt lit,_,e)) => getWeakExprFreeVars e - | T_Leaf (Some (DataAlt _ _ _ _ _ dc, vars,e)) => removeFromDistinctList' vars (getWeakExprFreeVars e) - | T_Branch b1 b2 => mergeDistinctLists (getWeakExprFreeVarsAlts b1) (getWeakExprFreeVarsAlts b2) - end) alts)) - | WELetRec mlr e => (fix removeVarsLetRec (mlr:Tree ??(CoreVar * WeakExpr))(cvl:list CoreVar) := - match mlr with - | T_Leaf None => cvl - | T_Leaf (Some (cv,e)) => removeFromDistinctList cv cvl - | T_Branch b1 b2 => removeVarsLetRec b1 (removeVarsLetRec b2 cvl) - end) mlr (mergeDistinctLists (getWeakExprFreeVars e) - ((fix getWeakExprFreeVarsLetRec (mlr:Tree ??(CoreVar * WeakExpr)) := - match mlr with - | T_Leaf None => nil - | T_Leaf (Some (cv,e)) => getWeakExprFreeVars e - | T_Branch b1 b2 => mergeDistinctLists (getWeakExprFreeVarsLetRec b1) (getWeakExprFreeVarsLetRec b2) - end) mlr)) - end. +(* The WeakType argument in WEBrak/WEEsc is used only when going back *) +(* from Weak to Core; it lets us dodge a possibly-failing type *) +(* calculation. The CoreVar argument is the GlobalVar for the hetmet_brak *) +(* or hetmet_esc identifier *) +| WEBrak : WeakExprVar -> WeakTypeVar -> WeakExpr -> WeakType -> WeakExpr +| WEEsc : WeakExprVar -> WeakTypeVar -> WeakExpr -> WeakType -> WeakExpr +| WECSP : WeakExprVar -> WeakTypeVar -> WeakExpr -> WeakType -> WeakExpr -(* wrap lambdas around an expression until it has no free variables *) -Definition makeClosedExpression : WeakExpr -> WeakExpr := - fun me => (fix closeExpression (me:WeakExpr)(cvl:list CoreVar) := - match cvl with - | nil => me - | cv::cvl' => WELam cv (closeExpression me cvl') - end) me (getWeakExprFreeVars me). +| WECase : forall (vscrut:WeakExprVar) + (scrutinee:WeakExpr) + (tbranches:WeakType) + (tc:TyCon) + (type_params:list WeakType) + (alts : Tree ??(WeakAltCon*list WeakTypeVar*list WeakCoerVar*list WeakExprVar*WeakExpr)), + WeakExpr. -(* messy first-order capture-avoiding substitution on CoreType's *) -Fixpoint replaceCoreVar (te:CoreType)(tv:CoreVar)(tsubst:CoreType) := - match te with - | TyVarTy tv' => if eqd_dec tv tv' then tsubst else te - | AppTy t1 t2 => AppTy (replaceCoreVar t1 tv tsubst) (replaceCoreVar t2 tv tsubst) - | FunTy t1 t2 => FunTy (replaceCoreVar t1 tv tsubst) (replaceCoreVar t2 tv tsubst) - | ForAllTy tv' t => if eqd_dec tv tv' then te else ForAllTy tv' (replaceCoreVar t tv tsubst) - | PredTy (EqPred t1 t2) => PredTy (EqPred (replaceCoreVar t1 tv tsubst) (replaceCoreVar t2 tv tsubst)) - | PredTy (IParam ip ty) => PredTy (IParam ip (replaceCoreVar ty tv tsubst)) - | PredTy (ClassP _ c lt) => PredTy (ClassP c ((fix replaceCoreDistinctList (lt:list CoreType) := - match lt with - | nil => nil - | h::t => (replaceCoreVar h tv tsubst)::(replaceCoreDistinctList t) - end) lt)) - | TyConApp _ tc lt => TyConApp tc ((fix replaceCoreDistinctList (lt:list CoreType) := - match lt with - | nil => nil - | h::t => (replaceCoreVar h tv tsubst)::(replaceCoreDistinctList t) - end) lt) - end. +Definition weakTypeOfLiteral (lit:HaskLiteral) : WeakType := + (WTyCon (haskLiteralToTyCon lit)). -(* calculate the CoreType of a WeakExpr *) -Fixpoint coreTypeOfWeakExpr (ce:WeakExpr) : ???CoreType := - match ce with - | WEVar v => match coreVarSort v with - | CoreExprVar t => OK t - | CoreTypeVar _ => Error "found tyvar in expression" - | CoreCoerVar _ => Error "found coercion variable in expression" - end - | WELit lit => OK (haskLiteralToCoreType lit) - | WEApp e1 e2 => coreTypeOfWeakExpr e1 >>= fun t' => - match t' with - | FunTy t1 t2 => OK t2 - | (TyConApp 2 tc (t1::t2::nil)) => - if (tyCon_eq tc ArrowTyCon) - then OK t2 - else Error ("found non-function type "+++(coreTypeToString t')+++" in EApp") - | _ => Error ("found non-function type "+++(coreTypeToString t')+++" in EApp") - end - | WETyApp e t => coreTypeOfWeakExpr e >>= fun te => - match te with - | ForAllTy v ct => match coreVarSort v with - | CoreExprVar _ => Error "found an expression variable inside an forall-type!" - | CoreTypeVar _ => OK (replaceCoreVar ct v t) - | CoreCoerVar _ => Error "found a coercion variable inside an forall-type!" - end - | _ => Error ("found non-forall type "+++(coreTypeToString te)+++" in ETyApp") - end - | WECoApp e co => coreTypeOfWeakExpr e >>= fun te => - match te with - | FunTy (PredTy (EqPred t1 t2)) t3 => OK t3 - | TyConApp 2 tc ((PredTy (EqPred t1 t2))::t3::nil) => - if (tyCon_eq tc ArrowTyCon) - then OK t3 - else Error ("found non-coercion type "+++(coreTypeToString te)+++" in ETyApp") - | _ => Error ("found non-coercion type "+++(coreTypeToString te)+++" in ETyApp") - end - | WELam ev e => coreTypeOfWeakExpr e >>= fun t' => match coreVarSort ev with - | CoreExprVar vt => OK (FunTy vt t') - | CoreTypeVar _ => Error "found a type variable in a WELam!" - | CoreCoerVar _ => Error "found a coercion variable in a WELam!" - end - | WETyLam tv e => coreTypeOfWeakExpr e >>= fun t' => OK (ForAllTy tv t') - | WECoLam cv e => coreTypeOfWeakExpr e >>= fun t' => match coreVarSort cv with - | CoreExprVar vt => Error "found an expression variable in a WECoLam!" - | CoreTypeVar _ => Error "found a type variable in a WECoLam!" - | CoreCoerVar (φ₁,φ₂) => OK (FunTy (PredTy (EqPred φ₁ φ₂)) t') - end - | WELet ev ve e => coreTypeOfWeakExpr e - | WELetRec rb e => coreTypeOfWeakExpr e - | WENote n e => coreTypeOfWeakExpr e - | WECast e co => OK (snd (coreCoercionKind co)) - | WECase scrutinee tbranches n tc type_params alts => OK tbranches - | WEBrak ec e => coreTypeOfWeakExpr e >>= fun t' => OK (TyConApp hetMetCodeTypeTyCon ((TyVarTy ec)::t'::nil)) - | WEEsc ec e => coreTypeOfWeakExpr e >>= fun t' => - match t' with - | (TyConApp 2 tc ((TyVarTy ec')::t''::nil)) => - if (tyCon_eq tc hetMetCodeTypeTyCon) - then if eqd_dec ec ec' then OK t'' - else Error "level mismatch in escapification" - else Error "ill-typed escapification" - | _ => Error "ill-typed escapification" - end +(* +Fixpoint weakExprVarOccursFree (wvf:WeakExprVar)(we:WeakExpr) : bool := + match we with + | WEVar wv => if eqd_dec (wvf:CoreVar) (wv:CoreVar) then true else false + | WELit lit => false + | WEApp e1 e2 => weakExprVarOccursFree wvf e1 || weakExprVarOccursFree wvf e2 + | WETyApp e t => weakExprVarOccursFree wvf e + | WECoApp e co => weakExprVarOccursFree wvf e + | WENote n e => weakExprVarOccursFree wvf e + | WELam ev e => if eqd_dec (wvf:CoreVar) (ev:CoreVar) then false else weakExprVarOccursFree wvf e + | WETyLam tv e => weakExprVarOccursFree wvf e + | WECoLam cv e => weakExprVarOccursFree wvf e + | WECast e co => weakExprVarOccursFree wvf e + | WEBrak v wtv e t => weakExprVarOccursFree wvf e + | WEEsc v wtv e t => weakExprVarOccursFree wvf e + | WECSP v wtv e t => weakExprVarOccursFree wvf e + | WELet v ebind ebody => weakExprVarOccursFree wvf ebind + || if eqd_dec (wvf:CoreVar) (v:CoreVar) then false else weakExprVarOccursFree wvf ebody + | WECase vs es tb tc tys alts => + if weakExprVarOccursFree wvf es + then true + else (fix weakExprVarOccursFreeBranches (alts:Tree ??(_)) : bool := + match alts with + | T_Leaf None => false + | T_Leaf (Some (_,_,_,v',e')) => + if fold_left bor (map (fun v'':WeakExprVar => if eqd_dec (wvf:CoreVar) (v'':CoreVar) then true else false ) v') false + then false + else weakExprVarOccursFree wvf e' + | T_Branch b1 b2 => weakExprVarOccursFreeBranches b1 || + weakExprVarOccursFreeBranches b2 + end) alts + | WELetRec mlr e => false end. +(* some very simple-minded cleanups to produce "prettier" expressions *) +Fixpoint simplifyWeakExpr (me:WeakExpr) : WeakExpr := + match me with + | WEVar wv => WEVar wv + | WELit lit => WELit lit + | WEApp e1 e2 => WEApp (simplifyWeakExpr e1) (simplifyWeakExpr e2) + | WETyApp e t => WETyApp (simplifyWeakExpr e ) t + | WECoApp e co => CoreEApp (simplifyWeakExpr e ) co + | WENote n e => CoreENote n (simplifyWeakExpr e ) + | WELam ev e => CoreELam ev (simplifyWeakExpr e ) + | WETyLam tv e => CoreELam tv (simplifyWeakExpr e ) + | WECoLam cv e => CoreELam cv (simplifyWeakExpr e ) + | WECast e co => CoreECast (simplifyWeakExpr e ) co + | WEBrak v wtv e t => WEBrak v wtv (simplifyWeakExpr e ) t + | WEEsc v wtv e t => WEEsc v wtv (simplifyWeakExpr e ) t + | WECSP v wtv e t => WECSP v wtv (simplifyWeakExpr e ) t + | WELet v ebind ebody => WELet v (simplifyWeakExpr ebind) (simplifyWeakExpr ebody) + | WECase vs es tb tc tys alts => WECase vs es tb tc tys (* FIXME alts *) + (* un-letrec-ify multi branch letrecs *) + | WELetRec mlr e => WELetRec mlr (simplifyWeakExpr e ) + end. +*)