Permalink
Cannot retrieve contributors at this time
Name already in use
A tag already exists with the provided branch name. Many Git commands accept both tag and branch names, so creating this branch may cause unexpected behavior. Are you sure you want to create this branch?
isabelle/src/HOL/List.thy
Go to fileThis commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
8354 lines (6697 sloc)
303 KB
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
(* Title: HOL/List.thy | |
Author: Tobias Nipkow; proofs tidied by LCP | |
*) | |
section \<open>The datatype of finite lists\<close> | |
theory List | |
imports Sledgehammer Lifting_Set | |
begin | |
datatype (set: 'a) list = | |
Nil ("[]") | |
| Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) | |
for | |
map: map | |
rel: list_all2 | |
pred: list_all | |
where | |
"tl [] = []" | |
datatype_compat list | |
lemma [case_names Nil Cons, cases type: list]: | |
\<comment> \<open>for backward compatibility -- names of variables differ\<close> | |
"(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P" | |
by (rule list.exhaust) | |
lemma [case_names Nil Cons, induct type: list]: | |
\<comment> \<open>for backward compatibility -- names of variables differ\<close> | |
"P [] \<Longrightarrow> (\<And>a list. P list \<Longrightarrow> P (a # list)) \<Longrightarrow> P list" | |
by (rule list.induct) | |
text \<open>Compatibility:\<close> | |
setup \<open>Sign.mandatory_path "list"\<close> | |
lemmas inducts = list.induct | |
lemmas recs = list.rec | |
lemmas cases = list.case | |
setup \<open>Sign.parent_path\<close> | |
lemmas set_simps = list.set (* legacy *) | |
syntax | |
\<comment> \<open>list Enumeration\<close> | |
"_list" :: "args => 'a list" ("[(_)]") | |
translations | |
"[x, xs]" == "x#[xs]" | |
"[x]" == "x#[]" | |
subsection \<open>Basic list processing functions\<close> | |
primrec (nonexhaustive) last :: "'a list \<Rightarrow> 'a" where | |
"last (x # xs) = (if xs = [] then x else last xs)" | |
primrec butlast :: "'a list \<Rightarrow> 'a list" where | |
"butlast [] = []" | | |
"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)" | |
lemma set_rec: "set xs = rec_list {} (\<lambda>x _. insert x) xs" | |
by (induct xs) auto | |
definition coset :: "'a list \<Rightarrow> 'a set" where | |
[simp]: "coset xs = - set xs" | |
primrec append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where | |
append_Nil: "[] @ ys = ys" | | |
append_Cons: "(x#xs) @ ys = x # xs @ ys" | |
primrec rev :: "'a list \<Rightarrow> 'a list" where | |
"rev [] = []" | | |
"rev (x # xs) = rev xs @ [x]" | |
primrec filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"filter P [] = []" | | |
"filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" | |
text \<open>Special input syntax for filter:\<close> | |
syntax (ASCII) | |
"_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])") | |
syntax | |
"_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_\<leftarrow>_ ./ _])") | |
translations | |
"[x<-xs . P]" \<rightharpoonup> "CONST filter (\<lambda>x. P) xs" | |
primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where | |
fold_Nil: "fold f [] = id" | | |
fold_Cons: "fold f (x # xs) = fold f xs \<circ> f x" | |
primrec foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where | |
foldr_Nil: "foldr f [] = id" | | |
foldr_Cons: "foldr f (x # xs) = f x \<circ> foldr f xs" | |
primrec foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where | |
foldl_Nil: "foldl f a [] = a" | | |
foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs" | |
primrec concat:: "'a list list \<Rightarrow> 'a list" where | |
"concat [] = []" | | |
"concat (x # xs) = x @ concat xs" | |
primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
drop_Nil: "drop n [] = []" | | |
drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)" | |
\<comment> \<open>Warning: simpset does not contain this definition, but separate | |
theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close> | |
primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
take_Nil:"take n [] = []" | | |
take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)" | |
\<comment> \<open>Warning: simpset does not contain this definition, but separate | |
theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close> | |
primrec (nonexhaustive) nth :: "'a list => nat => 'a" (infixl "!" 100) where | |
nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)" | |
\<comment> \<open>Warning: simpset does not contain this definition, but separate | |
theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close> | |
primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where | |
"list_update [] i v = []" | | |
"list_update (x # xs) i v = | |
(case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)" | |
nonterminal lupdbinds and lupdbind | |
syntax | |
"_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)") | |
"" :: "lupdbind => lupdbinds" ("_") | |
"_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _") | |
"_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [1000,0] 900) | |
translations | |
"_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs" | |
"xs[i:=x]" == "CONST list_update xs i x" | |
primrec takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"takeWhile P [] = []" | | |
"takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])" | |
primrec dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"dropWhile P [] = []" | | |
"dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)" | |
primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where | |
"zip xs [] = []" | | |
zip_Cons: "zip xs (y # ys) = | |
(case xs of [] \<Rightarrow> [] | z # zs \<Rightarrow> (z, y) # zip zs ys)" | |
\<comment> \<open>Warning: simpset does not contain this definition, but separate | |
theorems for \<open>xs = []\<close> and \<open>xs = z # zs\<close>\<close> | |
abbreviation map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where | |
"map2 f xs ys \<equiv> map (\<lambda>(x,y). f x y) (zip xs ys)" | |
primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where | |
"product [] _ = []" | | |
"product (x#xs) ys = map (Pair x) ys @ product xs ys" | |
hide_const (open) product | |
primrec product_lists :: "'a list list \<Rightarrow> 'a list list" where | |
"product_lists [] = [[]]" | | |
"product_lists (xs # xss) = concat (map (\<lambda>x. map (Cons x) (product_lists xss)) xs)" | |
primrec upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where | |
upt_0: "[i..<0] = []" | | |
upt_Suc: "[i..<(Suc j)] = (if i \<le> j then [i..<j] @ [j] else [])" | |
definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"insert x xs = (if x \<in> set xs then xs else x # xs)" | |
definition union :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"union = fold insert" | |
hide_const (open) insert union | |
hide_fact (open) insert_def union_def | |
primrec find :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a option" where | |
"find _ [] = None" | | |
"find P (x#xs) = (if P x then Some x else find P xs)" | |
text \<open>In the context of multisets, \<open>count_list\<close> is equivalent to | |
\<^term>\<open>count \<circ> mset\<close> and it it advisable to use the latter.\<close> | |
primrec count_list :: "'a list \<Rightarrow> 'a \<Rightarrow> nat" where | |
"count_list [] y = 0" | | |
"count_list (x#xs) y = (if x=y then count_list xs y + 1 else count_list xs y)" | |
definition | |
"extract" :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> ('a list * 'a * 'a list) option" | |
where "extract P xs = | |
(case dropWhile (Not \<circ> P) xs of | |
[] \<Rightarrow> None | | |
y#ys \<Rightarrow> Some(takeWhile (Not \<circ> P) xs, y, ys))" | |
hide_const (open) "extract" | |
primrec those :: "'a option list \<Rightarrow> 'a list option" | |
where | |
"those [] = Some []" | | |
"those (x # xs) = (case x of | |
None \<Rightarrow> None | |
| Some y \<Rightarrow> map_option (Cons y) (those xs))" | |
primrec remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"remove1 x [] = []" | | |
"remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)" | |
primrec removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"removeAll x [] = []" | | |
"removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)" | |
primrec distinct :: "'a list \<Rightarrow> bool" where | |
"distinct [] \<longleftrightarrow> True" | | |
"distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs" | |
fun successively :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where | |
"successively P [] = True" | | |
"successively P [x] = True" | | |
"successively P (x # y # xs) = (P x y \<and> successively P (y#xs))" | |
definition distinct_adj where | |
"distinct_adj = successively (\<noteq>)" | |
primrec remdups :: "'a list \<Rightarrow> 'a list" where | |
"remdups [] = []" | | |
"remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)" | |
fun remdups_adj :: "'a list \<Rightarrow> 'a list" where | |
"remdups_adj [] = []" | | |
"remdups_adj [x] = [x]" | | |
"remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))" | |
primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where | |
replicate_0: "replicate 0 x = []" | | |
replicate_Suc: "replicate (Suc n) x = x # replicate n x" | |
text \<open> | |
Function \<open>size\<close> is overloaded for all datatypes. Users may | |
refer to the list version as \<open>length\<close>.\<close> | |
abbreviation length :: "'a list \<Rightarrow> nat" where | |
"length \<equiv> size" | |
definition enumerate :: "nat \<Rightarrow> 'a list \<Rightarrow> (nat \<times> 'a) list" where | |
enumerate_eq_zip: "enumerate n xs = zip [n..<n + length xs] xs" | |
primrec rotate1 :: "'a list \<Rightarrow> 'a list" where | |
"rotate1 [] = []" | | |
"rotate1 (x # xs) = xs @ [x]" | |
definition rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"rotate n = rotate1 ^^ n" | |
definition nths :: "'a list => nat set => 'a list" where | |
"nths xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))" | |
primrec subseqs :: "'a list \<Rightarrow> 'a list list" where | |
"subseqs [] = [[]]" | | |
"subseqs (x#xs) = (let xss = subseqs xs in map (Cons x) xss @ xss)" | |
primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where | |
"n_lists 0 xs = [[]]" | | |
"n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))" | |
hide_const (open) n_lists | |
function splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
"splice [] ys = ys" | | |
"splice (x#xs) ys = x # splice ys xs" | |
by pat_completeness auto | |
termination | |
by(relation "measure(\<lambda>(xs,ys). size xs + size ys)") auto | |
function shuffles where | |
"shuffles [] ys = {ys}" | |
| "shuffles xs [] = {xs}" | |
| "shuffles (x # xs) (y # ys) = (#) x ` shuffles xs (y # ys) \<union> (#) y ` shuffles (x # xs) ys" | |
by pat_completeness simp_all | |
termination by lexicographic_order | |
text\<open>Use only if you cannot use \<^const>\<open>Min\<close> instead:\<close> | |
fun min_list :: "'a::ord list \<Rightarrow> 'a" where | |
"min_list (x # xs) = (case xs of [] \<Rightarrow> x | _ \<Rightarrow> min x (min_list xs))" | |
text\<open>Returns first minimum:\<close> | |
fun arg_min_list :: "('a \<Rightarrow> ('b::linorder)) \<Rightarrow> 'a list \<Rightarrow> 'a" where | |
"arg_min_list f [x] = x" | | |
"arg_min_list f (x#y#zs) = (let m = arg_min_list f (y#zs) in if f x \<le> f m then x else m)" | |
text\<open> | |
\begin{figure}[htbp] | |
\fbox{ | |
\begin{tabular}{l} | |
@{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\ | |
@{lemma "length [a,b,c] = 3" by simp}\\ | |
@{lemma "set [a,b,c] = {a,b,c}" by simp}\\ | |
@{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\ | |
@{lemma "rev [a,b,c] = [c,b,a]" by simp}\\ | |
@{lemma "hd [a,b,c,d] = a" by simp}\\ | |
@{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\ | |
@{lemma "last [a,b,c,d] = d" by simp}\\ | |
@{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\ | |
@{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\ | |
@{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\ | |
@{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\ | |
@{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\ | |
@{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\ | |
@{lemma "successively (\<noteq>) [True,False,True,False]" by simp}\\ | |
@{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\ | |
@{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\ | |
@{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\ | |
@{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\ | |
@{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\ | |
@{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\ | |
@{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\ | |
@{lemma "shuffles [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]}" | |
by (simp add: insert_commute)}\\ | |
@{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\ | |
@{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\ | |
@{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\ | |
@{lemma "drop 6 [a,b,c,d] = []" by simp}\\ | |
@{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\ | |
@{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\ | |
@{lemma "distinct [2,0,1::nat]" by simp}\\ | |
@{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\ | |
@{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\ | |
@{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\ | |
@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\ | |
@{lemma "List.union [2,3,4] [0::int,1,2] = [4,3,0,1,2]" by (simp add: List.insert_def List.union_def)}\\ | |
@{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\ | |
@{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\ | |
@{lemma "count_list [0,1,0,2::int] 0 = 2" by (simp)}\\ | |
@{lemma "List.extract (%i::int. i>0) [0,0] = None" by(simp add: extract_def)}\\ | |
@{lemma "List.extract (%i::int. i>0) [0,1,0,2] = Some([0], 1, [0,2])" by(simp add: extract_def)}\\ | |
@{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\ | |
@{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\ | |
@{lemma "nth [a,b,c,d] 2 = c" by simp}\\ | |
@{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\ | |
@{lemma "nths [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:nths_def)}\\ | |
@{lemma "subseqs [a,b] = [[a, b], [a], [b], []]" by simp}\\ | |
@{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)}\\ | |
@{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\ | |
@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\ | |
@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\ | |
@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\ | |
@{lemma "min_list [3,1,-2::int] = -2" by (simp)}\\ | |
@{lemma "arg_min_list (\<lambda>i. i*i) [3,-1,1,-2::int] = -1" by (simp)} | |
\end{tabular}} | |
\caption{Characteristic examples} | |
\label{fig:Characteristic} | |
\end{figure} | |
Figure~\ref{fig:Characteristic} shows characteristic examples | |
that should give an intuitive understanding of the above functions. | |
\<close> | |
text\<open>The following simple sort(ed) functions are intended for proofs, | |
not for efficient implementations.\<close> | |
text \<open>A sorted predicate w.r.t. a relation:\<close> | |
fun sorted_wrt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where | |
"sorted_wrt P [] = True" | | |
"sorted_wrt P (x # ys) = ((\<forall>y \<in> set ys. P x y) \<and> sorted_wrt P ys)" | |
text \<open>A class-based sorted predicate:\<close> | |
context linorder | |
begin | |
abbreviation sorted :: "'a list \<Rightarrow> bool" where | |
"sorted \<equiv> sorted_wrt (\<le>)" | |
lemma sorted_simps: "sorted [] = True" "sorted (x # ys) = ((\<forall>y \<in> set ys. x\<le>y) \<and> sorted ys)" | |
by auto | |
lemma strict_sorted_simps: "sorted_wrt (<) [] = True" "sorted_wrt (<) (x # ys) = ((\<forall>y \<in> set ys. x<y) \<and> sorted_wrt (<) ys)" | |
by auto | |
primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where | |
"insort_key f x [] = [x]" | | |
"insort_key f x (y#ys) = | |
(if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))" | |
definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where | |
"sort_key f xs = foldr (insort_key f) xs []" | |
definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where | |
"insort_insert_key f x xs = | |
(if f x \<in> f ` set xs then xs else insort_key f x xs)" | |
abbreviation "sort \<equiv> sort_key (\<lambda>x. x)" | |
abbreviation "insort \<equiv> insort_key (\<lambda>x. x)" | |
abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)" | |
definition stable_sort_key :: "(('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list) \<Rightarrow> bool" where | |
"stable_sort_key sk = | |
(\<forall>f xs k. filter (\<lambda>y. f y = k) (sk f xs) = filter (\<lambda>y. f y = k) xs)" | |
lemma strict_sorted_iff: "sorted_wrt (<) l \<longleftrightarrow> sorted l \<and> distinct l" | |
by (induction l) (auto iff: antisym_conv1) | |
lemma strict_sorted_imp_sorted: "sorted_wrt (<) xs \<Longrightarrow> sorted xs" | |
by (auto simp: strict_sorted_iff) | |
end | |
subsubsection \<open>List comprehension\<close> | |
text\<open>Input syntax for Haskell-like list comprehension notation. | |
Typical example: \<open>[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]\<close>, | |
the list of all pairs of distinct elements from \<open>xs\<close> and \<open>ys\<close>. | |
The syntax is as in Haskell, except that \<open>|\<close> becomes a dot | |
(like in Isabelle's set comprehension): \<open>[e. x \<leftarrow> xs, \<dots>]\<close> rather than | |
\verb![e| x <- xs, ...]!. | |
The qualifiers after the dot are | |
\begin{description} | |
\item[generators] \<open>p \<leftarrow> xs\<close>, | |
where \<open>p\<close> is a pattern and \<open>xs\<close> an expression of list type, or | |
\item[guards] \<open>b\<close>, where \<open>b\<close> is a boolean expression. | |
%\item[local bindings] @ {text"let x = e"}. | |
\end{description} | |
Just like in Haskell, list comprehension is just a shorthand. To avoid | |
misunderstandings, the translation into desugared form is not reversed | |
upon output. Note that the translation of \<open>[e. x \<leftarrow> xs]\<close> is | |
optmized to \<^term>\<open>map (%x. e) xs\<close>. | |
It is easy to write short list comprehensions which stand for complex | |
expressions. During proofs, they may become unreadable (and | |
mangled). In such cases it can be advisable to introduce separate | |
definitions for the list comprehensions in question.\<close> | |
nonterminal lc_qual and lc_quals | |
syntax | |
"_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list" ("[_ . __") | |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _") | |
"_lc_test" :: "bool \<Rightarrow> lc_qual" ("_") | |
(*"_lc_let" :: "letbinds => lc_qual" ("let _")*) | |
"_lc_end" :: "lc_quals" ("]") | |
"_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __") | |
syntax (ASCII) | |
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _") | |
parse_translation \<open> | |
let | |
val NilC = Syntax.const \<^const_syntax>\<open>Nil\<close>; | |
val ConsC = Syntax.const \<^const_syntax>\<open>Cons\<close>; | |
val mapC = Syntax.const \<^const_syntax>\<open>map\<close>; | |
val concatC = Syntax.const \<^const_syntax>\<open>concat\<close>; | |
val IfC = Syntax.const \<^const_syntax>\<open>If\<close>; | |
val dummyC = Syntax.const \<^const_syntax>\<open>Pure.dummy_pattern\<close> | |
fun single x = ConsC $ x $ NilC; | |
fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *) | |
let | |
(* FIXME proper name context!? *) | |
val x = | |
Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT); | |
val e = if opti then single e else e; | |
val case1 = Syntax.const \<^syntax_const>\<open>_case1\<close> $ p $ e; | |
val case2 = | |
Syntax.const \<^syntax_const>\<open>_case1\<close> $ dummyC $ NilC; | |
val cs = Syntax.const \<^syntax_const>\<open>_case2\<close> $ case1 $ case2; | |
in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end; | |
fun pair_pat_tr (x as Free _) e = Syntax_Trans.abs_tr [x, e] | |
| pair_pat_tr (_ $ p1 $ p2) e = | |
Syntax.const \<^const_syntax>\<open>case_prod\<close> $ pair_pat_tr p1 (pair_pat_tr p2 e) | |
| pair_pat_tr dummy e = Syntax_Trans.abs_tr [Syntax.const "_idtdummy", e] | |
fun pair_pat ctxt (Const (\<^const_syntax>\<open>Pair\<close>,_) $ s $ t) = | |
pair_pat ctxt s andalso pair_pat ctxt t | |
| pair_pat ctxt (Free (s,_)) = | |
let | |
val thy = Proof_Context.theory_of ctxt; | |
val s' = Proof_Context.intern_const ctxt s; | |
in not (Sign.declared_const thy s') end | |
| pair_pat _ t = (t = dummyC); | |
fun abs_tr ctxt p e opti = | |
let val p = Term_Position.strip_positions p | |
in if pair_pat ctxt p | |
then (pair_pat_tr p e, true) | |
else (pat_tr ctxt p e opti, false) | |
end | |
fun lc_tr ctxt [e, Const (\<^syntax_const>\<open>_lc_test\<close>, _) $ b, qs] = | |
let | |
val res = | |
(case qs of | |
Const (\<^syntax_const>\<open>_lc_end\<close>, _) => single e | |
| Const (\<^syntax_const>\<open>_lc_quals\<close>, _) $ q $ qs => lc_tr ctxt [e, q, qs]); | |
in IfC $ b $ res $ NilC end | |
| lc_tr ctxt | |
[e, Const (\<^syntax_const>\<open>_lc_gen\<close>, _) $ p $ es, | |
Const(\<^syntax_const>\<open>_lc_end\<close>, _)] = | |
(case abs_tr ctxt p e true of | |
(f, true) => mapC $ f $ es | |
| (f, false) => concatC $ (mapC $ f $ es)) | |
| lc_tr ctxt | |
[e, Const (\<^syntax_const>\<open>_lc_gen\<close>, _) $ p $ es, | |
Const (\<^syntax_const>\<open>_lc_quals\<close>, _) $ q $ qs] = | |
let val e' = lc_tr ctxt [e, q, qs]; | |
in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end; | |
in [(\<^syntax_const>\<open>_listcompr\<close>, lc_tr)] end | |
\<close> | |
ML_val \<open> | |
let | |
val read = Syntax.read_term \<^context> o Syntax.implode_input; | |
fun check s1 s2 = | |
read s1 aconv read s2 orelse | |
error ("Check failed: " ^ | |
quote (#1 (Input.source_content s1)) ^ Position.here_list [Input.pos_of s1, Input.pos_of s2]); | |
in | |
check \<open>[(x,y,z). b]\<close> \<open>if b then [(x, y, z)] else []\<close>; | |
check \<open>[(x,y,z). (x,_,y)\<leftarrow>xs]\<close> \<open>map (\<lambda>(x,_,y). (x, y, z)) xs\<close>; | |
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>; | |
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>; | |
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>; | |
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>; | |
check \<open>[(x,y). Cons True x \<leftarrow> xs]\<close> | |
\<open>concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)\<close>; | |
check \<open>[(x,y,z). Cons x [] \<leftarrow> xs]\<close> | |
\<open>concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)\<close>; | |
check \<open>[(x,y,z). x<a, x>b, x=d]\<close> | |
\<open>if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []\<close>; | |
check \<open>[(x,y,z). x<a, x>b, y\<leftarrow>ys]\<close> | |
\<open>if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []\<close>; | |
check \<open>[(x,y,z). x<a, (_,x)\<leftarrow>xs,y>b]\<close> | |
\<open>if x < a then concat (map (\<lambda>(_,x). if b < y then [(x, y, z)] else []) xs) else []\<close>; | |
check \<open>[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]\<close> | |
\<open>if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []\<close>; | |
check \<open>[(x,y,z). x\<leftarrow>xs, x>b, y<a]\<close> | |
\<open>concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)\<close>; | |
check \<open>[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]\<close> | |
\<open>concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)\<close>; | |
check \<open>[(x,y,z). x\<leftarrow>xs, (y,_)\<leftarrow>ys,y>x]\<close> | |
\<open>concat (map (\<lambda>x. concat (map (\<lambda>(y,_). if x < y then [(x, y, z)] else []) ys)) xs)\<close>; | |
check \<open>[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]\<close> | |
\<open>concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)\<close> | |
end; | |
\<close> | |
ML \<open> | |
(* Simproc for rewriting list comprehensions applied to List.set to set | |
comprehension. *) | |
signature LIST_TO_SET_COMPREHENSION = | |
sig | |
val simproc : Proof.context -> cterm -> thm option | |
end | |
structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION = | |
struct | |
(* conversion *) | |
fun all_exists_conv cv ctxt ct = | |
(case Thm.term_of ct of | |
Const (\<^const_name>\<open>Ex\<close>, _) $ Abs _ => | |
Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct | |
| _ => cv ctxt ct) | |
fun all_but_last_exists_conv cv ctxt ct = | |
(case Thm.term_of ct of | |
Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (_, _, Const (\<^const_name>\<open>Ex\<close>, _) $ _) => | |
Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct | |
| _ => cv ctxt ct) | |
fun Collect_conv cv ctxt ct = | |
(case Thm.term_of ct of | |
Const (\<^const_name>\<open>Collect\<close>, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct | |
| _ => raise CTERM ("Collect_conv", [ct])) | |
fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th) | |
fun conjunct_assoc_conv ct = | |
Conv.try_conv | |
(rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct | |
fun right_hand_set_comprehension_conv conv ctxt = | |
HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv | |
(Collect_conv (all_exists_conv conv o #2) ctxt)) | |
(* term abstraction of list comprehension patterns *) | |
datatype termlets = If | Case of typ * int | |
local | |
val set_Nil_I = @{lemma "set [] = {x. False}" by (simp add: empty_def [symmetric])} | |
val set_singleton = @{lemma "set [a] = {x. x = a}" by simp} | |
val inst_Collect_mem_eq = @{lemma "set A = {x. x \<in> set A}" by simp} | |
val del_refl_eq = @{lemma "(t = t \<and> P) \<equiv> P" by simp} | |
fun mk_set T = Const (\<^const_name>\<open>set\<close>, HOLogic.listT T --> HOLogic.mk_setT T) | |
fun dest_set (Const (\<^const_name>\<open>set\<close>, _) $ xs) = xs | |
fun dest_singleton_list (Const (\<^const_name>\<open>Cons\<close>, _) $ t $ (Const (\<^const_name>\<open>Nil\<close>, _))) = t | |
| dest_singleton_list t = raise TERM ("dest_singleton_list", [t]) | |
(*We check that one case returns a singleton list and all other cases | |
return [], and return the index of the one singleton list case.*) | |
fun possible_index_of_singleton_case cases = | |
let | |
fun check (i, case_t) s = | |
(case strip_abs_body case_t of | |
(Const (\<^const_name>\<open>Nil\<close>, _)) => s | |
| _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE)) | |
in | |
fold_index check cases (SOME NONE) |> the_default NONE | |
end | |
(*returns condition continuing term option*) | |
fun dest_if (Const (\<^const_name>\<open>If\<close>, _) $ cond $ then_t $ Const (\<^const_name>\<open>Nil\<close>, _)) = | |
SOME (cond, then_t) | |
| dest_if _ = NONE | |
(*returns (case_expr type index chosen_case constr_name) option*) | |
fun dest_case ctxt case_term = | |
let | |
val (case_const, args) = strip_comb case_term | |
in | |
(case try dest_Const case_const of | |
SOME (c, T) => | |
(case Ctr_Sugar.ctr_sugar_of_case ctxt c of | |
SOME {ctrs, ...} => | |
(case possible_index_of_singleton_case (fst (split_last args)) of | |
SOME i => | |
let | |
val constr_names = map (fst o dest_Const) ctrs | |
val (Ts, _) = strip_type T | |
val T' = List.last Ts | |
in SOME (List.last args, T', i, nth args i, nth constr_names i) end | |
| NONE => NONE) | |
| NONE => NONE) | |
| NONE => NONE) | |
end | |
fun tac ctxt [] = | |
resolve_tac ctxt [set_singleton] 1 ORELSE | |
resolve_tac ctxt [inst_Collect_mem_eq] 1 | |
| tac ctxt (If :: cont) = | |
Splitter.split_tac ctxt @{thms if_split} 1 | |
THEN resolve_tac ctxt @{thms conjI} 1 | |
THEN resolve_tac ctxt @{thms impI} 1 | |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => | |
CONVERSION (right_hand_set_comprehension_conv (K | |
(HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv | |
then_conv | |
rewr_conv' @{lemma "(True \<and> P) = P" by simp})) ctxt') 1) ctxt 1 | |
THEN tac ctxt cont | |
THEN resolve_tac ctxt @{thms impI} 1 | |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => | |
CONVERSION (right_hand_set_comprehension_conv (K | |
(HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv | |
then_conv rewr_conv' @{lemma "(False \<and> P) = False" by simp})) ctxt') 1) ctxt 1 | |
THEN resolve_tac ctxt [set_Nil_I] 1 | |
| tac ctxt (Case (T, i) :: cont) = | |
let | |
val SOME {injects, distincts, case_thms, split, ...} = | |
Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T)) | |
in | |
(* do case distinction *) | |
Splitter.split_tac ctxt [split] 1 | |
THEN EVERY (map_index (fn (i', _) => | |
(if i' < length case_thms - 1 then resolve_tac ctxt @{thms conjI} 1 else all_tac) | |
THEN REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1) | |
THEN resolve_tac ctxt @{thms impI} 1 | |
THEN (if i' = i then | |
(* continue recursively *) | |
Subgoal.FOCUS (fn {prems, context = ctxt', ...} => | |
CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K | |
((HOLogic.conj_conv | |
(HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv | |
(Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects)))) | |
Conv.all_conv) | |
then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq)) | |
then_conv conjunct_assoc_conv)) ctxt' | |
then_conv | |
(HOLogic.Trueprop_conv | |
(HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') => | |
Conv.repeat_conv | |
(all_but_last_exists_conv | |
(K (rewr_conv' | |
@{lemma "(\<exists>x. x = t \<and> P x) = P t" by simp})) ctxt'')) ctxt')))) 1) ctxt 1 | |
THEN tac ctxt cont | |
else | |
Subgoal.FOCUS (fn {prems, context = ctxt', ...} => | |
CONVERSION | |
(right_hand_set_comprehension_conv (K | |
(HOLogic.conj_conv | |
((HOLogic.eq_conv Conv.all_conv | |
(rewr_conv' (List.last prems))) then_conv | |
(Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts))) | |
Conv.all_conv then_conv | |
(rewr_conv' @{lemma "(False \<and> P) = False" by simp}))) ctxt' then_conv | |
HOLogic.Trueprop_conv | |
(HOLogic.eq_conv Conv.all_conv | |
(Collect_conv (fn (_, ctxt'') => | |
Conv.repeat_conv | |
(Conv.bottom_conv | |
(K (rewr_conv' @{lemma "(\<exists>x. P) = P" by simp})) ctxt'')) ctxt'))) 1) ctxt 1 | |
THEN resolve_tac ctxt [set_Nil_I] 1)) case_thms) | |
end | |
in | |
fun simproc ctxt redex = | |
let | |
fun make_inner_eqs bound_vs Tis eqs t = | |
(case dest_case ctxt t of | |
SOME (x, T, i, cont, constr_name) => | |
let | |
val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont) | |
val x' = incr_boundvars (length vs) x | |
val eqs' = map (incr_boundvars (length vs)) eqs | |
val constr_t = | |
list_comb | |
(Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0)) | |
val constr_eq = Const (\<^const_name>\<open>HOL.eq\<close>, T --> T --> \<^typ>\<open>bool\<close>) $ constr_t $ x' | |
in | |
make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body | |
end | |
| NONE => | |
(case dest_if t of | |
SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont | |
| NONE => | |
if null eqs then NONE (*no rewriting, nothing to be done*) | |
else | |
let | |
val Type (\<^type_name>\<open>list\<close>, [rT]) = fastype_of1 (map snd bound_vs, t) | |
val pat_eq = | |
(case try dest_singleton_list t of | |
SOME t' => | |
Const (\<^const_name>\<open>HOL.eq\<close>, rT --> rT --> \<^typ>\<open>bool\<close>) $ | |
Bound (length bound_vs) $ t' | |
| NONE => | |
Const (\<^const_name>\<open>Set.member\<close>, rT --> HOLogic.mk_setT rT --> \<^typ>\<open>bool\<close>) $ | |
Bound (length bound_vs) $ (mk_set rT $ t)) | |
val reverse_bounds = curry subst_bounds | |
((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)]) | |
val eqs' = map reverse_bounds eqs | |
val pat_eq' = reverse_bounds pat_eq | |
val inner_t = | |
fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t) | |
(rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq') | |
val lhs = Thm.term_of redex | |
val rhs = HOLogic.mk_Collect ("x", rT, inner_t) | |
val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) | |
in | |
SOME | |
((Goal.prove ctxt [] [] rewrite_rule_t | |
(fn {context = ctxt', ...} => tac ctxt' (rev Tis))) RS @{thm eq_reflection}) | |
end)) | |
in | |
make_inner_eqs [] [] [] (dest_set (Thm.term_of redex)) | |
end | |
end | |
end | |
\<close> | |
simproc_setup list_to_set_comprehension ("set xs") = | |
\<open>K List_to_Set_Comprehension.simproc\<close> | |
code_datatype set coset | |
hide_const (open) coset | |
subsubsection \<open>\<^const>\<open>Nil\<close> and \<^const>\<open>Cons\<close>\<close> | |
lemma not_Cons_self [simp]: | |
"xs \<noteq> x # xs" | |
by (induct xs) auto | |
lemma not_Cons_self2 [simp]: "x # xs \<noteq> xs" | |
by (rule not_Cons_self [symmetric]) | |
lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)" | |
by (induct xs) auto | |
lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (\<exists>x. xs = [x])" | |
by (cases xs) auto | |
lemmas Nil_tl = tl_Nil[THEN eq_iff_swap] | |
lemma length_induct: | |
"(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs" | |
by (fact measure_induct) | |
lemma induct_list012: | |
"\<lbrakk>P []; \<And>x. P [x]; \<And>x y zs. \<lbrakk> P zs; P (y # zs) \<rbrakk> \<Longrightarrow> P (x # y # zs)\<rbrakk> \<Longrightarrow> P xs" | |
by induction_schema (pat_completeness, lexicographic_order) | |
lemma list_nonempty_induct [consumes 1, case_names single cons]: | |
"\<lbrakk> xs \<noteq> []; \<And>x. P [x]; \<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)\<rbrakk> \<Longrightarrow> P xs" | |
by(induction xs rule: induct_list012) auto | |
lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X" | |
by (auto intro!: inj_onI) | |
lemma inj_on_Cons1 [simp]: "inj_on ((#) x) A" | |
by(simp add: inj_on_def) | |
subsubsection \<open>\<^const>\<open>length\<close>\<close> | |
text \<open> | |
Needs to come before \<open>@\<close> because of theorem \<open>append_eq_append_conv\<close>. | |
\<close> | |
lemma length_append [simp]: "length (xs @ ys) = length xs + length ys" | |
by (induct xs) auto | |
lemma length_map [simp]: "length (map f xs) = length xs" | |
by (induct xs) auto | |
lemma length_rev [simp]: "length (rev xs) = length xs" | |
by (induct xs) auto | |
lemma length_tl [simp]: "length (tl xs) = length xs - 1" | |
by (cases xs) auto | |
lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])" | |
by (induct xs) auto | |
lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])" | |
by (induct xs) auto | |
lemma length_pos_if_in_set: "x \<in> set xs \<Longrightarrow> length xs > 0" | |
by auto | |
lemma length_Suc_conv: "(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)" | |
by (induct xs) auto | |
lemmas Suc_length_conv = length_Suc_conv[THEN eq_iff_swap] | |
lemma Suc_le_length_iff: | |
"(Suc n \<le> length xs) = (\<exists>x ys. xs = x # ys \<and> n \<le> length ys)" | |
by (metis Suc_le_D[of n] Suc_le_mono[of n] Suc_length_conv[of _ xs]) | |
lemma impossible_Cons: "length xs \<le> length ys \<Longrightarrow> xs = x # ys = False" | |
by (induct xs) auto | |
lemma list_induct2 [consumes 1, case_names Nil Cons]: | |
"length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow> | |
(\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys)) | |
\<Longrightarrow> P xs ys" | |
proof (induct xs arbitrary: ys) | |
case (Cons x xs ys) then show ?case by (cases ys) simp_all | |
qed simp | |
lemma list_induct3 [consumes 2, case_names Nil Cons]: | |
"length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow> | |
(\<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)) | |
\<Longrightarrow> P xs ys zs" | |
proof (induct xs arbitrary: ys zs) | |
case Nil then show ?case by simp | |
next | |
case (Cons x xs ys zs) then show ?case by (cases ys, simp_all) | |
(cases zs, simp_all) | |
qed | |
lemma list_induct4 [consumes 3, case_names Nil Cons]: | |
"length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> | |
P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow> | |
length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow> | |
P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws" | |
proof (induct xs arbitrary: ys zs ws) | |
case Nil then show ?case by simp | |
next | |
case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all) | |
qed | |
lemma list_induct2': | |
"\<lbrakk> P [] []; | |
\<And>x xs. P (x#xs) []; | |
\<And>y ys. P [] (y#ys); | |
\<And>x xs y ys. P xs ys \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk> | |
\<Longrightarrow> P xs ys" | |
by (induct xs arbitrary: ys) (case_tac x, auto)+ | |
lemma list_all2_iff: | |
"list_all2 P xs ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y)" | |
by (induct xs ys rule: list_induct2') auto | |
lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False" | |
by (rule Eq_FalseI) auto | |
subsubsection \<open>\<open>@\<close> -- append\<close> | |
global_interpretation append: monoid append Nil | |
proof | |
fix xs ys zs :: "'a list" | |
show "(xs @ ys) @ zs = xs @ (ys @ zs)" | |
by (induct xs) simp_all | |
show "xs @ [] = xs" | |
by (induct xs) simp_all | |
qed simp | |
lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)" | |
by (fact append.assoc) | |
lemma append_Nil2: "xs @ [] = xs" | |
by (fact append.right_neutral) | |
lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])" | |
by (induct xs) auto | |
lemmas Nil_is_append_conv [iff] = append_is_Nil_conv[THEN eq_iff_swap] | |
lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])" | |
by (induct xs) auto | |
lemmas self_append_conv [iff] = append_self_conv[THEN eq_iff_swap] | |
lemma append_eq_append_conv [simp]: | |
"length xs = length ys \<or> length us = length vs | |
\<Longrightarrow> (xs@us = ys@vs) = (xs=ys \<and> us=vs)" | |
by (induct xs arbitrary: ys; case_tac ys; force) | |
lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) = | |
(\<exists>us. xs = zs @ us \<and> us @ ys = ts \<or> xs @ us = zs \<and> ys = us @ ts)" | |
proof (induct xs arbitrary: ys zs ts) | |
case (Cons x xs) | |
then show ?case | |
by (cases zs) auto | |
qed fastforce | |
lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)" | |
by simp | |
lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)" | |
by simp | |
lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)" | |
by simp | |
lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])" | |
using append_same_eq [of _ _ "[]"] by auto | |
lemmas self_append_conv2 [iff] = append_self_conv2[THEN eq_iff_swap] | |
lemma hd_Cons_tl: "xs \<noteq> [] \<Longrightarrow> hd xs # tl xs = xs" | |
by (fact list.collapse) | |
lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)" | |
by (induct xs) auto | |
lemma hd_append2 [simp]: "xs \<noteq> [] \<Longrightarrow> hd (xs @ ys) = hd xs" | |
by (simp add: hd_append split: list.split) | |
lemma tl_append: "tl (xs @ ys) = (case xs of [] \<Rightarrow> tl ys | z#zs \<Rightarrow> zs @ ys)" | |
by (simp split: list.split) | |
lemma tl_append2 [simp]: "xs \<noteq> [] \<Longrightarrow> tl (xs @ ys) = tl xs @ ys" | |
by (simp add: tl_append split: list.split) | |
lemma tl_append_if: "tl (xs @ ys) = (if xs = [] then tl ys else tl xs @ ys)" | |
by (simp) | |
lemma Cons_eq_append_conv: "x#xs = ys@zs = | |
(ys = [] \<and> x#xs = zs \<or> (\<exists>ys'. x#ys' = ys \<and> xs = ys'@zs))" | |
by(cases ys) auto | |
lemma append_eq_Cons_conv: "(ys@zs = x#xs) = | |
(ys = [] \<and> zs = x#xs \<or> (\<exists>ys'. ys = x#ys' \<and> ys'@zs = xs))" | |
by(cases ys) auto | |
lemma longest_common_prefix: | |
"\<exists>ps xs' ys'. xs = ps @ xs' \<and> ys = ps @ ys' | |
\<and> (xs' = [] \<or> ys' = [] \<or> hd xs' \<noteq> hd ys')" | |
by (induct xs ys rule: list_induct2') | |
(blast, blast, blast, | |
metis (no_types, opaque_lifting) append_Cons append_Nil list.sel(1)) | |
text \<open>Trivial rules for solving \<open>@\<close>-equations automatically.\<close> | |
lemma eq_Nil_appendI: "xs = ys \<Longrightarrow> xs = [] @ ys" | |
by simp | |
lemma Cons_eq_appendI: "\<lbrakk>x # xs1 = ys; xs = xs1 @ zs\<rbrakk> \<Longrightarrow> x # xs = ys @ zs" | |
by auto | |
lemma append_eq_appendI: "\<lbrakk>xs @ xs1 = zs; ys = xs1 @ us\<rbrakk> \<Longrightarrow> xs @ ys = zs @ us" | |
by auto | |
text \<open> | |
Simplification procedure for all list equalities. | |
Currently only tries to rearrange \<open>@\<close> to see if | |
- both lists end in a singleton list, | |
- or both lists end in the same list. | |
\<close> | |
simproc_setup list_eq ("(xs::'a list) = ys") = \<open> | |
let | |
fun last (cons as Const (\<^const_name>\<open>Cons\<close>, _) $ _ $ xs) = | |
(case xs of Const (\<^const_name>\<open>Nil\<close>, _) => cons | _ => last xs) | |
| last (Const(\<^const_name>\<open>append\<close>,_) $ _ $ ys) = last ys | |
| last t = t; | |
fun list1 (Const(\<^const_name>\<open>Cons\<close>,_) $ _ $ Const(\<^const_name>\<open>Nil\<close>,_)) = true | |
| list1 _ = false; | |
fun butlast ((cons as Const(\<^const_name>\<open>Cons\<close>,_) $ x) $ xs) = | |
(case xs of Const (\<^const_name>\<open>Nil\<close>, _) => xs | _ => cons $ butlast xs) | |
| butlast ((app as Const (\<^const_name>\<open>append\<close>, _) $ xs) $ ys) = app $ butlast ys | |
| butlast xs = Const(\<^const_name>\<open>Nil\<close>, fastype_of xs); | |
val rearr_ss = | |
simpset_of (put_simpset HOL_basic_ss \<^context> | |
addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]); | |
fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) = | |
let | |
val lastl = last lhs and lastr = last rhs; | |
fun rearr conv = | |
let | |
val lhs1 = butlast lhs and rhs1 = butlast rhs; | |
val Type(_,listT::_) = eqT | |
val appT = [listT,listT] ---> listT | |
val app = Const(\<^const_name>\<open>append\<close>,appT) | |
val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) | |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2)); | |
val thm = Goal.prove ctxt [] [] eq | |
(K (simp_tac (put_simpset rearr_ss ctxt) 1)); | |
in SOME ((conv RS (thm RS trans)) RS eq_reflection) end; | |
in | |
if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv} | |
else if lastl aconv lastr then rearr @{thm append_same_eq} | |
else NONE | |
end; | |
in K (fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct)) end | |
\<close> | |
subsubsection \<open>\<^const>\<open>map\<close>\<close> | |
lemma hd_map: "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)" | |
by (cases xs) simp_all | |
lemma map_tl: "map f (tl xs) = tl (map f xs)" | |
by (cases xs) simp_all | |
lemma map_ext: "(\<And>x. x \<in> set xs \<longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g xs" | |
by (induct xs) simp_all | |
lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)" | |
by (rule ext, induct_tac xs) auto | |
lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys" | |
by (induct xs) auto | |
lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs" | |
by (induct xs) auto | |
lemma map_comp_map[simp]: "((map f) \<circ> (map g)) = map(f \<circ> g)" | |
by (rule ext) simp | |
lemma rev_map: "rev (map f xs) = map f (rev xs)" | |
by (induct xs) auto | |
lemma map_eq_conv[simp]: "(map f xs = map g xs) = (\<forall>x \<in> set xs. f x = g x)" | |
by (induct xs) auto | |
lemma map_cong [fundef_cong]: | |
"xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys" | |
by simp | |
lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])" | |
by (rule list.map_disc_iff) | |
lemmas Nil_is_map_conv [iff] = map_is_Nil_conv[THEN eq_iff_swap] | |
lemma map_eq_Cons_conv: | |
"(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)" | |
by (cases xs) auto | |
lemma Cons_eq_map_conv: | |
"(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)" | |
by (cases ys) auto | |
lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1] | |
lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1] | |
declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] | |
lemma ex_map_conv: | |
"(\<exists>xs. ys = map f xs) = (\<forall>y \<in> set ys. \<exists>x. y = f x)" | |
by(induct ys, auto simp add: Cons_eq_map_conv) | |
lemma map_eq_imp_length_eq: | |
assumes "map f xs = map g ys" | |
shows "length xs = length ys" | |
using assms | |
proof (induct ys arbitrary: xs) | |
case Nil then show ?case by simp | |
next | |
case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto | |
from Cons xs have "map f zs = map g ys" by simp | |
with Cons have "length zs = length ys" by blast | |
with xs show ?case by simp | |
qed | |
lemma map_inj_on: | |
assumes map: "map f xs = map f ys" and inj: "inj_on f (set xs Un set ys)" | |
shows "xs = ys" | |
using map_eq_imp_length_eq [OF map] assms | |
proof (induct rule: list_induct2) | |
case (Cons x xs y ys) | |
then show ?case | |
by (auto intro: sym) | |
qed auto | |
lemma inj_on_map_eq_map: | |
"inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)" | |
by(blast dest:map_inj_on) | |
lemma map_injective: | |
"map f xs = map f ys \<Longrightarrow> inj f \<Longrightarrow> xs = ys" | |
by (induct ys arbitrary: xs) (auto dest!:injD) | |
lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)" | |
by(blast dest:map_injective) | |
lemma inj_mapI: "inj f \<Longrightarrow> inj (map f)" | |
by (iprover dest: map_injective injD intro: inj_onI) | |
lemma inj_mapD: "inj (map f) \<Longrightarrow> inj f" | |
by (metis (no_types, opaque_lifting) injI list.inject list.simps(9) the_inv_f_f) | |
lemma inj_map[iff]: "inj (map f) = inj f" | |
by (blast dest: inj_mapD intro: inj_mapI) | |
lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A" | |
by (blast intro:inj_onI dest:inj_onD map_inj_on) | |
lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs" | |
by (induct xs, auto) | |
lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs" | |
by (induct xs) auto | |
lemma map_fst_zip[simp]: | |
"length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs" | |
by (induct rule:list_induct2, simp_all) | |
lemma map_snd_zip[simp]: | |
"length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys" | |
by (induct rule:list_induct2, simp_all) | |
lemma map_fst_zip_take: | |
"map fst (zip xs ys) = take (min (length xs) (length ys)) xs" | |
by (induct xs ys rule: list_induct2') simp_all | |
lemma map_snd_zip_take: | |
"map snd (zip xs ys) = take (min (length xs) (length ys)) ys" | |
by (induct xs ys rule: list_induct2') simp_all | |
lemma map2_map_map: "map2 h (map f xs) (map g xs) = map (\<lambda>x. h (f x) (g x)) xs" | |
by (induction xs) (auto) | |
functor map: map | |
by (simp_all add: id_def) | |
declare map.id [simp] | |
subsubsection \<open>\<^const>\<open>rev\<close>\<close> | |
lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs" | |
by (induct xs) auto | |
lemma rev_rev_ident [simp]: "rev (rev xs) = xs" | |
by (induct xs) auto | |
lemma rev_swap: "(rev xs = ys) = (xs = rev ys)" | |
by auto | |
lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])" | |
by (induct xs) auto | |
lemmas Nil_is_rev_conv [iff] = rev_is_Nil_conv[THEN eq_iff_swap] | |
lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])" | |
by (cases xs) auto | |
lemma singleton_rev_conv [simp]: "([x] = rev xs) = ([x] = xs)" | |
by (cases xs) auto | |
lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)" | |
proof (induct xs arbitrary: ys) | |
case Nil | |
then show ?case by force | |
next | |
case Cons | |
then show ?case by (cases ys) auto | |
qed | |
lemma inj_on_rev[iff]: "inj_on rev A" | |
by(simp add:inj_on_def) | |
lemma rev_induct [case_names Nil snoc]: | |
assumes "P []" and "\<And>x xs. P xs \<Longrightarrow> P (xs @ [x])" | |
shows "P xs" | |
proof - | |
have "P (rev (rev xs))" | |
by (rule_tac list = "rev xs" in list.induct, simp_all add: assms) | |
then show ?thesis by simp | |
qed | |
lemma rev_exhaust [case_names Nil snoc]: | |
"(xs = [] \<Longrightarrow> P) \<Longrightarrow>(\<And>ys y. xs = ys @ [y] \<Longrightarrow> P) \<Longrightarrow> P" | |
by (induct xs rule: rev_induct) auto | |
lemmas rev_cases = rev_exhaust | |
lemma rev_nonempty_induct [consumes 1, case_names single snoc]: | |
assumes "xs \<noteq> []" | |
and single: "\<And>x. P [x]" | |
and snoc': "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (xs@[x])" | |
shows "P xs" | |
using \<open>xs \<noteq> []\<close> proof (induct xs rule: rev_induct) | |
case (snoc x xs) then show ?case | |
proof (cases xs) | |
case Nil thus ?thesis by (simp add: single) | |
next | |
case Cons with snoc show ?thesis by (fastforce intro!: snoc') | |
qed | |
qed simp | |
lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])" | |
by(rule rev_cases[of xs]) auto | |
lemma length_Suc_conv_rev: "(length xs = Suc n) = (\<exists>y ys. xs = ys @ [y] \<and> length ys = n)" | |
by (induct xs rule: rev_induct) auto | |
subsubsection \<open>\<^const>\<open>set\<close>\<close> | |
declare list.set[code_post] \<comment> \<open>pretty output\<close> | |
lemma finite_set [iff]: "finite (set xs)" | |
by (induct xs) auto | |
lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)" | |
by (induct xs) auto | |
lemma hd_in_set[simp]: "xs \<noteq> [] \<Longrightarrow> hd xs \<in> set xs" | |
by(cases xs) auto | |
lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)" | |
by auto | |
lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs" | |
by auto | |
lemma set_empty [iff]: "(set xs = {}) = (xs = [])" | |
by (induct xs) auto | |
lemmas set_empty2[iff] = set_empty[THEN eq_iff_swap] | |
lemma set_rev [simp]: "set (rev xs) = set xs" | |
by (induct xs) auto | |
lemma set_map [simp]: "set (map f xs) = f`(set xs)" | |
by (induct xs) auto | |
lemma set_filter [simp]: "set (filter P xs) = {x. x \<in> set xs \<and> P x}" | |
by (induct xs) auto | |
lemma set_upt [simp]: "set[i..<j] = {i..<j}" | |
by (induct j) auto | |
lemma split_list: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs" | |
proof (induct xs) | |
case Nil thus ?case by simp | |
next | |
case Cons thus ?case by (auto intro: Cons_eq_appendI) | |
qed | |
lemma in_set_conv_decomp: "x \<in> set xs \<longleftrightarrow> (\<exists>ys zs. xs = ys @ x # zs)" | |
by (auto elim: split_list) | |
lemma split_list_first: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys" | |
proof (induct xs) | |
case Nil thus ?case by simp | |
next | |
case (Cons a xs) | |
show ?case | |
proof cases | |
assume "x = a" thus ?case using Cons by fastforce | |
next | |
assume "x \<noteq> a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI) | |
qed | |
qed | |
lemma in_set_conv_decomp_first: | |
"(x \<in> set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)" | |
by (auto dest!: split_list_first) | |
lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs" | |
proof (induct xs rule: rev_induct) | |
case Nil thus ?case by simp | |
next | |
case (snoc a xs) | |
show ?case | |
proof cases | |
assume "x = a" thus ?case using snoc by (auto intro!: exI) | |
next | |
assume "x \<noteq> a" thus ?case using snoc by fastforce | |
qed | |
qed | |
lemma in_set_conv_decomp_last: | |
"(x \<in> set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)" | |
by (auto dest!: split_list_last) | |
lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs \<and> P x" | |
proof (induct xs) | |
case Nil thus ?case by simp | |
next | |
case Cons thus ?case | |
by(simp add:Bex_def)(metis append_Cons append.simps(1)) | |
qed | |
lemma split_list_propE: | |
assumes "\<exists>x \<in> set xs. P x" | |
obtains ys x zs where "xs = ys @ x # zs" and "P x" | |
using split_list_prop [OF assms] by blast | |
lemma split_list_first_prop: | |
"\<exists>x \<in> set xs. P x \<Longrightarrow> | |
\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y)" | |
proof (induct xs) | |
case Nil thus ?case by simp | |
next | |
case (Cons x xs) | |
show ?case | |
proof cases | |
assume "P x" | |
hence "x # xs = [] @ x # xs \<and> P x \<and> (\<forall>y\<in>set []. \<not> P y)" by simp | |
thus ?thesis by fast | |
next | |
assume "\<not> P x" | |
hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp | |
thus ?thesis using \<open>\<not> P x\<close> Cons(1) by (metis append_Cons set_ConsD) | |
qed | |
qed | |
lemma split_list_first_propE: | |
assumes "\<exists>x \<in> set xs. P x" | |
obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>y \<in> set ys. \<not> P y" | |
using split_list_first_prop [OF assms] by blast | |
lemma split_list_first_prop_iff: | |
"(\<exists>x \<in> set xs. P x) \<longleftrightarrow> | |
(\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y))" | |
by (rule, erule split_list_first_prop) auto | |
lemma split_list_last_prop: | |
"\<exists>x \<in> set xs. P x \<Longrightarrow> | |
\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z)" | |
proof(induct xs rule:rev_induct) | |
case Nil thus ?case by simp | |
next | |
case (snoc x xs) | |
show ?case | |
proof cases | |
assume "P x" thus ?thesis by (auto intro!: exI) | |
next | |
assume "\<not> P x" | |
hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp | |
thus ?thesis using \<open>\<not> P x\<close> snoc(1) by fastforce | |
qed | |
qed | |
lemma split_list_last_propE: | |
assumes "\<exists>x \<in> set xs. P x" | |
obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>z \<in> set zs. \<not> P z" | |
using split_list_last_prop [OF assms] by blast | |
lemma split_list_last_prop_iff: | |
"(\<exists>x \<in> set xs. P x) \<longleftrightarrow> | |
(\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))" | |
by rule (erule split_list_last_prop, auto) | |
lemma finite_list: "finite A \<Longrightarrow> \<exists>xs. set xs = A" | |
by (erule finite_induct) (auto simp add: list.set(2)[symmetric] simp del: list.set(2)) | |
lemma card_length: "card (set xs) \<le> length xs" | |
by (induct xs) (auto simp add: card_insert_if) | |
lemma set_minus_filter_out: | |
"set xs - {y} = set (filter (\<lambda>x. \<not> (x = y)) xs)" | |
by (induct xs) auto | |
lemma append_Cons_eq_iff: | |
"\<lbrakk> x \<notin> set xs; x \<notin> set ys \<rbrakk> \<Longrightarrow> | |
xs @ x # ys = xs' @ x # ys' \<longleftrightarrow> (xs = xs' \<and> ys = ys')" | |
by(auto simp: append_eq_Cons_conv Cons_eq_append_conv append_eq_append_conv2) | |
subsubsection \<open>\<^const>\<open>concat\<close>\<close> | |
lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys" | |
by (induct xs) auto | |
lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])" | |
by (induct xss) auto | |
lemmas Nil_eq_concat_conv [simp] = concat_eq_Nil_conv[THEN eq_iff_swap] | |
lemma set_concat [simp]: "set (concat xs) = (\<Union>x\<in>set xs. set x)" | |
by (induct xs) auto | |
lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs" | |
by (induct xs) auto | |
lemma map_concat: "map f (concat xs) = concat (map (map f) xs)" | |
by (induct xs) auto | |
lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))" | |
by (induct xs) auto | |
lemma length_concat_rev[simp]: "length (concat (rev xs)) = length (concat xs)" | |
by (induction xs) auto | |
lemma concat_eq_concat_iff: "\<forall>(x, y) \<in> set (zip xs ys). length x = length y \<Longrightarrow> length xs = length ys \<Longrightarrow> (concat xs = concat ys) = (xs = ys)" | |
proof (induct xs arbitrary: ys) | |
case (Cons x xs ys) | |
thus ?case by (cases ys) auto | |
qed (auto) | |
lemma concat_injective: "concat xs = concat ys \<Longrightarrow> length xs = length ys \<Longrightarrow> \<forall>(x, y) \<in> set (zip xs ys). length x = length y \<Longrightarrow> xs = ys" | |
by (simp add: concat_eq_concat_iff) | |
lemma concat_eq_appendD: | |
assumes "concat xss = ys @ zs" "xss \<noteq> []" | |
shows "\<exists>xss1 xs xs' xss2. xss = xss1 @ (xs @ xs') # xss2 \<and> ys = concat xss1 @ xs \<and> zs = xs' @ concat xss2" | |
using assms | |
proof(induction xss arbitrary: ys) | |
case (Cons xs xss) | |
from Cons.prems consider | |
us where "xs @ us = ys" "concat xss = us @ zs" | | |
us where "xs = ys @ us" "us @ concat xss = zs" | |
by(auto simp add: append_eq_append_conv2) | |
then show ?case | |
proof cases | |
case 1 | |
then show ?thesis using Cons.IH[OF 1(2)] | |
by(cases xss)(auto intro: exI[where x="[]"], metis append.assoc append_Cons concat.simps(2)) | |
qed(auto intro: exI[where x="[]"]) | |
qed simp | |
lemma concat_eq_append_conv: | |
"concat xss = ys @ zs \<longleftrightarrow> | |
(if xss = [] then ys = [] \<and> zs = [] | |
else \<exists>xss1 xs xs' xss2. xss = xss1 @ (xs @ xs') # xss2 \<and> ys = concat xss1 @ xs \<and> zs = xs' @ concat xss2)" | |
by(auto dest: concat_eq_appendD) | |
lemma hd_concat: "\<lbrakk>xs \<noteq> []; hd xs \<noteq> []\<rbrakk> \<Longrightarrow> hd (concat xs) = hd (hd xs)" | |
by (metis concat.simps(2) hd_Cons_tl hd_append2) | |
simproc_setup list_neq ("(xs::'a list) = ys") = \<open> | |
(* | |
Reduces xs=ys to False if xs and ys cannot be of the same length. | |
This is the case if the atomic sublists of one are a submultiset | |
of those of the other list and there are fewer Cons's in one than the other. | |
*) | |
let | |
fun len (Const(\<^const_name>\<open>Nil\<close>,_)) acc = acc | |
| len (Const(\<^const_name>\<open>Cons\<close>,_) $ _ $ xs) (ts,n) = len xs (ts,n+1) | |
| len (Const(\<^const_name>\<open>append\<close>,_) $ xs $ ys) acc = len xs (len ys acc) | |
| len (Const(\<^const_name>\<open>rev\<close>,_) $ xs) acc = len xs acc | |
| len (Const(\<^const_name>\<open>map\<close>,_) $ _ $ xs) acc = len xs acc | |
| len (Const(\<^const_name>\<open>concat\<close>,T) $ (Const(\<^const_name>\<open>rev\<close>,_) $ xss)) acc | |
= len (Const(\<^const_name>\<open>concat\<close>,T) $ xss) acc | |
| len t (ts,n) = (t::ts,n); | |
val ss = simpset_of \<^context>; | |
fun list_neq ctxt ct = | |
let | |
val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct; | |
val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0); | |
fun prove_neq() = | |
let | |
val Type(_,listT::_) = eqT; | |
val size = HOLogic.size_const listT; | |
val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs); | |
val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len); | |
val thm = Goal.prove ctxt [] [] neq_len | |
(K (simp_tac (put_simpset ss ctxt) 1)); | |
in SOME (thm RS @{thm neq_if_length_neq}) end | |
in | |
if m < n andalso submultiset (op aconv) (ls,rs) orelse | |
n < m andalso submultiset (op aconv) (rs,ls) | |
then prove_neq() else NONE | |
end; | |
in K list_neq end | |
\<close> | |
subsubsection \<open>\<^const>\<open>filter\<close>\<close> | |
lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys" | |
by (induct xs) auto | |
lemma rev_filter: "rev (filter P xs) = filter P (rev xs)" | |
by (induct xs) simp_all | |
lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs" | |
by (induct xs) auto | |
lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)" | |
by (induct xs) auto | |
lemma length_filter_le [simp]: "length (filter P xs) \<le> length xs" | |
by (induct xs) (auto simp add: le_SucI) | |
lemma sum_length_filter_compl: | |
"length(filter P xs) + length(filter (\<lambda>x. \<not>P x) xs) = length xs" | |
by(induct xs) simp_all | |
lemma filter_True [simp]: "\<forall>x \<in> set xs. P x \<Longrightarrow> filter P xs = xs" | |
by (induct xs) auto | |
lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x \<Longrightarrow> filter P xs = []" | |
by (induct xs) auto | |
lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)" | |
by (induct xs) simp_all | |
lemmas empty_filter_conv = filter_empty_conv[THEN eq_iff_swap] | |
lemma filter_id_conv: "(filter P xs = xs) = (\<forall>x\<in>set xs. P x)" | |
proof (induct xs) | |
case (Cons x xs) | |
then show ?case | |
using length_filter_le | |
by (simp add: impossible_Cons) | |
qed auto | |
lemma filter_map: "filter P (map f xs) = map f (filter (P \<circ> f) xs)" | |
by (induct xs) simp_all | |
lemma length_filter_map[simp]: | |
"length (filter P (map f xs)) = length(filter (P \<circ> f) xs)" | |
by (simp add:filter_map) | |
lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs" | |
by auto | |
lemma length_filter_less: | |
"\<lbrakk> x \<in> set xs; \<not> P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs" | |
proof (induct xs) | |
case Nil thus ?case by simp | |
next | |
case (Cons x xs) thus ?case | |
using Suc_le_eq by fastforce | |
qed | |
lemma length_filter_conv_card: | |
"length(filter p xs) = card{i. i < length xs \<and> p(xs!i)}" | |
proof (induct xs) | |
case Nil thus ?case by simp | |
next | |
case (Cons x xs) | |
let ?S = "{i. i < length xs \<and> p(xs!i)}" | |
have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite) | |
show ?case (is "?l = card ?S'") | |
proof (cases) | |
assume "p x" | |
hence eq: "?S' = insert 0 (Suc ` ?S)" | |
by(auto simp: image_def split:nat.split dest:gr0_implies_Suc) | |
have "length (filter p (x # xs)) = Suc(card ?S)" | |
using Cons \<open>p x\<close> by simp | |
also have "\<dots> = Suc(card(Suc ` ?S))" using fin | |
by (simp add: card_image) | |
also have "\<dots> = card ?S'" using eq fin | |
by (simp add:card_insert_if) | |
finally show ?thesis . | |
next | |
assume "\<not> p x" | |
hence eq: "?S' = Suc ` ?S" | |
by(auto simp add: image_def split:nat.split elim:lessE) | |
have "length (filter p (x # xs)) = card ?S" | |
using Cons \<open>\<not> p x\<close> by simp | |
also have "\<dots> = card(Suc ` ?S)" using fin | |
by (simp add: card_image) | |
also have "\<dots> = card ?S'" using eq fin | |
by (simp add:card_insert_if) | |
finally show ?thesis . | |
qed | |
qed | |
lemma Cons_eq_filterD: | |
"x#xs = filter P ys \<Longrightarrow> | |
\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs" | |
(is "_ \<Longrightarrow> \<exists>us vs. ?P ys us vs") | |
proof(induct ys) | |
case Nil thus ?case by simp | |
next | |
case (Cons y ys) | |
show ?case (is "\<exists>x. ?Q x") | |
proof cases | |
assume Py: "P y" | |
show ?thesis | |
proof cases | |
assume "x = y" | |
with Py Cons.prems have "?Q []" by simp | |
then show ?thesis .. | |
next | |
assume "x \<noteq> y" | |
with Py Cons.prems show ?thesis by simp | |
qed | |
next | |
assume "\<not> P y" | |
with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce | |
then have "?Q (y#us)" by simp | |
then show ?thesis .. | |
qed | |
qed | |
lemma filter_eq_ConsD: | |
"filter P ys = x#xs \<Longrightarrow> | |
\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs" | |
by(rule Cons_eq_filterD) simp | |
lemma filter_eq_Cons_iff: | |
"(filter P ys = x#xs) = | |
(\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)" | |
by(auto dest:filter_eq_ConsD) | |
lemmas Cons_eq_filter_iff = filter_eq_Cons_iff[THEN eq_iff_swap] | |
lemma inj_on_filter_key_eq: | |
assumes "inj_on f (insert y (set xs))" | |
shows "filter (\<lambda>x. f y = f x) xs = filter (HOL.eq y) xs" | |
using assms by (induct xs) auto | |
lemma filter_cong[fundef_cong]: | |
"xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys" | |
by (induct ys arbitrary: xs) auto | |
subsubsection \<open>List partitioning\<close> | |
primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where | |
"partition P [] = ([], [])" | | |
"partition P (x # xs) = | |
(let (yes, no) = partition P xs | |
in if P x then (x # yes, no) else (yes, x # no))" | |
lemma partition_filter1: "fst (partition P xs) = filter P xs" | |
by (induct xs) (auto simp add: Let_def split_def) | |
lemma partition_filter2: "snd (partition P xs) = filter (Not \<circ> P) xs" | |
by (induct xs) (auto simp add: Let_def split_def) | |
lemma partition_P: | |
assumes "partition P xs = (yes, no)" | |
shows "(\<forall>p \<in> set yes. P p) \<and> (\<forall>p \<in> set no. \<not> P p)" | |
proof - | |
from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)" | |
by simp_all | |
then show ?thesis by (simp_all add: partition_filter1 partition_filter2) | |
qed | |
lemma partition_set: | |
assumes "partition P xs = (yes, no)" | |
shows "set yes \<union> set no = set xs" | |
proof - | |
from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)" | |
by simp_all | |
then show ?thesis by (auto simp add: partition_filter1 partition_filter2) | |
qed | |
lemma partition_filter_conv[simp]: | |
"partition f xs = (filter f xs,filter (Not \<circ> f) xs)" | |
unfolding partition_filter2[symmetric] | |
unfolding partition_filter1[symmetric] by simp | |
declare partition.simps[simp del] | |
subsubsection \<open>\<^const>\<open>nth\<close>\<close> | |
lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x" | |
by auto | |
lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n" | |
by auto | |
declare nth.simps [simp del] | |
lemma nth_Cons_pos[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)" | |
by(auto simp: Nat.gr0_conv_Suc) | |
lemma nth_append: | |
"(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))" | |
proof (induct xs arbitrary: n) | |
case (Cons x xs) | |
then show ?case | |
using less_Suc_eq_0_disj by auto | |
qed simp | |
lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x" | |
by (induct xs) auto | |
lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n" | |
by (induct xs) auto | |
lemma nth_map [simp]: "n < length xs \<Longrightarrow> (map f xs)!n = f(xs!n)" | |
proof (induct xs arbitrary: n) | |
case (Cons x xs) | |
then show ?case | |
using less_Suc_eq_0_disj by auto | |
qed simp | |
lemma nth_tl: "n < length (tl xs) \<Longrightarrow> tl xs ! n = xs ! Suc n" | |
by (induction xs) auto | |
lemma hd_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd xs = xs!0" | |
by(cases xs) simp_all | |
lemma list_eq_iff_nth_eq: | |
"(xs = ys) = (length xs = length ys \<and> (\<forall>i<length xs. xs!i = ys!i))" | |
proof (induct xs arbitrary: ys) | |
case (Cons x xs ys) | |
show ?case | |
proof (cases ys) | |
case (Cons y ys) | |
with Cons.hyps show ?thesis by fastforce | |
qed simp | |
qed force | |
lemma map_equality_iff: | |
"map f xs = map g ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>i<length ys. f (xs!i) = g (ys!i))" | |
by (fastforce simp: list_eq_iff_nth_eq) | |
lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}" | |
proof (induct xs) | |
case (Cons x xs) | |
have "insert x {xs ! i |i. i < length xs} = {(x # xs) ! i |i. i < Suc (length xs)}" (is "?L=?R") | |
proof | |
show "?L \<subseteq> ?R" | |
by force | |
show "?R \<subseteq> ?L" | |
using less_Suc_eq_0_disj by auto | |
qed | |
with Cons show ?case | |
by simp | |
qed simp | |
lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)" | |
by(auto simp:set_conv_nth) | |
lemma nth_equal_first_eq: | |
assumes "x \<notin> set xs" | |
assumes "n \<le> length xs" | |
shows "(x # xs) ! n = x \<longleftrightarrow> n = 0" (is "?lhs \<longleftrightarrow> ?rhs") | |
proof | |
assume ?lhs | |
show ?rhs | |
proof (rule ccontr) | |
assume "n \<noteq> 0" | |
then have "n > 0" by simp | |
with \<open>?lhs\<close> have "xs ! (n - 1) = x" by simp | |
moreover from \<open>n > 0\<close> \<open>n \<le> length xs\<close> have "n - 1 < length xs" by simp | |
ultimately have "\<exists>i<length xs. xs ! i = x" by auto | |
with \<open>x \<notin> set xs\<close> in_set_conv_nth [of x xs] show False by simp | |
qed | |
next | |
assume ?rhs then show ?lhs by simp | |
qed | |
lemma nth_non_equal_first_eq: | |
assumes "x \<noteq> y" | |
shows "(x # xs) ! n = y \<longleftrightarrow> xs ! (n - 1) = y \<and> n > 0" (is "?lhs \<longleftrightarrow> ?rhs") | |
proof | |
assume "?lhs" with assms have "n > 0" by (cases n) simp_all | |
with \<open>?lhs\<close> show ?rhs by simp | |
next | |
assume "?rhs" then show "?lhs" by simp | |
qed | |
lemma list_ball_nth: "\<lbrakk>n < length xs; \<forall>x \<in> set xs. P x\<rbrakk> \<Longrightarrow> P(xs!n)" | |
by (auto simp add: set_conv_nth) | |
lemma nth_mem [simp]: "n < length xs \<Longrightarrow> xs!n \<in> set xs" | |
by (auto simp add: set_conv_nth) | |
lemma all_nth_imp_all_set: | |
"\<lbrakk>\<forall>i < length xs. P(xs!i); x \<in> set xs\<rbrakk> \<Longrightarrow> P x" | |
by (auto simp add: set_conv_nth) | |
lemma all_set_conv_all_nth: | |
"(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs \<longrightarrow> P (xs ! i))" | |
by (auto simp add: set_conv_nth) | |
lemma rev_nth: | |
"n < size xs \<Longrightarrow> rev xs ! n = xs ! (length xs - Suc n)" | |
proof (induct xs arbitrary: n) | |
case Nil thus ?case by simp | |
next | |
case (Cons x xs) | |
hence n: "n < Suc (length xs)" by simp | |
moreover | |
{ assume "n < length xs" | |
with n obtain n' where n': "length xs - n = Suc n'" | |
by (cases "length xs - n", auto) | |
moreover | |
from n' have "length xs - Suc n = n'" by simp | |
ultimately | |
have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp | |
} | |
ultimately | |
show ?case by (clarsimp simp add: Cons nth_append) | |
qed | |
lemma Skolem_list_nth: | |
"(\<forall>i<k. \<exists>x. P i x) = (\<exists>xs. size xs = k \<and> (\<forall>i<k. P i (xs!i)))" | |
(is "_ = (\<exists>xs. ?P k xs)") | |
proof(induct k) | |
case 0 show ?case by simp | |
next | |
case (Suc k) | |
show ?case (is "?L = ?R" is "_ = (\<exists>xs. ?P' xs)") | |
proof | |
assume "?R" thus "?L" using Suc by auto | |
next | |
assume "?L" | |
with Suc obtain x xs where "?P k xs \<and> P k x" by (metis less_Suc_eq) | |
hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq) | |
thus "?R" .. | |
qed | |
qed | |
subsubsection \<open>\<^const>\<open>list_update\<close>\<close> | |
lemma length_list_update [simp]: "length(xs[i:=x]) = length xs" | |
by (induct xs arbitrary: i) (auto split: nat.split) | |
lemma nth_list_update: | |
"i < length xs\<Longrightarrow> (xs[i:=x])!j = (if i = j then x else xs!j)" | |
by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) | |
lemma nth_list_update_eq [simp]: "i < length xs \<Longrightarrow> (xs[i:=x])!i = x" | |
by (simp add: nth_list_update) | |
lemma nth_list_update_neq [simp]: "i \<noteq> j \<Longrightarrow> xs[i:=x]!j = xs!j" | |
by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) | |
lemma list_update_id[simp]: "xs[i := xs!i] = xs" | |
by (induct xs arbitrary: i) (simp_all split:nat.splits) | |
lemma list_update_beyond[simp]: "length xs \<le> i \<Longrightarrow> xs[i:=x] = xs" | |
proof (induct xs arbitrary: i) | |
case (Cons x xs i) | |
then show ?case | |
by (metis leD length_list_update list_eq_iff_nth_eq nth_list_update_neq) | |
qed simp | |
lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]" | |
by (simp only: length_0_conv[symmetric] length_list_update) | |
lemma list_update_same_conv: | |
"i < length xs \<Longrightarrow> (xs[i := x] = xs) = (xs!i = x)" | |
by (induct xs arbitrary: i) (auto split: nat.split) | |
lemma list_update_append1: | |
"i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys" | |
by (induct xs arbitrary: i)(auto split:nat.split) | |
lemma list_update_append: | |
"(xs @ ys) [n:= x] = | |
(if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))" | |
by (induct xs arbitrary: n) (auto split:nat.splits) | |
lemma list_update_length [simp]: | |
"(xs @ x # ys)[length xs := y] = (xs @ y # ys)" | |
by (induct xs, auto) | |
lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]" | |
by(induct xs arbitrary: k)(auto split:nat.splits) | |
lemma rev_update: | |
"k < length xs \<Longrightarrow> rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]" | |
by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits) | |
lemma update_zip: | |
"(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])" | |
by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split) | |
lemma set_update_subset_insert: "set(xs[i:=x]) \<le> insert x (set xs)" | |
by (induct xs arbitrary: i) (auto split: nat.split) | |
lemma set_update_subsetI: "\<lbrakk>set xs \<subseteq> A; x \<in> A\<rbrakk> \<Longrightarrow> set(xs[i := x]) \<subseteq> A" | |
by (blast dest!: set_update_subset_insert [THEN subsetD]) | |
lemma set_update_memI: "n < length xs \<Longrightarrow> x \<in> set (xs[n := x])" | |
by (induct xs arbitrary: n) (auto split:nat.splits) | |
lemma list_update_overwrite[simp]: | |
"xs [i := x, i := y] = xs [i := y]" | |
by (induct xs arbitrary: i) (simp_all split: nat.split) | |
lemma list_update_swap: | |
"i \<noteq> i' \<Longrightarrow> xs [i := x, i' := x'] = xs [i' := x', i := x]" | |
by (induct xs arbitrary: i i') (simp_all split: nat.split) | |
lemma list_update_code [code]: | |
"[][i := y] = []" | |
"(x # xs)[0 := y] = y # xs" | |
"(x # xs)[Suc i := y] = x # xs[i := y]" | |
by simp_all | |
subsubsection \<open>\<^const>\<open>last\<close> and \<^const>\<open>butlast\<close>\<close> | |
lemma hd_Nil_eq_last: "hd Nil = last Nil" | |
unfolding hd_def last_def by simp | |
lemma last_snoc [simp]: "last (xs @ [x]) = x" | |
by (induct xs) auto | |
lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs" | |
by (induct xs) auto | |
lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x" | |
by simp | |
lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs" | |
by simp | |
lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)" | |
by (induct xs) (auto) | |
lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs" | |
by(simp add:last_append) | |
lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys" | |
by(simp add:last_append) | |
lemma last_tl: "xs = [] \<or> tl xs \<noteq> [] \<Longrightarrow>last (tl xs) = last xs" | |
by (induct xs) simp_all | |
lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)" | |
by (induct xs) simp_all | |
lemma hd_rev: "hd(rev xs) = last xs" | |
by (metis hd_Cons_tl hd_Nil_eq_last last_snoc rev_eq_Cons_iff rev_is_Nil_conv) | |
lemma last_rev: "last(rev xs) = hd xs" | |
by (metis hd_rev rev_swap) | |
lemma last_in_set[simp]: "as \<noteq> [] \<Longrightarrow> last as \<in> set as" | |
by (induct as) auto | |
lemma length_butlast [simp]: "length (butlast xs) = length xs - 1" | |
by (induct xs rule: rev_induct) auto | |
lemma butlast_append: | |
"butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)" | |
by (induct xs arbitrary: ys) auto | |
lemma append_butlast_last_id [simp]: | |
"xs \<noteq> [] \<Longrightarrow> butlast xs @ [last xs] = xs" | |
by (induct xs) auto | |
lemma in_set_butlastD: "x \<in> set (butlast xs) \<Longrightarrow> x \<in> set xs" | |
by (induct xs) (auto split: if_split_asm) | |
lemma in_set_butlast_appendI: | |
"x \<in> set (butlast xs) \<or> x \<in> set (butlast ys) \<Longrightarrow> x \<in> set (butlast (xs @ ys))" | |
by (auto dest: in_set_butlastD simp add: butlast_append) | |
lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs" | |
by (induct xs arbitrary: n)(auto split:nat.split) | |
lemma nth_butlast: | |
assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n" | |
proof (cases xs) | |
case (Cons y ys) | |
moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n" | |
by (simp add: nth_append) | |
ultimately show ?thesis using append_butlast_last_id by simp | |
qed simp | |
lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)" | |
by(induct xs)(auto simp:neq_Nil_conv) | |
lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs" | |
by (induction xs rule: induct_list012) simp_all | |
lemma last_list_update: | |
"xs \<noteq> [] \<Longrightarrow> last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)" | |
by (auto simp: last_conv_nth) | |
lemma butlast_list_update: | |
"butlast(xs[k:=x]) = | |
(if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])" | |
by(cases xs rule:rev_cases)(auto simp: list_update_append split: nat.splits) | |
lemma last_map: "xs \<noteq> [] \<Longrightarrow> last (map f xs) = f (last xs)" | |
by (cases xs rule: rev_cases) simp_all | |
lemma map_butlast: "map f (butlast xs) = butlast (map f xs)" | |
by (induct xs) simp_all | |
lemma snoc_eq_iff_butlast: | |
"xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] \<and> butlast ys = xs \<and> last ys = x)" | |
by fastforce | |
corollary longest_common_suffix: | |
"\<exists>ss xs' ys'. xs = xs' @ ss \<and> ys = ys' @ ss | |
\<and> (xs' = [] \<or> ys' = [] \<or> last xs' \<noteq> last ys')" | |
using longest_common_prefix[of "rev xs" "rev ys"] | |
unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv) | |
lemma butlast_rev [simp]: "butlast (rev xs) = rev (tl xs)" | |
by (cases xs) simp_all | |
subsubsection \<open>\<^const>\<open>take\<close> and \<^const>\<open>drop\<close>\<close> | |
lemma take_0: "take 0 xs = []" | |
by (induct xs) auto | |
lemma drop_0: "drop 0 xs = xs" | |
by (induct xs) auto | |
lemma take0[simp]: "take 0 = (\<lambda>xs. [])" | |
by(rule ext) (rule take_0) | |
lemma drop0[simp]: "drop 0 = (\<lambda>x. x)" | |
by(rule ext) (rule drop_0) | |
lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs" | |
by simp | |
lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs" | |
by simp | |
declare take_Cons [simp del] and drop_Cons [simp del] | |
lemma take_Suc: "xs \<noteq> [] \<Longrightarrow> take (Suc n) xs = hd xs # take n (tl xs)" | |
by(clarsimp simp add:neq_Nil_conv) | |
lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)" | |
by(cases xs, simp_all) | |
lemma hd_take[simp]: "j > 0 \<Longrightarrow> hd (take j xs) = hd xs" | |
by (metis gr0_conv_Suc list.sel(1) take.simps(1) take_Suc) | |
lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)" | |
by (induct xs arbitrary: n) simp_all | |
lemma drop_tl: "drop n (tl xs) = tl(drop n xs)" | |
by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split) | |
lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)" | |
by (cases n, simp, cases xs, auto) | |
lemma tl_drop: "tl (drop n xs) = drop n (tl xs)" | |
by (simp only: drop_tl) | |
lemma nth_via_drop: "drop n xs = y#ys \<Longrightarrow> xs!n = y" | |
by (induct xs arbitrary: n, simp)(auto simp: drop_Cons nth_Cons split: nat.splits) | |
lemma take_Suc_conv_app_nth: | |
"i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]" | |
proof (induct xs arbitrary: i) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases i) auto | |
qed | |
lemma Cons_nth_drop_Suc: | |
"i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs" | |
proof (induct xs arbitrary: i) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases i) auto | |
qed | |
lemma length_take [simp]: "length (take n xs) = min (length xs) n" | |
by (induct n arbitrary: xs) (auto, case_tac xs, auto) | |
lemma length_drop [simp]: "length (drop n xs) = (length xs - n)" | |
by (induct n arbitrary: xs) (auto, case_tac xs, auto) | |
lemma take_all [simp]: "length xs \<le> n \<Longrightarrow> take n xs = xs" | |
by (induct n arbitrary: xs) (auto, case_tac xs, auto) | |
lemma drop_all [simp]: "length xs \<le> n \<Longrightarrow> drop n xs = []" | |
by (induct n arbitrary: xs) (auto, case_tac xs, auto) | |
lemma take_all_iff [simp]: "take n xs = xs \<longleftrightarrow> length xs \<le> n" | |
by (metis length_take min.order_iff take_all) | |
(* Looks like a good simp rule but can cause looping; | |
too much interaction between take and length | |
lemmas take_all_iff2[simp] = take_all_iff[THEN eq_iff_swap] | |
*) | |
lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \<or> xs = [])" | |
by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split) | |
lemmas take_eq_Nil2[simp] = take_eq_Nil[THEN eq_iff_swap] | |
lemma drop_eq_Nil [simp]: "drop n xs = [] \<longleftrightarrow> length xs \<le> n" | |
by (metis diff_is_0_eq drop_all length_drop list.size(3)) | |
lemmas drop_eq_Nil2 [simp] = drop_eq_Nil[THEN eq_iff_swap] | |
lemma take_append [simp]: | |
"take n (xs @ ys) = (take n xs @ take (n - length xs) ys)" | |
by (induct n arbitrary: xs) (auto, case_tac xs, auto) | |
lemma drop_append [simp]: | |
"drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys" | |
by (induct n arbitrary: xs) (auto, case_tac xs, auto) | |
lemma take_take [simp]: "take n (take m xs) = take (min n m) xs" | |
proof (induct m arbitrary: xs n) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs; cases n) simp_all | |
qed | |
lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs" | |
proof (induct m arbitrary: xs) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs) simp_all | |
qed | |
lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)" | |
proof (induct m arbitrary: xs n) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs; cases n) simp_all | |
qed | |
lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)" | |
by(induct xs arbitrary: m n)(auto simp: take_Cons drop_Cons split: nat.split) | |
lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs" | |
proof (induct n arbitrary: xs) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs) simp_all | |
qed | |
lemma take_map: "take n (map f xs) = map f (take n xs)" | |
proof (induct n arbitrary: xs) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs) simp_all | |
qed | |
lemma drop_map: "drop n (map f xs) = map f (drop n xs)" | |
proof (induct n arbitrary: xs) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs) simp_all | |
qed | |
lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)" | |
proof (induct xs arbitrary: i) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases i) auto | |
qed | |
lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)" | |
proof (induct xs arbitrary: i) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases i) auto | |
qed | |
lemma drop_rev: "drop n (rev xs) = rev (take (length xs - n) xs)" | |
by (cases "length xs < n") (auto simp: rev_take) | |
lemma take_rev: "take n (rev xs) = rev (drop (length xs - n) xs)" | |
by (cases "length xs < n") (auto simp: rev_drop) | |
lemma nth_take [simp]: "i < n \<Longrightarrow> (take n xs)!i = xs!i" | |
proof (induct xs arbitrary: i n) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases n; cases i) simp_all | |
qed | |
lemma nth_drop [simp]: | |
"n \<le> length xs \<Longrightarrow> (drop n xs)!i = xs!(n + i)" | |
proof (induct n arbitrary: xs) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs) simp_all | |
qed | |
lemma butlast_take: | |
"n \<le> length xs \<Longrightarrow> butlast (take n xs) = take (n - 1) xs" | |
by (simp add: butlast_conv_take) | |
lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)" | |
by (simp add: butlast_conv_take drop_take ac_simps) | |
lemma take_butlast: "n < length xs \<Longrightarrow> take n (butlast xs) = take n xs" | |
by (simp add: butlast_conv_take) | |
lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)" | |
by (simp add: butlast_conv_take drop_take ac_simps) | |
lemma butlast_power: "(butlast ^^ n) xs = take (length xs - n) xs" | |
by (induct n) (auto simp: butlast_take) | |
lemma hd_drop_conv_nth: "n < length xs \<Longrightarrow> hd(drop n xs) = xs!n" | |
by(simp add: hd_conv_nth) | |
lemma set_take_subset_set_take: | |
"m \<le> n \<Longrightarrow> set(take m xs) \<le> set(take n xs)" | |
proof (induct xs arbitrary: m n) | |
case (Cons x xs m n) then show ?case | |
by (cases n) (auto simp: take_Cons) | |
qed simp | |
lemma set_take_subset: "set(take n xs) \<subseteq> set xs" | |
by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split) | |
lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs" | |
by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split) | |
lemma set_drop_subset_set_drop: | |
"m \<ge> n \<Longrightarrow> set(drop m xs) \<le> set(drop n xs)" | |
proof (induct xs arbitrary: m n) | |
case (Cons x xs m n) | |
then show ?case | |
by (clarsimp simp: drop_Cons split: nat.split) (metis set_drop_subset subset_iff) | |
qed simp | |
lemma in_set_takeD: "x \<in> set(take n xs) \<Longrightarrow> x \<in> set xs" | |
using set_take_subset by fast | |
lemma in_set_dropD: "x \<in> set(drop n xs) \<Longrightarrow> x \<in> set xs" | |
using set_drop_subset by fast | |
lemma append_eq_conv_conj: | |
"(xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)" | |
proof (induct xs arbitrary: zs) | |
case (Cons x xs zs) then show ?case | |
by (cases zs, auto) | |
qed auto | |
lemma map_eq_append_conv: | |
"map f xs = ys @ zs \<longleftrightarrow> (\<exists>us vs. xs = us @ vs \<and> ys = map f us \<and> zs = map f vs)" | |
proof - | |
have "map f xs \<noteq> ys @ zs \<and> map f xs \<noteq> ys @ zs \<or> map f xs \<noteq> ys @ zs \<or> map f xs = ys @ zs \<and> | |
(\<exists>bs bsa. xs = bs @ bsa \<and> ys = map f bs \<and> zs = map f bsa)" | |
by (metis append_eq_conv_conj append_take_drop_id drop_map take_map) | |
then show ?thesis | |
using map_append by blast | |
qed | |
lemmas append_eq_map_conv = map_eq_append_conv[THEN eq_iff_swap] | |
lemma take_add: "take (i+j) xs = take i xs @ take j (drop i xs)" | |
proof (induct xs arbitrary: i) | |
case (Cons x xs i) then show ?case | |
by (cases i, auto) | |
qed auto | |
lemma append_eq_append_conv_if: | |
"(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) = | |
(if size xs\<^sub>1 \<le> size ys\<^sub>1 | |
then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \<and> xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2 | |
else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \<and> drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)" | |
proof (induct xs\<^sub>1 arbitrary: ys\<^sub>1) | |
case (Cons a xs\<^sub>1 ys\<^sub>1) then show ?case | |
by (cases ys\<^sub>1, auto) | |
qed auto | |
lemma take_hd_drop: | |
"n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs" | |
by (induct xs arbitrary: n) (simp_all add:drop_Cons split:nat.split) | |
lemma id_take_nth_drop: | |
"i < length xs \<Longrightarrow> xs = take i xs @ xs!i # drop (Suc i) xs" | |
proof - | |
assume si: "i < length xs" | |
hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto | |
moreover | |
from si have "take (Suc i) xs = take i xs @ [xs!i]" | |
using take_Suc_conv_app_nth by blast | |
ultimately show ?thesis by auto | |
qed | |
lemma take_update_cancel[simp]: "n \<le> m \<Longrightarrow> take n (xs[m := y]) = take n xs" | |
by(simp add: list_eq_iff_nth_eq) | |
lemma drop_update_cancel[simp]: "n < m \<Longrightarrow> drop m (xs[n := x]) = drop m xs" | |
by(simp add: list_eq_iff_nth_eq) | |
lemma upd_conv_take_nth_drop: | |
"i < length xs \<Longrightarrow> xs[i:=a] = take i xs @ a # drop (Suc i) xs" | |
proof - | |
assume i: "i < length xs" | |
have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]" | |
by(rule arg_cong[OF id_take_nth_drop[OF i]]) | |
also have "\<dots> = take i xs @ a # drop (Suc i) xs" | |
using i by (simp add: list_update_append) | |
finally show ?thesis . | |
qed | |
lemma take_update_swap: "take m (xs[n := x]) = (take m xs)[n := x]" | |
proof (cases "n \<ge> length xs") | |
case False | |
then show ?thesis | |
by (simp add: upd_conv_take_nth_drop take_Cons drop_take min_def diff_Suc split: nat.split) | |
qed auto | |
lemma drop_update_swap: | |
assumes "m \<le> n" shows "drop m (xs[n := x]) = (drop m xs)[n-m := x]" | |
proof (cases "n \<ge> length xs") | |
case False | |
with assms show ?thesis | |
by (simp add: upd_conv_take_nth_drop drop_take) | |
qed auto | |
lemma nth_image: "l \<le> size xs \<Longrightarrow> nth xs ` {0..<l} = set(take l xs)" | |
by (simp add: set_conv_nth) force | |
subsubsection \<open>\<^const>\<open>takeWhile\<close> and \<^const>\<open>dropWhile\<close>\<close> | |
lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs" | |
by (induct xs) auto | |
lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs" | |
by (induct xs) auto | |
lemma takeWhile_append1 [simp]: | |
"\<lbrakk>x \<in> set xs; \<not>P(x)\<rbrakk> \<Longrightarrow> takeWhile P (xs @ ys) = takeWhile P xs" | |
by (induct xs) auto | |
lemma takeWhile_append2 [simp]: | |
"(\<And>x. x \<in> set xs \<Longrightarrow> P x) \<Longrightarrow> takeWhile P (xs @ ys) = xs @ takeWhile P ys" | |
by (induct xs) auto | |
lemma takeWhile_append: | |
"takeWhile P (xs @ ys) = (if \<forall>x\<in>set xs. P x then xs @ takeWhile P ys else takeWhile P xs)" | |
using takeWhile_append1[of _ xs P ys] takeWhile_append2[of xs P ys] by auto | |
lemma takeWhile_tail: "\<not> P x \<Longrightarrow> takeWhile P (xs @ (x#l)) = takeWhile P xs" | |
by (induct xs) auto | |
lemma takeWhile_eq_Nil_iff: "takeWhile P xs = [] \<longleftrightarrow> xs = [] \<or> \<not>P (hd xs)" | |
by (cases xs) auto | |
lemma takeWhile_nth: "j < length (takeWhile P xs) \<Longrightarrow> takeWhile P xs ! j = xs ! j" | |
by (metis nth_append takeWhile_dropWhile_id) | |
lemma takeWhile_takeWhile: "takeWhile Q (takeWhile P xs) = takeWhile (\<lambda>x. P x \<and> Q x) xs" | |
by(induct xs) simp_all | |
lemma dropWhile_nth: "j < length (dropWhile P xs) \<Longrightarrow> | |
dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))" | |
by (metis add.commute nth_append_length_plus takeWhile_dropWhile_id) | |
lemma length_dropWhile_le: "length (dropWhile P xs) \<le> length xs" | |
by (induct xs) auto | |
lemma dropWhile_append1 [simp]: | |
"\<lbrakk>x \<in> set xs; \<not>P(x)\<rbrakk> \<Longrightarrow> dropWhile P (xs @ ys) = (dropWhile P xs)@ys" | |
by (induct xs) auto | |
lemma dropWhile_append2 [simp]: | |
"(\<And>x. x \<in> set xs \<Longrightarrow> P(x)) \<Longrightarrow> dropWhile P (xs @ ys) = dropWhile P ys" | |
by (induct xs) auto | |
lemma dropWhile_append3: | |
"\<not> P y \<Longrightarrow>dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys" | |
by (induct xs) auto | |
lemma dropWhile_append: | |
"dropWhile P (xs @ ys) = (if \<forall>x\<in>set xs. P x then dropWhile P ys else dropWhile P xs @ ys)" | |
using dropWhile_append1[of _ xs P ys] dropWhile_append2[of xs P ys] by auto | |
lemma dropWhile_last: | |
"x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> last (dropWhile P xs) = last xs" | |
by (auto simp add: dropWhile_append3 in_set_conv_decomp) | |
lemma set_dropWhileD: "x \<in> set (dropWhile P xs) \<Longrightarrow> x \<in> set xs" | |
by (induct xs) (auto split: if_split_asm) | |
lemma set_takeWhileD: "x \<in> set (takeWhile P xs) \<Longrightarrow> x \<in> set xs \<and> P x" | |
by (induct xs) (auto split: if_split_asm) | |
lemma takeWhile_eq_all_conv[simp]: | |
"(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)" | |
by(induct xs, auto) | |
lemma dropWhile_eq_Nil_conv[simp]: | |
"(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)" | |
by(induct xs, auto) | |
lemma dropWhile_eq_Cons_conv: | |
"(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \<and> \<not> P y)" | |
by(induct xs, auto) | |
lemma dropWhile_eq_self_iff: "dropWhile P xs = xs \<longleftrightarrow> xs = [] \<or> \<not>P (hd xs)" | |
by (cases xs) (auto simp: dropWhile_eq_Cons_conv) | |
lemma dropWhile_dropWhile1: "(\<And>x. Q x \<Longrightarrow> P x) \<Longrightarrow> dropWhile Q (dropWhile P xs) = dropWhile P xs" | |
by(induct xs) simp_all | |
lemma dropWhile_dropWhile2: "(\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> takeWhile P (takeWhile Q xs) = takeWhile P xs" | |
by(induct xs) simp_all | |
lemma dropWhile_takeWhile: | |
"(\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> dropWhile P (takeWhile Q xs) = takeWhile Q (dropWhile P xs)" | |
by (induction xs) auto | |
lemma distinct_takeWhile[simp]: "distinct xs \<Longrightarrow> distinct (takeWhile P xs)" | |
by (induct xs) (auto dest: set_takeWhileD) | |
lemma distinct_dropWhile[simp]: "distinct xs \<Longrightarrow> distinct (dropWhile P xs)" | |
by (induct xs) auto | |
lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)" | |
by (induct xs) auto | |
lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \<circ> f) xs)" | |
by (induct xs) auto | |
lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs" | |
by (induct xs) auto | |
lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs" | |
by (induct xs) auto | |
lemma hd_dropWhile: "dropWhile P xs \<noteq> [] \<Longrightarrow> \<not> P (hd (dropWhile P xs))" | |
by (induct xs) auto | |
lemma takeWhile_eq_filter: | |
assumes "\<And> x. x \<in> set (dropWhile P xs) \<Longrightarrow> \<not> P x" | |
shows "takeWhile P xs = filter P xs" | |
proof - | |
have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)" | |
by simp | |
have B: "filter P (dropWhile P xs) = []" | |
unfolding filter_empty_conv using assms by blast | |
have "filter P xs = takeWhile P xs" | |
unfolding A filter_append B | |
by (auto simp add: filter_id_conv dest: set_takeWhileD) | |
thus ?thesis .. | |
qed | |
lemma takeWhile_eq_take_P_nth: | |
"\<lbrakk> \<And> i. \<lbrakk> i < n ; i < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) ; n < length xs \<Longrightarrow> \<not> P (xs ! n) \<rbrakk> \<Longrightarrow> | |
takeWhile P xs = take n xs" | |
proof (induct xs arbitrary: n) | |
case Nil | |
thus ?case by simp | |
next | |
case (Cons x xs) | |
show ?case | |
proof (cases n) | |
case 0 | |
with Cons show ?thesis by simp | |
next | |
case [simp]: (Suc n') | |
have "P x" using Cons.prems(1)[of 0] by simp | |
moreover have "takeWhile P xs = take n' xs" | |
proof (rule Cons.hyps) | |
fix i | |
assume "i < n'" "i < length xs" | |
thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp | |
next | |
assume "n' < length xs" | |
thus "\<not> P (xs ! n')" using Cons by auto | |
qed | |
ultimately show ?thesis by simp | |
qed | |
qed | |
lemma nth_length_takeWhile: | |
"length (takeWhile P xs) < length xs \<Longrightarrow> \<not> P (xs ! length (takeWhile P xs))" | |
by (induct xs) auto | |
lemma length_takeWhile_less_P_nth: | |
assumes all: "\<And> i. i < j \<Longrightarrow> P (xs ! i)" and "j \<le> length xs" | |
shows "j \<le> length (takeWhile P xs)" | |
proof (rule classical) | |
assume "\<not> ?thesis" | |
hence "length (takeWhile P xs) < length xs" using assms by simp | |
thus ?thesis using all \<open>\<not> ?thesis\<close> nth_length_takeWhile[of P xs] by auto | |
qed | |
lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow> | |
takeWhile (\<lambda>y. y \<noteq> x) (rev xs) = rev (tl (dropWhile (\<lambda>y. y \<noteq> x) xs))" | |
by(induct xs) (auto simp: takeWhile_tail[where l="[]"]) | |
lemma dropWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow> | |
dropWhile (\<lambda>y. y \<noteq> x) (rev xs) = x # rev (takeWhile (\<lambda>y. y \<noteq> x) xs)" | |
proof (induct xs) | |
case (Cons a xs) | |
then show ?case | |
by(auto, subst dropWhile_append2, auto) | |
qed simp | |
lemma takeWhile_not_last: | |
"distinct xs \<Longrightarrow> takeWhile (\<lambda>y. y \<noteq> last xs) xs = butlast xs" | |
by(induction xs rule: induct_list012) auto | |
lemma takeWhile_cong [fundef_cong]: | |
"\<lbrakk>l = k; \<And>x. x \<in> set l \<Longrightarrow> P x = Q x\<rbrakk> | |
\<Longrightarrow> takeWhile P l = takeWhile Q k" | |
by (induct k arbitrary: l) (simp_all) | |
lemma dropWhile_cong [fundef_cong]: | |
"\<lbrakk>l = k; \<And>x. x \<in> set l \<Longrightarrow> P x = Q x\<rbrakk> | |
\<Longrightarrow> dropWhile P l = dropWhile Q k" | |
by (induct k arbitrary: l, simp_all) | |
lemma takeWhile_idem [simp]: | |
"takeWhile P (takeWhile P xs) = takeWhile P xs" | |
by (induct xs) auto | |
lemma dropWhile_idem [simp]: | |
"dropWhile P (dropWhile P xs) = dropWhile P xs" | |
by (induct xs) auto | |
subsubsection \<open>\<^const>\<open>zip\<close>\<close> | |
lemma zip_Nil [simp]: "zip [] ys = []" | |
by (induct ys) auto | |
lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys" | |
by simp | |
declare zip_Cons [simp del] | |
lemma [code]: | |
"zip [] ys = []" | |
"zip xs [] = []" | |
"zip (x # xs) (y # ys) = (x, y) # zip xs ys" | |
by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+ | |
lemma zip_Cons1: | |
"zip (x#xs) ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x,y)#zip xs ys)" | |
by(auto split:list.split) | |
lemma length_zip [simp]: | |
"length (zip xs ys) = min (length xs) (length ys)" | |
by (induct xs ys rule:list_induct2') auto | |
lemma zip_obtain_same_length: | |
assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys) | |
\<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)" | |
shows "P (zip xs ys)" | |
proof - | |
let ?n = "min (length xs) (length ys)" | |
have "P (zip (take ?n xs) (take ?n ys))" | |
by (rule assms) simp_all | |
moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)" | |
proof (induct xs arbitrary: ys) | |
case Nil then show ?case by simp | |
next | |
case (Cons x xs) then show ?case by (cases ys) simp_all | |
qed | |
ultimately show ?thesis by simp | |
qed | |
lemma zip_append1: | |
"zip (xs @ ys) zs = | |
zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)" | |
by (induct xs zs rule:list_induct2') auto | |
lemma zip_append2: | |
"zip xs (ys @ zs) = | |
zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs" | |
by (induct xs ys rule:list_induct2') auto | |
lemma zip_append [simp]: | |
"\<lbrakk>length xs = length us\<rbrakk> \<Longrightarrow> | |
zip (xs@ys) (us@vs) = zip xs us @ zip ys vs" | |
by (simp add: zip_append1) | |
lemma zip_rev: | |
"length xs = length ys \<Longrightarrow> zip (rev xs) (rev ys) = rev (zip xs ys)" | |
by (induct rule:list_induct2, simp_all) | |
lemma zip_map_map: | |
"zip (map f xs) (map g ys) = map (\<lambda> (x, y). (f x, g y)) (zip xs ys)" | |
proof (induct xs arbitrary: ys) | |
case (Cons x xs) note Cons_x_xs = Cons.hyps | |
show ?case | |
proof (cases ys) | |
case (Cons y ys') | |
show ?thesis unfolding Cons using Cons_x_xs by simp | |
qed simp | |
qed simp | |
lemma zip_map1: | |
"zip (map f xs) ys = map (\<lambda>(x, y). (f x, y)) (zip xs ys)" | |
using zip_map_map[of f xs "\<lambda>x. x" ys] by simp | |
lemma zip_map2: | |
"zip xs (map f ys) = map (\<lambda>(x, y). (x, f y)) (zip xs ys)" | |
using zip_map_map[of "\<lambda>x. x" xs f ys] by simp | |
lemma map_zip_map: | |
"map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)" | |
by (auto simp: zip_map1) | |
lemma map_zip_map2: | |
"map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)" | |
by (auto simp: zip_map2) | |
text\<open>Courtesy of Andreas Lochbihler:\<close> | |
lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs" | |
by(induct xs) auto | |
lemma nth_zip [simp]: | |
"\<lbrakk>i < length xs; i < length ys\<rbrakk> \<Longrightarrow> (zip xs ys)!i = (xs!i, ys!i)" | |
proof (induct ys arbitrary: i xs) | |
case (Cons y ys) | |
then show ?case | |
by (cases xs) (simp_all add: nth.simps split: nat.split) | |
qed auto | |
lemma set_zip: | |
"set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}" | |
by(simp add: set_conv_nth cong: rev_conj_cong) | |
lemma zip_same: "((a,b) \<in> set (zip xs xs)) = (a \<in> set xs \<and> a = b)" | |
by(induct xs) auto | |
lemma zip_update: "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]" | |
by (simp add: update_zip) | |
lemma zip_replicate [simp]: | |
"zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)" | |
proof (induct i arbitrary: j) | |
case (Suc i) | |
then show ?case | |
by (cases j, auto) | |
qed auto | |
lemma zip_replicate1: "zip (replicate n x) ys = map (Pair x) (take n ys)" | |
by(induction ys arbitrary: n)(case_tac [2] n, simp_all) | |
lemma take_zip: "take n (zip xs ys) = zip (take n xs) (take n ys)" | |
proof (induct n arbitrary: xs ys) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs; cases ys) simp_all | |
qed | |
lemma drop_zip: "drop n (zip xs ys) = zip (drop n xs) (drop n ys)" | |
proof (induct n arbitrary: xs ys) | |
case 0 | |
then show ?case by simp | |
next | |
case Suc | |
then show ?case by (cases xs; cases ys) simp_all | |
qed | |
lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \<circ> fst) (zip xs ys)" | |
proof (induct xs arbitrary: ys) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases ys) auto | |
qed | |
lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \<circ> snd) (zip xs ys)" | |
proof (induct xs arbitrary: ys) | |
case Nil | |
then show ?case by simp | |
next | |
case Cons | |
then show ?case by (cases ys) auto | |
qed | |
lemma set_zip_leftD: "(x,y)\<in> set (zip xs ys) \<Longrightarrow> x \<in> set xs" | |
by (induct xs ys rule:list_induct2') auto | |
lemma set_zip_rightD: "(x,y)\<in> set (zip xs ys) \<Longrightarrow> y \<in> set ys" | |
by (induct xs ys rule:list_induct2') auto | |
lemma in_set_zipE: | |
"(x,y) \<in> set(zip xs ys) \<Longrightarrow> (\<lbrakk> x \<in> set xs; y \<in> set ys \<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R" | |
by(blast dest: set_zip_leftD set_zip_rightD) | |
lemma zip_map_fst_snd: "zip (map fst zs) (map snd zs) = zs" | |
by (induct zs) simp_all | |
lemma zip_eq_conv: | |
"length xs = length ys \<Longrightarrow> zip xs ys = zs \<longleftrightarrow> map fst zs = xs \<and> map snd zs = ys" | |
by (auto simp add: zip_map_fst_snd) | |
lemma in_set_zip: | |
"p \<in> set (zip xs ys) \<longleftrightarrow> (\<exists>n. xs ! n = fst p \<and> ys ! n = snd p | |
\<and> n < length xs \<and> n < length ys)" | |
by (cases p) (auto simp add: set_zip) | |
lemma in_set_impl_in_set_zip1: | |
assumes "length xs = length ys" | |
assumes "x \<in> set xs" | |
obtains y where "(x, y) \<in> set (zip xs ys)" | |
proof - | |
from assms have "x \<in> set (map fst (zip xs ys))" by simp | |
from this that show ?thesis by fastforce | |
qed | |
lemma in_set_impl_in_set_zip2: | |
assumes "length xs = length ys" | |
assumes "y \<in> set ys" | |
obtains x where "(x, y) \<in> set (zip xs ys)" | |
proof - | |
from assms have "y \<in> set (map snd (zip xs ys))" by simp | |
from this that show ?thesis by fastforce | |
qed | |
lemma zip_eq_Nil_iff[simp]: | |
"zip xs ys = [] \<longleftrightarrow> xs = [] \<or> ys = []" | |
by (cases xs; cases ys) simp_all | |
lemmas Nil_eq_zip_iff[simp] = zip_eq_Nil_iff[THEN eq_iff_swap] | |
lemma zip_eq_ConsE: | |
assumes "zip xs ys = xy # xys" | |
obtains x xs' y ys' where "xs = x # xs'" | |
and "ys = y # ys'" and "xy = (x, y)" | |
and "xys = zip xs' ys'" | |
proof - | |
from assms have "xs \<noteq> []" and "ys \<noteq> []" | |
using zip_eq_Nil_iff [of xs ys] by simp_all | |
then obtain x xs' y ys' where xs: "xs = x # xs'" | |
and ys: "ys = y # ys'" | |
by (cases xs; cases ys) auto | |
with assms have "xy = (x, y)" and "xys = zip xs' ys'" | |
by simp_all | |
with xs ys show ?thesis .. | |
qed | |
lemma semilattice_map2: | |
"semilattice (map2 (\<^bold>*))" if "semilattice (\<^bold>*)" | |
for f (infixl "\<^bold>*" 70) | |
proof - | |
from that interpret semilattice f . | |
show ?thesis | |
proof | |
show "map2 (\<^bold>*) (map2 (\<^bold>*) xs ys) zs = map2 (\<^bold>*) xs (map2 (\<^bold>*) ys zs)" | |
for xs ys zs :: "'a list" | |
proof (induction "zip xs (zip ys zs)" arbitrary: xs ys zs) | |
case Nil | |
from Nil [symmetric] show ?case | |
by auto | |
next | |
case (Cons xyz xyzs) | |
from Cons.hyps(2) [symmetric] show ?case | |
by (rule zip_eq_ConsE) (erule zip_eq_ConsE, | |
auto intro: Cons.hyps(1) simp add: ac_simps) | |
qed | |
show "map2 (\<^bold>*) xs ys = map2 (\<^bold>*) ys xs" | |
for xs ys :: "'a list" | |
proof (induction "zip xs ys" arbitrary: xs ys) | |
case Nil | |
then show ?case | |
by auto | |
next | |
case (Cons xy xys) | |
from Cons.hyps(2) [symmetric] show ?case | |
by (rule zip_eq_ConsE) (auto intro: Cons.hyps(1) simp add: ac_simps) | |
qed | |
show "map2 (\<^bold>*) xs xs = xs" | |
for xs :: "'a list" | |
by (induction xs) simp_all | |
qed | |
qed | |
lemma pair_list_eqI: | |
assumes "map fst xs = map fst ys" and "map snd xs = map snd ys" | |
shows "xs = ys" | |
proof - | |
from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq) | |
from this assms show ?thesis | |
by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI) | |
qed | |
lemma hd_zip: | |
\<open>hd (zip xs ys) = (hd xs, hd ys)\<close> if \<open>xs \<noteq> []\<close> and \<open>ys \<noteq> []\<close> | |
using that by (cases xs; cases ys) simp_all | |
lemma last_zip: | |
\<open>last (zip xs ys) = (last xs, last ys)\<close> if \<open>xs \<noteq> []\<close> and \<open>ys \<noteq> []\<close> | |
and \<open>length xs = length ys\<close> | |
using that by (cases xs rule: rev_cases; cases ys rule: rev_cases) simp_all | |
subsubsection \<open>\<^const>\<open>list_all2\<close>\<close> | |
lemma list_all2_lengthD [intro?]: | |
"list_all2 P xs ys \<Longrightarrow> length xs = length ys" | |
by (simp add: list_all2_iff) | |
lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])" | |
by (simp add: list_all2_iff) | |
lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])" | |
by (simp add: list_all2_iff) | |
lemma list_all2_Cons [iff, code]: | |
"list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)" | |
by (auto simp add: list_all2_iff) | |
lemma list_all2_Cons1: | |
"list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)" | |
by (cases ys) auto | |
lemma list_all2_Cons2: | |
"list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)" | |
by (cases xs) auto | |
lemma list_all2_induct | |
[consumes 1, case_names Nil Cons, induct set: list_all2]: | |
assumes P: "list_all2 P xs ys" | |
assumes Nil: "R [] []" | |
assumes Cons: "\<And>x xs y ys. | |
\<lbrakk>P x y; list_all2 P xs ys; R xs ys\<rbrakk> \<Longrightarrow> R (x # xs) (y # ys)" | |
shows "R xs ys" | |
using P | |
by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons) | |
lemma list_all2_rev [iff]: | |
"list_all2 P (rev xs) (rev ys) = list_all2 P xs ys" | |
by (simp add: list_all2_iff zip_rev cong: conj_cong) | |
lemma list_all2_rev1: | |
"list_all2 P (rev xs) ys = list_all2 P xs (rev ys)" | |
by (subst list_all2_rev [symmetric]) simp | |
lemma list_all2_append1: | |
"list_all2 P (xs @ ys) zs = | |
(\<exists>us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and> | |
list_all2 P xs us \<and> list_all2 P ys vs)" (is "?lhs = ?rhs") | |
proof | |
assume ?lhs | |
then show ?rhs | |
apply (rule_tac x = "take (length xs) zs" in exI) | |
apply (rule_tac x = "drop (length xs) zs" in exI) | |
apply (force split: nat_diff_split simp add: list_all2_iff zip_append1) | |
done | |
next | |
assume ?rhs | |
then show ?lhs | |
by (auto simp add: list_all2_iff) | |
qed | |
lemma list_all2_append2: | |
"list_all2 P xs (ys @ zs) = | |
(\<exists>us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and> | |
list_all2 P us ys \<and> list_all2 P vs zs)" (is "?lhs = ?rhs") | |
proof | |
assume ?lhs | |
then show ?rhs | |
apply (rule_tac x = "take (length ys) xs" in exI) | |
apply (rule_tac x = "drop (length ys) xs" in exI) | |
apply (force split: nat_diff_split simp add: list_all2_iff zip_append2) | |
done | |
next | |
assume ?rhs | |
then show ?lhs | |
by (auto simp add: list_all2_iff) | |
qed | |
lemma list_all2_append: | |
"length xs = length ys \<Longrightarrow> | |
list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \<and> list_all2 P us vs)" | |
by (induct rule:list_induct2, simp_all) | |
lemma list_all2_appendI [intro?, trans]: | |
"\<lbrakk> list_all2 P a b; list_all2 P c d \<rbrakk> \<Longrightarrow> list_all2 P (a@c) (b@d)" | |
by (simp add: list_all2_append list_all2_lengthD) | |
lemma list_all2_conv_all_nth: | |
"list_all2 P xs ys = | |
(length xs = length ys \<and> (\<forall>i < length xs. P (xs!i) (ys!i)))" | |
by (force simp add: list_all2_iff set_zip) | |
lemma list_all2_trans: | |
assumes tr: "!!a b c. P1 a b \<Longrightarrow> P2 b c \<Longrightarrow> P3 a c" | |
shows "!!bs cs. list_all2 P1 as bs \<Longrightarrow> list_all2 P2 bs cs \<Longrightarrow> list_all2 P3 as cs" | |
(is "!!bs cs. PROP ?Q as bs cs") | |
proof (induct as) | |
fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs" | |
show "!!cs. PROP ?Q (x # xs) bs cs" | |
proof (induct bs) | |
fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs" | |
show "PROP ?Q (x # xs) (y # ys) cs" | |
by (induct cs) (auto intro: tr I1 I2) | |
qed simp | |
qed simp | |
lemma list_all2_all_nthI [intro?]: | |
"length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> P (a!n) (b!n)) \<Longrightarrow> list_all2 P a b" | |
by (simp add: list_all2_conv_all_nth) | |
lemma list_all2I: | |
"\<forall>x \<in> set (zip a b). case_prod P x \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b" | |
by (simp add: list_all2_iff) | |
lemma list_all2_nthD: | |
"\<lbrakk> list_all2 P xs ys; p < size xs \<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)" | |
by (simp add: list_all2_conv_all_nth) | |
lemma list_all2_nthD2: | |
"\<lbrakk>list_all2 P xs ys; p < size ys\<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)" | |
by (frule list_all2_lengthD) (auto intro: list_all2_nthD) | |
lemma list_all2_map1: | |
"list_all2 P (map f as) bs = list_all2 (\<lambda>x y. P (f x) y) as bs" | |
by (simp add: list_all2_conv_all_nth) | |
lemma list_all2_map2: | |
"list_all2 P as (map f bs) = list_all2 (\<lambda>x y. P x (f y)) as bs" | |
by (auto simp add: list_all2_conv_all_nth) | |
lemma list_all2_refl [intro?]: | |
"(\<And>x. P x x) \<Longrightarrow> list_all2 P xs xs" | |
by (simp add: list_all2_conv_all_nth) | |
lemma list_all2_update_cong: | |
"\<lbrakk> list_all2 P xs ys; P x y \<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])" | |
by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update) | |
lemma list_all2_takeI [simp,intro?]: | |
"list_all2 P xs ys \<Longrightarrow> list_all2 P (take n xs) (take n ys)" | |
proof (induct xs arbitrary: n ys) | |
case (Cons x xs) | |
then show ?case | |
by (cases n) (auto simp: list_all2_Cons1) | |
qed auto | |
lemma list_all2_dropI [simp,intro?]: | |
"list_all2 P xs ys \<Longrightarrow> list_all2 P (drop n xs) (drop n ys)" | |
proof (induct xs arbitrary: n ys) | |
case (Cons x xs) | |
then show ?case | |
by (cases n) (auto simp: list_all2_Cons1) | |
qed auto | |
lemma list_all2_mono [intro?]: | |
"list_all2 P xs ys \<Longrightarrow> (\<And>xs ys. P xs ys \<Longrightarrow> Q xs ys) \<Longrightarrow> list_all2 Q xs ys" | |
by (rule list.rel_mono_strong) | |
lemma list_all2_eq: | |
"xs = ys \<longleftrightarrow> list_all2 (=) xs ys" | |
by (induct xs ys rule: list_induct2') auto | |
lemma list_eq_iff_zip_eq: | |
"xs = ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x,y) \<in> set (zip xs ys). x = y)" | |
by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong) | |
lemma list_all2_same: "list_all2 P xs xs \<longleftrightarrow> (\<forall>x\<in>set xs. P x x)" | |
by(auto simp add: list_all2_conv_all_nth set_conv_nth) | |
lemma zip_assoc: | |
"zip xs (zip ys zs) = map (\<lambda>((x, y), z). (x, y, z)) (zip (zip xs ys) zs)" | |
by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all | |
lemma zip_commute: "zip xs ys = map (\<lambda>(x, y). (y, x)) (zip ys xs)" | |
by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all | |
lemma zip_left_commute: | |
"zip xs (zip ys zs) = map (\<lambda>(y, (x, z)). (x, y, z)) (zip ys (zip xs zs))" | |
by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all | |
lemma zip_replicate2: "zip xs (replicate n y) = map (\<lambda>x. (x, y)) (take n xs)" | |
by(subst zip_commute)(simp add: zip_replicate1) | |
subsubsection \<open>\<^const>\<open>List.product\<close> and \<^const>\<open>product_lists\<close>\<close> | |
lemma product_concat_map: | |
"List.product xs ys = concat (map (\<lambda>x. map (\<lambda>y. (x,y)) ys) xs)" | |
by(induction xs) (simp)+ | |
lemma set_product[simp]: "set (List.product xs ys) = set xs \<times> set ys" | |
by (induct xs) auto | |
lemma length_product [simp]: | |
"length (List.product xs ys) = length xs * length ys" | |
by (induct xs) simp_all | |
lemma product_nth: | |
assumes "n < length xs * length ys" | |
shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))" | |
using assms proof (induct xs arbitrary: n) | |
case Nil then show ?case by simp | |
next | |
case (Cons x xs n) | |
then have "length ys > 0" by auto | |
with Cons show ?case | |
by (auto simp add: nth_append not_less le_mod_geq le_div_geq) | |
qed | |
lemma in_set_product_lists_length: | |
"xs \<in> set (product_lists xss) \<Longrightarrow> length xs = length xss" | |
by (induct xss arbitrary: xs) auto | |
lemma product_lists_set: | |
"set (product_lists xss) = {xs. list_all2 (\<lambda>x ys. x \<in> set ys) xs xss}" (is "?L = Collect ?R") | |
proof (intro equalityI subsetI, unfold mem_Collect_eq) | |
fix xs assume "xs \<in> ?L" | |
then have "length xs = length xss" by (rule in_set_product_lists_length) | |
from this \<open>xs \<in> ?L\<close> show "?R xs" by (induct xs xss rule: list_induct2) auto | |
next | |
fix xs assume "?R xs" | |
then show "xs \<in> ?L" by induct auto | |
qed | |
subsubsection \<open>\<^const>\<open>fold\<close> with natural argument order\<close> | |
lemma fold_simps [code]: \<comment> \<open>eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala\<close> | |
"fold f [] s = s" | |
"fold f (x # xs) s = fold f xs (f x s)" | |
by simp_all | |
lemma fold_remove1_split: | |
"\<lbrakk> \<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x; | |
x \<in> set xs \<rbrakk> | |
\<Longrightarrow> fold f xs = fold f (remove1 x xs) \<circ> f x" | |
by (induct xs) (auto simp add: comp_assoc) | |
lemma fold_cong [fundef_cong]: | |
"a = b \<Longrightarrow> xs = ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> f x = g x) | |
\<Longrightarrow> fold f xs a = fold g ys b" | |
by (induct ys arbitrary: a b xs) simp_all | |
lemma fold_id: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = id) \<Longrightarrow> fold f xs = id" | |
by (induct xs) simp_all | |
lemma fold_commute: | |
"(\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h) \<Longrightarrow> h \<circ> fold g xs = fold f xs \<circ> h" | |
by (induct xs) (simp_all add: fun_eq_iff) | |
lemma fold_commute_apply: | |
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h" | |
shows "h (fold g xs s) = fold f xs (h s)" | |
proof - | |
from assms have "h \<circ> fold g xs = fold f xs \<circ> h" by (rule fold_commute) | |
then show ?thesis by (simp add: fun_eq_iff) | |
qed | |
lemma fold_invariant: | |
"\<lbrakk> \<And>x. x \<in> set xs \<Longrightarrow> Q x; P s; \<And>x s. Q x \<Longrightarrow> P s \<Longrightarrow> P (f x s) \<rbrakk> | |
\<Longrightarrow> P (fold f xs s)" | |
by (induct xs arbitrary: s) simp_all | |
lemma fold_append [simp]: "fold f (xs @ ys) = fold f ys \<circ> fold f xs" | |
by (induct xs) simp_all | |
lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g \<circ> f) xs" | |
by (induct xs) simp_all | |
lemma fold_filter: | |
"fold f (filter P xs) = fold (\<lambda>x. if P x then f x else id) xs" | |
by (induct xs) simp_all | |
lemma fold_rev: | |
"(\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y) | |
\<Longrightarrow> fold f (rev xs) = fold f xs" | |
by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff) | |
lemma fold_Cons_rev: "fold Cons xs = append (rev xs)" | |
by (induct xs) simp_all | |
lemma rev_conv_fold [code]: "rev xs = fold Cons xs []" | |
by (simp add: fold_Cons_rev) | |
lemma fold_append_concat_rev: "fold append xss = append (concat (rev xss))" | |
by (induct xss) simp_all | |
text \<open>\<^const>\<open>Finite_Set.fold\<close> and \<^const>\<open>fold\<close>\<close> | |
lemma (in comp_fun_commute_on) fold_set_fold_remdups: | |
assumes "set xs \<subseteq> S" | |
shows "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" | |
by (rule sym, use assms in \<open>induct xs arbitrary: y\<close>) | |
(simp_all add: insert_absorb fold_fun_left_comm) | |
lemma (in comp_fun_idem_on) fold_set_fold: | |
assumes "set xs \<subseteq> S" | |
shows "Finite_Set.fold f y (set xs) = fold f xs y" | |
by (rule sym, use assms in \<open>induct xs arbitrary: y\<close>) (simp_all add: fold_fun_left_comm) | |
lemma union_set_fold [code]: "set xs \<union> A = fold Set.insert xs A" | |
proof - | |
interpret comp_fun_idem Set.insert | |
by (fact comp_fun_idem_insert) | |
show ?thesis by (simp add: union_fold_insert fold_set_fold) | |
qed | |
lemma union_coset_filter [code]: | |
"List.coset xs \<union> A = List.coset (List.filter (\<lambda>x. x \<notin> A) xs)" | |
by auto | |
lemma minus_set_fold [code]: "A - set xs = fold Set.remove xs A" | |
proof - | |
interpret comp_fun_idem Set.remove | |
by (fact comp_fun_idem_remove) | |
show ?thesis | |
by (simp add: minus_fold_remove [of _ A] fold_set_fold) | |
qed | |
lemma minus_coset_filter [code]: | |
"A - List.coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)" | |
by auto | |
lemma inter_set_filter [code]: | |
"A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)" | |
by auto | |
lemma inter_coset_fold [code]: | |
"A \<inter> List.coset xs = fold Set.remove xs A" | |
by (simp add: Diff_eq [symmetric] minus_set_fold) | |
lemma (in semilattice_set) set_eq_fold [code]: | |
"F (set (x # xs)) = fold f xs x" | |
proof - | |
interpret comp_fun_idem f | |
by standard (simp_all add: fun_eq_iff left_commute) | |
show ?thesis by (simp add: eq_fold fold_set_fold) | |
qed | |
lemma (in complete_lattice) Inf_set_fold: | |
"Inf (set xs) = fold inf xs top" | |
proof - | |
interpret comp_fun_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
by (fact comp_fun_idem_inf) | |
show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute) | |
qed | |
declare Inf_set_fold [where 'a = "'a set", code] | |
lemma (in complete_lattice) Sup_set_fold: | |
"Sup (set xs) = fold sup xs bot" | |
proof - | |
interpret comp_fun_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
by (fact comp_fun_idem_sup) | |
show ?thesis by (simp add: Sup_fold_sup fold_set_fold sup_commute) | |
qed | |
declare Sup_set_fold [where 'a = "'a set", code] | |
lemma (in complete_lattice) INF_set_fold: | |
"\<Sqinter>(f ` set xs) = fold (inf \<circ> f) xs top" | |
using Inf_set_fold [of "map f xs"] by (simp add: fold_map) | |
lemma (in complete_lattice) SUP_set_fold: | |
"\<Squnion>(f ` set xs) = fold (sup \<circ> f) xs bot" | |
using Sup_set_fold [of "map f xs"] by (simp add: fold_map) | |
subsubsection \<open>Fold variants: \<^const>\<open>foldr\<close> and \<^const>\<open>foldl\<close>\<close> | |
text \<open>Correspondence\<close> | |
lemma foldr_conv_fold [code_abbrev]: "foldr f xs = fold f (rev xs)" | |
by (induct xs) simp_all | |
lemma foldl_conv_fold: "foldl f s xs = fold (\<lambda>x s. f s x) xs s" | |
by (induct xs arbitrary: s) simp_all | |
lemma foldr_conv_foldl: \<comment> \<open>The ``Third Duality Theorem'' in Bird \& Wadler:\<close> | |
"foldr f xs a = foldl (\<lambda>x y. f y x) a (rev xs)" | |
by (simp add: foldr_conv_fold foldl_conv_fold) | |
lemma foldl_conv_foldr: | |
"foldl f a xs = foldr (\<lambda>x y. f y x) (rev xs) a" | |
by (simp add: foldr_conv_fold foldl_conv_fold) | |
lemma foldr_fold: | |
"(\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y) | |
\<Longrightarrow> foldr f xs = fold f xs" | |
unfolding foldr_conv_fold by (rule fold_rev) | |
lemma foldr_cong [fundef_cong]: | |
"a = b \<Longrightarrow> l = k \<Longrightarrow> (\<And>a x. x \<in> set l \<Longrightarrow> f x a = g x a) \<Longrightarrow> foldr f l a = foldr g k b" | |
by (auto simp add: foldr_conv_fold intro!: fold_cong) | |
lemma foldl_cong [fundef_cong]: | |
"a = b \<Longrightarrow> l = k \<Longrightarrow> (\<And>a x. x \<in> set l \<Longrightarrow> f a x = g a x) \<Longrightarrow> foldl f a l = foldl g b k" | |
by (auto simp add: foldl_conv_fold intro!: fold_cong) | |
lemma foldr_append [simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)" | |
by (simp add: foldr_conv_fold) | |
lemma foldl_append [simp]: "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys" | |
by (simp add: foldl_conv_fold) | |
lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g \<circ> f) xs a" | |
by (simp add: foldr_conv_fold fold_map rev_map) | |
lemma foldr_filter: | |
"foldr f (filter P xs) = foldr (\<lambda>x. if P x then f x else id) xs" | |
by (simp add: foldr_conv_fold rev_filter fold_filter) | |
lemma foldl_map [code_unfold]: | |
"foldl g a (map f xs) = foldl (\<lambda>a x. g a (f x)) a xs" | |
by (simp add: foldl_conv_fold fold_map comp_def) | |
lemma concat_conv_foldr [code]: | |
"concat xss = foldr append xss []" | |
by (simp add: fold_append_concat_rev foldr_conv_fold) | |
subsubsection \<open>\<^const>\<open>upt\<close>\<close> | |
lemma upt_rec[code]: "[i..<j] = (if i<j then i#[Suc i..<j] else [])" | |
\<comment> \<open>simp does not terminate!\<close> | |
by (induct j) auto | |
lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n | |
lemma upt_conv_Nil [simp]: "j \<le> i \<Longrightarrow> [i..<j] = []" | |
by (subst upt_rec) simp | |
lemma upt_eq_Nil_conv[simp]: "([i..<j] = []) = (j = 0 \<or> j \<le> i)" | |
by(induct j)simp_all | |
lemma upt_eq_Cons_conv: | |
"([i..<j] = x#xs) = (i < j \<and> i = x \<and> [i+1..<j] = xs)" | |
proof (induct j arbitrary: x xs) | |
case (Suc j) | |
then show ?case | |
by (simp add: upt_rec) | |
qed simp | |
lemma upt_Suc_append: "i \<le> j \<Longrightarrow> [i..<(Suc j)] = [i..<j]@[j]" | |
\<comment> \<open>Only needed if \<open>upt_Suc\<close> is deleted from the simpset.\<close> | |
by simp | |
lemma upt_conv_Cons: "i < j \<Longrightarrow> [i..<j] = i # [Suc i..<j]" | |
by (simp add: upt_rec) | |
lemma upt_conv_Cons_Cons: \<comment> \<open>no precondition\<close> | |
"m # n # ns = [m..<q] \<longleftrightarrow> n # ns = [Suc m..<q]" | |
proof (cases "m < q") | |
case False then show ?thesis by simp | |
next | |
case True then show ?thesis by (simp add: upt_conv_Cons) | |
qed | |
lemma upt_add_eq_append: "i<=j \<Longrightarrow> [i..<j+k] = [i..<j]@[j..<j+k]" | |
\<comment> \<open>LOOPS as a simprule, since \<open>j \<le> j\<close>.\<close> | |
by (induct k) auto | |
lemma length_upt [simp]: "length [i..<j] = j - i" | |
by (induct j) (auto simp add: Suc_diff_le) | |
lemma nth_upt [simp]: "i + k < j \<Longrightarrow> [i..<j] ! k = i + k" | |
by (induct j) (auto simp add: less_Suc_eq nth_append split: nat_diff_split) | |
lemma hd_upt[simp]: "i < j \<Longrightarrow> hd[i..<j] = i" | |
by(simp add:upt_conv_Cons) | |
lemma tl_upt [simp]: "tl [m..<n] = [Suc m..<n]" | |
by (simp add: upt_rec) | |
lemma last_upt[simp]: "i < j \<Longrightarrow> last[i..<j] = j - 1" | |
by(cases j)(auto simp: upt_Suc_append) | |
lemma take_upt [simp]: "i+m \<le> n \<Longrightarrow> take m [i..<n] = [i..<i+m]" | |
proof (induct m arbitrary: i) | |
case (Suc m) | |
then show ?case | |
by (subst take_Suc_conv_app_nth) auto | |
qed simp | |
lemma drop_upt[simp]: "drop m [i..<j] = [i+m..<j]" | |
by(induct j) auto | |
lemma map_Suc_upt: "map Suc [m..<n] = [Suc m..<Suc n]" | |
by (induct n) auto | |
lemma map_add_upt: "map (\<lambda>i. i + n) [0..<m] = [n..<m + n]" | |
by (induct m) simp_all | |
lemma nth_map_upt: "i < n-m \<Longrightarrow> (map f [m..<n]) ! i = f(m+i)" | |
proof (induct n m arbitrary: i rule: diff_induct) | |
case (3 x y) | |
then show ?case | |
by (metis add.commute length_upt less_diff_conv nth_map nth_upt) | |
qed auto | |
lemma map_decr_upt: "map (\<lambda>n. n - Suc 0) [Suc m..<Suc n] = [m..<n]" | |
by (induct n) simp_all | |
lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\<lambda>i. f (Suc i)) [0 ..< n]" | |
by (induct n arbitrary: f) auto | |
lemma nth_take_lemma: | |
"k \<le> length xs \<Longrightarrow> k \<le> length ys \<Longrightarrow> | |
(\<And>i. i < k \<longrightarrow> xs!i = ys!i) \<Longrightarrow> take k xs = take k ys" | |
proof (induct k arbitrary: xs ys) | |
case (Suc k) | |
then show ?case | |
apply (simp add: less_Suc_eq_0_disj) | |
by (simp add: Suc.prems(3) take_Suc_conv_app_nth) | |
qed simp | |
lemma nth_equalityI: | |
"\<lbrakk>length xs = length ys; \<And>i. i < length xs \<Longrightarrow> xs!i = ys!i\<rbrakk> \<Longrightarrow> xs = ys" | |
by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all | |
lemma map_nth: | |
"map (\<lambda>i. xs ! i) [0..<length xs] = xs" | |
by (rule nth_equalityI, auto) | |
lemma list_all2_antisym: | |
"\<lbrakk> (\<And>x y. \<lbrakk>P x y; Q y x\<rbrakk> \<Longrightarrow> x = y); list_all2 P xs ys; list_all2 Q ys xs \<rbrakk> | |
\<Longrightarrow> xs = ys" | |
by (simp add: list_all2_conv_all_nth nth_equalityI) | |