Theory Automatic_Refinement.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_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_inverse_image 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_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) auto
lemma map_zip2: "map (λx. (k,x)) l = zip (replicate (length l) k) l" by (induct l) auto
lemmas map_zip=map_zip1 map_zip2
lemma map_eq_nth_eq:
assumes A: "map f l = map f l'"
shows "f (l!i) = f (l'!i)"
proof -
from A have "length l = length l'"
by (metis length_map)
thus ?thesis using A
apply (induct arbitrary: i rule: list_induct2)
apply simp
apply (simp add: nth_def split: nat.split)
done
qed
lemma map_upd_eq:
"⟦i<length l ⟹ f (l!i) = f x⟧ ⟹ map f (l[i:=x]) = map f l"
by (metis list_update_beyond list_update_id map_update not_le_imp_less)
lemma inj_map_inv_f [simp]: "inj f ⟹ map (inv f) (map f l) = l"
by (simp)
lemma inj_on_map_the: "⟦D ⊆ dom m; inj_on m D⟧ ⟹ inj_on (the∘m) D"
apply (rule inj_onI)
apply simp
apply (case_tac "m x")
apply (case_tac "m y")
apply (auto intro: inj_onD) [1]
apply (auto intro: inj_onD) [1]
apply (case_tac "m y")
apply (auto intro: inj_onD) [1]
apply simp
apply (rule inj_onD)
apply assumption
apply auto
done
lemma map_consI:
"w=map f ww ⟹ f a#w = map f (a#ww)"
"w@l=map f ww@l ⟹ f a#w@l = map f (a#ww)@l"
by auto
lemma restrict_map_subset_eq:
fixes R
shows "⟦m |` R = m'; R'⊆R⟧ ⟹ m|` R' = m' |` R'"
by (auto simp add: Int_absorb1)
lemma restrict_map_self[simp]: "m |` dom m = m"
apply (rule ext)
apply (case_tac "m x")
apply (auto simp add: restrict_map_def)
done
lemma restrict_map_UNIV[simp]: "f |` UNIV = f"
by (auto simp add: restrict_map_def)
lemma restrict_map_inv[simp]: "f |` (- dom f) = Map.empty"
by (auto simp add: restrict_map_def intro: ext)
lemma restrict_map_upd: "(f |` S)(k ↦ v) = f(k↦v) |` (insert k S)"
by (auto simp add: restrict_map_def intro: ext)
lemma map_upd_eq_restrict: "m (x:=None) = m |` (-{x})"
by (auto intro: ext)
declare Map.finite_dom_map_of [simp, intro!]
lemma dom_const'[simp]: "dom (λx. Some (f x)) = UNIV"
by auto
lemma restrict_map_eq :
"((m |` A) k = None) ⟷ (k ∉ dom m ∩ A)"
"((m |` A) k = Some v) ⟷ (m k = Some v ∧ k ∈ A)"
unfolding restrict_map_def
by (simp_all add: dom_def)
definition "rel_of m P == {(k,v). m k = Some v ∧ P (k, v)}"
lemma rel_of_empty[simp]: "rel_of Map.empty P = {}"
by (auto simp add: rel_of_def)
lemma remove1_tl: "xs ≠ [] ⟹ remove1 (hd xs) xs = tl xs"
by (cases xs) auto
lemma set_oo_map_alt: "(set ∘∘ map) f = (λl. f ` set l)" by auto
subsubsection "Filter and Revert"
primrec filter_rev_aux where
"filter_rev_aux a P [] = a"
| "filter_rev_aux a P (x#xs) = (
if P x then filter_rev_aux (x#a) P xs else filter_rev_aux a P xs)"
lemma filter_rev_aux_alt: "filter_rev_aux a P l = filter P (rev l) @ a"
by (induct l arbitrary: a) auto
definition "filter_rev == filter_rev_aux []"
lemma filter_rev_alt: "filter_rev P l = filter P (rev l)"
unfolding filter_rev_def by (simp add: filter_rev_aux_alt)
definition "remove_rev x == filter_rev (Not o (=) x)"
lemma remove_rev_alt_def :
"remove_rev x xs = (filter (λy. y ≠ x) (rev xs))"
unfolding remove_rev_def
apply (simp add: filter_rev_alt comp_def)
by metis
subsubsection "zip"
declare zip_map_fst_snd[simp]
lemma pair_list_split: "⟦ !!l1 l2. ⟦ l = zip l1 l2; length l1=length l2; length l=length l2 ⟧ ⟹ P ⟧ ⟹ P"
proof (induct l arbitrary: P)
case Nil thus ?case by auto
next
case (Cons a l) from Cons.hyps obtain l1 l2 where IHAPP: "l=zip l1 l2" "length l1 = length l2" "length l=length l2" .
obtain a1 a2 where [simp]: "a=(a1,a2)" by (cases a) auto
from IHAPP have "a#l = zip (a1#l1) (a2#l2)" "length (a1#l1) = length (a2#l2)" "length (a#l) = length (a2#l2)"
by (simp_all only:) (simp_all (no_asm_use))
with Cons.prems show ?case by blast
qed
lemma set_zip_cart: "x∈set (zip l l') ⟹ x∈set l × set l'"
by (auto simp add: set_zip)
lemma map_prod_fun_zip: "map (λ(x, y). (f x, g y)) (zip xs ys) = zip (map f xs) (map g ys)"
proof(induct xs arbitrary: ys)
case Nil thus ?case by simp
next
case (Cons x xs) thus ?case by(cases ys) simp_all
qed
subsubsection ‹Generalized Zip›
text ‹Zip two lists element-wise, where the combination of two elements is specified by a function. Note that this function is underdefined for lists of different length.›
fun zipf :: "('a⇒'b⇒'c) ⇒ 'a list ⇒ 'b list ⇒ 'c list" where
"zipf f [] [] = []" |
"zipf f (a#as) (b#bs) = f a b # zipf f as bs"
lemma zipf_zip: "⟦length l1 = length l2⟧ ⟹ zipf Pair l1 l2 = zip l1 l2"
apply (induct l1 arbitrary: l2)
apply auto
apply (case_tac l2)
apply auto
done
fun list_all_zip where
"list_all_zip P [] [] ⟷ True" |
"list_all_zip P (a#as) (b#bs) ⟷ P a b ∧ list_all_zip P as bs" |
"list_all_zip P _ _ ⟷ False"
lemma list_all_zip_alt: "list_all_zip P as bs ⟷ length as = length bs ∧ (∀i<length as. P (as!i) (bs!i))"
apply (induct P≡P as bs rule: list_all_zip.induct)
apply auto
apply (case_tac i)
apply auto
done
lemma list_all_zip_map1: "list_all_zip P (List.map f as) bs ⟷ list_all_zip (λa b. P (f a) b) as bs"
apply (induct as arbitrary: bs)
apply (case_tac bs)
apply auto [2]
apply (case_tac bs)
apply auto [2]
done
lemma list_all_zip_map2: "list_all_zip P as (List.map f bs) ⟷ list_all_zip (λa b. P a (f b)) as bs"
apply (induct as arbitrary: bs)
apply (case_tac bs)
apply auto [2]
apply (case_tac bs)
apply auto [2]
done
declare list_all_zip_alt[mono]
lemma lazI[intro?]: "⟦ length a = length b; !!i. i<length b ⟹ P (a!i) (b!i) ⟧
⟹ list_all_zip P a b"
by (auto simp add: list_all_zip_alt)
lemma laz_conj[simp]: "list_all_zip (λx y. P x y ∧ Q x y) a b
⟷ list_all_zip P a b ∧ list_all_zip Q a b"
by (auto simp add: list_all_zip_alt)
lemma laz_len: "list_all_zip P a b ⟹ length a = length b"
by (simp add: list_all_zip_alt)
lemma laz_eq: "list_all_zip (=) a b ⟷ a=b"
apply (induct a arbitrary: b)
apply (case_tac b)
apply simp
apply simp
apply (case_tac b)
apply simp
apply simp
done
lemma laz_swap_ex:
assumes A: "list_all_zip (λa b. ∃c. P a b c) A B"
obtains C where
"list_all_zip (λa c. ∃b. P a b c) A C"
"list_all_zip (λb c. ∃a. P a b c) B C"
proof -
from A have
[simp]: "length A = length B" and
IC: "∀i<length B. ∃ci. P (A!i) (B!i) ci"
by (auto simp add: list_all_zip_alt)
from obtain_list_from_elements[OF IC] obtain C where
"length C = length B"
"∀i<length B. P (A!i) (B!i) (C!i)" .
thus ?thesis
by (rule_tac that) (auto simp add: list_all_zip_alt)
qed
lemma laz_weak_Pa[simp]:
"list_all_zip (λa b. P a) A B ⟷ (length A = length B) ∧ (∀a∈set A. P a)"
by (auto simp add: list_all_zip_alt set_conv_nth)
lemma laz_weak_Pb[simp]:
"list_all_zip (λa b. P b) A B ⟷ (length A = length B) ∧ (∀b∈set B. P b)"
by (force simp add: list_all_zip_alt set_conv_nth)
subsubsection "Collecting Sets over Lists"
definition "list_collect_set f l == ⋃{ f a | a. a∈set l }"
lemma list_collect_set_simps[simp]:
"list_collect_set f [] = {}"
"list_collect_set f [a] = f a"
"list_collect_set f (a#l) = f a ∪ list_collect_set f l"
"list_collect_set f (l@l') = list_collect_set f l ∪ list_collect_set f l'"
by (unfold list_collect_set_def) auto
lemma list_collect_set_map_simps[simp]:
"list_collect_set f (map x []) = {}"
"list_collect_set f (map x [a]) = f (x a)"
"list_collect_set f (map x (a#l)) = f (x a) ∪ list_collect_set f (map x l)"
"list_collect_set f (map x (l@l')) = list_collect_set f (map x l) ∪ list_collect_set f (map x l')"
by simp_all
lemma list_collect_set_alt: "list_collect_set f l = ⋃{ f (l!i) | i. i<length l }"
apply (induct l)
apply simp
apply safe
apply auto
apply (rule_tac x="f (l!i)" in exI)
apply simp
apply (rule_tac x="Suc i" in exI)
apply simp
apply (case_tac i)
apply auto
done
lemma list_collect_set_as_map: "list_collect_set f l = ⋃(set (map f l))"
by (unfold list_collect_set_def) auto
subsubsection ‹Sorted List with arbitrary Relations›
lemma (in linorder) sorted_wrt_rev_linord [simp] :
"sorted_wrt (≥) l ⟷ sorted (rev l)"
by (simp add: sorted_wrt_rev)
lemma (in linorder) sorted_wrt_map_linord [simp] :
"sorted_wrt (λ(x::'a × 'b) y. fst x ≤ fst y) l
⟷ sorted (map fst l)"
by (simp add: sorted_wrt_map)
lemma (in linorder) sorted_wrt_map_rev_linord [simp] :
"sorted_wrt (λ(x::'a × 'b) y. fst x ≥ fst y) l
⟷ sorted (rev (map fst l))"
by (induct l) (auto simp add: sorted_append)
subsubsection ‹Take and Drop›
lemma take_update[simp]: "take n (l[i:=x]) = (take n l)[i:=x]"
apply (induct l arbitrary: n i)
apply (auto split: nat.split)
apply (case_tac n)
apply simp_all
apply (case_tac n)
apply simp_all
done
lemma take_update_last: "length list>n ⟹ (take (Suc n) list) [n:=x] = take n list @ [x]"
by (induct list arbitrary: n)
(auto split: nat.split)
lemma drop_upd_irrelevant: "m < n ⟹ drop n (l[m:=x]) = drop n l"
apply (induct n arbitrary: l m)
apply simp
apply (case_tac l)
apply (auto split: nat.split)
done
lemma set_drop_conv:
"set (drop n l) = { l!i | i. n≤i ∧ i < length l }" (is "?L=?R")
proof (intro equalityI subsetI)
fix x
assume "x∈?L"
then obtain i where L: "i<length l - n" and X: "x = drop n l!i"
by (auto simp add: in_set_conv_nth)
note X
also have "… = l!(n+i)" using L by simp
finally show "x∈?R" using L by auto
next
fix x
assume "x∈?R"
then obtain i where L: "n≤i" "i<length l" and X: "x=l!i" by blast
note X
moreover have "l!i = drop n l ! (i - n)" and "(i-n) < length l - n" using L
by (auto)
ultimately show "x∈?L"
by (auto simp add: in_set_conv_nth)
qed
lemma filter_upt_take_conv:
"[i←[n..<m]. P (take m l ! i) ] = [i←[n..<m]. P (l ! i) ]"
by (rule filter_cong) (simp_all)
lemma in_set_drop_conv_nth: "x∈set (drop n l) ⟷ (∃i. n≤i ∧ i<length l ∧ x = l!i)"
apply (clarsimp simp: in_set_conv_nth)
apply safe
apply simp
apply (metis le_add2 less_diff_conv add.commute)
apply (rule_tac x="i-n" in exI)
apply auto []
done
lemma Union_take_drop_id: "⋃(set (drop n l)) ∪ ⋃(set (take n l)) = ⋃(set l)"
by (metis Union_Un_distrib append_take_drop_id set_union_code sup_commute)
lemma Un_set_drop_extend: "⟦j≥Suc 0; j < length l⟧
⟹ l ! (j - Suc 0) ∪ ⋃(set (drop j l)) = ⋃(set (drop (j - Suc 0) l))"
apply safe
apply simp_all
apply (metis diff_Suc_Suc diff_zero gr0_implies_Suc in_set_drop_conv_nth
le_refl less_eq_Suc_le order.strict_iff_order)
apply (metis Nat.diff_le_self set_drop_subset_set_drop subset_code(1))
by (metis diff_Suc_Suc gr0_implies_Suc in_set_drop_conv_nth
less_eq_Suc_le order.strict_iff_order minus_nat.diff_0)
lemma drop_take_drop_unsplit:
"i≤j ⟹ drop i (take j l) @ drop j l = drop i l"
proof -
assume "i ≤ j"
then obtain skf where "i + skf = j"
by (metis le_iff_add)
thus "drop i (take j l) @ drop j l = drop i l"
by (metis append_take_drop_id diff_add_inverse drop_drop drop_take
add.commute)
qed
lemma drop_last_conv[simp]: "l≠[] ⟹ drop (length l - Suc 0) l = [last l]"
by (cases l rule: rev_cases) auto
lemma take_butlast_conv[simp]: "take (length l - Suc 0) l = butlast l"
by (cases l rule: rev_cases) auto
lemma drop_takeWhile:
assumes "i≤length (takeWhile P l)"
shows "drop i (takeWhile P l) = takeWhile P (drop i l)"
using assms
proof (induction l arbitrary: i)
case Nil thus ?case by auto
next
case (Cons x l) thus ?case
by (cases i) auto
qed
lemma less_length_takeWhile_conv: "i < length (takeWhile P l) ⟷ (i<length l ∧ (∀j≤i. P (l!j)))"
apply safe
subgoal using length_takeWhile_le less_le_trans by blast
subgoal by (metis dual_order.strict_trans2 nth_mem set_takeWhileD takeWhile_nth)
subgoal by (meson less_le_trans not_le_imp_less nth_length_takeWhile)
done
lemma eq_len_takeWhile_conv: "i=length (takeWhile P l)
⟷ i≤length l ∧ (∀j<i. P (l!j)) ∧ (i<length l ⟶ ¬P (l!i))"
apply safe
subgoal using length_takeWhile_le less_le_trans by blast
subgoal by (auto simp: less_length_takeWhile_conv)
subgoal using nth_length_takeWhile by blast
subgoal by (metis length_takeWhile_le nth_length_takeWhile order.order_iff_strict)
subgoal by (metis dual_order.strict_trans2 leI less_length_takeWhile_conv linorder_neqE_nat nth_length_takeWhile)
done
subsubsection ‹Up-to›
lemma upt_eq_append_conv: "i≤j ⟹ [i..<j] = xs@ys ⟷ (∃k. i≤k ∧ k≤j ∧ [i..<k] = xs ∧ [k..<j] = ys)"
proof (rule iffI)
assume "[i..<j] = xs @ ys"
and "i≤j"
thus "∃k≥i. k ≤ j ∧ [i..<k] = xs ∧ [k..<j] = ys"
apply (induction xs arbitrary: i)
apply (auto; fail)
apply (clarsimp simp: upt_eq_Cons_conv)
by (meson Suc_le_eq less_imp_le_nat)
qed auto
lemma map_nth_upt_drop_take_conv: "N ≤ length l ⟹ map (nth l) [M..<N] = drop M (take N l)"
by (induction N) (auto simp: take_Suc_conv_app_nth)
lemma upt_eq_lel_conv:
"[l..<h] = is1@i#is2 ⟷ is1 = [l..<i] ∧ is2 = [Suc i..<h] ∧ l≤i ∧ i<h"
apply (rule)
subgoal
apply (induction is1 arbitrary: l)
apply (auto simp: upt_eq_Cons_conv) []
apply (clarsimp simp: upt_eq_Cons_conv)
using Suc_le_eq upt_rec by auto
subgoal by (auto simp: upt_conv_Cons[symmetric])
done
lemma map_add_upt': "map (λi. i + ofs) [a..<b] = [a+ofs..<b + ofs]"
by (induct b) simp_all
subsubsection ‹Last and butlast›
lemma butlast_upt: "butlast [m..<n] = [m..<n - 1]"
apply (cases "m<n")
apply (cases n)
apply simp
apply simp
apply simp
done
lemma butlast_update': "(butlast l) [i:=x] = butlast (l[i:=x])"
by (metis butlast_conv_take butlast_list_update length_butlast take_update)
lemma take_minus_one_conv_butlast:
"n≤length l ⟹ take (n - Suc 0) l = butlast (take n l)"
by (simp add: butlast_take)
lemma butlast_eq_cons_conv: "butlast l = x#xs ⟷ (∃xl. l=x#xs@[xl])"
by (metis Cons_eq_appendI append_butlast_last_id butlast.simps
butlast_snoc eq_Nil_appendI)
lemma butlast_eq_consE:
assumes "butlast l = x#xs"
obtains xl where "l=x#xs@[xl]"
using assms
by (auto simp: butlast_eq_cons_conv)
lemma drop_eq_ConsD: "drop n xs = x # xs' ⟹ drop (Suc n) xs = xs'"
by(induct xs arbitrary: n)(simp_all add: drop_Cons split: nat.split_asm)
subsubsection ‹List Slices›
text ‹Based on Lars Hupel's code.›
definition slice :: "nat ⇒ nat ⇒ 'a list ⇒ 'a list" where
"slice from to list = take (to - from) (drop from list)"
lemma slice_len[simp]: "⟦ from ≤ to; to ≤ length xs ⟧ ⟹ length (slice from to xs) = to - from"
unfolding slice_def
by simp
lemma slice_head: "⟦ from < to; to ≤ length xs ⟧ ⟹ hd (slice from to xs) = xs ! from"
unfolding slice_def
proof -
assume a1: "from < to"
assume "to ≤ length xs"
then have "⋀n. to - (to - n) ≤ length (take to xs)"
by (metis (no_types) slice_def diff_le_self drop_take length_drop slice_len)
then show "hd (take (to - from) (drop from xs)) = xs ! from"
using a1 by (metis diff_diff_cancel drop_take hd_drop_conv_nth leI le_antisym less_or_eq_imp_le nth_take)
qed
lemma slice_eq_bounds_empty[simp]: "slice i i xs = []"
unfolding slice_def by auto
lemma slice_nth: "⟦ from < to; to ≤ length xs; i < to - from ⟧ ⟹ slice from to xs ! i = xs ! (from + i)"
unfolding slice_def
by (induction "to - from" arbitrary: "from" to i) simp+
lemma slice_prepend: "⟦ i ≤ k; k ≤ length xs ⟧ ⟹ let p = length ys in slice i k xs = slice (i + p) (k + p) (ys @ xs)"
unfolding slice_def Let_def
by force
lemma slice_Nil[simp]: "slice begin end [] = []"
unfolding slice_def by auto
lemma slice_Cons: "slice begin end (x#xs)
= (if begin=0 ∧ end>0 then x#slice begin (end-1) xs else slice (begin - 1) (end - 1) xs)"
unfolding slice_def
by (auto simp: take_Cons' drop_Cons')
lemma slice_complete[simp]: "slice 0 (length xs) xs = xs"
unfolding slice_def
by simp
subsubsection ‹Miscellaneous›
lemma length_compl_induct[case_names Nil Cons]: "⟦P []; !! e l . ⟦!! ll . length ll <= length l ⟹ P ll⟧ ⟹ P (e#l)⟧ ⟹ P l"
apply(induct_tac l rule: length_induct)
apply(case_tac "xs")
apply(auto)
done
lemma in_set_list_format: "⟦ e∈set l; !!l1 l2. l=l1@e#l2 ⟹ P ⟧ ⟹ P"
proof (induct l arbitrary: P)
case Nil thus ?case by auto
next
case (Cons a l) show ?case proof (cases "a=e")
case True with Cons show ?thesis by force
next
case False with Cons.prems(1) have "e∈set l" by auto
with Cons.hyps obtain l1 l2 where "l=l1@e#l2" by blast
hence "a#l = (a#l1)@e#l2" by simp
with Cons.prems(2) show P by blast
qed
qed
lemma in_set_upd_cases:
assumes "x∈set (l[i:=y])"
obtains "i<length l" and "x=y" | "x∈set l"
by (metis assms in_set_conv_nth length_list_update nth_list_update_eq
nth_list_update_neq)
lemma in_set_upd_eq_aux:
assumes "i<length l"
shows "x∈set (l[i:=y]) ⟷ x=y ∨ (∀y. x∈set (l[i:=y]))"
by (metis in_set_upd_cases assms list_update_overwrite
set_update_memI)
lemma in_set_upd_eq:
assumes "i<length l"
shows "x∈set (l[i:=y]) ⟷ x=y ∨ (x∈set l ∧ (∀y. x∈set (l[i:=y])))"
by (metis in_set_upd_cases in_set_upd_eq_aux assms)
text ‹Simultaneous induction over two lists, prepending an element to one of the lists in each step›
lemma list_2pre_induct[case_names base left right]: assumes BASE: "P [] []" and LEFT: "!!e w1' w2. P w1' w2 ⟹ P (e#w1') w2" and RIGHT: "!!e w1 w2'. P w1 w2' ⟹ P w1 (e#w2')" shows "P w1 w2"
proof -
{
fix n
have "!!w1 w2. ⟦length w1 + length w2 = n; P [] []; !!e w1' w2. P w1' w2 ⟹ P (e#w1') w2; !!e w1 w2'. P w1 w2' ⟹ P w1 (e#w2') ⟧ ⟹ P w1 w2 "
apply (induct n)
apply simp
apply (case_tac w1)
apply auto
apply (case_tac w2)
apply auto
done
} from this[OF _ BASE LEFT RIGHT] show ?thesis by blast
qed
lemma list_decomp_1: "length l=1 ⟹ ∃a. l=[a]"
by (case_tac l, auto)
lemma list_decomp_2: "length l=2 ⟹ ∃a b. l=[a,b]"
by (case_tac l, auto simp add: list_decomp_1)
lemma list_rest_coinc: "⟦length s2 ≤ length s1; s1@r1 = s2@r2⟧ ⟹ ∃r1p. r2=r1p@r1"
by (metis append_eq_append_conv_if)
lemma list_tail_coinc: "n1#r1 = n2#r2 ⟹ n1=n2 & r1=r2"
by (auto)
lemma last_in_set[intro]: "⟦l≠[]⟧ ⟹ last l ∈ set l"
by (induct l) auto
lemma empty_append_eq_id[simp]: "(@) [] = (λx. x)" by auto
lemma op_conc_empty_img_id[simp]: "((@) [] ` L) = L" by auto
lemma distinct_match: "⟦ distinct (al@e#bl) ⟧ ⟹ (al@e#bl = al'@e#bl') ⟷ (al=al' ∧ bl=bl')"
proof (rule iffI, induct al arbitrary: al')
case Nil thus ?case by (cases al') auto
next
case (Cons a al) note Cprems=Cons.prems note Chyps=Cons.hyps
show ?case proof (cases al')
case Nil with Cprems have False by auto
thus ?thesis ..
next
case [simp]: (Cons a' all')
with Cprems have [simp]: "a=a'" and P: "al@e#bl = all'@e#bl'" by auto
from Cprems(1) have D: "distinct (al@e#bl)" by auto
from Chyps[OF D P] have [simp]: "al=all'" "bl=bl'" by auto
show ?thesis by simp
qed
qed simp
lemma prop_match: "⟦ list_all P al; ¬P e; ¬P e'; list_all P bl ⟧ ⟹ (al@e#bl = al'@e'#bl') ⟷ (al=al' ∧ e=e' ∧ bl=bl')"
apply (rule iffI, induct al arbitrary: al')
apply (case_tac al', fastforce, fastforce)+
done
lemmas prop_matchD = rev_iffD1[OF _ prop_match[where P=P]] for P
declare distinct_tl[simp]
lemma list_se_match[simp]:
"l1 ≠ [] ⟹ l1@l2 = [a] ⟷ l1 = [a] ∧ l2 = []"
"l2 ≠ [] ⟹ l1@l2 = [a] ⟷ l1 = [] ∧ l2 = [a]"
"l1 ≠ [] ⟹ [a] = l1@l2 ⟷ l1 = [a] ∧ l2 = []"
"l2 ≠ [] ⟹ [a] = l1@l2 ⟷ l1 = [] ∧ l2 = [a]"
apply (cases l1, simp_all)
apply (cases l1, simp_all)
apply (cases l1, auto) []
apply (cases l1, auto) []
done
lemma distinct_map_eq: "⟦ distinct (List.map f l); f x = f y; x∈set l; y∈set l ⟧ ⟹ x=y"
by (erule (2) xy_in_set_cases) auto
lemma upt_append:
assumes "i<j"
shows "[0..<i]@[i..<j] = [0..<j]"
using assms
apply (induct j)
apply simp
apply (case_tac "i=j")
apply auto
done
lemma upt_filter_extend:
assumes LE: "u≤u'"
assumes NP: "∀i. u≤i ∧ i<u' ⟶ ¬P i"
shows "[i←[0..<u]. P i] = [i←[0..<u']. P i]"
proof (cases "u=u'")
case True thus ?thesis by simp
next
case False hence "u<u'" using LE by simp
hence "[0..<u'] = [0..<u]@[u ..<u']"
by (simp add: upt_append)
hence "[i←[0..<u']. P i] = [i←[0..<u]. P i] @ [i←[u..<u']. P i]"
by simp
also have "[i←[u..<u']. P i] = []" using NP
by (auto simp: filter_empty_conv)
finally show ?thesis by simp
qed
lemma filter_upt_last:
assumes E: "[k←[0..<length l] . P (l!k)] = js @ [j]"
assumes "j<i" and "i<length l"
shows "¬ P (l!i)"
proof
assume A: "P (l!i)"
have "[0..<length l] = [0..<i]@[i..<length l]" using ‹i<length l›
by (simp add: upt_append)
also have "[i..<length l] = i#[Suc i..<length l]" using ‹i<length l›
by (auto simp: upt_conv_Cons)
finally
have "[k←[0..<i] . P (l!k)]@i#[k←[Suc i..<length l] . P (l!k)] = js@[j]"
unfolding E[symmetric]
using ‹P (l!i)› by simp
hence "j = last (i#[k←[Suc i..<length l] . P (l!k)])"
by (metis last_appendR last_snoc list.distinct(1))
also have "… ≥ i"
proof -
have "sorted (i#[k←[Suc i..<length l] . P (l!k)])" (is "sorted ?l")
by (simp add: sorted_filter[where f=id, simplified])
hence "hd ?l ≤ last ?l"
by (rule sorted_hd_last) simp
thus ?thesis by simp
qed
finally have "i≤j" . thus False using ‹j<i› by simp
qed
lemma all_set_conv_nth: "(∀x∈set l. P x) ⟷ (∀i<length l. P (l!i))"
by (auto intro: all_nth_imp_all_set)
lemma upt_0_eq_Nil_conv[simp]: "[0..<j] = [] ⟷ j=0"
by auto
lemma filter_eq_snocD: "filter P l = l'@[x] ⟹ x∈set l ∧ P x"
proof -
assume A: "filter P l = l'@[x]"
hence "x∈set (filter P l)" by simp
thus ?thesis by simp
qed
lemma lists_image_witness:
assumes A: "x∈lists (f`Q)"
obtains xo where "xo∈lists Q" "x=map f xo"
proof -
have "⟦ x∈lists (f`Q) ⟧ ⟹ ∃xo∈lists Q. x=map f xo"
proof (induct x)
case Nil thus ?case by auto
next
case (Cons x xs)
then obtain xos where "xos∈lists Q" "xs=map f xos" by force
moreover from Cons.prems have "x∈f`Q" by auto
then obtain xo where "xo∈Q" "x=f xo" by auto
ultimately show ?case
by (rule_tac x="xo#xos" in bexI) auto
qed
thus ?thesis
apply (simp_all add: A)
apply (erule_tac bexE)
apply (rule_tac that)
apply assumption+
done
qed
lemma map_of_None_filterD:
"map_of xs x = None ⟹ map_of (filter P xs) x = None"
by(induct xs) auto
lemma map_of_concat: "map_of (concat xss) = foldr (λxs f. f ++ map_of xs) xss Map.empty"
by(induct xss) simp_all
lemma map_of_Some_split:
"map_of xs k = Some v ⟹ ∃ys zs. xs = ys @ (k, v) # zs ∧ map_of ys k = None"
proof(induct xs)
case (Cons x xs)
obtain k' v' where x: "x = (k', v')" by(cases x)
show ?case
proof(cases "k' = k")
case True
with ‹map_of (x # xs) k = Some v› x have "x # xs = [] @ (k, v) # xs" "map_of [] k = None" by simp_all
thus ?thesis by blast
next
case False
with ‹map_of (x # xs) k = Some v› x
have "map_of xs k = Some v" by simp
from ‹map_of xs k = Some v ⟹ ∃ys zs. xs = ys @ (k, v) # zs ∧ map_of ys k = None›[OF this]
obtain ys zs where "xs = ys @ (k, v) # zs" "map_of ys k = None" by blast
with False x have "x # xs = (x # ys) @ (k, v) # zs" "map_of (x # ys) k = None" by simp_all
thus ?thesis by blast
qed
qed simp
lemma map_add_find_left:
"g k = None ⟹ (f ++ g) k = f k"
by(simp add: map_add_def)
lemma map_add_left_None:
"f k = None ⟹ (f ++ g) k = g k"
by(simp add: map_add_def split: option.split)
lemma map_of_Some_filter_not_in:
"⟦ map_of xs k = Some v; ¬ P (k, v); distinct (map fst xs) ⟧ ⟹ map_of (filter P xs) k = None"
apply(induct xs)
apply(auto)
apply(auto simp add: map_of_eq_None_iff)
done
lemma distinct_map_fst_filterI: "distinct (map fst xs) ⟹ distinct (map fst (filter P xs))"
by(induct xs) auto
lemma distinct_map_fstD: "⟦ distinct (map fst xs); (x, y) ∈ set xs; (x, z) ∈ set xs ⟧ ⟹ y = z"
by(induct xs)(fastforce elim: notE rev_image_eqI)+
lemma concat_filter_neq_Nil:
"concat [ys←xs. ys ≠ Nil] = concat xs"
by(induct xs) simp_all
lemma distinct_concat':
"⟦distinct [ys←xs. ys ≠ Nil]; ⋀ys. ys ∈ set xs ⟹ distinct ys;
⋀ys zs. ⟦ys ∈ set xs; zs ∈ set xs; ys ≠ zs⟧ ⟹ set ys ∩ set zs = {}⟧
⟹ distinct (concat xs)"
by(erule distinct_concat[of "[ys←xs. ys ≠ Nil]", unfolded concat_filter_neq_Nil]) auto
lemma distinct_idx:
assumes "distinct (map f l)"
assumes "i<length l"
assumes "j<length l"
assumes "f (l!i) = f (l!j)"
shows "i=j"
by (metis assms distinct_conv_nth length_map nth_map)
lemma replicate_Suc_conv_snoc:
"replicate (Suc n) x = replicate n x @ [x]"
by (metis replicate_Suc replicate_append_same)
lemma filter_nth_ex_nth:
assumes "n < length (filter P xs)"
shows "∃m. n ≤ m ∧ m < length xs ∧ filter P xs ! n = xs ! m ∧ filter P (take m xs) = take n (filter P xs)"
using assms
proof(induct xs rule: rev_induct)
case Nil thus ?case by simp
next
case (snoc x xs)
show ?case
proof(cases "P x")
case [simp]: False
from ‹n < length (filter P (xs @ [x]))› have "n < length (filter P xs)" by simp
hence "∃m≥n. m < length xs ∧ filter P xs ! n = xs ! m ∧ filter P (take m xs) = take n (filter P xs)" by(rule snoc)
thus ?thesis by(auto simp add: nth_append)
next
case [simp]: True
show ?thesis
proof(cases "n = length (filter P xs)")
case False
with ‹n < length (filter P (xs @ [x]))› have "n < length (filter P xs)" by simp
moreover hence "∃m≥n. m < length xs ∧ filter P xs ! n = xs ! m ∧ filter P (take m xs) = take n (filter P xs)"
by(rule snoc)
ultimately show ?thesis by(auto simp add: nth_append)
next
case [simp]: True
hence "filter P (xs @ [x]) ! n = (xs @ [x]) ! length xs" by simp
moreover have "length xs < length (xs @ [x])" by simp
moreover have "length xs ≥ n" by simp
moreover have "filter P (take (length xs) (xs @ [x])) = take n (filter P (xs @ [x]))" by simp
ultimately show ?thesis by blast
qed
qed
qed
lemma set_map_filter:
"set (List.map_filter g xs) = {y. ∃x. x ∈ set xs ∧ g x = Some y}"
by (induct xs) (auto simp add: List.map_filter_def set_eq_iff)
subsection ‹Quicksort by Relation›
text ‹A functional implementation of quicksort on lists. It it similar to the
one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary
relations.›
fun partition_rev :: "('a ⇒ bool) ⇒ ('a list × 'a list) ⇒ 'a list ⇒ ('a list × 'a list)" where
"partition_rev P (yes, no) [] = (yes, no)"
| "partition_rev P (yes, no) (x # xs) =
partition_rev P (if P x then (x # yes, no) else (yes, x # no)) xs"
lemma partition_rev_filter_conv :
"partition_rev P (yes, no) xs = (rev (filter P xs) @ yes, rev (filter (Not ∘ P) xs) @ no)"
by (induct xs arbitrary: yes no) (simp_all)
function quicksort_by_rel :: "('a ⇒ 'a ⇒ bool) ⇒ 'a list ⇒ 'a list ⇒ 'a list" where
"quicksort_by_rel R sl [] = sl"
| "quicksort_by_rel R sl (x#xs) =
(let (xs_s, xs_b) = partition_rev (λy. R y x) ([],[]) xs in
quicksort_by_rel R (x # (quicksort_by_rel R sl xs_b)) xs_s)"
by pat_completeness simp_all
termination
by (relation "measure (λ(_, _, xs). length xs)")
(simp_all add: partition_rev_filter_conv less_Suc_eq_le)
lemma quicksort_by_rel_remove_acc :
"quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)"
proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"])
case (less xs)
note ind_hyp = this
show ?case
proof (cases xs)
case Nil thus ?thesis by simp
next
case (Cons x xs') note xs_eq[simp] = this
obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (λy. R y x) ([], []) xs' = (xs1, xs2)"
by (rule prod.exhaust)
from part_rev_eq[symmetric]
have length_le: "length xs1 < length xs" "length xs2 < length xs"
unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le)
note ind_hyp1a = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2"]
note ind_hyp1b = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2 @ sl"]
note ind_hyp2 = ind_hyp[OF length_le(2), of sl]
show ?thesis by (simp add: ind_hyp1a ind_hyp1b ind_hyp2)
qed
qed
lemma quicksort_by_rel_remove_acc_guared :
"sl ≠ [] ⟹ quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)"
by (metis quicksort_by_rel_remove_acc)
lemma quicksort_by_rel_permutes [simp]:
"mset (quicksort_by_rel R sl xs) = mset (xs @ sl)"
proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"])
case (less xs)
note ind_hyp = this
show ?case
proof (cases xs)
case Nil thus ?thesis by simp
next
case (Cons x xs') note xs_eq[simp] = this
obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (λy. R y x) ([], []) xs' = (xs1, xs2)"
by (rule prod.exhaust)
from part_rev_eq[symmetric] have xs'_multi_eq : "mset xs' = mset xs1 + mset xs2"
unfolding partition_rev_filter_conv
by (simp add: mset_filter)
from part_rev_eq[symmetric]
have length_le: "length xs1 < length xs" "length xs2 < length xs"
unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le)
note ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)]
thus ?thesis by (simp add: xs'_multi_eq union_assoc)
qed
qed
lemma set_quicksort_by_rel [simp]: "set (quicksort_by_rel R sl xs) = set (xs @ sl)"
unfolding set_mset_comp_mset [symmetric] o_apply by simp
lemma sorted_wrt_quicksort_by_rel:
fixes R:: "'x ⇒ 'x ⇒ bool"
assumes lin : "⋀x y. (R x y) ∨ (R y x)"
and trans_R: "⋀x y z. R x y ⟹ R y z ⟹ R x z"
shows "sorted_wrt R (quicksort_by_rel R [] xs)"
proof (induct xs rule: measure_induct_rule[of "length"])
case (less xs)
note ind_hyp = this
show ?case
proof (cases xs)
case Nil thus ?thesis by simp
next
case (Cons x xs') note xs_eq[simp] = this
obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (λy. R y x) ([], []) xs' = (xs1, xs2)"
by (rule prod.exhaust)
from part_rev_eq[symmetric] have xs1_props: "⋀y. y ∈ set xs1 ⟹ (R y x)" and
xs2_props: "⋀y. y ∈ set xs2 ⟹ ¬(R y x)"
unfolding partition_rev_filter_conv
by simp_all
from xs2_props lin have xs2_props': "⋀y. y ∈ set xs2 ⟹ (R x y)" by blast
from xs2_props' xs1_props trans_R have xs1_props':
"⋀y1 y2. y1 ∈ set xs1 ⟹ y2 ∈ set xs2 ⟹ (R y1 y2)"
by metis
from part_rev_eq[symmetric]
have length_le: "length xs1 < length xs" "length xs2 < length xs"
unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le)
note ind_hyps = ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)]
thus ?thesis
by (simp add: quicksort_by_rel_remove_acc_guared sorted_wrt_append Ball_def
xs1_props xs2_props' xs1_props')
qed
qed
lemma sorted_quicksort_by_rel:
"sorted (quicksort_by_rel (≤) [] xs)"
by (rule sorted_wrt_quicksort_by_rel) auto
lemma sort_quicksort_by_rel:
"sort = quicksort_by_rel (≤) []"
apply (rule ext, rule properties_for_sort)
apply(simp_all add: sorted_quicksort_by_rel)
done
lemma [code]: "quicksort = quicksort_by_rel (≤) []"
apply (subst sort_quicksort[symmetric])
by (rule sort_quicksort_by_rel)
subsection ‹Mergesort by Relation›
text ‹A functional implementation of mergesort on lists. It it similar to the
one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary
relations.›
fun mergesort_by_rel_split :: "('a list × 'a list) ⇒ 'a list ⇒ ('a list × 'a list)" where
"mergesort_by_rel_split (xs1, xs2) [] = (xs1, xs2)"
| "mergesort_by_rel_split (xs1, xs2) [x] = (x # xs1, xs2)"
| "mergesort_by_rel_split (xs1, xs2) (x1 # x2 # xs) =
mergesort_by_rel_split (x1 # xs1, x2 # xs2) xs"
lemma list_induct_first2 [consumes 0, case_names Nil Sing Cons2]:
assumes "P []" "⋀x. P [x]" "⋀x1 x2 xs. P xs ⟹ P (x1 # x2 #xs)"
shows "P xs"
proof (induct xs rule: length_induct)
case (1 xs) note ind_hyp = this
show ?case
proof (cases xs)
case Nil thus ?thesis using assms(1) by simp
next
case (Cons x1 xs') note xs_eq[simp] = this
thus ?thesis
proof (cases xs')
case Nil thus ?thesis using assms(2) by simp
next
case (Cons x2 xs'') note xs'_eq[simp] = this
show ?thesis
by (simp add: ind_hyp assms(3))
qed
qed
qed
lemma mergesort_by_rel_split_length :
"length (fst (mergesort_by_rel_split (xs1, xs2) xs)) = length xs1 + (length xs div 2) + (length xs mod 2) ∧
length (snd (mergesort_by_rel_split (xs1, xs2) xs)) = length xs2 + (length xs div 2)"
by (induct xs arbitrary: xs1 xs2 rule: list_induct_first2)
(simp_all)
lemma mset_mergesort_by_rel_split [simp]:
"mset (fst (mergesort_by_rel_split (xs1, xs2) xs)) +
mset (snd (mergesort_by_rel_split (xs1, xs2) xs)) =
mset xs + mset xs1 + mset xs2"
apply (induct xs arbitrary: xs1 xs2 rule: list_induct_first2)
apply (simp_all add: ac_simps)
done
fun mergesort_by_rel_merge :: "('a ⇒ 'a ⇒ bool) ⇒ 'a list ⇒ 'a list ⇒ 'a list"
where
"mergesort_by_rel_merge R (x#xs) (y#ys) =
(if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)"
| "mergesort_by_rel_merge R xs [] = xs"
| "mergesort_by_rel_merge R [] ys = ys"
declare mergesort_by_rel_merge.simps [simp del]
lemma mergesort_by_rel_merge_simps[simp] :
"mergesort_by_rel_merge R (x#xs) (y#ys) =
(if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)"
"mergesort_by_rel_merge R xs [] = xs"
"mergesort_by_rel_merge R [] ys = ys"
apply (simp_all add: mergesort_by_rel_merge.simps)
apply (cases ys)
apply (simp_all add: mergesort_by_rel_merge.simps)
done
lemma mergesort_by_rel_merge_induct [consumes 0, case_names Nil1 Nil2 Cons1 Cons2]:
assumes "⋀xs::'a list. P xs []" "⋀ys::'b list. P [] ys"
"⋀x xs y ys. R x y ⟹ P xs (y # ys) ⟹ P (x # xs) (y # ys)"
"⋀x xs y ys. ¬(R x y) ⟹ P (x # xs) ys ⟹ P (x # xs) (y # ys)"
shows "P xs ys"
proof (induct xs arbitrary: ys)
case Nil thus ?case using assms(2) by simp
next
case (Cons x xs) note P_xs = this
show ?case
proof (induct ys)
case Nil thus ?case using assms(1) by simp
next
case (Cons y ys) note P_x_xs_ys = this
show ?case using assms(3,4)[of x y xs ys] P_x_xs_ys P_xs by metis
qed
qed
lemma mset_mergesort_by_rel_merge [simp]:
"mset (mergesort_by_rel_merge R xs ys) = mset xs + mset ys"
by (induct R xs ys rule: mergesort_by_rel_merge.induct) (simp_all add: ac_simps)
lemma set_mergesort_by_rel_merge [simp]:
"set (mergesort_by_rel_merge R xs ys) = set xs ∪ set ys"
by (induct R xs ys rule: mergesort_by_rel_merge.induct) auto
lemma sorted_wrt_mergesort_by_rel_merge [simp]:
assumes lin : "⋀x y. (R x y) ∨ (R y x)"
and trans_R: "⋀x y z. R x y ⟹ R y z ⟹ R x z"
shows "sorted_wrt R (mergesort_by_rel_merge R xs ys) ⟷
sorted_wrt R xs ∧ sorted_wrt R ys"
proof (induct xs ys rule: mergesort_by_rel_merge_induct[where R = R])
case Nil1 thus ?case by simp
next
case Nil2 thus ?case by simp
next
case (Cons1 x xs y ys) thus ?case
by (simp add: Ball_def) (metis trans_R)
next
case (Cons2 x xs y ys) thus ?case
apply (auto simp add: Ball_def)
apply (metis lin)
apply (metis lin trans_R)
done
qed
function mergesort_by_rel :: "('a ⇒ 'a ⇒ bool) ⇒ 'a list ⇒ 'a list"
where
"mergesort_by_rel R xs =
(if length xs < 2 then xs else
(mergesort_by_rel_merge R
(mergesort_by_rel R (fst (mergesort_by_rel_split ([], []) xs)))
(mergesort_by_rel R (snd (mergesort_by_rel_split ([], []) xs)))))"
by auto
termination
by (relation "measure (λ(_, xs). length xs)")
(simp_all add: mergesort_by_rel_split_length not_less minus_div_mult_eq_mod [symmetric])
declare mergesort_by_rel.simps [simp del]
lemma mergesort_by_rel_simps [simp, code] :
"mergesort_by_rel R [] = []"
"mergesort_by_rel R [x] = [x]"
"mergesort_by_rel R (x1 # x2 # xs) =
(let (xs1, xs2) = (mergesort_by_rel_split ([x1], [x2]) xs) in
mergesort_by_rel_merge R (mergesort_by_rel R xs1) (mergesort_by_rel R xs2))"
apply (simp add: mergesort_by_rel.simps)
apply (simp add: mergesort_by_rel.simps)
apply (simp add: mergesort_by_rel.simps[of _ "x1 # x2 # xs"] split: prod.splits)
done
lemma mergesort_by_rel_permutes [simp]:
"mset (mergesort_by_rel R xs) = mset xs"
proof (induct xs rule: length_induct)
case (1 xs) note ind_hyp = this
show ?case
proof (cases xs)
case Nil thus ?thesis by simp
next
case (Cons x1 xs') note xs_eq[simp] = this
show ?thesis
proof (cases xs')
case Nil thus ?thesis by simp
next
case (Cons x2 xs'') note xs'_eq[simp] = this
have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs"
"length (snd (mergesort_by_rel_split ([], []) xs)) < length xs"
by (simp_all add: mergesort_by_rel_split_length)
with ind_hyp show ?thesis
unfolding mergesort_by_rel.simps[of _ xs]
by (simp add: ac_simps)
qed
qed
qed
lemma set_mergesort_by_rel [simp]: "set (mergesort_by_rel R xs) = set xs"
unfolding set_mset_comp_mset [symmetric] o_apply by simp
lemma sorted_wrt_mergesort_by_rel:
fixes R:: "'x ⇒ 'x ⇒ bool"
assumes lin : "⋀x y. (R x y) ∨ (R y x)"
and trans_R: "⋀x y z. R x y ⟹ R y z ⟹ R x z"
shows "sorted_wrt R (mergesort_by_rel R xs)"
proof (induct xs rule: measure_induct_rule[of "length"])
case (less xs)
note ind_hyp = this
show ?case
proof (cases xs)
case Nil thus ?thesis by simp
next
case (Cons x xs') note xs_eq[simp] = this
thus ?thesis
proof (cases xs')
case Nil thus ?thesis by simp
next
case (Cons x2 xs'') note xs'_eq[simp] = this
have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs"
"length (snd (mergesort_by_rel_split ([], []) xs)) < length xs"
by (simp_all add: mergesort_by_rel_split_length)
with ind_hyp show ?thesis
unfolding mergesort_by_rel.simps[of _ xs]
by (simp add: sorted_wrt_mergesort_by_rel_merge[OF lin trans_R])
qed
qed
qed
lemma sorted_mergesort_by_rel:
"sorted (mergesort_by_rel (≤) xs)"
by (rule sorted_wrt_mergesort_by_rel) auto
lemma sort_mergesort_by_rel:
"sort = mergesort_by_rel (≤)"
apply (rule ext, rule properties_for_sort)
apply(simp_all add: sorted_mergesort_by_rel)
done
definition "mergesort = mergesort_by_rel (≤)"
lemma sort_mergesort: "sort = mergesort"
unfolding mergesort_def by (rule sort_mergesort_by_rel)
subsubsection ‹Mergesort with Remdup›
term merge
fun merge :: "'a::{linorder} list ⇒ 'a list ⇒ 'a list" where
"merge [] l2 = l2"
| "merge l1 [] = l1"
| "merge (x1 # l1) (x2 # l2) =
(if (x1 < x2) then x1 # (merge l1 (x2 # l2)) else
(if (x1 = x2) then x1 # (merge l1 l2) else x2 # (merge (x1 # l1) l2)))"
lemma merge_correct :
assumes l1_OK: "distinct l1 ∧ sorted l1"
assumes l2_OK: "distinct l2 ∧ sorted l2"
shows "distinct (merge l1 l2) ∧ sorted (merge l1 l2) ∧ set (merge l1 l2) = set l1 ∪ set l2"
using assms
proof (induct l1 arbitrary: l2)
case Nil thus ?case by simp
next
case (Cons x1 l1 l2)
note x1_l1_props = Cons(2)
note l2_props = Cons(3)
from x1_l1_props have l1_props: "distinct l1 ∧ sorted l1"
and x1_nin_l1: "x1 ∉ set l1"
and x1_le: "⋀x. x ∈ set l1 ⟹ x1 ≤ x"
by (simp_all add: Ball_def)
note ind_hyp_l1 = Cons(1)[OF l1_props]
show ?case
using l2_props
proof (induct l2)
case Nil with x1_l1_props show ?case by simp
next
case (Cons x2 l2)
note x2_l2_props = Cons(2)
from x2_l2_props have l2_props: "distinct l2 ∧ sorted l2"
and x2_nin_l2: "x2 ∉ set l2"
and x2_le: "⋀x. x ∈ set l2 ⟹ x2 ≤ x"
by (simp_all add: Ball_def)
note ind_hyp_l2 = Cons(1)[OF l2_props]
show ?case
proof (cases "x1 < x2")
case True note x1_less_x2 = this
from ind_hyp_l1[OF x2_l2_props] x1_less_x2 x1_nin_l1 x1_le x2_le
show ?thesis
apply (auto simp add: Ball_def)
apply (metis linorder_not_le)
apply (metis linorder_not_less xt1(6) xt1(9))
done
next
case False note x2_le_x1 = this
show ?thesis
proof (cases "x1 = x2")
case True note x1_eq_x2 = this
from ind_hyp_l1[OF l2_props] x1_le x2_le x2_nin_l2 x1_eq_x2 x1_nin_l1
show ?thesis by (simp add: x1_eq_x2 Ball_def)
next
case False note x1_neq_x2 = this
with x2_le_x1 have x2_less_x1 : "x2 < x1" by auto
from ind_hyp_l2 x2_le_x1 x1_neq_x2 x2_le x2_nin_l2 x1_le
show ?thesis
apply (simp add: x2_less_x1 Ball_def)
apply (metis linorder_not_le x2_less_x1 xt1(7))
done
qed
qed
qed
qed
function merge_list :: "'a::{linorder} list list ⇒ 'a list list ⇒ 'a list" where
"merge_list [] [] = []"
| "merge_list [] [l] = l"
| "merge_list (la # acc2) [] = merge_list [] (la # acc2)"
| "merge_list (la # acc2) [l] = merge_list [] (l # la # acc2)"
| "merge_list acc2 (l1 # l2 # ls) =
merge_list ((merge l1 l2) # acc2) ls"
by pat_completeness simp_all
termination
by (relation "measure (λ(acc, ls). 3 * length acc + 2 * length ls)") (simp_all)
lemma merge_list_correct :
assumes ls_OK: "⋀l. l ∈ set ls ⟹ distinct l ∧ sorted l"
assumes as_OK: "⋀l. l ∈ set as ⟹ distinct l ∧ sorted l"
shows "distinct (merge_list as ls) ∧ sorted (merge_list as ls) ∧
set (merge_list as ls) = set (concat (as @ ls))"
using assms
proof (induct as ls rule: merge_list.induct)
case 1 thus ?case by simp
next
case 2 thus ?case by simp
next
case 3 thus ?case by simp
next
case 4 thus ?case by auto
next
case (5 acc l1 l2 ls)
note ind_hyp = 5(1)
note l12_l_OK = 5(2)
note acc_OK = 5(3)
from l12_l_OK acc_OK merge_correct[of l1 l2]
have set_merge_eq: "set (merge l1 l2) = set l1 ∪ set l2" by auto
from l12_l_OK acc_OK merge_correct[of l1 l2]
have "distinct (merge_list (merge l1 l2 # acc) ls) ∧
sorted (merge_list (merge l1 l2 # acc) ls) ∧
set (merge_list (merge l1 l2 # acc) ls) =
set (concat ((merge l1 l2 # acc) @ ls))"
by (rule_tac ind_hyp) auto
with set_merge_eq show ?case by auto
qed
definition mergesort_remdups where
"mergesort_remdups xs = merge_list [] (map (λx. [x]) xs)"
lemma mergesort_remdups_correct :
"distinct (mergesort_remdups l)
∧ sorted (mergesort_remdups l)
∧ (set (mergesort_remdups l) = set l)"
proof -
let ?l' = "map (λx. [x]) l"
{ fix xs
assume "xs ∈ set ?l'"
then obtain x where xs_eq: "xs = [x]" by auto
hence "distinct xs ∧ sorted xs" by simp
} note l'_OK = this
from merge_list_correct[of "?l'" "[]", OF l'_OK]
show ?thesis unfolding mergesort_remdups_def by simp
qed
lemma ex1_eqI: "⟦∃!x. P x; P a; P b⟧ ⟹ a=b"
by blast
lemma remdup_sort_mergesort_remdups:
"remdups o sort = mergesort_remdups" (is "?lhs=?rhs")
proof
fix l
have "set (?lhs l) = set l" and "sorted (?lhs l)" and "distinct (?lhs l)"
by simp_all
moreover note mergesort_remdups_correct
ultimately show "?lhs l = ?rhs l"
by (auto intro!: ex1_eqI[OF finite_sorted_distinct_unique[OF finite_set]])
qed
subsection ‹Native Integers›
lemma int_of_integer_less_iff: "int_of_integer x < int_of_integer y ⟷ x<y"
by (simp add: less_integer_def)
lemma nat_of_integer_less_iff: "x≥0 ⟹ y≥0 ⟹ nat_of_integer x < nat_of_integer y ⟷ x<y"
unfolding nat_of_integer.rep_eq
by (auto simp: int_of_integer_less_iff nat_less_eq_zless int_of_integer_less_iff[of 0, simplified])
subsection "Natural Numbers"
lemma exists_leI:
assumes hyp: "(∀n' < n. ¬ P n') ⟹ P (n::nat)"
shows "∃n' ≤ n. P n'"
proof (rule classical)
assume contra: "¬ (∃n'≤n. P n')"
hence "∀n' < n. ¬ P n'" by auto
hence "P n" by (rule hyp)
thus "∃n'≤n. P n'" by auto
qed
subsubsection ‹Induction on nat›
lemma nat_compl_induct[case_names 0 Suc]: "⟦P 0; ⋀n . ∀nn. nn ≤ n ⟶ P nn ⟹ P (Suc n)⟧ ⟹ P n"
apply(induct_tac n rule: nat_less_induct)
apply(case_tac n)
apply(auto)
done
lemma nat_compl_induct'[case_names 0 Suc]: "⟦P 0; !! n . ⟦!! nn . nn ≤ n ⟹ P nn⟧ ⟹ P (Suc n)⟧ ⟹ P n"
apply(induct_tac n rule: nat_less_induct)
apply(case_tac n)
apply(auto)
done
lemma nz_le_conv_less: "0<k ⟹ k ≤ m ⟹ k - Suc 0 < m"
by auto
lemma min_Suc_gt[simp]:
"a<b ⟹ min (Suc a) b = Suc a"
"a<b ⟹ min b (Suc a) = Suc a"
by auto
subsection ‹Integer›
text ‹Some setup from ‹int› transferred to ‹integer››
lemma atLeastLessThanPlusOne_atLeastAtMost_integer: "{l..<u+1} = {l..(u::integer)}"
apply (auto simp add: atLeastAtMost_def atLeastLessThan_def)
including integer.lifting
apply transfer
apply simp
done
lemma atLeastPlusOneAtMost_greaterThanAtMost_integer: "{l+1..u} = {l<..(u::integer)}"
including integer.lifting
by (auto simp add: atLeastAtMost_def greaterThanAtMost_def, transfer, simp)
lemma atLeastPlusOneLessThan_greaterThanLessThan_integer:
"{l+1..<u} = {l<..<u::integer}"
including integer.lifting
by (auto simp add: atLeastLessThan_def greaterThanLessThan_def, transfer, simp)
lemma image_atLeastZeroLessThan_integer: "0 ≤ u ⟹
{(0::integer)..<u} = of_nat ` {..<nat_of_integer u}"
including integer.lifting
apply (unfold image_def lessThan_def)
apply auto
apply (rule_tac x = "nat_of_integer x" in exI)
apply transfer
apply (auto simp add: zless_nat_eq_int_zless [THEN sym])
apply transfer
apply simp
done
lemma image_add_integer_atLeastLessThan:
"(%x. x + (l::integer)) ` {0..<u-l} = {l..<u}"
apply (auto simp add: image_def)
apply (rule_tac x = "x - l" in bexI)
apply auto
done
lemma finite_atLeastZeroLessThan_integer: "finite {(0::integer)..<u}"
apply (cases "0 ≤ u")
apply (subst image_atLeastZeroLessThan_integer, assumption)
apply (rule finite_imageI)
apply auto
done
lemma finite_atLeastLessThan_integer [iff]: "finite {l..<u::integer}"
apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
apply (erule subst)
apply (rule finite_imageI)
apply (rule finite_atLeastZeroLessThan_integer)
apply (rule image_add_integer_atLeastLessThan)
done
lemma finite_atLeastAtMost_integer [iff]: "finite {l..(u::integer)}"
by (subst atLeastLessThanPlusOne_atLeastAtMost_integer [THEN sym], simp)
lemma finite_greaterThanAtMost_integer [iff]: "finite {l<..(u::integer)}"
by (subst atLeastPlusOneAtMost_greaterThanAtMost_integer [THEN sym], simp)
lemma finite_greaterThanLessThan_integer [iff]: "finite {l<..<u::integer}"
by (subst atLeastPlusOneLessThan_greaterThanLessThan_integer [THEN sym], simp)
subsection ‹Functions of type @{typ "bool⇒bool"}›
lemma boolfun_cases_helper: "g=(λx. False) | g=(λx. x) | g=(λx. True) | g= (λx. ¬x)"
proof -
{ assume "g False" "g True"
hence "g = (λx. True)" by (rule_tac ext, case_tac x, auto)
} moreover {
assume "g False" "¬g True"
hence "g = (λx. ¬x)" by (rule_tac ext, case_tac x, auto)
} moreover {
assume "¬g False" "g True"
hence "g = (λx. x)" by (rule_tac ext, case_tac x, auto)
} moreover {
assume "¬g False" "¬g True"
hence "g = (λx. False)" by (rule_tac ext, case_tac x, auto)
} ultimately show ?thesis by fast
qed
lemma boolfun_cases[case_names False Id True Neg]: "⟦g=(λx. False) ⟹ P; g=(λx. x) ⟹ P; g=(λx. True) ⟹ P; g=(λx. ¬x) ⟹ P⟧ ⟹ P"
proof -
note boolfun_cases_helper[of g]
moreover assume "g=(λx. False) ⟹ P" "g=(λx. x) ⟹ P" "g=(λx. True) ⟹ P" "g=(λx. ¬x) ⟹ P"
ultimately show ?thesis by fast
qed
subsection ‹Definite and indefinite description›
text "Combined definite and indefinite description for binary predicate"
lemma some_theI: assumes EX: "∃a b . P a b" and BUN: "!! b1 b2 . ⟦∃a . P a b1; ∃a . P a b2⟧ ⟹ b1=b2"
shows "P (SOME a . ∃b . P a b) (THE b . ∃a . P a b)"
proof -
from EX have "∃b. P (SOME a. ∃b. P a b) b" by (rule someI_ex)
moreover from EX have "∃b. ∃a. P a b" by blast
with BUN theI'[of "λb. ∃a. P a b"] have "∃a. P a (THE b. ∃a. P a b)" by (unfold Ex1_def, blast)
moreover note BUN
ultimately show ?thesis by (fast)
qed
lemma some_insert_self[simp]: "S≠{} ⟹ insert (SOME x. x∈S) S = S"
by (auto intro: someI)
lemma some_elem[simp]: "S≠{} ⟹ (SOME x. x∈S) ∈ S"
by (auto intro: someI)
subsubsection‹Hilbert Choice with option›
definition Eps_Opt where
"Eps_Opt P = (if (∃x. P x) then Some (SOME x. P x) else None)"
lemma some_opt_eq_trivial[simp] :
"Eps_Opt (λy. y = x) = Some x"
unfolding Eps_Opt_def by simp
lemma some_opt_sym_eq_trivial[simp] :
"Eps_Opt ((=) x) = Some x"
unfolding Eps_Opt_def by simp
lemma some_opt_false_trivial[simp] :
"Eps_Opt (λ_. False) = None"
unfolding Eps_Opt_def by simp
lemma Eps_Opt_eq_None[simp] :
"Eps_Opt P = None ⟷ ¬(Ex P)"
unfolding Eps_Opt_def by simp
lemma Eps_Opt_eq_Some_implies :
"Eps_Opt P = Some x ⟹ P x"
unfolding Eps_Opt_def
by (metis option.inject option.simps(2) someI_ex)
lemma Eps_Opt_eq_Some :
assumes P_prop: "⋀x'. P x ⟹ P x' ⟹ x' = x"
shows "Eps_Opt P = Some x ⟷ P x"
using P_prop
unfolding Eps_Opt_def
by (metis option.inject option.simps(2) someI_ex)
subsection ‹Product Type›
lemma nested_case_prod_simp: "(λ(a,b) c. f a b c) x y =
(case_prod (λa b. f a b y) x)"
by (auto split: prod.split)
lemma fn_fst_conv: "(λx. (f (fst x))) = (λ(a,_). f a)"
by auto
lemma fn_snd_conv: "(λx. (f (snd x))) = (λ(_,b). f b)"
by auto
fun pairself where
"pairself f (a,b) = (f a, f b)"
lemma pairself_image_eq[simp]:
"pairself f ` {(a,b). P a b} = {(f a, f b)| a b. P a b}"
by force
lemma pairself_image_cart[simp]: "pairself f ` (A×B) = f`A × f`B"
by (auto simp: image_def)
lemma in_prod_fst_sndI: "fst x ∈ A ⟹ snd x ∈ B ⟹ x∈A×B"
by (cases x) auto
lemma inj_Pair[simp]:
"inj_on (λx. (x,c x)) S"
"inj_on (λx. (c x,x)) S"
by (auto intro!: inj_onI)
declare Product_Type.swap_inj_on[simp]
lemma img_fst [intro]:
assumes "(a,b) ∈ S"
shows "a ∈ fst ` S"
by (rule image_eqI[OF _ assms]) simp
lemma img_snd [intro]:
assumes "(a,b) ∈ S"
shows "b ∈ snd ` S"
by (rule image_eqI[OF _ assms]) simp
lemma range_prod:
"range f ⊆ (range (fst ∘ f)) × (range (snd ∘ f))"
proof
fix y
assume "y ∈ range f"
then obtain x where y: "y = f x" by auto
hence "y = (fst(f x), snd(f x))"
by simp
thus "y ∈ (range (fst ∘ f)) × (range (snd ∘ f))"
by (fastforce simp add: image_def)
qed
lemma finite_range_prod:
assumes fst: "finite (range (fst ∘ f))"
and snd: "finite (range (snd ∘ f))"
shows "finite (range f)"
proof -
from fst snd have "finite (range (fst ∘ f) × range (snd ∘ f))"
by (rule finite_SigmaI)
thus ?thesis
by (rule finite_subset[OF range_prod])
qed
lemma fstE:
"x = (a,b) ⟹ P (fst x) ⟹ P a"
by (metis fst_conv)
lemma sndE:
"x = (a,b) ⟹ P (snd x) ⟹ P b"
by (metis snd_conv)
subsubsection ‹Uncurrying›
definition uncurry :: "('a ⇒ 'b ⇒ 'c) ⇒ 'a × 'b ⇒ 'c" where
"uncurry f ≡ λ(a,b). f a b"
lemma uncurry_apply[simp]: "uncurry f (a,b) = f a b"
unfolding uncurry_def
by simp
lemma curry_uncurry_id[simp]: "curry (uncurry f) = f"
unfolding uncurry_def
by simp
lemma uncurry_curry_id[simp]: "uncurry (curry f) = f"
unfolding uncurry_def
by simp
lemma do_curry: "f (a,b) = curry f a b" by simp
lemma do_uncurry: "f a b = uncurry f (a,b)" by simp
subsection ‹Sum Type›
lemma map_sum_Inr_conv: "map_sum fl fr s = Inr y ⟷ (∃x. s=Inr x ∧ y = fr x)"
by (cases s) auto
lemma map_sum_Inl_conv: "map_sum fl fr s = Inl y ⟷ (∃x. s=Inl x ∧ y = fl x)"
by (cases s) auto
subsection ‹Directed Graphs and Relations›
subsubsection "Reflexive-Transitive Closure"
lemma r_le_rtrancl[simp]: "S⊆S⇧*" by auto
lemma rtrancl_mono_rightI: "S⊆S' ⟹ S⊆S'⇧*" by auto
lemma trancl_sub:
"R ⊆ R⇧+"
by auto
lemma trancl_single[simp]:
"{(a,b)}⇧+ = {(a,b)}"
by (auto simp: trancl_insert)
text ‹Pick first non-reflexive step›
lemma converse_rtranclE'[consumes 1, case_names base step]:
assumes "(u,v)∈R⇧*"
obtains "u=v"
| vh where "u≠vh" and "(u,vh)∈R" and "(vh,v)∈R⇧*"
using assms
apply (induct rule: converse_rtrancl_induct)
apply auto []
apply (case_tac "y=z")
apply auto
done
lemma in_rtrancl_insert: "x∈R⇧* ⟹ x∈(insert r R)⇧*"
by (metis in_mono rtrancl_mono subset_insertI)
lemma rtrancl_apply_insert: "R⇧*``(insert x S) = insert x (R⇧*``(S∪R``{x}))"
apply (auto)
apply (erule converse_rtranclE)
apply auto [2]
apply (erule converse_rtranclE)
apply (auto intro: converse_rtrancl_into_rtrancl) [2]
done
text ‹A path in a graph either does not use nodes from S at all, or it has a prefix leading to a node in S and a suffix that does not use nodes in S›
lemma rtrancl_last_visit[cases set, case_names no_visit last_visit_point]:
shows
"⟦ (q,q')∈R⇧*;
(q,q')∈(R-UNIV×S)⇧* ⟹ P;
!!qt. ⟦ qt∈S; (q,qt)∈R⇧+; (qt,q')∈(R-UNIV×S)⇧* ⟧ ⟹ P
⟧ ⟹ P"
proof (induct rule: converse_rtrancl_induct[case_names refl step])
case refl thus ?case by auto
next
case (step q qh)
show P proof (rule step.hyps(3))
assume A: "(qh,q')∈(R-UNIV×S)⇧*"
show P proof (cases "qh∈S")
case False
with step.hyps(1) A have "(q,q')∈(R-UNIV×S)⇧*"
by (auto intro: converse_rtrancl_into_rtrancl)
with step.prems(1) show P .
next
case True
from step.hyps(1) have "(q,qh)∈R⇧+" by auto
with step.prems(2) True A show P by blast
qed
next
fix qt
assume A: "qt∈S" "(qh,qt)∈R⇧+" "(qt,q')∈(R-UNIV×S)⇧*"
with step.hyps(1) have "(q,qt)∈R⇧+" by auto
with step.prems(2) A(1,3) show P by blast
qed
qed
text ‹Less general version of ‹rtrancl_last_visit›, but there's a short automatic proof›
lemma rtrancl_last_visit': "⟦ (q,q')∈R⇧*; (q,q')∈(R-UNIV×S)⇧* ⟹ P; !!qt. ⟦ qt∈S; (q,qt)∈R⇧*; (qt,q')∈(R-UNIV×S)⇧* ⟧ ⟹ P ⟧ ⟹ P"
by (induct rule: converse_rtrancl_induct) (auto intro: converse_rtrancl_into_rtrancl)
lemma rtrancl_last_visit_node:
assumes "(s,s')∈R⇧*"
shows "s≠sh ∧ (s,s')∈(R ∩ (UNIV × (-{sh})))⇧* ∨
(s,sh)∈R⇧* ∧ (sh,s')∈(R ∩ (UNIV × (-{sh})))⇧*"
using assms
proof (induct rule: converse_rtrancl_induct)
case base thus ?case by auto
next
case (step s st)
moreover {
assume P: "(st,s')∈ (R ∩ UNIV × - {sh})⇧*"
{
assume "st=sh" with step have ?case
by auto
} moreover {
assume "st≠sh"
with ‹(s,st)∈R› have "(s,st)∈(R ∩ UNIV × - {sh})⇧*" by auto
also note P
finally have ?case by blast
} ultimately have ?case by blast
} moreover {
assume P: "(st, sh) ∈ R⇧* ∧ (sh, s') ∈ (R ∩ UNIV × - {sh})⇧*"
with step(1) have ?case
by (auto dest: converse_rtrancl_into_rtrancl)
} ultimately show ?case by blast
qed
text ‹Find last point where a path touches a set›
lemma rtrancl_last_touch: "⟦ (q,q')∈R⇧*; q∈S; !!qt. ⟦ qt∈S; (q,qt)∈R⇧*; (qt,q')∈(R-UNIV×S)⇧* ⟧ ⟹ P ⟧ ⟹ P"
by (erule rtrancl_last_visit') auto
text ‹A path either goes over edge once, or not at all›
lemma trancl_over_edgeE:
assumes "(u,w)∈(insert (v1,v2) E)⇧+"
obtains "(u,w)∈E⇧+"
| "(u,v1)∈E⇧*" and "(v2,w)∈E⇧*"
using assms
proof induct
case (base z) thus ?thesis
by (metis insertE prod.inject r_into_trancl' rtrancl_eq_or_trancl)
next
case (step y z) thus ?thesis
by (metis (opaque_lifting, no_types)
Pair_inject insertE rtrancl.simps trancl.simps trancl_into_rtrancl)
qed
lemma rtrancl_image_advance: "⟦q∈R⇧* `` Q0; (q,x)∈R⟧ ⟹ x∈R⇧* `` Q0"
by (auto intro: rtrancl_into_rtrancl)
lemma trancl_image_by_rtrancl: "(E⇧+)``Vi ∪ Vi = (E⇧*)``Vi"
by (metis Image_Id Un_Image rtrancl_trancl_reflcl)
lemma reachable_mono: "⟦R⊆R'; X⊆X'⟧ ⟹ R⇧*``X ⊆ R'⇧*``X'"
by (metis Image_mono rtrancl_mono)
lemma finite_reachable_advance:
"⟦ finite (E⇧*``{v0}); (v0,v)∈E⇧* ⟧ ⟹ finite (E⇧*``{v})"
by (erule finite_subset[rotated]) auto
lemma rtrancl_mono_mp: "U⊆V ⟹ x∈U⇧* ⟹ x∈V⇧*" by (metis in_mono rtrancl_mono)
lemma trancl_mono_mp: "U⊆V ⟹ x∈U⇧+ ⟹ x∈V⇧+" by (metis trancl_mono)
lemma rtrancl_mapI: "(a,b)∈E⇧* ⟹ (f a, f b)∈(pairself f `E)⇧*"
apply (induction rule: rtrancl_induct)
apply (force intro: rtrancl.intros)+
done
lemma rtrancl_image_advance_rtrancl:
assumes "q ∈ R⇧*``Q0"
assumes "(q,x) ∈ R⇧*"
shows "x ∈ R⇧*``Q0"
using assms
by (metis rtrancl_idemp rtrancl_image_advance)
lemma nth_step_trancl:
"⋀n m. ⟦ ⋀ n. n < length xs - 1 ⟹ (xs ! Suc n, xs ! n) ∈ R ⟧ ⟹ n < length xs ⟹ m < n ⟹ (xs ! n, xs ! m) ∈ R⇧+"
proof (induction xs)
case (Cons x xs)
hence "⋀n. n < length xs - 1 ⟹ (xs ! Suc n, xs ! n) ∈ R"
apply clarsimp
by (metis One_nat_def diff_Suc_eq_diff_pred nth_Cons_Suc zero_less_diff)
note IH = this[THEN Cons.IH]
from Cons obtain n' where n': "Suc n' = n" by (cases n) blast+
show ?case
proof (cases m)
case "0" with Cons have "xs ≠ []" by auto
with "0" Cons.prems(1)[of m] have "(xs ! 0, x) ∈ R" by simp
moreover from IH[where m = 0] have "⋀n. n < length xs ⟹ n > 0 ⟹ (xs ! n, xs ! 0) ∈ R⇧+" by simp
ultimately have "⋀n. n < length xs ⟹ (xs ! n, x) ∈ R⇧+" by (metis trancl_into_trancl gr0I r_into_trancl')
with Cons "0" show ?thesis by auto
next
case (Suc m') with Cons.prems n' have "n' < length xs" "m' < n'" by auto
with IH have "(xs ! n', xs ! m') ∈ R⇧+" by simp
with Suc n' show ?thesis by auto
qed
qed simp
lemma Image_empty_trancl_Image_empty:
"R `` {v} = {} ⟹ R⇧+ `` {v} = {}"
unfolding Image_def
by (auto elim: converse_tranclE)
lemma Image_empty_rtrancl_Image_id:
"R `` {v} = {} ⟹ R⇧* `` {v} = {v}"
unfolding Image_def
by (auto elim: converse_rtranclE)
lemma trans_rtrancl_eq_reflcl:
"trans A ⟹ A^* = A^="
by (simp add: rtrancl_trancl_reflcl)
lemma refl_on_reflcl_Image:
"refl_on B A ⟹ C ⊆ B ⟹ A^= `` C = A `` C"
by (auto simp add: Image_def dest: refl_onD)
lemma Image_absorb_rtrancl:
"⟦ trans A; refl_on B A; C ⊆ B ⟧ ⟹ A^* `` C = A `` C"
by (simp add: trans_rtrancl_eq_reflcl refl_on_reflcl_Image)
lemma trancl_Image_unfold_left: "E⇧+``S = E⇧*``E``S"
by (auto simp: trancl_unfold_left)
lemma trancl_Image_unfold_right: "E⇧+``S = E``E⇧*``S"
by (auto simp: trancl_unfold_right)
lemma trancl_Image_advance_ss: "(u,v)∈E ⟹ E⇧+``{v} ⊆ E⇧+``{u}"
by auto
lemma rtrancl_Image_advance_ss: "(u,v)∈E ⟹ E⇧*``{v} ⊆ E⇧*``{u}"
by auto
lemma trancl_union_outside:
assumes "(v,w) ∈ (E∪U)⇧+"
and "(v,w) ∉ E⇧+"
shows "∃x y. (v,x) ∈ (E∪U)⇧* ∧ (x,y) ∈ U ∧ (y,w) ∈ (E∪U)⇧*"
using assms
proof (induction)
case base thus ?case by auto
next
case (step w x)
show ?case
proof (cases "(v,w)∈E⇧+")
case True
from step have "(v,w)∈(E∪U)⇧*" by simp
moreover from True step have "(w,x) ∈ U" by (metis Un_iff trancl.simps)
moreover have "(x,x) ∈ (E∪U)⇧*" by simp
ultimately show ?thesis by blast
next
case False with step.IH obtain a b where "(v,a) ∈ (E∪U)⇧*" "(a,b) ∈ U" "(b,w) ∈ (E∪U)⇧*" by blast
moreover with step have "(b,x) ∈ (E∪U)⇧*" by (metis rtrancl_into_rtrancl)
ultimately show ?thesis by blast
qed
qed
lemma trancl_restrict_reachable:
assumes "(u,v) ∈ E⇧+"
assumes "E``S ⊆ S"
assumes "u∈S"
shows "(u,v) ∈ (E∩S×S)⇧+"
using assms
by (induction rule: converse_trancl_induct)
(auto intro: trancl_into_trancl2)
lemma rtrancl_image_unfold_right: "E``E⇧*``V ⊆ E⇧*``V"
by (auto intro: rtrancl_into_rtrancl)
lemma trancl_Image_in_Range:
"R⇧+ `` V ⊆ Range R"
by (auto elim: trancl.induct)
lemma rtrancl_Image_in_Field:
"R⇧* `` V ⊆ Field R ∪ V"
proof -
from trancl_Image_in_Range have "R⇧+ `` V ⊆ Field R"
unfolding Field_def by fast
hence "R⇧+ `` V ∪ V ⊆ Field R ∪ V" by blast
with trancl_image_by_rtrancl show ?thesis by metis
qed
lemma rtrancl_sub_insert_rtrancl:
"R⇧* ⊆ (insert x R)⇧*"
by (auto elim: rtrancl.induct rtrancl_into_rtrancl)
lemma trancl_sub_insert_trancl:
"R⇧+ ⊆ (insert x R)⇧+"
by (auto elim: trancl.induct trancl_into_trancl)
lemma Restr_rtrancl_mono:
"(v,w) ∈ (Restr E U)⇧* ⟹ (v,w) ∈ E⇧*"
by (metis inf_le1 rtrancl_mono subsetCE)
lemma Restr_trancl_mono:
"(v,w) ∈ (Restr E U)⇧+ ⟹ (v,w) ∈ E⇧+"
by (metis inf_le1 trancl_mono)
subsubsection "Converse Relation"
lemmas converse_add_simps = converse_Times trancl_converse[symmetric] converse_Un converse_Int
lemma dom_ran_disj_comp[simp]: "Domain R ∩ Range R = {} ⟹ R O R = {}"
by auto
lemma below_Id_inv[simp]: "R¯⊆Id ⟷ R⊆Id" by (auto)
subsubsection "Cyclicity"
lemma cyclicE: "⟦¬acyclic g; !!x. (x,x)∈g⇧+ ⟹ P⟧ ⟹ P"
by (unfold acyclic_def) blast
lemma acyclic_insert_cyclic: "⟦acyclic g; ¬acyclic (insert (x,y) g)⟧ ⟹ (y,x)∈g⇧*"
by (unfold acyclic_def) (auto simp add: trancl_insert)
text ‹
This lemma makes a case distinction about a path in a graph where a couple of edges with the same
endpoint have been inserted: If there is a path from a to b, then there's such a path in the original graph, or
there's a path that uses an inserted edge only once.
Originally, this lemma was used to reason about the graph of an updated acquisition history. Any path in
this graph is either already contained in the original graph, or passes via an
inserted edge. Because all the inserted edges point to the same target node, in the
second case, the path can be short-circuited to use exactly one inserted edge.
›
lemma trancl_multi_insert[cases set, case_names orig via]:
"⟦ (a,b)∈(r∪X×{m})⇧+;
(a,b)∈r⇧+ ⟹ P;
!!x. ⟦ x∈X; (a,x)∈r⇧*; (m,b)∈r⇧* ⟧ ⟹ P
⟧ ⟹ P"
proof (induct arbitrary: P rule: trancl_induct)
case (base b) thus ?case by auto
next
case (step b c) show ?case proof (rule step.hyps(3))
assume A: "(a,b)∈r⇧+"
note step.hyps(2)
moreover {
assume "(b,c)∈r"
with A have "(a,c)∈r⇧+" by auto
with step.prems have P by blast
} moreover {
assume "b∈X" "c=m"
with A have P by (rule_tac step.prems(2)) simp+
} ultimately show P by auto
next
fix x
assume A: "x ∈ X" "(a, x) ∈ r⇧*" "(m, b) ∈ r⇧*"
note step.hyps(2)
moreover {
assume "(b,c)∈r"
with A(3) have "(m,c)∈r⇧*" by auto
with step.prems(2)[OF A(1,2)] have P by blast
} moreover {
assume "b∈X" "c=m"
with A have P by (rule_tac step.prems(2)) simp+
} ultimately show P by auto
qed
qed
text ‹
Version of @{thm [source] trancl_multi_insert} for inserted edges with the same startpoint.
›
lemma trancl_multi_insert2[cases set, case_names orig via]:
"⟦(a,b)∈(r∪{m}×X)⇧+; (a,b)∈r⇧+ ⟹ P; !!x. ⟦ x∈X; (a,m)∈r⇧*; (x,b)∈r⇧* ⟧ ⟹ P ⟧ ⟹ P"
proof goal_cases
case prems: 1 from prems(1) have "(b,a)∈((r∪{m}×X)⇧+)¯" by simp
also have "((r∪{m}×X)⇧+)¯ = (r¯∪X×{m})⇧+" by (simp add: converse_add_simps)
finally have "(b, a) ∈ (r¯ ∪ X × {m})⇧+" .
thus ?case
by (auto elim!: trancl_multi_insert
intro: prems(2,3)
simp add: trancl_converse rtrancl_converse
)
qed
lemma cyclic_subset:
"⟦ ¬ acyclic R; R ⊆ S ⟧ ⟹ ¬ acyclic S"
unfolding acyclic_def
by (blast intro: trancl_mono)
subsubsection ‹Wellfoundedness›
lemma wf_min: assumes A: "wf R" "R≠{}" "!!m. m∈Domain R - Range R ⟹ P" shows P proof -
have H: "!!x. wf R ⟹ ∀y. (x,y)∈R ⟶ x∈Domain R - Range R ∨ (∃m. m∈Domain R - Range R)"
by (erule_tac wf_induct_rule[where P="λx. ∀y. (x,y)∈R ⟶ x∈Domain R - Range R ∨ (∃m. m∈Domain R - Range R)"]) auto
from A(2) obtain x y where "(x,y)∈R" by auto
with A(1,3) H show ?thesis by blast
qed
lemma finite_wf_eq_wf_converse: "finite R ⟹ wf (R¯) ⟷ wf R"
by (metis acyclic_converse finite_acyclic_wf finite_acyclic_wf_converse wf_acyclic)
lemma wf_max: assumes A: "wf (R¯)" "R≠{}" and C: "!!m. m∈Range R - Domain R ⟹ P" shows "P"
proof -
from A(2) have NE: "R¯≠{}" by auto
from wf_min[OF A(1) NE] obtain m where "m∈Range R - Domain R" by auto
thus P by (blast intro: C)
qed
lemma wf_bounded_supset: "finite S ⟹ wf {(Q',Q). Q'⊃Q ∧ Q'⊆ S}"
proof -
assume [simp]: "finite S"
hence [simp]: "!!x. finite (S-x)" by auto
have "{(Q',Q). Q⊂Q' ∧ Q'⊆ S} ⊆ inv_image ({(s'::nat,s). s'<s}) (λQ. card (S-Q))"
proof (intro subsetI, case_tac x, simp)
fix a b
assume A: "b⊂a ∧ a⊆S"
hence "S-a ⊂ S-b" by blast
thus "card (S-a) < card (S-b)" by (auto simp add: psubset_card_mono)
qed
moreover have "wf ({(s'::nat,s). s'<s})" by (rule wf_less)
ultimately show ?thesis by (blast intro: wf_inv_image wf_subset)
qed
lemma wf_no_path: "Domain R ∩ Range R = {} ⟹ wf R"
apply (rule wf_no_loop)
by simp
text ‹Extend a wf-relation by a break-condition›
definition "brk_rel R ≡
{((False,x),(False,y)) | x y. (x,y)∈R}
∪ {((True,x),(False,y)) | x y. True}"
lemma brk_rel_wf[simp,intro!]:
assumes WF[simp]: "wf R"
shows "wf (brk_rel R)"
proof -
have "wf {((False,x),(False,y)) | x y. (x,y)∈R}"
proof -
have "{((False,x),(False,y)) | x y. (x,y)∈R} ⊆ inv_image R snd"
by auto
from wf_subset[OF wf_inv_image[OF WF] this] show ?thesis .
qed
moreover have "wf {((True,x),(False,y)) | x y. True}"
by (rule wf_no_path) auto
ultimately show ?thesis
unfolding brk_rel_def
apply (subst Un_commute)
by (blast intro: wf_Un)
qed
subsubsection ‹Restrict Relation›
definition rel_restrict :: "('a × 'a) set ⇒ 'a set ⇒ ('a × 'a) set"
where
"rel_restrict R A ≡ {(v,w). (v,w) ∈ R ∧ v ∉ A ∧ w ∉ A}"
lemma rel_restrict_alt_def:
"rel_restrict R A = R ∩ (-A) × (-A)"
unfolding rel_restrict_def
by auto
lemma rel_restrict_empty[simp]:
"rel_restrict R {} = R"
by (simp add: rel_restrict_def)
lemma rel_restrict_notR:
assumes "(x,y) ∈ rel_restrict A R"
shows "x ∉ R" and "y ∉ R"
using assms
unfolding rel_restrict_def
by auto
lemma rel_restrict_sub:
"rel_restrict R A ⊆ R"
unfolding rel_restrict_def
by auto
lemma rel_restrict_Int_empty:
"A ∩ Field R = {} ⟹ rel_restrict R A = R"
unfolding rel_restrict_def Field_def
by auto
lemma Domain_rel_restrict:
"Domain (rel_restrict R A) ⊆ Domain R - A"
unfolding rel_restrict_def
by auto
lemma Range_rel_restrict:
"Range (rel_restrict R A) ⊆ Range R - A"
unfolding rel_restrict_def
by auto
lemma Field_rel_restrict:
"Field (rel_restrict R A) ⊆ Field R - A"
unfolding rel_restrict_def Field_def
by auto
lemma rel_restrict_compl:
"rel_restrict R A ∩ rel_restrict R (-A) = {}"
unfolding rel_restrict_def
by auto
lemma finite_rel_restrict:
"finite R ⟹ finite (rel_restrict R A)"
by (metis finite_subset rel_restrict_sub)
lemma R_subset_Field: "R ⊆ Field R × Field R"
unfolding Field_def
by auto
lemma homo_rel_restrict_mono:
"R ⊆ B × B ⟹ rel_restrict R A ⊆ (B - A) × (B - A)"
proof -
assume A: "R ⊆ B × B"
hence "Field R ⊆ B" unfolding Field_def by auto
with Field_rel_restrict have "Field (rel_restrict R A) ⊆ B - A"
by (metis Diff_mono order_refl order_trans)
with R_subset_Field show ?thesis by blast
qed
lemma rel_restrict_union:
"rel_restrict R (A ∪ B) = rel_restrict (rel_restrict R A) B"
unfolding rel_restrict_def
by auto
lemma rel_restrictI:
"x ∉ R ⟹ y ∉ R ⟹ (x,y) ∈ E ⟹ (x,y) ∈ rel_restrict E R"
unfolding rel_restrict_def
by auto
lemma rel_restrict_lift:
"(x,y) ∈ rel_restrict E R ⟹ (x,y) ∈ E"
unfolding rel_restrict_def
by simp
lemma rel_restrict_trancl_mem:
"(a,b) ∈ (rel_restrict A R)⇧+ ⟹ (a,b) ∈ rel_restrict (A⇧+) R"
by (induction rule: trancl_induct) (auto simp add: rel_restrict_def)
lemma rel_restrict_trancl_sub:
"(rel_restrict A R)⇧+ ⊆ rel_restrict (A⇧+) R"
by (metis subrelI rel_restrict_trancl_mem)
lemma rel_restrict_mono:
"A ⊆ B ⟹ rel_restrict A R ⊆ rel_restrict B R"
unfolding rel_restrict_def by auto
lemma rel_restrict_mono2:
"R ⊆ S ⟹ rel_restrict A S ⊆ rel_restrict A R"
unfolding rel_restrict_def by auto
lemma rel_restrict_Sigma_sub:
"rel_restrict ((A×A)⇧+) R ⊆ ((A - R) × (A - R))⇧+"
unfolding rel_restrict_def
by auto (metis DiffI converse_tranclE mem_Sigma_iff r_into_trancl tranclE)
lemma finite_reachable_restrictedI:
assumes F: "finite Q"
assumes I: "I⊆Q"
assumes R: "Range E ⊆ Q"
shows "finite (E⇧*``I)"
proof -
from I R have "E⇧*``I ⊆ Q"
by (force elim: rtrancl.cases)
also note F
finally (finite_subset) show ?thesis .
qed
context begin
private lemma rtrancl_restrictI_aux:
assumes "(u,v)∈(E-UNIV×R)⇧*"
assumes "u∉R"
shows "(u,v)∈(rel_restrict E R)⇧* ∧ v∉R"
using assms
by (induction) (auto simp: rel_restrict_def intro: rtrancl.intros)
corollary rtrancl_restrictI:
assumes "(u,v)∈(E-UNIV×R)⇧*"
assumes "u∉R"
shows "(u,v)∈(rel_restrict E R)⇧*"
using rtrancl_restrictI_aux[OF assms] ..
end
lemma E_closed_restr_reach_cases:
assumes P: "(u,v)∈E⇧*"
assumes CL: "E``R ⊆ R"
obtains "v∈R" | "u∉R" "(u,v)∈(rel_restrict E R)⇧*"
using P
proof (cases rule: rtrancl_last_visit[where S=R])
case no_visit
show ?thesis proof (cases "u∈R")
case True with P have "v∈R"
using rtrancl_reachable_induct[OF _ CL, where I="{u}"]
by auto
thus ?thesis ..
next
case False with no_visit have "(u,v)∈(rel_restrict E R)⇧*"
by (rule rtrancl_restrictI)
with False show ?thesis ..
qed
next
case (last_visit_point x)
from ‹(x, v) ∈ (E - UNIV × R)⇧*› have "(x,v)∈E⇧*"
by (rule rtrancl_mono_mp[rotated]) auto
with ‹x∈R› have "v∈R"
using rtrancl_reachable_induct[OF _ CL, where I="{x}"]
by auto
thus ?thesis ..
qed
lemma rel_restrict_trancl_notR:
assumes "(v,w) ∈ (rel_restrict E R)⇧+"
shows "v ∉ R" and "w ∉ R"
using assms
by (metis rel_restrict_trancl_mem rel_restrict_notR)+
lemma rel_restrict_tranclI:
assumes "(x,y) ∈ E⇧+"
and "x ∉ R" "y ∉ R"
and "E `` R ⊆ R"
shows "(x,y) ∈ (rel_restrict E R)⇧+"
using assms
proof (induct)
case base thus ?case by (metis r_into_trancl rel_restrictI)
next
case (step y z) hence "y ∉ R" by auto
with step show ?case by (metis trancl_into_trancl rel_restrictI)
qed
subsubsection ‹Single-Valued Relations›
lemma single_valued_inter1: "single_valued R ⟹ single_valued (R∩S)"
by (auto intro: single_valuedI dest: single_valuedD)
lemma single_valued_inter2: "single_valued R ⟹ single_valued (S∩R)"
by (auto intro: single_valuedI dest: single_valuedD)
lemma single_valued_below_Id: "R⊆Id ⟹ single_valued R"
by (auto intro: single_valuedI)
subsubsection ‹Bijective Relations›
definition "bijective R ≡
(∀x y z. (x,y)∈R ∧ (x,z)∈R ⟶ y=z) ∧
(∀x y z. (x,z)∈R ∧ (y,z)∈R ⟶ x=y)"
lemma bijective_alt: "bijective R ⟷ single_valued R ∧ single_valued (R¯)"
unfolding bijective_def single_valued_def by blast
lemma bijective_Id[simp, intro!]: "bijective Id"
by (auto simp: bijective_def)
lemma bijective_Empty[simp, intro!]: "bijective {}"
by (auto simp: bijective_def)
subsubsection ‹Miscellaneous›
lemma pair_vimage_is_Image[simp]: "(Pair u -` E) = E``{u}"
by auto
lemma fst_in_Field: "fst ` R ⊆ Field R"
by (simp add: Field_def fst_eq_Domain)
lemma snd_in_Field: "snd ` R ⊆ Field R"
by (simp add: Field_def snd_eq_Range)
lemma ran_map_of:
"ran (map_of xs) ⊆ snd ` set (xs)"
by (induct xs) (auto simp add: ran_def)
lemma Image_subset_snd_image:
"A `` B ⊆ snd ` A"
unfolding Image_def image_def
by force
lemma finite_Image_subset:
"finite (A `` B) ⟹ C ⊆ A ⟹ finite (C `` B)"
by (metis Image_mono order_refl rev_finite_subset)
lemma finite_Field_eq_finite[simp]: "finite (Field R) ⟷ finite R"
by (metis finite_cartesian_product finite_subset R_subset_Field finite_Field)
definition "fun_of_rel R x ≡ SOME y. (x,y)∈R"
lemma for_in_RI: "x∈Domain R ⟹ (x,fun_of_rel R x)∈R"
unfolding fun_of_rel_def
by (auto intro: someI)
lemma Field_not_elem:
"v ∉ Field R ⟹ ∀(x,y) ∈ R. x ≠ v ∧ y ≠ v"
unfolding Field_def by auto
lemma Sigma_UNIV_cancel[simp]: "(A × X - A × UNIV) = {}" by auto
lemma same_fst_trancl[simp]: "(same_fst P R)⇧+ = same_fst P (λx. (R x)⇧+)"
proof -
{
fix x y
assume "(x,y)∈(same_fst P R)⇧+"
hence "(x,y)∈same_fst P (λx. (R x)⇧+)"
by induction (auto simp: same_fst_def)
} moreover {
fix f f' x y
assume "((f,x),(f',y))∈same_fst P (λx. (R x)⇧+)"
hence [simp]: "f'=f" "P f" and 1: "(x,y)∈(R f)⇧+" by (auto simp: same_fst_def)
from 1 have "((f,x),(f',y))∈(same_fst P R)⇧+"
apply induction
subgoal by (rule r_into_trancl) auto
subgoal by (erule trancl_into_trancl) auto
done
} ultimately show ?thesis by auto
qed
subsection ‹‹option› Datatype›
lemma le_some_optE: "⟦Some m≤x; !!m'. ⟦x=Some m'; m≤m'⟧ ⟹ P⟧ ⟹ P"
by (cases x) auto
lemma not_Some_eq2[simp]: "(∀x y. v ≠ Some (x,y)) = (v = None)"
by (cases v) auto
subsection "Maps"
primrec the_default where
"the_default _ (Some x) = x"
| "the_default x None = x"
declare map_add_dom_app_simps[simp]
declare map_add_upd_left[simp]
lemma ran_add[simp]: "dom f ∩ dom g = {} ⟹ ran (f++g) = ran f ∪ ran g" by (fastforce simp add: ran_def map_add_def split: option.split_asm option.split)
lemma nempty_dom: "⟦e≠Map.empty; !!m. m∈dom e ⟹ P ⟧ ⟹ P"
by (subgoal_tac "dom e ≠ {}") (blast, auto)
lemma le_map_dom_mono: "m≤m' ⟹ dom m ⊆ dom m'"
apply (safe)
apply (drule_tac x=x in le_funD)
apply simp
apply (erule le_some_optE)
apply simp
done
lemma map_add_first_le: fixes m::"'a⇀('b::order)" shows "⟦ m≤m' ⟧ ⟹ m++n ≤ m'++n"
apply (rule le_funI)
apply (auto simp add: map_add_def split: option.split elim: le_funE)
done
lemma map_add_distinct_le: shows "⟦ m≤m'; n≤n'; dom m' ∩ dom n' = {} ⟧ ⟹ m++n ≤ m'++n'"
apply (rule le_funI)
apply (auto simp add: map_add_def split: option.split)
apply (fastforce elim: le_funE)
apply (drule le_map_dom_mono)
apply (drule le_map_dom_mono)
apply (case_tac "m x")
apply simp
apply (force)
apply (fastforce dest!: le_map_dom_mono)
apply (erule le_funE)
apply (erule_tac x=x in le_funE)
apply simp
done
lemma map_add_left_comm: assumes A: "dom A ∩ dom B = {}" shows "A ++ (B ++ C) = B ++ (A ++ C)"
proof -
have "A ++ (B ++ C) = (A++B)++C" by simp
also have "… = (B++A)++C" by (simp add: map_add_comm[OF A])
also have "… = B++(A++C)" by simp
finally show ?thesis .
qed
lemmas map_add_ac = map_add_assoc map_add_comm map_add_left_comm
lemma le_map_restrict[simp]: fixes m :: "'a ⇀ ('b::order)" shows "m |` X ≤ m"
by (rule le_funI) (simp add: restrict_map_def)
lemma map_of_distinct_upd:
"x ∉ set (map fst xs) ⟹ [x ↦ y] ++ map_of xs = (map_of xs) (x ↦ y)"
by (induct xs) (auto simp add: fun_upd_twist)
lemma map_of_distinct_upd2:
assumes "x ∉ set(map fst xs)"
"x ∉ set (map fst ys)"
shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x ↦ y)"
apply(insert assms)
apply(induct xs)
apply (auto intro: ext)
done
lemma map_of_distinct_upd3:
assumes "x ∉ set(map fst xs)"
"x ∉ set (map fst ys)"
shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x ↦ y)"
apply(insert assms)
apply(induct xs)
apply (auto intro: ext)
done
lemma map_of_distinct_upd4:
assumes "x ∉ set(map fst xs)"
"x ∉ set (map fst ys)"
shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)"
using assms by (induct xs) (auto simp: map_of_eq_None_iff)
lemma map_of_distinct_lookup:
assumes "x ∉ set(map fst xs)"
"x ∉ set (map fst ys)"
shows "map_of (xs @ (x,y) # ys) x = Some y"
proof -
have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x ↦ y)"
using assms map_of_distinct_upd2 by simp
thus ?thesis
by simp
qed
lemma ran_distinct:
assumes dist: "distinct (map fst al)"
shows "ran (map_of al) = snd ` set al"
using assms proof (induct al)
case Nil then show ?case by simp
next
case (Cons kv al)
then have "ran (map_of al) = snd ` set al" by simp
moreover from Cons.prems have "map_of al (fst kv) = None"
by (simp add: map_of_eq_None_iff)
ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp
qed
lemma ran_is_image:
"ran M = (the ∘ M) ` (dom M)"
unfolding ran_def dom_def image_def
by auto
lemma map_card_eq_iff:
assumes finite: "finite (dom M)"
and card_eq: "card (dom M) = card (ran M)"
and indom: "x ∈ dom M"
shows "(M x = M y) ⟷ (x = y)"
proof -
from ran_is_image finite card_eq have *: "inj_on (the ∘ M) (dom M)" using eq_card_imp_inj_on by metis
thus ?thesis
proof (cases "y ∈ dom M")
case False with indom show ?thesis by auto
next
case True with indom have "the (M x) = the (M y) ⟷ (x = y)" using inj_on_eq_iff[OF *] by auto
thus ?thesis by auto
qed
qed
lemma map_dom_ran_finite:
"finite (dom M) ⟹ finite (ran M)"
by (simp add: ran_is_image)
lemma map_update_eta_repair[simp]:
"dom (λx. if x=k then Some v else m x) = insert k (dom m)"
"m k = None ⟹ ran (λx. if x=k then Some v else m x) = insert v (ran m)"
apply auto []
apply (force simp: ran_def)
done
lemma map_leI[intro?]: "⟦⋀x v. m1 x = Some v ⟹ m2 x = Some v⟧ ⟹ m1⊆⇩mm2"
unfolding map_le_def by force
lemma map_leD: "m1⊆⇩mm2 ⟹ m1 k = Some v ⟹ m2 k = Some v"
unfolding map_le_def by force
lemma map_restrict_insert_none_simp: "m x = None ⟹ m|`(-insert x s) = m|`(-s)"
by (auto intro!: ext simp:restrict_map_def)
lemma eq_f_restr_conv: "s⊆dom (f A) ∧ A = f A |` (-s) ⟷ A ⊆⇩m f A ∧ s = dom (f A) - dom A"
apply auto
subgoal by (metis map_leI restrict_map_eq(2))
subgoal by (metis ComplD restrict_map_eq(2))
subgoal by (metis Compl_iff restrict_in)
subgoal by (force simp: map_le_def restrict_map_def)
done
corollary eq_f_restr_ss_eq: "⟦ s⊆dom (f A) ⟧ ⟹ A = f A |` (-s) ⟷ A ⊆⇩m f A ∧ s = dom (f A) - dom A"
using eq_f_restr_conv by blast
subsubsection ‹Simultaneous Map Update›
definition "map_mmupd m K v k ≡ if k∈K then Some v else m k"
lemma map_mmupd_empty[simp]: "map_mmupd m {} v = m"
by (auto simp: map_mmupd_def)
lemma mmupd_in_upd[simp]: "k∈K ⟹ map_mmupd m K v k = Some v"
by (auto simp: map_mmupd_def)
lemma mmupd_notin_upd[simp]: "k∉K ⟹ map_mmupd m K v k = m k"
by (auto simp: map_mmupd_def)
lemma map_mmupdE:
assumes "map_mmupd m K v k = Some x"
obtains "k∉K" "m k = Some x"
| "k∈K" "x=v"
using assms by (auto simp: map_mmupd_def split: if_split_asm)
lemma dom_mmupd[simp]: "dom (map_mmupd m K v) = dom m ∪ K"
by (auto simp: map_mmupd_def split: if_split_asm)
lemma le_map_mmupd_not_dom[simp, intro!]: "m ⊆⇩m map_mmupd m (K-dom m) v"
by (auto simp: map_le_def)
lemma map_mmupd_update_less: "K⊆K' ⟹ map_mmupd m (K - dom m) v ⊆⇩m map_mmupd m (K'-dom m) v"
by (auto simp: map_le_def map_mmupd_def)
subsection‹Connection between Maps and Sets of Key-Value Pairs›
definition map_to_set where
"map_to_set m = {(k, v) . m k = Some v}"
definition set_to_map where
"set_to_map S k = Eps_Opt (λv. (k, v) ∈ S)"
lemma set_to_map_simp :
assumes inj_on_fst: "inj_on fst S"
shows "(set_to_map S k = Some v) ⟷ (k, v) ∈ S"
proof (cases "∃v. (k, v) ∈ S")
case True
note kv_ex = this
then obtain v' where kv'_in: "(k, v') ∈ S" by blast
with inj_on_fst have kv''_in: "⋀v''. (k, v'') ∈ S ⟷ v' = v''"
unfolding inj_on_def Ball_def
by auto
show ?thesis
unfolding set_to_map_def
by (simp add: kv_ex kv''_in)
next
case False
hence kv''_nin: "⋀v''. (k, v'') ∉ S" by simp
thus ?thesis
by (simp add: set_to_map_def)
qed
lemma inj_on_fst_map_to_set :
"inj_on fst (map_to_set m)"
unfolding map_to_set_def inj_on_def by simp
lemma map_to_set_inverse :
"set_to_map (map_to_set m) = m"
proof
fix k
show "set_to_map (map_to_set m) k = m k"
proof (cases "m k")
case None note mk_eq = this
hence "⋀v. (k, v) ∉ map_to_set m"
unfolding map_to_set_def by simp
with set_to_map_simp [OF inj_on_fst_map_to_set, of m k]
show ?thesis unfolding mk_eq by auto
next
case (Some v) note mk_eq = this
hence "(k, v) ∈ map_to_set m"
unfolding map_to_set_def by simp
with set_to_map_simp [OF inj_on_fst_map_to_set, of m k v]
show ?thesis unfolding mk_eq by auto
qed
qed
lemma set_to_map_inverse :
assumes inj_on_fst_S: "inj_on fst S"
shows "map_to_set (set_to_map S) = S"
proof (rule set_eqI)
fix kv
from set_to_map_simp [OF inj_on_fst_S, of "fst kv" "snd kv"]
show "(kv ∈ map_to_set (set_to_map S)) = (kv ∈ S)"
unfolding map_to_set_def
by auto
qed
lemma map_to_set_empty[simp]: "map_to_set Map.empty = {}"
unfolding map_to_set_def by simp
lemma set_to_map_empty[simp]: "set_to_map {} = Map.empty"
unfolding set_to_map_def[abs_def] by simp
lemma map_to_set_empty_iff: "map_to_set m = {} ⟷ m = Map.empty"
"{} = map_to_set m ⟷ m = Map.empty"
unfolding map_to_set_def by auto
lemma set_to_map_empty_iff: "set_to_map S = Map.empty ⟷ S = {}" (is ?T1)
"Map.empty = set_to_map S ⟷ S = {}" (is ?T2)
proof -
show T1: ?T1
apply (simp only: set_eq_iff)
apply (simp only: fun_eq_iff)
apply (simp add: set_to_map_def)
apply auto
done
from T1 show ?T2 by auto
qed
lemma map_to_set_upd[simp]: "map_to_set (m (k ↦ v)) = insert (k, v) (map_to_set m - {(k, v') |v'. True})"
unfolding map_to_set_def
apply (simp add: set_eq_iff)
apply metis
done
lemma set_to_map_insert:
assumes k_nin: "fst kv ∉ fst ` S"
shows "set_to_map (insert kv S) = (set_to_map S) (fst kv ↦ snd kv)"
proof
fix k'
obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust)
from k_nin have k_nin': "⋀v'. (k, v') ∉ S"
by (auto simp add: image_iff Ball_def)
show "set_to_map (insert kv S) k' = ((set_to_map S)(fst kv ↦ snd kv)) k'"
by (simp add: set_to_map_def k_nin')
qed
lemma map_to_set_dom :
"dom m = fst ` (map_to_set m)"
unfolding dom_def map_to_set_def
by (auto simp add: image_iff)
lemma map_to_set_ran :
"ran m = snd ` (map_to_set m)"
unfolding ran_def map_to_set_def
by (auto simp add: image_iff)
lemma set_to_map_dom :
"dom (set_to_map S) = fst ` S"
unfolding set_to_map_def[abs_def] dom_def
by (auto simp add: image_iff Bex_def)
lemma set_to_map_ran :
"ran (set_to_map S) ⊆ snd ` S"
unfolding set_to_map_def[abs_def] ran_def subset_iff
by (auto simp add: image_iff Bex_def)
(metis Eps_Opt_eq_Some)
lemma finite_map_to_set:
"finite (map_to_set m) = finite (dom m)"
unfolding map_to_set_def map_to_set_dom
apply (intro iffI finite_imageI)
apply assumption
apply (rule finite_imageD[of fst])
apply assumption
apply (simp add: inj_on_def)
done
lemma card_map_to_set :
"card (map_to_set m) = card (dom m)"
unfolding map_to_set_def map_to_set_dom
apply (rule card_image[symmetric])
apply (simp add: inj_on_def)
done
lemma map_of_map_to_set :
"distinct (map fst l) ⟹
map_of l = m ⟷ set l = map_to_set m"
proof (induct l arbitrary: m)
case Nil thus ?case by (simp add: map_to_set_empty_iff) blast
next
case (Cons kv l m)
obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust)
from Cons(2) have dist_l: "distinct (map fst l)" and kv'_nin: "⋀v'. (k, v') ∉ set l"
by (auto simp add: image_iff)
note ind_hyp = Cons(1)[OF dist_l]
from kv'_nin have l_eq: "set (kv # l) = map_to_set m ⟷ (set l = map_to_set (m (k := None))) ∧ m k = Some v"
apply (simp add: map_to_set_def restrict_map_def set_eq_iff)
apply (auto)
apply (metis)
apply (metis option.inject)
done
from kv'_nin have m_eq: "map_of (kv # l) = m ⟷ map_of l = (m (k := None)) ∧ m k = Some v"
apply (simp add: fun_eq_iff restrict_map_def map_of_eq_None_iff image_iff Ball_def)
apply metis
done
show ?case
unfolding m_eq l_eq
using ind_hyp[of "m (k := None)"]
by metis
qed
lemma map_to_set_map_of :
"distinct (map fst l) ⟹ map_to_set (map_of l) = set l"
by (metis map_of_map_to_set)
subsubsection ‹Mapping empty set to None›
definition "dflt_None_set S ≡ if S={} then None else Some S"
lemma the_dflt_None_empty[simp]: "dflt_None_set {} = None"
unfolding dflt_None_set_def by simp
lemma the_dflt_None_nonempty[simp]: "S≠{} ⟹ dflt_None_set S = Some S"
unfolding dflt_None_set_def by simp
lemma the_dflt_None_set[simp]: "the_default {} (dflt_None_set x) = x"
unfolding dflt_None_set_def by auto
subsection ‹Orderings›
lemma (in order) min_arg_le[simp]:
"n ≤ min m n ⟷ min m n = n"
"m ≤ min m n ⟷ min m n = m"
by (auto simp: min_def)
lemma (in linorder) min_arg_not_ge[simp]:
"¬ min m n < m ⟷ min m n = m"
"¬ min m n < n ⟷ min m n = n"
by (auto simp: min_def)
lemma (in linorder) min_eq_arg[simp]:
"min m n = m ⟷ m≤n"
"min m n = n ⟷ n≤m"
by (auto simp: min_def)
lemma min_simps[simp]:
"a<(b::'a::order) ⟹ min a b = a"
"b<(a::'a::order) ⟹ min a b = b"
by (auto simp add: min_def dest: less_imp_le)
lemma (in -) min_less_self_conv[simp]:
"min a b < a ⟷ b < (a::_::linorder)"
"min a b < b ⟷ a < (b::_::linorder)"
by (auto simp: min_def)
lemma ord_eq_le_eq_trans: "⟦ a=b; b≤c; c=d ⟧ ⟹ a≤d" by auto
lemma zero_comp_diff_simps[simp]:
"(0::'a::linordered_idom) ≤ a - b ⟷ b ≤ a"
"(0::'a::linordered_idom) < a - b ⟷ b < a"
by auto
subsubsection ‹Termination Measures›
text ‹Lexicographic measure, assuming upper bound for second component›
lemma mlex_fst_decrI:
fixes a a' b b' N :: nat
assumes "a<a'"
assumes "b<N" "b'<N"
shows "a*N + b < a'*N + b'"
proof -
have "a*N + b + 1 ≤ a*N + N" using ‹b<N› by linarith
also have "… ≤ a'*N" using ‹a<a'›
by (metis Suc_leI ab_semigroup_add_class.add.commute
ab_semigroup_mult_class.mult.commute mult_Suc_right mult_le_mono2)
also have "… ≤ a'*N + b'" by auto
finally show ?thesis by auto
qed
lemma mlex_leI:
fixes a a' b b' N :: nat
assumes "a≤a'"
assumes "b≤b'"
shows "a*N + b ≤ a'*N + b'"
using assms
by (auto intro!: add_mono)
lemma mlex_snd_decrI:
fixes a a' b b' N :: nat
assumes "a=a'"
assumes "b<b'"
shows "a*N + b < a'*N + b'"
using assms
by (auto)
lemma mlex_bound:
fixes a b :: nat
assumes "a<A"
assumes "b<N"
shows "a*N + b < A*N"
using assms
using mlex_fst_decrI by fastforce
subsection ‹CCPOs›
context ccpo
begin
lemma ccpo_Sup_mono:
assumes C: "Complete_Partial_Order.chain (≤) A"
"Complete_Partial_Order.chain (≤) B"
assumes B: "∀x∈A. ∃y∈B. x≤y"
shows "Sup A ≤ Sup B"
proof (rule ccpo_Sup_least)
fix x
assume "x∈A"
with B obtain y where I: "y∈B" and L: "x≤y" by blast
note L
also from I ccpo_Sup_upper have "y≤Sup B" by (blast intro: C)
finally show "x≤Sup B" .
qed (rule C)
lemma fixp_mono:
assumes M: "monotone (≤) (≤) f" "monotone (≤) (≤) g"
assumes LE: "⋀Z. f Z ≤ g Z"
shows "ccpo_class.fixp f ≤ ccpo_class.fixp g"
unfolding fixp_def[abs_def]
apply (rule ccpo_Sup_mono)
apply (rule chain_iterates M)+
proof rule
fix x
assume "x∈ccpo_class.iterates f"
thus "∃y∈ccpo_class.iterates g. x≤y"
proof (induct)
case (step x)
then obtain y where I: "y∈ccpo_class.iterates g" and L: "x≤y" by blast
hence "g y ∈ ccpo_class.iterates g" and "f x ≤ g y"
apply -
apply (erule iterates.step)
apply (rule order_trans)
apply (erule monotoneD[OF M(1)])
apply (rule LE)
done
thus "∃y∈ccpo_class.iterates g. f x ≤ y" ..
next
case (Sup M)
define N where "N = {SOME y. y∈ccpo_class.iterates g ∧ x≤y | x. x∈M}"
have N1: "∀y∈N. y∈ccpo_class.iterates g ∧ (∃x∈M. x≤y)"
unfolding N_def
apply auto
apply (metis (lifting) Sup.hyps(2) tfl_some)
by (metis (lifting) Sup.hyps(2) tfl_some)
have N2: "∀x∈M. ∃y∈N. x≤y"
unfolding N_def
apply auto
by (metis (lifting) Sup.hyps(2) tfl_some)
have "N ⊆ ccpo_class.iterates g" using Sup
using N1 by auto
hence C_chain: "Complete_Partial_Order.chain (≤) N"
using chain_iterates[OF M(2)]
unfolding chain_def by auto
have "Sup N ∈ ccpo_class.iterates g" and "Sup M ≤ Sup N"
apply -
apply (rule iterates.Sup[OF C_chain])
using N1 apply blast
apply (rule ccpo_Sup_mono)
apply (rule Sup.hyps)
apply (rule C_chain)
apply (rule N2)
done
thus ?case by blast
qed
qed
end
subsection ‹Code›
text ‹Constant for code-abort. If that gets executed, an abort-exception is raised.›
definition [simp]: "CODE_ABORT f = f ()"
declare [[code abort: CODE_ABORT]]
end