1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/Fsub/util".
16 include "logic/equality.ma".
17 include "nat/compare.ma".
18 include "list/list.ma".
20 (*** useful definitions and lemmas not really related to Fsub ***)
26 |false \Rightarrow m].
28 lemma le_n_max_m_n: \forall m,n:nat. n \le max m n.
32 [simplify.intros.apply le_n
33 |simplify.intros.autobatch
37 lemma le_m_max_m_n: \forall m,n:nat. m \le max m n.
41 [simplify.intro.assumption
42 |simplify.intros.autobatch
46 inductive in_list (A:Type): A → (list A) → Prop ≝
47 | in_Base : ∀ x,l.(in_list A x (x::l))
48 | in_Skip : ∀ x,y,l.(in_list A x l) → (x ≠ y) → (in_list A x (y::l)).
50 notation > "hvbox(x ∈ l)"
51 non associative with precedence 30 for @{ 'inlist $x $l }.
52 notation < "hvbox(x \nbsp ∈ \nbsp l)"
53 non associative with precedence 30 for @{ 'inlist $x $l }.
54 interpretation "item in list" 'inlist x l =
55 (cic:/matita/Fsub/util/in_list.ind#xpointer(1/1) _ x l).
57 lemma in_Skip2 : ∀x,y,l.(in_list nat x l) → (in_list nat x (y::l)).
58 intros;elim (decidable_eq_nat x y)
59 [rewrite > H1;apply in_Base
60 |apply (in_Skip ? ? ? ? H H1)]
63 definition incl : \forall A.(list A) \to (list A) \to Prop \def
64 \lambda A,l,m.\forall x.(in_list A x l) \to (in_list A x m).
66 (* FIXME: use the map in library/list (when there will be one) *)
67 definition map : \forall A,B,f.((list A) \to (list B)) \def
68 \lambda A,B,f.let rec map (l : (list A)) : (list B) \def
69 match l in list return \lambda l0:(list A).(list B) with
70 [nil \Rightarrow (nil B)
71 |(cons (a:A) (t:(list A))) \Rightarrow
72 (cons B (f a) (map t))] in map.
74 lemma in_list_nil : \forall A,x.\lnot (in_list A x []).
75 intros.unfold.intro.inversion H
76 [intros;lapply (sym_eq ? ? ? H2);destruct Hletin
80 lemma in_cons_case : ∀A.∀x,h:A.∀t:list A.x ∈ h::t → x = h ∨ (x ≠ h ∧ (x ∈ t)).
81 intros;inversion H;intros
82 [destruct H2;left;symmetry;assumption
83 |destruct H5;right;split [rewrite > Hcut|rewrite > Hcut1] assumption ]
86 lemma max_case : \forall m,n.(max m n) = match (leb m n) with
88 | true \Rightarrow m ].
89 intros;elim m;simplify;reflexivity;