Theory Misc
section ‹Miscellaneous Definitions and Lemmas›
theory Misc
imports Main
"HOL-Library.Multiset"
"HOL-ex.Quicksort"
"HOL-Library.Option_ord"
"HOL-Library.Infinite_Set"
"HOL-Eisbach.Eisbach"
begin
text_raw ‹\label{thy:Misc}›
subsection ‹Isabelle Distribution Move Proposals›
subsubsection ‹Pure›
lemma TERMI: "TERM x" unfolding Pure.term_def .
subsubsection ‹HOL›
lemma disjE1: "⟦ P ∨ Q; P ⟹ R; ⟦¬P;Q⟧ ⟹ R ⟧ ⟹ R"
by metis
lemma disjE2: "⟦ P ∨ Q; ⟦P; ¬Q⟧ ⟹ R; Q ⟹ R ⟧ ⟹ R"
by blast
lemma imp_mp_iff[simp]:
"a ∧ (a ⟶ b) ⟷ a ∧ b"
"(a ⟶ b) ∧ a ⟷ a ∧ b"
by blast+
lemma atomize_Trueprop_eq[atomize]: "(Trueprop x ≡ Trueprop y) ≡ Trueprop (x=y)"
apply rule
apply (rule)
apply (erule equal_elim_rule1)
apply assumption
apply (erule equal_elim_rule2)
apply assumption
apply simp
done
subsubsection ‹Set›
lemma inter_compl_diff_conv[simp]: "A ∩ -B = A - B" by auto
lemma pair_set_inverse[simp]: "{(a,b). P a b}¯ = {(b,a). P a b}"
by auto
lemma card_doubleton_eq_2_iff[simp]: "card {a,b} = 2 ⟷ a≠b" by auto
subsubsection ‹List›
thm List.length_greater_0_conv
lemma length_ge_1_conv[iff]: "Suc 0 ≤ length l ⟷ l≠[]"
by (cases l) auto
lemma obtain_list_from_elements:
assumes A: "∀i<n. (∃li. P li i)"
obtains l where
"length l = n"
"∀i<n. P (l!i) i"
proof -
from A have "∃l. length l=n ∧ (∀i<n. P (l!i) i)"
proof (induct n)
case 0 thus ?case by simp
next
case (Suc n)
then obtain l where IH: "length l = n" "(∀i<n. P(l!i) i)" by auto
moreover from Suc.prems obtain ln where "P ln n" by auto
ultimately have "length (l@[ln]) = Suc n" "(∀i<Suc n. P((l@[ln])!i) i)"
by (auto simp add: nth_append dest: less_antisym)
thus ?case by blast
qed
thus ?thesis using that by (blast)
qed
lemma distinct_sorted_mono:
assumes S: "sorted l"
assumes D: "distinct l"
assumes B: "i<j" "j<length l"
shows "l!i < l!j"
proof -
from S B have "l!i ≤ l!j"
by (simp add: sorted_iff_nth_mono)
also from nth_eq_iff_index_eq[OF D] B have "l!i ≠ l!j"
by auto
finally show ?thesis .
qed
lemma distinct_sorted_strict_mono_iff:
assumes "distinct l" "sorted l"
assumes "i<length l" "j<length l"
shows "l!i<l!j ⟷ i<j"
using assms
by (metis distinct_sorted_mono leI less_le_not_le
order.strict_iff_order)
lemma distinct_sorted_mono_iff:
assumes "distinct l" "sorted l"
assumes "i<length l" "j<length l"
shows "l!i≤l!j ⟷ i≤j"
by (metis assms distinct_sorted_strict_mono_iff leD le_less_linear)
lemma map_eq_appendE:
assumes "map f ls = fl@fl'"
obtains l l' where "ls=l@l'" and "map f l=fl" and "map f l' = fl'"
using assms
proof (induction fl arbitrary: ls thesis)
case (Cons x xs)
then obtain l ls' where [simp]: "ls = l#ls'" "f l = x" by force
with Cons.prems(2) have "map f ls' = xs @ fl'" by simp
from Cons.IH[OF _ this] obtain ll ll' where "ls' = ll @ ll'" "map f ll = xs" "map f ll' = fl'" .
with Cons.prems(1)[of "l#ll" ll'] show thesis by simp
qed simp
lemma map_eq_append_conv: "map f ls = fl@fl' ⟷ (∃l l'. ls = l@l' ∧ map f l = fl ∧ map f l' = fl')"
by (auto elim!: map_eq_appendE)
lemmas append_eq_mapE = map_eq_appendE[OF sym]
lemma append_eq_map_conv: "fl@fl' = map f ls ⟷ (∃l l'. ls = l@l' ∧ map f l = fl ∧ map f l' = fl')"
by (auto elim!: append_eq_mapE)
lemma distinct_mapI: "distinct (map f l) ⟹ distinct l"
by (induct l) auto
lemma map_distinct_upd_conv:
"⟦i<length l; distinct l⟧ ⟹ (map f l)[i := x] = map (f(l!i := x)) l"
by (auto simp: nth_eq_iff_index_eq intro: nth_equalityI)
lemma zip_inj: "⟦length a = length b; length a' = length b'; zip a b = zip a' b'⟧ ⟹ a=a' ∧ b=b'"
proof (induct a b arbitrary: a' b' rule: list_induct2)
case Nil
then show ?case by (cases a'; cases b'; auto)
next
case (Cons x xs y ys)
then show ?case by (cases a'; cases b'; auto)
qed
lemma zip_eq_zip_same_len[simp]:
"⟦ length a = length b; length a' = length b' ⟧ ⟹
zip a b = zip a' b' ⟷ a=a' ∧ b=b'"
by (auto dest: zip_inj)
lemma upt_merge[simp]: "i≤j ∧ j≤k ⟹ [i..<j]@[j..<k] = [i..<k]"
by (metis le_Suc_ex upt_add_eq_append)
lemma snoc_eq_iff_butlast':
"(ys = xs @ [x]) ⟷ (ys ≠ [] ∧ butlast ys = xs ∧ last ys = x)"
by auto
lemma list_match_lel_lel:
assumes "c1 @ qs # c2 = c1' @ qs' # c2'"
obtains
(left) c21' where "c1 = c1' @ qs' # c21'" "c2' = c21' @ qs # c2"
| (center) "c1' = c1" "qs' = qs" "c2' = c2"
| (right) c21 where "c1' = c1 @ qs # c21" "c2 = c21 @ qs' # c2'"
using assms
apply (clarsimp simp: append_eq_append_conv2)
subgoal for us by (cases us) auto
done
lemma xy_in_set_cases[consumes 2, case_names EQ XY YX]:
assumes A: "x∈set l" "y∈set l"
and C:
"!!l1 l2. ⟦ x=y; l=l1@y#l2 ⟧ ⟹ P"
"!!l1 l2 l3. ⟦ x≠y; l=l1@x#l2@y#l3 ⟧ ⟹ P"
"!!l1 l2 l3. ⟦ x≠y; l=l1@y#l2@x#l3 ⟧ ⟹ P"
shows P
proof (cases "x=y")
case True with A(1) obtain l1 l2 where "l=l1@y#l2" by (blast dest: split_list)
with C(1) True show ?thesis by blast
next
case False
from A(1) obtain l1 l2 where S1: "l=l1@x#l2" by (blast dest: split_list)
from A(2) obtain l1' l2' where S2: "l=l1'@y#l2'" by (blast dest: split_list)
from S1 S2 have M: "l1@x#l2 = l1'@y#l2'" by simp
thus P proof (cases rule: list_match_lel_lel[consumes 1, case_names 1 2 3])
case (1 c) with S1 have "l=l1'@y#c@x#l2" by simp
with C(3) False show ?thesis by blast
next
case 2 with False have False by blast
thus ?thesis ..
next
case (3 c) with S1 have "l=l1@x#c@y#l2'" by simp
with C(2) False show ?thesis by blast
qed
qed
lemma list_e_eq_lel[simp]:
"[e] = l1@e'#l2 ⟷ l1=[] ∧ e'=e ∧ l2=[]"
"l1@e'#l2 = [e] ⟷ l1=[] ∧ e'=e ∧ l2=[]"
apply (cases l1, auto) []
apply (cases l1, auto) []
done
lemma list_ee_eq_leel[simp]:
"([e1,e2] = l1@e1'#e2'#l2) ⟷ (l1=[] ∧ e1=e1' ∧ e2=e2' ∧ l2=[])"
"(l1@e1'#e2'#l2 = [e1,e2]) ⟷ (l1=[] ∧ e1=e1' ∧ e2=e2' ∧ l2=[])"
apply (cases l1, auto) []
apply (cases l1, auto) []
done
subsubsection ‹Transitive Closure›
text ‹A point-free induction rule for elements reachable from an initial set›
lemma rtrancl_reachable_induct[consumes 0, case_names base step]:
assumes I0: "I ⊆ INV"
assumes IS: "E``INV ⊆ INV"
shows "E⇧*``I ⊆ INV"
by (metis I0 IS Image_closed_trancl Image_mono subset_refl)
lemma acyclic_empty[simp, intro!]: "acyclic {}" by (unfold acyclic_def) auto
lemma acyclic_union:
"acyclic (A∪B) ⟹ acyclic A"
"acyclic (A∪B) ⟹ acyclic B"
by (metis Un_upper1 Un_upper2 acyclic_subset)+
text ‹Here we provide a collection of miscellaneous definitions and helper lemmas›
subsection "Miscellaneous (1)"
text ‹This stuff is used in this theory itself, and thus occurs in first place or is simply not sorted into any other section of this theory.›
lemma IdD: "(a,b)∈Id ⟹ a=b" by simp
text ‹Conversion Tag›
definition [simp]: "CNV x y ≡ x=y"
lemma CNV_I: "CNV x x" by simp
lemma CNV_eqD: "CNV x y ⟹ x=y" by simp
lemma CNV_meqD: "CNV x y ⟹ x≡y" by simp
lemma ex_b_b_and_simp[simp]: "(∃b. b ∧ Q b) ⟷ Q True" by auto
lemma ex_b_not_b_and_simp[simp]: "(∃b. ¬b ∧ Q b) ⟷ Q False" by auto
method repeat_all_new methods m = m;(repeat_all_new ‹m›)?
subsubsection "AC-operators"
text ‹Locale to declare AC-laws as simplification rules›
locale Assoc =
fixes f
assumes assoc[simp]: "f (f x y) z = f x (f y z)"
locale AC = Assoc +
assumes commute[simp]: "f x y = f y x"
lemma (in AC) left_commute[simp]: "f x (f y z) = f y (f x z)"
by (simp only: assoc[symmetric]) simp
lemmas (in AC) AC_simps = commute assoc left_commute
text ‹Locale to define functions from surjective, unique relations›
locale su_rel_fun =
fixes F and f
assumes unique: "⟦(A,B)∈F; (A,B')∈F⟧ ⟹ B=B'"
assumes surjective: "⟦!!B. (A,B)∈F ⟹ P⟧ ⟹ P"
assumes f_def: "f A == THE B. (A,B)∈F"
lemma (in su_rel_fun) repr1: "(A,f A)∈F" proof (unfold f_def)
obtain B where "(A,B)∈F" by (rule surjective)
with theI[where P="λB. (A,B)∈F", OF this] show "(A, THE x. (A, x) ∈ F) ∈ F" by (blast intro: unique)
qed
lemma (in su_rel_fun) repr2: "(A,B)∈F ⟹ B=f A" using repr1
by (blast intro: unique)
lemma (in su_rel_fun) repr: "(f A = B) = ((A,B)∈F)" using repr1 repr2
by (blast)
lemma Ex_prod_contract: "(∃a b. P a b) ⟷ (∃z. P (fst z) (snd z))"
by auto
lemma All_prod_contract: "(∀a b. P a b) ⟷ (∀z. P (fst z) (snd z))"
by auto
lemma nat_geq_1_eq_neqz: "x≥1 ⟷ x≠(0::nat)"
by auto
lemma nat_in_between_eq:
"(a<b ∧ b≤Suc a) ⟷ b = Suc a"
"(a≤b ∧ b<Suc a) ⟷ b = a"
by auto
lemma Suc_n_minus_m_eq: "⟦ n≥m; m>1 ⟧ ⟹ Suc (n - m) = n - (m - 1)"
by simp
lemma Suc_to_right: "Suc n = m ⟹ n = m - Suc 0" by simp
lemma Suc_diff[simp]: "⋀n m. n≥m ⟹ m≥1 ⟹ Suc (n - m) = n - (m - 1)"
by simp
lemma if_not_swap[simp]: "(if ¬c then a else b) = (if c then b else a)" by auto
lemma all_to_meta: "Trueprop (∀a. P a) ≡ (⋀a. P a)"
apply rule
by auto
lemma imp_to_meta: "Trueprop (P⟶Q) ≡ (P⟹Q)"
apply rule
by auto
lemma iffI2: "⟦P ⟹ Q; ¬ P ⟹ ¬ Q⟧ ⟹ P ⟷ Q"
by metis
lemma iffExI:
"⟦ ⋀x. P x ⟹ Q x; ⋀x. Q x ⟹ P x ⟧ ⟹ (∃x. P x) ⟷ (∃x. Q x)"
by metis
lemma bex2I[intro?]: "⟦ (a,b)∈S; (a,b)∈S ⟹ P a b ⟧ ⟹ ∃a b. (a,b)∈S ∧ P a b"
by blast
lemma cnv_conj_to_meta: "(P ∧ Q ⟹ PROP X) ≡ (⟦P;Q⟧ ⟹ PROP X)"
by (rule BNF_Fixpoint_Base.conj_imp_eq_imp_imp)
subsection ‹Sets›
lemma remove_subset: "x∈S ⟹ S-{x} ⊂ S" by auto
lemma subset_minus_empty: "A⊆B ⟹ A-B = {}" by auto
lemma insert_minus_eq: "x≠y ⟹ insert x A - {y} = insert x (A - {y})" by auto
lemma set_notEmptyE: "⟦S≠{}; !!x. x∈S ⟹ P⟧ ⟹ P"
by (metis equals0I)
lemma subset_Collect_conv: "S ⊆ Collect P ⟷ (∀x∈S. P x)"
by auto
lemma memb_imp_not_empty: "x∈S ⟹ S≠{}"
by auto
lemma disjoint_mono: "⟦ a⊆a'; b⊆b'; a'∩b'={} ⟧ ⟹ a∩b={}" by auto
lemma disjoint_alt_simp1: "A-B = A ⟷ A∩B = {}" by auto
lemma disjoint_alt_simp2: "A-B ≠ A ⟷ A∩B ≠ {}" by auto
lemma disjoint_alt_simp3: "A-B ⊂ A ⟷ A∩B ≠ {}" by auto
lemma disjointI[intro?]: "⟦ ⋀x. ⟦x∈a; x∈b⟧ ⟹ False ⟧ ⟹ a∩b={}"
by auto
lemmas set_simps = subset_minus_empty disjoint_alt_simp1 disjoint_alt_simp2 disjoint_alt_simp3 Un_absorb1 Un_absorb2
lemma set_minus_singleton_eq: "x∉X ⟹ X-{x} = X"
by auto
lemma set_diff_diff_left: "A-B-C = A-(B∪C)"
by auto
lemma image_update[simp]: "x∉A ⟹ f(x:=n)`A = f`A"
by auto
lemma eq_or_mem_image_simp[simp]: "{f l |l. l = a ∨ l∈B} = insert (f a) {f l|l. l∈B}" by blast
lemma set_union_code [code_unfold]:
"set xs ∪ set ys = set (xs @ ys)"
by auto
lemma in_fst_imageE:
assumes "x ∈ fst`S"
obtains y where "(x,y)∈S"
using assms by auto
lemma in_snd_imageE:
assumes "y ∈ snd`S"
obtains x where "(x,y)∈S"
using assms by auto
lemma fst_image_mp: "⟦fst`A ⊆ B; (x,y)∈A ⟧ ⟹ x∈B"
by (metis Domain.DomainI fst_eq_Domain in_mono)
lemma snd_image_mp: "⟦snd`A ⊆ B; (x,y)∈A ⟧ ⟹ y∈B"
by (metis Range.intros rev_subsetD snd_eq_Range)
lemma inter_eq_subsetI: "⟦ S⊆S'; A∩S' = B∩S' ⟧ ⟹ A∩S = B∩S"
by auto
text ‹
Decompose general union over sum types.
›
lemma Union_plus:
"(⋃ x ∈ A <+> B. f x) = (⋃ a ∈ A. f (Inl a)) ∪ (⋃b ∈ B. f (Inr b))"
by auto
lemma Union_sum:
"(⋃x. f (x::'a+'b)) = (⋃l. f (Inl l)) ∪ (⋃r. f (Inr r))"
(is "?lhs = ?rhs")
proof -
have "?lhs = (⋃x ∈ UNIV <+> UNIV. f x)"
by simp
thus ?thesis
by (simp only: Union_plus)
qed
subsubsection ‹Finite Sets›
lemma card_1_singletonI: "⟦finite S; card S = 1; x∈S⟧ ⟹ S={x}"
proof (safe, rule ccontr, goal_cases)
case prems: (1 x')
hence "finite (S-{x})" "S-{x} ≠ {}" by auto
hence "card (S-{x}) ≠ 0" by auto
moreover from prems(1-3) have "card (S-{x}) = 0" by auto
ultimately have False by simp
thus ?case ..
qed
lemma card_insert_disjoint': "⟦finite A; x ∉ A⟧ ⟹ card (insert x A) - Suc 0 = card A"
by (drule (1) card_insert_disjoint) auto
lemma card_eq_UNIV[simp]: "card (S::'a::finite set) = card (UNIV::'a set) ⟷ S=UNIV"
proof (auto)
fix x
assume A: "card S = card (UNIV::'a set)"
show "x∈S" proof (rule ccontr)
assume "x∉S" hence "S⊂UNIV" by auto
with psubset_card_mono[of UNIV S] have "card S < card (UNIV::'a set)" by auto
with A show False by simp
qed
qed
lemma card_eq_UNIV2[simp]: "card (UNIV::'a set) = card (S::'a::finite set) ⟷ S=UNIV"
using card_eq_UNIV[of S] by metis
lemma card_ge_UNIV[simp]: "card (UNIV::'a::finite set) ≤ card (S::'a set) ⟷ S=UNIV"
using card_mono[of "UNIV::'a::finite set" S, simplified]
by auto
lemmas length_remdups_card = length_remdups_concat[of "[l]", simplified] for l
lemma fs_contract: "fst ` { p | p. f (fst p) (snd p) ∈ S } = { a . ∃b. f a b ∈ S }"
by (simp add: image_Collect)
lemma finite_Collect: "finite S ⟹ inj f ⟹ finite {a. f a : S}"
by(simp add: finite_vimageI vimage_def[symmetric])
lemma finite_imp_inj_to_nat_seg':
fixes A :: "'a set"
assumes A: "finite A"
obtains f::"'a ⇒ nat" and n::"nat" where
"f`A = {i. i<n}"
"inj_on f A"
by (metis A finite_imp_inj_to_nat_seg)
lemma lists_of_len_fin1: "finite P ⟹ finite (lists P ∩ { l. length l = n })"
proof (induct n)
case 0 thus ?case by auto
next
case (Suc n)
have "lists P ∩ { l. length l = Suc n }
= (λ(a,l). a#l) ` (P × (lists P ∩ {l. length l = n}))"
apply auto
apply (case_tac x)
apply auto
done
moreover from Suc have "finite …" by auto
ultimately show ?case by simp
qed
lemma lists_of_len_fin2: "finite P ⟹ finite (lists P ∩ { l. n = length l })"
proof -
assume A: "finite P"
have S: "{ l. n = length l } = { l. length l = n }" by auto
have "finite (lists P ∩ { l. n = length l })
⟷ finite (lists P ∩ { l. length l = n })"
by (subst S) simp
thus ?thesis using lists_of_len_fin1[OF A] by auto
qed
lemmas lists_of_len_fin = lists_of_len_fin1 lists_of_len_fin2
lemmas cset_fin_simps = Ex_prod_contract fs_contract[symmetric] image_Collect[symmetric]
lemmas cset_fin_intros = finite_imageI finite_Collect inj_onI
lemma Un_interval:
fixes b1 :: "'a::linorder"
assumes "b1≤b2" and "b2≤b3"
shows "{ f i | i. b1≤i ∧ i<b2 } ∪ { f i | i. b2≤i ∧ i<b3 }
= {f i | i. b1≤i ∧ i<b3}"
using assms
apply -
apply rule
apply safe []
apply (rule_tac x=i in exI, auto) []
apply (rule_tac x=i in exI, auto) []
apply rule
apply simp
apply (elim exE, simp)
apply (case_tac "i<b2")
apply (rule disjI1)
apply (rule_tac x=i in exI, auto) []
apply (rule disjI2)
apply (rule_tac x=i in exI, auto) []
done
text ‹
The standard library proves that a generalized union is finite
if the index set is finite and if for every index the component
set is itself finite. Conversely, we show that every component
set must be finite when the union is finite.
›
lemma finite_UNION_then_finite:
"finite (⋃(B ` A)) ⟹ a ∈ A ⟹ finite (B a)"
by (metis Set.set_insert UN_insert Un_infinite)
lemma finite_if_eq_beyond_finite: "finite S ⟹ finite {s. s - S = s' - S}"
proof (rule finite_subset[where B="(λs. s ∪ (s' - S)) ` Pow S"], clarsimp)
fix s
have "s = (s ∩ S) ∪ (s - S)"
by auto
also assume "s - S = s' - S"
finally show "s ∈ (λs. s ∪ (s' - S)) ` Pow S" by blast
qed blast
lemma distinct_finite_subset:
assumes "finite x"
shows "finite {ys. set ys ⊆ x ∧ distinct ys}" (is "finite ?S")
proof (rule finite_subset)
from assms show "?S ⊆ {ys. set ys ⊆ x ∧ length ys ≤ card x}"
by clarsimp (metis distinct_card card_mono)
from assms show "finite ..." by (rule finite_lists_length_le)
qed
lemma distinct_finite_set:
shows "finite {ys. set ys = x ∧ distinct ys}" (is "finite ?S")
proof (cases "finite x")
case False hence "{ys. set ys = x} = {}" by auto
thus ?thesis by simp
next
case True show ?thesis
proof (rule finite_subset)
show "?S ⊆ {ys. set ys ⊆ x ∧ length ys ≤ card x}"
using distinct_card by force
from True show "finite ..." by (rule finite_lists_length_le)
qed
qed
lemma finite_set_image:
assumes f: "finite (set ` A)"
and dist: "⋀xs. xs ∈ A ⟹ distinct xs"
shows "finite A"
proof (rule finite_subset)
from f show "finite (set -` (set ` A) ∩ {xs. distinct xs})"
proof (induct rule: finite_induct)
case (insert x F)
have "finite (set -` {x} ∩ {xs. distinct xs})"
apply (simp add: vimage_def)
by (metis Collect_conj_eq distinct_finite_set)
with insert show ?case
apply (subst vimage_insert)
apply (subst Int_Un_distrib2)
apply (rule finite_UnI)
apply simp_all
done
qed simp
moreover from dist show "A ⊆ ..."
by (auto simp add: vimage_image_eq)
qed
subsubsection ‹Infinite Set›
lemma INFM_nat_inductI:
assumes P0: "P (0::nat)"
assumes PS: "⋀i. P i ⟹ ∃j>i. P j ∧ Q j"
shows "∃⇩∞i. Q i"
proof -
have "∀i. ∃j>i. P j ∧ Q j" proof
fix i
show "∃j>i. P j ∧ Q j"
apply (induction i)
using PS[OF P0] apply auto []
by (metis PS Suc_lessI)
qed
thus ?thesis unfolding INFM_nat by blast
qed
subsection ‹Functions›
lemma fun_neq_ext_iff: "m≠m' ⟷ (∃x. m x ≠ m' x)" by auto
definition "inv_on f A x == SOME y. y∈A ∧ f y = x"
lemma inv_on_f_f[simp]: "⟦inj_on f A; x∈A⟧ ⟹ inv_on f A (f x) = x"
by (auto simp add: inv_on_def inj_on_def)
lemma f_inv_on_f: "⟦ y∈f`A ⟧ ⟹ f (inv_on f A y) = y"
by (auto simp add: inv_on_def intro: someI2)
lemma inv_on_f_range: "⟦ y ∈ f`A ⟧ ⟹ inv_on f A y ∈ A"
by (auto simp add: inv_on_def intro: someI2)
lemma inj_on_map_inv_f [simp]: "⟦set l ⊆ A; inj_on f A⟧ ⟹ map (inv_on f A) (map f l) = l"
apply (simp)
apply (induct l)
apply auto
done
lemma comp_cong_right: "x = y ⟹ f o x = f o y" by (simp)
lemma comp_cong_left: "x = y ⟹ x o f = y o f" by (simp)
lemma fun_comp_eq_conv: "f o g = fg ⟷ (∀x. f (g x) = fg x)"
by auto
abbreviation comp2 (infixl "oo" 55) where "f oo g ≡ λx. f o (g x)"
abbreviation comp3 (infixl "ooo" 55) where "f ooo g ≡ λx. f oo (g x)"
notation
comp2 (infixl "∘∘" 55) and
comp3 (infixl "∘∘∘" 55)
definition [code_unfold, simp]: "swap_args2 f x y ≡ f y x"
subsection ‹Multisets›
lemma count_mset_set_finite_iff:
"finite S ⟹ count (mset_set S) a = (if a ∈ S then 1 else 0)"
by simp
lemma ex_Melem_conv: "(∃x. x ∈# A) = (A ≠ {#})"
by (simp add: ex_in_conv)
subsubsection ‹Count›
lemma count_ne_remove: "⟦ x ~= t⟧ ⟹ count S x = count (S-{#t#}) x"
by (auto)
lemma mset_empty_count[simp]: "(∀p. count M p = 0) = (M={#})"
by (auto simp add: multiset_eq_iff)
subsubsection ‹Union, difference and intersection›
lemma size_diff_se: "t ∈# S ⟹ size S = size (S - {#t#}) + 1" proof (unfold size_multiset_overloaded_eq)
let ?SIZE = "sum (count S) (set_mset S)"
assume A: "t ∈# S"
from A have SPLITPRE: "finite (set_mset S) & {t}⊆(set_mset S)" by auto
hence "?SIZE = sum (count S) (set_mset S - {t}) + sum (count S) {t}" by (blast dest: sum.subset_diff)
hence "?SIZE = sum (count S) (set_mset S - {t}) + count (S) t" by auto
moreover with A have "count S t = count (S-{#t#}) t + 1" by auto
ultimately have D: "?SIZE = sum (count S) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (arith)
moreover have "sum (count S) (set_mset S - {t}) = sum (count (S-{#t#})) (set_mset S - {t})" proof -
have "∀x∈(set_mset S - {t}). count S x = count (S-{#t#}) x" by (auto iff add: count_ne_remove)
thus ?thesis by simp
qed
ultimately have D: "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (simp)
moreover
{ assume CASE: "t ∉# S - {#t#}"
from CASE have "set_mset S - {t} = set_mset (S - {#t#})"
by (auto simp add: in_diff_count split: if_splits)
with CASE D have "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1"
by (simp add: not_in_iff)
}
moreover
{ assume CASE: "t ∈# S - {#t#}"
from CASE have "t ∈# S" by (rule in_diffD)
with CASE have 1: "set_mset S = set_mset (S-{#t#})"
by (auto simp add: in_diff_count split: if_splits)
moreover from D have "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t} + 1" by simp
moreover from SPLITPRE sum.subset_diff have "sum (count (S-{#t#})) (set_mset S) = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t}" by (blast)
ultimately have "?SIZE = sum (count (S-{#t#})) (set_mset (S-{#t#})) + 1" by simp
}
ultimately show "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by blast
qed
lemma mset_union_diff_comm: "t ∈# S ⟹ T + (S - {#t#}) = (T + S) - {#t#}" proof -
assume "t ∈# S"
then obtain S' where S: "S = add_mset t S'"
by (metis insert_DiffM)
then show ?thesis
by auto
qed
lemma mset_right_cancel_union: "⟦a ∈# A+B; ~(a ∈# B)⟧ ⟹ a∈#A"
by (simp)
lemma mset_left_cancel_union: "⟦a ∈# A+B; ~(a ∈# A)⟧ ⟹ a∈#B"
by (simp)
lemmas mset_cancel_union = mset_right_cancel_union mset_left_cancel_union
lemma mset_right_cancel_elem: "⟦a ∈# A+{#b#}; a~=b⟧ ⟹ a∈#A"
by simp
lemma mset_left_cancel_elem: "⟦a ∈# {#b#}+A; a~=b⟧ ⟹ a∈#A"
by simp
lemmas mset_cancel_elem = mset_right_cancel_elem mset_left_cancel_elem
lemma mset_diff_cancel1elem[simp]: "~(a ∈# B) ⟹ {#a#}-B = {#a#}"
by (auto simp add: not_in_iff intro!: multiset_eqI)
lemma union_diff_assoc: "C-B={#} ⟹ (A+B)-C = A + (B-C)"
by (simp add: multiset_eq_iff)
lemmas mset_neutral_cancel1 = union_left_cancel[where N="{#}", simplified] union_right_cancel[where N="{#}", simplified]
declare mset_neutral_cancel1[simp]
lemma mset_union_2_elem: "{#a, b#} = add_mset c M ⟹ {#a#}=M & b=c | a=c & {#b#}=M"
by (auto simp: add_eq_conv_diff)
lemma mset_un_cases[cases set, case_names left right]:
"⟦a ∈# A + B; a ∈# A ⟹ P; a ∈# B ⟹ P⟧ ⟹ P"
by (auto)
lemma mset_unplusm_dist_cases[cases set, case_names left right]:
assumes A: "{#s#}+A = B+C"
assumes L: "⟦B={#s#}+(B-{#s#}); A=(B-{#s#})+C⟧ ⟹ P"
assumes R: "⟦C={#s#}+(C-{#s#}); A=B+(C-{#s#})⟧ ⟹ P"
shows P
proof -
from A[symmetric] have "s ∈# B+C" by simp
thus ?thesis proof (cases rule: mset_un_cases)
case left hence 1: "B={#s#}+(B-{#s#})" by simp
with A have "{#s#}+A = {#s#}+((B-{#s#})+C)" by (simp add: union_ac)
hence 2: "A = (B-{#s#})+C" by (simp)
from L[OF 1 2] show ?thesis .
next
case right hence 1: "C={#s#}+(C-{#s#})" by simp
with A have "{#s#}+A = {#s#}+(B+(C-{#s#}))" by (simp add: union_ac)
hence 2: "A = B+(C-{#s#})" by (simp)
from R[OF 1 2] show ?thesis .
qed
qed
lemma mset_unplusm_dist_cases2[cases set, case_names left right]:
assumes A: "B+C = {#s#}+A"
assumes L: "⟦B={#s#}+(B-{#s#}); A=(B-{#s#})+C⟧ ⟹ P"
assumes R: "⟦C={#s#}+(C-{#s#}); A=B+(C-{#s#})⟧ ⟹ P"
shows P
using mset_unplusm_dist_cases[OF A[symmetric]] L R by blast
lemma mset_single_cases[cases set, case_names loc env]:
assumes A: "add_mset s c = add_mset r' c'"
assumes CASES: "⟦s=r'; c=c'⟧ ⟹ P" "⟦c'={#s#}+(c'-{#s#}); c={#r'#}+(c-{#r'#}); c-{#r'#} = c'-{#s#} ⟧ ⟹ P"
shows "P"
proof -
{ assume CASE: "s=r'"
with A have "c=c'" by simp
with CASE CASES have ?thesis by auto
} moreover {
assume CASE: "s≠r'"
have "s ∈# {#s#}+c" by simp
with A have "s ∈# {#r'#}+c'" by simp
with CASE have "s ∈# c'" by simp
hence 1: "c' = {#s#} + (c' - {#s#})" by simp
with A have "{#s#}+c = {#s#}+({#r'#}+(c' - {#s#}))" by (auto simp add: union_ac)
hence 2: "c={#r'#}+(c' - {#s#})" by (auto)
hence 3: "c-{#r'#} = (c' - {#s#})" by auto
from 1 2 3 CASES have ?thesis by auto
} ultimately show ?thesis by blast
qed
lemma mset_single_cases'[cases set, case_names loc env]:
assumes A: "add_mset s c = add_mset r' c'"
assumes CASES: "⟦s=r'; c=c'⟧ ⟹ P" "!!cc. ⟦c'={#s#}+cc; c={#r'#}+cc; c'-{#s#}=cc; c-{#r'#}=cc⟧ ⟹ P"
shows "P"
using A CASES by (auto elim!: mset_single_cases)
lemma mset_single_cases2[cases set, case_names loc env]:
assumes A: "add_mset s c = add_mset r' c'"
assumes CASES: "⟦s=r'; c=c'⟧ ⟹ P" "⟦c'=(c'-{#s#})+{#s#}; c=(c-{#r'#})+{#r'#}; c-{#r'#} = c'-{#s#} ⟧ ⟹ P"
shows "P"
proof -
from A have "add_mset s c = add_mset r' c'" by (simp add: union_ac)
thus ?thesis using CASES by (cases rule: mset_single_cases) simp_all
qed
lemma mset_single_cases2'[cases set, case_names loc env]:
assumes A: "add_mset s c = add_mset r' c'"
assumes CASES: "⟦s=r'; c=c'⟧ ⟹ P" "!!cc. ⟦c'=cc+{#s#}; c=cc+{#r'#}; c'-{#s#}=cc; c-{#r'#}=cc⟧ ⟹ P"
shows "P"
using A CASES by (auto elim!: mset_single_cases2)
lemma mset_un_single_un_cases [consumes 1, case_names left right]:
assumes A: "add_mset a A = B + C"
and CASES: "a ∈# B ⟹ A = (B - {#a#}) + C ⟹ P"
"a ∈# C ⟹ A = B + (C - {#a#}) ⟹ P" shows "P"
proof -
have "a ∈# A+{#a#}" by simp
with A have "a ∈# B+C" by auto
thus ?thesis proof (cases rule: mset_un_cases)
case left hence "B=B-{#a#}+{#a#}" by auto
with A have "A+{#a#} = (B-{#a#})+C+{#a#}" by (auto simp add: union_ac)
hence "A=(B-{#a#})+C" by simp
with CASES(1)[OF left] show ?thesis by blast
next
case right hence "C=C-{#a#}+{#a#}" by auto
with A have "A+{#a#} = B+(C-{#a#})+{#a#}" by (auto simp add: union_ac)
hence "A=B+(C-{#a#})" by simp
with CASES(2)[OF right] show ?thesis by blast
qed
qed
lemma mset_distrib[consumes 1, case_names dist]: assumes A: "(A::'a multiset)+B = M+N" "!!Am An Bm Bn. ⟦A=Am+An; B=Bm+Bn; M=Am+Bm; N=An+Bn⟧ ⟹ P" shows "P"
proof -
have BN_MA: "B - N = M - A"
by (metis (no_types) add_diff_cancel_right assms(1) union_commute)
have H: "A = A∩# C + (A - C) ∩# D" if "A + B = C + D" for A B C D :: "'a multiset"
by (metis add.commute diff_intersect_left_idem mset_subset_eq_add_left subset_eq_diff_conv
subset_mset.add_diff_inverse subset_mset.inf_absorb1 subset_mset.inf_le1 that)
have A': "A = A∩# M + (A - M) ∩# N"
using A(1) H by blast
moreover have B': "B = (B - N) ∩# M + B∩# N"
using A(1) H[of B A N M] by (auto simp: ac_simps)
moreover have "M = A ∩# M + (B - N) ∩# M"
using H[of M N A B] BN_MA[symmetric] A(1) by (metis (no_types) diff_intersect_left_idem
diff_union_cancelR multiset_inter_commute subset_mset.diff_add subset_mset.inf.cobounded1
union_commute)
moreover have "N = (A - M) ∩# N + B ∩# N"
by (metis A' assms(1) diff_union_cancelL inter_union_distrib_left inter_union_distrib_right
mset_subset_eq_multiset_union_diff_commute subset_mset.inf.cobounded1 subset_mset.inf.commute)
ultimately show P
using A(2) by blast
qed
subsubsection ‹Singleton multisets›
lemma mset_size_le1_cases[case_names empty singleton,consumes 1]: "⟦ size M ≤ Suc 0; M={#} ⟹ P; !!m. M={#m#} ⟹ P ⟧ ⟹ P"
by (cases M) auto
lemma diff_union_single_conv2: "a ∈# J ⟹ J + I - {#a#} = (J - {#a#}) + I"
by simp
lemmas diff_union_single_convs = diff_union_single_conv diff_union_single_conv2
lemma mset_contains_eq: "(m ∈# M) = ({#m#}+(M-{#m#})=M)"
using diff_single_trivial by fastforce
subsubsection ‹Pointwise ordering›
lemma mset_le_incr_right1: "a⊆#(b::'a multiset) ⟹ a⊆#b+c" using mset_subset_eq_mono_add[of a b "{#}" c, simplified] .
lemma mset_le_incr_right2: "a⊆#(b::'a multiset) ⟹ a⊆#c+b" using mset_le_incr_right1
by (auto simp add: union_commute)
lemmas mset_le_incr_right = mset_le_incr_right1 mset_le_incr_right2
lemma mset_le_decr_left1: "a+c⊆#(b::'a multiset) ⟹ a⊆#b" using mset_le_incr_right1 mset_subset_eq_mono_add_right_cancel
by blast
lemma mset_le_decr_left2: "c+a⊆#(b::'a multiset) ⟹ a⊆#b" using mset_le_decr_left1
by (auto simp add: union_ac)
lemma mset_le_add_mset_decr_left1: "add_mset c a⊆#(b::'a multiset) ⟹ a⊆#b"
by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order)
lemma mset_le_add_mset_decr_left2: "add_mset c a⊆#(b::'a multiset) ⟹ {#c#}⊆#b"
by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order)
lemmas mset_le_decr_left = mset_le_decr_left1 mset_le_decr_left2 mset_le_add_mset_decr_left1
mset_le_add_mset_decr_left2
lemma mset_le_subtract: "A⊆#B ⟹ A-C ⊆# B-(C::'a multiset)"
apply (unfold subseteq_mset_def count_diff)
apply clarify
apply (subgoal_tac "count A a ≤ count B a")
apply arith
apply simp
done
lemma mset_union_subset: "A+B ⊆# C ⟹ A⊆#C ∧ B⊆#(C::'a multiset)"
by (auto dest: mset_le_decr_left)
lemma mset_le_add_mset: "add_mset x B ⊆# C ⟹ {#x#}⊆#C ∧ B⊆#(C::'a multiset)"
by (auto dest: mset_le_decr_left)
lemma mset_le_subtract_left: "A+B ⊆# (X::'a multiset) ⟹ B ⊆# X-A ∧ A⊆#X"
by (auto dest: mset_le_subtract[of "A+B" "X" "A"] mset_union_subset)
lemma mset_le_subtract_right: "A+B ⊆# (X::'a multiset) ⟹ A ⊆# X-B ∧ B⊆#X"
by (auto dest: mset_le_subtract[of "A+B" "X" "B"] mset_union_subset)
lemma mset_le_subtract_add_mset_left: "add_mset x B ⊆# (X::'a multiset) ⟹ B ⊆# X-{#x#} ∧ {#x#}⊆#X"
by (auto dest: mset_le_subtract[of "add_mset x B" "X" "{#x#}"] mset_le_add_mset)
lemma mset_le_subtract_add_mset_right: "add_mset x B ⊆# (X::'a multiset) ⟹ {#x#} ⊆# X-B ∧ B⊆#X"
by (auto dest: mset_le_subtract[of "add_mset x B" "X" "B"] mset_le_add_mset)
lemma mset_le_addE: "⟦ xs ⊆# (ys::'a multiset); !!zs. ys=xs+zs ⟹ P ⟧ ⟹ P" using mset_subset_eq_exists_conv
by blast
declare subset_mset.diff_add[simp, intro]
lemma mset_2dist2_cases:
assumes A: "{#a#}+{#b#} ⊆# A+B"
assumes CASES: "{#a#}+{#b#} ⊆# A ⟹ P" "{#a#}+{#b#} ⊆# B ⟹ P" "⟦a ∈# A; b ∈# B⟧ ⟹ P" "⟦a ∈# B; b ∈# A⟧ ⟹ P"
shows "P"
proof -
{ assume C: "a ∈# A" "b ∈# A-{#a#}"
with mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "A-{#a#}"] have "{#a#}+{#b#} ⊆# A" by auto
} moreover {
assume C: "a ∈# A" "¬ (b ∈# A-{#a#})"
with A have "b ∈# B"
by (metis diff_union_single_conv2 mset_le_subtract_left mset_subset_eq_insertD mset_un_cases)
} moreover {
assume C: "¬ (a ∈# A)" "b ∈# B-{#a#}"
with A have "a ∈# B"
by (auto dest: mset_subset_eqD)
with C mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "B-{#a#}"] have "{#a#}+{#b#} ⊆# B" by auto
} moreover {
assume C: "¬ (a ∈# A)" "¬ (b ∈# B-{#a#})"
with A have "a ∈# B ∧ b ∈# A"
apply (intro conjI)
apply (auto dest!: mset_subset_eq_insertD simp: insert_union_subset_iff; fail)[]
by (metis mset_diff_cancel1elem mset_le_subtract_left multiset_diff_union_assoc
single_subset_iff subset_eq_diff_conv)
} ultimately show P using CASES by blast
qed
lemma mset_union_subset_s: "{#a#}+B ⊆# C ⟹ a ∈# C ∧ B ⊆# C"
by (drule mset_union_subset) simp
lemma mset_le_single_cases[consumes 1, case_names empty singleton]: "⟦M⊆#{#a#}; M={#} ⟹ P; M={#a#} ⟹ P⟧ ⟹ P"
by (induct M) auto
lemma mset_le_distrib[consumes 1, case_names dist]: "⟦(X::'a multiset)⊆#A+B; !!Xa Xb. ⟦X=Xa+Xb; Xa⊆#A; Xb⊆#B⟧ ⟹ P ⟧ ⟹ P"
by (auto elim!: mset_le_addE mset_distrib)
lemma mset_le_mono_add_single: "⟦a ∈# ys; b ∈# ws⟧ ⟹ {#a#} + {#b#} ⊆# ys + ws"
by (meson mset_subset_eq_mono_add single_subset_iff)
lemma mset_size1elem: "⟦size P ≤ 1; q ∈# P⟧ ⟹ P={#q#}"
by (auto elim: mset_size_le1_cases)
lemma mset_size2elem: "⟦size P ≤ 2; {#q#}+{#q'#} ⊆# P⟧ ⟹ P={#q#}+{#q'#}"
by (auto elim: mset_le_addE)
subsubsection ‹Image under function›
notation image_mset (infixr "`#" 90)
lemma mset_map_split_orig: "!!M1 M2. ⟦f `# P = M1+M2; !!P1 P2. ⟦P=P1+P2; f `# P1 = M1; f `# P2 = M2⟧ ⟹ Q ⟧ ⟹ Q"
by (induct P) (force elim!: mset_un_single_un_cases)+
lemma mset_map_id: "⟦!!x. f (g x) = x⟧ ⟹ f `# g `# X = X"
by (induct X) auto
text ‹The following is a very specialized lemma. Intuitively, it splits the original multiset
by a splitting of some pointwise supermultiset of its image.
Application:
This lemma came in handy when proving the correctness of a constraint system that collects at most k sized submultisets of the sets of spawned threads.
›
lemma mset_map_split_orig_le: assumes A: "f `# P ⊆# M1+M2" and EX: "!!P1 P2. ⟦P=P1+P2; f `# P1 ⊆# M1; f `# P2 ⊆# M2⟧ ⟹ Q" shows "Q"
using A EX by (auto elim: mset_le_distrib mset_map_split_orig)
subsection ‹Lists›
lemma len_greater_imp_nonempty[simp]: "length l > x ⟹ l≠[]"
by auto
lemma list_take_induct_tl2:
"⟦length xs = length ys; ∀n<length xs. P (ys ! n) (xs ! n)⟧
⟹ ∀n < length (tl xs). P ((tl ys) ! n) ((tl xs) ! n)"
by (induct xs ys rule: list_induct2) auto
lemma not_distinct_split_distinct:
assumes "¬ distinct xs"
obtains y ys zs where "distinct ys" "y ∈ set ys" "xs = ys@[y]@zs"
using assms
proof (induct xs rule: rev_induct)
case Nil thus ?case by simp
next
case (snoc x xs) thus ?case by (cases "distinct xs") auto
qed
lemma distinct_length_le:
assumes d: "distinct ys"
and eq: "set ys = set xs"
shows "length ys ≤ length xs"
proof -
from d have "length ys = card (set ys)" by (simp add: distinct_card)
also from eq List.card_set have "card (set ys) = length (remdups xs)" by simp
also have "... ≤ length xs" by simp
finally show ?thesis .
qed
lemma find_SomeD:
"List.find P xs = Some x ⟹ P x"
"List.find P xs = Some x ⟹ x∈set xs"
by (auto simp add: find_Some_iff)
lemma in_hd_or_tl_conv[simp]: "l≠[] ⟹ x=hd l ∨ x∈set (tl l) ⟷ x∈set l"
by (cases l) auto
lemma length_dropWhile_takeWhile:
assumes "x < length (dropWhile P xs)"
shows "x + length (takeWhile P xs) < length xs"
using assms
by (induct xs) auto
text ‹Elim-version of @{thm neq_Nil_conv}.›
lemma neq_NilE:
assumes "l≠[]"
obtains x xs where "l=x#xs"
using assms by (metis list.exhaust)
lemma length_Suc_rev_conv: "length xs = Suc n ⟷ (∃ys y. xs=ys@[y] ∧ length ys = n)"
by (cases xs rule: rev_cases) auto
subsubsection ‹List Destructors›
lemma not_hd_in_tl:
"x ≠ hd xs ⟹ x ∈ set xs ⟹ x ∈ set (tl xs)"
by (induct xs) simp_all
lemma distinct_hd_tl:
"distinct xs ⟹ x = hd xs ⟹ x ∉ set (tl (xs))"
by (induct xs) simp_all
lemma in_set_tlD: "x ∈ set (tl xs) ⟹ x ∈ set xs"
by (induct xs) simp_all
lemma nth_tl: "xs ≠ [] ⟹ tl xs ! n = xs ! Suc n"
by (induct xs) simp_all
lemma tl_subset:
"xs ≠ [] ⟹ set xs ⊆ A ⟹ set (tl xs) ⊆ A"
by (metis in_set_tlD rev_subsetD subsetI)
lemma tl_last:
"tl xs ≠ [] ⟹ last xs = last (tl xs)"
by (induct xs) simp_all
lemma tl_obtain_elem:
assumes "xs ≠ []" "tl xs = []"
obtains e where "xs = [e]"
using assms
by (induct xs rule: list_nonempty_induct) simp_all
lemma butlast_subset:
"xs ≠ [] ⟹ set xs ⊆ A ⟹ set (butlast xs) ⊆ A"
by (metis in_set_butlastD rev_subsetD subsetI)
lemma butlast_rev_tl:
"xs ≠ [] ⟹ butlast (rev xs) = rev (tl xs)"
by (induct xs rule: rev_induct) simp_all
lemma hd_butlast:
"length xs > 1 ⟹ hd (butlast xs) = hd xs"
by (induct xs) simp_all
lemma butlast_upd_last_eq[simp]: "length l ≥ 2 ⟹
(butlast l) [ length l - 2 := x ] = take (length l - 2) l @ [x]"
apply (case_tac l rule: rev_cases)
apply simp
apply simp
apply (case_tac ys rule: rev_cases)
apply simp
apply simp
done
lemma distinct_butlast_swap[simp]:
"distinct pq ⟹ distinct (butlast (pq[i := last pq]))"
apply (cases pq rule: rev_cases)
apply (auto simp: list_update_append distinct_list_update split: nat.split)
done
subsubsection ‹Splitting list according to structure of other list›
context begin
private definition "SPLIT_ACCORDING m l ≡ length l = length m"
private lemma SPLIT_ACCORDINGE:
assumes "length m = length l"
obtains "SPLIT_ACCORDING m l"
unfolding SPLIT_ACCORDING_def using assms by auto
private lemma SPLIT_ACCORDING_simp:
"SPLIT_ACCORDING m (l1@l2) ⟷ (∃m1 m2. m=m1@m2 ∧ SPLIT_ACCORDING m1 l1 ∧ SPLIT_ACCORDING m2 l2)"
"SPLIT_ACCORDING m (x#l') ⟷ (∃y m'. m=y#m' ∧ SPLIT_ACCORDING m' l')"
apply (fastforce simp: SPLIT_ACCORDING_def intro: exI[where x = "take (length l1) m"] exI[where x = "drop (length l1) m"])
apply (cases m;auto simp: SPLIT_ACCORDING_def)
done
text ‹Split structure of list @{term m} according to structure of list @{term l}.›
method split_list_according for m :: "'a list" and l :: "'b list" =
(rule SPLIT_ACCORDINGE[of m l],
(simp; fail),
( simp only: SPLIT_ACCORDING_simp,
elim exE conjE,
simp only: SPLIT_ACCORDING_def
)
)
end
subsubsection ‹‹list_all2››
lemma list_all2_induct[consumes 1, case_names Nil Cons]:
assumes "list_all2 P l l'"
assumes "Q [] []"
assumes "⋀x x' ls ls'. ⟦ P x x'; list_all2 P ls ls'; Q ls ls' ⟧
⟹ Q (x#ls) (x'#ls')"
shows "Q l l'"
using list_all2_lengthD[OF assms(1)] assms
apply (induct rule: list_induct2)
apply auto
done
subsubsection ‹Indexing›
lemma ran_nth_set_encoding_conv[simp]:
"ran (λi. if i<length l then Some (l!i) else None) = set l"
apply safe
apply (auto simp: ran_def split: if_split_asm) []
apply (auto simp: in_set_conv_nth intro: ranI) []
done
lemma nth_image_indices[simp]: "(!) l ` {0..<length l} = set l"
by (auto simp: in_set_conv_nth)
lemma nth_update_invalid[simp]:"¬i<length l ⟹ l[j:=x]!i = l!i"
apply (induction l arbitrary: i j)
apply (auto split: nat.splits)
done
lemma nth_list_update': "l[i:=x]!j = (if i=j ∧ i<length l then x else l!j)"
by auto
lemma last_take_nth_conv: "n ≤ length l ⟹ n≠0 ⟹ last (take n l) = l!(n - 1)"
apply (induction l arbitrary: n)
apply (auto simp: take_Cons split: nat.split)
done
lemma nth_append_first[simp]: "i<length l ⟹ (l@l')!i = l!i"
by (simp add: nth_append)
lemma in_set_image_conv_nth: "f x ∈ f`set l ⟷ (∃i<length l. f (l!i) = f x)"
by (auto simp: in_set_conv_nth) (metis image_eqI nth_mem)
lemma set_image_eq_pointwiseI:
assumes "length l = length l'"
assumes "⋀i. i<length l ⟹ f (l!i) = f (l'!i)"
shows "f`set l = f`set l'"
using assms
by (fastforce simp: in_set_conv_nth in_set_image_conv_nth)
lemma insert_swap_set_eq: "i<length l ⟹ insert (l!i) (set (l[i:=x])) = insert x (set l)"
by (auto simp: in_set_conv_nth nth_list_update split: if_split_asm)
subsubsection ‹Reverse lists›
lemma neq_Nil_revE:
assumes "l≠[]"
obtains ll e where "l = ll@[e]"
using assms by (cases l rule: rev_cases) auto
lemma neq_Nil_rev_conv: "l≠[] ⟷ (∃xs x. l = xs@[x])"
by (cases l rule: rev_cases) auto
text ‹Caution: Same order of case variables in snoc-case as @{thm [source] rev_exhaust}, the other way round than @{thm [source] rev_induct} !›
lemma length_compl_rev_induct[case_names Nil snoc]: "⟦P []; !! l e . ⟦!! ll . length ll <= length l ⟹ P ll⟧ ⟹ P (l@[e])⟧ ⟹ P l"
apply(induct_tac l rule: length_induct)
apply(case_tac "xs" rule: rev_cases)
apply(auto)
done
lemma list_append_eq_Cons_cases[consumes 1]: "⟦ys@zs = x#xs; ⟦ys=[]; zs=x#xs⟧ ⟹ P; !!ys'. ⟦ ys=x#ys'; ys'@zs=xs ⟧ ⟹ P ⟧ ⟹ P"
by (auto iff add: append_eq_Cons_conv)
lemma list_Cons_eq_append_cases[consumes 1]: "⟦x#xs = ys@zs; ⟦ys=[]; zs=x#xs⟧ ⟹ P; !!ys'. ⟦ ys=x#ys'; ys'@zs=xs ⟧ ⟹ P ⟧ ⟹ P"
by (auto iff add: Cons_eq_append_conv)
lemma map_of_rev_distinct[simp]:
"distinct (map fst m) ⟹ map_of (rev m) = map_of m"
apply (induct m)
apply simp
apply simp
apply (subst map_add_comm)
apply force
apply simp
done
fun revg where
"revg [] b = b" |
"revg (a#as) b = revg as (a#b)"
lemma revg_fun[simp]: "revg a b = rev a @ b"
by (induct a arbitrary: b)
auto
lemma rev_split_conv[simp]:
"l ≠ [] ⟹ rev (tl l) @ [hd l] = rev l"
by (induct l) simp_all
lemma rev_butlast_is_tl_rev: "rev (butlast l) = tl (rev l)"
by (induct l) auto
lemma hd_last_singletonI:
"⟦xs ≠ []; hd xs = last xs; distinct xs⟧ ⟹ xs = [hd xs]"
by (induct xs rule: list_nonempty_induct) auto
lemma last_filter:
"⟦xs ≠ []; P (last xs)⟧ ⟹ last (filter P xs) = last xs"
by (induct xs rule: rev_nonempty_induct) simp_all
lemma rev_induct2' [case_names empty snocl snocr snoclr]:
assumes empty: "P [] []"
and snocl: "⋀x xs. P (xs@[x]) []"
and snocr: "⋀y ys. P [] (ys@[y])"
and snoclr: "⋀x xs y ys. P xs ys ⟹ P (xs@[x]) (ys@[y])"
shows "P xs ys"
proof (induct xs arbitrary: ys rule: rev_induct)
case Nil thus ?case using empty snocr
by (cases ys rule: rev_exhaust) simp_all
next
case snoc thus ?case using snocl snoclr
by (cases ys rule: rev_exhaust) simp_all
qed
lemma rev_nonempty_induct2' [case_names single snocl snocr snoclr, consumes 2]:
assumes "xs ≠ []" "ys ≠ []"
assumes single': "⋀x y. P [x] [y]"
and snocl: "⋀x xs y. xs ≠ [] ⟹ P (xs@[x]) [y]"
and snocr: "⋀x y ys. ys ≠ [] ⟹ P [x] (ys@[y])"
and snoclr: "⋀x xs y ys. ⟦P xs ys; xs ≠ []; ys≠[]⟧ ⟹ P (xs@[x]) (ys@[y])"
shows "P xs ys"
using assms(1,2)
proof (induct xs arbitrary: ys rule: rev_nonempty_induct)
case single then obtain z zs where "ys = zs@[z]" by (metis rev_exhaust)
thus ?case using single' snocr
by (cases "zs = []") simp_all
next
case (snoc x xs) then obtain z zs where zs: "ys = zs@[z]" by (metis rev_exhaust)
thus ?case using snocl snoclr snoc
by (cases "zs = []") simp_all
qed
subsubsection "Folding"
text "Ugly lemma about foldl over associative operator with left and right neutral element"
lemma foldl_A1_eq: "!!i. ⟦ !! e. f n e = e; !! e. f e n = e; !!a b c. f a (f b c) = f (f a b) c ⟧ ⟹ foldl f i ww = f i (foldl f n ww)"
proof (induct ww)
case Nil thus ?case by simp
next
case (Cons a ww i) note IHP[simplified]=this
have "foldl f i (a # ww) = foldl f (f i a) ww" by simp
also from IHP have "… = f (f i a) (foldl f n ww)" by blast
also from IHP(4) have "… = f i (f a (foldl f n ww))" by simp
also from IHP(1)[OF IHP(2,3,4), where i=a] have "… = f i (foldl f a ww)" by simp
also from IHP(2)[of a] have "… = f i (foldl f (f n a) ww)" by simp
also have "… = f i (foldl f n (a#ww))" by simp
finally show ?case .
qed
lemmas foldl_conc_empty_eq = foldl_A1_eq[of "(@)" "[]", simplified]
lemmas foldl_un_empty_eq = foldl_A1_eq[of "(∪)" "{}", simplified, OF Un_assoc[symmetric]]
lemma foldl_set: "foldl (∪) {} l = ⋃{x. x∈set l}"
apply (induct l)
apply simp_all
apply (subst foldl_un_empty_eq)
apply auto
done
lemma (in monoid_mult) foldl_absorb1: "x*foldl (*) 1 zs = foldl (*) x zs"
apply (rule sym)
apply (rule foldl_A1_eq)
apply (auto simp add: mult.assoc)
done
text ‹Towards an invariant rule for foldl›
lemma foldl_rule_aux:
fixes I :: "'σ ⇒ 'a list ⇒ bool"
assumes initial: "I σ0 l0"
assumes step: "!!l1 l2 x σ. ⟦ l0=l1@x#l2; I σ (x#l2) ⟧ ⟹ I (f σ x) l2"
shows "I (foldl f σ0 l0) []"
using initial step
apply (induct l0 arbitrary: σ0)
apply auto
done
lemma foldl_rule_aux_P:
fixes I :: "'σ ⇒ 'a list ⇒ bool"
assumes initial: "I σ0 l0"
assumes step: "!!l1 l2 x σ. ⟦ l0=l1@x#l2; I σ (x#l2) ⟧ ⟹ I (f σ x) l2"
assumes final: "!!σ. I σ [] ⟹ P σ"
shows "P (foldl f σ0 l0)"
using foldl_rule_aux[of I σ0 l0, OF initial, OF step] final
by simp
lemma foldl_rule:
fixes I :: "'σ ⇒ 'a list ⇒ 'a list ⇒ bool"
assumes initial: "I σ0 [] l0"
assumes step: "!!l1 l2 x σ. ⟦ l0=l1@x#l2; I σ l1 (x#l2) ⟧ ⟹ I (f σ x) (l1@[x]) l2"
shows "I (foldl f σ0 l0) l0 []"
using initial step
apply (rule_tac I="λσ lr. ∃ll. l0=ll@lr ∧ I σ ll lr" in foldl_rule_aux_P)
apply auto
done
text ‹
Invariant rule for foldl. The invariant is parameterized with
the state, the list of items that have already been processed and
the list of items that still have to be processed.
›
lemma foldl_rule_P:
fixes I :: "'σ ⇒ 'a list ⇒ 'a list ⇒ bool"
assumes initial: "I σ0 [] l0"
assumes step: "!!l1 l2 x σ. ⟦ l0=l1@x#l2; I σ l1 (x#l2) ⟧ ⟹ I (f σ x) (l1@[x]) l2"
assumes final: "!!σ. I σ l0 [] ⟹ P σ"
shows "P (foldl f σ0 l0)"
using foldl_rule[of I, OF initial step] by (simp add: final)
text ‹Invariant reasoning over @{const foldl} for distinct lists. Invariant rule makes no
assumptions about ordering.›
lemma distinct_foldl_invar:
"⟦ distinct S; I (set S) σ0;
⋀x it σ. ⟦x ∈ it; it ⊆ set S; I it σ⟧ ⟹ I (it - {x}) (f σ x)
⟧ ⟹ I {} (foldl f σ0 S)"
proof (induct S arbitrary: σ0)
case Nil thus ?case by auto
next
case (Cons x S)
note [simp] = Cons.prems(1)[simplified]
show ?case
apply simp
apply (rule Cons.hyps)
proof -
from Cons.prems(1) show "distinct S" by simp
from Cons.prems(3)[of x "set (x#S)", simplified,
OF Cons.prems(2)[simplified]]
show "I (set S) (f σ0 x)" .
fix xx it σ
assume A: "xx∈it" "it ⊆ set S" "I it σ"
show "I (it - {xx}) (f σ xx)" using A(2)
apply (rule_tac Cons.prems(3))
apply (simp_all add: A(1,3))
apply blast
done
qed
qed
lemma foldl_length_aux: "foldl (λi x. Suc i) a l = a + length l"
by (induct l arbitrary: a) auto
lemmas foldl_length[simp] = foldl_length_aux[where a=0, simplified]
lemma foldr_length_aux: "foldr (λx i. Suc i) l a = a + length l"
by (induct l arbitrary: a rule: rev_induct) auto
lemmas foldr_length[simp] = foldr_length_aux[where a=0, simplified]
context comp_fun_commute begin
lemma foldl_f_commute: "f a (foldl (λa b. f b a) b xs) = foldl (λa b. f b a) (f a b) xs"
by(induct xs arbitrary: b)(simp_all add: fun_left_comm)
lemma foldr_conv_foldl: "foldr f xs a = foldl (λa b. f b a) a xs"
by(induct xs arbitrary: a)(simp_all add: foldl_f_commute)
end
lemma filter_conv_foldr:
"filter P xs = foldr (λx xs. if P x then x # xs else xs) xs []"
by(induct xs) simp_all
lemma foldr_Cons: "foldr Cons xs [] = xs"
by(induct xs) simp_all
lemma foldr_snd_zip:
"length xs ≥ length ys ⟹ foldr (λ(x, y). f y) (zip xs ys) b = foldr f ys b"
proof(induct ys arbitrary: xs)
case (Cons y ys) thus ?case by(cases xs) simp_all
qed simp
lemma foldl_snd_zip:
"length xs ≥ length ys ⟹ foldl (λb (x, y). f b y) b (zip xs ys) = foldl f b ys"
proof(induct ys arbitrary: xs b)
case (Cons y ys) thus ?case by(cases xs) simp_all
qed simp
lemma fst_foldl: "fst (foldl (λ(a, b) x. (f a x, g a b x)) (a, b) xs) = foldl f a xs"
by(induct xs arbitrary: a b) simp_all
lemma foldl_foldl_conv_concat: "foldl (foldl f) a xs = foldl f a (concat xs)"
by(induct xs arbitrary: a) simp_all
lemma foldl_list_update:
"n < length xs ⟹ foldl f a (xs[n := x]) = foldl f (f (foldl f a (take n xs)) x) (drop (Suc n) xs)"
by(simp add: upd_conv_take_nth_drop)
lemma map_by_foldl:
fixes l :: "'a list" and f :: "'a ⇒ 'b"
shows "foldl (λl x. l@[f x]) [] l = map f l"
proof -
{
fix l'
have "foldl (λl x. l@[f x]) l' l = l'@map f l"
by (induct l arbitrary: l') auto
} thus ?thesis by simp
qed
subsubsection ‹Sorting›
lemma sorted_in_between:
assumes A: "0≤i" "i<j" "j<length l"
assumes S: "sorted l"
assumes E: "l!i ≤ x" "x<l!j"
obtains k where "i≤k" and "k<j" and "l!k≤x" and "x<l!(k+1)"
proof -
from A E have "∃k. i≤k ∧ k<j ∧ l!k≤x ∧ x<l!(k+1)"
proof (induct "j-i" arbitrary: i j)
case (Suc d)
show ?case proof (cases "l!(i+1) ≤ x")
case True
from True Suc.hyps have "d = j - (i + 1)" by simp
moreover from True have "i+1 < j"
by (metis Suc.prems Suc_eq_plus1 Suc_lessI not_less)
moreover from True have "0≤i+1" by simp
ultimately obtain k where
"i+1≤k" "k<j" "l!k ≤ x" "x<l!(k+1)"
using Suc.hyps(1)[of j "i+1"] Suc.prems True
by auto
thus ?thesis by (auto dest: Suc_leD)
next
case False
show ?thesis proof (cases "x<(l!(j - 1))")
case True
from True Suc.hyps have "d = j - (i + 1)" by simp
moreover from True Suc.prems have "i < j - 1"
by (metis Suc_eq_plus1 Suc_lessI diff_Suc_1 less_diff_conv not_le)
moreover from True Suc.prems have "j - 1 < length l" by simp
ultimately obtain k where
"i≤k" "k<j - 1" "l!k ≤ x" "x<l!(k+1)"
using Suc.hyps(1)[of "j - 1" i] Suc.prems True
by auto
thus ?thesis by (auto dest: Suc_leD)
next
case False thus ?thesis using Suc
apply clarsimp
by (metis Suc_leI add_0_iff add_diff_inverse diff_Suc_1 le_add2 lessI
not0_implies_Suc not_less)
qed
qed
qed simp
thus ?thesis by (blast intro: that)
qed
lemma sorted_hd_last:
"⟦sorted l; l≠[]⟧ ⟹ hd l ≤ last l"
by (metis eq_iff hd_Cons_tl last_in_set not_hd_in_tl sorted_wrt.simps(2))
lemma (in linorder) sorted_hd_min:
"⟦xs ≠ []; sorted xs⟧ ⟹ ∀x ∈ set xs. hd xs ≤ x"
by (induct xs, auto)
lemma sorted_append_bigger:
"⟦sorted xs; ∀x ∈ set xs. x ≤ y⟧ ⟹ sorted (xs @ [y])"
proof (induct xs)
case Nil
then show ?case by simp
next
case (Cons x xs)
then have s: "sorted xs" by (cases xs) simp_all
from Cons have a: "∀x∈set xs. x ≤ y" by simp
from Cons(1)[OF s a] Cons(2-) show ?case by (cases xs) simp_all
qed
lemma sorted_filter':
"sorted l ⟹ sorted (filter P l)"
using sorted_filter[where f=id, simplified] .
subsubsection ‹Map›
lemma map_eq_consE: "⟦map f ls = fa#fl; !!a l. ⟦ ls=a#l; f a=fa; map f l = fl ⟧ ⟹ P⟧ ⟹ P"
by auto
lemma map_fst_mk_snd[simp]: "map fst (map (λx. (x,k)) l) = l" by (induct l) auto
lemma map_snd_mk_fst[simp]: "map snd (map (λx. (k,x)) l) = l" by (induct l) auto
lemma map_fst_mk_fst[simp]: "map fst (map (λx. (k,x)) l) = replicate (length l) k" by (induct l) auto
lemma map_snd_mk_snd[simp]: "map snd (map (λx. (x,k)) l) = replicate (length l) k" by (induct l) auto
lemma map_zip1: "map (λx. (x,k)) l = zip l (replicate (length l) k)" by (induct l)