593bccf77b5d2de3a138cc61f0facefea7d2dfef
[coq-hetmet.git] / src / HaskWeakTypes.v
1 (*********************************************************************************************************************************)
2 (* HaskWeakTypes: types 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
15 (* TO DO: mark the TyCon used for hetmet as a "type family" so GHC keeps it fully-applied-at-all-times *)
16 Variable tyConToCoreTyCon : TyCon -> CoreTyCon.           Extract Inlined Constant tyConToCoreTyCon  => "(\x -> x)".
17 Variable tyFunToCoreTyCon : TyFun -> CoreTyCon.           Extract Inlined Constant tyFunToCoreTyCon  => "(\x -> x)".
18 Coercion tyConToCoreTyCon : TyCon >-> CoreTyCon.
19 Coercion tyFunToCoreTyCon : TyFun >-> CoreTyCon.
20
21 Instance TyConToString : ToString TyCon := { toString := tyConToString }.
22 Instance TyFunToString : ToString TyFun := { toString := tyConToString }.
23
24 (* a WeakTypeVar merely wraps a CoreVar and includes its Kind *)
25 Inductive WeakTypeVar := weakTypeVar : CoreVar -> Kind -> WeakTypeVar.
26
27 (*
28  * WeakType is much like CoreType, but:
29  *   1. avoids mutually-inductive definitions
30  *   2. gives special cases for the tycons which have their own typing rules so we can pattern-match on them
31  *   3. separates type functions from type constructors, and uses a normal "AppTy" for applying the latter
32  *)
33 Inductive WeakType :=
34 | WTyVarTy  : WeakTypeVar                                      -> WeakType
35 | WAppTy    : WeakType            ->                  WeakType -> WeakType
36 | WTyFunApp : TyFun               ->             list WeakType -> WeakType
37 | WTyCon    : TyCon                                            -> WeakType
38 | WFunTyCon :                                                     WeakType    (* never use (WTyCon ArrowCon);    always use this! *)
39 | WCodeTy   : WeakTypeVar         ->                  WeakType -> WeakType    (* never use the raw tycon *)
40 | WCoFunTy  : WeakType            -> WeakType      -> WeakType -> WeakType
41 | WForAllTy : WeakTypeVar         ->                  WeakType -> WeakType
42 | WClassP   : Class_              ->             list WeakType -> WeakType
43 | WIParam   : CoreIPName CoreName ->                  WeakType -> WeakType.
44
45 (* EqDecidable instances for WeakType *)
46 Instance WeakTypeVarEqDecidable : EqDecidable WeakTypeVar.
47   apply Build_EqDecidable.
48   intros.
49   destruct v1 as [cv1 k1].
50   destruct v2 as [cv2 k2].
51   destruct (eqd_dec cv1 cv2); subst.
52     destruct (eqd_dec k1 k2); subst.
53     left; auto.
54     right; intro; apply n; inversion H; subst; auto.
55     right; intro; apply n; inversion H; subst; auto.
56     Defined.
57
58 (* TO DO: write a proper EqDecidable instance for WeakType and then move the rest of this into HaskWeakToCore *)
59 Variable ModalBoxTyCon   : TyCon.        Extract Inlined Constant ModalBoxTyCon => "TysWiredIn.hetMetCodeTypeTyCon".
60 Variable ArrowTyCon      : TyCon.        Extract Constant ArrowTyCon    => "Type.funTyCon".
61
62 (* if this is a (WAppTy (WAppTy ... (WTyCon tc) .. ) .. ), return (tc,[...]) *)
63 Fixpoint isTyConApp (wt:WeakType)(acc:list WeakType) : ??(TyCon * list WeakType) :=
64   match wt with
65     | WTyCon tc    => Some (tc,acc)
66     | WAppTy t1 t2 => isTyConApp t1 (t2::acc)
67     | _            => None
68   end.
69
70 (* we try to normalize the representation of a type as much as possible before feeding it back to GHCs type-comparison function *)
71 Definition normalizeWeakType (wt:WeakType) : WeakType := wt.
72
73 Fixpoint weakTypeToCoreType' (wt:WeakType) : CoreType :=
74   match wt with
75     | WTyVarTy  (weakTypeVar v _)     => TyVarTy v
76     | WAppTy (WAppTy WFunTyCon t1) t2 => FunTy (weakTypeToCoreType' t1) (weakTypeToCoreType' t2)
77     | WAppTy  t1 t2                   => match (weakTypeToCoreType' t1) with
78                                            | TyConApp tc tys => TyConApp tc (app tys ((weakTypeToCoreType' t2)::nil))
79                                            | t1'             => AppTy t1' (weakTypeToCoreType' t2)
80                                          end
81     | WTyCon    tc                    => TyConApp tc nil
82     | WTyFunApp tf lt                 => TyConApp tf (map weakTypeToCoreType' lt)
83     | WClassP c lt                    => PredTy (ClassP c (map weakTypeToCoreType' lt))
84     | WIParam n ty                    => PredTy (IParam n (weakTypeToCoreType' ty))
85     | WForAllTy (weakTypeVar wtv _) t => ForAllTy wtv (weakTypeToCoreType' t)
86     | WFunTyCon                       => TyConApp ArrowTyCon nil
87     | WCodeTy  (weakTypeVar ec _) t   => TyConApp ModalBoxTyCon ((TyVarTy ec)::(weakTypeToCoreType' t)::nil)
88     | WCoFunTy t1 t2 t3               => FunTy (PredTy (EqPred (weakTypeToCoreType' t1) (weakTypeToCoreType' t2)))
89                                             (weakTypeToCoreType' t3)
90   end.
91
92 Definition weakTypeToCoreType (wt:WeakType) :=
93   weakTypeToCoreType' (normalizeWeakType wt).
94
95 Definition compare_weakTypes (w1 w2:WeakType) :=
96   if coretype_eq_dec (weakTypeToCoreType w1) (weakTypeToCoreType w2)
97     then true
98     else false.
99
100 (*
101 Instance EqDecidableWeakType : EqDecidable WeakType.
102   apply Build_EqDecidable.
103   intros.
104   set (compare_weakTypes_axiom v1 v2) as x.
105   set (compare_weakTypes v1 v2) as y.
106   assert (y=compare_weakTypes v1 v2). reflexivity.
107   destruct y; rewrite <- H in x.
108   left; auto.
109   right; auto.
110   Defined.
111 *)
112
113 Instance WeakTypeToString : ToString WeakType :=
114   { toString := coreTypeToString ○ weakTypeToCoreType }.
115