]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/apps_2/etc/models/model_nlift.etc
update in models and ground_2
[helm.git] / matita / matita / contribs / lambdadelta / apps_2 / etc / models / model_nlift.etc
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/lift_neg.ma".
16 include "apps_2/models/model_sound.ma".
17
18 (* MODEL ********************************************************************)
19
20 lemma sound_ti_eq_nlift: ∀M. sound M → ∀sv,gv,U,lv1,lv2.
21                          (∀i. (∀T. ⬆[i, 1]T ≡ U → ⊥) → lv1 i = lv2 i) →
22                          〚U〛⦋M⦌{sv, gv, lv1} = 〚U〛⦋M⦌{sv, gv, lv2}.
23 #M #HM #sv #gv #U elim U -U * [||| #a * | * ]
24 [ #k #lv1 #lv2 #H >(m1 … HM) >(m1 … HM) -HM //
25 | #i #lv1 #lv2 #H >(m2 … HM) >(m2 … HM) -HM
26   /3 width=7 by lift_inv_lref2_be/
27 | #p #lv1 #lv2 #H >(m3 … HM) >(m3 … HM) -HM //
28 | #W #U #IHW #IHU #lv1 #lv2 #H >(m4 … HM) >(m4 … HM)
29   @IHU -IHU #i @(nat_ind_plus … i) -i
30   [ /5 width=9 by nlift_bind_sn/
31   | #i #_ #H0 >raise_gt // >raise_gt //
32     /4 width=9 by nlift_bind_dx/
33   ]
34 | #W #U #IHW #IHU #lv1 #lv2 #H @(m5 … HM) /5 width=9 by nlift_bind_sn/
35   #d @IHU -IHU #i @(nat_ind_plus … i) -i //
36   #i #_ #H0 lapply (nlift_bind_dx … H0) -H0
37   #H0 lapply (H i ?) -H /2 width=5 by/ -H0 #H
38   >raise_gt // >raise_gt //
39 | #W #U #IHW #IHU #lv1 #lv2 #H >(m6 … HM) >(m6 … HM)
40   /6 width=8 by nlift_flat_sn, nlift_flat_dx, eq_f2/
41 | #W #U #_ #IHU #lv1 #lv2 #H >(m7 … HM) >(m7 … HM)
42   /5 width=8 by nlift_flat_dx/
43 ]
44 qed-.