X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskStrongTypes.v;h=2ee78b2223952536632ce1ce45f8db96e9bba0da;hp=f849be13c22818178855627bf377228ac2b8104e;hb=61fb093700aab006b77998d1cbd30235e144ca1f;hpb=3d56944e3882ec751fa99b4476a013c4d86fd0f8 diff --git a/src/HaskStrongTypes.v b/src/HaskStrongTypes.v index f849be1..2ee78b2 100644 --- a/src/HaskStrongTypes.v +++ b/src/HaskStrongTypes.v @@ -21,7 +21,7 @@ Variable dataConExVars_ : CoreDataCon -> list CoreVar. Extract Inlined Const Variable dataConEqTheta_ : CoreDataCon -> list PredType. Extract Inlined Constant dataConEqTheta_ => "DataCon.dataConEqTheta". Variable dataConOrigArgTys_: CoreDataCon -> list CoreType. Extract Inlined Constant dataConOrigArgTys_=>"DataCon.dataConOrigArgTys". -(* FIXME: might be a better idea to panic here than simply drop things that look wrong *) +(* TODO: might be a better idea to panic here than simply drop things that look wrong *) Definition dataConExTyVars cdc := filter (map (fun x => match coreVarToWeakVar x with WTypeVar v => Some v | _ => None end) (dataConExVars_ cdc)). Opaque dataConExTyVars. @@ -53,27 +53,6 @@ Inductive DataCon : TyCon -> Type := Coercion dataConToCoreDataCon : DataCon >-> CoreDataCon. (*Opaque DataCon.*) -Definition compare_datacons : forall tc, forall dc1 dc2:DataCon tc, bool. - intros. - destruct dc1. - destruct dc2. - exact (if eqd_dec cdc cdc0 then true else false). - Defined. - -Axiom trust_compare_datacons : forall tc dc1 dc2, if compare_datacons tc dc1 dc2 then dc1=dc2 else not (dc1=dc2). - -Instance DataConEqDecidable : forall tc, EqDecidable (@DataCon tc). - intros. - apply Build_EqDecidable. - intros. - set (trust_compare_datacons _ v1 v2) as x. - set (compare_datacons tc v1 v2) as y. - assert (y=compare_datacons tc v1 v2). reflexivity. - destruct y; rewrite <- H in x. - left; auto. - right; auto. - Defined. - Definition tyConKind' tc := fold_right KindArrow ★ (tyConKind tc). (* types prefixed with "Raw" are NOT binder-polymorphic; they have had their PHOAS parameter instantiated already *) @@ -385,13 +364,35 @@ Notation "a ∼∼∼ b" := (@mkHaskCoercionKind _ _ a b) (at level 18). Fixpoint update_ξ `{EQD_VV:EqDecidable VV}{Γ} (ξ:VV -> LeveledHaskType Γ ★) - (vt:list (VV * LeveledHaskType Γ ★)) + (lev:HaskLevel Γ) + (vt:list (VV * HaskType Γ ★)) : VV -> LeveledHaskType Γ ★ := match vt with | nil => ξ - | (v,τ)::tl => fun v' => if eqd_dec v v' then τ else (update_ξ ξ tl) v' + | (v,τ)::tl => fun v' => if eqd_dec v v' then τ @@ lev else (update_ξ ξ lev tl) v' end. +Lemma update_ξ_lemma0 `{EQD_VV:EqDecidable VV} : forall Γ ξ (lev:HaskLevel Γ)(varstypes:list (VV*_)) v, + not (In v (map (@fst _ _) varstypes)) -> + (update_ξ ξ lev varstypes) v = ξ v. + intros. + induction varstypes. + reflexivity. + simpl. + destruct a. + destruct (eqd_dec v0 v). + subst. + simpl in H. + assert False. + apply H. + auto. + inversion H0. + apply IHvarstypes. + unfold not; intros. + apply H. + simpl. + auto. + Defined. (***************************************************************************************************) @@ -546,7 +547,7 @@ Instance haskLevelEqDecidable Γ : EqDecidable (HaskLevel Γ). (* ToString instance for PHOAS types *) Fixpoint typeToString' (needparens:bool)(n:nat){κ}(t:RawHaskType (fun _ => nat) κ) {struct t} : string := match t with - | TVar _ v => "tv" +++ v + | TVar _ v => "tv" +++ toString v | TCon tc => toString tc | TCoerc _ t1 t2 t => "("+++typeToString' false n t1+++"~" +++typeToString' false n t2+++")=>" @@ -563,11 +564,12 @@ Fixpoint typeToString' (needparens:bool)(n:nat){κ}(t:RawHaskType (fun _ => nat) else (typeToString' true n t1)+++" "+++(typeToString' false n t2) end | TArrow => "(->)" - | TAll k f => let alpha := "tv"+++n - in "(forall "+++ alpha +++ ":"+++ k +++")"+++ + | TAll k f => let alpha := "tv"+++ toString n + in "(forall "+++ alpha +++ ":"+++ toString k +++")"+++ typeToString' false (S n) (f n) | TCode ec t => "<["+++(typeToString' true n t)+++"]>@"+++(typeToString' false n ec) - | TyFunApp tfc lt => tfc+++"_"+++n+++" ["+++(fold_left (fun x y => " \ "+++x+++y) (typeList2string false n lt) "")+++"]" + | TyFunApp tfc lt => toString tfc+++ "_" +++ toString n+++" ["+++ + (fold_left (fun x y => " \ "+++x+++y) (typeList2string false n lt) "")+++"]" end with typeList2string (needparens:bool)(n:nat){κ}(t:RawHaskTypeList κ) {struct t} : list string := match t with