]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/apps_2/models/veq.ma
update in models and ground_2
[helm.git] / matita / matita / contribs / lambdadelta / apps_2 / models / veq.ma
index a64dc2b1937c1360196bd3fe48a3e630c006c870..11f07c404293f8df145f2828635c0885659af203 100644 (file)
@@ -27,10 +27,41 @@ interpretation "evaluation equivalence (model)"
 lemma veq_refl (M): is_model M →
                     reflexive … (veq M).
 /2 width=1 by mq/ qed.
+
+lemma veq_repl (M): is_model M →
+                    replace_2 … (veq M) (veq M) (veq M).
+/2 width=5 by mr/ qed-.
+
+(* Properties with evaluation push ******************************************)
+
+lemma push_comp (M): ∀i. compatible_3 … (push 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
+[ >(push_lt … Hij) >(push_lt … Hij) //
+| >(push_eq …) >(push_eq …) //
+| >(push_gt … Hij) >(push_gt … Hij) //
+]
+qed.
+
+(* Inversion lemmas with evaluation push *************************************)
+
+axiom veq_inv_push_sn: ∀M,lv1,y2,d1,i. ⫯[i←d1]lv1 ≗{M} y2 →
+                       ∃∃lv2,d2. lv1 ≗ lv2 & d1 ≗ d2 & ⫯[i←d2]lv2 = y2.   
 (*
-lemma veq_sym: ∀M. symmetric … (veq M).
-// qed-.
+#M #lv1 #y2 #d1 #i #H 
+*)
+(* Properies with term interpretation ***************************************) 
 
-theorem veq_trans: ∀M. transitive … (veq M).
-// qed-.
-*)
\ No newline at end of file
+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 push_comp, seq_sym, md, mr/
+| /5 width=1 by push_comp, mi, mq/
+| /4 width=5 by seq_sym, ma, mc, mr/
+| /4 width=5 by seq_sym, me, mr/
+]
+qed.