]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/delayed_updating/syntax/path_proper.ma
update in delayed_updating
[helm.git] / matita / matita / contribs / lambdadelta / delayed_updating / syntax / path_proper.ma
index e519394f4e98d76220a7a3000ff6c7aeb96a28cd..66de8a4f6c37a0e4a3c869316c20f88690fbff93 100644 (file)
@@ -13,7 +13,8 @@
 (**************************************************************************)
 
 include "delayed_updating/syntax/path.ma".
-include "delayed_updating/notation/relations/predicate_p_tail_1.ma".
+include "delayed_updating/notation/functions/class_p_0.ma".
+include "ground/lib/subset.ma".
 include "ground/xoa/ex_1_2.ma".
 
 (* PROPER CONDITION FOR PATH ************************************************)
@@ -24,20 +25,25 @@ definition ppc: predicate path ≝
 
 interpretation
   "proper condition (path)"
-  'PredicatePTail p = (ppc p).
+  'ClassP = (ppc).
 
 (* Basic constructions ******************************************************)
 
-lemma ppc_lcons (l) (q): Ꝕ(l◗q).
+lemma ppc_lcons (l) (q): l◗q ϵ 𝐏.
 #l #p #H destruct
 qed.
 
 (* Basic inversions ********************************************************)
 
+lemma ppc_inv_empty:
+      (𝐞) ϵ 𝐏 → ⊥.
+#H0 @H0 -H0 //
+qed-.
+
 lemma ppc_inv_lcons (p):
-      Ꝕp → ∃∃l,q. l◗q = p.
+      p ϵ 𝐏 → ∃∃l,q. l◗q = p.
 *
-[ #H elim H -H //
+[ #H0 elim (ppc_inv_empty … H0)
 | #l #q #_ /2 width=3 by ex1_2_intro/
 ]
 qed-.