X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskWeakVars.v;h=e7ab943afe134d56fa4e4ff9d008429e089e54c4;hp=c43842fe02d85716697f8a88e589899a897ae9ab;hb=4c5c94487aa2bf5489371f112607f0a4c4f01a94;hpb=5a0761840d89b82cdacb0bf9215fd41aba847b68 diff --git a/src/HaskWeakVars.v b/src/HaskWeakVars.v index c43842f..e7ab943 100644 --- a/src/HaskWeakVars.v +++ b/src/HaskWeakVars.v @@ -4,85 +4,56 @@ Generalizable All Variables. Require Import Preamble. -Require Import General. Require Import Coq.Strings.String. Require Import Coq.Lists.List. -Require Import HaskGeneral. +Require Import General. +Require Import HaskKinds. Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskCoreVars. Require Import HaskCoreTypes. +Require Import HaskWeakTypes. -Inductive WeakTypeVar := weakTypeVar : CoreVar -> Kind -> WeakTypeVar. - -Inductive WeakCoercion : Type := weakCoercion : CoreType -> CoreType -> CoreCoercion -> WeakCoercion. - -Inductive WeakCoerVar := weakCoerVar : CoreVar -> CoreType -> CoreType -> WeakCoerVar. +Inductive CoreVarToWeakVarResult : Type := +| CVTWVR_EVar : CoreType -> CoreVarToWeakVarResult +| CVTWVR_TyVar : Kind -> CoreVarToWeakVarResult +| CVTWVR_CoVar : CoreType -> CoreType -> CoreVarToWeakVarResult. -Inductive WeakExprVar := weakExprVar : CoreVar -> CoreType -> WeakExprVar. +(* a WeakExprVar just wraps a CoreVar and tags it with the type of its value *) +Inductive WeakExprVar := weakExprVar : CoreVar -> WeakType -> WeakExprVar. +(* a WeakVar is one of the three sorts *) Inductive WeakVar : Type := | WExprVar : WeakExprVar -> WeakVar | WTypeVar : WeakTypeVar -> WeakVar | WCoerVar : WeakCoerVar -> WeakVar. + Coercion WExprVar : WeakExprVar >-> WeakVar. + Coercion WTypeVar : WeakTypeVar >-> WeakVar. + Coercion WCoerVar : WeakCoerVar >-> WeakVar. -Definition haskLiteralToCoreType lit : CoreType := - TyConApp (haskLiteralToTyCon lit) nil. - -Definition weakTypeToCoreType (wt:CoreType) : CoreType := wt. - -Definition weakTypeToString : CoreType -> string := coreTypeToString ○ weakTypeToCoreType. - -(* EqDecidable instances for all of the above *) -Instance CoreTypeVarEqDecidable : EqDecidable WeakTypeVar. - 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 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. +Definition weakTypeVarToKind (tv:WeakTypeVar) : Kind := + match tv with weakTypeVar _ k => k end. + Coercion weakTypeVarToKind : WeakTypeVar >-> Kind. -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. +Definition weakVarToCoreVar (wv:WeakVar) : CoreVar := + match wv with + | WExprVar (weakExprVar v _ ) => v + | WTypeVar (weakTypeVar v _ ) => v + | WCoerVar (weakCoerVar v _ _) => v + end. + Coercion weakVarToCoreVar : WeakVar >-> CoreVar. -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 haskLiteralToWeakType lit : WeakType := + WTyCon (haskLiteralToTyCon lit). + Coercion haskLiteralToWeakType : HaskLiteral >-> WeakType. +Variable coreVarToWeakVar : CoreVar -> CoreVarToWeakVarResult. Extract Inlined Constant coreVarToWeakVar => "coreVarToWeakVar". +Variable getTyConTyVars_ : CoreTyCon -> list CoreVar. Extract Inlined Constant getTyConTyVars_ => "getTyConTyVars". +Variable rawTyFunKind : CoreTyCon -> ((list Kind) * Kind). Extract Inlined Constant rawTyFunKind => "rawTyFunKind". +Definition tyFunKind (tc:TyFun) : ((list Kind) * Kind) := + rawTyFunKind tc. +Instance WeakVarToString : ToString WeakVar := + { toString := fun x => toString (weakVarToCoreVar x) }.