give HaskWeak its own type representation, fix numerous bugs
[coq-hetmet.git] / src / HaskStrong.v
1 (*********************************************************************************************************************************)
2 (* HaskStrong: a dependently-typed version of CoreSyn                                                                            *)
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 HaskCoreTypes.
12 Require Import HaskCoreLiterals.
13 Require Import HaskStrongTypes.
14
15 Section HaskStrong.
16
17   (* any type with decidable equality may be used to represent value variables *)
18   Context `{EQD_VV:EqDecidable VV}.
19
20   (* a StrongCaseBranchWithVVs contains all the data found in a case branch except the expression itself *)
21   Record StrongCaseBranchWithVVs {tc:TyCon}{Γ}{atypes:vec (HaskType Γ) tc} :=
22   { scbwv_sac       :  @StrongAltCon tc
23   ; scbwv_exprvars  :  vec VV (sac_numExprVars scbwv_sac)
24   ; scbwv_varstypes := vec_zip scbwv_exprvars (sac_types scbwv_sac Γ atypes)
25   ; scbwv_ξ         := fun ξ lev =>  update_ξ (weakLT'○ξ) (vec2list
26                                         (vec_map (fun x => ((fst x),(snd x @@ weakL' lev))) scbwv_varstypes))
27   }.
28   Implicit Arguments StrongCaseBranchWithVVs [[Γ]].
29   Coercion scbwv_sac : StrongCaseBranchWithVVs >-> StrongAltCon.
30
31   Inductive Expr : forall Γ (Δ:CoercionEnv Γ), (VV -> LeveledHaskType Γ) -> LeveledHaskType Γ -> Type :=
32   | EVar   : ∀ Γ Δ ξ ev,                                                                             Expr Γ Δ ξ (ξ ev)
33   | ELit   : ∀ Γ Δ ξ lit   l,                                                                        Expr Γ Δ ξ (literalType lit@@l)
34   | EApp   : ∀ Γ Δ ξ t1 t2 l,        Expr Γ Δ ξ (t2--->t1 @@ l)   -> Expr Γ Δ ξ (t2 @@ l)         -> Expr Γ Δ ξ (t1 @@ l)
35   | ELam   : ∀ Γ Δ ξ t1 t2 l ev, Γ ⊢ᴛy t1:★ ->Expr Γ Δ (update_ξ ξ ((ev,t1@@l)::nil)) (t2@@l)     -> Expr Γ Δ ξ (t1--->t2@@l)
36   | ELet   : ∀ Γ Δ ξ tv t  l ev,Expr Γ Δ ξ (tv@@l)->Expr Γ Δ (update_ξ ξ ((ev,tv@@l)::nil))(t@@l) -> Expr Γ Δ ξ (t@@l)
37   | EEsc   : ∀ Γ Δ ξ ec t  l,     Expr Γ Δ ξ (<[ ec |- t ]> @@ l)                                 -> Expr Γ Δ ξ (t @@ (ec::l))
38   | EBrak  : ∀ Γ Δ ξ ec t  l,     Expr Γ Δ ξ (t @@ (ec::l))                                       -> Expr Γ Δ ξ (<[ ec |- t ]> @@ l)
39   | ECast  : ∀ Γ Δ ξ γ t1 t2 l, Δ ⊢ᴄᴏ γ : t1 ∼ t2 ->  Expr Γ Δ ξ (t1 @@ l)                        -> Expr Γ Δ ξ (t2 @@ l)
40   | ENote  : ∀ Γ Δ ξ t, Note                      -> Expr Γ Δ ξ t                                 -> Expr Γ Δ ξ t
41   | ETyApp : ∀ Γ Δ κ σ τ ξ l,     Γ ⊢ᴛy τ : κ -> Expr Γ Δ ξ (HaskTAll κ σ @@ l)                   -> Expr Γ Δ ξ (substT σ τ @@ l)
42   | ECoLam : ∀ Γ Δ κ σ σ₁ σ₂ ξ l,   Γ ⊢ᴛy σ₁:κ -> Γ ⊢ᴛy σ₂:κ -> Expr Γ (σ₁∼∼∼σ₂::Δ) ξ (σ @@ l)    -> Expr Γ Δ ξ (σ₁∼∼σ₂    ⇒ σ @@ l)
43   | ECoApp : ∀ Γ Δ   γ σ₁ σ₂ σ ξ l, Δ ⊢ᴄᴏ γ  : σ₁∼σ₂ -> Expr Γ Δ ξ (σ₁ ∼∼ σ₂ ⇒ σ @@ l)            -> Expr Γ Δ ξ (σ        @@l)
44   | ETyLam : ∀ Γ Δ ξ κ σ l,
45     Expr (κ::Γ) (weakCE Δ) (weakLT○ξ) (HaskTApp (weakF σ) (FreshHaskTyVar _)@@(weakL l))-> Expr Γ Δ ξ (HaskTAll κ σ @@ l)
46
47   | ECase    : forall Γ Δ ξ l tc atypes tbranches,
48                Expr Γ Δ ξ (caseType tc atypes @@ l) ->
49                Tree ??{ scb : StrongCaseBranchWithVVs tc atypes
50                             & Expr (sac_Γ scb Γ)
51                                    (sac_Δ scb Γ atypes (weakCK'' Δ))
52                                    (scbwv_ξ scb ξ l)
53                                    (weakLT' (tbranches@@l)) } ->
54                Expr Γ Δ ξ (tbranches @@ l)
55
56   | ELetRec  : ∀ Γ Δ ξ l τ vars, let ξ' := update_ξ ξ (map (fun x => ((fst x),(snd x @@ l))) (leaves vars)) in
57     ELetRecBindings Γ Δ ξ'     l vars ->
58     Expr            Γ Δ ξ' (τ@@l) ->
59     Expr            Γ Δ ξ  (τ@@l)
60
61   (* can't avoid having an additional inductive: it is a tree of Expr's, each of whose ξ depends on the type of the entire tree *)
62   with ELetRecBindings : forall Γ (Δ:CoercionEnv Γ), (VV -> LeveledHaskType Γ) -> HaskLevel Γ -> Tree ??(VV*HaskType Γ) -> Type :=
63   | ELR_nil    : ∀ Γ Δ ξ l  ,                                                                 ELetRecBindings Γ Δ ξ l []
64   | ELR_leaf   : ∀ Γ Δ ξ t l v,                                        Expr Γ Δ ξ (t @@ l) -> ELetRecBindings Γ Δ ξ l [(v,t)]
65   | ELR_branch : ∀ Γ Δ ξ l t1 t2, ELetRecBindings Γ Δ ξ l t1 -> ELetRecBindings Γ Δ ξ l t2 -> ELetRecBindings Γ Δ ξ l (t1,,t2)
66   .
67
68 End HaskStrong.
69 Implicit Arguments StrongCaseBranchWithVVs [[Γ]].