]> 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, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / term_simple.ma
index 4bef6764dcdbb5675777f4628b198881b26afae2..50e1a4d22fdd22b16d77e4a70b244d31939e4787 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground_2/xoa/ex_1_4.ma".
 include "static_2/notation/relations/simple_1.ma".
 include "static_2/syntax/term.ma".
 
 (* SIMPLE (NEUTRAL) TERMS ***************************************************)
 
 inductive simple: predicate term ≝
-   | simple_atom: ∀I. simple (⓪{I})
-   | simple_flat: ∀I,V,T. simple (ⓕ{I}V.T)
+   | simple_atom: ∀I. simple (⓪[I])
+   | simple_flat: ∀I,V,T. simple (ⓕ[I]V.T)
 .
 
 interpretation "simple (term)" 'Simple T = (simple T).
 
 (* Basic inversion lemmas ***************************************************)
 
-fact simple_inv_bind_aux: â\88\80T. ð\9d\90\92â¦\83Tâ¦\84 â\86\92 â\88\80p,J,W,U. T = â\93\91{p,J}W.U → ⊥.
+fact simple_inv_bind_aux: â\88\80T. ð\9d\90\92â\9dªTâ\9d« â\86\92 â\88\80p,J,W,U. T = â\93\91[p,J]W.U → ⊥.
 #T * -T
 [ #I #p #J #W #U #H destruct
 | #I #V #T #a #J #W #U #H destruct
 ]
 qed-.
 
-lemma simple_inv_bind: â\88\80p,I,V,T. ð\9d\90\92â¦\83â\93\91{p,I} V. Tâ¦\84 → ⊥.
+lemma simple_inv_bind: â\88\80p,I,V,T. ð\9d\90\92â\9dªâ\93\91[p,I] V. Tâ\9d« → ⊥.
 /2 width=7 by simple_inv_bind_aux/ qed-.
 
-lemma simple_inv_pair: â\88\80I,V,T. ð\9d\90\92â¦\83â\91¡{I}V.Tâ¦\84 → ∃J. I = Flat2 J.
+lemma simple_inv_pair: â\88\80I,V,T. ð\9d\90\92â\9dªâ\91¡[I]V.Tâ\9d« → ∃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-.