1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
17 (* POINTER ORDER ************************************************************)
19 (* Note: precedence relation on redex pointers: p ≺ q
20 to serve as base for the order relations: p < q and p ≤ q *)
21 inductive pprec: relation ptr ≝
22 | pprec_nil : ∀c,q. pprec (◊) (c::q)
23 | ppprc_cons: ∀p,q. pprec (dx::p) (sn::q)
24 | pprec_comp: ∀c,p,q. pprec p q → pprec (c::p) (c::q)
25 | pprec_skip: pprec (dx::◊) ◊
28 interpretation "'precedes' on pointers"
29 'prec p q = (pprec p q).
31 (* Note: this should go to core_notation *)
32 notation "hvbox(a break ≺ b)"
33 non associative with precedence 45
36 lemma pprec_fwd_in_whd: ∀p,q. p ≺ q → in_whd q → in_whd p.
37 #p #q #H elim H -p -q // /2 width=1/
39 | #c #p #q #_ #IHpq * #H destruct /3 width=1/
43 (* Note: this is p < q *)
44 definition plt: relation ptr ≝ TC … pprec.
46 interpretation "'less than' on redex pointers"
49 lemma plt_step_rc: ∀p,q. p ≺ q → p < q.
53 lemma plt_nil: ∀c,p. ◊ < c::p.
57 lemma plt_skip: dx::◊ < ◊.
61 lemma plt_comp: ∀c,p,q. p < q → c::p < c::q.
62 #c #p #q #H elim H -q /3 width=1/ /3 width=3/
65 theorem plt_trans: transitive … plt.
69 lemma plt_refl: ∀p. p < p.
70 #p elim p -p /2 width=1/
71 @(plt_trans … (dx::◊)) //
74 (* Note: this is p ≤ q *)
75 definition ple: relation ptr ≝ star … pprec.
77 interpretation "'less or equal to' on redex pointers"
80 lemma ple_step_rc: ∀p,q. p ≺ q → p ≤ q.
84 lemma ple_step_sn: ∀p1,p,p2. p1 ≺ p → p ≤ p2 → p1 ≤ p2.
88 lemma ple_cons: ∀p,q. dx::p ≤ sn::q.
92 lemma ple_skip: dx::◊ ≤ ◊.
96 lemma ple_nil: ∀p. ◊ ≤ p.
100 lemma ple_comp: ∀p1,p2. p1 ≤ p2 → ∀c. (c::p1) ≤ (c::p2).
101 #p1 #p2 #H elim H -p2 // /3 width=3/
104 lemma ple_skip_ple: ∀p. p ≤ ◊ → dx::p ≤ ◊.
105 #p #H @(star_ind_l ??????? H) -p //
106 #p #q #Hpq #_ #H @(ple_step_sn … H) -H /2 width=1/
109 lemma ple_dichotomy: ∀p1,p2:ptr. p1 ≤ p2 ∨ p2 ≤ p1.
112 | #c1 #p1 #IHp1 * /2 width=1/
113 * #p2 cases c1 -c1 /2 width=1/
114 elim (IHp1 p2) -IHp1 /3 width=1/
118 lemma in_whd_ple_nil: ∀p. in_whd p → p ≤ ◊.
119 #p #H @(in_whd_ind … H) -p // /2 width=1/
122 theorem in_whd_ple: ∀p. in_whd p → ∀q. p ≤ q.
123 #p #H @(in_whd_ind … H) -p //
124 #p #_ #IHp * /3 width=1/ * #q /2 width=1/
127 lemma ple_nil_inv_in_whd: ∀p. p ≤ ◊ → in_whd p.
128 #p #H @(star_ind_l ??????? H) -p // /2 width=3 by pprec_fwd_in_whd/
131 lemma ple_inv_in_whd: ∀p. (∀q. p ≤ q) → in_whd p.
132 /2 width=1 by ple_nil_inv_in_whd/