]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda/paths/alternative_standard_order.ma
e8eca4f11e88a217f1e7550bdca125344ea3be93
[helm.git] / matita / matita / contribs / lambda / paths / alternative_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 (* ALTERNATIVE STANDARD ORDER ***********************************************)
18
19 (* Note: this is p < q *)
20 definition slt: relation path ≝ TC … sprec.
21
22 interpretation "standard 'less than' on paths"
23    'lt p q = (slt p q).
24
25 lemma slt_step_rc: ∀p,q. p ≺ q → p < q.
26 /2 width=1/
27 qed.
28
29 lemma slt_nil: ∀o,p. ◊ < o::p.
30 /2 width=1/
31 qed.
32
33 lemma slt_skip: dx::◊ < ◊.
34 /2 width=1/
35 qed.
36
37 lemma slt_comp: ∀o,p,q. p < q → o::p < o::q.
38 #o #p #q #H elim H -q /3 width=1/ /3 width=3/
39 qed.
40
41 theorem slt_trans: transitive … slt.
42 /2 width=3/
43 qed-.
44
45 lemma slt_refl: ∀p. p < p.
46 #p elim p -p /2 width=1/
47 @(slt_trans … (dx::◊)) //
48 qed.