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 (**************************************************************************)
15 include "basic_2/static/lsubr.ma".
17 (* RESTRICTED REFINEMENT FOR LOCAL ENVIRONMENTS *****************************)
19 (* Auxiliary inversion lemmas ***********************************************)
21 fact lsubr_inv_pair1_aux: ∀L1,L2. L1 ⫃ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X →
23 | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
24 | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
27 [ #L #J #K1 #X #H destruct /2 width=1 by or3_intro0/
28 | #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3 by or3_intro1, ex2_intro/
29 | #L1 #L2 #V #W #HL12 #J #K1 #X #H destruct /3 width=6 by or3_intro2, ex4_3_intro/
33 lemma lsubr_inv_pair1: ∀I,K1,L2,X. K1.ⓑ{I}X ⫃ L2 →
35 | ∃∃K2. K1 ⫃ K2 & L2 = K2.ⓑ{I}X
36 | ∃∃K2,V,W. K1 ⫃ K2 & L2 = K2.ⓛW &
38 /2 width=3 by lsubr_inv_pair1_aux/ qed-.
40 (* Main properties **********************************************************)
42 theorem lsubr_trans: Transitive … lsubr.
43 #L1 #L #H elim H -L1 -L
45 lapply (lsubr_inv_atom1 … H) -H //
46 | #I #L1 #L #V #_ #IHL1 #X #H
47 elim (lsubr_inv_pair1 … H) -H // *
48 #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1 by lsubr_pair, lsubr_beta/
49 | #L1 #L #V1 #W #_ #IHL1 #X #H
50 elim (lsubr_inv_abst1 … H) -H // *
51 #L2 #HL2 #H destruct /3 width=1 by lsubr_beta/