]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/grammar/term_simple.ma
- more properties on strongly normalizing terms
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / grammar / term_simple.ma
index cd387ee0270eb31ad5780510bfe05b429d254813..557299afd7f4176ac394a29c3dc79e708395302b 100644 (file)
@@ -16,21 +16,21 @@ include "Basic_2/grammar/term.ma".
 
 (* SIMPLE (NEUTRAL) TERMS ***************************************************)
 
-inductive simple: term → Prop ≝
-   | simple_atom: ∀I. simple (𝕒{I})
-   | simple_flat: ∀I,V,T. simple (𝕗{I} V. T)
+inductive simple: predicate term ≝
+   | 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\95\8a[T] â\86\92 â\88\80J,W,U. T = ð\9d\95\93{J} W. U → False.
+fact simple_inv_bind_aux: â\88\80T. ð\9d\90\92[T] â\86\92 â\88\80J,W,U. T = â\93\91{J} W. U → False.
 #T * -T
 [ #I #J #W #U #H destruct
 | #I #V #T #J #W #U #H destruct
 ]
 qed.
 
-lemma simple_inv_bind: â\88\80I,V,T. ð\9d\95\8a\9d\95\93{I} V. T] → False.
+lemma simple_inv_bind: â\88\80I,V,T. ð\9d\90\92\93\91{I} V. T] → False.
 /2 width=6/ qed-.