X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakVars.v;h=f1745639fb9489a7a07201a04cd995e590959532;hp=ee8d3cf9d8e0a7db004bdbf4fa47261cccd66409;hb=ab2e0681a81695cc2380b007f2a3314005ec1c99;hpb=8c26722a1ee110077968a8a166eb7130266b2035 diff --git a/src/HaskWeakVars.v b/src/HaskWeakVars.v index ee8d3cf..f174563 100644 --- a/src/HaskWeakVars.v +++ b/src/HaskWeakVars.v @@ -13,12 +13,6 @@ Require Import HaskCoreVars. Require Import HaskCoreTypes. Require Import HaskWeakTypes. -(* TO DO: finish this *) -Inductive WeakCoercion : Type := weakCoercion : WeakType -> WeakType -> CoreCoercion -> WeakCoercion. - -(* a WeakCoerVar just wraps a CoreVar and tags it with the pair of types amongst which it coerces *) -Inductive WeakCoerVar := weakCoerVar : CoreVar -> WeakType -> WeakType -> WeakCoerVar. - (* a WeakExprVar just wraps a CoreVar and tags it with the type of its value *) Inductive WeakExprVar := weakExprVar : CoreVar -> WeakType -> WeakExprVar. @@ -37,9 +31,9 @@ Definition weakTypeVarToKind (tv:WeakTypeVar) : Kind := Definition weakVarToCoreVar (wv:WeakVar) : CoreVar := match wv with - | WExprVar (weakExprVar v _ ) => v - | WTypeVar (weakTypeVar v _ ) => v - | WCoerVar (weakCoerVar v _ _ ) => v + | WExprVar (weakExprVar v _ ) => v + | WTypeVar (weakTypeVar v _ ) => v + | WCoerVar (weakCoerVar v _ _ _) => v end. Coercion weakVarToCoreVar : WeakVar >-> CoreVar. @@ -58,6 +52,7 @@ Variable tyFunResultKind : CoreTyCon -> Kind. Extract Inlined Constant tyFunResu Definition tyFunKind (tc:TyFun) : ((list Kind) * Kind) := ((map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc)) , (tyFunResultKind tc)). +(* (* EqDecidable instances for all of the above *) Instance WeakCoerVarEqDecidable : EqDecidable WeakCoerVar. apply Build_EqDecidable. @@ -96,8 +91,7 @@ Instance WeakVarEqDecidable : EqDecidable WeakVar. left; auto. right; intro X; apply n; inversion X; auto. Defined. - - +*) Instance WeakVarToString : ToString WeakVar := { toString := fun x => toString (weakVarToCoreVar x) }. \ No newline at end of file