X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FPCF.v;h=373519b72a1431f4b989a701480c22254fdfc427;hp=16dceef2efbde4e3205d8385ba48978504ae2361;hb=1a2754d2e135ef3c5fd7ef817e1129af93b533a5;hpb=d6342fb07462cc126df948459ce98ea9caadb95c diff --git a/src/PCF.v b/src/PCF.v index 16dceef..373519b 100644 --- a/src/PCF.v +++ b/src/PCF.v @@ -29,7 +29,8 @@ Require Import Coherence_ch7_8. Require Import HaskKinds. Require Import HaskCoreTypes. -Require Import HaskLiteralsAndTyCons. +Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskStrongTypes. Require Import HaskProof. Require Import NaturalDeduction. @@ -60,9 +61,9 @@ Section PCF. Context {ndr_systemfc:@ND_Relation _ Rule}. Context Γ (Δ:CoercionEnv Γ). - Definition PCFJudg (ec:HaskTyVar Γ ★) := + Definition PCFJudg (ec:HaskTyVar Γ ECKind) := @prod (Tree ??(HaskType Γ ★)) (Tree ??(HaskType Γ ★)). - Definition pcfjudg (ec:HaskTyVar Γ ★) := + Definition pcfjudg (ec:HaskTyVar Γ ECKind) := @pair (Tree ??(HaskType Γ ★)) (Tree ??(HaskType Γ ★)). (* given an PCFJudg at depth (ec::depth) we can turn it into an PCFJudg @@ -73,7 +74,7 @@ Section PCF. (Σ,τ) => Γ > Δ > (Σ@@@(ec::nil)) |- (mapOptionTree (fun t => HaskBrak ec t) τ @@@ nil) end. - Definition pcf_vars {Γ}(ec:HaskTyVar Γ ★)(t:Tree ??(LeveledHaskType Γ ★)) : Tree ??(HaskType Γ ★) + Definition pcf_vars {Γ}(ec:HaskTyVar Γ ECKind)(t:Tree ??(LeveledHaskType Γ ★)) : Tree ??(HaskType Γ ★) := mapOptionTreeAndFlatten (fun lt => match lt with t @@ l => match l with | ec'::nil => if eqd_dec ec ec' then [t] else [] @@ -90,34 +91,13 @@ Section PCF. [((pcf_vars ec Σ) , τ )] [Γ > Δ > Σ |- (mapOptionTree (HaskBrak ec) τ @@@ lev)]. - Definition fc_vars {Γ}(ec:HaskTyVar Γ ★)(t:Tree ??(LeveledHaskType Γ ★)) : Tree ??(HaskType Γ ★) - := mapOptionTreeAndFlatten (fun lt => - match lt with t @@ l => match l with - | ec'::nil => if eqd_dec ec ec' then [] else [t] - | _ => [] - end - end) t. - - Definition FCJudg := - @prod (Tree ??(LeveledHaskType Γ ★)) (Tree ??(LeveledHaskType Γ ★)). - Definition fcjudg2judg (fc:FCJudg) := - match fc with - (x,y) => Γ > Δ > x |- y - end. - Coercion fcjudg2judg : FCJudg >-> Judg. - Definition pcfjudg2judg ec (cj:PCFJudg ec) := match cj with (Σ,τ) => Γ > Δ > (Σ @@@ (ec::nil)) |- (τ @@@ (ec::nil)) end. - Definition pcfjudg2fcjudg ec (fc:PCFJudg ec) : FCJudg := - match fc with - (x,y) => (x @@@ (ec::nil),y @@@ (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 (ec:HaskTyVar Γ ★) + Inductive Rule_PCF (ec:HaskTyVar Γ ECKind) : 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)) @@ -171,13 +151,6 @@ Section PCF. apply (Prelude_error "mkBrak got multi-leaf succedent"). Defined. - (* - Definition Partition {Γ} ec (Σ:Tree ??(LeveledHaskType Γ ★)) := - { vars:(_ * _) | - fc_vars ec Σ = fst vars /\ - 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). intros. @@ -192,120 +165,6 @@ Section PCF. { ndr_eqv := fun a b f g => (pcfToND _ _ _ _ _ f) === (pcfToND _ _ _ _ _ g) }. Admitted. - (* - * An intermediate representation necessitated by Coq's termination - * conditions. This is basically a tree where each node is a - * subproof which is either entirely level-1 or entirely level-0 - *) - Inductive Alternating : Tree ??Judg -> Type := - - | alt_nil : Alternating [] - - | alt_branch : forall a b, - Alternating a -> Alternating b -> Alternating (a,,b) - - | alt_fc : forall h c, - Alternating h -> - ND Rule h c -> - Alternating c - - | alt_pcf : forall Γ Δ ec h c h' 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 : - 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:ClosedSIND(Rule:=Rule) [Γ > Δ > Σ |- τ]), Alternating [Γ > Δ > Σ |- τ]. - - refine ( - fix orgify_fc' Γ Δ Σ τ (pf:ClosedSIND [Γ > Δ > Σ |- τ]) {struct pf} : Alternating [Γ > Δ > Σ |- τ] := - let case_main := tt in _ - 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:ClosedSIND (mapOptionTree (pcfjudg2judg Γ Δ ec) pcfj)) {struct pf} : Alternating j := - let case_pcf := tt in _ - for orgify_fc'). - - destruct case_main. - inversion pf; subst. - set (alt_fc _ _ (orgify_fc _ X) (nd_rule X0)) as backup. - refine (match X0 as R in Rule H C return - match C with - | T_Leaf (Some (Γ > Δ > Σ |- τ)) => - h=H -> Alternating [Γ > Δ > Σ |- τ] -> Alternating [Γ > Δ > Σ |- τ] - | _ => True - end - with - | RBrak Σ a b c n m => let case_RBrak := tt in fun pf' backup => _ - | REsc Σ a b c n m => let case_REsc := tt in fun pf' backup => _ - | _ => fun pf' x => x - end (refl_equal _) backup). - clear backup0 backup. - - destruct case_RBrak. - rename c into ec. - set (@match_leaf Σ0 a ec n [b] m) as q. - set (orgify_pcf Σ0 a ec _ _ q) as q'. - apply q'. - simpl. - rewrite pf' in X. - apply magic in X. - apply X. - - destruct case_REsc. - apply (Prelude_error "encountered Esc in wrong side of mkalt"). - - destruct case_leaf. - apply orgify_fc'. - rewrite eqpf. - apply pf. - - destruct case_branch. - rewrite <- eqpf in pf. - inversion pf; subst. - apply no_rules_with_multiple_conclusions in X0. - inversion X0. - exists b1. exists b2. - auto. - apply (alt_branch _ _ (orgify_fc _ X) (orgify_fc _ X0)). - - destruct case_pcf. - Admitted. - - Definition pcfify Γ Δ ec : forall Σ τ, - ClosedSIND(Rule:=Rule) [ Γ > Δ > Σ@@@(ec::nil) |- τ @@@ (ec::nil)] - -> ND (PCFRule Γ Δ ec) [] [(Σ,τ)]. - - refine (( - 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 _ - end (refl_equal _)))). - intros. - inversion H. - intros. - destruct c; try destruct o; inversion H. - destruct j. - Admitted. -*) - Hint Constructors Rule_Flat. Definition PCF_Arrange {Γ}{Δ}{lev} : forall x y z, Arrange x y -> ND (PCFRule Γ Δ lev) [(x,z)] [(y,z)]. @@ -324,8 +183,8 @@ Section PCF. eapply nd_prod. apply nd_id. apply (PCF_Arrange [h] ([],,[h]) [h0]). - apply RuCanL. - eapply nd_comp; [ idtac | apply (PCF_Arrange ([],,a) a [h0]); apply RCanL ]. + apply AuCanL. + eapply nd_comp; [ idtac | apply (PCF_Arrange ([],,a) a [h0]); apply ACanL ]. apply nd_rule. (* set (@RLet Γ Δ [] (a@@@(ec::nil)) h0 h (ec::nil)) as q. @@ -382,27 +241,27 @@ Section PCF. ; cnd_expand_right := fun a b c => PCF_right Γ Δ lev c a b }. intros; apply nd_rule. unfold PCFRule. simpl. - exists (RArrange _ _ _ _ _ (RCossa _ _ _)). + exists (RArrange _ _ _ _ _ (AuAssoc _ _ _)). apply (PCF_RArrange _ _ lev ((a,,b),,c) (a,,(b,,c)) x). intros; apply nd_rule. unfold PCFRule. simpl. - exists (RArrange _ _ _ _ _ (RAssoc _ _ _)). + exists (RArrange _ _ _ _ _ (AAssoc _ _ _)). apply (PCF_RArrange _ _ lev (a,,(b,,c)) ((a,,b),,c) x). intros; apply nd_rule. unfold PCFRule. simpl. - exists (RArrange _ _ _ _ _ (RCanL _)). + exists (RArrange _ _ _ _ _ (ACanL _)). apply (PCF_RArrange _ _ lev ([],,a) _ _). intros; apply nd_rule. unfold PCFRule. simpl. - exists (RArrange _ _ _ _ _ (RCanR _)). + exists (RArrange _ _ _ _ _ (ACanR _)). apply (PCF_RArrange _ _ lev (a,,[]) _ _). intros; apply nd_rule. unfold PCFRule. simpl. - exists (RArrange _ _ _ _ _ (RuCanL _)). + exists (RArrange _ _ _ _ _ (AuCanL _)). apply (PCF_RArrange _ _ lev _ ([],,a) _). intros; apply nd_rule. unfold PCFRule. simpl. - exists (RArrange _ _ _ _ _ (RuCanR _)). + exists (RArrange _ _ _ _ _ (AuCanR _)). apply (PCF_RArrange _ _ lev _ (a,,[]) _). Defined.