]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/list/in.ma
Some notation and additional lemmata.
[helm.git] / helm / software / matita / library / list / in.ma
index 310c3f369513d8eac9e8866219d1ae08670c117a..3285fef2e0967f75a5bc601953eadba75bef148c 100644 (file)
@@ -23,6 +23,13 @@ inductive in_list (A:Type): A → (list A) → Prop ≝
 definition incl : \forall A.(list A) \to (list A) \to Prop \def
   \lambda A,l,m.\forall x.in_list A x l \to in_list A x m.
   
+notation "hvbox(a break ∉ b)" non associative with precedence 45
+for @{ 'notmem $a $b }. 
+  
+interpretation "list member" 'mem x l = (in_list ? x l).
+interpretation "list not member" 'notmem x l = (Not (in_list ? x l)).
+interpretation "list inclusion" 'subseteq l1 l2 = (incl ? l1 l2).
+  
 lemma not_in_list_nil : \forall A,x.\lnot in_list A x [].
 intros.unfold.intro.inversion H
   [intros;lapply (sym_eq ? ? ? H2);destruct Hletin
@@ -164,4 +171,5 @@ qed.
 
 lemma incl_cons : ∀T,A,B,x.incl T A B → incl T (x::A) (x::B).
 unfold incl; intros;elim (in_list_cons_case ? ? ? ? H1);autobatch.
-qed.
\ No newline at end of file
+qed.
+