]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/syntax/teqx.ma
update in ground_2, static_2, basic_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / teqx.ma
index 6c1ddd17efa049a5d101af6e7c78f1ab09faafb3..fbe372f561474ced7b5689c859a5bd4290f8c16d 100644 (file)
@@ -12,6 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground_2/xoa/ex_3_2.ma".
 include "static_2/notation/relations/stareq_2.ma".
 include "static_2/syntax/term.ma".
 
@@ -173,7 +174,7 @@ qed-.
 (* Negated inversion lemmas *************************************************)
 
 lemma tneqx_inv_pair: ∀I1,I2,V1,V2,T1,T2.
-                      (②{I1}V1.T1 ≛ ②{I2}V2.T2 → ⊥) → 
+                      (②{I1}V1.T1 ≛ ②{I2}V2.T2 → ⊥) →
                       ∨∨ I1 = I2 → ⊥
                       |  (V1 ≛ V2 → ⊥)
                       |  (T1 ≛ T2 → ⊥).