author | nipkow |
Thu, 17 Aug 2017 07:27:17 +0200 | |
changeset 66442 | 050bc74d55ed |
parent 66441 | b9468503742a |
child 66508 | 5df7a346f07b |
permissions | -rw-r--r-- |
13462 | 1 |
(* Title: HOL/List.thy |
2 |
Author: Tobias Nipkow |
|
923 | 3 |
*) |
4 |
||
60758 | 5 |
section \<open>The datatype of finite lists\<close> |
13122 | 6 |
|
15131 | 7 |
theory List |
58916 | 8 |
imports Sledgehammer Code_Numeral Lifting_Set |
15131 | 9 |
begin |
923 | 10 |
|
58310 | 11 |
datatype (set: 'a) list = |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57198
diff
changeset
|
12 |
Nil ("[]") |
55405
0a155051bd9d
use new selector support to define 'the', 'hd', 'tl'
blanchet
parents:
55404
diff
changeset
|
13 |
| Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) |
57206
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57200
diff
changeset
|
14 |
for |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57200
diff
changeset
|
15 |
map: map |
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents:
57200
diff
changeset
|
16 |
rel: list_all2 |
62328 | 17 |
pred: list_all |
57200
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57198
diff
changeset
|
18 |
where |
aab87ffa60cc
use 'where' clause for selector default value syntax
blanchet
parents:
57198
diff
changeset
|
19 |
"tl [] = []" |
57123
b5324647e0f1
tuned whitespace, to make datatype definitions slightly less intimidating
blanchet
parents:
57091
diff
changeset
|
20 |
|
55531
601ca8efa000
renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents:
55525
diff
changeset
|
21 |
datatype_compat list |
55404
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
22 |
|
55406 | 23 |
lemma [case_names Nil Cons, cases type: list]: |
61799 | 24 |
\<comment> \<open>for backward compatibility -- names of variables differ\<close> |
55406 | 25 |
"(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P" |
26 |
by (rule list.exhaust) |
|
27 |
||
28 |
lemma [case_names Nil Cons, induct type: list]: |
|
61799 | 29 |
\<comment> \<open>for backward compatibility -- names of variables differ\<close> |
55406 | 30 |
"P [] \<Longrightarrow> (\<And>a list. P list \<Longrightarrow> P (a # list)) \<Longrightarrow> P list" |
31 |
by (rule list.induct) |
|
32 |
||
60758 | 33 |
text \<open>Compatibility:\<close> |
34 |
||
35 |
setup \<open>Sign.mandatory_path "list"\<close> |
|
55404
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
36 |
|
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
37 |
lemmas inducts = list.induct |
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
38 |
lemmas recs = list.rec |
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
39 |
lemmas cases = list.case |
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
40 |
|
60758 | 41 |
setup \<open>Sign.parent_path\<close> |
55404
5cb95b79a51f
transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
blanchet
parents:
55148
diff
changeset
|
42 |
|
57816
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
blanchet
parents:
57587
diff
changeset
|
43 |
lemmas set_simps = list.set (* legacy *) |
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
blanchet
parents:
57587
diff
changeset
|
44 |
|
34941 | 45 |
syntax |
61799 | 46 |
\<comment> \<open>list Enumeration\<close> |
35115 | 47 |
"_list" :: "args => 'a list" ("[(_)]") |
34941 | 48 |
|
49 |
translations |
|
50 |
"[x, xs]" == "x#[xs]" |
|
51 |
"[x]" == "x#[]" |
|
52 |
||
35115 | 53 |
|
60758 | 54 |
subsection \<open>Basic list processing functions\<close> |
15302 | 55 |
|
58137 | 56 |
primrec (nonexhaustive) last :: "'a list \<Rightarrow> 'a" where |
50548 | 57 |
"last (x # xs) = (if xs = [] then x else last xs)" |
58 |
||
59 |
primrec butlast :: "'a list \<Rightarrow> 'a list" where |
|
57816
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
blanchet
parents:
57587
diff
changeset
|
60 |
"butlast [] = []" | |
50548 | 61 |
"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)" |
62 |
||
55584 | 63 |
lemma set_rec: "set xs = rec_list {} (\<lambda>x _. insert x) xs" |
64 |
by (induct xs) auto |
|
50548 | 65 |
|
66 |
definition coset :: "'a list \<Rightarrow> 'a set" where |
|
67 |
[simp]: "coset xs = - set xs" |
|
68 |
||
69 |
primrec append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where |
|
70 |
append_Nil: "[] @ ys = ys" | |
|
71 |
append_Cons: "(x#xs) @ ys = x # xs @ ys" |
|
72 |
||
73 |
primrec rev :: "'a list \<Rightarrow> 'a list" where |
|
74 |
"rev [] = []" | |
|
75 |
"rev (x # xs) = rev xs @ [x]" |
|
76 |
||
77 |
primrec filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
78 |
"filter P [] = []" | |
|
79 |
"filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" |
|
34941 | 80 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
81 |
text \<open>Special syntax for filter:\<close> |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
82 |
syntax (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
83 |
"_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])") |
34941 | 84 |
syntax |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
85 |
"_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_\<leftarrow>_ ./ _])") |
34941 | 86 |
translations |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
87 |
"[x<-xs . P]" \<rightleftharpoons> "CONST filter (\<lambda>x. P) xs" |
34941 | 88 |
|
50548 | 89 |
primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where |
90 |
fold_Nil: "fold f [] = id" | |
|
91 |
fold_Cons: "fold f (x # xs) = fold f xs \<circ> f x" |
|
92 |
||
93 |
primrec foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where |
|
94 |
foldr_Nil: "foldr f [] = id" | |
|
95 |
foldr_Cons: "foldr f (x # xs) = f x \<circ> foldr f xs" |
|
96 |
||
97 |
primrec foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where |
|
98 |
foldl_Nil: "foldl f a [] = a" | |
|
99 |
foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs" |
|
100 |
||
101 |
primrec concat:: "'a list list \<Rightarrow> 'a list" where |
|
102 |
"concat [] = []" | |
|
103 |
"concat (x # xs) = x @ concat xs" |
|
104 |
||
105 |
primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
106 |
drop_Nil: "drop n [] = []" | |
|
107 |
drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)" |
|
61799 | 108 |
\<comment> \<open>Warning: simpset does not contain this definition, but separate |
109 |
theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close> |
|
34941 | 110 |
|
50548 | 111 |
primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
112 |
take_Nil:"take n [] = []" | |
|
113 |
take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)" |
|
61799 | 114 |
\<comment> \<open>Warning: simpset does not contain this definition, but separate |
115 |
theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close> |
|
34941 | 116 |
|
58137 | 117 |
primrec (nonexhaustive) nth :: "'a list => nat => 'a" (infixl "!" 100) where |
50548 | 118 |
nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)" |
61799 | 119 |
\<comment> \<open>Warning: simpset does not contain this definition, but separate |
120 |
theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close> |
|
34941 | 121 |
|
50548 | 122 |
primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where |
123 |
"list_update [] i v = []" | |
|
124 |
"list_update (x # xs) i v = |
|
125 |
(case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)" |
|
923 | 126 |
|
41229
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents:
41075
diff
changeset
|
127 |
nonterminal lupdbinds and lupdbind |
5077
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents:
4643
diff
changeset
|
128 |
|
923 | 129 |
syntax |
13366 | 130 |
"_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)") |
131 |
"" :: "lupdbind => lupdbinds" ("_") |
|
132 |
"_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _") |
|
133 |
"_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [900,0] 900) |
|
5077
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
nipkow
parents:
4643
diff
changeset
|
134 |
|
923 | 135 |
translations |
35115 | 136 |
"_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs" |
34941 | 137 |
"xs[i:=x]" == "CONST list_update xs i x" |
138 |
||
50548 | 139 |
primrec takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
140 |
"takeWhile P [] = []" | |
|
141 |
"takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])" |
|
142 |
||
143 |
primrec dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
144 |
"dropWhile P [] = []" | |
|
145 |
"dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)" |
|
146 |
||
147 |
primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where |
|
148 |
"zip xs [] = []" | |
|
149 |
zip_Cons: "zip xs (y # ys) = |
|
150 |
(case xs of [] => [] | z # zs => (z, y) # zip zs ys)" |
|
61799 | 151 |
\<comment> \<open>Warning: simpset does not contain this definition, but separate |
152 |
theorems for \<open>xs = []\<close> and \<open>xs = z # zs\<close>\<close> |
|
34941 | 153 |
|
50548 | 154 |
primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where |
155 |
"product [] _ = []" | |
|
156 |
"product (x#xs) ys = map (Pair x) ys @ product xs ys" |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
157 |
|
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
158 |
hide_const (open) product |
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
159 |
|
53721
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
160 |
primrec product_lists :: "'a list list \<Rightarrow> 'a list list" where |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
161 |
"product_lists [] = [[]]" | |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
162 |
"product_lists (xs # xss) = concat (map (\<lambda>x. map (Cons x) (product_lists xss)) xs)" |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
163 |
|
50548 | 164 |
primrec upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where |
165 |
upt_0: "[i..<0] = []" | |
|
166 |
upt_Suc: "[i..<(Suc j)] = (if i <= j then [i..<j] @ [j] else [])" |
|
167 |
||
168 |
definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
169 |
"insert x xs = (if x \<in> set xs then xs else x # xs)" |
|
34978
874150ddd50a
canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents:
34942
diff
changeset
|
170 |
|
57198 | 171 |
definition union :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
172 |
"union = fold insert" |
|
173 |
||
174 |
hide_const (open) insert union |
|
175 |
hide_fact (open) insert_def union_def |
|
34978
874150ddd50a
canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents:
34942
diff
changeset
|
176 |
|
47122 | 177 |
primrec find :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a option" where |
50548 | 178 |
"find _ [] = None" | |
179 |
"find P (x#xs) = (if P x then Some x else find P xs)" |
|
47122 | 180 |
|
61799 | 181 |
text \<open>In the context of multisets, \<open>count_list\<close> is equivalent to |
60541 | 182 |
@{term "count o mset"} and it it advisable to use the latter.\<close> |
183 |
primrec count_list :: "'a list \<Rightarrow> 'a \<Rightarrow> nat" where |
|
184 |
"count_list [] y = 0" | |
|
185 |
"count_list (x#xs) y = (if x=y then count_list xs y + 1 else count_list xs y)" |
|
59728 | 186 |
|
55807 | 187 |
definition |
188 |
"extract" :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> ('a list * 'a * 'a list) option" |
|
189 |
where "extract P xs = |
|
190 |
(case dropWhile (Not o P) xs of |
|
191 |
[] \<Rightarrow> None | |
|
192 |
y#ys \<Rightarrow> Some(takeWhile (Not o P) xs, y, ys))" |
|
193 |
||
194 |
hide_const (open) "extract" |
|
195 |
||
51096 | 196 |
primrec those :: "'a option list \<Rightarrow> 'a list option" |
197 |
where |
|
198 |
"those [] = Some []" | |
|
199 |
"those (x # xs) = (case x of |
|
200 |
None \<Rightarrow> None |
|
55466 | 201 |
| Some y \<Rightarrow> map_option (Cons y) (those xs))" |
51096 | 202 |
|
50548 | 203 |
primrec remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
204 |
"remove1 x [] = []" | |
|
205 |
"remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)" |
|
206 |
||
207 |
primrec removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
208 |
"removeAll x [] = []" | |
|
209 |
"removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)" |
|
210 |
||
211 |
primrec distinct :: "'a list \<Rightarrow> bool" where |
|
212 |
"distinct [] \<longleftrightarrow> True" | |
|
213 |
"distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs" |
|
214 |
||
215 |
primrec remdups :: "'a list \<Rightarrow> 'a list" where |
|
216 |
"remdups [] = []" | |
|
217 |
"remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)" |
|
218 |
||
53721
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
219 |
fun remdups_adj :: "'a list \<Rightarrow> 'a list" where |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
220 |
"remdups_adj [] = []" | |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
221 |
"remdups_adj [x] = [x]" | |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
222 |
"remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))" |
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
223 |
|
50548 | 224 |
primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where |
225 |
replicate_0: "replicate 0 x = []" | |
|
226 |
replicate_Suc: "replicate (Suc n) x = x # replicate n x" |
|
3342
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
227 |
|
60758 | 228 |
text \<open> |
61799 | 229 |
Function \<open>size\<close> is overloaded for all datatypes. Users may |
230 |
refer to the list version as \<open>length\<close>.\<close> |
|
13142 | 231 |
|
50548 | 232 |
abbreviation length :: "'a list \<Rightarrow> nat" where |
233 |
"length \<equiv> size" |
|
15307 | 234 |
|
51173 | 235 |
definition enumerate :: "nat \<Rightarrow> 'a list \<Rightarrow> (nat \<times> 'a) list" where |
236 |
enumerate_eq_zip: "enumerate n xs = zip [n..<n + length xs] xs" |
|
237 |
||
46440
d4994e2e7364
use 'primrec' to define "rotate1", for uniformity (and to help first-order tools that rely on "Spec_Rules")
blanchet
parents:
46439
diff
changeset
|
238 |
primrec rotate1 :: "'a list \<Rightarrow> 'a list" where |
50548 | 239 |
"rotate1 [] = []" | |
240 |
"rotate1 (x # xs) = xs @ [x]" |
|
241 |
||
242 |
definition rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
243 |
"rotate n = rotate1 ^^ n" |
|
244 |
||
65956
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
245 |
definition nths :: "'a list => nat set => 'a list" where |
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
246 |
"nths xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))" |
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
247 |
|
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
248 |
primrec subseqs :: "'a list \<Rightarrow> 'a list list" where |
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
249 |
"subseqs [] = [[]]" | |
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
250 |
"subseqs (x#xs) = (let xss = subseqs xs in map (Cons x) xss @ xss)" |
50548 | 251 |
|
252 |
primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where |
|
253 |
"n_lists 0 xs = [[]]" | |
|
254 |
"n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))" |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
255 |
|
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
256 |
hide_const (open) n_lists |
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
257 |
|
40593
1e57b18d27b1
code eqn for slice was missing; redefined splice with fun
nipkow
parents:
40365
diff
changeset
|
258 |
fun splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
1e57b18d27b1
code eqn for slice was missing; redefined splice with fun
nipkow
parents:
40365
diff
changeset
|
259 |
"splice [] ys = ys" | |
1e57b18d27b1
code eqn for slice was missing; redefined splice with fun
nipkow
parents:
40365
diff
changeset
|
260 |
"splice xs [] = xs" | |
1e57b18d27b1
code eqn for slice was missing; redefined splice with fun
nipkow
parents:
40365
diff
changeset
|
261 |
"splice (x#xs) (y#ys) = x # y # splice xs ys" |
21061
580dfc999ef6
added normal post setup; cleaned up "execution" constants
haftmann
parents:
21046
diff
changeset
|
262 |
|
65348
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
263 |
function shuffle where |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
264 |
"shuffle [] ys = {ys}" |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
265 |
| "shuffle xs [] = {xs}" |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
266 |
| "shuffle (x # xs) (y # ys) = op # x ` shuffle xs (y # ys) \<union> op # y ` shuffle (x # xs) ys" |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
267 |
by pat_completeness simp_all |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
268 |
termination by lexicographic_order |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
269 |
|
60758 | 270 |
text\<open> |
26771 | 271 |
\begin{figure}[htbp] |
272 |
\fbox{ |
|
273 |
\begin{tabular}{l} |
|
27381 | 274 |
@{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\ |
275 |
@{lemma "length [a,b,c] = 3" by simp}\\ |
|
276 |
@{lemma "set [a,b,c] = {a,b,c}" by simp}\\ |
|
277 |
@{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\ |
|
278 |
@{lemma "rev [a,b,c] = [c,b,a]" by simp}\\ |
|
279 |
@{lemma "hd [a,b,c,d] = a" by simp}\\ |
|
280 |
@{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\ |
|
281 |
@{lemma "last [a,b,c,d] = d" by simp}\\ |
|
282 |
@{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\ |
|
283 |
@{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\ |
|
284 |
@{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\ |
|
46133
d9fe85d3d2cd
incorporated canonical fold combinator on lists into body of List theory; refactored passages on List.fold(l/r)
haftmann
parents:
46125
diff
changeset
|
285 |
@{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\ |
47397
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
haftmann
parents:
47131
diff
changeset
|
286 |
@{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\ |
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
haftmann
parents:
47131
diff
changeset
|
287 |
@{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\ |
27381 | 288 |
@{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\ |
289 |
@{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\ |
|
51173 | 290 |
@{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\ |
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
291 |
@{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\ |
53721
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
292 |
@{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\ |
27381 | 293 |
@{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\ |
294 |
@{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\ |
|
65348
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
295 |
@{lemma "shuffle [a,b] [c,d] = {[a,b,c,d],[a,c,b,d],[a,c,d,b],[c,a,b,d],[c,a,d,b],[c,d,a,b]}" |
b149abe619f7
added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents:
64966
diff
changeset
|
296 |
by (simp add: insert_commute)}\\ |
27381 | 297 |
@{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\ |
298 |
@{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\ |
|
299 |
@{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\ |
|
300 |
@{lemma "drop 6 [a,b,c,d] = []" by simp}\\ |
|
301 |
@{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\ |
|
302 |
@{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\ |
|
303 |
@{lemma "distinct [2,0,1::nat]" by simp}\\ |
|
304 |
@{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\ |
|
53721
ccaceea6c768
added two functions to List (one contributed by Manuel Eberl)
traytel
parents:
53689
diff
changeset
|
305 |
@{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\ |
34978
874150ddd50a
canonical insert operation; generalized lemma foldl_apply_inv to foldl_apply
haftmann
parents:
34942
diff
changeset
|
306 |
@{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\ |
35295 | 307 |
@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\ |
57198 | 308 |
@{lemma "List.union [2,3,4] [0::int,1,2] = [4,3,0,1,2]" by (simp add: List.insert_def List.union_def)}\\ |
47122 | 309 |
@{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\ |
310 |
@{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\ |
|
60541 | 311 |
@{lemma "count_list [0,1,0,2::int] 0 = 2" by (simp)}\\ |
55807 | 312 |
@{lemma "List.extract (%i::int. i>0) [0,0] = None" by(simp add: extract_def)}\\ |
313 |
@{lemma "List.extract (%i::int. i>0) [0,1,0,2] = Some([0], 1, [0,2])" by(simp add: extract_def)}\\ |
|
27381 | 314 |
@{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\ |
27693 | 315 |
@{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\ |
27381 | 316 |
@{lemma "nth [a,b,c,d] 2 = c" by simp}\\ |
317 |
@{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\ |
|
65956
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
318 |
@{lemma "nths [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:nths_def)}\\ |
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
65348
diff
changeset
|
319 |
@{lemma "subseqs [a,b] = [[a, b], [a], [b], []]" by simp}\\ |
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
320 |
@{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)}\\ |
46440
d4994e2e7364
use 'primrec' to define "rotate1", for uniformity (and to help first-order tools that rely on "Spec_Rules")
blanchet
parents:
46439
diff
changeset
|
321 |
@{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\ |
d4994e2e7364
use 'primrec' to define "rotate1", for uniformity (and to help first-order tools that rely on "Spec_Rules")
blanchet
parents:
46439
diff
changeset
|
322 |
@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\ |
40077 | 323 |
@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\ |
58101 | 324 |
@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)} |
26771 | 325 |
\end{tabular}} |
326 |
\caption{Characteristic examples} |
|
327 |
\label{fig:Characteristic} |
|
328 |
\end{figure} |
|
29927 | 329 |
Figure~\ref{fig:Characteristic} shows characteristic examples |
26771 | 330 |
that should give an intuitive understanding of the above functions. |
60758 | 331 |
\<close> |
332 |
||
333 |
text\<open>The following simple sort functions are intended for proofs, |
|
334 |
not for efficient implementations.\<close> |
|
24616 | 335 |
|
66436
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
336 |
text \<open>A sorted predicate w.r.t. a relation:\<close> |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
337 |
|
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
338 |
fun sorted_wrt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
339 |
"sorted_wrt P [] = True" | |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
340 |
"sorted_wrt P [x] = True" | |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
341 |
"sorted_wrt P (x # y # zs) = (P x y \<and> sorted_wrt P (y # zs))" |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
342 |
|
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
343 |
(* FIXME: define sorted in terms of sorted_wrt *) |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
344 |
|
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
345 |
text \<open>A class-based sorted predicate:\<close> |
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents:
66360
diff
changeset
|
346 |
|
25221
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
347 |
context linorder |
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
348 |
begin |
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
349 |
|
39915
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
350 |
inductive sorted :: "'a list \<Rightarrow> bool" where |
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
351 |
Nil [iff]: "sorted []" |
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
352 |
| Cons: "\<forall>y\<in>set xs. x \<le> y \<Longrightarrow> sorted xs \<Longrightarrow> sorted (x # xs)" |
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
353 |
|
58807 | 354 |
lemma sorted_single [iff]: "sorted [x]" |
355 |
by (rule sorted.Cons) auto |
|
356 |
||
357 |
lemma sorted_many: "x \<le> y \<Longrightarrow> sorted (y # zs) \<Longrightarrow> sorted (x # y # zs)" |
|
358 |
by (rule sorted.Cons) (cases "y # zs" rule: sorted.cases, auto) |
|
39915
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
359 |
|
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
360 |
lemma sorted_many_eq [simp, code]: |
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
361 |
"sorted (x # y # zs) \<longleftrightarrow> x \<le> y \<and> sorted (y # zs)" |
58807 | 362 |
by (auto intro: sorted_many elim: sorted.cases) |
39915
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
363 |
|
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
364 |
lemma [code]: |
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
365 |
"sorted [] \<longleftrightarrow> True" |
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents:
39774
diff
changeset
|
366 |
"sorted [x] \<longleftrightarrow> True" |
58807 | 367 |
by simp_all |
24697 | 368 |
|
33639
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
369 |
primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where |
50548 | 370 |
"insort_key f x [] = [x]" | |
371 |
"insort_key f x (y#ys) = |
|
372 |
(if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))" |
|
33639
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
373 |
|
35195 | 374 |
definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where |
50548 | 375 |
"sort_key f xs = foldr (insort_key f) xs []" |
33639
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
376 |
|
40210
aee7ef725330
sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents:
40195
diff
changeset
|
377 |
definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where |
50548 | 378 |
"insort_insert_key f x xs = |
379 |
(if f x \<in> f ` set xs then xs else insort_key f x xs)" |
|
40210
aee7ef725330
sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents:
40195
diff
changeset
|
380 |
|
33639
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
381 |
abbreviation "sort \<equiv> sort_key (\<lambda>x. x)" |
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
382 |
abbreviation "insort \<equiv> insort_key (\<lambda>x. x)" |
40210
aee7ef725330
sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents:
40195
diff
changeset
|
383 |
abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)" |
35608 | 384 |
|
25221
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
385 |
end |
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
386 |
|
24616 | 387 |
|
60758 | 388 |
subsubsection \<open>List comprehension\<close> |
389 |
||
390 |
text\<open>Input syntax for Haskell-like list comprehension notation. |
|
61799 | 391 |
Typical example: \<open>[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]\<close>, |
392 |
the list of all pairs of distinct elements from \<open>xs\<close> and \<open>ys\<close>. |
|
393 |
The syntax is as in Haskell, except that \<open>|\<close> becomes a dot |
|
394 |
(like in Isabelle's set comprehension): \<open>[e. x \<leftarrow> xs, \<dots>]\<close> rather than |
|
24349 | 395 |
\verb![e| x <- xs, ...]!. |
396 |
||
397 |
The qualifiers after the dot are |
|
398 |
\begin{description} |
|
61799 | 399 |
\item[generators] \<open>p \<leftarrow> xs\<close>, |
400 |
where \<open>p\<close> is a pattern and \<open>xs\<close> an expression of list type, or |
|
401 |
\item[guards] \<open>b\<close>, where \<open>b\<close> is a boolean expression. |
|
24476
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
402 |
%\item[local bindings] @ {text"let x = e"}. |
24349 | 403 |
\end{description} |
23240 | 404 |
|
24476
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
405 |
Just like in Haskell, list comprehension is just a shorthand. To avoid |
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
406 |
misunderstandings, the translation into desugared form is not reversed |
61799 | 407 |
upon output. Note that the translation of \<open>[e. x \<leftarrow> xs]\<close> is |
24476
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
408 |
optmized to @{term"map (%x. e) xs"}. |
23240 | 409 |
|
24349 | 410 |
It is easy to write short list comprehensions which stand for complex |
411 |
expressions. During proofs, they may become unreadable (and |
|
412 |
mangled). In such cases it can be advisable to introduce separate |
|
60758 | 413 |
definitions for the list comprehensions in question.\<close> |
24349 | 414 |
|
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
415 |
nonterminal lc_qual and lc_quals |
23192 | 416 |
|
417 |
syntax |
|
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
418 |
"_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list" ("[_ . __") |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
419 |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _") |
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
420 |
"_lc_test" :: "bool \<Rightarrow> lc_qual" ("_") |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
421 |
(*"_lc_let" :: "letbinds => lc_qual" ("let _")*) |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
422 |
"_lc_end" :: "lc_quals" ("]") |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
423 |
"_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __") |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
424 |
"_lc_abs" :: "'a => 'b list => 'b list" |
23192 | 425 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
426 |
syntax (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
427 |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _") |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
428 |
|
24476
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
429 |
(* These are easier than ML code but cannot express the optimized |
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
430 |
translation of [e. p<-xs] |
23192 | 431 |
translations |
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
432 |
"[e. p<-xs]" => "concat(map (_lc_abs p [e]) xs)" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
433 |
"_listcompr e (_lc_gen p xs) (_lc_quals Q Qs)" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
434 |
=> "concat (map (_lc_abs p (_listcompr e Q Qs)) xs)" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
435 |
"[e. P]" => "if P then [e] else []" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
436 |
"_listcompr e (_lc_test P) (_lc_quals Q Qs)" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
437 |
=> "if P then (_listcompr e Q Qs) else []" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
438 |
"_listcompr e (_lc_let b) (_lc_quals Q Qs)" |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
439 |
=> "_Let b (_listcompr e Q Qs)" |
24476
f7ad9fbbeeaa
turned list comprehension translations into ML to optimize base case
nipkow
parents:
24471
diff
changeset
|
440 |
*) |
23240 | 441 |
|
60758 | 442 |
parse_translation \<open> |
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
443 |
let |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
444 |
val NilC = Syntax.const @{const_syntax Nil}; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
445 |
val ConsC = Syntax.const @{const_syntax Cons}; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
446 |
val mapC = Syntax.const @{const_syntax map}; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
447 |
val concatC = Syntax.const @{const_syntax concat}; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
448 |
val IfC = Syntax.const @{const_syntax If}; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
449 |
|
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
450 |
fun single x = ConsC $ x $ NilC; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
451 |
|
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
452 |
fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *) |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
453 |
let |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
454 |
(* FIXME proper name context!? *) |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
455 |
val x = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
456 |
Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT); |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
457 |
val e = if opti then single e else e; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
458 |
val case1 = Syntax.const @{syntax_const "_case1"} $ p $ e; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
459 |
val case2 = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
460 |
Syntax.const @{syntax_const "_case1"} $ |
56241 | 461 |
Syntax.const @{const_syntax Pure.dummy_pattern} $ NilC; |
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
462 |
val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2; |
51678
1e33b81c328a
allow redundant cases in the list comprehension translation
traytel
parents:
51673
diff
changeset
|
463 |
in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end; |
46138
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
464 |
|
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
465 |
fun abs_tr ctxt p e opti = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
466 |
(case Term_Position.strip_positions p of |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
467 |
Free (s, T) => |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
468 |
let |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
469 |
val thy = Proof_Context.theory_of ctxt; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
470 |
val s' = Proof_Context.intern_const ctxt s; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
471 |
in |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
472 |
if Sign.declared_const thy s' |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
473 |
then (pat_tr ctxt p e opti, false) |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
474 |
else (Syntax_Trans.abs_tr [p, e], true) |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
475 |
end |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
476 |
| _ => (pat_tr ctxt p e opti, false)); |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
477 |
|
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
478 |
fun lc_tr ctxt [e, Const (@{syntax_const "_lc_test"}, _) $ b, qs] = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
479 |
let |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
480 |
val res = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
481 |
(case qs of |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
482 |
Const (@{syntax_const "_lc_end"}, _) => single e |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
483 |
| Const (@{syntax_const "_lc_quals"}, _) $ q $ qs => lc_tr ctxt [e, q, qs]); |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
484 |
in IfC $ b $ res $ NilC end |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
485 |
| lc_tr ctxt |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
486 |
[e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es, |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
487 |
Const(@{syntax_const "_lc_end"}, _)] = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
488 |
(case abs_tr ctxt p e true of |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
489 |
(f, true) => mapC $ f $ es |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
490 |
| (f, false) => concatC $ (mapC $ f $ es)) |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
491 |
| lc_tr ctxt |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
492 |
[e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es, |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
493 |
Const (@{syntax_const "_lc_quals"}, _) $ q $ qs] = |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
494 |
let val e' = lc_tr ctxt [e, q, qs]; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
495 |
in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end; |
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
496 |
|
85f8d8a8c711
improved list comprehension syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
46133
diff
changeset
|
497 |
in [(@{syntax_const "_listcompr"}, lc_tr)] end |
60758 | 498 |
\<close> |
499 |
||
500 |
ML_val \<open> |
|
42167 | 501 |
let |
60160 | 502 |
val read = Syntax.read_term @{context} o Syntax.implode_input; |
503 |
fun check s1 s2 = |
|
504 |
read s1 aconv read s2 orelse |
|
505 |
error ("Check failed: " ^ |
|
506 |
quote (Input.source_content s1) ^ Position.here_list [Input.pos_of s1, Input.pos_of s2]); |
|
42167 | 507 |
in |
60160 | 508 |
check \<open>[(x,y,z). b]\<close> \<open>if b then [(x, y, z)] else []\<close>; |
509 |
check \<open>[(x,y,z). x\<leftarrow>xs]\<close> \<open>map (\<lambda>x. (x, y, z)) xs\<close>; |
|
510 |
check \<open>[e x y. x\<leftarrow>xs, y\<leftarrow>ys]\<close> \<open>concat (map (\<lambda>x. map (\<lambda>y. e x y) ys) xs)\<close>; |
|
511 |
check \<open>[(x,y,z). x<a, x>b]\<close> \<open>if x < a then if b < x then [(x, y, z)] else [] else []\<close>; |
|
512 |
check \<open>[(x,y,z). x\<leftarrow>xs, x>b]\<close> \<open>concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)\<close>; |
|
513 |
check \<open>[(x,y,z). x<a, x\<leftarrow>xs]\<close> \<open>if x < a then map (\<lambda>x. (x, y, z)) xs else []\<close>; |
|
514 |
check \<open>[(x,y). Cons True x \<leftarrow> xs]\<close> |
|
515 |
\<open>concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)\<close>; |
|
516 |
check \<open>[(x,y,z). Cons x [] \<leftarrow> xs]\<close> |
|
517 |
\<open>concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)\<close>; |
|
518 |
check \<open>[(x,y,z). x<a, x>b, x=d]\<close> |
|
519 |
\<open>if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []\<close>; |
|
520 |
check \<open>[(x,y,z). x<a, x>b, y\<leftarrow>ys]\<close> |
|
521 |
\<open>if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []\<close>; |
|
522 |
check \<open>[(x,y,z). x<a, x\<leftarrow>xs,y>b]\<close> |
|
523 |
\<open>if x < a then concat (map (\<lambda>x. if b < y then [(x, y, z)] else []) xs) else []\<close>; |
|
524 |
check \<open>[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]\<close> |
|
525 |
\<open>if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []\<close>; |
|
526 |
check \<open>[(x,y,z). x\<leftarrow>xs, x>b, y<a]\<close> |
|
527 |
\<open>concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)\<close>; |
|
528 |
check \<open>[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]\<close> |
|
529 |
\<open>concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)\<close>; |
|
530 |
check \<open>[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]\<close> |
|
531 |
\<open>concat (map (\<lambda>x. concat (map (\<lambda>y. if x < y then [(x, y, z)] else []) ys)) xs)\<close>; |
|
532 |
check \<open>[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]\<close> |
|
533 |
\<open>concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)\<close> |
|
42167 | 534 |
end; |
60758 | 535 |
\<close> |
42167 | 536 |
|
35115 | 537 |
(* |
24349 | 538 |
term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]" |
23192 | 539 |
*) |
540 |
||
42167 | 541 |
|
60758 | 542 |
ML \<open> |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
543 |
(* Simproc for rewriting list comprehensions applied to List.set to set |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
544 |
comprehension. *) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
545 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
546 |
signature LIST_TO_SET_COMPREHENSION = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
547 |
sig |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
548 |
val simproc : Proof.context -> cterm -> thm option |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
549 |
end |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
550 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
551 |
structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
552 |
struct |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
553 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
554 |
(* conversion *) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
555 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
556 |
fun all_exists_conv cv ctxt ct = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
557 |
(case Thm.term_of ct of |
60156 | 558 |
Const (@{const_name Ex}, _) $ Abs _ => |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
559 |
Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
560 |
| _ => cv ctxt ct) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
561 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
562 |
fun all_but_last_exists_conv cv ctxt ct = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
563 |
(case Thm.term_of ct of |
60156 | 564 |
Const (@{const_name Ex}, _) $ Abs (_, _, Const (@{const_name Ex}, _) $ _) => |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
565 |
Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
566 |
| _ => cv ctxt ct) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
567 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
568 |
fun Collect_conv cv ctxt ct = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
569 |
(case Thm.term_of ct of |
60156 | 570 |
Const (@{const_name Collect}, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
571 |
| _ => raise CTERM ("Collect_conv", [ct])) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
572 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
573 |
fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
574 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
575 |
fun conjunct_assoc_conv ct = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
576 |
Conv.try_conv |
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
577 |
(rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
578 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
579 |
fun right_hand_set_comprehension_conv conv ctxt = |
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
580 |
HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
581 |
(Collect_conv (all_exists_conv conv o #2) ctxt)) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
582 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
583 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
584 |
(* term abstraction of list comprehension patterns *) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
585 |
|
60156 | 586 |
datatype termlets = If | Case of typ * int |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
587 |
|
60158 | 588 |
local |
589 |
||
590 |
val set_Nil_I = @{lemma "set [] = {x. False}" by (simp add: empty_def [symmetric])} |
|
591 |
val set_singleton = @{lemma "set [a] = {x. x = a}" by simp} |
|
592 |
val inst_Collect_mem_eq = @{lemma "set A = {x. x \<in> set A}" by simp} |
|
593 |
val del_refl_eq = @{lemma "(t = t \<and> P) \<equiv> P" by simp} |
|
594 |
||
595 |
fun mk_set T = Const (@{const_name set}, HOLogic.listT T --> HOLogic.mk_setT T) |
|
596 |
fun dest_set (Const (@{const_name set}, _) $ xs) = xs |
|
597 |
||
598 |
fun dest_singleton_list (Const (@{const_name Cons}, _) $ t $ (Const (@{const_name Nil}, _))) = t |
|
599 |
| dest_singleton_list t = raise TERM ("dest_singleton_list", [t]) |
|
600 |
||
601 |
(*We check that one case returns a singleton list and all other cases |
|
602 |
return [], and return the index of the one singleton list case.*) |
|
603 |
fun possible_index_of_singleton_case cases = |
|
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
604 |
let |
60158 | 605 |
fun check (i, case_t) s = |
606 |
(case strip_abs_body case_t of |
|
607 |
(Const (@{const_name Nil}, _)) => s |
|
608 |
| _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE)) |
|
609 |
in |
|
610 |
fold_index check cases (SOME NONE) |> the_default NONE |
|
611 |
end |
|
612 |
||
613 |
(*returns condition continuing term option*) |
|
614 |
fun dest_if (Const (@{const_name If}, _) $ cond $ then_t $ Const (@{const_name Nil}, _)) = |
|
615 |
SOME (cond, then_t) |
|
616 |
| dest_if _ = NONE |
|
617 |
||
618 |
(*returns (case_expr type index chosen_case constr_name) option*) |
|
619 |
fun dest_case ctxt case_term = |
|
620 |
let |
|
621 |
val (case_const, args) = strip_comb case_term |
|
622 |
in |
|
623 |
(case try dest_Const case_const of |
|
624 |
SOME (c, T) => |
|
625 |
(case Ctr_Sugar.ctr_sugar_of_case ctxt c of |
|
626 |
SOME {ctrs, ...} => |
|
627 |
(case possible_index_of_singleton_case (fst (split_last args)) of |
|
628 |
SOME i => |
|
629 |
let |
|
630 |
val constr_names = map (fst o dest_Const) ctrs |
|
631 |
val (Ts, _) = strip_type T |
|
632 |
val T' = List.last Ts |
|
633 |
in SOME (List.last args, T', i, nth args i, nth constr_names i) end |
|
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
634 |
| NONE => NONE) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
635 |
| NONE => NONE) |
60158 | 636 |
| NONE => NONE) |
637 |
end |
|
638 |
||
60752 | 639 |
fun tac ctxt [] = |
640 |
resolve_tac ctxt [set_singleton] 1 ORELSE |
|
641 |
resolve_tac ctxt [inst_Collect_mem_eq] 1 |
|
60158 | 642 |
| tac ctxt (If :: cont) = |
62381 | 643 |
Splitter.split_tac ctxt @{thms if_split} 1 |
60752 | 644 |
THEN resolve_tac ctxt @{thms conjI} 1 |
645 |
THEN resolve_tac ctxt @{thms impI} 1 |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
646 |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => |
60158 | 647 |
CONVERSION (right_hand_set_comprehension_conv (K |
648 |
(HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv |
|
649 |
then_conv |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
650 |
rewr_conv' @{lemma "(True \<and> P) = P" by simp})) ctxt') 1) ctxt 1 |
60158 | 651 |
THEN tac ctxt cont |
60752 | 652 |
THEN resolve_tac ctxt @{thms impI} 1 |
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
653 |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => |
60158 | 654 |
CONVERSION (right_hand_set_comprehension_conv (K |
655 |
(HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
656 |
then_conv rewr_conv' @{lemma "(False \<and> P) = False" by simp})) ctxt') 1) ctxt 1 |
60752 | 657 |
THEN resolve_tac ctxt [set_Nil_I] 1 |
60158 | 658 |
| tac ctxt (Case (T, i) :: cont) = |
659 |
let |
|
660 |
val SOME {injects, distincts, case_thms, split, ...} = |
|
661 |
Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T)) |
|
662 |
in |
|
663 |
(* do case distinction *) |
|
664 |
Splitter.split_tac ctxt [split] 1 |
|
665 |
THEN EVERY (map_index (fn (i', _) => |
|
60752 | 666 |
(if i' < length case_thms - 1 then resolve_tac ctxt @{thms conjI} 1 else all_tac) |
667 |
THEN REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1) |
|
668 |
THEN resolve_tac ctxt @{thms impI} 1 |
|
60158 | 669 |
THEN (if i' = i then |
670 |
(* continue recursively *) |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
671 |
Subgoal.FOCUS (fn {prems, context = ctxt', ...} => |
60158 | 672 |
CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K |
673 |
((HOLogic.conj_conv |
|
674 |
(HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv |
|
675 |
(Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects)))) |
|
676 |
Conv.all_conv) |
|
677 |
then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq)) |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
678 |
then_conv conjunct_assoc_conv)) ctxt' |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
679 |
then_conv |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
680 |
(HOLogic.Trueprop_conv |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
681 |
(HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') => |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
682 |
Conv.repeat_conv |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
683 |
(all_but_last_exists_conv |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
684 |
(K (rewr_conv' |
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
685 |
@{lemma "(\<exists>x. x = t \<and> P x) = P t" by simp})) ctxt'')) ctxt')))) 1) ctxt 1 |
60158 | 686 |
THEN tac ctxt cont |
687 |
else |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
688 |
Subgoal.FOCUS (fn {prems, context = ctxt', ...} => |
60158 | 689 |
CONVERSION |
690 |
(right_hand_set_comprehension_conv (K |
|
691 |
(HOLogic.conj_conv |
|
692 |
((HOLogic.eq_conv Conv.all_conv |
|
693 |
(rewr_conv' (List.last prems))) then_conv |
|
694 |
(Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts))) |
|
695 |
Conv.all_conv then_conv |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
696 |
(rewr_conv' @{lemma "(False \<and> P) = False" by simp}))) ctxt' then_conv |
60158 | 697 |
HOLogic.Trueprop_conv |
698 |
(HOLogic.eq_conv Conv.all_conv |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
699 |
(Collect_conv (fn (_, ctxt'') => |
60158 | 700 |
Conv.repeat_conv |
701 |
(Conv.bottom_conv |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
702 |
(K (rewr_conv' @{lemma "(\<exists>x. P) = P" by simp})) ctxt'')) ctxt'))) 1) ctxt 1 |
60752 | 703 |
THEN resolve_tac ctxt [set_Nil_I] 1)) case_thms) |
60158 | 704 |
end |
705 |
||
706 |
in |
|
707 |
||
708 |
fun simproc ctxt redex = |
|
709 |
let |
|
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
710 |
fun make_inner_eqs bound_vs Tis eqs t = |
60158 | 711 |
(case dest_case ctxt t of |
54404
9f0f1152c875
port list comprehension simproc to 'Ctr_Sugar' abstraction
blanchet
parents:
54295
diff
changeset
|
712 |
SOME (x, T, i, cont, constr_name) => |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
713 |
let |
52131 | 714 |
val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont) |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
715 |
val x' = incr_boundvars (length vs) x |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
716 |
val eqs' = map (incr_boundvars (length vs)) eqs |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
717 |
val constr_t = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
718 |
list_comb |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
719 |
(Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0)) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
720 |
val constr_eq = Const (@{const_name HOL.eq}, T --> T --> @{typ bool}) $ constr_t $ x' |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
721 |
in |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
722 |
make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
723 |
end |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
724 |
| NONE => |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
725 |
(case dest_if t of |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
726 |
SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
727 |
| NONE => |
60158 | 728 |
if null eqs then NONE (*no rewriting, nothing to be done*) |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
729 |
else |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
730 |
let |
60156 | 731 |
val Type (@{type_name list}, [rT]) = fastype_of1 (map snd bound_vs, t) |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
732 |
val pat_eq = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
733 |
(case try dest_singleton_list t of |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
734 |
SOME t' => |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
735 |
Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool}) $ |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
736 |
Bound (length bound_vs) $ t' |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
737 |
| NONE => |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
738 |
Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool}) $ |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
739 |
Bound (length bound_vs) $ (mk_set rT $ t)) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
740 |
val reverse_bounds = curry subst_bounds |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
741 |
((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)]) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
742 |
val eqs' = map reverse_bounds eqs |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
743 |
val pat_eq' = reverse_bounds pat_eq |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
744 |
val inner_t = |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
745 |
fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
746 |
(rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq') |
59582 | 747 |
val lhs = Thm.term_of redex |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
748 |
val rhs = HOLogic.mk_Collect ("x", rT, inner_t) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
749 |
val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
750 |
in |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
751 |
SOME |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
752 |
((Goal.prove ctxt [] [] rewrite_rule_t |
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
753 |
(fn {context = ctxt', ...} => tac ctxt' (rev Tis))) RS @{thm eq_reflection}) |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
754 |
end)) |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
755 |
in |
59582 | 756 |
make_inner_eqs [] [] [] (dest_set (Thm.term_of redex)) |
50422
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
757 |
end |
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
758 |
|
ee729dbd1b7f
avoid ML_file in large theory files to improve performance of dependency discovery of main HOL (approx. 1s CPU time) -- relevant for any application using it, e.g. small paper sessions;
wenzelm
parents:
50134
diff
changeset
|
759 |
end |
60158 | 760 |
|
761 |
end |
|
60758 | 762 |
\<close> |
41463
edbf0a86fb1c
adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents:
41372
diff
changeset
|
763 |
|
60159
879918f4ee0f
tuned -- avoid odd rebinding of "ctxt" and "context";
wenzelm
parents:
60158
diff
changeset
|
764 |
simproc_setup list_to_set_comprehension ("set xs") = |
60758 | 765 |
\<open>K List_to_Set_Comprehension.simproc\<close> |
41463
edbf0a86fb1c
adding simproc to rewrite list comprehensions to set comprehensions; adopting proofs
bulwahn
parents:
41372
diff
changeset
|
766 |
|
46133
d9fe85d3d2cd
incorporated canonical fold combinator on lists into body of List theory; refactored passages on List.fold(l/r)
haftmann
parents:
46125
diff
changeset
|
767 |
code_datatype set coset |
d9fe85d3d2cd
incorporated canonical fold combinator on lists into body of List theory; refactored passages on List.fold(l/r)
haftmann
parents:
46125
diff
changeset
|
768 |
hide_const (open) coset |
35115 | 769 |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
770 |
|
60758 | 771 |
subsubsection \<open>@{const Nil} and @{const Cons}\<close> |
21061
580dfc999ef6
added normal post setup; cleaned up "execution" constants
haftmann
parents:
21046
diff
changeset
|
772 |
|
580dfc999ef6
added normal post setup; cleaned up "execution" constants
haftmann
parents:
21046
diff
changeset
|
773 |
lemma not_Cons_self [simp]: |
580dfc999ef6
added normal post setup; cleaned up "execution" constants
haftmann
parents:
21046
diff
changeset
|
774 |
"xs \<noteq> x # xs" |
13145 | 775 |
by (induct xs) auto |
13114 | 776 |
|
58807 | 777 |
lemma not_Cons_self2 [simp]: "x # xs \<noteq> xs" |
41697 | 778 |
by (rule not_Cons_self [symmetric]) |
13114 | 779 |
|
13142 | 780 |
lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)" |
13145 | 781 |
by (induct xs) auto |
13114 | 782 |
|
53689 | 783 |
lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])" |
784 |
by (cases xs) auto |
|
785 |
||
786 |
lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])" |
|
787 |
by (cases xs) auto |
|
788 |
||
13142 | 789 |
lemma length_induct: |
21061
580dfc999ef6
added normal post setup; cleaned up "execution" constants
haftmann
parents:
21046
diff
changeset
|
790 |
"(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs" |
53689 | 791 |
by (fact measure_induct) |
13114 | 792 |
|
37289 | 793 |
lemma list_nonempty_induct [consumes 1, case_names single cons]: |
794 |
assumes "xs \<noteq> []" |
|
795 |
assumes single: "\<And>x. P [x]" |
|
796 |
assumes cons: "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)" |
|
797 |
shows "P xs" |
|
60758 | 798 |
using \<open>xs \<noteq> []\<close> proof (induct xs) |
37289 | 799 |
case Nil then show ?case by simp |
800 |
next |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
801 |
case (Cons x xs) |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
802 |
show ?case |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
803 |
proof (cases xs) |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
804 |
case Nil |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
805 |
with single show ?thesis by simp |
37289 | 806 |
next |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
807 |
case Cons |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
808 |
show ?thesis |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
809 |
proof (rule cons) |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
810 |
from Cons show "xs \<noteq> []" by simp |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
811 |
with Cons.hyps show "P xs" . |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
812 |
qed |
37289 | 813 |
qed |
814 |
qed |
|
815 |
||
45714 | 816 |
lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X" |
817 |
by (auto intro!: inj_onI) |
|
13114 | 818 |
|
61629 | 819 |
lemma inj_on_Cons1 [simp]: "inj_on (op # x) A" |
820 |
by(simp add: inj_on_def) |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
821 |
|
60758 | 822 |
subsubsection \<open>@{const length}\<close> |
823 |
||
824 |
text \<open> |
|
61799 | 825 |
Needs to come before \<open>@\<close> because of theorem \<open>append_eq_append_conv\<close>. |
60758 | 826 |
\<close> |
13114 | 827 |
|
13142 | 828 |
lemma length_append [simp]: "length (xs @ ys) = length xs + length ys" |
13145 | 829 |
by (induct xs) auto |
13114 | 830 |
|
13142 | 831 |
lemma length_map [simp]: "length (map f xs) = length xs" |
13145 | 832 |
by (induct xs) auto |
13114 | 833 |
|
13142 | 834 |
lemma length_rev [simp]: "length (rev xs) = length xs" |
13145 | 835 |
by (induct xs) auto |
13114 | 836 |
|
13142 | 837 |
lemma length_tl [simp]: "length (tl xs) = length xs - 1" |
13145 | 838 |
by (cases xs) auto |
13114 | 839 |
|
13142 | 840 |
lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])" |
13145 | 841 |
by (induct xs) auto |
13114 | 842 |
|
13142 | 843 |
lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])" |
13145 | 844 |
by (induct xs) auto |
13114 | 845 |
|
23479 | 846 |
lemma length_pos_if_in_set: "x : set xs \<Longrightarrow> length xs > 0" |
847 |
by auto |
|
848 |
||
13114 | 849 |
lemma length_Suc_conv: |
13145 | 850 |
"(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)" |
851 |
by (induct xs) auto |
|
13142 | 852 |
|
14025 | 853 |
lemma Suc_length_conv: |
58807 | 854 |
"(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)" |
14208 | 855 |
apply (induct xs, simp, simp) |
14025 | 856 |
apply blast |
857 |
done |
|
858 |
||
25221
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
859 |
lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False" |
58807 | 860 |
by (induct xs) auto |
25221
5ded95dda5df
append/member: more light-weight way to declare authentic syntax;
wenzelm
parents:
25215
diff
changeset
|
861 |
|
26442
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
862 |
lemma list_induct2 [consumes 1, case_names Nil Cons]: |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
863 |
"length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow> |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
864 |
(\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys)) |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
865 |
\<Longrightarrow> P xs ys" |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
866 |
proof (induct xs arbitrary: ys) |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
867 |
case Nil then show ?case by simp |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
868 |
next |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
869 |
case (Cons x xs ys) then show ?case by (cases ys) simp_all |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
870 |
qed |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
871 |
|
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
872 |
lemma list_induct3 [consumes 2, case_names Nil Cons]: |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
873 |
"length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow> |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
874 |
(\<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)) |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
875 |
\<Longrightarrow> P xs ys zs" |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
876 |
proof (induct xs arbitrary: ys zs) |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
877 |
case Nil then show ?case by simp |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
878 |
next |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
879 |
case (Cons x xs ys zs) then show ?case by (cases ys, simp_all) |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
880 |
(cases zs, simp_all) |
57fb6a8b099e
restructuring; explicit case names for rule list_induct2
haftmann
parents:
26300
diff
changeset
|
881 |
qed |
13114 | 882 |
|
36154
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
883 |
lemma list_induct4 [consumes 3, case_names Nil Cons]: |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
884 |
"length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
885 |
P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow> |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
886 |
length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow> |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
887 |
P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws" |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
888 |
proof (induct xs arbitrary: ys zs ws) |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
889 |
case Nil then show ?case by simp |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
890 |
next |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
891 |
case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all) |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
892 |
qed |
11c6106d7787
Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35828
diff
changeset
|
893 |
|
64963 | 894 |
lemma list_induct2': |
22493
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
895 |
"\<lbrakk> P [] []; |
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
896 |
\<And>x xs. P (x#xs) []; |
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
897 |
\<And>y ys. P [] (y#ys); |
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
898 |
\<And>x xs y ys. P xs ys \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk> |
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
899 |
\<Longrightarrow> P xs ys" |
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
900 |
by (induct xs arbitrary: ys) (case_tac x, auto)+ |
db930e490fe5
added another rule for simultaneous induction, and lemmas for zip
krauss
parents:
22422
diff
changeset
|
901 |
|
55524
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents:
55473
diff
changeset
|
902 |
lemma list_all2_iff: |
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents:
55473
diff
changeset
|
903 |
"list_all2 P xs ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y)" |
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents:
55473
diff
changeset
|
904 |
by (induct xs ys rule: list_induct2') auto |
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents:
55473
diff
changeset
|
905 |
|
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
906 |
lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False" |
24349 | 907 |
by (rule Eq_FalseI) auto |
24037 | 908 |
|
60758 | 909 |
simproc_setup list_neq ("(xs::'a list) = ys") = \<open> |
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
910 |
(* |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
911 |
Reduces xs=ys to False if xs and ys cannot be of the same length. |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
912 |
This is the case if the atomic sublists of one are a submultiset |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
913 |
of those of the other list and there are fewer Cons's in one than the other. |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
914 |
*) |
24037 | 915 |
|
916 |
let |
|
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
917 |
|
29856 | 918 |
fun len (Const(@{const_name Nil},_)) acc = acc |
919 |
| len (Const(@{const_name Cons},_) $ _ $ xs) (ts,n) = len xs (ts,n+1) |
|
920 |
| len (Const(@{const_name append},_) $ xs $ ys) acc = len xs (len ys acc) |
|
921 |
| len (Const(@{const_name rev},_) $ xs) acc = len xs acc |
|
922 |
| len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc |
|
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
923 |
| len t (ts,n) = (t::ts,n); |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
924 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
925 |
val ss = simpset_of @{context}; |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
926 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
927 |
fun list_neq ctxt ct = |
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
928 |
let |
24037 | 929 |
val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct; |
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
930 |
val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0); |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
931 |
fun prove_neq() = |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
932 |
let |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
933 |
val Type(_,listT::_) = eqT; |
22994 | 934 |
val size = HOLogic.size_const listT; |
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
935 |
val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs); |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
936 |
val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len); |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
937 |
val thm = Goal.prove ctxt [] [] neq_len |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
938 |
(K (simp_tac (put_simpset ss ctxt) 1)); |
22633 | 939 |
in SOME (thm RS @{thm neq_if_length_neq}) end |
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
940 |
in |
23214 | 941 |
if m < n andalso submultiset (op aconv) (ls,rs) orelse |
942 |
n < m andalso submultiset (op aconv) (rs,ls) |
|
22143
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
943 |
then prove_neq() else NONE |
cf58486ca11b
Added simproc list_neq (prompted by an application)
nipkow
parents:
21911
diff
changeset
|
944 |
end; |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
945 |
in K list_neq end; |
60758 | 946 |
\<close> |
947 |
||
948 |
||
61799 | 949 |
subsubsection \<open>\<open>@\<close> -- append\<close> |
13114 | 950 |
|
63662 | 951 |
global_interpretation append: monoid append Nil |
952 |
proof |
|
953 |
fix xs ys zs :: "'a list" |
|
954 |
show "(xs @ ys) @ zs = xs @ (ys @ zs)" |
|
955 |
by (induct xs) simp_all |
|
956 |
show "xs @ [] = xs" |
|
957 |
by (induct xs) simp_all |
|
958 |
qed simp |
|
959 |
||
13142 | 960 |
lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)" |
63662 | 961 |
by (fact append.assoc) |
962 |
||
963 |
lemma append_Nil2: "xs @ [] = xs" |
|
964 |
by (fact append.right_neutral) |
|
3507 | 965 |
|
13142 | 966 |
lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])" |
13145 | 967 |
by (induct xs) auto |
13114 | 968 |
|
13142 | 969 |
lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])" |
13145 | 970 |
by (induct xs) auto |
13114 | 971 |
|
13142 | 972 |
lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])" |
13145 | 973 |
by (induct xs) auto |
13114 | 974 |
|
13142 | 975 |
lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])" |
13145 | 976 |
by (induct xs) auto |
13114 | 977 |
|
54147
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
blanchet
parents:
53954
diff
changeset
|
978 |
lemma append_eq_append_conv [simp]: |
58807 | 979 |
"length xs = length ys \<or> length us = length vs |
980 |
==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)" |
|
24526 | 981 |
apply (induct xs arbitrary: ys) |
14208 | 982 |
apply (case_tac ys, simp, force) |
983 |
apply (case_tac ys, force, simp) |
|
13145 | 984 |
done |
13142 | 985 |
|
24526 | 986 |
lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) = |
987 |
(EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)" |
|
988 |
apply (induct xs arbitrary: ys zs ts) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44635
diff
changeset
|
989 |
apply fastforce |
14495 | 990 |
apply(case_tac zs) |
991 |
apply simp |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44635
diff
changeset
|
992 |
apply fastforce |
14495 | 993 |
done |
994 |
||
34910
b23bd3ee4813
same_append_eq / append_same_eq are now used for simplifying induction rules.
berghofe
parents:
34064
diff
changeset
|
995 |
lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)" |
13145 | 996 |
by simp |
13142 | 997 |
|
998 |
lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)" |
|
13145 | 999 |
by simp |
13114 | 1000 |
|
34910
b23bd3ee4813
same_append_eq / append_same_eq are now used for simplifying induction rules.
berghofe
parents:
34064
diff
changeset
|
1001 |
lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)" |
13145 | 1002 |
by simp |
13114 | 1003 |
|
13142 | 1004 |
lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])" |
13145 | 1005 |
using append_same_eq [of _ _ "[]"] by auto |
3507 | 1006 |
|
13142 | 1007 |
lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])" |
13145 | 1008 |
using append_same_eq [of "[]"] by auto |
13114 | 1009 |
|
63662 | 1010 |
lemma hd_Cons_tl: "xs \<noteq> [] ==> hd xs # tl xs = xs" |
1011 |
by (fact list.collapse) |
|
13114 | 1012 |
|
13142 | 1013 |
lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)" |
13145 | 1014 |
by (induct xs) auto |
13114 | 1015 |
|
13142 | 1016 |
lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs" |
13145 | 1017 |
by (simp add: hd_append split: list.split) |
13114 | 1018 |
|
13142 | 1019 |
lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)" |
13145 | 1020 |
by (simp split: list.split) |
13114 | 1021 |
|
13142 | 1022 |
lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys" |
13145 | 1023 |
by (simp add: tl_append split: list.split) |
13114 | 1024 |
|
1025 |
||
14300 | 1026 |
lemma Cons_eq_append_conv: "x#xs = ys@zs = |
1027 |
(ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))" |
|
1028 |
by(cases ys) auto |
|
1029 |
||
15281 | 1030 |
lemma append_eq_Cons_conv: "(ys@zs = x#xs) = |
1031 |
(ys = [] & zs = x#xs | (EX ys'. ys = x#ys' & ys'@zs = xs))" |
|
1032 |
by(cases ys) auto |
|
1033 |
||
63173 | 1034 |
lemma longest_common_prefix: |
1035 |
"\<exists>ps xs' ys'. xs = ps @ xs' \<and> ys = ps @ ys' |
|
1036 |
\<and> (xs' = [] \<or> ys' = [] \<or> hd xs' \<noteq> hd ys')" |
|
1037 |
by (induct xs ys rule: list_induct2') |
|
1038 |
(blast, blast, blast, |
|
1039 |
metis (no_types, hide_lams) append_Cons append_Nil list.sel(1)) |
|
14300 | 1040 |
|
61799 | 1041 |
text \<open>Trivial rules for solving \<open>@\<close>-equations automatically.\<close> |
13114 | 1042 |
|
1043 |
lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys" |
|
13145 | 1044 |
by simp |
13114 | 1045 |
|
13142 | 1046 |
lemma Cons_eq_appendI: |
13145 | 1047 |
"[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs" |
1048 |
by (drule sym) simp |
|
13114 | 1049 |
|
13142 | 1050 |
lemma append_eq_appendI: |
13145 | 1051 |
"[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us" |
1052 |
by (drule sym) simp |
|
13114 | 1053 |
|
1054 |
||
60758 | 1055 |
text \<open> |
13145 | 1056 |
Simplification procedure for all list equalities. |
61799 | 1057 |
Currently only tries to rearrange \<open>@\<close> to see if |
13145 | 1058 |
- both lists end in a singleton list, |
1059 |
- or both lists end in the same list. |
|
60758 | 1060 |
\<close> |
1061 |
||
1062 |
simproc_setup list_eq ("(xs::'a list) = ys") = \<open> |
|
13462 | 1063 |
let |
43594 | 1064 |
fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) = |
1065 |
(case xs of Const (@{const_name Nil}, _) => cons | _ => last xs) |
|
1066 |
| last (Const(@{const_name append},_) $ _ $ ys) = last ys |
|
1067 |
| last t = t; |
|
64963 | 1068 |
|
43594 | 1069 |
fun list1 (Const(@{const_name Cons},_) $ _ $ Const(@{const_name Nil},_)) = true |
1070 |
| list1 _ = false; |
|
64963 | 1071 |
|
43594 | 1072 |
fun butlast ((cons as Const(@{const_name Cons},_) $ x) $ xs) = |
1073 |
(case xs of Const (@{const_name Nil}, _) => xs | _ => cons $ butlast xs) |
|
1074 |
| butlast ((app as Const (@{const_name append}, _) $ xs) $ ys) = app $ butlast ys |
|
1075 |
| butlast xs = Const(@{const_name Nil}, fastype_of xs); |
|
64963 | 1076 |
|
43594 | 1077 |
val rearr_ss = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
1078 |
simpset_of (put_simpset HOL_basic_ss @{context} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
1079 |
addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]); |
64963 | 1080 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
1081 |
fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) = |
13462 | 1082 |
let |
43594 | 1083 |
val lastl = last lhs and lastr = last rhs; |
1084 |
fun rearr conv = |
|
1085 |
let |
|
1086 |
val lhs1 = butlast lhs and rhs1 = butlast rhs; |
|
1087 |
val Type(_,listT::_) = eqT |
|
1088 |
val appT = [listT,listT] ---> listT |
|
1089 |
val app = Const(@{const_name append},appT) |
|
1090 |
val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) |
|
1091 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2)); |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
1092 |
val thm = Goal.prove ctxt [] [] eq |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51678
diff
changeset
|
1093 |
(K (simp_tac (put_simpset rearr_ss ctxt) 1)); |
43594 | 1094 |
in SOME ((conv RS (thm RS trans)) RS eq_reflection) end; |
1095 |
in |
|
1096 |
if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv} |
|
1097 |
else if lastl aconv lastr then rearr @{thm append_same_eq} |
|
1098 |
else NONE |
|
1099 |
end; |
|
59582 | 1100 |
in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end; |
60758 | 1101 |
\<close> |
1102 |
||
1103 |
||
1104 |
subsubsection \<open>@{const map}\<close> |
|
13114 | 1105 |
|
58807 | 1106 |
lemma hd_map: "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)" |
1107 |
by (cases xs) simp_all |
|
1108 |
||
1109 |
lemma map_tl: "map f (tl xs) = tl (map f xs)" |
|
1110 |
by (cases xs) simp_all |
|
40210
aee7ef725330
sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
haftmann
parents:
40195
diff
changeset
|
1111 |
|
13142 | 1112 |
lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs" |
13145 | 1113 |
by (induct xs) simp_all |
13114 | 1114 |
|
13142 | 1115 |
lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)" |
13145 | 1116 |
by (rule ext, induct_tac xs) auto |
13114 | 1117 |
|
13142 | 1118 |
lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys" |
13145 | 1119 |
by (induct xs) auto |
13114 | 1120 |
|
33639
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
1121 |
lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs" |
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
1122 |
by (induct xs) auto |
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents:
33593
diff
changeset
|
1123 |
|
35208 | 1124 |
lemma map_comp_map[simp]: "((map f) o (map g)) = map(f o g)" |
58807 | 1125 |
by (rule ext) simp |
35208 | 1126 |
|
13142 | 1127 |
lemma rev_map: "rev (map f xs) = map f (rev xs)" |
13145 | 1128 |
by (induct xs) auto |
13114 | 1129 |
|
13737 | 1130 |
lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)" |
1131 |
by (induct xs) auto |
|
1132 |
||
44013
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents:
43594
diff
changeset
|
1133 |
lemma map_cong [fundef_cong]: |
40122
1d8ad2ff3e01
dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents:
40077
diff
changeset
|
1134 |
"xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys" |
58807 | 1135 |
by simp |
13114 | 1136 |
|
13142 | 1137 |
lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])" |
13145 | 1138 |
by (cases xs) auto |
13114 | 1139 |
|
13142 | 1140 |
lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])" |
13145 | 1141 |
by (cases xs) auto |
13114 | 1142 |
|
18447 | 1143 |
lemma map_eq_Cons_conv: |
58807 | 1144 |
"(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)" |
13145 | 1145 |
by (cases xs) auto |
13114 | 1146 |
|
18447 | 1147 |
lemma Cons_eq_map_conv: |
58807 | 1148 |
"(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)" |
14025 | 1149 |
by (cases ys) auto |
1150 |
||
18447 | 1151 |
lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1] |
1152 |
lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1] |
|
1153 |
declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] |
|
1154 |
||
14111 | 1155 |
lemma ex_map_conv: |
1156 |
"(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)" |
|
18447 | 1157 |
by(induct ys, auto simp add: Cons_eq_map_conv) |
14111 | 1158 |
|
15110
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1159 |
lemma map_eq_imp_length_eq: |
35510 | 1160 |
assumes "map f xs = map g ys" |
26734 | 1161 |
shows "length xs = length ys" |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
1162 |
using assms |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
1163 |
proof (induct ys arbitrary: xs) |
26734 | 1164 |
case Nil then show ?case by simp |
1165 |
next |
|
1166 |
case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto |
|
35510 | 1167 |
from Cons xs have "map f zs = map g ys" by simp |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53017
diff
changeset
|
1168 |
with Cons have "length zs = length ys" by blast |
26734 | 1169 |
with xs show ?case by simp |
1170 |
qed |
|
64963 | 1171 |
|
15110
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1172 |
lemma map_inj_on: |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1173 |
"[| map f xs = map f ys; inj_on f (set xs Un set ys) |] |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1174 |
==> xs = ys" |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1175 |
apply(frule map_eq_imp_length_eq) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1176 |
apply(rotate_tac -1) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1177 |
apply(induct rule:list_induct2) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1178 |
apply simp |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1179 |
apply(simp) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1180 |
apply (blast intro:sym) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1181 |
done |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1182 |
|
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1183 |
lemma inj_on_map_eq_map: |
58807 | 1184 |
"inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)" |
15110
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1185 |
by(blast dest:map_inj_on) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1186 |
|
13114 | 1187 |
lemma map_injective: |
58807 | 1188 |
"map f xs = map f ys ==> inj f ==> xs = ys" |
24526 | 1189 |
by (induct ys arbitrary: xs) (auto dest!:injD) |
13114 | 1190 |
|
14339 | 1191 |
lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)" |
1192 |
by(blast dest:map_injective) |
|
1193 |
||
13114 | 1194 |
lemma inj_mapI: "inj f ==> inj (map f)" |
17589 | 1195 |
by (iprover dest: map_injective injD intro: inj_onI) |
13114 | 1196 |
|
1197 |
lemma inj_mapD: "inj (map f) ==> inj f" |
|
64966
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents:
64963
diff
changeset
|
1198 |
apply (unfold inj_def) |
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents:
64963
diff
changeset
|
1199 |
apply clarify |
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents:
64963
diff
changeset
|
1200 |
apply (erule_tac x = "[x]" in allE) |
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents:
64963
diff
changeset
|
1201 |
apply (erule_tac x = "[y]" in allE) |
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents:
64963
diff
changeset
|
1202 |
apply auto |
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
wenzelm
parents:
64963
diff
changeset
|
1203 |
done |
13114 | 1204 |
|
14339 | 1205 |
lemma inj_map[iff]: "inj (map f) = inj f" |
13145 | 1206 |
by (blast dest: inj_mapD intro: inj_mapI) |
13114 | 1207 |
|
15303 | 1208 |
lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A" |
1209 |
apply(rule inj_onI) |
|
1210 |
apply(erule map_inj_on) |
|
1211 |
apply(blast intro:inj_onI dest:inj_onD) |
|
1212 |
done |
|
1213 |
||
14343 | 1214 |
lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs" |
1215 |
by (induct xs, auto) |
|
13114 | 1216 |
|
14402
4201e1916482
moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents:
14395
diff
changeset
|
1217 |
lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs" |
4201e1916482
moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents:
14395
diff
changeset
|
1218 |
by (induct xs) auto |
4201e1916482
moved lemmas from MicroJava/Comp/AuxLemmas.thy to List.thy
nipkow
parents:
14395
diff
changeset
|
1219 |
|
15110
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1220 |
lemma map_fst_zip[simp]: |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1221 |
"length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs" |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1222 |
by (induct rule:list_induct2, simp_all) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1223 |
|
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1224 |
lemma map_snd_zip[simp]: |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1225 |
"length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys" |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1226 |
by (induct rule:list_induct2, simp_all) |
78b5636eabc7
Added a number of new thms and the new function remove1
nipkow
parents:
15072
diff
changeset
|
1227 |
|
55467
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents:
55466
diff
changeset
|
1228 |
functor map: map |
47122 | 1229 |
by (simp_all add: id_def) |
1230 |
||
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
1231 |
declare map.id [simp] |
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
1232 |
|
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49808
diff
changeset
|
1233 |
|
60758 | 1234 |
subsubsection \<open>@{const rev}\<close> |
13114 | 1235 |
|
13142 | 1236 |
lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs" |
13145 | 1237 |
by (induct xs) auto |
13114 | 1238 |
|
13142 | 1239 |
lemma rev_rev_ident [simp]: "rev (rev xs) = xs" |
13145 | 1240 |
by (induct xs) auto |
13114 | 1241 |
|
15870 | 1242 |
lemma rev_swap: "(rev xs = ys) = (xs = rev ys)" |
1243 |
by auto |
|
1244 |
||
13142 | 1245 |
lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])" |
13145 | 1246 |
by (induct xs) auto |
13114 | 1247 |
|
13142 | 1248 |
lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])" |
13145 | 1249 |
by (induct xs) auto |
13114 | 1250 |
|
15870 | 1251 |
lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])" |
1252 |
by (cases xs) auto |
|
1253 |
||
1254 |
lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])" |
|
1255 |
by (cases xs) auto |
|
1256 |
||
54147
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
blanchet
parents:
53954
diff
changeset
|
1257 |
lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)" |
21061
580dfc999ef6
added normal post setup; cleaned up "execution" constants
haftmann
parents:
21046
diff
changeset
|
1258 |
apply (induct xs arbitrary: ys, force) |
14208 | 1259 |
apply (case_tac ys, simp, force) |
13145 | 1260 |
done |
13114 | 1261 |
|
15439 | 1262 |
lemma inj_on_rev[iff]: "inj_on rev A" |
1263 |
by(simp add:inj_on_def) |
|
1264 |
||
13366 | 1265 |
lemma rev_induct [case_names Nil snoc]: |
1266 |
"[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs" |
|
15489
d136af442665
Replaced application of subst by simplesubst in proof of rev_induct
berghofe
parents:
15439
diff
changeset
|
1267 |
apply(simplesubst rev_rev_ident[symmetric]) |
13145 | 1268 |
apply(rule_tac list = "rev xs" in list.induct, simp_all) |
1269 |
done |
|
13114 | 1270 |
|
13366 | 1271 |
lemma rev_exhaust [case_names Nil snoc]: |
1272 |
"(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P" |
|
13145 | 1273 |
by (induct xs rule: rev_induct) auto |
13114 | 1274 |
|
13366 | 1275 |
lemmas rev_cases = rev_exhaust |
1276 |
||
57577 | 1277 |
lemma rev_nonempty_induct [consumes 1, case_names single snoc]: |
1278 |
assumes "xs \<noteq> []" |
|
1279 |
and single: "\<And>x. P [x]" |
|
1280 |
and snoc': "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (xs@[x])" |
|
1281 |
shows "P xs" |
|
60758 | 1282 |
using \<open>xs \<noteq> []\<close> proof (induct xs rule: rev_induct) |
57577 | 1283 |
case (snoc x xs) then show ?case |
1284 |
proof (cases xs) |
|
1285 |
case Nil thus ?thesis by (simp add: single) |
|
1286 |
next |
|
1287 |
case Cons with snoc show ?thesis by (fastforce intro!: snoc') |
|
1288 |
qed |
|
1289 |
qed simp |
|
1290 |
||
18423 | 1291 |
lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])" |
d7859164447f
new lemmas
n
|