X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakVars.v;h=3fb7a4bb6f775d367abd7b3e7265ea0097f1b3fc;hp=f1745639fb9489a7a07201a04cd995e590959532;hb=5deda3b8240059e9969a31706d89b8a3818b184c;hpb=ab2e0681a81695cc2380b007f2a3314005ec1c99 diff --git a/src/HaskWeakVars.v b/src/HaskWeakVars.v index f174563..3fb7a4b 100644 --- a/src/HaskWeakVars.v +++ b/src/HaskWeakVars.v @@ -4,11 +4,12 @@ Generalizable All Variables. Require Import Preamble. -Require Import General. Require Import Coq.Strings.String. Require Import Coq.Lists.List. +Require Import General. Require Import HaskKinds. -Require Import HaskCoreLiterals. +Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskCoreVars. Require Import HaskCoreTypes. Require Import HaskWeakTypes. @@ -31,9 +32,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. @@ -44,54 +45,14 @@ Definition haskLiteralToWeakType lit : WeakType := Variable coreVarToWeakVar : CoreVar -> WeakVar. Extract Inlined Constant coreVarToWeakVar => "coreVarToWeakVar". Variable getTyConTyVars_ : CoreTyCon -> list CoreVar. Extract Inlined Constant getTyConTyVars_ => "getTyConTyVars". Definition tyConTyVars (tc:CoreTyCon) := - General.filter (map (fun x => match coreVarToWeakVar x with WTypeVar v => Some v | _ => None end) (getTyConTyVars_ tc)). + filter (map (fun x => match coreVarToWeakVar x with WTypeVar v => Some v | _ => None end) (getTyConTyVars_ tc)). 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". -Definition tyFunKind (tc:TyFun) : ((list Kind) * Kind) := - ((map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc)) , (tyFunResultKind tc)). +Definition tyConKind (tc:TyCon) : list Kind := map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc). -(* -(* EqDecidable instances for all of the above *) -Instance WeakCoerVarEqDecidable : EqDecidable WeakCoerVar. - apply Build_EqDecidable. - intros. - destruct v1 as [cv1 t1a t1b]. - destruct v2 as [cv2 t2a t2b]. - destruct (eqd_dec cv1 cv2); subst. - destruct (eqd_dec t1a t2a); subst. - destruct (eqd_dec t1b t2b); subst. - left; auto. - right; intro; apply n; inversion H; subst; auto. - right; intro; apply n; inversion H; subst; auto. - right; intro; apply n; inversion H; subst; auto. - Defined. +Variable rawTyFunKind : CoreTyCon -> ((list Kind) * Kind). Extract Inlined Constant rawTyFunKind => "rawTyFunKind". -Instance WeakExprVarEqDecidable : EqDecidable WeakExprVar. - apply Build_EqDecidable. - intros. - destruct v1 as [cv1 k1]. - destruct v2 as [cv2 k2]. - destruct (eqd_dec cv1 cv2); subst. - destruct (eqd_dec k1 k2); subst. - left; auto. - right; intro; apply n; inversion H; subst; auto. - right; intro; apply n; inversion H; subst; auto. - Defined. - -Instance WeakVarEqDecidable : EqDecidable WeakVar. - apply Build_EqDecidable. - induction v1; destruct v2; try (right; intro q; inversion q; fail) ; auto; - destruct (eqd_dec w w0); subst. - left; auto. - right; intro X; apply n; inversion X; auto. - left; auto. - right; intro X; apply n; inversion X; auto. - left; auto. - right; intro X; apply n; inversion X; auto. - Defined. -*) +Definition tyFunKind (tc:TyFun) : ((list Kind) * Kind) := + rawTyFunKind tc. Instance WeakVarToString : ToString WeakVar := - { toString := fun x => toString (weakVarToCoreVar x) }. \ No newline at end of file + { toString := fun x => toString (weakVarToCoreVar x) }.