]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/apps_2/models/veq.ma
update in apps_2
[helm.git] / matita / matita / contribs / lambdadelta / apps_2 / models / veq.ma
index a64dc2b1937c1360196bd3fe48a3e630c006c870..87a0800e75746a5481bc7e83207c9f159a64bed0 100644 (file)
@@ -27,10 +27,62 @@ interpretation "evaluation equivalence (model)"
 lemma veq_refl (M): is_model M →
                     reflexive … (veq M).
 /2 width=1 by mq/ qed.
-(*
-lemma veq_sym: ∀M. symmetric … (veq M).
-// qed-.
 
-theorem veq_trans: ∀M. transitive … (veq M).
+lemma veq_repl (M): is_model M →
+                    replace_2 … (veq M) (veq M) (veq M).
+/2 width=5 by mr/ qed-.
+
+lemma ext_veq (M): is_model M →
+                   ∀lv1,lv2. lv1 ≐ lv2 → lv1 ≗{M} lv2.
+/2 width=1 by mq/ qed.
+
+lemma exteq_veq_trans (M): ∀lv1,lv. lv1 ≐ lv →
+                           ∀lv2. lv ≗{M} lv2 → lv1 ≗ lv2.
 // qed-.
-*)
\ No newline at end of file
+
+(* Properties with evaluation evaluation lift *******************************)
+
+theorem vlift_swap (M): ∀i1,i2. i1 ≤ i2 →
+                        ∀lv,d1,d2. ⫯[i1←d1] ⫯[i2←d2] lv ≐{?,dd M} ⫯[↑i2←d2] ⫯[i1←d1] lv.
+#M #i1 #i2 #Hi12 #lv #d1 #d2 #j
+elim (lt_or_eq_or_gt j i1) #Hji1 destruct
+[ >vlift_lt // >vlift_lt /2 width=3 by lt_to_le_to_lt/
+  >vlift_lt /3 width=3 by lt_S, lt_to_le_to_lt/ >vlift_lt //
+| >vlift_eq >vlift_lt /2 width=1 by monotonic_le_plus_l/ >vlift_eq //
+| >vlift_gt // elim (lt_or_eq_or_gt (↓j) i2) #Hji2 destruct
+  [ >vlift_lt // >vlift_lt /2 width=1 by lt_minus_to_plus/ >vlift_gt //
+  | >vlift_eq <(lt_succ_pred … Hji1) >vlift_eq //
+  | >vlift_gt // >vlift_gt /2 width=1 by lt_minus_to_plus_r/ >vlift_gt /2 width=3 by le_to_lt_to_lt/
+  ]
+]
+qed-.
+
+lemma vlift_comp (M): ∀i. compatible_3 … (vlift M i) (sq M) (veq M) (veq M).
+#m #i #d1 #d2 #Hd12 #lv1 #lv2 #HLv12 #j
+elim (lt_or_eq_or_gt j i) #Hij destruct
+[ >vlift_lt // >vlift_lt //
+| >vlift_eq >vlift_eq //
+| >vlift_gt // >vlift_gt //
+]
+qed-.
+
+(* Properies with term interpretation ***************************************) 
+
+lemma ti_comp_l (M): is_model M →
+                     ∀T,gv,lv1,lv2. lv1 ≗{M} lv2 →
+                     ⟦T⟧[gv, lv1] ≗ ⟦T⟧[gv, lv2].
+#M #HM #T elim T -T * [||| #p * | * ]
+[ /4 width=3 by seq_trans, seq_sym, ms/
+| /4 width=5 by seq_sym, ml, mr/
+| /4 width=3 by seq_trans, seq_sym, mg/
+| /5 width=5 by vlift_comp, seq_sym, md, mr/
+| /5 width=1 by vlift_comp, mi, mq/
+| /4 width=5 by seq_sym, ma, mc, mr/
+| /4 width=5 by seq_sym, me, mr/
+]
+qed.
+
+lemma ti_ext_l (M): is_model M →
+                    ∀T,gv,lv1,lv2. lv1 ≐ lv2 →
+                    ⟦T⟧[gv, lv1] ≗{M} ⟦T⟧[gv, lv2].
+/3 width=1 by ti_comp_l, ext_veq/ qed.