X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskProof.v;h=d9c6a7eef94758a8327913aa74aa3b2c71a4afda;hp=a5e4abd42709b1a70b22584475550be8b4d0429b;hb=93ac0d63048027161f816c451a7954fb8a6470c0;hpb=6ef9f270b138fc7aab48013d55a8192ff022c0f1 diff --git a/src/HaskProof.v b/src/HaskProof.v index a5e4abd..d9c6a7e 100644 --- a/src/HaskProof.v +++ b/src/HaskProof.v @@ -14,7 +14,8 @@ Require Import Coq.Strings.String. Require Import Coq.Lists.List. Require Import HaskKinds. Require Import HaskCoreTypes. -Require Import HaskLiteralsAndTyCons. +Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskStrongTypes. Require Import HaskWeakVars. @@ -44,6 +45,7 @@ Implicit Arguments ProofCaseBranch [ ]. (* Figure 3, production $\vdash_E$, Uniform rules *) Inductive Arrange {T} : Tree ??T -> Tree ??T -> Type := +| RId : forall a , Arrange a a | RCanL : forall a , Arrange ( [],,a ) ( a ) | RCanR : forall a , Arrange ( a,,[] ) ( a ) | RuCanL : forall a , Arrange ( a ) ( [],,a ) @@ -73,14 +75,14 @@ Inductive Rule : Tree ??Judg -> Tree ??Judg -> Type := (* SystemFC rules *) | RVar : ∀ Γ Δ σ l, Rule [ ] [Γ>Δ> [σ@@l] |- [σ @@l]] -| RGlobal : ∀ Γ Δ τ l, WeakExprVar -> Rule [ ] [Γ>Δ> [] |- [τ @@l]] +| RGlobal : forall Γ Δ l (g:Global Γ) v, Rule [ ] [Γ>Δ> [] |- [g v @@l]] | RLam : forall Γ Δ Σ (tx:HaskType Γ ★) te l, Rule [Γ>Δ> Σ,,[tx@@l]|- [te@@l] ] [Γ>Δ> Σ |- [tx--->te @@l]] | RCast : forall Γ Δ Σ (σ₁ σ₂:HaskType Γ ★) l, HaskCoercion Γ Δ (σ₁∼∼∼σ₂) -> Rule [Γ>Δ> Σ |- [σ₁@@l] ] [Γ>Δ> Σ |- [σ₂ @@l]] | RJoin : ∀ Γ Δ Σ₁ Σ₂ τ₁ τ₂ , Rule ([Γ > Δ > Σ₁ |- τ₁ ],,[Γ > Δ > Σ₂ |- τ₂ ]) [Γ>Δ> Σ₁,,Σ₂ |- τ₁,,τ₂ ] -| RApp : ∀ Γ Δ Σ₁ Σ₂ tx te l, Rule ([Γ>Δ> Σ₁ |- [tx--->te @@l]],,[Γ>Δ> Σ₂ |- [tx@@l]]) [Γ>Δ> Σ₁,,Σ₂ |- [te @@l]] +| RApp : ∀ Γ Δ Σ₁ Σ₂ tx te l, Rule ([Γ>Δ> Σ₁ |- [tx@@l]],,[Γ>Δ> Σ₂ |- [tx--->te @@l]]) [Γ>Δ> Σ₁,,Σ₂ |- [te @@l]] | RLet : ∀ Γ Δ Σ₁ Σ₂ σ₁ σ₂ l, Rule ([Γ>Δ> Σ₂ |- [σ₂@@l]],,[Γ>Δ> Σ₁,,[σ₂@@l] |- [σ₁@@l] ]) [Γ>Δ> Σ₁,,Σ₂ |- [σ₁ @@l]] @@ -168,4 +170,41 @@ Lemma systemfc_all_rules_one_conclusion : forall h c1 c2 (r:Rule h (c1,,c2)), Fa auto. Qed. - +(* "Arrange" objects are parametric in the type of the leaves of the tree *) +Definition arrangeMap : + forall {T} (Σ₁ Σ₂:Tree ??T) {R} (f:T -> R), + Arrange Σ₁ Σ₂ -> + Arrange (mapOptionTree f Σ₁) (mapOptionTree f Σ₂). + intros. + induction X; simpl. + apply RId. + apply RCanL. + apply RCanR. + apply RuCanL. + apply RuCanR. + apply RAssoc. + apply RCossa. + apply RExch. + apply RWeak. + apply RCont. + apply RLeft; auto. + apply RRight; auto. + eapply RComp; [ apply IHX1 | apply IHX2 ]. + Defined. + +(* a frequently-used Arrange *) +Definition arrangeSwapMiddle {T} (a b c d:Tree ??T) : + Arrange ((a,,b),,(c,,d)) ((a,,c),,(b,,d)). + eapply RComp. + apply RCossa. + eapply RComp. + eapply RLeft. + eapply RComp. + eapply RAssoc. + eapply RRight. + apply RExch. + eapply RComp. + eapply RLeft. + eapply RCossa. + eapply RAssoc. + Defined.