X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakToStrong.v;h=122db2afc13408cffed68494a968c0a7c7bd4abc;hp=b2521e388917551cb135b3d1c0966b89b274a37b;hb=b4857a6f575dffd5c9c9d5decbc21ff63a338270;hpb=97552c1a6dfb32098d4491951929ab1d4aca96a0 diff --git a/src/HaskWeakToStrong.v b/src/HaskWeakToStrong.v index b2521e3..122db2a 100644 --- a/src/HaskWeakToStrong.v +++ b/src/HaskWeakToStrong.v @@ -19,9 +19,6 @@ Require Import HaskStrongTypes. Require Import HaskStrong. Require Import HaskCoreVars. -(* can remove *) -Require Import HaskStrongToWeak. - Open Scope string_scope. Definition TyVarResolver Γ := forall wt:WeakTypeVar, ???(HaskTyVar Γ wt). Definition CoVarResolver Γ Δ := forall wt:WeakCoerVar, ???(HaskCoVar Γ Δ). @@ -127,7 +124,8 @@ Definition weakTypeToType : forall {Γ:TypeEnv}(φ:TyVarResolver Γ)(t:WeakType) | WAppTy t1 t2 => let case_WAppTy := tt in weakTypeToType _ φ t1 >>= fun t1' => weakTypeToType _ φ t2 >>= fun t2' => _ | WTyVarTy v => let case_WTyVarTy := tt in φ v >>= fun v' => _ | WForAllTy wtv t => let case_WForAllTy := tt in weakTypeToType _ (upφ wtv φ) t >>= fun t => _ - | WCodeTy ec tbody => let case_WCodeTy := tt in weakTypeToType _ φ tbody >>= fun tbody' => φ (@fixkind ★ ec) >>= fun ec' => _ + | WCodeTy ec tbody => let case_WCodeTy := tt in weakTypeToType _ φ tbody + >>= fun tbody' => φ (@fixkind ECKind ec) >>= fun ec' => _ | WCoFunTy t1 t2 t3 => let case_WCoFunTy := tt in weakTypeToType _ φ t1 >>= fun t1' => weakTypeToType _ φ t2 >>= fun t2' => @@ -179,7 +177,7 @@ Definition weakTypeToType : forall {Γ:TypeEnv}(φ:TyVarResolver Γ)(t:WeakType) apply OK. eapply haskTypeOfSomeKind. unfold HaskType; intros. - apply TyFunApp. + apply (TyFunApp tc (fst (tyFunKind tc)) (snd (tyFunKind tc))). apply lt'. apply X. @@ -513,7 +511,22 @@ Fixpoint doesWeakVarOccurAlts (wev:WeakExprVar) | T_Branch b1 b2 => doesWeakVarOccurAlts wev b1 || doesWeakVarOccurAlts wev b2 end. -(*Definition ensureCaseBindersAreNotUsed (we:WeakExpr) : UniqM WeakExpr := FIXME *) +Definition checkDistinct : + forall {V}(EQ:EqDecidable V)(lv:list V), ???(distinct lv). + intros. + set (distinct_decidable lv) as q. + destruct q. + exact (OK d). + exact (Error "checkDistinct failed"). + Defined. + +(* FIXME: check the kind of the type of the weakexprvar to support >0 *) +Definition mkGlobal Γ (τ:HaskType Γ ★) (wev:WeakExprVar) : Global Γ. + refine {| glob_kinds := nil |}. + apply wev. + intros. + apply τ. + Defined. Definition weakExprToStrongExpr : forall (Γ:TypeEnv) @@ -540,7 +553,7 @@ Definition weakExprToStrongExpr : forall match we with | WEVar v => if ig v - then OK (EGlobal Γ Δ ξ (τ@@lev) v) + then OK ((EGlobal Γ Δ ξ (mkGlobal Γ τ v) INil lev) : Expr Γ Δ ξ (τ@@lev)) else castExpr we ("WEVar "+++toString (v:CoreVar)) (τ @@ lev) (EVar Γ Δ ξ v) | WELit lit => castExpr we ("WELit "+++toString lit) (τ @@ lev) (ELit Γ Δ ξ lit lev) @@ -647,8 +660,9 @@ Definition weakExprToStrongExpr : forall OK (ELR_branch Γ Δ ξ' lev _ _ b1' b2') end) rb in binds >>= fun binds' => + checkDistinct CoreVarEqDecidable (map (@fst _ _) (leaves (varsTypes rb φ))) >>= fun rb_distinct => weakExprToStrongExpr Γ Δ φ ψ ξ' ig' τ lev e >>= fun e' => - OK (ELetRec Γ Δ ξ lev τ _ binds' e') + OK (ELetRec Γ Δ ξ lev τ _ _ binds' e') | WECase vscrut escrut tbranches tc avars alts => weakTypeOfWeakExpr escrut >>= fun tscrut => @@ -703,6 +717,8 @@ Definition weakExprToStrongExpr : forall destruct (ξ c). simpl. apply e1. + rewrite mapleaves. + apply rb_distinct. destruct case_pf. set (distinct_decidable (vec2list exprvars')) as dec.