]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma
- intermediate commit to allow debugging of auto tactic in xprs.ma
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / grammar / term_simple.ma
index 9b6c0827066b39fee9b56cbb8356303b17a62a1b..83aff07c497541a142bbfaabfbb04fdbcd11cc86 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic_2/grammar/term.ma".
+include "basic_2/grammar/term.ma".
 
 (* SIMPLE (NEUTRAL) TERMS ***************************************************)
 
@@ -25,17 +25,22 @@ interpretation "simple (term)" 'Simple T = (simple T).
 
 (* Basic inversion lemmas ***************************************************)
 
-fact simple_inv_bind_aux: ∀T. 𝐒[T] → ∀J,W,U. T = ⓑ{J} W. U → False.
+fact simple_inv_bind_aux: ∀T. 𝐒⦃T⦄ → ∀J,W,U. T = ⓑ{J} W. U → ⊥.
 #T * -T
 [ #I #J #W #U #H destruct
 | #I #V #T #J #W #U #H destruct
 ]
 qed.
 
-lemma simple_inv_bind: ∀I,V,T. 𝐒[ⓑ{I} V. T] → False.
+lemma simple_inv_bind: ∀I,V,T. 𝐒⦃ⓑ{I} V. T⦄ → ⊥.
 /2 width=6/ qed-.
 
-lemma simple_inv_pair: ∀I,V,T.  𝐒[②{I}V.T] → ∃J. I = Flat2 J.
+lemma simple_inv_pair: ∀I,V,T.  𝐒⦃②{I}V.T⦄ → ∃J. I = Flat2 J.
 * /2 width=2/ #I #V #T #H
 elim (simple_inv_bind … H)
 qed-.
+
+(*
+lemma mt: ∀R1,R2:Prop. (R1 → R2) → (R2 → ⊥) → R1 → ⊥.
+/3 width=1/ qed-.
+*)