(* Title: HOL/List.thy
Author: Tobias Nipkow; proofs tidied by LCP
*)
section \<open>The datatype of finite lists\<close>
theory List
imports Sledgehammer Code_Numeral 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 = [] \ P) \ (\a list. y = a # list \ P) \ P"
by (rule list.exhaust)
lemma [case_names Nil Cons, induct type: list]:
\<comment> \<open>for backward compatibility -- names of variables differ\<close>
"P [] \ (\a list. P list \ P (a # list)) \ 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 \ 'a" where
"last (x # xs) = (if xs = [] then x else last xs)"
primrec butlast :: "'a list \ 'a list" where
"butlast [] = []" |
"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
lemma set_rec: "set xs = rec_list {} (\x _. insert x) xs"
by (induct xs) auto
definition coset :: "'a list \ 'a set" where
[simp]: "coset xs = - set xs"
primrec append :: "'a list \ 'a list \ 'a list" (infixr "@" 65) where
append_Nil: "[] @ ys = ys" |
append_Cons: "(x#xs) @ ys = x # xs @ ys"
primrec rev :: "'a list \ 'a list" where
"rev [] = []" |
"rev (x # xs) = rev xs @ [x]"
primrec filter:: "('a \ bool) \ 'a list \ '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[_\_ ./ _])")
translations
"[x<-xs . P]" \<rightharpoonup> "CONST filter (\<lambda>x. P) xs"
primrec fold :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where
fold_Nil: "fold f [] = id" |
fold_Cons: "fold f (x # xs) = fold f xs \ f x"
primrec foldr :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where
foldr_Nil: "foldr f [] = id" |
foldr_Cons: "foldr f (x # xs) = f x \ foldr f xs"
primrec foldl :: "('b \ 'a \ 'b) \ 'b \ 'a list \ '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 \ 'a list" where
"concat [] = []" |
"concat (x # xs) = x @ concat xs"
primrec drop:: "nat \ 'a list \ 'a list" where
drop_Nil: "drop n [] = []" |
drop_Cons: "drop n (x # xs) = (case n of 0 \ x # xs | Suc m \ 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 \ 'a list \ 'a list" where
take_Nil:"take n [] = []" |
take_Cons: "take n (x # xs) = (case n of 0 \ [] | Suc m \ 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 \ x | Suc k \ 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 \ nat \ 'a \ '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 \ bool) \ 'a list \ 'a list" where
"takeWhile P [] = []" |
"takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
primrec dropWhile :: "('a \ bool) \ 'a list \ 'a list" where
"dropWhile P [] = []" |
"dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)"
primrec zip :: "'a list \ 'b list \ ('a \ '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 \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where
"map2 f xs ys \ map (\(x,y). f x y) (zip xs ys)"
primrec product :: "'a list \ 'b list \ ('a \ '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 \ 'a list list" where
"product_lists [] = [[]]" |
"product_lists (xs # xss) = concat (map (\x. map (Cons x) (product_lists xss)) xs)"
primrec upt :: "nat \ nat \ nat list" ("(1[_..
upt_0: "[i..<0] = []" |
upt_Suc: "[i..<(Suc j)] = (if i \ j then [i..
definition insert :: "'a \ 'a list \ 'a list" where
"insert x xs = (if x \ set xs then xs else x # xs)"
definition union :: "'a list \ 'a list \ 'a list" where
"union = fold insert"
hide_const (open) insert union
hide_fact (open) insert_def union_def
primrec find :: "('a \ bool) \ 'a list \ '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 \ 'a \ 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 \ bool) \ 'a list \ ('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 \ '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 \ 'a list \ 'a list" where
"remove1 x [] = []" |
"remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)"
primrec removeAll :: "'a \ 'a list \ '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 \ bool" where
"distinct [] \ True" |
"distinct (x # xs) \ x \ set xs \ distinct xs"
fun successively :: "('a \ 'a \ bool) \ 'a list \ bool" where
"successively P [] = True" |
"successively P [x] = True" |
"successively P (x # y # xs) = (P x y \ successively P (y#xs))"
definition distinct_adj where
"distinct_adj = successively (\)"
primrec remdups :: "'a list \ 'a list" where
"remdups [] = []" |
"remdups (x # xs) = (if x \ set xs then remdups xs else x # remdups xs)"
fun remdups_adj :: "'a list \ '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 \ 'a \ '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 \ nat" where
"length \ size"
definition enumerate :: "nat \ 'a list \ (nat \ 'a) list" where
enumerate_eq_zip: "enumerate n xs = zip [n..
primrec rotate1 :: "'a list \ 'a list" where
"rotate1 [] = []" |
"rotate1 (x # xs) = xs @ [x]"
definition rotate :: "nat \ 'a list \ 'a list" where
"rotate n = rotate1 ^^ n"
definition nths :: "'a list => nat set => 'a list" where
"nths xs A = map fst (filter (\p. snd p \ A) (zip xs [0..
primrec subseqs :: "'a list \ 'a list list" where
"subseqs [] = [[]]" |
"subseqs (x#xs) = (let xss = subseqs xs in map (Cons x) xss @ xss)"
primrec n_lists :: "nat \ 'a list \ 'a list list" where
"n_lists 0 xs = [[]]" |
"n_lists (Suc n) xs = concat (map (\ys. map (\y. y # ys) xs) (n_lists n xs))"
hide_const (open) n_lists
function splice :: "'a list \ 'a list \ 'a list" where
"splice [] ys = ys" |
"splice (x#xs) ys = x # splice ys xs"
by pat_completeness auto
termination
by(relation "measure(\(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) \ (#) 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 \ 'a" where
"min_list (x # xs) = (case xs of [] \ x | _ \ min x (min_list xs))"
text\<open>Returns first minimum:\<close>
fun arg_min_list :: "('a \ ('b::linorder)) \ 'a list \ '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 \ 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 (\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 (\) [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 (\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 \ 'a \ bool) \ 'a list \ bool" where
"sorted_wrt P [] = True" |
"sorted_wrt P (x # ys) = ((\y \ set ys. P x y) \ sorted_wrt P ys)"
text \<open>A class-based sorted predicate:\<close>
context linorder
begin
fun sorted :: "'a list \ bool" where
"sorted [] = True" |
"sorted (x # ys) = ((\y \ set ys. x \ y) \ sorted ys)"
fun strict_sorted :: "'a list \ bool" where
"strict_sorted [] = True" |
"strict_sorted (x # ys) = ((\y \ List.set ys. x < y) \ strict_sorted ys)"
lemma sorted_sorted_wrt: "sorted = sorted_wrt (\)"
proof (rule ext)
fix xs show "sorted xs = sorted_wrt (\) xs"
by(induction xs rule: sorted.induct) auto
qed
lemma strict_sorted_sorted_wrt: "strict_sorted = sorted_wrt (<)"
proof (rule ext)
fix xs show "strict_sorted xs = sorted_wrt (<) xs"
by(induction xs rule: strict_sorted.induct) auto
qed
primrec insort_key :: "('b \ 'a) \ 'b \ 'b list \ '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 \ 'a) \ 'b list \ 'b list" where
"sort_key f xs = foldr (insort_key f) xs []"
definition insort_insert_key :: "('b \ 'a) \ 'b \ 'b list \ '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 \ sort_key (\x. x)"
abbreviation "insort \ insort_key (\x. x)"
abbreviation "insort_insert \ insort_insert_key (\x. x)"
definition stable_sort_key :: "(('b \ 'a) \ 'b list \ 'b list) \ 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: "strict_sorted l \ sorted l \ distinct l"
by (induction l) (auto iff: antisym_conv1)
lemma strict_sorted_imp_sorted: "strict_sorted xs \ 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): \[e. x \ xs, \]\ 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 \ lc_qual \ lc_quals \ 'a list" ("[_ . __")
"_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ \ _")
"_lc_test" :: "bool \ lc_qual" ("_")
(*"_lc_let" :: "letbinds => lc_qual" ("let _")*)
"_lc_end" :: "lc_quals" ("]")
"_lc_quals" :: "lc_qual \ lc_quals \ lc_quals" (", __")
syntax (ASCII)
"_lc_gen" :: "'a \ 'a list \ 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 \ set A}" by simp}
val del_refl_eq = @{lemma "(t = t \ P) \ 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 \ 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 \ 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 "(\x. x = t \ 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 \ 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 "(\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 \ x # xs"
by (induct xs) auto
lemma not_Cons_self2 [simp]: "x # xs \ xs"
by (rule not_Cons_self [symmetric])
lemma neq_Nil_conv: "(xs \ []) = (\y ys. xs = y # ys)"
by (induct xs) auto
lemma tl_Nil: "tl xs = [] \ xs = [] \ (\x. xs = [x])"
by (cases xs) auto
lemma Nil_tl: "[] = tl xs \ xs = [] \ (\x. xs = [x])"
by (cases xs) auto
lemma length_induct:
"(\xs. \ys. length ys < length xs \ P ys \ P xs) \ P xs"
by (fact measure_induct)
lemma induct_list012:
"\P []; \x. P [x]; \x y zs. \ P zs; P (y # zs) \ \ P (x # y # zs)\ \ P xs"
by induction_schema (pat_completeness, lexicographic_order)
lemma list_nonempty_induct [consumes 1, case_names single cons]:
"\ xs \ []; \x. P [x]; \x xs. xs \ [] \ P xs \ P (x # xs)\ \ P xs"
by(induction xs rule: induct_list012) auto
lemma inj_split_Cons: "inj_on (\(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 \ [])"
by (induct xs) auto
lemma length_pos_if_in_set: "x \ set xs \ length xs > 0"
by auto
lemma length_Suc_conv:
"(length xs = Suc n) = (\y ys. xs = y # ys \ length ys = n)"
by (induct xs) auto
lemma Suc_length_conv:
"(Suc n = length xs) = (\y ys. xs = y # ys \ length ys = n)"
by (induct xs; simp; blast)
lemma Suc_le_length_iff:
"(Suc n \ length xs) = (\x ys. xs = x # ys \ n \ length ys)"
by (metis Suc_le_D[of n] Suc_le_mono[of n] Suc_length_conv[of _ xs])
lemma impossible_Cons: "length xs \ length ys \ xs = x # ys = False"
by (induct xs) auto
lemma list_induct2 [consumes 1, case_names Nil Cons]:
"length xs = length ys \ P [] [] \
(\<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 \ length ys = length zs \ P [] [] [] \
(\<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 \ length ys = length zs \ length zs = length ws \
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':
"\ 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 \ length xs = length ys \ (\(x, y) \ set (zip xs ys). P x y)"
by (induct xs ys rule: list_induct2') auto
lemma neq_if_length_neq: "length xs \ length ys \ (xs = ys) == False"
by (rule Eq_FalseI) auto
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 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>\<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 = [] \ ys = [])"
by (induct xs) auto
lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \ ys = [])"
by (induct xs) auto
lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
by (induct xs) auto
lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
by (induct xs) auto
lemma append_eq_append_conv [simp]:
"length xs = length ys \ 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 \ 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
lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
using append_same_eq [of "[]"] by auto
lemma hd_Cons_tl: "xs \ [] \ 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 \ [] \ hd (xs @ ys) = hd xs"
by (simp add: hd_append split: list.split)
lemma tl_append: "tl (xs @ ys) = (case xs of [] \ tl ys | z#zs \ zs @ ys)"
by (simp split: list.split)
lemma tl_append2 [simp]: "xs \ [] \ tl (xs @ ys) = tl xs @ ys"
by (simp add: tl_append split: list.split)
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:
"\ps xs' ys'. xs = ps @ xs' \ 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, hide_lams) append_Cons append_Nil list.sel(1))
text \<open>Trivial rules for solving \<open>@\<close>-equations automatically.\<close>
lemma eq_Nil_appendI: "xs = ys \ xs = [] @ ys"
by simp
lemma Cons_eq_appendI: "\x # xs1 = ys; xs = xs1 @ zs\ \ x # xs = ys @ zs"
by auto
lemma append_eq_appendI: "\xs @ xs1 = zs; ys = xs1 @ us\ \ 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 fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end
\<close>
subsubsection \<open>\<^const>\<open>map\<close>\<close>
lemma hd_map: "xs \ [] \ 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: "(\x. x \ set xs \ f x = g x) \ map f xs = map g xs"
by (induct xs) simp_all
lemma map_ident [simp]: "map (\x. x) = (\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 \ g) xs"
by (induct xs) auto
lemma map_comp_map[simp]: "((map f) \ (map g)) = map(f \ 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) = (\x \ set xs. f x = g x)"
by (induct xs) auto
lemma map_cong [fundef_cong]:
"xs = ys \ (\x. x \ set ys \ f x = g x) \ map f xs = map g ys"
by simp
lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
by (cases xs) auto
lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
by (cases xs) auto
lemma map_eq_Cons_conv:
"(map f xs = y#ys) = (\z zs. xs = z#zs \ f z = y \ map f zs = ys)"
by (cases xs) auto
lemma Cons_eq_map_conv:
"(x#xs = map f ys) = (\z zs. ys = z#zs \ x = f z \ 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:
"(\xs. ys = map f xs) = (\y \ set ys. \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) \ (map f xs = map f ys) = (xs = ys)"
by(blast dest:map_inj_on)
lemma map_injective:
"map f xs = map f ys \ inj f \ xs = ys"
by (induct ys arbitrary: xs) (auto dest!:injD)
lemma inj_map_eq_map[simp]: "inj f \ (map f xs = map f ys) = (xs = ys)"
by(blast dest:map_injective)
lemma inj_mapI: "inj f \ inj (map f)"
by (iprover dest: map_injective injD intro: inj_onI)
lemma inj_mapD: "inj (map f) \ inj f"
by (metis (no_types, hide_lams) 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 (\(set ` A)) \ inj_on (map f) A"
by (blast intro:inj_onI dest:inj_onD map_inj_on)
lemma map_idI: "(\x. x \ set xs \ f x = x) \ map f xs = xs"
by (induct xs, auto)
lemma map_fun_upd [simp]: "y \ set xs \ map (f(y:=v)) xs = map f xs"
by (induct xs) auto
lemma map_fst_zip[simp]:
"length xs = length ys \ map fst (zip xs ys) = xs"
by (induct rule:list_induct2, simp_all)
lemma map_snd_zip[simp]:
"length xs = length ys \ 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 (\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
lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
by (induct xs) auto
lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])"
by (cases xs) auto
lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
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 "\x xs. P xs \ 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 = [] \ P) \(\ys y. xs = ys @ [y] \ P) \ 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 \ []"
and single: "\x. P [x]"
and snoc': "\x xs. xs \ [] \ P xs \ 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
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 \ set ys)"
by (induct xs) auto
lemma hd_in_set[simp]: "xs \ [] \ hd xs \ set xs"
by(cases xs) auto
lemma set_subset_Cons: "set xs \ set (x # xs)"
by auto
lemma set_ConsD: "y \ set (x # xs) \ y=x \ y \ set xs"
by auto
lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
by (induct xs) auto
lemma set_empty2[iff]: "({} = set xs) = (xs = [])"
by(induct xs) auto
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 \ set xs \ P x}"
by (induct xs) auto
lemma set_upt [simp]: "set[i..
by (induct j) auto
lemma split_list: "x \ set xs \ \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 \ set xs \ (\ys zs. xs = ys @ x # zs)"
by (auto elim: split_list)
lemma split_list_first: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ 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 \ a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI)
qed
qed
lemma in_set_conv_decomp_first:
"(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set ys)"
by (auto dest!: split_list_first)
lemma split_list_last: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ 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 \ a" thus ?case using snoc by fastforce
qed
qed
lemma in_set_conv_decomp_last:
"(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set zs)"
by (auto dest!: split_list_last)
lemma split_list_prop: "\x \ set xs. P x \ \ys x zs. xs = ys @ x # zs \ 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 "\x \ 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:
"\x \ set xs. P x \
\<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 \ P x \ (\y\set []. \ P y)" by simp
thus ?thesis by fast
next
assume "\ P x"
hence "\x\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 "\x \ set xs. P x"
obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\y \ set ys. \ P y"
using split_list_first_prop [OF assms] by blast
lemma split_list_first_prop_iff:
"(\x \ set xs. P x) \
(\<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:
"\x \ set xs. P x \
\<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 "\ P x"
hence "\x\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 "\x \ set xs. P x"
obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\z \ set zs. \ P z"
using split_list_last_prop [OF assms] by blast
lemma split_list_last_prop_iff:
"(\x \ set xs. P x) \
(\<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 \ \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) \ length xs"
by (induct xs) (auto simp add: card_insert_if)
lemma set_minus_filter_out:
"set xs - {y} = set (filter (\x. \ (x = y)) xs)"
by (induct xs) auto
lemma append_Cons_eq_iff:
"\ x \ set xs; x \ set ys \ \
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>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 (\x. Q x \ P x) xs"
by (induct xs) auto
lemma length_filter_le [simp]: "length (filter P xs) \ length xs"
by (induct xs) (auto simp add: le_SucI)
lemma sum_length_filter_compl:
"length(filter P xs) + length(filter (\x. \P x) xs) = length xs"
by(induct xs) simp_all
lemma filter_True [simp]: "\x \ set xs. P x \ filter P xs = xs"
by (induct xs) auto
lemma filter_False [simp]: "\x \ set xs. \ P x \ filter P xs = []"
by (induct xs) auto
lemma filter_empty_conv: "(filter P xs = []) = (\x\set xs. \ P x)"
by (induct xs) simp_all
lemma filter_id_conv: "(filter P xs = xs) = (\x\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 \ f) xs)"
by (induct xs) simp_all
lemma length_filter_map[simp]:
"length (filter P (map f xs)) = length(filter (P \ f) xs)"
by (simp add:filter_map)
lemma filter_is_subset [simp]: "set (filter P xs) \ set xs"
by auto
lemma length_filter_less:
"\ x \ set xs; \ P x \ \ 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 \ p(xs!i)}"
proof (induct xs)
case Nil thus ?case by simp
next
case (Cons x xs)
let ?S = "{i. i < length xs \ 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 "\ = Suc(card(Suc ` ?S))" using fin
by (simp add: card_image)
also have "\ = card ?S'" using eq fin
by (simp add:card_insert_if)
finally show ?thesis .
next
assume "\ 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 "\ = card(Suc ` ?S)" using fin
by (simp add: card_image)
also have "\ = 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 \
\<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 "_ \ \us vs. ?P ys us vs")
proof(induct ys)
case Nil thus ?case by simp
next
case (Cons y ys)
show ?case (is "\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 \ y"
with Py Cons.prems show ?thesis by simp
qed
next
assume "\ 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 \
\<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)
lemma Cons_eq_filter_iff:
"(x#xs = filter P ys) =
(\<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:Cons_eq_filterD)
lemma inj_on_filter_key_eq:
assumes "inj_on f (insert y (set xs))"
shows "filter (\x. f y = f x) xs = filter (HOL.eq y) xs"
using assms by (induct xs) auto
lemma filter_cong[fundef_cong]:
"xs = ys \ (\x. x \ set ys \ P x = Q x) \ filter P xs = filter Q ys"
by (induct ys arbitrary: xs) auto
subsubsection \<open>List partitioning\<close>
primrec partition :: "('a \ bool) \'a list \ 'a list \ '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 \ P) xs"
by (induct xs) (auto simp add: Let_def split_def)
lemma partition_P:
assumes "partition P xs = (yes, no)"
shows "(\p \ set yes. P p) \ (\p \ set no. \ 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 \ 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 \ f) xs)"
unfolding partition_filter2[symmetric]
unfolding partition_filter1[symmetric] by simp
declare partition.simps[simp del]
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 = []) = (\xs \ set xss. xs = [])"
by (induct xss) auto
lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\xs \ set xss. xs = [])"
by (induct xss) auto
lemma set_concat [simp]: "set (concat xs) = (\x\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 filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
by (induct xs) auto
lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
by (induct xs) auto
lemma concat_eq_concat_iff: "\(x, y) \ set (zip xs ys). length x = length y \ length xs = length ys \ (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 \ length xs = length ys \ \(x, y) \ set (zip xs ys). length x = length y \ xs = ys"
by (simp add: concat_eq_concat_iff)
lemma concat_eq_appendD:
assumes "concat xss = ys @ zs" "xss \ []"
--> --------------------
--> maximum size reached
--> --------------------
¤ Dauer der Verarbeitung: 0.32 Sekunden
(vorverarbeitet)
¤
|
Haftungshinweis
Die Informationen auf dieser Webseite wurden
nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit,
noch Qualität der bereit gestellten Informationen zugesichert.
Bemerkung:
Die farbliche Syntaxdarstellung ist noch experimentell.
|