]> matita.cs.unibo.it Git - helm.git/blob - matita/library/Fsub/util.ma
Qualche semplificazione.
[helm.git] / matita / library / Fsub / util.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/Fsub/util".
16 include "logic/equality.ma".
17 include "nat/compare.ma".
18 include "list/list.ma".
19
20 (*** useful definitions and lemmas not really related to Fsub ***)
21
22 definition max \def 
23 \lambda m,n.
24   match (leb m n) with
25      [true \Rightarrow n
26      |false \Rightarrow m]. 
27      
28 lemma le_n_max_m_n: \forall m,n:nat. n \le max m n.
29 intros.
30 unfold max.
31 apply (leb_elim m n)
32   [simplify.intros.apply le_n
33   |simplify.intros.autobatch
34   ]
35 qed.
36   
37 lemma le_m_max_m_n: \forall m,n:nat. m \le max m n.
38 intros.
39 unfold max.
40 apply (leb_elim m n)
41   [simplify.intro.assumption
42   |simplify.intros.autobatch
43   ]
44 qed.  
45
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)).
49
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).
56
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)]
61 qed.
62
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).               
65               
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.
73
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
77   |intros;destruct H5]
78 qed.
79
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 ]
84 qed.
85
86 lemma max_case : \forall m,n.(max m n) = match (leb m n) with
87       [ false \Rightarrow n
88       | true \Rightarrow m ].
89 intros;elim m;simplify;reflexivity;
90 qed.