X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskProofStratified.v;h=177bd6d3911e9bacc798cf47077a93d71976fef1;hp=8f70b3113bd9da4cab63f7c572410a212b61b8be;hb=9e7ea73d3a6f4bbfba279164a806490cf95efec4;hpb=e3e2ce9cb83acdd8191049b4e9bd3d4fcf6a4db4 diff --git a/src/HaskProofStratified.v b/src/HaskProofStratified.v index 8f70b31..177bd6d 100644 --- a/src/HaskProofStratified.v +++ b/src/HaskProofStratified.v @@ -55,18 +55,22 @@ Open Scope nd_scope. *) Section HaskProofStratified. + Section PCF. + Context (ndr_systemfc:@ND_Relation _ Rule). - Inductive PCFJudg Γ (Δ:CoercionEnv Γ) (ec:HaskTyVar Γ ★) := - pcfjudg : Tree ??(HaskType Γ ★) -> Tree ??(HaskType Γ ★) -> PCFJudg Γ Δ ec. - Implicit Arguments pcfjudg [ [Γ] [Δ] [ec] ]. + Context Γ (Δ:CoercionEnv Γ). + Definition PCFJudg (ec:HaskTyVar Γ ★) := + @prod (Tree ??(HaskType Γ ★)) (Tree ??(HaskType Γ ★)). + Definition pcfjudg (ec:HaskTyVar Γ ★) := + @pair (Tree ??(HaskType Γ ★)) (Tree ??(HaskType Γ ★)). (* given an PCFJudg at depth (ec::depth) we can turn it into an PCFJudg * from depth (depth) by wrapping brackets around everything in the * succedent and repopulating *) - Definition brakify {Γ}{Δ}{ec} (j:PCFJudg Γ Δ ec) : Judg := + Definition brakify {ec} (j:PCFJudg ec) : Judg := match j with - pcfjudg Σ τ => Γ > Δ > (Σ@@@(ec::nil)) |- (mapOptionTree (fun t => HaskBrak ec t) τ @@@ nil) + (Σ,τ) => Γ > Δ > (Σ@@@(ec::nil)) |- (mapOptionTree (fun t => HaskBrak ec t) τ @@@ nil) end. Definition pcf_vars {Γ}(ec:HaskTyVar Γ ★)(t:Tree ??(LeveledHaskType Γ ★)) : Tree ??(HaskType Γ ★) @@ -77,13 +81,13 @@ Section HaskProofStratified. end end) t. - Inductive MatchingJudgments {Γ}{Δ}{ec} : Tree ??(PCFJudg Γ Δ ec) -> Tree ??Judg -> Type := + Inductive MatchingJudgments {ec} : Tree ??(PCFJudg ec) -> Tree ??Judg -> Type := | match_nil : MatchingJudgments [] [] | match_branch : forall a b c d, MatchingJudgments a b -> MatchingJudgments c d -> MatchingJudgments (a,,c) (b,,d) | match_leaf : forall Σ τ lev, MatchingJudgments - [pcfjudg (pcf_vars ec Σ) τ ] + [((pcf_vars ec Σ) , τ )] [Γ > Δ > Σ |- (mapOptionTree (HaskBrak ec) τ @@@ lev)]. Definition fc_vars {Γ}(ec:HaskTyVar Γ ★)(t:Tree ??(LeveledHaskType Γ ★)) : Tree ??(HaskType Γ ★) @@ -94,56 +98,69 @@ Section HaskProofStratified. end end) t. - Definition pcfjudg2judg {Γ}{Δ:CoercionEnv Γ} ec (cj:PCFJudg Γ Δ ec) := - match cj with pcfjudg Σ τ => Γ > Δ > (Σ @@@ (ec::nil)) |- (τ @@@ (ec::nil)) end. + Definition pcfjudg2judg ec (cj:PCFJudg ec) := + match cj with (Σ,τ) => Γ > Δ > (Σ @@@ (ec::nil)) |- (τ @@@ (ec::nil)) end. (* Rules allowed in PCF; i.e. rules we know how to turn into GArrows *) (* Rule_PCF consists of the rules allowed in flat PCF: everything except *) (* AppT, AbsT, AppC, AbsC, Cast, Global, and some Case statements *) - Inductive Rule_PCF {Γ}{Δ:CoercionEnv Γ} (ec:HaskTyVar Γ ★) - : forall (h c:Tree ??(PCFJudg Γ Δ ec)), Rule (mapOptionTree (pcfjudg2judg ec) h) (mapOptionTree (pcfjudg2judg ec) c) -> Type := - | PCF_RArrange : ∀ x y t a, Rule_PCF ec [pcfjudg _ _ ] [ pcfjudg _ _ ] (RArrange Γ Δ (x@@@(ec::nil)) (y@@@(ec::nil)) (t@@@(ec::nil)) a) - | PCF_RLit : ∀ lit , Rule_PCF ec [ ] [ pcfjudg [] [_] ] (RLit Γ Δ lit (ec::nil)) - | PCF_RNote : ∀ Σ τ n , Rule_PCF ec [pcfjudg _ [_]] [ pcfjudg _ [_] ] (RNote Γ Δ (Σ@@@(ec::nil)) τ (ec::nil) n) - | PCF_RVar : ∀ σ , Rule_PCF ec [ ] [ pcfjudg [_] [_] ] (RVar Γ Δ σ (ec::nil) ) - | PCF_RLam : ∀ Σ tx te , Rule_PCF ec [pcfjudg (_,,[_]) [_] ] [ pcfjudg _ [_] ] (RLam Γ Δ (Σ@@@(ec::nil)) tx te (ec::nil) ) + Inductive Rule_PCF (ec:HaskTyVar Γ ★) + : forall (h c:Tree ??(PCFJudg ec)), Rule (mapOptionTree (pcfjudg2judg ec) h) (mapOptionTree (pcfjudg2judg ec) c) -> Type := + | PCF_RArrange : ∀ x y t a, Rule_PCF ec [(_, _)] [(_, _)] (RArrange Γ Δ (x@@@(ec::nil)) (y@@@(ec::nil)) (t@@@(ec::nil)) a) + | PCF_RLit : ∀ lit , Rule_PCF ec [ ] [ ([],[_]) ] (RLit Γ Δ lit (ec::nil)) + | PCF_RNote : ∀ Σ τ n , Rule_PCF ec [(_,[_])] [(_,[_])] (RNote Γ Δ (Σ@@@(ec::nil)) τ (ec::nil) n) + | PCF_RVar : ∀ σ , Rule_PCF ec [ ] [([_],[_])] (RVar Γ Δ σ (ec::nil) ) + | PCF_RLam : ∀ Σ tx te , Rule_PCF ec [((_,,[_]),[_])] [(_,[_])] (RLam Γ Δ (Σ@@@(ec::nil)) tx te (ec::nil) ) | PCF_RApp : ∀ Σ Σ' tx te , - Rule_PCF ec ([pcfjudg _ [_]],,[pcfjudg _ [_]]) [pcfjudg (_,,_) [_]] + Rule_PCF ec ([(_,[_])],,[(_,[_])]) [((_,,_),[_])] (RApp Γ Δ (Σ@@@(ec::nil))(Σ'@@@(ec::nil)) tx te (ec::nil)) | PCF_RLet : ∀ Σ Σ' σ₂ p, - Rule_PCF ec ([pcfjudg _ [_]],,[pcfjudg (_,,[_]) [_]]) [pcfjudg (_,,_) [_]] + Rule_PCF ec ([(_,[_])],,[((_,,[_]),[_])]) [((_,,_),[_])] (RLet Γ Δ (Σ@@@(ec::nil)) (Σ'@@@(ec::nil)) σ₂ p (ec::nil)) - | PCF_REmptyGroup : Rule_PCF ec [ ] [ pcfjudg [] [] ] (REmptyGroup Γ Δ ) + | PCF_RVoid : Rule_PCF ec [ ] [([],[])] (RVoid Γ Δ ) (*| PCF_RLetRec : ∀ Σ₁ τ₁ τ₂ , Rule_PCF (ec::nil) _ _ (RLetRec Γ Δ Σ₁ τ₁ τ₂ (ec::nil) )*) - | PCF_RBindingGroup : ∀ Σ₁ Σ₂ τ₁ τ₂, Rule_PCF ec ([pcfjudg _ _],,[pcfjudg _ _]) [pcfjudg (_,,_) (_,,_)] - (RBindingGroup Γ Δ (Σ₁@@@(ec::nil)) (Σ₂@@@(ec::nil)) (τ₁@@@(ec::nil)) (τ₂@@@(ec::nil))). + | PCF_RJoin : ∀ Σ₁ Σ₂ τ₁ τ₂, Rule_PCF ec ([(_,_)],,[(_,_)]) [((_,,_),(_,,_))] + (RJoin Γ Δ (Σ₁@@@(ec::nil)) (Σ₂@@@(ec::nil)) (τ₁@@@(ec::nil)) (τ₂@@@(ec::nil))). (* need int/boolean case *) Implicit Arguments Rule_PCF [ ]. - Definition PCFRule Γ Δ lev h c := { r:_ & @Rule_PCF Γ Δ lev h c r }. + Definition PCFRule lev h c := { r:_ & @Rule_PCF lev h c r }. + End PCF. + + Definition FCJudg Γ (Δ:CoercionEnv Γ) := + @prod (Tree ??(LeveledHaskType Γ ★)) (Tree ??(LeveledHaskType Γ ★)). + Definition fcjudg2judg {Γ}{Δ}(fc:FCJudg Γ Δ) := + match fc with + (x,y) => Γ > Δ > x |- y + end. + Coercion fcjudg2judg : FCJudg >-> Judg. + + Definition pcfjudg2fcjudg {Γ}{Δ} ec (fc:PCFJudg Γ ec) : FCJudg Γ Δ := + match fc with + (x,y) => (x @@@ (ec::nil),y @@@ (ec::nil)) + end. (* An organized deduction has been reorganized into contiguous blocks whose * hypotheses (if any) and conclusion have the same Γ and Δ and a fixed nesting depth. The boolean * indicates if non-PCF rules have been used *) - Inductive OrgR : Tree ??Judg -> Tree ??Judg -> Type := + Inductive OrgR Γ Δ : Tree ??(FCJudg Γ Δ) -> Tree ??(FCJudg Γ Δ) -> Type := - | org_fc : forall h c (r:Rule h c), + | org_fc : forall (h c:Tree ??(FCJudg Γ Δ)) + (r:Rule (mapOptionTree fcjudg2judg h) (mapOptionTree fcjudg2judg c)), Rule_Flat r -> - OrgR h c + OrgR _ _ h c - | org_pcf : forall Γ Δ ec h h' c c', - MatchingJudgments h h' -> - MatchingJudgments c c' -> - ND (PCFRule Γ Δ ec) h c -> - OrgR h' c'. + | org_pcf : forall ec h c, + ND (PCFRule Γ Δ ec) h c -> + OrgR Γ Δ (mapOptionTree (pcfjudg2fcjudg ec) h) (mapOptionTree (pcfjudg2fcjudg ec) c). - Definition mkEsc {Γ}{Δ}{ec}(h:Tree ??(PCFJudg Γ Δ ec)) + Definition mkEsc Γ Δ ec (h:Tree ??(PCFJudg Γ ec)) : ND Rule - (mapOptionTree brakify h) - (mapOptionTree (pcfjudg2judg ec) h). + (mapOptionTree (brakify Γ Δ) h) + (mapOptionTree (pcfjudg2judg Γ Δ ec) h). apply nd_replicate; intros. destruct o; simpl in *. induction t0. @@ -154,10 +171,10 @@ Section HaskProofStratified. apply (Prelude_error "mkEsc got multi-leaf succedent"). Defined. - Definition mkBrak {Γ}{Δ}{ec}(h:Tree ??(PCFJudg Γ Δ ec)) + Definition mkBrak Γ Δ ec (h:Tree ??(PCFJudg Γ ec)) : ND Rule - (mapOptionTree (pcfjudg2judg ec) h) - (mapOptionTree brakify h). + (mapOptionTree (pcfjudg2judg Γ Δ ec) h) + (mapOptionTree (brakify Γ Δ) h). apply nd_replicate; intros. destruct o; simpl in *. induction t0. @@ -175,8 +192,8 @@ Section HaskProofStratified. pcf_vars ec Σ = snd vars }. *) - Definition pcfToND : forall Γ Δ ec h c, - ND (PCFRule Γ Δ ec) h c -> ND Rule (mapOptionTree (pcfjudg2judg ec) h) (mapOptionTree (pcfjudg2judg ec) c). + Definition pcfToND Γ Δ : forall ec h c, + ND (PCFRule Γ Δ ec) h c -> ND Rule (mapOptionTree (pcfjudg2judg Γ Δ ec) h) (mapOptionTree (pcfjudg2judg Γ Δ ec) c). intros. eapply (fun q => nd_map' _ q X). intros. @@ -186,21 +203,8 @@ Section HaskProofStratified. Defined. Instance OrgPCF Γ Δ lev : @ND_Relation _ (PCFRule Γ Δ lev) := - { ndr_eqv := fun a b f g => (pcfToND _ _ _ _ _ f) === (pcfToND _ _ _ _ _ g) }. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - Defined. + { ndr_eqv := fun a b f g => (pcfToND _ _ _ _ _ f) === (pcfToND _ _ _ _ _ g) }. + Admitted. (* * An intermediate representation necessitated by Coq's termination @@ -220,33 +224,33 @@ Section HaskProofStratified. Alternating c | alt_pcf : forall Γ Δ ec h c h' c', - MatchingJudgments h h' -> - MatchingJudgments c c' -> + MatchingJudgments Γ Δ h h' -> + MatchingJudgments Γ Δ c c' -> Alternating h' -> ND (PCFRule Γ Δ ec) h c -> Alternating c'. Require Import Coq.Logic.Eqdep. - +(* Lemma magic a b c d ec e : - ClosedND(Rule:=Rule) [a > b > c |- [d @@ (ec :: e)]] -> - ClosedND(Rule:=Rule) [a > b > pcf_vars ec c @@@ (ec :: nil) |- [d @@ (ec :: nil)]]. - admit. - Defined. + ClosedSIND(Rule:=Rule) [a > b > c |- [d @@ (ec :: e)]] -> + ClosedSIND(Rule:=Rule) [a > b > pcf_vars ec c @@@ (ec :: nil) |- [d @@ (ec :: nil)]]. + admit. + Defined. - Definition orgify : forall Γ Δ Σ τ (pf:ClosedND(Rule:=Rule) [Γ > Δ > Σ |- τ]), Alternating [Γ > Δ > Σ |- τ]. + Definition orgify : forall Γ Δ Σ τ (pf:ClosedSIND(Rule:=Rule) [Γ > Δ > Σ |- τ]), Alternating [Γ > Δ > Σ |- τ]. refine ( - fix orgify_fc' Γ Δ Σ τ (pf:ClosedND [Γ > Δ > Σ |- τ]) {struct pf} : Alternating [Γ > Δ > Σ |- τ] := + fix orgify_fc' Γ Δ Σ τ (pf:ClosedSIND [Γ > Δ > Σ |- τ]) {struct pf} : Alternating [Γ > Δ > Σ |- τ] := let case_main := tt in _ - with orgify_fc c (pf:ClosedND c) {struct pf} : Alternating c := + with orgify_fc c (pf:ClosedSIND c) {struct pf} : Alternating c := (match c as C return C=c -> Alternating C with | T_Leaf None => fun _ => alt_nil | T_Leaf (Some (Γ > Δ > Σ |- τ)) => let case_leaf := tt in fun eqpf => _ | T_Branch b1 b2 => let case_branch := tt in fun eqpf => _ end (refl_equal _)) - with orgify_pcf Γ Δ ec pcfj j (m:MatchingJudgments pcfj j) - (pf:ClosedND (mapOptionTree (pcfjudg2judg ec) pcfj)) {struct pf} : Alternating j := + with orgify_pcf Γ Δ ec pcfj j (m:MatchingJudgments Γ Δ pcfj j) + (pf:ClosedSIND (mapOptionTree (pcfjudg2judg Γ Δ ec) pcfj)) {struct pf} : Alternating j := let case_pcf := tt in _ for orgify_fc'). @@ -297,13 +301,13 @@ Section HaskProofStratified. Admitted. Definition pcfify Γ Δ ec : forall Σ τ, - ClosedND(Rule:=Rule) [ Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)] - -> ND (PCFRule Γ Δ ec) [] [pcfjudg Σ τ]. + ClosedSIND(Rule:=Rule) [ Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)] + -> ND (PCFRule Γ Δ ec) [] [(Σ,τ)]. refine (( - fix pcfify Σ τ (pn:@ClosedND _ Rule [ Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)]) {struct pn} - : ND (PCFRule Γ Δ ec) [] [pcfjudg Σ τ] := - (match pn in @ClosedND _ _ J return J=[Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)] -> _ with + fix pcfify Σ τ (pn:@ClosedSIND _ Rule [ Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)]) {struct pn} + : ND (PCFRule Γ Δ ec) [] [(Σ,τ)] := + (match pn in @ClosedSIND _ _ J return J=[Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)] -> _ with | cnd_weak => let case_nil := tt in _ | cnd_rule h c cnd' r => let case_rule := tt in _ | cnd_branch _ _ c1 c2 => let case_branch := tt in _ @@ -314,15 +318,14 @@ Section HaskProofStratified. destruct c; try destruct o; inversion H. destruct j. Admitted. - +*) (* any proof in organized form can be "dis-organized" *) - Definition unOrgR : forall h c, OrgR h c -> ND Rule h c. + (* + Definition unOrgR : forall Γ Δ h c, OrgR Γ Δ h c -> ND Rule h c. intros. - induction X. apply nd_rule. apply r. - eapply nd_comp. (* apply (mkEsc h). @@ -331,52 +334,16 @@ Section HaskProofStratified. apply n. *) Admitted. - - Definition unOrgND h c : ND OrgR h c -> ND Rule h c := nd_map unOrgR. + Definition unOrgND Γ Δ h c : ND (OrgR Γ Δ) h c -> ND Rule h c := nd_map (unOrgR Γ Δ). + *) - Instance OrgNDR : @ND_Relation _ OrgR := - { ndr_eqv := fun a b f g => (unOrgND _ _ f) === (unOrgND _ _ g) }. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - admit. - Defined. - Hint Constructors Rule_Flat. - Instance PCF_sequents Γ Δ lev : @SequentCalculus _ (PCFRule Γ Δ lev) _ pcfjudg. - apply Build_SequentCalculus. - intros. - induction a. - destruct a; simpl. - apply nd_rule. - exists (RVar _ _ _ _). - apply PCF_RVar. - apply nd_rule. - exists (REmptyGroup _ _ ). - apply PCF_REmptyGroup. - eapply nd_comp. - eapply nd_comp; [ apply nd_llecnac | idtac ]. - apply (nd_prod IHa1 IHa2). - apply nd_rule. - exists (RBindingGroup _ _ _ _ _ _). - apply PCF_RBindingGroup. - Defined. - - Definition PCF_Arrange {Γ}{Δ}{lev} : forall x y z, Arrange x y -> ND (PCFRule Γ Δ lev) [pcfjudg x z] [pcfjudg y z]. + Definition PCF_Arrange {Γ}{Δ}{lev} : forall x y z, Arrange x y -> ND (PCFRule Γ Δ lev) [(x,z)] [(y,z)]. admit. Defined. - Definition PCF_cut Γ Δ lev : forall a b c, ND (PCFRule Γ Δ lev) ([ pcfjudg a b ],,[ pcfjudg b c ]) [ pcfjudg a c ]. + Definition PCF_cut Γ Δ lev : forall a b c, ND (PCFRule Γ Δ lev) ([(a,b)],,[(b,c)]) [(a,c)]. intros. destruct b. destruct o. @@ -391,7 +358,7 @@ Section HaskProofStratified. apply RuCanL. eapply nd_comp; [ idtac | apply (PCF_Arrange ([],,a) a [h0]); apply RCanL ]. apply nd_rule. -(* + (* set (@RLet Γ Δ [] (a@@@(ec::nil)) h0 h (ec::nil)) as q. exists q. apply (PCF_RLet _ [] a h0 h). @@ -399,103 +366,93 @@ Section HaskProofStratified. apply (Prelude_error "cut rule invoked with [a|=[b]] [[b]|=[x,,y]]"). apply (Prelude_error "cut rule invoked with [a|=[]] [[]|=c]"). apply (Prelude_error "cut rule invoked with [a|=[b,,c]] [[b,,c]|=z]"). -*) - admit. - admit. - admit. - admit. - admit. - Defined. + *) + Admitted. - Instance PCF_cutrule Γ Δ lev : CutRule (PCF_sequents Γ Δ lev) := - { nd_cut := PCF_cut Γ Δ lev }. - admit. - admit. - admit. - Defined. + Instance PCF_sequents Γ Δ lev ec : @SequentND _ (PCFRule Γ Δ lev) _ (pcfjudg Γ ec) := + { snd_cut := PCF_cut Γ Δ lev }. + apply Build_SequentND. + intros. + induction a. + destruct a; simpl. + apply nd_rule. + exists (RVar _ _ _ _). + apply PCF_RVar. + apply nd_rule. + exists (RVoid _ _ ). + apply PCF_RVoid. + eapply nd_comp. + eapply nd_comp; [ apply nd_llecnac | idtac ]. + apply (nd_prod IHa1 IHa2). + apply nd_rule. + exists (RJoin _ _ _ _ _ _). + apply PCF_RJoin. + admit. + Defined. - Definition PCF_left Γ Δ lev a b c : ND (PCFRule Γ Δ lev) [pcfjudg b c] [pcfjudg (a,,b) (a,,c)]. + Definition PCF_left Γ Δ lev a b c : ND (PCFRule Γ Δ lev) [(b,c)] [((a,,b),(a,,c))]. eapply nd_comp; [ apply nd_llecnac | eapply nd_comp; [ idtac | idtac ] ]. - eapply nd_prod; [ apply nd_seq_reflexive | apply nd_id ]. + eapply nd_prod; [ apply snd_initial | apply nd_id ]. apply nd_rule. - set (@PCF_RBindingGroup Γ Δ lev a b a c) as q'. + set (@PCF_RJoin Γ Δ lev a b a c) as q'. refine (existT _ _ _). apply q'. - Defined. + Admitted. - Definition PCF_right Γ Δ lev a b c : ND (PCFRule Γ Δ lev) [pcfjudg b c] [pcfjudg (b,,a) (c,,a)]. + Definition PCF_right Γ Δ lev a b c : ND (PCFRule Γ Δ lev) [(b,c)] [((b,,a),(c,,a))]. eapply nd_comp; [ apply nd_rlecnac | eapply nd_comp; [ idtac | idtac ] ]. - eapply nd_prod; [ apply nd_id | apply nd_seq_reflexive ]. + eapply nd_prod; [ apply nd_id | apply snd_initial ]. apply nd_rule. - set (@PCF_RBindingGroup Γ Δ lev b a c a) as q'. + set (@PCF_RJoin Γ Δ lev b a c a) as q'. refine (existT _ _ _). apply q'. - Defined. - - Instance PCF_sequent_join Γ Δ lev : @SequentExpansion _ _ _ _ _ (PCF_sequents Γ Δ lev) (PCF_cutrule Γ Δ lev) := - { se_expand_left := PCF_left Γ Δ lev - ; se_expand_right := PCF_right Γ Δ lev }. - admit. - admit. - admit. - admit. - Defined. + Admitted. - (* 5.1.3 *) - Instance PCF Γ Δ lev : @ProgrammingLanguage _ _ pcfjudg (PCFRule Γ Δ lev) := - { pl_eqv := OrgPCF Γ Δ lev - ; pl_sc := PCF_sequents Γ Δ lev - ; pl_subst := PCF_cutrule Γ Δ lev - ; pl_sequent_join := PCF_sequent_join Γ Δ lev - }. - apply Build_TreeStructuralRules; intros; unfold eqv; unfold hom; simpl. + Instance PCF_sequent_join Γ Δ lev : @ContextND _ (PCFRule Γ Δ lev) _ (pcfjudg Γ lev) _ := + { cnd_expand_left := fun a b c => PCF_left Γ Δ lev c a b + ; cnd_expand_right := fun a b c => PCF_right Γ Δ lev c a b }. - apply nd_rule. unfold PCFRule. simpl. + intros; apply nd_rule. unfold PCFRule. simpl. exists (RArrange _ _ _ _ _ (RCossa _ _ _)). - apply (PCF_RArrange lev ((a,,b),,c) (a,,(b,,c)) x). + apply (PCF_RArrange _ _ lev ((a,,b),,c) (a,,(b,,c)) x). - apply nd_rule. unfold PCFRule. simpl. + intros; apply nd_rule. unfold PCFRule. simpl. exists (RArrange _ _ _ _ _ (RAssoc _ _ _)). - apply (PCF_RArrange lev (a,,(b,,c)) ((a,,b),,c) x). + apply (PCF_RArrange _ _ lev (a,,(b,,c)) ((a,,b),,c) x). - apply nd_rule. unfold PCFRule. simpl. + intros; apply nd_rule. unfold PCFRule. simpl. exists (RArrange _ _ _ _ _ (RCanL _)). - apply (PCF_RArrange lev ([],,a) _ _). + apply (PCF_RArrange _ _ lev ([],,a) _ _). - apply nd_rule. unfold PCFRule. simpl. + intros; apply nd_rule. unfold PCFRule. simpl. exists (RArrange _ _ _ _ _ (RCanR _)). - apply (PCF_RArrange lev (a,,[]) _ _). + apply (PCF_RArrange _ _ lev (a,,[]) _ _). - apply nd_rule. unfold PCFRule. simpl. + intros; apply nd_rule. unfold PCFRule. simpl. exists (RArrange _ _ _ _ _ (RuCanL _)). - apply (PCF_RArrange lev _ ([],,a) _). + apply (PCF_RArrange _ _ lev _ ([],,a) _). - apply nd_rule. unfold PCFRule. simpl. + intros; apply nd_rule. unfold PCFRule. simpl. exists (RArrange _ _ _ _ _ (RuCanR _)). - apply (PCF_RArrange lev _ (a,,[]) _). + apply (PCF_RArrange _ _ lev _ (a,,[]) _). Defined. - Instance SystemFCa_sequents Γ Δ : @SequentCalculus _ OrgR _ (mkJudg Γ Δ). - apply Build_SequentCalculus. - intros. - induction a. - destruct a; simpl. - apply nd_rule. - destruct l. - apply org_fc with (r:=RVar _ _ _ _). - auto. - apply nd_rule. - apply org_fc with (r:=REmptyGroup _ _ ). - auto. - eapply nd_comp. - eapply nd_comp; [ apply nd_llecnac | idtac ]. - apply (nd_prod IHa1 IHa2). - apply nd_rule. - apply org_fc with (r:=RBindingGroup _ _ _ _ _ _). - auto. - Defined. + Instance OrgPCF_SequentND_Relation Γ Δ lev : SequentND_Relation (PCF_sequent_join Γ Δ lev) (OrgPCF Γ Δ lev). + admit. + Defined. + + Definition OrgPCF_ContextND_Relation Γ Δ lev + : @ContextND_Relation _ _ _ _ _ (PCF_sequent_join Γ Δ lev) (OrgPCF Γ Δ lev) (OrgPCF_SequentND_Relation Γ Δ lev). + admit. + Defined. - Definition SystemFCa_cut Γ Δ : forall a b c, ND OrgR ([ Γ > Δ > a |- b ],,[ Γ > Δ > b |- c ]) [ Γ > Δ > a |- c ]. + (* 5.1.3 *) + Instance PCF Γ Δ lev : ProgrammingLanguage := + { pl_cnd := PCF_sequent_join Γ Δ lev + ; pl_eqv := OrgPCF_ContextND_Relation Γ Δ lev + }. + + Definition SystemFCa_cut Γ Δ : forall a b c, ND (OrgR Γ Δ) ([(a,b)],,[(b,c)]) [(a,c)]. intros. destruct b. destruct o. @@ -503,11 +460,14 @@ Section HaskProofStratified. destruct o. (* when the cut is a single leaf and the RHS is a single leaf: *) + (* eapply nd_comp. eapply nd_prod. apply nd_id. eapply nd_rule. - apply org_fc with (r:=RArrange _ _ _ _ _ (RuCanL [l])). + set (@org_fc) as ofc. + set (RArrange Γ Δ _ _ _ (RuCanL [l0])) as rule. + apply org_fc with (r:=RArrange _ _ _ _ _ (RuCanL [_])). auto. eapply nd_comp; [ idtac | eapply nd_rule; apply org_fc with (r:=RArrange _ _ _ _ _ (RCanL _)) ]. apply nd_rule. @@ -518,58 +478,111 @@ Section HaskProofStratified. apply org_fc with (r:=@RLet Γ Δ [] a h1 h h2). auto. auto. + *) + admit. apply (Prelude_error "systemfc cut rule invoked with [a|=[b]] [[b]|=[]]"). apply (Prelude_error "systemfc cut rule invoked with [a|=[b]] [[b]|=[x,,y]]"). apply (Prelude_error "systemfc rule invoked with [a|=[]] [[]|=c]"). apply (Prelude_error "systemfc rule invoked with [a|=[b,,c]] [[b,,c]|=z]"). Defined. - Instance SystemFCa_cutrule Γ Δ : CutRule (SystemFCa_sequents Γ Δ) := - { nd_cut := SystemFCa_cut Γ Δ }. - admit. - admit. - admit. - Defined. + Instance SystemFCa_sequents Γ Δ : @SequentND _ (OrgR Γ Δ) _ _ := + { snd_cut := SystemFCa_cut Γ Δ }. + apply Build_SequentND. + intros. + induction a. + destruct a; simpl. + (* + apply nd_rule. + destruct l. + apply org_fc with (r:=RVar _ _ _ _). + auto. + apply nd_rule. + apply org_fc with (r:=RVoid _ _ ). + auto. + eapply nd_comp. + eapply nd_comp; [ apply nd_llecnac | idtac ]. + apply (nd_prod IHa1 IHa2). + apply nd_rule. + apply org_fc with (r:=RJoin _ _ _ _ _ _). + auto. + admit. + *) + admit. + admit. + admit. + admit. + Defined. - Definition SystemFCa_left Γ Δ a b c : ND OrgR [Γ > Δ > b |- c] [Γ > Δ > (a,,b) |- (a,,c)]. + Definition SystemFCa_left Γ Δ a b c : ND (OrgR Γ Δ) [(b,c)] [((a,,b),(a,,c))]. + admit. + (* eapply nd_comp; [ apply nd_llecnac | eapply nd_comp; [ idtac | idtac ] ]. - eapply nd_prod; [ apply nd_seq_reflexive | apply nd_id ]. + eapply nd_prod; [ apply snd_initial | apply nd_id ]. apply nd_rule. - apply org_fc with (r:=RBindingGroup Γ Δ a b a c). + apply org_fc with (r:=RJoin Γ Δ a b a c). auto. + *) Defined. - Definition SystemFCa_right Γ Δ a b c : ND OrgR [Γ > Δ > b |- c] [Γ > Δ > (b,,a) |- (c,,a)]. + Definition SystemFCa_right Γ Δ a b c : ND (OrgR Γ Δ) [(b,c)] [((b,,a),(c,,a))]. + admit. + (* eapply nd_comp; [ apply nd_rlecnac | eapply nd_comp; [ idtac | idtac ] ]. - eapply nd_prod; [ apply nd_id | apply nd_seq_reflexive ]. + eapply nd_prod; [ apply nd_id | apply snd_initial ]. apply nd_rule. - apply org_fc with (r:=RBindingGroup Γ Δ b a c a). + apply org_fc with (r:=RJoin Γ Δ b a c a). auto. + *) Defined. - Instance SystemFCa_sequent_join Γ Δ : @SequentExpansion _ _ _ _ _ (SystemFCa_sequents Γ Δ) (SystemFCa_cutrule Γ Δ) := - { se_expand_left := SystemFCa_left Γ Δ - ; se_expand_right := SystemFCa_right Γ Δ }. - admit. - admit. + Instance SystemFCa_sequent_join Γ Δ : @ContextND _ _ _ _ (SystemFCa_sequents Γ Δ) := + { cnd_expand_left := fun a b c => SystemFCa_left Γ Δ c a b + ; cnd_expand_right := fun a b c => SystemFCa_right Γ Δ c a b }. + (* + intros; apply nd_rule. simpl. + apply (org_fc _ _ _ _ ((RArrange _ _ _ _ _ (RCossa _ _ _)))). + auto. + + intros; apply nd_rule. simpl. + apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RAssoc _ _ _))); auto. + + intros; apply nd_rule. simpl. + apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RCanL _))); auto. + + intros; apply nd_rule. simpl. + apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RCanR _))); auto. + + intros; apply nd_rule. simpl. + apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RuCanL _))); auto. + + intros; apply nd_rule. simpl. + apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RuCanR _))); auto. + *) + admit. + admit. + admit. + admit. + admit. + admit. + Defined. + + Instance OrgFC Γ Δ : @ND_Relation _ (OrgR Γ Δ). + Admitted. + + Instance OrgFC_SequentND_Relation Γ Δ : SequentND_Relation (SystemFCa_sequent_join Γ Δ) (OrgFC Γ Δ). admit. + Defined. + + Definition OrgFC_ContextND_Relation Γ Δ + : @ContextND_Relation _ _ _ _ _ (SystemFCa_sequent_join Γ Δ) (OrgFC Γ Δ) (OrgFC_SequentND_Relation Γ Δ). admit. Defined. (* 5.1.2 *) - Instance SystemFCa Γ Δ : @ProgrammingLanguage _ _ (mkJudg Γ Δ) OrgR := - { pl_eqv := OrgNDR - ; pl_sc := SystemFCa_sequents Γ Δ - ; pl_subst := SystemFCa_cutrule Γ Δ - ; pl_sequent_join := SystemFCa_sequent_join Γ Δ + Instance SystemFCa Γ Δ : @ProgrammingLanguage (LeveledHaskType Γ ★) _ := + { pl_eqv := OrgFC_ContextND_Relation Γ Δ + ; pl_snd := SystemFCa_sequents Γ Δ }. - apply Build_TreeStructuralRules; intros; unfold eqv; unfold hom; simpl. - apply nd_rule. apply (org_fc _ _ (RArrange _ _ _ _ _ (RCossa a b c))). apply Flat_RArrange. - apply nd_rule. apply (org_fc _ _ (RArrange _ _ _ _ _ (RAssoc a b c))). apply Flat_RArrange. - apply nd_rule. apply (org_fc _ _ (RArrange _ _ _ _ _ (RCanL a ))). apply Flat_RArrange. - apply nd_rule. apply (org_fc _ _ (RArrange _ _ _ _ _ (RCanR a ))). apply Flat_RArrange. - apply nd_rule. apply (org_fc _ _ (RArrange _ _ _ _ _ (RuCanL a ))). apply Flat_RArrange. - apply nd_rule. apply (org_fc _ _ (RArrange _ _ _ _ _ (RuCanR a ))). apply Flat_RArrange. - Defined. End HaskProofStratified.