X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambda%2Fpointer_order.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda%2Fpointer_order.ma;h=0b738215a59b3d6764293418065d2595986dd445;hb=bde429ac54e48de74b3d8b1df72dfcb86aa9bae5;hp=92618412e635092d8043edaef6227a64d2660a10;hpb=4bb6799d029b7b377f7aa28b0e90f0a69c149a9c;p=helm.git diff --git a/matita/matita/contribs/lambda/pointer_order.ma b/matita/matita/contribs/lambda/pointer_order.ma index 92618412e..0b738215a 100644 --- a/matita/matita/contribs/lambda/pointer_order.ma +++ b/matita/matita/contribs/lambda/pointer_order.ma @@ -20,7 +20,8 @@ include "pointer.ma". to serve as base for the order relations: p < q and p ≤ q *) inductive pprec: relation ptr ≝ | pprec_nil : ∀c,q. pprec (◊) (c::q) -| ppprc_cons: ∀p,q. pprec (dx::p) (sn::q) +| ppprc_rc : ∀p,q. pprec (dx::p) (rc::q) +| ppprc_sn : ∀p,q. pprec (rc::p) (sn::q) | pprec_comp: ∀c,p,q. pprec p q → pprec (c::p) (c::q) | pprec_skip: pprec (dx::◊) ◊ . @@ -36,6 +37,7 @@ notation "hvbox(a break ≺ b)" lemma pprec_fwd_in_whd: ∀p,q. p ≺ q → in_whd q → in_whd p. #p #q #H elim H -p -q // /2 width=1/ [ #p #q * #H destruct +| #p #q * #H destruct | #c #p #q #_ #IHpq * #H destruct /3 width=1/ ] qed-. @@ -85,7 +87,11 @@ lemma ple_step_sn: ∀p1,p,p2. p1 ≺ p → p ≤ p2 → p1 ≤ p2. /2 width=3/ qed-. -lemma ple_cons: ∀p,q. dx::p ≤ sn::q. +lemma ple_rc: ∀p,q. dx::p ≤ rc::q. +/2 width=1/ +qed. + +lemma ple_sn: ∀p,q. rc::p ≤ sn::q. /2 width=1/ qed. @@ -106,6 +112,15 @@ lemma ple_skip_ple: ∀p. p ≤ ◊ → dx::p ≤ ◊. #p #q #Hpq #_ #H @(ple_step_sn … H) -H /2 width=1/ qed. +theorem ple_trans: transitive … ple. +/2 width=3/ +qed-. + +lemma ple_cons: ∀p,q. dx::p ≤ sn::q. +#p #q +@(ple_trans … (rc::q)) /2 width=1/ +qed. + lemma ple_dichotomy: ∀p1,p2:ptr. p1 ≤ p2 ∨ p2 ≤ p1. #p1 elim p1 -p1 [ * /2 width=1/