X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;f=src%2FHaskWeakVars.v;h=44e267d0ab21bc49d2cfb98e735ef686a486f24b;hb=273645efdb974dd04042e6c59bbedbe0ad658298;hp=8d8461048ecdd97bb9c7d96dd7a32af508fd545b;hpb=5c493a75fbaf8454d8a21e55edc5b193e2c5879c;p=coq-hetmet.git diff --git a/src/HaskWeakVars.v b/src/HaskWeakVars.v index 8d84610..44e267d 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. @@ -54,9 +48,11 @@ Definition tyConTyVars (tc:CoreTyCon) := Opaque tyConTyVars. Definition tyConKind (tc:TyCon) : list Kind := map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc). -Variable tyFunResultKind : CoreTyCon -> Kind. Extract Inlined Constant tyFunResultKind => "tyFunResultKind". + +Variable rawTyFunKind : CoreTyCon -> Kind. Extract Inlined Constant rawTyFunKind => "(coreKindToKind . TyCon.tyConKind)". + Definition tyFunKind (tc:TyFun) : ((list Kind) * Kind) := - ((map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc)) , (tyFunResultKind tc)). + splitKind (rawTyFunKind tc). (* (* EqDecidable instances for all of the above *)