]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr_lsubr.ma
426166ab6657b3eedf3de9747d07978d737d6391
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / substitution / lsubr_lsubr.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 "basic_2/substitution/lsubr.ma".
16
17 (* LOCAL ENVIRONMENT REFINEMENT FOR SUBSTITUTION ****************************)
18
19 fact lsubr_inv_abbr1_aux: ∀L1,L2. L1 ⊑ L2 → ∀K1,W. L1 = K1.ⓓW →
20                           ∨∨ L2 = ⋆
21                            | ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓓW
22                            | ∃∃K2,W2. K1 ⊑ K2 & L2 = K2.ⓛW2.
23 #L1 #L2 * -L1 -L2
24 [ #L #K1 #W #H destruct /2 width=1/
25 | #L1 #L2 #V #HL12 #K1 #W #H destruct /3 width=3/
26 | #I #L1 #L2 #V1 #V2 #HL12 #K1 #W #H destruct /3 width=4/
27 ]
28 qed-.
29
30 lemma lsubr_inv_abbr1: ∀K1,L2,W. K1.ⓓW ⊑ L2 →
31                        ∨∨ L2 = ⋆
32                         | ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓓW
33                         | ∃∃K2,W2. K1 ⊑ K2 & L2 = K2.ⓛW2.
34 /2 width=3 by lsubr_inv_abbr1_aux/ qed-.
35
36 fact lsubr_inv_abst1_aux: ∀L1,L2. L1 ⊑ L2 → ∀K1,W1. L1 = K1.ⓛW1 →
37                           L2 = ⋆ ∨
38                           ∃∃K2,W2. K1 ⊑ K2 & L2 = K2.ⓛW2.
39 #L1 #L2 * -L1 -L2
40 [ #L #K1 #W1 #H destruct /2 width=1/
41 | #L1 #L2 #V #_ #K1 #W1 #H destruct
42 | #I #L1 #L2 #V1 #V2 #HL12 #K1 #W1 #H destruct /3 width=4/
43 ]
44 qed-.
45
46 lemma lsubr_inv_abst1: ∀K1,L2,W1. K1.ⓛW1 ⊑ L2 →
47                        L2 = ⋆ ∨
48                        ∃∃K2,W2. K1 ⊑ K2 & L2 = K2.ⓛW2.
49 /2 width=4 by lsubr_inv_abst1_aux/ qed-.
50
51 (* Main properties **********************************************************)
52
53 theorem lsubr_trans: Transitive … lsubr.
54 #L1 #L #H elim H -L1 -L
55 [ #L1 #X #H
56   lapply (lsubr_inv_atom1 … H) -H //
57 | #L1 #L #V #_ #IHL1 #X #H
58   elim (lsubr_inv_abbr1 … H) -H // *
59   #L2 [2: #V2 ] #HL2 #H destruct /3 width=1/
60 | #I #L1 #L #V1 #V #_ #IHL1 #X #H
61   elim (lsubr_inv_abst1 … H) -H // *
62   #L2 #V2 #HL2 #H destruct /3 width=1/
63 ]
64 qed-.