]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda/paths/standard_order.ma
- some additions and renaming ...
[helm.git] / matita / matita / contribs / lambda / paths / standard_order.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "paths/standard_precedence.ma".
16
17 (* STANDARD ORDER ************************************************************)
18
19 (* Note: this is p ≤ q *)
20 definition sle: relation path ≝ star … sprec.
21
22 interpretation "standard 'less or equal to' on paths"
23    'leq p q = (sle p q).
24
25 lemma sle_step_rc: ∀p,q. p ≺ q → p ≤ q.
26 /2 width=1/
27 qed.
28
29 lemma sle_step_sn: ∀p1,p,p2. p1 ≺ p → p ≤ p2 → p1 ≤ p2.
30 /2 width=3/
31 qed-.
32
33 lemma sle_rc: ∀p,q. dx::p ≤ rc::q.
34 /2 width=1/
35 qed.
36
37 lemma sle_sn: ∀p,q. rc::p ≤ sn::q.
38 /2 width=1/
39 qed.
40
41 lemma sle_skip: dx::◊ ≤ ◊.
42 /2 width=1/
43 qed.
44
45 lemma sle_nil: ∀p. ◊ ≤ p.
46 * // /2 width=1/
47 qed.
48
49 lemma sle_comp: ∀p1,p2. p1 ≤ p2 → ∀o. (o::p1) ≤ (o::p2).
50 #p1 #p2 #H elim H -p2 // /3 width=3/
51 qed.
52
53 lemma sle_skip_sle: ∀p. p ≤ ◊ → dx::p ≤ ◊.
54 #p #H @(star_ind_l ??????? H) -p //
55 #p #q #Hpq #_ #H @(sle_step_sn … H) -H /2 width=1/
56 qed.
57
58 theorem sle_trans: transitive … sle.
59 /2 width=3/
60 qed-.
61
62 lemma sle_cons: ∀p,q. dx::p ≤ sn::q.
63 #p #q
64 @(sle_trans … (rc::q)) /2 width=1/
65 qed.
66
67 lemma sle_dichotomy: ∀p1,p2:path. p1 ≤ p2 ∨ p2 ≤ p1.
68 #p1 elim p1 -p1
69 [ * /2 width=1/
70 | #o1 #p1 #IHp1 * /2 width=1/
71   * #p2 cases o1 -o1 /2 width=1/
72   elim (IHp1 p2) -IHp1 /3 width=1/
73 ]
74 qed-.
75
76 lemma sle_inv_sn: ∀p,q. p ≤ q → ∀p0. sn::p0 = p →
77                   ∃∃q0. p0 ≤ q0 & sn::q0 = q.
78 #p #q #H elim H -q /2 width=3/
79 #q #q0 #_ #Hq0 #IHpq #p0 #H destruct
80 elim (IHpq p0 ?) -IHpq // #p1 #Hp01 #H
81 elim (sprec_inv_sn … Hq0 … H) -q /3 width=3/
82 qed-.
83
84 lemma in_whd_sle_nil: ∀p. in_whd p → p ≤ ◊.
85 #p #H @(in_whd_ind … H) -p // /2 width=1/
86 qed.
87
88 theorem in_whd_sle: ∀p. in_whd p → ∀q. p ≤ q.
89 #p #H @(in_whd_ind … H) -p //
90 #p #_ #IHp * /3 width=1/ * #q /2 width=1/
91 qed.
92
93 lemma sle_nil_inv_in_whd: ∀p. p ≤ ◊ → in_whd p.
94 #p #H @(star_ind_l ??????? H) -p // /2 width=3 by sprec_fwd_in_whd/
95 qed-.
96
97 lemma sle_inv_in_whd: ∀p. (∀q. p ≤ q) → in_whd p.
98 /2 width=1 by sle_nil_inv_in_whd/
99 qed-.