]> matita.cs.unibo.it Git - helm.git/blob - matita/library/Fsub/util.ma
41089172e56d6c76ece6e834aa665db4ba8feee9
[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) → (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 definition incl : \forall A.(list A) \to (list A) \to Prop \def
58   \lambda A,l,m.\forall x.(in_list A x l) \to (in_list A x m).               
59               
60 (* FIXME: use the map in library/list (when there will be one) *)
61 definition map : \forall A,B,f.((list A) \to (list B)) \def
62   \lambda A,B,f.let rec map (l : (list A)) : (list B) \def
63     match l in list return \lambda l0:(list A).(list B) with
64       [nil \Rightarrow (nil B)
65       |(cons (a:A) (t:(list A))) \Rightarrow 
66         (cons B (f a) (map t))] in map.
67
68 lemma in_list_nil : \forall A,x.\lnot (in_list A x []).
69 intros.unfold.intro.inversion H
70   [intros;lapply (sym_eq ? ? ? H2);destruct Hletin
71   |intros;destruct H4]
72 qed.
73
74 lemma in_cons_case : ∀A.∀x,h:A.∀t:list A.x ∈ h::t → x = h ∨ (x ∈ t).
75 intros;inversion H;intros
76   [destruct H2;left;symmetry;assumption
77   |destruct H4;right;applyS H1]
78 qed.
79
80 lemma append_nil:\forall A:Type.\forall l:list A. l@[]=l.
81 intros.
82 elim l;intros;autobatch.
83 qed.
84
85 lemma append_cons:\forall A.\forall a:A.\forall l,l1. 
86 l@(a::l1)=(l@[a])@l1.
87 intros.
88 rewrite > associative_append.
89 reflexivity.
90 qed.
91
92 lemma in_list_append1: \forall A.\forall x:A.
93 \forall l1,l2. x \in l1 \to x \in l1@l2.
94 intros.
95 elim H
96   [simplify.apply in_Base
97   |simplify.apply in_Skip.assumption
98   ]
99 qed.
100
101 lemma in_list_append2: \forall A.\forall x:A.
102 \forall l1,l2. x \in l2 \to x \in l1@l2.
103 intros 3.
104 elim l1
105   [simplify.assumption
106   |simplify.apply in_Skip.apply H.assumption
107   ]
108 qed.
109
110 theorem append_to_or_in_list: \forall A:Type.\forall x:A.
111 \forall l,l1. x \in l@l1 \to (x \in l) \lor (x \in l1).
112 intros 3.
113 elim l
114   [right.apply H
115   |simplify in H1.inversion H1;intros
116     [destruct H3.left.rewrite < Hcut.
117      apply in_Base
118     |destruct H5.
119      elim (H l2)
120       [left.apply in_Skip.
121        rewrite < H4.assumption
122       |right.rewrite < H4.assumption
123       |rewrite > Hcut1.rewrite > H4.assumption
124       ]
125     ]
126   ]
127 qed.
128
129 lemma max_case : \forall m,n.(max m n) = match (leb m n) with
130       [ true \Rightarrow n
131       | false \Rightarrow m ].
132 intros;elim m;simplify;reflexivity;
133 qed.
134
135 lemma incl_A_A: ∀T,A.incl T A A.
136 intros.unfold incl.intros.assumption.
137 qed.