]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/grammar/term_simple.ma
- nDestructTac: Sys.break handled in two places
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / grammar / term_simple.ma
index f32ae0e7ec6cb3f896d8be5519ae135ae2213d29..4c02d521d248770e2546c13dd834b2d608894da2 100644 (file)
@@ -25,17 +25,17 @@ interpretation "simple (term)" 'Simple T = (simple T).
 
 (* Basic inversion lemmas ***************************************************)
 
-fact simple_inv_bind_aux: ∀T. 𝐒[T] → ∀J,W,U. T = ⓑ{J} W. U → ⊥.
+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] → ⊥.
+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-.