X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;f=src%2FHaskWeakTypes.v;h=cdbc9e7211e420f7bd45fab9c6b6855ecf286e3e;hb=273645efdb974dd04042e6c59bbedbe0ad658298;hp=82cf8ad2624b44b41627f1968d7b41d6d9f2d7e9;hpb=8c26722a1ee110077968a8a166eb7130266b2035;p=coq-hetmet.git diff --git a/src/HaskWeakTypes.v b/src/HaskWeakTypes.v index 82cf8ad..cdbc9e7 100644 --- a/src/HaskWeakTypes.v +++ b/src/HaskWeakTypes.v @@ -55,6 +55,39 @@ Instance WeakTypeVarEqDecidable : EqDecidable WeakTypeVar. right; intro; apply n; inversion H; subst; auto. Defined. +(* a WeakCoerVar just wraps a CoreVar and tags it with the pair of types amongst which it coerces *) +Inductive WeakCoerVar := weakCoerVar : CoreVar -> Kind -> WeakType -> WeakType -> WeakCoerVar. + +Inductive WeakCoercion : Type := +| WCoVar : WeakCoerVar -> WeakCoercion (* g *) +| WCoType : WeakType -> WeakCoercion (* τ *) +| WCoApp : WeakCoercion -> WeakCoercion -> WeakCoercion (* γ γ *) +| WCoAppT : WeakCoercion -> WeakType -> WeakCoercion (* γ@v *) +| WCoAll : Kind -> (WeakTypeVar -> WeakCoercion) -> WeakCoercion (* ∀a:κ.γ *) +| WCoSym : WeakCoercion -> WeakCoercion (* sym *) +| WCoComp : WeakCoercion -> WeakCoercion -> WeakCoercion (* ◯ *) +| WCoLeft : WeakCoercion -> WeakCoercion (* left *) +| WCoRight : WeakCoercion -> WeakCoercion (* right *) +| WCoUnsafe : WeakType -> WeakType -> WeakCoercion (* unsafe *) +(*| WCoCFApp : ∀ n, CoFunConst n -> vec WeakCoercion n -> WeakCoercion (* C γⁿ *)*) +(*| WCoTFApp : ∀ n, TyFunConst n -> vec WeakCoercion n -> WeakCoercion (* S_n γⁿ *)*) +. + +Variable Prelude_error : forall {A}, string -> A. Extract Inlined Constant Prelude_error => "Prelude.error". +Fixpoint weakCoercionTypes (wc:WeakCoercion) : WeakType * WeakType := +match wc with +| WCoVar (weakCoerVar _ _ t1 t2) => (t1,t2) +| WCoType t => Prelude_error "FIXME WCoType" +| WCoApp c1 c2 => Prelude_error "FIXME WCoApp" +| WCoAppT c t => Prelude_error "FIXME WCoAppT" +| WCoAll k f => Prelude_error "FIXME WCoAll" +| WCoSym c => let (t2,t1) := weakCoercionTypes c in (t1,t2) +| WCoComp c1 c2 => Prelude_error "FIXME WCoComp" +| WCoLeft c => Prelude_error "FIXME WCoLeft" +| WCoRight c => Prelude_error "FIXME WCoRight" +| WCoUnsafe t1 t2 => (t1,t2) +end. + (* TO DO: write a proper EqDecidable instance for WeakType and then move the rest of this into HaskWeakToCore *) Variable ModalBoxTyCon : TyCon. Extract Inlined Constant ModalBoxTyCon => "TysWiredIn.hetMetCodeTypeTyCon". Variable ArrowTyCon : TyCon. Extract Constant ArrowTyCon => "Type.funTyCon". @@ -67,36 +100,13 @@ Fixpoint isTyConApp (wt:WeakType)(acc:list WeakType) : ??(TyCon * list WeakType) | _ => None end. -(* messy first-order NON-CAPTURE-AVOIDING substitution on WeakType's *) -Fixpoint replaceWeakTypeVar (te:WeakType)(tv:WeakTypeVar)(tsubst:WeakType) : WeakType := - match te with - | WTyVarTy tv' => if eqd_dec tv tv' then tsubst else te - | WAppTy t1 t2 => WAppTy (replaceWeakTypeVar t1 tv tsubst) (replaceWeakTypeVar t2 tv tsubst) - | WForAllTy tv' t => if eqd_dec tv tv' then te else WForAllTy tv' (replaceWeakTypeVar t tv tsubst) - | WCoFunTy t1 t2 t => WCoFunTy (replaceWeakTypeVar t1 tv tsubst) - (replaceWeakTypeVar t2 tv tsubst) (replaceWeakTypeVar t tv tsubst) - | WIParam ip ty => WIParam ip (replaceWeakTypeVar ty tv tsubst) - | WClassP c lt => WClassP c ((fix replaceCoreDistinctList (lt:list WeakType) := - match lt with - | nil => nil - | h::t => (replaceWeakTypeVar h tv tsubst)::(replaceCoreDistinctList t) - end) lt) - | WTyFunApp tc lt => WTyFunApp tc ((fix replaceCoreDistinctList (lt:list WeakType) := - match lt with - | nil => nil - | h::t => (replaceWeakTypeVar h tv tsubst)::(replaceCoreDistinctList t) - end) lt) - | WTyCon tc => WTyCon tc - | WFunTyCon => WFunTyCon - | WModalBoxTyCon => WModalBoxTyCon - end. - (* we try to normalize the representation of a type as much as possible before feeding it back to GHCs type-comparison function *) Definition normalizeWeakType (wt:WeakType) : WeakType := wt. Fixpoint weakTypeToCoreType' (wt:WeakType) : CoreType := match wt with | WTyVarTy (weakTypeVar v _) => TyVarTy v + | WAppTy (WAppTy WFunTyCon t1) t2 => FunTy (weakTypeToCoreType' t1) (weakTypeToCoreType' t2) | WAppTy t1 t2 => match (weakTypeToCoreType' t1) with | TyConApp tc tys => TyConApp tc (app tys ((weakTypeToCoreType' t2)::nil)) | t1' => AppTy t1' (weakTypeToCoreType' t2) @@ -120,9 +130,7 @@ Definition compare_weakTypes (w1 w2:WeakType) := then true else false. -(* Coq's "decide equality" can't cope here; we have to cheat for now *) -Axiom compare_weakTypes_axiom : forall w1 w2, if compare_weakTypes w1 w2 then w1=w2 else not (w1=w2). - +(* Instance EqDecidableWeakType : EqDecidable WeakType. apply Build_EqDecidable. intros. @@ -133,6 +141,7 @@ Instance EqDecidableWeakType : EqDecidable WeakType. left; auto. right; auto. Defined. +*) Instance WeakTypeToString : ToString WeakType := { toString := coreTypeToString ○ weakTypeToCoreType }.