first pass at proper handling of coercions in HaskWeak
[coq-hetmet.git] / src / HaskWeakVars.v
index 782c2a2..f174563 100644 (file)
@@ -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.
 
@@ -47,6 +41,18 @@ 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.
@@ -85,3 +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