X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2FFsub%2Futil.ma;h=4f5e1542232f753e7fbb209ecc3fbb0a7ec69760;hb=947ee89dec9e60561dfac3ce7e1842f35f178cb8;hp=5a81f7ec576ca0eaa9cc4e0681feeac1333194a9;hpb=5f6974da8825bf7a7b23a5a9c7b051656d03aa37;p=helm.git diff --git a/helm/software/matita/library/Fsub/util.ma b/helm/software/matita/library/Fsub/util.ma index 5a81f7ec5..4f5e15422 100644 --- a/helm/software/matita/library/Fsub/util.ma +++ b/helm/software/matita/library/Fsub/util.ma @@ -19,14 +19,33 @@ include "list/list.ma". (*** useful definitions and lemmas not really related to Fsub ***) -let rec max m n \def +definition max \def +\lambda m,n. match (leb m n) with [true \Rightarrow n |false \Rightarrow m]. + +lemma le_n_max_m_n: \forall m,n:nat. n \le max m n. +intros. +unfold max. +apply (leb_elim m n) + [simplify.intros.apply le_n + |simplify.intros.autobatch + ] +qed. + +lemma le_m_max_m_n: \forall m,n:nat. m \le max m n. +intros. +unfold max. +apply (leb_elim m n) + [simplify.intro.assumption + |simplify.intros.autobatch + ] +qed. inductive in_list (A:Type): A → (list A) → Prop ≝ | in_Base : ∀ x,l.(in_list A x (x::l)) -| in_Skip : ∀ x,y,l.(in_list A x l) → (x ≠ y) → (in_list A x (y::l)). +| in_Skip : ∀ x,y,l.(in_list A x l) → (in_list A x (y::l)). notation > "hvbox(x ∈ l)" non associative with precedence 30 for @{ 'inlist $x $l }. @@ -35,12 +54,6 @@ notation < "hvbox(x \nbsp ∈ \nbsp l)" interpretation "item in list" 'inlist x l = (cic:/matita/Fsub/util/in_list.ind#xpointer(1/1) _ x l). -lemma in_Skip2 : ∀x,y,l.(in_list nat x l) → (in_list nat x (y::l)). -intros;elim (decidable_eq_nat x y) - [rewrite > H1;apply in_Base - |apply (in_Skip ? ? ? ? H H1)] -qed. - 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). @@ -55,17 +68,67 @@ definition map : \forall A,B,f.((list A) \to (list B)) \def lemma in_list_nil : \forall A,x.\lnot (in_list A x []). intros.unfold.intro.inversion H [intros;lapply (sym_eq ? ? ? H2);destruct Hletin - |intros;destruct H5] + |intros;destruct H4] qed. -lemma in_cons_case : ∀A.∀x,h:A.∀t:list A.x ∈ h::t → x = h ∨ (x ≠ h ∧ (x ∈ t)). +lemma in_cons_case : ∀A.∀x,h:A.∀t:list A.x ∈ h::t → x = h ∨ (x ∈ t). intros;inversion H;intros - [destruct H2;left;symmetry;assumption - |destruct H5;right;split [rewrite > Hcut|rewrite > Hcut1] assumption ] + [destruct H2;left;symmetry;reflexivity + |destruct H4;right;applyS H1] +qed. + +lemma append_nil:\forall A:Type.\forall l:list A. l@[]=l. +intros. +elim l;intros;autobatch. +qed. + +lemma append_cons:\forall A.\forall a:A.\forall l,l1. +l@(a::l1)=(l@[a])@l1. +intros. +rewrite > associative_append. +reflexivity. +qed. + +lemma in_list_append1: \forall A.\forall x:A. +\forall l1,l2. x \in l1 \to x \in l1@l2. +intros. +elim H + [simplify.apply in_Base + |simplify.apply in_Skip.assumption + ] +qed. + +lemma in_list_append2: \forall A.\forall x:A. +\forall l1,l2. x \in l2 \to x \in l1@l2. +intros 3. +elim l1 + [simplify.assumption + |simplify.apply in_Skip.apply H.assumption + ] +qed. + +theorem append_to_or_in_list: \forall A:Type.\forall x:A. +\forall l,l1. x \in l@l1 \to (x \in l) \lor (x \in l1). +intros 3. +elim l + [right.apply H + |simplify in H1.inversion H1;intros; destruct; + [left.apply in_Base + | elim (H l2) + [left.apply in_Skip. assumption + |right.assumption + |assumption + ] + ] + ] qed. lemma max_case : \forall m,n.(max m n) = match (leb m n) with - [ false \Rightarrow n - | true \Rightarrow m ]. + [ true \Rightarrow n + | false \Rightarrow m ]. intros;elim m;simplify;reflexivity; +qed. + +lemma incl_A_A: ∀T,A.incl T A A. +intros.unfold incl.intros.assumption. qed. \ No newline at end of file