]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda/pointer_order.ma
- list.ma: improved notation for constant lists (a "term 19" was missing)
[helm.git] / matita / matita / contribs / lambda / pointer_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 "pointer.ma".
16
17 (* POINTER ORDER ************************************************************)
18
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::◊) ◊
26 .
27
28 interpretation "'precedes' on pointers"
29    'prec p q = (pprec p q).
30
31 (* Note: this should go to core_notation *)
32 notation "hvbox(a break ≺ b)"
33    non associative with precedence 45
34    for @{ 'prec $a $b }.
35
36 lemma pprec_fwd_in_head: ∀p,q. p ≺ q → in_head q → in_head p.
37 #p #q #H elim H -p -q // /2 width=1/
38 [ #p #q * #H destruct
39 | #c #p #q #_ #IHpq * #H destruct /3 width=1/
40 ]
41 qed-.
42
43 (* Note: this is p < q *)
44 definition plt: relation ptr ≝ TC … pprec.
45
46 interpretation "'less than' on redex pointers"
47    'lt p q = (plt p q).
48
49 (* Note: this is p ≤ q *)
50 definition ple: relation ptr ≝ star … pprec.
51
52 interpretation "'less or equal to' on redex pointers"
53    'leq p q = (ple p q).
54
55 lemma ple_step_rc: ∀p,q. p ≺ q → p ≤ q.
56 /2 width=1/
57 qed.
58
59 lemma ple_step_sn: ∀p1,p,p2. p1 ≺ p → p ≤ p2 → p1 ≤ p2.
60 /2 width=3/
61 qed-.
62
63 lemma ple_skip: dx::◊ ≤ ◊.
64 /2 width=1/
65 qed.
66
67 lemma ple_nil: ∀p. ◊ ≤ p.
68 * // /2 width=1/
69 qed.
70
71 lemma ple_comp: ∀p1,p2. p1 ≤ p2 → ∀c. (c::p1) ≤ (c::p2).
72 #p1 #p2 #H elim H -p2 // /3 width=3/
73 qed.
74
75 lemma ple_skip_ple: ∀p. p ≤ ◊ → dx::p ≤ ◊.
76 #p #H @(star_ind_l ??????? H) -p //
77 #p #q #Hpq #_ #H @(ple_step_sn … H) -H /2 width=1/
78 qed.
79
80 lemma in_head_ple_nil: ∀p. in_head p → p ≤ ◊.
81 #p #H @(in_head_ind … H) -p // /2 width=1/
82 qed.
83
84 theorem in_head_ple: ∀p. in_head p → ∀q. in_head q → p ≤ q.
85 #p #H @(in_head_ind … H) -p //
86 #p #Hp #IHp #q #H @(in_head_ind … H) -q /3 width=1/
87 qed.
88
89 lemma ple_nil_inv_in_head: ∀p. p ≤ ◊ → in_head p.
90 #p #H @(star_ind_l ??????? H) -p // /2 width=3 by pprec_fwd_in_head/
91 qed-.