wenzelm@13462
|
1 |
(* Title: HOL/List.thy
|
wenzelm@13462
|
2 |
Author: Tobias Nipkow
|
clasohm@923
|
3 |
*)
|
clasohm@923
|
4 |
|
wenzelm@13114
|
5 |
header {* The datatype of finite lists *}
|
wenzelm@13122
|
6 |
|
nipkow@15131
|
7 |
theory List
|
kuncar@53012
|
8 |
imports Presburger Code_Numeral Quotient ATP Lifting_Set Lifting_Option Lifting_Product
|
nipkow@15131
|
9 |
begin
|
clasohm@923
|
10 |
|
wenzelm@13142
|
11 |
datatype 'a list =
|
wenzelm@13366
|
12 |
Nil ("[]")
|
wenzelm@13366
|
13 |
| Cons 'a "'a list" (infixr "#" 65)
|
clasohm@923
|
14 |
|
haftmann@34941
|
15 |
syntax
|
haftmann@34941
|
16 |
-- {* list Enumeration *}
|
wenzelm@35115
|
17 |
"_list" :: "args => 'a list" ("[(_)]")
|
haftmann@34941
|
18 |
|
haftmann@34941
|
19 |
translations
|
haftmann@34941
|
20 |
"[x, xs]" == "x#[xs]"
|
haftmann@34941
|
21 |
"[x]" == "x#[]"
|
haftmann@34941
|
22 |
|
wenzelm@35115
|
23 |
|
wenzelm@35115
|
24 |
subsection {* Basic list processing functions *}
|
nipkow@15302
|
25 |
|
nipkow@50548
|
26 |
primrec hd :: "'a list \<Rightarrow> 'a" where
|
nipkow@50548
|
27 |
"hd (x # xs) = x"
|
nipkow@50548
|
28 |
|
nipkow@50548
|
29 |
primrec tl :: "'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
30 |
"tl [] = []" |
|
nipkow@50548
|
31 |
"tl (x # xs) = xs"
|
nipkow@50548
|
32 |
|
nipkow@50548
|
33 |
primrec last :: "'a list \<Rightarrow> 'a" where
|
nipkow@50548
|
34 |
"last (x # xs) = (if xs = [] then x else last xs)"
|
nipkow@50548
|
35 |
|
nipkow@50548
|
36 |
primrec butlast :: "'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
37 |
"butlast []= []" |
|
nipkow@50548
|
38 |
"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
|
nipkow@50548
|
39 |
|
nipkow@50548
|
40 |
primrec set :: "'a list \<Rightarrow> 'a set" where
|
nipkow@50548
|
41 |
"set [] = {}" |
|
nipkow@50548
|
42 |
"set (x # xs) = insert x (set xs)"
|
nipkow@50548
|
43 |
|
nipkow@50548
|
44 |
definition coset :: "'a list \<Rightarrow> 'a set" where
|
nipkow@50548
|
45 |
[simp]: "coset xs = - set xs"
|
nipkow@50548
|
46 |
|
nipkow@50548
|
47 |
primrec map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
|
nipkow@50548
|
48 |
"map f [] = []" |
|
nipkow@50548
|
49 |
"map f (x # xs) = f x # map f xs"
|
nipkow@50548
|
50 |
|
nipkow@50548
|
51 |
primrec append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where
|
nipkow@50548
|
52 |
append_Nil: "[] @ ys = ys" |
|
nipkow@50548
|
53 |
append_Cons: "(x#xs) @ ys = x # xs @ ys"
|
nipkow@50548
|
54 |
|
nipkow@50548
|
55 |
primrec rev :: "'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
56 |
"rev [] = []" |
|
nipkow@50548
|
57 |
"rev (x # xs) = rev xs @ [x]"
|
nipkow@50548
|
58 |
|
nipkow@50548
|
59 |
primrec filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
60 |
"filter P [] = []" |
|
nipkow@50548
|
61 |
"filter P (x # xs) = (if P x then x # filter P xs else filter P xs)"
|
haftmann@34941
|
62 |
|
haftmann@34941
|
63 |
syntax
|
haftmann@34941
|
64 |
-- {* Special syntax for filter *}
|
wenzelm@35115
|
65 |
"_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])")
|
haftmann@34941
|
66 |
|
haftmann@34941
|
67 |
translations
|
haftmann@34941
|
68 |
"[x<-xs . P]"== "CONST filter (%x. P) xs"
|
haftmann@34941
|
69 |
|
haftmann@34941
|
70 |
syntax (xsymbols)
|
wenzelm@35115
|
71 |
"_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
|
haftmann@34941
|
72 |
syntax (HTML output)
|
wenzelm@35115
|
73 |
"_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
|
haftmann@34941
|
74 |
|
nipkow@50548
|
75 |
primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
|
nipkow@50548
|
76 |
fold_Nil: "fold f [] = id" |
|
nipkow@50548
|
77 |
fold_Cons: "fold f (x # xs) = fold f xs \<circ> f x"
|
nipkow@50548
|
78 |
|
nipkow@50548
|
79 |
primrec foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
|
nipkow@50548
|
80 |
foldr_Nil: "foldr f [] = id" |
|
nipkow@50548
|
81 |
foldr_Cons: "foldr f (x # xs) = f x \<circ> foldr f xs"
|
nipkow@50548
|
82 |
|
nipkow@50548
|
83 |
primrec foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where
|
nipkow@50548
|
84 |
foldl_Nil: "foldl f a [] = a" |
|
nipkow@50548
|
85 |
foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs"
|
nipkow@50548
|
86 |
|
nipkow@50548
|
87 |
primrec concat:: "'a list list \<Rightarrow> 'a list" where
|
nipkow@50548
|
88 |
"concat [] = []" |
|
nipkow@50548
|
89 |
"concat (x # xs) = x @ concat xs"
|
nipkow@50548
|
90 |
|
nipkow@50548
|
91 |
definition (in monoid_add) listsum :: "'a list \<Rightarrow> 'a" where
|
nipkow@50548
|
92 |
"listsum xs = foldr plus xs 0"
|
nipkow@50548
|
93 |
|
nipkow@50548
|
94 |
primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
95 |
drop_Nil: "drop n [] = []" |
|
nipkow@50548
|
96 |
drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)"
|
haftmann@34941
|
97 |
-- {*Warning: simpset does not contain this definition, but separate
|
haftmann@34941
|
98 |
theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
|
haftmann@34941
|
99 |
|
nipkow@50548
|
100 |
primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
101 |
take_Nil:"take n [] = []" |
|
nipkow@50548
|
102 |
take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)"
|
haftmann@34941
|
103 |
-- {*Warning: simpset does not contain this definition, but separate
|
haftmann@34941
|
104 |
theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
|
haftmann@34941
|
105 |
|
nipkow@50548
|
106 |
primrec nth :: "'a list => nat => 'a" (infixl "!" 100) where
|
nipkow@50548
|
107 |
nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)"
|
haftmann@34941
|
108 |
-- {*Warning: simpset does not contain this definition, but separate
|
haftmann@34941
|
109 |
theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
|
haftmann@34941
|
110 |
|
nipkow@50548
|
111 |
primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
|
nipkow@50548
|
112 |
"list_update [] i v = []" |
|
nipkow@50548
|
113 |
"list_update (x # xs) i v =
|
nipkow@50548
|
114 |
(case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)"
|
clasohm@923
|
115 |
|
wenzelm@41229
|
116 |
nonterminal lupdbinds and lupdbind
|
nipkow@5077
|
117 |
|
clasohm@923
|
118 |
syntax
|
wenzelm@13366
|
119 |
"_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)")
|
wenzelm@13366
|
120 |
"" :: "lupdbind => lupdbinds" ("_")
|
wenzelm@13366
|
121 |
"_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _")
|
wenzelm@13366
|
122 |
"_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [900,0] 900)
|
nipkow@5077
|
123 |
|
clasohm@923
|
124 |
translations
|
wenzelm@35115
|
125 |
"_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"
|
haftmann@34941
|
126 |
"xs[i:=x]" == "CONST list_update xs i x"
|
haftmann@34941
|
127 |
|
nipkow@50548
|
128 |
primrec takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
129 |
"takeWhile P [] = []" |
|
nipkow@50548
|
130 |
"takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
|
nipkow@50548
|
131 |
|
nipkow@50548
|
132 |
primrec dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
133 |
"dropWhile P [] = []" |
|
nipkow@50548
|
134 |
"dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)"
|
nipkow@50548
|
135 |
|
nipkow@50548
|
136 |
primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
|
nipkow@50548
|
137 |
"zip xs [] = []" |
|
nipkow@50548
|
138 |
zip_Cons: "zip xs (y # ys) =
|
nipkow@50548
|
139 |
(case xs of [] => [] | z # zs => (z, y) # zip zs ys)"
|
haftmann@34941
|
140 |
-- {*Warning: simpset does not contain this definition, but separate
|
haftmann@34941
|
141 |
theorems for @{text "xs = []"} and @{text "xs = z # zs"} *}
|
haftmann@34941
|
142 |
|
nipkow@50548
|
143 |
primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
|
nipkow@50548
|
144 |
"product [] _ = []" |
|
nipkow@50548
|
145 |
"product (x#xs) ys = map (Pair x) ys @ product xs ys"
|
haftmann@49948
|
146 |
|
haftmann@49948
|
147 |
hide_const (open) product
|
haftmann@49948
|
148 |
|
traytel@53721
|
149 |
primrec product_lists :: "'a list list \<Rightarrow> 'a list list" where
|
traytel@53721
|
150 |
"product_lists [] = [[]]" |
|
traytel@53721
|
151 |
"product_lists (xs # xss) = concat (map (\<lambda>x. map (Cons x) (product_lists xss)) xs)"
|
traytel@53721
|
152 |
|
nipkow@50548
|
153 |
primrec upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where
|
nipkow@50548
|
154 |
upt_0: "[i..<0] = []" |
|
nipkow@50548
|
155 |
upt_Suc: "[i..<(Suc j)] = (if i <= j then [i..<j] @ [j] else [])"
|
nipkow@50548
|
156 |
|
nipkow@50548
|
157 |
definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
158 |
"insert x xs = (if x \<in> set xs then xs else x # xs)"
|
haftmann@34978
|
159 |
|
wenzelm@36176
|
160 |
hide_const (open) insert
|
wenzelm@36176
|
161 |
hide_fact (open) insert_def
|
haftmann@34978
|
162 |
|
nipkow@47122
|
163 |
primrec find :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a option" where
|
nipkow@50548
|
164 |
"find _ [] = None" |
|
nipkow@50548
|
165 |
"find P (x#xs) = (if P x then Some x else find P xs)"
|
nipkow@47122
|
166 |
|
nipkow@47122
|
167 |
hide_const (open) find
|
nipkow@47122
|
168 |
|
haftmann@51096
|
169 |
primrec those :: "'a option list \<Rightarrow> 'a list option"
|
haftmann@51096
|
170 |
where
|
haftmann@51096
|
171 |
"those [] = Some []" |
|
haftmann@51096
|
172 |
"those (x # xs) = (case x of
|
haftmann@51096
|
173 |
None \<Rightarrow> None
|
haftmann@51096
|
174 |
| Some y \<Rightarrow> Option.map (Cons y) (those xs))"
|
haftmann@51096
|
175 |
|
nipkow@50548
|
176 |
primrec remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
177 |
"remove1 x [] = []" |
|
nipkow@50548
|
178 |
"remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)"
|
nipkow@50548
|
179 |
|
nipkow@50548
|
180 |
primrec removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
181 |
"removeAll x [] = []" |
|
nipkow@50548
|
182 |
"removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)"
|
nipkow@50548
|
183 |
|
nipkow@50548
|
184 |
primrec distinct :: "'a list \<Rightarrow> bool" where
|
nipkow@50548
|
185 |
"distinct [] \<longleftrightarrow> True" |
|
nipkow@50548
|
186 |
"distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
|
nipkow@50548
|
187 |
|
nipkow@50548
|
188 |
primrec remdups :: "'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
189 |
"remdups [] = []" |
|
nipkow@50548
|
190 |
"remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)"
|
nipkow@50548
|
191 |
|
traytel@53721
|
192 |
fun remdups_adj :: "'a list \<Rightarrow> 'a list" where
|
traytel@53721
|
193 |
"remdups_adj [] = []" |
|
traytel@53721
|
194 |
"remdups_adj [x] = [x]" |
|
traytel@53721
|
195 |
"remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))"
|
traytel@53721
|
196 |
|
nipkow@50548
|
197 |
primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
|
nipkow@50548
|
198 |
replicate_0: "replicate 0 x = []" |
|
nipkow@50548
|
199 |
replicate_Suc: "replicate (Suc n) x = x # replicate n x"
|
paulson@3342
|
200 |
|
wenzelm@13142
|
201 |
text {*
|
wenzelm@14589
|
202 |
Function @{text size} is overloaded for all datatypes. Users may
|
wenzelm@13366
|
203 |
refer to the list version as @{text length}. *}
|
wenzelm@13142
|
204 |
|
nipkow@50548
|
205 |
abbreviation length :: "'a list \<Rightarrow> nat" where
|
nipkow@50548
|
206 |
"length \<equiv> size"
|
paulson@15307
|
207 |
|
haftmann@51173
|
208 |
definition enumerate :: "nat \<Rightarrow> 'a list \<Rightarrow> (nat \<times> 'a) list" where
|
haftmann@51173
|
209 |
enumerate_eq_zip: "enumerate n xs = zip [n..<n + length xs] xs"
|
haftmann@51173
|
210 |
|
blanchet@46440
|
211 |
primrec rotate1 :: "'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
212 |
"rotate1 [] = []" |
|
nipkow@50548
|
213 |
"rotate1 (x # xs) = xs @ [x]"
|
nipkow@50548
|
214 |
|
nipkow@50548
|
215 |
definition rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@50548
|
216 |
"rotate n = rotate1 ^^ n"
|
nipkow@50548
|
217 |
|
nipkow@50548
|
218 |
definition list_all2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool" where
|
nipkow@50548
|
219 |
"list_all2 P xs ys =
|
nipkow@50548
|
220 |
(length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y))"
|
nipkow@50548
|
221 |
|
nipkow@50548
|
222 |
definition sublist :: "'a list => nat set => 'a list" where
|
nipkow@50548
|
223 |
"sublist xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
|
nipkow@50548
|
224 |
|
nipkow@50548
|
225 |
primrec sublists :: "'a list \<Rightarrow> 'a list list" where
|
nipkow@50548
|
226 |
"sublists [] = [[]]" |
|
nipkow@50548
|
227 |
"sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)"
|
nipkow@50548
|
228 |
|
nipkow@50548
|
229 |
primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
|
nipkow@50548
|
230 |
"n_lists 0 xs = [[]]" |
|
nipkow@50548
|
231 |
"n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))"
|
haftmann@49948
|
232 |
|
haftmann@49948
|
233 |
hide_const (open) n_lists
|
haftmann@49948
|
234 |
|
nipkow@40593
|
235 |
fun splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
|
nipkow@40593
|
236 |
"splice [] ys = ys" |
|
nipkow@40593
|
237 |
"splice xs [] = xs" |
|
nipkow@40593
|
238 |
"splice (x#xs) (y#ys) = x # y # splice xs ys"
|
haftmann@21061
|
239 |
|
nipkow@26771
|
240 |
text{*
|
nipkow@26771
|
241 |
\begin{figure}[htbp]
|
nipkow@26771
|
242 |
\fbox{
|
nipkow@26771
|
243 |
\begin{tabular}{l}
|
wenzelm@27381
|
244 |
@{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\
|
wenzelm@27381
|
245 |
@{lemma "length [a,b,c] = 3" by simp}\\
|
wenzelm@27381
|
246 |
@{lemma "set [a,b,c] = {a,b,c}" by simp}\\
|
wenzelm@27381
|
247 |
@{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\
|
wenzelm@27381
|
248 |
@{lemma "rev [a,b,c] = [c,b,a]" by simp}\\
|
wenzelm@27381
|
249 |
@{lemma "hd [a,b,c,d] = a" by simp}\\
|
wenzelm@27381
|
250 |
@{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\
|
wenzelm@27381
|
251 |
@{lemma "last [a,b,c,d] = d" by simp}\\
|
wenzelm@27381
|
252 |
@{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\
|
wenzelm@27381
|
253 |
@{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\
|
wenzelm@27381
|
254 |
@{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\
|
haftmann@46133
|
255 |
@{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\
|
haftmann@47397
|
256 |
@{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\
|
haftmann@47397
|
257 |
@{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\
|
wenzelm@27381
|
258 |
@{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\
|
wenzelm@27381
|
259 |
@{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\
|
haftmann@51173
|
260 |
@{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\
|
haftmann@49948
|
261 |
@{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\
|
traytel@53721
|
262 |
@{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\
|
wenzelm@27381
|
263 |
@{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\
|
wenzelm@27381
|
264 |
@{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\
|
wenzelm@27381
|
265 |
@{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\
|
wenzelm@27381
|
266 |
@{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\
|
wenzelm@27381
|
267 |
@{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\
|
wenzelm@27381
|
268 |
@{lemma "drop 6 [a,b,c,d] = []" by simp}\\
|
wenzelm@27381
|
269 |
@{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\
|
wenzelm@27381
|
270 |
@{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\
|
wenzelm@27381
|
271 |
@{lemma "distinct [2,0,1::nat]" by simp}\\
|
wenzelm@27381
|
272 |
@{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
|
traytel@53721
|
273 |
@{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\
|
haftmann@34978
|
274 |
@{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
|
haftmann@35295
|
275 |
@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
|
nipkow@47122
|
276 |
@{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\
|
nipkow@47122
|
277 |
@{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\
|
wenzelm@27381
|
278 |
@{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
|
nipkow@27693
|
279 |
@{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
|
wenzelm@27381
|
280 |
@{lemma "nth [a,b,c,d] 2 = c" by simp}\\
|
wenzelm@27381
|
281 |
@{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\
|
wenzelm@27381
|
282 |
@{lemma "sublist [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:sublist_def)}\\
|
haftmann@49948
|
283 |
@{lemma "sublists [a,b] = [[a, b], [a], [b], []]" by simp}\\
|
haftmann@49948
|
284 |
@{lemma "List.n_lists 2 [a,b,c] = [[a, a], [b, a], [c, a], [a, b], [b, b], [c, b], [a, c], [b, c], [c, c]]" by (simp add: eval_nat_numeral)}\\
|
blanchet@46440
|
285 |
@{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\
|
blanchet@46440
|
286 |
@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\
|
nipkow@40077
|
287 |
@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\
|
nipkow@40077
|
288 |
@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\
|
haftmann@47397
|
289 |
@{lemma "listsum [1,2,3::nat] = 6" by (simp add: listsum_def)}
|
nipkow@26771
|
290 |
\end{tabular}}
|
nipkow@26771
|
291 |
\caption{Characteristic examples}
|
nipkow@26771
|
292 |
\label{fig:Characteristic}
|
nipkow@26771
|
293 |
\end{figure}
|
blanchet@29927
|
294 |
Figure~\ref{fig:Characteristic} shows characteristic examples
|
nipkow@26771
|
295 |
that should give an intuitive understanding of the above functions.
|
nipkow@26771
|
296 |
*}
|
nipkow@26771
|
297 |
|
nipkow@24616
|
298 |
text{* The following simple sort functions are intended for proofs,
|
nipkow@24616
|
299 |
not for efficient implementations. *}
|
nipkow@24616
|
300 |
|
wenzelm@25221
|
301 |
context linorder
|
wenzelm@25221
|
302 |
begin
|
wenzelm@25221
|
303 |
|
haftmann@39915
|
304 |
inductive sorted :: "'a list \<Rightarrow> bool" where
|
haftmann@39915
|
305 |
Nil [iff]: "sorted []"
|
haftmann@39915
|
306 |
| Cons: "\<forall>y\<in>set xs. x \<le> y \<Longrightarrow> sorted xs \<Longrightarrow> sorted (x # xs)"
|
haftmann@39915
|
307 |
|
haftmann@39915
|
308 |
lemma sorted_single [iff]:
|
haftmann@39915
|
309 |
"sorted [x]"
|
haftmann@39915
|
310 |
by (rule sorted.Cons) auto
|
haftmann@39915
|
311 |
|
haftmann@39915
|
312 |
lemma sorted_many:
|
haftmann@39915
|
313 |
"x \<le> y \<Longrightarrow> sorted (y # zs) \<Longrightarrow> sorted (x # y # zs)"
|
haftmann@39915
|
314 |
by (rule sorted.Cons) (cases "y # zs" rule: sorted.cases, auto)
|
haftmann@39915
|
315 |
|
haftmann@39915
|
316 |
lemma sorted_many_eq [simp, code]:
|
haftmann@39915
|
317 |
"sorted (x # y # zs) \<longleftrightarrow> x \<le> y \<and> sorted (y # zs)"
|
haftmann@39915
|
318 |
by (auto intro: sorted_many elim: sorted.cases)
|
haftmann@39915
|
319 |
|
haftmann@39915
|
320 |
lemma [code]:
|
haftmann@39915
|
321 |
"sorted [] \<longleftrightarrow> True"
|
haftmann@39915
|
322 |
"sorted [x] \<longleftrightarrow> True"
|
haftmann@39915
|
323 |
by simp_all
|
nipkow@24697
|
324 |
|
hoelzl@33639
|
325 |
primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
|
nipkow@50548
|
326 |
"insort_key f x [] = [x]" |
|
nipkow@50548
|
327 |
"insort_key f x (y#ys) =
|
nipkow@50548
|
328 |
(if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))"
|
hoelzl@33639
|
329 |
|
haftmann@35195
|
330 |
definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where
|
nipkow@50548
|
331 |
"sort_key f xs = foldr (insort_key f) xs []"
|
hoelzl@33639
|
332 |
|
haftmann@40210
|
333 |
definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
|
nipkow@50548
|
334 |
"insort_insert_key f x xs =
|
nipkow@50548
|
335 |
(if f x \<in> f ` set xs then xs else insort_key f x xs)"
|
haftmann@40210
|
336 |
|
hoelzl@33639
|
337 |
abbreviation "sort \<equiv> sort_key (\<lambda>x. x)"
|
hoelzl@33639
|
338 |
abbreviation "insort \<equiv> insort_key (\<lambda>x. x)"
|
haftmann@40210
|
339 |
abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)"
|
haftmann@35608
|
340 |
|
wenzelm@25221
|
341 |
end
|
wenzelm@25221
|
342 |
|
nipkow@24616
|
343 |
|
wenzelm@23388
|
344 |
subsubsection {* List comprehension *}
|
nipkow@23192
|
345 |
|
nipkow@24349
|
346 |
text{* Input syntax for Haskell-like list comprehension notation.
|
nipkow@24349
|
347 |
Typical example: @{text"[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]"},
|
nipkow@24349
|
348 |
the list of all pairs of distinct elements from @{text xs} and @{text ys}.
|
nipkow@24349
|
349 |
The syntax is as in Haskell, except that @{text"|"} becomes a dot
|
nipkow@24349
|
350 |
(like in Isabelle's set comprehension): @{text"[e. x \<leftarrow> xs, \<dots>]"} rather than
|
nipkow@24349
|
351 |
\verb![e| x <- xs, ...]!.
|
nipkow@24349
|
352 |
|
nipkow@24349
|
353 |
The qualifiers after the dot are
|
nipkow@24349
|
354 |
\begin{description}
|
nipkow@24349
|
355 |
\item[generators] @{text"p \<leftarrow> xs"},
|
nipkow@24476
|
356 |
where @{text p} is a pattern and @{text xs} an expression of list type, or
|
nipkow@24476
|
357 |
\item[guards] @{text"b"}, where @{text b} is a boolean expression.
|
nipkow@24476
|
358 |
%\item[local bindings] @ {text"let x = e"}.
|
nipkow@24349
|
359 |
\end{description}
|
nipkow@23240
|
360 |
|
nipkow@24476
|
361 |
Just like in Haskell, list comprehension is just a shorthand. To avoid
|
nipkow@24476
|
362 |
misunderstandings, the translation into desugared form is not reversed
|
nipkow@24476
|
363 |
upon output. Note that the translation of @{text"[e. x \<leftarrow> xs]"} is
|
nipkow@24476
|
364 |
optmized to @{term"map (%x. e) xs"}.
|
nipkow@23240
|
365 |
|
nipkow@24349
|
366 |
It is easy to write short list comprehensions which stand for complex
|
nipkow@24349
|
367 |
expressions. During proofs, they may become unreadable (and
|
nipkow@24349
|
368 |
mangled). In such cases it can be advisable to introduce separate
|
nipkow@24349
|
369 |
definitions for the list comprehensions in question. *}
|
nipkow@24349
|
370 |
|
wenzelm@46138
|
371 |
nonterminal lc_qual and lc_quals
|
nipkow@23192
|
372 |
|
nipkow@23192
|
373 |
syntax
|
wenzelm@46138
|
374 |
"_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list" ("[_ . __")
|
wenzelm@46138
|
375 |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _")
|
wenzelm@46138
|
376 |
"_lc_test" :: "bool \<Rightarrow> lc_qual" ("_")
|
wenzelm@46138
|
377 |
(*"_lc_let" :: "letbinds => lc_qual" ("let _")*)
|
wenzelm@46138
|
378 |
"_lc_end" :: "lc_quals" ("]")
|
wenzelm@46138
|
379 |
"_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __")
|
wenzelm@46138
|
380 |
"_lc_abs" :: "'a => 'b list => 'b list"
|
nipkow@23192
|
381 |
|
nipkow@24476
|
382 |
(* These are easier than ML code but cannot express the optimized
|
nipkow@24476
|
383 |
translation of [e. p<-xs]
|
nipkow@23192
|
384 |
translations
|
wenzelm@46138
|
385 |
"[e. p<-xs]" => "concat(map (_lc_abs p [e]) xs)"
|
wenzelm@46138
|
386 |
"_listcompr e (_lc_gen p xs) (_lc_quals Q Qs)"
|
wenzelm@46138
|
387 |
=> "concat (map (_lc_abs p (_listcompr e Q Qs)) xs)"
|
wenzelm@46138
|
388 |
"[e. P]" => "if P then [e] else []"
|
wenzelm@46138
|
389 |
"_listcompr e (_lc_test P) (_lc_quals Q Qs)"
|
wenzelm@46138
|
390 |
=> "if P then (_listcompr e Q Qs) else []"
|
wenzelm@46138
|
391 |
"_listcompr e (_lc_let b) (_lc_quals Q Qs)"
|
wenzelm@46138
|
392 |
=> "_Let b (_listcompr e Q Qs)"
|
nipkow@24476
|
393 |
*)
|
nipkow@23240
|
394 |
|
nipkow@23279
|
395 |
syntax (xsymbols)
|
wenzelm@46138
|
396 |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
|
nipkow@23279
|
397 |
syntax (HTML output)
|
wenzelm@46138
|
398 |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
|
nipkow@24349
|
399 |
|
wenzelm@52143
|
400 |
parse_translation {*
|
wenzelm@46138
|
401 |
let
|
wenzelm@46138
|
402 |
val NilC = Syntax.const @{const_syntax Nil};
|
wenzelm@46138
|
403 |
val ConsC = Syntax.const @{const_syntax Cons};
|
wenzelm@46138
|
404 |
val mapC = Syntax.const @{const_syntax map};
|
wenzelm@46138
|
405 |
val concatC = Syntax.const @{const_syntax concat};
|
wenzelm@46138
|
406 |
val IfC = Syntax.const @{const_syntax If};
|
wenzelm@46138
|
407 |
|
wenzelm@46138
|
408 |
fun single x = ConsC $ x $ NilC;
|
wenzelm@46138
|
409 |
|
wenzelm@46138
|
410 |
fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
|
wenzelm@46138
|
411 |
let
|
wenzelm@46138
|
412 |
(* FIXME proper name context!? *)
|
wenzelm@46138
|
413 |
val x =
|
wenzelm@46138
|
414 |
Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT);
|
wenzelm@46138
|
415 |
val e = if opti then single e else e;
|
wenzelm@46138
|
416 |
val case1 = Syntax.const @{syntax_const "_case1"} $ p $ e;
|
wenzelm@46138
|
417 |
val case2 =
|
wenzelm@46138
|
418 |
Syntax.const @{syntax_const "_case1"} $
|
wenzelm@46138
|
419 |
Syntax.const @{const_syntax dummy_pattern} $ NilC;
|
wenzelm@46138
|
420 |
val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2;
|
traytel@51678
|
421 |
in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end;
|
wenzelm@46138
|
422 |
|
wenzelm@46138
|
423 |
fun abs_tr ctxt p e opti =
|
wenzelm@46138
|
424 |
(case Term_Position.strip_positions p of
|
wenzelm@46138
|
425 |
Free (s, T) =>
|
wenzelm@46138
|
426 |
let
|
wenzelm@46138
|
427 |
val thy = Proof_Context.theory_of ctxt;
|
wenzelm@46138
|
428 |
val s' = Proof_Context.intern_const ctxt s;
|
wenzelm@46138
|
429 |
in
|
wenzelm@46138
|
430 |
if Sign.declared_const thy s'
|
wenzelm@46138
|
431 |
then (pat_tr ctxt p e opti, false)
|
wenzelm@46138
|
432 |
else (Syntax_Trans.abs_tr [p, e], true)
|
wenzelm@46138
|
433 |
end
|
wenzelm@46138
|
434 |
| _ => (pat_tr ctxt p e opti, false));
|
wenzelm@46138
|
435 |
|
wenzelm@46138
|
436 |
fun lc_tr ctxt [e, Const (@{syntax_const "_lc_test"}, _) $ b, qs] =
|
wenzelm@46138
|
437 |
let
|
wenzelm@46138
|
438 |
val res =
|
wenzelm@46138
|
439 |
(case qs of
|
wenzelm@46138
|
440 |
Const (@{syntax_const "_lc_end"}, _) => single e
|
wenzelm@46138
|
441 |
| Const (@{syntax_const "_lc_quals"}, _) $ q $ qs => lc_tr ctxt [e, q, qs]);
|
wenzelm@46138
|
442 |
in IfC $ b $ res $ NilC end
|
wenzelm@46138
|
443 |
| lc_tr ctxt
|
wenzelm@46138
|
444 |
[e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
|
wenzelm@46138
|
445 |
Const(@{syntax_const "_lc_end"}, _)] =
|
wenzelm@46138
|
446 |
(case abs_tr ctxt p e true of
|
wenzelm@46138
|
447 |
(f, true) => mapC $ f $ es
|
wenzelm@46138
|
448 |
| (f, false) => concatC $ (mapC $ f $ es))
|
wenzelm@46138
|
449 |
| lc_tr ctxt
|
wenzelm@46138
|
450 |
[e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
|
wenzelm@46138
|
451 |
Const (@{syntax_const "_lc_quals"}, _) $ q $ qs] =
|
wenzelm@46138
|
452 |
let val e' = lc_tr ctxt [e, q, qs];
|
wenzelm@46138
|
453 |
in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end;
|
wenzelm@46138
|
454 |
|
wenzelm@46138
|
455 |
in [(@{syntax_const "_listcompr"}, lc_tr)] end
|
nipkow@24349
|
456 |
*}
|
nipkow@23279
|
457 |
|
wenzelm@51272
|
458 |
ML_val {*
|
wenzelm@42167
|
459 |
let
|
wenzelm@42167
|
460 |
val read = Syntax.read_term @{context};
|
wenzelm@42167
|
461 |
fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote s1);
|
wenzelm@42167
|
462 |
in
|
wenzelm@42167
|
463 |
check "[(x,y,z). b]" "if b then [(x, y, z)] else []";
|
wenzelm@42167
|
464 |
check "[(x,y,z). x\<leftarrow>xs]" "map (\<lambda>x. (x, y, z)) xs";
|
wenzelm@42167
|
465 |
check "[e x y. x\<leftarrow>xs, y\<leftarrow>ys]" "concat (map (\<lambda>x. map (\<lambda>y. e x y) ys) xs)";
|
wenzelm@42167
|
466 |
check "[(x,y,z). x<a, x>b]" "if x < a then if b < x then [(x, y, z)] else [] else []";
|
wenzelm@42167
|
467 |
check "[(x,y,z). x\<leftarrow>xs, x>b]" "concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)";
|
wenzelm@42167
|
468 |
check "[(x,y,z). x<a, x\<leftarrow>xs]" "if x < a then map (\<lambda>x. (x, y, z)) xs else []";
|
wenzelm@42167
|
469 |
check "[(x,y). Cons True x \<leftarrow> xs]"
|
wenzelm@42167
|
470 |
"concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)";
|
wenzelm@42167
|
471 |
check "[(x,y,z). Cons x [] \<leftarrow> xs]"
|
wenzelm@42167
|
472 |
"concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)";
|
wenzelm@42167
|
473 |
check "[(x,y,z). x<a, x>b, x=d]"
|
wenzelm@42167
|
474 |
"if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []";
|
wenzelm@42167
|
475 |
check "[(x,y,z). x<a, x>b, y\<leftarrow>ys]"
|
wenzelm@42167
|
476 |
"if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []";
|
wenzelm@42167
|
477 |
check "[(x,y,z). x<a, x\<leftarrow>xs,y>b]"
|
wenzelm@42167
|
478 |
"if x < a then concat (map (\<lambda>x. if b < y then [(x, y, z)] else []) xs) else []";
|
wenzelm@42167
|
479 |
check "[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]"
|
wenzelm@42167
|
480 |
"if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []";
|
wenzelm@42167
|
481 |
check "[(x,y,z). x\<leftarrow>xs, x>b, y<a]"
|
wenzelm@42167
|
482 |
"concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)";
|
wenzelm@42167
|
483 |
check "[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]"
|
wenzelm@42167
|
484 |
"concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)";
|
wenzelm@42167
|
485 |
check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]"
|
wenzelm@42167
|
486 |
"concat (map (\<lambda>x. concat (map (\<lambda>y. if x < y then [(x, y, z)] else []) ys)) xs)";
|
wenzelm@42167
|
487 |
check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]"
|
wenzelm@42167
|
488 |
"concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)"
|
wenzelm@42167
|
489 |
end;
|
wenzelm@42167
|
490 |
*}
|
wenzelm@42167
|
491 |
|
wenzelm@35115
|
492 |
(*
|
nipkow@24349
|
493 |
term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]"
|
nipkow@23192
|
494 |
*)
|
nipkow@23192
|
495 |
|
wenzelm@42167
|
496 |
|
wenzelm@50422
|
497 |
ML {*
|
wenzelm@50422
|
498 |
(* Simproc for rewriting list comprehensions applied to List.set to set
|
wenzelm@50422
|
499 |
comprehension. *)
|
wenzelm@50422
|
500 |
|
wenzelm@50422
|
501 |
signature LIST_TO_SET_COMPREHENSION =
|
wenzelm@50422
|
502 |
sig
|
wenzelm@51717
|
503 |
val simproc : Proof.context -> cterm -> thm option
|
wenzelm@50422
|
504 |
end
|
wenzelm@50422
|
505 |
|
wenzelm@50422
|
506 |
structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION =
|
wenzelm@50422
|
507 |
struct
|
wenzelm@50422
|
508 |
|
wenzelm@50422
|
509 |
(* conversion *)
|
wenzelm@50422
|
510 |
|
wenzelm@50422
|
511 |
fun all_exists_conv cv ctxt ct =
|
wenzelm@50422
|
512 |
(case Thm.term_of ct of
|
wenzelm@50422
|
513 |
Const (@{const_name HOL.Ex}, _) $ Abs _ =>
|
wenzelm@50422
|
514 |
Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct
|
wenzelm@50422
|
515 |
| _ => cv ctxt ct)
|
wenzelm@50422
|
516 |
|
wenzelm@50422
|
517 |
fun all_but_last_exists_conv cv ctxt ct =
|
wenzelm@50422
|
518 |
(case Thm.term_of ct of
|
wenzelm@50422
|
519 |
Const (@{const_name HOL.Ex}, _) $ Abs (_, _, Const (@{const_name HOL.Ex}, _) $ _) =>
|
wenzelm@50422
|
520 |
Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct
|
wenzelm@50422
|
521 |
| _ => cv ctxt ct)
|
wenzelm@50422
|
522 |
|
wenzelm@50422
|
523 |
fun Collect_conv cv ctxt ct =
|
wenzelm@50422
|
524 |
(case Thm.term_of ct of
|
wenzelm@50422
|
525 |
Const (@{const_name Set.Collect}, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct
|
wenzelm@50422
|
526 |
| _ => raise CTERM ("Collect_conv", [ct]))
|
wenzelm@50422
|
527 |
|
wenzelm@50422
|
528 |
fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th)
|
wenzelm@50422
|
529 |
|
wenzelm@50422
|
530 |
fun conjunct_assoc_conv ct =
|
wenzelm@50422
|
531 |
Conv.try_conv
|
wenzelm@51315
|
532 |
(rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct
|
wenzelm@50422
|
533 |
|
wenzelm@50422
|
534 |
fun right_hand_set_comprehension_conv conv ctxt =
|
wenzelm@51315
|
535 |
HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv
|
wenzelm@50422
|
536 |
(Collect_conv (all_exists_conv conv o #2) ctxt))
|
wenzelm@50422
|
537 |
|
wenzelm@50422
|
538 |
|
wenzelm@50422
|
539 |
(* term abstraction of list comprehension patterns *)
|
wenzelm@50422
|
540 |
|
wenzelm@50422
|
541 |
datatype termlets = If | Case of (typ * int)
|
wenzelm@50422
|
542 |
|
wenzelm@51717
|
543 |
fun simproc ctxt redex =
|
wenzelm@50422
|
544 |
let
|
wenzelm@50422
|
545 |
val set_Nil_I = @{thm trans} OF [@{thm set.simps(1)}, @{thm empty_def}]
|
wenzelm@50422
|
546 |
val set_singleton = @{lemma "set [a] = {x. x = a}" by simp}
|
wenzelm@50422
|
547 |
val inst_Collect_mem_eq = @{lemma "set A = {x. x : set A}" by simp}
|
wenzelm@50422
|
548 |
val del_refl_eq = @{lemma "(t = t & P) == P" by simp}
|
wenzelm@50422
|
549 |
fun mk_set T = Const (@{const_name List.set}, HOLogic.listT T --> HOLogic.mk_setT T)
|
wenzelm@50422
|
550 |
fun dest_set (Const (@{const_name List.set}, _) $ xs) = xs
|
wenzelm@50422
|
551 |
fun dest_singleton_list (Const (@{const_name List.Cons}, _)
|
wenzelm@50422
|
552 |
$ t $ (Const (@{const_name List.Nil}, _))) = t
|
wenzelm@50422
|
553 |
| dest_singleton_list t = raise TERM ("dest_singleton_list", [t])
|
wenzelm@50422
|
554 |
(* We check that one case returns a singleton list and all other cases
|
wenzelm@50422
|
555 |
return [], and return the index of the one singleton list case *)
|
wenzelm@50422
|
556 |
fun possible_index_of_singleton_case cases =
|
wenzelm@50422
|
557 |
let
|
wenzelm@50422
|
558 |
fun check (i, case_t) s =
|
wenzelm@50422
|
559 |
(case strip_abs_body case_t of
|
wenzelm@50422
|
560 |
(Const (@{const_name List.Nil}, _)) => s
|
traytel@53412
|
561 |
| _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE))
|
wenzelm@50422
|
562 |
in
|
traytel@53412
|
563 |
fold_index check cases (SOME NONE) |> the_default NONE
|
wenzelm@50422
|
564 |
end
|
blanchet@54404
|
565 |
(* returns (case_expr type index chosen_case constr_name) option *)
|
wenzelm@50422
|
566 |
fun dest_case case_term =
|
wenzelm@50422
|
567 |
let
|
wenzelm@50422
|
568 |
val (case_const, args) = strip_comb case_term
|
wenzelm@50422
|
569 |
in
|
wenzelm@50422
|
570 |
(case try dest_Const case_const of
|
wenzelm@50422
|
571 |
SOME (c, T) =>
|
blanchet@54404
|
572 |
(case Ctr_Sugar.ctr_sugar_of_case ctxt c of
|
blanchet@54404
|
573 |
SOME {ctrs, ...} =>
|
wenzelm@50422
|
574 |
(case possible_index_of_singleton_case (fst (split_last args)) of
|
wenzelm@50422
|
575 |
SOME i =>
|
wenzelm@50422
|
576 |
let
|
blanchet@54404
|
577 |
val constr_names = map (fst o dest_Const) ctrs
|
wenzelm@50422
|
578 |
val (Ts, _) = strip_type T
|
wenzelm@50422
|
579 |
val T' = List.last Ts
|
blanchet@54404
|
580 |
in SOME (List.last args, T', i, nth args i, nth constr_names i) end
|
wenzelm@50422
|
581 |
| NONE => NONE)
|
wenzelm@50422
|
582 |
| NONE => NONE)
|
wenzelm@50422
|
583 |
| NONE => NONE)
|
wenzelm@50422
|
584 |
end
|
wenzelm@50422
|
585 |
(* returns condition continuing term option *)
|
wenzelm@50422
|
586 |
fun dest_if (Const (@{const_name If}, _) $ cond $ then_t $ Const (@{const_name Nil}, _)) =
|
wenzelm@50422
|
587 |
SOME (cond, then_t)
|
wenzelm@50422
|
588 |
| dest_if _ = NONE
|
wenzelm@50422
|
589 |
fun tac _ [] = rtac set_singleton 1 ORELSE rtac inst_Collect_mem_eq 1
|
wenzelm@50422
|
590 |
| tac ctxt (If :: cont) =
|
wenzelm@50422
|
591 |
Splitter.split_tac [@{thm split_if}] 1
|
wenzelm@50422
|
592 |
THEN rtac @{thm conjI} 1
|
wenzelm@50422
|
593 |
THEN rtac @{thm impI} 1
|
wenzelm@50422
|
594 |
THEN Subgoal.FOCUS (fn {prems, context, ...} =>
|
wenzelm@50422
|
595 |
CONVERSION (right_hand_set_comprehension_conv (K
|
wenzelm@51315
|
596 |
(HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
|
wenzelm@50422
|
597 |
then_conv
|
wenzelm@50422
|
598 |
rewr_conv' @{lemma "(True & P) = P" by simp})) context) 1) ctxt 1
|
wenzelm@50422
|
599 |
THEN tac ctxt cont
|
wenzelm@50422
|
600 |
THEN rtac @{thm impI} 1
|
wenzelm@50422
|
601 |
THEN Subgoal.FOCUS (fn {prems, context, ...} =>
|
wenzelm@50422
|
602 |
CONVERSION (right_hand_set_comprehension_conv (K
|
wenzelm@51315
|
603 |
(HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
|
wenzelm@50422
|
604 |
then_conv rewr_conv' @{lemma "(False & P) = False" by simp})) context) 1) ctxt 1
|
wenzelm@50422
|
605 |
THEN rtac set_Nil_I 1
|
wenzelm@50422
|
606 |
| tac ctxt (Case (T, i) :: cont) =
|
wenzelm@50422
|
607 |
let
|
blanchet@54404
|
608 |
val SOME {injects, distincts, case_thms, split, ...} =
|
blanchet@54404
|
609 |
Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T))
|
wenzelm@50422
|
610 |
in
|
wenzelm@50422
|
611 |
(* do case distinction *)
|
blanchet@54404
|
612 |
Splitter.split_tac [split] 1
|
wenzelm@50422
|
613 |
THEN EVERY (map_index (fn (i', _) =>
|
blanchet@54404
|
614 |
(if i' < length case_thms - 1 then rtac @{thm conjI} 1 else all_tac)
|
wenzelm@50422
|
615 |
THEN REPEAT_DETERM (rtac @{thm allI} 1)
|
wenzelm@50422
|
616 |
THEN rtac @{thm impI} 1
|
wenzelm@50422
|
617 |
THEN (if i' = i then
|
wenzelm@50422
|
618 |
(* continue recursively *)
|
wenzelm@50422
|
619 |
Subgoal.FOCUS (fn {prems, context, ...} =>
|
wenzelm@50422
|
620 |
CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K
|
wenzelm@51315
|
621 |
((HOLogic.conj_conv
|
wenzelm@51315
|
622 |
(HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv
|
blanchet@54404
|
623 |
(Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects))))
|
wenzelm@50422
|
624 |
Conv.all_conv)
|
wenzelm@50422
|
625 |
then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq))
|
wenzelm@50422
|
626 |
then_conv conjunct_assoc_conv)) context
|
wenzelm@51315
|
627 |
then_conv (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt) =>
|
wenzelm@50422
|
628 |
Conv.repeat_conv
|
wenzelm@50422
|
629 |
(all_but_last_exists_conv
|
wenzelm@50422
|
630 |
(K (rewr_conv'
|
wenzelm@50422
|
631 |
@{lemma "(EX x. x = t & P x) = P t" by simp})) ctxt)) context)))) 1) ctxt 1
|
wenzelm@50422
|
632 |
THEN tac ctxt cont
|
wenzelm@50422
|
633 |
else
|
wenzelm@50422
|
634 |
Subgoal.FOCUS (fn {prems, context, ...} =>
|
wenzelm@50422
|
635 |
CONVERSION
|
wenzelm@50422
|
636 |
(right_hand_set_comprehension_conv (K
|
wenzelm@51315
|
637 |
(HOLogic.conj_conv
|
wenzelm@51315
|
638 |
((HOLogic.eq_conv Conv.all_conv
|
wenzelm@50422
|
639 |
(rewr_conv' (List.last prems))) then_conv
|
blanchet@54404
|
640 |
(Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts)))
|
wenzelm@50422
|
641 |
Conv.all_conv then_conv
|
wenzelm@50422
|
642 |
(rewr_conv' @{lemma "(False & P) = False" by simp}))) context then_conv
|
wenzelm@51314
|
643 |
HOLogic.Trueprop_conv
|
wenzelm@51315
|
644 |
(HOLogic.eq_conv Conv.all_conv
|
wenzelm@50422
|
645 |
(Collect_conv (fn (_, ctxt) =>
|
wenzelm@50422
|
646 |
Conv.repeat_conv
|
wenzelm@50422
|
647 |
(Conv.bottom_conv
|
wenzelm@50422
|
648 |
(K (rewr_conv'
|
wenzelm@50422
|
649 |
@{lemma "(EX x. P) = P" by simp})) ctxt)) context))) 1) ctxt 1
|
blanchet@54404
|
650 |
THEN rtac set_Nil_I 1)) case_thms)
|
wenzelm@50422
|
651 |
end
|
wenzelm@50422
|
652 |
fun make_inner_eqs bound_vs Tis eqs t =
|
wenzelm@50422
|
653 |
(case dest_case t of
|
blanchet@54404
|
654 |
SOME (x, T, i, cont, constr_name) =>
|
wenzelm@50422
|
655 |
let
|
wenzelm@52131
|
656 |
val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont)
|
wenzelm@50422
|
657 |
val x' = incr_boundvars (length vs) x
|
wenzelm@50422
|
658 |
val eqs' = map (incr_boundvars (length vs)) eqs
|
wenzelm@50422
|
659 |
val constr_t =
|
wenzelm@50422
|
660 |
list_comb
|
wenzelm@50422
|
661 |
(Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0))
|
wenzelm@50422
|
662 |
val constr_eq = Const (@{const_name HOL.eq}, T --> T --> @{typ bool}) $ constr_t $ x'
|
wenzelm@50422
|
663 |
in
|
wenzelm@50422
|
664 |
make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body
|
wenzelm@50422
|
665 |
end
|
wenzelm@50422
|
666 |
| NONE =>
|
wenzelm@50422
|
667 |
(case dest_if t of
|
wenzelm@50422
|
668 |
SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont
|
wenzelm@50422
|
669 |
| NONE =>
|
wenzelm@50422
|
670 |
if eqs = [] then NONE (* no rewriting, nothing to be done *)
|
wenzelm@50422
|
671 |
else
|
wenzelm@50422
|
672 |
let
|
wenzelm@50422
|
673 |
val Type (@{type_name List.list}, [rT]) = fastype_of1 (map snd bound_vs, t)
|
wenzelm@50422
|
674 |
val pat_eq =
|
wenzelm@50422
|
675 |
(case try dest_singleton_list t of
|
wenzelm@50422
|
676 |
SOME t' =>
|
wenzelm@50422
|
677 |
Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool}) $
|
wenzelm@50422
|
678 |
Bound (length bound_vs) $ t'
|
wenzelm@50422
|
679 |
| NONE =>
|
wenzelm@50422
|
680 |
Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool}) $
|
wenzelm@50422
|
681 |
Bound (length bound_vs) $ (mk_set rT $ t))
|
wenzelm@50422
|
682 |
val reverse_bounds = curry subst_bounds
|
wenzelm@50422
|
683 |
((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)])
|
wenzelm@50422
|
684 |
val eqs' = map reverse_bounds eqs
|
wenzelm@50422
|
685 |
val pat_eq' = reverse_bounds pat_eq
|
wenzelm@50422
|
686 |
val inner_t =
|
wenzelm@50422
|
687 |
fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t)
|
wenzelm@50422
|
688 |
(rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
|
wenzelm@50422
|
689 |
val lhs = term_of redex
|
wenzelm@50422
|
690 |
val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
|
wenzelm@50422
|
691 |
val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
|
wenzelm@50422
|
692 |
in
|
wenzelm@50422
|
693 |
SOME
|
wenzelm@50422
|
694 |
((Goal.prove ctxt [] [] rewrite_rule_t
|
wenzelm@50422
|
695 |
(fn {context, ...} => tac context (rev Tis))) RS @{thm eq_reflection})
|
wenzelm@50422
|
696 |
end))
|
wenzelm@50422
|
697 |
in
|
wenzelm@50422
|
698 |
make_inner_eqs [] [] [] (dest_set (term_of redex))
|
wenzelm@50422
|
699 |
end
|
wenzelm@50422
|
700 |
|
wenzelm@50422
|
701 |
end
|
wenzelm@50422
|
702 |
*}
|
bulwahn@41463
|
703 |
|
bulwahn@41463
|
704 |
simproc_setup list_to_set_comprehension ("set xs") = {* K List_to_Set_Comprehension.simproc *}
|
bulwahn@41463
|
705 |
|
haftmann@46133
|
706 |
code_datatype set coset
|
haftmann@46133
|
707 |
|
haftmann@46133
|
708 |
hide_const (open) coset
|
wenzelm@35115
|
709 |
|
haftmann@49948
|
710 |
|
haftmann@21061
|
711 |
subsubsection {* @{const Nil} and @{const Cons} *}
|
haftmann@21061
|
712 |
|
haftmann@21061
|
713 |
lemma not_Cons_self [simp]:
|
haftmann@21061
|
714 |
"xs \<noteq> x # xs"
|
nipkow@13145
|
715 |
by (induct xs) auto
|
wenzelm@13114
|
716 |
|
wenzelm@41697
|
717 |
lemma not_Cons_self2 [simp]:
|
wenzelm@41697
|
718 |
"x # xs \<noteq> xs"
|
wenzelm@41697
|
719 |
by (rule not_Cons_self [symmetric])
|
wenzelm@13114
|
720 |
|
wenzelm@13142
|
721 |
lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
|
nipkow@13145
|
722 |
by (induct xs) auto
|
wenzelm@13114
|
723 |
|
nipkow@53689
|
724 |
lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
|
nipkow@53689
|
725 |
by (cases xs) auto
|
nipkow@53689
|
726 |
|
nipkow@53689
|
727 |
lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
|
nipkow@53689
|
728 |
by (cases xs) auto
|
nipkow@53689
|
729 |
|
wenzelm@13142
|
730 |
lemma length_induct:
|
haftmann@21061
|
731 |
"(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs"
|
nipkow@53689
|
732 |
by (fact measure_induct)
|
wenzelm@13114
|
733 |
|
haftmann@37289
|
734 |
lemma list_nonempty_induct [consumes 1, case_names single cons]:
|
haftmann@37289
|
735 |
assumes "xs \<noteq> []"
|
haftmann@37289
|
736 |
assumes single: "\<And>x. P [x]"
|
haftmann@37289
|
737 |
assumes cons: "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)"
|
haftmann@37289
|
738 |
shows "P xs"
|
haftmann@37289
|
739 |
using `xs \<noteq> []` proof (induct xs)
|
haftmann@37289
|
740 |
case Nil then show ?case by simp
|
haftmann@37289
|
741 |
next
|
wenzelm@53374
|
742 |
case (Cons x xs)
|
wenzelm@53374
|
743 |
show ?case
|
wenzelm@53374
|
744 |
proof (cases xs)
|
wenzelm@53374
|
745 |
case Nil
|
wenzelm@53374
|
746 |
with single show ?thesis by simp
|
haftmann@37289
|
747 |
next
|
wenzelm@53374
|
748 |
case Cons
|
wenzelm@53374
|
749 |
show ?thesis
|
wenzelm@53374
|
750 |
proof (rule cons)
|
wenzelm@53374
|
751 |
from Cons show "xs \<noteq> []" by simp
|
wenzelm@53374
|
752 |
with Cons.hyps show "P xs" .
|
wenzelm@53374
|
753 |
qed
|
haftmann@37289
|
754 |
qed
|
haftmann@37289
|
755 |
qed
|
haftmann@37289
|
756 |
|
hoelzl@45714
|
757 |
lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X"
|
hoelzl@45714
|
758 |
by (auto intro!: inj_onI)
|
wenzelm@13114
|
759 |
|
haftmann@49948
|
760 |
|
haftmann@21061
|
761 |
subsubsection {* @{const length} *}
|
wenzelm@13114
|
762 |
|
wenzelm@13142
|
763 |
text {*
|
haftmann@21061
|
764 |
Needs to come before @{text "@"} because of theorem @{text
|
haftmann@21061
|
765 |
append_eq_append_conv}.
|
wenzelm@13142
|
766 |
*}
|
wenzelm@13114
|
767 |
|
wenzelm@13142
|
768 |
lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
|
nipkow@13145
|
769 |
by (induct xs) auto
|
wenzelm@13114
|
770 |
|
wenzelm@13142
|
771 |
lemma length_map [simp]: "length (map f xs) = length xs"
|
nipkow@13145
|
772 |
by (induct xs) auto
|
wenzelm@13114
|
773 |
|
wenzelm@13142
|
774 |
lemma length_rev [simp]: "length (rev xs) = length xs"
|
nipkow@13145
|
775 |
by (induct xs) auto
|
wenzelm@13114
|
776 |
|
wenzelm@13142
|
777 |
lemma length_tl [simp]: "length (tl xs) = length xs - 1"
|
nipkow@13145
|
778 |
by (cases xs) auto
|
wenzelm@13114
|
779 |
|
wenzelm@13142
|
780 |
lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])"
|
nipkow@13145
|
781 |
by (induct xs) auto
|
wenzelm@13114
|
782 |
|
wenzelm@13142
|
783 |
lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])"
|
nipkow@13145
|
784 |
by (induct xs) auto
|
wenzelm@13114
|
785 |
|
nipkow@23479
|
786 |
lemma length_pos_if_in_set: "x : set xs \<Longrightarrow> length xs > 0"
|
nipkow@23479
|
787 |
by auto
|
nipkow@23479
|
788 |
|
wenzelm@13114
|
789 |
lemma length_Suc_conv:
|
nipkow@13145
|
790 |
"(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
|
nipkow@13145
|
791 |
by (induct xs) auto
|
wenzelm@13142
|
792 |
|
nipkow@14025
|
793 |
lemma Suc_length_conv:
|
nipkow@14025
|
794 |
"(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
|
paulson@14208
|
795 |
apply (induct xs, simp, simp)
|
nipkow@14025
|
796 |
apply blast
|
nipkow@14025
|
797 |
done
|
nipkow@14025
|
798 |
|
wenzelm@25221
|
799 |
lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False"
|
wenzelm@25221
|
800 |
by (induct xs) auto
|
wenzelm@25221
|
801 |
|
haftmann@26442
|
802 |
lemma list_induct2 [consumes 1, case_names Nil Cons]:
|
haftmann@26442
|
803 |
"length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow>
|
haftmann@26442
|
804 |
(\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys))
|
haftmann@26442
|
805 |
\<Longrightarrow> P xs ys"
|
haftmann@26442
|
806 |
proof (induct xs arbitrary: ys)
|
haftmann@26442
|
807 |
case Nil then show ?case by simp
|
haftmann@26442
|
808 |
next
|
haftmann@26442
|
809 |
case (Cons x xs ys) then show ?case by (cases ys) simp_all
|
haftmann@26442
|
810 |
qed
|
haftmann@26442
|
811 |
|
haftmann@26442
|
812 |
lemma list_induct3 [consumes 2, case_names Nil Cons]:
|
haftmann@26442
|
813 |
"length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow>
|
haftmann@26442
|
814 |
(\<And>x xs y ys z zs. length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P xs ys zs \<Longrightarrow> P (x#xs) (y#ys) (z#zs))
|
haftmann@26442
|
815 |
\<Longrightarrow> P xs ys zs"
|
haftmann@26442
|
816 |
proof (induct xs arbitrary: ys zs)
|
haftmann@26442
|
817 |
case Nil then show ?case by simp
|
haftmann@26442
|
818 |
next
|
haftmann@26442
|
819 |
case (Cons x xs ys zs) then show ?case by (cases ys, simp_all)
|
haftmann@26442
|
820 |
(cases zs, simp_all)
|
haftmann@26442
|
821 |
qed
|
wenzelm@13114
|
822 |
|
kaliszyk@36154
|
823 |
lemma list_induct4 [consumes 3, case_names Nil Cons]:
|
kaliszyk@36154
|
824 |
"length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow>
|
kaliszyk@36154
|
825 |
P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow>
|
kaliszyk@36154
|
826 |
length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow>
|
kaliszyk@36154
|
827 |
P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws"
|
kaliszyk@36154
|
828 |
proof (induct xs arbitrary: ys zs ws)
|
kaliszyk@36154
|
829 |
case Nil then show ?case by simp
|
kaliszyk@36154
|
830 |
next
|
kaliszyk@36154
|
831 |
case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all)
|
kaliszyk@36154
|
832 |
qed
|
kaliszyk@36154
|
833 |
|
krauss@22493
|
834 |
lemma list_induct2':
|
krauss@22493
|
835 |
"\<lbrakk> P [] [];
|
krauss@22493
|
836 |
\<And>x xs. P (x#xs) [];
|
krauss@22493
|
837 |
\<And>y ys. P [] (y#ys);
|
krauss@22493
|
838 |
\<And>x xs y ys. P xs ys \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
|
krauss@22493
|
839 |
\<Longrightarrow> P xs ys"
|
krauss@22493
|
840 |
by (induct xs arbitrary: ys) (case_tac x, auto)+
|
krauss@22493
|
841 |
|
nipkow@22143
|
842 |
lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False"
|
nipkow@24349
|
843 |
by (rule Eq_FalseI) auto
|
wenzelm@24037
|
844 |
|
wenzelm@24037
|
845 |
simproc_setup list_neq ("(xs::'a list) = ys") = {*
|
nipkow@22143
|
846 |
(*
|
nipkow@22143
|
847 |
Reduces xs=ys to False if xs and ys cannot be of the same length.
|
nipkow@22143
|
848 |
This is the case if the atomic sublists of one are a submultiset
|
nipkow@22143
|
849 |
of those of the other list and there are fewer Cons's in one than the other.
|
nipkow@22143
|
850 |
*)
|
wenzelm@24037
|
851 |
|
wenzelm@24037
|
852 |
let
|
nipkow@22143
|
853 |
|
huffman@29856
|
854 |
fun len (Const(@{const_name Nil},_)) acc = acc
|
huffman@29856
|
855 |
| len (Const(@{const_name Cons},_) $ _ $ xs) (ts,n) = len xs (ts,n+1)
|
huffman@29856
|
856 |
| len (Const(@{const_name append},_) $ xs $ ys) acc = len xs (len ys acc)
|
huffman@29856
|
857 |
| len (Const(@{const_name rev},_) $ xs) acc = len xs acc
|
huffman@29856
|
858 |
| len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc
|
nipkow@22143
|
859 |
| len t (ts,n) = (t::ts,n);
|
nipkow@22143
|
860 |
|
wenzelm@51717
|
861 |
val ss = simpset_of @{context};
|
wenzelm@51717
|
862 |
|
wenzelm@51717
|
863 |
fun list_neq ctxt ct =
|
nipkow@22143
|
864 |
let
|
wenzelm@24037
|
865 |
val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct;
|
nipkow@22143
|
866 |
val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0);
|
nipkow@22143
|
867 |
fun prove_neq() =
|
nipkow@22143
|
868 |
let
|
nipkow@22143
|
869 |
val Type(_,listT::_) = eqT;
|
haftmann@22994
|
870 |
val size = HOLogic.size_const listT;
|
nipkow@22143
|
871 |
val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs);
|
nipkow@22143
|
872 |
val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len);
|
wenzelm@51717
|
873 |
val thm = Goal.prove ctxt [] [] neq_len
|
wenzelm@51717
|
874 |
(K (simp_tac (put_simpset ss ctxt) 1));
|
haftmann@22633
|
875 |
in SOME (thm RS @{thm neq_if_length_neq}) end
|
nipkow@22143
|
876 |
in
|
wenzelm@23214
|
877 |
if m < n andalso submultiset (op aconv) (ls,rs) orelse
|
wenzelm@23214
|
878 |
n < m andalso submultiset (op aconv) (rs,ls)
|
nipkow@22143
|
879 |
then prove_neq() else NONE
|
nipkow@22143
|
880 |
end;
|
wenzelm@51717
|
881 |
in K list_neq end;
|
nipkow@22143
|
882 |
*}
|
nipkow@22143
|
883 |
|
nipkow@22143
|
884 |
|
nipkow@15392
|
885 |
subsubsection {* @{text "@"} -- append *}
|
wenzelm@13114
|
886 |
|
wenzelm@13142
|
887 |
lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
|
nipkow@13145
|
888 |
by (induct xs) auto
|
wenzelm@13114
|
889 |
|
wenzelm@13142
|
890 |
lemma append_Nil2 [simp]: "xs @ [] = xs"
|
nipkow@13145
|
891 |
by (induct xs) auto
|
nipkow@3507
|
892 |
|
wenzelm@13142
|
893 |
lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
|
nipkow@13145
|
894 |
by (induct xs) auto
|
wenzelm@13114
|
895 |
|
wenzelm@13142
|
896 |
lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])"
|
nipkow@13145
|
897 |
by (induct xs) auto
|
wenzelm@13114
|
898 |
|
wenzelm@13142
|
899 |
lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
|
nipkow@13145
|
900 |
by (induct xs) auto
|
wenzelm@13114
|
901 |
|
wenzelm@13142
|
902 |
lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
|
nipkow@13145
|
903 |
by (induct xs) auto
|
wenzelm@13114
|
904 |
|
blanchet@54147
|
905 |
lemma append_eq_append_conv [simp]:
|
nipkow@24526
|
906 |
"length xs = length ys \<or> length us = length vs
|
berghofe@13883
|
907 |
==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
|
nipkow@24526
|
908 |
apply (induct xs arbitrary: ys)
|
paulson@14208
|
909 |
apply (case_tac ys, simp, force)
|
paulson@14208
|
910 |
apply (case_tac ys, force, simp)
|
nipkow@13145
|
911 |
done
|
wenzelm@13142
|
912 |
|
nipkow@24526
|
913 |
lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
|
nipkow@24526
|
914 |
(EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)"
|
nipkow@24526
|
915 |
apply (induct xs arbitrary: ys zs ts)
|
nipkow@44890
|
916 |
apply fastforce
|
nipkow@14495
|
917 |
apply(case_tac zs)
|
nipkow@14495
|
918 |
apply simp
|
nipkow@44890
|
919 |
apply fastforce
|
nipkow@14495
|
920 |
done
|
nipkow@14495
|
921 |
|
berghofe@34910
|
922 |
lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)"
|
nipkow@13145
|
923 |
by simp
|
wenzelm@13142
|
924 |
|
wenzelm@13142
|
925 |
lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)"
|
nipkow@13145
|
926 |
by simp
|
wenzelm@13114
|
927 |
|
berghofe@34910
|
928 |
lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)"
|
nipkow@13145
|
929 |
by simp
|
wenzelm@13114
|
930 |
|
wenzelm@13142
|
931 |
lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])"
|
nipkow@13145
|
932 |
using append_same_eq [of _ _ "[]"] by auto
|
nipkow@3507
|
933 |
|
wenzelm@13142
|
934 |
lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
|
nipkow@13145
|
935 |
using append_same_eq [of "[]"] by auto
|
wenzelm@13114
|
936 |
|
blanchet@54147
|
937 |
lemma hd_Cons_tl [simp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
|
nipkow@13145
|
938 |
by (induct xs) auto
|
wenzelm@13114
|
939 |
|
wenzelm@13142
|
940 |
lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
|
nipkow@13145
|
941 |
by (induct xs) auto
|
wenzelm@13114
|
942 |
|
wenzelm@13142
|
943 |
lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
|
nipkow@13145
|
944 |
by (simp add: hd_append split: list.split)
|
wenzelm@13114
|
945 |
|
wenzelm@13142
|
946 |
lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)"
|
nipkow@13145
|
947 |
by (simp split: list.split)
|
wenzelm@13114
|
948 |
|
wenzelm@13142
|
949 |
lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
|
nipkow@13145
|
950 |
by (simp add: tl_append split: list.split)
|
wenzelm@13114
|
951 |
|
wenzelm@13114
|
952 |
|
nipkow@14300
|
953 |
lemma Cons_eq_append_conv: "x#xs = ys@zs =
|
nipkow@14300
|
954 |
(ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))"
|
nipkow@14300
|
955 |
by(cases ys) auto
|
nipkow@14300
|
956 |
|
nipkow@15281
|
957 |
lemma append_eq_Cons_conv: "(ys@zs = x#xs) =
|
nipkow@15281
|
958 |
(ys = [] & zs = x#xs | (EX ys'. ys = x#ys' & ys'@zs = xs))"
|
nipkow@15281
|
959 |
by(cases ys) auto
|
nipkow@15281
|
960 |
|
nipkow@14300
|
961 |
|
wenzelm@13142
|
962 |
text {* Trivial rules for solving @{text "@"}-equations automatically. *}
|
wenzelm@13114
|
963 |
|
wenzelm@13114
|
964 |
lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
|
nipkow@13145
|
965 |
by simp
|
wenzelm@13114
|
966 |
|
wenzelm@13142
|
967 |
lemma Cons_eq_appendI:
|
nipkow@13145
|
968 |
"[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
|
nipkow@13145
|
969 |
by (drule sym) simp
|
wenzelm@13114
|
970 |
|
wenzelm@13142
|
971 |
lemma append_eq_appendI:
|
nipkow@13145
|
972 |
"[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
|
nipkow@13145
|
973 |
by (drule sym) simp
|
wenzelm@13114
|
974 |
|
wenzelm@13114
|
975 |
|
wenzelm@13142
|
976 |
text {*
|
nipkow@13145
|
977 |
Simplification procedure for all list equalities.
|
nipkow@13145
|
978 |
Currently only tries to rearrange @{text "@"} to see if
|
nipkow@13145
|
979 |
- both lists end in a singleton list,
|
nipkow@13145
|
980 |
- or both lists end in the same list.
|
wenzelm@13142
|
981 |
*}
|
wenzelm@13142
|
982 |
|
wenzelm@43594
|
983 |
simproc_setup list_eq ("(xs::'a list) = ys") = {*
|
wenzelm@13462
|
984 |
let
|
wenzelm@43594
|
985 |
fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) =
|
wenzelm@43594
|
986 |
(case xs of Const (@{const_name Nil}, _) => cons | _ => last xs)
|
wenzelm@43594
|
987 |
| last (Const(@{const_name append},_) $ _ $ ys) = last ys
|
wenzelm@43594
|
988 |
| last t = t;
|
wenzelm@43594
|
989 |
|
wenzelm@43594
|
990 |
fun list1 (Const(@{const_name Cons},_) $ _ $ Const(@{const_name Nil},_)) = true
|
wenzelm@43594
|
991 |
| list1 _ = false;
|
wenzelm@43594
|
992 |
|
wenzelm@43594
|
993 |
fun butlast ((cons as Const(@{const_name Cons},_) $ x) $ xs) =
|
wenzelm@43594
|
994 |
(case xs of Const (@{const_name Nil}, _) => xs | _ => cons $ butlast xs)
|
wenzelm@43594
|
995 |
| butlast ((app as Const (@{const_name append}, _) $ xs) $ ys) = app $ butlast ys
|
wenzelm@43594
|
996 |
| butlast xs = Const(@{const_name Nil}, fastype_of xs);
|
wenzelm@43594
|
997 |
|
wenzelm@43594
|
998 |
val rearr_ss =
|
wenzelm@51717
|
999 |
simpset_of (put_simpset HOL_basic_ss @{context}
|
wenzelm@51717
|
1000 |
addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]);
|
wenzelm@43594
|
1001 |
|
wenzelm@51717
|
1002 |
fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) =
|
wenzelm@13462
|
1003 |
let
|
wenzelm@43594
|
1004 |
val lastl = last lhs and lastr = last rhs;
|
wenzelm@43594
|
1005 |
fun rearr conv =
|
wenzelm@43594
|
1006 |
let
|
wenzelm@43594
|
1007 |
val lhs1 = butlast lhs and rhs1 = butlast rhs;
|
wenzelm@43594
|
1008 |
val Type(_,listT::_) = eqT
|
wenzelm@43594
|
1009 |
val appT = [listT,listT] ---> listT
|
wenzelm@43594
|
1010 |
val app = Const(@{const_name append},appT)
|
wenzelm@43594
|
1011 |
val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
|
wenzelm@43594
|
1012 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
|
wenzelm@51717
|
1013 |
val thm = Goal.prove ctxt [] [] eq
|
wenzelm@51717
|
1014 |
(K (simp_tac (put_simpset rearr_ss ctxt) 1));
|
wenzelm@43594
|
1015 |
in SOME ((conv RS (thm RS trans)) RS eq_reflection) end;
|
wenzelm@43594
|
1016 |
in
|
wenzelm@43594
|
1017 |
if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv}
|
wenzelm@43594
|
1018 |
else if lastl aconv lastr then rearr @{thm append_same_eq}
|
wenzelm@43594
|
1019 |
else NONE
|
wenzelm@43594
|
1020 |
end;
|
wenzelm@51717
|
1021 |
in fn _ => fn ctxt => fn ct => list_eq ctxt (term_of ct) end;
|
wenzelm@13114
|
1022 |
*}
|
wenzelm@13114
|
1023 |
|
wenzelm@13114
|
1024 |
|
haftmann@49948
|
1025 |
subsubsection {* @{const map} *}
|
wenzelm@13114
|
1026 |
|
haftmann@40210
|
1027 |
lemma hd_map:
|
haftmann@40210
|
1028 |
"xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
|
haftmann@40210
|
1029 |
by (cases xs) simp_all
|
haftmann@40210
|
1030 |
|
haftmann@40210
|
1031 |
lemma map_tl:
|
haftmann@40210
|
1032 |
"map f (tl xs) = tl (map f xs)"
|
haftmann@40210
|
1033 |
by (cases xs) simp_all
|
haftmann@40210
|
1034 |
|
wenzelm@13142
|
1035 |
lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
|
nipkow@13145
|
1036 |
by (induct xs) simp_all
|
wenzelm@13114
|
1037 |
|
wenzelm@13142
|
1038 |
lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
|
nipkow@13145
|
1039 |
by (rule ext, induct_tac xs) auto
|
wenzelm@13114
|
1040 |
|
wenzelm@13142
|
1041 |
lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys"
|
nipkow@13145
|
1042 |
by (induct xs) auto
|
wenzelm@13114
|
1043 |
|
hoelzl@33639
|
1044 |
lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs"
|
hoelzl@33639
|
1045 |
by (induct xs) auto
|
hoelzl@33639
|
1046 |
|
nipkow@35208
|
1047 |
lemma map_comp_map[simp]: "((map f) o (map g)) = map(f o g)"
|
nipkow@35208
|
1048 |
apply(rule ext)
|
nipkow@35208
|
1049 |
apply(simp)
|
nipkow@35208
|
1050 |
done
|
nipkow@35208
|
1051 |
|
wenzelm@13142
|
1052 |
lemma rev_map: "rev (map f xs) = map f (rev xs)"
|
nipkow@13145
|
1053 |
by (induct xs) auto
|
wenzelm@13114
|
1054 |
|
nipkow@13737
|
1055 |
lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
|
nipkow@13737
|
1056 |
by (induct xs) auto
|
nipkow@13737
|
1057 |
|
krauss@44013
|
1058 |
lemma map_cong [fundef_cong]:
|
haftmann@40122
|
1059 |
"xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys"
|
haftmann@40122
|
1060 |
by simp
|
wenzelm@13114
|
1061 |
|
wenzelm@13142
|
1062 |
lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
|
nipkow@13145
|
1063 |
by (cases xs) auto
|
wenzelm@13114
|
1064 |
|
wenzelm@13142
|
1065 |
lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
|
nipkow@13145
|
1066 |
by (cases xs) auto
|
wenzelm@13114
|
1067 |
|
paulson@18447
|
1068 |
lemma map_eq_Cons_conv:
|
nipkow@14025
|
1069 |
"(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)"
|
nipkow@13145
|
1070 |
by (cases xs) auto
|
wenzelm@13114
|
1071 |
|
paulson@18447
|
1072 |
lemma Cons_eq_map_conv:
|
nipkow@14025
|
1073 |
"(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)"
|
nipkow@14025
|
1074 |
by (cases ys) auto
|
nipkow@14025
|
1075 |
|
paulson@18447
|
1076 |
lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1]
|
paulson@18447
|
1077 |
lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1]
|
paulson@18447
|
1078 |
declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!]
|
paulson@18447
|
1079 |
|
nipkow@14111
|
1080 |
lemma ex_map_conv:
|
nipkow@14111
|
1081 |
"(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)"
|
paulson@18447
|
1082 |
by(induct ys, auto simp add: Cons_eq_map_conv)
|
nipkow@14111
|
1083 |
|
nipkow@15110
|
1084 |
lemma map_eq_imp_length_eq:
|
paulson@35510
|
1085 |
assumes "map f xs = map g ys"
|
haftmann@26734
|
1086 |
shows "length xs = length ys"
|
wenzelm@53374
|
1087 |
using assms
|
wenzelm@53374
|
1088 |
proof (induct ys arbitrary: xs)
|
haftmann@26734
|
1089 |
case Nil then show ?case by simp
|
haftmann@26734
|
1090 |
next
|
haftmann@26734
|
1091 |
case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto
|
paulson@35510
|
1092 |
from Cons xs have "map f zs = map g ys" by simp
|
wenzelm@53374
|
1093 |
with Cons have "length zs = length ys" by blast
|
haftmann@26734
|
1094 |
with xs show ?case by simp
|
haftmann@26734
|
1095 |
qed
|
haftmann@26734
|
1096 |
|
nipkow@15110
|
1097 |
lemma map_inj_on:
|
nipkow@15110
|
1098 |
"[| map f xs = map f ys; inj_on f (set xs Un set ys) |]
|
nipkow@15110
|
1099 |
==> xs = ys"
|
nipkow@15110
|
1100 |
apply(frule map_eq_imp_length_eq)
|
nipkow@15110
|
1101 |
apply(rotate_tac -1)
|
nipkow@15110
|
1102 |
apply(induct rule:list_induct2)
|
nipkow@15110
|
1103 |
apply simp
|
nipkow@15110
|
1104 |
apply(simp)
|
nipkow@15110
|
1105 |
apply (blast intro:sym)
|
nipkow@15110
|
1106 |
done
|
nipkow@15110
|
1107 |
|
nipkow@15110
|
1108 |
lemma inj_on_map_eq_map:
|
nipkow@15110
|
1109 |
"inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
|
nipkow@15110
|
1110 |
by(blast dest:map_inj_on)
|
nipkow@15110
|
1111 |
|
wenzelm@13114
|
1112 |
lemma map_injective:
|
nipkow@24526
|
1113 |
"map f xs = map f ys ==> inj f ==> xs = ys"
|
nipkow@24526
|
1114 |
by (induct ys arbitrary: xs) (auto dest!:injD)
|
wenzelm@13114
|
1115 |
|
nipkow@14339
|
1116 |
lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
|
nipkow@14339
|
1117 |
by(blast dest:map_injective)
|
nipkow@14339
|
1118 |
|
wenzelm@13114
|
1119 |
lemma inj_mapI: "inj f ==> inj (map f)"
|
nipkow@17589
|
1120 |
by (iprover dest: map_injective injD intro: inj_onI)
|
wenzelm@13114
|
1121 |
|
wenzelm@13114
|
1122 |
lemma inj_mapD: "inj (map f) ==> inj f"
|
paulson@14208
|
1123 |
apply (unfold inj_on_def, clarify)
|
nipkow@13145
|
1124 |
apply (erule_tac x = "[x]" in ballE)
|
paulson@14208
|
1125 |
apply (erule_tac x = "[y]" in ballE, simp, blast)
|
nipkow@13145
|
1126 |
apply blast
|
nipkow@13145
|
1127 |
done
|
wenzelm@13114
|
1128 |
|
nipkow@14339
|
1129 |
lemma inj_map[iff]: "inj (map f) = inj f"
|
nipkow@13145
|
1130 |
by (blast dest: inj_mapD intro: inj_mapI)
|
wenzelm@13114
|
1131 |
|
nipkow@15303
|
1132 |
lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A"
|
nipkow@15303
|
1133 |
apply(rule inj_onI)
|
nipkow@15303
|
1134 |
apply(erule map_inj_on)
|
nipkow@15303
|
1135 |
apply(blast intro:inj_onI dest:inj_onD)
|
nipkow@15303
|
1136 |
done
|
nipkow@15303
|
1137 |
|
kleing@14343
|
1138 |
lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs"
|
kleing@14343
|
1139 |
by (induct xs, auto)
|
wenzelm@13114
|
1140 |
|
nipkow@14402
|
1141 |
lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
|
nipkow@14402
|
1142 |
by (induct xs) auto
|
nipkow@14402
|
1143 |
|
nipkow@15110
|
1144 |
lemma map_fst_zip[simp]:
|
nipkow@15110
|
1145 |
"length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs"
|
nipkow@15110
|
1146 |
by (induct rule:list_induct2, simp_all)
|
nipkow@15110
|
1147 |
|
nipkow@15110
|
1148 |
lemma map_snd_zip[simp]:
|
nipkow@15110
|
1149 |
"length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
|
nipkow@15110
|
1150 |
by (induct rule:list_induct2, simp_all)
|
nipkow@15110
|
1151 |
|
haftmann@41505
|
1152 |
enriched_type map: map
|
nipkow@47122
|
1153 |
by (simp_all add: id_def)
|
nipkow@47122
|
1154 |
|
haftmann@49948
|
1155 |
declare map.id [simp]
|
haftmann@49948
|
1156 |
|
haftmann@49948
|
1157 |
|
haftmann@49948
|
1158 |
subsubsection {* @{const rev} *}
|
wenzelm@13114
|
1159 |
|
wenzelm@13142
|
1160 |
lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
|
nipkow@13145
|
1161 |
by (induct xs) auto
|
wenzelm@13114
|
1162 |
|
wenzelm@13142
|
1163 |
lemma rev_rev_ident [simp]: "rev (rev xs) = xs"
|
nipkow@13145
|
1164 |
by (induct xs) auto
|
wenzelm@13114
|
1165 |
|
kleing@15870
|
1166 |
lemma rev_swap: "(rev xs = ys) = (xs = rev ys)"
|
kleing@15870
|
1167 |
by auto
|
kleing@15870
|
1168 |
|
wenzelm@13142
|
1169 |
lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])"
|
nipkow@13145
|
1170 |
by (induct xs) auto
|
wenzelm@13114
|
1171 |
|
wenzelm@13142
|
1172 |
lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
|
nipkow@13145
|
1173 |
by (induct xs) auto
|
wenzelm@13114
|
1174 |
|
kleing@15870
|
1175 |
lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])"
|
kleing@15870
|
1176 |
by (cases xs) auto
|
kleing@15870
|
1177 |
|
kleing@15870
|
1178 |
lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
|
kleing@15870
|
1179 |
by (cases xs) auto
|
kleing@15870
|
1180 |
|
blanchet@54147
|
1181 |
lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
|
haftmann@21061
|
1182 |
apply (induct xs arbitrary: ys, force)
|
paulson@14208
|
1183 |
apply (case_tac ys, simp, force)
|
nipkow@13145
|
1184 |
done
|
wenzelm@13114
|
1185 |
|
nipkow@15439
|
1186 |
lemma inj_on_rev[iff]: "inj_on rev A"
|
nipkow@15439
|
1187 |
by(simp add:inj_on_def)
|
nipkow@15439
|
1188 |
|
wenzelm@13366
|
1189 |
lemma rev_induct [case_names Nil snoc]:
|
wenzelm@13366
|
1190 |
"[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
|
berghofe@15489
|
1191 |
apply(simplesubst rev_rev_ident[symmetric])
|
nipkow@13145
|
1192 |
apply(rule_tac list = "rev xs" in list.induct, simp_all)
|
nipkow@13145
|
1193 |
done
|
wenzelm@13114
|
1194 |
|
wenzelm@13366
|
1195 |
lemma rev_exhaust [case_names Nil snoc]:
|
wenzelm@13366
|
1196 |
"(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
|
nipkow@13145
|
1197 |
by (induct xs rule: rev_induct) auto
|
wenzelm@13114
|
1198 |
|
wenzelm@13366
|
1199 |
lemmas rev_cases = rev_exhaust
|
wenzelm@13366
|
1200 |
|
nipkow@18423
|
1201 |
lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
|
nipkow@18423
|
1202 |
by(rule rev_cases[of xs]) auto
|
nipkow@18423
|
1203 |
|
wenzelm@13114
|
1204 |
|
haftmann@49948
|
1205 |
subsubsection {* @{const set} *}
|
wenzelm@13114
|
1206 |
|
nipkow@46698
|
1207 |
declare set.simps [code_post] --"pretty output"
|
nipkow@46698
|
1208 |
|
wenzelm@13142
|
1209 |
lemma finite_set [iff]: "finite (set xs)"
|
nipkow@13145
|
1210 |
by (induct xs) auto
|
wenzelm@13114
|
1211 |
|
wenzelm@13142
|
1212 |
lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)"
|
nipkow@13145
|
1213 |
by (induct xs) auto
|
wenzelm@13114
|
1214 |
|
nipkow@17830
|
1215 |
lemma hd_in_set[simp]: "xs \<noteq> [] \<Longrightarrow> hd xs : set xs"
|
nipkow@17830
|
1216 |
by(cases xs) auto
|
oheimb@14099
|
1217 |
|
wenzelm@13142
|
1218 |
lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)"
|
nipkow@13145
|
1219 |
by auto
|
wenzelm@13114
|
1220 |
|
oheimb@14099
|
1221 |
lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs"
|
oheimb@14099
|
1222 |
by auto
|
oheimb@14099
|
1223 |
|
wenzelm@13142
|
1224 |
lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
|
nipkow@13145
|
1225 |
by (induct xs) auto
|
wenzelm@13114
|
1226 |
|
nipkow@15245
|
1227 |
lemma set_empty2[iff]: "({} = set xs) = (xs = [])"
|
nipkow@15245
|
1228 |
by(induct xs) auto
|
nipkow@15245
|
1229 |
|
wenzelm@13142
|
1230 |
lemma set_rev [simp]: "set (rev xs) = set xs"
|
nipkow@13145
|
1231 |
by (induct xs) auto
|
wenzelm@13114
|
1232 |
|
wenzelm@13142
|
1233 |
lemma set_map [simp]: "set (map f xs) = f`(set xs)"
|
nipkow@13145
|
1234 |
by (induct xs) auto
|
wenzelm@13114
|
1235 |
|
wenzelm@13142
|
1236 |
lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \<and> P x}"
|
nipkow@13145
|
1237 |
by (induct xs) auto
|
wenzelm@13114
|
1238 |
|
nipkow@32417
|
1239 |
lemma set_upt [simp]: "set[i..<j] = {i..<j}"
|
bulwahn@41463
|
1240 |
by (induct j) auto
|
wenzelm@13114
|
1241 |
|
wenzelm@13142
|
1242 |
|
wenzelm@25221
|
1243 |
lemma split_list: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs"
|
nipkow@18049
|
1244 |
proof (induct xs)
|
nipkow@26073
|
1245 |
case Nil thus ?case by simp
|
nipkow@26073
|
1246 |
next
|
nipkow@26073
|
1247 |
case Cons thus ?case by (auto intro: Cons_eq_appendI)
|
nipkow@26073
|
1248 |
qed
|
nipkow@26073
|
1249 |
|
haftmann@26734
|
1250 |
lemma in_set_conv_decomp: "x \<in> set xs \<longleftrightarrow> (\<exists>ys zs. xs = ys @ x # zs)"
|
haftmann@26734
|
1251 |
by (auto elim: split_list)
|
nipkow@26073
|
1252 |
|
nipkow@26073
|
1253 |
lemma split_list_first: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys"
|
nipkow@26073
|
1254 |
proof (induct xs)
|
nipkow@26073
|
1255 |
case Nil thus ?case by simp
|
nipkow@18049
|
1256 |
next
|
nipkow@18049
|
1257 |
case (Cons a xs)
|
nipkow@18049
|
1258 |
show ?case
|
nipkow@18049
|
1259 |
proof cases
|
nipkow@44890
|
1260 |
assume "x = a" thus ?case using Cons by fastforce
|
nipkow@18049
|
1261 |
next
|
nipkow@44890
|
1262 |
assume "x \<noteq> a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI)
|
nipkow@26073
|
1263 |
qed
|
nipkow@26073
|
1264 |
qed
|
nipkow@26073
|
1265 |
|
nipkow@26073
|
1266 |
lemma in_set_conv_decomp_first:
|
nipkow@26073
|
1267 |
"(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)"
|
haftmann@26734
|
1268 |
by (auto dest!: split_list_first)
|
nipkow@26073
|
1269 |
|
haftmann@40122
|
1270 |
lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
|
haftmann@40122
|
1271 |
proof (induct xs rule: rev_induct)
|
nipkow@26073
|
1272 |
case Nil thus ?case by simp
|
nipkow@26073
|
1273 |
next
|
nipkow@26073
|
1274 |
case (snoc a xs)
|
nipkow@26073
|
1275 |
show ?case
|
nipkow@26073
|
1276 |
proof cases
|
haftmann@40122
|
1277 |
assume "x = a" thus ?case using snoc by (metis List.set.simps(1) emptyE)
|
nipkow@26073
|
1278 |
next
|
nipkow@44890
|
1279 |
assume "x \<noteq> a" thus ?case using snoc by fastforce
|
nipkow@18049
|
1280 |
qed
|
nipkow@18049
|
1281 |
qed
|
nipkow@18049
|
1282 |
|
nipkow@26073
|
1283 |
lemma in_set_conv_decomp_last:
|
nipkow@26073
|
1284 |
"(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)"
|
haftmann@26734
|
1285 |
by (auto dest!: split_list_last)
|
nipkow@26073
|
1286 |
|
nipkow@26073
|
1287 |
lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs & P x"
|
nipkow@26073
|
1288 |
proof (induct xs)
|
nipkow@26073
|
1289 |
case Nil thus ?case by simp
|
nipkow@26073
|
1290 |
next
|
nipkow@26073
|
1291 |
case Cons thus ?case
|
nipkow@26073
|
1292 |
by(simp add:Bex_def)(metis append_Cons append.simps(1))
|
nipkow@26073
|
1293 |
qed
|
nipkow@26073
|
1294 |
|
nipkow@26073
|
1295 |
lemma split_list_propE:
|
haftmann@26734
|
1296 |
assumes "\<exists>x \<in> set xs. P x"
|
haftmann@26734
|
1297 |
obtains ys x zs where "xs = ys @ x # zs" and "P x"
|
haftmann@26734
|
1298 |
using split_list_prop [OF assms] by blast
|
nipkow@26073
|
1299 |
|
nipkow@26073
|
1300 |
lemma split_list_first_prop:
|
nipkow@26073
|
1301 |
"\<exists>x \<in> set xs. P x \<Longrightarrow>
|
nipkow@26073
|
1302 |
\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y)"
|
haftmann@26734
|
1303 |
proof (induct xs)
|
nipkow@26073
|
1304 |
case Nil thus ?case by simp
|
nipkow@26073
|
1305 |
next
|
nipkow@26073
|
1306 |
case (Cons x xs)
|
nipkow@26073
|
1307 |
show ?case
|
nipkow@26073
|
1308 |
proof cases
|
nipkow@26073
|
1309 |
assume "P x"
|
haftmann@40122
|
1310 |
thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append)
|
nipkow@26073
|
1311 |
next
|
nipkow@26073
|
1312 |
assume "\<not> P x"
|
nipkow@26073
|
1313 |
hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
|
nipkow@26073
|
1314 |
thus ?thesis using `\<not> P x` Cons(1) by (metis append_Cons set_ConsD)
|
nipkow@26073
|
1315 |
qed
|
nipkow@26073
|
1316 |
qed
|
nipkow@26073
|
1317 |
|
nipkow@26073
|
1318 |
lemma split_list_first_propE:
|
haftmann@26734
|
1319 |
assumes "\<exists>x \<in> set xs. P x"
|
haftmann@26734
|
1320 |
obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>y \<in> set ys. \<not> P y"
|
haftmann@26734
|
1321 |
using split_list_first_prop [OF assms] by blast
|
nipkow@26073
|
1322 |
|
nipkow@26073
|
1323 |
lemma split_list_first_prop_iff:
|
nipkow@26073
|
1324 |
"(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
|
nipkow@26073
|
1325 |
(\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y))"
|
haftmann@26734
|
1326 |
by (rule, erule split_list_first_prop) auto
|
nipkow@26073
|
1327 |
|
nipkow@26073
|
1328 |
lemma split_list_last_prop:
|
nipkow@26073
|
1329 |
"\<exists>x \<in> set xs. P x \<Longrightarrow>
|
nipkow@26073
|
1330 |
\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z)"
|
nipkow@26073
|
1331 |
proof(induct xs rule:rev_induct)
|
nipkow@26073
|
1332 |
case Nil thus ?case by simp
|
nipkow@26073
|
1333 |
next
|
nipkow@26073
|
1334 |
case (snoc x xs)
|
nipkow@26073
|
1335 |
show ?case
|
nipkow@26073
|
1336 |
proof cases
|
nipkow@26073
|
1337 |
assume "P x" thus ?thesis by (metis emptyE set_empty)
|
nipkow@26073
|
1338 |
next
|
nipkow@26073
|
1339 |
assume "\<not> P x"
|
nipkow@26073
|
1340 |
hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
|
nipkow@44890
|
1341 |
thus ?thesis using `\<not> P x` snoc(1) by fastforce
|
nipkow@26073
|
1342 |
qed
|
nipkow@26073
|
1343 |
qed
|
nipkow@26073
|
1344 |
|
nipkow@26073
|
1345 |
lemma split_list_last_propE:
|
haftmann@26734
|
1346 |
assumes "\<exists>x \<in> set xs. P x"
|
haftmann@26734
|
1347 |
obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>z \<in> set zs. \<not> P z"
|
haftmann@26734
|
1348 |
using split_list_last_prop [OF assms] by blast
|
nipkow@26073
|
1349 |
|
nipkow@26073
|
1350 |
lemma split_list_last_prop_iff:
|
nipkow@26073
|
1351 |
"(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
|
nipkow@26073
|
1352 |
(\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))"
|
haftmann@26734
|
1353 |
by (metis split_list_last_prop [where P=P] in_set_conv_decomp)
|
nipkow@26073
|
1354 |
|
nipkow@26073
|
1355 |
lemma finite_list: "finite A ==> EX xs. set xs = A"
|
haftmann@26734
|
1356 |
by (erule finite_induct)
|
haftmann@26734
|
1357 |
(auto simp add: set.simps(2) [symmetric] simp del: set.simps(2))
|
paulson@13508
|
1358 |
|
kleing@14388
|
1359 |
lemma card_length: "card (set xs) \<le> length xs"
|
kleing@14388
|
1360 |
by (induct xs) (auto simp add: card_insert_if)
|
wenzelm@13114
|
1361 |
|
haftmann@26442
|
1362 |
lemma set_minus_filter_out:
|
haftmann@26442
|
1363 |
"set xs - {y} = set (filter (\<lambda>x. \<not> (x = y)) xs)"
|
haftmann@26442
|
1364 |
by (induct xs) auto
|
paulson@15168
|
1365 |
|
wenzelm@35115
|
1366 |
|
haftmann@49948
|
1367 |
subsubsection {* @{const filter} *}
|
wenzelm@13114
|
1368 |
|
wenzelm@13142
|
1369 |
lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
|
nipkow@13145
|
1370 |
by (induct xs) auto
|
wenzelm@13114
|
1371 |
|
nipkow@15305
|
1372 |
lemma rev_filter: "rev (filter P xs) = filter P (rev xs)"
|
nipkow@15305
|
1373 |
by (induct xs) simp_all
|
nipkow@15305
|
1374 |
|
wenzelm@13142
|
1375 |
lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs"
|
nipkow@13145
|
1376 |
by (induct xs) auto
|
wenzelm@13114
|
1377 |
|
nipkow@16998
|
1378 |
lemma length_filter_le [simp]: "length (filter P xs) \<le> length xs"
|
nipkow@16998
|
1379 |
by (induct xs) (auto simp add: le_SucI)
|
nipkow@16998
|
1380 |
|
nipkow@18423
|
1381 |
lemma sum_length_filter_compl:
|
nipkow@18423
|
1382 |
"length(filter P xs) + length(filter (%x. ~P x) xs) = length xs"
|
nipkow@18423
|
1383 |
by(induct xs) simp_all
|
nipkow@18423
|
1384 |
|
wenzelm@13142
|
1385 |
lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
|
nipkow@13145
|
1386 |
by (induct xs) auto
|
wenzelm@13114
|
1387 |
|
wenzelm@13142
|
1388 |
lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
|
nipkow@13145
|
1389 |
by (induct xs) auto
|
wenzelm@13114
|
1390 |
|
nipkow@16998
|
1391 |
lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)"
|
nipkow@24349
|
1392 |
by (induct xs) simp_all
|
nipkow@16998
|
1393 |
|
nipkow@16998
|
1394 |
lemma filter_id_conv: "(filter P xs = xs) = (\<forall>x\<in>set xs. P x)"
|
nipkow@16998
|
1395 |
apply (induct xs)
|
nipkow@16998
|
1396 |
apply auto
|
nipkow@16998
|
1397 |
apply(cut_tac P=P and xs=xs in length_filter_le)
|
nipkow@16998
|
1398 |
apply simp
|
nipkow@16998
|
1399 |
done
|
wenzelm@13114
|
1400 |
|
nipkow@16965
|
1401 |
lemma filter_map:
|
nipkow@16965
|
1402 |
"filter P (map f xs) = map f (filter (P o f) xs)"
|
nipkow@16965
|
1403 |
by (induct xs) simp_all
|
nipkow@16965
|
1404 |
|
nipkow@16965
|
1405 |
lemma length_filter_map[simp]:
|
nipkow@16965
|
1406 |
"length (filter P (map f xs)) = length(filter (P o f) xs)"
|
nipkow@16965
|
1407 |
by (simp add:filter_map)
|
nipkow@16965
|
1408 |
|
wenzelm@13142
|
1409 |
lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
|
nipkow@13145
|
1410 |
by auto
|
wenzelm@13114
|
1411 |
|
nipkow@15246
|
1412 |
lemma length_filter_less:
|
nipkow@15246
|
1413 |
"\<lbrakk> x : set xs; ~ P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
|
nipkow@15246
|
1414 |
proof (induct xs)
|
nipkow@15246
|
1415 |
case Nil thus ?case by simp
|
nipkow@15246
|
1416 |
next
|
nipkow@15246
|
1417 |
case (Cons x xs) thus ?case
|
nipkow@15246
|
1418 |
apply (auto split:split_if_asm)
|
nipkow@15246
|
1419 |
using length_filter_le[of P xs] apply arith
|
nipkow@15246
|
1420 |
done
|
nipkow@15246
|
1421 |
qed
|
wenzelm@13114
|
1422 |
|
nipkow@15281
|
1423 |
lemma length_filter_conv_card:
|
nipkow@15281
|
1424 |
"length(filter p xs) = card{i. i < length xs & p(xs!i)}"
|
nipkow@15281
|
1425 |
proof (induct xs)
|
nipkow@15281
|
1426 |
case Nil thus ?case by simp
|
nipkow@15281
|
1427 |
next
|
nipkow@15281
|
1428 |
case (Cons x xs)
|
nipkow@15281
|
1429 |
let ?S = "{i. i < length xs & p(xs!i)}"
|
nipkow@15281
|
1430 |
have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite)
|
nipkow@15281
|
1431 |
show ?case (is "?l = card ?S'")
|
nipkow@15281
|
1432 |
proof (cases)
|
nipkow@15281
|
1433 |
assume "p x"
|
nipkow@15281
|
1434 |
hence eq: "?S' = insert 0 (Suc ` ?S)"
|
nipkow@25162
|
1435 |
by(auto simp: image_def split:nat.split dest:gr0_implies_Suc)
|
nipkow@15281
|
1436 |
have "length (filter p (x # xs)) = Suc(card ?S)"
|
wenzelm@23388
|
1437 |
using Cons `p x` by simp
|
nipkow@15281
|
1438 |
also have "\<dots> = Suc(card(Suc ` ?S))" using fin
|
huffman@44921
|
1439 |
by (simp add: card_image)
|
nipkow@15281
|
1440 |
also have "\<dots> = card ?S'" using eq fin
|
nipkow@15281
|
1441 |
by (simp add:card_insert_if) (simp add:image_def)
|
nipkow@15281
|
1442 |
finally show ?thesis .
|
nipkow@15281
|
1443 |
next
|
nipkow@15281
|
1444 |
assume "\<not> p x"
|
nipkow@15281
|
1445 |
hence eq: "?S' = Suc ` ?S"
|
nipkow@25162
|
1446 |
by(auto simp add: image_def split:nat.split elim:lessE)
|
nipkow@15281
|
1447 |
have "length (filter p (x # xs)) = card ?S"
|
wenzelm@23388
|
1448 |
using Cons `\<not> p x` by simp
|
nipkow@15281
|
1449 |
also have "\<dots> = card(Suc ` ?S)" using fin
|
huffman@44921
|
1450 |
by (simp add: card_image)
|
nipkow@15281
|
1451 |
also have "\<dots> = card ?S'" using eq fin
|
nipkow@15281
|
1452 |
by (simp add:card_insert_if)
|
nipkow@15281
|
1453 |
finally show ?thesis .
|
nipkow@15281
|
1454 |
qed
|
nipkow@15281
|
1455 |
qed
|
nipkow@15281
|
1456 |
|
nipkow@17629
|
1457 |
lemma Cons_eq_filterD:
|
nipkow@17629
|
1458 |
"x#xs = filter P ys \<Longrightarrow>
|
nipkow@17629
|
1459 |
\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
|
wenzelm@19585
|
1460 |
(is "_ \<Longrightarrow> \<exists>us vs. ?P ys us vs")
|
nipkow@17629
|
1461 |
proof(induct ys)
|
nipkow@17629
|
1462 |
case Nil thus ?case by simp
|
nipkow@17629
|
1463 |
next
|
nipkow@17629
|
1464 |
case (Cons y ys)
|
nipkow@17629
|
1465 |
show ?case (is "\<exists>x. ?Q x")
|
nipkow@17629
|
1466 |
proof cases
|
nipkow@17629
|
1467 |
assume Py: "P y"
|
nipkow@17629
|
1468 |
show ?thesis
|
nipkow@17629
|
1469 |
proof cases
|
wenzelm@25221
|
1470 |
assume "x = y"
|
wenzelm@25221
|
1471 |
with Py Cons.prems have "?Q []" by simp
|
wenzelm@25221
|
1472 |
then show ?thesis ..
|
nipkow@17629
|
1473 |
next
|
wenzelm@25221
|
1474 |
assume "x \<noteq> y"
|
wenzelm@25221
|
1475 |
with Py Cons.prems show ?thesis by simp
|
nipkow@17629
|
1476 |
qed
|
nipkow@17629
|
1477 |
next
|
wenzelm@25221
|
1478 |
assume "\<not> P y"
|
nipkow@44890
|
1479 |
with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce
|
wenzelm@25221
|
1480 |
then have "?Q (y#us)" by simp
|
wenzelm@25221
|
1481 |
then show ?thesis ..
|
nipkow@17629
|
1482 |
qed
|
nipkow@17629
|
1483 |
qed
|
nipkow@17629
|
1484 |
|
nipkow@17629
|
1485 |
lemma filter_eq_ConsD:
|
nipkow@17629
|
1486 |
"filter P ys = x#xs \<Longrightarrow>
|
nipkow@17629
|
1487 |
\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
|
nipkow@17629
|
1488 |
by(rule Cons_eq_filterD) simp
|
nipkow@17629
|
1489 |
|
nipkow@17629
|
1490 |
lemma filter_eq_Cons_iff:
|
nipkow@17629
|
1491 |
"(filter P ys = x#xs) =
|
nipkow@17629
|
1492 |
(\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
|
nipkow@17629
|
1493 |
by(auto dest:filter_eq_ConsD)
|
nipkow@17629
|
1494 |
|
nipkow@17629
|
1495 |
lemma Cons_eq_filter_iff:
|
nipkow@17629
|
1496 |
"(x#xs = filter P ys) =
|
nipkow@17629
|
1497 |
(\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
|
nipkow@17629
|
1498 |
by(auto dest:Cons_eq_filterD)
|
nipkow@17629
|
1499 |
|
krauss@44013
|
1500 |
lemma filter_cong[fundef_cong]:
|
nipkow@17501
|
1501 |
"xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
|
nipkow@17501
|
1502 |
apply simp
|
nipkow@17501
|
1503 |
apply(erule thin_rl)
|
nipkow@17501
|
1504 |
by (induct ys) simp_all
|
nipkow@17501
|
1505 |
|
nipkow@15281
|
1506 |
|
haftmann@26442
|
1507 |
subsubsection {* List partitioning *}
|
haftmann@26442
|
1508 |
|
haftmann@26442
|
1509 |
primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where
|
nipkow@50548
|
1510 |
"partition P [] = ([], [])" |
|
nipkow@50548
|
1511 |
"partition P (x # xs) =
|
nipkow@50548
|
1512 |
(let (yes, no) = partition P xs
|
nipkow@50548
|
1513 |
in if P x then (x # yes, no) else (yes, x # no))"
|
haftmann@26442
|
1514 |
|
haftmann@26442
|
1515 |
lemma partition_filter1:
|
haftmann@26442
|
1516 |
"fst (partition P xs) = filter P xs"
|
haftmann@26442
|
1517 |
by (induct xs) (auto simp add: Let_def split_def)
|
haftmann@26442
|
1518 |
|
haftmann@26442
|
1519 |
lemma partition_filter2:
|
haftmann@26442
|
1520 |
"snd (partition P xs) = filter (Not o P) xs"
|
haftmann@26442
|
1521 |
by (induct xs) (auto simp add: Let_def split_def)
|
haftmann@26442
|
1522 |
|
haftmann@26442
|
1523 |
lemma partition_P:
|
haftmann@26442
|
1524 |
assumes "partition P xs = (yes, no)"
|
haftmann@26442
|
1525 |
shows "(\<forall>p \<in> set yes. P p) \<and> (\<forall>p \<in> set no. \<not> P p)"
|
haftmann@26442
|
1526 |
proof -
|
haftmann@26442
|
1527 |
from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
|
haftmann@26442
|
1528 |
by simp_all
|
haftmann@26442
|
1529 |
then show ?thesis by (simp_all add: partition_filter1 partition_filter2)
|
haftmann@26442
|
1530 |
qed
|
haftmann@26442
|
1531 |
|
haftmann@26442
|
1532 |
lemma partition_set:
|
haftmann@26442
|
1533 |
assumes "partition P xs = (yes, no)"
|
haftmann@26442
|
1534 |
shows "set yes \<union> set no = set xs"
|
haftmann@26442
|
1535 |
proof -
|
haftmann@26442
|
1536 |
from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
|
haftmann@26442
|
1537 |
by simp_all
|
haftmann@26442
|
1538 |
then show ?thesis by (auto simp add: partition_filter1 partition_filter2)
|
haftmann@26442
|
1539 |
qed
|
haftmann@26442
|
1540 |
|
hoelzl@33639
|
1541 |
lemma partition_filter_conv[simp]:
|
hoelzl@33639
|
1542 |
"partition f xs = (filter f xs,filter (Not o f) xs)"
|
hoelzl@33639
|
1543 |
unfolding partition_filter2[symmetric]
|
hoelzl@33639
|
1544 |
unfolding partition_filter1[symmetric] by simp
|
hoelzl@33639
|
1545 |
|
hoelzl@33639
|
1546 |
declare partition.simps[simp del]
|
haftmann@26442
|
1547 |
|
wenzelm@35115
|
1548 |
|
haftmann@49948
|
1549 |
subsubsection {* @{const concat} *}
|
wenzelm@13114
|
1550 |
|
wenzelm@13142
|
1551 |
lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
|
nipkow@13145
|
1552 |
by (induct xs) auto
|
wenzelm@13114
|
1553 |
|
paulson@18447
|
1554 |
lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])"
|
nipkow@13145
|
1555 |
by (induct xss) auto
|
wenzelm@13114
|
1556 |
|
paulson@18447
|
1557 |
lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\<forall>xs \<in> set xss. xs = [])"
|
nipkow@13145
|
1558 |
by (induct xss) auto
|
wenzelm@13114
|
1559 |
|
nipkow@24308
|
1560 |
lemma set_concat [simp]: "set (concat xs) = (UN x:set xs. set x)"
|
nipkow@13145
|
1561 |
by (induct xs) auto
|
wenzelm@13114
|
1562 |
|
nipkow@24476
|
1563 |
lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs"
|
nipkow@24349
|
1564 |
by (induct xs) auto
|
nipkow@24349
|
1565 |
|
wenzelm@13142
|
1566 |
lemma map_concat: "map f (concat xs) = concat (map (map f) xs)"
|
nipkow@13145
|
1567 |
by (induct xs) auto
|
wenzelm@13114
|
1568 |
|
wenzelm@13142
|
1569 |
lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
|
nipkow@13145
|
1570 |
by (induct xs) auto
|
wenzelm@13114
|
1571 |
|
wenzelm@13142
|
1572 |
lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
|
nipkow@13145
|
1573 |
by (induct xs) auto
|
wenzelm@13114
|
1574 |
|
bulwahn@40365
|
1575 |
lemma concat_eq_concat_iff: "\<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> length xs = length ys ==> (concat xs = concat ys) = (xs = ys)"
|
bulwahn@40365
|
1576 |
proof (induct xs arbitrary: ys)
|
bulwahn@40365
|
1577 |
case (Cons x xs ys)
|
bulwahn@40365
|
1578 |
thus ?case by (cases ys) auto
|
bulwahn@40365
|
1579 |
qed (auto)
|
bulwahn@40365
|
1580 |
|
bulwahn@40365
|
1581 |
lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> xs = ys"
|
bulwahn@40365
|
1582 |
by (simp add: concat_eq_concat_iff)
|
bulwahn@40365
|
1583 |
|
wenzelm@13114
|
1584 |
|
haftmann@49948
|
1585 |
subsubsection {* @{const nth} *}
|
wenzelm@13114
|
1586 |
|
haftmann@29827
|
1587 |
lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x"
|
nipkow@13145
|
1588 |
by auto
|
wenzelm@13114
|
1589 |
|
haftmann@29827
|
1590 |
lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n"
|
nipkow@13145
|
1591 |
by auto
|
wenzelm@13114
|
1592 |
|
wenzelm@13142
|
1593 |
declare nth.simps [simp del]
|
wenzelm@13114
|
1594 |
|
nipkow@41842
|
1595 |
lemma nth_Cons_pos[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
|
nipkow@41842
|
1596 |
by(auto simp: Nat.gr0_conv_Suc)
|
nipkow@41842
|
1597 |
|
wenzelm@13114
|
1598 |
lemma nth_append:
|
nipkow@24526
|
1599 |
"(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"
|
nipkow@24526
|
1600 |
apply (induct xs arbitrary: n, simp)
|
paulson@14208
|
1601 |
apply (case_tac n, auto)
|
nipkow@13145
|
1602 |
done
|
wenzelm@13114
|
1603 |
|
nipkow@14402
|
1604 |
lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x"
|
wenzelm@25221
|
1605 |
by (induct xs) auto
|
nipkow@14402
|
1606 |
|
nipkow@14402
|
1607 |
lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
|
wenzelm@25221
|
1608 |
by (induct xs) auto
|
nipkow@14402
|
1609 |
|
nipkow@24526
|
1610 |
lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)"
|
nipkow@24526
|
1611 |
apply (induct xs arbitrary: n, simp)
|
paulson@14208
|
1612 |
apply (case_tac n, auto)
|
nipkow@13145
|
1613 |
done
|
wenzelm@13114
|
1614 |
|
noschinl@45841
|
1615 |
lemma nth_tl:
|
noschinl@45841
|
1616 |
assumes "n < length (tl x)" shows "tl x ! n = x ! Suc n"
|
noschinl@45841
|
1617 |
using assms by (induct x) auto
|
noschinl@45841
|
1618 |
|
nipkow@18423
|
1619 |
lemma hd_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd xs = xs!0"
|
nipkow@18423
|
1620 |
by(cases xs) simp_all
|
nipkow@18423
|
1621 |
|
nipkow@18049
|
1622 |
|
nipkow@18049
|
1623 |
lemma list_eq_iff_nth_eq:
|
nipkow@24526
|
1624 |
"(xs = ys) = (length xs = length ys \<and> (ALL i<length xs. xs!i = ys!i))"
|
nipkow@24526
|
1625 |
apply(induct xs arbitrary: ys)
|
paulson@24632
|
1626 |
apply force
|
nipkow@18049
|
1627 |
apply(case_tac ys)
|
nipkow@18049
|
1628 |
apply simp
|
nipkow@18049
|
1629 |
apply(simp add:nth_Cons split:nat.split)apply blast
|
nipkow@18049
|
1630 |
done
|
nipkow@18049
|
1631 |
|
wenzelm@13142
|
1632 |
lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
|
paulson@15251
|
1633 |
apply (induct xs, simp, simp)
|
nipkow@13145
|
1634 |
apply safe
|
paulson@24632
|
1635 |
apply (metis nat_case_0 nth.simps zero_less_Suc)
|
paulson@24632
|
1636 |
apply (metis less_Suc_eq_0_disj nth_Cons_Suc)
|
paulson@14208
|
1637 |
apply (case_tac i, simp)
|
paulson@24632
|
1638 |
apply (metis diff_Suc_Suc nat_case_Suc nth.simps zero_less_diff)
|
nipkow@13145
|
1639 |
done
|
wenzelm@13114
|
1640 |
|
nipkow@17501
|
1641 |
lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)"
|
nipkow@17501
|
1642 |
by(auto simp:set_conv_nth)
|
nipkow@17501
|
1643 |
|
haftmann@51160
|
1644 |
lemma nth_equal_first_eq:
|
haftmann@51160
|
1645 |
assumes "x \<notin> set xs"
|
haftmann@51160
|
1646 |
assumes "n \<le> length xs"
|
haftmann@51160
|
1647 |
shows "(x # xs) ! n = x \<longleftrightarrow> n = 0" (is "?lhs \<longleftrightarrow> ?rhs")
|
haftmann@51160
|
1648 |
proof
|
haftmann@51160
|
1649 |
assume ?lhs
|
haftmann@51160
|
1650 |
show ?rhs
|
haftmann@51160
|
1651 |
proof (rule ccontr)
|
haftmann@51160
|
1652 |
assume "n \<noteq> 0"
|
haftmann@51160
|
1653 |
then have "n > 0" by simp
|
haftmann@51160
|
1654 |
with `?lhs` have "xs ! (n - 1) = x" by simp
|
haftmann@51160
|
1655 |
moreover from `n > 0` `n \<le> length xs` have "n - 1 < length xs" by simp
|
haftmann@51160
|
1656 |
ultimately have "\<exists>i<length xs. xs ! i = x" by auto
|
haftmann@51160
|
1657 |
with `x \<notin> set xs` in_set_conv_nth [of x xs] show False by simp
|
haftmann@51160
|
1658 |
qed
|
haftmann@51160
|
1659 |
next
|
haftmann@51160
|
1660 |
assume ?rhs then show ?lhs by simp
|
haftmann@51160
|
1661 |
qed
|
haftmann@51160
|
1662 |
|
haftmann@51160
|
1663 |
lemma nth_non_equal_first_eq:
|
haftmann@51160
|
1664 |
assumes "x \<noteq> y"
|
haftmann@51160
|
1665 |
shows "(x # xs) ! n = y \<longleftrightarrow> xs ! (n - 1) = y \<and> n > 0" (is "?lhs \<longleftrightarrow> ?rhs")
|
haftmann@51160
|
1666 |
proof
|
haftmann@51160
|
1667 |
assume "?lhs" with assms have "n > 0" by (cases n) simp_all
|
haftmann@51160
|
1668 |
with `?lhs` show ?rhs by simp
|
haftmann@51160
|
1669 |
next
|
haftmann@51160
|
1670 |
assume "?rhs" then show "?lhs" by simp
|
haftmann@51160
|
1671 |
qed
|
haftmann@51160
|
1672 |
|
nipkow@13145
|
1673 |
lemma list_ball_nth: "[| n < length xs; !x : set xs. P x|] ==> P(xs!n)"
|
nipkow@13145
|
1674 |
by (auto simp add: set_conv_nth)
|
wenzelm@13114
|
1675 |
|
wenzelm@13142
|
1676 |
lemma nth_mem [simp]: "n < length xs ==> xs!n : set xs"
|
nipkow@13145
|
1677 |
by (auto simp add: set_conv_nth)
|
wenzelm@13114
|
1678 |
|
wenzelm@13114
|
1679 |
lemma all_nth_imp_all_set:
|
nipkow@13145
|
1680 |
"[| !i < length xs. P(xs!i); x : set xs|] ==> P x"
|
nipkow@13145
|
1681 |
by (auto simp add: set_conv_nth)
|
wenzelm@13114
|
1682 |
|
wenzelm@13114
|
1683 |
lemma all_set_conv_all_nth:
|
nipkow@13145
|
1684 |
"(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs --> P (xs ! i))"
|
nipkow@13145
|
1685 |
by (auto simp add: set_conv_nth)
|
wenzelm@13114
|
1686 |
|
kleing@25296
|
1687 |
lemma rev_nth:
|
kleing@25296
|
1688 |
"n < size xs \<Longrightarrow> rev xs ! n = xs ! (length xs - Suc n)"
|
kleing@25296
|
1689 |
proof (induct xs arbitrary: n)
|
kleing@25296
|
1690 |
case Nil thus ?case by simp
|
kleing@25296
|
1691 |
next
|
kleing@25296
|
1692 |
case (Cons x xs)
|
kleing@25296
|
1693 |
hence n: "n < Suc (length xs)" by simp
|
kleing@25296
|
1694 |
moreover
|
kleing@25296
|
1695 |
{ assume "n < length xs"
|
wenzelm@53374
|
1696 |
with n obtain n' where n': "length xs - n = Suc n'"
|
kleing@25296
|
1697 |
by (cases "length xs - n", auto)
|
kleing@25296
|
1698 |
moreover
|
wenzelm@53374
|
1699 |
from n' have "length xs - Suc n = n'" by simp
|
kleing@25296
|
1700 |
ultimately
|
kleing@25296
|
1701 |
have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp
|
kleing@25296
|
1702 |
}
|
kleing@25296
|
1703 |
ultimately
|
kleing@25296
|
1704 |
show ?case by (clarsimp simp add: Cons nth_append)
|
kleing@25296
|
1705 |
qed
|
wenzelm@13114
|
1706 |
|
nipkow@31159
|
1707 |
lemma Skolem_list_nth:
|
nipkow@31159
|
1708 |
"(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
|
nipkow@31159
|
1709 |
(is "_ = (EX xs. ?P k xs)")
|
nipkow@31159
|
1710 |
proof(induct k)
|
nipkow@31159
|
1711 |
case 0 show ?case by simp
|
nipkow@31159
|
1712 |
next
|
nipkow@31159
|
1713 |
case (Suc k)
|
nipkow@31159
|
1714 |
show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
|
nipkow@31159
|
1715 |
proof
|
nipkow@31159
|
1716 |
assume "?R" thus "?L" using Suc by auto
|
nipkow@31159
|
1717 |
next
|
nipkow@31159
|
1718 |
assume "?L"
|
nipkow@31159
|
1719 |
with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
|
nipkow@31159
|
1720 |
hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
|
nipkow@31159
|
1721 |
thus "?R" ..
|
nipkow@31159
|
1722 |
qed
|
nipkow@31159
|
1723 |
qed
|
nipkow@31159
|
1724 |
|
nipkow@31159
|
1725 |
|
haftmann@49948
|
1726 |
subsubsection {* @{const list_update} *}
|
wenzelm@13114
|
1727 |
|
nipkow@24526
|
1728 |
lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
|
nipkow@24526
|
1729 |
by (induct xs arbitrary: i) (auto split: nat.split)
|
wenzelm@13114
|
1730 |
|
wenzelm@13114
|
1731 |
lemma nth_list_update:
|
nipkow@24526
|
1732 |
"i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
|
nipkow@24526
|
1733 |
by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
|
wenzelm@13114
|
1734 |
|
wenzelm@13142
|
1735 |
lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
|
nipkow@13145
|
1736 |
by (simp add: nth_list_update)
|
wenzelm@13114
|
1737 |
|
nipkow@24526
|
1738 |
lemma nth_list_update_neq [simp]: "i \<noteq> j ==> xs[i:=x]!j = xs!j"
|
nipkow@24526
|
1739 |
by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
|
wenzelm@13114
|
1740 |
|
nipkow@24526
|
1741 |
lemma list_update_id[simp]: "xs[i := xs!i] = xs"
|
nipkow@24526
|
1742 |
by (induct xs arbitrary: i) (simp_all split:nat.splits)
|
nipkow@24526
|
1743 |
|
nipkow@24526
|
1744 |
lemma list_update_beyond[simp]: "length xs \<le> i \<Longrightarrow> xs[i:=x] = xs"
|
nipkow@24526
|
1745 |
apply (induct xs arbitrary: i)
|
nipkow@17501
|
1746 |
apply simp
|
nipkow@17501
|
1747 |
apply (case_tac i)
|
nipkow@17501
|
1748 |
apply simp_all
|
nipkow@17501
|
1749 |
done
|
nipkow@17501
|
1750 |
|
nipkow@31077
|
1751 |
lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]"
|
nipkow@31077
|
1752 |
by(metis length_0_conv length_list_update)
|
nipkow@31077
|
1753 |
|
wenzelm@13114
|
1754 |
lemma list_update_same_conv:
|
nipkow@24526
|
1755 |
"i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
|
nipkow@24526
|
1756 |
by (induct xs arbitrary: i) (auto split: nat.split)
|
wenzelm@13114
|
1757 |
|
nipkow@14187
|
1758 |
lemma list_update_append1:
|
nipkow@24526
|
1759 |
"i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys"
|
nipkow@24526
|
1760 |
apply (induct xs arbitrary: i, simp)
|
nipkow@14187
|
1761 |
apply(simp split:nat.split)
|
nipkow@14187
|
1762 |
done
|
nipkow@14187
|
1763 |
|
kleing@15868
|
1764 |
lemma list_update_append:
|
nipkow@24526
|
1765 |
"(xs @ ys) [n:= x] =
|
kleing@15868
|
1766 |
(if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))"
|
nipkow@24526
|
1767 |
by (induct xs arbitrary: n) (auto split:nat.splits)
|
kleing@15868
|
1768 |
|
nipkow@14402
|
1769 |
lemma list_update_length [simp]:
|
nipkow@14402
|
1770 |
"(xs @ x # ys)[length xs := y] = (xs @ y # ys)"
|
nipkow@14402
|
1771 |
by (induct xs, auto)
|
nipkow@14402
|
1772 |
|
nipkow@31264
|
1773 |
lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]"
|
nipkow@31264
|
1774 |
by(induct xs arbitrary: k)(auto split:nat.splits)
|
nipkow@31264
|
1775 |
|
nipkow@31264
|
1776 |
lemma rev_update:
|
nipkow@31264
|
1777 |
"k < length xs \<Longrightarrow> rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]"
|
nipkow@31264
|
1778 |
by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits)
|
nipkow@31264
|
1779 |
|
wenzelm@13114
|
1780 |
lemma update_zip:
|
nipkow@31080
|
1781 |
"(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"
|
nipkow@24526
|
1782 |
by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split)
|
nipkow@24526
|
1783 |
|
nipkow@24526
|
1784 |
lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)"
|
nipkow@24526
|
1785 |
by (induct xs arbitrary: i) (auto split: nat.split)
|
wenzelm@13114
|
1786 |
|
wenzelm@13114
|
1787 |
lemma set_update_subsetI: "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A"
|
nipkow@13145
|
1788 |
by (blast dest!: set_update_subset_insert [THEN subsetD])
|
wenzelm@13114
|
1789 |
|
nipkow@24526
|
1790 |
lemma set_update_memI: "n < length xs \<Longrightarrow> x \<in> set (xs[n := x])"
|
nipkow@24526
|
1791 |
by (induct xs arbitrary: n) (auto split:nat.splits)
|
kleing@15868
|
1792 |
|
nipkow@31077
|
1793 |
lemma list_update_overwrite[simp]:
|
haftmann@24796
|
1794 |
"xs [i := x, i := y] = xs [i := y]"
|
nipkow@31077
|
1795 |
apply (induct xs arbitrary: i) apply simp
|
nipkow@31077
|
1796 |
apply (case_tac i, simp_all)
|
haftmann@24796
|
1797 |
done
|
haftmann@24796
|
1798 |
|
haftmann@24796
|
1799 |
lemma list_update_swap:
|
haftmann@24796
|
1800 |
"i \<noteq> i' \<Longrightarrow> xs [i := x, i' := x'] = xs [i' := x', i := x]"
|
haftmann@24796
|
1801 |
apply (induct xs arbitrary: i i')
|
haftmann@24796
|
1802 |
apply simp
|
haftmann@24796
|
1803 |
apply (case_tac i, case_tac i')
|
haftmann@24796
|
1804 |
apply auto
|
haftmann@24796
|
1805 |
apply (case_tac i')
|
haftmann@24796
|
1806 |
apply auto
|
haftmann@24796
|
1807 |
done
|
haftmann@24796
|
1808 |
|
haftmann@29827
|
1809 |
lemma list_update_code [code]:
|
haftmann@29827
|
1810 |
"[][i := y] = []"
|
haftmann@29827
|
1811 |
"(x # xs)[0 := y] = y # xs"
|
haftmann@29827
|
1812 |
"(x # xs)[Suc i := y] = x # xs[i := y]"
|
haftmann@29827
|
1813 |
by simp_all
|
haftmann@29827
|
1814 |
|
wenzelm@13114
|
1815 |
|
haftmann@49948
|
1816 |
subsubsection {* @{const last} and @{const butlast} *}
|
wenzelm@13114
|
1817 |
|
wenzelm@13142
|
1818 |
lemma last_snoc [simp]: "last (xs @ [x]) = x"
|
nipkow@13145
|
1819 |
by (induct xs) auto
|
wenzelm@13114
|
1820 |
|
wenzelm@13142
|
1821 |
lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs"
|
nipkow@13145
|
1822 |
by (induct xs) auto
|
wenzelm@13114
|
1823 |
|
nipkow@14302
|
1824 |
lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x"
|
huffman@44921
|
1825 |
by simp
|
nipkow@14302
|
1826 |
|
nipkow@14302
|
1827 |
lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs"
|
huffman@44921
|
1828 |
by simp
|
nipkow@14302
|
1829 |
|
nipkow@14302
|
1830 |
lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)"
|
nipkow@14302
|
1831 |
by (induct xs) (auto)
|
nipkow@14302
|
1832 |
|
nipkow@14302
|
1833 |
lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs"
|
nipkow@14302
|
1834 |
by(simp add:last_append)
|
nipkow@14302
|
1835 |
|
nipkow@14302
|
1836 |
lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys"
|
nipkow@14302
|
1837 |
by(simp add:last_append)
|
nipkow@14302
|
1838 |
|
noschinl@45841
|
1839 |
lemma last_tl: "xs = [] \<or> tl xs \<noteq> [] \<Longrightarrow>last (tl xs) = last xs"
|
noschinl@45841
|
1840 |
by (induct xs) simp_all
|
noschinl@45841
|
1841 |
|
noschinl@45841
|
1842 |
lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)"
|
noschinl@45841
|
1843 |
by (induct xs) simp_all
|
noschinl@45841
|
1844 |
|
nipkow@17762
|
1845 |
lemma hd_rev: "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"
|
nipkow@17762
|
1846 |
by(rule rev_exhaust[of xs]) simp_all
|
nipkow@17762
|
1847 |
|
nipkow@17762
|
1848 |
lemma last_rev: "xs \<noteq> [] \<Longrightarrow> last(rev xs) = hd xs"
|
nipkow@17762
|
1849 |
by(cases xs) simp_all
|
nipkow@17762
|
1850 |
|
nipkow@17765
|
1851 |
lemma last_in_set[simp]: "as \<noteq> [] \<Longrightarrow> last as \<in> set as"
|
nipkow@17765
|
1852 |
by (induct as) auto
|
nipkow@17762
|
1853 |
|
wenzelm@13142
|
1854 |
lemma length_butlast [simp]: "length (butlast xs) = length xs - 1"
|
nipkow@13145
|
1855 |
by (induct xs rule: rev_induct) auto
|
wenzelm@13114
|
1856 |
|
wenzelm@13114
|
1857 |
lemma butlast_append:
|
nipkow@24526
|
1858 |
"butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)"
|
nipkow@24526
|
1859 |
by (induct xs arbitrary: ys) auto
|
wenzelm@13114
|
1860 |
|
wenzelm@13142
|
1861 |
lemma append_butlast_last_id [simp]:
|
nipkow@13145
|
1862 |
"xs \<noteq> [] ==> butlast xs @ [last xs] = xs"
|
nipkow@13145
|
1863 |
by (induct xs) auto
|
wenzelm@13114
|
1864 |
|
wenzelm@13142
|
1865 |
lemma in_set_butlastD: "x : set (butlast xs) ==> x : set xs"
|
nipkow@13145
|
1866 |
by (induct xs) (auto split: split_if_asm)
|
wenzelm@13114
|
1867 |
|
wenzelm@13114
|
1868 |
lemma in_set_butlast_appendI:
|
nipkow@13145
|
1869 |
"x : set (butlast xs) | x : set (butlast ys) ==> x : set (butlast (xs @ ys))"
|
nipkow@13145
|
1870 |
by (auto dest: in_set_butlastD simp add: butlast_append)
|
wenzelm@13114
|
1871 |
|
nipkow@24526
|
1872 |
lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs"
|
nipkow@24526
|
1873 |
apply (induct xs arbitrary: n)
|
nipkow@17501
|
1874 |
apply simp
|
nipkow@17501
|
1875 |
apply (auto split:nat.split)
|
nipkow@17501
|
1876 |
done
|
nipkow@17501
|
1877 |
|
noschinl@45841
|
1878 |
lemma nth_butlast:
|
noschinl@45841
|
1879 |
assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n"
|
noschinl@45841
|
1880 |
proof (cases xs)
|
noschinl@45841
|
1881 |
case (Cons y ys)
|
noschinl@45841
|
1882 |
moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n"
|
noschinl@45841
|
1883 |
by (simp add: nth_append)
|
noschinl@45841
|
1884 |
ultimately show ?thesis using append_butlast_last_id by simp
|
noschinl@45841
|
1885 |
qed simp
|
noschinl@45841
|
1886 |
|
huffman@30128
|
1887 |
lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)"
|
nipkow@17589
|
1888 |
by(induct xs)(auto simp:neq_Nil_conv)
|
nipkow@17589
|
1889 |
|
huffman@30128
|
1890 |
lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs"
|
huffman@26584
|
1891 |
by (induct xs, simp, case_tac xs, simp_all)
|
huffman@26584
|
1892 |
|
nipkow@31077
|
1893 |
lemma last_list_update:
|
nipkow@31077
|
1894 |
"xs \<noteq> [] \<Longrightarrow> last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)"
|
nipkow@31077
|
1895 |
by (auto simp: last_conv_nth)
|
nipkow@31077
|
1896 |
|
nipkow@31077
|
1897 |
lemma butlast_list_update:
|
nipkow@31077
|
1898 |
"butlast(xs[k:=x]) =
|
nipkow@31077
|
1899 |
(if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])"
|
nipkow@31077
|
1900 |
apply(cases xs rule:rev_cases)
|
nipkow@31077
|
1901 |
apply simp
|
nipkow@31077
|
1902 |
apply(simp add:list_update_append split:nat.splits)
|
nipkow@31077
|
1903 |
done
|
nipkow@31077
|
1904 |
|
haftmann@36851
|
1905 |
lemma last_map:
|
haftmann@36851
|
1906 |
"xs \<noteq> [] \<Longrightarrow> last (map f xs) = f (last xs)"
|
haftmann@36851
|
1907 |
by (cases xs rule: rev_cases) simp_all
|
haftmann@36851
|
1908 |
|
haftmann@36851
|
1909 |
lemma map_butlast:
|
haftmann@36851
|
1910 |
"map f (butlast xs) = butlast (map f xs)"
|
haftmann@36851
|
1911 |
by (induct xs) simp_all
|
haftmann@36851
|
1912 |
|
nipkow@40230
|
1913 |
lemma snoc_eq_iff_butlast:
|
nipkow@40230
|
1914 |
"xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
|
nipkow@40230
|
1915 |
by (metis append_butlast_last_id append_is_Nil_conv butlast_snoc last_snoc not_Cons_self)
|
nipkow@40230
|
1916 |
|
haftmann@24796
|
1917 |
|
haftmann@49948
|
1918 |
subsubsection {* @{const take} and @{const drop} *}
|
wenzelm@13114
|
1919 |
|
wenzelm@13142
|
1920 |
lemma take_0 [simp]: "take 0 xs = []"
|
nipkow@13145
|
1921 |
by (induct xs) auto
|
wenzelm@13114
|
1922 |
|
wenzelm@13142
|
1923 |
lemma drop_0 [simp]: "drop 0 xs = xs"
|
nipkow@13145
|
1924 |
by (induct xs) auto
|
wenzelm@13114
|
1925 |
|
wenzelm@13142
|
1926 |
lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs"
|
nipkow@13145
|
1927 |
by simp
|
wenzelm@13114
|
1928 |
|
wenzelm@13142
|
1929 |
lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs"
|
nipkow@13145
|
1930 |
by simp
|
wenzelm@13114
|
1931 |
|
wenzelm@13142
|
1932 |
declare take_Cons [simp del] and drop_Cons [simp del]
|
wenzelm@13114
|
1933 |
|
huffman@30128
|
1934 |
lemma take_1_Cons [simp]: "take 1 (x # xs) = [x]"
|
huffman@30128
|
1935 |
unfolding One_nat_def by simp
|
huffman@30128
|
1936 |
|
huffman@30128
|
1937 |
lemma drop_1_Cons [simp]: "drop 1 (x # xs) = xs"
|
huffman@30128
|
1938 |
unfolding One_nat_def by simp
|
huffman@30128
|
1939 |
|
nipkow@15110
|
1940 |
lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)"
|
nipkow@15110
|
1941 |
by(clarsimp simp add:neq_Nil_conv)
|
nipkow@15110
|
1942 |
|
nipkow@14187
|
1943 |
lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
|
nipkow@14187
|
1944 |
by(cases xs, simp_all)
|
nipkow@14187
|
1945 |
|
huffman@26584
|
1946 |
lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)"
|
huffman@26584
|
1947 |
by (induct xs arbitrary: n) simp_all
|
huffman@26584
|
1948 |
|
nipkow@24526
|
1949 |
lemma drop_tl: "drop n (tl xs) = tl(drop n xs)"
|
nipkow@24526
|
1950 |
by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split)
|
nipkow@24526
|
1951 |
|
huffman@26584
|
1952 |
lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)"
|
huffman@26584
|
1953 |
by (cases n, simp, cases xs, auto)
|
huffman@26584
|
1954 |
|
huffman@26584
|
1955 |
lemma tl_drop: "tl (drop n xs) = drop n (tl xs)"
|
huffman@26584
|
1956 |
by (simp only: drop_tl)
|
huffman@26584
|
1957 |
|
nipkow@24526
|
1958 |
lemma nth_via_drop: "drop n xs = y#ys \<Longrightarrow> xs!n = y"
|
nipkow@24526
|
1959 |
apply (induct xs arbitrary: n, simp)
|
nipkow@14187
|
1960 |
apply(simp add:drop_Cons nth_Cons split:nat.splits)
|
nipkow@14187
|
1961 |
done
|
nipkow@14187
|
1962 |
|
nipkow@13913
|
1963 |
lemma take_Suc_conv_app_nth:
|
nipkow@24526
|
1964 |
"i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]"
|
nipkow@24526
|
1965 |
apply (induct xs arbitrary: i, simp)
|
paulson@14208
|
1966 |
apply (case_tac i, auto)
|
nipkow@13913
|
1967 |
done
|
nipkow@13913
|
1968 |
|
mehta@14591
|
1969 |
lemma drop_Suc_conv_tl:
|
nipkow@24526
|
1970 |
"i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs"
|
nipkow@24526
|
1971 |
apply (induct xs arbitrary: i, simp)
|
mehta@14591
|
1972 |
apply (case_tac i, auto)
|
mehta@14591
|
1973 |
done
|
mehta@14591
|
1974 |
|
nipkow@24526
|
1975 |
lemma length_take [simp]: "length (take n xs) = min (length xs) n"
|
nipkow@24526
|
1976 |
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
|
nipkow@24526
|
1977 |
|
nipkow@24526
|
1978 |
lemma length_drop [simp]: "length (drop n xs) = (length xs - n)"
|
nipkow@24526
|
1979 |
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
|
nipkow@24526
|
1980 |
|
nipkow@24526
|
1981 |
lemma take_all [simp]: "length xs <= n ==> take n xs = xs"
|
nipkow@24526
|
1982 |
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
|
nipkow@24526
|
1983 |
|
nipkow@24526
|
1984 |
lemma drop_all [simp]: "length xs <= n ==> drop n xs = []"
|
nipkow@24526
|
1985 |
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
|
wenzelm@13114
|
1986 |
|
wenzelm@13142
|
1987 |
lemma take_append [simp]:
|
nipkow@24526
|
1988 |
"take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"
|
nipkow@24526
|
1989 |
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
|
wenzelm@13114
|
1990 |
|
wenzelm@13142
|
1991 |
lemma drop_append [simp]:
|
nipkow@24526
|
1992 |
"drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys"
|
nipkow@24526
|
1993 |
by (induct n arbitrary: xs) (auto, case_tac xs, auto)
|
nipkow@24526
|
1994 |
|
nipkow@24526
|
1995 |
lemma take_take [simp]: "take n (take m xs) = take (min n m) xs"
|
nipkow@24526
|
1996 |
apply (induct m arbitrary: xs n, auto)
|
paulson@14208
|
1997 |
apply (case_tac xs, auto)
|
nipkow@15236
|
1998 |
apply (case_tac n, auto)
|
nipkow@13145
|
1999 |
done
|
wenzelm@13114
|
2000 |
|
nipkow@24526
|
2001 |
lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs"
|
nipkow@24526
|
2002 |
apply (induct m arbitrary: xs, auto)
|
paulson@14208
|
2003 |
apply (case_tac xs, auto)
|
nipkow@13145
|
2004 |
done
|
wenzelm@13114
|
2005 |
|
nipkow@24526
|
2006 |
lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)"
|
nipkow@24526
|
2007 |
apply (induct m arbitrary: xs n, auto)
|
paulson@14208
|
2008 |
apply (case_tac xs, auto)
|
nipkow@13145
|
2009 |
done
|
wenzelm@13114
|
2010 |
|
nipkow@24526
|
2011 |
lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)"
|
nipkow@24526
|
2012 |
apply(induct xs arbitrary: m n)
|
nipkow@14802
|
2013 |
apply simp
|
nipkow@14802
|
2014 |
apply(simp add: take_Cons drop_Cons split:nat.split)
|
nipkow@14802
|
2015 |
done
|
nipkow@14802
|
2016 |
|
nipkow@24526
|
2017 |
lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs"
|
nipkow@24526
|
2018 |
apply (induct n arbitrary: xs, auto)
|
paulson@14208
|
2019 |
apply (case_tac xs, auto)
|
nipkow@13145
|
2020 |
done
|
wenzelm@13114
|
2021 |
|
nipkow@24526
|
2022 |
lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \<or> xs = [])"
|
nipkow@24526
|
2023 |
apply(induct xs arbitrary: n)
|
nipkow@15110
|
2024 |
apply simp
|
nipkow@15110
|
2025 |
apply(simp add:take_Cons split:nat.split)
|
nipkow@15110
|
2026 |
done
|
nipkow@15110
|
2027 |
|
nipkow@24526
|
2028 |
lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs <= n)"
|
nipkow@24526
|
2029 |
apply(induct xs arbitrary: n)
|
nipkow@15110
|
2030 |
apply simp
|
nipkow@15110
|
2031 |
apply(simp add:drop_Cons split:nat.split)
|
nipkow@15110
|
2032 |
done
|
nipkow@15110
|
2033 |
|
nipkow@24526
|
2034 |
lemma take_map: "take n (map f xs) = map f (take n xs)"
|
nipkow@24526
|
2035 |
apply (induct n arbitrary: xs, auto)
|
paulson@14208
|
2036 |
apply (case_tac xs, auto)
|
nipkow@13145
|
2037 |
done
|
wenzelm@13114
|
2038 |
|
nipkow@24526
|
2039 |
lemma drop_map: "drop n (map f xs) = map f (drop n xs)"
|
nipkow@24526
|
2040 |
apply (induct n arbitrary: xs, auto)
|
paulson@14208
|
2041 |
apply (case_tac xs, auto)
|
nipkow@13145
|
2042 |
done
|
wenzelm@13114
|
2043 |
|
nipkow@24526
|
2044 |
lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)"
|
nipkow@24526
|
2045 |
apply (induct xs arbitrary: i, auto)
|
paulson@14208
|
2046 |
apply (case_tac i, auto)
|
nipkow@13145
|
2047 |
done
|
wenzelm@13114
|
2048 |
|
nipkow@24526
|
2049 |
lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)"
|
nipkow@24526
|
2050 |
apply (induct xs arbitrary: i, auto)
|
paulson@14208
|
2051 |
apply (case_tac i, auto)
|
nipkow@13145
|
2052 |
done
|
wenzelm@13114
|
2053 |
|
nipkow@24526
|
2054 |
lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i"
|
nipkow@24526
|
2055 |
apply (induct xs arbitrary: i n, auto)
|
paulson@14208
|
2056 |
apply (case_tac n, blast)
|
paulson@14208
|
2057 |
apply (case_tac i, auto)
|
nipkow@13145
|
2058 |
done
|
wenzelm@13114
|
2059 |
|
wenzelm@13142
|
2060 |
lemma nth_drop [simp]:
|
nipkow@24526
|
2061 |
"n + i <= length xs ==> (drop n xs)!i = xs!(n + i)"
|
nipkow@24526
|
2062 |
apply (induct n arbitrary: xs i, auto)
|
paulson@14208
|
2063 |
apply (case_tac xs, auto)
|
nipkow@13145
|
2064 |
done
|
nipkow@3507
|
2065 |
|
huffman@26584
|
2066 |
lemma butlast_take:
|
huffman@30128
|
2067 |
"n <= length xs ==> butlast (take n xs) = take (n - 1) xs"
|
huffman@26584
|
2068 |
by (simp add: butlast_conv_take min_max.inf_absorb1 min_max.inf_absorb2)
|
huffman@26584
|
2069 |
|
huffman@26584
|
2070 |
lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
|
huffman@30128
|
2071 |
by (simp add: butlast_conv_take drop_take add_ac)
|
huffman@26584
|
2072 |
|
huffman@26584
|
2073 |
lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs"
|
huffman@26584
|
2074 |
by (simp add: butlast_conv_take min_max.inf_absorb1)
|
huffman@26584
|
2075 |
|
huffman@26584
|
2076 |
lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)"
|
huffman@30128
|
2077 |
by (simp add: butlast_conv_take drop_take add_ac)
|
huffman@26584
|
2078 |
|
bulwahn@46500
|
2079 |
lemma hd_drop_conv_nth: "n < length xs \<Longrightarrow> hd(drop n xs) = xs!n"
|
nipkow@18423
|
2080 |
by(simp add: hd_conv_nth)
|
nipkow@18423
|
2081 |
|
nipkow@35248
|
2082 |
lemma set_take_subset_set_take:
|
nipkow@35248
|
2083 |
"m <= n \<Longrightarrow> set(take m xs) <= set(take n xs)"
|
bulwahn@41463
|
2084 |
apply (induct xs arbitrary: m n)
|
bulwahn@41463
|
2085 |
apply simp
|
bulwahn@41463
|
2086 |
apply (case_tac n)
|
bulwahn@41463
|
2087 |
apply (auto simp: take_Cons)
|
bulwahn@41463
|
2088 |
done
|
nipkow@35248
|
2089 |
|
nipkow@24526
|
2090 |
lemma set_take_subset: "set(take n xs) \<subseteq> set xs"
|
nipkow@24526
|
2091 |
by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split)
|
nipkow@24526
|
2092 |
|
nipkow@24526
|
2093 |
lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs"
|
nipkow@24526
|
2094 |
by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split)
|
nipkow@14025
|
2095 |
|
nipkow@35248
|
2096 |
lemma set_drop_subset_set_drop:
|
nipkow@35248
|
2097 |
"m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
|
nipkow@35248
|
2098 |
apply(induct xs arbitrary: m n)
|
nipkow@35248
|
2099 |
apply(auto simp:drop_Cons split:nat.split)
|
nipkow@35248
|
2100 |
apply (metis set_drop_subset subset_iff)
|
nipkow@35248
|
2101 |
done
|
nipkow@35248
|
2102 |
|
nipkow@14187
|
2103 |
lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
|
nipkow@14187
|
2104 |
using set_take_subset by fast
|
nipkow@14187
|
2105 |
|
nipkow@14187
|
2106 |
lemma in_set_dropD: "x : set(drop n xs) \<Longrightarrow> x : set xs"
|
nipkow@14187
|
2107 |
using set_drop_subset by fast
|
nipkow@14187
|
2108 |
|
wenzelm@13114
|
2109 |
lemma append_eq_conv_conj:
|
nipkow@24526
|
2110 |
"(xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)"
|
nipkow@24526
|
2111 |
apply (induct xs arbitrary: zs, simp, clarsimp)
|
paulson@14208
|
2112 |
apply (case_tac zs, auto)
|
nipkow@13145
|
2113 |
done
|
wenzelm@13142
|
2114 |
|
nipkow@24526
|
2115 |
lemma take_add:
|
noschinl@42713
|
2116 |
"take (i+j) xs = take i xs @ take j (drop i xs)"
|
nipkow@24526
|
2117 |
apply (induct xs arbitrary: i, auto)
|
nipkow@24526
|
2118 |
apply (case_tac i, simp_all)
|
paulson@14050
|
2119 |
done
|
paulson@14050
|
2120 |
|
nipkow@14300
|
2121 |
lemma append_eq_append_conv_if:
|
wenzelm@53015
|
2122 |
"(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) =
|
wenzelm@53015
|
2123 |
(if size xs\<^sub>1 \<le> size ys\<^sub>1
|
wenzelm@53015
|
2124 |
then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \<and> xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2
|
wenzelm@53015
|
2125 |
else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \<and> drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)"
|
wenzelm@53015
|
2126 |
apply(induct xs\<^sub>1 arbitrary: ys\<^sub>1)
|
nipkow@14300
|
2127 |
apply simp
|
wenzelm@53015
|
2128 |
apply(case_tac ys\<^sub>1)
|
nipkow@14300
|
2129 |
apply simp_all
|
nipkow@14300
|
2130 |
done
|
nipkow@14300
|
2131 |
|
nipkow@15110
|
2132 |
lemma take_hd_drop:
|
huffman@30079
|
2133 |
"n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs"
|
nipkow@24526
|
2134 |
apply(induct xs arbitrary: n)
|
nipkow@15110
|
2135 |
apply simp
|
nipkow@15110
|
2136 |
apply(simp add:drop_Cons split:nat.split)
|
nipkow@15110
|
2137 |
done
|
nipkow@15110
|
2138 |
|
nipkow@17501
|
2139 |
lemma id_take_nth_drop:
|
nipkow@17501
|
2140 |
"i < length xs \<Longrightarrow> xs = take i xs @ xs!i # drop (Suc i) xs"
|
nipkow@17501
|
2141 |
proof -
|
nipkow@17501
|
2142 |
assume si: "i < length xs"
|
nipkow@17501
|
2143 |
hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto
|
nipkow@17501
|
2144 |
moreover
|
nipkow@17501
|
2145 |
from si have "take (Suc i) xs = take i xs @ [xs!i]"
|
nipkow@17501
|
2146 |
apply (rule_tac take_Suc_conv_app_nth) by arith
|
nipkow@17501
|
2147 |
ultimately show ?thesis by auto
|
nipkow@17501
|
2148 |
qed
|
nipkow@17501
|
2149 |
|
nipkow@17501
|
2150 |
lemma upd_conv_take_nth_drop:
|
nipkow@17501
|
2151 |
"i < length xs \<Longrightarrow> xs[i:=a] = take i xs @ a # drop (Suc i) xs"
|
nipkow@17501
|
2152 |
proof -
|
nipkow@17501
|
2153 |
assume i: "i < length xs"
|
nipkow@17501
|
2154 |
have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]"
|
nipkow@17501
|
2155 |
by(rule arg_cong[OF id_take_nth_drop[OF i]])
|
nipkow@17501
|
2156 |
also have "\<dots> = take i xs @ a # drop (Suc i) xs"
|
nipkow@17501
|
2157 |
using i by (simp add: list_update_append)
|
nipkow@17501
|
2158 |
finally show ?thesis .
|
nipkow@17501
|
2159 |
qed
|
nipkow@17501
|
2160 |
|
haftmann@24796
|
2161 |
lemma nth_drop':
|
haftmann@24796
|
2162 |
"i < length xs \<Longrightarrow> xs ! i # drop (Suc i) xs = drop i xs"
|
haftmann@24796
|
2163 |
apply (induct i arbitrary: xs)
|
haftmann@24796
|
2164 |
apply (simp add: neq_Nil_conv)
|
haftmann@24796
|
2165 |
apply (erule exE)+
|
haftmann@24796
|
2166 |
apply simp
|
haftmann@24796
|
2167 |
apply (case_tac xs)
|
haftmann@24796
|
2168 |
apply simp_all
|
haftmann@24796
|
2169 |
done
|
haftmann@24796
|
2170 |
|
wenzelm@13114
|
2171 |
|
haftmann@49948
|
2172 |
subsubsection {* @{const takeWhile} and @{const dropWhile} *}
|
wenzelm@13114
|
2173 |
|
hoelzl@33639
|
2174 |
lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs"
|
hoelzl@33639
|
2175 |
by (induct xs) auto
|
hoelzl@33639
|
2176 |
|
wenzelm@13142
|
2177 |
lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs"
|
nipkow@13145
|
2178 |
by (induct xs) auto
|
wenzelm@13114
|
2179 |
|
wenzelm@13142
|
2180 |
lemma takeWhile_append1 [simp]:
|
nipkow@13145
|
2181 |
"[| x:set xs; ~P(x)|] ==> takeWhile P (xs @ ys) = takeWhile P xs"
|
nipkow@13145
|
2182 |
by (induct xs) auto
|
wenzelm@13114
|
2183 |
|
wenzelm@13142
|
2184 |
lemma takeWhile_append2 [simp]:
|
nipkow@13145
|
2185 |
"(!!x. x : set xs ==> P x) ==> takeWhile P (xs @ ys) = xs @ takeWhile P ys"
|
nipkow@13145
|
2186 |
by (induct xs) auto
|
wenzelm@13114
|
2187 |
|
wenzelm@13142
|
2188 |
lemma takeWhile_tail: "\<not> P x ==> takeWhile P (xs @ (x#l)) = takeWhile P xs"
|
nipkow@13145
|
2189 |
by (induct xs) auto
|
wenzelm@13114
|
2190 |
|
hoelzl@33639
|
2191 |
lemma takeWhile_nth: "j < length (takeWhile P xs) \<Longrightarrow> takeWhile P xs ! j = xs ! j"
|
hoelzl@33639
|
2192 |
apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
|
hoelzl@33639
|
2193 |
|
hoelzl@33639
|
2194 |
lemma dropWhile_nth: "j < length (dropWhile P xs) \<Longrightarrow> dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))"
|
hoelzl@33639
|
2195 |
apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
|
hoelzl@33639
|
2196 |
|
hoelzl@33639
|
2197 |
lemma length_dropWhile_le: "length (dropWhile P xs) \<le> length xs"
|
hoelzl@33639
|
2198 |
by (induct xs) auto
|
hoelzl@33639
|
2199 |
|
wenzelm@13142
|
2200 |
lemma dropWhile_append1 [simp]:
|
nipkow@13145
|
2201 |
"[| x : set xs; ~P(x)|] ==> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
|
nipkow@13145
|
2202 |
by (induct xs) auto
|
wenzelm@13114
|
2203 |
|
wenzelm@13142
|
2204 |
lemma dropWhile_append2 [simp]:
|
nipkow@13145
|
2205 |
"(!!x. x:set xs ==> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
|
nipkow@13145
|
2206 |
by (induct xs) auto
|
wenzelm@13114
|
2207 |
|
noschinl@45841
|
2208 |
lemma dropWhile_append3:
|
noschinl@45841
|
2209 |
"\<not> P y \<Longrightarrow>dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys"
|
noschinl@45841
|
2210 |
by (induct xs) auto
|
noschinl@45841
|
2211 |
|
noschinl@45841
|
2212 |
lemma dropWhile_last:
|
noschinl@45841
|
2213 |
"x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> last (dropWhile P xs) = last xs"
|
noschinl@45841
|
2214 |
by (auto simp add: dropWhile_append3 in_set_conv_decomp)
|
noschinl@45841
|
2215 |
|
noschinl@45841
|
2216 |
lemma set_dropWhileD: "x \<in> set (dropWhile P xs) \<Longrightarrow> x \<in> set xs"
|
noschinl@45841
|
2217 |
by (induct xs) (auto split: split_if_asm)
|
noschinl@45841
|
2218 |
|
krauss@23971
|
2219 |
lemma set_takeWhileD: "x : set (takeWhile P xs) ==> x : set xs \<and> P x"
|
nipkow@13145
|
2220 |
by (induct xs) (auto split: split_if_asm)
|
wenzelm@13114
|
2221 |
|
nipkow@13913
|
2222 |
lemma takeWhile_eq_all_conv[simp]:
|
nipkow@13913
|
2223 |
"(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)"
|
nipkow@13913
|
2224 |
by(induct xs, auto)
|
nipkow@13913
|
2225 |
|
nipkow@13913
|
2226 |
lemma dropWhile_eq_Nil_conv[simp]:
|
nipkow@13913
|
2227 |
"(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)"
|
nipkow@13913
|
2228 |
by(induct xs, auto)
|
nipkow@13913
|
2229 |
|
nipkow@13913
|
2230 |
lemma dropWhile_eq_Cons_conv:
|
nipkow@13913
|
2231 |
"(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys & \<not> P y)"
|
nipkow@13913
|
2232 |
by(induct xs, auto)
|
nipkow@13913
|
2233 |
|
nipkow@31077
|
2234 |
lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
|
nipkow@31077
|
2235 |
by (induct xs) (auto dest: set_takeWhileD)
|
nipkow@31077
|
2236 |
|
nipkow@31077
|
2237 |
lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)"
|
nipkow@31077
|
2238 |
by (induct xs) auto
|
nipkow@31077
|
2239 |
|
hoelzl@33639
|
2240 |
lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)"
|
hoelzl@33639
|
2241 |
by (induct xs) auto
|
hoelzl@33639
|
2242 |
|
hoelzl@33639
|
2243 |
lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \<circ> f) xs)"
|
hoelzl@33639
|
2244 |
by (induct xs) auto
|
hoelzl@33639
|
2245 |
|
hoelzl@33639
|
2246 |
lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs"
|
hoelzl@33639
|
2247 |
by (induct xs) auto
|
hoelzl@33639
|
2248 |
|
hoelzl@33639
|
2249 |
lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs"
|
hoelzl@33639
|
2250 |
by (induct xs) auto
|
hoelzl@33639
|
2251 |
|
hoelzl@33639
|
2252 |
lemma hd_dropWhile:
|
hoelzl@33639
|
2253 |
"dropWhile P xs \<noteq> [] \<Longrightarrow> \<not> P (hd (dropWhile P xs))"
|
hoelzl@33639
|
2254 |
using assms by (induct xs) auto
|
hoelzl@33639
|
2255 |
|
hoelzl@33639
|
2256 |
lemma takeWhile_eq_filter:
|
hoelzl@33639
|
2257 |
assumes "\<And> x. x \<in> set (dropWhile P xs) \<Longrightarrow> \<not> P x"
|
hoelzl@33639
|
2258 |
shows "takeWhile P xs = filter P xs"
|
hoelzl@33639
|
2259 |
proof -
|
hoelzl@33639
|
2260 |
have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)"
|
hoelzl@33639
|
2261 |
by simp
|
hoelzl@33639
|
2262 |
have B: "filter P (dropWhile P xs) = []"
|
hoelzl@33639
|
2263 |
unfolding filter_empty_conv using assms by blast
|
hoelzl@33639
|
2264 |
have "filter P xs = takeWhile P xs"
|
hoelzl@33639
|
2265 |
unfolding A filter_append B
|
hoelzl@33639
|
2266 |
by (auto simp add: filter_id_conv dest: set_takeWhileD)
|
hoelzl@33639
|
2267 |
thus ?thesis ..
|
hoelzl@33639
|
2268 |
qed
|
hoelzl@33639
|
2269 |
|
hoelzl@33639
|
2270 |
lemma takeWhile_eq_take_P_nth:
|
hoelzl@33639
|
2271 |
"\<lbrakk> \<And> i. \<lbrakk> i < n ; i < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) ; n < length xs \<Longrightarrow> \<not> P (xs ! n) \<rbrakk> \<Longrightarrow>
|
hoelzl@33639
|
2272 |
takeWhile P xs = take n xs"
|
hoelzl@33639
|
2273 |
proof (induct xs arbitrary: n)
|
hoelzl@33639
|
2274 |
case (Cons x xs)
|
hoelzl@33639
|
2275 |
thus ?case
|
hoelzl@33639
|
2276 |
proof (cases n)
|
hoelzl@33639
|
2277 |
case (Suc n') note this[simp]
|
hoelzl@33639
|
2278 |
have "P x" using Cons.prems(1)[of 0] by simp
|
hoelzl@33639
|
2279 |
moreover have "takeWhile P xs = take n' xs"
|
hoelzl@33639
|
2280 |
proof (rule Cons.hyps)
|
hoelzl@33639
|
2281 |
case goal1 thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp
|
hoelzl@33639
|
2282 |
next case goal2 thus ?case using Cons by auto
|
hoelzl@33639
|
2283 |
qed
|
hoelzl@33639
|
2284 |
ultimately show ?thesis by simp
|
hoelzl@33639
|
2285 |
qed simp
|
hoelzl@33639
|
2286 |
qed simp
|
hoelzl@33639
|
2287 |
|
hoelzl@33639
|
2288 |
lemma nth_length_takeWhile:
|
hoelzl@33639
|
2289 |
"length (takeWhile P xs) < length xs \<Longrightarrow> \<not> P (xs ! length (takeWhile P xs))"
|
hoelzl@33639
|
2290 |
by (induct xs) auto
|
hoelzl@33639
|
2291 |
|
hoelzl@33639
|
2292 |
lemma length_takeWhile_less_P_nth:
|
hoelzl@33639
|
2293 |
assumes all: "\<And> i. i < j \<Longrightarrow> P (xs ! i)" and "j \<le> length xs"
|
hoelzl@33639
|
2294 |
shows "j \<le> length (takeWhile P xs)"
|
hoelzl@33639
|
2295 |
proof (rule classical)
|
hoelzl@33639
|
2296 |
assume "\<not> ?thesis"
|
hoelzl@33639
|
2297 |
hence "length (takeWhile P xs) < length xs" using assms by simp
|
hoelzl@33639
|
2298 |
thus ?thesis using all `\<not> ?thesis` nth_length_takeWhile[of P xs] by auto
|
hoelzl@33639
|
2299 |
qed
|
nipkow@31077
|
2300 |
|
nipkow@17501
|
2301 |
text{* The following two lemmmas could be generalized to an arbitrary
|
nipkow@17501
|
2302 |
property. *}
|
nipkow@17501
|
2303 |
|
nipkow@17501
|
2304 |
lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
|
nipkow@17501
|
2305 |
takeWhile (\<lambda>y. y \<noteq> x) (rev xs) = rev (tl (dropWhile (\<lambda>y. y \<noteq> x) xs))"
|
nipkow@17501
|
2306 |
by(induct xs) (auto simp: takeWhile_tail[where l="[]"])
|
nipkow@17501
|
2307 |
|
nipkow@17501
|
2308 |
lemma dropWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
|
nipkow@17501
|
2309 |
dropWhile (\<lambda>y. y \<noteq> x) (rev xs) = x # rev (takeWhile (\<lambda>y. y \<noteq> x) xs)"
|
nipkow@17501
|
2310 |
apply(induct xs)
|
nipkow@17501
|
2311 |
apply simp
|
nipkow@17501
|
2312 |
apply auto
|
nipkow@17501
|
2313 |
apply(subst dropWhile_append2)
|
nipkow@17501
|
2314 |
apply auto
|
nipkow@17501
|
2315 |
done
|
nipkow@17501
|
2316 |
|
nipkow@18423
|
2317 |
lemma takeWhile_not_last:
|
bulwahn@46500
|
2318 |
"distinct xs \<Longrightarrow> takeWhile (\<lambda>y. y \<noteq> last xs) xs = butlast xs"
|
nipkow@18423
|
2319 |
apply(induct xs)
|
nipkow@18423
|
2320 |
apply simp
|
nipkow@18423
|
2321 |
apply(case_tac xs)
|
nipkow@18423
|
2322 |
apply(auto)
|
nipkow@18423
|
2323 |
done
|
nipkow@18423
|
2324 |
|
krauss@44013
|
2325 |
lemma takeWhile_cong [fundef_cong]:
|
krauss@18336
|
2326 |
"[| l = k; !!x. x : set l ==> P x = Q x |]
|
krauss@18336
|
2327 |
==> takeWhile P l = takeWhile Q k"
|
nipkow@24349
|
2328 |
by (induct k arbitrary: l) (simp_all)
|
krauss@18336
|
2329 |
|
krauss@44013
|
2330 |
lemma dropWhile_cong [fundef_cong]:
|
krauss@18336
|
2331 |
"[| l = k; !!x. x : set l ==> P x = Q x |]
|
krauss@18336
|
2332 |
==> dropWhile P l = dropWhile Q k"
|
nipkow@24349
|
2333 |
by (induct k arbitrary: l, simp_all)
|
krauss@18336
|
2334 |
|
haftmann@52380
|
2335 |
lemma takeWhile_idem [simp]:
|
haftmann@52380
|
2336 |
"takeWhile P (takeWhile P xs) = takeWhile P xs"
|
haftmann@52380
|
2337 |
by (induct xs) auto
|
haftmann@52380
|
2338 |
|
haftmann@52380
|
2339 |
lemma dropWhile_idem [simp]:
|
haftmann@52380
|
2340 |
"dropWhile P (dropWhile P xs) = dropWhile P xs"
|
haftmann@52380
|
2341 |
by (induct xs) auto
|
haftmann@52380
|
2342 |
|
wenzelm@13114
|
2343 |
|
haftmann@49948
|
2344 |
subsubsection {* @{const zip} *}
|
wenzelm@13114
|
2345 |
|
wenzelm@13142
|
2346 |
lemma zip_Nil [simp]: "zip [] ys = []"
|
nipkow@13145
|
2347 |
by (induct ys) auto
|
wenzelm@13114
|
2348 |
|
wenzelm@13142
|
2349 |
lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
|
nipkow@13145
|
2350 |
by simp
|
wenzelm@13114
|
2351 |
|
wenzelm@13142
|
2352 |
declare zip_Cons [simp del]
|
wenzelm@13114
|
2353 |
|
haftmann@36198
|
2354 |
lemma [code]:
|
haftmann@36198
|
2355 |
"zip [] ys = []"
|
haftmann@36198
|
2356 |
"zip xs [] = []"
|
haftmann@36198
|
2357 |
"zip (x # xs) (y # ys) = (x, y) # zip xs ys"
|
haftmann@36198
|
2358 |
by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+
|
haftmann@36198
|
2359 |
|
nipkow@15281
|
2360 |
lemma zip_Cons1:
|
nipkow@15281
|
2361 |
"zip (x#xs) ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x,y)#zip xs ys)"
|
nipkow@15281
|
2362 |
by(auto split:list.split)
|
nipkow@15281
|
2363 |
|
wenzelm@13142
|
2364 |
lemma length_zip [simp]:
|
krauss@22493
|
2365 |
"length (zip xs ys) = min (length xs) (length ys)"
|
krauss@22493
|
2366 |
by (induct xs ys rule:list_induct2') auto
|
wenzelm@13114
|
2367 |
|
haftmann@34978
|
2368 |
lemma zip_obtain_same_length:
|
haftmann@34978
|
2369 |
assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys)
|
haftmann@34978
|
2370 |
\<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)"
|
haftmann@34978
|
2371 |
shows "P (zip xs ys)"
|
haftmann@34978
|
2372 |
proof -
|
haftmann@34978
|
2373 |
let ?n = "min (length xs) (length ys)"
|
haftmann@34978
|
2374 |
have "P (zip (take ?n xs) (take ?n ys))"
|
haftmann@34978
|
2375 |
by (rule assms) simp_all
|
haftmann@34978
|
2376 |
moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)"
|
haftmann@34978
|
2377 |
proof (induct xs arbitrary: ys)
|
haftmann@34978
|
2378 |
case Nil then show ?case by simp
|
haftmann@34978
|
2379 |
next
|
haftmann@34978
|
2380 |
case (Cons x xs) then show ?case by (cases ys) simp_all
|
haftmann@34978
|
2381 |
qed
|
haftmann@34978
|
2382 |
ultimately show ?thesis by simp
|
haftmann@34978
|
2383 |
qed
|
haftmann@34978
|
2384 |
|
wenzelm@13114
|
2385 |
lemma zip_append1:
|
krauss@22493
|
2386 |
"zip (xs @ ys) zs =
|
nipkow@13145
|
2387 |
zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
|
krauss@22493
|
2388 |
by (induct xs zs rule:list_induct2') auto
|
wenzelm@13114
|
2389 |
|
wenzelm@13114
|
2390 |
lemma zip_append2:
|
krauss@22493
|
2391 |
"zip xs (ys @ zs) =
|
nipkow@13145
|
2392 |
zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"
|
krauss@22493
|
2393 |
by (induct xs ys rule:list_induct2') auto
|
wenzelm@13114
|
2394 |
|
wenzelm@13142
|
2395 |
lemma zip_append [simp]:
|
bulwahn@46500
|
2396 |
"[| length xs = length us |] ==>
|
nipkow@13145
|
2397 |
zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
|
nipkow@13145
|
2398 |
by (simp add: zip_append1)
|
wenzelm@13114
|
2399 |
|
wenzelm@13114
|
2400 |
lemma zip_rev:
|
nipkow@14247
|
2401 |
"length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
|
nipkow@14247
|
2402 |
by (induct rule:list_induct2, simp_all)
|
wenzelm@13114
|
2403 |
|
hoelzl@33639
|
2404 |
lemma zip_map_map:
|
hoelzl@33639
|
2405 |
"zip (map f xs) (map g ys) = map (\<lambda> (x, y). (f x, g y)) (zip xs ys)"
|
hoelzl@33639
|
2406 |
proof (induct xs arbitrary: ys)
|
hoelzl@33639
|
2407 |
case (Cons x xs) note Cons_x_xs = Cons.hyps
|
hoelzl@33639
|
2408 |
show ?case
|
hoelzl@33639
|
2409 |
proof (cases ys)
|
hoelzl@33639
|
2410 |
case (Cons y ys')
|
hoelzl@33639
|
2411 |
show ?thesis unfolding Cons using Cons_x_xs by simp
|
hoelzl@33639
|
2412 |
qed simp
|
hoelzl@33639
|
2413 |
qed simp
|
hoelzl@33639
|
2414 |
|
hoelzl@33639
|
2415 |
lemma zip_map1:
|
hoelzl@33639
|
2416 |
"zip (map f xs) ys = map (\<lambda>(x, y). (f x, y)) (zip xs ys)"
|
hoelzl@33639
|
2417 |
using zip_map_map[of f xs "\<lambda>x. x" ys] by simp
|
hoelzl@33639
|
2418 |
|
hoelzl@33639
|
2419 |
lemma zip_map2:
|
hoelzl@33639
|
2420 |
"zip xs (map f ys) = map (\<lambda>(x, y). (x, f y)) (zip xs ys)"
|
hoelzl@33639
|
2421 |
using zip_map_map[of "\<lambda>x. x" xs f ys] by simp
|
hoelzl@33639
|
2422 |
|
nipkow@23096
|
2423 |
lemma map_zip_map:
|
hoelzl@33639
|
2424 |
"map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)"
|
hoelzl@33639
|
2425 |
unfolding zip_map1 by auto
|
nipkow@23096
|
2426 |
|
nipkow@23096
|
2427 |
lemma map_zip_map2:
|
hoelzl@33639
|
2428 |
"map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)"
|
hoelzl@33639
|
2429 |
unfolding zip_map2 by auto
|
nipkow@23096
|
2430 |
|
nipkow@31080
|
2431 |
text{* Courtesy of Andreas Lochbihler: *}
|
nipkow@31080
|
2432 |
lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs"
|
nipkow@31080
|
2433 |
by(induct xs) auto
|
nipkow@31080
|
2434 |
|
wenzelm@13142
|
2435 |
lemma nth_zip [simp]:
|
nipkow@24526
|
2436 |
"[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
|
nipkow@24526
|
2437 |
apply (induct ys arbitrary: i xs, simp)
|
nipkow@13145
|
2438 |
apply (case_tac xs)
|
nipkow@13145
|
2439 |
apply (simp_all add: nth.simps split: nat.split)
|
nipkow@13145
|
2440 |
done
|
wenzelm@13114
|
2441 |
|
wenzelm@13114
|
2442 |
lemma set_zip:
|
nipkow@13145
|
2443 |
"set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}"
|
nipkow@31080
|
2444 |
by(simp add: set_conv_nth cong: rev_conj_cong)
|
wenzelm@13114
|
2445 |
|
hoelzl@33639
|
2446 |
lemma zip_same: "((a,b) \<in> set (zip xs xs)) = (a \<in> set xs \<and> a = b)"
|
hoelzl@33639
|
2447 |
by(induct xs) auto
|
hoelzl@33639
|
2448 |
|
wenzelm@13114
|
2449 |
lemma zip_update:
|
nipkow@31080
|
2450 |
"zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"
|
nipkow@31080
|
2451 |
by(rule sym, simp add: update_zip)
|
wenzelm@13114
|
2452 |
|
wenzelm@13142
|
2453 |
lemma zip_replicate [simp]:
|
nipkow@24526
|
2454 |
"zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"
|
nipkow@24526
|
2455 |
apply (induct i arbitrary: j, auto)
|
paulson@14208
|
2456 |
apply (case_tac j, auto)
|
nipkow@13145
|
2457 |
done
|
wenzelm@13114
|
2458 |
|
nipkow@19487
|
2459 |
lemma take_zip:
|
nipkow@24526
|
2460 |
"take n (zip xs ys) = zip (take n xs) (take n ys)"
|
nipkow@24526
|
2461 |
apply (induct n arbitrary: xs ys)
|
nipkow@19487
|
2462 |
apply simp
|
nipkow@19487
|
2463 |
apply (case_tac xs, simp)
|
nipkow@19487
|
2464 |
apply (case_tac ys, simp_all)
|
nipkow@19487
|
2465 |
done
|
nipkow@19487
|
2466 |
|
nipkow@19487
|
2467 |
lemma drop_zip:
|
nipkow@24526
|
2468 |
"drop n (zip xs ys) = zip (drop n xs) (drop n ys)"
|
nipkow@24526
|
2469 |
apply (induct n arbitrary: xs ys)
|
nipkow@19487
|
2470 |
apply simp
|
nipkow@19487
|
2471 |
apply (case_tac xs, simp)
|
nipkow@19487
|
2472 |
apply (case_tac ys, simp_all)
|
nipkow@19487
|
2473 |
done
|
nipkow@19487
|
2474 |
|
hoelzl@33639
|
2475 |
lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \<circ> fst) (zip xs ys)"
|
hoelzl@33639
|
2476 |
proof (induct xs arbitrary: ys)
|
hoelzl@33639
|
2477 |
case (Cons x xs) thus ?case by (cases ys) auto
|
hoelzl@33639
|
2478 |
qed simp
|
hoelzl@33639
|
2479 |
|
hoelzl@33639
|
2480 |
lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \<circ> snd) (zip xs ys)"
|
hoelzl@33639
|
2481 |
proof (induct xs arbitrary: ys)
|
hoelzl@33639
|
2482 |
case (Cons x xs) thus ?case by (cases ys) auto
|
hoelzl@33639
|
2483 |
qed simp
|
hoelzl@33639
|
2484 |
|
krauss@22493
|
2485 |
lemma set_zip_leftD:
|
krauss@22493
|
2486 |
"(x,y)\<in> set (zip xs ys) \<Longrightarrow> x \<in> set xs"
|
krauss@22493
|
2487 |
by (induct xs ys rule:list_induct2') auto
|
krauss@22493
|
2488 |
|
krauss@22493
|
2489 |
lemma set_zip_rightD:
|
krauss@22493
|
2490 |
"(x,y)\<in> set (zip xs ys) \<Longrightarrow> y \<in> set ys"
|
krauss@22493
|
2491 |
by (induct xs ys rule:list_induct2') auto
|
wenzelm@13142
|
2492 |
|
nipkow@23983
|
2493 |
lemma in_set_zipE:
|
nipkow@23983
|
2494 |
"(x,y) : set(zip xs ys) \<Longrightarrow> (\<lbrakk> x : set xs; y : set ys \<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R"
|
nipkow@23983
|
2495 |
by(blast dest: set_zip_leftD set_zip_rightD)
|
nipkow@23983
|
2496 |
|
haftmann@29829
|
2497 |
lemma zip_map_fst_snd:
|
haftmann@29829
|
2498 |
"zip (map fst zs) (map snd zs) = zs"
|
haftmann@29829
|
2499 |
by (induct zs) simp_all
|
haftmann@29829
|
2500 |
|
haftmann@29829
|
2501 |
lemma zip_eq_conv:
|
haftmann@29829
|
2502 |
"length xs = length ys \<Longrightarrow> zip xs ys = zs \<longleftrightarrow> map fst zs = xs \<and> map snd zs = ys"
|
haftmann@29829
|
2503 |
by (auto simp add: zip_map_fst_snd)
|
haftmann@29829
|
2504 |
|
haftmann@51173
|
2505 |
lemma in_set_zip:
|
haftmann@51173
|
2506 |
"p \<in> set (zip xs ys) \<longleftrightarrow> (\<exists>n. xs ! n = fst p \<and> ys ! n = snd p
|
haftmann@51173
|
2507 |
\<and> n < length xs \<and> n < length ys)"
|
haftmann@51173
|
2508 |
by (cases p) (auto simp add: set_zip)
|
haftmann@51173
|
2509 |
|
haftmann@51173
|
2510 |
lemma pair_list_eqI:
|
haftmann@51173
|
2511 |
assumes "map fst xs = map fst ys" and "map snd xs = map snd ys"
|
haftmann@51173
|
2512 |
shows "xs = ys"
|
haftmann@51173
|
2513 |
proof -
|
haftmann@51173
|
2514 |
from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq)
|
haftmann@51173
|
2515 |
from this assms show ?thesis
|
haftmann@51173
|
2516 |
by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI)
|
haftmann@51173
|
2517 |
qed
|
haftmann@51173
|
2518 |
|
wenzelm@35115
|
2519 |
|
haftmann@49948
|
2520 |
subsubsection {* @{const list_all2} *}
|
wenzelm@13114
|
2521 |
|
kleing@14316
|
2522 |
lemma list_all2_lengthD [intro?]:
|
kleing@14316
|
2523 |
"list_all2 P xs ys ==> length xs = length ys"
|
nipkow@24349
|
2524 |
by (simp add: list_all2_def)
|
haftmann@19607
|
2525 |
|
haftmann@19787
|
2526 |
lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])"
|
nipkow@24349
|
2527 |
by (simp add: list_all2_def)
|
haftmann@19607
|
2528 |
|
haftmann@19787
|
2529 |
lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])"
|
nipkow@24349
|
2530 |
by (simp add: list_all2_def)
|
haftmann@19607
|
2531 |
|
haftmann@19607
|
2532 |
lemma list_all2_Cons [iff, code]:
|
haftmann@19607
|
2533 |
"list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)"
|
nipkow@24349
|
2534 |
by (auto simp add: list_all2_def)
|
wenzelm@13114
|
2535 |
|
wenzelm@13114
|
2536 |
lemma list_all2_Cons1:
|
nipkow@13145
|
2537 |
"list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)"
|
nipkow@13145
|
2538 |
by (cases ys) auto
|
wenzelm@13114
|
2539 |
|
wenzelm@13114
|
2540 |
lemma list_all2_Cons2:
|
nipkow@13145
|
2541 |
"list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)"
|
nipkow@13145
|
2542 |
by (cases xs) auto
|
wenzelm@13114
|
2543 |
|
huffman@45794
|
2544 |
lemma list_all2_induct
|
huffman@45794
|
2545 |
[consumes 1, case_names Nil Cons, induct set: list_all2]:
|
huffman@45794
|
2546 |
assumes P: "list_all2 P xs ys"
|
huffman@45794
|
2547 |
assumes Nil: "R [] []"
|
huffman@47640
|
2548 |
assumes Cons: "\<And>x xs y ys.
|
huffman@47640
|
2549 |
\<lbrakk>P x y; list_all2 P xs ys; R xs ys\<rbrakk> \<Longrightarrow> R (x # xs) (y # ys)"
|
huffman@45794
|
2550 |
shows "R xs ys"
|
huffman@45794
|
2551 |
using P
|
huffman@45794
|
2552 |
by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons)
|
huffman@45794
|
2553 |
|
wenzelm@13142
|
2554 |
lemma list_all2_rev [iff]:
|
nipkow@13145
|
2555 |
"list_all2 P (rev xs) (rev ys) = list_all2 P xs ys"
|
nipkow@13145
|
2556 |
by (simp add: list_all2_def zip_rev cong: conj_cong)
|
wenzelm@13114
|
2557 |
|
kleing@13863
|
2558 |
lemma list_all2_rev1:
|
kleing@13863
|
2559 |
"list_all2 P (rev xs) ys = list_all2 P xs (rev ys)"
|
kleing@13863
|
2560 |
by (subst list_all2_rev [symmetric]) simp
|
kleing@13863
|
2561 |
|
wenzelm@13114
|
2562 |
lemma list_all2_append1:
|
nipkow@13145
|
2563 |
"list_all2 P (xs @ ys) zs =
|
nipkow@13145
|
2564 |
(EX us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
|
nipkow@13145
|
2565 |
list_all2 P xs us \<and> list_all2 P ys vs)"
|
nipkow@13145
|
2566 |
apply (simp add: list_all2_def zip_append1)
|
nipkow@13145
|
2567 |
apply (rule iffI)
|
nipkow@13145
|
2568 |
apply (rule_tac x = "take (length xs) zs" in exI)
|
nipkow@13145
|
2569 |
apply (rule_tac x = "drop (length xs) zs" in exI)
|
paulson@14208
|
2570 |
apply (force split: nat_diff_split simp add: min_def, clarify)
|
nipkow@13145
|
2571 |
apply (simp add: ball_Un)
|
nipkow@13145
|
2572 |
done
|
wenzelm@13114
|
2573 |
|
wenzelm@13114
|
2574 |
lemma list_all2_append2:
|
nipkow@13145
|
2575 |
"list_all2 P xs (ys @ zs) =
|
nipkow@13145
|
2576 |
(EX us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
|
nipkow@13145
|
2577 |
list_all2 P us ys \<and> list_all2 P vs zs)"
|
nipkow@13145
|
2578 |
apply (simp add: list_all2_def zip_append2)
|
nipkow@13145
|
2579 |
apply (rule iffI)
|
nipkow@13145
|
2580 |
apply (rule_tac x = "take (length ys) xs" in exI)
|
nipkow@13145
|
2581 |
apply (rule_tac x = "drop (length ys) xs" in exI)
|
|