X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;ds=sidebyside;f=src%2FHaskStrong.v;h=611a4c8aec813069ef2b3653d59141aca1f92650;hb=30cc675d57492799644506f3632625f371a3e89a;hp=b7229d058dba92b8dd075c300156558c5fc7e1e9;hpb=8c26722a1ee110077968a8a166eb7130266b2035;p=coq-hetmet.git diff --git a/src/HaskStrong.v b/src/HaskStrong.v index b7229d0..611a4c8 100644 --- a/src/HaskStrong.v +++ b/src/HaskStrong.v @@ -8,10 +8,10 @@ Require Import General. Require Import Coq.Strings.String. Require Import Coq.Lists.List. Require Import HaskKinds. -Require Import HaskWeakVars. Require Import HaskCoreTypes. -Require Import HaskCoreLiterals. +Require Import HaskLiteralsAndTyCons. Require Import HaskStrongTypes. +Require Import HaskWeakVars. Section HaskStrong. @@ -20,8 +20,8 @@ Section HaskStrong. (* a StrongCaseBranchWithVVs contains all the data found in a case branch except the expression itself *) - Record StrongCaseBranchWithVVs {tc:TyCon}{Γ}{sac:@StrongAltCon tc}{atypes:IList _ (HaskType Γ) (tyConKind sac)} := - { scbwv_sac := sac + Record StrongCaseBranchWithVVs {tc:TyCon}{Γ}{atypes:IList _ (HaskType Γ) (tyConKind tc)} := + { scbwv_sac : @StrongAltCon tc ; scbwv_exprvars : vec VV (sac_numExprVars scbwv_sac) ; scbwv_varstypes := vec_zip scbwv_exprvars (sac_types scbwv_sac Γ atypes) ; scbwv_ξ := fun ξ lev => update_ξ (weakLT'○ξ) (vec2list @@ -31,6 +31,10 @@ Section HaskStrong. Coercion scbwv_sac : StrongCaseBranchWithVVs >-> StrongAltCon. Inductive Expr : forall Γ (Δ:CoercionEnv Γ), (VV -> LeveledHaskType Γ ★) -> LeveledHaskType Γ ★ -> Type := + + (* an "EGlobal" is any variable which is free in the expression which was passed to -fcoqpass (ie bound outside it) *) + | EGlobal: ∀ Γ Δ ξ t, WeakExprVar -> Expr Γ Δ ξ t + | EVar : ∀ Γ Δ ξ ev, Expr Γ Δ ξ (ξ ev) | ELit : ∀ Γ Δ ξ lit l, Expr Γ Δ ξ (literalType lit@@l) | EApp : ∀ Γ Δ ξ t1 t2 l, Expr Γ Δ ξ (t2--->t1 @@ l) -> Expr Γ Δ ξ (t2 @@ l) -> Expr Γ Δ ξ (t1 @@ l) @@ -48,9 +52,9 @@ Section HaskStrong. Expr Γ Δ ξ (σ₁ ∼∼ σ₂ ⇒ σ @@ l) -> Expr Γ Δ ξ (σ @@l) | ETyLam : ∀ Γ Δ ξ κ σ l, Expr (κ::Γ) (weakCE Δ) (weakLT○ξ) (HaskTApp (weakF σ) (FreshHaskTyVar _)@@(weakL l))-> Expr Γ Δ ξ (HaskTAll κ σ @@ l) - | ECase : forall Γ Δ ξ l tc tbranches sac atypes, + | ECase : forall Γ Δ ξ l tc tbranches atypes, Expr Γ Δ ξ (caseType tc atypes @@ l) -> - Tree ??{ scb : StrongCaseBranchWithVVs tc sac atypes + Tree ??{ scb : StrongCaseBranchWithVVs tc atypes & Expr (sac_Γ scb Γ) (sac_Δ scb Γ atypes (weakCK'' Δ)) (scbwv_ξ scb ξ l) @@ -65,7 +69,7 @@ Section HaskStrong. (* 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 *) with ELetRecBindings : ∀ Γ, CoercionEnv Γ -> (VV -> LeveledHaskType Γ ★) -> HaskLevel Γ -> Tree ??(VV*HaskType Γ ★) -> Type := | ELR_nil : ∀ Γ Δ ξ l , ELetRecBindings Γ Δ ξ l [] - | ELR_leaf : ∀ Γ Δ ξ t l v, Expr Γ Δ ξ (t @@ l) -> ELetRecBindings Γ Δ ξ l [(v,t)] + | ELR_leaf : ∀ Γ Δ ξ l v, Expr Γ Δ ξ (unlev (ξ v) @@ l) -> ELetRecBindings Γ Δ ξ l [(v,unlev (ξ v))] | ELR_branch : ∀ Γ Δ ξ l t1 t2, ELetRecBindings Γ Δ ξ l t1 -> ELetRecBindings Γ Δ ξ l t2 -> ELetRecBindings Γ Δ ξ l (t1,,t2) .