]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/syntax/term_simple.ma
update in ground_2 static_2 basic_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / term_simple.ma
index 4bef6764dcdbb5675777f4628b198881b26afae2..a32b352688a1863982fca23f4848448ac68e04c2 100644 (file)
@@ -12,6 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground_2/xoa/ex_1_4.ma".
 include "static_2/notation/relations/simple_1.ma".
 include "static_2/syntax/term.ma".
 
@@ -40,3 +41,11 @@ lemma simple_inv_pair: ∀I,V,T. 𝐒⦃②{I}V.T⦄ → ∃J. I = Flat2 J.
 * /2 width=2 by ex_intro/
 #p #I #V #T #H elim (simple_inv_bind … H)
 qed-.
+
+(* Basic properties *********************************************************)
+
+lemma simple_dec_ex (X): ∨∨ 𝐒⦃X⦄ | ∃∃p,I,T,U. X = ⓑ{p,I}T.U.
+* [ /2 width=1 by simple_atom, or_introl/ ]
+* [| /2 width=1 by simple_flat, or_introl/ ]
+/3 width=5 by ex1_4_intro, or_intror/
+qed-.