X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskStrongTypes.v;h=f849be13c22818178855627bf377228ac2b8104e;hp=e79927ab5ba4ecf6e69cb82d9adb867170935581;hb=3d56944e3882ec751fa99b4476a013c4d86fd0f8;hpb=ab2e0681a81695cc2380b007f2a3314005ec1c99 diff --git a/src/HaskStrongTypes.v b/src/HaskStrongTypes.v index e79927a..f849be1 100644 --- a/src/HaskStrongTypes.v +++ b/src/HaskStrongTypes.v @@ -8,7 +8,7 @@ Require Import Coq.Strings.String. Require Import Coq.Lists.List. Require Import General. Require Import HaskKinds. -Require Import HaskCoreLiterals. +Require Import HaskLiteralsAndTyCons. Require Import HaskCoreTypes. Require Import HaskCoreVars. Require Import HaskWeakTypes. @@ -74,7 +74,7 @@ Instance DataConEqDecidable : forall tc, EqDecidable (@DataCon tc). right; auto. Defined. -Definition tyConKind' tc := fold_right KindTypeFunction ★ (tyConKind tc). +Definition tyConKind' tc := fold_right KindArrow ★ (tyConKind tc). (* types prefixed with "Raw" are NOT binder-polymorphic; they have had their PHOAS parameter instantiated already *) Section Raw. @@ -185,7 +185,7 @@ Definition HaskTApp {Γ}{κ}(σ:forall TV (env:@InstantiatedTypeEnv TV Γ), TV := fun TV env => σ TV env (cv TV env). Definition HaskBrak {Γ}(v:HaskTyVar Γ ★)(t:HaskType Γ ★) : HaskType Γ ★:= fun TV env => @TCode TV (TVar (v TV env)) (t TV env). -Definition HaskTCon {Γ}(tc:TyCon) : HaskType Γ (fold_right KindTypeFunction ★ (tyConKind tc)) +Definition HaskTCon {Γ}(tc:TyCon) : HaskType Γ (fold_right KindArrow ★ (tyConKind tc)) := fun TV ite => TCon tc. Definition HaskAppT {Γ}{κ₁}{κ₂}(t1:HaskType Γ (κ₂⇛κ₁))(t2:HaskType Γ κ₂) : HaskType Γ κ₁ := fun TV ite => TApp (t1 TV ite) (t2 TV ite). @@ -336,11 +336,11 @@ Definition weakF {Γ:TypeEnv}{κ}{κ₂}(f:forall TV (env:@InstantiatedTypeEnv T Fixpoint caseType0 {Γ}(lk:list Kind) : IList _ (HaskType Γ) lk -> - HaskType Γ (fold_right KindTypeFunction ★ lk) -> + HaskType Γ (fold_right KindArrow ★ lk) -> HaskType Γ ★ := match lk as LK return IList _ (HaskType Γ) LK -> - HaskType Γ (fold_right KindTypeFunction ★ LK) -> + HaskType Γ (fold_right KindArrow ★ LK) -> HaskType Γ ★ with | nil => fun _ ht => ht @@ -353,7 +353,7 @@ Definition caseType {Γ}(tc:TyCon)(atypes:IList _ (HaskType Γ) (tyConKind tc)) (* like a GHC DataCon, but using PHOAS representation for types and coercions *) Record StrongAltCon {tc:TyCon} := { sac_tc := tc -; sac_altcon : AltCon +; sac_altcon : WeakAltCon ; sac_numExTyVars : nat ; sac_numCoerVars : nat ; sac_numExprVars : nat @@ -365,7 +365,7 @@ Record StrongAltCon {tc:TyCon} := ; sac_Δ := fun Γ (atypes:IList _ (HaskType Γ) (tyConKind tc)) Δ => app (vec2list (sac_coercions Γ atypes)) Δ }. Coercion sac_tc : StrongAltCon >-> TyCon. -Coercion sac_altcon : StrongAltCon >-> AltCon. +Coercion sac_altcon : StrongAltCon >-> WeakAltCon. Definition kindOfType {Γ}{κ}(ht:@HaskType Γ κ) : ???Kind := OK κ.