first pass at proper handling of coercions in HaskWeak
[coq-hetmet.git] / src / HaskWeakVars.v
1 (*********************************************************************************************************************************)
2 (* HaskWeakVars: types and variables for HaskWeak                                                                               *)
3 (*********************************************************************************************************************************)
4
5 Generalizable All Variables.
6 Require Import Preamble.
7 Require Import General.
8 Require Import Coq.Strings.String.
9 Require Import Coq.Lists.List.
10 Require Import HaskKinds.
11 Require Import HaskCoreLiterals.
12 Require Import HaskCoreVars.
13 Require Import HaskCoreTypes.
14 Require Import HaskWeakTypes.
15
16 (* a WeakExprVar just wraps a CoreVar and tags it with the type of its value *)
17 Inductive WeakExprVar := weakExprVar : CoreVar -> WeakType -> WeakExprVar.
18
19 (* a WeakVar is one of the three sorts *)
20 Inductive WeakVar : Type :=
21 | WExprVar : WeakExprVar -> WeakVar
22 | WTypeVar : WeakTypeVar -> WeakVar
23 | WCoerVar : WeakCoerVar -> WeakVar.
24   Coercion WExprVar : WeakExprVar >-> WeakVar.
25   Coercion WTypeVar : WeakTypeVar >-> WeakVar.
26   Coercion WCoerVar : WeakCoerVar >-> WeakVar.
27
28 Definition weakTypeVarToKind (tv:WeakTypeVar) : Kind :=
29   match tv with weakTypeVar _ k => k end.
30   Coercion weakTypeVarToKind : WeakTypeVar >-> Kind.
31
32 Definition weakVarToCoreVar (wv:WeakVar) : CoreVar :=
33   match wv with
34   | WExprVar (weakExprVar v _    ) => v
35   | WTypeVar (weakTypeVar v _    ) => v
36   | WCoerVar (weakCoerVar v _ _ _) => v
37  end.
38  Coercion weakVarToCoreVar : WeakVar >-> CoreVar.
39
40 Definition haskLiteralToWeakType lit : WeakType :=
41   WTyCon (haskLiteralToTyCon lit).
42   Coercion haskLiteralToWeakType : HaskLiteral >-> WeakType.
43
44 Variable coreVarToWeakVar  : CoreVar  -> WeakVar.   Extract Inlined Constant coreVarToWeakVar    => "coreVarToWeakVar".
45 Variable getTyConTyVars_   : CoreTyCon   -> list CoreVar.  Extract Inlined Constant getTyConTyVars_   => "getTyConTyVars".
46 Definition tyConTyVars (tc:CoreTyCon) :=
47   General.filter (map (fun x => match coreVarToWeakVar x with WTypeVar v => Some v | _ => None end) (getTyConTyVars_ tc)).
48   Opaque tyConTyVars.
49 Definition tyConKind (tc:TyCon) : list Kind :=
50   map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc).
51 Variable tyFunResultKind : CoreTyCon -> Kind. Extract Inlined Constant tyFunResultKind => "tyFunResultKind".
52 Definition tyFunKind (tc:TyFun) : ((list Kind) * Kind) :=
53   ((map (fun (x:WeakTypeVar) => x:Kind) (tyConTyVars tc)) , (tyFunResultKind tc)).
54
55 (*
56 (* EqDecidable instances for all of the above *)
57 Instance WeakCoerVarEqDecidable : EqDecidable WeakCoerVar.
58   apply Build_EqDecidable.
59   intros.
60   destruct v1 as [cv1 t1a t1b].
61   destruct v2 as [cv2 t2a t2b].
62   destruct (eqd_dec cv1 cv2); subst.
63     destruct (eqd_dec t1a t2a); subst.
64     destruct (eqd_dec t1b t2b); subst.
65     left; auto.
66     right; intro; apply n; inversion H; subst; auto.
67     right; intro; apply n; inversion H; subst; auto.
68     right; intro; apply n; inversion H; subst; auto.
69     Defined.
70
71 Instance WeakExprVarEqDecidable : EqDecidable WeakExprVar.
72   apply Build_EqDecidable.
73   intros.
74   destruct v1 as [cv1 k1].
75   destruct v2 as [cv2 k2].
76   destruct (eqd_dec cv1 cv2); subst.
77     destruct (eqd_dec k1 k2); subst.
78     left; auto.
79     right; intro; apply n; inversion H; subst; auto.
80     right; intro; apply n; inversion H; subst; auto.
81     Defined.
82
83 Instance WeakVarEqDecidable : EqDecidable WeakVar.
84   apply Build_EqDecidable.
85   induction v1; destruct v2; try (right; intro q; inversion q; fail) ; auto;
86      destruct (eqd_dec w w0); subst.
87      left; auto.
88      right; intro X; apply n; inversion X; auto.
89      left; auto.
90      right; intro X; apply n; inversion X; auto.
91      left; auto.
92      right; intro X; apply n; inversion X; auto.
93   Defined.
94 *)
95
96 Instance WeakVarToString : ToString WeakVar :=
97   { toString := fun x => toString (weakVarToCoreVar x) }.