X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakTypes.v;h=9ec126eb69607f52924330adbc336a62133d5870;hp=4d6500d5607d0fb0debcae08a7ef63c3468df510;hb=de8bb8acbfd411d7494c1b35f55cd7edba003b7f;hpb=2ec43bc871b579bac89707988c4855ee1d6c8eda diff --git a/src/HaskWeakTypes.v b/src/HaskWeakTypes.v index 4d6500d..9ec126e 100644 --- a/src/HaskWeakTypes.v +++ b/src/HaskWeakTypes.v @@ -8,7 +8,8 @@ Require Import General. Require Import Coq.Strings.String. Require Import Coq.Lists.List. Require Import HaskKinds. -Require Import HaskLiteralsAndTyCons. +Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskCoreVars. (* a WeakTypeVar merely wraps a CoreVar and includes its Kind *) @@ -46,7 +47,7 @@ Instance WeakTypeVarEqDecidable : EqDecidable WeakTypeVar. 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 WeakCoerVar := weakCoerVar : CoreVar -> WeakType -> WeakType -> WeakCoerVar. Inductive WeakCoercion : Type := | WCoVar : WeakCoerVar -> WeakCoercion (* g *) @@ -63,18 +64,17 @@ Inductive WeakCoercion : Type := (*| 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" +| WCoVar (weakCoerVar _ t1 t2) => (WFunTyCon,WFunTyCon) (* FIXME!!! *) +| WCoType t => (WFunTyCon,WFunTyCon) (* FIXME!!! *) +| WCoApp c1 c2 => (WFunTyCon,WFunTyCon) (* FIXME!!! *) +| WCoAppT c t => (WFunTyCon,WFunTyCon) (* FIXME!!! *) +| WCoAll k f => (WFunTyCon,WFunTyCon) (* FIXME!!! *) | 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" +| WCoComp c1 c2 => (WFunTyCon,WFunTyCon) (* FIXME!!! *) +| WCoLeft c => (WFunTyCon,WFunTyCon) (* FIXME!!! *) +| WCoRight c => (WFunTyCon,WFunTyCon) (* FIXME!!! *) | WCoUnsafe t1 t2 => (t1,t2) end.