]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda/paths/standard_order.ma
refactoring ...
[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/path.ma".
16
17 (* STANDARD ORDER ************************************************************)
18
19 (* Note: standard precedence relation on paths: p ≺ q
20          to serve as base for the order relations: p < q and p ≤ q *)
21 inductive sprec: relation path ≝
22 | sprec_nil : ∀c,q.   sprec (◊) (c::q)
23 | sprec_rc  : ∀p,q.   sprec (dx::p) (rc::q)
24 | sprec_sn  : ∀p,q.   sprec (rc::p) (sn::q)
25 | sprec_comp: ∀c,p,q. sprec p q → sprec (c::p) (c::q)
26 | sprec_skip:         sprec (dx::◊) ◊
27 .
28
29 interpretation "standard 'precedes' on paths"
30    'prec p q = (sprec p q).
31
32 lemma sprec_inv_sn: ∀p,q. p ≺ q → ∀p0. sn::p0 = p →
33                     ∃∃q0. p0 ≺ q0 & sn::q0 = q.
34 #p #q * -p -q
35 [ #c #q #p0 #H destruct
36 | #p #q #p0 #H destruct
37 | #p #q #p0 #H destruct
38 | #c #p #q #Hpq #p0 #H destruct /2 width=3/
39 | #p0 #H destruct
40 ]
41 qed-.
42
43 lemma sprec_fwd_in_whd: ∀p,q. p ≺ q → in_whd q → in_whd p.
44 #p #q #H elim H -p -q // /2 width=1/
45 [ #p #q * #H destruct
46 | #p #q * #H destruct
47 | #c #p #q #_ #IHpq * #H destruct /3 width=1/
48 ]
49 qed-.
50
51 (* Note: this is p < q *)
52 definition slt: relation path ≝ TC … sprec.
53
54 interpretation "standard 'less than' on paths"
55    'lt p q = (slt p q).
56
57 lemma slt_step_rc: ∀p,q. p ≺ q → p < q.
58 /2 width=1/
59 qed.
60
61 lemma slt_nil: ∀c,p. ◊ < c::p.
62 /2 width=1/
63 qed.
64
65 lemma slt_skip: dx::◊ < ◊.
66 /2 width=1/
67 qed.
68
69 lemma slt_comp: ∀c,p,q. p < q → c::p < c::q.
70 #c #p #q #H elim H -q /3 width=1/ /3 width=3/
71 qed.
72
73 theorem slt_trans: transitive … slt.
74 /2 width=3/
75 qed-.
76
77 lemma slt_refl: ∀p. p < p.
78 #p elim p -p /2 width=1/
79 @(slt_trans … (dx::◊)) //
80 qed.
81
82 (* Note: this is p ≤ q *)
83 definition sle: relation path ≝ star … sprec.
84
85 interpretation "standard 'less or equal to' on paths"
86    'leq p q = (sle p q).
87
88 lemma sle_step_rc: ∀p,q. p ≺ q → p ≤ q.
89 /2 width=1/
90 qed.
91
92 lemma sle_step_sn: ∀p1,p,p2. p1 ≺ p → p ≤ p2 → p1 ≤ p2.
93 /2 width=3/
94 qed-.
95
96 lemma sle_rc: ∀p,q. dx::p ≤ rc::q.
97 /2 width=1/
98 qed.
99
100 lemma sle_sn: ∀p,q. rc::p ≤ sn::q.
101 /2 width=1/
102 qed.
103
104 lemma sle_skip: dx::◊ ≤ ◊.
105 /2 width=1/
106 qed.
107
108 lemma sle_nil: ∀p. ◊ ≤ p.
109 * // /2 width=1/
110 qed.
111
112 lemma sle_comp: ∀p1,p2. p1 ≤ p2 → ∀c. (c::p1) ≤ (c::p2).
113 #p1 #p2 #H elim H -p2 // /3 width=3/
114 qed.
115
116 lemma sle_skip_sle: ∀p. p ≤ ◊ → dx::p ≤ ◊.
117 #p #H @(star_ind_l ??????? H) -p //
118 #p #q #Hpq #_ #H @(sle_step_sn … H) -H /2 width=1/
119 qed.
120
121 theorem sle_trans: transitive … sle.
122 /2 width=3/
123 qed-.
124
125 lemma sle_cons: ∀p,q. dx::p ≤ sn::q.
126 #p #q
127 @(sle_trans … (rc::q)) /2 width=1/
128 qed.
129
130 lemma sle_dichotomy: ∀p1,p2:path. p1 ≤ p2 ∨ p2 ≤ p1.
131 #p1 elim p1 -p1
132 [ * /2 width=1/
133 | #c1 #p1 #IHp1 * /2 width=1/
134   * #p2 cases c1 -c1 /2 width=1/
135   elim (IHp1 p2) -IHp1 /3 width=1/
136 ]
137 qed-.
138
139 lemma sle_inv_sn: ∀p,q. p ≤ q → ∀p0. sn::p0 = p →
140                   ∃∃q0. p0 ≤ q0 & sn::q0 = q.
141 #p #q #H elim H -q /2 width=3/
142 #q #q0 #_ #Hq0 #IHpq #p0 #H destruct
143 elim (IHpq p0 ?) -IHpq // #p1 #Hp01 #H
144 elim (sprec_inv_sn … Hq0 … H) -q /3 width=3/
145 qed-.
146
147 lemma in_whd_sle_nil: ∀p. in_whd p → p ≤ ◊.
148 #p #H @(in_whd_ind … H) -p // /2 width=1/
149 qed.
150
151 theorem in_whd_sle: ∀p. in_whd p → ∀q. p ≤ q.
152 #p #H @(in_whd_ind … H) -p //
153 #p #_ #IHp * /3 width=1/ * #q /2 width=1/
154 qed.
155
156 lemma sle_nil_inv_in_whd: ∀p. p ≤ ◊ → in_whd p.
157 #p #H @(star_ind_l ??????? H) -p // /2 width=3 by sprec_fwd_in_whd/
158 qed-.
159
160 lemma sle_inv_in_whd: ∀p. (∀q. p ≤ q) → in_whd p.
161 /2 width=1 by sle_nil_inv_in_whd/
162 qed-.