clean up demo code
[coq-hetmet.git] / src / HaskWeakVars.v
index 8d84610..e7ab943 100644 (file)
@@ -4,20 +4,20 @@
 
 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.
 
-(* 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.
+Inductive CoreVarToWeakVarResult : Type :=
+| CVTWVR_EVar  : CoreType ->             CoreVarToWeakVarResult
+| CVTWVR_TyVar : Kind     ->             CoreVarToWeakVarResult
+| CVTWVR_CoVar : CoreType -> CoreType -> CoreVarToWeakVarResult.
 
 (* a WeakExprVar just wraps a CoreVar and tags it with the type of its value *)
 Inductive WeakExprVar := weakExprVar : CoreVar -> WeakType -> WeakExprVar.
@@ -37,9 +37,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.
 
@@ -47,57 +47,13 @@ Definition haskLiteralToWeakType lit : WeakType :=
   WTyCon (haskLiteralToTyCon lit).
   Coercion haskLiteralToWeakType : HaskLiteral >-> 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)).
-  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)).
-
-(*
-(* 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 coreVarToWeakVar  : CoreVar     -> CoreVarToWeakVarResult. Extract Inlined Constant coreVarToWeakVar => "coreVarToWeakVar".
+Variable getTyConTyVars_   : CoreTyCon   -> list CoreVar.           Extract Inlined Constant getTyConTyVars_  => "getTyConTyVars".
 
-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.
+Variable rawTyFunKind : CoreTyCon -> ((list Kind) * Kind). Extract Inlined Constant rawTyFunKind => "rawTyFunKind".
 
-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) }.