diff --git a/thys/ABY3_Protocols/Finite_Number_Type.thy b/thys/ABY3_Protocols/Finite_Number_Type.thy --- a/thys/ABY3_Protocols/Finite_Number_Type.thy +++ b/thys/ABY3_Protocols/Finite_Number_Type.thy @@ -1,102 +1,102 @@ theory Finite_Number_Type imports "HOL-Library.Cardinality" begin text \ To avoid carrying the modulo all the time, we introduce a new type for integers @{term "{0.. consts L :: nat specification (L) L_gt_1: "L > 1" by auto typedef natL = "{0.. natL" is "\x. (-x) mod int L" by simp lift_definition plus_natL :: "natL \ natL \ natL" is "\x y. (x + y) mod int L" by simp lift_definition minus_natL :: "natL \ natL \ natL" is "\x y. (x - y) mod int L" by simp lift_definition times_natL :: "natL \ natL \ natL" is "\x y. (x * y) mod int L" by simp instance by (standard; (transfer, simp add: algebra_simps mod_simps)) end typ int instantiation natL :: "{distrib_lattice, bounded_lattice, linorder}" begin lift_definition inf_natL :: "natL \ natL \ natL" is inf unfolding inf_int_def by auto lift_definition sup_natL :: "natL \ natL \ natL" is sup unfolding sup_int_def by auto lift_definition less_eq_natL :: "natL \ natL \ bool" is less_eq . lift_definition less_natL :: "natL \ natL \ bool" is less . lift_definition top_natL :: "natL" is "int L-1" using L_gt_1 by simp lift_definition bot_natL :: "natL" is 0 using L_gt_1 by simp instance by (standard; (transfer, simp add: inf_int_def sup_int_def))+ end instantiation natL :: semiring_modulo begin lift_definition divide_natL :: "natL \ natL \ natL" is divide apply (auto simp: div_int_pos_iff) - by (smt div_by_0 div_by_1 zdiv_mono2) + by (smt (verit) div_by_0 div_by_1 zdiv_mono2) lift_definition modulo_natL :: "natL \ natL \ natL" is modulo apply (auto simp: mod_int_pos_iff) - by (smt zmod_le_nonneg_dividend) + by (smt (verit) zmod_le_nonneg_dividend) instance by (standard; (transfer, simp add: mod_simps)) end instance natL :: finite apply standard unfolding type_definition.univ[OF type_definition_natL] by simp lemma natL_card[simp]: "CARD(natL) = L" unfolding type_definition.univ[OF type_definition_natL] apply (subst card_image) subgoal by (meson Abs_natL_inject inj_onI) subgoal by simp done end \ No newline at end of file diff --git a/thys/ABY3_Protocols/Shuffle.thy b/thys/ABY3_Protocols/Shuffle.thy --- a/thys/ABY3_Protocols/Shuffle.thy +++ b/thys/ABY3_Protocols/Shuffle.thy @@ -1,2343 +1,2343 @@ theory Shuffle imports CryptHOL.CryptHOL Additive_Sharing Spmf_Common Sharing_Lemmas begin text \ This is the formalization of the array shuffling protocol defined in \cite{laud2016secure} adapted for the ABY3 sharing scheme. For the moment, we assume an oracle that generates uniformly distributed permutations, instead of instantiating it with e.g. Fischer-Yates algorithm. \ no_notation (ASCII) comp (infixl "o" 55) no_notation m_inv ("inv\ _" [81] 80) no_adhoc_overloading Monad_Syntax.bind bind_pmf fun shuffleF :: "natL sharing list \ natL sharing list spmf" where "shuffleF xsl = spmf_of_set (permutations_of_multiset (mset xsl))" type_synonym zero_sharing = "natL sharing list" type_synonym party2_data = "natL list" type_synonym party01_permutation = "nat \ nat" type_synonym phase_msg = "zero_sharing \ party2_data \ party01_permutation" type_synonym role_msg = "(natL list \ natL list \ natL list) \ party2_data \ (party01_permutation \ party01_permutation)" (* (a, b, c) \ (a, b+c, 0) *) definition aby3_stack_sharing :: "Role \ natL sharing \ natL sharing" where "aby3_stack_sharing r s = make_sharing' r (next_role r) (prev_role r) (get_party r s) (get_party (next_role r) s + get_party (prev_role r) s) 0" (* one permutation step *) definition aby3_do_permute :: "Role \ natL sharing list \ (phase_msg \ natL sharing list) spmf" where "aby3_do_permute r x = (do { let n = length x; \ \ sequence_spmf (replicate n zero_sharing); \ \ spmf_of_set {\. \ permutes {.. y') \; let msg = (\, x2, \); return_spmf (msg, y) })" (* the shuffling protocol, consisting of three shuffling steps *) definition aby3_shuffleR :: "natL sharing list \ (role_msg sharing \ natL sharing list) spmf" where "aby3_shuffleR x = (do { ((\a,x',\a), a) \ aby3_do_permute Party1 x; \ \1st round\ ((\b,a',\b), b) \ aby3_do_permute Party2 a; \ \2nd round\ ((\c,b',\c), c) \ aby3_do_permute Party3 b; \ \3rd round\ let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (msg, c) })" (* the ideal functionality of shuffling *) definition aby3_shuffleF :: "natL sharing list \ natL sharing list spmf" where "aby3_shuffleF x = (do { \ \ spmf_of_set {\. \ permutes {.. = permute_list \ x1; y \ sequence_spmf (map share_nat x\); return_spmf y })" (* the simulator for party 1 *) definition S1 :: "natL list \ natL list \ role_msg spmf" where "S1 x1 yc1 = (do { let n = length x1; \a \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); yb1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); yb2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); \ \round 1\ let \a1 = map2 (-) ya1 (permute_list \a x1); \ \round 2\ let \b1 = yb1; \ \round 3\ let b' = yb2; let \c1 = map2 (-) (yc1) (permute_list \c (map2 (+) yb1 yb2)); \ \non-free message\ let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf msg1 })" (* the simulator for party 2 *) definition S2 :: "natL list \ natL list \ role_msg spmf" where "S2 x2 yc2 = (do { let n = length x2; x3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); yb2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); \ \round 1\ let x' = x3; let \a2 = map2 (-) ya2 (permute_list \a (map2 (+) x2 x3)); \ \round 2\ let \b2 = map2 (-) yb2 (permute_list \b ya2); \ \round 3\ let \c2 = yc2; \ \non-free message\ let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf msg2 })" (* the simulator for party 3 *) definition S3 :: "natL list \ natL list \ role_msg spmf" where "S3 x3 yc3 = (do { let n = length x3; \b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); ya1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); yb3::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); \ \round 1\ let \a3 = ya3; \ \round 2\ let a' = ya1; let \b3 = map2 (-) yb3 (permute_list \b (map2 (+) ya3 ya1)); \ \round 3\ let \c3 = map2 (-) yc3 (permute_list \c yb3); \ \non-free message\ let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf msg3 })" definition S :: "Role \ natL list \ natL list \ role_msg spmf" where "S r = get_party r (make_sharing S1 S2 S3)" definition is_uniform_sharing_list :: "natL sharing list spmf \ bool" where "is_uniform_sharing_list xss \ (\xs. xss = bind_spmf xs (sequence_spmf \ map share_nat))" lemma case_prod_nesting_same: "case_prod (\a b. f (case_prod g x) a b ) x = case_prod (\a b. f (g a b) a b ) x" by (cases x) simp lemma zip_map_map_same: "map (\x. (f x, g x)) xs = zip (map f xs) (map g xs)" unfolding zip_map_map unfolding zip_same_conv_map by simp lemma dup_map_eq: "length xs = length ys \ (xs, map2 f ys xs) = (\xys. (map fst xys, map snd xys)) (map2 (\x y. (y, f x y)) ys xs)" by (auto simp: map_snd_zip[unfolded snd_def]) abbreviation "map2_spmf f xs ys \ map_spmf (case_prod f) (pair_spmf xs ys)" abbreviation "map3_spmf f xs ys zs \ map2_spmf (\a. case_prod (f a)) xs (pair_spmf ys zs)" lemma map_spmf_cong2: assumes "p = map_spmf m q" "\x. x\set_spmf q \ f (m x) = g x" shows "map_spmf f p = map_spmf g q" using assms by (simp add: spmf.map_comp cong: map_spmf_cong) lemma bind_spmf_cong2: assumes "p = map_spmf m q" "\x. x\set_spmf q \ f (m x) = g x" shows "bind_spmf p f = bind_spmf q g" using assms by (simp add: map_spmf_conv_bind_spmf cong: bind_spmf_cong) lemma map2_spmf_map2_sequence: "length xss = length yss \ map2_spmf (map2 f) (sequence_spmf xss) (sequence_spmf yss) = sequence_spmf (map2 (map2_spmf f) xss yss)" apply (induction xss yss rule: list_induct2) subgoal by simp subgoal premises IH for x xs y ys apply simp apply (fold IH) apply (unfold pair_map_spmf) apply (unfold spmf.map_comp) apply (rule map_spmf_cong2[where m="\((x,y),(xs,ys)). ((x,xs),(y,ys))"]) subgoal unfolding pair_spmf_alt_def apply (simp add: map_spmf_conv_bind_spmf) apply (subst bind_commute_spmf[where q=y]) .. subgoal by auto done done abbreviation map3 :: "('a \ 'b \ 'c \ 'd) \ 'a list \ 'b list \ 'c list \ 'd list" where "map3 f a b c \ map2 (\a (b,c). f a b c) a (zip b c)" lemma map3_spmf_map3_sequence: "length xss = length yss \ length yss = length zss \ map3_spmf (map3 f) (sequence_spmf xss) (sequence_spmf yss) (sequence_spmf zss) = sequence_spmf (map3 (map3_spmf f) xss yss zss)" apply (induction xss yss zss rule: list_induct3) subgoal by simp subgoal premises IH for x xs y ys z zs apply simp apply (fold IH) apply (unfold pair_map_spmf) apply (unfold spmf.map_comp) apply (rule map_spmf_cong2[where m="\((x,y,z),(xs,ys,zs)). ((x,xs),(y,ys),(z,zs))"]) subgoal unfolding pair_spmf_alt_def apply (simp add: map_spmf_conv_bind_spmf) apply (subst bind_commute_spmf[where q=y]) apply (subst bind_commute_spmf[where q=z]) apply (subst bind_commute_spmf[where q=z]) .. subgoal by auto done done lemma in_pairD2: "x \ A \ B \ snd x \ B" by auto lemma list_map_cong2: "x = map m y \ (\z. z\set y =simp=> f (m z) = g z) \ map f x = map g y" unfolding simp_implies_def by simp lemma map_swap_zip: "map prod.swap (zip xs ys) = zip ys xs" apply (induction xs arbitrary: ys) subgoal by simp subgoal for x xs ys by (cases ys) auto done lemma inv_zero_sharing_sequence: "n = length x \ map_spmf (\\s. (\s, map2 (map_sharing2 (+)) x \s)) (sequence_spmf (replicate n zero_sharing)) = map_spmf (\ys. (map2 (map_sharing2 (-)) ys x, ys)) (sequence_spmf (map (share_nat \reconstruct) x))" proof - assume n: "n = length x" have "map_spmf (\\s. (\s, map2 (map_sharing2 (+)) x \s)) (sequence_spmf (replicate n zero_sharing)) = map2_spmf (\\s x. (\s, map2 (map_sharing2 (+)) x \s)) (sequence_spmf (replicate n zero_sharing)) (sequence_spmf (map return_spmf x))" unfolding sequence_map_return_spmf apply (rule map_spmf_cong2[where m="fst"]) subgoal by simp subgoal by (auto simp: case_prod_unfold dest: in_pairD2) done also have "... = map_spmf (\\xs. (map fst \xs, map snd \xs)) (map2_spmf (map2 (\\ x. (\, map_sharing2 (+) x \))) (sequence_spmf (replicate n zero_sharing)) (sequence_spmf (map return_spmf x)))" apply (unfold spmf.map_comp) apply (rule map_spmf_cong[OF refl]) using n by (auto simp: case_prod_unfold comp_def set_sequence_spmf list_all2_iff map_swap_zip intro: list_map_cong2[where m=prod.swap]) also have "... = map_spmf (\\xs. (map fst \xs, map snd \xs)) (map2_spmf (map2 (\y x. (map_sharing2 (-) y x, y))) (sequence_spmf (map (share_nat \ reconstruct) x)) (sequence_spmf (map return_spmf x)))" apply (rule arg_cong[where f="map_spmf _"]) using n apply (simp add: map2_spmf_map2_sequence) apply (rule arg_cong[where f=sequence_spmf]) apply (unfold list_eq_iff_nth_eq) apply safe subgoal by simp apply (simp add: ) apply (subst map_spmf_cong2[where p="pair_spmf _ (return_spmf _)"]) apply (rule pair_spmf_return_spmf2) apply simp apply (subst map_spmf_cong2[where p="pair_spmf _ (return_spmf _)"]) apply (rule pair_spmf_return_spmf2) apply simp using inv_zero_sharing . also have "... = map2_spmf (\ys x. (map2 (map_sharing2 (-)) ys x, ys)) (sequence_spmf (map (share_nat \reconstruct) x)) (sequence_spmf (map return_spmf x))" apply (unfold spmf.map_comp) apply (rule map_spmf_cong[OF refl]) using n by (auto simp: case_prod_unfold comp_def set_sequence_spmf list_all2_iff map_swap_zip intro: list_map_cong2[where m=prod.swap]) also have "... = map_spmf (\ys. (map2 (map_sharing2 (-)) ys x, ys)) (sequence_spmf (map (share_nat \reconstruct) x))" unfolding sequence_map_return_spmf apply (rule map_spmf_cong2[where m="fst", symmetric]) subgoal by simp subgoal by (auto simp: case_prod_unfold dest: in_pairD2) done finally show ?thesis . qed lemma get_party_map_sharing2: "get_party p \ (case_prod (map_sharing2 f)) = case_prod f \ map_prod (get_party p) (get_party p)" by auto lemma map_map_prod_zip: "map (map_prod f g) (zip xs ys) = zip (map f xs) (map g ys)" by (simp add: map_prod_def zip_map_map) lemma map_map_prod_zip': "map (h \ map_prod f g) (zip xs ys) = map h (zip (map f xs) (map g ys))" by (simp add: map_prod_def zip_map_map) lemma eq_map_spmf_conv: assumes "\x. f (f' x) = x" "f' = inv f" "map_spmf f' x = y" shows "x = map_spmf f y" proof - have surj: "surj f" apply (rule surjI) using assms(1) . have "map_spmf f (map_spmf f' x) = map_spmf f y" unfolding assms(3) .. thus ?thesis using assms(1) by (simp add: spmf.map_comp surj_iff comp_def) qed lemma lift_map_spmf_pairs: "map2_spmf f = F \ map_spmf (case_prod f) (pair_spmf A B) = F A B" by auto lemma measure_pair_spmf_times': "C = A \ B \ measure (measure_spmf (pair_spmf p q)) C = measure (measure_spmf p) A * measure (measure_spmf q) B" by (simp add: measure_pair_spmf_times) lemma map_spmf_pairs_tmp: "map_spmf (\(a,b,c,d,e,f,g). (a,e,b,f,c,g,d)) (pair_spmf A (pair_spmf B (pair_spmf C (pair_spmf D (pair_spmf E (pair_spmf F G)))))) = (pair_spmf A (pair_spmf E (pair_spmf B (pair_spmf F (pair_spmf C (pair_spmf G D))))))" apply (rule spmf_eqI) apply (clarsimp simp add: spmf_map) subgoal for a e b f c g d apply (subst measure_pair_spmf_times'[where A="{a}"]) defer apply (subst measure_pair_spmf_times'[where A="{b}"]) defer apply (subst measure_pair_spmf_times'[where A="{c}"]) defer apply (subst measure_pair_spmf_times'[where A="{d}"]) defer apply (subst measure_pair_spmf_times'[where A="{e}"]) defer apply (subst measure_pair_spmf_times'[where A="{f}" and B="{g}"]) defer apply (auto simp: spmf_conv_measure_spmf) done done lemma case_case_prods_tmp: "(case case x of (a, b, c, d, e, f, g) \ (a, e, b, f, c, g, d) of (ya, yb, yc, yd, ye, yf, yg) \ F ya yb yc yd ye yf yg) = (case x of (a,b,c,d,e,f,g) \ F a e b f c g d)" by (cases x) simp lemma bind_spmf_permutes_cong: "(\\. \ permutes {..<(x::nat)} \ f \ = g \) \ bind_spmf (spmf_of_set {\. \ permutes {... \ permutes {.. list_all2 (\x y. f x = g y) xs ys" apply (induction xs arbitrary: ys) subgoal by auto subgoal for x xs ys by (cases ys) (auto) done lemma bind_spmf_sequence_map_share_nat_cong: "(\l. map reconstruct l = x \ f l = g l) \ bind_spmf (sequence_spmf (map share_nat x)) f = bind_spmf (sequence_spmf (map share_nat x)) g" subgoal premises prems apply (rule bind_spmf_cong[OF refl]) apply (rule prems) unfolding set_sequence_spmf mem_Collect_eq apply (simp add: map_eq_iff_list_all2[where g=id, simplified]) apply (simp add: list_all2_map2) apply (erule list_all2_mono) unfolding share_nat_def by simp done lemma map_reconstruct_comp_eq_iff: "(\x. x\set xs \ reconstruct (f x) = reconstruct x) \ map (reconstruct \ f) xs = map reconstruct xs" by (induction xs) auto lemma permute_list_replicate: "p permutes {.. permute_list p (replicate n x) = replicate n x" apply (fold map_replicate_const[where lst="[0.. (\y::natL. y\set ys \ y = 0) \ map2 (-) xs ys = xs" by (induction xs ys rule: list_induct2) auto lemma permute_comp_left_inj: "p permutes {.. inj (\p'. p \ p')" by (rule fun.inj_map) (rule permutes_inj_on) lemma permute_comp_left_inj_on: "p permutes {.. inj_on (\p'. p \ p') A" using permute_comp_left_inj inj_on_subset by blast lemma permute_comp_right_inj: "p permutes {.. inj (\p'. p' \ p)" using inj_onI comp_id o_assoc permutes_surj surj_iff - by (smt) + by (smt (verit)) lemma permute_comp_right_inj_on: "p permutes {.. inj_on (\p'. p' \ p) A" using permute_comp_right_inj inj_on_subset by blast lemma permutes_inv_comp_left: "p permutes {.. inv (\p'. p \ p') = (\p'. inv p \ p')" by (rule inv_unique_comp; rule ext, simp add: permutes_inv_o comp_assoc[symmetric]) lemma permutes_inv_comp_right: "p permutes {.. inv (\p'. p' \ p) = (\p'. p' \ inv p)" by (rule inv_unique_comp; rule ext, simp add: permutes_inv_o comp_assoc) lemma permutes_inv_comp_left_right: "\a permutes {.. \b permutes {.. inv (\p'. \a \ p' \ \b) = (\p'. inv \a \ p' \ inv \b)" by (rule inv_unique_comp; rule ext, simp add: permutes_inv_o comp_assoc, simp add: permutes_inv_o comp_assoc[symmetric]) lemma permutes_inv_comp_left_left: "\a permutes {.. \b permutes {.. inv (\p'. \a \ \b \ p') = (\p'. inv \b \ inv \a \ p')" by (rule inv_unique_comp; rule ext, simp add: permutes_inv_o comp_assoc, simp add: permutes_inv_o comp_assoc[symmetric]) lemma permutes_inv_comp_right_right: "\a permutes {.. \b permutes {.. inv (\p'. p' \ \a \ \b) = (\p'. p' \ inv \b \ inv \a)" by (rule inv_unique_comp; rule ext, simp add: permutes_inv_o comp_assoc, simp add: permutes_inv_o comp_assoc[symmetric]) lemma image_compose_permutations_left_right: fixes S assumes "\a permutes S" "\b permutes S" shows "{\a \ \ \ \b |\. \ permutes S} = {\. \ permutes S}" proof - have *: "(\\. \a \ \ \ \b) = (\\'. \a \ \') \ (\\. \ \ \b)" by (simp add: comp_def) then show ?thesis apply (fold image_Collect) apply (unfold *) apply (fold image_comp) apply (subst image_Collect) apply (unfold image_compose_permutations_right[OF assms(2)]) apply (subst image_Collect) apply (unfold image_compose_permutations_left[OF assms(1)]) .. qed lemma image_compose_permutations_left_left: fixes S assumes "\a permutes S" "\b permutes S" shows "{\a \ \b \ \ |\. \ permutes S} = {\. \ permutes S}" using image_compose_permutations_left image_compose_permutations_right proof - have *: "(\\. \a \ \b \ \) = (\\'. \a \ \') \ (\\. \b \ \)" by (simp add: comp_def) show ?thesis apply (fold image_Collect) apply (unfold *) apply (fold image_comp) apply (subst image_Collect) apply (unfold image_compose_permutations_left[OF assms(2)]) apply (subst image_Collect) apply (unfold image_compose_permutations_left[OF assms(1)]) .. qed lemma image_compose_permutations_right_right: fixes S assumes "\a permutes S" "\b permutes S" shows "{\ \ \a \ \b |\. \ permutes S} = {\. \ permutes S}" using image_compose_permutations_left image_compose_permutations_right proof - have *: "(\\. \ \ \a \ \b) = (\\. \ \ \b) \ (\\'. \' \ \a)" by (simp add: comp_def) show ?thesis apply (fold image_Collect) apply (unfold *) apply (fold image_comp) apply (subst image_Collect) apply (unfold image_compose_permutations_right[OF assms(1)]) apply (subst image_Collect) apply (unfold image_compose_permutations_right[OF assms(2)]) .. qed lemma random_perm_middle: defines "random_perm n \ spmf_of_set {\. \ permutes {..(\a,\b,\c). ((\a,\b,\c),\a \ \b \ \c)) (pair_spmf (random_perm n) (pair_spmf (random_perm n) (random_perm n))) = map_spmf (\(\,\a,\c). ((\a,inv \a \ \ \ inv \c,\c),\)) (pair_spmf (random_perm n) (pair_spmf (random_perm n) (random_perm n)))" (is "?lhs = ?rhs") proof - have "?lhs = (do {\a \ random_perm n; \c \ random_perm n; (\b,p) \ map_spmf (\\b. (\b,\a \ \b \ \c)) (random_perm n); return_spmf ((\a, \b, \c), p)})" unfolding map_spmf_conv_bind_spmf pair_spmf_alt_def apply simp apply (subst (4) bind_commute_spmf) .. also have "\ = (do {\a \ random_perm n; \c \ random_perm n; map_spmf (\p. ((\a, inv \a \ p \ inv \c,\c),p)) (random_perm n)})" unfolding random_perm_def supply [intro!] = bind_spmf_permutes_cong apply rule+ apply (subst inv_uniform) subgoal for \a \c apply (rule inj_compose[unfolded comp_def, where f="\p. p \ \c"]) subgoal by (rule permute_comp_right_inj_on) subgoal by (rule permute_comp_left_inj_on) done apply (simp add: permutes_inv_comp_left_right map_spmf_conv_bind_spmf image_Collect image_compose_permutations_left_right) done also have "\ = ?rhs" unfolding map_spmf_conv_bind_spmf pair_spmf_alt_def apply (subst (2) bind_commute_spmf) apply (subst (1) bind_commute_spmf) apply simp done finally show ?thesis . qed lemma random_perm_right: defines "random_perm n \ spmf_of_set {\. \ permutes {..(\a,\b,\c). ((\a,\b,\c),\a \ \b \ \c)) (pair_spmf (random_perm n) (pair_spmf (random_perm n) (random_perm n))) = map_spmf (\(\,\a,\b). ((\a,\b,inv \b \ inv \a \ \),\)) (pair_spmf (random_perm n) (pair_spmf (random_perm n) (random_perm n)))" (is "?lhs = ?rhs") proof - have "?lhs = (do {\a \ random_perm n; \b \ random_perm n; (\c,\) \ map_spmf (\\c. (\c,\a \ \b \ \c)) (random_perm n); return_spmf ((\a, \b, \c), \)})" unfolding map_spmf_conv_bind_spmf pair_spmf_alt_def by simp also have "\ = (do {\a \ random_perm n; \b \ random_perm n; map_spmf (\\. ((\a, \b, inv \b \ inv \a \ \),\)) (random_perm n)})" unfolding random_perm_def supply [intro!] = bind_spmf_permutes_cong apply rule+ apply (subst inv_uniform) subgoal apply (rule permute_comp_left_inj_on) using permutes_compose . apply (simp add: permutes_inv_comp_left_left map_spmf_conv_bind_spmf image_Collect image_compose_permutations_left_left) done also have "\ = ?rhs" unfolding map_spmf_conv_bind_spmf pair_spmf_alt_def apply (subst (2) bind_commute_spmf) apply (subst (1) bind_commute_spmf) apply simp done finally show ?thesis . qed lemma random_perm_left: defines "random_perm n \ spmf_of_set {\. \ permutes {..(\a,\b,\c). ((\a,\b,\c),\a \ \b \ \c)) (pair_spmf (random_perm n) (pair_spmf (random_perm n) (random_perm n))) = map_spmf (\(\,\b,\c). ((\ \ inv \c \ inv \b,\b,\c),\)) (pair_spmf (random_perm n) (pair_spmf (random_perm n) (random_perm n)))" (is "?lhs = ?rhs") proof - have "?lhs = (do {\b \ random_perm n; \c \ random_perm n; (\a,\) \ map_spmf (\\a. (\a,\a \ \b \ \c)) (random_perm n); return_spmf ((\a, \b, \c), \)})" unfolding map_spmf_conv_bind_spmf pair_spmf_alt_def apply simp apply (subst (4) bind_commute_spmf) apply (subst (3) bind_commute_spmf) .. also have "\ = (do {\b \ random_perm n; \c \ random_perm n; map_spmf (\\. ((\ \ inv \c \ inv \b, \b, \c),\)) (random_perm n)})" unfolding random_perm_def supply [intro!] = bind_spmf_permutes_cong apply rule+ apply (subst inv_uniform) subgoal apply (unfold comp_assoc) apply (rule permute_comp_right_inj_on) using permutes_compose . apply (simp add: permutes_inv_comp_right_right map_spmf_conv_bind_spmf image_Collect image_compose_permutations_right_right) done also have "\ = ?rhs" unfolding map_spmf_conv_bind_spmf pair_spmf_alt_def apply (subst (2) bind_commute_spmf) apply (subst (1) bind_commute_spmf) apply simp done finally show ?thesis . qed lemma case_prod_return_spmf: "case_prod (\a b. return_spmf (f a b)) = (\x. return_spmf (case_prod f x))" by auto lemma sequence_share_nat_calc': assumes "r1\r2" "r2\r3" "r3\r1" shows "sequence_spmf (map share_nat xs) = (do { let n = length xs; let random_seq = sequence_spmf (replicate n (spmf_of_set UNIV)); (dp, dpn) \ (pair_spmf random_seq random_seq); return_spmf (map3 (\x a b. make_sharing' r1 r2 r3 a b (x - (a + b))) xs dp dpn) })" (is "_ = ?rhs") proof - have "sequence_spmf (map share_nat xs) = (do { let n = length xs; let random_seq = sequence_spmf (replicate n (spmf_of_set UNIV)); (xs, dp, dpn) \ pair_spmf (sequence_spmf (map return_spmf xs)) (pair_spmf random_seq random_seq); return_spmf (map3 (\x a b. make_sharing' r1 r2 r3 a b (x - (a + b))) xs dp dpn) })" apply (unfold Let_def) apply (unfold case_prod_return_spmf) apply (fold map_spmf_conv_bind_spmf) apply (subst map3_spmf_map3_sequence) subgoal by simp subgoal by simp apply (rule arg_cong[where f=sequence_spmf]) apply (unfold map_eq_iff_list_all2) apply (rule list_all2_all_nthI) subgoal by simp unfolding share_nat_def_calc'[OF assms] apply (auto simp: map_spmf_conv_bind_spmf pair_spmf_alt_def) done also have "\ = ?rhs" by (auto simp: pair_spmf_alt_def sequence_map_return_spmf) finally show ?thesis . qed lemma reconstruct_stack_sharing_eq_reconstruct: "reconstruct \ aby3_stack_sharing r = reconstruct" unfolding aby3_stack_sharing_def reconstruct_def by (cases r) (auto simp: make_sharing'_sel) lemma map2_ignore1: "length xs = length ys \ map2 (\_. f) xs ys = map f ys" apply (unfold map_eq_iff_list_all2) apply (rule list_all2_all_nthI) by auto lemma map2_ignore2: "length xs = length ys \ map2 (\a b. f a) xs ys = map f xs" apply (unfold map_eq_iff_list_all2) apply (rule list_all2_all_nthI) by auto lemma map_sequence_share_nat_reconstruct: "map_spmf (\x. (x, map reconstruct x)) (sequence_spmf (map share_nat y)) = map_spmf (\x. (x, y)) (sequence_spmf (map share_nat y))" apply (unfold map_spmf_conv_bind_spmf) apply (rule bind_spmf_cong[OF refl]) apply (auto simp: set_sequence_spmf list_eq_iff_nth_eq list_all2_conv_all_nth share_nat_def) done (* the main theorem of security of the shuffling protocol *) theorem shuffle_secrecy: assumes "is_uniform_sharing_list x_dist" shows "(do { x \ x_dist; (msg, y) \ aby3_shuffleR x; return_spmf (map (get_party r) x, get_party r msg, y) }) = (do { x \ x_dist; y \ aby3_shuffleF x; let xr = map (get_party r) x; let yr = map (get_party r) y; msg \ S r xr yr; return_spmf (xr, msg, y) })" (is "?lhs = ?rhs") proof - obtain xs where xs: "x_dist = xs \ (sequence_spmf \ map share_nat)" using assms unfolding is_uniform_sharing_list_def by auto have left_unfolded: "(do { x \ x_dist; (msg, y) \ aby3_shuffleR x; return_spmf (map (get_party r) x, get_party r msg, y)}) = (do { xs \ xs; x \ sequence_spmf (map share_nat xs); \ \round 1\ let n = length x; \a \ sequence_spmf (replicate n zero_sharing); \a \ spmf_of_set {\. \ permutes {..a y') \a; \ \round 2\ let n = length a; \b \ sequence_spmf (replicate n zero_sharing); \b \ spmf_of_set {\. \ permutes {..b y') \b; \ \round 3\ let n = length b; \c \ sequence_spmf (replicate n zero_sharing); \c \ spmf_of_set {\. \ permutes {..c y') \c; let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (map (get_party r) x, get_party r msg, c) })" unfolding xs aby3_shuffleR_def aby3_do_permute_def by (auto simp: case_prod_unfold Let_def) also have clarify_length: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \round 1\ \a \ sequence_spmf (replicate n zero_sharing); \a \ spmf_of_set {\. \ permutes {..a y') \a; \ \round 2\ \b \ sequence_spmf (replicate n zero_sharing); \b \ spmf_of_set {\. \ permutes {..b y') \b; \ \round 3\ \c \ sequence_spmf (replicate n zero_sharing); \c \ spmf_of_set {\. \ permutes {..c y') \c; let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (map (get_party r) x, get_party r msg, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong by (auto simp: Let_def) also have hoist_permutations: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. \round 1\ let x' = map (get_party (prev_role Party1)) x; let y' = map (aby3_stack_sharing Party1) x; \a \ sequence_spmf (replicate n zero_sharing); let a = map2 (map_sharing2 (+)) (permute_list \a y') \a; \ \round 2\ let a' = map (get_party (prev_role Party2)) a; let y' = map (aby3_stack_sharing Party2) a; \b \ sequence_spmf (replicate n zero_sharing); let b = map2 (map_sharing2 (+)) (permute_list \b y') \b; \ \round 3\ let b' = map (get_party (prev_role Party3)) b; let y' = map (aby3_stack_sharing Party3) b; \c \ sequence_spmf (replicate n zero_sharing); let c = map2 (map_sharing2 (+)) (permute_list \c y') \c; let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (map (get_party r) x, get_party r msg, c) })" apply (simp add: Let_def) apply (subst (1) bind_commute_spmf[where q="spmf_of_set _"]) apply (subst (2) bind_commute_spmf[where q="spmf_of_set _"]) apply (subst (2) bind_commute_spmf[where q="spmf_of_set _"]) apply (subst (3) bind_commute_spmf[where q="spmf_of_set _"]) apply (subst (3) bind_commute_spmf[where q="spmf_of_set _"]) apply (subst (3) bind_commute_spmf[where q="spmf_of_set _"]) by simp also have hoist_permutations: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. \round 1\ let x' = map (get_party (prev_role Party1)) x; let y' = map (aby3_stack_sharing Party1) x; a \ sequence_spmf (map (share_nat \ reconstruct) (permute_list \a y')); let \a = map2 (map_sharing2 (-)) a (permute_list \a y'); \ \round 2\ let a' = map (get_party (prev_role Party2)) a; let y' = map (aby3_stack_sharing Party2) a; b \ sequence_spmf (map (share_nat \ reconstruct) (permute_list \b y')); let \b = map2 (map_sharing2 (-)) b (permute_list \b y'); \ \round 3\ let b' = map (get_party (prev_role Party3)) b; let y' = map (aby3_stack_sharing Party3) b; c \ sequence_spmf (map (share_nat \ reconstruct) (permute_list \c y')); let \c = map2 (map_sharing2 (-)) c (permute_list \c y'); let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (map (get_party r) x, get_party r msg, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst hoist_map_spmf[where s="sequence_spmf (replicate _ _)" and f = "map2 (map_sharing2 (+)) _"]) apply (subst hoist_map_spmf'[where s="sequence_spmf (map _ _)" and f = "\ys. map2 (map_sharing2 (-)) ys _"]) apply (subst inv_zero_sharing_sequence) subgoal by simp apply (unfold map_spmf_conv_bind_spmf) apply (unfold bind_spmf_assoc) apply (unfold return_bind_spmf) apply rule+ apply (subst (1 12) Let_def) apply rule+ apply (subst hoist_map_spmf[where s="sequence_spmf (replicate _ _)" and f = "map2 (map_sharing2 (+)) _"]) apply (subst hoist_map_spmf'[where s="sequence_spmf (map _ _)" and f = "\ys. map2 (map_sharing2 (-)) ys _"]) apply (subst inv_zero_sharing_sequence) subgoal by simp apply (unfold map_spmf_conv_bind_spmf) apply (unfold bind_spmf_assoc) apply (unfold return_bind_spmf) apply rule+ apply (subst (1 9) Let_def) apply rule+ apply (subst hoist_map_spmf[where s="sequence_spmf (replicate _ _)" and f = "map2 (map_sharing2 (+)) _"]) apply (subst hoist_map_spmf'[where s="sequence_spmf (map _ _)" and f = "\ys. map2 (map_sharing2 (-)) ys _"]) apply (subst inv_zero_sharing_sequence) subgoal by simp apply (unfold map_spmf_conv_bind_spmf) apply (unfold bind_spmf_assoc) apply (unfold return_bind_spmf) apply rule+ apply (subst (1 6) Let_def) apply rule+ done also have reconstruct: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. \round 1\ let x' = map (get_party (prev_role Party1)) x; let y' = map (aby3_stack_sharing Party1) x; a \ sequence_spmf (map share_nat (permute_list \a xs)); let \a = map2 (map_sharing2 (-)) a (permute_list \a y'); \ \round 2\ let a' = map (get_party (prev_role Party2)) a; let y' = map (aby3_stack_sharing Party2) a; b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); let \b = map2 (map_sharing2 (-)) b (permute_list \b y'); \ \round 3\ let b' = map (get_party (prev_role Party3)) b; let y' = map (aby3_stack_sharing Party3) b; c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); let \c = map2 (map_sharing2 (-)) c (permute_list \c y'); let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (map (get_party r) x, get_party r msg, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong bind_spmf_sequence_map_share_nat_cong apply rule+ apply (subst list.map_comp[symmetric]) apply (rule bind_spmf_cong) subgoal by (auto simp:permute_list_map[symmetric] map_reconstruct_comp_eq_iff reconstruct_def set_sequence_spmf[unfolded list_all2_iff] make_sharing'_sel reconstruct_stack_sharing_eq_reconstruct comp_assoc) apply rule+ apply (subst list.map_comp[symmetric]) apply (rule bind_spmf_cong) subgoal for x l xa \a \b \c xb xc xd xe xf xg apply (subst permute_list_map[symmetric] ) subgoal by (auto simp add: set_sequence_spmf[unfolded list_all2_iff]) apply simp apply (subst map_reconstruct_comp_eq_iff) subgoal by (simp add: reconstruct_def make_sharing'_sel aby3_stack_sharing_def) unfolding set_sequence_spmf mem_Collect_eq unfolding list_all2_map2 apply (subst map_eq_iff_list_all2[where f=reconstruct and g=id and xs=xd and ys="permute_list \a x", simplified, THEN iffD2]) subgoal by (erule list_all2_mono) (simp add: share_nat_def) apply (subst permute_list_compose) subgoal by auto .. apply rule+ apply (subst list.map_comp[symmetric]) apply (rule bind_spmf_cong) subgoal for x l xa \a \b \c xb xc xd xe xf xg xh xi xj xk apply (subst permute_list_map[symmetric] ) subgoal by (auto simp add: set_sequence_spmf[unfolded list_all2_iff]) apply simp apply (subst map_reconstruct_comp_eq_iff) subgoal by (simp add: reconstruct_def make_sharing'_sel aby3_stack_sharing_def) unfolding set_sequence_spmf mem_Collect_eq unfolding list_all2_map2 apply (subst map_eq_iff_list_all2[where f=reconstruct and g=id and xs=xh and ys="permute_list (\a \ \b) x", simplified, THEN iffD2]) subgoal by (erule list_all2_mono) (simp add: share_nat_def) apply (subst permute_list_compose[symmetric]) subgoal by auto .. .. also have hoist: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party (prev_role Party1)) x; let y' = map (aby3_stack_sharing Party1) x; let \a = map2 (map_sharing2 (-)) a (permute_list \a y'); \ \round 2\ let a' = map (get_party (prev_role Party2)) a; let y' = map (aby3_stack_sharing Party2) a; let \b = map2 (map_sharing2 (-)) b (permute_list \b y'); \ \round 3\ let b' = map (get_party (prev_role Party3)) b; let y' = map (aby3_stack_sharing Party3) b; let \c = map2 (map_sharing2 (-)) c (permute_list \c y'); let msg1 = ((map (get_party Party1) \a, map (get_party Party1) \b, map (get_party Party1) \c), b', \a, \c); let msg2 = ((map (get_party Party2) \a, map (get_party Party2) \b, map (get_party Party2) \c), x', \a, \b); let msg3 = ((map (get_party Party3) \a, map (get_party Party3) \b, map (get_party Party3) \c), a', \b, \c); let msg = make_sharing msg1 msg2 msg3; return_spmf (map (get_party r) x, get_party r msg, c) })" unfolding Let_def .. finally have hoisted_save: "?lhs = \" . let ?hoisted = \ { assume r: "r = Party1" have project_to_Party1: "?hoisted = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let y' = map (aby3_stack_sharing Party1) x; let \a1 = map (case_prod (-)) (zip (map (get_party Party1) a) (map (get_party Party1) (permute_list \a y'))); \ \round 2\ let a' = map (get_party Party1) a; let y' = map (aby3_stack_sharing Party2) a; let \b1 = map (case_prod (-)) (zip (map (get_party Party1) b) (map (get_party Party1) (permute_list \b y'))); \ \round 3\ let b' = map (get_party Party2) b; let y' = map (aby3_stack_sharing Party3) b; let \c1 = map (case_prod (-)) (zip (map (get_party Party1) c) (map (get_party Party1) (permute_list \c y'))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" by (simp add: r Let_def get_party_map_sharing2 map_map_prod_zip') also have project_to_Party1: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let y' = map (aby3_stack_sharing Party1) x; let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) y')); \ \round 2\ let a' = map (get_party Party1) a; let y' = map (aby3_stack_sharing Party2) a; let \b1 = map2 (-) (map (get_party Party1) b) (permute_list \b (map (get_party Party1) y')); \ \round 3\ let b' = map (get_party Party2) b; let y' = map (aby3_stack_sharing Party3) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map (get_party Party1) y')); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp .. also have reduce_Lets: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = map2 (-) (map (get_party Party1) b) (permute_list \b (replicate (length a) 0)); \ \round 3\ let b' = map (get_party Party2) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) (map (get_party Party1) b) (map (get_party Party2) b))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" unfolding Let_def unfolding aby3_stack_sharing_def by (simp add: comp_def make_sharing'_sel map_replicate_const zip_map_map_same[symmetric]) also have simplify_minus_zero: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = (map (get_party Party1) b); \ \round 3\ let b' = map (get_party Party2) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) (map (get_party Party1) b) (map (get_party Party2) b))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst permute_list_replicate) subgoal by simp apply (subst map2_minus_zero) subgoal by simp subgoal by simp .. also have break_perms_1: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); (\a,\b,\c) \ pair_spmf (spmf_of_set {\. \ permutes {... \ permutes {... \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = (map (get_party Party1) b); \ \round 3\ let b' = map (get_party Party2) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) (map (get_party Party1) b) (map (get_party Party2) b))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" unfolding pair_spmf_alt_def by simp also have break_perms_2: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); ((\a,\b,\c),\) \ map_spmf (\(\a,\b,\c). ((\a,\b,\c), \a \ \b \ \c)) (pair_spmf (spmf_of_set {\. \ permutes {... \ permutes {... \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = (map (get_party Party1) b); \ \round 3\ let b' = map (get_party Party2) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) (map (get_party Party1) b) (map (get_party Party2) b))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" unfolding pair_spmf_alt_def map_spmf_conv_bind_spmf by simp also have break_perms_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..a \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {..b = inv \a \ \ \ inv \c; a \ sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = (map (get_party Party1) b); \ \round 3\ let b' = map (get_party Party2) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) (map (get_party Party1) b) (map (get_party Party2) b))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" apply (unfold random_perm_middle) apply (unfold map_spmf_conv_bind_spmf pair_spmf_alt_def) by simp also have break_seqs_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..a \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {..b = inv \a \ \ \ inv \c; a1 \ sequence_spmf (replicate n (spmf_of_set UNIV)); a2 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let a = map3 (\a b c. make_sharing b c (a - (b + c))) (permute_list \a xs) a1 a2; b1 \ sequence_spmf (replicate n (spmf_of_set UNIV)); b2 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let b = map3 (\a b c. make_sharing b c (a - (b + c))) (permute_list (\a \ \b) xs) b1 b2; c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a1 = map2 (-) (map (get_party Party1) a) (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = (map (get_party Party1) b); \ \round 3\ let b' = map (get_party Party2) b; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) (map (get_party Party1) b) (map (get_party Party2) b))); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" apply (unfold sequence_share_nat_calc'[of Party1 Party2 Party3, simplified]) apply (simp add: pair_spmf_alt_def Let_def) done also have break_seqs_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..a \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); a2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); b1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); b2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a1 = map2 (-) a1 (permute_list \a (map (get_party Party1) x)); \ \round 2\ let \b1 = b1; \ \round 3\ let b' = b2; let \c1 = map2 (-) (map (get_party Party1) c) (permute_list \c (map2 (+) b1 b2)); let msg1 = ((\a1, \b1, \c1), b', \a, \c); return_spmf (map (get_party Party1) x, msg1, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong unfolding Let_def apply rule+ apply (auto simp: map2_ignore1 map2_ignore2 comp_def prod.case_distrib bind_spmf_const) done also have "\ = (do {x \ x_dist; y \ aby3_shuffleF x; let xr = map (get_party r) x; let yr = map (get_party r) y; msg \ S r xr yr; return_spmf (xr, msg, y)})" unfolding xs unfolding aby3_shuffleF_def apply (simp add: bind_spmf_const map_spmf_conv_bind_spmf) apply (subst lossless_sequence_spmf[unfolded lossless_spmf_def]) subgoal by simp apply simp apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst (3) hoist_map_spmf[where s="sequence_spmf (map share_nat _)" and f="map reconstruct"]) apply (subst map_sequence_share_nat_reconstruct) apply (simp add: map_spmf_conv_bind_spmf) apply (subst Let_def) supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong supply [simp] = finite_permutations apply rule apply rule apply simp apply rule apply rule unfolding S_def S1_def r apply (simp add: Let_def) done finally have "?hoisted = \" . } note simulate_party1 = this { assume r: "r = Party2" have project_to_Party2: "?hoisted = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let y' = map (aby3_stack_sharing Party1) x; let \a2 = map (case_prod (-)) (zip (map (get_party Party2) a) (map (get_party Party2) (permute_list \a y'))); \ \round 2\ let a' = map (get_party Party1) a; let y' = map (aby3_stack_sharing Party2) a; let \b2 = map (case_prod (-)) (zip (map (get_party Party2) b) (map (get_party Party2) (permute_list \b y'))); \ \round 3\ let b' = map (get_party Party2) b; let y' = map (aby3_stack_sharing Party3) b; let \c2 = map (case_prod (-)) (zip (map (get_party Party2) c) (map (get_party Party2) (permute_list \c y'))); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" by (simp add: r Let_def get_party_map_sharing2 map_map_prod_zip') also have project_to_Party2: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let y' = map (aby3_stack_sharing Party1) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map (get_party Party2) y')); \ \round 2\ let a' = map (get_party Party1) a; let y' = map (aby3_stack_sharing Party2) a; let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) y')); \ \round 3\ let b' = map (get_party Party2) b; let y' = map (aby3_stack_sharing Party3) b; let \c2 = map2 (-) (map (get_party Party2) c) (permute_list \c (map (get_party Party2) y')); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp .. also have reduce_Lets: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map2 (+) (map (get_party Party2) x) (map (get_party Party3) x))); \ \round 2\ let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) a)); \ \round 3\ let \c2 = map2 (-) (map (get_party Party2) c) (permute_list \c (replicate (length b) 0)); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" unfolding Let_def unfolding aby3_stack_sharing_def by (simp add: comp_def make_sharing'_sel map_replicate_const zip_map_map_same[symmetric]) also have simplify_minus_zero: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map2 (+) (map (get_party Party2) x) (map (get_party Party3) x))); \ \round 2\ let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) a)); \ \round 3\ let \c2 = (map (get_party Party2) c); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst permute_list_replicate) subgoal by simp apply (subst map2_minus_zero) subgoal by simp subgoal by simp .. also have break_perms_1: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); (\a,\b,\c) \ pair_spmf (spmf_of_set {\. \ permutes {... \ permutes {... \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map2 (+) (map (get_party Party2) x) (map (get_party Party3) x))); \ \round 2\ let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) a)); \ \round 3\ let \c2 = (map (get_party Party2) c); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" unfolding pair_spmf_alt_def by simp also have break_perms_2: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); ((\a,\b,\c),\) \ map_spmf (\(\a,\b,\c). ((\a,\b,\c), \a \ \b \ \c)) (pair_spmf (spmf_of_set {\. \ permutes {... \ permutes {... \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let x' = map (get_party Party3) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map2 (+) (map (get_party Party2) x) (map (get_party Party3) x))); \ \round 2\ let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) a)); \ \round 3\ let \c2 = (map (get_party Party2) c); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" unfolding pair_spmf_alt_def map_spmf_conv_bind_spmf by simp also have break_perms_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c = inv \b \ inv \a \ \; a \ sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let x' = map (get_party Party3) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map2 (+) (map (get_party Party2) x) (map (get_party Party3) x))); \ \round 2\ let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) a)); \ \round 3\ let \c2 = (map (get_party Party2) c); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" apply (unfold random_perm_right) apply (unfold map_spmf_conv_bind_spmf pair_spmf_alt_def) by simp also have break_seqs_3: "\ = (do { xs \ xs; let n = length xs; x2 \ sequence_spmf (replicate n (spmf_of_set UNIV)); x3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let x = map3 (\a b c. make_sharing' Party2 Party3 Party1 b c (a - (b + c))) xs x2 x3; \ \ spmf_of_set {\. \ permutes {..a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); a3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let a = map3 (\a b c. make_sharing' Party2 Party3 Party1 b c (a - (b + c))) (permute_list \a xs) a2 a3; b2 \ sequence_spmf (replicate n (spmf_of_set UNIV)); b3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let b = map3 (\a b c. make_sharing' Party2 Party3 Party1 b c (a - (b + c))) (permute_list (\a \ \b) xs) b2 b3; c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let x' = map (get_party Party3) x; let \a2 = map2 (-) (map (get_party Party2) a) (permute_list \a (map2 (+) (map (get_party Party2) x) (map (get_party Party3) x))); \ \round 2\ let \b2 = map2 (-) (map (get_party Party2) b) (permute_list \b (map (get_party Party2) a)); \ \round 3\ let \c2 = (map (get_party Party2) c); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (map (get_party Party2) x, msg2, c) })" apply (unfold sequence_share_nat_calc'[of Party2 Party3 Party1, simplified]) apply (simp add: pair_spmf_alt_def Let_def) done also have break_seqs_3: "\ = (do { xs \ xs; let n = length xs; x2 \ sequence_spmf (replicate n (spmf_of_set UNIV)); x3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); \ \ spmf_of_set {\. \ permutes {..a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); a3::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); b2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); b3::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let x' = x3; let \a2 = map2 (-) a2 (permute_list \a (map2 (+) x2 x3)); \ \round 2\ let \b2 = map2 (-) b2 (permute_list \b a2); \ \round 3\ let \c2 = (map (get_party Party2) c); let msg2 = ((\a2, \b2, \c2), x', \a, \b); return_spmf (x2, msg2, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong unfolding Let_def apply rule+ apply (auto simp: map2_ignore1 map2_ignore2 comp_def prod.case_distrib bind_spmf_const make_sharing'_sel) done also have "\ = (do {x \ x_dist; y \ aby3_shuffleF x; let xr = map (get_party r) x; let yr = map (get_party r) y; msg \ S r xr yr; return_spmf (xr, msg, y)})" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong unfolding xs unfolding aby3_shuffleF_def apply (simp add: bind_spmf_const map_spmf_conv_bind_spmf) apply (subst lossless_sequence_spmf[unfolded lossless_spmf_def]) subgoal by simp apply (subst lossless_sequence_spmf[unfolded lossless_spmf_def]) subgoal by simp apply simp apply rule apply (subst (2) sequence_share_nat_calc'[of Party2 Party3 Party1, simplified]) apply (subst (2) Let_def) apply (simp add: pair_spmf_alt_def) apply (subst Let_def) unfolding S_def r apply (simp add: pair_spmf_alt_def comp_def prod.case_distrib map2_ignore2 make_sharing'_sel) apply (rule trans[rotated]) apply (rule bind_spmf_sequence_replicate_cong) apply (rule bind_spmf_sequence_replicate_cong) apply (simp add: map2_ignore1 map2_ignore2) apply (subst bind_spmf_const) apply (subst lossless_sequence_spmf[unfolded lossless_spmf_def]) subgoal by simp apply simp apply (subst (2) bind_commute_spmf[where p="sequence_spmf (replicate _ _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply rule apply rule apply rule unfolding S2_def Let_def apply simp done finally have "?hoisted = \" . } note simulate_party2 = this { assume r: "r = Party3" have project_to_Party3: "?hoisted = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let y' = map (aby3_stack_sharing Party1) x; let \a3 = map (case_prod (-)) (zip (map (get_party Party3) a) (map (get_party Party3) (permute_list \a y'))); \ \round 2\ let a' = map (get_party Party1) a; let y' = map (aby3_stack_sharing Party2) a; let \b3 = map (case_prod (-)) (zip (map (get_party Party3) b) (map (get_party Party3) (permute_list \b y'))); \ \round 3\ let b' = map (get_party Party2) b; let y' = map (aby3_stack_sharing Party3) b; let \c3 = map (case_prod (-)) (zip (map (get_party Party3) c) (map (get_party Party3) (permute_list \c y'))); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" by (simp add: r Let_def get_party_map_sharing2 map_map_prod_zip') also have project_to_Party1: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let x' = map (get_party Party3) x; let y' = map (aby3_stack_sharing Party1) x; let \a3 = map2 (-) (map (get_party Party3) a) (permute_list \a (map (get_party Party3) y')); \ \round 2\ let a' = map (get_party Party1) a; let y' = map (aby3_stack_sharing Party2) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map (get_party Party3) y')); \ \round 3\ let b' = map (get_party Party2) b; let y' = map (aby3_stack_sharing Party3) b; let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) y')); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp apply rule+ apply (subst permute_list_map[symmetric]) subgoal by simp .. also have reduce_Lets: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let \a3 = map2 (-) (map (get_party Party3) a) (permute_list \a (replicate (length x) 0)); \ \round 2\ let a' = map (get_party Party1) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map2 (+) (map (get_party Party3) a) (map (get_party Party1) a))); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) b)); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" unfolding Let_def unfolding aby3_stack_sharing_def by (simp add: comp_def make_sharing'_sel map_replicate_const zip_map_map_same[symmetric]) also have simplify_minus_zero: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let \a3 = (map (get_party Party3) a); \ \round 2\ let a' = map (get_party Party1) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map2 (+) (map (get_party Party3) a) (map (get_party Party1) a))); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) b)); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong apply rule+ apply (subst permute_list_replicate) subgoal by simp apply (subst map2_minus_zero) subgoal by simp subgoal by simp .. also have break_perms_1: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); (\a,\b,\c) \ pair_spmf (spmf_of_set {\. \ permutes {... \ permutes {... \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list (\a \ \b \ \c) xs)); \ \round 1\ let \a3 = (map (get_party Party3) a); \ \round 2\ let a' = map (get_party Party1) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map2 (+) (map (get_party Party3) a) (map (get_party Party1) a))); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) b)); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" unfolding pair_spmf_alt_def by simp also have break_perms_2: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); ((\a,\b,\c),\) \ map_spmf (\(\a,\b,\c). ((\a,\b,\c), \a \ \b \ \c)) (pair_spmf (spmf_of_set {\. \ permutes {... \ permutes {... \ permutes {.. sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a3 = (map (get_party Party3) a); \ \round 2\ let a' = map (get_party Party1) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map2 (+) (map (get_party Party3) a) (map (get_party Party1) a))); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) b)); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" unfolding pair_spmf_alt_def map_spmf_conv_bind_spmf by simp also have break_perms_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {..a = \ \ inv \c \ inv \b; a \ sequence_spmf (map share_nat (permute_list \a xs)); b \ sequence_spmf (map share_nat (permute_list (\a \ \b) xs)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a3 = (map (get_party Party3) a); \ \round 2\ let a' = map (get_party Party1) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map2 (+) (map (get_party Party3) a) (map (get_party Party1) a))); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) b)); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" apply (unfold random_perm_left) apply (unfold map_spmf_conv_bind_spmf pair_spmf_alt_def) by (simp add: Let_def) also have break_seqs_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {..a = \ \ inv \c \ inv \b; a3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); a1 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let a = map3 (\a b c. make_sharing' Party3 Party1 Party2 b c (a - (b + c))) (permute_list \a xs) a3 a1; b3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); b1 \ sequence_spmf (replicate n (spmf_of_set UNIV)); let b = map3 (\a b c. make_sharing' Party3 Party1 Party2 b c (a - (b + c))) (permute_list (\a \ \b) xs) b3 b1; c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a3 = (map (get_party Party3) a); \ \round 2\ let a' = map (get_party Party1) a; let \b3 = map2 (-) (map (get_party Party3) b) (permute_list \b (map2 (+) (map (get_party Party3) a) (map (get_party Party1) a))); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c (map (get_party Party3) b)); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" apply (unfold sequence_share_nat_calc'[of Party3 Party1 Party2, simplified]) apply (simp add: pair_spmf_alt_def Let_def) done also have break_seqs_3: "\ = (do { xs \ xs; let n = length xs; x \ sequence_spmf (map share_nat xs); \ \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {..a = \ \ inv \c \ inv \b; a3::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); a1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); b3::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); b1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); c \ sequence_spmf (map share_nat (permute_list \ xs)); \ \round 1\ let \a3 = a3; \ \round 2\ let a' = a1; let \b3 = map2 (-) b3 (permute_list \b (map2 (+) a3 a1)); \ \round 3\ let \c3 = map2 (-) (map (get_party Party3) c) (permute_list \c b3); let msg3 = ((\a3, \b3, \c3), a', \b, \c); return_spmf (map (get_party Party3) x, msg3, c) })" supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong unfolding Let_def apply rule+ apply (auto simp: map2_ignore1 map2_ignore2 comp_def prod.case_distrib bind_spmf_const make_sharing'_sel) done also have "\ = (do {x \ x_dist; y \ aby3_shuffleF x; let xr = map (get_party r) x; let yr = map (get_party r) y; msg \ S r xr yr; return_spmf (xr, msg, y)})" unfolding xs unfolding aby3_shuffleF_def apply (subst bind_spmf_const) apply (subst lossless_sequence_spmf[unfolded lossless_spmf_def]) subgoal by simp apply (simp add: map_spmf_conv_bind_spmf) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst bind_commute_spmf[where q="sequence_spmf (map share_nat _)"]) apply (subst (3) hoist_map_spmf[where s="sequence_spmf (map share_nat _)" and f="map reconstruct"]) apply (subst map_sequence_share_nat_reconstruct) apply (simp add: map_spmf_conv_bind_spmf) apply (subst Let_def) supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong supply [simp] = finite_permutations apply rule apply rule apply simp apply rule apply rule unfolding S_def S3_def r apply (simp add: Let_def) done finally have "?hoisted = \" . } note simulate_party3 = this show ?thesis unfolding hoisted_save apply (cases r) subgoal using simulate_party1 . subgoal using simulate_party2 . subgoal using simulate_party3 . done qed lemma Collect_case_prod: "{f x y | x y. P x y} = (case_prod f) ` (Collect (case_prod P))" by auto lemma inj_split_Cons': "inj_on (\(n, xs). n#xs) X" by (auto intro!: inj_onI) lemma finite_indicator_eq_sum: "finite A \ indicat_real A x = sum (indicat_real {x}) A" by (induction rule: finite_induct) (auto simp: indicator_def) lemma spmf_of_set_Cons: "spmf_of_set (set_Cons A B) = map2_spmf (#) (spmf_of_set A) (spmf_of_set B)" unfolding set_Cons_def pair_spmf_of_set apply (rule spmf_eq_iff_set) subgoal unfolding Collect_case_prod apply (auto simp: set_spmf_of_set ) apply (subst (asm) finite_image_iff) subgoal by (rule inj_split_Cons') subgoal by (auto simp: finite_cartesian_product_iff) done subgoal unfolding Collect_case_prod apply (auto simp: spmf_of_set map_spmf_conv_bind_spmf spmf_bind integral_spmf_of_set) apply (subst card_image) subgoal by (rule inj_split_Cons') apply (auto simp: card_eq_0_iff indicator_single_Some) apply (subst (asm) finite_indicator_eq_sum) subgoal by (simp add: finite_image_iff inj_split_Cons') apply (subst (asm) sum.reindex) subgoal by (simp add: finite_image_iff inj_split_Cons') apply (auto) done done lemma sequence_spmf_replicate: "sequence_spmf (replicate n (spmf_of_set A)) = spmf_of_set (listset (replicate n A))" apply (induction n) subgoal by (auto simp: spmf_of_set_singleton) subgoal by (auto simp: spmf_of_set_Cons) done lemma listset_replicate: "listset (replicate n A) = {l. length l = n \ set l \ A}" apply (induction n) apply (auto simp: set_Cons_def) subgoal for n x by (cases x; simp) done lemma map2_map2_map3: "map2 f (map2 g x y) z = map3 (\x y. f (g x y)) x y z" by (auto simp: zip_assoc map_zip_map) lemma inv_add_sequence: assumes "n = length x" shows " map_spmf (\\::natL list. (\, map2 (+) \ x)) (sequence_spmf (replicate n (spmf_of_set UNIV))) = map_spmf (\y. (map2 (-) y x, y)) (sequence_spmf (replicate n (spmf_of_set UNIV)))" unfolding sequence_spmf_replicate apply (subst map_spmf_of_set_inj_on) subgoal unfolding inj_on_def by simp apply (subst map_spmf_of_set_inj_on) subgoal unfolding inj_on_def by simp apply (rule arg_cong[where f="spmf_of_set"]) using assms apply (auto simp: image_def listset_replicate map2_map2_map3 zip_same_conv_map map_zip_map2 map2_ignore2) done lemma S1_def_simplified: "S1 x1 yc1 = (do { let n = length x1; \a \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {..a1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); yb1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); yb2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); let \c1 = map2 (-) (yc1) (permute_list \c (map2 (+) yb1 yb2)); return_spmf ((\a1, yb1, \c1), yb2, \a, \c) })" unfolding S1_def supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong bind_spmf_sequence_map_share_nat_cong apply rule apply rule apply rule apply (subst hoist_map_spmf'[where s="sequence_spmf _" and f="\x. map2 (-) x _"]) apply (subst inv_add_sequence[symmetric]) subgoal by simp unfolding map_spmf_conv_bind_spmf apply simp done lemma S2_def_simplified: "S2 x2 yc2 = (do { let n = length x2; x3 \ sequence_spmf (replicate n (spmf_of_set UNIV)); \a \ spmf_of_set {\. \ permutes {..b \ spmf_of_set {\. \ permutes {..a2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); \b2::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); let msg2 = ((\a2, \b2, yc2), x3, \a, \b); return_spmf msg2 })" unfolding S2_def supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong bind_spmf_sequence_map_share_nat_cong apply rule apply rule apply rule apply rule apply (rule trans) apply (rule bind_spmf_sequence_replicate_cong) apply (subst hoist_map_spmf'[where s="sequence_spmf _" and f="\x. map2 (-) x _"]) apply (subst inv_add_sequence[symmetric]) subgoal by simp apply (rule refl) apply (unfold map_spmf_conv_bind_spmf) apply simp apply (subst hoist_map_spmf'[where s="sequence_spmf _" and f="\x. map2 (-) x _"]) apply (subst inv_add_sequence[symmetric]) subgoal by simp apply (unfold map_spmf_conv_bind_spmf) apply simp done lemma S3_def_simplified: "S3 x3 yc3 = (do { let n = length x3; \b \ spmf_of_set {\. \ permutes {..c \ spmf_of_set {\. \ permutes {.. sequence_spmf (replicate n (spmf_of_set UNIV)); ya1::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); \b3::natL list \ sequence_spmf (replicate n (spmf_of_set UNIV)); let \c3 = map2 (-) yc3 (permute_list \c (map2 (+) \b3 (permute_list \b (map2 (+) ya3 ya1)))); return_spmf ((ya3, \b3, \c3), ya1, \b, \c) })" unfolding S3_def supply [intro!] = bind_spmf_cong[OF refl] let_cong[OF refl] prod.case_cong[OF refl] bind_spmf_sequence_map_cong bind_spmf_sequence_replicate_cong bind_spmf_permutes_cong bind_spmf_sequence_map_share_nat_cong apply rule apply rule apply rule apply rule apply rule apply (subst hoist_map_spmf'[where s="sequence_spmf _" and f="\x. map2 (-) x _"]) apply (subst inv_add_sequence[symmetric]) subgoal by simp apply (unfold map_spmf_conv_bind_spmf) apply simp done end diff --git a/thys/ABY3_Protocols/Spmf_Common.thy b/thys/ABY3_Protocols/Spmf_Common.thy --- a/thys/ABY3_Protocols/Spmf_Common.thy +++ b/thys/ABY3_Protocols/Spmf_Common.thy @@ -1,437 +1,437 @@ theory Spmf_Common imports CryptHOL.CryptHOL begin no_adhoc_overloading Monad_Syntax.bind bind_pmf lemma mk_lossless_back_eq: "scale_spmf (weight_spmf s) (mk_lossless s) = s" unfolding mk_lossless_def unfolding scale_scale_spmf by (auto simp: field_simps weight_spmf_eq_0) lemma cond_spmf_enforce: "cond_spmf sx (Collect A) = mk_lossless (enforce_spmf A sx)" unfolding enforce_spmf_def unfolding cond_spmf_alt unfolding restrict_spmf_def unfolding enforce_option_alt_def apply (rule arg_cong[where f="mk_lossless"]) apply (rule arg_cong[where f="\x. map_pmf x sx"]) apply (intro ext) apply (rule arg_cong[where f="Option.bind _"]) apply auto done definition "rel_scale_spmf s t \ (mk_lossless s = mk_lossless t)" lemma rel_scale_spmf_refl: "rel_scale_spmf s s" unfolding rel_scale_spmf_def .. lemma rel_scale_spmf_sym: "rel_scale_spmf s t \ rel_scale_spmf t s" unfolding rel_scale_spmf_def by simp lemma rel_scale_spmf_trans: "rel_scale_spmf s t \ rel_scale_spmf t u \ rel_scale_spmf s u" unfolding rel_scale_spmf_def by simp lemma rel_scale_spmf_equiv: "equivp rel_scale_spmf" using rel_scale_spmf_refl rel_scale_spmf_sym by (auto intro!: equivpI reflpI sympI transpI dest: rel_scale_spmf_trans) lemma spmf_eq_iff: "p = q \ (\i. spmf p i = spmf q i)" using spmf_eqI by auto lemma spmf_eq_iff_set: "set_spmf a = set_spmf b \ (\x. x \ set_spmf b \ spmf a x = spmf b x) \ a = b" using in_set_spmf_iff_spmf spmf_eq_iff by (metis) lemma rel_scale_spmf_None: "rel_scale_spmf s t \ s = return_pmf None \ t = return_pmf None" unfolding rel_scale_spmf_def by auto lemma rel_scale_spmf_def_alt: "rel_scale_spmf s t \ (\k>0. s = scale_spmf k t)" proof assume rel: "rel_scale_spmf s t" then consider (isNone) "s = return_pmf None \ t = return_pmf None" | (notNone) "weight_spmf s > 0 \ weight_spmf t > 0" using rel_scale_spmf_None weight_spmf_eq_0 zero_less_measure_iff by blast then show "\k>0. s = scale_spmf k t" proof cases case isNone show ?thesis apply (rule exI[of _ 1]) using isNone by simp next case notNone have "scale_spmf (weight_spmf s) (mk_lossless s) = scale_spmf (weight_spmf s / weight_spmf t) t" unfolding rel[unfolded rel_scale_spmf_def] unfolding mk_lossless_def unfolding scale_scale_spmf by (auto simp: field_simps) then show "\k>0. s = scale_spmf k t" apply (unfold mk_lossless_back_eq) using notNone divide_pos_pos by blast qed next assume "\k>0. s = scale_spmf k t" then obtain k where kpos: "k>0" and st: "s = scale_spmf k t" by blast then consider (isNone) "weight_spmf s = 0 \ weight_spmf t = 0" | (notNone) "weight_spmf s > 0 \ weight_spmf t > 0" using zero_less_measure_iff mult_pos_pos zero_less_measure_iff by (fastforce simp: weight_scale_spmf) then show "rel_scale_spmf s t" proof cases case isNone then show ?thesis unfolding rel_scale_spmf_def weight_spmf_eq_0 by simp next case notNone then show ?thesis unfolding rel_scale_spmf_def unfolding mk_lossless_def unfolding st by (cases "k < inverse (weight_spmf t)") (auto simp: weight_scale_spmf scale_scale_spmf field_simps) qed qed lemma rel_scale_spmf_def_alt2: "rel_scale_spmf s t \ (s = return_pmf None \ t = return_pmf None) | (weight_spmf s > 0 \ weight_spmf t > 0 \ s = scale_spmf (weight_spmf s / weight_spmf t) t)" (is "?lhs \ ?rhs") proof assume rel: ?lhs then consider (isNone) "s = return_pmf None \ t = return_pmf None" | (notNone) "weight_spmf s > 0 \ weight_spmf t > 0" using rel_scale_spmf_None weight_spmf_eq_0 zero_less_measure_iff by blast thus ?rhs proof cases case notNone have "scale_spmf (weight_spmf s) (mk_lossless s) = scale_spmf (weight_spmf s / weight_spmf t) t" unfolding rel[unfolded rel_scale_spmf_def] unfolding mk_lossless_def unfolding scale_scale_spmf by (auto simp: field_simps) thus ?thesis apply (unfold mk_lossless_back_eq) using notNone by simp qed simp next assume ?rhs then show ?lhs proof cases case right then have gt0: "weight_spmf s > 0" "weight_spmf t > 0" and st: "s = scale_spmf (weight_spmf s / weight_spmf t) t" by auto then have "(1 / weight_spmf t) \ (weight_spmf s / weight_spmf t)" using weight_spmf_le_1 divide_le_cancel by fastforce then show ?thesis unfolding rel_scale_spmf_def mk_lossless_def apply (subst (3) st) using gt0 by (auto simp: scale_scale_spmf field_simps) qed (simp add: rel_scale_spmf_refl) qed lemma rel_scale_spmf_scale: "r > 0 \ rel_scale_spmf s t \ rel_scale_spmf s (scale_spmf r t)" apply (unfold rel_scale_spmf_def_alt) by (metis rel_scale_spmf_def rel_scale_spmf_def_alt) lemma rel_scale_spmf_mk_lossless: "rel_scale_spmf s t \ rel_scale_spmf s (mk_lossless t)" unfolding rel_scale_spmf_def_alt unfolding mk_lossless_def apply (cases "weight_spmf t = 0") subgoal by(simp add: weight_spmf_eq_0) subgoal apply (auto simp: weight_spmf_eq_0 field_simps scale_scale_spmf) using rel_scale_spmf_def_alt rel_scale_spmf_def_alt2 by blast done lemma rel_scale_spmf_eq_iff: "rel_scale_spmf s t \ weight_spmf s = weight_spmf t \ s = t" unfolding rel_scale_spmf_def_alt2 by auto lemma rel_scale_spmf_set_restrict: "finite A \ rel_scale_spmf (restrict_spmf (spmf_of_set A) B) (spmf_of_set (A \ B))" apply (unfold rel_scale_spmf_def) apply (fold cond_spmf_alt) apply (subst cond_spmf_spmf_of_set) subgoal . apply (unfold mk_lossless_spmf_of_set) .. lemma spmf_of_set_restrict_empty: "A \ B = {} \ restrict_spmf (spmf_of_set A) B = return_pmf None" unfolding spmf_of_set_def by simp lemma spmf_of_set_restrict_scale: "finite A \ restrict_spmf (spmf_of_set A) B = scale_spmf (card (A\B) / card A) (spmf_of_set (A\B))" apply (rule rel_scale_spmf_eq_iff) subgoal apply (cases "A\B = {}") subgoal by (auto simp: spmf_of_set_restrict_empty intro: rel_scale_spmf_refl) subgoal apply (rule rel_scale_spmf_scale) subgoal by (metis card_gt_0_iff divide_pos_pos finite_Int inf_bot_left of_nat_0_less_iff) subgoal by (rule rel_scale_spmf_set_restrict) done done subgoal apply (auto simp add: weight_scale_spmf measure_spmf_of_set) by (smt (verit, best) card_gt_0_iff card_mono disjoint_notin1 divide_le_eq_1_pos finite_Int inf_le1 of_nat_0_less_iff of_nat_le_iff) done lemma min_em2: "min a b = c \ a\c \ b=c" unfolding min_def by auto lemma weight_0_spmf: "weight_spmf s = 0 \ spmf s i = 0" using order_trans[OF spmf_le_weight, of s 0 i] by simp lemma mk_lossless_scale_absorb: "r > 0 \ mk_lossless (scale_spmf r s) = mk_lossless s" apply (rule rel_scale_spmf_eq_iff) subgoal apply (rule rel_scale_spmf_trans[where t=s]) subgoal apply (rule rel_scale_spmf_sym) apply (rule rel_scale_spmf_mk_lossless) apply (rule rel_scale_spmf_scale) subgoal . subgoal by (rule rel_scale_spmf_refl) done subgoal apply (rule rel_scale_spmf_mk_lossless) apply (rule rel_scale_spmf_refl) done done subgoal unfolding weight_mk_lossless by (auto simp flip: weight_spmf_eq_0 simp: weight_scale_spmf dest: min_em2) done lemma scale_spmf_None_iff: "scale_spmf k s = return_pmf None \ k\0 \ s=return_pmf None" apply (auto simp: spmf_eq_iff spmf_scale_spmf) using inverse_nonpositive_iff_nonpositive weight_0_spmf measure_le_0_iff - by smt + by (smt (verit)) lemma spmf_of_pmf_the: "lossless_spmf s \ spmf_of_pmf (map_pmf the s) = s" unfolding lossless_spmf_conv_spmf_of_pmf by auto lemma lossless_mk_lossless: "s \ return_pmf None \ lossless_spmf (mk_lossless s)" unfolding lossless_spmf_def unfolding weight_mk_lossless by simp definition pmf_of_spmf where "pmf_of_spmf p = map_pmf the (mk_lossless p)" lemma scale_weight_spmf_of_pmf: "p = scale_spmf (weight_spmf p) (spmf_of_pmf (pmf_of_spmf p))" unfolding pmf_of_spmf_def apply (cases "p = return_pmf None") subgoal by simp subgoal apply (subst mk_lossless_back_eq[of p, symmetric]) apply (rule arg_cong[where f="scale_spmf _ "]) apply (rule spmf_of_pmf_the[symmetric]) by (fact lossless_mk_lossless) done lemma pmf_spmf: "pmf_of_spmf (spmf_of_pmf p) = p" unfolding pmf_of_spmf_def unfolding lossless_spmf_spmf_of_spmf[THEN mk_lossless_lossless] unfolding map_the_spmf_of_pmf .. lemma spmf_pmf: "lossless_spmf p \ spmf_of_pmf (pmf_of_spmf p) = p" unfolding pmf_of_spmf_def by (simp add: spmf_of_pmf_the) lemma pmf_of_spmf_scale_spmf: "r > 0 \ pmf_of_spmf (scale_spmf r p) = pmf_of_spmf p" unfolding pmf_of_spmf_def by (simp add: mk_lossless_scale_absorb) lemma nonempty_spmf_weight: "p \ return_pmf None \ weight_spmf p > 0" apply (fold weight_spmf_eq_0) using dual_order.not_eq_order_implies_strict[OF _ weight_spmf_nonneg[of p]] by auto lemma pmf_of_spmf_mk_lossless: "pmf_of_spmf (mk_lossless p) = pmf_of_spmf p" apply (cases "p = return_pmf None") subgoal by auto apply (unfold mk_lossless_def) apply (subst pmf_of_spmf_scale_spmf) subgoal by (simp add: nonempty_spmf_weight) .. lemma spmf_pmf': "p \ return_pmf None \ spmf_of_pmf (pmf_of_spmf p) = mk_lossless p" apply (subst spmf_pmf[of "mk_lossless p", symmetric]) apply (unfold pmf_of_spmf_mk_lossless) subgoal using lossless_mk_lossless . subgoal .. done lemma rel_scale_spmf_cond_UNIV: "rel_scale_spmf s (cond_spmf s UNIV)" unfolding cond_spmf_UNIV by (rule rel_scale_spmf_mk_lossless) (rule rel_scale_spmf_refl) lemma "set_pmf p \ g \ {} \ pmf_prob p (f \ g) = pmf_prob (cond_pmf p g) f * pmf_prob p g" using measure_cond_pmf unfolding pmf_prob_def by (metis Int_commute divide_eq_eq measure_measure_pmf_not_zero) lemma bayes: "set_pmf p \ A \ {} \ set_pmf p \ B \ {} \ measure_pmf.prob (cond_pmf p A) B = measure_pmf.prob (cond_pmf p B) A * measure_pmf.prob p B / measure_pmf.prob p A" unfolding measure_cond_pmf by (subst inf_commute) (simp add: measure_pmf_zero_iff) definition spmf_prob :: "'a spmf \ 'a set \ real" where "spmf_prob p = Sigma_Algebra.measure (measure_spmf p)" lemma "spmf_prob = measure_measure_spmf" unfolding spmf_prob_def measure_measure_spmf_def .. lemma spmf_prob_pmf: "spmf_prob p A = pmf_prob p (Some ` A)" unfolding spmf_prob_def pmf_prob_def unfolding measure_measure_spmf_conv_measure_pmf .. lemma bayes_spmf: "spmf_prob (cond_spmf p A) B = spmf_prob (cond_spmf p B) A * spmf_prob p B / spmf_prob p A" unfolding spmf_prob_def unfolding measure_cond_spmf by (subst inf_commute) (auto simp: measure_spmf_zero_iff) lemma spmf_prob_pmf_of_spmf: "spmf_prob p A = weight_spmf p * pmf_prob (pmf_of_spmf p) A" apply (subst scale_weight_spmf_of_pmf) apply (unfold spmf_prob_def) apply (subst measure_spmf_scale_spmf') subgoal using weight_spmf_le_1 . by (simp add: pmf_prob_def) lemma cond_spmf_Int: "cond_spmf (cond_spmf p A) B = cond_spmf p (A \ B)" apply (rule spmf_eqI) apply (unfold spmf_cond_spmf) apply(auto simp add: measure_cond_spmf split: if_split_asm) using Int_lower1[THEN measure_spmf.finite_measure_mono[simplified]] measure_le_0_iff by metis lemma cond_spmf_prob: "spmf_prob p (A \ B) = spmf_prob (cond_spmf p A) B * spmf_prob p A" unfolding spmf_prob_def measure_cond_spmf using Int_lower1[THEN measure_spmf.finite_measure_mono[simplified]] measure_le_0_iff by (metis mult_eq_0_iff nonzero_eq_divide_eq) definition "empty_spmf = return_pmf None" lemma spmf_prob_empty: "spmf_prob empty_spmf A = 0" unfolding spmf_prob_def empty_spmf_def by simp definition le_spmf :: "'a spmf \ 'a spmf \ bool" where "le_spmf p q \ (\k\1. p = scale_spmf k q)" definition lt_spmf :: "'a spmf \ 'a spmf \ bool" where "lt_spmf p q \ (\k<1. p = scale_spmf k q)" lemma "class.order_bot empty_spmf le_spmf lt_spmf" oops lemma spmf_prob_cond_Int: "spmf_prob (cond_spmf p C) (A \ B) = spmf_prob (cond_spmf p (B \ C)) A * spmf_prob (cond_spmf p C) B" apply (subst Int_commute[of B C]) apply (subst Int_commute[of A B]) apply (fold cond_spmf_Int) using cond_spmf_prob . lemma cond_spmf_mk_lossless: "cond_spmf (mk_lossless p) A = cond_spmf p A" apply (fold cond_spmf_UNIV) apply (unfold cond_spmf_Int) by simp primrec sequence_spmf :: "'a spmf list \ 'a list spmf" where "sequence_spmf [] = return_spmf []" | "sequence_spmf (x#xs) = map_spmf (case_prod Cons) (pair_spmf x (sequence_spmf xs))" lemma set_sequence_spmf: "set_spmf (sequence_spmf xs) = {l. list_all2 (\x s. x \ set_spmf s) l xs}" by (induction xs) (auto simp: list_all2_Cons2) lemma map_spmf_map_sequence: "map_spmf (map f) (sequence_spmf xs) = sequence_spmf (map (map_spmf f) xs)" apply (induction xs) subgoal by simp subgoal premises IH unfolding list.map unfolding sequence_spmf.simps apply (fold IH) apply (unfold pair_map_spmf) apply (unfold spmf.map_comp) by (simp add: comp_def case_prod_map_prod prod.case_distrib) done lemma sequence_map_return_spmf: "sequence_spmf (map return_spmf xs) = return_spmf xs" by (induction xs) auto lemma sequence_bind_cong: "\xs=ys; \y. length y = length ys \ f y = g y\ \ bind_spmf (sequence_spmf xs) f = bind_spmf (sequence_spmf ys) g" apply (rule bind_spmf_cong) subgoal by simp subgoal unfolding set_sequence_spmf list_all2_iff by simp done lemma bind_spmf_sequence_replicate_cong: "(\l. length l = n \ f l = g l) \ bind_spmf (sequence_spmf (replicate n x)) f = bind_spmf (sequence_spmf (replicate n x)) g" by (rule bind_spmf_cong[OF refl]) (simp add: set_spmf_of_set finite_permutations set_sequence_spmf[unfolded list_all2_iff]) lemma bind_spmf_sequence_map_cong: "(\l. length l = length x \ f l = g l) \ bind_spmf (sequence_spmf (map m x)) f = bind_spmf (sequence_spmf (map m x)) g" by (rule bind_spmf_cong[OF refl]) (simp add: set_spmf_of_set finite_permutations set_sequence_spmf[unfolded list_all2_iff]) lemma lossless_pair_spmf_iff: "lossless_spmf (pair_spmf a b) \ lossless_spmf a \ lossless_spmf b" unfolding pair_spmf_alt_def by (auto simp: set_spmf_eq_empty) lemma lossless_sequence_spmf: "(\x. x\set xs \ lossless_spmf x) \ lossless_spmf (sequence_spmf xs)" by (induction xs) (auto simp: lossless_pair_spmf_iff) end \ No newline at end of file diff --git a/thys/Approximation_Algorithms/Approx_BP_Hoare.thy b/thys/Approximation_Algorithms/Approx_BP_Hoare.thy --- a/thys/Approximation_Algorithms/Approx_BP_Hoare.thy +++ b/thys/Approximation_Algorithms/Approx_BP_Hoare.thy @@ -1,1708 +1,1708 @@ section \Bin Packing\ theory Approx_BP_Hoare imports Complex_Main "HOL-Hoare.Hoare_Logic" "HOL-Library.Disjoint_Sets" begin text \The algorithm and proofs are based on the work by Berghammer and Reuter \<^cite>\BerghammerR03\.\ subsection \Formalization of a Correct Bin Packing\ text \Definition of the unary operator \\\\\ from the article. \B\ will only be wrapped into a set if it is non-empty.\ definition wrap :: "'a set \ 'a set set" where "wrap B = (if B = {} then {} else {B})" lemma wrap_card: "card (wrap B) \ 1" unfolding wrap_def by auto text \If \M\ and \N\ are pairwise disjoint with \V\ and not yet contained in V, then the union of \M\ and \N\ is also pairwise disjoint with \V\.\ lemma pairwise_disjnt_Un: assumes "pairwise disjnt ({M} \ {N} \ V)" "M \ V" "N \ V" shows "pairwise disjnt ({M \ N} \ V)" using assms unfolding pairwise_def by auto text \A Bin Packing Problem is defined like in the article:\ locale BinPacking = fixes U :: "'a set" \ \A finite, non-empty set of objects\ and w :: "'a \ real" \ \A mapping from objects to their respective weights (positive real numbers)\ and c :: nat \ \The maximum capacity of a bin (a natural number)\ and S :: "'a set" \ \The set of \small\ objects (weight no larger than \1/2\ of \c\)\ and L :: "'a set" \ \The set of \large\ objects (weight larger than \1/2\ of \c\)\ assumes weight: "\u \ U. 0 < w(u) \ w(u) \ c" and U_Finite: "finite U" and U_NE: "U \ {}" and S_def: "S = {u \ U. w(u) \ c / 2}" and L_def: "L = U - S" begin text \In the article, this is defined as \w\ as well. However, to avoid ambiguity, we will abbreviate the weight of a bin as \W\.\ abbreviation W :: "'a set \ real" where "W B \ (\u \ B. w(u))" text \\P\ constitutes as a correct bin packing if \P\ is a partition of \U\ (as defined in @{thm [source] partition_on_def}) and the weights of the bins do not exceed their maximum capacity \c\.\ definition bp :: "'a set set \ bool" where "bp P \ partition_on U P \ (\B \ P. W(B) \ c)" lemma bpE: assumes "bp P" shows "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" using assms unfolding bp_def partition_on_def by blast+ lemma bpI: assumes "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" shows "bp P" using assms unfolding bp_def partition_on_def by blast text \Although we assume the \S\ and \L\ sets as given, manually obtaining them from \U\ is trivial and can be achieved in linear time. Proposed by the article \<^cite>\"BerghammerR03"\.\ lemma S_L_set_generation: "VARS S L W u {True} S := {}; L := {}; W := U; WHILE W \ {} INV {W \ U \ S = {v \ U - W. w(v) \ c / 2} \ L = {v \ U - W. w(v) > c / 2}} DO u := (SOME u. u \ W); IF 2 * w(u) \ c THEN S := S \ {u} ELSE L := L \ {u} FI; W := W - {u} OD {S = {v \ U. w(v) \ c / 2} \ L = {v \ U. w(v) > c / 2}}" by vcg (auto simp: some_in_eq) subsection \The Proposed Approximation Algorithm\ subsubsection \Functional Correctness\ text \According to the article, \inv\<^sub>1\ holds if \P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}\ is a correct solution for the bin packing problem \<^cite>\BerghammerR03\. However, various assumptions made in the article seem to suggest that more information is demanded from this invariant and, indeed, mere correctness (as defined in @{thm [source] bp_def}) does not appear to suffice. To amend this, four additional conjuncts have been added to this invariant, whose necessity will be explained in the following proofs. It should be noted that there may be other (shorter) ways to amend this invariant. This approach, however, makes for rather straight-forward proofs, as these conjuncts can be utilized and proved in relatively few steps.\ definition inv\<^sub>1 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ \A correct solution to the bin packing problem\ \ \(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V \ \The partial solution does not contain objects that have not yet been assigned\ \ B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ \\B\<^sub>1\ is distinct from all the other bins\ \ B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2) \ \\B\<^sub>2\ is distinct from all the other bins\ \ (P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {} \ \The first and second partial solutions are disjoint from each other.\" (* lemma "partition_on U (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ u \ V \ partition_on U (P\<^sub>1 \ wrap (insert u B\<^sub>1) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ (V-{u})})" nitpick*) lemma inv\<^sub>1E: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using assms unfolding inv\<^sub>1_def by auto lemma inv\<^sub>1I: assumes "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>1_def by blast lemma wrap_Un [simp]: "wrap (M \ {x}) = {M \ {x}}" unfolding wrap_def by simp lemma wrap_empty [simp]: "wrap {} = {}" unfolding wrap_def by simp lemma wrap_not_empty [simp]: "M \ {} \ wrap M = {M}" unfolding wrap_def by simp text \If \inv\<^sub>1\ holds for the current partial solution, and the weight of an object \u \ V\ added to \B\<^sub>1\ does not exceed its capacity, then \inv\<^sub>1\ also holds if \B\<^sub>1\ and \{u}\ are replaced by \B\<^sub>1 \ {u}\.\ lemma inv\<^sub>1_stepA: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W(B\<^sub>1) + w(u) \ c" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] text \In the proof for \Theorem 3.2\ of the article it is erroneously argued that if \P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}\ is a partition of \U\, then the same holds if \B\<^sub>1\ is replaced by \B\<^sub>1 \ {u}\. This is, however, not necessarily the case if \B\<^sub>1\ or \{u}\ are already contained in the partial solution. Suppose \P\<^sub>1\ contains the non-empty bin \B\<^sub>1\, then \P\<^sub>1 \ wrap B\<^sub>1\ would still be pairwise disjoint, provided \P\<^sub>1\ was pairwise disjoint before, as the union simply ignores the duplicate \B\<^sub>1\. Now, if the algorithm modifies \B\<^sub>1\ by adding an element from \V\ such that \B\<^sub>1\ becomes some non-empty \B\<^sub>1'\ with \B\<^sub>1 \ B\<^sub>1' \ \\ and \B\<^sub>1' \ P\<^sub>1\, one can see that this property would no longer be preserved. To avoid such a situation, we will use the first additional conjunct in \inv\<^sub>1\ to ensure that \{u}\ is not yet contained in the partial solution, and the second additional conjunct to ensure that \B\<^sub>1\ is not yet contained in the partial solution.\ \ \Rule 1: Pairwise Disjoint\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ ({{u}} \ {{v} |v. v \ V - {u}}))" using bprules(1) assms(2) by simp then have "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" by (simp add: Un_commute) then have assm: "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc) have "pairwise disjnt ({B\<^sub>1 \ {u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>1 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>1} \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,3) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>1" using bprules(3) U_Finite by (cases \B\<^sub>1 = {}\) auto then have "W (B\<^sub>1 \ {u}) \ W B\<^sub>1 + w u" using \0 < w u\ by (cases \u \ B\<^sub>1\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>1 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>1 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" by blast from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = U - (V - {u})" unfolding L R invrules(2) .. have 3: "B\<^sub>1 \ {u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2" using NOTIN by auto have 4: "B\<^sub>2 \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2" using invrules(4) NOTIN unfolding wrap_def by fastforce have 5: "(P\<^sub>1 \ wrap (B\<^sub>1 \ {u})) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) NOTIN unfolding wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \If \inv\<^sub>1\ holds for the current partial solution, and the weight of an object \u \ V\ added to \B\<^sub>2\ does not exceed its capacity, then \inv\<^sub>1\ also holds if \B\<^sub>2\ and \{u}\ are replaced by \B\<^sub>2 \ {u}\.\ lemma inv\<^sub>1_stepB: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W B\<^sub>2 + w u \ c" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] text \The argumentation here is similar to the one in @{thm [source] inv\<^sub>1_stepA} with \B\<^sub>1\ replaced with \B\<^sub>2\ and using the first and third additional conjuncts of \inv\<^sub>1\ to amend the issue, instead of the first and second.\ \ \Rule 1: Pairwise Disjoint\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}})" using bprules(1) assms(2) by simp then have assm: "pairwise disjnt (wrap B\<^sub>2 \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc Un_commute) have "pairwise disjnt ({B\<^sub>2 \ {u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>2 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>2} \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,4) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>2" using bprules(3) U_Finite by (cases \B\<^sub>2 = {}\) auto then have "W (B\<^sub>2 \ {u}) \ W B\<^sub>2 + w u" using \0 < w u\ by (cases \u \ B\<^sub>2\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>2 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>2 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}. W B \ c" by auto from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = U - (V - {u})" unfolding L R using invrules(2) by simp have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})" using bpE(2)[OF 1] by simp have 4: "B\<^sub>2 \ {u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = {}" using invrules(5) NOTIN unfolding wrap_empty wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \If \inv\<^sub>1\ holds for the current partial solution, then \inv\<^sub>1\ also holds if \B\<^sub>1\ and \B\<^sub>2\ are added to \P\<^sub>1\ and \P\<^sub>2\ respectively, \B\<^sub>1\ is emptied and \B\<^sub>2\ initialized with \u \ V\.\ lemma inv\<^sub>1_stepC: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] \ \Rule 1-4: Correct Bin Packing\ have "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}" by (metis (no_types, lifting) Un_assoc Un_empty_right insert_not_empty wrap_empty wrap_not_empty) also have "... = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" using assms(2) by auto finally have EQ: "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" . from invrules(1) have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}})" unfolding EQ . \ \Auxiliary information is preserved\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "u \ U" using assms(2) bpE(3)[OF invrules(1)] by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = U - (V - {u})" unfolding L R using invrules(2) by auto have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}" using bpE(2)[OF 1] by simp have 4: "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}) = {}" using invrules(5) NOTIN unfolding wrap_def by force from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \A simplified version of the bin packing algorithm proposed in the article. It serves as an introduction into the approach taken, and, while it does not provide the desired approximation factor, it does ensure that \P\ is a correct solution of the bin packing problem.\ lemma simple_bp_correct: "VARS P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u {True} P\<^sub>1 := {}; P\<^sub>2 := {}; B\<^sub>1 := {}; B\<^sub>2 := {}; V := U; WHILE V \ S \ {} INV {inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V} DO u := (SOME u. u \ V); V := V - {u}; IF W(B\<^sub>1) + w(u) \ c THEN B\<^sub>1 := B\<^sub>1 \ {u} ELSE IF W(B\<^sub>2) + w(u) \ c THEN B\<^sub>2 := B\<^sub>2 \ {u} ELSE P\<^sub>2 := P\<^sub>2 \ wrap B\<^sub>2; B\<^sub>2 := {u} FI; P\<^sub>1 := P\<^sub>1 \ wrap B\<^sub>1; B\<^sub>1 := {} FI OD; P := P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} | v. v \ V} {bp P}" proof (vcg, goal_cases) case (1 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) show ?case unfolding bp_def partition_on_def pairwise_def wrap_def inv\<^sub>1_def using weight by auto next case (2 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then have INV: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" .. from 2 have "V \ {}" by blast then have IN: "(SOME u. u \ V) \ V" by (simp add: some_in_eq) from inv\<^sub>1_stepA[OF INV IN] inv\<^sub>1_stepB[OF INV IN] inv\<^sub>1_stepC[OF INV IN] show ?case by blast next case (3 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then show ?case unfolding inv\<^sub>1_def by blast qed subsubsection \Lower Bounds for the Bin Packing Problem\ lemma bp_bins_finite [simp]: assumes "bp P" shows "\B \ P. finite B" using bpE(3)[OF assms] U_Finite by (meson Sup_upper finite_subset) lemma bp_sol_finite [simp]: assumes "bp P" shows "finite P" using bpE(3)[OF assms] U_Finite by (simp add: finite_UnionD) text \If \P\ is a solution of the bin packing problem, then no bin in \P\ may contain more than one large object.\ lemma only_one_L_per_bin: assumes "bp P" "B \ P" shows "\x \ B. \y \ B. x \ y \ x \ L \ y \ L" proof (rule ccontr, simp) assume "\x\B. \y\B. x \ y \ x \ L \ y \ L" then obtain x y where *: "x \ B" "y \ B" "x \ y" "x \ L" "y \ L" by blast then have "c < w x + w y" using L_def S_def by force have "finite B" using assms by simp have "y \ B - {x}" using *(2,3) by blast have "W B = W (B - {x}) + w x" using *(1) \finite B\ by (simp add: sum.remove) also have "... = W (B - {x} - {y}) + w x + w y" using \y \ B - {x}\ \finite B\ by (simp add: sum.remove) finally have *: "W B = W (B - {x} - {y}) + w x + w y" . have "\u \ B. 0 < w u" using bpE(3)[OF assms(1)] assms(2) weight by blast - then have "0 \ W (B - {x} - {y})" by (smt DiffD1 sum_nonneg) + then have "0 \ W (B - {x} - {y})" by (smt (verit) DiffD1 sum_nonneg) with * have "c < W B" using \c < w x + w y\ by simp then show False using bpE(4)[OF assms(1)] assms(2) by fastforce qed text \If \P\ is a solution of the bin packing problem, then the amount of large objects is a lower bound for the amount of bins in P.\ lemma L_lower_bound_card: assumes "bp P" shows "card L \ card P" proof - have "\x \ L. \B \ P. x \ B" using bpE(3)[OF assms] L_def by blast then obtain f where f_def: "\u \ L. u \ f u \ f u \ P" by metis then have "inj_on f L" unfolding inj_on_def using only_one_L_per_bin[OF assms] by blast then have card_eq: "card L = card (f ` L)" by (simp add: card_image) have "f ` L \ P" using f_def by blast moreover have "finite P" using assms by simp ultimately have "card (f ` L) \ card P" by (simp add: card_mono) then show ?thesis unfolding card_eq . qed text \If \P\ is a solution of the bin packing problem, then the amount of bins of a subset of P in which every bin contains a large object is a lower bound on the amount of large objects.\ lemma subset_bp_card: assumes "bp P" "M \ P" "\B \ M. B \ L \ {}" shows "card M \ card L" proof - have "\B \ M. \u \ L. u \ B" using assms(3) by fast then have "\f. \B \ M. f B \ L \ f B \ B" by metis then obtain f where f_def: "\B \ M. f B \ L \ f B \ B" .. have "inj_on f M" proof (rule ccontr) assume "\ inj_on f M" then have "\x \ M. \y \ M. x \ y \ f x = f y" unfolding inj_on_def by blast then obtain x y where *: "x \ M" "y \ M" "x \ y" "f x = f y" by blast then have "\u. u \ x \ u \ y" using f_def by metis then have "x \ y \ {}" by blast moreover have "pairwise disjnt M" using pairwise_subset[OF bpE(1)[OF assms(1)] assms(2)] . ultimately show False using * unfolding pairwise_def disjnt_def by simp qed moreover have "finite L" using L_def U_Finite by blast moreover have "f ` M \ L" using f_def by blast ultimately show ?thesis using card_inj_on_le by blast qed text \If \P\ is a correct solution of the bin packing problem, \inv\<^sub>1\ holds for the partial solution, and every bin in \P\<^sub>1 \ wrap B\<^sub>1\ contains a large object, then the amount of bins in \P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}\ is a lower bound for the amount of bins in \P\.\ lemma L_bins_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}) \ card P" proof - note invrules = inv\<^sub>1E[OF assms(2)] have "\B \ {{v} |v. v \ V \ L}. B \ L \ {}" by blast with assms(3) have "P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" "\B \ P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}. B \ L \ {}" by blast+ from subset_bp_card[OF invrules(1) this] show ?thesis using L_lower_bound_card[OF assms(1)] by linarith qed text \If \P\ is a correct solution of the bin packing problem, then the sum of the weights of the objects is equal to the sum of the weights of the bins in \P\.\ lemma sum_Un_eq_sum_sum: assumes "bp P" shows "(\u \ U. w u) = (\B \ P. W B)" proof - have FINITE: "\B \ P. finite B" using assms by simp have DISJNT: "\A \ P. \B \ P. A \ B \ A \ B = {}" using bpE(1)[OF assms] unfolding pairwise_def disjnt_def . have "(\u \ (\P). w u) = (\B \ P. W B)" using sum.Union_disjoint[OF FINITE DISJNT] by auto then show ?thesis unfolding bpE(3)[OF assms] . qed text \If \P\ is a correct solution of the bin packing problem, then the sum of the weights of the items is a lower bound of amount of bins in \P\ multiplied by their maximum capacity.\ lemma sum_lower_bound_card: assumes "bp P" shows "(\u \ U. w u) \ c * card P" proof - have *: "\B \ P. 0 < W B \ W B \ c" using bpE(2-4)[OF assms] weight by (metis UnionI assms bp_bins_finite sum_pos) have "(\u \ U. w u) = (\B \ P. W B)" using sum_Un_eq_sum_sum[OF assms] . also have "... \ (\B \ P. c)" using sum_mono * by fastforce also have "... = c * card P" by simp finally show ?thesis . qed lemma bp_NE: assumes "bp P" shows "P \ {}" using U_NE bpE(3)[OF assms] by blast lemma sum_Un_ge: fixes f :: "_ \ real" assumes "finite M" "finite N" "\B \ M \ N. 0 < f B" shows "sum f M \ sum f (M \ N)" proof - have "0 \ sum f N - sum f (M \ N)" - using assms by (smt DiffD1 inf.cobounded2 UnCI sum_mono2) + using assms by (smt (verit) DiffD1 inf.cobounded2 UnCI sum_mono2) then have "sum f M \ sum f M + sum f N - sum f (M \ N)" by simp also have "... = sum f (M \ N)" using sum_Un[OF assms(1,2), symmetric] . finally show ?thesis . qed text \If \bij_exists\ holds, one can obtain a function which is bijective between the bins in \P\ and the objects in \V\ such that an object returned by the function would cause the bin to exceed its capacity.\ definition bij_exists :: "'a set set \ 'a set \ bool" where "bij_exists P V = (\f. bij_betw f P V \ (\B \ P. W B + w (f B) > c))" text \If \P\ is a functionally correct solution of the bin packing problem, \inv\<^sub>1\ holds for the partial solution, and such a bijective function exists between the bins in \P\<^sub>1\ and the objects in @{term "P\<^sub>2 \ wrap B\<^sub>2"}, the following strict lower bound can be shown:\ lemma P\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card P\<^sub>1 + 1 \ card P" proof (cases \P\<^sub>1 = {}\) case True have "finite P" using assms(1) by simp then have "1 \ card P" using bp_NE[OF assms(1)] by (metis Nat.add_0_right Suc_diff_1 Suc_le_mono card_gt_0_iff le0 mult_Suc_right nat_mult_1) then show ?thesis unfolding True by simp next note invrules = inv\<^sub>1E[OF assms(2)] case False obtain f where f_def: "bij_betw f P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" "\B \ P\<^sub>1. W B + w (f B) > c" using assms(3) unfolding bij_exists_def by blast have FINITE: "finite P\<^sub>1" "finite (P\<^sub>2 \ wrap B\<^sub>2)" "finite (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" "finite (wrap B\<^sub>1 \ {{v} |v. v \ V})" using inv\<^sub>1E(1)[OF assms(2)] bp_sol_finite by blast+ have F: "\B \ P\<^sub>2 \ wrap B\<^sub>2. finite B" using invrules(1) by simp have D: "\A \ P\<^sub>2 \ wrap B\<^sub>2. \B \ P\<^sub>2 \ wrap B\<^sub>2. A \ B \ A \ B = {}" using bpE(1)[OF invrules(1)] unfolding pairwise_def disjnt_def by auto have sum_eq: "W (\ (P\<^sub>2 \ wrap B\<^sub>2)) = (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.Union_disjoint[OF F D] by auto have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}. 0 < W B" using bpE(2,3)[OF invrules(1)] weight by (metis (no_types, lifting) UnionI bp_bins_finite invrules(1) sum_pos) then have "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ (wrap B\<^sub>1 \ {{v} |v. v \ V}). W B)" using sum_Un_ge[OF FINITE(3,4), of W] by blast - also have "... = (\B \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}. W B)" by (smt Un_assoc Un_commute) + also have "... = (\B \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}. W B)" by (smt (verit) Un_assoc Un_commute) also have "... = W U" using sum_Un_eq_sum_sum[OF invrules(1), symmetric] . finally have *: "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ W U" . \ \This follows from the fourth and final additional conjunct of \inv\<^sub>1\ and is necessary to combine the sums of the bins of the two partial solutions. This does not inherently follow from the union being a correct solution, as this need not be the case if \P\<^sub>1\ and \P\<^sub>2 \ wrap B\<^sub>2\ happened to be equal.\ have DISJNT: "P\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) by blast \ \This part of the proof is based on the proof on page 72 of the article \<^cite>\BerghammerR03\.\ have "c * card P\<^sub>1 = (\B \ P\<^sub>1. c)" by simp also have "... < (\B \ P\<^sub>1. W B + w (f B))" using f_def(2) sum_strict_mono[OF FINITE(1) False] by fastforce also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>1. w (f B))" by (simp add: Groups_Big.comm_monoid_add_class.sum.distrib) also have "... = (\B \ P\<^sub>1. W B) + W (\ (P\<^sub>2 \ wrap B\<^sub>2))" unfolding sum.reindex_bij_betw[OF f_def(1), of w] .. also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" unfolding sum_eq .. also have "... = (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.union_disjoint[OF FINITE(1,2) DISJNT, of W] by (simp add: Un_assoc) also have "... \ (\u \ U. w u)" using * . also have "... \ c * card P" using sum_lower_bound_card[OF assms(1)] . finally show ?thesis by (simp flip: of_nat_mult) qed text \As @{thm wrap_card} holds, it follows that the amount of bins in \P\<^sub>1 \ wrap B\<^sub>1\ are a lower bound for the amount of bins in \P\.\ lemma P\<^sub>1_B\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P" proof - have "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P\<^sub>1 + card (wrap B\<^sub>1)" using card_Un_le by blast also have "... \ card P\<^sub>1 + 1" using wrap_card by simp also have "... \ card P" using P\<^sub>1_lower_bound_card[OF assms] . finally show ?thesis . qed text \If \inv\<^sub>1\ holds, there are at most half as many bins in \P\<^sub>2\ as there are objects in \P\<^sub>2\, and we can again obtain a bijective function between the bins in \P\<^sub>1\ and the objects of the second partial solution, then the amount of bins in the second partial solution are a strict lower bound for half the bins of the first partial solution.\ lemma P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "2 * card P\<^sub>2 \ card (\P\<^sub>2)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card P\<^sub>1 + 1" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using invrules(4) by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using bp_bins_finite[OF invrules(1)] wrap_not_empty by blast have "finite P\<^sub>2" "finite (wrap B\<^sub>2)" using bp_sol_finite[OF invrules(1)] by blast+ have DISJNT2: "P\<^sub>2 \ wrap B\<^sub>2 = {}" unfolding wrap_def using \B\<^sub>2 \ P\<^sub>2\ by auto have "card (wrap B\<^sub>2) \ card B\<^sub>2" proof (cases \B\<^sub>2 = {}\) case False then have "1 \ card B\<^sub>2" by (simp add: leI \finite B\<^sub>2\) then show ?thesis using wrap_card[of B\<^sub>2] by linarith qed simp \ \This part of the proof is based on the proof on page 73 of the article \<^cite>\BerghammerR03\.\ from assms(2) have "2 * card P\<^sub>2 + 2 * card (wrap B\<^sub>2) \ card (\P\<^sub>2) + card (wrap B\<^sub>2) + 1" using wrap_card[of B\<^sub>2] by linarith then have "2 * (card P\<^sub>2 + card (wrap B\<^sub>2)) \ card (\P\<^sub>2) + card B\<^sub>2 + 1" using \card (wrap B\<^sub>2) \ card B\<^sub>2\ by simp then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\P\<^sub>2 \ B\<^sub>2) + 1" using card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\ DISJNT] and card_Un_disjoint[OF \finite P\<^sub>2\ \finite (wrap B\<^sub>2)\ DISJNT2] by argo then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\(P\<^sub>2 \ wrap B\<^sub>2)) + 1" by (cases \B\<^sub>2 = {}\) (auto simp: Un_commute) then show "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card P\<^sub>1 + 1" using assms(3) bij_betw_same_card unfolding bij_exists_def by metis qed subsubsection \Proving the Approximation Factor\ text \We define \inv\<^sub>2\ as it is defined in the article. These conjuncts allow us to prove the desired approximation factor.\ definition inv\<^sub>2 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ \\inv\<^sub>1\ holds for the partial solution\ \ (V \ L \ {} \ (\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {})) \ \If there are still large objects left, then every bin of the first partial solution must contain a large object\ \ bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2)) \ \There exists a bijective function between the bins of the first partial solution and the objects of the second one\ \ (2 * card P\<^sub>2 \ card (\P\<^sub>2)) \ \There are at most twice as many bins in \P\<^sub>2\ as there are objects in \P\<^sub>2\\" lemma inv\<^sub>2E: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "V \ L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" using assms unfolding inv\<^sub>2_def by blast+ lemma inv\<^sub>2I: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "V \ L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>2_def by blast text \If \P\ is a correct solution of the bin packing problem, \inv\<^sub>2\ holds for the partial solution, and there are no more small objects left to be distributed, then the amount of bins of the partial solution is no larger than \3 / 2\ of the amount of bins in \P\. This proof strongly follows the proof in \Theorem 4.1\ of the article \<^cite>\BerghammerR03\.\ lemma bin_packing_lower_bound_card: assumes "V \ S = {}" "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "bp P" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ 3 / 2 * card P" proof (cases \V = {}\) note invrules = inv\<^sub>2E[OF assms(2)] case True then have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" by simp also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using P\<^sub>1_B\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by simp also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith next note invrules = inv\<^sub>2E[OF assms(2)] case False have "U = S \ L" using S_def L_def by blast then have *: "V = V \ L" using bpE(3)[OF inv\<^sub>1E(1)[OF invrules(1)]] and assms(1) by blast with False have NE: "V \ L \ {}" by simp have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L} \ P\<^sub>2 \ wrap B\<^sub>2)" using * by (simp add: Un_commute Un_assoc) also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using L_bins_lower_bound_card[OF assms(3) invrules(1) invrules(2)[OF NE]] by linarith also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith qed text \We define \inv\<^sub>3\ as it is defined in the article. This final conjunct allows us to prove that the invariant will be maintained by the algorithm.\ definition inv\<^sub>3 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ B\<^sub>2 \ S" lemma inv\<^sub>3E: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "B\<^sub>2 \ S" using assms unfolding inv\<^sub>3_def by blast+ lemma inv\<^sub>3I: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "B\<^sub>2 \ S" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>3_def by blast lemma loop_init: "inv\<^sub>3 {} {} {} {} U" proof - have *: "inv\<^sub>1 {} {} {} {} U" unfolding bp_def partition_on_def pairwise_def wrap_def inv\<^sub>1_def using weight by auto have "bij_exists {} (\ ({} \ wrap {}))" using bij_betwI' unfolding bij_exists_def by fastforce from inv\<^sub>2I[OF * _ this] have "inv\<^sub>2 {} {} {} {} U" by auto from inv\<^sub>3I[OF this] show ?thesis by blast qed text \If \B\<^sub>1\ is empty and there are no large objects left, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ is initialized with \u \ V \ S\.\ lemma loop_stepA: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 = {}" "V \ L = {}" "u \ V \ S" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using S_def assms(2,4) by simp from assms(4) have "u \ V" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2) have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by simp have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L \ {}" using assms(3) by blast from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed text \If \B\<^sub>1\ is empty and there are large objects left, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ is initialized with \u \ V \ L\.\ lemma loop_stepB: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 = {}" "u \ V \ L" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using L_def weight assms(2,3) by simp from assms(3) have "u \ V" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2) have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by simp have "\B\P\<^sub>1. B \ L \ {}" using assms(3) invrules(2) by blast then have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L \ {}" using assms(3) by (metis Int_iff UnE empty_iff insertE singletonI wrap_not_empty) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed text \If \B\<^sub>1\ is not empty and \u \ V \ S\ does not exceed its maximum capacity, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ and \{u}\ are replaced with \B\<^sub>1 \ {u}\.\ lemma loop_stepC: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 \ {}" "u \ V \ S" "W B\<^sub>1 + w(u) \ c" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] from assms(3) have "u \ V" by blast from inv\<^sub>1_stepA[OF invrules(1) this assms(4)] have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" . have "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" using invrules(2) by blast then have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}). B \ L \ {}" by (metis Int_commute Un_empty_right Un_insert_right assms(2) disjoint_insert(2) insert_iff wrap_not_empty) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed text \If \B\<^sub>1\ is not empty and \u \ V \ S\ does exceed its maximum capacity but not the capacity of \B\<^sub>2\, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ is added to \P\<^sub>1\ and emptied, and \B\<^sub>2\ and \{u}\ are replaced with \B\<^sub>2 \ {u}\.\ lemma loop_stepD: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 \ {}" "u \ V \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) \ c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] from assms(3) have "u \ V" by blast from inv\<^sub>1_stepB[OF invrules(1) this assms(5)] have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" . have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast then have "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" by (metis Sup_empty Un_assoc Union_Un_distrib ccpo_Sup_singleton wrap_empty wrap_not_empty) also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding bij_exists_def by blast from inv\<^sub>2I[OF 1 2 3] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" using invrules(4) by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] assms(3) by blast qed text \If the maximum capacity of \B\<^sub>2\ is exceeded by \u \ V \ S\, then \B\<^sub>2\ must contain at least two objects.\ lemma B\<^sub>2_at_least_two_objects: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V \ S" "W B\<^sub>2 + w(u) > c" shows "2 \ card B\<^sub>2" proof (rule ccontr, simp add: not_le) have FINITE: "finite B\<^sub>2" using inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF inv\<^sub>3E(1)[OF assms(1)]]] by (metis (no_types, lifting) Finite_Set.finite.simps U_Finite Union_Un_distrib bpE(3) ccpo_Sup_singleton finite_Un wrap_not_empty) assume "card B\<^sub>2 < 2" then consider (0) "card B\<^sub>2 = 0" | (1) "card B\<^sub>2 = 1" by linarith then show False proof cases case 0 then have "B\<^sub>2 = {}" using FINITE by simp then show ?thesis using assms(2,3) S_def by simp next case 1 then obtain v where "B\<^sub>2 = {v}" using card_1_singletonE by auto with inv\<^sub>3E(2)[OF assms(1)] have "2 * w v \ c" using S_def by simp moreover from \B\<^sub>2 = {v}\ have "W B\<^sub>2 = w v" by simp ultimately show ?thesis using assms(2,3) S_def by simp qed qed text \If \B\<^sub>1\ is not empty and \u \ V \ S\ exceeds the maximum capacity of both \B\<^sub>1\ and \B\<^sub>2\, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ and \B\<^sub>2\ are added to \P\<^sub>1\ and \P\<^sub>2\ respectively, emptied, and \B\<^sub>2\ initialized with \u\.\ lemma loop_stepE: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 \ {}" "u \ V \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) > c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] from assms(3) have "u \ V" by blast from inv\<^sub>1_stepC[OF invrules(1) this] have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" . have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast have "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" unfolding wrap_def by simp also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding bij_exists_def by blast have 4: "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card (\ (P\<^sub>2 \ wrap B\<^sub>2))" proof - note bprules = bpE[OF inv\<^sub>1E(1)[OF invrules(1)]] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using inv\<^sub>1E(4)[OF invrules(1)] by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using inv\<^sub>1E(1)[OF invrules(1)] bp_bins_finite wrap_not_empty by blast have "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ 2 * (card P\<^sub>2 + card (wrap B\<^sub>2))" using card_Un_le[of P\<^sub>2 \wrap B\<^sub>2\] by simp also have "... \ 2 * card P\<^sub>2 + 2" using wrap_card by auto also have "... \ card (\ P\<^sub>2) + 2" using invrules(4) by simp also have "... \ card (\ P\<^sub>2) + card B\<^sub>2" using B\<^sub>2_at_least_two_objects[OF assms(1,3,5)] by simp also have "... = card (\ (P\<^sub>2 \ {B\<^sub>2}))" using DISJNT card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\] by (simp add: Un_commute) also have "... = card (\ (P\<^sub>2 \ wrap B\<^sub>2))" by (cases \B\<^sub>2 = {}\) auto finally show ?thesis . qed from inv\<^sub>2I[OF 1 2 3 4] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" . from inv\<^sub>3I[OF this] show ?thesis using assms(3) by blast qed text \The bin packing algorithm as it is proposed in the article \<^cite>\BerghammerR03\. \P\ will not only be a correct solution of the bin packing problem, but the amount of bins will be a lower bound for \3 / 2\ of the amount of bins of any correct solution \Q\, and thus guarantee an approximation factor of \3 / 2\ for the optimum.\ lemma bp_approx: "VARS P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u {True} P\<^sub>1 := {}; P\<^sub>2 := {}; B\<^sub>1 := {}; B\<^sub>2 := {}; V := U; WHILE V \ S \ {} INV {inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V} DO IF B\<^sub>1 \ {} THEN u := (SOME u. u \ V \ S) ELSE IF V \ L \ {} THEN u := (SOME u. u \ V \ L) ELSE u := (SOME u. u \ V \ S) FI FI; V := V - {u}; IF W(B\<^sub>1) + w(u) \ c THEN B\<^sub>1 := B\<^sub>1 \ {u} ELSE IF W(B\<^sub>2) + w(u) \ c THEN B\<^sub>2 := B\<^sub>2 \ {u} ELSE P\<^sub>2 := P\<^sub>2 \ wrap B\<^sub>2; B\<^sub>2 := {u} FI; P\<^sub>1 := P\<^sub>1 \ wrap B\<^sub>1; B\<^sub>1 := {} FI OD; P := P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} | v. v \ V} {bp P \ (\Q. bp Q \ card P \ 3 / 2 * card Q)}" proof (vcg, goal_cases) case (1 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then show ?case by (simp add: loop_init) next case (2 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then have INV: "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" .. let ?s = "SOME u. u \ V \ S" let ?l = "SOME u. u \ V \ L" have LIN: "V \ L \ {} \ ?l \ V \ L" using some_in_eq by metis then have LWEIGHT: "V \ L \ {} \ w ?l \ c" using L_def weight by blast from 2 have "V \ S \ {}" .. then have IN: "?s \ V \ S" using some_in_eq by metis then have "w ?s \ c" using S_def by simp then show ?case using LWEIGHT loop_stepA[OF INV _ _ IN] loop_stepB[OF INV _ LIN] loop_stepC[OF INV _ IN] and loop_stepD[OF INV _ IN] loop_stepE[OF INV _ IN] by (cases \B\<^sub>1 = {}\, cases \V \ L = {}\) auto next case (3 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then have INV: "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and EMPTY: "V \ S = {}" by blast+ from inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF inv\<^sub>3E(1)[OF INV]]] and bin_packing_lower_bound_card[OF EMPTY inv\<^sub>3E(1)[OF INV]] show ?case by blast qed end (* BinPacking *) subsection \The Full Linear Time Version of the Proposed Algorithm\ text \Finally, we prove the Algorithm proposed on page 78 of the article \<^cite>\BerghammerR03\. This version generates the S and L sets beforehand and uses them directly to calculate the solution, thus removing the need for intersection operations, and ensuring linear time if we can perform \insertion, removal, and selection of an element, the union of two sets, and the emptiness test in constant time\ \<^cite>\BerghammerR03\.\ locale BinPacking_Complete = fixes U :: "'a set" \ \A finite, non-empty set of objects\ and w :: "'a \ real" \ \A mapping from objects to their respective weights (positive real numbers)\ and c :: nat \ \The maximum capacity of a bin (as a natural number)\ assumes weight: "\u \ U. 0 < w(u) \ w(u) \ c" and U_Finite: "finite U" and U_NE: "U \ {}" begin text \The correctness proofs will be identical to the ones of the simplified algorithm.\ abbreviation W :: "'a set \ real" where "W B \ (\u \ B. w(u))" definition bp :: "'a set set \ bool" where "bp P \ partition_on U P \ (\B \ P. W(B) \ c)" lemma bpE: assumes "bp P" shows "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" using assms unfolding bp_def partition_on_def by blast+ lemma bpI: assumes "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" shows "bp P" using assms unfolding bp_def partition_on_def by blast definition inv\<^sub>1 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ \A correct solution to the bin packing problem\ \ \(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V \ \The partial solution does not contain objects that have not yet been assigned\ \ B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ \\B\<^sub>1\ is distinct from all the other bins\ \ B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2) \ \\B\<^sub>2\ is distinct from all the other bins\ \ (P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {} \ \The first and second partial solutions are disjoint from each other.\" lemma inv\<^sub>1E: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using assms unfolding inv\<^sub>1_def by auto lemma inv\<^sub>1I: assumes "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>1_def by blast lemma wrap_Un [simp]: "wrap (M \ {x}) = {M \ {x}}" unfolding wrap_def by simp lemma wrap_empty [simp]: "wrap {} = {}" unfolding wrap_def by simp lemma wrap_not_empty [simp]: "M \ {} \ wrap M = {M}" unfolding wrap_def by simp lemma inv\<^sub>1_stepA: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W(B\<^sub>1) + w(u) \ c" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] \ \Rule 1: Pairwise Disjoint\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ ({{u}} \ {{v} |v. v \ V - {u}}))" using bprules(1) assms(2) by simp then have "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" by (simp add: Un_commute) then have assm: "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc) have "pairwise disjnt ({B\<^sub>1 \ {u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>1 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>1} \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,3) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>1" using bprules(3) U_Finite by (cases \B\<^sub>1 = {}\) auto then have "W (B\<^sub>1 \ {u}) \ W B\<^sub>1 + w u" using \0 < w u\ by (cases \u \ B\<^sub>1\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>1 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>1 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" by blast from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = U - (V - {u})" unfolding L R invrules(2) .. have 3: "B\<^sub>1 \ {u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2" using NOTIN by auto have 4: "B\<^sub>2 \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2" using invrules(4) NOTIN unfolding wrap_def by fastforce have 5: "(P\<^sub>1 \ wrap (B\<^sub>1 \ {u})) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) NOTIN unfolding wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed lemma inv\<^sub>1_stepB: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W B\<^sub>2 + w u \ c" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}})" using bprules(1) assms(2) by simp then have assm: "pairwise disjnt (wrap B\<^sub>2 \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc Un_commute) have "pairwise disjnt ({B\<^sub>2 \ {u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>2 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>2} \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,4) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>2" using bprules(3) U_Finite by (cases \B\<^sub>2 = {}\) auto then have "W (B\<^sub>2 \ {u}) \ W B\<^sub>2 + w u" using \0 < w u\ by (cases \u \ B\<^sub>2\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>2 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>2 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}. W B \ c" by auto from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = U - (V - {u})" unfolding L R using invrules(2) by simp have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})" using bpE(2)[OF 1] by simp have 4: "B\<^sub>2 \ {u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = {}" using invrules(5) NOTIN unfolding wrap_empty wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed lemma inv\<^sub>1_stepC: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] \ \Rule 1-4: Correct Bin Packing\ have "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}" by (metis (no_types, lifting) Un_assoc Un_empty_right insert_not_empty wrap_empty wrap_not_empty) also have "... = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" using assms(2) by auto finally have EQ: "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" . from invrules(1) have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}})" unfolding EQ . \ \Auxiliary information is preserved\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "u \ U" using assms(2) bpE(3)[OF invrules(1)] by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = U - (V - {u})" unfolding L R using invrules(2) by auto have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}" using bpE(2)[OF 1] by simp have 4: "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}) = {}" using invrules(5) NOTIN unfolding wrap_def by force from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \From this point onward, we will require a different approach for proving lower bounds. Instead of fixing and assuming the definitions of the \S\ and \L\ sets, we will introduce the abbreviations \S\<^sub>U\ and \L\<^sub>U\ for any occurrences of the original \S\ and \L\ sets. The union of \S\ and \L\ can be interpreted as \V\. As a result, occurrences of \V \ S\ become \(S \ L) \ S = S\, and \V \ L\ become \(S \ L) \ L = L\. Occurrences of these sets will have to be replaced appropriately.\ abbreviation S\<^sub>U where "S\<^sub>U \ {u \ U. w u \ c / 2}" abbreviation L\<^sub>U where "L\<^sub>U \ {u \ U. c / 2 < w u}" text \As we will remove elements from \S\ and \L\, we will only be able to show that they remain subsets of \S\<^sub>U\ and \L\<^sub>U\ respectively.\ abbreviation SL where "SL S L \ S \ S\<^sub>U \ L \ L\<^sub>U" lemma bp_bins_finite [simp]: assumes "bp P" shows "\B \ P. finite B" using bpE(3)[OF assms] U_Finite by (meson Sup_upper finite_subset) lemma bp_sol_finite [simp]: assumes "bp P" shows "finite P" using bpE(3)[OF assms] U_Finite by (simp add: finite_UnionD) lemma only_one_L_per_bin: assumes "bp P" "B \ P" shows "\x \ B. \y \ B. x \ y \ x \ L\<^sub>U \ y \ L\<^sub>U" proof (rule ccontr, simp) assume "\x\B. \y\B. x \ y \ y \ U \ x \ U \ real c < w x * 2 \ real c < w y * 2" then obtain x y where *: "x \ B" "y \ B" "x \ y" "x \ L\<^sub>U" "y \ L\<^sub>U" by auto then have "c < w x + w y" by force have "finite B" using assms by simp have "y \ B - {x}" using *(2,3) by blast have "W B = W (B - {x}) + w x" using *(1) \finite B\ by (simp add: sum.remove) also have "... = W (B - {x} - {y}) + w x + w y" using \y \ B - {x}\ \finite B\ by (simp add: sum.remove) finally have *: "W B = W (B - {x} - {y}) + w x + w y" . have "\u \ B. 0 < w u" using bpE(3)[OF assms(1)] assms(2) weight by blast - then have "0 \ W (B - {x} - {y})" by (smt DiffD1 sum_nonneg) + then have "0 \ W (B - {x} - {y})" by (smt (verit) DiffD1 sum_nonneg) with * have "c < W B" using \c < w x + w y\ by simp then show False using bpE(4)[OF assms(1)] assms(2) by fastforce qed lemma L_lower_bound_card: assumes "bp P" shows "card L\<^sub>U \ card P" proof - have "\x \ L\<^sub>U. \B \ P. x \ B" using bpE(3)[OF assms] by blast then obtain f where f_def: "\u \ L\<^sub>U. u \ f u \ f u \ P" by metis then have "inj_on f L\<^sub>U" unfolding inj_on_def using only_one_L_per_bin[OF assms] by blast then have card_eq: "card L\<^sub>U = card (f ` L\<^sub>U)" by (simp add: card_image) have "f ` L\<^sub>U \ P" using f_def by blast moreover have "finite P" using assms by simp ultimately have "card (f ` L\<^sub>U) \ card P" by (simp add: card_mono) then show ?thesis unfolding card_eq . qed lemma subset_bp_card: assumes "bp P" "M \ P" "\B \ M. B \ L\<^sub>U \ {}" shows "card M \ card L\<^sub>U" proof - have "\B \ M. \u \ L\<^sub>U. u \ B" using assms(3) by fast then have "\f. \B \ M. f B \ L\<^sub>U \ f B \ B" by metis then obtain f where f_def: "\B \ M. f B \ L\<^sub>U \ f B \ B" .. have "inj_on f M" proof (rule ccontr) assume "\ inj_on f M" then have "\x \ M. \y \ M. x \ y \ f x = f y" unfolding inj_on_def by blast then obtain x y where *: "x \ M" "y \ M" "x \ y" "f x = f y" by blast then have "\u. u \ x \ u \ y" using f_def by metis then have "x \ y \ {}" by blast moreover have "pairwise disjnt M" using pairwise_subset[OF bpE(1)[OF assms(1)] assms(2)] . ultimately show False using * unfolding pairwise_def disjnt_def by simp qed moreover have "finite L\<^sub>U" using U_Finite by auto moreover have "f ` M \ L\<^sub>U" using f_def by blast ultimately show ?thesis using card_inj_on_le by blast qed lemma L_bins_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" and SL_def: "SL S L" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}) \ card P" proof - note invrules = inv\<^sub>1E[OF assms(2)] have "\B \ {{v} |v. v \ L}. B \ L\<^sub>U \ {}" using SL_def by blast with assms(3) have "P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}" "\B \ P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}. B \ L\<^sub>U \ {}" by blast+ from subset_bp_card[OF invrules(1) this] show ?thesis using L_lower_bound_card[OF assms(1)] by linarith qed lemma sum_Un_eq_sum_sum: assumes "bp P" shows "(\u \ U. w u) = (\B \ P. W B)" proof - have FINITE: "\B \ P. finite B" using assms by simp have DISJNT: "\A \ P. \B \ P. A \ B \ A \ B = {}" using bpE(1)[OF assms] unfolding pairwise_def disjnt_def . have "(\u \ (\P). w u) = (\B \ P. W B)" using sum.Union_disjoint[OF FINITE DISJNT] by auto then show ?thesis unfolding bpE(3)[OF assms] . qed lemma sum_lower_bound_card: assumes "bp P" shows "(\u \ U. w u) \ c * card P" proof - have *: "\B \ P. 0 < W B \ W B \ c" using bpE(2-4)[OF assms] weight by (metis UnionI assms bp_bins_finite sum_pos) have "(\u \ U. w u) = (\B \ P. W B)" using sum_Un_eq_sum_sum[OF assms] . also have "... \ (\B \ P. c)" using sum_mono * by fastforce also have "... = c * card P" by simp finally show ?thesis . qed lemma bp_NE: assumes "bp P" shows "P \ {}" using U_NE bpE(3)[OF assms] by blast lemma sum_Un_ge: fixes f :: "_ \ real" assumes "finite M" "finite N" "\B \ M \ N. 0 < f B" shows "sum f M \ sum f (M \ N)" proof - have "0 \ sum f N - sum f (M \ N)" - using assms by (smt DiffD1 inf.cobounded2 UnCI sum_mono2) + using assms by (smt (verit) DiffD1 inf.cobounded2 UnCI sum_mono2) then have "sum f M \ sum f M + sum f N - sum f (M \ N)" by simp also have "... = sum f (M \ N)" using sum_Un[OF assms(1,2), symmetric] . finally show ?thesis . qed definition bij_exists :: "'a set set \ 'a set \ bool" where "bij_exists P V = (\f. bij_betw f P V \ (\B \ P. W B + w (f B) > c))" lemma P\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card P\<^sub>1 + 1 \ card P" proof (cases \P\<^sub>1 = {}\) case True have "finite P" using assms(1) by simp then have "1 \ card P" using bp_NE[OF assms(1)] by (metis Nat.add_0_right Suc_diff_1 Suc_le_mono card_gt_0_iff le0 mult_Suc_right nat_mult_1) then show ?thesis unfolding True by simp next note invrules = inv\<^sub>1E[OF assms(2)] case False obtain f where f_def: "bij_betw f P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" "\B \ P\<^sub>1. W B + w (f B) > c" using assms(3) unfolding bij_exists_def by blast have FINITE: "finite P\<^sub>1" "finite (P\<^sub>2 \ wrap B\<^sub>2)" "finite (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" "finite (wrap B\<^sub>1 \ {{v} |v. v \ S \ L})" using inv\<^sub>1E(1)[OF assms(2)] bp_sol_finite by blast+ have F: "\B \ P\<^sub>2 \ wrap B\<^sub>2. finite B" using invrules(1) by simp have D: "\A \ P\<^sub>2 \ wrap B\<^sub>2. \B \ P\<^sub>2 \ wrap B\<^sub>2. A \ B \ A \ B = {}" using bpE(1)[OF invrules(1)] unfolding pairwise_def disjnt_def by auto have sum_eq: "W (\ (P\<^sub>2 \ wrap B\<^sub>2)) = (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.Union_disjoint[OF F D] by auto have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}. 0 < W B" using bpE(2,3)[OF invrules(1)] weight by (metis (no_types, lifting) UnionI bp_bins_finite invrules(1) sum_pos) then have "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ (wrap B\<^sub>1 \ {{v} |v. v \ S \ L}). W B)" using sum_Un_ge[OF FINITE(3,4), of W] by blast - also have "... = (\B \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}. W B)" by (smt Un_assoc Un_commute) + also have "... = (\B \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}. W B)" by (smt (verit) Un_assoc Un_commute) also have "... = W U" using sum_Un_eq_sum_sum[OF invrules(1), symmetric] . finally have *: "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ W U" . have DISJNT: "P\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) by blast \ \This part of the proof is based on the proof on page 72 of the article \<^cite>\BerghammerR03\.\ have "c * card P\<^sub>1 = (\B \ P\<^sub>1. c)" by simp also have "... < (\B \ P\<^sub>1. W B + w (f B))" using f_def(2) sum_strict_mono[OF FINITE(1) False] by fastforce also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>1. w (f B))" by (simp add: Groups_Big.comm_monoid_add_class.sum.distrib) also have "... = (\B \ P\<^sub>1. W B) + W (\ (P\<^sub>2 \ wrap B\<^sub>2))" unfolding sum.reindex_bij_betw[OF f_def(1), of w] .. also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" unfolding sum_eq .. also have "... = (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.union_disjoint[OF FINITE(1,2) DISJNT, of W] by (simp add: Un_assoc) also have "... \ (\u \ U. w u)" using * . also have "... \ c * card P" using sum_lower_bound_card[OF assms(1)] . finally show ?thesis by (simp flip: of_nat_mult) qed lemma P\<^sub>1_B\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P" proof - have "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P\<^sub>1 + card (wrap B\<^sub>1)" using card_Un_le by blast also have "... \ card P\<^sub>1 + 1" using wrap_card by simp also have "... \ card P" using P\<^sub>1_lower_bound_card[OF assms] . finally show ?thesis . qed lemma P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "2 * card P\<^sub>2 \ card (\P\<^sub>2)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card P\<^sub>1 + 1" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using invrules(4) by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using bp_bins_finite[OF invrules(1)] wrap_not_empty by blast have "finite P\<^sub>2" "finite (wrap B\<^sub>2)" using bp_sol_finite[OF invrules(1)] by blast+ have DISJNT2: "P\<^sub>2 \ wrap B\<^sub>2 = {}" unfolding wrap_def using \B\<^sub>2 \ P\<^sub>2\ by auto have "card (wrap B\<^sub>2) \ card B\<^sub>2" proof (cases \B\<^sub>2 = {}\) case False then have "1 \ card B\<^sub>2" by (simp add: leI \finite B\<^sub>2\) then show ?thesis using wrap_card[of B\<^sub>2] by linarith qed simp \ \This part of the proof is based on the proof on page 73 of the article \<^cite>\BerghammerR03\.\ from assms(2) have "2 * card P\<^sub>2 + 2 * card (wrap B\<^sub>2) \ card (\P\<^sub>2) + card (wrap B\<^sub>2) + 1" using wrap_card[of B\<^sub>2] by linarith then have "2 * (card P\<^sub>2 + card (wrap B\<^sub>2)) \ card (\P\<^sub>2) + card B\<^sub>2 + 1" using \card (wrap B\<^sub>2) \ card B\<^sub>2\ by simp then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\P\<^sub>2 \ B\<^sub>2) + 1" using card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\ DISJNT] and card_Un_disjoint[OF \finite P\<^sub>2\ \finite (wrap B\<^sub>2)\ DISJNT2] by argo then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\(P\<^sub>2 \ wrap B\<^sub>2)) + 1" by (cases \B\<^sub>2 = {}\) (auto simp: Un_commute) then show "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card P\<^sub>1 + 1" using assms(3) bij_betw_same_card unfolding bij_exists_def by metis qed text \We add \SL S L\ to \inv\<^sub>2\ to ensure that the \S\ and \L\ sets only contain objects with correct weights.\ definition inv\<^sub>2 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L) \ \\inv\<^sub>1\ holds for the partial solution\ \ (L \ {} \ (\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {})) \ \If there are still large objects left, then every bin of the first partial solution must contain a large object\ \ bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2)) \ \There exists a bijective function between the bins of the first partial solution and the objects of the second one\ \ (2 * card P\<^sub>2 \ card (\P\<^sub>2)) \ \There are at most twice as many bins in \P\<^sub>2\ as there are objects in \P\<^sub>2\\ \ SL S L \ \\S\ and \L\ are subsets of \S\<^sub>U\ and \L\<^sub>U\\" lemma inv\<^sub>2E: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" and "L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" and "SL S L" using assms unfolding inv\<^sub>2_def by blast+ lemma inv\<^sub>2I: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" and "L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" and "SL S L" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" using assms unfolding inv\<^sub>2_def by blast lemma bin_packing_lower_bound_card: assumes "S = {}" "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "bp P" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}) \ 3 / 2 * card P" proof (cases \L = {}\) note invrules = inv\<^sub>2E[OF assms(2)] case True then have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" using assms(1) by simp also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using P\<^sub>1_B\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by simp also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith next note invrules = inv\<^sub>2E[OF assms(2)] case False have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L} \ P\<^sub>2 \ wrap B\<^sub>2)" using assms(1) by (simp add: Un_commute Un_assoc) also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using L_bins_lower_bound_card[OF assms(3) invrules(1) invrules(2)[OF False] invrules(5)] by linarith also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith qed definition inv\<^sub>3 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ B\<^sub>2 \ S\<^sub>U" lemma inv\<^sub>3E: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" and "B\<^sub>2 \ S\<^sub>U" using assms unfolding inv\<^sub>3_def by blast+ lemma inv\<^sub>3I: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" and "B\<^sub>2 \ S\<^sub>U" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" using assms unfolding inv\<^sub>3_def by blast lemma loop_init: "inv\<^sub>3 {} {} {} {} S\<^sub>U L\<^sub>U" proof - have "S\<^sub>U \ L\<^sub>U = U" by auto then have *: "inv\<^sub>1 {} {} {} {} (S\<^sub>U \ L\<^sub>U)" unfolding bp_def partition_on_def pairwise_def wrap_def inv\<^sub>1_def using weight by auto have "bij_exists {} (\ ({} \ wrap {}))" using bij_betwI' unfolding bij_exists_def by fastforce from inv\<^sub>2I[OF * _ this] have "inv\<^sub>2 {} {} {} {} S\<^sub>U L\<^sub>U" by auto from inv\<^sub>3I[OF this] show ?thesis by blast qed lemma loop_stepA: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 = {}" "L = {}" "u \ S" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using invrules(5) assms(2,4) by fastforce from assms(4) have "u \ S \ L" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2,3) have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S - {u} \ L)" by simp have 2: "L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L\<^sub>U \ {}" using assms(3) by blast from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S - {u}) L" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed lemma loop_stepB: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 = {}" "u \ L" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 S (L - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using weight invrules(5) assms(2,3) by fastforce \ \This observation follows from the fact that the \S\ and \L\ sets have to be disjoint from each other, and allows us to reuse our proofs of the preservation of \inv\<^sub>1\ by simply replacing \V\ with \S \ L\\ have *: "S \ L - {u} = S \ (L - {u})" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2) * have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S \ (L - {u}))" by simp have "\B\P\<^sub>1. B \ L\<^sub>U \ {}" "{u} \ L\<^sub>U \ {}" using assms(3) invrules(2,5) by blast+ then have 2: "L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L\<^sub>U \ {}" using assms(3) by (metis (full_types) Un_iff empty_iff insert_iff wrap_not_empty) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 S (L - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed lemma loop_stepC: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 \ {}" "u \ S" "W B\<^sub>1 + w(u) \ c" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] \ \Same approach, but removing \{u}\ from \S\ instead of \L\\ have *: "S \ L - {u} = (S - {u}) \ L" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepA[OF invrules(1) this assms(4)] * have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (S - {u} \ L)" by simp have "L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" using invrules(2) by blast then have 2: "L \ {} \ \B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}). B \ L\<^sub>U \ {}" - by (smt Int_insert_left Un_empty_right Un_iff Un_insert_right assms(2) insert_not_empty singletonD singletonI wrap_def) + by (smt (verit) Int_insert_left Un_empty_right Un_iff Un_insert_right assms(2) insert_not_empty singletonD singletonI wrap_def) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (S - {u}) L" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed lemma loop_stepD: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 \ {}" "u \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) \ c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have *: "S \ L - {u} = (S - {u}) \ L" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepB[OF invrules(1) this assms(5)] * have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (S - {u} \ L)" by simp have 2: "L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L\<^sub>U \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast then have "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" by (metis Sup_empty Un_assoc Union_Un_distrib ccpo_Sup_singleton wrap_empty wrap_not_empty) also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding bij_exists_def by blast from inv\<^sub>2I[OF 1 2 3] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (S - {u}) L" using invrules(4,5) by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] assms(3) invrules(5) by blast qed lemma B\<^sub>2_at_least_two_objects: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "u \ S" "W B\<^sub>2 + w(u) > c" shows "2 \ card B\<^sub>2" proof (rule ccontr, simp add: not_le) have FINITE: "finite B\<^sub>2" using inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF inv\<^sub>3E(1)[OF assms(1)]]] by (metis (no_types, lifting) Finite_Set.finite.simps U_Finite Union_Un_distrib bpE(3) ccpo_Sup_singleton finite_Un wrap_not_empty) assume "card B\<^sub>2 < 2" then consider (0) "card B\<^sub>2 = 0" | (1) "card B\<^sub>2 = 1" by linarith then show False proof cases case 0 then have "B\<^sub>2 = {}" using FINITE by simp then show ?thesis using assms(2,3) inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF assms(1)]] by force next case 1 then obtain v where "B\<^sub>2 = {v}" using card_1_singletonE by auto with inv\<^sub>3E(2)[OF assms(1)] have "2 * w v \ c" using inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF assms(1)]] by simp moreover from \B\<^sub>2 = {v}\ have "W B\<^sub>2 = w v" by simp ultimately show ?thesis using assms(2,3) inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF assms(1)]] by force qed qed lemma loop_stepE: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 \ {}" "u \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) > c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have *: "S \ L - {u} = (S - {u}) \ L" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepC[OF invrules(1) this] * have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (S - {u} \ L)" by simp have 2: "L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L\<^sub>U \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast have "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" unfolding wrap_def by simp also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding bij_exists_def by blast have 4: "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card (\ (P\<^sub>2 \ wrap B\<^sub>2))" proof - note bprules = bpE[OF inv\<^sub>1E(1)[OF invrules(1)]] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using inv\<^sub>1E(4)[OF invrules(1)] by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using inv\<^sub>1E(1)[OF invrules(1)] bp_bins_finite wrap_not_empty by blast have "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ 2 * (card P\<^sub>2 + card (wrap B\<^sub>2))" using card_Un_le[of P\<^sub>2 \wrap B\<^sub>2\] by simp also have "... \ 2 * card P\<^sub>2 + 2" using wrap_card by auto also have "... \ card (\ P\<^sub>2) + 2" using invrules(4) by simp also have "... \ card (\ P\<^sub>2) + card B\<^sub>2" using B\<^sub>2_at_least_two_objects[OF assms(1,3,5)] by simp also have "... = card (\ (P\<^sub>2 \ {B\<^sub>2}))" using DISJNT card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\] by (simp add: Un_commute) also have "... = card (\ (P\<^sub>2 \ wrap B\<^sub>2))" by (cases \B\<^sub>2 = {}\) auto finally show ?thesis . qed from inv\<^sub>2I[OF 1 2 3 4] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (S - {u}) L" using invrules(5) by blast from inv\<^sub>3I[OF this] show ?thesis using assms(3) invrules(5) by blast qed text \The bin packing algorithm as it is proposed on page 78 of the article \<^cite>\BerghammerR03\. \P\ will not only be a correct solution of the bin packing problem, but the amount of bins will be a lower bound for \3 / 2\ of the amount of bins of any correct solution \Q\, and thus guarantee an approximation factor of \3 / 2\ for the optimum.\ lemma bp_approx: "VARS P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u {True} S := {}; L:= {}; V := U; WHILE V \ {} INV {V \ U \ S = {u \ U - V. w(u) \ c / 2} \ L = {u \ U - V. c / 2 < w(u)}} DO u := (SOME u. u \ V); IF w(u) \ c / 2 THEN S := S \ {u} ELSE L := L \ {u} FI; V := V - {u} OD; P\<^sub>1 := {}; P\<^sub>2 := {}; B\<^sub>1 := {}; B\<^sub>2 := {}; WHILE S \ {} INV {inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L} DO IF B\<^sub>1 \ {} THEN u := (SOME u. u \ S); S := S - {u} ELSE IF L \ {} THEN u := (SOME u. u \ L); L := L - {u} ELSE u := (SOME u. u \ S); S := S - {u} FI FI; IF W(B\<^sub>1) + w(u) \ c THEN B\<^sub>1 := B\<^sub>1 \ {u} ELSE IF W(B\<^sub>2) + w(u) \ c THEN B\<^sub>2 := B\<^sub>2 \ {u} ELSE P\<^sub>2 := P\<^sub>2 \ wrap B\<^sub>2; B\<^sub>2 := {u} FI; P\<^sub>1 := P\<^sub>1 \ wrap B\<^sub>1; B\<^sub>1 := {} FI OD; P := P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2; V := L; WHILE V \ {} INV {S = {} \ inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ V \ L \ P = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v}|v. v \ L - V}} DO u := (SOME u. u \ V); P := P \ {{u}}; V := V - {u} OD {bp P \ (\Q. bp Q \ card P \ 3 / 2 * card Q)}" proof (vcg, goal_cases) case (1 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case by blast next case (2 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case by (auto simp: some_in_eq) next case (3 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case using loop_init by force next case (4 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then have INV: "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" .. let ?s = "SOME u. u \ S" let ?l = "SOME u. u \ L" note SL_def = inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF INV]] have LIN: "L \ {} \ ?l \ L" using some_in_eq by metis then have LWEIGHT: "L \ {} \ w ?l \ c" using weight SL_def by blast from 4 have "S \ {}" .. then have IN: "?s \ S" using some_in_eq by metis then have "w ?s \ c" using SL_def by auto then show ?case using LWEIGHT loop_stepA[OF INV _ _ IN] loop_stepB[OF INV _ LIN] loop_stepC[OF INV _ IN] and loop_stepD[OF INV _ IN] loop_stepE[OF INV _ IN] by (cases \B\<^sub>1 = {}\, cases \L = {}\) auto next case (5 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case by blast next case (6 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then have *: "(SOME u. u \ V) \ V" "(SOME u. u \ V) \ L" by (auto simp add: some_in_eq) then have "P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ L - (V - {SOME u. u \ V})} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ L - V \ {SOME u. u \ V}}" by blast with 6 * show ?case by blast next case (7 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then have *: "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" using inv\<^sub>3E(1) by blast from inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF *]] 7 have "bp P" by fastforce with bin_packing_lower_bound_card[OF _ *] 7 show ?case by fastforce qed end (* BinPacking_Complete *) end (* Theory *) \ No newline at end of file diff --git a/thys/Approximation_Algorithms/Center_Selection.thy b/thys/Approximation_Algorithms/Center_Selection.thy --- a/thys/Approximation_Algorithms/Center_Selection.thy +++ b/thys/Approximation_Algorithms/Center_Selection.thy @@ -1,457 +1,457 @@ (* Author: Ujkan Sulejmani *) section \Center Selection\ theory Center_Selection imports Complex_Main "HOL-Hoare.Hoare_Logic" begin text \The Center Selection (or metric k-center) problem. Given a set of \textit{sites} \S\ in a metric space, find a subset \C \ S\ that minimizes the maximal distance from any \s \ S\ to some \c \ C\. This theory presents a verified 2-approximation algorithm. It is based on Section 11.2 in the book by Kleinberg and Tardos \<^cite>\"KleinbergT06"\. In contrast to the proof in the book, our proof is a standard invariant proof.\ locale Center_Selection = fixes S :: "('a :: metric_space) set" and k :: nat assumes finite_sites: "finite S" and non_empty_sites: "S \ {}" and non_zero_k: "k > 0" begin definition distance :: "('a::metric_space) set \ ('a::metric_space) \ real" where "distance C s = Min (dist s ` C)" definition radius :: "('a :: metric_space) set \ real" where "radius C = Max (distance C ` S)" lemma distance_mono: assumes "C\<^sub>1 \ C\<^sub>2" and "C\<^sub>1 \ {}" and "finite C\<^sub>2" shows "distance C\<^sub>1 s \ distance C\<^sub>2 s" by (simp add: Min.subset_imp assms distance_def image_mono) lemma finite_distances: "finite (distance C ` S)" using finite_sites by simp lemma non_empty_distances: "distance C ` S \ {}" using non_empty_sites by simp lemma radius_contained: "radius C \ distance C ` S" using finite_distances non_empty_distances Max_in radius_def by simp lemma radius_def2: "\s \ S. distance C s = radius C" using radius_contained image_iff by metis lemma dist_lemmas_aux: assumes "finite C" and "C \ {}" shows "finite (dist s ` C)" and "finite (dist s ` C) \ distance C s \ dist s ` C" and "distance C s \ dist s ` C \ \c \ C. dist s c = distance C s" and "\c \ C. dist s c = distance C s \ distance C s \ 0" proof show "finite C" using assms(1) by simp next assume "finite (dist s ` C)" then show "distance C s \ dist s ` C" using distance_def eq_Min_iff assms(2) by blast next assume "distance C s \ dist s ` C" then show "\c \ C. dist s c = distance C s" by auto next assume "\c \ C. dist s c = distance C s" then show "distance C s \ 0" by (metis zero_le_dist) qed lemma dist_lemmas: assumes "finite C" and "C \ {}" shows "finite (dist s ` C)" and "distance C s \ dist s ` C" and "\c \ C. dist s c = distance C s" and "distance C s \ 0" using dist_lemmas_aux assms by auto lemma radius_max_prop: "(\s \ S. distance C s \ r) \ (radius C \ r)" by (metis image_iff radius_contained) lemma dist_ins: assumes "\c\<^sub>1 \ C. \c\<^sub>2 \ C. c\<^sub>1 \ c\<^sub>2 \ x < dist c\<^sub>1 c\<^sub>2" and "distance C s > x" and "finite C" and "C \ {}" shows "\c\<^sub>1 \ (C \ {s}). \c\<^sub>2 \ (C \ {s}). c\<^sub>1 \ c\<^sub>2 \ x < dist c\<^sub>1 c\<^sub>2" proof (rule+) fix c\<^sub>1 c\<^sub>2 assume local_assms: "c\<^sub>1\C \ {s}" "c\<^sub>2\C \ {s}" "c\<^sub>1 \ c\<^sub>2" then have "c\<^sub>1 \ C \ c\<^sub>2 \ C \ c\<^sub>1 \C \ c\<^sub>2\ {s} \ c\<^sub>2\C \ c\<^sub>1 \ {s} \ c\<^sub>1 \ {s} \ c\<^sub>2\ {s}" by auto then show "x < dist c\<^sub>1 c\<^sub>2" proof (elim disjE) assume "c\<^sub>1 \C \ c\<^sub>2\C" then show ?thesis using assms(1) local_assms(3) by simp next assume case_assm: "c\<^sub>1 \ C \ c\<^sub>2 \ {s}" have "x < distance C c\<^sub>2" using assms(2) case_assm by simp also have " ... \ dist c\<^sub>2 c\<^sub>1" using Min.coboundedI distance_def assms(3,4) dist_lemmas(1, 2) case_assm by simp also have " ... = dist c\<^sub>1 c\<^sub>2" using dist_commute by metis finally show ?thesis . next assume case_assm: "c\<^sub>2 \ C \ c\<^sub>1 \ {s}" have "x < distance C c\<^sub>1" using assms(2) case_assm by simp also have " ... \ dist c\<^sub>1 c\<^sub>2" using Min.coboundedI distance_def assms(3,4) dist_lemmas(1, 2) case_assm by simp finally show ?thesis . next assume "c\<^sub>1 \ {s} \ c\<^sub>2 \ {s}" then have False using local_assms by simp then show ?thesis by simp qed qed subsection \A Preliminary Algorithm and Proof\ text \This subsection verifies an auxiliary algorithm by Kleinberg and Tardos. Our proof of the main algorithm does not does not rely on this auxiliary algorithm at all but we do reuse part off its invariant proof later on.\ definition inv :: "('a :: metric_space) set \ ('a :: metric_space set) \ real \ bool" where "inv S' C r = ((\s \ (S - S'). distance C s \ 2*r) \ S' \ S \ C \ S \ (\c \ C. \s \ S'. S' \ {} \ dist c s > 2 * r) \ (S' = S \ C \ {}) \ (\c\<^sub>1 \ C. \c\<^sub>2 \ C. c\<^sub>1 \ c\<^sub>2 \ dist c\<^sub>1 c\<^sub>2 > 2 * r))" lemma inv_init: "inv S {} r" unfolding inv_def non_empty_sites by simp lemma inv_step: assumes "S' \ {}" and IH: "inv S' C r" defines[simp]: "s \ (SOME s. s \ S')" shows "inv (S' - {s' . s' \ S' \ dist s s' \ 2*r}) (C \ {s}) r" proof - have s_def: "s \ S'" using assms(1) some_in_eq by auto have "finite (C \ {s})" using IH finite_subset[OF _ finite_sites] by (simp add: inv_def) moreover have "(\s' \ (S - (S' - {s' . s' \ S' \ dist s s' \ 2*r})). distance (C \ {s}) s' \ 2*r)" proof fix s'' assume "s'' \ S - (S' - {s' . s' \ S' \ dist s s' \ 2*r})" then have "s'' \ S - S' \ s'' \ {s' . s' \ S' \ dist s s' \ 2*r}" by simp then show "distance (C \ {s}) s'' \ 2 * r" proof (elim disjE) assume local_assm: "s'' \ S - S'" have "S' = S \ C \ {}" using IH by (simp add: inv_def) then show ?thesis proof (elim disjE) assume "S' = S" then have "s'' \ {}" using local_assm by simp then show ?thesis by simp next assume C_not_empty: "C \ {}" have "finite C" using IH finite_subset[OF _ finite_sites] by (simp add: inv_def) then have "distance (C \ {s}) s'' \ distance C s''" using distance_mono C_not_empty by (meson Un_upper1 calculation) also have " ... \ 2 * r" using IH local_assm inv_def by simp finally show ?thesis . qed next assume local_assm: "s'' \ {s' . s' \ S' \ dist s s' \ 2*r}" then have "distance (C \ {s}) s'' \ dist s'' s" using Min.coboundedI distance_def dist_lemmas calculation by auto - also have " ... \ 2 * r" using local_assm by (smt dist_self dist_triangle2 mem_Collect_eq) + also have " ... \ 2 * r" using local_assm by (smt (verit) dist_self dist_triangle2 mem_Collect_eq) finally show ?thesis . qed qed moreover have "S' - {s' . s' \ S' \ dist s s' \ 2*r} \ S" using IH by (auto simp: inv_def) moreover { have "s \ S" using IH inv_def s_def by auto then have "C \ {s} \ S" using IH by (simp add: inv_def) } moreover have "(\c\C \ {s}. \c\<^sub>2\C \ {s}. c \ c\<^sub>2 \ 2 * r < dist c c\<^sub>2)" proof (rule+) fix c\<^sub>1 c\<^sub>2 assume local_assms: "c\<^sub>1 \ C \ {s}" "c\<^sub>2 \ C \ {s}" "c\<^sub>1 \ c\<^sub>2" then have "(c\<^sub>1 \ C \ c\<^sub>2 \ C) \ (c\<^sub>1 = s \ c\<^sub>2 \ C) \ (c\<^sub>1 \ C \ c\<^sub>2 = s) \ (c\<^sub>1 = s \ c\<^sub>2 = s)" using assms by auto then show "2 * r < dist c\<^sub>1 c\<^sub>2" proof (elim disjE) assume "c\<^sub>1 \ C \ c\<^sub>2 \ C" then show "2 * r < dist c\<^sub>1 c\<^sub>2" using IH inv_def local_assms by simp next assume case_assm: "c\<^sub>1 = s \ c\<^sub>2 \ C" have "(\c \ C. \s\S'. S' \ {} \ 2 * r < dist c s)" using IH inv_def by simp - then show ?thesis by (smt case_assm s_def assms(1) dist_self dist_triangle3 singletonD) + then show ?thesis by (smt (verit) case_assm s_def assms(1) dist_self dist_triangle3 singletonD) next assume case_assm: "c\<^sub>1 \ C \ c\<^sub>2 = s" have "(\c \ C. \s\S'. S' \ {} \ 2 * r < dist c s)" using IH inv_def by simp - then show ?thesis by (smt case_assm s_def assms(1) dist_self dist_triangle3 singletonD) + then show ?thesis by (smt (verit) case_assm s_def assms(1) dist_self dist_triangle3 singletonD) next assume "c\<^sub>1 = s \ c\<^sub>2 = s" then have False using local_assms(3) by simp then show ?thesis by simp qed qed moreover have "(\c\C \ {s}. \s'' \ S' - {s' \ S'. dist s s' \ 2 * r}. S' - {s' \ S'. dist s s' \ 2 * r} \ {} \ 2 * r < dist c s'')" using IH inv_def by fastforce moreover have "(S' - {s' \ S'. dist s s' \ 2 * r} = S \ C \ {s} \ {})" by simp ultimately show ?thesis unfolding inv_def by blast qed lemma inv_last_1: assumes "\s \ (S - S'). distance C s \ 2*r" and "S' = {}" shows "radius C \ 2*r" by (metis Diff_empty assms image_iff radius_contained) lemma inv_last_2: assumes "finite C" and "card C > n" and "C \ S" and "\c\<^sub>1 \ C. \c\<^sub>2 \ C. c\<^sub>1 \ c\<^sub>2 \ dist c\<^sub>1 c\<^sub>2 > 2*r" shows "\C'. card C' \ n \ card C' > 0 \ radius C' > r" (is ?P) proof (rule ccontr) assume "\ ?P" then obtain C' where card_C': "card C' \ n \ card C' > 0" and radius_C': "radius C' \ r" by auto have "\c \ C. (\c'. c' \ C' \ dist c c' \ r)" proof fix c assume "c \ C" then have "c \ S" using assms(3) by blast then have "distance C' c \ radius C'" using finite_distances by (simp add: radius_def) then have "distance C' c \ r" using radius_C' by simp then show "\c'. c' \ C' \ dist c c' \ r" using dist_lemmas by (metis card_C' card_gt_0_iff) qed then obtain f where f: "\c\C. f c \ C' \ dist c (f c) \ r" by metis have "\inj_on f C" proof assume "inj_on f C" then have "card C' \ card C" using \inj_on f C\ card_inj_on_le card_ge_0_finite card_C' f by blast then show False using card_C' \n < card C\ by linarith qed then obtain c1 c2 where defs: "c1 \ C \ c2 \ C \ c1 \ c2 \ f c1 = f c2" using inj_on_def by blast then have *: "dist c1 (f c1) \ r \ dist c2 (f c1) \ r" using f by auto have "2 * r < dist c1 c2" using assms defs by simp also have " ... \ dist c1 (f c1) + dist (f c1) c2" by(rule dist_triangle) also have " ... = dist c1 (f c1) + dist c2 (f c1)" using dist_commute by simp also have " ... \ 2 * r" using * by simp finally show False by simp qed lemma inv_last: assumes "inv {} C r" shows "(card C \ k \ radius C \ 2*r) \ (card C > k \ (\C'. card C' > 0 \ card C' \ k \ radius C' > r))" using assms inv_def inv_last_1 inv_last_2 finite_subset[OF _ finite_sites] by auto theorem Center_Selection_r: "VARS (S' :: ('a :: metric_space) set) (C :: ('a :: metric_space) set) (r :: real) (s :: 'a) {True} S' := S; C := {}; WHILE S' \ {} INV {inv S' C r} DO s := (SOME s. s \ S'); C := C \ {s}; S' := S' - {s' . s' \ S' \ dist s s' \ 2*r} OD {(card C \ k \ radius C \ 2*r) \ (card C > k \ (\C'. card C' > 0 \ card C' \ k \ radius C' > r))}" proof (vcg, goal_cases) case (1 S' C r) then show ?case using inv_init by simp next case (2 S' C r) then show ?case using inv_step by simp next case (3 S' C r) then show ?case using inv_last by blast qed subsection \The Main Algorithm\ definition invar :: "('a :: metric_space) set \ bool" where "invar C = (C \ {} \ card C \ k \ C \ S \ (\C'. (\c\<^sub>1 \ C. \c\<^sub>2 \ C. c\<^sub>1 \ c\<^sub>2 \ dist c\<^sub>1 c\<^sub>2 > 2 * radius C') \ (\s \ S. distance C s \ 2 * radius C')))" abbreviation some where "some A \ (SOME s. s \ A)" lemma invar_init: "invar {some S}" proof - let ?s = "some S" have s_in_S: "?s \ S" using some_in_eq non_empty_sites by blast have "{?s} \ {}" by simp moreover have "{SOME s. s \ S} \ S" using s_in_S by simp moreover have "card {SOME s. s \ S} \ k" using non_zero_k by simp ultimately show ?thesis by (auto simp: invar_def) qed abbreviation furthest_from where "furthest_from C \ (SOME s. s \ S \ distance C s = Max (distance C ` S))" lemma invar_step: assumes "invar C" and "card C < k" shows "invar (C \ {furthest_from C})" proof - have furthest_from_C_props: "furthest_from C \ S \ distance C (furthest_from C) = radius C " using someI_ex[of "\x. x \ S \ distance C x = radius C"] radius_def2 radius_def by auto have C_props: "finite C \ C \ {}" using finite_subset[OF _ finite_sites] assms(1) unfolding invar_def by blast { have "card (C \ {furthest_from C}) \ card C + 1" using assms(1) C_props unfolding invar_def by (simp add: card_insert_if) then have "card (C \ {furthest_from C}) < k + 1" using assms(2) by simp then have "card (C \ {furthest_from C}) \ k" by simp } moreover have "C \ {furthest_from C} \ {}" by simp moreover have "(C \ {furthest_from C}) \ S" using assms(1) furthest_from_C_props unfolding invar_def by simp moreover have "\C'. (\s \ S. distance (C \ {furthest_from C}) s \ 2 * radius C') \ (\c\<^sub>1 \ C \ {furthest_from C}. \c\<^sub>2 \ C \ {furthest_from C}. c\<^sub>1 \ c\<^sub>2 \ 2 * radius C' < dist c\<^sub>1 c\<^sub>2)" proof fix C' have "distance C (furthest_from C) > 2 * radius C' \ distance C (furthest_from C) \ 2 * radius C'" by auto then show "(\s \ S. distance (C \ {furthest_from C}) s \ 2 * radius C') \ (\c\<^sub>1 \ C \ {furthest_from C}. \c\<^sub>2 \ C \ {furthest_from C}. c\<^sub>1 \ c\<^sub>2 \ 2 * radius C' < dist c\<^sub>1 c\<^sub>2)" proof (elim disjE) assume asm: "distance C (furthest_from C) > 2 * radius C'" then have "\(\s \ S. distance C s \ 2 * radius C')" using furthest_from_C_props by force then have IH: "\c\<^sub>1 \ C. \c\<^sub>2 \ C. c\<^sub>1 \ c\<^sub>2 \ 2 * radius C' < dist c\<^sub>1 c\<^sub>2" using assms(1) unfolding invar_def by blast have "(\c\<^sub>1 \ C \ {furthest_from C}. (\c\<^sub>2 \ C \ {furthest_from C}. c\<^sub>1 \ c\<^sub>2 \ 2 * radius C' < dist c\<^sub>1 c\<^sub>2))" using dist_ins[of "C" "2 * radius C'" "furthest_from C"] IH C_props asm by simp then show ?thesis by simp next assume main_assm: "2 * radius C' \ distance C (furthest_from C)" have "(\s \ S. distance (C \ {furthest_from C}) s \ 2 * radius C')" proof fix s assume local_assm: "s \ S" then show "distance (C \ {furthest_from C}) s \ 2 * radius C'" proof - have "distance (C \ {furthest_from C}) s \ distance C s" using distance_mono[of C "C \ {furthest_from C}"] C_props by auto also have " ... \ distance C (furthest_from C)" using Max.coboundedI local_assm finite_distances radius_def furthest_from_C_props by auto also have " ... \ 2 * radius C'" using main_assm by simp finally show ?thesis . qed qed then show ?thesis by blast qed qed ultimately show ?thesis unfolding invar_def by blast qed lemma invar_last: assumes "invar C" and "\card C < k" shows "card C = k" and "card C' > 0 \ card C' \ k \ radius C \ 2 * radius C'" proof - show "card C = k" using assms(1, 2) unfolding invar_def by simp next have C_props: "finite C \ C \ {}" using finite_sites assms(1) unfolding invar_def by (meson finite_subset) show "card C' > 0 \ card C' \ k \ radius C \ 2 * radius C'" proof (rule impI) assume C'_assms: "0 < card (C' :: 'a set) \ card C' \ k" let ?r = "radius C'" have "(\c\<^sub>1 \ C. \c\<^sub>2 \ C. c\<^sub>1 \ c\<^sub>2 \ 2 * ?r < dist c\<^sub>1 c\<^sub>2) \ (\s \ S. distance C s \ 2 * ?r)" using assms(1) unfolding invar_def by simp then show "radius C \ 2 * ?r" proof assume case_assm: "\c\<^sub>1\C. \c\<^sub>2\C. c\<^sub>1 \ c\<^sub>2 \ 2 * ?r < dist c\<^sub>1 c\<^sub>2" obtain s where s_def: "radius C = distance C s \ s \ S" using radius_def2 by metis show ?thesis proof (rule ccontr) assume contr_assm: "\ radius C \ 2 * ?r" then have s_prop: "distance C s > 2 * ?r" using s_def by simp then have \\c\<^sub>1 \ C \ {s}. \c\<^sub>2 \ C \ {s}. c\<^sub>1 \ c\<^sub>2 \ dist c\<^sub>1 c\<^sub>2 > 2 * ?r\ using C_props dist_ins[of "C" "2*?r" "s"] case_assm by blast moreover { have "s \ C" proof assume "s \ C" then have "distance C s \ dist s s" using Min.coboundedI[of "distance C ` S" "dist s s"] by (simp add: distance_def C_props) also have " ... = 0" by simp - finally have "distance C s = 0" using dist_lemmas(4) by (smt C_props) + finally have "distance C s = 0" using dist_lemmas(4) by (smt (verit) C_props) then have radius_le_zero: "2 * ?r < 0" using contr_assm s_def by simp obtain x where x_def: "?r = distance C' x" using radius_def2 by metis obtain l where l_def: "distance C' x = dist x l" using dist_lemmas(3) by (metis C'_assms card_gt_0_iff) then have "dist x l = ?r" by (simp add: x_def) also have "... < 0" using C'_assms radius_le_zero by simp finally show False by simp qed then have "card (C \ {s}) > k" using assms(1,2) C_props unfolding invar_def by simp } moreover have "C \ {s} \ S" using assms(1) s_def unfolding invar_def by simp moreover have "finite (C \ {s})" using calculation(3) finite_subset finite_sites by auto ultimately have "\C. card C \ k \ card C > 0 \ radius C > ?r" using inv_last_2 by metis then have "?r > ?r" using C'_assms by blast then show False by simp qed next assume "\s\S. distance C s \ 2 * radius C'" then show ?thesis by (metis image_iff radius_contained) qed qed qed theorem Center_Selection: "VARS (C :: ('a :: metric_space) set) (s :: ('a :: metric_space)) {k \ card S} C := {some S}; WHILE card C < k INV {invar C} DO C := C \ {furthest_from C} OD {card C = k \ (\C'. card C' > 0 \ card C' \ k \ radius C \ 2 * radius C')}" proof (vcg, goal_cases) case (1 C s) show ?case using invar_init by simp next case (2 C s) then show ?case using invar_step by blast next case (3 C s) then show ?case using invar_last by blast qed end end \ No newline at end of file diff --git a/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy b/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy --- a/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy +++ b/thys/Arith_Prog_Rel_Primes/Arith_Prog_Rel_Primes.thy @@ -1,955 +1,955 @@ (* File: Arith_Prog_Rel_Primes.thy Author: Jose Manuel Rodriguez Caballero, University of Tartu *) section \Problem ARITHMETIC PROGRESSIONS (Putnam exam problems 2002)\ theory Arith_Prog_Rel_Primes imports Complex_Main "HOL-Number_Theory.Number_Theory" begin text \ Statement of the problem (from ~\<^cite>\"putnam"\): For which integers $n>1$ does the set of positive integers less than and relatively prime to $n$ constitute an arithmetic progression? The solution of the above problem is theorem @{text arith_prog_rel_primes_solution}. First, we will require some auxiliary material before we get started with the actual solution. \ subsection \Auxiliary results\ lemma even_and_odd_parts: fixes n::nat assumes \n \ 0\ shows \\ k q::nat. n = (2::nat)^k*q \ odd q\ proof- have \prime (2::nat)\ by simp thus ?thesis using prime_power_canonical[where p = "2" and m = "n"] assms semiring_normalization_rules(7) by auto qed lemma only_one_odd_div_power2: fixes n::nat assumes \n \ 0\ and \\ x. x dvd n \ odd x \ x = 1\ shows \\ k. n = (2::nat)^k\ by (metis even_and_odd_parts assms(1) assms(2) dvd_triv_left semiring_normalization_rules(11) semiring_normalization_rules(7)) lemma coprime_power2: fixes n::nat assumes \n \ 0\ and \\ x. x < n \ (coprime x n \ odd x)\ shows \\ k. n = (2::nat)^k\ proof- have \x dvd n \ odd x \ x = 1\ for x by (metis neq0_conv One_nat_def Suc_1 Suc_lessI assms(1) assms(2) coprime_left_2_iff_odd dvd_refl linorder_neqE_nat nat_dvd_1_iff_1 nat_dvd_not_less not_coprimeI) thus ?thesis using assms(1) only_one_odd_div_power2 by auto qed subsection \Main result\ text \ The solution to the problem ARITHMETIC PROGRESSIONS (Putnam exam problems 2002) \ theorem arith_prog_rel_primes_solution: fixes n :: nat assumes \n > 1\ shows \(prime n \ (\ k. n = 2^k) \ n = 6) \ (\ a b m. m \ 0 \ {x | x. x < n \ coprime x n} = {a+j*b| j::nat. j < m})\ proof- have \ (prime n \ (\ k. n = 2^k) \ n = 6) \ (\ b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m})\ proof show "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" if "prime n \ (\k. n = 2 ^ k) \ n = 6" proof- have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" if "prime n" proof- have \\m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m}\ proof- have \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x \ 0 \ x < n}\ proof show "{x |x. x < n \ coprime x n} \ {x |x. x \ 0 \ x < n}" - by (smt Collect_mono not_le ord_0_nat ord_eq_0 order_refl prime_gt_1_nat that zero_neq_one) + using prime_nat_iff'' that by fastforce show "{x |x. x \ 0 \ x < n} \ {x |x. x < n \ coprime x n}" using coprime_commute prime_nat_iff'' that by fastforce qed obtain m where \m+1 = n\ using add.commute assms less_imp_add_positive by blast have \{1+j| j::nat. j < (m::nat)} = {x | x :: nat. x \ 0 \ x < m+1}\ by (metis Suc_eq_plus1 \m + 1 = n\ gr0_implies_Suc le_simps(3) less_nat_zero_code linorder_not_less nat.simps(3) nat_neq_iff plus_1_eq_Suc ) hence \{x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < (m::nat)}\ using \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x \ 0 \ x < n}\ \m+1 = n\ by auto from \n > 1\ have \m \ 0\ using \m + 1 = n\ by linarith have \{x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m}\ using Suc_eq_plus1 \1 < n\ \{x |x. x < n \ coprime x n} = {1 + j |j. j < m}\ by auto hence \(\ m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j| j::nat. j < m})\ using \m \ 0\ by blast thus ?thesis by blast qed hence \\m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*1| j::nat. j < m}\ by auto thus ?thesis by blast qed moreover have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" if "\k. n = 2 ^ k" proof- obtain k where \n = 2 ^ k\ using \\k. n = 2 ^ k\ by auto have \k \ 0\ by (metis \1 < n\ \n = 2 ^ k\ nat_less_le power.simps(1)) obtain t where \Suc t = k\ by (metis \k \ 0\ fib.cases) have \n = 2^(Suc t)\ by (simp add: \Suc t = k\ \n = 2 ^ k\) have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < 2^t}\ proof show "{x |x. x < n \ coprime x n} \ {1 + j * 2 |j. j < 2^t}" proof fix x assume \x \ {x |x. x < n \ coprime x n}\ hence \x < n\ by blast have \coprime x n\ using \x \ {x |x. x < n \ coprime x n}\ by blast hence \coprime x (2^(Suc k))\ by (simp add: \k \ 0\ \n = 2 ^ k\) have \odd x\ using \coprime x n\ \k \ 0\ \n = 2 ^ k\ by auto then obtain j where \x = 1+j*2\ by (metis add.commute add.left_commute left_add_twice mult_2_right oddE) have \x < 2^k\ using \n = 2 ^ k\ \x < n\ \x = 1+j*2\ by linarith hence \1+j*2 < 2^k\ using \x = 1+j*2\ by blast hence \j < 2^t\ using \Suc t = k\ by auto thus \x \ {1 + j * 2 |j. j < 2^t}\ using \x = 1+j*2\ by blast qed show "{1 + j * 2 |j. j < 2 ^ t} \ {x |x. x < n \ coprime x n}" proof fix x::nat assume \x \ {1 + j * 2 |j. j < 2 ^ t}\ then obtain j where \x = 1 + j * 2\ and \j < 2 ^ t\ by blast have \x < 2*(2^t)\ using \x = 1 + j * 2\ \j < 2 ^ t\ by linarith hence \x < n\ by (simp add: \n = 2 ^ Suc t\) moreover have \coprime x n\ by (metis (no_types) \\thesis. (\j. \x = 1 + j * 2; j < 2 ^ t\ \ thesis) \ thesis\ \n = 2 ^ k\ coprime_Suc_left_nat coprime_mult_right_iff coprime_power_right_iff plus_1_eq_Suc) ultimately show \x \ {x |x. x < n \ coprime x n}\ by blast qed qed have \(2::nat)^(t::nat) \ 0\ by simp obtain m where \m = (2::nat)^t\ by blast have \m \ 0\ using \m = 2 ^ t\ by auto have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ using \m = 2 ^ t\ \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < 2 ^ t}\ by auto from \m \ 0\ \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ show ?thesis by blast qed moreover have "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" if "n = 6" proof- have \{x | x. x < 6 \ coprime x 6} = {1+j*4| j::nat. j < 2}\ proof- have \{x | x::nat. x < 6 \ coprime x 6} = {1, 5}\ proof show "{u. \x. u = (x::nat) \ x < 6 \ coprime x 6} \ {1, 5}" proof fix u::nat assume \u \ {u. \x. u = x \ x < 6 \ coprime x 6}\ hence \coprime u 6\ by blast have \u < 6\ using \u \ {u. \x. u = x \ x < 6 \ coprime x 6}\ by blast moreover have \u \ 0\ using \coprime u 6\ ord_eq_0 by fastforce moreover have \u \ 2\ using \coprime u 6\ by auto moreover have \u \ 3\ proof- have \gcd (3::nat) 6 = 3\ by auto thus ?thesis by (metis (no_types) \coprime u 6\ \gcd 3 6 = 3\ coprime_iff_gcd_eq_1 numeral_eq_one_iff semiring_norm(86)) qed moreover have \u \ 4\ proof- have \gcd (4::nat) 6 = 2\ by (metis (no_types, lifting) add_numeral_left gcd_add1 gcd_add2 gcd_nat.idem numeral_Bit0 numeral_One one_plus_numeral semiring_norm(4) semiring_norm(5)) thus ?thesis using \coprime u 6\ coprime_iff_gcd_eq_1 by auto qed ultimately have \u = 1 \ u = 5\ by auto thus \u \ {1, 5}\ by blast qed show "{1::nat, 5} \ {x |x. x < 6 \ coprime x 6}" proof- have \(1::nat) \ {x |x. x < 6 \ coprime x 6}\ by simp moreover have \(5::nat) \ {x |x. x < 6 \ coprime x 6}\ by (metis Suc_numeral coprime_Suc_right_nat less_add_one mem_Collect_eq numeral_plus_one semiring_norm(5) semiring_norm(8)) ultimately show ?thesis by blast qed qed moreover have \{1+j*4| j::nat. j < 2} = {1, 5}\ by auto ultimately show ?thesis by auto qed moreover have \(2::nat) \ 0\ by simp ultimately have \\ m. m \ 0 \ {x | x :: nat. x < 6 \ coprime x 6} = {1+j*4| j::nat. j < m}\ by blast thus ?thesis using that by auto qed ultimately show ?thesis using that by blast qed show "prime n \ (\k. n = 2 ^ k) \ n = 6" if "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" proof- obtain b m where \m \ 0\ and \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ using \\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by auto show ?thesis proof(cases \n = 2\) case True thus ?thesis by auto next case False have \b \ 4\ proof(cases \odd b\) case True show ?thesis proof(rule classical) assume \\(b \ 4)\ hence \b > 4\ using le_less_linear by blast obtain m where \m \ 0\ and \{x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast have \b \ 0\ using \4 < b\ by linarith have \n = 2 + (m-1)*b\ proof- have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 \ {1+j*b| j::nat. j < m}\ using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ by blast have \i \ m-1\ using \i < m\ by linarith have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ using \m \ 0\ by auto hence \1 + (m-1)*b \ {x | x::nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \1 + (m-1)*b < n\ by blast hence \1 + (m-1)*b \ n-1\ by linarith hence \1 + (m-1)*b \ 1+i*b\ using \n - 1 = 1 + i * b\ by linarith hence \(m-1)*b \ i*b\ by linarith hence \m-1 \ i\ using \b \ 0\ by auto hence \m-1 = i\ using \i \ m - 1\ le_antisym by blast thus ?thesis using \m \ 0\ \n - 1 = 1 + i * b\ by auto qed have \m \ 2\ using \n = 2 + (m - 1)*b\ \n \ 2\ by auto hence \1+b \ {1+j*b| j. j < m}\ by fastforce hence \1+b \ {x | x::nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \coprime (1+b) n\ by blast have \(2::nat) dvd (1+b)\ using \odd b\ by simp hence \coprime (2::nat) n\ using \coprime (1 + b) n\ coprime_common_divisor coprime_left_2_iff_odd odd_one by blast have \(2::nat) < n\ using \1 < n\ \n \ 2\ by linarith have \2 \ {x | x :: nat. x < n \ coprime x n}\ using \2 < n\ \coprime 2 n\ by blast hence \2 \ {1+j*b| j::nat. j < m}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain j::nat where \2 = 1+j*b\ by blast have \1 = j*b\ using \2 = 1+j*b\ by linarith thus ?thesis by simp qed next case False hence \even b\ by simp show ?thesis proof(rule classical) assume \\(b \ 4)\ hence \b > 4\ using le_less_linear by blast obtain m where \ m \ 0\ and \{x | x::nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast have \b \ 0\ using \4 < b\ by linarith have \n = 2 + (m-1)*b\ proof- have \n-1 \ {x | x::nat. x < n \ coprime x n}\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 \ {1+j*b| j::nat. j < m}\ using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ by blast have \i \ m-1\ using \i < m\ by linarith have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ using \m \ 0\ by auto hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \1 + (m-1)*b < n\ by blast hence \1 + (m-1)*b \ n-1\ by linarith hence \1 + (m-1)*b \ 1+i*b\ using \n - 1 = 1 + i * b\ by linarith hence \(m-1)*b \ i*b\ by linarith hence \m-1 \ i\ using \b \ 0\ by auto hence \m-1 = i\ using \i \ m - 1\ le_antisym by blast thus ?thesis using \m \ 0\ \n - 1 = 1 + i * b\ by auto qed obtain k :: nat where \b = 2*k\ using \even b\ by blast have \n = 2*(1 + (m-1)*k)\ using \n = 2 + (m-1)*b\ \b = 2*k\ by simp show ?thesis proof (cases \odd m\) case True hence \odd m\ by blast then obtain t::nat where \m-1 = 2*t\ by (metis odd_two_times_div_two_nat) have \n = 2*(1 + b*t)\ using \m - 1 = 2 * t\ \n = 2 + (m - 1) * b\ by auto have \t < m\ using \m - 1 = 2 * t\ \m \ 0\ by linarith have \1 + b*t \ {1+j*b| j::nat. j < m}\ using \t < m\ by auto hence \1 + b*t \ {x | x::nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \coprime (1 + b*t) n\ by auto thus ?thesis by (metis (no_types, lifting) \b = 2 * k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 * (1 + b * t)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_cancel_right_right coprime_mult_right_iff coprime_self mult_cancel_left mult_is_0 nat_dvd_1_iff_1) next case False thus ?thesis proof(cases \odd k\) case True hence \odd k\ by blast have \even (1 + (m - 1) * k)\ by (simp add: False True \m \ 0\) have \coprime (2 + (m - 1) * k) (1 + (m - 1) * k)\ by simp have \coprime (2 + (m - 1) * k) n\ using \coprime (2 + (m - 1) * k) (1 + (m - 1) * k)\ \even (1 + (m - 1) * k)\ \n = 2 * (1 + (m - 1) * k)\ coprime_common_divisor coprime_mult_right_iff coprime_right_2_iff_odd odd_one by blast have \2 + (m - 1) * k < n\ by (metis (no_types, lifting) \even (1 + (m - 1) * k)\ \n = 2 * (1 + (m - 1) * k)\ add_gr_0 add_mono_thms_linordered_semiring(1) dvd_add_left_iff dvd_add_triv_left_iff dvd_imp_le le_add2 le_neq_implies_less less_numeral_extra(1) mult_2 odd_one) have \2 + (m - 1) * k \ {x | x :: nat. x < n \ coprime x n}\ using \2 + (m - 1) * k < n\ \coprime (2 + (m - 1) * k) n\ by blast hence \2 + (m - 1) * k \ {1 + j * b |j. j < m}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain j::nat where \2 + (m - 1) * k = 1 + j * b\ and \j < m\ by blast have \1 + (m - 1) * k = j * b\ using \2 + (m - 1) * k = 1 + j * b\ by simp hence \1 + (m - 1) * k = j * (2*k)\ using \b = 2 * k\ by blast thus ?thesis by (metis \b = 2 * k\ \even b\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ dvd_add_times_triv_right_iff dvd_antisym dvd_imp_le dvd_triv_right even_numeral mult_2 zero_less_numeral) next case False hence \even k\ by auto have \odd (1 + (m - 1) * k)\ by (simp add: \even k\ ) hence \coprime (3 + (m - 1) * k) (1 + (m - 1) * k)\ - by (smt add_numeral_left coprime_common_divisor coprime_right_2_iff_odd dvd_add_left_iff not_coprimeE numeral_Bit1 numeral_One numeral_plus_one one_add_one) + by (smt (verit) add_numeral_left coprime_common_divisor coprime_right_2_iff_odd dvd_add_left_iff not_coprimeE numeral_Bit1 numeral_One numeral_plus_one one_add_one) hence \coprime (3 + (m - 1) * k) n\ by (metis \even k\ \n = 2 * (1 + (m - 1) * k)\ coprime_mult_right_iff coprime_right_2_iff_odd even_add even_mult_iff odd_numeral) have \3 + (m - 1) * k < n\ - by (smt Groups.add_ac(2) \even k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_Suc_right add_cancel_right_right add_mono_thms_linordered_semiring(1) dvd_imp_le even_add even_mult_iff le_add2 le_neq_implies_less left_add_twice mult_2 neq0_conv numeral_Bit1 numeral_One odd_numeral one_add_one plus_1_eq_Suc) + by (smt (verit) Groups.add_ac(2) \even k\ \n = 2 * (1 + (m - 1) * k)\ \n = 2 + (m - 1) * b\ \n \ 2\ add_Suc_right add_cancel_right_right add_mono_thms_linordered_semiring(1) dvd_imp_le even_add even_mult_iff le_add2 le_neq_implies_less left_add_twice mult_2 neq0_conv numeral_Bit1 numeral_One odd_numeral one_add_one plus_1_eq_Suc) have \3 + (m - 1) * k \ {x |x. x < n \ coprime x n}\ using \3 + (m - 1) * k < n\ \coprime (3 + (m - 1) * k) n\ by blast hence \3 + (m - 1) * k \ {1 + j * b |j. j < m}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain j::nat where \3 + (m - 1) * k = 1 + j * b\ by blast have \2 + (m - 1) * k = j * b\ using \3 + (m - 1) * k = 1 + j * b\ by simp hence \2 + (m - 1) * k = j * 2*k\ by (simp add: \b = 2 * k\) thus ?thesis by (metis \4 < b\ \b = 2 * k\ \even k\ dvd_add_times_triv_right_iff dvd_antisym dvd_triv_right mult_2 nat_neq_iff numeral_Bit0) qed qed qed qed moreover have \b \ 3\ proof (rule classical) assume \\ (b \ 3)\ hence \b = 3\ by blast obtain m where \m \ 0\ and \{x | x::nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}\ using \m \ 0\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast have \b \ 0\ by (simp add: \b = 3\) have \n = 2 + (m-1)*b\ proof- have \n-1 \ {x | x::nat. x < n \ coprime x n}\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 \ {1+j*b| j::nat. j < m}\ using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ by blast have \i \ m-1\ using \i < m\ by linarith have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ using \m \ 0\ by auto hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \1 + (m-1)*b < n\ by blast hence \1 + (m-1)*b \ n-1\ by linarith hence \1 + (m-1)*b \ 1+i*b\ using \n - 1 = 1 + i * b\ by linarith hence \(m-1)*b \ i*b\ by linarith hence \m-1 \ i\ using \b \ 0\ by auto hence \m-1 = i\ using \i \ m - 1\ le_antisym by blast thus ?thesis using \m \ 0\ \n - 1 = 1 + i * b\ by auto qed have \n > 2\ using \1 < n\ \n \ 2\ by linarith hence \ m \ 2 \ using \n = 2 + (m-1)*b\ \b = 3\ by simp have \4 \ {1+j*(b::nat)| j::nat. j < m}\ using \2 \ m\ \b = 3\ by force hence \(4::nat) \ {x | x :: nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by auto hence \coprime (4::nat) n\ by blast have \(2::nat) dvd 4\ by auto hence \coprime (2::nat) n\ using \coprime (4::nat) n\ coprime_divisors dvd_refl by blast have \4 < n\ using \4 \ {x |x. x < n \ coprime x n}\ by blast have \2 < (4::nat)\ by auto have \2 < n\ by (simp add: \2 < n\) hence \2 \ {x | x :: nat. x < n \ coprime x n}\ using \coprime (2::nat) n\ by blast hence \2 \ {1+j*(b::nat)| j::nat. j < m}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain j::nat where \2 = 1+j*3\ using \b = 3\ by blast from \2 = 1+j*3\ have \1 = j*3\ by auto hence \3 dvd 1\ by auto thus ?thesis using nat_dvd_1_iff_1 numeral_eq_one_iff by blast qed ultimately have \b = 0 \ b = 1 \ b = 2 \ b = 4\ by auto moreover have \b = 0 \ \k. n = 2^k\ proof- assume \b = 0\ have \{1 + j * b |j. j < m} = {1}\ using \b = 0\ \m \ 0\ by auto hence \{x |x. x < n \ coprime x n} = {1}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \n = 2\ proof- have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 \ {1}\ using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1}\ by blast hence \n-1 = 1\ by blast hence \n = 2\ by simp thus ?thesis by blast qed hence \n = 2^1\ by auto thus ?thesis by blast qed moreover have \b = 1 \ prime n\ proof- assume \b = 1\ have \x < n \ x \ 0 \ coprime x n\ for x proof- assume \x < n\ and \x \ 0\ have \{1+j| j::nat. j < m} = {x | x::nat. x < m+1 \ x \ 0}\ by (metis (full_types) Suc_eq_plus1 add_mono1 less_Suc_eq_0_disj nat.simps(3) plus_1_eq_Suc ) hence \{x | x :: nat. x < n \ coprime x n} = {x | x :: nat. x < m+1 \ x \ 0}\ using \b = 1\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by auto have \coprime (n-1) n\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 < n\ using \1 < n\ by auto have \n-1 \ {x |x. x < n \ coprime x n}\ using \coprime (n - 1) n\ \n - 1 < n\ by blast have \n-1 \ m\ by (metis (no_types, lifting) CollectD Suc_eq_plus1 Suc_less_eq2 \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ leD le_less_linear not_less_eq_eq ) have \m \ {x | x :: nat. x < m+1 \ x \ 0}\ using \m \ 0\ by auto have \m \ {x |x. x < n \ coprime x n} \ using \m \ {x |x. x < m + 1 \ x \ 0}\ \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ by blast have \m < n\ using \m \ {x |x. x < n \ coprime x n}\ by blast have \m+1 = n\ using \m < n\ \n - 1 \ m\ by linarith have \x \ {x | x :: nat. x < m+1 \ x \ 0}\ using \m + 1 = n\ \x < n\ \x \ 0\ by blast hence \x \ {x |x. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {x |x. x < m + 1 \ x \ 0}\ by blast thus ?thesis by blast qed thus ?thesis using assms coprime_commute nat_neq_iff prime_nat_iff'' by auto qed moreover have \b = 2 \ \ k. n = 2^k\ proof- assume \b = 2\ have \{x | x :: nat. x < n \ coprime x n} = {1+j*2| j::nat. j < m}\ using \b = 2\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by auto have \x < n \ coprime x n \ odd x\ for x::nat proof- assume \x < n\ have \coprime x n \ odd x\ proof- assume \coprime x n\ have \x \ {x | x :: nat. x < n \ coprime x n}\ by (simp add: \coprime x n\ \x < n\) hence \x \ {1+j*2| j::nat. j < m}\ using \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < m}\ by blast then obtain j where \x = 1+j*2\ by blast thus ?thesis by simp qed moreover have \odd x \ coprime x n\ proof- assume \odd x\ obtain j::nat where \x = 1+j*2\ by (metis \odd x\ add.commute mult_2_right odd_two_times_div_two_succ one_add_one semiring_normalization_rules(16)) have \j < (n-1)/2\ using \x < n\ \x = 1 + j * 2\ by linarith have \n = 2*m\ proof- have \(2::nat) \ 0\ by auto have \n = 2+(m-1)*2\ proof- have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 \ {1+j*b| j::nat. j < m}\ using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ by blast have \i \ m-1\ using \i < m\ by linarith have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ using \m \ 0\ by auto hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \1 + (m-1)*b < n\ by blast hence \1 + (m-1)*b \ n-1\ by linarith hence \1 + (m-1)*b \ 1+i*b\ using \n - 1 = 1 + i * b\ by linarith hence \(m-1)*b \ i*b\ by linarith hence \m-1 \ i\ proof- have \b \ 0\ using \b = 2\ by simp thus ?thesis using \(m - 1) * b \ i * b\ mult_le_cancel2 by blast qed hence \m-1 = i\ using \i \ m - 1\ le_antisym by blast thus ?thesis using \m \ 0\ \n - 1 = 1 + i * b\ by (simp add: \b = 2\) qed thus ?thesis by (simp add: \m \ 0\ \n = 2 + (m - 1) * 2\ mult.commute mult_eq_if) qed hence \j < m\ using \x < n\ \x = 1 + j * 2\ by linarith hence \x \ {1+j*2| j::nat. j < m}\ using \x = 1 + j * 2\ by blast hence \x \ {x | x :: nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * 2 |j. j < m}\ by blast thus ?thesis by blast qed ultimately show ?thesis by blast qed thus ?thesis using coprime_power2 assms by auto qed moreover have \b = 4 \ n = 6\ proof- assume \b = 4\ have \n = 2 \ n = 6\ proof(rule classical) assume \\ (n = 2 \ n = 6)\ have \(4::nat) \ 0\ by auto have \n = 2+(m-1)*4\ proof- have \n-1 \ {x | x :: nat. x < n \ coprime x n}\ using \1 < n\ coprime_diff_one_left_nat by auto have \n-1 \ {1+j*b| j::nat. j < m}\ using \n - 1 \ {x |x. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain i::nat where \n-1 = 1+i*b\ and \i < m\ by blast have \i \ m-1\ using \i < m\ by linarith have \1 + (m-1)*b \ {1+j*b| j::nat. j < m}\ using \m \ 0\ by auto hence \1 + (m-1)*b \ {x | x :: nat. x < n \ coprime x n}\ using \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast hence \1 + (m-1)*b < n\ by blast hence \1 + (m-1)*b \ n-1\ by linarith hence \1 + (m-1)*b \ 1+i*b\ using \n - 1 = 1 + i * b\ by linarith hence \(m-1)*b \ i*b\ by linarith hence \m-1 \ i\ proof- have \b \ 0\ using \b = 4\ by auto thus ?thesis using \(m - 1) * b \ i * b\ mult_le_cancel2 by blast qed hence \m-1 = i\ using \i \ m - 1\ le_antisym by blast thus ?thesis using \m \ 0\ \n - 1 = 1 + i * b\ by (simp add: \b = 4\) qed hence \n = 4*m - 2\ by (simp add: \m \ 0\ mult.commute mult_eq_if) have \m \ 3\ using \\ (n = 2 \ n = 6)\ \n = 2 + (m - 1) * 4\ by auto hence \ {1+j*4| j::nat. j < 3} \ {1+j*4| j::nat. j < m}\ by force hence \9 \ {1+j*4| j::nat. j < 3}\ by force hence \9 \ {1+j*4| j::nat. j < m}\ using \ {1+j*4| j::nat. j < 3} \ {1+j*4| j::nat. j < m}\ by blast hence \9 \ {x | x :: nat. x < n \ coprime x n}\ using \b = 4\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by auto hence \coprime (9::nat) n\ by blast have \(3::nat) dvd 9\ by auto have \coprime (3::nat) n\ using \coprime (9::nat) n\ \(3::nat) dvd 9\ by (metis coprime_commute coprime_mult_right_iff dvd_def) have \(3::nat) < n\ by (metis One_nat_def Suc_lessI \1 < n\ \\ (n = 2 \ n = 6)\ \coprime 3 n\ coprime_self numeral_2_eq_2 numeral_3_eq_3 less_numeral_extra(1) nat_dvd_not_less) have \3 \ {x | x :: nat. x < n \ coprime x n}\ using \3 < n\ \coprime 3 n\ by blast hence \(3::nat) \ {1+j*4| j::nat. j < m}\ using \b = 4\ \{x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}\ by blast then obtain j::nat where \(3::nat) = 1 + j*4\ by blast have \2 = j*4\ using numeral_3_eq_3 \(3::nat) = 1 + j*4\ by linarith hence \1 = j*2\ by linarith hence \even 1\ by simp thus ?thesis using odd_one by blast qed thus ?thesis by (simp add: False) qed ultimately show ?thesis by blast qed qed qed moreover have \(\ b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {1+j*b| j::nat. j < m}) \ (\ a b m. m \ 0 \ {x | x :: nat. x < n \ coprime x n} = {a+j*b| j::nat. j < m})\ proof show "\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}" if "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" using that by blast show "\b m. m \ 0 \ {x |x. x < n \ coprime x n} = {1 + j * b |j. j < m}" if "\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}" proof- obtain a b m::nat where \m \ 0\ and \{x | x :: nat. x < n \ coprime x n} = {a+j*b| j::nat. j < m}\ using \\a b m. m \ 0 \ {x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ by auto have \a = 1\ proof- have \{x | x :: nat. x < n \ coprime x n} = {(a::nat)+j*(b::nat)| j::nat. j < m} \ a = 1\ proof- have \Min {x | x :: nat. x < n \ coprime x n} = Min {a+j*b| j::nat. j < m}\ using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ by auto have \Min {x | x :: nat. x < n \ coprime x n} = 1\ proof- have \finite {x | x :: nat. x < n \ coprime x n}\ by auto have \{x | x :: nat. x < n \ coprime x n} \ {}\ using \1 < n\ by auto have \1 \ {x | x :: nat. x < n \ coprime x n}\ using \1 < n\ by auto have \\ x. coprime x n \ x \ 1\ using \1 < n\ le_less_linear by fastforce hence \\ x. x < n \ coprime x n \ x \ 1\ by blast hence \\ x \ {x | x :: nat. x < n \ coprime x n}. x \ 1\ by blast hence \Min {x | x :: nat. x < n \ coprime x n} \ 1\ using \finite {x | x :: nat. x < n \ coprime x n}\ \{x |x. x < n \ coprime x n} \ {}\ by auto thus ?thesis using Min_le \1 \ {x |x. x < n \ coprime x n}\ \finite {x |x. x < n \ coprime x n}\ antisym by blast qed have \Min {a+j*b| j::nat. j < m} = a\ proof - have f1: "\n. a = a + n * b \ n < m" using \m \ 0\ by auto have f2: "\n. 1 = a + n * b \ n < m" using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ assms coprime_1_left by blast have f3: "\na. a = na \ na < n \ coprime na n" using f1 \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ by blast have "n \ 1" by (metis (lifting) assms less_irrefl_nat) then have "\ coprime 0 n" by simp then show ?thesis using f3 f2 by (metis \Min {x |x. x < n \ coprime x n} = 1\ \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ less_one linorder_neqE_nat not_add_less1) qed hence \Min {a+j*b| j::nat. j < m} = a\ by blast thus ?thesis using \Min {x | x :: nat. x < n \ coprime x n} = 1\ \Min {x | x :: nat. x < n \ coprime x n} = Min {a+j*b| j::nat. j < m}\ by linarith qed thus ?thesis using \{x |x. x < n \ coprime x n} = {a + j * b |j. j < m}\ by blast qed thus ?thesis using \m \ 0\ \{x | x. x < n \ coprime x n} = {a+j*b| j::nat. j < m}\ by auto qed qed ultimately show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/Arith_Prog_Rel_Primes/ROOT b/thys/Arith_Prog_Rel_Primes/ROOT --- a/thys/Arith_Prog_Rel_Primes/ROOT +++ b/thys/Arith_Prog_Rel_Primes/ROOT @@ -1,10 +1,9 @@ chapter AFP - session Arith_Prog_Rel_Primes = "HOL-Number_Theory" + options [timeout = 300] theories Arith_Prog_Rel_Primes document_files "root.tex" "root.bib" diff --git a/thys/Auto2_HOL/HOL/Arith_Thms.thy b/thys/Auto2_HOL/HOL/Arith_Thms.thy --- a/thys/Auto2_HOL/HOL/Arith_Thms.thy +++ b/thys/Auto2_HOL/HOL/Arith_Thms.thy @@ -1,252 +1,252 @@ (* File: Arith_Thms.thy Author: Bohua Zhan Setup for proof steps related to arithmetic, mostly on natural numbers. *) section \Setup for arithmetic\ theory Arith_Thms imports Order_Thms HOL.Binomial begin (* Reducing inequality on natural numbers. *) theorem reduce_le_plus_consts: "(x::nat) + n1 \ y + n2 \ x \ y + (n2-n1)" by simp theorem reduce_le_plus_consts': "n1 \ n2 \ (x::nat) + n1 \ y + n2 \ x + (n1-n2) \ y" by simp theorem reduce_less_plus_consts: "(x::nat) + n1 < y + n2 \ x < y + (n2-n1)" by simp theorem reduce_less_plus_consts': "n1 \ n2 \ (x::nat) + n1 < y + n2 \ x + (n1-n2) < y" by simp (* To normal form. *) theorem norm_less_lminus: "(x::nat) - n < y \ x \ y + (n-1)" by simp theorem norm_less_lplus: "(x::nat) + n < y \ x + (n+1) \ y" by simp theorem norm_less_rminus: "(x::nat) < y - n \ x + (n+1) \ y" by simp theorem norm_less_rplus: "(x::nat) < y + n \ x \ y + (n-1)" by simp theorem norm_less: "(x::nat) < y \ x + 1 \ y" by simp theorem norm_le_lminus: "(x::nat) - n \ y \ x \ y + n" by simp theorem norm_le_rminus: "(x::nat) \ y - n \ x \ y + 0" by simp theorem norm_le: "(x::nat) \ y \ x \ y + 0" by simp theorem norm_le_lplus0: "(x::nat) + 0 \ y \ x \ y + 0" by simp (* Transitive resolve. *) theorem trans_resolve1: "n1 > 0 \ (x::nat) + n1 \ y \ (y::nat) + n2 \ x \ False" by simp theorem trans_resolve2: "n1 > n2 \ (x::nat) + n1 \ y \ (y::nat) \ x + n2 \ False" by simp (* Transitive. *) theorem trans1: "(x::nat) + n1 \ y \ y + n2 \ z \ x + (n1+n2) \ z" by simp theorem trans2: "(x::nat) \ y + n1 \ y \ z + n2 \ x \ z + (n1+n2)" by simp theorem trans3: "(x::nat) + n1 \ y \ y \ z + n2 \ x \ z + (n2-n1)" by simp theorem trans4: "n1 > n2 \ (x::nat) + n1 \ y \ y \ z + n2 \ x + (n1-n2) \ z" by simp theorem trans5: "(x::nat) \ y + n1 \ y + n2 \ z \ x \ z + (n1-n2)" by simp theorem trans6: "n2 > n1 \ (x::nat) \ y + n1 \ y + n2 \ z \ x + (n2-n1) \ z" by simp (* Resolve. *) theorem single_resolve: "n > 0 \ (x::nat) + n \ x \ False" by simp theorem single_resolve_const: "n > 0 \ (x::nat) + n \ 0 \ False" by simp (* Comparison with constants. *) theorem cv_const1: "(x::nat) + n \ y \ 0 + (x+n) \ y" by simp (* x is const *) theorem cv_const2: "(x::nat) + n \ y \ x \ 0 + (y-n)" by simp (* y is const *) theorem cv_const3: "y < n \ (x::nat) + n \ y \ x + (n-y) \ 0" by simp (* y is const (contradiction with 0 \ x) *) theorem cv_const4: "(x::nat) \ y + n \ 0 + (x-n) \ y" by simp (* x is const *) theorem cv_const5: "(x::nat) \ y + n \ 0 \ y + (n-x)" by simp (* x is const (trivial) *) theorem cv_const6: "(x::nat) \ y + n \ x \ 0 + (y+n)" by simp (* y is const *) (* Misc *) theorem nat_eq_to_ineqs: "(x::nat) = y + n \ x \ y + n \ x \ y + n" by simp theorem nat_ineq_impl_not_eq: "(x::nat) + n \ y \ n > 0 \ x \ y" by simp theorem eq_to_ineqs: "(x::nat) \ y \ x \ y \ y \ x" by simp theorem ineq_to_eqs1: "(x::nat) \ y + 0 \ y \ x + 0 \ x = y" by simp ML_file \arith.ML\ ML_file \order.ML\ ML_file \order_test.ML\ setup \register_wellform_data ("(a::nat) - b", ["a \ b"])\ setup \add_prfstep_check_req ("(a::nat) - b", "(a::nat) \ b")\ (* Normalize any expression to "a - b" form. *) lemma nat_sub_norm: "(a::nat) = a - 0 \ a \ 0" by simp (* Adding and subtracting two normalized expressions. *) lemma nat_sub1: "(a::nat) \ b \ c \ d \ (a - b) + (c - d) = (a + c) - (b + d) \ a + c \ b + d" by simp lemma nat_sub2: "(a::nat) \ b \ c \ d \ a - b \ c - d \ (a - b) - (c - d) = (a + d) - (b + c) \ a + d \ b + c" by simp lemma nat_sub3: "(a::nat) \ b \ c \ d \ (a - b) * (c - d) = (a * c + b * d) - (a * d + b * c) \ a * c + b * d \ a * d + b * c" - by (smt diff_mult_distrib mult.commute mult_le_mono1 nat_sub2) + by (smt (verit) diff_mult_distrib mult.commute mult_le_mono1 nat_sub2) (* Cancel identical terms on two sides, yielding a normalized expression. *) lemma nat_sub_combine: "(a::nat) + b \ c + b \ (a + b) - (c + b) = a - c \ a \ c" by simp lemma nat_sub_combine2: "n \ m \ (a::nat) + b * n \ c + b * m \ (a + b * n) - (c + b * m) = (a + b * (n - m)) - c \ a + b * (n - m) \ c \ n \ m" by (simp add: add.commute right_diff_distrib') lemma nat_sub_combine3: "n \ m \ (a::nat) + b * n \ c + b * m \ (a + b * n) - (c + b * m) = a - (c + b * (m - n)) \ a \ c + b * (m - n) \ m \ n" - by (smt add.commute mult.commute nat_diff_add_eq2 nat_le_add_iff1) + by (smt (verit) add.commute mult.commute nat_diff_add_eq2 nat_le_add_iff1) ML_file \nat_sub.ML\ ML_file \nat_sub_test.ML\ (* Ordering on Nats. *) lemma le_neq_implies_less' [forward]: "(m::nat) \ n \ m \ n \ m < n" by simp lemma le_zero_to_equal_zero [forward]: "(n::nat) \ 0 \ n = 0" by simp lemma less_one_to_equal_zero [forward]: "(n::nat) < 1 \ n = 0" by simp setup \add_backward_prfstep_cond @{thm Nat.mult_le_mono1} [with_cond "?k \ 1"]\ setup \add_resolve_prfstep @{thm Nat.not_add_less1}\ lemma not_minus_less [resolve]: "\(i::nat) < (i - j)" by simp lemma nat_le_prod_with_same [backward]: "m \ 0 \ (n::nat) \ m * n" by simp lemma nat_le_prod_with_le [backward1]: "k \ 0 \ (n::nat) \ m \ (n::nat) \ k * m" using le_trans nat_le_prod_with_same by blast lemma nat_plus_le_to_less [backward1]: "b \ 0 \ (a::nat) + b \ c \ a < c" by simp lemma nat_plus_le_to_less2 [backward1]: "a \ 0 \ (a::nat) + b \ c \ b < c" by simp setup \add_forward_prfstep @{thm add_right_imp_eq}\ setup \add_forward_prfstep @{thm add_left_imp_eq}\ setup \add_rewrite_rule_cond @{thm Nat.le_diff_conv2} [with_term "?i + ?k"]\ lemma nat_less_diff_conv: "(i::nat) < j - k \ i + k < j" by simp setup \add_forward_prfstep_cond @{thm nat_less_diff_conv} [with_cond "?k \ ?NUMC", with_term "?i + ?k"]\ lemma Nat_le_diff_conv2_same [forward]: "i \ j \ (i::nat) \ i - j \ j = 0" by simp lemma nat_gt_zero [forward]: "b - a > 0 \ b > (a::nat)" by simp lemma n_minus_1_less_n: "(n::nat) \ 1 \ n - 1 < n" by simp setup \add_forward_prfstep_cond @{thm n_minus_1_less_n} [with_term "?n - 1"]\ (* Monotonicity of ordering *) setup \add_backward_prfstep @{thm Nat.diff_le_mono}\ setup \add_backward2_prfstep @{thm Nat.diff_less_mono}\ setup \add_backward_prfstep @{thm Nat.mult_le_mono2}\ setup \add_resolve_prfstep @{thm Nat.le_add1}\ setup \add_resolve_prfstep @{thm Nat.le_add2}\ setup \add_backward_prfstep @{thm add_left_mono}\ setup \add_backward_prfstep @{thm add_right_mono}\ lemma add_mono_neutr [backward]: "(0::'a::linordered_ring) \ b \ a \ a + b" by simp lemma add_mono_neutl [backward]: "(0::'a::linordered_ring) \ b \ a \ b + a" by simp setup \add_forward_prfstep @{thm add_less_imp_less_left}\ lemma sum_le_zero1 [forward]: "(a::'a::linordered_ring) + b < 0 \ a \ 0 \ b < 0" by (meson add_less_same_cancel1 less_le_trans) lemma less_sum1 [backward]: "b > 0 \ a < a + (b::nat)" by auto setup \add_backward_prfstep @{thm Nat.trans_less_add2}\ setup \add_backward_prfstep @{thm Nat.add_less_mono1}\ setup \add_backward1_prfstep @{thm Nat.add_less_mono}\ setup \add_backward1_prfstep @{thm Nat.add_le_mono}\ setup \add_backward1_prfstep @{thm add_increasing2}\ setup \add_backward1_prfstep @{thm add_mono}\ setup \add_backward_prfstep @{thm add_strict_left_mono}\ setup \add_backward1_prfstep @{thm Nat.mult_le_mono}\ (* Addition. *) theorem nat_add_eq_self_zero [forward]: "(m::nat) = m + n \ n = 0" by simp theorem nat_add_eq_self_zero' [forward]: "(m::nat) = n + m \ n = 0" by simp theorem nat_mult_2: "(a::nat) + a = 2 * a" by simp setup \add_rewrite_rule_cond @{thm nat_mult_2} [with_cond "?a \ 0"]\ theorem plus_one_non_zero [resolve]: "\(n::nat) + 1 = 0" by simp (* Diff. *) lemma nat_same_minus_ge [forward]: "(c::nat) - a \ c - b \ a \ c \ a \ b" by arith lemma diff_eq_zero [forward]: "(k::nat) \ j \ j - k = 0 \ j = k" by simp lemma diff_eq_zero' [forward]: "(k::nat) \ j \ j - k + i = j \ k = i" by simp (* Divides. *) theorem dvd_defD1 [resolve]: "(a::nat) dvd b \ \k. b = a * k" using dvdE by blast theorem dvd_defD2 [resolve]: "(a::nat) dvd b \ \k. b = k * a" by (metis dvd_mult_div_cancel mult.commute) setup \add_forward_prfstep @{thm Nat.dvd_imp_le}\ theorem dvd_ineq2 [forward]: "(k::nat) dvd n \ n > 0 \ k \ 1" by (simp add: Suc_leI dvd_pos_nat) setup \add_forward_prfstep_cond @{thm dvd_trans} (with_conds ["?a \ ?b", "?b \ ?c", "?a \ ?c"])\ setup \add_forward_prfstep_cond @{thm Nat.dvd_antisym} [with_cond "?m \ ?n"]\ theorem dvd_cancel [backward1]: "c > 0 \ (a::nat) * c dvd b * c \ a dvd b" by simp setup \add_forward_prfstep (equiv_forward_th @{thm dvd_add_right_iff})\ (* Divisibility: existence. *) setup \add_resolve_prfstep @{thm dvd_refl}\ theorem exists_n_dvd_n [backward]: "P (n::nat) \ \k. k dvd n \ P k" using dvd_refl by blast setup \add_resolve_prfstep @{thm one_dvd}\ theorem any_n_dvd_0 [forward]: "\ (\ k. k dvd (0::nat) \ P k) \ \ (\ k. P k)" by simp theorem n_dvd_one: "(n::nat) dvd 1 \ n = 1" by simp setup \add_forward_prfstep_cond @{thm n_dvd_one} [with_cond "?n \ 1"]\ (* Products. *) setup \add_rewrite_rule @{thm mult_zero_left}\ lemma prod_ineqs1 [forward]: "(m::nat) * k > 0 \ m > 0 \ k > 0" by simp lemma prod_ineqs2 [backward]: "(k::nat) > 0 \ m \ m * k" by simp theorem prod_cancel: "(a::nat) * b = a * c \ a > 0 \ b = c" by auto setup \add_forward_prfstep_cond @{thm prod_cancel} [with_cond "?b \ ?c"]\ theorem mult_n1n: "(n::nat) = m * n \ n > 0 \ m = 1" by auto setup \add_forward_prfstep_cond @{thm mult_n1n} [with_cond "?m \ 1"]\ theorem prod_is_one [forward]: "(x::nat) * y = 1 \ x = 1" by simp theorem prod_dvd_intro [backward]: "(k::nat) dvd m \ k dvd n \ k dvd m * n" using dvd_mult dvd_mult2 by blast (* Definition of gcd. *) setup \add_forward_prfstep_cond @{thm gcd_dvd1} [with_term "gcd ?a ?b"]\ setup \add_forward_prfstep_cond @{thm gcd_dvd2} [with_term "gcd ?a ?b"]\ (* Coprimality. *) setup \add_rewrite_rule_bidir @{thm coprime_iff_gcd_eq_1}\ lemma coprime_exp [backward]: "coprime d a \ coprime (d::nat) (a ^ n)" by simp setup \add_backward_prfstep @{thm coprime_exp}\ setup \add_rewrite_rule @{thm gcd.commute}\ lemma coprime_dvd_mult [backward1]: "coprime (a::nat) b \ a dvd c * b \ a dvd c" by (metis coprime_dvd_mult_left_iff) lemma coprime_dvd_mult' [backward1]: "coprime (a::nat) b \ a dvd b * c \ a dvd c" by (metis coprime_dvd_mult_right_iff) theorem coprime_dvd [forward]: "coprime (a::nat) b \ p dvd a \ p > 1 \ \ p dvd b" using coprime_common_divisor_nat by blast (* Powers. *) setup \add_rewrite_rule @{thm power_0}\ theorem power_ge_0 [rewrite]: "m \ 0 \ p ^ m = p * (p ^ (m - 1))" by (simp add: power_eq_if) setup \add_rewrite_rule_cond @{thm power_one} [with_cond "?n \ 0"]\ setup \add_rewrite_rule_cond @{thm power_one_right} [with_cond "?a \ 1"]\ setup \add_gen_prfstep ("power_case_intro", [WithTerm @{term_pat "?p ^ (?FREE::nat)"}, CreateCase @{term_pat "(?FREE::nat) = 0"}])\ lemma one_is_power_of_any [resolve]: "\i. (1::nat) = a ^ i" by (metis power.simps(1)) setup \add_rewrite_rule @{thm power_Suc}\ theorem power_dvd [forward]: "(p::nat)^n dvd a \ n \ 0 \ p dvd a" using dvd_power dvd_trans by blast theorem power_eq_one: "(b::nat) ^ n = 1 \ b = 1 \ n = 0" by (metis One_nat_def Suc_lessI nat_zero_less_power_iff power_0 power_inject_exp) setup \add_forward_prfstep_cond @{thm power_eq_one} (with_conds ["?b \ 1", "?n \ 0"])\ (* Factorial. *) theorem fact_ge_1_nat: "fact (n::nat) \ (1::nat)" by simp setup \add_forward_prfstep_cond @{thm fact_ge_1_nat} [with_term "fact ?n"]\ setup \add_backward1_prfstep @{thm dvd_fact}\ (* Successor function. *) setup \add_rewrite_rule @{thm Nat.Suc_eq_plus1}\ setup \add_backward_prfstep @{thm Nat.gr0_implies_Suc}\ (* Cases *) setup \fold add_rewrite_rule @{thms Nat.nat.case}\ (* Induction. *) lemma nat_cases: "P 0 \ (\n. P (Suc n)) \ P n" using nat_induct by auto (* div *) declare times_div_less_eq_dividend [resolve] setup \ add_var_induct_rule @{thm nat_induct} #> add_strong_induct_rule @{thm nat_less_induct} #> add_cases_rule @{thm nat_cases} \ end diff --git a/thys/Banach_Steinhaus/Banach_Steinhaus.thy b/thys/Banach_Steinhaus/Banach_Steinhaus.thy --- a/thys/Banach_Steinhaus/Banach_Steinhaus.thy +++ b/thys/Banach_Steinhaus/Banach_Steinhaus.thy @@ -1,481 +1,482 @@ (* File: Banach_Steinhaus.thy Author: Dominique Unruh, University of Tartu Author: Jose Manuel Rodriguez Caballero, University of Tartu *) section \Banach-Steinhaus theorem\ theory Banach_Steinhaus imports Banach_Steinhaus_Missing begin text \ We formalize Banach-Steinhaus theorem as theorem @{text banach_steinhaus}. This theorem was originally proved in Banach-Steinhaus's paper~\<^cite>\"banach1927principe"\. For the proof, we follow Sokal's approach~\<^cite>\"sokal2011really"\. Furthermore, we prove as a corollary a result about pointwise convergent sequences of bounded operators whose domain is a Banach space. \ subsection \Preliminaries for Sokal's proof of Banach-Steinhaus theorem\ lemma linear_plus_norm: includes notation_norm assumes \linear f\ shows \\f \\ \ max \f (x + \)\ \f (x - \)\\ text \ Explanation: For arbitrary \<^term>\x\ and a linear operator \<^term>\f\, \<^term>\norm (f \)\ is upper bounded by the maximum of the norms of the shifts of \<^term>\f\ (i.e., \<^term>\f (x + \)\ and \<^term>\f (x - \)\). \ proof- have \norm (f \) = norm ( (inverse (of_nat 2)) *\<^sub>R (f (x + \) - f (x - \)) )\ - by (smt add_diff_cancel_left' assms diff_add_cancel diff_diff_add linear_diff midpoint_def - midpoint_plus_self of_nat_1 of_nat_add one_add_one scaleR_half_double) + by (metis (no_types, opaque_lifting) add.commute assms diff_diff_eq2 group_cancel.sub1 + linear_cmul linear_diff of_nat_numeral real_vector_affinity_eq scaleR_2 + scaleR_right_diff_distrib zero_neq_numeral) also have \\ = inverse (of_nat 2) * norm (f (x + \) - f (x - \))\ using Real_Vector_Spaces.real_normed_vector_class.norm_scaleR by simp also have \\ \ inverse (of_nat 2) * (norm (f (x + \)) + norm (f (x - \)))\ by (simp add: norm_triangle_ineq4) also have \\ \ max (norm (f (x + \))) (norm (f (x - \)))\ by auto finally show ?thesis by blast qed lemma onorm_Sup_on_ball: includes notation_norm assumes \r > 0\ shows "\f\ \ Sup ( (\x. \f *\<^sub>v x\) ` (ball x r) ) / r" text \ Explanation: Let \<^term>\f\ be a bounded operator and let \<^term>\x\ be a point. For any \<^term>\r > 0\, the operator norm of \<^term>\f\ is bounded above by the supremum of $f$ applied to the open ball of radius \<^term>\r\ around \<^term>\x\, divided by \<^term>\r\. \ proof- have bdd_above_3: \bdd_above ((\x. \f *\<^sub>v x\) ` (ball 0 r))\ proof - obtain M where \\ \. \f *\<^sub>v \\ \ M * norm \\ and \M \ 0\ using norm_blinfun norm_ge_zero by blast hence \\ \. \ \ ball 0 r \ \f *\<^sub>v \\ \ M * r\ - using \r > 0\ by (smt mem_ball_0 mult_left_mono) + using \r > 0\ by (smt (verit) mem_ball_0 mult_left_mono) thus ?thesis by (meson bdd_aboveI2) qed have bdd_above_2: \bdd_above ((\ \. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ proof- have \bdd_above ((\ \. \f *\<^sub>v x\) ` (ball 0 r))\ by auto moreover have \bdd_above ((\ \. \f *\<^sub>v \\) ` (ball 0 r))\ using bdd_above_3 by blast ultimately have \bdd_above ((\ \. \f *\<^sub>v x\ + \f *\<^sub>v \\) ` (ball 0 r))\ by (rule bdd_above_plus) then obtain M where \\ \. \ \ ball 0 r \ \f *\<^sub>v x\ + \f *\<^sub>v \\ \ M\ unfolding bdd_above_def by (meson image_eqI) moreover have \\f *\<^sub>v (x + \)\ \ \f *\<^sub>v x\ + \f *\<^sub>v \\\ for \ by (simp add: blinfun.add_right norm_triangle_ineq) ultimately have \\ \. \ \ ball 0 r \ \f *\<^sub>v (x + \)\ \ M\ by (simp add: blinfun.add_right norm_triangle_le) thus ?thesis by (meson bdd_aboveI2) qed have bdd_above_4: \bdd_above ((\ \. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ proof- obtain K where K_def: \\ \. \ \ ball 0 r \ \f *\<^sub>v (x + \)\ \ K\ using \bdd_above ((\ \. norm (f (x + \))) ` (ball 0 r))\ unfolding bdd_above_def by (meson image_eqI) have \\ \ ball (0::'a) r \ -\ \ ball 0 r\ for \ by auto thus ?thesis by (metis K_def ab_group_add_class.ab_diff_conv_add_uminus bdd_aboveI2) qed have bdd_above_1: \bdd_above ((\ \. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ proof- have \bdd_above ((\ \. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ using bdd_above_2 by blast moreover have \bdd_above ((\ \. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ using bdd_above_4 by blast ultimately show ?thesis unfolding max_def apply auto apply (meson bdd_above_Int1 bdd_above_mono image_Int_subset) by (meson bdd_above_Int1 bdd_above_mono image_Int_subset) qed have bdd_above_6: \bdd_above ((\t. \f *\<^sub>v t\) ` ball x r)\ proof- have \bounded (ball x r)\ by simp hence \bounded ((\t. \f *\<^sub>v t\) ` ball x r)\ by (metis (no_types) add.left_neutral bdd_above_2 bdd_above_norm bounded_norm_comp image_add_ball image_image) thus ?thesis by (simp add: bounded_imp_bdd_above) qed have norm_1: \(\\. \f *\<^sub>v (x + \)\) ` ball 0 r = (\t. \f *\<^sub>v t\) ` ball x r\ by (metis add.right_neutral ball_translation image_image) have bdd_above_5: \bdd_above ((\\. norm (f (x + \))) ` ball 0 r)\ by (simp add: bdd_above_2) have norm_2: \\\\ < r \ \f *\<^sub>v (x - \)\ \ (\\. \f *\<^sub>v (x + \)\) ` ball 0 r\ for \ proof- assume \\\\ < r\ hence \\ \ ball (0::'a) r\ by auto hence \-\ \ ball (0::'a) r\ by auto thus ?thesis by (metis (no_types, lifting) ab_group_add_class.ab_diff_conv_add_uminus image_iff) qed have norm_2': \\\\ < r \ \f *\<^sub>v (x + \)\ \ (\\. \f *\<^sub>v (x - \)\) ` ball 0 r\ for \ proof- assume \norm \ < r\ hence \\ \ ball (0::'a) r\ by auto hence \-\ \ ball (0::'a) r\ by auto thus ?thesis by (metis (no_types, lifting) diff_minus_eq_add image_iff) qed have bdd_above_6: \bdd_above ((\\. \f *\<^sub>v (x - \)\) ` ball 0 r)\ by (simp add: bdd_above_4) have Sup_2: \(SUP \\ball 0 r. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) = max (SUP \\ball 0 r. \f *\<^sub>v (x + \)\) (SUP \\ball 0 r. \f *\<^sub>v (x - \)\)\ proof- have \ball (0::'a) r \ {}\ using \r > 0\ by auto moreover have \bdd_above ((\\. \f *\<^sub>v (x + \)\) ` ball 0 r)\ using bdd_above_5 by blast moreover have \bdd_above ((\\. \f *\<^sub>v (x - \)\) ` ball 0 r)\ using bdd_above_6 by blast ultimately show ?thesis using max_Sup by (metis (mono_tags, lifting) Banach_Steinhaus_Missing.pointwise_max_def image_cong) qed have Sup_3': \\\\ < r \ \f *\<^sub>v (x + \)\ \ (\\. \f *\<^sub>v (x - \)\) ` ball 0 r\ for \::'a by (simp add: norm_2') have Sup_3'': \\\\ < r \ \f *\<^sub>v (x - \)\ \ (\\. \f *\<^sub>v (x + \)\) ` ball 0 r\ for \::'a by (simp add: norm_2) have Sup_3: \max (SUP \\ball 0 r. \f *\<^sub>v (x + \)\) (SUP \\ball 0 r. \f *\<^sub>v (x - \)\) = (SUP \\ball 0 r. \f *\<^sub>v (x + \)\)\ proof- have \(\\. \f *\<^sub>v (x + \)\) ` (ball 0 r) = (\\. \f *\<^sub>v (x - \)\) ` (ball 0 r)\ apply auto using Sup_3' apply auto using Sup_3'' by blast hence \Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))=Sup ((\\. \f *\<^sub>v (x - \)\) ` (ball 0 r))\ by simp thus ?thesis by simp qed have Sup_1: \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) \ Sup ( (\\. \f *\<^sub>v \\) ` (ball x r) )\ proof- have \(\t. \f *\<^sub>v t\) \ \ max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\\ for \ apply(rule linear_plus_norm) apply (rule bounded_linear.linear) by (simp add: blinfun.bounded_linear_right) moreover have \bdd_above ((\ \. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ using bdd_above_1 by blast moreover have \ball (0::'a) r \ {}\ using \r > 0\ by auto ultimately have \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) \ Sup ((\\. max \f *\<^sub>v (x + \)\ \f *\<^sub>v (x - \)\) ` (ball 0 r))\ - using cSUP_mono by smt + using cSUP_mono by (smt (verit)) also have \\ = max (Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))) (Sup ((\\. \f *\<^sub>v (x - \)\) ` (ball 0 r)))\ using Sup_2 by blast also have \\ = Sup ((\\. \f *\<^sub>v (x + \)\) ` (ball 0 r))\ using Sup_3 by blast also have \\ = Sup ((\\. \f *\<^sub>v \\) ` (ball x r))\ by (metis add.right_neutral ball_translation image_image) finally show ?thesis by blast qed have \\f\ = (SUP x\ball 0 r. \f *\<^sub>v x\) / r\ using \0 < r\ onorm_r by blast moreover have \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) / r \ Sup ((\\. \f *\<^sub>v \\) ` (ball x r)) / r\ using Sup_1 \0 < r\ divide_right_mono by fastforce ultimately have \\f\ \ Sup ((\t. \f *\<^sub>v t\) ` ball x r) / r\ by simp thus ?thesis by simp qed lemma onorm_Sup_on_ball': includes notation_norm assumes \r > 0\ and \\ < 1\ shows \\\\ball x r. \ * r * \f\ \ \f *\<^sub>v \\\ text \ In the proof of Banach-Steinhaus theorem, we will use this variation of the lemma @{text onorm_Sup_on_ball}. Explanation: Let \<^term>\f\ be a bounded operator, let \<^term>\x\ be a point and let \<^term>\r\ be a positive real number. For any real number \<^term>\\ < 1\, there is a point \<^term>\\\ in the open ball of radius \<^term>\r\ around \<^term>\x\ such that \<^term>\\ * r * \f\ \ \f *\<^sub>v \\\. \ proof(cases \f = 0\) case True thus ?thesis by (metis assms(1) centre_in_ball mult_zero_right norm_zero order_refl zero_blinfun.rep_eq) next case False have bdd_above_1: \bdd_above ((\t. \(*\<^sub>v) f t\) ` ball x r)\ for f::\'a \\<^sub>L 'b\ using assms(1) bounded_linear_image by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) have \norm f > 0\ using \f \ 0\ by auto have \norm f \ Sup ( (\\. \(*\<^sub>v) f \\) ` (ball x r) ) / r\ using \r > 0\ by (simp add: onorm_Sup_on_ball) hence \r * norm f \ Sup ( (\\. \(*\<^sub>v) f \\) ` (ball x r) )\ - using \0 < r\ by (smt divide_strict_right_mono nonzero_mult_div_cancel_left) + using \0 < r\ by (smt (verit) divide_strict_right_mono nonzero_mult_div_cancel_left) moreover have \\ * r * norm f < r * norm f\ using \\ < 1\ using \0 < norm f\ \0 < r\ by auto ultimately have \\ * r * norm f < Sup ( (norm \ ((*\<^sub>v) f)) ` (ball x r) )\ by simp moreover have \(norm \ ( (*\<^sub>v) f)) ` (ball x r) \ {}\ using \0 < r\ by auto moreover have \bdd_above ((norm \ ( (*\<^sub>v) f)) ` (ball x r))\ using bdd_above_1 apply transfer by simp ultimately have \\t \ (norm \ ( (*\<^sub>v) f)) ` (ball x r). \ * r * norm f < t\ by (simp add: less_cSup_iff) - thus ?thesis by (smt comp_def image_iff) + thus ?thesis by (smt (verit) comp_def image_iff) qed subsection \Banach-Steinhaus theorem\ theorem banach_steinhaus: fixes f::\'c \ ('a::banach \\<^sub>L 'b::real_normed_vector)\ assumes \\x. bounded (range (\n. (f n) *\<^sub>v x))\ shows \bounded (range f)\ text\ This is Banach-Steinhaus Theorem. Explanation: If a family of bounded operators on a Banach space is pointwise bounded, then it is uniformly bounded. \ proof(rule classical) assume \\(bounded (range f))\ have sum_1: \\K. \n. sum (\k. inverse (real_of_nat 3^k)) {0..n} \ K\ proof- have \summable (\n. inverse ((3::real) ^ n))\ by (simp flip: power_inverse) hence \bounded (range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. using summable_imp_sums_bounded[where f = "(\n. inverse (real_of_nat 3^n))"] lessThan_atLeast0 by auto hence \\M. \h\(range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. M\ using bounded_iff by blast then obtain M where \h\range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. norm h \ M\ for h by blast have sum_2: \sum (\k. inverse (real_of_nat 3^k)) {0..n} \ M\ for n proof- have \norm (sum (\ k. inverse (real 3 ^ k)) {0..< Suc n}) \ M\ using \\h. h\(range (\n. sum (\ k. inverse (real 3 ^ k)) {0.. norm h \ M\ by blast hence \norm (sum (\ k. inverse (real 3 ^ k)) {0..n}) \ M\ by (simp add: atLeastLessThanSuc_atLeastAtMost) hence \(sum (\ k. inverse (real 3 ^ k)) {0..n}) \ M\ by auto thus ?thesis by blast qed have \sum (\k. inverse (real_of_nat 3^k)) {0..n} \ M\ for n using sum_2 by blast thus ?thesis by blast qed have \of_rat 2/3 < (1::real)\ by auto hence \\g::'a \\<^sub>L 'b. \x. \r. \\. g \ 0 \ r > 0 \ (\\ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g \))\ using onorm_Sup_on_ball' by blast hence \\\. \g::'a \\<^sub>L 'b. \x. \r. g \ 0 \ r > 0 \ ((\ g x r)\ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g (\ g x r)))\ by metis then obtain \ where f1: \\g \ 0; r > 0\ \ \ g x r \ ball x r \ (of_rat 2/3) * r * norm g \ norm ((*\<^sub>v) g (\ g x r))\ for g::\'a \\<^sub>L 'b\ and x and r by blast have \\n. \k. norm (f k) \ 4^n\ using \\(bounded (range f))\ by (metis (mono_tags, opaque_lifting) boundedI image_iff linear) hence \\k. \n. norm (f (k n)) \ 4^n\ by metis hence \\k. \n. norm ((f \ k) n) \ 4^n\ by simp then obtain k where \norm ((f \ k) n) \ 4^n\ for n by blast define T where \T = f \ k\ have \T n \ range f\ for n unfolding T_def by simp have \norm (T n) \ of_nat (4^n)\ for n unfolding T_def using \\ n. norm ((f \ k) n) \ 4^n\ by auto hence \T n \ 0\ for n - by (smt T_def \\n. 4 ^ n \ norm ((f \ k) n)\ norm_zero power_not_zero zero_le_power) + by (smt (verit) T_def \\n. 4 ^ n \ norm ((f \ k) n)\ norm_zero power_not_zero zero_le_power) have \inverse (of_nat 3^n) > (0::real)\ for n by auto define y::\nat \ 'a\ where \y = rec_nat 0 (\n x. \ (T n) x (inverse (of_nat 3^n)))\ have \y (Suc n) \ ball (y n) (inverse (of_nat 3^n))\ for n using f1 \\ n. T n \ 0\ \\ n. inverse (of_nat 3^n) > 0\ unfolding y_def by auto hence \norm (y (Suc n) - y n) \ inverse (of_nat 3^n)\ for n - unfolding ball_def apply auto using dist_norm by (smt norm_minus_commute) + unfolding ball_def apply auto using dist_norm by (smt (verit) norm_minus_commute) moreover have \\K. \n. sum (\k. inverse (real_of_nat 3^k)) {0..n} \ K\ using sum_1 by blast moreover have \Cauchy y\ using convergent_series_Cauchy[where a = "\n. inverse (of_nat 3^n)" and \ = y] dist_norm by (metis calculation(1) calculation(2)) hence \\ x. y \ x\ by (simp add: convergent_eq_Cauchy) then obtain x where \y \ x\ by blast have norm_2: \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ for n proof- have \inverse (real_of_nat 3) < 1\ by simp moreover have \y 0 = 0\ using y_def by auto ultimately have \norm (x - y (Suc n)) \ (inverse (of_nat 3)) * inverse (1 - (inverse (of_nat 3))) * ((inverse (of_nat 3)) ^ n)\ using bound_Cauchy_to_lim[where c = "inverse (of_nat 3)" and y = y and x = x] power_inverse semiring_norm(77) \y \ x\ \\ n. norm (y (Suc n) - y n) \ inverse (of_nat 3^n)\ by (metis divide_inverse) moreover have \inverse (real_of_nat 3) * inverse (1 - (inverse (of_nat 3))) = inverse (of_nat 2)\ by auto ultimately show ?thesis by (metis power_inverse) qed have \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ for n using norm_2 by blast have \\ M. \ n. norm ((*\<^sub>v) (T n) x) \ M\ unfolding T_def apply auto by (metis \\x. bounded (range (\n. (*\<^sub>v) (f n) x))\ bounded_iff rangeI) then obtain M where \norm ((*\<^sub>v) (T n) x) \ M\ for n by blast have norm_1: \norm (T n) * norm (y (Suc n) - x) + norm ((*\<^sub>v) (T n) x) \ inverse (real 2) * inverse (real 3 ^ n) * norm (T n) + norm ((*\<^sub>v) (T n) x)\ for n proof- have \norm (y (Suc n) - x) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ using \norm (x - y (Suc n)) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))\ by (simp add: norm_minus_commute) moreover have \norm (T n) \ 0\ by auto ultimately have \norm (T n) * norm (y (Suc n) - x) \ (inverse (of_nat 2))*(inverse (of_nat 3^n))*norm (T n)\ by (simp add: \\n. T n \ 0\) thus ?thesis by simp qed have inverse_2: \(inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ for n proof- have \(of_rat 2/3)*(inverse (of_nat 3^n))*norm (T n) \ norm ((*\<^sub>v) (T n) (y (Suc n)))\ using f1 \\ n. T n \ 0\ \\ n. inverse (of_nat 3^n) > 0\ unfolding y_def by auto also have \\ = norm ((*\<^sub>v) (T n) ((y (Suc n) - x) + x))\ by auto also have \\ = norm ((*\<^sub>v) (T n) (y (Suc n) - x) + (*\<^sub>v) (T n) x)\ apply transfer apply auto by (metis diff_add_cancel linear_simps(1)) also have \\ \ norm ((*\<^sub>v) (T n) (y (Suc n) - x)) + norm ((*\<^sub>v) (T n) x)\ by (simp add: norm_triangle_ineq) also have \\ \ norm (T n) * norm (y (Suc n) - x) + norm ((*\<^sub>v) (T n) x)\ apply transfer apply auto using onorm by auto also have \\ \ (inverse (of_nat 2))*(inverse (of_nat 3^n))*norm (T n) + norm ((*\<^sub>v) (T n) x)\ using norm_1 by blast finally have \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) \ inverse (real 2) * inverse (real 3 ^ n) * norm (T n) + norm ((*\<^sub>v) (T n) x)\ by blast hence \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) - inverse (real 2) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ by linarith moreover have \(of_rat 2/3) * inverse (real 3 ^ n) * norm (T n) - inverse (real 2) * inverse (real 3 ^ n) * norm (T n) = (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ by fastforce ultimately show \(inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n) \ norm ((*\<^sub>v) (T n) x)\ by linarith qed have inverse_3: \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ for n proof- have \of_rat (4/3)^n = inverse (real 3 ^ n) * (of_nat 4^n)\ apply auto by (metis divide_inverse_commute of_rat_divide power_divide of_rat_numeral_eq) also have \\ \ inverse (real 3 ^ n) * norm (T n)\ using \\n. norm (T n) \ of_nat (4^n)\ by simp finally have \of_rat (4/3)^n \ inverse (real 3 ^ n) * norm (T n)\ by blast moreover have \inverse (of_nat 6) > (0::real)\ by auto ultimately show ?thesis by auto qed have inverse_1: \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ M\ for n proof- have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ (inverse (of_nat 6)) * inverse (real 3 ^ n) * norm (T n)\ using inverse_3 by blast also have \\ \ norm ((*\<^sub>v) (T n) x)\ using inverse_2 by blast finally have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ norm ((*\<^sub>v) (T n) x)\ by auto - thus ?thesis using \\ n. norm ((*\<^sub>v) (T n) x) \ M\ by smt + thus ?thesis using \\ n. norm ((*\<^sub>v) (T n) x) \ M\ by (smt (verit)) qed have \\n. M < (inverse (of_nat 6)) * (of_rat (4/3)^n)\ using Real.real_arch_pow by auto moreover have \(inverse (of_nat 6)) * (of_rat (4/3)^n) \ M\ for n using inverse_1 by blast - ultimately show ?thesis by smt + ultimately show ?thesis by (smt (verit)) qed subsection \A consequence of Banach-Steinhaus theorem\ corollary bounded_linear_limit_bounded_linear: fixes f::\nat \ ('a::banach \\<^sub>L 'b::real_normed_vector)\ assumes \\x. convergent (\n. (f n) *\<^sub>v x)\ shows \\g. (\n. (*\<^sub>v) (f n)) \pointwise\ (*\<^sub>v) g\ text\ Explanation: If a sequence of bounded operators on a Banach space converges pointwise, then the limit is also a bounded operator. \ proof- have \\l. (\n. (*\<^sub>v) (f n) x) \ l\ for x by (simp add: \\x. convergent (\n. (*\<^sub>v) (f n) x)\ convergentD) hence \\F. (\n. (*\<^sub>v) (f n)) \pointwise\ F\ unfolding pointwise_convergent_to_def by metis obtain F where \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ using \\F. (\n. (*\<^sub>v) (f n)) \pointwise\ F\ by auto have \\x. (\ n. (*\<^sub>v) (f n) x) \ F x\ using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ apply transfer by (simp add: pointwise_convergent_to_def) have \bounded (range f)\ using \\x. convergent (\n. (*\<^sub>v) (f n) x)\ banach_steinhaus \\x. \l. (\n. (*\<^sub>v) (f n) x) \ l\ convergent_imp_bounded by blast have norm_f_n: \\ M. \ n. norm (f n) \ M\ unfolding bounded_def by (meson UNIV_I \bounded (range f)\ bounded_iff image_eqI) have \isCont (\ t::'b. norm t) y\ for y::'b using Limits.isCont_norm by simp hence \(\ n. norm ((*\<^sub>v) (f n) x)) \ (norm (F x))\ for x using \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ by (simp add: tendsto_norm) hence norm_f_n_x: \\ M. \ n. norm ((*\<^sub>v) (f n) x) \ M\ for x using Elementary_Metric_Spaces.convergent_imp_bounded by (metis UNIV_I \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ bounded_iff image_eqI) have norm_f: \\K. \n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * K\ proof- have \\ M. \ n. norm ((*\<^sub>v) (f n) x) \ M\ for x using norm_f_n_x \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ by blast hence \\ M. \ n. norm (f n) \ M\ using norm_f_n by simp then obtain M::real where \\ M. \ n. norm (f n) \ M\ by blast have \\ n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * norm (f n)\ apply transfer apply auto by (metis mult.commute onorm) thus ?thesis using \\ M. \ n. norm (f n) \ M\ by (metis (no_types, opaque_lifting) dual_order.trans norm_eq_zero order_refl mult_le_cancel_left_pos vector_space_over_itself.scale_zero_left zero_less_norm_iff) qed have norm_F_x: \\K. \x. norm (F x) \ norm x * K\ proof- have "\K. \n. \x. norm ((*\<^sub>v) (f n) x) \ norm x * K" using norm_f \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ by auto thus ?thesis using \\ x::'a. (\ n. (*\<^sub>v) (f n) x) \ F x\ apply transfer by (metis Lim_bounded tendsto_norm) qed have \linear F\ proof(rule linear_limit_linear) show \linear ((*\<^sub>v) (f n))\ for n apply transfer apply auto by (simp add: bounded_linear.linear) show \f \pointwise\ F\ using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ by auto qed moreover have \bounded_linear_axioms F\ using norm_F_x by (simp add: \\x. (\n. (*\<^sub>v) (f n) x) \ F x\ bounded_linear_axioms_def) ultimately have \bounded_linear F\ unfolding bounded_linear_def by blast hence \\g. (*\<^sub>v) g = F\ using bounded_linear_Blinfun_apply by auto thus ?thesis using \(\n. (*\<^sub>v) (f n)) \pointwise\ F\ apply transfer by auto qed end diff --git a/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy b/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy --- a/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy +++ b/thys/Banach_Steinhaus/Banach_Steinhaus_Missing.thy @@ -1,898 +1,898 @@ (* File: Banach_Steinhaus_Missing.thy Author: Dominique Unruh, University of Tartu Author: Jose Manuel Rodriguez Caballero, University of Tartu *) section \Missing results for the proof of Banach-Steinhaus theorem\ theory Banach_Steinhaus_Missing imports "HOL-Analysis.Bounded_Linear_Function" "HOL-Analysis.Line_Segment" begin subsection \Results missing for the proof of Banach-Steinhaus theorem\ text \ The results proved here are preliminaries for the proof of Banach-Steinhaus theorem using Sokal's approach, but they do not explicitly appear in Sokal's paper \<^cite>\sokal2011really\. \ text\Notation for the norm\ bundle notation_norm begin notation norm ("\_\") end bundle no_notation_norm begin no_notation norm ("\_\") end unbundle notation_norm text\Notation for apply bilinear function\ bundle notation_blinfun_apply begin notation blinfun_apply (infixr "*\<^sub>v" 70) end bundle no_notation_blinfun_apply begin no_notation blinfun_apply (infixr "*\<^sub>v" 70) end unbundle notation_blinfun_apply lemma bdd_above_plus: fixes f::\'a \ real\ assumes \bdd_above (f ` S)\ and \bdd_above (g ` S)\ shows \bdd_above ((\ x. f x + g x) ` S)\ text \ Explanation: If the images of two real-valued functions \<^term>\f\,\<^term>\g\ are bounded above on a set \<^term>\S\, then the image of their sum is bounded on \<^term>\S\. \ proof- obtain M where \\ x. x\S \ f x \ M\ using \bdd_above (f ` S)\ unfolding bdd_above_def by blast obtain N where \\ x. x\S \ g x \ N\ using \bdd_above (g ` S)\ unfolding bdd_above_def by blast have \\ x. x\S \ f x + g x \ M + N\ using \\x. x \ S \ f x \ M\ \\x. x \ S \ g x \ N\ by fastforce thus ?thesis unfolding bdd_above_def by blast qed text\The maximum of two functions\ definition pointwise_max:: "('a \ 'b::ord) \ ('a \ 'b) \ ('a \ 'b)" where \pointwise_max f g = (\x. max (f x) (g x))\ lemma max_Sup_absorb_left: fixes f g::\'a \ real\ assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ shows \Sup ((pointwise_max f g) ` X) = Sup (f ` X)\ text \Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\, if the supremum of \<^term>\f\ is greater-equal the supremum of \<^term>\g\, then the supremum of \<^term>\max f g\ equals the supremum of \<^term>\f\. (Under some technical conditions.)\ proof- have y_Sup: \y \ ((\ x. max (f x) (g x)) ` X) \ y \ Sup (f ` X)\ for y proof- assume \y \ ((\ x. max (f x) (g x)) ` X)\ then obtain x where \y = max (f x) (g x)\ and \x \ X\ by blast have \f x \ Sup (f ` X)\ by (simp add: \x \ X\ \bdd_above (f ` X)\ cSUP_upper) moreover have \g x \ Sup (g ` X)\ by (simp add: \x \ X\ \bdd_above (g ` X)\ cSUP_upper) ultimately have \max (f x) (g x) \ Sup (f ` X)\ using \Sup (f ` X) \ Sup (g ` X)\ by auto thus ?thesis by (simp add: \y = max (f x) (g x)\) qed have y_f_X: \y \ f ` X \ y \ Sup ((\ x. max (f x) (g x)) ` X)\ for y proof- assume \y \ f ` X\ then obtain x where \x \ X\ and \y = f x\ by blast have \bdd_above ((\ \. max (f \) (g \)) ` X)\ by (metis (no_types) \bdd_above (f ` X)\ \bdd_above (g ` X)\ bdd_above_image_sup sup_max) moreover have \e > 0 \ \ k \ (\ \. max (f \) (g \)) ` X. y \ k + e\ for e::real using \Sup (f ` X) \ Sup (g ` X)\ by (smt (verit, best) \x \ X\ \y = f x\ imageI) ultimately show ?thesis using \x \ X\ \y = f x\ cSUP_upper by fastforce qed have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ using y_Sup by (simp add: \X \ {}\ cSup_least) moreover have \Sup ((\ x. max (f x) (g x)) ` X) \ Sup (f ` X)\ using y_f_X by (metis (mono_tags) cSup_least calculation empty_is_image) ultimately show ?thesis unfolding pointwise_max_def by simp qed lemma max_Sup_absorb_right: fixes f g::\'a \ real\ assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ and \Sup (f ` X) \ Sup (g ` X)\ shows \Sup ((pointwise_max f g) ` X) = Sup (g ` X)\ text \ Explanation: For real-valued functions \<^term>\f\ and \<^term>\g\ and a nonempty set \<^term>\X\, such that the \<^term>\f\ and \<^term>\g\ are bounded above on \<^term>\X\, if the supremum of \<^term>\f\ on \<^term>\X\ is lower-equal the supremum of \<^term>\g\ on \<^term>\X\, then the supremum of \<^term>\pointwise_max f g\ on \<^term>\X\ equals the supremum of \<^term>\g\. This is the right analog of @{text max_Sup_absorb_left}. \ proof- have \Sup ((pointwise_max g f) ` X) = Sup (g ` X)\ using assms by (simp add: max_Sup_absorb_left) moreover have \pointwise_max g f = pointwise_max f g\ unfolding pointwise_max_def by auto ultimately show ?thesis by simp qed lemma max_Sup: fixes f g::\'a \ real\ assumes \X \ {}\ and \bdd_above (f ` X)\ and \bdd_above (g ` X)\ shows \Sup ((pointwise_max f g) ` X) = max (Sup (f ` X)) (Sup (g ` X))\ text \ Explanation: Let \<^term>\X\ be a nonempty set. Two supremum over \<^term>\X\ of the maximum of two real-value functions is equal to the maximum of their suprema over \<^term>\X\, provided that the functions are bounded above on \<^term>\X\. \ proof(cases \Sup (f ` X) \ Sup (g ` X)\) case True thus ?thesis by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_left) next case False have f1: "\ 0 \ Sup (f ` X) + - 1 * Sup (g ` X)" using False by linarith hence "Sup (Banach_Steinhaus_Missing.pointwise_max f g ` X) = Sup (g ` X)" by (simp add: assms(1) assms(2) assms(3) max_Sup_absorb_right) thus ?thesis using f1 by linarith qed lemma identity_telescopic: fixes x :: \_ \ 'a::real_normed_vector\ assumes \x \ l\ shows \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) \ l - x n\ text\ Expression of a limit as a telescopic series. Explanation: If \<^term>\x\ converges to \<^term>\l\ then the sum \<^term>\sum (\ k. x (Suc k) - x k) {n..N}\ converges to \<^term>\l - x n\ as \<^term>\N\ goes to infinity. \ proof- have \(\ p. x (p + Suc n)) \ l\ using \x \ l\ by (rule LIMSEQ_ignore_initial_segment) hence \(\ p. x (Suc n + p)) \ l\ by (simp add: add.commute) hence \(\ p. x (Suc (n + p))) \ l\ by simp hence \(\ t. (- (x n)) + (\ p. x (Suc (n + p))) t ) \ (- (x n)) + l\ using tendsto_add_const_iff by metis hence f1: \(\ p. x (Suc (n + p)) - x n)\ l - x n\ by simp have \sum (\ k. x (Suc k) - x k) {n..n+p} = x (Suc (n+p)) - x n\ for p by (simp add: sum_Suc_diff) moreover have \(\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + t) = (\ p. sum (\ k. x (Suc k) - x k) {n..n+p}) t\ for t by blast ultimately have \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (n + p)) \ l - x n\ using f1 by simp hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) (p + n)) \ l - x n\ by (simp add: add.commute) hence \(\ p. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) p) \ l - x n\ using Topological_Spaces.LIMSEQ_offset[where f = "(\ N. sum (\ k. x (Suc k) - x k) {n..N})" and a = "l - x n" and k = n] by blast hence \(\ M. (\ N. sum (\ k. x (Suc k) - x k) {n..N}) M) \ l - x n\ by simp thus ?thesis by blast qed lemma bound_Cauchy_to_lim: assumes \y \ x\ and \\n. \y (Suc n) - y n\ \ c^n\ and \y 0 = 0\ and \c < 1\ shows \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ text\ Inequality about a sequence of approximations assuming that the sequence of differences is bounded by a geometric progression. Explanation: Let \<^term>\y\ be a sequence converging to \<^term>\x\. If \<^term>\y\ satisfies the inequality \\y (Suc n) - y n\ \ c ^ n\ for some \<^term>\c < 1\ and assuming \<^term>\y 0 = 0\ then the inequality \\x - y (Suc n)\ \ (c / (1 - c)) * c ^ n\ holds. \ proof- have \c \ 0\ using \\ n. \y (Suc n) - y n\ \ c^n\ by (metis dual_order.trans norm_ge_zero power_one_right) have norm_1: \norm (\k = Suc n..N. y (Suc k) - y k) \ (c ^ Suc n)/(1 - c)\ for N proof(cases \N < Suc n\) case True hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ = 0\ by auto thus ?thesis using \c \ 0\ \c < 1\ by auto next case False hence \N \ Suc n\ by auto have \c^(Suc N) \ 0\ using \c \ 0\ by auto have \1 - c > 0\ by (simp add: \c < 1\) hence \(1 - c)/(1 - c) = 1\ by auto have \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (\k. \y (Suc k) - y k\) {Suc n .. N})\ by (simp add: sum_norm_le) hence \\sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (sum (power c) {Suc n .. N})\ by (simp add: assms(2) sum_norm_le) hence \(1 - c) * \sum (\k. y (Suc k) - y k) {Suc n .. N}\ \ (1 - c) * (sum (power c) {Suc n .. N})\ using \0 < 1 - c\ mult_le_cancel_left_pos by blast also have \\ = c^(Suc n) - c^(Suc N)\ using Set_Interval.sum_gp_multiplied \Suc n \ N\ by blast also have \\ \ c^(Suc n)\ using \c^(Suc N) \ 0\ by auto finally have \(1 - c) * \\k = Suc n..N. y (Suc k) - y k\ \ c ^ Suc n\ by blast hence \((1 - c) * \\k = Suc n..N. y (Suc k) - y k\)/(1 - c) \ (c ^ Suc n)/(1 - c)\ using \0 < 1 - c\ divide_le_cancel by fastforce thus \\\k = Suc n..N. y (Suc k) - y k\ \ (c ^ Suc n)/(1 - c)\ using \0 < 1 - c\ by auto qed have \(\ N. (sum (\k. y (Suc k) - y k) {Suc n .. N})) \ x - y (Suc n)\ by (metis (no_types) \y \ x\ identity_telescopic) hence \(\ N. \sum (\k. y (Suc k) - y k) {Suc n .. N}\) \ \x - y (Suc n)\\ using tendsto_norm by blast hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ using norm_1 Lim_bounded by blast hence \\x - y (Suc n)\ \ (c ^ Suc n)/(1 - c)\ by auto moreover have \(c ^ Suc n)/(1 - c) = (c / (1 - c)) * (c ^ n)\ by (simp add: divide_inverse_commute) ultimately show \\x - y (Suc n)\ \ (c / (1 - c)) * (c ^ n)\ by linarith qed lemma onorm_open_ball: includes notation_norm shows \\f\ = Sup { \f *\<^sub>v x\ | x. \x\ < 1 }\ text \ Explanation: Let \<^term>\f\ be a bounded linear operator. The operator norm of \<^term>\f\ is the supremum of \<^term>\norm (f x)\ for \<^term>\x\ such that \<^term>\norm x < 1\. \ proof(cases \(UNIV::'a set) = 0\) case True hence \x = 0\ for x::'a by auto hence \f *\<^sub>v x = 0\ for x by (metis (full_types) blinfun.zero_right) hence \\f\ = 0\ by (simp add: blinfun_eqI zero_blinfun.rep_eq) have \{ \f *\<^sub>v x\ | x. \x\ < 1} = {0}\ by (smt (verit, ccfv_SIG) Collect_cong \\x. f *\<^sub>v x = 0\ norm_zero singleton_conv) hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} = 0\ by simp thus ?thesis using \\f\ = 0\ by auto next case False hence \(UNIV::'a set) \ 0\ by simp have nonnegative: \\f *\<^sub>v x\ \ 0\ for x by simp have \\ x::'a. x \ 0\ using \UNIV \ 0\ by auto then obtain x::'a where \x \ 0\ by blast hence \\x\ \ 0\ by auto define y where \y = x /\<^sub>R \x\\ have \norm y = \ x /\<^sub>R \x\ \\ unfolding y_def by auto also have \\ = \x\ /\<^sub>R \x\\ by auto also have \\ = 1\ using \\x\ \ 0\ by auto finally have \\y\ = 1\ by blast hence norm_1_non_empty: \{ \f *\<^sub>v x\ | x. \x\ = 1} \ {}\ by blast have norm_1_bounded: \bdd_above { \f *\<^sub>v x\ | x. \x\ = 1}\ unfolding bdd_above_def apply auto by (metis norm_blinfun) have norm_less_1_non_empty: \{\f *\<^sub>v x\ | x. \x\ < 1} \ {}\ by (metis (mono_tags, lifting) Collect_empty_eq_bot bot_empty_eq empty_iff norm_zero zero_less_one) have norm_less_1_bounded: \bdd_above {\f *\<^sub>v x\ | x. \x\ < 1}\ proof- have \\r. \a r\ < 1 \ \f *\<^sub>v (a r)\ \ r\ for a :: "real \ 'a" proof- obtain r :: "('a \\<^sub>L 'b) \ real" where "\f x. 0 \ r f \ (bounded_linear f \ \f *\<^sub>v x\ \ \x\ * r f)" by (metis mult.commute norm_blinfun norm_ge_zero) have \\ \f\ < 0\ by simp hence "(\r. \f\ * \a r\ \ r) \ (\r. \a r\ < 1 \ \f *\<^sub>v a r\ \ r)" by (meson less_eq_real_def mult_le_cancel_left2) thus ?thesis using dual_order.trans norm_blinfun by blast qed hence \\ M. \ x. \x\ < 1 \ \f *\<^sub>v x\ \ M\ by metis thus ?thesis by auto qed have Sup_non_neg: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} \ 0\ by (metis (mono_tags, lifting) \\y\ = 1\ cSup_upper2 mem_Collect_eq norm_1_bounded norm_ge_zero) have \{0::real} \ {}\ by simp have \bdd_above {0::real}\ by simp show \\f\ = Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ proof(cases \\x. f *\<^sub>v x = 0\) case True have \\f *\<^sub>v x\ = 0\ for x by (simp add: True) hence \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ by blast moreover have \{\f *\<^sub>v x\ | x. \x\ < 1 } \ {0}\ using calculation norm_less_1_non_empty by fastforce ultimately have \{\f *\<^sub>v x\ | x. \x\ < 1 } = {0}\ by blast hence Sup1: \Sup {\f *\<^sub>v x\ | x. \x\ < 1 } = 0\ by simp have \\f\ = 0\ by (simp add: True blinfun_eqI) moreover have \Sup {\f *\<^sub>v x\ | x. \x\ < 1} = 0\ using Sup1 by blast ultimately show ?thesis by simp next case False have norm_f_eq_leq: \y \ {\f *\<^sub>v x\ | x. \x\ = 1} \ y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for y proof- assume \y \ {\f *\<^sub>v x\ | x. \x\ = 1}\ hence \\ x. y = \f *\<^sub>v x\ \ \x\ = 1\ by blast then obtain x where \y = \f *\<^sub>v x\\ and \\x\ = 1\ by auto define y' where \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R y\ for n have \y' n \ {\f *\<^sub>v x\ | x. \x\ < 1}\ for n proof- have \y' n = (1 - (inverse (real (Suc n)))) *\<^sub>R \f *\<^sub>v x\\ using y'_def \y = \f *\<^sub>v x\\ by blast also have \... = \(1 - (inverse (real (Suc n))))\ *\<^sub>R \f *\<^sub>v x\\ by (metis (mono_tags, opaque_lifting) \y = \f *\<^sub>v x\\ abs_1 abs_le_self_iff abs_of_nat abs_of_nonneg add_diff_cancel_left' add_eq_if cancel_comm_monoid_add_class.diff_cancel diff_ge_0_iff_ge eq_iff_diff_eq_0 inverse_1 inverse_le_iff_le nat.distinct(1) of_nat_0 of_nat_Suc of_nat_le_0_iff zero_less_abs_iff zero_neq_one) also have \... = \f *\<^sub>v ((1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ by (simp add: blinfun.scaleR_right) finally have y'_1: \y' n = \f *\<^sub>v ( (1 - (inverse (real (Suc n)))) *\<^sub>R x)\\ by blast have \\(1 - (inverse (Suc n))) *\<^sub>R x\ = (1 - (inverse (real (Suc n)))) * \x\\ by (simp add: linordered_field_class.inverse_le_1_iff) hence \\(1 - (inverse (Suc n))) *\<^sub>R x\ < 1\ by (simp add: \\x\ = 1\) thus ?thesis using y'_1 by blast qed have \(\n. (1 - (inverse (real (Suc n)))) ) \ 1\ using Limits.LIMSEQ_inverse_real_of_nat_add_minus by simp hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ 1 *\<^sub>R y\ using Limits.tendsto_scaleR by blast hence \(\n. (1 - (inverse (real (Suc n)))) *\<^sub>R y) \ y\ by simp hence \(\n. y' n) \ y\ using y'_def by simp hence \y' \ y\ by simp have \y' n \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ for n using cSup_upper \\n. y' n \ {\f *\<^sub>v x\ |x. \x\ < 1}\ norm_less_1_bounded by blast hence \y \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ using \y' \ y\ Topological_Spaces.Sup_lim by (meson LIMSEQ_le_const2) thus ?thesis by blast qed hence \Sup {\f *\<^sub>v x\ | x. \x\ = 1} \ Sup {\f *\<^sub>v x\ | x. \x\ < 1}\ by (metis (lifting) cSup_least norm_1_non_empty) have \y \ {\f *\<^sub>v x\ | x. \x\ < 1} \ y \ Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ for y proof(cases \y = 0\) case True thus ?thesis by (simp add: Sup_non_neg) next case False hence \y \ 0\ by blast assume \y \ {\f *\<^sub>v x\ | x. \x\ < 1}\ hence \\ x. y = \f *\<^sub>v x\ \ \x\ < 1\ by blast then obtain x where \y = \f *\<^sub>v x\\ and \\x\ < 1\ by blast have \(1/\x\) * y = (1/\x\) * \f x\\ by (simp add: \y = \f *\<^sub>v x\\) also have \... = \1/\x\\ * \f *\<^sub>v x\\ by simp also have \... = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ by simp also have \... = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ by (simp add: blinfun.scaleR_right) finally have \(1/\x\) * y = \f *\<^sub>v ((1/\x\) *\<^sub>R x)\\ by blast have \x \ 0\ using \y \ 0\ \y = \f *\<^sub>v x\\ blinfun.zero_right by auto have \\ (1/\x\) *\<^sub>R x \ = \ (1/\x\) \ * \x\\ by simp also have \... = (1/\x\) * \x\\ by simp finally have \\(1/\x\) *\<^sub>R x\ = 1\ using \x \ 0\ by simp hence \(1/\x\) * y \ { \f *\<^sub>v x\ | x. \x\ = 1}\ using \1 / \x\ * y = \f *\<^sub>v (1 / \x\) *\<^sub>R x\\ by blast hence \(1/\x\) * y \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ by (simp add: cSup_upper norm_1_bounded) moreover have \y \ (1/\x\) * y\ by (metis \\x\ < 1\ \y = \f *\<^sub>v x\\ mult_le_cancel_right1 norm_not_less_zero order.strict_implies_order \x \ 0\ less_divide_eq_1_pos zero_less_norm_iff) ultimately show ?thesis by linarith qed hence \Sup { \f *\<^sub>v x\ | x. \x\ < 1} \ Sup { \f *\<^sub>v x\ | x. \x\ = 1}\ by (smt (verit, del_insts) less_cSupD norm_less_1_non_empty) hence \Sup { \f *\<^sub>v x\ | x. \x\ = 1} = Sup { \f *\<^sub>v x\ | x. \x\ < 1}\ using \Sup {\f *\<^sub>v x\ |x. norm x = 1} \ Sup { \f *\<^sub>v x\ |x. \x\ < 1}\ by linarith have f1: \(SUP x. \f *\<^sub>v x\ / \x\) = Sup { \f *\<^sub>v x\ / \x\ | x. True}\ by (simp add: full_SetCompr_eq) have \y \ { \f *\<^sub>v x\ / \x\ |x. True} \ y \ { \f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ for y proof- assume \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ show ?thesis proof(cases \y = 0\) case True thus ?thesis by simp next case False have \\ x. y = \f *\<^sub>v x\ / \x\\ using \y \ { \f *\<^sub>v x\ / \x\ |x. True}\ by auto then obtain x where \y = \f *\<^sub>v x\ / \x\\ by blast hence \y = \(1/\x\)\ * \ f *\<^sub>v x \\ by simp hence \y = \(1/\x\) *\<^sub>R (f *\<^sub>v x)\\ by simp hence \y = \f ((1/\x\) *\<^sub>R x)\\ by (simp add: blinfun.scaleR_right) moreover have \\ (1/\x\) *\<^sub>R x \ = 1\ using False \y = \f *\<^sub>v x\ / \x\\ by auto ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ by blast thus ?thesis by blast qed qed moreover have \y \ {\f x\ |x. \x\ = 1} \ {0} \ y \ {\f *\<^sub>v x\ / \x\ |x. True}\ for y proof(cases \y = 0\) case True thus ?thesis by auto next case False hence \y \ {0}\ by simp moreover assume \y \ {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ ultimately have \y \ {\f *\<^sub>v x\ |x. \x\ = 1}\ by simp then obtain x where \\x\ = 1\ and \y = \f *\<^sub>v x\\ by auto have \y = \f *\<^sub>v x\ / \x\\ using \\x\ = 1\ \y = \f *\<^sub>v x\\ by simp thus ?thesis by auto qed ultimately have \{\f *\<^sub>v x\ / \x\ |x. True} = {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}\ by blast hence \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ by simp have "\r s. \ (r::real) \ s \ sup r s = s" by (metis (lifting) sup.absorb_iff1 sup_commute) hence \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {(0::real)}) = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0::real})\ using \0 \ Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ \bdd_above {0}\ \{0} \ {}\ cSup_singleton cSup_union_distrib max.absorb_iff1 sup_commute norm_1_bounded norm_1_non_empty by (metis (no_types, lifting) ) moreover have \Sup {(0::real)} = (0::real)\ by simp ultimately have \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = Sup {\f *\<^sub>v x\ |x. \x\ = 1}\ using Sup_non_neg by linarith moreover have \Sup ( {\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0}) \ using Sup_non_neg \Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0}) = max (Sup {\f *\<^sub>v x\ |x. \x\ = 1}) (Sup {0})\ by auto ultimately have f2: \Sup {\f *\<^sub>v x\ / \x\ | x. True} = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ using \Sup {\f *\<^sub>v x\ / \x\ |x. True} = Sup ({\f *\<^sub>v x\ |x. \x\ = 1} \ {0})\ by linarith have \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ = 1}\ using f1 f2 by linarith hence \(SUP x. \f *\<^sub>v x\ / \x\) = Sup {\f *\<^sub>v x\ | x. \x\ < 1 }\ by (simp add: \Sup {\f *\<^sub>v x\ |x. \x\ = 1} = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\) thus ?thesis apply transfer by (simp add: onorm_def) qed qed lemma onorm_r: includes notation_norm assumes \r > 0\ shows \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ text \ Explanation: The norm of \<^term>\f\ is \<^term>\1/r\ of the supremum of the norm of \<^term>\f *\<^sub>v x\ for \<^term>\x\ in the ball of radius \<^term>\r\ centered at the origin. \ proof- have \\f\ = Sup {\f *\<^sub>v x\ |x. \x\ < 1}\ using onorm_open_ball by blast moreover have *: \{\f *\<^sub>v x\ |x. \x\ < 1} = (\x. \f *\<^sub>v x\) ` (ball 0 1)\ unfolding ball_def by auto ultimately have onorm_f: \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 1))\ by simp have s2: \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1 \ x \ r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1)\ for x proof- assume \x \ (\t. r *\<^sub>R \f *\<^sub>v t\) ` ball 0 1\ hence \\ t. x = r *\<^sub>R \f *\<^sub>v t\ \ \t\ < 1\ by auto then obtain t where t: \x = r *\<^sub>R \f *\<^sub>v t\\ \\t\ < 1\ by blast define y where \y = x /\<^sub>R r\ have \x = r * (inverse r * x)\ using \x = r *\<^sub>R norm (f t)\ by auto hence \x - (r * (inverse r * x)) \ 0\ by linarith hence \x \ r * (x /\<^sub>R r)\ by auto have \y \ (\k. \f *\<^sub>v k\) ` ball 0 1\ unfolding y_def using assms t * by fastforce moreover have \x \ r * y\ using \x \ r * (x /\<^sub>R r)\ y_def by blast ultimately have y_norm_f: \y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ by blast have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ by simp moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) moreover have \\ y. y \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ r * y\ using y_norm_f by blast ultimately show ?thesis by (meson assms cSup_upper dual_order.trans mult_le_cancel_left_pos) qed have s3: \(\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y) \ r * Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y\ for y proof- assume \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ have x_leq: \x \ (\t. \f *\<^sub>v t\) ` ball 0 1 \ x \ y / r\ for x proof- assume \x \ (\t. \f *\<^sub>v t\) ` ball 0 1\ then obtain t where \t \ ball (0::'a) 1\ and \x = \f *\<^sub>v t\\ by auto define x' where \x' = r *\<^sub>R x\ have \x' = r * \f *\<^sub>v t\\ by (simp add: \x = \f *\<^sub>v t\\ x'_def) hence \x' \ (\t. r * \f *\<^sub>v t\) ` ball 0 1\ using \t \ ball (0::'a) 1\ by auto hence \x' \ y\ using \\x. x \ (\t. r * \f *\<^sub>v t\) ` ball 0 1 \ x \ y\ by blast thus \x \ y / r\ unfolding x'_def using \r > 0\ by (simp add: mult.commute pos_le_divide_eq) qed have \(\t. \f *\<^sub>v t\) ` ball 0 1 \ {}\ by simp moreover have \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ by (simp add: bounded_linear_image blinfun.bounded_linear_right bounded_imp_bdd_above bounded_norm_comp) ultimately have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 1) \ y/r\ using x_leq by (simp add: \bdd_above ((\t. \f *\<^sub>v t\) ` ball 0 1)\ cSup_least) thus ?thesis using \r > 0\ by (simp add: mult.commute pos_le_divide_eq) qed have norm_scaleR: \norm \ ((*\<^sub>R) r) = ((*\<^sub>R) \r\) \ (norm::'a \ real)\ by auto have f_x1: \f (r *\<^sub>R x) = r *\<^sub>R f x\ for x by (simp add: blinfun.scaleR_right) have \ball (0::'a) r = ((*\<^sub>R) r) ` (ball 0 1)\ - by (smt assms ball_scale nonzero_mult_div_cancel_left right_inverse_eq scale_zero_right) + by (smt (verit) assms ball_scale nonzero_mult_div_cancel_left right_inverse_eq scale_zero_right) hence \Sup ((\t. \f *\<^sub>v t\) ` (ball 0 r)) = Sup ((\t. \f *\<^sub>v t\) ` (((*\<^sub>R) r) ` (ball 0 1)))\ by simp also have \\ = Sup (((\t. \f *\<^sub>v t\) \ ((*\<^sub>R) r)) ` (ball 0 1))\ using Sup.SUP_image by auto also have \\ = Sup ((\t. \f *\<^sub>v (r *\<^sub>R t)\) ` (ball 0 1))\ using f_x1 by (simp add: comp_assoc) also have \\ = Sup ((\t. \r\ *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ using norm_scaleR f_x1 by auto also have \\ = Sup ((\t. r *\<^sub>R \f *\<^sub>v t\) ` (ball 0 1))\ using \r > 0\ by auto also have \\ = r * Sup ((\t. \f *\<^sub>v t\) ` (ball 0 1))\ apply (rule cSup_eq_non_empty) apply simp using s2 apply auto using s3 by auto also have \\ = r * \f\\ using onorm_f by auto finally have \Sup ((\t. \f *\<^sub>v t\) ` ball 0 r) = r * \f\\ by blast thus \\f\ = Sup ((\x. \f *\<^sub>v x\) ` (ball 0 r)) / r\ using \r > 0\ by simp qed text\Pointwise convergence\ definition pointwise_convergent_to :: \( nat \ ('a \ 'b::topological_space) ) \ ('a \ 'b) \ bool\ (\((_)/ \pointwise\ (_))\ [60, 60] 60) where \pointwise_convergent_to x l = (\ t::'a. (\ n. (x n) t) \ l t)\ lemma linear_limit_linear: fixes f :: \_ \ ('a::real_vector \ 'b::real_normed_vector)\ assumes \\n. linear (f n)\ and \f \pointwise\ F\ shows \linear F\ text\ Explanation: If a family of linear operators converges pointwise, then the limit is also a linear operator. \ proof show "F (x + y) = F x + F y" for x y proof- have "\a. F a = lim (\n. f n a)" using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by (metis (full_types) limI) moreover have "\f b c g. (lim (\n. g n + f n) = (b::'b) + c \ \ f \ c) \ \ g \ b" by (metis (no_types) limI tendsto_add) moreover have "\a. (\n. f n a) \ F a" using assms(2) pointwise_convergent_to_def by force ultimately have lim_sum: \lim (\ n. (f n) x + (f n) y) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ by metis have \(f n) (x + y) = (f n) x + (f n) y\ for n using \\ n. linear (f n)\ unfolding linear_def using Real_Vector_Spaces.linear_iff assms(1) by auto hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x + (f n) y)\ by simp hence \lim (\ n. (f n) (x + y)) = lim (\ n. (f n) x) + lim (\ n. (f n) y)\ using lim_sum by simp moreover have \(\ n. (f n) (x + y)) \ F (x + y)\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast moreover have \(\ n. (f n) x) \ F x\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast moreover have \(\ n. (f n) y) \ F y\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast ultimately show ?thesis by (metis limI) qed show "F (r *\<^sub>R x) = r *\<^sub>R F x" for r and x proof- have \(f n) (r *\<^sub>R x) = r *\<^sub>R (f n) x\ for n using \\ n. linear (f n)\ by (simp add: Real_Vector_Spaces.linear_def real_vector.linear_scale) hence \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ by simp have \convergent (\ n. (f n) x)\ by (metis assms(2) convergentI pointwise_convergent_to_def) moreover have \isCont (\ t::'b. r *\<^sub>R t) tt\ for tt by (simp add: bounded_linear_scaleR_right) ultimately have \lim (\ n. r *\<^sub>R ((f n) x)) = r *\<^sub>R lim (\ n. (f n) x)\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by (metis (mono_tags) isCont_tendsto_compose limI) hence \lim (\ n. (f n) (r *\<^sub>R x)) = r *\<^sub>R lim (\ n. (f n) x)\ using \lim (\ n. (f n) (r *\<^sub>R x)) = lim (\ n. r *\<^sub>R (f n) x)\ by simp moreover have \(\ n. (f n) x) \ F x\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast moreover have \(\ n. (f n) (r *\<^sub>R x)) \ F (r *\<^sub>R x)\ using \f \pointwise\ F\ unfolding pointwise_convergent_to_def by blast ultimately show ?thesis by (metis limI) qed qed lemma non_Cauchy_unbounded: fixes a ::\_ \ real\ assumes \\n. a n \ 0\ and \e > 0\ and \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ shows \(\n. (sum a {0..n})) \ \\ text\ Explanation: If the sequence of partial sums of nonnegative terms is not Cauchy, then it converges to infinite. \ proof- define S::"ereal set" where \S = range (\n. sum a {0..n})\ have \\s\S. k*e \ s\ for k::nat proof(induction k) case 0 from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ obtain m n where \m \ 0\ and \n \ 0\ and \m > n\ and \sum a {Suc n..m} \ e\ by blast have \n < Suc n\ by simp hence \{0..n} \ {Suc n..m} = {0..m}\ using Set_Interval.ivl_disj_un(7) \n < m\ by auto moreover have \finite {0..n}\ by simp moreover have \finite {Suc n..m}\ by simp moreover have \{0..n} \ {Suc n..m} = {}\ by simp ultimately have \sum a {0..n} + sum a {Suc n..m} = sum a {0..m}\ by (metis sum.union_disjoint) moreover have \sum a {Suc n..m} > 0\ using \e > 0\ \sum a {Suc n..m} \ e\ by linarith moreover have \sum a {0..n} \ 0\ by (simp add: assms(1) sum_nonneg) ultimately have \sum a {0..m} > 0\ by linarith moreover have \sum a {0..m} \ S\ unfolding S_def by blast ultimately have \\s\S. 0 \ s\ using ereal_less_eq(5) by fastforce thus ?case by (simp add: zero_ereal_def) next case (Suc k) assume \\s\S. k*e \ s\ then obtain s where \s\S\ and \ereal (k * e) \ s\ by blast have \\N. s = sum a {0..N}\ using \s\S\ unfolding S_def by blast then obtain N where \s = sum a {0..N}\ by blast from \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ obtain m n where \m \ Suc N\ and \n \ Suc N\ and \m > n\ and \sum a {Suc n..m} \ e\ by blast have \finite {Suc N..n}\ by simp moreover have \finite {Suc n..m}\ by simp moreover have \{Suc N..n} \ {Suc n..m} = {Suc N..m}\ using Set_Interval.ivl_disj_un by (metis \Suc N \ n\ \n < m\ atLeastSucAtMost_greaterThanAtMost order_less_imp_le) moreover have \{} = {Suc N..n} \ {Suc n..m}\ by simp ultimately have \sum a {Suc N..m} = sum a {Suc N..n} + sum a {Suc n..m}\ by (metis sum.union_disjoint) moreover have \sum a {Suc N..n} \ 0\ using \\n. a n \ 0\ by (simp add: sum_nonneg) ultimately have \sum a {Suc N..m} \ e\ using \e \ sum a {Suc n..m}\ by linarith have \finite {0..N}\ by simp have \finite {Suc N..m}\ by simp moreover have \{0..N} \ {Suc N..m} = {0..m}\ using Set_Interval.ivl_disj_un(7) \Suc N \ m\ by auto moreover have \{0..N} \ {Suc N..m} = {}\ by simp ultimately have \sum a {0..N} + sum a {Suc N..m} = sum a {0..m}\ by (metis \finite {0..N}\ sum.union_disjoint) hence \e + k * e \ sum a {0..m}\ using \ereal (real k * e) \ s\ \s = ereal (sum a {0..N})\ \e \ sum a {Suc N..m}\ by auto moreover have \e + k * e = (Suc k) * e\ by (simp add: semiring_normalization_rules(3)) ultimately have \(Suc k) * e \ sum a {0..m}\ by linarith hence \ereal ((Suc k) * e) \ sum a {0..m}\ by auto moreover have \sum a {0..m}\S\ unfolding S_def by blast ultimately show ?case by blast qed hence \\s\S. (real n) \ s\ for n by (meson assms(2) ereal_le_le ex_less_of_nat_mult less_le_not_le) hence \Sup S = \\ using Sup_le_iff Sup_subset_mono dual_order.strict_trans1 leD less_PInf_Ex_of_nat subsetI by metis hence Sup: \Sup ((range (\ n. (sum a {0..n})))::ereal set) = \\ using S_def by blast have \incseq (\n. (sum a {.. using \\n. a n \ 0\ using Extended_Real.incseq_sumI by auto hence \incseq (\n. (sum a {..< Suc n}))\ by (meson incseq_Suc_iff) hence \incseq (\n. (sum a {0..n})::ereal)\ using incseq_ereal by (simp add: atLeast0AtMost lessThan_Suc_atMost) hence \(\n. sum a {0..n}) \ Sup (range (\n. (sum a {0..n})::ereal))\ using LIMSEQ_SUP by auto thus ?thesis using Sup PInfty_neq_ereal by auto qed lemma sum_Cauchy_positive: fixes a ::\_ \ real\ assumes \\n. a n \ 0\ and \\K. \n. (sum a {0..n}) \ K\ shows \Cauchy (\n. sum a {0..n})\ text\ Explanation: If a series of nonnegative reals is bounded, then the series is Cauchy. \ proof (unfold Cauchy_altdef2, rule, rule) fix e::real assume \e>0\ have \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ proof(rule classical) assume \\(\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e)\ hence \\M. \m. \n. m \ M \ n \ M \ m > n \ \(sum a {Suc n..m} < e)\ by blast hence \\M. \m. \n. m \ M \ n \ M \ m > n \ sum a {Suc n..m} \ e\ by fastforce hence \(\n. (sum a {0..n}) ) \ \\ using non_Cauchy_unbounded \0 < e\ assms(1) by blast from \\K. \n. sum a {0..n} \ K\ obtain K where \\n. sum a {0..n} \ K\ by blast from \(\n. sum a {0..n}) \ \\ have \\B. \N. \n\N. (\ n. (sum a {0..n}) ) n \ B\ using Lim_PInfty by simp hence \\n. (sum a {0..n}) \ K+1\ using ereal_less_eq(3) by blast thus ?thesis using \\n. (sum a {0..n}) \ K\ by (smt (verit, best)) qed have \sum a {Suc n..m} = sum a {0..m} - sum a {0..n}\ if "m > n" for m n by (metis add_diff_cancel_left' atLeast0AtMost less_imp_add_positive sum_up_index_split that) hence \\M. \m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ using \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ by presburger then obtain M where \\m\M. \n\M. m > n \ sum a {0..m} - sum a {0..n} < e\ by blast moreover have \m > n \ sum a {0..m} \ sum a {0..n}\ for m n using \\ n. a n \ 0\ by (simp add: sum_mono2) ultimately have \\M. \m\M. \n\M. m > n \ \sum a {0..m} - sum a {0..n}\ < e\ by auto hence \\M. \m\M. \n\M. m \ n \ \sum a {0..m} - sum a {0..n}\ < e\ by (metis \0 < e\ abs_zero cancel_comm_monoid_add_class.diff_cancel diff_is_0_eq' less_irrefl_nat linorder_neqE_nat zero_less_diff) hence \\M. \m\M. \n\M. \sum a {0..m} - sum a {0..n}\ < e\ by (metis abs_minus_commute nat_le_linear) hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ by (simp add: dist_real_def) hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ by blast thus \\N. \n\N. dist (sum a {0..n}) (sum a {0..N}) < e\ by auto qed lemma convergent_series_Cauchy: fixes a::\nat \ real\ and \::\nat \ 'a::metric_space\ assumes \\M. \n. sum a {0..n} \ M\ and \\n. dist (\ (Suc n)) (\ n) \ a n\ shows \Cauchy \\ text\ Explanation: Let \<^term>\a\ be a real-valued sequence and let \<^term>\\\ be sequence in a metric space. If the partial sums of \<^term>\a\ are uniformly bounded and the distance between consecutive terms of \<^term>\\\ are bounded by the sequence \<^term>\a\, then \<^term>\\\ is Cauchy.\ proof (unfold Cauchy_altdef2, rule, rule) fix e::real assume \e > 0\ have \\k. a k \ 0\ using \\n. dist (\ (Suc n)) (\ n) \ a n\ dual_order.trans zero_le_dist by blast hence \Cauchy (\k. sum a {0..k})\ using \\M. \n. sum a {0..n} \ M\ sum_Cauchy_positive by blast hence \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\ unfolding Cauchy_def using \e > 0\ by blast hence \\M. \m\M. \n\M. m > n \ dist (sum a {0..m}) (sum a {0..n}) < e\ by blast have \dist (sum a {0..m}) (sum a {0..n}) = sum a {Suc n..m}\ if \n for m n proof - have \n < Suc n\ by simp have \finite {0..n}\ by simp moreover have \finite {Suc n..m}\ by simp moreover have \{0..n} \ {Suc n..m} = {0..m}\ using \n < Suc n\ \n < m\ by auto moreover have \{0..n} \ {Suc n..m} = {}\ by simp ultimately have sum_plus: \(sum a {0..n}) + sum a {Suc n..m} = (sum a {0..m})\ by (metis sum.union_disjoint) have \dist (sum a {0..m}) (sum a {0..n}) = \(sum a {0..m}) - (sum a {0..n})\\ using dist_real_def by blast moreover have \(sum a {0..m}) - (sum a {0..n}) = sum a {Suc n..m}\ using sum_plus by linarith ultimately show ?thesis by (simp add: \\k. 0 \ a k\ sum_nonneg) qed hence sum_a: \\M. \m\M. \n\M. m > n \ sum a {Suc n..m} < e\ by (metis \\M. \m\M. \n\M. dist (sum a {0..m}) (sum a {0..n}) < e\) obtain M where \\m\M. \n\M. m > n \ sum a {Suc n..m} < e\ using sum_a \e > 0\ by blast hence \\m. \n. Suc m \ Suc M \ Suc n \ Suc M \ Suc m > Suc n \ sum a {Suc n..Suc m - 1} < e\ by simp hence \\m\1. \n\1. m \ Suc M \ n \ Suc M \ m > n \ sum a {n..m - 1} < e\ by (metis Suc_le_D) hence sum_a2: \\M. \m\M. \n\M. m > n \ sum a {n..m-1} < e\ by (meson add_leE) have \dist (\ (n+p+1)) (\ n) \ sum a {n..n+p}\ for p n :: nat proof(induction p) case 0 thus ?case by (simp add: assms(2)) next case (Suc p) thus ?case by (smt(verit, ccfv_SIG) Suc_eq_plus1 add_Suc_right add_less_same_cancel1 assms(2) dist_self dist_triangle2 gr_implies_not0 sum.cl_ivl_Suc) qed hence \m > n \ dist (\ m) (\ n) \ sum a {n..m-1}\ for m n :: nat by (metis Suc_eq_plus1 Suc_le_D diff_Suc_1 gr0_implies_Suc less_eq_Suc_le less_imp_Suc_add zero_less_Suc) hence \\M. \m\M. \n\M. m > n \ dist (\ m) (\ n) < e\ - using sum_a2 \e > 0\ by smt + using sum_a2 \e > 0\ by (smt (verit)) thus "\N. \n\N. dist (\ n) (\ N) < e" using \0 < e\ by fastforce qed unbundle notation_blinfun_apply unbundle no_notation_norm end diff --git a/thys/Boolos_Curious_Inference_Automated/ROOT b/thys/Boolos_Curious_Inference_Automated/ROOT --- a/thys/Boolos_Curious_Inference_Automated/ROOT +++ b/thys/Boolos_Curious_Inference_Automated/ROOT @@ -1,12 +1,9 @@ chapter AFP session "Boolos_Curious_Inference_Automated" = HOL + - options [timeout = 600] - theories - Boolos_Curious_Inference_Automated - + Boolos_Curious_Inference_Automated document_files "root.bib" "root.tex" diff --git a/thys/DigitsInBase/DigitsInBase.thy b/thys/DigitsInBase/DigitsInBase.thy --- a/thys/DigitsInBase/DigitsInBase.thy +++ b/thys/DigitsInBase/DigitsInBase.thy @@ -1,920 +1,918 @@ theory DigitsInBase imports "HOL-Computational_Algebra.Computational_Algebra" "HOL-Number_Theory.Number_Theory" begin section\Infinite sums\ text\In this section, it is shown that an infinite series \emph{of natural numbers} converges if and only if its terms are eventually zero. Additionally, the notion of a summation starting from an index other than zero is defined. A few obvious lemmas about these notions are established.\ definition eventually_zero :: "(nat \ _) \ bool" where "eventually_zero (D :: nat \ _) \ (\\<^sub>\ n. D n = 0)" lemma eventually_zero_char: shows "eventually_zero D \ (\s. \i\s. D i = 0)" unfolding eventually_zero_def using MOST_nat_le . text\There's a lot of commonality between this setup and univariate polynomials, but drawing out the similarities and proving them is beyond the scope of the current version of this theory except for the following lemma.\ lemma eventually_zero_poly: shows "eventually_zero D \ D = poly.coeff (Abs_poly D)" by (metis Abs_poly_inverse MOST_coeff_eq_0 eventually_zero_def mem_Collect_eq) lemma eventually_zero_imp_summable [simp]: assumes "eventually_zero D" shows "summable D" using summable_finite assms eventually_zero_char by (metis (mono_tags) atMost_iff finite_atMost nat_le_linear) lemma summable_bounded: fixes my_seq :: "nat \ nat" and n :: nat assumes "\ i . i \ n \ my_seq i = 0" shows "summable my_seq" using assms eventually_zero_char eventually_zero_imp_summable by blast lemma sum_bounded: fixes my_seq :: "nat \ nat" and n :: nat assumes "\ i . i \ n \ my_seq i = 0" shows "(\i. my_seq i) = (\i nat" assumes "eventually_zero seq1" shows "eventually_zero (\ i. seq1 i * seq2 i)" using mult_0 eventually_zero_char by (metis (no_types, lifting) assms) abbreviation upper_sum where "upper_sum seq n \ \i. seq (i + n)" syntax "_from_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\i\n. t" == "CONST upper_sum (\i. t) n" text \The following two statements are proved as a sanity check. They are not intended to be used anywhere.\ corollary fixes seq :: "nat \ nat" and a :: nat assumes seq_def: "\ i. seq i = (if i = 0 then a else 0)" shows "(\i\0. seq i) = upper_sum (\ i. seq i) 0" by simp corollary fixes seq :: "nat \ nat" and a :: nat assumes seq_def: "\ i. seq i = (if i = 0 then a else 0)" shows "(\i\0. seq i) = a" by (smt (verit) group_cancel.rule0 lessI lessThan_0 linorder_not_less seq_def sum.empty sum.lessThan_Suc_shift sum_bounded) lemma bounded_sum_from: fixes seq :: "nat \ nat" and n s :: nat assumes "\i>s. seq i = 0" and "n \ s" shows "(\i\n. seq i) = (\i=n..s. seq i)" proof - have "\i. i > (s - n) \ seq (i + n) = 0" using assms by (meson less_diff_conv2) then have "(\i\n. seq i) = (\i\s-n. seq (i + n))" by (meson atMost_iff finite_atMost leI suminf_finite) also have "\ = (\i=n..s. seq i)" proof - have "\na. (\na\na. seq (na + n)) = sum seq {0 + n..na + n}" by (metis (no_types) atLeast0AtMost sum.shift_bounds_cl_nat_ivl) then show ?thesis by (simp add: assms(2)) qed finally show ?thesis . qed lemma split_suminf: fixes seq :: "nat \ nat" and n :: nat assumes "eventually_zero seq" shows "(\i. seq i) = (\ii\n. seq i)" proof - obtain s where s: "\i. i\s \ seq i = 0" using assms unfolding eventually_zero_char by presburger then have sum_s: "(\i. seq i) = (\ii. seq i) = (\ii\(n). seq i)" proof (cases "n \ s") case True then have "(\i\(n). seq i) = 0" using s by force moreover have "(\iii=n..s. seq i) = (\i\n. seq i)" by (metis False bounded_sum_from le_eq_less_or_eq nle_le s) from False have "n \ s" by simp then have "(\iii=n..s. seq i)" by (metis add_cancel_left_right nat_le_iff_add s sum.atLeastLessThan_concat add_0 lessThan_atLeast0 sum.last_plus) then show ?thesis using 0 sum_s by presburger qed qed lemma dvd_suminf: fixes seq :: "nat \ nat" and b :: nat assumes "eventually_zero seq" and "\i. b dvd seq i" shows "b dvd (\i. seq i)" proof - obtain s::nat where s: "i \ s \ seq i = 0" for i using assms(1) eventually_zero_char by blast then have "(\i. seq i) = (\ii i. seq (i + n))" by (meson assms eventually_zero_char trans_le_add1) section\Modular arithmetic\ text\This section establishes a number of lemmas about modular arithmetic, including that modular division distributes over an ``infinite'' sum whose terms are eventually zero.\ lemma pmod_int_char: fixes a b d :: int shows "[a = b] (mod d) \ (\(n::int). a = b + n*d)" by (metis cong_iff_lin cong_sym mult.commute) lemma equiv_conj_if: assumes "P \ Q" and "P \ R" and "Q \ R \ P" shows "P \ Q \ R" using assms by auto lemma mod_successor_char: fixes k k' i b :: nat assumes "(b::nat) \ 2" shows "[k = k'] (mod b^(Suc i)) \ [k div b^i = k' div b^i] (mod b) \ [k = k'] (mod b^i)" proof (rule equiv_conj_if) assume kk'_cong: "[k = k'] (mod b ^ Suc i)" then show "[k div b^i = k' div b^i] (mod b)" by (smt (verit, ccfv_SIG) Groups.mult_ac(2) add_diff_cancel_right' cong_def div_mult_mod_eq mod_mult2_eq mod_mult_self4 mult_cancel1 power_Suc2) from kk'_cong show "[k = k'] (mod b ^ i)" using Cong.cong_dvd_modulus_nat by (meson Suc_n_not_le_n le_imp_power_dvd nat_le_linear) next assume "[k div b^i = k' div b^i] (mod b)" moreover assume "[k = k'] (mod b^i)" ultimately show "[k = k'] (mod b ^ Suc i)" by (metis (mono_tags, lifting) cong_def mod_mult2_eq power_Suc2) qed lemma mod_1: fixes k k' b :: nat shows "[k = k'] (mod b^0)" by simp lemma mod_distributes: fixes seq :: "nat \ nat" and b :: nat assumes "\n. \i\n. seq i = 0" shows "[(\i. seq i) = (\i. seq i mod b)] (mod b)" proof - obtain n where n: "\i. i\n \ seq i = 0" using assms by presburger from n have "(\i. seq i) = (\ii. seq i mod b) = (\i 0" and "[m = a + b] (mod c * d)" and "a div d = 0" and "d dvd b" shows "[m div d = b div d] (mod c)" proof (subst pmod_int_char) obtain k::int where k: "m = a + b + k*c*d" using pmod_int_char assms(2) by (metis mult.assoc) have "d dvd (b + k*c*d)" using \d dvd b\ by simp from k have "m div d = (a + b + k*c*d) div d" by presburger also have "\ = (b + k*c*d) div d" using \a div d = 0\ \d dvd (b + k*c*d)\ by fastforce also have "\ = (b div d) + k*c" using \d dvd b\ \d > 0\ by auto finally show "\n. m div d = b div d + n * c" by blast qed lemma another_mod_cancellation: fixes a b c d m :: nat assumes "d > 0" and "[m = a + b] (mod c * d)" and "a div d = 0" and "d dvd b" shows "[m div d = b div d] (mod c)" by (smt (verit) another_mod_cancellation_int assms cong_int_iff of_nat_0 of_nat_0_less_iff of_nat_add of_nat_dvd_iff of_nat_mult zdiv_int) section\Digits as sequence\ text\Rules are introduced for computing the $i^{\text{th}}$ digit of a base-$b$ representation and the number of digits required to represent a given number. (The latter is essentially an integer version of the base-$b$ logarithm.) It is shown that the sum of the terms $d_i b^i$ converges to $m$ if $d_i$ is the $i^{\text{th}}$ digit $m$. It is shown that the sequence of digits defined is the unique sequence of digits less than $b$ with this property Additionally, the \texttt{digits\_in\_base} locale is introduced, which specifies a single symbol @{term b} referring to a natural number greater than one (the base of the digits). Consequently this symbol is omitted from many of the following lemmas and definitions. \ locale digits_in_base = fixes b :: nat assumes b_gte_2: "b \ 2" begin lemma b_facts [simp]: shows "b > 1" and "b > 0" and "b \ 1" and "b \ 0" and "1 mod b = 1" and "1 div b = 0" using b_gte_2 by force+ text\Definition based on @{cite ThreeDivides}.\ abbreviation ith_digit :: "nat \ nat \ nat" where "ith_digit m i \ (m div b^i) mod b" lemma ith_digit_lt_base: fixes m i :: nat shows "0 \ ith_digit m i" and "ith_digit m i < b" apply (rule Nat.le0) using b_facts(2) mod_less_divisor by presburger definition num_digits :: "nat \ nat" where "num_digits m = (LEAST i. m < b^i)" lemma num_digits_works: shows "m < b ^ (num_digits m)" by (metis LeastI One_nat_def b_facts(1) num_digits_def power_gt_expt) lemma num_digits_le: assumes "m < b^i" shows "num_digits m \ i" using assms num_digits_works[of m] Least_le num_digits_def by metis lemma num_digits_zero: fixes m :: nat assumes "num_digits m = 0" shows "m = 0" using num_digits_works[of m] unfolding assms by simp lemma num_digits_gt: assumes "m \ b^i" shows "num_digits m > i" by (meson assms b_facts(2) dual_order.strict_trans2 nat_power_less_imp_less num_digits_works) lemma num_digits_eqI [intro]: assumes "m \ b^i" and "m < b^(i+1)" shows "num_digits m = i + 1" proof - { fix j :: nat assume "j < i + 1" then have "m \ b^j" by (metis Suc_eq_plus1 assms(1) b_facts(1) less_Suc_eq_le order_trans power_increasing_iff) } then show ?thesis using num_digits_works unfolding num_digits_def by (meson assms(2) leD linorder_neqE_nat not_less_Least) qed lemma num_digits_char: assumes "m \ 1" shows "num_digits m = i + 1 \ m \ b^i \ m < b^(i+1)" by (metis add_diff_cancel_right' assms b_gte_2 ex_power_ivl1 num_digits_eqI) text\Statement based on @{cite ThreeDivides}.\ lemma num_digits_recurrence: fixes m :: nat assumes "m \ 1" shows "num_digits m = num_digits (m div b) + 1" proof - define nd where "nd = num_digits m" then have lb: "m \ b^(nd-1)" and ub: "m < b^nd" using num_digits_char[OF assms] apply (metis assms diff_is_0_eq le_add_diff_inverse2 nat_le_linear power_0) using nd_def num_digits_works by presburger from ub have ub2: "m div b < b^(nd-1)" by (metis Suc_eq_plus1 add.commute add_diff_inverse_nat assms less_mult_imp_div_less less_one linorder_not_less mult.commute power.simps(2) power_0) from lb have lb2: "m div b \ b^(nd - 1) div b" using div_le_mono by presburger show ?thesis proof (cases "m \ b") assume "m \ b" then have "nd \ 2" unfolding nd_def by (metis One_nat_def assms less_2_cases_iff linorder_not_le nd_def power_0 power_one_right ub) then have "m div b \ b^(nd-2)" using lb2 by (smt (verit) One_nat_def add_le_imp_le_diff b_facts(4) diff_diff_left le_add_diff_inverse2 nonzero_mult_div_cancel_left numeral_2_eq_2 plus_1_eq_Suc power_add power_commutes power_one_right) then show ?thesis using ub2 num_digits_char assms nd_def by (smt (verit) \2 \ nd\ add_diff_cancel_right' add_leD2 add_le_imp_le_diff diff_diff_left eq_diff_iff le_add2 nat_1_add_1 num_digits_eqI) next assume "\ b \ m" then have "m < b" by simp then have "num_digits m = 1" using assms by (metis One_nat_def Suc_eq_plus1 num_digits_char power_0 power_one_right) from \m < b\ have "m div b = 0" using div_less by presburger then have "num_digits (m div b) = 0" using Least_eq_0 num_digits_def by presburger show ?thesis using \num_digits (m div b) = 0\ \num_digits m = 1\ by presburger qed qed lemma num_digits_zero_2 [simp]: "num_digits 0 = 0" by (simp add: num_digits_def) end (* digits_in_base *) locale base_10 begin text\As a sanity check, the number of digits in base ten is computed for several natural numbers.\ sublocale digits_in_base 10 by (unfold_locales, simp) corollary shows "num_digits 0 = 0" and "num_digits 1 = 1" and "num_digits 9 = 1" and "num_digits 10 = 2" and "num_digits 99 = 2" and "num_digits 100 = 3" by (simp_all add: num_digits_recurrence) end (* base_10 *) context digits_in_base begin lemma high_digits_zero_helper: fixes m i :: nat shows "i < num_digits m \ ith_digit m i = 0" proof (cases "i < num_digits m") case True then show ?thesis by meson next case False then have "i \ num_digits m" by force then have "m < b^i" by (meson b_facts(1) num_digits_works order_less_le_trans power_increasing_iff) then show ?thesis by simp qed lemma high_digits_zero: fixes m i :: nat assumes "i \ num_digits m" shows "ith_digit m i = 0" using high_digits_zero_helper assms leD by blast lemma digit_expansion_bound: fixes i :: nat and A :: "nat \ nat" assumes "\j. A j < b" shows "(\j (b-1) * b^i" using assms by (metis One_nat_def Suc_pred b_facts(2) le_simps(2) mult_le_mono1) then have "(\j \ b ^ Suc i" using assms(1) mult_eq_if by auto finally show "(\jStatement and proof based on @{cite ThreeDivides}.\ lemma num_digits_suc: fixes n m :: nat assumes "Suc n = num_digits m" shows "n = num_digits (m div b)" using num_digits_recurrence assms by (metis One_nat_def Suc_eq_plus1 Suc_le_lessD le_add2 linorder_not_less num_digits_le old.nat.inject power_0) text\Proof (and to some extent statement) based on @{cite ThreeDivides}.\ lemma digit_expansion_bounded_seq: fixes m :: nat shows "m = (\jjjjj = (\j = (\j = (\j=Suc 0.. = (\jjA natural number can be obtained from knowing all its base-$b$ digits by the formula $\sum_j d_j b^j$.\ theorem digit_expansion_seq: fixes m :: nat shows "m = (\j. ith_digit m j * b^j)" using digit_expansion_bounded_seq[of m] high_digits_zero[of m] sum_bounded mult_0 by (metis (no_types, lifting)) lemma lower_terms: fixes a c i :: nat assumes "c < b^i" and "a < b" shows "ith_digit (a * b^i + c) i = a" using assms by force lemma upper_terms: fixes A a i :: nat assumes "b*b^i dvd A" and "a < b" shows "ith_digit (A + a * b^i) i = a" using assms by force lemma current_term: fixes A a c i :: nat assumes "b*b^i dvd A" and "c < b^i" and "a < b" shows "ith_digit (A + a*b^i + c) i = a" proof - have "(A + a*b^i + c) div b^i mod b = (a*b^i + c) div b^i mod b" using assms(1) by (metis (no_types, lifting) div_eq_0_iff add_cancel_right_right assms(2) assms(3) div_plus_div_distrib_dvd_left dvd_add_times_triv_right_iff dvd_mult_right lower_terms upper_terms) also have "\ = a" using assms by force finally show "(A + a*b^i + c) div b^i mod b = a" . qed text\Given that \begin{equation*} m = \sum_i d_i b^i \end{equation*} where the $d_i$ are all natural numbers less than $b$, it follows that $d_j$ is the $j^{\text{th}}$ digit of $m$.\ theorem seq_uniqueness: fixes m j :: nat and D :: "nat \ nat" assumes "eventually_zero D" and "m = (\i. D i * b^i)" and "\i. D i < b" shows "D j = ith_digit m j" proof - have "eventually_zero (ith_digit m)" using high_digits_zero by (meson eventually_zero_char) then have term_eventually_zero: "eventually_zero (\ i. D i * b^i)" using product_seq_eventually_zero assms(1) by auto then have shifted_term_eventually_zero: "eventually_zero (\ i. D (i + n) * b ^ (i + n))" for n using eventually_zero_shifted by blast note \m = (\i. D i * b^i)\ then have two_sums: "m = (\ii\Suc j. D i * b^i)" using split_suminf[OF term_eventually_zero] by presburger have "i\Suc j \ b*b^j dvd (D i * b^i)" for i by (metis dvd_mult2 le_imp_power_dvd mult.commute power_Suc) then have "b*b^j dvd (\i\Suc j. D i * b^i)" using dvd_suminf shifted_term_eventually_zero le_add2 by presburger with two_sums have "[m = (\iiiLittle Endian notation\ text\In this section we begin to define finite digit expansions. Ultimately we want to write digit expansions in ``big endian'' notation, by which we mean with the highest place digit on the left and the ones digit on the write, since this ordering is standard in informal mathematics. However, it is easier to first define ``little endian'' expansions with the ones digit on the left since that way the list indexing agrees with the sequence indexing used in the previous section (whenever both are defined). Notation, definitions, and lemmas in this section typically start with the prefix \texttt{LE} (for ``little endian'') to distinguish them from the big endian versions in the next section. \ fun LEeval_as_base ("_\<^bsub>LEbase _\<^esub>" [65, 65] 70) where "[] \<^bsub>LEbase b\<^esub> = 0" | "(d # d_list) \<^bsub>LEbase b\<^esub> = d + b * (d_list\<^bsub>LEbase b\<^esub>)" corollary shows "[2, 4] \<^bsub>LEbase 5\<^esub> = (22::nat)" by simp lemma LEbase_one_digit [simp]: shows "[a::nat] \<^bsub>LEbase b\<^esub> = a" by simp lemma LEbase_two_digits [simp]: shows "[a\<^sub>0::nat, a\<^sub>1] \<^bsub>LEbase b\<^esub> = a\<^sub>0 + a\<^sub>1 * b" by simp lemma LEbase_three_digits [simp]: shows "[a\<^sub>0::nat, a\<^sub>1, a\<^sub>2] \<^bsub>LEbase b\<^esub> = a\<^sub>0 + a\<^sub>1*b + a\<^sub>2*b^2" proof - have "[a\<^sub>0::nat, a\<^sub>1, a\<^sub>2] \<^bsub>LEbase b\<^esub> = a\<^sub>0 + ([a\<^sub>1, a\<^sub>2] \<^bsub>LEbase b\<^esub>) * b" by simp also have "\ = a\<^sub>0 + (a\<^sub>1 + a\<^sub>2*b) * b" by simp also have "\ = a\<^sub>0 + a\<^sub>1*b + a\<^sub>2*b^2" by (simp add: add_mult_distrib power2_eq_square) finally show ?thesis . qed lemma LEbase_closed_form: shows "(A :: nat list) \<^bsub>LEbase b\<^esub> = (\i < length A . A!i * b^i)" proof (induct A) case Nil show ?case by simp next case (Cons a A) show ?case proof - have "(a # A)\<^bsub>LEbase b\<^esub> = a + b * (A\<^bsub>LEbase b\<^esub>)" by simp also have "\ = a + b * (\i = a + (\i = a + (\i = a + (\i = (a#A)!0 * b^0 + (\i = (\iLEbase b\<^esub> = (A\<^bsub>LEbase b\<^esub>) + b^(length A) * (D\<^bsub>LEbase b\<^esub>)" proof (induct A) case Nil show ?case by simp next case (Cons a A) show ?case proof - have "((a # A) @ D)\<^bsub>LEbase b\<^esub> = ((a # (A @ D)) \<^bsub>LEbase b\<^esub>)" by simp also have "\ = a + b * ((A @ D) \<^bsub>LEbase b\<^esub>)" by simp also have "\ = a + b * (A\<^bsub>LEbase b\<^esub> + b ^ length A * (D\<^bsub>LEbase b\<^esub>))" unfolding Cons by rule also have "\ = (a + b * (A\<^bsub>LEbase b\<^esub>)) + b ^ (length (a#A)) * (D\<^bsub>LEbase b\<^esub>)" by (simp add: distrib_left) also have "\ = ((a # A)\<^bsub>LEbase b\<^esub>) + b ^ length (a # A) * (D\<^bsub>LEbase b\<^esub>)" by simp finally show ?thesis . qed qed context digits_in_base begin definition LEdigits :: "nat \ nat list" where "LEdigits m = [ith_digit m i. i \ [0..<(num_digits m)]]" lemma length_is_num_digits: fixes m :: nat shows "length (LEdigits m) = num_digits m" unfolding LEdigits_def by simp lemma ith_list_element [simp]: assumes "(i::nat) < length (LEdigits m)" shows "(LEdigits m) ! i = ith_digit m i" using assms by (simp add: length_is_num_digits LEdigits_def) lemma LEbase_infinite_sum: fixes m :: nat shows "(\i. ith_digit m i * b^i) = (LEdigits m)\<^bsub>LEbase b\<^esub>" proof (unfold LEdigits_def LEbase_closed_form) have "(\ii= (\i=(\i. ith_digit m i * b ^ i)" using sum_bounded high_digits_zero mult_0 by (metis (no_types, lifting)) finally show "(\i. ith_digit m i * b ^ i) = (\iLEbase b\<^esub> = m" using digit_expansion_seq LEbase_infinite_sum by presburger lemma LElist_uniqueness: fixes D :: "nat list" assumes "\ i < length D. D!i < b" and "D = [] \ last D \ 0" shows "LEdigits (D\<^bsub>LEbase b\<^esub>) = D" proof - define seq where "seq i = (if i < length D then D!i else 0)" for i then have seq_bound: "i \ length D \ seq i = 0" for i by simp then have seq_eventually_zero: "eventually_zero seq" using eventually_zero_char by blast have ith_digit_connection: "i < num_digits m \ (LEdigits m)!i = ith_digit m i" for m i unfolding LEdigits_def by simp let ?m = "D\<^bsub>LEbase b\<^esub>" have length_bounded_sum: "D\<^bsub>LEbase b\<^esub> = (\i = (\i. seq i * b^i)" using seq_bound sum_bounded by fastforce finally have seq_is_digits: "seq j = ith_digit ?m j" for j using seq_uniqueness[OF seq_eventually_zero] assms(1) by (metis b_facts(2) seq_def) then have "i < length D \ ith_digit ?m i = D!i" for i using seq_def by presburger then have "i < length D \ i < num_digits ?m \ (LEdigits ?m)!i = D!i" for i using ith_digit_connection[of i ?m] by presburger moreover have "length D = num_digits ?m" proof (rule le_antisym) show "length D \ num_digits ?m" proof (cases "D = []") assume "D \ []" then have "last D \ 0" using assms(2) by auto then have "last D \ 1" by simp have "?m \ seq (length D - 1) * b^(length D - 1)" using length_bounded_sum - by (smt (z3) One_nat_def Suc_pred diff_is_0_eq diff_le_self dual_order.strict_trans1 le_add2 - linorder_le_less_linear mult.right_neutral mult_cancel1 seq_def sum.lessThan_Suc - zero_less_diff) + by (metis b_facts(2) less_eq_div_iff_mult_less_eq mod_less_eq_dividend seq_is_digits zero_less_power) then have "?m \ (last D) * b^(length D - 1)" by (simp add: \D \ []\ last_conv_nth seq_def) with \last D \ 1\ have "?m \ b^(length D - 1)" by (metis le_trans mult_1 mult_le_mono1) then show "num_digits ?m \ length D" using num_digits_gt not_less_eq by (metis One_nat_def Suc_pred \D \ []\ bot_nat_0.extremum_uniqueI leI length_0_conv) qed simp show "num_digits ?m \ length D" by (metis length_bounded_sum seq_is_digits digit_expansion_bound ith_digit_lt_base(2) num_digits_le) qed ultimately show ?thesis by (simp add: length_is_num_digits list_eq_iff_nth_eq) qed lemma LE_digits_zero [simp]: "LEdigits 0 = []" using LEdigits_def by auto lemma LE_units_digit [simp]: assumes "(m::nat) \ {1..Big Endian notation\ text\In this section the desired representation of natural numbers, as finite lists of digits with the highest place on the left, is finally realized.\ definition BEeval_as_base ("_\<^bsub>base _\<^esub>" [65, 65] 70) where [simp]: "D\<^bsub>base b\<^esub> = (rev D)\<^bsub>LEbase b\<^esub>" corollary shows "[4, 2]\<^bsub>base 5\<^esub> = (22::nat)" by simp lemma BEbase_one_digit [simp]: shows "[a::nat] \<^bsub>base b\<^esub> = a" by simp lemma BEbase_two_digits [simp]: shows "[a\<^sub>1::nat, a\<^sub>0] \<^bsub>base b\<^esub> = a\<^sub>1*b + a\<^sub>0" by simp lemma BEbase_three_digits [simp]: shows "[a\<^sub>2::nat, a\<^sub>1, a\<^sub>0] \<^bsub>base b\<^esub> = a\<^sub>2*b^2 + a\<^sub>1*b + a\<^sub>0" proof - have "b * (a\<^sub>1 + b * a\<^sub>2) = a\<^sub>2 * b\<^sup>2 + a\<^sub>1 * b" apply (subst mult.commute) unfolding add_mult_distrib power2_eq_square by simp then show ?thesis by simp qed lemma BEbase_closed_form: fixes A :: "nat list" and b :: nat shows "A\<^bsub>base b\<^esub> = (\ibase b\<^esub> = (A\<^bsub>base b\<^esub>)*b^(length D) + (D\<^bsub>base b\<^esub>)" using LEbase_concatenate by simp context digits_in_base begin definition digits :: "nat \ nat list" where "digits m = rev (LEdigits m)" lemma length_is_num_digits_2: fixes m :: nat shows "length (digits m) = num_digits m" using length_is_num_digits digits_def by simp lemma LE_BE_equivalence: fixes m :: nat shows "(digits m) \<^bsub>base b\<^esub> = (LEdigits m) \<^bsub>LEbase b\<^esub>" by (simp add: digits_def) lemma BEbase_infinite_sum: fixes m :: nat shows "(\i. ith_digit m i * b^i) = (digits m)\<^bsub>base b\<^esub>" using LE_BE_equivalence LEbase_infinite_sum by presburger text\Every natural number can be represented in base $b$, specifically by the digits sequence defined earlier.\ theorem digit_expansion_list: fixes m :: nat shows "(digits m)\<^bsub>base b\<^esub> = m" using LE_BE_equivalence digit_expansion_LElist by auto text\If two natural numbers have the same base-$b$ representation, then they are equal.\ lemma digits_cancellation: fixes k m :: nat assumes "digits k = digits m" shows "k = m" by (metis assms digit_expansion_list) text\Suppose we have a finite (possibly empty) sequence $D_1, \dotsc, D_n$ of natural numbers such that $0 \le D_i < b$ for all $i$ and such that $D_1$, if it exists, is nonzero. Then this sequence is the base-$b$ representation for $\sum_i D_i b^{n-i}$.\ theorem list_uniqueness: fixes D :: "nat list" assumes "\ d \ set D. d < b" and "D = [] \ D!0 \ 0" shows "digits (D\<^bsub>base b\<^esub>) = D" unfolding digits_def BEeval_as_base_def using LElist_uniqueness by (metis Nil_is_rev_conv One_nat_def assms last_conv_nth length_greater_0_conv nth_mem rev_nth rev_swap set_rev) text\We now prove some simplification rules (including a reccurrence relation) to make it easier for Isabelle/HOL to compute the base-$b$ representation of a natural number.\ text\The base-b representation of $0$ is empty, at least following the conventions of this theory file.\ lemma digits_zero [simp]: shows "digits 0 = []" by (simp add: digits_def) text\If $0 < m < b$, then the base-$b$ representation of $m$ consists of a single digit, namely $m$ itself.\ lemma single_digit_number [simp]: assumes "m \ {0<..For all $m \ge b$, the base-$b$ representation of $m$ consists of the base-$b$ representation of $\lfloor m / b \rfloor$ followed by (as the last digit) the remainder of $m$ when divided by $b$.\ lemma digits_recurrence [simp]: assumes "m \ b" shows "digits m = (digits (m div b)) @ [m mod b]" proof - have "num_digits m > 1" using assms by (simp add: num_digits_gt) then have "num_digits m > 0" by simp then have "num_digits (m div b) = num_digits m - 1" by (metis Suc_diff_1 num_digits_suc) have "k > 0 \ last (rev [0.. [0.. rev [0.. rev [Suc 0.. rev [0.. 0 \ [f i. i \ rev [1.. rev [0..<(k-1)]]" for f and k::nat by (metis One_nat_def Suc_diff_1) have digit_down: "ith_digit m (Suc i) = ith_digit (m div b) i" for i::nat by (simp add: div_mult2_eq) have "digits m = rev [ith_digit m i. i \ [0.. = [ith_digit m i. i \ rev [0.. = [ith_digit m i. i \ butlast (rev [0..1 < num_digits m\ bot_nat_0.extremum_strict dual_order.strict_trans1 last_map map_butlast snoc_eq_iff_butlast upt_eq_Nil_conv) also have "\ = [ith_digit m i. i \ rev [1..1 < num_digits m\ \\k. 0 < k \ last (rev [0.. by fastforce also have "\ = [ith_digit m (Suc i). i \ rev [0..<(num_digits m - 1)]] @ [ith_digit m 0]" using map_shift[OF \num_digits m > 0\] by blast also have "\ = [ith_digit (m div b) i. i \ rev [0..<(num_digits m - 1)]] @ [ith_digit m 0]" using digit_down by presburger also have "\ = (digits (m div b)) @ [ith_digit m 0]" by (simp add: LEdigits_def \num_digits (m div b) = num_digits m - 1\ digits_def rev_map) also have "\ = (digits (m div b)) @ [m mod b]" by simp finally show ?thesis . qed end (* digits_in_base *) section\Exercises\ text\This section contains demonstrations of how to denote certain facts with the notation of the previous sections, and how to quickly prove those facts using the lemmas and theorems above. \ text\The base-5 representation of $22$ is $42_5$.\ corollary "digits_in_base.digits 5 22 = [4, 2]" proof - interpret digits_in_base 5 by (simp add: digits_in_base.intro) show "digits 22 = [4, 2]" by simp qed text\A different proof of the same statement.\ corollary "digits_in_base.digits 5 22 = [4, 2]" proof - interpret digits_in_base 5 by (simp add: digits_in_base.intro) have "[4, 2]\<^bsub>base 5\<^esub> = (22::nat)" by simp have "d \ set [4, 2] \ d < 5" for d::nat by fastforce then show ?thesis using list_uniqueness by (metis \[4, 2]\<^bsub>base 5\<^esub> = 22\ nth_Cons_0 numeral_2_eq_2 zero_neq_numeral) qed end diff --git a/thys/FOL_Seq_Calc1/Sequent2.thy b/thys/FOL_Seq_Calc1/Sequent2.thy --- a/thys/FOL_Seq_Calc1/Sequent2.thy +++ b/thys/FOL_Seq_Calc1/Sequent2.thy @@ -1,44 +1,44 @@ theory Sequent2 imports Sequent begin section \Completeness Revisited\ lemma \\p. q = compl p\ by (metis compl.simps(1)) definition compl' where \compl' = (\q. (SOME p. q = compl p))\ lemma comp'_sem: \eval e f g (compl' p) \ \ eval e f g p\ by (smt compl'_def compl.simps(1) compl eval.simps(7) someI_ex) lemma comp'_sem_list: \list_ex (\p. \ eval e f g p) (map compl' ps) \ list_ex (eval e f g) ps\ by (induct ps) (use comp'_sem in auto) theorem SC_completeness': fixes ps :: \(nat, nat) form list\ assumes \\(e :: nat \ nat hterm) f g. list_ex (eval e f g) (p # ps)\ shows \\ p # ps\ proof - define ps' where \ps' = map compl' ps\ then have \ps = map compl ps'\ - by (induct ps arbitrary: ps') (simp, smt compl'_def compl.simps(1) list.simps(9) someI_ex) + by (induct ps arbitrary: ps') (simp, smt (verit) compl'_def compl.simps(1) list.simps(9) someI_ex) from assms have \\(e :: nat \ nat hterm) f g. (list_ex (eval e f g) ps) \ eval e f g p\ by auto then have \\(e :: nat \ nat hterm) f g. (list_ex (\p. \ eval e f g p) ps') \ eval e f g p\ unfolding ps'_def using comp'_sem_list by blast then have \\(e :: nat \ nat hterm) f g. list_all (eval e f g) ps' \ eval e f g p\ by (metis Ball_set Bex_set) then have \\ p # map compl ps'\ using SC_completeness by blast then show ?thesis using \ps = map compl ps'\ by auto qed corollary fixes ps :: \(nat, nat) form list\ assumes \\(e :: nat \ nat hterm) f g. list_ex (eval e f g) ps\ shows \\ ps\ using assms SC_completeness' by (cases ps) auto end diff --git a/thys/GaleStewart_Games/AlternatingLists.thy b/thys/GaleStewart_Games/AlternatingLists.thy --- a/thys/GaleStewart_Games/AlternatingLists.thy +++ b/thys/GaleStewart_Games/AlternatingLists.thy @@ -1,144 +1,144 @@ section \ Alternating lists \ text \In lists where even and odd elements play different roles, it helps to define functions to take out the even elements. We defined the function (l)alternate on (coinductive) lists to do exactly this, and define certain properties.\ theory AlternatingLists imports MoreCoinductiveList2 (* for notation and lemmas like infinite_small_llength *) begin text \The functions ``alternate" and ``lalternate" are our main workhorses: they take every other item, so every item at even indices.\ fun alternate where "alternate Nil = Nil" | "alternate (Cons x xs) = Cons x (alternate (tl xs))" text \``lalternate" takes every other item from a co-inductive list.\ primcorec lalternate :: "'a llist \ 'a llist" where "lalternate xs = (case xs of LNil \ LNil | (LCons x xs) \ LCons x (lalternate (ltl xs)))" lemma lalternate_ltake: "ltake (enat n) (lalternate xs) = lalternate (ltake (2*n) xs)" proof(induct n arbitrary:xs) case 0 then show ?case by (metis LNil_eq_ltake_iff enat_defs(1) lalternate.ctr(1) lnull_def mult_zero_right) next case (Suc n) hence lt:"ltake (enat n) (lalternate (ltl (ltl xs))) = lalternate (ltake (enat (2 * n)) (ltl (ltl xs)))". show ?case proof(cases "lalternate xs") case LNil then show ?thesis by(metis lalternate.disc(2) lnull_def ltake_LNil) next case (LCons x21 x22) thus ?thesis unfolding ltake_ltl mult_Suc_right add_2_eq_Suc using eSuc_enat lalternate.code lalternate.ctr(1) lhd_LCons_ltl llist.sel(1) - by (smt (z3) lt ltake_ltl llist.simps(3) llist.simps(5) ltake_eSuc_LCons) + by (smt (verit) lt ltake_ltl llist.simps(3) llist.simps(5) ltake_eSuc_LCons) qed qed lemma lalternate_llist_of[simp]: "lalternate (llist_of xs) = llist_of (alternate xs)" proof(induct "alternate xs" arbitrary:xs) case Nil then show ?case by (metis alternate.elims lalternate.ctr(1) list.simps(3) llist_of.simps(1) lnull_llist_of) next case (Cons a xs) then show ?case by(cases xs,auto simp: lalternate.ctr) qed lemma lalternate_finite_helper: (* The other direction is proved later, added as SIMP rule *) assumes "lfinite (lalternate xs)" shows "lfinite xs" using assms proof(induct "lalternate xs" arbitrary:xs rule:lfinite_induct) case LNil then show ?case unfolding lalternate.code[of xs] by(cases xs;auto) next case (LCons xs) then show ?case unfolding lalternate.code[of xs] by(cases "xs";cases "ltl xs";auto) qed lemma alternate_list_of: (* Note that this only holds for finite lists, as the other direction is left undefined with arguments (not just undefined) *) assumes "lfinite xs" shows "alternate (list_of xs) = list_of (lalternate xs)" using assms by (metis lalternate_llist_of list_of_llist_of llist_of_list_of) lemma alternate_length: "length (alternate xs) = (1+length xs) div 2" by (induct xs rule:induct_list012;simp) lemma lalternate_llength: "llength (lalternate xs) * 2 = (1+llength xs) \ llength (lalternate xs) * 2 = llength xs" proof(cases "lfinite xs") case True let ?xs = "list_of xs" have "length (alternate ?xs) = (1+length ?xs) div 2" using alternate_length by auto hence "length (alternate ?xs) * 2 = (1+length ?xs) \ length (alternate ?xs) * 2 = length ?xs" by auto then show ?thesis using alternate_list_of[OF True] lalternate_llist_of True length_list_of_conv_the_enat[OF True] llist_of_list_of[OF True] by (metis llength_llist_of numeral_One of_nat_eq_enat of_nat_mult of_nat_numeral plus_enat_simps(1)) next case False have "\ lfinite (lalternate xs)" using False lalternate_finite_helper by auto hence l1:"llength (lalternate xs) = \" by(rule not_lfinite_llength) from False have l2:"llength xs = \" using not_lfinite_llength by auto show ?thesis using l1 l2 by (simp add: mult_2_right) qed lemma lalternate_finite[simp]: shows "lfinite (lalternate xs) = lfinite xs" proof(cases "lfinite xs") case True then show ?thesis proof(cases "lfinite (lalternate xs)") case False hence False using not_lfinite_llength[OF False] True[unfolded lfinite_conv_llength_enat] lalternate_llength[of xs] by (auto simp:one_enat_def numeral_eq_enat) thus ?thesis by metis qed auto next case False then show ?thesis using lalternate_finite_helper by blast qed lemma nth_alternate: assumes "2*n < length xs" shows "alternate xs ! n = xs ! (2 * n)" using assms proof (induct xs arbitrary:n rule:induct_list012) case (3 x y zs) then show ?case proof(cases n) case (Suc nat) show ?thesis using "3.hyps"(1) "3.prems" Suc by force qed simp qed auto lemma lnth_lalternate: assumes "2*n < llength xs" shows "lalternate xs $ n = xs $ (2 * n)" proof - let ?xs = "ltake (2*Suc n) xs" have "lalternate ?xs $ n = ?xs $ (2 * n)" using assms alternate_list_of[of "ltake (2*Suc n) xs"] nth_alternate[of n "list_of ?xs"] - by (smt (z3) Suc_1 Suc_mult_less_cancel1 enat_ord_simps(2) infinite_small_llength lalternate_ltake length_list_of lessI llength_eq_enat_lfiniteD llength_ltake' ltake_all not_less nth_list_of numeral_eq_enat the_enat.simps times_enat_simps(1)) + by (smt (verit) Suc_1 Suc_mult_less_cancel1 enat_ord_simps(2) infinite_small_llength lalternate_ltake length_list_of lessI llength_eq_enat_lfiniteD llength_ltake' ltake_all not_less nth_list_of numeral_eq_enat the_enat.simps times_enat_simps(1)) thus ?thesis by (metis Suc_1 Suc_mult_less_cancel1 enat_ord_simps(2) lalternate_ltake lessI lnth_ltake) qed lemma lnth_lalternate2[simp]: assumes "n < llength (lalternate xs)" shows "lalternate xs $ n = xs $ (2 * n)" proof - from assms have "2*enat n < llength xs" by (metis enat_numeral lalternate_ltake leI linorder_neq_iff llength_ltake' ltake_all times_enat_simps(1)) from lnth_lalternate[OF this] show ?thesis. qed end \ No newline at end of file diff --git a/thys/GaleStewart_Games/FilteredList.thy b/thys/GaleStewart_Games/FilteredList.thy --- a/thys/GaleStewart_Games/FilteredList.thy +++ b/thys/GaleStewart_Games/FilteredList.thy @@ -1,310 +1,310 @@ theory FilteredList imports MoreCoinductiveList2 begin subsection \More on filtered lists\ text \We will reason about (co-inductive) lists with distinct elements. However, for our setting, this 'distinct' property only holds on the list after filtering. For this reason, we need some additional lemmas.\ text \Taking a sublist preserves distinctness after filtering.\ lemma ldistinct_lfilter_ltake[intro]: assumes "ldistinct (lfilter P xs)" shows "ldistinct (lfilter P (ltake x xs))" using assms by(induct xs,force,force ,(* sledgehammer found this gem to prove the inductive step via lfilter_lappend_lfinite! We will use this strategy ourselves later on *) metis lappend_ltake_ldrop ldistinct_lappend lfilter_lappend_lfinite lfinite_LConsI lfinite_ltake) text \The function lfind is used in multiple proofs, all are introduced to prove ltake_lfilter.\ definition lfind where "lfind P lst = (LEAST i. P (lst $ i))" lemma lfilter_lfind: assumes "lfilter P lst \ LNil" shows "P (lst $ lfind P lst)" (is ?g1) "P (lst $ y) \ lfind P lst \ y" (is "?a \ ?g2") "lfind P lst < llength lst" (is ?g3) proof - let ?I = "{n. enat n < llength lst \ P (lst $ n)}" let ?xs = lst from assms[unfolded lfilter_conv_lnths] lset_LNil have "lset (lnths lst {n. enat n < llength lst \ P (lst $ n)}) \ {}" by auto hence "{?xs $ i |i. enat i < llength ?xs \ i \ ?I} \ {}" using lset_lnths[of ?xs ?I] by metis then obtain i where p:"P (lst $ i)" "i < llength lst" by auto from p show ?g1 using LeastI lfind_def by metis from p show "?a \ ?g2" using Least_le lfind_def by metis from p show ?g3 using Least_le lfind_def by (metis enat_ord_simps(1) le_less_trans) qed lemma ltake_lfind_lset: assumes "x \ lset (ltake (enat (lfind P lst)) lst)" shows "\ P x" proof(cases "lfilter P (ltake (enat (lfind P lst)) lst) = LNil") case True then show ?thesis using assms unfolding lfilter_eq_LNil by auto next case False from assms[unfolded in_lset_conv_lnth] obtain n where n:"enat n < llength (ltake (enat (lfind P lst)) lst)" "ltake (enat (lfind P lst)) lst $ n = x" by blast { assume a:"P x" (* The idea of this {block} is that the element n must come after (lfind P lst) by lfilter_lfind(2) but this contradicts n(1). However, in the last step when writing this proof, sledgehammer found one that didn't use any of my previous steps, so here's a one-liner: *) from n Coinductive_List.lset_ltake False a enat_ord_simps(1) leD lfilter_empty_conv lfilter_lfind(2,3) llength_ltake' lnth_ltake subset_eq have False by metis } then show ?thesis by blast qed lemma ltake_lfind_conv: assumes "lfilter P lst \ LNil" shows "ltake (lfind P lst) lst = ltakeWhile (Not o P) lst" (is "?t1 = ?t2") "ldrop (lfind P lst) lst = ldropWhile (Not o P) lst" (is "?t3 = ?t4") proof - have lfin:"lfinite ?t1" by simp have [simp]:"min (enat (lfind P lst)) (llength lst) = (lfind P lst)" using lfilter_lfind(3)[OF assms] by (metis min.strict_order_iff) have l1:"llength ?t1 = lfind P lst" by simp from ltake_lfind_lset ltakeWhile_all have t:"ltakeWhile (Not o P) ?t1 = ?t1" unfolding o_def by metis have inset:"lset (ltake (enat (lfind P lst)) lst) \ {x. (Not \ P) x}" using ltake_lfind_lset[of _ P lst] by auto (* for use in ltakeWhile_lappend2 *) have isnull:"ltakeWhile (Not \ P) (ldrop (enat (lfind P lst)) lst) = LNil" apply(cases "ldrop (enat (lfind P lst)) lst") using lfilter_lfind(1)[OF assms] lhd_ldrop[OF lfilter_lfind(3)[OF assms]] by auto have "ltakeWhile (Not o P) ?t1 = ltakeWhile (Not o P) (lappend ?t1 ?t3)" unfolding ltakeWhile_lappend2[OF inset] isnull lappend_LNil2 t.. hence leneq:"llength ?t1 = llength ?t2" using t l1 lappend_ltake_ldrop by metis have "lappend ?t1 ?t3 = lappend ?t2 ?t4" unfolding lappend_ltakeWhile_ldropWhile[of "Not \ P" lst] lappend_ltake_ldrop[of "lfind P lst" lst] by simp from this[unfolded lappend_eq_lappend_conv[OF leneq]] lfin show "?t1 = ?t2" "?t3 = ?t4" by auto qed lemma lfilter_hdtl: assumes "lfilter P lst \ LNil" shows "\ n. LCons (lhd (lfilter P lst)) LNil = lfilter P (ltake (enat n) lst) \ ltl (lfilter P lst) = lfilter P (ldrop (enat n) lst)" proof(standard,standard) note * = lfilter_lfind[OF assms] let ?n = "Suc (lfind P lst)" let ?ltake = "ltake (enat ?n) lst" have ltake:"lappend (ltakeWhile (Not \ P) ?ltake) (ldropWhile (Not \ P) ?ltake) = ?ltake" (is "lappend ?ltw ?ldw = _") using lappend_ltakeWhile_ldropWhile by blast have "llength ?ldw \ 1" unfolding ldropWhile_lappend ltake_Suc_conv_snoc_lnth[OF *(3)] using ltake_lfind_lset[of _ P lst] by (auto intro:* simp:one_eSuc) hence null:"lnull (ltl (ldropWhile (Not \ P) ?ltake))" unfolding llength_eq_0[symmetric] llength_ltl by (metis dual_order.order_iff_strict enat_ile epred_0 epred_1 iless_Suc_eq le_zero_eq one_eSuc one_enat_def) have e:"enat (lfind P lst) < enat (Suc (lfind P lst))" by auto from * have "P (?ltake $ lfind P lst)" using lnth_ltake[OF e] by metis hence nonnull:"\ lnull (lfilter P ?ltake)" unfolding lnull_lfilter by (metis "*"(3) e in_lset_conv_lnth leI llength_ltake' ltake_all) show a:"LCons (lhd (lfilter P lst)) LNil = lfilter P ?ltake" (is "?lhs = ?rhs") proof - have "lhd (lfilter P ?ltake) = lhd (lfilter P lst)" by(rule lprefix_lhdD[OF lprefix_lfilterI[OF ltake_is_lprefix] nonnull]) hence h:"lhd ?lhs = lhd ?rhs" by simp have "ltl ?rhs = LNil" unfolding ltl_lfilter using null by (metis lfilter_LNil llist.collapse(1)) hence t:"ltl ?lhs = ltl ?rhs" by simp have flt:"?rhs \ LNil" using nonnull by fastforce show ?thesis by(rule llist_eq_lcons[of ?lhs ?rhs,OF _ flt h t],auto) qed from lappend_ltake_ldrop[of ?n lst] lappend_ltakeWhile_ldropWhile[of "Not \ P" lst] have "lappend (ltake ?n lst) (ldrop ?n lst) = lappend (ltakeWhile (Not \ P) lst) (ldropWhile (Not \ P) lst)" by auto from ltake_lfind_conv(2)[OF assms] have "ltl (ldropWhile (Not \ P) lst) = ldrop (enat (Suc (lfind P lst))) lst" unfolding ldrop_eSuc_conv_ltl eSuc_enat[symmetric] by simp thus "ltl (lfilter P lst) = lfilter P (ldrop (enat ?n) lst)" unfolding ltl_lfilter by metis qed lemma ltake_lfilter: shows "\ n. ltake (enat x) (lfilter P lst) = lfilter P (ltake (enat n) lst) \ ldrop (enat x) (lfilter P lst) = lfilter P (ldrop (enat n) lst)" proof(induct x) case 0 then show ?case by (metis LNil_eq_ltake_iff ldrop_enat ldropn_0 lfilter_code(1) zero_enat_def) next let ?fP = "lfilter P" case (Suc x) then obtain n where n:"ltake (enat x) (?fP lst) = ?fP (ltake (enat n) lst)" "ldrop (enat x) (lfilter P lst) = lfilter P (ldrop (enat n) lst)" by blast consider "lfilter P (ldrop (enat n) lst) \ LNil \ x < llength (?fP lst)" | "lfilter P (ldrop (enat n) lst) = LNil" | "x \ llength (?fP lst)" by force then show ?case proof(cases) case 1 hence *:"lfilter P (ldrop (enat n) lst) \ LNil" "enat x < llength (lfilter P lst)" by auto from lappend_ltake_ldrop have "lst = lappend (ltake (enat n) lst) (ldrop (enat n) lst)" by metis from lfilter_hdtl[OF *(1)] obtain delta where delta:"LCons (lhd (?fP (ldrop (enat n) lst))) LNil = ?fP (ltake (enat delta) (ldrop (enat n) lst))" "ltl (lfilter P (ldrop (enat n) lst)) = lfilter P (ldrop (enat delta) (ldrop (enat n) lst))" by blast have "ltake (enat (Suc x)) (?fP lst) = lappend (?fP (ltake (enat n) lst)) (LCons (?fP lst $ x) LNil)" using n ltake_Suc_conv_snoc_lnth * by metis also have "?fP lst $ x = ?fP lst $ (the_enat x)" by auto also have "\ = lhd (ldrop x (?fP lst))" using lhd_ldrop[symmetric] *(2) by metis also have "\ = lhd (?fP (ldrop (enat n) lst))" using n by metis also note delta(1) finally have take_part:"ltake (enat (Suc x)) (?fP lst) = ?fP (ltake (enat (n + delta)) lst)" using ltake_plus_conv_lappend by (metis infinite_small_llength lfilter_lappend_lfinite llength_ltake' ltake_all min.strict_order_iff not_less plus_enat_simps(1)) have "ldrop (enat (Suc x)) (?fP lst) = ltl (ldrop x (?fP lst))" by (simp add: ltl_ldropn ldrop_eSuc_ltl ldrop_enat) also have "ldrop x (?fP lst) = ?fP (ldrop (enat n) lst)" using n by metis also note delta(2) also have "lfilter P (ldrop (enat delta) (ldrop (enat n) lst)) = lfilter P (ldrop (enat delta + enat n) lst)" by simp also have "(enat delta + enat n) = enat (n + delta)" by simp finally have drop_part:"ldrop (enat (Suc x)) (?fP lst) = ?fP (ldrop (enat (n + delta)) lst)". from take_part drop_part show ?thesis by blast next case 2 note * = 2 lappend_ltake_ldrop[of "enat n" lst] Suc_llength infinite_small_llength lappend_LNil2 leI lfilter_lappend_lfinite llength_ltake' min.strict_order_iff n have take_part:"ltake (enat (Suc x)) (?fP lst) = ?fP (ltake (enat n) lst)" - using * by (smt (z3) ltake_all) + using * by (smt (verit) ltake_all) from 2 have drop_part:"ldrop (enat (Suc x)) (?fP lst) = ?fP (ldrop (enat n) lst)" - using * by (smt (z3) ldrop_all) + using * by (smt (verit) ldrop_all) from take_part drop_part show ?thesis by blast next case 3 then show ?thesis using n dual_order.order_iff_strict eSuc_enat iless_Suc_eq le_less_trans ltake_all ldrop_all by metis qed qed lemma filter_obtain_two: assumes "i < j" "j < length (filter P lst)" shows "\ i2 j2. i2 < j2 \ j2 < length lst \ lst ! i2 = filter P lst ! i \ lst ! j2 = filter P lst ! j" using assms proof(induct lst arbitrary: i j) case (Cons a lst) then obtain jprev where jprev:"j = Suc jprev" using lessE by metis show ?case proof(cases "P a") case True hence lnth[simp]:"length (filter P (a # lst)) = Suc (length (filter P lst))" by auto show ?thesis proof(cases i) case 0 from jprev True Cons(3) have "jprev < length (filter P lst) " by auto from nth_mem[OF this] have "filter P lst ! jprev \ set lst" by auto from this[unfolded in_set_conv_nth] obtain j2 where "j2 i j. i < llength lst \ j < llength lst \ lst $ i = lst $ j \ P (lst $ i) \ i = j" shows "ldistinct (lfilter P lst)" proof - { fix i j assume *: "enat i < llength (lfilter P lst)" "enat j < llength (lfilter P lst)" "lfilter P lst $ i = lfilter P lst $ j" "i < j" hence "lfilter P lst $ i \ lset (lfilter P lst)" unfolding in_lset_conv_lnth by auto with lset_lfilter have P:"P (lfilter P lst $ i)" by auto let ?maxij = "Suc (max i j)" from ltake_lfilter obtain maxij where maxij:"ltake ?maxij (lfilter P lst) = lfilter P (ltake (enat maxij) lst)" by blast let ?lst = "ltake (enat maxij) lst" have "lfinite ?lst" by auto define flst where "flst = list_of ?lst" hence flst:"llist_of flst = ?lst" by auto let ?flst = "llist_of flst" from * P have "enat i < llength (lfilter P ?flst)" "enat j < llength (lfilter P ?flst)" "lfilter P ?flst $ i = lfilter P ?flst $ j" and P2:"P (lfilter P ?lst $ i)" unfolding maxij[symmetric] flst by (auto simp add: lnth_ltake) hence "i < length (filter P flst)" "j < length (filter P flst)" and eq_ij: "filter P flst ! i = filter P flst ! j" unfolding llength_llist_of lfilter_llist_of lnth_llist_of by auto with filter_obtain_two[OF *(4) this(2)] obtain i2 j2 where "i2 < length flst" "j2 < length flst" "flst ! i2 = filter P flst ! i" "flst ! j2 = filter P flst ! j" and ineq:"i2 i j. i < llength lst \ j < llength lst \ P (lst $ i) \ lst $ i = lst $ j \ i = j)" proof show "ldistinct (lfilter P lst) \ \i j. enat i < llength lst \ enat j < llength lst \ P (lst $ i) \ lst $ i = lst $ j \ i = j" by (auto simp add: ldistinct_lfilterE) show "\i j. enat i < llength lst \ enat j < llength lst \ P (lst $ i) \ lst $ i = lst $ j \ i = j \ ldistinct (lfilter P lst) " using ldistinct_lfilterI by blast qed end \ No newline at end of file diff --git a/thys/IMO2019/IMO2019_Q1.thy b/thys/IMO2019/IMO2019_Q1.thy --- a/thys/IMO2019/IMO2019_Q1.thy +++ b/thys/IMO2019/IMO2019_Q1.thy @@ -1,66 +1,66 @@ (* File: IMO2019_Q1.thy Author: Manuel Eberl, TU München *) section \Q1\ theory IMO2019_Q1 imports Main begin text \ Consider a function \f : \ \ \\ that fulfils the functional equation \f(2a) + 2f(b) = f(f(a+b))\ for all \a, b \ \\. Then \f\ is either identically 0 or of the form \f(x) = 2x + c\ for some constant \c \ \\. \ context fixes f :: "int \ int" and m :: int assumes f_eq: "f (2 * a) + 2 * f b = f (f (a + b))" defines "m \ (f 0 - f (-2)) div 2" begin text \ We first show that \f\ is affine with slope \(f(0) - f(-2)) / 2\. This follows from plugging in \(0, b)\ and \(-1, b + 1)\ into the functional equation. \ lemma f_eq': "f x = m * x + f 0" proof - have rec: "f (b + 1) = f b + m" for b using f_eq[of 0 b] f_eq[of "-1" "b + 1"] by (simp add: m_def) moreover have "f (b - 1) = f b - m" for b using rec[of "b - 1"] by simp ultimately show ?thesis by (induction x rule: int_induct[of _ 0]) (auto simp: algebra_simps) qed text \ This version is better for the simplifier because it prevents it from looping. \ lemma f_eq'_aux [simp]: "NO_MATCH 0 x \ f x = m * x + f 0" by (rule f_eq') text \ Plugging in \(0, 0)\ and \(0, 1)\. \ lemma f_classification: "(\x. f x = 0) \ (\x. f x = 2 * x + f 0)" using f_eq[of 0 0] f_eq[of 0 1] by auto end text \ It is now easy to derive the full characterisation of the functions we considered: \ theorem fixes f :: "int \ int" shows "(\a b. f (2 * a) + 2 * f b = f (f (a + b))) \ (\x. f x = 0) \ (\x. f x = 2 * x + f 0)" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs using f_classification[of f] by blast next assume ?rhs - thus ?lhs by smt + thus ?lhs by (smt (verit, ccfv_threshold) mult_2) qed end \ No newline at end of file diff --git a/thys/Isabelle_hoops/Ordinal_Sums.thy b/thys/Isabelle_hoops/Ordinal_Sums.thy --- a/thys/Isabelle_hoops/Ordinal_Sums.thy +++ b/thys/Isabelle_hoops/Ordinal_Sums.thy @@ -1,778 +1,778 @@ section\Ordinal sums\ text\We define @{text "tower of hoops"}, a family of almost disjoint hoops indexed by a total order. This is based on the definition of @{text "bounded tower of irreducible hoops"} in \<^cite>\"BUSANICHE2005"\ (see paragraph after Lemma 3.3). Parting from a tower of hoops we can define a hoop known as @{text "ordinal sum"}. Ordinal sums are a fundamental tool in the study of totally ordered hoops.\ theory Ordinal_Sums imports Hoops begin subsection\Tower of hoops\ locale tower_of_hoops = fixes index_set :: "'b set" ("I") fixes index_lesseq :: "'b \ 'b \ bool" (infix "\\<^sup>I" 60) fixes index_less :: "'b \ 'b \ bool" (infix "<\<^sup>I" 60) fixes universes :: "'b \ ('a set)" ("UNI") fixes multiplications :: "'b \ ('a \ 'a \ 'a)" ("MUL") fixes implications :: "'b \ ('a \ 'a \ 'a)" ("IMP") fixes sum_one :: 'a ("1\<^sup>S") assumes index_set_total_order: "total_poset_on I (\\<^sup>I) (<\<^sup>I)" and almost_disjoint: "i \ I \ j \ I \ i \ j \ UNI i \ UNI j = {1\<^sup>S}" and family_of_hoops: "i \ I \ hoop (UNI i) (MUL i) (IMP i) 1\<^sup>S" begin sublocale total_poset_on "I" "(\\<^sup>I)" "(<\<^sup>I)" using index_set_total_order by simp abbreviation (uni_i) uni_i :: "['b] \ ('a set)" ("(\(\<^sub>_))" [61] 60) where "\\<^sub>i \ UNI i" abbreviation (mult_i) mult_i :: "['b] \ ('a \ 'a \ 'a)" ("(*(\<^sup>_))" [61] 60) where "*\<^sup>i \ MUL i" abbreviation (imp_i) imp_i :: "['b] \ ('a \ 'a \ 'a)" ("(\(\<^sup>_))" [61] 60) where "\\<^sup>i \ IMP i" abbreviation (mult_i_xy) mult_i_xy :: "['a, 'b, 'a] \ 'a" ("((_)/ *(\<^sup>_) / (_))" [61, 50, 61] 60) where "x *\<^sup>i y \ MUL i x y" abbreviation (imp_i_xy) imp_i_xy :: "['a, 'b, 'a] \ 'a" ("((_)/ \(\<^sup>_) / (_))" [61, 50, 61] 60) where "x \\<^sup>i y \ IMP i x y" subsection\Ordinal sum universe\ definition sum_univ :: "'a set" ("S") where "S = {x. \ i \ I. x \ \\<^sub>i}" lemma sum_one_closed [simp]: "1\<^sup>S \ S" using family_of_hoops hoop.one_closed not_empty sum_univ_def by fastforce lemma sum_subsets: assumes "i \ I" shows "\\<^sub>i \ S" using sum_univ_def assms by blast subsection\Floor function: definition and properties\ lemma floor_unique: assumes "a \ S-{1\<^sup>S}" shows "\! i. i \ I \ a \ \\<^sub>i" using assms sum_univ_def almost_disjoint by blast function floor :: "'a \ 'b" where "floor x = (THE i. i \ I \ x \ \\<^sub>i)" if "x \ S-{1\<^sup>S}" | "floor x = undefined" if "x = 1\<^sup>S \ x \ S" by auto termination by lexicographic_order abbreviation (uni_floor) uni_floor :: "['a] \ ('a set)" ("(\\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r (\<^sub>_))" [61] 60) where "\\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>x \ UNI (floor x)" abbreviation (mult_floor) mult_floor :: "['a] \ ('a \ 'a \ 'a)" ("(*\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r (\<^sup>_))" [61] 60) where "*\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a \ MUL (floor a)" abbreviation (imp_floor) imp_floor :: "['a] \ ('a \ 'a \ 'a)" ("(\\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r (\<^sup>_))" [61] 60) where "\\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a \ IMP (floor a)" abbreviation (mult_floor_xy) mult_floor_xy :: "['a, 'a, 'a] \ 'a" ("((_)/ *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r (\<^sup>_) / (_))" [61, 50, 61] 60) where "x *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>y z \ MUL (floor y) x z" abbreviation (imp_floor_xy) imp_floor_xy :: "['a, 'a, 'a] \ 'a" ("((_)/ \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r (\<^sup>_) / (_))" [61, 50, 61] 60) where "x \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>y z \ IMP (floor y) x z" lemma floor_prop: assumes "a \ S-{1\<^sup>S}" shows "floor a \ I \ a \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" proof - have "floor a = (THE i. i \ I \ a \ \\<^sub>i)" using assms by auto then show ?thesis using assms theI_unique floor_unique by (metis (mono_tags, lifting)) qed lemma floor_one_closed: assumes "i \ I" shows "1\<^sup>S \ \\<^sub>i" using assms floor_prop family_of_hoops hoop.one_closed by metis lemma floor_mult_closed: assumes "i \ I" "a \ \\<^sub>i" "b \ \\<^sub>i" shows "a *\<^sup>i b \ \\<^sub>i" using assms family_of_hoops hoop.mult_closed by meson lemma floor_imp_closed: assumes "i \ I" "a \ \\<^sub>i" "b \ \\<^sub>i" shows "a \\<^sup>i b \ \\<^sub>i" using assms family_of_hoops hoop.imp_closed by meson subsection\Ordinal sum multiplication and implication\ function sum_mult :: "'a \ 'a \ 'a" (infix "*\<^sup>S" 60) where "x *\<^sup>S y = x *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>x y" if "x \ S-{1\<^sup>S}" "y \ S-{1\<^sup>S}" "floor x = floor y" | "x *\<^sup>S y = x" if "x \ S-{1\<^sup>S}" "y \ S-{1\<^sup>S}" "floor x <\<^sup>I floor y" | "x *\<^sup>S y = y" if "x \ S-{1\<^sup>S}" "y \ S-{1\<^sup>S}" "floor y <\<^sup>I floor x" | "x *\<^sup>S y = y" if "x = 1\<^sup>S" "y \ S-{1\<^sup>S}" | "x *\<^sup>S y = x" if "x \ S-{1\<^sup>S}" "y = 1\<^sup>S" | "x *\<^sup>S y = 1\<^sup>S" if "x = 1\<^sup>S" "y = 1\<^sup>S" | "x *\<^sup>S y = undefined" if "x \ S \ y \ S" apply auto - using floor.cases floor.simps(1) floor_prop trichotomy apply smt + using floor.cases floor.simps(1) floor_prop trichotomy apply (smt (verit)) using floor_prop strict_iff_order apply force using floor_prop strict_iff_order apply force using floor_prop trichotomy by auto termination by lexicographic_order function sum_imp :: "'a \ 'a \ 'a" (infix "\\<^sup>S" 60) where "x \\<^sup>S y = x \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>x y" if "x \ S-{1\<^sup>S}" "y \ S-{1\<^sup>S}" "floor x = floor y" | "x \\<^sup>S y = 1\<^sup>S" if "x \ S-{1\<^sup>S}" "y \ S-{1\<^sup>S}" "floor x <\<^sup>I floor y" | "x \\<^sup>S y = y" if "x \ S-{1\<^sup>S}" "y \ S-{1\<^sup>S}" "floor y <\<^sup>I floor x" | "x \\<^sup>S y = y" if "x = 1\<^sup>S" "y \ S-{1\<^sup>S}" | "x \\<^sup>S y = 1\<^sup>S" if "x \ S-{1\<^sup>S}" "y = 1\<^sup>S" | "x \\<^sup>S y = 1\<^sup>S" if "x = 1\<^sup>S" "y = 1\<^sup>S" | "x \\<^sup>S y = undefined" if "x \ S \ y \ S" apply auto - using floor.cases floor.simps(1) floor_prop trichotomy apply smt + using floor.cases floor.simps(1) floor_prop trichotomy apply (smt (verit)) using floor_prop strict_iff_order apply force using floor_prop strict_iff_order apply force using floor_prop trichotomy by auto termination by lexicographic_order subsubsection\Some multiplication properties\ lemma sum_mult_not_one_aux: assumes "a \ S-{1\<^sup>S}" "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" shows "a *\<^sup>S b \ (\\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a)-{1\<^sup>S}" proof - consider (1) "b \ S-{1\<^sup>S}" | (2) "b = 1\<^sup>S" using sum_subsets assms floor_prop by blast then show ?thesis proof(cases) case 1 then have same_floor: "floor a = floor b" using assms floor_prop floor_unique by metis moreover have "a *\<^sup>S b = a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b" using "1" assms(1) same_floor by simp moreover have "a \ (\\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a)-{1\<^sup>S} \ b \ (\\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a)-{1\<^sup>S}" using "1" assms floor_prop by simp ultimately show ?thesis using assms(1) family_of_hoops floor_prop hoop.mult_C by metis next case 2 then show ?thesis using assms(1) floor_prop by auto qed qed corollary sum_mult_not_one: assumes "a \ S-{1\<^sup>S}" "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" shows "a *\<^sup>S b \ S-{1\<^sup>S} \ floor (a *\<^sup>S b) = floor a" proof - have "a *\<^sup>S b \ (\\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a)-{1\<^sup>S}" using sum_mult_not_one_aux assms by meson then have "a *\<^sup>S b \ S-{1\<^sup>S} \ a *\<^sup>S b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" using sum_subsets assms(1) floor_prop by fastforce then show ?thesis using assms(1) floor_prop floor_unique by metis qed lemma sum_mult_A: assumes "a \ S-{1\<^sup>S}" "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" shows "a *\<^sup>S b = a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b \ b *\<^sup>S a = b *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a a" proof - consider (1) "b \ S-{1\<^sup>S}" | (2) "b = 1\<^sup>S" using sum_subsets assms floor_prop by blast then show ?thesis proof(cases) case 1 then have "floor a = floor b" using assms floor.cases floor_prop floor_unique by metis then show ?thesis using "1" assms by auto next case 2 then show ?thesis using assms(1) family_of_hoops floor_prop hoop.mult_neutr hoop.mult_neutr_2 by fastforce qed qed subsubsection\Some implication properties\ lemma sum_imp_floor: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" "floor a = floor b" "a \\<^sup>S b \ S-{1\<^sup>S}" shows "floor (a \\<^sup>S b) = floor a" proof - have "a \\<^sup>S b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" using sum_imp.simps(1) assms(1-3) floor_imp_closed floor_prop by metis then show ?thesis using assms(1,4) floor_prop floor_unique by blast qed lemma sum_imp_A: assumes "a \ S-{1\<^sup>S}" "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" shows "a \\<^sup>S b = a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b" proof - consider (1) "b \ S-{1\<^sup>S}" | (2) "b = 1\<^sup>S" using sum_subsets assms floor_prop by blast then show ?thesis proof(cases) case 1 then show ?thesis using sum_imp.simps(1) assms floor_prop floor_unique by metis next case 2 then show ?thesis using sum_imp.simps(5) assms(1) family_of_hoops floor_prop hoop.imp_one_top by metis qed qed lemma sum_imp_B: assumes "a \ S-{1\<^sup>S}" "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" shows "b \\<^sup>S a = b \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a a" proof - consider (1) "b \ S-{1\<^sup>S}" | (2) "b = 1\<^sup>S" using sum_subsets assms floor_prop by blast then show ?thesis proof(cases) case 1 then show ?thesis using sum_imp.simps(1) assms floor_prop floor_unique by metis next case 2 then show ?thesis using sum_imp.simps(4) assms(1) family_of_hoops floor_prop hoop.imp_one_C by metis qed qed lemma sum_imp_floor_antisymm: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" "floor a = floor b" "a \\<^sup>S b = 1\<^sup>S" "b \\<^sup>S a = 1\<^sup>S" shows "a = b" proof - have "a \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a \ b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a \ floor a \ I" using floor_prop assms by metis moreover have "a \\<^sup>S b = a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b \ b \\<^sup>S a = b \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a a" using assms by auto ultimately show ?thesis using assms(4,5) family_of_hoops hoop.ord_antisymm_equiv by metis qed corollary sum_imp_C: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" "a \ b" "floor a = floor b" "a \\<^sup>S b = 1\<^sup>S" shows "b \\<^sup>S a \ 1\<^sup>S" using sum_imp_floor_antisymm assms by blast lemma sum_imp_D: assumes "a \ S" shows "1\<^sup>S \\<^sup>S a = a" using sum_imp.simps(4,6) assms by blast lemma sum_imp_E: assumes "a \ S" shows "a \\<^sup>S 1\<^sup>S = 1\<^sup>S" using sum_imp.simps(5,6) assms by blast subsection\The ordinal sum of a tower of hoops is a hoop\ subsubsection\@{term S} is not empty\ lemma sum_not_empty: "S \ \" using sum_one_closed by blast subsubsection\@{term sum_mult} and @{term sum_imp} are well defined\ lemma sum_mult_closed_one: assumes "a \ S" "b \ S" "a = 1\<^sup>S \ b = 1\<^sup>S" shows "a *\<^sup>S b \ S" using sum_mult.simps(4-6) assms floor.cases by metis lemma sum_mult_closed_not_one: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" shows "a *\<^sup>S b \ S-{1\<^sup>S}" proof - from assms consider (1) "floor a = floor b" | (2) "floor a <\<^sup>I floor b \ floor b <\<^sup>I floor a" using trichotomy floor_prop by blast then show ?thesis proof(cases) case 1 then show ?thesis using sum_mult_not_one assms floor_prop by metis next case 2 then show ?thesis using assms by auto qed qed lemma sum_mult_closed: assumes "a \ S" "b \ S" shows "a *\<^sup>S b \ S" using sum_mult_closed_not_one sum_mult_closed_one assms by auto lemma sum_imp_closed_one: assumes "a \ S" "b \ S" "a = 1\<^sup>S \ b = 1\<^sup>S" shows "a \\<^sup>S b \ S" using sum_imp.simps(4-6) assms floor.cases by metis lemma sum_imp_closed_not_one: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" shows "a \\<^sup>S b \ S" proof - from assms consider (1) "floor a = floor b" | (2) "floor a <\<^sup>I floor b \ floor b <\<^sup>I floor a" using trichotomy floor_prop by blast then show "a \\<^sup>S b \ S" proof(cases) case 1 then have "a \\<^sup>S b = a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b" using assms by auto moreover have "a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" using "1" assms floor_imp_closed floor_prop by metis ultimately show ?thesis using sum_subsets assms(1) floor_prop by auto next case 2 then show ?thesis using assms by auto qed qed lemma sum_imp_closed: assumes "a \ S" "b \ S" shows "a \\<^sup>S b \ S" using sum_imp_closed_one sum_imp_closed_not_one assms by auto subsubsection\Neutrality of @{term sum_one}\ lemma sum_mult_neutr: assumes "a \ S" shows "a *\<^sup>S 1\<^sup>S = a \ 1\<^sup>S *\<^sup>S a = a" using assms sum_mult.simps(4-6) by blast subsubsection\Commutativity of @{term sum_mult}\ text\Now we prove @{term "x *\<^sup>S y = y *\<^sup>S x"} by showing that it holds when one of the variables is equal to @{term "1\<^sup>S"}. Then we consider when none of them is @{term "1\<^sup>S"}.\ lemma sum_mult_comm_one: assumes "a \ S" "b \ S" "a = 1\<^sup>S \ b = 1\<^sup>S" shows "a *\<^sup>S b = b *\<^sup>S a" using sum_mult_neutr assms by auto lemma sum_mult_comm_not_one: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" shows "a *\<^sup>S b = b *\<^sup>S a" proof - from assms consider (1) "floor a = floor b" | (2) "floor a <\<^sup>I floor b \ floor b <\<^sup>I floor a" using trichotomy floor_prop by blast then show ?thesis proof(cases) case 1 then have same_floor: "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" using assms(2) floor_prop by simp then have "a *\<^sup>S b = a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b" using sum_mult_A assms(1) by blast also have "\ = b *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a a" using assms(1) family_of_hoops floor_prop hoop.mult_comm same_floor by meson also have "\ = b *\<^sup>S a" using sum_mult_A assms(1) same_floor by simp finally show ?thesis by auto next case 2 then show ?thesis using assms by auto qed qed lemma sum_mult_comm: assumes "a \ S" "b \ S" shows "a *\<^sup>S b = b *\<^sup>S a" using assms sum_mult_comm_one sum_mult_comm_not_one by auto subsubsection\Associativity of @{term sum_mult}\ text\Next we prove @{term "x *\<^sup>S (y *\<^sup>S z) = (x *\<^sup>S y) *\<^sup>S z"}.\ lemma sum_mult_assoc_one: assumes "a \ S" "b \ S" "c \ S" "a = 1\<^sup>S \ b = 1\<^sup>S \ c = 1\<^sup>S" shows "a *\<^sup>S (b *\<^sup>S c) = (a *\<^sup>S b) *\<^sup>S c" using sum_mult_neutr assms sum_mult_closed by metis lemma sum_mult_assoc_not_one: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" "c \ S-{1\<^sup>S}" shows "a *\<^sup>S (b *\<^sup>S c) = (a *\<^sup>S b) *\<^sup>S c" proof - from assms consider (1) "floor a = floor b" "floor b = floor c" | (2) "floor a = floor b" "floor b <\<^sup>I floor c" | (3) "floor a = floor b" "floor c <\<^sup>I floor b" | (4) "floor a <\<^sup>I floor b" "floor b = floor c" | (5) "floor a <\<^sup>I floor b" "floor b <\<^sup>I floor c" | (6) "floor a <\<^sup>I floor b" "floor c <\<^sup>I floor b" | (7) "floor b <\<^sup>I floor a" "floor b = floor c" | (8) "floor b <\<^sup>I floor a" "floor b <\<^sup>I floor c" | (9) "floor b <\<^sup>I floor a" "floor c <\<^sup>I floor b" using trichotomy floor_prop by meson then show ?thesis proof(cases) case 1 then have "a *\<^sup>S (b *\<^sup>S c) = a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a (b *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a c)" using sum_mult_A assms floor_mult_closed floor_prop by metis also have "\ = (a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b) *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a c" using "1" assms family_of_hoops floor_prop hoop.mult_assoc by metis also have "\ = (a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>b b) *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>b c" using "1" by simp also have "\ = (a *\<^sup>S b) *\<^sup>S c" using "1" sum_mult_A assms floor_mult_closed floor_prop by metis finally show ?thesis by auto next case 2 then show ?thesis using sum_mult.simps(2,3) sum_mult_not_one assms floor_prop by metis next case 3 then show ?thesis using sum_mult.simps(3) sum_mult_not_one assms floor_prop by metis next case 4 then show ?thesis using sum_mult.simps(2) sum_mult_not_one assms floor_prop by metis next case 5 then show ?thesis using sum_mult.simps(2) assms floor_prop strict_trans by metis next case 6 then show ?thesis using sum_mult.simps(2,3) assms by metis next case 7 then show ?thesis using sum_mult.simps(3) sum_mult_not_one assms floor_prop by metis next case 8 then show ?thesis using sum_mult.simps(2,3) assms by metis next case 9 then show ?thesis using sum_mult.simps(3) assms floor_prop strict_trans by metis qed qed lemma sum_mult_assoc: assumes "a \ S" "b \ S" "c \ S" shows "a *\<^sup>S (b *\<^sup>S c) = (a *\<^sup>S b) *\<^sup>S c" using assms sum_mult_assoc_one sum_mult_assoc_not_one by blast subsubsection\Reflexivity of @{term sum_imp}\ lemma sum_imp_reflex: assumes "a \ S" shows "a \\<^sup>S a = 1\<^sup>S" proof - consider (1) "a \ S-{1\<^sup>S}" | (2) "a = 1\<^sup>S" using assms by blast then show ?thesis proof(cases) case 1 then have "a \\<^sup>S a = a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a a" by simp then show ?thesis using "1" family_of_hoops floor_prop hoop.imp_reflex by metis next case 2 then show ?thesis by simp qed qed subsubsection\Divisibility\ text\We prove @{term "x *\<^sup>S (x \\<^sup>S y) = y *\<^sup>S (y \\<^sup>S x)"} using the same methods as before.\ lemma sum_divisibility_one: assumes "a \ S" "b \ S" "a = 1\<^sup>S \ b = 1\<^sup>S" shows "a *\<^sup>S (a \\<^sup>S b) = b *\<^sup>S (b \\<^sup>S a)" proof - have "x \\<^sup>S y = y \ y \\<^sup>S x = 1\<^sup>S" if "x = 1\<^sup>S" "y \ S" for x y using sum_imp_D sum_imp_E that by simp then show ?thesis using assms sum_mult_neutr by metis qed lemma sum_divisibility_aux: assumes "a \ S-{1\<^sup>S}" "b \ \\<^sub>f\<^sub>l\<^sub>o\<^sub>o\<^sub>r \<^sub>a" shows "a *\<^sup>S (a \\<^sup>S b) = a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a (a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b)" using sum_imp_A sum_mult_A assms floor_imp_closed floor_prop by metis lemma sum_divisibility_not_one: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" shows "a *\<^sup>S (a \\<^sup>S b) = b *\<^sup>S (b \\<^sup>S a)" proof - from assms consider (1) "floor a = floor b" | (2) "floor a <\<^sup>I floor b \ floor b <\<^sup>I floor a" using trichotomy floor_prop by blast then show ?thesis proof(cases) case 1 then have "a *\<^sup>S (a \\<^sup>S b) = a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a (a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b)" using "1" sum_divisibility_aux assms floor_prop by metis also have "\ = b *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a (b \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a a)" using "1" assms family_of_hoops floor_prop hoop.divisibility by metis also have "\ = b *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>b (b \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>b a)" using "1" by simp also have "\ = b *\<^sup>S (b \\<^sup>S a)" using "1" sum_divisibility_aux assms floor_prop by metis finally show ?thesis by auto next case 2 then show ?thesis using assms by auto qed qed lemma sum_divisibility: assumes "a \ S" "b \ S" shows "a *\<^sup>S (a \\<^sup>S b) = b *\<^sup>S (b \\<^sup>S a)" using assms sum_divisibility_one sum_divisibility_not_one by auto subsubsection\Residuation\ text\Finally we prove @{term "(x *\<^sup>S y) \\<^sup>S z = (x \\<^sup>S (y \\<^sup>S z))"}.\ lemma sum_residuation_one: assumes "a \ S" "b \ S" "c \ S" "a = 1\<^sup>S \ b = 1\<^sup>S \ c = 1\<^sup>S" shows "(a *\<^sup>S b) \\<^sup>S c = a \\<^sup>S (b \\<^sup>S c)" using sum_imp_D sum_imp_E sum_imp_closed sum_mult_closed sum_mult_neutr assms by metis lemma sum_residuation_not_one: assumes "a \ S-{1\<^sup>S}" "b \ S-{1\<^sup>S}" "c \ S-{1\<^sup>S}" shows "(a *\<^sup>S b) \\<^sup>S c = a \\<^sup>S (b \\<^sup>S c)" proof - from assms consider (1) "floor a = floor b" "floor b = floor c" | (2) "floor a = floor b" "floor b <\<^sup>I floor c" | (3) "floor a = floor b" "floor c <\<^sup>I floor b" | (4) "floor a <\<^sup>I floor b" "floor b = floor c" | (5) "floor a <\<^sup>I floor b" "floor b <\<^sup>I floor c" | (6) "floor a <\<^sup>I floor b" "floor c <\<^sup>I floor b" | (7) "floor b <\<^sup>I floor a" "floor b = floor c" | (8) "floor b <\<^sup>I floor a" "floor b <\<^sup>I floor c" | (9) "floor b <\<^sup>I floor a" "floor c <\<^sup>I floor b" using trichotomy floor_prop by meson then show ?thesis proof(cases) case 1 then have "(a *\<^sup>S b) \\<^sup>S c = (a *\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a b) \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a c" using sum_imp_B sum_mult_A assms floor_mult_closed floor_prop by metis also have "\ = a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a (b \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>a c)" using "1" assms family_of_hoops floor_prop hoop.residuation by metis also have "\ = a \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>b (b \\<^sup>f\<^sup>l\<^sup>o\<^sup>o\<^sup>r \<^sup>b c)" using "1" by simp also have "\ = a \\<^sup>S (b \\<^sup>S c)" using "1" sum_imp_A assms floor_imp_closed floor_prop by metis finally show ?thesis by auto next case 2 then show ?thesis using sum_imp.simps(2,5) sum_mult_not_one assms floor_prop by metis next case 3 then show ?thesis using sum_imp.simps(3) sum_mult_not_one assms floor_prop by metis next case 4 then have "(a *\<^sup>S b) \\<^sup>S c = 1\<^sup>S" using "4" sum_imp.simps(2) sum_mult.simps(2) assms by metis moreover have "b \\<^sup>S c = 1\<^sup>S \ (b \\<^sup>S c \ S-{1\<^sup>S} \ floor (b \\<^sup>S c) = floor b)" using "4"(2) sum_imp_closed_not_one sum_imp_floor assms(2,3) by blast ultimately show ?thesis using "4"(1) sum_imp.simps(2,5) assms(1) by metis next case 5 then show ?thesis using sum_imp.simps(2,5) sum_mult.simps(2) assms floor_prop strict_trans by metis next case 6 then show ?thesis using assms by auto next case 7 then have "(a *\<^sup>S b) \\<^sup>S c = (b \\<^sup>S c)" using assms(1,2) by auto moreover have "b \\<^sup>S c = 1\<^sup>S \ (b \\<^sup>S c \ S-{1\<^sup>S} \ floor (b \\<^sup>S c) = floor b)" using "7"(2) sum_imp_closed_not_one sum_imp_floor assms(2,3) by blast ultimately show ?thesis using "7"(1) sum_imp.simps(3,5) assms(1) by metis next case 8 then show ?thesis using assms by auto next case 9 then show ?thesis using sum_imp.simps(3) sum_mult.simps(3) assms floor_prop strict_trans by metis qed qed lemma sum_residuation: assumes "a \ S" "b \ S" "c \ S" shows "(a *\<^sup>S b) \\<^sup>S c = a \\<^sup>S (b \\<^sup>S c)" using assms sum_residuation_one sum_residuation_not_one by blast subsubsection\Main result\ sublocale hoop "S" "(*\<^sup>S)" "(\\<^sup>S)" "1\<^sup>S" proof show "x *\<^sup>S y \ S" if "x \ S" "y \ S" for x y using that sum_mult_closed by simp next show "x \\<^sup>S y \ S" if "x \ S" "y \ S" for x y using that sum_imp_closed by simp next show "1\<^sup>S \ S" by simp next show "x *\<^sup>S y = y *\<^sup>S x" if "x \ S" "y \ S" for x y using that sum_mult_comm by simp next show "x *\<^sup>S (y *\<^sup>S z) = (x *\<^sup>S y) *\<^sup>S z" if "x \ S" "y \ S" "z \ S" for x y z using that sum_mult_assoc by simp next show "x *\<^sup>S 1\<^sup>S = x" if "x \ S" for x using that sum_mult_neutr by simp next show "x \\<^sup>S x = 1\<^sup>S" if "x \ S" for x using that sum_imp_reflex by simp next show "x *\<^sup>S (x \\<^sup>S y) = y *\<^sup>S (y \\<^sup>S x)" if "x \ S" "y \ S" for x y using that sum_divisibility by simp next show "x \\<^sup>S (y \\<^sup>S z) = (x *\<^sup>S y) \\<^sup>S z" if "x \ S" "y \ S" "z \ S" for x y z using that sum_residuation by simp qed end end \ No newline at end of file diff --git a/thys/Isabelle_hoops/Posets.thy b/thys/Isabelle_hoops/Posets.thy --- a/thys/Isabelle_hoops/Posets.thy +++ b/thys/Isabelle_hoops/Posets.thy @@ -1,57 +1,57 @@ section\Some order tools: posets with explicit universe\ theory Posets imports Main "HOL-Library.LaTeXsugar" begin locale poset_on = fixes P :: "'b set" fixes P_lesseq :: "'b \ 'b \ bool" (infix "\\<^sup>P" 60) fixes P_less :: "'b \ 'b \ bool" (infix "<\<^sup>P" 60) assumes not_empty [simp]: "P \ \" and reflex: "reflp_on P (\\<^sup>P)" and antisymm: "antisymp_on P (\\<^sup>P)" and trans: "transp_on P (\\<^sup>P)" and strict_iff_order: "x \ P \ y \ P \ x <\<^sup>P y = (x \\<^sup>P y \ x \ y)" begin lemma strict_trans: assumes "a \ P" "b \ P" "c \ P" "a <\<^sup>P b" "b <\<^sup>P c" shows "a <\<^sup>P c" using antisymm antisymp_onD assms trans strict_iff_order transp_onD - by smt + by (smt (verit, ccfv_SIG)) end locale bot_poset_on = poset_on + fixes bot :: "'b" ("0\<^sup>P") assumes bot_closed: "0\<^sup>P \ P" and bot_first: "x \ P \ 0\<^sup>P \\<^sup>P x" locale top_poset_on = poset_on + fixes top :: "'b" ("1\<^sup>P") assumes top_closed: "1\<^sup>P \ P" and top_last: "x \ P \ x \\<^sup>P 1\<^sup>P" locale bounded_poset_on = bot_poset_on + top_poset_on locale total_poset_on = poset_on + assumes total: "totalp_on P (\\<^sup>P)" begin lemma trichotomy: assumes "a \ P" "b \ P" shows "(a <\<^sup>P b \ \(a = b \ b <\<^sup>P a)) \ (a = b \ \(a <\<^sup>P b \ b <\<^sup>P a)) \ (b <\<^sup>P a \ \(a = b \ a <\<^sup>P b))" using antisymm antisymp_onD assms strict_iff_order total totalp_onD by metis lemma strict_order_equiv_not_converse: assumes "a \ P" "b \ P" shows "a <\<^sup>P b \ \(b \\<^sup>P a)" using assms strict_iff_order reflex reflp_onD strict_trans trichotomy by metis end end \ No newline at end of file diff --git a/thys/KD_Tree/Range_Search.thy b/thys/KD_Tree/Range_Search.thy --- a/thys/KD_Tree/Range_Search.thy +++ b/thys/KD_Tree/Range_Search.thy @@ -1,87 +1,88 @@ (* File: Range_Search.thy Author: Martin Rau, TU München *) section \Range Searching\ theory Range_Search imports KD_Tree begin text\ Given two \k\-dimensional points \p\<^sub>0\ and \p\<^sub>1\ which bound the search space, the search should return only the points which satisfy the following criteria: For every point p in the resulting set: \newline \hspace{1cm} For every axis @{term "k"}: \newline \hspace{2cm} @{term "p\<^sub>0$k \ p$k \ p$k \ p\<^sub>1$k"} \newline For a \2\-d tree a query corresponds to selecting all the points in the rectangle that has \p\<^sub>0\ and \p\<^sub>1\ as its defining edges. \ subsection \Rectangle Definition\ lemma cbox_point_def: fixes p\<^sub>0 :: "('k::finite) point" shows "cbox p\<^sub>0 p\<^sub>1 = { p. \k. p\<^sub>0$k \ p$k \ p$k \ p\<^sub>1$k }" proof - have "cbox p\<^sub>0 p\<^sub>1 = { p. \k. p\<^sub>0 \ axis k 1 \ p \ axis k 1 \ p \ axis k 1 \ p\<^sub>1 \ axis k 1 }" unfolding cbox_def using axis_inverse by auto also have "... = { p. \k. p\<^sub>0$k \ 1 \ p$k \ 1 \ p$k \ 1 \ p\<^sub>1$k \ 1 }" - using inner_axis[of _ _ 1] by (smt Collect_cong) + using inner_axis[of _ _ 1] + by (metis (mono_tags, opaque_lifting)) also have "... = { p. \k. p\<^sub>0$k \ p$k \ p$k \ p\<^sub>1$k }" by simp finally show ?thesis . qed subsection \Search Function\ fun search :: "('k::finite) point \ 'k point \ 'k kdt \ 'k point set" where "search p\<^sub>0 p\<^sub>1 (Leaf p) = (if p \ cbox p\<^sub>0 p\<^sub>1 then { p } else {})" | "search p\<^sub>0 p\<^sub>1 (Node k v l r) = ( if v < p\<^sub>0$k then search p\<^sub>0 p\<^sub>1 r else if p\<^sub>1$k < v then search p\<^sub>0 p\<^sub>1 l else search p\<^sub>0 p\<^sub>1 l \ search p\<^sub>0 p\<^sub>1 r )" subsection \Auxiliary Lemmas\ lemma l_empty: assumes "invar (Node k v l r)" "v < p\<^sub>0$k" shows "set_kdt l \ cbox p\<^sub>0 p\<^sub>1 = {}" proof - have "\p \ set_kdt l. p$k < p\<^sub>0$k" using assms by auto hence "\p \ set_kdt l. p \ cbox p\<^sub>0 p\<^sub>1" using cbox_point_def leD by blast thus ?thesis by blast qed lemma r_empty: assumes "invar (Node k v l r)" "p\<^sub>1$k < v" shows "set_kdt r \ cbox p\<^sub>0 p\<^sub>1 = {}" proof - have "\p \ set_kdt r. p\<^sub>1$k < p$k" using assms by auto hence "\p \ set_kdt r. p \ cbox p\<^sub>0 p\<^sub>1" using cbox_point_def leD by blast thus ?thesis by blast qed subsection \Main Theorem\ theorem search_cbox: assumes "invar kdt" shows "search p\<^sub>0 p\<^sub>1 kdt = set_kdt kdt \ cbox p\<^sub>0 p\<^sub>1" using assms l_empty r_empty by (induction kdt) (auto, blast+) end diff --git a/thys/Lucas_Theorem/Lucas_Theorem.thy b/thys/Lucas_Theorem/Lucas_Theorem.thy --- a/thys/Lucas_Theorem/Lucas_Theorem.thy +++ b/thys/Lucas_Theorem/Lucas_Theorem.thy @@ -1,372 +1,373 @@ (* Title: Lucas_Theorem.thy Author: Chelsea Edmonds, University of Cambridge *) theory Lucas_Theorem imports Main "HOL-Computational_Algebra.Computational_Algebra" begin notation fps_nth (infixl "$" 75) section \Extensions on Formal Power Series (FPS) Library\ text \This section presents a few extensions on the Formal Power Series (FPS) library, described in \<^cite>\"Chaieb2011"\ \ subsection \FPS Equivalence Relation \ text \ This proof requires reasoning around the equivalence of coefficients mod some prime number. This section defines an equivalence relation on FPS using the pattern described by Paulson in \<^cite>\"paulsonDefiningFunctionsEquivalence2006"\, as well as some basic lemmas for reasoning around how the equivalence holds after common operations are applied \ definition "fpsmodrel p \ { (f, g). \ n. (f $ n) mod p = (g $ n) mod p }" lemma fpsrel_iff [simp]: "(f, g) \ fpsmodrel p \ (\n. (f $ n) mod p = (g $ n) mod p)" by (simp add: fpsmodrel_def) lemma fps_equiv: "equiv UNIV (fpsmodrel p)" proof (rule equivI) show "refl (fpsmodrel p)" by (simp add: refl_on_def fpsmodrel_def) show "sym (fpsmodrel p)" by (simp add: sym_def fpsmodrel_def) show "trans (fpsmodrel p)" by (intro transI) (simp add: fpsmodrel_def) qed text \ Equivalence relation over multiplication \ lemma fps_mult_equiv_coeff: fixes f g :: "('a :: {euclidean_ring_cancel}) fps" assumes "(f, g) \ fpsmodrel p" shows "(f*h)$n mod p = (g*h)$n mod p" proof - have "((f*h) $ n) mod p =(\i=0..n. (f$i mod p * h$(n - i) mod p) mod p) mod p" using mod_sum_eq mod_mult_left_eq by (simp add: fps_mult_nth mod_sum_eq mod_mult_left_eq) also have "... = (\i=0..n. (g$i mod p * h$(n - i) mod p) mod p) mod p" using assms by auto also have "... = ((g*h) $ n) mod p" by (simp add: mod_mult_left_eq mod_sum_eq fps_mult_nth) thus ?thesis by (simp add: calculation) qed lemma fps_mult_equiv: fixes f g :: "('a :: {euclidean_ring_cancel}) fps" assumes "(f, g) \ fpsmodrel p" shows "(f*h, g*h) \ fpsmodrel p" using fpsmodrel_def fps_mult_equiv_coeff assms by blast text \ Equivalence relation over power operator \ lemma fps_power_equiv: fixes f g :: "('a :: {euclidean_ring_cancel}) fps" fixes x :: nat assumes "(f, g) \ fpsmodrel p" shows "(f^x, g^x) \ fpsmodrel p" using assms proof (induct x) case 0 thus ?case by (simp add: fpsmodrel_def) next case (Suc x) then have hyp: " \n. f^x $ n mod p = g ^x $ n mod p" using fpsrel_iff by blast thus ?case proof - have fact: "\n h. (g * h) $ n mod p = (f * h) $ n mod p" by (metis assms fps_mult_equiv_coeff) have "\n h. (g ^ x * h) $ n mod p = (f ^ x * h) $ n mod p" by (simp add: fps_mult_equiv_coeff hyp) then have "\n h. (h * g ^ x) $ n mod p = (h * f ^ x) $ n mod p" by (simp add: mult.commute) thus ?thesis using fact by force qed qed subsection \Binomial Coefficients \ text \The @{term "fps_binomial"} definition in the formal power series uses the @{term "n gchoose k"} operator. It's defined as being of type @{typ "'a :: field_char_0 fps"}, however the equivalence relation requires a type @{typ 'a} that supports the modulo operator. The proof of the binomial theorem based on FPS coefficients below uses the choose operator and does not put bounds on the type of @{term "fps_X"}.\ lemma binomial_coeffs_induct: fixes n k :: nat shows "(1 + fps_X)^n $ k = of_nat(n choose k)" proof (induct n arbitrary: k) case 0 thus ?case by (metis binomial_eq_0_iff binomial_n_0 fps_nth_of_nat not_gr_zero of_nat_0 of_nat_1 power_0) next case h: (Suc n) - fix k have start: "(1 + fps_X)^(n + 1) = (1 + fps_X) * (1 + fps_X)^n" by auto show ?case using One_nat_def Suc_eq_plus1 Suc_pred add.commute binomial_Suc_Suc binomial_n_0 - fps_mult_fps_X_plus_1_nth h.hyps neq0_conv start by (smt of_nat_add) + fps_mult_fps_X_plus_1_nth h.hyps neq0_conv start + by (smt (verit, del_insts) of_nat_add) qed subsection \Freshman's Dream Lemma on FPS \ text \ The Freshman's dream lemma modulo a prime number $p$ is a well known proof that $(1 + x^p) \equiv (1 + x)^p \mod p$\ text \ First prove that $\binom{p^n}{k} \equiv 0 \mod p$ for $k \ge 1$ and $k < p^n$. The eventual proof only ended up requiring this with $n = 1$\ lemma pn_choose_k_modp_0: fixes n k::nat assumes "prime p" "k \ 1 \ k \ p^n - 1" "n > 0" shows "(p^n choose k) mod p = 0" proof - have inequality: "k \ p^n" using assms (2) by arith have choose_take_1: "((p^n - 1) choose ( k - 1))= fact (p^n - 1) div (fact (k - 1) * fact (p^n - k))" using binomial_altdef_nat diff_le_mono inequality assms(2) by auto have "k * (p^n choose k) = k * ((fact (p^n)) div (fact k * fact((p^n) - k)))" using assms binomial_fact'[OF inequality] by auto also have "... = k * fact (p^n) div (fact k * fact((p^n) - k))" using binomial_fact_lemma div_mult_self_is_m fact_gt_zero inequality mult.assoc mult.commute - nat_0_less_mult_iff by smt + nat_0_less_mult_iff + by (simp add: choose_dvd div_mult_swap) also have "... = k * fact (p^n) div (k * fact (k - 1) * fact((p^n) - k))" by (metis assms(2) fact_nonzero fact_num_eq_if le0 le_antisym of_nat_id) also have "... = fact (p^n) div (fact (k - 1) * fact((p^n) - k))" using assms by auto also have "... = ((p^n) * fact (p^n - 1)) div (fact (k - 1) * fact((p^n) - k))" by (metis assms(2) fact_nonzero fact_num_eq_if inequality le0 le_antisym of_nat_id) also have "... = (p^n) * (fact (p^n - 1) div (fact (k - 1) * fact((p^n) - k)))" by (metis assms(2) calculation choose_take_1 neq0_conv not_one_le_zero times_binomial_minus1_eq) finally have equality: "k * (p^n choose k) = p^n * ((p^n - 1) choose (k - 1))" using assms(2) times_binomial_minus1_eq by auto then have dvd_result: "p^n dvd (k * (p^n choose k))" by simp have "\ (p^n dvd k)" using assms (2) binomial_n_0 diff_diff_cancel nat_dvd_not_less neq0_conv by auto then have "p dvd (p^n choose k)" using mult.commute prime_imp_prime_elem prime_power_dvd_multD assms dvd_result by metis thus "?thesis" by simp qed text \ Applying the above lemma to the coefficients of $(1 + X)^p$, it is easy to show that all coefficients other than the $0$th and $p$th will be $0$ \ lemma fps_middle_coeffs: assumes "prime p" "n \ 0 \ n \ p" shows "((1 + fps_X :: int fps) ^p) $ n mod p = 0 mod p" proof - let ?f = "(1 + fps_X :: int fps)^p" have "\ n. n > 0 \ n < p \ (p choose n) mod p = 0" using pn_choose_k_modp_0 [of p _ 1] \prime p\ by auto then have middle_0: "\ n. n > 0 \ n < p \ (?f $ n) mod p = 0" using binomial_coeffs_induct by (metis of_nat_0 zmod_int) have "\ n. n > p \ ?f $ n mod p = 0" using binomial_eq_0_iff binomial_coeffs_induct mod_0 by (metis of_nat_eq_0_iff) thus ?thesis using middle_0 assms(2) nat_neq_iff by auto qed text \It follows that $(1+ X)^p$ is equivalent to $(1 + X^p)$ under our equivalence relation, as required to prove the freshmans dream lemma. \ lemma fps_freshmans_dream: assumes "prime p" shows "(((1 + fps_X :: int fps ) ^p), (1 + (fps_X)^(p))) \ fpsmodrel p" proof - let ?f = "(1 + fps_X :: int fps)^p" let ?g = "(1 + (fps_X :: int fps)^p)" have all_f_coeffs: "\ n. n \ 0 \ n \ p \ ?f $ n mod p = 0 mod p" using fps_middle_coeffs assms by blast have "?g $ 0 = 1" using assms by auto then have "?g $ 0 mod p = 1 mod p" using int_ops(2) zmod_int assms by presburger then have "?g $ p mod p = 1 mod p" using assms by auto then have "\ n . ?f $ n mod p = ?g $ n mod p" using all_f_coeffs by (simp add: binomial_coeffs_induct) thus ?thesis using fpsrel_iff by blast qed section \Lucas's Theorem Proof\ text \A formalisation of Lucas's theorem based on a generating function proof using the existing formal power series (FPS) Isabelle library\ subsection \Reasoning about Coefficients Helpers\ text \A generating function proof of Lucas's theorem relies on direct comparison between coefficients of FPS which requires a number of helper lemmas to prove formally. In particular it compares the coefficients of $(1 + X)^n \mod p$ to $(1 + X^p)^N * (1 + X) ^rn \mod p$, where $N = n / p$, and $rn = n \mod p$. This section proves that the $k$th coefficient of $(1 + X^p)^N * (1 + X) ^rn = (N choose K) * (rn choose rk)$\ text \Applying the @{term "fps_compose"} operator enables reasoning about the coefficients of $(1 + X^p)^n$ using the existing binomial theorem proof with $X^p$ instead of $X$.\ lemma fps_binomial_p_compose: assumes "p \ 0" shows "(1 + (fps_X:: ('a :: {idom} fps))^p)^n = ((1 + fps_X)^n) oo (fps_X^p)" proof - have "(1::'a fps) + fps_X ^ p = 1 + fps_X oo fps_X ^ p" by (simp add: assms fps_compose_add_distrib) then show ?thesis by (simp add: assms fps_compose_power) qed text \ Next the proof determines the value of the $k$th coefficient of $(1 + X^p)^N$. \ lemma fps_X_pow_binomial_coeffs: assumes "prime p" shows "(1 + (fps_X ::int fps)^p)^N $k = (if p dvd k then (N choose (k div p)) else 0)" proof - let ?fx = "(fps_X :: int fps)" have "(1 + ?fx^p)^N $ k = (((1 + ?fx)^N) oo (?fx^p)) $k" by (metis assms fps_binomial_p_compose not_prime_0) also have "... = (\i=0..k.((1 + ?fx)^N)$i * ((?fx^p)^i$k))" by (simp add: fps_compose_nth) finally have coeffs: "(1 + ?fx^p)^N $ k = (\i=0..k. (N choose i) * ((?fx^(p*i))$k))" using binomial_coeffs_induct sum.cong by (metis (no_types, lifting) power_mult) thus ?thesis proof (cases "p dvd k") case False \ \$p$ does not divide $k$ implies the $k$th term has a coefficient of 0\ have "\ i. \(p dvd k) \ (?fx^(p*i)) $ k = 0" by auto thus ?thesis using coeffs by (simp add: False) next case True \ \$p$ divides $k$ implies the $k$th term has a non-zero coefficient\ have contained: "k div p \ {0.. k}" by simp have "\ i. i \ k div p \ (?fx^(p*i)) $ k = 0" using assms by auto then have notdivpis0: "\ i \ ({0 .. k} - {k div p}). (?fx^(p*i)) $ k = 0" by simp have "(1 + ?fx^p)^N $ k = (N choose (k div p)) * (?fx^(p * (k div p))) $ k + (\i\({0..k} -{k div p}). (N choose i) * ((?fx^(p*i))$k))" using contained coeffs sum.remove by (metis (no_types, lifting) finite_atLeastAtMost) thus ?thesis using notdivpis0 True by simp qed qed text \ The final helper lemma proves the $k$th coefficient is equivalent to $\binom{?N}{?K}*\binom{?rn}{?rk}$ as required.\ lemma fps_div_rep_coeffs: assumes "prime p" shows "((1 + (fps_X::int fps)^p)^(n div p) * (1 + fps_X)^(n mod p)) $ k = ((n div p) choose (k div p)) * ((n mod p) choose (k mod p))" (is "((1 + (fps_X::int fps)^p)^?N * (1 + fps_X)^?rn) $ k = (?N choose ?K) * (?rn choose ?rk)") proof - \ \Initial facts with results around representation and 0 valued terms\ let ?fx = "fps_X :: int fps" have krep: "k - ?rk = ?K*p" by (simp add: minus_mod_eq_mult_div) have rk_in_range: "?rk \ {0..k}" by simp have "\ i \ p. (?rn choose i) = 0" using binomial_eq_0_iff by (metis assms(1) leD le_less_trans linorder_cases mod_le_divisor mod_less_divisor prime_gt_0_nat) then have ptok0: "\ i \ {p..k}. ((?rn choose i) * (1 + ?fx^p)^?N $ (k - i)) = 0" by simp then have notrkis0: "\i \ {0.. k}. i \ ?rk \ (?rn choose i) * (1 + ?fx^p)^?N $ (k - i) = 0" proof (cases "k < p") case True \ \When $k < p$, it presents a side case with regards to range of reasoning\ then have k_value: "k = ?rk" by simp then have "\ i < k. \ (p dvd (k - i))" using True by (metis diff_diff_cancel diff_is_0_eq dvd_imp_mod_0 less_imp_diff_less less_irrefl_nat mod_less) then show ?thesis using fps_X_pow_binomial_coeffs assms(1) k_value by simp next case False then have "\ i < p. i \ ?rk \ \(p dvd (k - i))" using mod_nat_eqI by auto then have "\ i \ {0.. ?rk \ (1 + ?fx^p)^?N $ (k - i) = 0" using assms fps_X_pow_binomial_coeffs by simp then show ?thesis using ptok0 by auto qed \ \Main body of the proof, using helper facts above\ have "((1 + fps_X^p)^?N * (1 + fps_X)^?rn) $ k = (((1 + fps_X)^?rn) * (1 + fps_X^p)^?N) $ k" by (metis (no_types, opaque_lifting) distrib_left distrib_right fps_mult_fps_X_commute fps_one_mult(1) fps_one_mult(2) power_commuting_commutes) also have "... = (\i=0..k.(of_nat(?rn choose i)) * ((1 + (fps_X)^p)^?N $ (k - i)))" by (simp add: fps_mult_nth binomial_coeffs_induct) also have "... = ((?rn choose ?rk) * (1 + ?fx^p)^?N $ (k - ?rk)) + (\i\({0..k} - {?rk}). (?rn choose i) * (1 + ?fx^p)^?N $ (k - i))" using rk_in_range sum.remove by (metis (no_types, lifting) finite_atLeastAtMost) finally have "((1 + ?fx^p)^?N * (1 + ?fx)^?rn) $ k = ((?rn choose ?rk) * (1 + ?fx^p)^?N $ (k - ?rk))" using notrkis0 by simp thus ?thesis using fps_X_pow_binomial_coeffs assms krep by auto qed (* Lucas theorem proof *) subsection \Lucas Theorem Proof\ text \ The proof of Lucas's theorem combines a generating function approach, based off \<^cite>\"Fine"\ with induction. For formalisation purposes, it was easier to first prove a well known corollary of the main theorem (also often presented as an alternative statement for Lucas's theorem), which can itself be used to backwards prove the the original statement by induction. This approach was adapted from P. Cameron's lecture notes on combinatorics \<^cite>\"petercameronNotesCombinatorics2007"\ \ subsubsection \ Proof of the Corollary \ text \ This step makes use of the coefficient equivalence arguments proved in the previous sections \ corollary lucas_corollary: fixes n k :: nat assumes "prime p" shows "(n choose k) mod p = (((n div p) choose (k div p)) * ((n mod p) choose (k mod p))) mod p" (is "(n choose k) mod p = ((?N choose ?K) * (?rn choose ?rk)) mod p") proof - let ?fx = "fps_X :: int fps" have n_rep: "n = ?N * p + ?rn" by simp have k_rep: "k =?K * p + ?rk" by simp have rhs_coeffs: "((1 + ?fx^p)^(?N) * (1 + ?fx)^(?rn)) $ k = (?N choose ?K) * (?rn choose ?rk)" using assms fps_div_rep_coeffs k_rep n_rep by blast \ \Application of coefficient reasoning\ have "((((1 + ?fx)^p)^(?N) * (1 + ?fx)^(?rn)), ((1 + ?fx^p)^(?N) * (1 + ?fx)^(?rn))) \ fpsmodrel p" using fps_freshmans_dream assms fps_mult_equiv fps_power_equiv by blast \ \Application of equivalence facts and freshmans dream lemma\ then have modrel2: "((1 + ?fx)^n, ((1 + ?fx^p)^(?N) * (1 + ?fx)^(?rn))) \ fpsmodrel p" by (metis (mono_tags, opaque_lifting) mult_div_mod_eq power_add power_mult) thus ?thesis using fpsrel_iff binomial_coeffs_induct rhs_coeffs by (metis of_nat_eq_iff zmod_int) qed subsubsection \ Proof of the Theorem \ text \The theorem statement requires a formalised way of referring to the base $p$ representation of a number. We use a definition that specifies the $i$th digit of the base $p$ representation. This definition is originally from the Hilbert's 10th Problem Formalisation project \<^cite>\"bayerDPRMTheoremIsabelle2019"\ which this work contributes to.\ definition nth_digit_general :: "nat \ nat \ nat \ nat" where "nth_digit_general num i base = (num div (base ^ i)) mod base" text \Applying induction on $d$, where $d$ is the highest power required in either $n$ or $k$'s base $p$ representation, @{thm lucas_corollary} can be used to prove the original theorem.\ theorem lucas_theorem: fixes n k d::nat assumes "n < p ^ (Suc d)" assumes "k < p ^ (Suc d)" assumes "prime p" shows "(n choose k) mod p = (\i\d. ((nth_digit_general n i p) choose (nth_digit_general k i p))) mod p" using assms proof (induct d arbitrary: n k) case 0 thus ?case using nth_digit_general_def assms by simp next case (Suc d) \ \Representation Variables\ let ?N = "n div p" let ?K = "k div p" let ?nr = "n mod p" let ?kr = "k mod p" \ \Required assumption facts\ have Mlessthan: "?N < p ^ (Suc d)" using less_mult_imp_div_less power_Suc2 assms(3) prime_ge_2_nat Suc.prems(1) by metis have Nlessthan: "?K < p ^ (Suc d)" using less_mult_imp_div_less power_Suc2 prime_ge_2_nat Suc.prems(2) assms(3) by metis have shift_bounds_fact: "(\i=(Suc 0)..(Suc (d )). ((nth_digit_general n i p) choose (nth_digit_general k i p))) = (\i=0..(d). (nth_digit_general n (Suc i) p) choose (nth_digit_general k (Suc i) p))" using prod.shift_bounds_cl_Suc_ivl by blast \ \Product manipulation helper fact\ have "(n choose k ) mod p = ((?N choose ?K) * (?nr choose ?kr)) mod p" using lucas_corollary assms(3) by blast \ \Application of corollary\ also have "...= ((\i\d. ((nth_digit_general ?N i p) choose (nth_digit_general ?K i p))) * (?nr choose ?kr)) mod p" using Mlessthan Nlessthan Suc.hyps mod_mult_cong assms(3) by blast \ \Using Inductive Hypothesis\ \ \Product manipulation steps\ also have "... = ((\i=0..(d). (nth_digit_general n (Suc i) p) choose (nth_digit_general k (Suc i) p)) * (?nr choose ?kr)) mod p" using atMost_atLeast0 nth_digit_general_def div_mult2_eq by auto also have "... = ((\i=1..(d+1). (nth_digit_general n i p) choose (nth_digit_general k i p)) * ((nth_digit_general n 0 p) choose (nth_digit_general k 0 p))) mod p" using nth_digit_general_def shift_bounds_fact by simp finally have "(n choose k ) mod p = ((\i=0..(d+1). (nth_digit_general n i p) choose (nth_digit_general k i p))) mod p" using One_nat_def atMost_atLeast0 mult.commute prod.atLeast1_atMost_eq prod.atMost_shift - by (smt Suc_eq_plus1 shift_bounds_fact) + by (smt (verit, ccfv_threshold)) thus ?case using Suc_eq_plus1 atMost_atLeast0 by presburger qed end \ No newline at end of file diff --git a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy --- a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy +++ b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution_Prover.thy @@ -1,1360 +1,1360 @@ (* Title: An Ordered Resolution Prover for First-Order Clauses Author: Anders Schlichtkrull , 2016, 2017 Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2014 Maintainer: Anders Schlichtkrull *) section \An Ordered Resolution Prover for First-Order Clauses\ theory FO_Ordered_Resolution_Prover imports FO_Ordered_Resolution begin text \ This material is based on Section 4.3 (``A Simple Resolution Prover for First-Order Clauses'') of Bachmair and Ganzinger's chapter. Specifically, it formalizes the RP prover defined in Figure 5 and its related lemmas and theorems, including Lemmas 4.10 and 4.11 and Theorem 4.13 (completeness). \ definition is_least :: "(nat \ bool) \ nat \ bool" where "is_least P n \ P n \ (\n' < n. \ P n')" lemma least_exists: "P n \ \n. is_least P n" using exists_least_iff unfolding is_least_def by auto text \ The following corresponds to page 42 and 43 of Section 4.3, from the explanation of RP to Lemma 4.10. \ type_synonym 'a state = "'a clause set \ 'a clause set \ 'a clause set" locale FO_resolution_prover = FO_resolution subst_atm id_subst comp_subst renamings_apart atm_of_atms mgu less_atm + selection S for S :: "('a :: wellorder) clause \ 'a clause" and subst_atm :: "'a \ 's \ 'a" and id_subst :: "'s" and comp_subst :: "'s \ 's \ 's" and renamings_apart :: "'a clause list \ 's list" and atm_of_atms :: "'a list \ 'a" and mgu :: "'a set set \ 's option" and less_atm :: "'a \ 'a \ bool" + assumes sel_stable: "\\ C. is_renaming \ \ S (C \ \) = S C \ \" begin fun N_of_state :: "'a state \ 'a clause set" where "N_of_state (N, P, Q) = N" fun P_of_state :: "'a state \ 'a clause set" where "P_of_state (N, P, Q) = P" text \ \O\ denotes relation composition in Isabelle, so the formalization uses \Q\ instead. \ fun Q_of_state :: "'a state \ 'a clause set" where "Q_of_state (N, P, Q) = Q" abbreviation clss_of_state :: "'a state \ 'a clause set" where "clss_of_state St \ N_of_state St \ P_of_state St \ Q_of_state St" abbreviation grounding_of_state :: "'a state \ 'a clause set" where "grounding_of_state St \ grounding_of_clss (clss_of_state St)" interpretation ord_FO_resolution: inference_system "ord_FO_\ S" . text \ The following inductive predicate formalizes the resolution prover in Figure 5. \ inductive RP :: "'a state \ 'a state \ bool" (infix "\" 50) where tautology_deletion: "Neg A \# C \ Pos A \# C \ (N \ {C}, P, Q) \ (N, P, Q)" | forward_subsumption: "D \ P \ Q \ subsumes D C \ (N \ {C}, P, Q) \ (N, P, Q)" | backward_subsumption_P: "D \ N \ strictly_subsumes D C \ (N, P \ {C}, Q) \ (N, P, Q)" | backward_subsumption_Q: "D \ N \ strictly_subsumes D C \ (N, P, Q \ {C}) \ (N, P, Q)" | forward_reduction: "D + {#L'#} \ P \ Q \ - L = L' \l \ \ D \ \ \# C \ (N \ {C + {#L#}}, P, Q) \ (N \ {C}, P, Q)" | backward_reduction_P: "D + {#L'#} \ N \ - L = L' \l \ \ D \ \ \# C \ (N, P \ {C + {#L#}}, Q) \ (N, P \ {C}, Q)" | backward_reduction_Q: "D + {#L'#} \ N \ - L = L' \l \ \ D \ \ \# C \ (N, P, Q \ {C + {#L#}}) \ (N, P \ {C}, Q)" | clause_processing: "(N \ {C}, P, Q) \ (N, P \ {C}, Q)" | inference_computation: "N = concls_of (ord_FO_resolution.inferences_between Q C) \ ({}, P \ {C}, Q) \ (N, P, Q \ {C})" lemma final_RP: "\ ({}, {}, Q) \ St" by (auto elim: RP.cases) definition Sup_state :: "'a state llist \ 'a state" where "Sup_state Sts = (Sup_llist (lmap N_of_state Sts), Sup_llist (lmap P_of_state Sts), Sup_llist (lmap Q_of_state Sts))" definition Liminf_state :: "'a state llist \ 'a state" where "Liminf_state Sts = (Liminf_llist (lmap N_of_state Sts), Liminf_llist (lmap P_of_state Sts), Liminf_llist (lmap Q_of_state Sts))" context fixes Sts Sts' :: "'a state llist" assumes Sts: "lfinite Sts" "lfinite Sts'" "\ lnull Sts" "\ lnull Sts'" "llast Sts' = llast Sts" begin lemma N_of_Liminf_state_fin: "N_of_state (Liminf_state Sts') = N_of_state (Liminf_state Sts)" and P_of_Liminf_state_fin: "P_of_state (Liminf_state Sts') = P_of_state (Liminf_state Sts)" and Q_of_Liminf_state_fin: "Q_of_state (Liminf_state Sts') = Q_of_state (Liminf_state Sts)" using Sts by (simp_all add: Liminf_state_def lfinite_Liminf_llist llast_lmap) lemma Liminf_state_fin: "Liminf_state Sts' = Liminf_state Sts" using N_of_Liminf_state_fin P_of_Liminf_state_fin Q_of_Liminf_state_fin by (simp add: Liminf_state_def) end context fixes Sts Sts' :: "'a state llist" assumes Sts: "\ lfinite Sts" "emb Sts Sts'" begin lemma N_of_Liminf_state_inf: "N_of_state (Liminf_state Sts') \ N_of_state (Liminf_state Sts)" and P_of_Liminf_state_inf: "P_of_state (Liminf_state Sts') \ P_of_state (Liminf_state Sts)" and Q_of_Liminf_state_inf: "Q_of_state (Liminf_state Sts') \ Q_of_state (Liminf_state Sts)" using Sts by (simp_all add: Liminf_state_def emb_Liminf_llist_infinite emb_lmap) lemma clss_of_Liminf_state_inf: "clss_of_state (Liminf_state Sts') \ clss_of_state (Liminf_state Sts)" using N_of_Liminf_state_inf P_of_Liminf_state_inf Q_of_Liminf_state_inf by blast end definition fair_state_seq :: "'a state llist \ bool" where "fair_state_seq Sts \ N_of_state (Liminf_state Sts) = {} \ P_of_state (Liminf_state Sts) = {}" text \ The following formalizes Lemma 4.10. \ context fixes Sts :: "'a state llist" begin definition S_Q :: "'a clause \ 'a clause" where "S_Q = S_M S (Q_of_state (Liminf_state Sts))" interpretation sq: selection S_Q unfolding S_Q_def using S_M_selects_subseteq S_M_selects_neg_lits selection_axioms by unfold_locales auto interpretation gr: ground_resolution_with_selection S_Q by unfold_locales interpretation sr: standard_redundancy_criterion_reductive gr.ord_\ by unfold_locales interpretation sr: standard_redundancy_criterion_counterex_reducing gr.ord_\ "ground_resolution_with_selection.INTERP S_Q" by unfold_locales text \ The extension of ordered resolution mentioned in 4.10. We let it consist of all sound rules. \ definition ground_sound_\:: "'a inference set" where "ground_sound_\ = {Infer CC D E | CC D E. (\I. I \m CC \ I \ D \ I \ E)}" text \ We prove that we indeed defined an extension. \ lemma gd_ord_\_ngd_ord_\: "gr.ord_\ \ ground_sound_\" unfolding ground_sound_\_def using gr.ord_\_def gr.ord_resolve_sound by fastforce lemma sound_ground_sound_\: "sound_inference_system ground_sound_\" unfolding sound_inference_system_def ground_sound_\_def by auto lemma sat_preserving_ground_sound_\: "sat_preserving_inference_system ground_sound_\" using sound_ground_sound_\ sat_preserving_inference_system.intro sound_inference_system.\_sat_preserving by blast definition sr_ext_Ri :: "'a clause set \ 'a inference set" where "sr_ext_Ri N = sr.Ri N \ (ground_sound_\ - gr.ord_\)" interpretation sr_ext: sat_preserving_redundancy_criterion ground_sound_\ sr.Rf sr_ext_Ri unfolding sat_preserving_redundancy_criterion_def sr_ext_Ri_def using sat_preserving_ground_sound_\ redundancy_criterion_standard_extension gd_ord_\_ngd_ord_\ sr.redundancy_criterion_axioms by auto lemma strict_subset_subsumption_redundant_clause: assumes sub: "D \ \ \# C" and ground_\: "is_ground_subst \" shows "C \ sr.Rf (grounding_of_cls D)" proof - from sub have "\I. I \ D \ \ \ I \ C" unfolding true_cls_def by blast moreover have "C > D \ \" using sub by (simp add: subset_imp_less_mset) moreover have "D \ \ \ grounding_of_cls D" using ground_\ by (metis (mono_tags) mem_Collect_eq substitution_ops.grounding_of_cls_def) ultimately have "set_mset {#D \ \#} \ grounding_of_cls D" "(\I. I \m {#D \ \#} \ I \ C)" "(\D'. D' \# {#D \ \#} \ D' < C)" by auto then show ?thesis using sr.Rf_def by blast qed lemma strict_subset_subsumption_redundant_clss: assumes "D \ \ \# C" and "is_ground_subst \" and "D \ CC" shows "C \ sr.Rf (grounding_of_clss CC)" using assms proof - have "C \ sr.Rf (grounding_of_cls D)" using strict_subset_subsumption_redundant_clause assms by auto then show ?thesis using assms unfolding grounding_of_clss_def by (metis (no_types) sr.Rf_mono sup_ge1 SUP_absorb contra_subsetD) qed lemma strict_subset_subsumption_grounding_redundant_clss: assumes D\_subset_C: "D \ \ \# C" and D_in_St: "D \ CC" shows "grounding_of_cls C \ sr.Rf (grounding_of_clss CC)" proof fix C\ assume "C\ \ grounding_of_cls C" then obtain \ where \_p: "C\ = C \ \ \ is_ground_subst \" unfolding grounding_of_cls_def by auto have D\\C\: "D \ \ \ \ \# C \ \" using D\_subset_C subst_subset_mono by auto then show "C\ \ sr.Rf (grounding_of_clss CC)" using \_p strict_subset_subsumption_redundant_clss[of D "\ \ \" "C \ \"] D_in_St by auto qed lemma derive_if_remove_subsumed: assumes "D \ clss_of_state St" and "subsumes D C" shows "sr_ext.derive (grounding_of_state St \ grounding_of_cls C) (grounding_of_state St)" proof - from assms obtain \ where "D \ \ = C \ D \ \ \# C" by (auto simp: subsumes_def subset_mset_def) then have "D \ \ = C \ D \ \ \# C" by (simp add: subset_mset_def) then show ?thesis proof assume "D \ \ = C" then have "grounding_of_cls C \ grounding_of_cls D" using subst_cls_eq_grounding_of_cls_subset_eq by (auto dest: sym) then have "(grounding_of_state St \ grounding_of_cls C) = grounding_of_state St" using assms unfolding grounding_of_clss_def by auto then show ?thesis by (auto intro: sr_ext.derive.intros) next assume a: "D \ \ \# C" then have "grounding_of_cls C \ sr.Rf (grounding_of_state St)" using strict_subset_subsumption_grounding_redundant_clss assms by auto then show ?thesis unfolding grounding_of_clss_def by (force intro: sr_ext.derive.intros) qed qed lemma reduction_in_concls_of: assumes "C\ \ grounding_of_cls C" and "D + {#L'#} \ CC" and "- L = L' \l \" and "D \ \ \# C" shows "C\ \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" proof - from assms obtain \ where \_p: "C\ = C \ \ \ is_ground_subst \" unfolding grounding_of_cls_def by auto define \ where "\ = Infer {#(C + {#L#}) \ \#} ((D + {#L'#}) \ \ \ \) (C \ \)" have "(D + {#L'#}) \ \ \ \ \ grounding_of_clss (CC \ {C + {#L#}})" unfolding grounding_of_clss_def grounding_of_cls_def by (rule UN_I[of "D + {#L'#}"], use assms(2) in simp, metis (mono_tags, lifting) \_p is_ground_comp_subst mem_Collect_eq subst_cls_comp_subst) moreover have "(C + {#L#}) \ \ \ grounding_of_clss (CC \ {C + {#L#}})" using \_p unfolding grounding_of_clss_def grounding_of_cls_def by auto moreover have "\I. I \ D \ \ \ \ + {#- (L \l \)#} \ I \ C \ \ + {#L \l \#} \ I \ D \ \ \ \ + C \ \" by auto then have "\I. I \ (D + {#L'#}) \ \ \ \ \ I \ (C + {#L#}) \ \ \ I \ D \ \ \ \ + C \ \" using assms by (metis add_mset_add_single subst_cls_add_mset subst_cls_union subst_minus) then have "\I. I \ (D + {#L'#}) \ \ \ \ \ I \ (C + {#L#}) \ \ \ I \ C \ \" using assms by (metis (no_types, lifting) subset_mset.le_iff_add subst_cls_union true_cls_union) then have "\I. I \m {#(D + {#L'#}) \ \ \ \#} \ I \ (C + {#L#}) \ \ \ I \ C \ \" by (meson true_cls_mset_singleton) ultimately have "\ \ sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}}))" unfolding sr_ext.inferences_from_def unfolding ground_sound_\_def infer_from_def \_def by auto then have "C \ \ \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" using image_iff unfolding \_def by fastforce then show "C\ \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" using \_p by auto qed lemma reduction_derivable: assumes "D + {#L'#} \ CC" and "- L = L' \l \" and "D \ \ \# C" shows "sr_ext.derive (grounding_of_clss (CC \ {C + {#L#}})) (grounding_of_clss (CC \ {C}))" proof - from assms have "grounding_of_clss (CC \ {C}) - grounding_of_clss (CC \ {C + {#L#}}) \ concls_of (sr_ext.inferences_from (grounding_of_clss (CC \ {C + {#L#}})))" using reduction_in_concls_of unfolding grounding_of_clss_def by auto moreover have "grounding_of_cls (C + {#L#}) \ sr.Rf (grounding_of_clss (CC \ {C}))" using strict_subset_subsumption_grounding_redundant_clss[of C "id_subst"] by auto then have "grounding_of_clss (CC \ {C + {#L#}}) - grounding_of_clss (CC \ {C}) \ sr.Rf (grounding_of_clss (CC \ {C}))" unfolding grounding_of_clss_def by auto ultimately show "sr_ext.derive (grounding_of_clss (CC \ {C + {#L#}})) (grounding_of_clss (CC \ {C}))" using sr_ext.derive.intros[of "grounding_of_clss (CC \ {C})" "grounding_of_clss (CC \ {C + {#L#}})"] by auto qed text \ The following corresponds the part of Lemma 4.10 that states we have a theorem proving process: \ lemma RP_ground_derive: "St \ St' \ sr_ext.derive (grounding_of_state St) (grounding_of_state St')" proof (induction rule: RP.induct) case (tautology_deletion A C N P Q) { fix C\ assume "C\ \ grounding_of_cls C" then obtain \ where "C\ = C \ \" unfolding grounding_of_cls_def by auto then have "Neg (A \a \) \# C\ \ Pos (A \a \) \# C\" using tautology_deletion Neg_Melem_subst_atm_subst_cls Pos_Melem_subst_atm_subst_cls by auto then have "C\ \ sr.Rf (grounding_of_state (N, P, Q))" using sr.tautology_Rf by auto } then have "grounding_of_state (N \ {C}, P, Q) - grounding_of_state (N, P, Q) \ sr.Rf (grounding_of_state (N, P, Q))" unfolding grounding_of_clss_def by auto moreover have "grounding_of_state (N, P, Q) - grounding_of_state (N \ {C}, P, Q) = {}" unfolding grounding_of_clss_def by auto ultimately show ?case using sr_ext.derive.intros[of "grounding_of_state (N, P, Q)" "grounding_of_state (N \ {C}, P, Q)"] by auto next case (forward_subsumption D P Q C N) then show ?case using derive_if_remove_subsumed[of D "(N, P, Q)" C] unfolding grounding_of_clss_def by (simp add: sup_commute sup_left_commute) next case (backward_subsumption_P D N C P Q) then show ?case using derive_if_remove_subsumed[of D "(N, P, Q)" C] strictly_subsumes_def unfolding grounding_of_clss_def by (simp add: sup_commute sup_left_commute) next case (backward_subsumption_Q D N C P Q) then show ?case using derive_if_remove_subsumed[of D "(N, P, Q)" C] strictly_subsumes_def unfolding grounding_of_clss_def by (simp add: sup_commute sup_left_commute) next case (forward_reduction D L' P Q L \ C N) then show ?case using reduction_derivable[of _ _ "N \ P \ Q"] by force next case (backward_reduction_P D L' N L \ C P Q) then show ?case using reduction_derivable[of _ _ "N \ P \ Q"] by force next case (backward_reduction_Q D L' N L \ C P Q) then show ?case using reduction_derivable[of _ _ "N \ P \ Q"] by force next case (clause_processing N C P Q) then show ?case using sr_ext.derive.intros by auto next case (inference_computation N Q C P) { fix E\ assume "E\ \ grounding_of_clss N" then obtain \ E where E_\_p: "E\ = E \ \ \ E \ N \ is_ground_subst \" unfolding grounding_of_clss_def grounding_of_cls_def by auto then have E_concl: "E \ concls_of (ord_FO_resolution.inferences_between Q C)" using inference_computation by auto then obtain \ where \_p: "\ \ ord_FO_\ S \ infer_from (Q \ {C}) \ \ C \# prems_of \ \ concl_of \ = E" unfolding ord_FO_resolution.inferences_between_def by auto then obtain CC CAs D AAs As \ where \_p2: "\ = Infer CC D E \ ord_resolve_rename S CAs D AAs As \ E \ mset CAs = CC" unfolding ord_FO_\_def by auto define \ where "\ = hd (renamings_apart (D # CAs))" define \s where "\s = tl (renamings_apart (D # CAs))" define \_ground where "\_ground = Infer (mset (CAs \\cl \s) \cm \ \cm \) (D \ \ \ \ \ \) (E \ \)" have "\I. I \m mset (CAs \\cl \s) \cm \ \cm \ \ I \ D \ \ \ \ \ \ \ I \ E \ \" using ord_resolve_rename_ground_inst_sound[of _ _ _ _ _ _ _ _ _ _ \] \_def \s_def E_\_p \_p2 by auto then have "\_ground \ {Infer cc d e | cc d e. \I. I \m cc \ I \ d \ I \ e}" unfolding \_ground_def by auto moreover have "set_mset (prems_of \_ground) \ grounding_of_state ({}, P \ {C}, Q)" proof - have "D = C \ D \ Q" unfolding \_ground_def using E_\_p \_p2 \_p unfolding infer_from_def unfolding grounding_of_clss_def grounding_of_cls_def by simp then have "D \ \ \ \ \ \ \ grounding_of_cls C \ (\x \ Q. D \ \ \ \ \ \ \ grounding_of_cls x)" using E_\_p unfolding grounding_of_cls_def by (metis (mono_tags, lifting) is_ground_comp_subst mem_Collect_eq subst_cls_comp_subst) then have "(D \ \ \ \ \ \ \ grounding_of_cls C \ (\x \ P. D \ \ \ \ \ \ \ grounding_of_cls x) \ (\x \ Q. D \ \ \ \ \ \ \ grounding_of_cls x))" by metis moreover have "\i < length (CAs \\cl \s \cl \ \cl \). (CAs \\cl \s \cl \ \cl \) ! i \ {C \ \ |\. is_ground_subst \} \ ((\C \ P. {C \ \ | \. is_ground_subst \}) \ (\C\Q. {C \ \ | \. is_ground_subst \}))" proof (rule, rule) fix i assume "i < length (CAs \\cl \s \cl \ \cl \)" then have a: "i < length CAs \ i < length \s" by simp moreover from a have "CAs ! i \ {C} \ Q" using \_p2 \_p unfolding infer_from_def by (metis (no_types, lifting) Un_subset_iff inference.sel(1) set_mset_union sup_commute nth_mem_mset subsetCE) ultimately have "(CAs \\cl \s \cl \ \cl \) ! i \ {C \ \ |\. is_ground_subst \} \ ((CAs \\cl \s \cl \ \cl \) ! i \ (\C\P. {C \ \ |\. is_ground_subst \}) \ (CAs \\cl \s \cl \ \cl \) ! i \ (\C \ Q. {C \ \ | \. is_ground_subst \}))" using E_\_p \_p2 \_p unfolding \_ground_def infer_from_def grounding_of_clss_def grounding_of_cls_def apply - apply (cases "CAs ! i = C") subgoal apply (rule disjI1) apply (rule Set.CollectI) apply (rule_tac x = "(\s ! i) \ \ \ \" in exI) using \s_def using renamings_apart_length by (auto; fail) subgoal apply (rule disjI2) apply (rule disjI2) apply (rule_tac a = "CAs ! i" in UN_I) subgoal by blast subgoal apply (rule Set.CollectI) apply (rule_tac x = "(\s ! i) \ \ \ \" in exI) using \s_def using renamings_apart_length by (auto; fail) done done then show "(CAs \\cl \s \cl \ \cl \) ! i \ {C \ \ |\. is_ground_subst \} \ ((\C \ P. {C \ \ |\. is_ground_subst \}) \ (\C \ Q. {C \ \ |\. is_ground_subst \}))" by blast qed then have "\x \# mset (CAs \\cl \s \cl \ \cl \). x \ {C \ \ |\. is_ground_subst \} \ ((\C \ P. {C \ \ |\. is_ground_subst \}) \ (\C \ Q. {C \ \ |\. is_ground_subst \}))" by (metis (lifting) in_set_conv_nth set_mset_mset) then have "set_mset (mset (CAs \\cl \s) \cm \ \cm \) \ grounding_of_cls C \ grounding_of_clss P \ grounding_of_clss Q" unfolding grounding_of_cls_def grounding_of_clss_def using mset_subst_cls_list_subst_cls_mset by auto ultimately show ?thesis unfolding \_ground_def grounding_of_clss_def by auto qed ultimately have "E \ \ \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" unfolding sr_ext.inferences_from_def inference_system.inferences_from_def ground_sound_\_def infer_from_def using \_ground_def by (metis (mono_tags, lifting) image_eqI inference.sel(3) mem_Collect_eq) then have "E\ \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" using E_\_p by auto } then have "grounding_of_state (N, P, Q \ {C}) - grounding_of_state ({}, P \ {C}, Q) \ concls_of (sr_ext.inferences_from (grounding_of_state ({}, P \ {C}, Q)))" unfolding grounding_of_clss_def by auto moreover have "grounding_of_state ({}, P \ {C}, Q) - grounding_of_state (N, P, Q \ {C}) = {}" unfolding grounding_of_clss_def by auto ultimately show ?case using sr_ext.derive.intros[of "(grounding_of_state (N, P, Q \ {C}))" "(grounding_of_state ({}, P \ {C}, Q))"] by auto qed text \ A useful consequence: \ theorem RP_model: "St \ St' \ I \s grounding_of_state St' \ I \s grounding_of_state St" proof (drule RP_ground_derive, erule sr_ext.derive.cases, hypsubst) let ?gSt = "grounding_of_state St" and ?gSt' = "grounding_of_state St'" assume deduct: "?gSt' - ?gSt \ concls_of (sr_ext.inferences_from ?gSt)" (is "_ \ ?concls") and delete: "?gSt - ?gSt' \ sr.Rf ?gSt'" show "I \s ?gSt' \ I \s ?gSt" proof assume bef: "I \s ?gSt" then have "I \s ?concls" unfolding ground_sound_\_def inference_system.inferences_from_def true_clss_def true_cls_mset_def by (auto simp add: image_def infer_from_def dest!: spec[of _ I]) then have diff: "I \s ?gSt' - ?gSt" using deduct by (blast intro: true_clss_mono) then show "I \s ?gSt'" using bef unfolding true_clss_def by blast next assume aft: "I \s ?gSt'" have "I \s ?gSt' \ sr.Rf ?gSt'" - by (rule sr.Rf_model) (smt Diff_eq_empty_iff Diff_subset Un_Diff aft + by (rule sr.Rf_model) (smt (verit) Diff_eq_empty_iff Diff_subset Un_Diff aft standard_redundancy_criterion.Rf_mono sup_bot.right_neutral sup_ge1 true_clss_mono) then have "I \s sr.Rf ?gSt'" using true_clss_union by blast then have diff: "I \s ?gSt - ?gSt'" using delete by (blast intro: true_clss_mono) then show "I \s ?gSt" using aft unfolding true_clss_def by blast qed qed text \ Another formulation of the part of Lemma 4.10 that states we have a theorem proving process: \ lemma ground_derive_chain: "chain (\) Sts \ chain sr_ext.derive (lmap grounding_of_state Sts)" using RP_ground_derive by (simp add: chain_lmap[of "(\)"]) text \ The following is used prove to Lemma 4.11: \ lemma Sup_llist_grounding_of_state_ground: assumes "C \ Sup_llist (lmap grounding_of_state Sts)" shows "is_ground_cls C" proof - have "\j. enat j < llength (lmap grounding_of_state Sts) \ C \ lnth (lmap grounding_of_state Sts) j" using assms Sup_llist_imp_exists_index by fast then show ?thesis unfolding grounding_of_clss_def grounding_of_cls_def by auto qed lemma Liminf_grounding_of_state_ground: "C \ Liminf_llist (lmap grounding_of_state Sts) \ is_ground_cls C" using Liminf_llist_subset_Sup_llist[of "lmap grounding_of_state Sts"] Sup_llist_grounding_of_state_ground by blast lemma in_Sup_llist_in_Sup_state: assumes "C \ Sup_llist (lmap grounding_of_state Sts)" shows "\D \. D \ clss_of_state (Sup_state Sts) \ D \ \ = C \ is_ground_subst \" proof - from assms obtain i where i_p: "enat i < llength Sts \ C \ lnth (lmap grounding_of_state Sts) i" using Sup_llist_imp_exists_index by fastforce then obtain D \ where "D \ clss_of_state (lnth Sts i) \ D \ \ = C \ is_ground_subst \" using assms unfolding grounding_of_clss_def grounding_of_cls_def by fastforce then have "D \ clss_of_state (Sup_state Sts) \ D \ \ = C \ is_ground_subst \" using i_p unfolding Sup_state_def by (metis (no_types, lifting) UnCI UnE contra_subsetD N_of_state.simps P_of_state.simps Q_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist) then show ?thesis by auto qed lemma N_of_state_Liminf: "N_of_state (Liminf_state Sts) = Liminf_llist (lmap N_of_state Sts)" and P_of_state_Liminf: "P_of_state (Liminf_state Sts) = Liminf_llist (lmap P_of_state Sts)" unfolding Liminf_state_def by auto lemma eventually_removed_from_N: assumes d_in: "D \ N_of_state (lnth Sts i)" and fair: "fair_state_seq Sts" and i_Sts: "enat i < llength Sts" shows "\l. D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" proof (rule ccontr) assume a: "\ ?thesis" have "i \ l \ enat l < llength Sts \ D \ N_of_state (lnth Sts l)" for l using d_in by (induction l, blast, metis a Suc_ile_eq le_SucE less_imp_le) then have "D \ Liminf_llist (lmap N_of_state Sts)" unfolding Liminf_llist_def using i_Sts by auto then show False using fair unfolding fair_state_seq_def by (simp add: N_of_state_Liminf) qed lemma eventually_removed_from_P: assumes d_in: "D \ P_of_state (lnth Sts i)" and fair: "fair_state_seq Sts" and i_Sts: "enat i < llength Sts" shows "\l. D \ P_of_state (lnth Sts l) \ D \ P_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" proof (rule ccontr) assume a: "\ ?thesis" have "i \ l \ enat l < llength Sts \ D \ P_of_state (lnth Sts l)" for l using d_in by (induction l, blast, metis a Suc_ile_eq le_SucE less_imp_le) then have "D \ Liminf_llist (lmap P_of_state Sts)" unfolding Liminf_llist_def using i_Sts by auto then show False using fair unfolding fair_state_seq_def by (simp add: P_of_state_Liminf) qed lemma instance_if_subsumed_and_in_limit: assumes deriv: "chain (\) Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ clss_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" shows "\\. D \ \ = C \ is_ground_subst \" proof - let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" using ground_derive_chain deriv ns by auto have "\\. D \ \ = C" proof (rule ccontr) assume "\\. D \ \ = C" moreover from d(3) obtain \_proto where "D \ \_proto \# C" unfolding subsumes_def by blast then obtain \ where \_p: "D \ \ \# C \ is_ground_subst \" using ground_C by (metis is_ground_cls_mono make_ground_subst subset_mset.order_refl) ultimately have subsub: "D \ \ \# C" using subset_mset.le_imp_less_or_eq by auto moreover have "is_ground_subst \" using \_p by auto moreover have "D \ clss_of_state (lnth Sts i)" using d by auto ultimately have "C \ sr.Rf (grounding_of_state (lnth Sts i))" using strict_subset_subsumption_redundant_clss by auto then have "C \ sr.Rf (Sup_llist Gs)" - using d ns by (smt contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr.Rf_mono) + using d ns by (smt (verit) contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr.Rf_mono) then have "C \ sr.Rf (Liminf_llist Gs)" unfolding ns using local.sr_ext.Rf_limit_Sup derivns ns by auto then show False using c by auto qed then obtain \ where "D \ \ = C \ is_ground_subst \" using ground_C by (metis make_ground_subst) then show ?thesis by auto qed lemma from_Q_to_Q_inf: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ Q_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" and d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" shows "D \ Q_of_state (Liminf_state Sts)" proof - let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit[OF deriv] c d unfolding ns by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by auto have in_Sts_in_Sts_Suc: "\l \ i. enat (Suc l) < llength Sts \ D \ Q_of_state (lnth Sts l) \ D \ Q_of_state (lnth Sts (Suc l))" proof (rule, rule, rule, rule) fix l assume len: "i \ l" and llen: "enat (Suc l) < llength Sts" and d_in_q: "D \ Q_of_state (lnth Sts l)" have "lnth Sts l \ lnth Sts (Suc l)" using llen deriv chain_lnth_rel by blast then show "D \ Q_of_state (lnth Sts (Suc l))" proof (cases rule: RP.cases) case (backward_subsumption_Q D' N D_removed P Q) moreover { assume "D_removed = D" then obtain D_subsumes where D_subsumes_p: "D_subsumes \ N \ strictly_subsumes D_subsumes D" using backward_subsumption_Q by auto moreover from D_subsumes_p have "subsumes D_subsumes C" using d subsumes_trans unfolding strictly_subsumes_def by blast moreover from backward_subsumption_Q have "D_subsumes \ clss_of_state (Sup_state Sts)" using D_subsumes_p llen by (metis (no_types) UnI1 N_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist rev_subsetD Sup_state_def) ultimately have False using d_least unfolding subsumes_def by auto } ultimately show ?thesis using d_in_q by auto next case (backward_reduction_Q E L' N L \ D' P Q) { assume "D' + {#L#} = D" then have D'_p: "strictly_subsumes D' D \ D' \ ?Ps (Suc l)" using subset_strictly_subsumes[of D' D] backward_reduction_Q by auto then have subc: "subsumes D' C" using d(3) subsumes_trans unfolding strictly_subsumes_def by auto from D'_p have "D' \ clss_of_state (Sup_state Sts)" using llen by (metis (no_types) UnI1 P_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist subsetCE sup_ge2 Sup_state_def) then have False using d_least D'_p subc by auto } then show ?thesis using backward_reduction_Q d_in_q by auto qed (use d_in_q in auto) qed have D_in_Sts: "D \ Q_of_state (lnth Sts l)" and D_in_Sts_Suc: "D \ Q_of_state (lnth Sts (Suc l))" if l_i: "l \ i" and enat: "enat (Suc l) < llength Sts" for l proof - show "D \ Q_of_state (lnth Sts l)" using l_i enat apply (induction "l - i" arbitrary: l) subgoal using d by auto subgoal using d(1) in_Sts_in_Sts_Suc by (metis (no_types, lifting) Suc_ile_eq add_Suc_right add_diff_cancel_left' le_SucE le_Suc_ex less_imp_le) done then show "D \ Q_of_state (lnth Sts (Suc l))" using l_i enat in_Sts_in_Sts_Suc by blast qed have "i \ x \ enat x < llength Sts \ D \ Q_of_state (lnth Sts x)" for x apply (cases x) subgoal using d(1) by (auto intro!: exI[of _ i] simp: less_Suc_eq) subgoal for x' using d(1) D_in_Sts_Suc[of x'] by (cases \i \ x'\) (auto simp: not_less_eq_eq) done then have "D \ Liminf_llist (lmap Q_of_state Sts)" unfolding Liminf_llist_def by (auto intro!: exI[of _ i] simp: d) then show ?thesis unfolding Liminf_state_def by auto qed lemma from_P_to_Q: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ P_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" and d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" shows "\l. D \ Q_of_state (lnth Sts l) \ enat l < llength Sts" proof - let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit[OF deriv] ns c d by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by auto obtain l where l_p: "D \ P_of_state (lnth Sts l) \ D \ P_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" using fair using eventually_removed_from_P d unfolding ns by auto then have l_Gs: "enat (Suc l) < llength Gs" using ns by auto from l_p have "lnth Sts l \ lnth Sts (Suc l)" using deriv using chain_lnth_rel by auto then show ?thesis proof (cases rule: RP.cases) case (backward_subsumption_P D' N D_twin P Q) note lrhs = this(1,2) and D'_p = this(3,4) then have twins: "D_twin = D" "?Ns (Suc l) = N" "?Ns l = N" "?Ps (Suc l) = P" "?Ps l = P \ {D_twin}" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto note D'_p = D'_p[unfolded twins(1)] then have subc: "subsumes D' C" unfolding strictly_subsumes_def subsumes_def using \ by (metis subst_cls_comp_subst subst_cls_mono_mset) from D'_p have "D' \ clss_of_state (Sup_state Sts)" unfolding twins(2)[symmetric] using l_p by (metis (no_types) UnI1 N_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist subsetCE Sup_state_def) then have False using d_least D'_p subc by auto then show ?thesis by auto next case (backward_reduction_P E L' N L \ D' P Q) then have twins: "D' + {#L#} = D" "?Ns (Suc l) = N" "?Ns l = N" "?Ps (Suc l) = P \ {D'}" "?Ps l = P \ {D' + {#L#}}" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto then have D'_p: "strictly_subsumes D' D \ D' \ ?Ps (Suc l)" using subset_strictly_subsumes[of D' D] by auto then have subc: "subsumes D' C" using d(3) subsumes_trans unfolding strictly_subsumes_def by auto from D'_p have "D' \ clss_of_state (Sup_state Sts)" using l_p by (metis (no_types) UnI1 P_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist subsetCE sup_ge2 Sup_state_def) then have False using d_least D'_p subc by auto then show ?thesis by auto next case (inference_computation N Q D_twin P) then have twins: "D_twin = D" "?Ps (Suc l) = P" "?Ps l = P \ {D_twin}" "?Qs (Suc l) = Q \ {D_twin}" "?Qs l = Q" using l_p by auto then show ?thesis using d \ l_p by auto qed (use l_p in auto) qed lemma from_N_to_P_or_Q: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and d: "D \ N_of_state (lnth Sts i)" "enat i < llength Sts" "subsumes D C" and d_least: "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" shows "\l D' \'. D' \ P_of_state (lnth Sts l) \ Q_of_state (lnth Sts l) \ enat l < llength Sts \ (\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D') \ D' \ \' = C \ is_ground_subst \' \ subsumes D' C" proof - let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_C: "is_ground_cls C" using c using Liminf_grounding_of_state_ground ns by auto have derivns: "chain sr_ext.derive Gs" using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit[OF deriv] ns c d by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by auto from c have no_taut: "\ (\A. Pos A \# C \ Neg A \# C)" using sr.tautology_Rf by auto have "\l. D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" using fair using eventually_removed_from_N d unfolding ns by auto then obtain l where l_p: "D \ N_of_state (lnth Sts l) \ D \ N_of_state (lnth Sts (Suc l)) \ i \ l \ enat (Suc l) < llength Sts" by auto then have l_Gs: "enat (Suc l) < llength Gs" using ns by auto from l_p have "lnth Sts l \ lnth Sts (Suc l)" using deriv using chain_lnth_rel by auto then show ?thesis proof (cases rule: RP.cases) case (tautology_deletion A D_twin N P Q) then have "D_twin = D" using l_p by auto then have "Pos (A \a \) \# C \ Neg (A \a \) \# C" using tautology_deletion(3,4) \ by (metis Melem_subst_cls eql_neg_lit_eql_atm eql_pos_lit_eql_atm) then have False using no_taut by metis then show ?thesis by blast next case (forward_subsumption D' P Q D_twin N) note lrhs = this(1,2) and D'_p = this(3,4) then have twins: "D_twin = D" "?Ns (Suc l) = N" "?Ns l = N \ {D_twin}" "?Ps (Suc l) = P " "?Ps l = P" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto note D'_p = D'_p[unfolded twins(1)] from D'_p(2) have subs: "subsumes D' C" using d(3) by (blast intro: subsumes_trans) moreover have "D' \ clss_of_state (Sup_state Sts)" using twins D'_p l_p unfolding Sup_state_def by simp (metis (no_types) contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist) ultimately have "\ strictly_subsumes D' D" using d_least by auto then have "subsumes D D'" unfolding strictly_subsumes_def using D'_p by auto then have v: "variants D D'" using D'_p unfolding variants_iff_subsumes by auto then have mini: "\E \ {E \ clss_of_state (Sup_state Sts). subsumes E C}. \ strictly_subsumes E D'" using d_least D'_p neg_strictly_subsumes_variants[of _ D D'] by auto from v have "\\'. D' \ \' = C" using \ variants_imp_exists_substitution variants_sym by (metis subst_cls_comp_subst) then have "\\'. D' \ \' = C \ is_ground_subst \'" using ground_C by (meson make_ground_subst refl) then obtain \' where \'_p: "D' \ \' = C \ is_ground_subst \'" by metis show ?thesis using D'_p twins l_p subs mini \'_p by auto next case (forward_reduction E L' P Q L \ D' N) then have twins: "D' + {#L#} = D" "?Ns (Suc l) = N \ {D'}" "?Ns l = N \ {D' + {#L#}}" "?Ps (Suc l) = P " "?Ps l = P" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto then have D'_p: "strictly_subsumes D' D \ D' \ ?Ns (Suc l)" using subset_strictly_subsumes[of D' D] by auto then have subc: "subsumes D' C" using d(3) subsumes_trans unfolding strictly_subsumes_def by blast from D'_p have "D' \ clss_of_state (Sup_state Sts)" using l_p by (metis (no_types) UnI1 N_of_state.simps llength_lmap lnth_lmap lnth_subset_Sup_llist subsetCE Sup_state_def) then have False using d_least D'_p subc by auto then show ?thesis by auto next case (clause_processing N D_twin P Q) then have twins: "D_twin = D" "?Ns (Suc l) = N" "?Ns l = N \ {D}" "?Ps (Suc l) = P \ {D}" "?Ps l = P" "?Qs (Suc l) = Q" "?Qs l = Q" using l_p by auto then show ?thesis using d \ l_p d_least by blast qed (use l_p in auto) qed lemma eventually_in_Qinf: assumes deriv: "chain (\) Sts" and D_p: "D \ clss_of_state (Sup_state Sts)" "subsumes D C" "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D" and fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" and c: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" and ground_C: "is_ground_cls C" shows "\D' \'. D' \ Q_of_state (Liminf_state Sts) \ D' \ \' = C \ is_ground_subst \'" proof - let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" from D_p obtain i where i_p: "i < llength Sts" "D \ ?Ns i \ D \ ?Ps i \ D \ ?Qs i" unfolding Sup_state_def by simp_all (metis (no_types) Sup_llist_imp_exists_index llength_lmap lnth_lmap) have derivns: "chain sr_ext.derive Gs" using ground_derive_chain deriv ns by auto have "\\. D \ \ = C \ is_ground_subst \" using instance_if_subsumed_and_in_limit[OF deriv ns c] D_p i_p by blast then obtain \ where \: "D \ \ = C" "is_ground_subst \" by blast { assume a: "D \ ?Ns i" then obtain D' \' l where D'_p: "D' \ ?Ps l \ ?Qs l" "D' \ \' = C" "enat l < llength Sts" "is_ground_subst \'" "\E \ {E. E \ (clss_of_state (Sup_state Sts)) \ subsumes E C}. \ strictly_subsumes E D'" "subsumes D' C" using from_N_to_P_or_Q deriv fair ns c i_p(1) D_p(2) D_p(3) by blast then obtain l' where l'_p: "D' \ ?Qs l'" "l' < llength Sts" using from_P_to_Q[OF deriv fair ns c _ D'_p(3) D'_p(6) D'_p(5)] by blast then have "D' \ Q_of_state (Liminf_state Sts)" using from_Q_to_Q_inf[OF deriv fair ns c _ l'_p(2)] D'_p by auto then have ?thesis using D'_p by auto } moreover { assume a: "D \ ?Ps i" then obtain l' where l'_p: "D \ ?Qs l'" "l' < llength Sts" using from_P_to_Q[OF deriv fair ns c a i_p(1) D_p(2) D_p(3)] by auto then have "D \ Q_of_state (Liminf_state Sts)" using from_Q_to_Q_inf[OF deriv fair ns c l'_p(1) l'_p(2)] D_p(3) \(1) \(2) D_p(2) by auto then have ?thesis using D_p \ by auto } moreover { assume a: "D \ ?Qs i" then have "D \ Q_of_state (Liminf_state Sts)" using from_Q_to_Q_inf[OF deriv fair ns c a i_p(1)] \ D_p(2,3) by auto then have ?thesis using D_p \ by auto } ultimately show ?thesis using i_p by auto qed text \ The following corresponds to Lemma 4.11: \ lemma fair_imp_Liminf_minus_Rf_subset_ground_Liminf_state: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and ns: "Gs = lmap grounding_of_state Sts" shows "Liminf_llist Gs - sr.Rf (Liminf_llist Gs) \ grounding_of_clss (Q_of_state (Liminf_state Sts))" proof let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have SQinf: "clss_of_state (Liminf_state Sts) = Liminf_llist (lmap Q_of_state Sts)" using fair unfolding fair_state_seq_def Liminf_state_def by auto fix C assume C_p: "C \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" then have "C \ Sup_llist Gs" using Liminf_llist_subset_Sup_llist[of Gs] by blast then obtain D_proto where "D_proto \ clss_of_state (Sup_state Sts) \ subsumes D_proto C" using in_Sup_llist_in_Sup_state unfolding ns subsumes_def by blast then obtain D where D_p: "D \ clss_of_state (Sup_state Sts)" "subsumes D C" "\E \ {E. E \ clss_of_state (Sup_state Sts) \ subsumes E C}. \ strictly_subsumes E D" using strictly_subsumes_has_minimum[of "{E. E \ clss_of_state (Sup_state Sts) \ subsumes E C}"] by auto have ground_C: "is_ground_cls C" using C_p using Liminf_grounding_of_state_ground ns by auto have "\D' \'. D' \ Q_of_state (Liminf_state Sts) \ D' \ \' = C \ is_ground_subst \'" using eventually_in_Qinf[of D C Gs] using D_p(1-3) deriv fair ns C_p ground_C by auto then obtain D' \' where D'_p: "D' \ Q_of_state (Liminf_state Sts) \ D' \ \' = C \ is_ground_subst \'" by blast then have "D' \ clss_of_state (Liminf_state Sts)" by simp then have "C \ grounding_of_state (Liminf_state Sts)" unfolding grounding_of_clss_def grounding_of_cls_def using D'_p by auto then show "C \ grounding_of_clss (Q_of_state (Liminf_state Sts))" using SQinf fair fair_state_seq_def by auto qed text \ The following corresponds to (one direction of) Theorem 4.13: \ lemma subseteq_Liminf_state_eventually_always: fixes CC assumes "finite CC" and "CC \ {}" and "CC \ Q_of_state (Liminf_state Sts)" shows "\j. enat j < llength Sts \ (\j' \ enat j. j' < llength Sts \ CC \ Q_of_state (lnth Sts j'))" proof - from assms(3) have "\C \ CC. \j. enat j < llength Sts \ (\j' \ enat j. j' < llength Sts \ C \ Q_of_state (lnth Sts j'))" unfolding Liminf_state_def Liminf_llist_def by force then obtain f where f_p: "\C \ CC. f C < llength Sts \ (\j' \ enat (f C). j' < llength Sts \ C \ Q_of_state (lnth Sts j'))" by moura define j :: nat where "j = Max (f ` CC)" have "enat j < llength Sts" unfolding j_def using f_p assms(1) by (metis (mono_tags) Max_in assms(2) finite_imageI imageE image_is_empty) moreover have "\C j'. C \ CC \ enat j \ j' \ j' < llength Sts \ C \ Q_of_state (lnth Sts j')" proof (intro allI impI) fix C :: "'a clause" and j' :: nat assume a: "C \ CC" "enat j \ enat j'" "enat j' < llength Sts" then have "f C \ j'" unfolding j_def using assms(1) Max.bounded_iff by auto then show "C \ Q_of_state (lnth Sts j')" using f_p a by auto qed ultimately show ?thesis by auto qed lemma empty_clause_in_Q_of_Liminf_state: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and empty_in: "{#} \ Liminf_llist (lmap grounding_of_state Sts)" shows "{#} \ Q_of_state (Liminf_state Sts)" proof - define Gs :: "'a clause set llist" where ns: "Gs = lmap grounding_of_state Sts" from empty_in have in_Liminf_not_Rf: "{#} \ Liminf_llist Gs - sr.Rf (Liminf_llist Gs)" unfolding ns sr.Rf_def by auto then have "{#} \ grounding_of_clss (Q_of_state (Liminf_state Sts))" using fair_imp_Liminf_minus_Rf_subset_ground_Liminf_state[OF deriv fair ns] by auto then show ?thesis unfolding grounding_of_clss_def grounding_of_cls_def by auto qed lemma grounding_of_state_Liminf_state_subseteq: "grounding_of_state (Liminf_state Sts) \ Liminf_llist (lmap grounding_of_state Sts)" proof fix C :: "'a clause" assume "C \ grounding_of_state (Liminf_state Sts)" then obtain D \ where D_\_p: "D \ clss_of_state (Liminf_state Sts)" "D \ \ = C" "is_ground_subst \" unfolding grounding_of_clss_def grounding_of_cls_def by auto then have ii: "D \ Liminf_llist (lmap N_of_state Sts) \ D \ Liminf_llist (lmap P_of_state Sts) \ D \ Liminf_llist (lmap Q_of_state Sts)" unfolding Liminf_state_def by simp then have "C \ Liminf_llist (lmap grounding_of_clss (lmap N_of_state Sts)) \ C \ Liminf_llist (lmap grounding_of_clss (lmap P_of_state Sts)) \ C \ Liminf_llist (lmap grounding_of_clss (lmap Q_of_state Sts))" unfolding Liminf_llist_def grounding_of_clss_def grounding_of_cls_def using D_\_p apply - apply (erule disjE) subgoal apply (rule disjI1) using D_\_p by auto subgoal apply (erule disjE) subgoal apply (rule disjI2) apply (rule disjI1) using D_\_p by auto subgoal apply (rule disjI2) apply (rule disjI2) using D_\_p by auto done done then show "C \ Liminf_llist (lmap grounding_of_state Sts)" unfolding Liminf_llist_def grounding_of_clss_def by auto qed theorem RP_sound: assumes deriv: "chain (\) Sts" and "{#} \ clss_of_state (Liminf_state Sts)" shows "\ satisfiable (grounding_of_state (lhd Sts))" proof - from assms have "{#} \ grounding_of_state (Liminf_state Sts)" unfolding grounding_of_clss_def by (force intro: ex_ground_subst) then have "{#} \ Liminf_llist (lmap grounding_of_state Sts)" using grounding_of_state_Liminf_state_subseteq by auto then have "\ satisfiable (Liminf_llist (lmap grounding_of_state Sts))" using true_clss_def by auto then have "\ satisfiable (lhd (lmap grounding_of_state Sts))" using sr_ext.sat_limit_iff ground_derive_chain deriv by blast then show ?thesis using chain_not_lnull deriv by fastforce qed theorem RP_saturated_if_fair: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and empty_Q0: "Q_of_state (lhd Sts) = {}" shows "sr.saturated_upto (Liminf_llist (lmap grounding_of_state Sts))" proof - define Gs :: "'a clause set llist" where ns: "Gs = lmap grounding_of_state Sts" let ?N = "\i. grounding_of_state (lnth Sts i)" let ?Ns = "\i. N_of_state (lnth Sts i)" let ?Ps = "\i. P_of_state (lnth Sts i)" let ?Qs = "\i. Q_of_state (lnth Sts i)" have ground_ns_in_ground_limit_st: "Liminf_llist Gs - sr.Rf (Liminf_llist Gs) \ grounding_of_clss (Q_of_state (Liminf_state Sts))" using fair deriv fair_imp_Liminf_minus_Rf_subset_ground_Liminf_state ns by blast have derivns: "chain sr_ext.derive Gs" using ground_derive_chain deriv ns by auto { fix \ :: "'a inference" assume \_p: "\ \ gr.ord_\" let ?CC = "side_prems_of \" let ?DA = "main_prem_of \" let ?E = "concl_of \" assume a: "set_mset ?CC \ {?DA} \ Liminf_llist (lmap grounding_of_state Sts) - sr.Rf (Liminf_llist (lmap grounding_of_state Sts))" have ground_ground_Liminf: "is_ground_clss (Liminf_llist (lmap grounding_of_state Sts))" using Liminf_grounding_of_state_ground unfolding is_ground_clss_def by auto have ground_cc: "is_ground_clss (set_mset ?CC)" using a ground_ground_Liminf is_ground_clss_def by auto have ground_da: "is_ground_cls ?DA" using a grounding_ground singletonI ground_ground_Liminf by (simp add: Liminf_grounding_of_state_ground) from \_p obtain CAs AAs As where CAs_p: "gr.ord_resolve CAs ?DA AAs As ?E \ mset CAs = ?CC" unfolding gr.ord_\_def by auto have DA_CAs_in_ground_Liminf: "{?DA} \ set CAs \ grounding_of_clss (Q_of_state (Liminf_state Sts))" using a CAs_p fair unfolding fair_state_seq_def by (metis (no_types, lifting) Un_empty_left ground_ns_in_ground_limit_st a ns set_mset_mset subset_trans sup_commute) then have ground_cas: "is_ground_cls_list CAs" using CAs_p unfolding is_ground_cls_list_def by auto have "\\. ord_resolve S_Q CAs ?DA AAs As \ ?E" by (rule ground_ord_resolve_imp_ord_resolve[OF ground_da ground_cas gr.ground_resolution_with_selection_axioms CAs_p[THEN conjunct1]]) then obtain \ where \_p: "ord_resolve S_Q CAs ?DA AAs As \ ?E" by auto then obtain \s' \' \2' CAs' DA' AAs' As' \' E' where s_p: "is_ground_subst \'" "is_ground_subst_list \s'" "is_ground_subst \2'" "ord_resolve_rename S CAs' DA' AAs' As' \' E'" "CAs' \\cl \s' = CAs" "DA' \ \' = ?DA" "E' \ \2' = ?E" "{DA'} \ set CAs' \ Q_of_state (Liminf_state Sts)" using ord_resolve_rename_lifting[OF sel_stable, of "Q_of_state (Liminf_state Sts)" CAs ?DA] \_p[unfolded S_Q_def] selection_axioms DA_CAs_in_ground_Liminf by metis from this(8) have "\j. enat j < llength Sts \ (set CAs' \ {DA'} \ ?Qs j)" unfolding Liminf_llist_def using subseteq_Liminf_state_eventually_always[of "{DA'} \ set CAs'"] by auto then obtain j where j_p: "is_least (\j. enat j < llength Sts \ set CAs' \ {DA'} \ ?Qs j) j" using least_exists[of "\j. enat j < llength Sts \ set CAs' \ {DA'} \ ?Qs j"] by force then have j_p': "enat j < llength Sts" "set CAs' \ {DA'} \ ?Qs j" unfolding is_least_def by auto then have jn0: "j \ 0" using empty_Q0 by (metis bot_eq_sup_iff gr_implies_not_zero insert_not_empty llength_lnull lnth_0_conv_lhd sup.orderE) then have j_adds_CAs': "\ set CAs' \ {DA'} \ ?Qs (j - 1)" "set CAs' \ {DA'} \ ?Qs j" using j_p unfolding is_least_def apply (metis (no_types) One_nat_def Suc_diff_Suc Suc_ile_eq diff_diff_cancel diff_zero less_imp_le less_one neq0_conv zero_less_diff) using j_p'(2) by blast have "lnth Sts (j - 1) \ lnth Sts j" using j_p'(1) jn0 deriv chain_lnth_rel[of _ _ "j - 1"] by force then obtain C' where C'_p: "?Ns (j - 1) = {}" "?Ps (j - 1) = ?Ps j \ {C'}" "?Qs j = ?Qs (j - 1) \ {C'}" "?Ns j = concls_of (ord_FO_resolution.inferences_between (?Qs (j - 1)) C')" "C' \ set CAs' \ {DA'}" "C' \ ?Qs (j - 1)" using j_adds_CAs' by (induction rule: RP.cases) auto have "E' \ ?Ns j" proof - have "E' \ concls_of (ord_FO_resolution.inferences_between (Q_of_state (lnth Sts (j - 1))) C')" unfolding infer_from_def ord_FO_\_def inference_system.inferences_between_def apply (rule_tac x = "Infer (mset CAs') DA' E'" in image_eqI) subgoal by auto subgoal unfolding infer_from_def by (rule ord_resolve_rename.cases[OF s_p(4)]) (use s_p(4) C'_p(3,5) j_p'(2) in force) done then show ?thesis using C'_p(4) by auto qed then have "E' \ clss_of_state (lnth Sts j)" using j_p' by auto then have "?E \ grounding_of_state (lnth Sts j)" using s_p(7) s_p(3) unfolding grounding_of_clss_def grounding_of_cls_def by force then have "\ \ sr.Ri (grounding_of_state (lnth Sts j))" using sr.Ri_effective \_p by auto then have "\ \ sr_ext_Ri (?N j)" unfolding sr_ext_Ri_def by auto then have "\ \ sr_ext_Ri (Sup_llist (lmap grounding_of_state Sts))" - using j_p' contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr_ext.Ri_mono by smt + using j_p' contra_subsetD llength_lmap lnth_lmap lnth_subset_Sup_llist sr_ext.Ri_mono by (smt (verit)) then have "\ \ sr_ext_Ri (Liminf_llist (lmap grounding_of_state Sts))" using sr_ext.Ri_limit_Sup[of Gs] derivns ns by blast } then have "sr_ext.saturated_upto (Liminf_llist (lmap grounding_of_state Sts))" unfolding sr_ext.saturated_upto_def sr_ext.inferences_from_def infer_from_def sr_ext_Ri_def by auto then show ?thesis using gd_ord_\_ngd_ord_\ sr.redundancy_criterion_axioms redundancy_criterion_standard_extension_saturated_upto_iff[of gr.ord_\] unfolding sr_ext_Ri_def by auto qed corollary RP_complete_if_fair: assumes deriv: "chain (\) Sts" and fair: "fair_state_seq Sts" and empty_Q0: "Q_of_state (lhd Sts) = {}" and unsat: "\ satisfiable (grounding_of_state (lhd Sts))" shows "{#} \ Q_of_state (Liminf_state Sts)" proof - have "\ satisfiable (Liminf_llist (lmap grounding_of_state Sts))" using unsat sr_ext.sat_limit_iff[OF ground_derive_chain] chain_not_lnull deriv by fastforce moreover have "sr.saturated_upto (Liminf_llist (lmap grounding_of_state Sts))" by (rule RP_saturated_if_fair[OF deriv fair empty_Q0, simplified]) ultimately have "{#} \ Liminf_llist (lmap grounding_of_state Sts)" using sr.saturated_upto_complete_if by auto then show ?thesis using empty_clause_in_Q_of_Liminf_state[OF deriv fair] by auto qed end end end diff --git a/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy b/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy --- a/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy +++ b/thys/Ordered_Resolution_Prover/Lazy_List_Chain.thy @@ -1,707 +1,707 @@ (* Title: Relational Chains over Lazy Lists Author: Jasmin Blanchette , 2014, 2017 Author: Dmitriy Traytel , 2017 Author: Anders Schlichtkrull , 2017 Maintainer: Anders Schlichtkrull *) section \Relational Chains over Lazy Lists\ theory Lazy_List_Chain imports "HOL-Library.BNF_Corec" Lazy_List_Liminf begin text \ A chain is a lazy list of elements such that all pairs of consecutive elements are related by a given relation. A full chain is either an infinite chain or a finite chain that cannot be extended. The inspiration for this theory is Section 4.1 (``Theorem Proving Processes'') of Bachmair and Ganzinger's chapter. \ subsection \Chains\ coinductive chain :: "('a \ 'a \ bool) \ 'a llist \ bool" for R :: "'a \ 'a \ bool" where chain_singleton: "chain R (LCons x LNil)" | chain_cons: "chain R xs \ R x (lhd xs) \ chain R (LCons x xs)" lemma chain_LNil[simp]: "\ chain R LNil" and chain_not_lnull: "chain R xs \ \ lnull xs" by (auto elim: chain.cases) lemma chain_lappend: assumes r_xs: "chain R xs" and r_ys: "chain R ys" and mid: "R (llast xs) (lhd ys)" shows "chain R (lappend xs ys)" proof (cases "lfinite xs") case True then show ?thesis using r_xs mid proof (induct rule: lfinite.induct) case (lfinite_LConsI xs x) note fin = this(1) and ih = this(2) and r_xxs = this(3) and mid = this(4) show ?case proof (cases "xs = LNil") case True then show ?thesis using r_ys mid by simp (rule chain_cons) next case xs_nnil: False have r_xs: "chain R xs" by (metis chain.simps ltl_simps(2) r_xxs xs_nnil) have mid': "R (llast xs) (lhd ys)" by (metis llast_LCons lnull_def mid xs_nnil) have start: "R x (lhd (lappend xs ys))" by (metis (no_types) chain.simps lhd_LCons lhd_lappend chain_not_lnull ltl_simps(2) r_xxs xs_nnil) show ?thesis unfolding lappend_code(2) using ih[OF r_xs mid'] start by (rule chain_cons) qed qed simp qed (simp add: r_xs lappend_inf) lemma chain_length_pos: "chain R xs \ llength xs > 0" by (cases xs) simp+ lemma chain_ldropn: assumes "chain R xs" and "enat n < llength xs" shows "chain R (ldropn n xs)" using assms by (induct n arbitrary: xs, simp, metis chain.cases ldrop_eSuc_ltl ldropn_LNil ldropn_eq_LNil ltl_simps(2) not_less) lemma inf_chain_ldropn_chain: "chain R xs \ \ lfinite xs \ chain R (ldropn n xs)" using chain.simps[of R xs] by (simp add: chain_ldropn not_lfinite_llength) lemma inf_chain_ltl_chain: "chain R xs \ \ lfinite xs \ chain R (ltl xs)" by (metis inf_chain_ldropn_chain ldropn_0 ldropn_ltl) lemma chain_lnth_rel: assumes chain: "chain R xs" and len: "enat (Suc j) < llength xs" shows "R (lnth xs j) (lnth xs (Suc j))" proof - define ys where "ys = ldropn j xs" have "llength ys > 1" unfolding ys_def using len by (metis One_nat_def funpow_swap1 ldropn_0 ldropn_def ldropn_eq_LNil ldropn_ltl not_less one_enat_def) obtain y0 y1 ys' where ys: "ys = LCons y0 (LCons y1 ys')" unfolding ys_def by (metis Suc_ile_eq ldropn_Suc_conv_ldropn len less_imp_not_less not_less) have "chain R ys" unfolding ys_def using Suc_ile_eq chain chain_ldropn len less_imp_le by blast then have "R y0 y1" unfolding ys by (auto elim: chain.cases) then show ?thesis using ys_def unfolding ys by (metis ldropn_Suc_conv_ldropn ldropn_eq_LConsD llist.inject) qed lemma infinite_chain_lnth_rel: assumes "\ lfinite c" and "chain r c" shows "r (lnth c i) (lnth c (Suc i))" using assms chain_lnth_rel lfinite_conv_llength_enat by force lemma lnth_rel_chain: assumes "\ lnull xs" and "\j. enat (j + 1) < llength xs \ R (lnth xs j) (lnth xs (j + 1))" shows "chain R xs" using assms proof (coinduction arbitrary: xs rule: chain.coinduct) case chain note nnul = this(1) and nth_chain = this(2) show ?case proof (cases "lnull (ltl xs)") case True have "xs = LCons (lhd xs) LNil" using nnul True by (simp add: llist.expand) then show ?thesis by blast next case nnul': False moreover have "xs = LCons (lhd xs) (ltl xs)" using nnul by simp moreover have "\j. enat (j + 1) < llength (ltl xs) \ R (lnth (ltl xs) j) (lnth (ltl xs) (j + 1))" using nnul nth_chain by (metis Suc_eq_plus1 ldrop_eSuc_ltl ldropn_Suc_conv_ldropn ldropn_eq_LConsD lnth_ltl) moreover have "R (lhd xs) (lhd (ltl xs))" using nnul' nnul nth_chain[rule_format, of 0, simplified] by (metis ldropn_0 ldropn_Suc_conv_ldropn ldropn_eq_LConsD lhd_LCons_ltl lhd_conv_lnth lnth_Suc_LCons ltl_simps(2)) ultimately show ?thesis by blast qed qed lemma chain_lmap: assumes "\x y. R x y \ R' (f x) (f y)" and "chain R xs" shows "chain R' (lmap f xs)" using assms proof (coinduction arbitrary: xs) case chain then have "(\y. xs = LCons y LNil) \ (\ys x. xs = LCons x ys \ chain R ys \ R x (lhd ys))" using chain.simps[of R xs] by auto then show ?case proof assume "\ys x. xs = LCons x ys \ chain R ys \ R x (lhd ys)" then have "\ys x. lmap f xs = LCons x ys \ (\xs. ys = lmap f xs \ (\x y. R x y \ R' (f x) (f y)) \ chain R xs) \ R' x (lhd ys)" using chain by (metis (no_types) lhd_LCons llist.distinct(1) llist.exhaust_sel llist.map_sel(1) lmap_eq_LNil chain_not_lnull ltl_lmap ltl_simps(2)) then show ?thesis by auto qed auto qed lemma chain_mono: assumes "\x y. R x y \ R' x y" and "chain R xs" shows "chain R' xs" using assms by (rule chain_lmap[of _ _ "\x. x", unfolded llist.map_ident]) lemma chain_ldropnI: assumes rel: "\j. j \ i \ enat (Suc j) < llength xs \ R (lnth xs j) (lnth xs (Suc j))" and si_lt: "enat (Suc i) < llength xs" shows "chain R (ldropn i xs)" proof (rule lnth_rel_chain) show "\ lnull (ldropn i xs)" using si_lt by (simp add: Suc_ile_eq less_le_not_le) next show "\j. enat (j + 1) < llength (ldropn i xs) \ R (lnth (ldropn i xs) j) (lnth (ldropn i xs) (j + 1))" - using rel by (smt (z3) One_nat_def Suc_ile_eq add.commute add.right_neutral add_Suc_right - add_le_cancel_right ldropn_eq_LNil ldropn_ldropn less_le_not_le linorder_not_less - lnth_ldropn not_less_zero) + using rel + by (smt (verit, best) Suc_ile_eq add.commute ldropn_eq_LNil ldropn_ldropn leD + le_add1 linorder_le_less_linear lnth_ldropn order_less_imp_le plus_1_eq_Suc) qed lemma chain_ldropn_lmapI: assumes rel: "\j. j \ i \ enat (Suc j) < llength xs \ R (f (lnth xs j)) (f (lnth xs (Suc j)))" and si_lt: "enat (Suc i) < llength xs" shows "chain R (ldropn i (lmap f xs))" proof - have "chain R (lmap f (ldropn i xs))" using chain_lmap[of "\x y. R (f x) (f y)" R f, of "ldropn i xs"] chain_ldropnI[OF rel si_lt] by auto thus ?thesis by auto qed lemma lfinite_chain_imp_rtranclp_lhd_llast: "lfinite xs \ chain R xs \ R\<^sup>*\<^sup>* (lhd xs) (llast xs)" proof (induct rule: lfinite.induct) case (lfinite_LConsI xs x) note fin_xs = this(1) and ih = this(2) and r_x_xs = this(3) show ?case proof (cases "xs = LNil") case xs_nnil: False then have r_xs: "chain R xs" using r_x_xs by (blast elim: chain.cases) then show ?thesis using ih[OF r_xs] xs_nnil r_x_xs by (metis chain.cases converse_rtranclp_into_rtranclp lhd_LCons llast_LCons chain_not_lnull ltl_simps(2)) qed simp qed simp lemma tranclp_imp_exists_finite_chain_list: "R\<^sup>+\<^sup>+ x y \ \xs. chain R (llist_of (x # xs @ [y]))" proof (induct rule: tranclp.induct) case (r_into_trancl x y) then have "chain R (llist_of (x # [] @ [y]))" by (auto intro: chain.intros) then show ?case by blast next case (trancl_into_trancl x y z) note rstar_xy = this(1) and ih = this(2) and r_yz = this(3) obtain xs where xs: "chain R (llist_of (x # xs @ [y]))" using ih by blast define ys where "ys = xs @ [y]" have "chain R (llist_of (x # ys @ [z]))" unfolding ys_def using r_yz chain_lappend[OF xs chain_singleton, of z] by (auto simp: lappend_llist_of_LCons llast_LCons) then show ?case by blast qed inductive_cases chain_consE: "chain R (LCons x xs)" inductive_cases chain_nontrivE: "chain R (LCons x (LCons y xs))" subsection \A Coinductive Puzzle\ primrec prepend where "prepend [] ys = ys" | "prepend (x # xs) ys = LCons x (prepend xs ys)" lemma lnull_prepend[simp]: "lnull (prepend xs ys) = (xs = [] \ lnull ys)" by (induct xs) auto lemma lhd_prepend[simp]: "lhd (prepend xs ys) = (if xs \ [] then hd xs else lhd ys)" by (induct xs) auto lemma prepend_LNil[simp]: "prepend xs LNil = llist_of xs" by (induct xs) auto lemma lfinite_prepend[simp]: "lfinite (prepend xs ys) \ lfinite ys" by (induct xs) auto lemma llength_prepend[simp]: "llength (prepend xs ys) = length xs + llength ys" by (induct xs) (auto simp: enat_0 iadd_Suc eSuc_enat[symmetric]) lemma llast_prepend[simp]: "\ lnull ys \ llast (prepend xs ys) = llast ys" by (induct xs) (auto simp: llast_LCons) lemma prepend_prepend: "prepend xs (prepend ys zs) = prepend (xs @ ys) zs" by (induct xs) auto lemma chain_prepend: "chain R (llist_of zs) \ last zs = lhd xs \ chain R xs \ chain R (prepend zs (ltl xs))" by (induct zs; cases xs) (auto split: if_splits simp: lnull_def[symmetric] intro!: chain_cons elim!: chain_consE) lemma lmap_prepend[simp]: "lmap f (prepend xs ys) = prepend (map f xs) (lmap f ys)" by (induct xs) auto lemma lset_prepend[simp]: "lset (prepend xs ys) = set xs \ lset ys" by (induct xs) auto lemma prepend_LCons: "prepend xs (LCons y ys) = prepend (xs @ [y]) ys" by (induct xs) auto lemma lnth_prepend: "lnth (prepend xs ys) i = (if i < length xs then nth xs i else lnth ys (i - length xs))" by (induct xs arbitrary: i) (auto simp: lnth_LCons' nth_Cons') theorem lfinite_less_induct[consumes 1, case_names less]: assumes fin: "lfinite xs" and step: "\xs. lfinite xs \ (\zs. llength zs < llength xs \ P zs) \ P xs" shows "P xs" using fin proof (induct "the_enat (llength xs)" arbitrary: xs rule: less_induct) case (less xs) show ?case using less(2) by (intro step[OF less(2)] less(1)) (auto dest!: lfinite_llength_enat simp: eSuc_enat elim!: less_enatE llength_eq_enat_lfiniteD) qed theorem lfinite_prepend_induct[consumes 1, case_names LNil prepend]: assumes "lfinite xs" and LNil: "P LNil" and prepend: "\xs. lfinite xs \ (\zs. (\ys. xs = prepend ys zs \ ys \ []) \ P zs) \ P xs" shows "P xs" using assms(1) proof (induct xs rule: lfinite_less_induct) case (less xs) from less(1) show ?case by (cases xs) (force simp: LNil neq_Nil_conv dest: lfinite_llength_enat intro!: prepend[of "LCons _ _"] intro: less)+ qed coinductive emb :: "'a llist \ 'a llist \ bool" where "lfinite xs \ emb LNil xs" | "emb xs ys \ emb (LCons x xs) (prepend zs (LCons x ys))" inductive_cases emb_LConsE: "emb (LCons z zs) ys" inductive_cases emb_LNil1E: "emb LNil ys" inductive_cases emb_LNil2E: "emb xs LNil" lemma emb_lfinite: assumes "emb xs ys" shows "lfinite ys \ lfinite xs" proof assume "lfinite xs" then show "lfinite ys" using assms by (induct xs arbitrary: ys rule: lfinite_induct) (auto simp: lnull_def neq_LNil_conv elim!: emb_LNil1E emb_LConsE) next assume "lfinite ys" then show "lfinite xs" using assms proof (induction ys arbitrary: xs rule: lfinite_less_induct) case (less ys) from less.prems \lfinite ys\ show ?case by (cases xs) (auto simp: eSuc_enat elim!: emb_LNil1E emb_LConsE less.IH[rotated] dest!: lfinite_llength_enat) qed qed inductive prepend_cong1 for X where prepend_cong1_base: "X xs \ prepend_cong1 X xs" | prepend_cong1_prepend: "prepend_cong1 X ys \ prepend_cong1 X (prepend xs ys)" lemma prepend_cong1_alt: "prepend_cong1 X xs \ (\ys zs. xs = prepend ys zs \ X zs)" by (rule iffI, induct xs rule: prepend_cong1.induct) (force simp: prepend_prepend intro: prepend_cong1.intros exI[of _ "[]"])+ lemma emb_prepend_coinduct_cong[rotated, case_names emb]: assumes "(\x1 x2. X x1 x2 \ (\xs. x1 = LNil \ x2 = xs \ lfinite xs) \ (\xs ys x zs. x1 = LCons x xs \ x2 = prepend zs (LCons x ys) \ (prepend_cong1 (X xs) ys \ emb xs ys)))" (is "\x1 x2. X x1 x2 \ ?bisim x1 x2") shows "X x1 x2 \ emb x1 x2" proof (erule emb.coinduct[OF prepend_cong1_base]) fix xs zs assume "prepend_cong1 (X xs) zs" then show "?bisim xs zs" by (induct zs rule: prepend_cong1.induct) (erule assms, force simp: prepend_prepend) qed lemma emb_prepend: "emb xs ys \ emb xs (prepend zs ys)" by (coinduction arbitrary: xs zs ys rule: emb_prepend_coinduct_cong) (force elim: emb.cases simp: prepend_prepend) lemma prepend_cong1_emb: "prepend_cong1 (emb xs) ys = emb xs ys" by (rule iffI, induct ys rule: prepend_cong1.induct) (simp_all add: emb_prepend prepend_cong1_base) lemma prepend_cong_distrib: "prepend_cong1 (P \ Q) xs \ prepend_cong1 P xs \ prepend_cong1 Q xs" unfolding prepend_cong1_alt by auto lemma emb_prepend_coinduct_aux[case_names emb]: assumes "X x1 x2 " "(\x1 x2. X x1 x2 \ (\xs. x1 = LNil \ x2 = xs \ lfinite xs) \ (\xs ys x zs. x1 = LCons x xs \ x2 = prepend zs (LCons x ys) \ (prepend_cong1 (X xs \ emb xs) ys)))" shows "emb x1 x2" using assms unfolding prepend_cong_distrib prepend_cong1_emb by (rule emb_prepend_coinduct_cong) lemma emb_prepend_coinduct[rotated, case_names emb]: assumes "(\x1 x2. X x1 x2 \ (\xs. x1 = LNil \ x2 = xs \ lfinite xs) \ (\xs ys x zs zs'. x1 = LCons x xs \ x2 = prepend zs (LCons x (prepend zs' ys)) \ (X xs ys \ emb xs ys)))" shows "X x1 x2 \ emb x1 x2" by (erule emb_prepend_coinduct_aux[of X]) (force simp: prepend_cong1_alt dest: assms) context begin private coinductive chain' for R where "chain' R (LCons x LNil)" | "chain R (llist_of (x # zs @ [lhd xs])) \ chain' R xs \ chain' R (LCons x (prepend zs xs))" private lemma chain_imp_chain': "chain R xs \ chain' R xs" proof (coinduction arbitrary: xs rule: chain'.coinduct) case chain' then show ?case proof (cases rule: chain.cases) case (chain_cons zs z) then show ?thesis by (intro disjI2 exI[of _ z] exI[of _ "[]"] exI[of _ "zs"]) (auto intro: chain.intros) qed simp qed private lemma chain'_imp_chain: "chain' R xs \ chain R xs" proof (coinduction arbitrary: xs rule: chain.coinduct) case chain then show ?case proof (cases rule: chain'.cases) case (2 y zs ys) then show ?thesis by (intro disjI2 exI[of _ "prepend zs ys"] exI[of _ y]) (force dest!: neq_Nil_conv[THEN iffD1] elim: chain.cases chain_nontrivE intro: chain'.intros) qed simp qed private lemma chain_chain': "chain = chain'" unfolding fun_eq_iff by (metis chain_imp_chain' chain'_imp_chain) lemma chain_prepend_coinduct[case_names chain]: "X x \ (\x. X x \ (\z. x = LCons z LNil) \ (\y xs zs. x = LCons y (prepend zs xs) \ (X xs \ chain R xs) \ chain R (llist_of (y # zs @ [lhd xs])))) \ chain R x" by (subst chain_chain', erule chain'.coinduct) (force simp: chain_chain') end context fixes R :: "'a \ 'a \ bool" begin private definition pick where "pick x y = (SOME xs. chain R (llist_of (x # xs @ [y])))" private lemma pick[simp]: assumes "R\<^sup>+\<^sup>+ x y" shows "chain R (llist_of (x # pick x y @ [y]))" unfolding pick_def using tranclp_imp_exists_finite_chain_list[THEN someI_ex, OF assms] by auto private friend_of_corec prepend where "prepend xs ys = (case xs of [] \ (case ys of LNil \ LNil | LCons x xs \ LCons x xs) | x # xs' \ LCons x (prepend xs' ys))" by (simp split: list.splits llist.splits) transfer_prover private corec wit where "wit xs = (case xs of LCons x (LCons y xs) \ LCons x (prepend (pick x y) (wit (LCons y xs))) | _ \ xs)" private lemma wit_LNil[simp]: "wit LNil = LNil" and wit_lsingleton[simp]: "wit (LCons x LNil) = LCons x LNil" and wit_LCons2: "wit (LCons x (LCons y xs)) = (LCons x (prepend (pick x y) (wit (LCons y xs))))" by (subst wit.code; auto)+ private lemma lnull_wit[simp]: "lnull (wit xs) \ lnull xs" by (subst wit.code) (auto split: llist.splits simp: Let_def) private lemma lhd_wit[simp]: "chain R\<^sup>+\<^sup>+ xs \ lhd (wit xs) = lhd xs" by (erule chain.cases; subst wit.code) (auto split: llist.splits simp: Let_def) private lemma LNil_eq_iff_lnull: "LNil = xs \ lnull xs" by (cases xs) auto lemma emb_wit[simp]: "chain R\<^sup>+\<^sup>+ xs \ emb xs (wit xs)" proof (coinduction arbitrary: xs rule: emb_prepend_coinduct) case (emb xs) then show ?case proof (cases rule: chain.cases) case (chain_cons zs z) then show ?thesis by (subst (2) wit.code) (auto 0 3 split: llist.splits intro: exI[of _ "[]"] exI[of _ "pick z _"] intro!: exI[of _ "_ :: _ llist"]) qed (auto intro!: exI[of _ LNil] exI[of _ "[]"] emb.intros) qed private lemma lfinite_wit[simp]: assumes "chain R\<^sup>+\<^sup>+ xs" shows "lfinite (wit xs) \ lfinite xs" using emb_wit emb_lfinite assms by blast private lemma llast_wit[simp]: assumes "chain R\<^sup>+\<^sup>+ xs" shows "llast (wit xs) = llast xs" proof (cases "lfinite xs") case True from this assms show ?thesis proof (induct rule: lfinite.induct) case (lfinite_LConsI xs x) then show ?case by (cases xs) (auto simp: wit_LCons2 llast_LCons elim: chain_nontrivE) qed auto qed (auto simp: llast_linfinite assms) lemma chain_tranclp_imp_exists_chain: "chain R\<^sup>+\<^sup>+ xs \ \ys. chain R ys \ emb xs ys \ lhd ys = lhd xs \ llast ys = llast xs" proof (intro exI[of _ "wit xs"] conjI, coinduction arbitrary: xs rule: chain_prepend_coinduct) case chain then show ?case by (subst (1 2) wit.code) (erule chain.cases; force split: llist.splits dest: pick) qed auto lemma emb_lset_mono[rotated]: "x \ lset xs \ emb xs ys \ x \ lset ys" by (induct x xs arbitrary: ys rule: llist.set_induct) (auto elim!: emb_LConsE) lemma emb_Ball_lset_antimono: assumes "emb Xs Ys" shows "\Y \ lset Ys. x \ Y \ \X \ lset Xs. x \ X" using emb_lset_mono[OF assms] by blast lemma emb_lfinite_antimono[rotated]: "lfinite ys \ emb xs ys \ lfinite xs" by (induct ys arbitrary: xs rule: lfinite_prepend_induct) (force elim!: emb_LNil2E simp: LNil_eq_iff_lnull prepend_LCons elim: emb.cases)+ lemma emb_Liminf_llist_mono_aux: assumes "emb Xs Ys" and "\ lfinite Xs" and "\ lfinite Ys" and "\j\i. x \ lnth Ys j" shows "\j\i. x \ lnth Xs j" using assms proof (induct i arbitrary: Xs Ys rule: less_induct) case (less i) then show ?case proof (cases i) case 0 then show ?thesis using emb_Ball_lset_antimono[OF less(2), of x] less(5) unfolding Ball_def in_lset_conv_lnth simp_thms not_lfinite_llength[OF less(3)] not_lfinite_llength[OF less(4)] enat_ord_code subset_eq by blast next case [simp]: (Suc nat) from less(2,3) obtain xs as b bs where [simp]: "Xs = LCons b xs" "Ys = prepend as (LCons b bs)" and "emb xs bs" by (auto elim: emb.cases) have IH: "\k\j. x \ lnth xs k" if "\k\j. x \ lnth bs k" "j < i" for j using that less(1)[OF _ \emb xs bs\] less(3,4) by auto from less(5) have "\k\i - length as - 1. x \ lnth xs k" by (intro IH allI) (drule spec[of _ "_ + length as + 1"], auto simp: lnth_prepend lnth_LCons') then show ?thesis by (auto simp: lnth_LCons') qed qed lemma emb_Liminf_llist_infinite: assumes "emb Xs Ys" and "\ lfinite Xs" shows "Liminf_llist Ys \ Liminf_llist Xs" proof - from assms have "\ lfinite Ys" using emb_lfinite_antimono by blast with assms show ?thesis unfolding Liminf_llist_def by (auto simp: not_lfinite_llength dest: emb_Liminf_llist_mono_aux) qed lemma emb_lmap: "emb xs ys \ emb (lmap f xs) (lmap f ys)" proof (coinduction arbitrary: xs ys rule: emb.coinduct) case emb show ?case proof (cases xs) case xs: (LCons x xs') obtain ysa0 and zs0 where ys: "ys = prepend zs0 (LCons x ysa0)" and emb': "emb xs' ysa0" using emb_LConsE[OF emb[unfolded xs]] by metis let ?xa = "f x" let ?xsa = "lmap f xs'" let ?zs = "map f zs0" let ?ysa = "lmap f ysa0" have "lmap f xs = LCons ?xa ?xsa" unfolding xs by simp moreover have "lmap f ys = prepend ?zs (LCons ?xa ?ysa)" unfolding ys by simp moreover have "\xsa ysa. ?xsa = lmap f xsa \ ?ysa = lmap f ysa \ emb xsa ysa" using emb' by blast ultimately show ?thesis by blast qed (simp add: emb_lfinite[OF emb]) qed end lemma chain_inf_llist_if_infinite_chain_function: assumes "\i. r (f (Suc i)) (f i)" shows "\ lfinite (inf_llist f) \ chain r\\ (inf_llist f)" using assms by (simp add: lnth_rel_chain) lemma infinite_chain_function_iff_infinite_chain_llist: "(\f. \i. r (f (Suc i)) (f i)) \ (\c. \ lfinite c \ chain r\\ c)" using chain_inf_llist_if_infinite_chain_function infinite_chain_lnth_rel by blast lemma wfP_iff_no_infinite_down_chain_llist: "wfP r \ (\c. \ lfinite c \ chain r\\ c)" proof - have "wfP r \ wf {(x, y). r x y}" unfolding wfP_def by auto also have "\ \ (\f. \i. (f (Suc i), f i) \ {(x, y). r x y})" using wf_iff_no_infinite_down_chain by blast also have "\ \ (\f. \i. r (f (Suc i)) (f i))" by auto also have "\ \ (\c. \lfinite c \ chain r\\ c)" using infinite_chain_function_iff_infinite_chain_llist by blast finally show ?thesis by auto qed subsection \Full Chains\ coinductive full_chain :: "('a \ 'a \ bool) \ 'a llist \ bool" for R :: "'a \ 'a \ bool" where full_chain_singleton: "(\y. \ R x y) \ full_chain R (LCons x LNil)" | full_chain_cons: "full_chain R xs \ R x (lhd xs) \ full_chain R (LCons x xs)" lemma full_chain_LNil[simp]: "\ full_chain R LNil" and full_chain_not_lnull: "full_chain R xs \ \ lnull xs" by (auto elim: full_chain.cases) lemma full_chain_ldropn: assumes full: "full_chain R xs" and "enat n < llength xs" shows "full_chain R (ldropn n xs)" using assms by (induct n arbitrary: xs, simp, metis full_chain.cases ldrop_eSuc_ltl ldropn_LNil ldropn_eq_LNil ltl_simps(2) not_less) lemma full_chain_iff_chain: "full_chain R xs \ chain R xs \ (lfinite xs \ (\y. \ R (llast xs) y))" proof (intro iffI conjI impI allI; (elim conjE)?) assume full: "full_chain R xs" show chain: "chain R xs" using full by (coinduction arbitrary: xs) (auto elim: full_chain.cases) { fix y assume "lfinite xs" then obtain n where suc_n: "Suc n = llength xs" by (metis chain chain_length_pos lessE less_enatE lfinite_conv_llength_enat) have "full_chain R (ldropn n xs)" by (rule full_chain_ldropn[OF full]) (use suc_n Suc_ile_eq in force) moreover have "ldropn n xs = LCons (llast xs) LNil" using suc_n by (metis enat_le_plus_same(2) enat_ord_simps(2) gen_llength_def ldropn_Suc_conv_ldropn ldropn_all lessI llast_ldropn llast_singleton llength_code) ultimately show "\ R (llast xs) y" by (auto elim: full_chain.cases) } next assume "chain R xs" and "lfinite xs \ (\y. \ R (llast xs) y)" then show "full_chain R xs" by (coinduction arbitrary: xs) (erule chain.cases, simp, metis lfinite_LConsI llast_LCons) qed lemma full_chain_imp_chain: "full_chain R xs \ chain R xs" using full_chain_iff_chain by blast lemma full_chain_length_pos: "full_chain R xs \ llength xs > 0" by (fact chain_length_pos[OF full_chain_imp_chain]) lemma full_chain_lnth_rel: "full_chain R xs \ enat (Suc j) < llength xs \ R (lnth xs j) (lnth xs (Suc j))" by (fact chain_lnth_rel[OF full_chain_imp_chain]) lemma full_chain_lnth_not_rel: assumes full: "full_chain R xs" and sj: "enat (Suc j) = llength xs" shows "\ R (lnth xs j) y" proof - have "lfinite xs" by (metis llength_eq_enat_lfiniteD sj) hence "\ R (llast xs) y" using full_chain_iff_chain full by metis thus ?thesis by (metis eSuc_enat llast_conv_lnth sj) qed inductive_cases full_chain_consE: "full_chain R (LCons x xs)" inductive_cases full_chain_nontrivE: "full_chain R (LCons x (LCons y xs))" lemma full_chain_tranclp_imp_exists_full_chain: assumes full: "full_chain R\<^sup>+\<^sup>+ xs" shows "\ys. full_chain R ys \ emb xs ys \ lhd ys = lhd xs \ llast ys = llast xs" proof - obtain ys where ys: "chain R ys" "emb xs ys" "lhd ys = lhd xs" "llast ys = llast xs" using full_chain_imp_chain[OF full] chain_tranclp_imp_exists_chain by blast have "full_chain R ys" using ys(1,4) emb_lfinite[OF ys(2)] full unfolding full_chain_iff_chain by auto then show ?thesis using ys(2-4) by auto qed end diff --git a/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy b/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy --- a/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy +++ b/thys/Ordered_Resolution_Prover/Lazy_List_Liminf.thy @@ -1,389 +1,389 @@ (* Title: Supremum and Liminf of Lazy Lists Author: Jasmin Blanchette , 2014, 2017, 2020 Author: Dmitriy Traytel , 2014 Maintainer: Jasmin Blanchette *) section \Supremum and Liminf of Lazy Lists\ theory Lazy_List_Liminf imports Coinductive.Coinductive_List begin text \ Lazy lists, as defined in the \emph{Archive of Formal Proofs}, provide finite and infinite lists in one type, defined coinductively. The present theory introduces the concept of the union of all elements of a lazy list of sets and the limit of such a lazy list. The definitions are stated more generally in terms of lattices. The basis for this theory is Section 4.1 (``Theorem Proving Processes'') of Bachmair and Ganzinger's chapter. \ subsection \Library\ lemma less_llength_ltake: "i < llength (ltake k Xs) \ i < k \ i < llength Xs" by simp subsection \Supremum\ definition Sup_llist :: "'a set llist \ 'a set" where "Sup_llist Xs = (\i \ {i. enat i < llength Xs}. lnth Xs i)" lemma lnth_subset_Sup_llist: "enat i < llength Xs \ lnth Xs i \ Sup_llist Xs" unfolding Sup_llist_def by auto lemma Sup_llist_imp_exists_index: "x \ Sup_llist Xs \ \i. enat i < llength Xs \ x \ lnth Xs i" unfolding Sup_llist_def by auto lemma exists_index_imp_Sup_llist: "enat i < llength Xs \ x \ lnth Xs i \ x \ Sup_llist Xs" unfolding Sup_llist_def by auto lemma Sup_llist_LNil[simp]: "Sup_llist LNil = {}" unfolding Sup_llist_def by auto lemma Sup_llist_LCons[simp]: "Sup_llist (LCons X Xs) = X \ Sup_llist Xs" unfolding Sup_llist_def proof (intro subset_antisym subsetI) fix x assume "x \ (\i \ {i. enat i < llength (LCons X Xs)}. lnth (LCons X Xs) i)" then obtain i where len: "enat i < llength (LCons X Xs)" and nth: "x \ lnth (LCons X Xs) i" by blast from nth have "x \ X \ i > 0 \ x \ lnth Xs (i - 1)" by (metis lnth_LCons' neq0_conv) then have "x \ X \ (\i. enat i < llength Xs \ x \ lnth Xs i)" by (metis len Suc_pred' eSuc_enat iless_Suc_eq less_irrefl llength_LCons not_less order_trans) then show "x \ X \ (\i \ {i. enat i < llength Xs}. lnth Xs i)" by blast qed ((auto)[], metis i0_lb lnth_0 zero_enat_def, metis Suc_ile_eq lnth_Suc_LCons) lemma lhd_subset_Sup_llist: "\ lnull Xs \ lhd Xs \ Sup_llist Xs" by (cases Xs) simp_all subsection \Supremum up-to\ definition Sup_upto_llist :: "'a set llist \ enat \ 'a set" where "Sup_upto_llist Xs j = (\i \ {i. enat i < llength Xs \ enat i \ j}. lnth Xs i)" lemma Sup_upto_llist_eq_Sup_llist_ltake: "Sup_upto_llist Xs j = Sup_llist (ltake (eSuc j) Xs)" unfolding Sup_upto_llist_def Sup_llist_def - by (smt Collect_cong Sup.SUP_cong iless_Suc_eq lnth_ltake less_llength_ltake mem_Collect_eq) + by (smt (verit) Collect_cong Sup.SUP_cong iless_Suc_eq lnth_ltake less_llength_ltake mem_Collect_eq) lemma Sup_upto_llist_enat_0[simp]: "Sup_upto_llist Xs (enat 0) = (if lnull Xs then {} else lhd Xs)" proof (cases "lnull Xs") case True then show ?thesis unfolding Sup_upto_llist_def by auto next case False show ?thesis unfolding Sup_upto_llist_def image_def by (simp add: lhd_conv_lnth enat_0 enat_0_iff) qed lemma Sup_upto_llist_Suc[simp]: "Sup_upto_llist Xs (enat (Suc j)) = Sup_upto_llist Xs (enat j) \ (if enat (Suc j) < llength Xs then lnth Xs (Suc j) else {})" unfolding Sup_upto_llist_def image_def by (auto intro: le_SucI elim: le_SucE) lemma Sup_upto_llist_infinity[simp]: "Sup_upto_llist Xs \ = Sup_llist Xs" unfolding Sup_upto_llist_def Sup_llist_def by simp lemma Sup_upto_llist_0[simp]: "Sup_upto_llist Xs 0 = (if lnull Xs then {} else lhd Xs)" unfolding zero_enat_def by (rule Sup_upto_llist_enat_0) lemma Sup_upto_llist_eSuc[simp]: "Sup_upto_llist Xs (eSuc j) = (case j of enat k \ Sup_upto_llist Xs (enat (Suc k)) | \ \ Sup_llist Xs)" by (auto simp: eSuc_enat split: enat.split) lemma Sup_upto_llist_mono[simp]: "j \ k \ Sup_upto_llist Xs j \ Sup_upto_llist Xs k" unfolding Sup_upto_llist_def by auto lemma Sup_upto_llist_subset_Sup_llist: "Sup_upto_llist Xs j \ Sup_llist Xs" unfolding Sup_llist_def Sup_upto_llist_def by auto lemma elem_Sup_llist_imp_Sup_upto_llist: "x \ Sup_llist Xs \ \j < llength Xs. x \ Sup_upto_llist Xs (enat j)" unfolding Sup_llist_def Sup_upto_llist_def by blast lemma lnth_subset_Sup_upto_llist: "j < llength Xs \ lnth Xs j \ Sup_upto_llist Xs j" unfolding Sup_upto_llist_def by auto lemma finite_Sup_llist_imp_Sup_upto_llist: assumes "finite X" and "X \ Sup_llist Xs" shows "\k. X \ Sup_upto_llist Xs (enat k)" using assms proof induct case (insert x X) then have x: "x \ Sup_llist Xs" and X: "X \ Sup_llist Xs" by simp+ from x obtain k where k: "x \ Sup_upto_llist Xs (enat k)" using elem_Sup_llist_imp_Sup_upto_llist by fast from X obtain k' where k': "X \ Sup_upto_llist Xs (enat k')" using insert.hyps(3) by fast have "insert x X \ Sup_upto_llist Xs (max k k')" using k k' by (metis (mono_tags) Sup_upto_llist_mono enat_ord_simps(1) insert_subset max.cobounded1 max.cobounded2 subset_iff) then show ?case by fast qed simp subsection \Liminf\ definition Liminf_llist :: "'a set llist \ 'a set" where "Liminf_llist Xs = (\i \ {i. enat i < llength Xs}. \j \ {j. i \ j \ enat j < llength Xs}. lnth Xs j)" lemma Liminf_llist_LNil[simp]: "Liminf_llist LNil = {}" unfolding Liminf_llist_def by simp lemma Liminf_llist_LCons: "Liminf_llist (LCons X Xs) = (if lnull Xs then X else Liminf_llist Xs)" (is "?lhs = ?rhs") proof (cases "lnull Xs") case nnull: False show ?thesis proof { fix x assume "\i. enat i \ llength Xs \ (\j. i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j)" then have "\i. enat (Suc i) \ llength Xs \ (\j. Suc i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j)" by (cases "llength Xs", metis not_lnull_conv[THEN iffD1, OF nnull] Suc_le_D eSuc_enat eSuc_ile_mono llength_LCons not_less_eq_eq zero_enat_def zero_le, metis Suc_leD enat_ord_code(3)) then have "\i. enat i < llength Xs \ (\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j)" by (metis Suc_ile_eq Suc_n_not_le_n lift_Suc_mono_le lnth_Suc_LCons nat_le_linear) } then show "?lhs \ ?rhs" by (simp add: Liminf_llist_def nnull) (rule subsetI, simp) { fix x assume "\i. enat i < llength Xs \ (\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j)" then obtain i where i: "enat i < llength Xs" and j: "\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j" by blast have "enat (Suc i) \ llength Xs" using i by (simp add: Suc_ile_eq) moreover have "\j. Suc i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j" using Suc_ile_eq Suc_le_D j by force ultimately have "\i. enat i \ llength Xs \ (\j. i \ j \ enat j \ llength Xs \ x \ lnth (LCons X Xs) j)" by blast } then show "?rhs \ ?lhs" by (simp add: Liminf_llist_def nnull) (rule subsetI, simp) qed qed (simp add: Liminf_llist_def enat_0_iff(1)) lemma lfinite_Liminf_llist: "lfinite Xs \ Liminf_llist Xs = (if lnull Xs then {} else llast Xs)" proof (induction rule: lfinite_induct) case (LCons xs) then obtain y ys where xs: "xs = LCons y ys" by (meson not_lnull_conv) show ?case unfolding xs by (simp add: Liminf_llist_LCons LCons.IH[unfolded xs, simplified] llast_LCons) qed (simp add: Liminf_llist_def) lemma Liminf_llist_ltl: "\ lnull (ltl Xs) \ Liminf_llist Xs = Liminf_llist (ltl Xs)" by (metis Liminf_llist_LCons lhd_LCons_ltl lnull_ltlI) lemma Liminf_llist_subset_Sup_llist: "Liminf_llist Xs \ Sup_llist Xs" unfolding Liminf_llist_def Sup_llist_def by fast lemma image_Liminf_llist_subset: "f ` Liminf_llist Ns \ Liminf_llist (lmap ((`) f) Ns)" unfolding Liminf_llist_def by auto lemma Liminf_llist_imp_exists_index: "x \ Liminf_llist Xs \ \i. enat i < llength Xs \ x \ lnth Xs i" unfolding Liminf_llist_def by auto lemma not_Liminf_llist_imp_exists_index: "\ lnull Xs \ x \ Liminf_llist Xs \ enat i < llength Xs \ (\j. i \ j \ enat j < llength Xs \ x \ lnth Xs j)" unfolding Liminf_llist_def by auto lemma finite_subset_Liminf_llist_imp_exists_index: assumes nnil: "\ lnull Xs" and fin: "finite X" and in_lim: "X \ Liminf_llist Xs" shows "\i. enat i < llength Xs \ X \ (\j \ {j. i \ j \ enat j < llength Xs}. lnth Xs j)" proof - show ?thesis proof (cases "X = {}") case True then show ?thesis using nnil by (auto intro: exI[of _ 0] simp: zero_enat_def[symmetric]) next case nemp: False have in_lim': "\x \ X. \i. enat i < llength Xs \ x \ (\j \ {j. i \ j \ enat j < llength Xs}. lnth Xs j)" using in_lim[unfolded Liminf_llist_def] in_mono by fastforce obtain i_of where i_of_lt: "\x \ X. enat (i_of x) < llength Xs" and in_inter: "\x \ X. x \ (\j \ {j. i_of x \ j \ enat j < llength Xs}. lnth Xs j)" using bchoice[OF in_lim'] by blast define i_max where "i_max = Max (i_of ` X)" have "i_max \ i_of ` X" by (simp add: fin i_max_def nemp) then obtain x_max where x_max_in: "x_max \ X" and i_max_is: "i_max = i_of x_max" unfolding i_max_def by blast have le_i_max: "\x \ X. i_of x \ i_max" unfolding i_max_def by (simp add: fin) have "enat i_max < llength Xs" using i_of_lt x_max_in i_max_is by auto moreover have "X \ (\j \ {j. i_max \ j \ enat j < llength Xs}. lnth Xs j)" proof fix x assume x_in: "x \ X" then have x_in_inter: "x \ (\j \ {j. i_of x \ j \ enat j < llength Xs}. lnth Xs j)" using in_inter by auto moreover have "{j. i_max \ j \ enat j < llength Xs} \ {j. i_of x \ j \ enat j < llength Xs}" using x_in le_i_max by auto ultimately show "x \ (\j \ {j. i_max \ j \ enat j < llength Xs}. lnth Xs j)" by auto qed ultimately show ?thesis by auto qed qed lemma Liminf_llist_lmap_image: assumes f_inj: "inj_on f (Sup_llist (lmap g xs))" shows "Liminf_llist (lmap (\x. f ` g x) xs) = f ` Liminf_llist (lmap g xs)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof fix x assume "x \ Liminf_llist (lmap (\x. f ` g x) xs)" then obtain i where i_lt: "enat i < llength xs" and x_in_fgj: "\j. i \ j \ enat j < llength xs \ x \ f ` g (lnth xs j)" unfolding Liminf_llist_def by auto have ex_in_gi: "\y. y \ g (lnth xs i) \ x = f y" using f_inj i_lt x_in_fgj unfolding inj_on_def Sup_llist_def by auto have "\y. \j. i \ j \ enat j < llength xs \ y \ g (lnth xs j) \ x = f y" apply (rule exI[of _ "SOME y. y \ g (lnth xs i) \ x = f y"]) using someI_ex[OF ex_in_gi] x_in_fgj f_inj i_lt x_in_fgj unfolding inj_on_def Sup_llist_def by simp (metis (no_types, lifting) imageE) then show "x \ f ` Liminf_llist (lmap g xs)" using i_lt unfolding Liminf_llist_def by auto qed next show "?rhs \ ?lhs" using image_Liminf_llist_subset[of f "lmap g xs", unfolded llist.map_comp] by auto qed lemma Liminf_llist_lmap_union: assumes "\x \ lset xs. \Y \ lset xs. g x \ h Y = {}" shows "Liminf_llist (lmap (\x. g x \ h x) xs) = Liminf_llist (lmap g xs) \ Liminf_llist (lmap h xs)" (is "?lhs = ?rhs") proof (intro equalityI subsetI) fix x assume x_in: "x \ ?lhs" then obtain i where i_lt: "enat i < llength xs" and j: "\j. i \ j \ enat j < llength xs \ x \ g (lnth xs j) \ x \ h (lnth xs j)" using x_in[unfolded Liminf_llist_def, simplified] by blast then have "(\i'. enat i' < llength xs \ (\j. i' \ j \ enat j < llength xs \ x \ g (lnth xs j))) \ (\i'. enat i' < llength xs \ (\j. i' \ j \ enat j < llength xs \ x \ h (lnth xs j)))" using assms[unfolded disjoint_iff_not_equal] by (metis in_lset_conv_lnth) then show "x \ ?rhs" unfolding Liminf_llist_def by simp next fix x show "x \ ?rhs \ x \ ?lhs" using assms unfolding Liminf_llist_def by auto qed lemma Liminf_set_filter_commute: "Liminf_llist (lmap (\X. {x \ X. p x}) Xs) = {x \ Liminf_llist Xs. p x}" unfolding Liminf_llist_def by force subsection \Liminf up-to\ definition Liminf_upto_llist :: "'a set llist \ enat \ 'a set" where "Liminf_upto_llist Xs k = (\i \ {i. enat i < llength Xs \ enat i \ k}. \j \ {j. i \ j \ enat j < llength Xs \ enat j \ k}. lnth Xs j)" lemma Liminf_upto_llist_eq_Liminf_llist_ltake: "Liminf_upto_llist Xs j = Liminf_llist (ltake (eSuc j) Xs)" unfolding Liminf_upto_llist_def Liminf_llist_def by (smt Collect_cong Sup.SUP_cong iless_Suc_eq lnth_ltake less_llength_ltake mem_Collect_eq) lemma Liminf_upto_llist_enat[simp]: "Liminf_upto_llist Xs (enat k) = (if enat k < llength Xs then lnth Xs k else if lnull Xs then {} else llast Xs)" proof (cases "enat k < llength Xs") case True then show ?thesis unfolding Liminf_upto_llist_def by force next case k_ge: False show ?thesis proof (cases "lnull Xs") case nil: True then show ?thesis unfolding Liminf_upto_llist_def by simp next case nnil: False then obtain j where j: "eSuc (enat j) = llength Xs" using k_ge by (metis eSuc_enat_iff enat_ile le_less_linear lhd_LCons_ltl llength_LCons) have fin: "lfinite Xs" using k_ge not_lfinite_llength by fastforce have le_k: "enat i < llength Xs \ i \ k \ enat i < llength Xs" for i using k_ge linear order_le_less_subst2 by fastforce have "Liminf_upto_llist Xs (enat k) = llast Xs" using j nnil lfinite_Liminf_llist[OF fin] nnil unfolding Liminf_upto_llist_def Liminf_llist_def using llast_conv_lnth[OF j[symmetric]] by (simp add: le_k) then show ?thesis using k_ge nnil by simp qed qed lemma Liminf_upto_llist_infinity[simp]: "Liminf_upto_llist Xs \ = Liminf_llist Xs" unfolding Liminf_upto_llist_def Liminf_llist_def by simp lemma Liminf_upto_llist_0[simp]: "Liminf_upto_llist Xs 0 = (if lnull Xs then {} else lhd Xs)" unfolding Liminf_upto_llist_def image_def by (simp add: enat_0[symmetric]) (simp add: enat_0 lnth_0_conv_lhd) lemma Liminf_upto_llist_eSuc[simp]: "Liminf_upto_llist Xs (eSuc j) = (case j of enat k \ Liminf_upto_llist Xs (enat (Suc k)) | \ \ Liminf_llist Xs)" by (auto simp: eSuc_enat split: enat.split) lemma elem_Liminf_llist_imp_Liminf_upto_llist: "x \ Liminf_llist Xs \ \i < llength Xs. \j. i \ j \ j < llength Xs \ x \ Liminf_upto_llist Xs (enat j)" unfolding Liminf_llist_def Liminf_upto_llist_def using enat_ord_simps(1) by force end diff --git a/thys/Relational-Incorrectness-Logic/ROOT b/thys/Relational-Incorrectness-Logic/ROOT --- a/thys/Relational-Incorrectness-Logic/ROOT +++ b/thys/Relational-Incorrectness-Logic/ROOT @@ -1,10 +1,9 @@ chapter AFP session "Relational-Incorrectness-Logic" = "HOL-IMP" + options [timeout = 600] theories - RelationalIncorrectness - + RelationalIncorrectness document_files - "root.bib" - "root.tex" + "root.bib" + "root.tex" diff --git a/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy b/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy --- a/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy +++ b/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy @@ -1,834 +1,834 @@ theory RelationalIncorrectness imports "HOL-IMP.Big_Step" begin (* Author: Toby Murray *) section "Under-Approximate Relational Judgement" text \ This is the relational analogue of OHearn's~\<^cite>\"OHearn_19"\ and de Vries \& Koutavas'~\<^cite>\"deVries_Koutavas_11"\ judgements. Note that in our case it doesn't really make sense to talk about ``erroneous'' states: the presence of an error can be seen only by the violation of a relation. Unlike O'Hearn, we cannot encode it directly into the semantics of our programs, without giving them a relational semantics. We use the standard big step semantics of IMP unchanged. \ type_synonym rassn = "state \ state \ bool" definition ir_valid :: "rassn \ com \ com \ rassn \ bool" where "ir_valid P c c' Q \ (\ t t'. Q t t' \ (\s s'. P s s' \ (c,s) \ t \ (c',s') \ t'))" section "Rules of the Logic" definition flip :: "rassn \ rassn" where "flip P \ \s s'. P s' s" inductive ir_hoare :: "rassn \ com \ com \ rassn \ bool" where ir_Skip: "(\t t'. Q t t' \ \s'. P t s' \ (c',s') \ t') \ ir_hoare P SKIP c' Q" | ir_If_True: "ir_hoare (\s s'. P s s' \ bval b s) c\<^sub>1 c' Q \ ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q" | ir_If_False: "ir_hoare (\s s'. P s s' \ \ bval b s) c\<^sub>2 c' Q \ ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q" | ir_Seq1: "ir_hoare P c c' Q \ ir_hoare Q d SKIP R \ ir_hoare P (Seq c d) c' R" | ir_Assign: "ir_hoare (\t t'. \ v. P (t(x := v)) t' \ (t x) = aval e (t(x := v))) SKIP c' Q \ ir_hoare P (Assign x e) c' Q" | ir_While_False: "ir_hoare (\s s'. P s s' \ \ bval b s) SKIP c' Q \ ir_hoare P (WHILE b DO c) c' Q" | ir_While_True: "ir_hoare (\s s'. P s s' \ bval b s) (c;; WHILE b DO c) c' Q \ ir_hoare P (WHILE b DO c) c' Q" | ir_While_backwards_frontier: "(\n. ir_hoare (\ s s'. P n s s' \ bval b s) c SKIP (P (Suc n))) \ ir_hoare (\s s'. \n. P n s s') (WHILE b DO c) c' Q \ ir_hoare (P 0) (WHILE b DO c) c' Q" | ir_conseq: "ir_hoare P c c' Q \ (\s s'. P s s' \ P' s s') \ (\s s'. Q' s s' \ Q s s') \ ir_hoare P' c c' Q'" | ir_disj: "ir_hoare P\<^sub>1 c c' Q\<^sub>1 \ ir_hoare P\<^sub>2 c c' Q\<^sub>2 \ ir_hoare (\s s'. P\<^sub>1 s s' \ P\<^sub>2 s s') c c' (\ t t'. Q\<^sub>1 t t' \ Q\<^sub>2 t t')" | ir_sym: "ir_hoare (flip P) c c' (flip Q) \ ir_hoare P c' c Q" section "Simple Derived Rules" lemma meh_simp[simp]: "(SKIP, s') \ t' = (s' = t')" by auto lemma ir_pre: "ir_hoare P c c' Q \ (\s s'. P s s' \ P' s s') \ ir_hoare P' c c' Q" by(erule ir_conseq, blast+) lemma ir_post: "ir_hoare P c c' Q \ (\s s'. Q' s s' \ Q s s') \ ir_hoare P c c' Q'" by(erule ir_conseq, blast+) section "Soundness" lemma Skip_ir_valid[intro]: "(\t t'. Q t t' \ \s'. P t s' \ (c', s') \ t') \ ir_valid P SKIP c' Q" by(auto simp: ir_valid_def) lemma sym_ir_valid[intro]: "ir_valid (flip P) c' c (flip Q) \ ir_valid P c c' Q" by(fastforce simp: ir_valid_def flip_def) lemma If_True_ir_valid[intro]: "ir_valid (\a c. P a c \ bval b a) c\<^sub>1 c' Q \ ir_valid P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q" by(fastforce simp: ir_valid_def) lemma If_False_ir_valid[intro]: "ir_valid (\a c. P a c \ \ bval b a) c\<^sub>2 c' Q \ ir_valid P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q" by(fastforce simp: ir_valid_def) lemma Seq1_ir_valid[intro]: "ir_valid P c c' Q \ ir_valid Q d SKIP R \ ir_valid P (c;; d) c' R" by(fastforce simp: ir_valid_def) lemma Seq2_ir_valid[intro]: "ir_valid P c SKIP Q \ ir_valid Q d c' R \ ir_valid P (c;; d) c' R" by(fastforce simp: ir_valid_def) lemma Seq_ir_valid[intro]: "ir_valid P c c' Q \ ir_valid Q d d' R \ ir_valid P (c;; d) (c';; d') R" by(fastforce simp: ir_valid_def) lemma Assign_blah[intro]: "t x = aval e (t(x := v)) \ (x ::= e, t(x := v)) \ t" using Assign by (metis fun_upd_triv fun_upd_upd) lemma Assign_ir_valid[intro]: "ir_valid (\t t'. \ v. P (t(x := v)) t' \ (t x) = aval e (t(x := v))) SKIP c' Q \ ir_valid P (Assign x e) c' Q" by(fastforce simp: ir_valid_def) lemma While_False_ir_valid[intro]: "ir_valid (\s s'. P s s' \ \ bval b s) SKIP c' Q \ ir_valid P (WHILE b DO c) c' Q" by(fastforce simp: ir_valid_def) lemma While_True_ir_valid[intro]: "ir_valid (\s s'. P s s' \ bval b s) (Seq c (WHILE b DO c)) c' Q \ ir_valid P (WHILE b DO c) c' Q" by(clarsimp simp: ir_valid_def, blast) lemma While_backwards_frontier_ir_valid': assumes asm: "\n. \t t'. P (k + Suc n) t t' \ (\s. P (k + n) s t' \ bval b s \ (c, s) \ t)" assumes last: "\t t'. Q t t' \ (\s s'. (\n. P (k + n) s s') \ (WHILE b DO c, s) \ t \ (c', s') \ t')" assumes post: "Q t t'" shows "\s s'. P k s s' \ (WHILE b DO c, s) \ t \ (c', s') \ t'" proof - from post last obtain s s' n where "P (k + n) s s'" "(WHILE b DO c, s) \ t" "(c', s') \ t'" by blast with asm show ?thesis proof(induction n arbitrary: k t t') case 0 then show ?case by (metis WhileFalse WhileTrue add.right_neutral) next case (Suc n) from Suc obtain r r' where final_iteration: "P (Suc k) r r'" "(WHILE b DO c, r) \ t" "(c', r') \ t'" by (metis add_Suc_shift) from final_iteration(1) obtain q q' where "P k q r' \ bval b q \ (c, q) \ r" by (metis Nat.add_0_right Suc.prems(1) plus_1_eq_Suc semiring_normalization_rules(24)) with final_iteration show ?case by blast qed qed lemma While_backwards_frontier_ir_valid[intro]: "(\n. ir_valid (\ s s'. P n s s' \ bval b s) c SKIP (P (Suc n))) \ ir_valid (\s s'. \n. P n s s') (WHILE b DO c) c' Q \ ir_valid (P 0) (WHILE b DO c) c' Q" by(auto simp: meh_simp ir_valid_def intro: While_backwards_frontier_ir_valid') lemma conseq_ir_valid: "ir_valid P c c' Q \ (\s s'. P s s' \ P' s s') \ (\s s'. Q' s s' \ Q s s') \ ir_valid P' c c' Q'" by(clarsimp simp: ir_valid_def, blast) lemma disj_ir_valid[intro]: "ir_valid P\<^sub>1 c c' Q\<^sub>1 \ ir_valid P\<^sub>2 c c' Q\<^sub>2 \ ir_valid (\s s'. P\<^sub>1 s s' \ P\<^sub>2 s s') c c' (\ t t'. Q\<^sub>1 t t' \ Q\<^sub>2 t t')" by(fastforce simp: ir_valid_def) theorem soundness: "ir_hoare P c c' Q \ ir_valid P c c' Q" apply(induction rule: ir_hoare.induct) apply(blast intro!: Skip_ir_valid) by (blast intro: conseq_ir_valid)+ section "Completeness" lemma ir_Skip_Skip[intro]: "ir_hoare P SKIP SKIP P" by(rule ir_Skip, simp) lemma ir_hoare_Skip_Skip[simp]: "ir_hoare P SKIP SKIP Q = (\s s'. Q s s' \ P s s')" using soundness ir_valid_def meh_simp ir_conseq ir_Skip by metis lemma ir_valid_Seq1: "ir_valid P (c1;; c2) c' Q \ ir_valid P c1 c' (\t t'. \s s'. P s s' \ (c1,s) \ t \ (c',s') \ t' \ (\u. (c2,t) \ u \ Q u t'))" by(auto simp: ir_valid_def) lemma ir_valid_Seq1': "ir_valid P (c1;; c2) c' Q \ ir_valid (\t t'. \s s'. P s s' \ (c1,s) \ t \ (c',s') \ t' \ (\u. (c2,t) \ u \ Q u t')) c2 SKIP Q" by(clarsimp simp: ir_valid_def, meson SeqE) lemma ir_valid_track_history: "ir_valid P c c' Q \ ir_valid P c c' (\t t'. Q s s' \ (\s s'. P s s' \ (c,s) \ t \ (c',s') \ t'))" by(auto simp: ir_valid_def) lemma ir_valid_If: "ir_valid (\s s'. P s s') (IF b THEN c1 ELSE c2) c' Q \ ir_valid (\s s'. P s s' \ bval b s) c1 c' (\t t'. Q t t' \ (\s s'. P s s' \ (c1,s) \ t \ (c',s') \ t' \ bval b s)) \ ir_valid (\s s'. P s s' \ \ bval b s) c2 c' (\t t'. Q t t' \ (\s s'. P s s' \ (c2,s) \ t \ (c',s') \ t' \ \ bval b s))" by(clarsimp simp: ir_valid_def, blast) text \ Inspired by the ``@{text "p(n) = {\ | you can get back from \ to some state in p by executing C backwards n times}"}'' part of OHearn~\<^cite>\"OHearn_19"\. \ primrec get_back where "get_back P b c 0 = (\t t'. P t t')" | "get_back P b c (Suc n) = (\t t'. \s. (c,s) \ t \ bval b s \ get_back P b c n s t')" (* Currently not used anywhere *) lemma ir_valid_get_back: "ir_valid (get_back P b c (Suc k)) (WHILE b DO c) c' Q \ ir_valid (get_back P b c k) (WHILE b DO c) c' Q" proof(induct k) case 0 then show ?case by(clarsimp simp: ir_valid_def, blast) next case (Suc k) - then show ?case using WhileTrue get_back.simps(2) ir_valid_def by smt + then show ?case using WhileTrue get_back.simps(2) ir_valid_def by (smt (verit)) qed (* both this an the next one get used in the completeness proof *) lemma ir_valid_While1: "ir_valid (get_back P b c k) (WHILE b DO c) c' Q \ (ir_valid (\s s'. get_back P b c k s s' \ bval b s) c SKIP (\t t'. get_back P b c (Suc k) t t' \ (\u u'. (WHILE b DO c,t) \ u \ (c',t') \ u' \ Q u u')))" proof (subst ir_valid_def, clarsimp) fix t t' s u u' assume "ir_valid (get_back P b c k) (WHILE b DO c) c' Q" "(WHILE b DO c, t) \ u" "(c, s) \ t" "(c', t') \ u'" "Q u u'" "bval b s" "get_back P b c k s t'" thus "\s. get_back P b c k s t' \ bval b s \ (c, s) \ t" proof(induction k arbitrary: t t' s u u') case 0 then show ?case by auto next case (Suc k) show ?case using Suc.prems(3) Suc.prems(6) Suc.prems(7) by blast qed qed lemma ir_valid_While3: "ir_valid (get_back P b c k) (WHILE b DO c) c' Q \ (ir_valid (\s s'. get_back P b c k s s' \ bval b s) c c' (\t t'. (\s'. (c',s') \ t' \ get_back P b c (Suc k) t s' \ (\u. (WHILE b DO c,t) \ u \ Q u t'))))" apply(subst ir_valid_def, clarsimp) proof - fix t t' s' s u assume "ir_valid (get_back P b c k) (WHILE b DO c) c' Q" "(WHILE b DO c, t) \ u" "(c, s) \ t" "(c', s') \ t'" "Q u t'" "bval b s" "get_back P b c k s s'" thus "\s s'. get_back P b c k s s' \ bval b s \ (c, s) \ t \ (c',s') \ t'" proof(induction k arbitrary: t t' s' s u) case 0 then show ?case by auto next case (Suc k) show ?case using Suc.prems(3) Suc.prems(4) Suc.prems(6) Suc.prems(7) by blast qed qed (* not used anywhere *) lemma ir_valid_While2: "ir_valid P (WHILE b DO c) c' Q \ ir_valid (\s s'. P s s' \ \ bval b s) SKIP c' (\t t'. Q t t' \ (\s'. (c',s') \ t' \ P t s' \ \ bval b t))" by(clarsimp simp: ir_valid_def, blast) lemma assign_upd_blah: "(\a. if a = x1 then s x1 else (s(x1 := aval x2 s)) a) = s" by(rule ext, auto) lemma Assign_complete: assumes v: "ir_valid P (x1 ::= x2) c' Q" assumes q: "Q t t'" shows "\s'. (\v. P (t(x1 := v)) s' \ t x1 = aval x2 (t(x1 := v))) \ (c', s') \ t'" proof - from v and q obtain s s' where a: "P s s' \ (x1 ::= x2,s) \ t \ (c',s') \ t'" - using ir_valid_def by smt + using ir_valid_def by (smt (verit)) hence "P (\a. if a = x1 then s x1 else (s(x1 := aval x2 s)) a) s' \ aval x2 s = aval x2 (s(x1 := s x1))" using assign_upd_blah by simp thus ?thesis using assign_upd_blah a by (metis AssignE fun_upd_same fun_upd_triv fun_upd_upd) qed lemmas ir_Skip_sym = ir_sym[OF ir_Skip, simplified flip_def] theorem completeness: "ir_valid P c c' Q \ ir_hoare P c c' Q" proof(induct c arbitrary: P c' Q) case SKIP show ?case apply(rule ir_Skip) using SKIP by(fastforce simp: ir_valid_def) next case (Assign x1 x2) show ?case apply(rule ir_Assign) apply(rule ir_Skip) using Assign_complete Assign by blast next case (Seq c1 c2) have a: "ir_hoare P c1 c' (\t t'. \s s'. P s s' \ (c1, s) \ t \ (c', s') \ t' \ (\u. (c2, t) \ u \ Q u t'))" using ir_valid_Seq1 Seq by blast show ?case apply(rule ir_Seq1) apply (blast intro: a) apply(rule ir_Skip_sym) by(metis SeqE ir_valid_def Seq) next case (If x1 c1 c2) have t: "ir_hoare (\s s'. P s s' \ bval x1 s) c1 c' (\t t'. Q t t' \ (\s s'. P s s' \ (c1, s) \ t \ (c', s') \ t' \ bval x1 s))" and f: " ir_hoare (\s s'. P s s' \ \ bval x1 s) c2 c' (\t t'. Q t t' \ (\s s'. P s s' \ (c2, s) \ t \ (c', s') \ t' \ \ bval x1 s))" using ir_valid_If If by blast+ show ?case (* consider both cases of the if via conseq, disj, then _True and _False *) apply(rule ir_conseq) apply(rule ir_disj) apply(rule ir_If_True,fastforce intro: t) apply(rule ir_If_False, fastforce intro: f) apply blast - by (smt IfE ir_valid_def If) + by (smt (verit) IfE ir_valid_def If) next case (While x1 c) have a: "\n. ir_hoare (\s s'. get_back P x1 c n s s' \ bval x1 s) c SKIP (get_back P x1 c (Suc n))" - using ir_valid_While1 While - by (smt get_back.simps(2) ir_valid_def meh_simp) + using ir_valid_While1 While + by (smt (verit, ccfv_threshold) get_back.simps(2) ir_Skip_sym) have b: "ir_hoare (\s s'. P s s' \ bval x1 s) c c' (\t t'. \s'. (c', s') \ t' \ (\s. (c, s) \ t \ bval x1 s \ P s s') \ (\u. (WHILE x1 DO c, t) \ u \ Q u t'))" using ir_valid_While3[where k=0,simplified] While by blast have gb: "\t t'. Q t t' \ (\s'. (c', s') \ t' \ P t s' \ \ bval x1 t) \ \s'. ((\n. get_back P x1 c n t s') \ \ bval x1 t) \ (c', s') \ t'" by (meson get_back.simps(1)) show ?case (* use the frontier rule much as in OHearn POPL *) apply(rule ir_conseq) apply(rule_tac P="get_back P x1 c" in ir_While_backwards_frontier) apply(blast intro: a) (* consider both cases of the While via conseq, disj, then _True and _False *) apply(rule ir_conseq) apply(rule ir_disj) apply(rule_tac P="\s s'. \n. get_back P x1 c n s s'" and Q="(\t t'. Q t t' \ (\s'. (c', s') \ t' \ P t s' \ \ bval x1 t))" in ir_While_False) apply(rule ir_Skip, blast intro: gb) apply(rule ir_While_True) apply(rule ir_Seq1[OF b]) apply(rule ir_Skip_sym) apply(fastforce) apply (metis get_back.simps(1)) apply assumption apply simp by (metis While.prems WhileE ir_valid_def) qed section "A Decomposition Principle: Proofs via Under-Approximate Hoare Logic" text \ We show the under-approximate analogue holds for Beringer's~\<^cite>\"Beringer_11"\ decomposition principle for over-approximate relational logic. \ definition decomp :: "rassn \ com \ com \ rassn \ rassn" where "decomp P c c' Q \ \t s'. \s t'. P s s' \ (c,s) \ t \ (c',s') \ t' \ Q t t'" lemma ir_valid_decomp1: "ir_valid P c c' Q \ ir_valid P c SKIP (decomp P c c' Q) \ ir_valid (decomp P c c' Q) SKIP c' Q" by(fastforce simp: ir_valid_def meh_simp decomp_def) lemma ir_valid_decomp2: "ir_valid P c SKIP R \ ir_valid R SKIP c' Q \ ir_valid P c c' Q" by(fastforce simp: ir_valid_def meh_simp decomp_def) lemma ir_valid_decomp: "ir_valid P c c' Q = (ir_valid P c SKIP (decomp P c c' Q) \ ir_valid (decomp P c c' Q) SKIP c' Q)" using ir_valid_decomp1 ir_valid_decomp2 by blast text \ Completeness with soundness means we can prove proof rules about @{term ir_hoare} in terms of @{term ir_valid}. \ lemma ir_to_Skip: "ir_hoare P c c' Q = (ir_hoare P c SKIP (decomp P c c' Q) \ ir_hoare (decomp P c c' Q) SKIP c' Q)" using soundness completeness ir_valid_decomp by meson text \ O'Hearn's under-approximate Hoare triple, for the ``ok'' case (since that is the only case we have) This is also likely the same as from the "Reverse Hoare Logic" paper (SEFM). \ type_synonym assn = "state \ bool" definition ohearn :: "assn \ com \ assn \ bool" where "ohearn P c Q \ (\t. Q t \ (\s. P s \ (c,s) \ t))" lemma fold_ohearn1: "ir_valid P c SKIP Q = (\t'. ohearn (\t. P t t') c (\t. Q t t'))" by(fastforce simp add: ir_valid_def ohearn_def) lemma fold_ohearn2: "ir_valid P SKIP c' Q = (\t. ohearn (P t) c' (Q t))" by(simp add: ir_valid_def ohearn_def) theorem relational_via_hoare: "ir_hoare P c c' Q = ((\t'. ohearn (\t. P t t') c (\t. decomp P c c' Q t t')) \ (\t. ohearn (decomp P c c' Q t) c' (Q t)))" proof - have a: "\P c c' Q. ir_hoare P c c' Q = ir_valid P c c' Q" using soundness completeness by auto show ?thesis using ir_to_Skip a fold_ohearn1 fold_ohearn2 by metis qed section "Deriving Proof Rules from Completeness" text \ Note that we can more easily derive proof rules sometimes by appealing to the corresponding properties of @{term ir_valid} than from the proof rules directly. We see more examples of this later on when we consider examples. \ lemma ir_Seq2: "ir_hoare P c SKIP Q \ ir_hoare Q d c' R \ ir_hoare P (Seq c d) c' R" using soundness completeness Seq2_ir_valid by metis lemma ir_Seq: "ir_hoare P c c' Q \ ir_hoare Q d d' R \ ir_hoare P (Seq c d) (Seq c' d') R" using soundness completeness Seq_ir_valid by metis section "Examples" subsection "Some Derived Proof Rules" text \ First derive some proof rules -- here not by appealing to completeness but just using the existing rules \ lemma ir_If_True_False: "ir_hoare (\s s'. P s s' \ bval b s \ \ bval b' s') c\<^sub>1 c\<^sub>2' Q \ ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') Q" apply(rule ir_If_True) apply(rule ir_sym) apply(rule ir_If_False) apply(rule ir_sym) by(simp add: flip_def) lemma ir_Assign_Assign: "ir_hoare P (x ::= e) (x' ::= e') (\t t'. \v v'. P (t(x := v)) (t'(x' := v')) \ t x = aval e (t(x := v)) \ t' x' = aval e' (t'(x' := v')))" apply(rule ir_Assign) apply(rule ir_sym) apply(rule ir_Assign) by(simp add: flip_def, auto) subsection "prog1" text \ A tiny insecure program. Note that we only need to reason on one path through this program to detect that it is insecure. \ abbreviation low_eq :: rassn where "low_eq s s' \ s ''low'' = s' ''low''" abbreviation low_neq :: rassn where "low_neq s s' \ \ low_eq s s'" definition prog1 :: com where "prog1 \ (IF (Less (N 0) (V ''x'')) THEN (Assign ''low'' (N 1)) ELSE (Assign ''low'' (N 0)))" text \ We prove that @{term prog1} is definitely insecure. To do that, we need to find some non-empty post-relation that implies insecurity. The following property encodes the idea that the post-relation is non-empty, i.e. represents a feasible pair of execution paths. \ definition nontrivial :: "rassn \ bool" where "nontrivial Q \ (\t t'. Q t t')" text \ Note the property we prove here explicitly encodes the fact that the postcondition can be anything that implies insecurity, i.e. implies @{term low_neq}. In particular we should not necessarily expect it to cover the entirely of all states that satisfy @{term low_neq}. Also note that we also have to prove that the postcondition is non-trivial. This is necessary to make sure that the violation we find is not an infeasible path. \ lemma prog1: "\Q. ir_hoare low_eq prog1 prog1 Q \ (\s s'. Q s s' \ low_neq s s') \ nontrivial Q" apply(rule exI) apply(rule conjI) apply(simp add: prog1_def) apply(rule ir_If_True_False) apply(rule ir_Assign_Assign) apply(rule conjI) apply auto[1] apply(clarsimp simp: nontrivial_def) apply(rule_tac x="\v. 1" in exI) apply simp apply(rule_tac x="\v. 0" in exI) by auto subsection "More Derived Proof Rules for Examples" definition BEq :: "aexp \ aexp \ bexp" where "BEq a b \ And (Less a (Plus b (N 1))) (Less b (Plus a (N 1)))" lemma BEq_aval[simp]: "bval (BEq a b) s = ((aval a s) = (aval b s))" by(auto simp add: BEq_def) lemma ir_If_True_True: "ir_hoare (\s s'. P s s' \ bval b s \ bval b' s') c\<^sub>1 c\<^sub>1' Q\<^sub>1 \ ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') (\t t'. Q\<^sub>1 t t')" by(simp add: ir_If_True ir_sym flip_def) lemma ir_If_False_False: "ir_hoare (\s s'. P s s' \ \ bval b s \ \ bval b' s') c\<^sub>2 c\<^sub>2' Q\<^sub>2 \ ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') (\t t'. Q\<^sub>2 t t')" by(simp add: ir_If_False ir_sym flip_def) lemma ir_If': "ir_hoare (\s s'. P s s' \ bval b s \ bval b' s') c\<^sub>1 c\<^sub>1' Q\<^sub>1 \ ir_hoare (\s s'. P s s' \ \ bval b s \ \ bval b' s') c\<^sub>2 c\<^sub>2' Q\<^sub>2 \ ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') (\t t'. Q\<^sub>1 t t' \ Q\<^sub>2 t t')" apply(rule ir_pre) apply(rule ir_disj) apply(rule ir_If_True_True) apply assumption apply(rule ir_If_False_False) apply assumption apply blast done lemma ir_While_triv: "ir_hoare (\s s'. P s s' \ \ bval b s \ \ bval b' s') SKIP SKIP Q\<^sub>2 \ ir_hoare P (WHILE b DO c) (WHILE b' DO c') (\s s'. (Q\<^sub>2 s s'))" by(simp add: ir_While_False ir_sym flip_def) lemma ir_While': "ir_hoare (\s s'. P s s' \ bval b s \ bval b' s') (c;;WHILE b DO c) (c';;WHILE b' DO c') Q\<^sub>1 \ ir_hoare (\s s'. P s s' \ \ bval b s \ \ bval b' s') SKIP SKIP Q\<^sub>2 \ ir_hoare P (WHILE b DO c) (WHILE b' DO c') (\s s'. (Q\<^sub>1 s s' \ Q\<^sub>2 s s'))" apply(rule ir_pre) apply(rule ir_disj) apply(rule ir_While_True) apply(rule ir_sym) apply(simp add: flip_def) apply(rule ir_While_True) apply(rule ir_sym) apply(simp add: flip_def) apply(rule ir_While_triv) apply assumption apply simp done subsection "client0" definition low_eq_strong where "low_eq_strong s s' \ (s ''high'' \ s' ''high'') \ low_eq s s'" lemma low_eq_strong_upd[simp]: "var \ ''high'' \ var \ ''low'' \ low_eq_strong(s(var := v)) (s'(var := v')) = low_eq_strong s s'" by(auto simp: low_eq_strong_def) text \ A variation on client0 from O'Hearn~\<^cite>\"OHearn_19"\: how to reason about loops via a single unfolding \ definition client0 :: com where "client0 \ (Assign ''x'' (N 0);; (While (Less (N 0) (V ''n'')) ((Assign ''x'' (Plus (V ''x'') (V ''n'')));; (Assign ''n'' (V ''nondet''))));; (If (BEq (V ''x'') (N 2000000)) (Assign ''low'' (V ''high'')) SKIP))" lemma client0: "\Q. ir_hoare low_eq client0 client0 Q \ (\s s'. Q s s' \ low_neq s s') \ nontrivial Q" apply(rule exI, rule conjI, simp add: client0_def) apply(rule_tac P=low_eq_strong in ir_pre) apply(rule ir_Seq) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply clarsimp apply(rule ir_While') apply clarsimp apply(rule ir_Seq) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply(rule ir_Assign_Assign) apply clarsimp apply(rule ir_While_triv) apply clarsimp apply assumption apply clarsimp apply assumption apply(rule ir_If_True_True) apply(rule ir_Assign_Assign) apply(fastforce simp: low_eq_strong_def) apply(rule conjI) apply(clarsimp simp: low_eq_strong_def split: if_splits) (* ugh having to manually do constraint solving here... *) apply(clarsimp simp: low_eq_strong_def nontrivial_def) apply(rule_tac x="\v. if v = ''x'' then 2000000 else if v = ''high'' then 1 else if v = ''n'' then -1 else if v = ''nondet'' then -1 else if v = ''low'' then 1 else undefined" in exI) apply(rule_tac x="\v. if v = ''x'' then 2000000 else if v = ''high'' then 0 else if v = ''n'' then -1 else if v = ''nondet'' then -1 else if v = ''low'' then 0 else undefined" in exI) apply clarsimp done (* Not needed? *) lemma ir_While_backwards: "(\n. ir_hoare (\ s s'. P n s s' \ bval b s) c SKIP (P (Suc n))) \ ir_hoare (\s s'. \n. P n s s' \ \ bval b s) SKIP c' Q \ ir_hoare (P 0) (WHILE b DO c) c' Q" apply(rule ir_While_backwards_frontier) apply assumption apply(rule ir_While_False) apply auto done subsection "Derive a variant of the backwards variant rule" text \Here we appeal to completeness again to derive this rule from the corresponding property about @{term ir_valid}.\ subsection "A variant of the frontier rule" text \ Agin we derive this rule by appealing to completeness and the corresponding property of @{term ir_valid} \ lemma While_backwards_frontier_both_ir_valid': assumes asm: "\n. \t t'. P (k + Suc n) t t' \ (\s s'. P (k + n) s s' \ bval b s \ bval b' s' \ (c, s) \ t \ (c', s') \ t')" assumes last: "\t t'. Q t t' \ (\s s'. (\n. P (k + n) s s') \ (WHILE b DO c, s) \ t \ (WHILE b' DO c', s') \ t')" assumes post: "Q t t'" shows "\s s'. P k s s' \ (WHILE b DO c, s) \ t \ (WHILE b' DO c', s') \ t'" proof - from post last obtain s s' n where "P (k + n) s s'" "(WHILE b DO c, s) \ t" "(WHILE b' DO c', s') \ t'" by blast with asm show ?thesis proof(induction n arbitrary: k t t') case 0 then show ?case by (metis WhileFalse WhileTrue add.right_neutral) next case (Suc n) from Suc obtain r r' where final_iteration: "P (Suc k) r r'" "(WHILE b DO c, r) \ t" "(WHILE b' DO c', r') \ t'" by (metis add_Suc_shift) from final_iteration(1) obtain q q' where "P k q q' \ bval b q \ bval b' q' \ (c, q) \ r \ (c', q') \ r'" by (metis Nat.add_0_right Suc.prems(1) plus_1_eq_Suc semiring_normalization_rules(24)) with final_iteration show ?case by blast qed qed lemma While_backwards_frontier_both_ir_valid[intro]: "(\n. ir_valid (\ s s'. P n s s' \ bval b s \ bval b' s') c c' (P (Suc n))) \ ir_valid (\s s'. \n. P n s s') (WHILE b DO c) (WHILE b' DO c') Q \ ir_valid (P 0) (WHILE b DO c) (WHILE b' DO c') (\s s'. Q s s')" by(auto simp: ir_valid_def intro: While_backwards_frontier_both_ir_valid') lemma ir_While_backwards_frontier_both: "\\n. ir_hoare (\s s'. P n s s' \ bval b s \ bval b' s') c c' (P (Suc n)); ir_hoare (\s s'. \n. P n s s') (WHILE b DO c) (WHILE b' DO c') Q\ \ ir_hoare (P 0) (WHILE b DO c) (WHILE b' DO c') (\s s'. Q s s')" using soundness completeness While_backwards_frontier_both_ir_valid by auto text \ The following rule then follows easily as a special case \ lemma ir_While_backwards_both: "(\n. ir_hoare (\ s s'. P n s s' \ bval b s \ bval b' s') c c' (P (Suc n))) \ ir_hoare (P 0) (WHILE b DO c) (WHILE b' DO c') (\s s'. \n. P n s s' \ \ bval b s \ \ bval b' s')" apply(rule ir_While_backwards_frontier_both) apply blast by(simp add: ir_While_False ir_sym flip_def) subsection "client1" text \ An example roughly equivalent to cient1 from O'Hearn~\<^cite>\"OHearn_19"\0 In particular we use the backwards variant rule to reason about the loop. \ definition client1 :: com where "client1 \ (Assign ''x'' (N 0);; (While (Less (V ''x'') (V ''n'')) ((Assign ''x'' (Plus (V ''x'') (N 1)))));; (If (BEq (V ''x'') (N 2000000)) (Assign ''low'' (V ''high'')) SKIP))" lemma client1: "\Q. ir_hoare low_eq client1 client1 Q \ (\s s'. Q s s' \ low_neq s s') \ nontrivial Q" apply(rule exI, rule conjI, simp add: client1_def) apply(rule_tac P=low_eq_strong in ir_pre) apply(rule ir_Seq) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply clarsimp apply(rule ir_pre) apply(rule ir_While_backwards_both[where P="\n s s'. low_eq_strong s s' \ s ''x'' = int n \ s' ''x'' = int n \ int n \ s ''n'' \ int n \ s' ''n''"]) apply(rule ir_post) apply(rule ir_Assign_Assign) apply clarsimp apply clarsimp apply(rule ir_If_True_True) apply(rule ir_Assign_Assign) apply(fastforce simp: low_eq_strong_def) apply(rule conjI) apply(clarsimp simp: low_eq_strong_def split: if_splits) apply clarsimp (* ugh having to manually do constraint solving here... *) apply(clarsimp simp: low_eq_strong_def nontrivial_def) apply(rule_tac x="\v. if v = ''x'' then 2000000 else if v = ''high'' then 1 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 1 else undefined" in exI) apply(rule_tac x="\v. if v = ''x'' then 2000000 else if v = ''high'' then 0 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 0 else undefined" in exI) apply clarsimp done subsection "client2" text \ An example akin to client2 from O'Hearn~\<^cite>\"OHearn_19"\. Note that this example is carefully written to show use of the frontier rule first to reason up to the broken loop iteration, and then we unfold the loop at that point to reason about the broken iteration, and then use the plain backwards variant rule to reason over the remainder of the loop. \ definition client2 :: com where "client2 \ (Assign ''x'' (N 0);; (While (Less (V ''x'') (N 4000000)) ((Assign ''x'' (Plus (V ''x'') (N 1)));; (If (BEq (V ''x'') (N 2000000)) (Assign ''low'' (V ''high'')) SKIP)) ) )" lemma client2: "\Q. ir_hoare low_eq client2 client2 Q \ (\s s'. Q s s' \ low_neq s s') \ nontrivial Q" apply(rule exI, rule conjI, simp add: client2_def) apply(rule_tac P=low_eq_strong in ir_pre) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply clarsimp apply(rule ir_pre) apply(rule ir_While_backwards_frontier_both[where P="\n s s'. low_eq_strong s s' \ s ''x'' = int n \ s' ''x'' = int n \ s ''x'' \ 0 \ s ''x'' \ 2000000 - 1 \ s' ''x'' \ 0 \ s' ''x'' \ 2000000 - 1"]) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply clarsimp apply(rule ir_post) apply(rule ir_If') apply(rule ir_Assign_Assign) apply(rule ir_Skip_Skip) apply clarsimp apply clarsimp apply(rule ir_While') apply clarsimp apply(rule ir_Seq) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply(rule ir_If_True_True) apply(rule ir_Assign_Assign) apply clarsimp apply(rule ir_pre) apply(rule ir_While_backwards_both[where P="\n s s'. s ''x'' = 2000000 + int n \ s' ''x'' = 2000000 + int n \ s ''x'' \ 2000000 \ s ''x'' \ 4000000 \ s' ''x'' \ 2000000 \ s' ''x'' \ 4000000 \ s ''low'' = s ''high'' \ s' ''low'' = s' ''high'' \ s ''high'' \ s' ''high''"]) apply(rule ir_Seq) apply(rule ir_Assign_Assign) apply(rule ir_If_False_False) apply fastforce apply (fastforce simp: low_eq_strong_def) apply fastforce apply(fastforce simp: low_eq_strong_def) apply(fastforce simp: low_eq_strong_def) apply(rule conjI) apply(clarsimp simp: low_eq_strong_def split: if_splits) apply clarsimp (* ugh having to manually do constraint solving here... *) apply(clarsimp simp: low_eq_strong_def nontrivial_def) apply(rule_tac x="\v. if v = ''x'' then 4000000 else if v = ''high'' then 1 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 1 else undefined" in exI) apply(rule_tac x="\v. if v = ''x'' then 4000000 else if v = ''high'' then 0 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 0 else undefined" in exI) apply clarsimp done end diff --git a/thys/SCC_Bloemen_Sequential/SCC_Bloemen_Sequential.thy b/thys/SCC_Bloemen_Sequential/SCC_Bloemen_Sequential.thy --- a/thys/SCC_Bloemen_Sequential/SCC_Bloemen_Sequential.thy +++ b/thys/SCC_Bloemen_Sequential/SCC_Bloemen_Sequential.thy @@ -1,3439 +1,3438 @@ section \Overview\ text \ Computing the maximal strongly connected components (SCCs) of a finite directed graph is a celebrated problem in the theory of graph algorithms. Although Tarjan's algorithm~\<^cite>\"tarjan:depth-first"\ is perhaps the best-known solution, there are many others. In his PhD thesis, Bloemen~\<^cite>\"bloemen:strong"\ presents an algorithm that is itself based on earlier algorithms by Munro~\<^cite>\"munro:efficient"\ and Dijkstra~\<^cite>\"dijkstra:finding"\. Just like these algorithms, Bloemen's solution is based on enumerating SCCs in a depth-first traversal of the graph. Gabow's algorithm that has already been formalized in Isabelle~\<^cite>\"lammich:gabow"\ also falls into this category of solutions. Nevertheless, Bloemen goes on to present a parallel variant of the algorithm suitable for execution on multi-core processors, based on clever data structures that minimize locking. In the following, we encode the sequential version of the algorithm in the proof assistant Isabelle/HOL, and prove its correctness. Bloemen's thesis briefly and informally explains why the algorithm is correct. Our proof expands on these arguments, making them completely formal. The encoding is based on a direct representation of the algorithm as a pair of mutually recursive functions; we are not aiming at extracting executable code. \ theory SCC_Bloemen_Sequential imports Main begin text \ The record below represents the variables of the algorithm. Most variables correspond to those used in Bloemen's presentation. Thus, the variable @{text \} associates to every node the set of nodes that have already been determined to be part of the same SCC. A core invariant of the algorithm will be that this mapping represents equivalence classes of nodes: for all nodes @{text v} and @{text w}, we maintain the relationship @{text "v \ \ w \ \ v = \ w."} In an actual implementation of this algorithm, this variable could conveniently be represented by a union-find structure. Variable @{text stack} holds the list of roots of these (not yet maximal) SCCs, in depth-first order, @{text visited} and @{text explored} represent the nodes that have already been seen, respectively that have been completely explored, by the algorithm, and @{text sccs} is the set of maximal SCCs that the algorithm has found so far. Additionally, the record holds some auxiliary variables that are used in the proof of correctness. In particular, @{text root} denotes the node on which the algorithm was called, @{text cstack} represents the call stack of the recursion of function @{text dfs}, and @{text vsuccs} stores the successors of each node that have already been visited by the function @{text dfss} that loops over all successors of a given node. \ record 'v env = root :: "'v" \ :: "'v \ 'v set" explored :: "'v set" visited :: "'v set" vsuccs :: "'v \ 'v set" sccs :: "'v set set" stack :: "'v list" cstack :: "'v list" text \ The algorithm is initially called with an environment that initializes the root node and trivializes all other components. \ definition init_env where "init_env v = \ root = v, \ = (\u. {u}), explored = {}, visited = {}, vsuccs = (\u. {}), sccs = {}, stack = [], cstack = [] \" \ \Make the simplifier expand let-constructions automatically.\ declare Let_def[simp] section \Auxiliary lemmas about lists\ text \ We use the precedence order on the elements that appear in a list. In particular, stacks are represented as lists, and a node @{text x} precedes another node @{text y} on the stack if @{text x} was pushed on the stack later than @{text y}. \ definition precedes ("_ \ _ in _" [100,100,100] 39) where "x \ y in xs \ \l r. xs = l @ (x # r) \ y \ set (x # r)" lemma precedes_mem: assumes "x \ y in xs" shows "x \ set xs" "y \ set xs" using assms unfolding precedes_def by auto lemma head_precedes: assumes "y \ set (x # xs)" shows "x \ y in (x # xs)" using assms unfolding precedes_def by force lemma precedes_in_tail: assumes "x \ z" shows "x \ y in (z # zs) \ x \ y in zs" using assms unfolding precedes_def by (auto simp: Cons_eq_append_conv) lemma tail_not_precedes: assumes "y \ x in (x # xs)" "x \ set xs" shows "x = y" using assms unfolding precedes_def by (metis Cons_eq_append_conv Un_iff list.inject set_append) lemma split_list_precedes: assumes "y \ set (ys @ [x])" shows "y \ x in (ys @ x # xs)" using assms unfolding precedes_def by (metis append_Cons append_assoc in_set_conv_decomp rotate1.simps(2) set_ConsD set_rotate1) lemma precedes_refl [simp]: "(x \ x in xs) = (x \ set xs)" proof assume "x \ x in xs" thus "x \ set xs" by (simp add: precedes_mem) next assume "x \ set xs" from this[THEN split_list] show "x \ x in xs" unfolding precedes_def by auto qed lemma precedes_append_left: assumes "x \ y in xs" shows "x \ y in (ys @ xs)" using assms unfolding precedes_def by (metis append.assoc) lemma precedes_append_left_iff: assumes "x \ set ys" shows "x \ y in (ys @ xs) \ x \ y in xs" (is "?lhs = ?rhs") proof assume "?lhs" then obtain l r where lr: "ys @ xs = l @ (x # r)" "y \ set (x # r)" unfolding precedes_def by blast then obtain us where "(ys = l @ us \ us @ xs = x # r) \ (ys @ us = l \ xs = us @ (x # r))" by (auto simp: append_eq_append_conv2) thus ?rhs proof assume us: "ys = l @ us \ us @ xs = x # r" with assms have "us = []" by (metis Cons_eq_append_conv in_set_conv_decomp) with us lr show ?rhs unfolding precedes_def by auto next assume us: "ys @ us = l \ xs = us @ (x # r)" with \y \ set (x # r)\ show ?rhs unfolding precedes_def by blast qed next assume "?rhs" thus "?lhs" by (rule precedes_append_left) qed lemma precedes_append_right: assumes "x \ y in xs" shows "x \ y in (xs @ ys)" using assms unfolding precedes_def by force lemma precedes_append_right_iff: assumes "y \ set ys" shows "x \ y in (xs @ ys) \ x \ y in xs" (is "?lhs = ?rhs") proof assume ?lhs then obtain l r where lr: "xs @ ys = l @ (x # r)" "y \ set (x # r)" unfolding precedes_def by blast then obtain us where "(xs = l @ us \ us @ ys = x # r) \ (xs @ us = l \ ys = us @ (x # r))" by (auto simp: append_eq_append_conv2) thus ?rhs proof assume us: "xs = l @ us \ us @ ys = x # r" with \y \ set (x # r)\ assms show ?rhs unfolding precedes_def by (metis Cons_eq_append_conv Un_iff set_append) next assume us: "xs @ us = l \ ys = us @ (x # r)" with \y \ set (x # r)\ assms show ?rhs by auto \ \contradiction\ qed next assume ?rhs thus ?lhs by (rule precedes_append_right) qed text \ Precedence determines an order on the elements of a list, provided elements have unique occurrences. However, consider a list such as @{text "[2,3,1,2]"}: then $1$ precedes $2$ and $2$ precedes $3$, but $1$ does not precede $3$. \ lemma precedes_trans: assumes "x \ y in xs" and "y \ z in xs" and "distinct xs" shows "x \ z in xs" using assms unfolding precedes_def - by (smt Un_iff append.assoc append_Cons_eq_iff distinct_append - not_distinct_conv_prefix set_append split_list_last) + by (metis assms(2) disjoint_iff distinct_append precedes_append_left_iff precedes_mem(2)) lemma precedes_antisym: assumes "x \ y in xs" and "y \ x in xs" and "distinct xs" shows "x = y" proof - from \x \ y in xs\ \distinct xs\ obtain as bs where 1: "xs = as @ (x # bs)" "y \ set (x # bs)" "y \ set as" unfolding precedes_def by force from \y \ x in xs\ \distinct xs\ obtain cs ds where 2: "xs = cs @ (y # ds)" "x \ set (y # ds)" "x \ set cs" unfolding precedes_def by force from 1 2 have "as @ (x # bs) = cs @ (y # ds)" by simp then obtain zs where "(as = cs @ zs \ zs @ (x # bs) = y # ds) \ (as @ zs = cs \ x # bs = zs @ (y # ds))" (is "?P \ ?Q") by (auto simp: append_eq_append_conv2) then show ?thesis proof assume "?P" with \y \ set as\ show ?thesis by (cases "zs") auto next assume "?Q" with \x \ set cs\ show ?thesis by (cases "zs") auto qed qed section \Finite directed graphs\ text \ We represent a graph as an Isabelle locale that identifies a finite set of vertices (of some base type @{text "'v"}) and associates to each vertex its set of successor vertices. \ locale graph = fixes vertices :: "'v set" and successors :: "'v \ 'v set" assumes vfin: "finite vertices" and sclosed: "\x \ vertices. successors x \ vertices" context graph begin abbreviation edge where "edge x y \ y \ successors x" text \ We inductively define reachability of nodes in the graph. \ inductive reachable where reachable_refl[iff]: "reachable x x" | reachable_succ[elim]: "\edge x y; reachable y z\ \ reachable x z" lemma reachable_edge: "edge x y \ reachable x y" by auto lemma succ_reachable: assumes "reachable x y" and "edge y z" shows "reachable x z" using assms by induct auto lemma reachable_trans: assumes y: "reachable x y" and z: "reachable y z" shows "reachable x z" using assms by induct auto text \ In order to derive a ``reverse'' induction rule for @{const "reachable"}, we define an alternative reachability predicate and prove that the two coincide. \ inductive reachable_end where re_refl[iff]: "reachable_end x x" | re_succ: "\reachable_end x y; edge y z\ \ reachable_end x z" lemma succ_re: assumes y: "edge x y" and z: "reachable_end y z" shows "reachable_end x z" using z y by (induction) (auto intro: re_succ) lemma reachable_re: assumes "reachable x y" shows "reachable_end x y" using assms by (induction) (auto intro: succ_re) lemma re_reachable: assumes "reachable_end x y" shows "reachable x y" using assms by (induction) (auto intro: succ_reachable) lemma reachable_end_induct: assumes r: "reachable x y" and base: "\x. P x x" and step: "\x y z. \P x y; edge y z\ \ P x z" shows "P x y" using r[THEN reachable_re] proof (induction) case (re_refl x) from base show ?case . next case (re_succ x y z) with step show ?case by blast qed text \ We also need the following variant of reachability avoiding certain edges. More precisely, @{text y} is reachable from @{text x} avoiding a set @{text E} of edges if there exists a path such that no edge from @{text E} appears along the path. \ inductive reachable_avoiding where ra_refl[iff]: "reachable_avoiding x x E" | ra_succ[elim]: "\reachable_avoiding x y E; edge y z; (y,z) \ E\ \ reachable_avoiding x z E" lemma edge_ra: assumes "edge x y" and "(x,y) \ E" shows "reachable_avoiding x y E" using assms by (meson reachable_avoiding.simps) lemma ra_trans: assumes 1: "reachable_avoiding x y E" and 2: "reachable_avoiding y z E" shows "reachable_avoiding x z E" using 2 1 by induction auto lemma ra_cases: assumes "reachable_avoiding x y E" shows "x=y \ (\z. z \ successors x \ (x,z) \ E \ reachable_avoiding z y E)" using assms proof (induction) case (ra_refl x S) then show ?case by simp next case (ra_succ x y S z) then show ?case by (metis ra_refl reachable_avoiding.ra_succ) qed lemma ra_mono: assumes "reachable_avoiding x y E" and "E' \ E" shows "reachable_avoiding x y E'" using assms by induction auto lemma ra_add_edge: assumes "reachable_avoiding x y E" shows "reachable_avoiding x y (E \ {(v,w)}) \ (reachable_avoiding x v (E \ {(v,w)}) \ reachable_avoiding w y (E \ {(v,w)}))" using assms proof (induction) case (ra_refl x E) then show ?case by simp next case (ra_succ x y E z) then show ?case using reachable_avoiding.ra_succ by auto qed text \ Reachability avoiding some edges obviously implies reachability. Conversely, reachability implies reachability avoiding the empty set. \ lemma ra_reachable: "reachable_avoiding x y E \ reachable x y" by (induction rule: reachable_avoiding.induct) (auto intro: succ_reachable) lemma ra_empty: "reachable_avoiding x y {} = reachable x y" proof assume "reachable_avoiding x y {}" thus "reachable x y" by (rule ra_reachable) next assume "reachable x y" hence "reachable_end x y" by (rule reachable_re) thus "reachable_avoiding x y {}" by induction auto qed section \Strongly connected components\ text \ A strongly connected component is a set @{text S} of nodes such that any two nodes in @{text S} are reachable from each other. This concept is represented by the predicate @{text "is_subscc"} below. We are ultimately interested in non-empty, maximal strongly connected components, represented by the predicate @{text "is_scc"}. \ definition is_subscc where "is_subscc S \ \x \ S. \y \ S. reachable x y" definition is_scc where "is_scc S \ S \ {} \ is_subscc S \ (\S'. S \ S' \ is_subscc S' \ S' = S)" lemma subscc_add: assumes "is_subscc S" and "x \ S" and "reachable x y" and "reachable y x" shows "is_subscc (insert y S)" using assms unfolding is_subscc_def by (metis insert_iff reachable_trans) lemma sccE: \ \Two nodes that are reachable from each other are in the same SCC.\ assumes "is_scc S" and "x \ S" and "reachable x y" and "reachable y x" shows "y \ S" using assms unfolding is_scc_def by (metis insertI1 subscc_add subset_insertI) lemma scc_partition: \ \Two SCCs that contain a common element are identical.\ assumes "is_scc S" and "is_scc S'" and "x \ S \ S'" shows "S = S'" using assms unfolding is_scc_def is_subscc_def by (metis IntE assms(2) sccE subsetI) section \Algorithm for computing strongly connected components\ text \ We now introduce our representation of Bloemen's algorithm in Isabelle/HOL. The auxiliary function @{text unite} corresponds to the inner \textsf{while} loop in Bloemen's pseudo-code~\<^cite>\\p.32\ in "bloemen:strong"\. It is applied to two nodes @{text v} and @{text w} (and the environment @{text e} holding the current values of the program variables) when a loop is found, i.e.\ when @{text w} is a successor of @{text v} in the graph that has already been visited in the depth-first search. In that case, the root of the SCC of node @{text w} determined so far must appear below the root of @{text v}'s SCC in the @{text stack} maintained by the algorithm. The effect of the function is to merge the SCCs of all nodes on the top of the stack above (and including) @{text w}. Node @{text w}'s root will be the root of the merged SCC. \ definition unite :: "'v \ 'v \ 'v env \ 'v env" where "unite v w e \ let pfx = takeWhile (\x. w \ \ e x) (stack e); sfx = dropWhile (\x. w \ \ e x) (stack e); cc = \ { \ e x | x . x \ set pfx \ {hd sfx} } in e\\ := \x. if x \ cc then cc else \ e x, stack := sfx\" text \ We now represent the algorithm as two mutually recursive functions @{text dfs} and @{text dfss} in Isabelle/HOL. The function @{text dfs} corresponds to Bloemen's function \textsf{SetBased}, whereas @{text dfss} corresponds to the \textsf{forall} loop over the successors of the node on which @{text dfs} was called. Instead of using global program variables in imperative style, our functions explicitly pass environments that hold the current values of these variables. A technical complication in the development of the algorithm in Isabelle is the fact that the functions need not terminate when their pre-conditions (introduced below) are violated, for example when @{text dfs} is called for a node that was already visited previously. We therefore cannot prove termination at this point, but will later show that the explicitly given pre-conditions ensure termination. \ function (domintros) dfs :: "'v \ 'v env \ 'v env" and dfss :: "'v \ 'v env \ 'v env" where "dfs v e = (let e1 = e\visited := visited e \ {v}, stack := (v # stack e), cstack := (v # cstack e)\; e' = dfss v e1 in if v = hd(stack e') then e'\sccs := sccs e' \ {\ e' v}, explored := explored e' \ (\ e' v), stack := tl(stack e'), cstack := tl(cstack e')\ else e'\cstack := tl(cstack e')\)" | "dfss v e = (let vs = successors v - vsuccs e v in if vs = {} then e else let w = SOME x. x \ vs; e' = (if w \ explored e then e else if w \ visited e then dfs w e else unite v w e); e'' = (e'\vsuccs := (\x. if x=v then vsuccs e' v \ {w} else vsuccs e' x)\) in dfss v e'')" by pat_completeness (force+) section \Definition of the predicates used in the correctness proof\ text \ Environments are partially ordered according to the following definition. \ definition sub_env where "sub_env e e' \ root e' = root e \ visited e \ visited e' \ explored e \ explored e' \ (\v. vsuccs e v \ vsuccs e' v) \ (\ v. \ e v \ \ e' v) \ (\ {\ e v | v . v \ set (stack e)}) \ (\ {\ e' v | v . v \ set (stack e')}) " lemma sub_env_trans: assumes "sub_env e e'" and "sub_env e' e''" shows "sub_env e e''" using assms unfolding sub_env_def by (metis (no_types, lifting) subset_trans) text \ The set @{text "unvisited e u"} contains all edges @{text "(a,b)"} such that node @{text a} is in the same SCC as node @{text u} and the edge has not yet been followed, in the sense represented by variable @{text vsuccs}. \ definition unvisited where "unvisited e u \ {(a,b) | a b. a \ \ e u \ b \ successors a - vsuccs e a}" subsection \Main invariant\ text \ The following definition characterizes well-formed environments. This predicate will be shown to hold throughout the execution of the algorithm. In words, it asserts the following facts: \begin{itemize} \item Only nodes reachable from the root (for which the algorithm was originally called) are visited. \item The two stacks @{text stack} and @{text cstack} do not contain duplicate nodes, and @{text stack} contains a subset of the nodes on @{text cstack}, in the same order. \item Any node higher on the @{text stack} (i.e., that was pushed later) is reachable from nodes lower in the @{text stack}. This property also holds for nodes on the call stack, but this is not needed for the correctness proof. \item Every explored node, and every node on the call stack, has been visited. \item Nodes reachable from fully explored nodes have themselves been fully explored. \item The set @{text "vsuccs e n"}, for any node @{text n}, is a subset of @{text n}'s successors, and all these nodes are in @{text visited}. The set is empty if @{text "n \ visited"}, and it contains all successors if @{text n} has been fully explored or if @{text n} has been visited, but is no longer on the call stack. \item The sets @{text "\ e n"} represent an equivalence relation. The equivalence classes of nodes that have not yet been visited are singletons. Also, equivalence classes for two distinct nodes on the @{text stack} are disjoint because the stack only stores roots of SCCs, and the union of the equivalence classes for these root nodes corresponds to the set of live nodes, i.e. those nodes that have already been visited but not yet fully explored. \item More precisely, an equivalence class is represented on the stack by the oldest node in the sense of the call order: any node in the class that is still on the call stack precedes the representative on the call stack and was therefore pushed later. \item Equivalence classes represent the maximal available information about strong connectedness: nodes represented by some node @{text n} on the @{text stack} can reach some node @{text m} that is lower in the stack only by taking an edge from some node in @{text n}'s equivalence class that has not yet been followed. (Remember that @{text m} can reach @{text n} by one of the previous conjuncts.) \item Equivalence classes represent partial SCCs in the sense of the predicate @{text is_subscc}. Variable @{text sccs} holds maximal SCCs in the sense of the predicate @{text is_scc}, and their union corresponds to the set of explored nodes. \end{itemize} \ definition wf_env where "wf_env e \ (\n \ visited e. reachable (root e) n) \ distinct (stack e) \ distinct (cstack e) \ (\n m. n \ m in stack e \ n \ m in cstack e) \ (\n m. n \ m in stack e \ reachable m n) \ explored e \ visited e \ set (cstack e) \ visited e \ (\n \ explored e. \m. reachable n m \ m \ explored e) \ (\n. vsuccs e n \ successors n \ visited e) \ (\n. n \ visited e \ vsuccs e n = {}) \ (\n \ explored e. vsuccs e n = successors n) \ (\n \ visited e - set (cstack e). vsuccs e n = successors n) \ (\n m. m \ \ e n \ (\ e n = \ e m)) \ (\n. n \ visited e \ \ e n = {n}) \ (\n \ set (stack e). \m \ set (stack e). n \ m \ \ e n \ \ e m = {}) \ \ {\ e n | n. n \ set (stack e)} = visited e - explored e \ (\n \ set (stack e). \m \ \ e n. m \ set (cstack e) \ m \ n in cstack e) \ (\n m. n \ m in stack e \ n \ m \ (\u \ \ e n. \ reachable_avoiding u m (unvisited e n))) \ (\n. is_subscc (\ e n)) \ (\S \ sccs e. is_scc S) \ \ (sccs e) = explored e" subsection \Consequences of the invariant\ text \ Since every node on the call stack is an element of @{text visited} and every node on the @{text stack} also appears on @{text cstack}, all these nodes are also in @{text visited}. \ lemma stack_visited: assumes "wf_env e" "n \ set (stack e)" shows "n \ visited e" using assms unfolding wf_env_def by (meson precedes_refl subset_iff) text \ Classes represented on the stack consist of visited nodes that have not yet been fully explored. \ lemma stack_class: assumes "wf_env e" "n \ set (stack e)" "m \ \ e n" shows "m \ visited e - explored e" using assms unfolding wf_env_def by blast text \ Conversely, every such node belongs to some class represented on the stack. \ lemma visited_unexplored: assumes "wf_env e" "m \ visited e" "m \ explored e" obtains n where "n \ set (stack e)" "m \ \ e n" using assms unfolding wf_env_def by (smt (verit, ccfv_threshold) Diff_iff Union_iff mem_Collect_eq) text \ Every node belongs to its own equivalence class. \ lemma S_reflexive: assumes "wf_env e" shows "n \ \ e n" using assms by (auto simp: wf_env_def) text \ No node on the stack has been fully explored. \ lemma stack_unexplored: assumes 1: "wf_env e" and 2: "n \ set (stack e)" and 3: "n \ explored e" shows "P" using stack_class[OF 1 2] S_reflexive[OF 1] 3 by blast text \ If @{text w} is reachable from visited node @{text v}, but no unvisited successor of a node reachable from @{text v} can reach @{text w}, then @{text w} must be visited. \ lemma reachable_visited: assumes e: "wf_env e" and v: "v \ visited e" and w: "reachable v w" and s: "\n \ visited e. \m \ successors n - vsuccs e n. reachable v n \ \ reachable m w" shows "w \ visited e" using w v s proof (induction) case (reachable_refl x) then show ?case by simp next case (reachable_succ x y z) then have "y \ vsuccs e x" by blast with e have "y \ visited e" unfolding wf_env_def by (meson le_infE subset_eq) with reachable_succ reachable.reachable_succ show ?case by blast qed text \ Edges towards explored nodes do not contribute to reachability of unexplored nodes avoiding some set of edges. \ lemma avoiding_explored: assumes e: "wf_env e" and xy: "reachable_avoiding x y E" and y: "y \ explored e" and w: "w \ explored e" shows "reachable_avoiding x y (E \ {(v,w)})" using xy y proof (induction) case (ra_refl x E) then show ?case by simp next case (ra_succ x y E z) from e \z \ successors y\ \z \ explored e\ have "y \ explored e" unfolding wf_env_def by (meson reachable_edge) with ra_succ.IH have "reachable_avoiding x y (E \ {(v,w)})" . moreover from w \(y,z) \ E\ \z \ explored e\ have "(y,z) \ E \ {(v,w)}" by auto ultimately show ?case using \z \ successors y\ by auto qed subsection \Pre- and post-conditions of function @{text dfs}\ text \ Function @{text dfs} should be called for a well-formed environment and a node @{text v} that has not yet been visited and that is reachable from the root node, as well as from all nodes in the stack. No outgoing edges from node @{text v} have yet been followed. \ definition pre_dfs where "pre_dfs v e \ wf_env e \ v \ visited e \ reachable (root e) v \ vsuccs e v = {} \ (\n \ set (stack e). reachable n v)" text \ Function @{text dfs} maintains the invariant @{text wf_env} and returns an environment @{text e'} that extends the input environment @{text e}. Node @{text v} has been visited and all its outgoing edges have been followed. Because the algorithm works in depth-first fashion, no new outgoing edges of nodes that had already been visited in the input environment have been followed, and the stack of @{text e'} is a suffix of the one of @{text e} such that @{text v} is still reachable from all nodes on the stack. The stack may have been shortened because SCCs represented at the top of the stack may have been merged. The call stack is reestablished as it was in @{text e}. There are two possible outcomes of the algorithm: \begin{itemize} \item Either @{text v} has been fully explored, in which case the stacks of @{text e} and @{text e'} are the same, and the equivalence classes of all nodes represented on the stack are unchanged. This corresponds to the case where @{text v} is the root node of its (maximal) SCC. \item Alternatively, the stack of @{text e'} must be non-empty and @{text v} must be represented by the node at the top of the stack. The SCCs of the nodes lower on the stack are unchanged. This corresponds to the case where @{text v} is not the root node of its SCC, but some SCCs at the top of the stack may have been merged. \end{itemize} \ definition post_dfs where "post_dfs v e e' \ wf_env e' \ v \ visited e' \ sub_env e e' \ vsuccs e' v = successors v \ (\w \ visited e. vsuccs e' w = vsuccs e w) \ (\n \ set (stack e'). reachable n v) \ (\ns. stack e = ns @ (stack e')) \ ( (v \ explored e' \ stack e' = stack e \ (\n \ set (stack e'). \ e' n = \ e n)) \ (stack e' \ [] \ v \ \ e' (hd (stack e')) \ (\n \ set (tl (stack e')). \ e' n = \ e n))) \ cstack e' = cstack e" text \ The initial environment is easily seen to satisfy @{text dfs}'s pre-condition. \ lemma init_env_pre_dfs: "pre_dfs v (init_env v)" by (auto simp: pre_dfs_def wf_env_def init_env_def is_subscc_def dest: precedes_mem) text \ Any node represented by the top stack element of the input environment is still represented by the top element of the output stack. \ lemma dfs_S_hd_stack: assumes wf: "wf_env e" and post: "post_dfs v e e'" and n: "stack e \ []" "n \ \ e (hd (stack e))" shows "stack e' \ []" "n \ \ e' (hd (stack e'))" proof - have 1: "stack e' \ [] \ n \ \ e' (hd (stack e'))" proof (cases "stack e' = stack e \ (\n \ set (stack e'). \ e' n = \ e n)") case True with n show ?thesis by auto next case 2: False with post have "stack e' \ []" by (simp add: post_dfs_def) from n have "hd (stack e) \ set (stack e)" by simp with 2 n post obtain u where u: "u \ set (stack e')" "n \ \ e' u" unfolding post_dfs_def sub_env_def by blast show ?thesis proof (cases "u = hd (stack e')") case True with u \stack e' \ []\ show ?thesis by simp next case False with u have "u \ set (tl (stack e'))" by (metis empty_set equals0D list.collapse set_ConsD) with u 2 post have "u \ set (tl (stack e)) \ n \ \ e u" unfolding post_dfs_def by (metis Un_iff append_self_conv2 set_append tl_append2) with n wf \hd (stack e) \ set (stack e)\ show ?thesis unfolding wf_env_def by (metis (no_types, opaque_lifting) disjoint_iff_not_equal distinct.simps(2) list.collapse list.set_sel(2)) qed qed from 1 show "stack e' \ []" by simp from 1 show "n \ \ e' (hd (stack e'))" by simp qed text \ Function @{text dfs} leaves the SCCs represented by elements in the (new) tail of the @{text stack} unchanged. \ lemma dfs_S_tl_stack: assumes post: "post_dfs v e e'" and nempty: "stack e \ []" shows "stack e' \ []" "\n \ set (tl (stack e')). \ e' n = \ e n" proof - have 1: "stack e' \ [] \ (\n \ set (tl (stack e')). \ e' n = \ e n)" proof (cases "stack e' = stack e \ (\n \ set (stack e'). \ e' n = \ e n)") case True with nempty show ?thesis by (simp add: list.set_sel(2)) next case False with post show ?thesis by (auto simp: post_dfs_def) qed from 1 show "stack e' \ []" by simp from 1 show "\n \ set (tl (stack e')). \ e' n = \ e n" by simp qed subsection \Pre- and post-conditions of function @{text dfss}\ text \ The pre- and post-conditions of function @{text dfss} correspond to the invariant of the loop over all outgoing edges from node @{text v}. The environment must be well-formed, node @{text v} must be visited and represented by the top element of the (non-empty) stack. Node @{text v} must be reachable from all nodes on the stack, and it must be the top node on the call stack. All outgoing edges of node @{text v} that have already been followed must either lead to completely explored nodes (that are no longer represented on the stack) or to nodes that are part of the same SCC as @{text v}. \ definition pre_dfss where "pre_dfss v e \ wf_env e \ v \ visited e \ (stack e \ []) \ (v \ \ e (hd (stack e))) \ (\w \ vsuccs e v. w \ explored e \ \ e (hd (stack e))) \ (\n \ set (stack e). reachable n v) \ (\ns. cstack e = v # ns)" text \ The post-condition establishes that all outgoing edges of node @{text v} have been followed. As for function @{text dfs}, no new outgoing edges of previously visited nodes have been followed. Also as before, the new stack is a suffix of the old one, and the call stack is restored. In case node @{text v} is still on the stack (and therefore is the root node of its SCC), no node that is lower on the stack can be reachable from @{text v}. This condition guarantees the maximality of the computed SCCs. \ definition post_dfss where "post_dfss v e e' \ wf_env e' \ vsuccs e' v = successors v \ (\w \ visited e - {v}. vsuccs e' w = vsuccs e w) \ sub_env e e' \ (\w \ successors v. w \ explored e' \ \ e' (hd (stack e'))) \ (\n \ set (stack e'). reachable n v) \ (stack e' \ []) \ (\ns. stack e = ns @ (stack e')) \ v \ \ e' (hd (stack e')) \ (\n \ set (tl (stack e')). \ e' n = \ e n) \ (hd (stack e') = v \ (\n \ set (tl (stack e')). \ reachable v n)) \ cstack e' = cstack e" section \Proof of partial correctness\ subsection \Lemmas about function @{text unite}\ text \ We start by establishing a few lemmas about function @{text unite} in the context where it is called. \ lemma unite_stack: fixes e v w defines "e' \ unite v w e" assumes wf: "wf_env e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" obtains pfx where "stack e = pfx @ (stack e')" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" "w \ \ e' (hd (stack e'))" proof - define pfx where "pfx = takeWhile (\x. w \ \ e x) (stack e)" define sfx where "sfx = dropWhile (\x. w \ \ e x) (stack e)" define cc where "cc = \ {\ e x |x. x \ set pfx \ {hd sfx}}" have "stack e = pfx @ sfx" by (simp add: pfx_def sfx_def) moreover have "stack e' = sfx" by (simp add: e'_def unite_def sfx_def) moreover from wf w have "w \ \ {\ e n | n. n \ set (stack e)}" by (simp add: wf_env_def) then obtain n where "n \ set (stack e)" "w \ \ e n" by auto hence sfx: "sfx \ [] \ w \ \ e (hd sfx)" unfolding sfx_def by (metis dropWhile_eq_Nil_conv hd_dropWhile) moreover have "\ e' = (\x. if x \ cc then cc else \ e x)" by (rule, auto simp add: e'_def unite_def pfx_def sfx_def cc_def) moreover from sfx have "w \ cc" by (auto simp: cc_def) from S_reflexive[OF wf, of "hd sfx"] have "hd sfx \ cc" by (auto simp: cc_def) with \w \ cc\ \\ e' = (\x. if x \ cc then cc else \ e x)\ have "w \ \ e' (hd sfx)" by simp ultimately show ?thesis using that e'_def unite_def pfx_def sfx_def cc_def by meson qed text \ Function @{text unite} leaves intact the equivalence classes represented by the tail of the new stack. \ lemma unite_S_tl: fixes e v w defines "e' \ unite v w e" assumes wf: "wf_env e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" and n: "n \ set (tl (stack e'))" shows "\ e' n = \ e n" proof - from assms obtain pfx where pfx: "stack e = pfx @ (stack e')" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" by (blast dest: unite_stack) define cc where "cc \ \ {\ e n |n. n \ set pfx \ {hd (stack e')}}" have "n \ cc" proof assume "n \ cc" then obtain m where "m \ set pfx \ {hd (stack e')}" "n \ \ e m" by (auto simp: cc_def) with S_reflexive[OF wf, of n] n wf \stack e = pfx @ stack e'\ \stack e' \ []\ show "False" unfolding wf_env_def by (smt (z3) Diff_triv Un_iff Un_insert_right append.right_neutral disjoint_insert(1) distinct.simps(2) distinct_append empty_set insertE insert_Diff list.exhaust_sel list.simps(15) set_append) qed with pfx show "\ e' n = \ e n" by (auto simp add: cc_def) qed text \ The stack of the result of @{text unite} represents the same vertices as the input stack, potentially in fewer equivalence classes. \ lemma unite_S_equal: fixes e v w defines "e' \ unite v w e" assumes wf: "wf_env e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" shows "(\ {\ e' n | n. n \ set (stack e')}) = (\ {\ e n | n. n \ set (stack e)})" proof - from assms obtain pfx where pfx: "stack e = pfx @ (stack e')" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" by (blast dest: unite_stack) define cc where "cc \ \ {\ e n |n. n \ set pfx \ {hd (stack e')}}" from pfx have Se': "\x. \ e' x = (if x \ cc then cc else \ e x)" by (auto simp: cc_def) from S_reflexive[OF wf, of "hd (stack e')"] have S_hd: "\ e' (hd (stack e')) = cc" by (auto simp: Se' cc_def) from \stack e' \ []\ have ste': "set (stack e') = {hd (stack e')} \ set (tl (stack e'))" by (metis insert_is_Un list.exhaust_sel list.simps(15)) from \stack e = pfx @ stack e'\ \stack e' \ []\ have "stack e = pfx @ (hd (stack e') # tl (stack e'))" by auto hence "\ {\ e n | n. n \ set (stack e)} = cc \ (\ {\ e n | n. n \ set (tl (stack e'))})" by (auto simp add: cc_def) also from S_hd unite_S_tl[OF wf w] have "\ = \ e' (hd (stack e')) \ (\ {\ e' n | n. n \ set (tl (stack e'))})" by (auto simp: e'_def) also from ste' have "\ = \ {\ e' n | n. n \ set (stack e')}" by auto finally show ?thesis by simp qed text \ The head of the stack represents a (not necessarily maximal) SCC. \ lemma unite_subscc: fixes e v w defines "e' \ unite v w e" assumes pre: "pre_dfss v e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" shows "is_subscc (\ e' (hd (stack e')))" proof - from pre have wf: "wf_env e" by (simp add: pre_dfss_def) from assms obtain pfx where pfx: "stack e = pfx @ (stack e')" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" by (blast dest: unite_stack[OF wf]) define cc where "cc \ \ {\ e n |n. n \ set pfx \ {hd (stack e')}}" from wf w have "w \ \ {\ e n | n. n \ set (stack e)}" by (simp add: wf_env_def) hence "w \ \ e (hd (stack e'))" apply (simp add: e'_def unite_def) by (metis dropWhile_eq_Nil_conv hd_dropWhile) have "is_subscc cc" proof (clarsimp simp: is_subscc_def) fix x y assume "x \ cc" "y \ cc" then obtain nx ny where nx: "nx \ set pfx \ {hd (stack e')}" "x \ \ e nx" and ny: "ny \ set pfx \ {hd (stack e')}" "y \ \ e ny" by (auto simp: cc_def) with wf have "reachable x nx" "reachable ny y" by (auto simp: wf_env_def is_subscc_def) from w pre have "reachable v w" by (auto simp: pre_dfss_def) from pre have "reachable (hd (stack e)) v" by (auto simp: pre_dfss_def wf_env_def is_subscc_def) from pre have "stack e = hd (stack e) # tl (stack e)" by (auto simp: pre_dfss_def) with nx \stack e = pfx @ (stack e')\ \stack e' \ []\ have "hd (stack e) \ nx in stack e" by (metis Un_iff Un_insert_right head_precedes list.exhaust_sel list.simps(15) set_append sup_bot.right_neutral) with wf have "reachable nx (hd (stack e))" by (auto simp: wf_env_def) from \stack e = pfx @ (stack e')\ \stack e' \ []\ ny have "ny \ hd (stack e') in stack e" by (metis List.set_insert empty_set insert_Nil list.exhaust_sel set_append split_list_precedes) with wf have "reachable (hd (stack e')) ny" by (auto simp: wf_env_def is_subscc_def) from wf \stack e' \ []\ \w \ \ e (hd (stack e'))\ have "reachable w (hd (stack e'))" by (auto simp: wf_env_def is_subscc_def) from \reachable x nx\ \reachable nx (hd (stack e))\ \reachable (hd (stack e)) v\ \reachable v w\ \reachable w (hd (stack e'))\ \reachable (hd (stack e')) ny\ \reachable ny y\ show "reachable x y" using reachable_trans by meson qed with S_reflexive[OF wf, of "hd (stack e')"] pfx show ?thesis by (auto simp: cc_def) qed text \ The environment returned by function @{text unite} extends the input environment. \ lemma unite_sub_env: fixes e v w defines "e' \ unite v w e" assumes pre: "pre_dfss v e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" shows "sub_env e e'" proof - from pre have wf: "wf_env e" by (simp add: pre_dfss_def) from assms obtain pfx where pfx: "stack e = pfx @ (stack e')" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" by (blast dest: unite_stack[OF wf]) define cc where "cc \ \ {\ e n |n. n \ set pfx \ {hd (stack e')}}" have "\n. \ e n \ \ e' n" proof (clarify) fix n u assume u: "u \ \ e n" show "u \ \ e' n" proof (cases "n \ cc") case True then obtain m where m: "m \ set pfx \ {hd (stack e')}" "n \ \ e m" by (auto simp: cc_def) with wf S_reflexive[OF wf, of n] u have "u \ \ e m" by (auto simp: wf_env_def) with m pfx show ?thesis by (auto simp: cc_def) next case False with pfx u show ?thesis by (auto simp: cc_def) qed qed moreover have "root e' = root e \ visited e' = visited e \ explored e' = explored e \ vsuccs e' = vsuccs e" by (simp add: e'_def unite_def) ultimately show ?thesis using unite_S_equal[OF wf w] by (simp add: e'_def sub_env_def) qed text \ The environment returned by function @{text unite} is well-formed. \ lemma unite_wf_env: fixes e v w defines "e' \ unite v w e" assumes pre: "pre_dfss v e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" shows "wf_env e'" proof - from pre have wf: "wf_env e" by (simp add: pre_dfss_def) from assms obtain pfx where pfx: "stack e = pfx @ (stack e')" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" by (blast dest: unite_stack[OF wf]) define cc where "cc \ \ {\ e n |n. n \ set pfx \ {hd (stack e')}}" from pfx have Se': "\x. \ e' x = (if x \ cc then cc else \ e x)" by (auto simp add: cc_def) have cc_Un: "cc = \ {\ e x | x. x \ cc}" proof from S_reflexive[OF wf] show "cc \ \ {\ e x | x. x \ cc}" by (auto simp: cc_def) next { fix n x assume "x \ cc" "n \ \ e x" with wf have "n \ cc" unfolding wf_env_def cc_def by (smt (verit) Union_iff mem_Collect_eq) } thus "(\ {\ e x | x. x \ cc}) \ cc" by blast qed from S_reflexive[OF wf, of "hd (stack e')"] have hd_cc: "\ e' (hd (stack e')) = cc" by (auto simp: cc_def Se') { fix n m assume n: "n \ set (tl (stack e'))" and m: "m \ \ e n \ cc" from m obtain l where "l \ set pfx \ {hd (stack e')}" "m \ \ e l" by (auto simp: cc_def) with n m wf \stack e = pfx @ stack e'\ \stack e' \ []\ have "False" unfolding wf_env_def by (metis (no_types, lifting) Int_iff UnCI UnE disjoint_insert(1) distinct.simps(2) distinct_append emptyE hd_Cons_tl insert_iff list.set_sel(1) list.set_sel(2) mk_disjoint_insert set_append) } hence tl_cc: "\n \ set (tl (stack e')). \ e n \ cc = {}" by blast from wf have "\n \ visited e'. reachable (root e') n" "distinct (cstack e')" "explored e' \ visited e'" "set (cstack e') \ visited e'" "\n \ explored e'. \m. reachable n m \ m \ explored e'" "\n. vsuccs e' n \ successors n \ visited e'" "\n. n \ visited e' \ vsuccs e' n = {}" "\n \ explored e'. vsuccs e' n = successors n" "\n \ visited e' - set (cstack e'). vsuccs e' n = successors n" "\S \ sccs e'. is_scc S" "\ (sccs e') = explored e'" by (auto simp: wf_env_def e'_def unite_def) moreover from wf \stack e = pfx @ stack e'\ have "distinct (stack e')" by (auto simp: wf_env_def) moreover have "\n m. n \ m in stack e' \ n \ m in cstack e'" proof (clarify) fix n m assume "n \ m in stack e'" with \stack e = pfx @ stack e'\ wf have "n \ m in cstack e" unfolding wf_env_def by (metis precedes_append_left) thus "n \ m in cstack e'" by (simp add: e'_def unite_def) qed moreover from wf \stack e = pfx @ stack e'\ have "\n m. n \ m in stack e' \ reachable m n" unfolding wf_env_def by (metis precedes_append_left) moreover have "\n m. m \ \ e' n \ (\ e' n = \ e' m)" proof (clarify) fix n m show "m \ \ e' n \ (\ e' n = \ e' m)" proof assume l: "m \ \ e' n" show "\ e' n = \ e' m" proof (cases "n \ cc") case True with l show ?thesis by (simp add: Se') next case False with l wf have "\ e n = \ e m" by (simp add: wf_env_def Se') with False cc_Un wf have "m \ cc" unfolding wf_env_def e'_def by (smt (verit, best) Union_iff mem_Collect_eq) with \\ e n = \ e m\ False show ?thesis by (simp add: Se') qed next assume r: "\ e' n = \ e' m" show "m \ \ e' n" proof (cases "n \ cc") case True with r pfx have "\ e' m = cc" by (auto simp: cc_def) have "m \ cc" proof (rule ccontr) assume "m \ cc" with pfx have "\ e' m = \ e m" by (auto simp: cc_def) with S_reflexive[OF wf, of m] \\ e' m = cc\ \m \ cc\ show "False" by simp qed with pfx True show "m \ \ e' n" by (auto simp: cc_def) next case False hence "\ e' n = \ e n" by (simp add: Se') have "m \ cc" proof assume m: "m \ cc" with \\ e' n = \ e n\ r have "\ e n = cc" by (simp add: Se') with S_reflexive[OF wf, of n] have "n \ cc" by simp with \n \ cc\ show "False" .. qed with r \\ e' n = \ e n\ have "\ e m = \ e n" by (simp add: Se') with S_reflexive[OF wf, of m] have "m \ \ e n" by simp with \\ e' n = \ e n\ show ?thesis by simp qed qed qed moreover have "\n. n \ visited e' \ \ e' n = {n}" proof (clarify) fix n assume "n \ visited e'" hence "n \ visited e" by (simp add: e'_def unite_def) moreover have "n \ cc" proof assume "n \ cc" then obtain m where "m \ set pfx \ {hd (stack e')}" "n \ \ e m" by (auto simp: cc_def) with \stack e = pfx @ stack e'\ \stack e' \ []\ have "m \ set (stack e)" by auto with stack_class[OF wf this \n \ \ e m\] \n \ visited e\ show "False" by simp qed ultimately show "\ e' n = {n}" using wf by (auto simp: wf_env_def Se') qed moreover have "\n \ set (stack e'). \m \ set (stack e'). n \ m \ \ e' n \ \ e' m = {}" proof (clarify) fix n m assume "n \ set (stack e')" "m \ set (stack e')" "n \ m" show "\ e' n \ \ e' m = {}" proof (cases "n = hd (stack e')") case True with \m \ set (stack e')\ \n \ m\ \stack e' \ []\ have "m \ set (tl (stack e'))" by (metis hd_Cons_tl set_ConsD) with True hd_cc tl_cc unite_S_tl[OF wf w] show ?thesis by (auto simp: e'_def) next case False with \n \ set (stack e')\ \stack e' \ []\ have "n \ set (tl (stack e'))" by (metis hd_Cons_tl set_ConsD) show ?thesis proof (cases "m = hd (stack e')") case True with \n \ set (tl (stack e'))\ hd_cc tl_cc unite_S_tl[OF wf w] show ?thesis by (auto simp: e'_def) next case False with \m \ set (stack e')\ \stack e' \ []\ have "m \ set (tl (stack e'))" by (metis hd_Cons_tl set_ConsD) with \n \ set (tl (stack e'))\ have "\ e' m = \ e m \ \ e' n = \ e n" by (auto simp: e'_def unite_S_tl[OF wf w]) moreover from \m \ set (stack e')\ \n \ set (stack e')\ \stack e = pfx @ stack e'\ have "m \ set (stack e) \ n \ set (stack e)" by auto ultimately show ?thesis using wf \n \ m\ by (auto simp: wf_env_def) qed qed qed moreover { from unite_S_equal[OF wf w] have "\ {\ e' n | n. n \ set (stack e')} = \ {\ e n | n. n \ set (stack e)}" by (simp add: e'_def) with wf have "\ {\ e' n | n. n \ set (stack e')} = visited e - explored e" by (simp add: wf_env_def) } hence "\ {\ e' n | n. n \ set (stack e')} = visited e' - explored e'" by (simp add: e'_def unite_def) moreover have "\n \ set (stack e'). \m \ \ e' n. m \ set (cstack e') \ m \ n in cstack e'" proof (clarify) fix n m assume "n \ set (stack e')" "m \ \ e' n" "m \ set (cstack e')" from \m \ set (cstack e')\ have "m \ set (cstack e)" by (simp add: e'_def unite_def) have "m \ n in cstack e" proof (cases "n = hd (stack e')") case True with \m \ \ e' n\ have "m \ cc" by (simp add: hd_cc) then obtain l where "l \ set pfx \ {hd (stack e')}" "m \ \ e l" by (auto simp: cc_def) with \stack e = pfx @ stack e'\ \stack e' \ []\ have "l \ set (stack e)" by auto with \m \ \ e l\ \m \ set (cstack e)\ wf have "m \ l in cstack e" by (auto simp: wf_env_def) moreover from \l \ set pfx \ {hd (stack e')}\ True \stack e = pfx @ stack e'\ \stack e' \ []\ have "l \ n in stack e" by (metis List.set_insert empty_set hd_Cons_tl insert_Nil set_append split_list_precedes) with wf have "l \ n in cstack e" by (auto simp: wf_env_def) ultimately show ?thesis using wf unfolding wf_env_def by (meson precedes_trans) next case False with \n \ set (stack e')\ \stack e' \ []\ have "n \ set (tl (stack e'))" by (metis list.collapse set_ConsD) with unite_S_tl[OF wf w] \m \ \ e' n\ have "m \ \ e n" by (simp add: e'_def) with \n \ set (stack e')\ \stack e = pfx @ stack e'\ \m \ set (cstack e)\ wf show ?thesis by (auto simp: wf_env_def) qed thus "m \ n in cstack e'" by (simp add: e'_def unite_def) qed moreover have "\n m. n \ m in stack e' \ n \ m \ (\u \ \ e' n. \ reachable_avoiding u m (unvisited e' n))" proof (clarify) fix x y u assume xy: "x \ y in stack e'" "x \ y" and u: "u \ \ e' x" "reachable_avoiding u y (unvisited e' x)" show "False" proof (cases "x = hd (stack e')") case True hence "\ e' x = cc" by (simp add: hd_cc) with \u \ \ e' x\ obtain x' where x': "x' \ set pfx \ {hd (stack e')}" "u \ \ e x'" by (auto simp: cc_def) from \stack e = pfx @ stack e'\ \stack e' \ []\ have "stack e = pfx @ (hd (stack e') # tl (stack e'))" by auto with x' True have "x' \ x in stack e" by (simp add: split_list_precedes) moreover from xy \stack e = pfx @ stack e'\ have "x \ y in stack e" by (simp add: precedes_append_left) ultimately have "x' \ y in stack e" using wf by (auto simp: wf_env_def elim: precedes_trans) from \x' \ x in stack e\ \x \ y in stack e\ wf \x \ y\ have "x' \ y" by (auto simp: wf_env_def dest: precedes_antisym) let ?unv = "\ {unvisited e y | y. y \ set pfx \ {hd (stack e')}}" from \\ e' x = cc\ have "?unv = unvisited e' x" by (auto simp: unvisited_def cc_def e'_def unite_def) with \reachable_avoiding u y (unvisited e' x)\ have "reachable_avoiding u y ?unv" by simp with x' have "reachable_avoiding u y (unvisited e x')" by (blast intro: ra_mono) with \x' \ y in stack e\ \x' \ y\ \u \ \ e x'\ wf show ?thesis by (auto simp: wf_env_def) next case False with \x \ y in stack e'\ \stack e' \ []\ have "x \ set (tl (stack e'))" by (metis list.exhaust_sel precedes_mem(1) set_ConsD) with \u \ \ e' x\ have "u \ \ e x" by (auto simp add: unite_S_tl[OF wf w] e'_def) moreover from \x \ y in stack e'\ \stack e = pfx @ stack e'\ have "x \ y in stack e" by (simp add: precedes_append_left) moreover from unite_S_tl[OF wf w] \x \ set (tl (stack e'))\ have "unvisited e' x = unvisited e x" by (auto simp: unvisited_def e'_def unite_def) ultimately show ?thesis using \x \ y\ \reachable_avoiding u y (unvisited e' x)\ wf by (auto simp: wf_env_def) qed qed moreover have "\n. is_subscc (\ e' n)" proof fix n show "is_subscc (\ e' n)" proof (cases "n \ cc") case True hence "\ e' n = cc" by (simp add: Se') with unite_subscc[OF pre w] hd_cc show ?thesis by (auto simp: e'_def) next case False with wf show ?thesis by (simp add: Se' wf_env_def) qed qed ultimately show ?thesis unfolding wf_env_def by blast qed subsection \Lemmas establishing the pre-conditions\ text \ The precondition of function @{text dfs} ensures the precondition of @{text dfss} at the call of that function. \ lemma pre_dfs_pre_dfss: assumes "pre_dfs v e" shows "pre_dfss v (e\visited := visited e \ {v}, stack := v # stack e, cstack := v # cstack e\)" (is "pre_dfss v ?e'") proof - from assms have wf: "wf_env e" by (simp add: pre_dfs_def) from assms have v: "v \ visited e" by (simp add: pre_dfs_def) from assms stack_visited[OF wf] have "\n \ visited ?e'. reachable (root ?e') n" "distinct (stack ?e')" "distinct (cstack ?e')" "explored ?e' \ visited ?e'" "set (cstack ?e') \ visited ?e'" "\n \ explored ?e'. \m. reachable n m \ m \ explored ?e'" "\n. vsuccs ?e' n \ successors n" "\n \ explored ?e'. vsuccs ?e' n = successors n" "\n \ visited ?e' - set(cstack ?e'). vsuccs ?e' n = successors n" "\n. n \ visited ?e' \ vsuccs ?e' n = {}" "(\n m. m \ \ ?e' n \ (\ ?e' n = \ ?e' m))" "(\n. n \ visited ?e' \ \ ?e' n = {n})" "\n. is_subscc (\ ?e' n)" "\S \ sccs ?e'. is_scc S" "\ (sccs ?e') = explored ?e'" by (auto simp: pre_dfs_def wf_env_def) moreover have "\n m. n \ m in stack ?e' \ reachable m n" proof (clarify) fix x y assume "x \ y in stack ?e'" show "reachable y x" proof (cases "x=v") assume "x=v" with \x \ y in stack ?e'\ assms show ?thesis apply (simp add: pre_dfs_def) by (metis insert_iff list.simps(15) precedes_mem(2) reachable_refl) next assume "x \ v" with \x \ y in stack ?e'\ wf show ?thesis by (simp add: pre_dfs_def wf_env_def precedes_in_tail) qed qed moreover from wf v have "\n. vsuccs ?e' n \ visited ?e'" by (auto simp: wf_env_def) moreover from wf v have "(\n \ set (stack ?e'). \ m \ set (stack ?e'). n \ m \ \ ?e' n \ \ ?e' m = {})" apply (simp add: wf_env_def) by (metis singletonD) moreover have "\ {\ ?e' v | v . v \ set (stack ?e')} = visited ?e' - explored ?e'" proof - have "\ {\ ?e' v | v . v \ set (stack ?e')} = (\ {\ e v | v . v \ set (stack e)}) \ \ e v" by auto also from wf v have "\ = visited ?e' - explored ?e'" by (auto simp: wf_env_def) finally show ?thesis . qed moreover have "\n m. n \ m in stack ?e' \ n \ m \ (\u \ \ ?e' n. \ reachable_avoiding u m (unvisited ?e' n))" proof (clarify) fix x y u assume asm: "x \ y in stack ?e'" "x \ y" "u \ \ ?e' x" "reachable_avoiding u y (unvisited ?e' x)" show "False" proof (cases "x = v") case True with wf v \u \ \ ?e' x\ have "u = v" "vsuccs ?e' v = {}" by (auto simp: wf_env_def) with \reachable_avoiding u y (unvisited ?e' x)\[THEN ra_cases] True \x \ y\ wf show ?thesis by (auto simp: wf_env_def unvisited_def) next case False with asm wf show ?thesis by (auto simp: precedes_in_tail wf_env_def unvisited_def) qed qed moreover have "\n m. n \ m in stack ?e' \ n \ m in cstack ?e'" proof (clarsimp) fix n m assume "n \ m in (v # stack e)" with assms show "n \ m in (v # cstack e)" unfolding pre_dfs_def wf_env_def by (metis head_precedes insertI1 list.simps(15) precedes_in_tail precedes_mem(2) precedes_refl) qed moreover have "\n \ set (stack ?e'). \m \ \ ?e' n. m \ set (cstack ?e') \ m \ n in cstack ?e'" proof (clarify) fix n m assume "n \ set (stack ?e')" "m \ \ ?e' n" "m \ set (cstack ?e')" show "m \ n in cstack ?e'" proof (cases "n = v") case True with wf v \m \ \ ?e' n\ show ?thesis by (auto simp: wf_env_def) next case False with \n \ set (stack ?e')\ \m \ \ ?e' n\ have "n \ set (stack e)" "m \ \ e n" by auto with wf v False \m \ \ e n\ \m \ set (cstack ?e')\ show ?thesis apply (simp add: wf_env_def) by (metis (mono_tags, lifting) precedes_in_tail singletonD) qed qed ultimately have "wf_env ?e'" unfolding wf_env_def by (meson le_inf_iff) moreover from assms have "\w \ vsuccs ?e' v. w \ explored ?e' \ \ ?e' (hd (stack ?e'))" by (simp add: pre_dfs_def) moreover from \\n m. n \ m in stack ?e' \ reachable m n\ have "\n \ set (stack ?e'). reachable n v" by (simp add: head_precedes) moreover from wf v have "\ ?e' (hd (stack ?e')) = {v}" by (simp add: pre_dfs_def wf_env_def) ultimately show ?thesis by (auto simp: pre_dfss_def) qed text \ Similarly, we now show that the pre-conditions of the different function calls in the body of function @{text dfss} are satisfied. First, it is very easy to see that the pre-condition of @{text dfs} holds at the call of that function. \ lemma pre_dfss_pre_dfs: assumes "pre_dfss v e" and "w \ visited e" and "w \ successors v" shows "pre_dfs w e" using assms unfolding pre_dfss_def pre_dfs_def wf_env_def by (meson succ_reachable) text \ The pre-condition of @{text dfss} holds when the successor considered in the current iteration has already been explored. \ lemma pre_dfss_explored_pre_dfss: fixes e v w defines "e'' \ e\vsuccs := (\x. if x=v then vsuccs e v \ {w} else vsuccs e x)\" assumes 1: "pre_dfss v e" and 2: "w \ successors v" and 3: "w \ explored e" shows "pre_dfss v e''" proof - from 1 have v: "v \ visited e" by (simp add: pre_dfss_def) have "wf_env e''" proof - from 1 have wf: "wf_env e" by (simp add: pre_dfss_def) hence "\v \ visited e''. reachable (root e'') v" "distinct (stack e'')" "distinct (cstack e'')" "\n m. n \ m in stack e'' \ n \ m in cstack e''" "\n m. n \ m in stack e'' \ reachable m n" "explored e'' \ visited e''" "set (cstack e'') \ visited e''" "\n \ explored e''. \m. reachable n m \ m \ explored e''" "\n m. m \ \ e'' n \ (\ e'' n = \ e'' m)" "\n. n \ visited e'' \ \ e'' n = {n}" "\n \ set (stack e''). \ m \ set (stack e''). n \ m \ \ e'' n \ \ e'' m = {}" "\ {\ e'' n | n. n \ set (stack e'')} = visited e'' - explored e''" "\n \ set (stack e''). \m \ \ e'' n. m \ set (cstack e'') \ m \ n in cstack e''" "\n. is_subscc (\ e'' n)" "\S \ sccs e''. is_scc S" "\ (sccs e'') = explored e''" by (auto simp: wf_env_def e''_def) moreover from wf 2 3 have "\v. vsuccs e'' v \ successors v \ visited e''" by (auto simp: wf_env_def e''_def) moreover from wf v have "\n. n \ visited e'' \ vsuccs e'' n = {}" by (auto simp: wf_env_def e''_def) moreover from wf 2 have "\v. v \ explored e'' \ vsuccs e'' v = successors v" by (auto simp: wf_env_def e''_def) moreover have "\x y. x \ y in stack e'' \ x \ y \ (\u \ \ e'' x. \ reachable_avoiding u y (unvisited e'' x))" proof (clarify) fix x y u assume "x \ y in stack e''" "x \ y" "u \ \ e'' x" "reachable_avoiding u y (unvisited e'' x)" hence prec: "x \ y in stack e" "u \ \ e x" by (auto simp: e''_def) with stack_unexplored[OF wf] have "y \ explored e" by (blast dest: precedes_mem) have "(unvisited e x = unvisited e'' x) \ (unvisited e x = unvisited e'' x \ {(v,w)})" by (auto simp: e''_def unvisited_def split: if_splits) thus "False" proof assume "unvisited e x = unvisited e'' x" with prec \x \ y\ \reachable_avoiding u y (unvisited e'' x)\ wf show ?thesis unfolding wf_env_def by metis next assume "unvisited e x = unvisited e'' x \ {(v,w)}" with wf \reachable_avoiding u y (unvisited e'' x)\ \y \ explored e\ \w \ explored e\ prec \x \ y\ show ?thesis using avoiding_explored[OF wf] unfolding wf_env_def by (metis (no_types, lifting)) qed qed moreover from wf 2 have "\n \ visited e'' - set (cstack e''). vsuccs e'' n = successors n" by (auto simp: e''_def wf_env_def) ultimately show ?thesis unfolding wf_env_def by meson qed with 1 3 show ?thesis by (auto simp: pre_dfss_def e''_def) qed text \ The call to @{text dfs} establishes the pre-condition for the recursive call to @{text dfss} in the body of @{text dfss}. \ lemma pre_dfss_post_dfs_pre_dfss: fixes e v w defines "e' \ dfs w e" defines "e'' \ e'\vsuccs := (\x. if x=v then vsuccs e' v \ {w} else vsuccs e' x)\" assumes pre: "pre_dfss v e" and w: "w \ successors v" "w \ visited e" and post: "post_dfs w e e'" shows "pre_dfss v e''" proof - from pre have "wf_env e" "v \ visited e" "stack e \ []" "v \ \ e (hd (stack e))" by (auto simp: pre_dfss_def) with post have "stack e' \ []" "v \ \ e' (hd (stack e'))" by (auto dest: dfs_S_hd_stack) from post have "w \ visited e'" by (simp add: post_dfs_def) have "wf_env e''" proof - from post have wf': "wf_env e'" by (simp add: post_dfs_def) hence "\n \ visited e''. reachable (root e'') n" "distinct (stack e'')" "distinct (cstack e'')" "\n m. n \ m in stack e'' \ n \ m in cstack e''" "\n m. n \ m in stack e'' \ reachable m n" "explored e'' \ visited e''" "set (cstack e'') \ visited e''" "\n \ explored e''. \m. reachable n m \ m \ explored e''" "\n m. m \ \ e'' n \ (\ e'' n = \ e'' m)" "\n. n \ visited e'' \ \ e'' n = {n}" "\n \ set (stack e''). \ m \ set (stack e''). n \ m \ \ e'' n \ \ e'' m = {}" "\ {\ e'' n | n. n \ set (stack e'')} = visited e'' - explored e''" "\n \ set (stack e''). \ m \ \ e'' n. m \ set (cstack e'') \ m \ n in cstack e''" "\n. is_subscc (\ e'' n)" "\S \ sccs e''. is_scc S" "\ (sccs e'') = explored e''" by (auto simp: wf_env_def e''_def) moreover from wf' w have "\n. vsuccs e'' n \ successors n" by (auto simp: wf_env_def e''_def) moreover from wf' \w \ visited e'\ have "\n. vsuccs e'' n \ visited e''" by (auto simp: wf_env_def e''_def) moreover from post \v \ visited e\ have "\n. n \ visited e'' \ vsuccs e'' n = {}" apply (simp add: post_dfs_def wf_env_def sub_env_def e''_def) by (meson subsetD) moreover from wf' w have "\n \ explored e''. vsuccs e'' n = successors n" by (auto simp: wf_env_def e''_def) moreover have "\n m. n \ m in stack e'' \ n \ m \ (\u \ \ e'' n. \ reachable_avoiding u m (unvisited e'' n))" proof (clarify) fix x y u assume "x \ y in stack e''" "x \ y" "u \ \ e'' x" "reachable_avoiding u y (unvisited e'' x)" hence 1: "x \ y in stack e'" "u \ \ e' x" by (auto simp: e''_def) with stack_unexplored[OF wf'] have "y \ explored e'" by (auto dest: precedes_mem) have "(unvisited e' x = unvisited e'' x) \ (unvisited e' x = unvisited e'' x \ {(v,w)})" by (auto simp: e''_def unvisited_def split: if_splits) thus "False" proof assume "unvisited e' x = unvisited e'' x" with 1 \x \ y\ \reachable_avoiding u y (unvisited e'' x)\ wf' show ?thesis unfolding wf_env_def by metis next assume unv: "unvisited e' x = unvisited e'' x \ {(v,w)}" from post have "w \ explored e' \ (w \ \ e' (hd (stack e')) \ (\n \ set (tl (stack e')). \ e' n = \ e n))" by (auto simp: post_dfs_def) thus ?thesis proof assume "w \ explored e'" with wf' unv \reachable_avoiding u y (unvisited e'' x)\ \y \ explored e'\ 1 \x \ y\ show ?thesis using avoiding_explored[OF wf'] unfolding wf_env_def by (metis (no_types, lifting)) next assume w: "w \ \ e' (hd (stack e')) \ (\n \ set (tl (stack e')). \ e' n = \ e n)" from \reachable_avoiding u y (unvisited e'' x)\[THEN ra_add_edge] unv have "reachable_avoiding u y (unvisited e' x) \ reachable_avoiding w y (unvisited e' x)" by auto thus ?thesis proof assume "reachable_avoiding u y (unvisited e' x)" with \x \ y in stack e''\ \x \ y\ \u \ \ e'' x\ wf' show ?thesis by (auto simp: e''_def wf_env_def) next assume "reachable_avoiding w y (unvisited e' x)" from unv have "v \ \ e' x" by (auto simp: unvisited_def) from \x \ y in stack e''\ have "x \ set (stack e')" by (simp add: e''_def precedes_mem) have "x = hd (stack e')" proof (rule ccontr) assume "x \ hd (stack e')" with \x \ set (stack e')\ \stack e' \ []\ have "x \ set (tl (stack e'))" by (metis hd_Cons_tl set_ConsD) with w \v \ \ e' x\ have "v \ \ e x" by auto moreover from post \stack e' \ []\ \x \ set (stack e')\ \x \ set (tl (stack e'))\ have "x \ set (tl (stack e))" unfolding post_dfs_def by (metis Un_iff self_append_conv2 set_append tl_append2) moreover from pre have "wf_env e" "stack e \ []" "v \ \ e (hd (stack e))" by (auto simp: pre_dfss_def) ultimately show "False" unfolding wf_env_def by (metis (no_types, lifting) distinct.simps(2) hd_Cons_tl insert_disjoint(2) list.set_sel(1) list.set_sel(2) mk_disjoint_insert) qed with \reachable_avoiding w y (unvisited e' x)\ \x \ y in stack e''\ \x \ y\ w wf' show ?thesis by (auto simp add: e''_def wf_env_def) qed qed qed qed moreover from wf' \\n. vsuccs e'' n \ successors n\ have "\n \ visited e'' - set (cstack e''). vsuccs e'' n = successors n" by (auto simp: wf_env_def e''_def split: if_splits) ultimately show ?thesis unfolding wf_env_def by (meson le_inf_iff) qed show "pre_dfss v e''" proof - from pre post have "v \ visited e''" by (auto simp: pre_dfss_def post_dfs_def sub_env_def e''_def) moreover { fix u assume u: "u \ vsuccs e'' v" have "u \ explored e'' \ \ e'' (hd (stack e''))" proof (cases "u = w") case True with post show ?thesis by (auto simp: post_dfs_def e''_def) next case False with u pre post have "u \ explored e \ \ e (hd (stack e))" by (auto simp: pre_dfss_def post_dfs_def e''_def) then show ?thesis proof assume "u \ explored e" with post show ?thesis by (auto simp: post_dfs_def sub_env_def e''_def) next assume "u \ \ e (hd (stack e))" with \wf_env e\ post \stack e \ []\ show ?thesis by (auto simp: e''_def dest: dfs_S_hd_stack) qed qed } moreover from pre post have "\n \ set (stack e''). reachable n v" unfolding pre_dfss_def post_dfs_def using e''_def by force moreover from \stack e' \ []\ have "stack e'' \ []" by (simp add: e''_def) moreover from \v \ \ e' (hd (stack e'))\ have "v \ \ e'' (hd (stack e''))" by (simp add: e''_def) moreover from pre post have "\ns. cstack e'' = v # ns" by (auto simp: pre_dfss_def post_dfs_def e''_def) ultimately show ?thesis using \wf_env e''\ unfolding pre_dfss_def by blast qed qed text \ Finally, the pre-condition for the recursive call to @{text dfss} at the end of the body of function @{text dfss} also holds if @{text unite} was applied. \ lemma pre_dfss_unite_pre_dfss: fixes e v w defines "e' \ unite v w e" defines "e'' \ e'\vsuccs := (\x. if x=v then vsuccs e' v \ {w} else vsuccs e' x)\" assumes pre: "pre_dfss v e" and w: "w \ successors v" "w \ vsuccs e v" "w \ visited e" "w \ explored e" shows "pre_dfss v e''" proof - from pre have wf: "wf_env e" by (simp add: pre_dfss_def) from pre have "v \ visited e" by (simp add: pre_dfss_def) from pre w have "v \ explored e" unfolding pre_dfss_def wf_env_def by (meson reachable_edge) from unite_stack[OF wf w] obtain pfx where pfx: "stack e = pfx @ stack e'" "stack e' \ []" "let cc = \ {\ e n |n. n \ set pfx \ {hd (stack e')}} in \ e' = (\x. if x \ cc then cc else \ e x)" "w \ \ e' (hd (stack e'))" by (auto simp: e'_def) define cc where "cc \ \ {\ e n |n. n \ set pfx \ {hd (stack e')}}" from unite_wf_env[OF pre w] have wf': "wf_env e'" by (simp add: e'_def) from \stack e = pfx @ stack e'\ \stack e' \ []\ have "hd (stack e) \ set pfx \ {hd (stack e')}" by (simp add: hd_append) with pre have "v \ cc" by (auto simp: pre_dfss_def cc_def) from S_reflexive[OF wf, of "hd (stack e')"] have "hd (stack e') \ cc" by (auto simp: cc_def) with pfx \v \ cc\ have "v \ \ e' (hd (stack e'))" by (auto simp: cc_def) from unite_sub_env[OF pre w] have "sub_env e e'" by (simp add: e'_def) have "wf_env e''" proof - from wf' have "\n \ visited e''. reachable (root e'') n" "distinct (stack e'')" "distinct (cstack e'')" "\n m. n \ m in stack e'' \ n \ m in cstack e''" "\n m. n \ m in stack e'' \ reachable m n" "explored e'' \ visited e''" "set (cstack e'') \ visited e''" "\n \ explored e''. \m. reachable n m \ m \ explored e''" "\n m. m \ \ e'' n \ (\ e'' n = \ e'' m)" "\n. n \ visited e'' \ \ e'' n = {n}" "\n \ set (stack e''). \m \ set (stack e''). n \ m \ \ e'' n \ \ e'' m = {}" "\ {\ e'' n | n. n \ set (stack e'')} = visited e'' - explored e''" "\n \ set (stack e''). \m \ \ e'' n. m \ set (cstack e'') \ m \ n in cstack e''" "\n. is_subscc (\ e'' n)" "\S \ sccs e''. is_scc S" "\ (sccs e'') = explored e''" by (auto simp: wf_env_def e''_def) moreover from wf' w \sub_env e e'\ have "\n. vsuccs e'' n \ successors n \ visited e''" by (auto simp: wf_env_def sub_env_def e''_def) moreover from wf' \v \ visited e\ \sub_env e e'\ have "\n. n \ visited e'' \ vsuccs e'' n = {}" by (auto simp: wf_env_def sub_env_def e''_def) moreover from wf' \v \ explored e\ have "\n \ explored e''. vsuccs e'' n = successors n" by (auto simp: wf_env_def e''_def e'_def unite_def) moreover from wf' \w \ successors v\ have "\n \ visited e'' - set (cstack e''). vsuccs e'' n = successors n" by (auto simp: wf_env_def e''_def e'_def unite_def) moreover have "\x y. x \ y in stack e'' \ x \ y \ (\u \ \ e'' x. \ reachable_avoiding u y (unvisited e'' x))" proof (clarify) fix x y u assume xy: "x \ y in stack e''" "x \ y" and u: "u \ \ e'' x" "reachable_avoiding u y (unvisited e'' x)" hence prec: "x \ y in stack e'" "u \ \ e' x" by (simp add: e''_def)+ show "False" proof (cases "x = hd (stack e')") case True with \v \ \ e' (hd (stack e'))\ have "unvisited e' x = unvisited e'' x \ (unvisited e' x = unvisited e'' x \ {(v,w)})" by (auto simp: e''_def unvisited_def split: if_splits) thus "False" proof assume "unvisited e' x = unvisited e'' x" with prec \x \ y\ \reachable_avoiding u y (unvisited e'' x)\ wf' show ?thesis unfolding wf_env_def by metis next assume "unvisited e' x = unvisited e'' x \ {(v,w)}" with \reachable_avoiding u y (unvisited e'' x)\[THEN ra_add_edge] have "reachable_avoiding u y (unvisited e' x) \ reachable_avoiding w y (unvisited e' x)" by auto thus ?thesis proof assume "reachable_avoiding u y (unvisited e' x)" with prec \x \ y\ wf' show ?thesis by (auto simp: wf_env_def) next assume "reachable_avoiding w y (unvisited e' x)" with \x = hd (stack e')\ \w \ \ e' (hd (stack e'))\ \x \ y in stack e'\ \x \ y\ wf' show ?thesis by (auto simp: wf_env_def) qed qed next case False with \x \ y in stack e'\ \stack e' \ []\ have "x \ set (tl (stack e'))" by (metis list.exhaust_sel precedes_mem(1) set_ConsD) with unite_S_tl[OF wf w] \u \ \ e' x\ have "u \ \ e x" by (simp add: e'_def) moreover from \x \ y in stack e'\ \stack e = pfx @ stack e'\ have "x \ y in stack e" by (simp add: precedes_append_left) moreover from \v \ \ e' (hd (stack e'))\ \x \ set (tl (stack e'))\ \stack e' \ []\ wf' have "v \ \ e' x" unfolding wf_env_def by (metis (no_types, lifting) Diff_cancel Diff_triv distinct.simps(2) insert_not_empty list.exhaust_sel list.set_sel(1) list.set_sel(2) mk_disjoint_insert) hence "unvisited e'' x = unvisited e' x" by (auto simp: unvisited_def e''_def split: if_splits) moreover from \x \ set (tl (stack e'))\ unite_S_tl[OF wf w] have "unvisited e' x = unvisited e x" by (simp add: unvisited_def e'_def unite_def) ultimately show ?thesis using \x \ y\ \reachable_avoiding u y (unvisited e'' x)\ wf by (auto simp: wf_env_def) qed qed ultimately show ?thesis unfolding wf_env_def by meson qed show "pre_dfss v e''" proof - from pre have "v \ visited e''" by (simp add: pre_dfss_def e''_def e'_def unite_def) moreover { fix u assume u: "u \ vsuccs e'' v" have "u \ explored e'' \ \ e'' (hd (stack e''))" proof (cases "u = w") case True with \w \ \ e' (hd (stack e'))\ show ?thesis by (simp add: e''_def) next case False with u have "u \ vsuccs e v" by (simp add: e''_def e'_def unite_def) with pre have "u \ explored e \ \ e (hd (stack e))" by (auto simp: pre_dfss_def) then show ?thesis proof assume "u \ explored e" thus ?thesis by (simp add: e''_def e'_def unite_def) next assume "u \ \ e (hd (stack e))" with \hd (stack e) \ set pfx \ {hd (stack e')}\ have "u \ cc" by (auto simp: cc_def) moreover from S_reflexive[OF wf, of "hd (stack e')"] pfx have "\ e' (hd (stack e')) = cc" by (auto simp: cc_def) ultimately show ?thesis by (simp add: e''_def) qed qed } hence "\w \ vsuccs e'' v. w \ explored e'' \ \ e'' (hd (stack e''))" by blast moreover from pre \stack e = pfx @ stack e'\ have "\n \ set (stack e''). reachable n v" by (auto simp: pre_dfss_def e''_def) moreover from \stack e' \ []\ have "stack e'' \ []" by (simp add: e''_def) moreover from \v \ \ e' (hd (stack e'))\ have "v \ \ e'' (hd (stack e''))" by (simp add: e''_def) moreover from pre have "\ns. cstack e'' = v # ns" by (auto simp: pre_dfss_def e''_def e'_def unite_def) ultimately show ?thesis using \wf_env e''\ unfolding pre_dfss_def by blast qed qed subsection \Lemmas establishing the post-conditions\ text \ Assuming the pre-condition of function @{text dfs} and the post-condition of the call to @{text dfss} in the body of that function, the post-condition of @{text dfs} is established. \ lemma pre_dfs_implies_post_dfs: fixes v e defines "e1 \ e\visited := visited e \ {v}, stack := (v # stack e), cstack:=(v # cstack e)\" defines "e' \ dfss v e1" defines "e'' \ e'\ cstack := tl(cstack e')\" assumes 1: "pre_dfs v e" and 2: "dfs_dfss_dom (Inl(v, e))" and 3: "post_dfss v e1 e'" shows "post_dfs v e (dfs v e)" proof - from 1 have wf: "wf_env e" by (simp add: pre_dfs_def) from 1 have v: "v \ visited e" by (simp add: pre_dfs_def) from 3 have wf': "wf_env e'" by (simp add: post_dfss_def) from 3 have cst': "cstack e' = v # cstack e" by (simp add: post_dfss_def e1_def) show ?thesis proof (cases "v = hd(stack e')") case True have notempty: "stack e' = v # stack e" proof - from 3 obtain ns where ns: "stack e1 = ns @ (stack e')" "stack e' \ []" by (auto simp: post_dfss_def) have "ns = []" proof (rule ccontr) assume "ns \ []" with ns have "hd ns = v" apply (simp add: e1_def) by (metis hd_append2 list.sel(1)) with True ns \ns \ []\ have "\ distinct (stack e1)" by (metis disjoint_iff_not_equal distinct_append hd_in_set) with wf v stack_visited[OF wf] show False by (auto simp: wf_env_def e1_def) qed with ns show ?thesis by (simp add: e1_def) qed have e2: "dfs v e = e'\sccs := sccs e' \ {\ e' v}, explored := explored e' \ (\ e' v), stack := tl (stack e'), cstack := tl (cstack e')\" (is "_ = ?e2") using True 2 dfs.psimps[of v e] unfolding e1_def e'_def by (fastforce simp: e1_def e'_def) have sub: "sub_env e e1" by (auto simp: sub_env_def e1_def) from notempty have stack2: "stack ?e2 = stack e" by (simp add: e1_def) moreover from 3 have "v \ visited ?e2" by (auto simp: post_dfss_def sub_env_def e1_def) moreover from sub 3 have "sub_env e e'" unfolding post_dfss_def by (auto elim: sub_env_trans) with stack2 have subenv: "sub_env e ?e2" by (fastforce simp: sub_env_def) moreover have "wf_env ?e2" proof - from wf' have "\n \ visited ?e2. reachable (root ?e2) n" "distinct (stack ?e2)" "\n. vsuccs ?e2 n \ successors n \ visited ?e2" "\n. n \ visited ?e2 \ vsuccs ?e2 n = {}" "\n m. m \ \ ?e2 n \ (\ ?e2 n = \ ?e2 m)" "\n. n \ visited ?e2 \ \ ?e2 n = {n}" "\n. is_subscc (\ ?e2 n)" "\ (sccs ?e2) = explored ?e2" by (auto simp: wf_env_def distinct_tl) moreover from 1 cst' have "distinct (cstack ?e2)" by (auto simp: pre_dfs_def wf_env_def) moreover from 1 stack2 have "\n m. n \ m in stack ?e2 \ reachable m n" by (auto simp: pre_dfs_def wf_env_def) moreover from 1 stack2 cst' have "\n m. n \ m in stack ?e2 \ n \ m in cstack ?e2" by (auto simp: pre_dfs_def wf_env_def) moreover from notempty wf' have "explored ?e2 \ visited ?e2" apply (simp add: wf_env_def) using stack_class[OF wf'] by (smt (verit, del_insts) Diff_iff insert_subset list.simps(15) subset_eq) moreover from 3 cst' have "set (cstack ?e2) \ visited ?e2" by (simp add: post_dfss_def wf_env_def e1_def) moreover { fix u assume "u \ explored ?e2" have "vsuccs ?e2 u = successors u" proof (cases "u \ explored e'") case True with wf' show ?thesis by (auto simp: wf_env_def) next case False with \u \ explored ?e2\ have "u \ \ e' v" by simp show ?thesis proof (cases "u = v") case True with 3 show ?thesis by (auto simp: post_dfss_def) next case False have "u \ visited e' - set (cstack e')" proof from notempty \u \ \ e' v\ stack_class[OF wf'] False show "u \ visited e'" by auto next show "u \ set (cstack e')" proof assume u: "u \ set (cstack e')" with notempty \u \ \ e' v\ \wf_env e'\ have "u \ v in cstack e'" by (auto simp: wf_env_def) with cst' u False wf' show "False" unfolding wf_env_def by (metis head_precedes precedes_antisym) qed qed with 3 show ?thesis by (auto simp: post_dfss_def wf_env_def) qed qed } note explored_vsuccs = this moreover have "\n \ explored ?e2. \m. reachable n m \ m \ explored ?e2" proof (clarify) fix x y assume asm: "x \ explored ?e2" "reachable x y" show "y \ explored ?e2" proof (cases "x \ explored e'") case True with \wf_env e'\ \reachable x y\ show ?thesis by (simp add: wf_env_def) next case False with asm have "x \ \ e' v" by simp with \explored ?e2 \ visited ?e2\ have "x \ visited e'" by auto from \x \ \ e' v\ wf' have "reachable v x" by (auto simp: wf_env_def is_subscc_def) have "y \ visited e'" proof (rule ccontr) assume "y \ visited e'" with reachable_visited[OF wf' \x \ visited e'\ \reachable x y\] obtain n m where "n \ visited e'" "m \ successors n - vsuccs e' n" "reachable x n" "reachable m y" by blast from wf' \m \ successors n - vsuccs e' n\ have "n \ explored e'" by (auto simp: wf_env_def) obtain n' where "n' \ set (stack e')" "n \ \ e' n'" by (rule visited_unexplored[OF wf' \n \ visited e'\ \n \ explored e'\]) have "n' = v" proof (rule ccontr) assume "n' \ v" with \n' \ set (stack e')\ \v = hd (stack e')\ have "n' \ set (tl (stack e'))" by (metis emptyE hd_Cons_tl set_ConsD set_empty) moreover from \n \ \ e' n'\ \wf_env e'\ have "reachable n n'" by (auto simp: wf_env_def is_subscc_def) with \reachable v x\ \reachable x n\ reachable_trans have "reachable v n'" by blast ultimately show "False" using 3 \v = hd (stack e')\ by (auto simp: post_dfss_def) qed with \n \ \ e' n'\ \m \ successors n - vsuccs e' n\ explored_vsuccs show "False" by auto qed show ?thesis proof (cases "y \ explored e'") case True then show ?thesis by simp next case False obtain n where ndef: "n \ set (stack e')" "(y \ \ e' n)" by (rule visited_unexplored[OF wf' \y \ visited e'\ False]) show ?thesis proof (cases "n = v") case True with ndef show ?thesis by simp next case False with ndef notempty have "n \ set (tl (stack e'))" by simp moreover from wf' ndef have "reachable y n" by (auto simp: wf_env_def is_subscc_def) with \reachable v x\ \reachable x y\ have "reachable v n" by (meson reachable_trans) ultimately show ?thesis using \v = hd (stack e')\ 3 by (simp add: post_dfss_def) qed qed qed qed moreover from 3 cst' have "\n \ visited ?e2 - set (cstack ?e2). vsuccs ?e2 n = successors n" apply (simp add: post_dfss_def wf_env_def) by (metis (no_types, lifting) Diff_empty Diff_iff empty_set insertE list.exhaust_sel list.sel(1) list.simps(15)) moreover from wf' notempty have "\n m. n \ set (stack ?e2) \ m \ set (stack ?e2) \ n \ m \ (\ ?e2 n \ \ ?e2 m = {})" by (simp add: wf_env_def) moreover have "\ {\ ?e2 n | n . n \ set (stack ?e2)} = visited ?e2 - explored ?e2" proof - from wf' notempty have "(\ {\ ?e2 n | n . n \ set (stack ?e2)}) \ \ e' v = {}" by (auto simp: wf_env_def) with notempty have "\ {\ ?e2 n | n . n \ set (stack ?e2)} = (\ {\ e' n | n . n \ set (stack e')}) - \ e' v" by auto also from wf' have "\ = (visited e' - explored e') - \ e' v" by (simp add: wf_env_def) finally show ?thesis by auto qed moreover have "\n \ set (stack ?e2). \m \ \ ?e2 n. m \ set (cstack ?e2) \ m \ n in cstack ?e2" proof (clarsimp simp: cst') fix n m assume "n \ set (tl (stack e'))" "m \ \ e' n" "m \ set (cstack e)" with 3 have "m \ \ e n" by (auto simp: post_dfss_def e1_def) with wf notempty \n \ set (tl (stack e'))\ \m \ set (cstack e)\ show "m \ n in cstack e" by (auto simp: wf_env_def) qed moreover { fix x y u assume xy: "x \ y in stack ?e2" "x \ y" and u: "u \ \ ?e2 x" "reachable_avoiding u y (unvisited ?e2 x)" from xy notempty stack2 have "x \ y in stack e'" by (metis head_precedes insert_iff list.simps(15) precedes_in_tail precedes_mem(2)) with wf' \x \ y\ u have "False" by (auto simp: wf_env_def unvisited_def) } moreover have "\S \ sccs ?e2. is_scc S" proof (clarify) fix S assume asm: "S \ sccs ?e2" show "is_scc S" proof (cases "S = \ e' v") case True with S_reflexive[OF wf'] have "S \ {}" by blast from wf' True have subscc: "is_subscc S" by (simp add: wf_env_def) { assume "\ is_scc S" with \S \ {}\ \is_subscc S\ obtain S' where S'_def: "S' \ S" "S \ S'" "is_subscc S'" unfolding is_scc_def by blast then obtain x where "x \ S' \ x \ S" by blast with True S'_def wf' have xv: "reachable v x \ reachable x v" unfolding wf_env_def is_subscc_def by (metis in_mono) from \\v w. w \ \ ?e2 v \ (\ ?e2 v = \ ?e2 w)\ have "v \ explored ?e2" by auto with \\x \ explored ?e2. \y. reachable x y \ y \ explored ?e2\ xv \S = \ e' v\ \x \ S' \ x \ S\ have "x \ explored e'" by auto with wf' xv have "v \ explored e'" by (auto simp: wf_env_def) with notempty have "False" by (auto intro: stack_unexplored[OF wf']) } then show ?thesis by blast next case False with asm wf' show ?thesis by (auto simp: wf_env_def) qed qed ultimately show ?thesis unfolding wf_env_def by meson qed moreover from \wf_env ?e2\ have "v \ explored ?e2" by (auto simp: wf_env_def) moreover from 3 have "vsuccs ?e2 v = successors v" by (simp add: post_dfss_def) moreover from 1 3 have "\w \ visited e. vsuccs ?e2 w = vsuccs e w" by (auto simp: pre_dfs_def post_dfss_def e1_def) moreover from stack2 1 have "\n \ set (stack ?e2). reachable n v" by (simp add: pre_dfs_def) moreover from stack2 have "\ns. stack e = ns @ (stack ?e2)" by auto moreover from 3 have "\n \ set (stack ?e2). \ ?e2 n = \ e n" by (auto simp: post_dfss_def e1_def) moreover from cst' have "cstack ?e2 = cstack e" by simp ultimately show ?thesis unfolding post_dfs_def using e2 by simp next case False with 2 have e': "dfs v e = e''" by (simp add: dfs.psimps e''_def e'_def e1_def) moreover have "wf_env e''" proof - from wf' have "\n \ visited e''. reachable (root e'') n" "distinct (stack e'')" "distinct (cstack e'')" "\n m. n \ m in stack e'' \ reachable m n" "explored e'' \ visited e''" "\n \ explored e''. \m. reachable n m \ m \ explored e''" "\n. vsuccs e'' n \ successors n \ visited e''" "\n. n \ visited e'' \ vsuccs e'' n = {}" "\n \ explored e''. vsuccs e'' n = successors n" "\n m. m \ \ e'' n \ (\ e'' n = \ e'' m)" "\n. n \ visited e'' \ \ e'' n = {n}" "\n \ set (stack e''). \m \ set (stack e''). n \ m \ \ e'' n \ \ e'' m = {}" "\ {\ e'' n | n. n \ set (stack e'')} = visited e'' - explored e''" "\n. is_subscc (\ e'' n)" "\S \ sccs e''. is_scc S" "\ (sccs e'') = explored e''" by (auto simp: e''_def wf_env_def distinct_tl) moreover have "\n m. n \ m in stack e'' \ n \ m in cstack e''" proof (clarsimp simp add: e''_def) fix n m assume nm: "n \ m in stack e'" with 3 have "n \ m in cstack e'" unfolding post_dfss_def wf_env_def by meson moreover have "n \ v" proof assume "n = v" with nm have "n \ set (stack e')" by (simp add: precedes_mem) with 3 \n = v\ have "v = hd (stack e')" unfolding post_dfss_def wf_env_def by (metis (no_types, opaque_lifting) IntI equals0D list.set_sel(1)) with \v \ hd (stack e')\ show "False" by simp qed ultimately show "n \ m in tl (cstack e')" by (simp add: cst' precedes_in_tail) qed moreover from 3 have "set (cstack e'') \ visited e''" by (simp add: post_dfss_def wf_env_def e''_def e1_def subset_eq) moreover from 3 have "\n \ visited e'' - set (cstack e''). vsuccs e'' n = successors n" apply (simp add: post_dfss_def wf_env_def e''_def e1_def) by (metis (no_types, opaque_lifting) DiffE DiffI set_ConsD) moreover have "\n \ set (stack e''). \m \ \ e'' n. m \ set (cstack e'') \ m \ n in cstack e''" proof (clarsimp simp: e''_def) fix m n assume asm: "n \ set (stack e')" "m \ \ e' n" "m \ set (tl (cstack e'))" with wf' cst' have "m \ v" "m \ n in cstack e'" by (auto simp: wf_env_def) with cst' show "m \ n in tl (cstack e')" by (simp add: precedes_in_tail) qed moreover from wf' have "(\x y. x \ y in stack e'' \ x \ y \ (\u \ \ e'' x. \ reachable_avoiding u y (unvisited e'' x)))" by (force simp: e''_def wf_env_def unvisited_def) ultimately show ?thesis unfolding wf_env_def by blast qed moreover from 3 have "v \ visited e''" by (auto simp: post_dfss_def sub_env_def e''_def e1_def) moreover have subenv: "sub_env e e''" proof - have "sub_env e e1" by (auto simp: sub_env_def e1_def) with 3 have "sub_env e e'" by (auto simp: post_dfss_def elim: sub_env_trans) thus ?thesis by (auto simp add: sub_env_def e''_def) qed moreover from 3 have "vsuccs e'' v = successors v" by (simp add: post_dfss_def e''_def) moreover from 1 3 have "\w \ visited e. vsuccs e'' w = vsuccs e w" by (auto simp: pre_dfs_def post_dfss_def e1_def e''_def) moreover from 3 have "\n \ set (stack e''). reachable n v" by (auto simp: e''_def post_dfss_def) moreover from 3 \v \ hd (stack e')\ have "\ns. stack e = ns @ (stack e'')" apply (simp add: post_dfss_def e''_def e1_def) by (metis append_Nil list.sel(1) list.sel(3) tl_append2) moreover from 3 have "stack e'' \ []" "v \ \ e'' (hd (stack e''))" "\n \ set (tl (stack e'')). \ e'' n = \ e n" by (auto simp: post_dfss_def e1_def e''_def) moreover from cst' have "cstack e'' = cstack e" by (simp add: e''_def) ultimately show ?thesis unfolding post_dfs_def by blast qed qed text \ The following lemma is central for proving partial correctness: assuming termination (represented by the predicate @{text dfs_dfss_dom}) and the pre-condition of the functions, both @{text dfs} and @{text dfss} establish their post-conditions. The first part of the theorem follows directly from the preceding lemma and the computational induction rule generated by Isabelle, the second part is proved directly, distinguishing the different cases in the definition of function @{text dfss}. \ lemma pre_post: shows "\dfs_dfss_dom (Inl(v,e)); pre_dfs v e\ \ post_dfs v e (dfs v e)" "\dfs_dfss_dom (Inr(v,e)); pre_dfss v e\ \ post_dfss v e (dfss v e)" proof (induct rule: dfs_dfss.pinduct) fix v e assume dom: "dfs_dfss_dom (Inl(v,e))" and predfs: "pre_dfs v e" and prepostdfss: "\e1. \ e1 = e \visited := visited e \ {v}, stack := v # stack e, cstack := v # cstack e\; pre_dfss v e1 \ \ post_dfss v e1 (dfss v e1)" then show "post_dfs v e (dfs v e)" using pre_dfs_implies_post_dfs pre_dfs_pre_dfss by auto next fix v e assume dom: "dfs_dfss_dom (Inr(v,e))" and predfss: "pre_dfss v e" and prepostdfs: "\vs w. \ vs = successors v - vsuccs e v; vs \ {}; w = (SOME x. x \ vs); w \ explored e; w \ visited e; pre_dfs w e \ \ post_dfs w e (dfs w e)" and prepostdfss: "\vs w e' e''. \ vs = successors v - vsuccs e v; vs \ {}; w = (SOME x. x \ vs); e' = (if w \ explored e then e else if w \ visited e then dfs w e else unite v w e); e'' = e'\vsuccs := \x. if x = v then vsuccs e' v \ {w} else vsuccs e' x\ ; pre_dfss v e'' \ \ post_dfss v e'' (dfss v e'')" show "post_dfss v e (dfss v e)" proof - let ?vs = "successors v - vsuccs e v" from predfss have wf: "wf_env e" by (simp add: pre_dfss_def) from predfss have "v \ visited e" by (simp add: pre_dfss_def) from predfss have "v \ explored e" by (meson DiffD2 list.set_sel(1) pre_dfss_def stack_class) show ?thesis proof (cases "?vs = {}") case True with dom have "dfss v e = e" by (simp add: dfss.psimps) moreover from True wf have "vsuccs e v = successors v" unfolding wf_env_def by (meson Diff_eq_empty_iff le_infE subset_antisym) moreover have "sub_env e e" by (simp add: sub_env_def) moreover from predfss \vsuccs e v = successors v\ have "\w \ successors v. w \ explored e \ \ e (hd (stack e))" "\n \ set (stack e). reachable n v" "stack e \ []" "v \ \ e (hd (stack e))" by (auto simp: pre_dfss_def) moreover have "\ns. stack e = ns @ (stack e)" by simp moreover { fix n assume asm: "hd (stack e) = v" "n \ set (tl (stack e))" "reachable v n" with \stack e \ []\ have "v \ n in stack e" by (metis head_precedes hd_Cons_tl list.set_sel(2)) moreover from wf \stack e \ []\ asm have "v \ n" unfolding wf_env_def by (metis distinct.simps(2) list.exhaust_sel) moreover from wf have "v \ \ e v" by (rule S_reflexive) moreover { fix a b assume "a \ \ e v" "b \ successors a - vsuccs e a" with \vsuccs e v = successors v\ have "a \ v" by auto from \stack e \ []\ \hd (stack e) = v\ have "v \ set (stack e)" by auto with \a \ v\ \a \ \ e v\ wf have "a \ visited e" unfolding wf_env_def by (metis singletonD) have "False" proof (cases "a \ set (cstack e)") case True with \v \ set (stack e)\ \a \ \ e v\ \wf_env e\ have "a \ v in cstack e" by (auto simp: wf_env_def) moreover from predfss obtain ns where "cstack e = v # ns" by (auto simp: pre_dfss_def) moreover from wf have "distinct (cstack e)" by (simp add: wf_env_def) ultimately have "a = v" using tail_not_precedes by force with \a \ v\ show ?thesis .. next case False with \a \ visited e\ wf have "vsuccs e a = successors a" by (auto simp: wf_env_def) with \b \ successors a - vsuccs e a\ show ?thesis by simp qed } hence "unvisited e v = {}" by (auto simp: unvisited_def) ultimately have "\ reachable_avoiding v n {}" using wf unfolding wf_env_def by metis with \reachable v n\ have "False" by (simp add: ra_empty) } ultimately show ?thesis using wf by (auto simp: post_dfss_def) next case vs_case: False define w where "w = (SOME x. x \ ?vs)" define e' where "e' = (if w \ explored e then e else if w \ visited e then dfs w e else unite v w e)" define e'' where "e'' = (e'\vsuccs := \x. if x=v then vsuccs e' v \ {w} else vsuccs e' x\)" from dom vs_case have dfss: "dfss v e = dfss v e''" apply (simp add: dfss.psimps e''_def) using e'_def w_def by auto from vs_case have wvs: "w \ ?vs" unfolding w_def by (metis some_in_eq) show ?thesis proof (cases "w \ explored e") case True hence e': "e' = e" by (simp add: e'_def) with predfss wvs True have "pre_dfss v e''" by (auto simp: e''_def pre_dfss_explored_pre_dfss) with prepostdfss vs_case have post'': "post_dfss v e'' (dfss v e'')" by (auto simp: w_def e'_def e''_def) moreover from post'' have "\u \ visited e - {v}. vsuccs (dfss v e'') u = vsuccs e u" by (auto simp: post_dfss_def e' e''_def) moreover have "sub_env e e''" by (auto simp: sub_env_def e' e''_def) with post'' have "sub_env e (dfss v e'')" by (auto simp: post_dfss_def elim: sub_env_trans) moreover from e' have "stack e'' = stack e" "\ e'' = \ e" by (auto simp add: e''_def) moreover have "cstack e'' = cstack e" by (simp add: e''_def e') ultimately show ?thesis by (auto simp: dfss post_dfss_def) next case notexplored: False then show ?thesis proof (cases "w \ visited e") case True with e'_def notexplored have "e' = dfs w e" by auto with True notexplored pre_dfss_pre_dfs predfss prepostdfs vs_case w_def have postdfsw: "post_dfs w e e'" by (metis DiffD1 some_in_eq) with predfss wvs True \e' = dfs w e\ have "pre_dfss v e''" by (auto simp: e''_def pre_dfss_post_dfs_pre_dfss) with prepostdfss vs_case have post'': "post_dfss v e'' (dfss v e'')" by (auto simp: w_def e'_def e''_def) moreover have "\u \ visited e - {v}. vsuccs (dfss v e'') u = vsuccs e u" proof fix u assume "u \ visited e - {v}" with postdfsw have u: "vsuccs e' u = vsuccs e u" "u \ visited e'' - {v}" by (auto simp: post_dfs_def sub_env_def e''_def) with post'' have "vsuccs (dfss v e'') u = vsuccs e'' u" by (auto simp: post_dfss_def) with u show "vsuccs (dfss v e'') u = vsuccs e u" by (simp add: e''_def) qed moreover have "sub_env e (dfss v e'')" proof - from postdfsw have "sub_env e e'" by (simp add: post_dfs_def) moreover have "sub_env e' e''" by (auto simp: sub_env_def e''_def) moreover from post'' have "sub_env e'' (dfss v e'')" by (simp add: post_dfss_def) ultimately show ?thesis by (metis sub_env_trans) qed moreover from postdfsw post'' have "\ns. stack e = ns @ (stack (dfss v e''))" by (auto simp: post_dfs_def post_dfss_def e''_def) moreover { fix n assume n: "n \ set (tl (stack (dfss v e'')))" with post'' have "\ (dfss v e'') n = \ e' n" by (simp add: post_dfss_def e''_def) moreover from \pre_dfss v e''\ n post'' have "stack e' \ [] \ n \ set (tl (stack e''))" apply (simp add: pre_dfss_def post_dfss_def e''_def) by (metis (no_types, lifting) Un_iff list.set_sel(2) self_append_conv2 set_append tl_append2) with postdfsw have "\ e' n = \ e n" apply (simp add: post_dfs_def e''_def) by (metis list.set_sel(2)) ultimately have "\ (dfss v e'') n = \ e n" by simp } moreover from postdfsw have "cstack e'' = cstack e" by (auto simp: post_dfs_def e''_def) ultimately show ?thesis by (auto simp: dfss post_dfss_def) next case False hence e': "e' = unite v w e" using notexplored using e'_def by simp from False have "w \ visited e" by simp from wf wvs notexplored False obtain pfx where pfx: "stack e = pfx @ (stack e')" "stack e' \ []" unfolding e' by (blast dest: unite_stack) from predfss wvs notexplored False \e' = unite v w e\ have "pre_dfss v e''" by (auto simp: e''_def pre_dfss_unite_pre_dfss) with prepostdfss vs_case \e' = unite v w e\ \w \ explored e\ \w \ visited e\ have post'': "post_dfss v e'' (dfss v e'')" by (auto simp: w_def e''_def) moreover from post'' have "\u \ visited e - {v}. vsuccs (dfss v e'') u = vsuccs e u" by (auto simp: post_dfss_def e''_def e' unite_def) moreover have "sub_env e (dfss v e'')" proof - from predfss wvs \w \ visited e\ notexplored have "sub_env e e'" unfolding e' by (blast dest: unite_sub_env) moreover have "sub_env e' e''" by (auto simp: sub_env_def e''_def) moreover from post'' have "sub_env e'' (dfss v e'')" by (simp add: post_dfss_def) ultimately show ?thesis by (metis sub_env_trans) qed moreover from post'' \stack e = pfx @ stack e'\ have "\ns. stack e = ns @ (stack (dfss v e''))" by (auto simp: post_dfss_def e''_def) moreover { fix n assume n: "n \ set (tl (stack (dfss v e'')))" with post'' have "\ (dfss v e'') n = \ e'' n" by (simp add: post_dfss_def) moreover from n post'' \stack e' \ []\ have "n \ set (tl (stack e''))" apply (simp add: post_dfss_def e''_def) by (metis (no_types, lifting) Un_iff list.set_sel(2) self_append_conv2 set_append tl_append2) with wf wvs \w \ visited e\ notexplored have "\ e'' n = \ e n" by (auto simp: e''_def e' dest: unite_S_tl) ultimately have "\ (dfss v e'') n = \ e n" by simp } moreover from post'' have "cstack (dfss v e'') = cstack e" by (simp add: post_dfss_def e''_def e' unite_def) ultimately show ?thesis by (simp add: dfss post_dfss_def) qed qed qed qed qed text \ We can now show partial correctness of the algorithm: applied to some node @{text "v"} and the empty environment, it computes the set of strongly connected components in the subgraph reachable from node @{text "v"}. In particular, if @{text "v"} is a root of the graph, the algorithm computes the set of SCCs of the graph. \ theorem partial_correctness: fixes v defines "e \ dfs v (init_env v)" assumes "dfs_dfss_dom (Inl (v, init_env v))" shows "sccs e = {S . is_scc S \ (\n\S. reachable v n)}" (is "_ = ?rhs") proof - from assms init_env_pre_dfs[of v] have post: "post_dfs v (init_env v) e" by (auto dest: pre_post) hence wf: "wf_env e" by (simp add: post_dfs_def) from post have "cstack e = []" by (auto simp: post_dfs_def init_env_def) have "stack e = []" proof (rule ccontr) assume "stack e \ []" hence "hd (stack e) \ hd (stack e) in stack e" by simp with wf \cstack e = []\ show "False" unfolding wf_env_def by (metis empty_iff empty_set precedes_mem(2)) qed with post have vexp: "v \ explored e" by (simp add: post_dfs_def) from wf \stack e = []\ have "explored e = visited e" by (auto simp: wf_env_def) have "sccs e \ ?rhs" proof fix S assume S: "S \ sccs e" with wf have "is_scc S" by (simp add: wf_env_def) moreover from S wf have "S \ explored e" unfolding wf_env_def by blast with post \explored e = visited e\ have "\n\S. reachable v n" by (auto simp: post_dfs_def wf_env_def sub_env_def init_env_def) ultimately show "S \ ?rhs" by auto qed moreover { fix S assume "is_scc S" "\n\S. reachable v n" from \\n\S. reachable v n\ vexp wf have "S \ \ (sccs e)" unfolding wf_env_def by (metis subset_eq) with \is_scc S\ obtain S' where S': "S' \ sccs e \ S \ S' \ {}" unfolding is_scc_def by (metis Union_disjoint inf.absorb_iff2 inf_commute) with wf have "is_scc S'" by (simp add: wf_env_def) with S' \is_scc S\ have "S \ sccs e" by (auto dest: scc_partition) } ultimately show ?thesis by blast qed section \Proof of termination and total correctness\ text \ We define a binary relation on the arguments of functions @{text dfs} and @{text dfss}, and prove that this relation is well-founded and that all calls within the function bodies respect the relation, assuming that the pre-conditions of the initial function call are satisfied. By well-founded induction, we conclude that the pre-conditions of the functions are sufficient to ensure termination. Following the internal representation of the two mutually recursive functions in Isabelle as a single function on the disjoint sum of the types of arguments, our relation is defined as a set of argument pairs injected into the sum type. The left injection @{text Inl} takes arguments of function @{text dfs}, the right injection @{text Inr} takes arguments of function @{text dfss}.\footnote{Note that the types of the arguments of @{text dfs} and @{text dfss} are actually identical. We nevertheless use the sum type in order to remember the function that was called.} The conditions on the arguments in the definition of the relation overapproximate the arguments in the actual calls. \ definition dfs_dfss_term::"(('v \ 'v env + 'v \ 'v env) \ ('v \ 'v env + 'v \ 'v env)) set" where "dfs_dfss_term \ { (Inr(v, e1), Inl(v, e)) | v e e1. v \ vertices - visited e \ visited e1 = visited e \ {v} } \ { (Inl(w, e), Inr(v, e)) | v w e. v \ vertices} \ { (Inr(v, e''), Inr(v, e)) | v e e''. v \ vertices \ sub_env e e'' \ (\w \ vertices. w \ vsuccs e v \ w \ vsuccs e'' v)}" text \ Informally, termination is ensured because at each call, either a new vertex is visited (hence the complement of the set of visited nodes w.r.t. the finite set of vertices decreases) or a new successor is added to the set @{text "vsuccs e v"} of some vertex @{text v}. In order to make this argument formal, we inject the argument tuples that appear in our relation into tuples consisting of the sets mentioned in the informal argument. However, there is one added complication because the call of @{text dfs} from @{text dfss} does not immediately add the vertex to the set of visited nodes (this happens only at the beginning of function @{text dfs}). We therefore add a third component of $0$ or $1$ to these tuples, reflecting the fact that there can only be one call of @{text dfs} from @{text dfss} for a given vertex @{text v}. \ fun dfs_dfss_to_tuple where "dfs_dfss_to_tuple (Inl(v::'v, e::'v env)) = (vertices - visited e, vertices \ vertices - {(u,u') | u u'. u' \ vsuccs e u}, 0)" | "dfs_dfss_to_tuple (Inr(v::'v, e::'v env)) = (vertices - visited e, vertices \ vertices - {(u,u') | u u'. u' \ vsuccs e u}, 1::nat)" text \ The triples defined in this way can be ordered lexicographically (with the first two components ordered as finite subsets and the third one following the predecessor relation on natural numbers). We prove that the injection of the above relation into sets of triples respects the lexicographic ordering and conclude that our relation is well-founded. \ lemma wf_term: "wf dfs_dfss_term" proof - let ?r = "(finite_psubset :: ('v set \ 'v set) set) <*lex*> (finite_psubset :: ((('v \ 'v) set) \ ('v \ 'v) set) set) <*lex*> pred_nat" have "wf (finite_psubset :: ('v set \ 'v set) set)" by (rule wf_finite_psubset) moreover have "wf (finite_psubset :: ((('v \ 'v) set) \ ('v \ 'v) set) set)" by (rule wf_finite_psubset) ultimately have "wf ?r" using wf_pred_nat by blast moreover have "dfs_dfss_term \ inv_image ?r dfs_dfss_to_tuple" proof (clarify) fix a b assume "(a,b) \ dfs_dfss_term" hence "(\v w e e''. a = Inr(v,e'') \ b = Inr(v,e) \ v \ vertices \ sub_env e e'' \ w \ vertices \ w \ vsuccs e v \ w \ vsuccs e'' v) \ (\v e e1. a = Inr(v,e1) \ b = Inl(v,e) \ v \ vertices - visited e \ visited e1 = visited e \ {v}) \ (\v w e. a = Inl(w,e) \ b = Inr(v,e))" (is "?c1 \ ?c2 \ ?c3") by (auto simp: dfs_dfss_term_def) then show "(a,b) \ inv_image ?r dfs_dfss_to_tuple" proof assume "?c1" then obtain v w e e'' where ab: "a = Inr(v, e'')" "b = Inr(v,e)" and vw: "v \ vertices" "w \ vertices" "w \ vsuccs e'' v" "w \ vsuccs e v" and sub: "sub_env e e''" by blast from sub have "vertices - visited e'' \ vertices - visited e" by (auto simp: sub_env_def) moreover from sub vw have "(vertices \ vertices - {(u,u') | u u'. u' \ vsuccs e'' u}) \ (vertices \ vertices - {(u,u') | u u'. u' \ vsuccs e u})" by (auto simp: sub_env_def) ultimately show ?thesis using vfin ab by auto next assume "?c2 \ ?c3" with vfin show ?thesis by (auto simp: pred_nat_def) qed qed ultimately show ?thesis using wf_inv_image wf_subset by blast qed text \ The following theorem establishes sufficient conditions that ensure termination of the two functions @{text dfs} and @{text dfss}. The proof proceeds by well-founded induction using the relation @{text dfs_dfss_term}. Isabelle represents the termination domains of the functions by the predicate @{text dfs_dfss_dom} and generates a theorem @{text dfs_dfss.domintros} for proving membership of arguments in the termination domains. The actual formulation is a litte technical because the mutual induction must again be encoded in a single induction argument over the sum type representing the arguments of both functions. \ theorem dfs_dfss_termination: "\v \ vertices ; pre_dfs v e\ \ dfs_dfss_dom(Inl(v, e))" "\v \ vertices ; pre_dfss v e\ \ dfs_dfss_dom(Inr(v, e))" proof - { fix args have "(case args of Inl(v,e) \ v \ vertices \ pre_dfs v e | Inr(v,e) \ v \ vertices \ pre_dfss v e) \ dfs_dfss_dom args" (is "?P args \ ?Q args") proof (rule wf_induct[OF wf_term]) fix arg :: "('v \ 'v env) + ('v \ 'v env)" assume ih: "\ arg'. (arg', arg) \ dfs_dfss_term \ (?P arg' \ ?Q arg')" show "?P arg \ ?Q arg" proof assume P: "?P arg" show "?Q arg" proof (cases arg) case (Inl a) then obtain v e where a: "arg = Inl(v, e)" using dfs.cases by metis with P have pre: "v \ vertices \ pre_dfs v e" by simp let ?e1 = "e\visited := visited e \ {v}, stack := v # stack e, cstack := v # cstack e\" let ?recarg = "Inr(v, ?e1)" from a pre have "(?recarg, arg) \ dfs_dfss_term" by (auto simp: pre_dfs_def dfs_dfss_term_def) moreover from pre have "?P ?recarg" by (auto dest: pre_dfs_pre_dfss) ultimately have "?Q ?recarg" using ih a by auto then have "?Q (Inl(v, e))" by (auto intro: dfs_dfss.domintros) then show ?thesis by (simp add: a) next case (Inr b) then obtain v e where b: "arg = Inr(v, e)" using dfs.cases by metis with P have pre: "v \ vertices \ pre_dfss v e" by simp let ?sw = "SOME w. w \ successors v \ w \ vsuccs e v" have "?Q (Inr(v, e))" proof (rule dfs_dfss.domintros) fix w assume "w \ successors v" "?sw \ explored e" "?sw \ visited e" "\ dfs_dfss_dom (Inl (?sw, e))" show "w \ vsuccs e v" proof (rule ccontr) assume "w \ vsuccs e v" with \w \ successors v\ have sw: "?sw \ successors v - vsuccs e v" by (metis (mono_tags, lifting) Diff_iff some_eq_imp) with pre \?sw \ visited e\ have "pre_dfs ?sw e" by (blast intro: pre_dfss_pre_dfs) moreover from pre sw sclosed have "?sw \ vertices" by blast moreover from pre have "(Inl(?sw,e), Inr(v,e)) \ dfs_dfss_term" by (simp add: dfs_dfss_term_def) ultimately have "dfs_dfss_dom (Inl(?sw,e))" using ih b by auto with \\ dfs_dfss_dom (Inl (?sw, e))\ show "False" .. qed next let ?e' = "dfs ?sw e" let ?e''= "?e'\vsuccs := \x. if x = v then vsuccs ?e' v \ {?sw} else vsuccs ?e' x\" fix w assume asm: "w \ successors v" "w \ vsuccs e v" "?sw \ visited e" "?sw \ explored e" from \w \ successors v\ \w \ vsuccs e v\ have sw: "?sw \ successors v - vsuccs e v" by (metis (no_types, lifting) Diff_iff some_eq_imp) with pre \?sw \ visited e\ have "pre_dfs ?sw e" by (blast intro: pre_dfss_pre_dfs) moreover from pre sw sclosed have "?sw \ vertices" by blast moreover from pre have "(Inl(?sw, e), Inr(v,e)) \ dfs_dfss_term" by (simp add: dfs_dfss_term_def) ultimately have "dfs_dfss_dom (Inl(?sw, e))" using ih b by auto from this \pre_dfs ?sw e\ have post: "post_dfs ?sw e ?e'" by (rule pre_post) hence "sub_env e ?e'" by (simp add: post_dfs_def) moreover have "sub_env ?e' ?e''" by (auto simp: sub_env_def) ultimately have "sub_env e ?e''" by (rule sub_env_trans) with pre \?sw \ vertices\ sw have "(Inr(v, ?e''), Inr(v, e)) \ dfs_dfss_term" by (auto simp: dfs_dfss_term_def) moreover from pre post sw \?sw \ visited e\ have "pre_dfss v ?e''" by (blast intro: pre_dfss_post_dfs_pre_dfss) ultimately show "dfs_dfss_dom(Inr(v, ?e''))" using pre ih b by auto next let ?e'' = "e\vsuccs := \x. if x = v then vsuccs e v \ {?sw} else vsuccs e x\" fix w assume "w \ successors v" "w \ vsuccs e v" "?sw \ visited e" "?sw \ explored e" with pre have "False" unfolding pre_dfss_def wf_env_def by (meson subsetD) thus "?Q (Inr(v, ?e''))" by simp next fix w assume asm: "w \ successors v" "w \ vsuccs e v" "?sw \ visited e" "?sw \ explored e" let ?e'' = "e\vsuccs := \x. if x = v then vsuccs e v \ {?sw} else vsuccs e x\" let ?recarg = "Inr(v, ?e'')" from \w \ successors v\ \w \ vsuccs e v\ have sw: "?sw \ successors v - vsuccs e v" by (metis (no_types, lifting) Diff_iff some_eq_imp) have "(?recarg, arg) \ dfs_dfss_term" proof - have "sub_env e ?e''" by (auto simp: sub_env_def) moreover from sw pre sclosed have "\u \ vertices. u \ vsuccs e v \ u \ vsuccs ?e'' v" by auto ultimately show ?thesis using pre b unfolding dfs_dfss_term_def by blast qed moreover from pre sw \?sw \ explored e\ have "?P ?recarg" by (auto dest: pre_dfss_explored_pre_dfss) ultimately show "?Q ?recarg" using ih b by blast next fix w assume asm: "w \ successors v" "w \ vsuccs e v" "?sw \ visited e" "?sw \ explored e" let ?eu = "unite v ?sw e" let ?e'' = "?eu\vsuccs := \x. if x = v then vsuccs ?eu v \ {?sw} else vsuccs ?eu x\" let ?recarg = "Inr(v, ?e'')" from \w \ successors v\ \w \ vsuccs e v\ have sw: "?sw \ successors v - vsuccs e v" by (metis (no_types, lifting) Diff_iff some_eq_imp) have "(?recarg, arg) \ dfs_dfss_term" proof - from pre asm sw have "sub_env e ?eu" by (blast dest: unite_sub_env) hence "sub_env e ?e''" by (auto simp: sub_env_def) moreover from sw pre sclosed have "\u \ vertices. u \ vsuccs e v \ u \ vsuccs ?e'' v" by auto ultimately show ?thesis using pre b unfolding dfs_dfss_term_def by blast qed moreover from pre sw \?sw \ visited e\ \?sw \ explored e\ have "?P ?recarg" by (auto dest: pre_dfss_unite_pre_dfss) ultimately show "?Q ?recarg" using ih b by auto qed then show ?thesis by (simp add: b) qed qed qed } note dom=this from dom show "\ v \ vertices ; pre_dfs v e\ \ dfs_dfss_dom(Inl(v, e))" by auto from dom show "\ v \ vertices ; pre_dfss v e\ \ dfs_dfss_dom(Inr(v, e))" by auto qed text \ Putting everything together, we prove the total correctness of the algorithm when applied to some (root) vertex. \ theorem correctness: assumes "v \ vertices" shows "sccs (dfs v (init_env v)) = {S . is_scc S \ (\n\S. reachable v n)}" using assms init_env_pre_dfs[of v] by (simp add: dfs_dfss_termination partial_correctness) end end diff --git a/thys/Stalnaker_Logic/Stalnaker_Logic.thy b/thys/Stalnaker_Logic/Stalnaker_Logic.thy --- a/thys/Stalnaker_Logic/Stalnaker_Logic.thy +++ b/thys/Stalnaker_Logic/Stalnaker_Logic.thy @@ -1,603 +1,603 @@ (* File: Stalnaker_Logic.thy Author: Laura Gamboa Guzman This work is a formalization of Stalnaker's epistemic logic and its soundness and completeness theorems, as well as the equivalence between the axiomatization of S4 available in the Epistemic Logic theory and the topological one. It builds on the Epistemic_Logic theory by A.H. From. *) theory Stalnaker_Logic imports "Epistemic_Logic.Epistemic_Logic" begin section \Utility\ subsection \Some properties of Normal Modal Logics\ lemma duality_taut: \tautology (((K i p) \<^bold>\ K i (\<^bold>\q))\<^bold>\ ((L i q) \<^bold>\ (\<^bold>\ K i p)))\ by force lemma K_imp_trans: assumes \A \ (p \<^bold>\ q)\ \A \ (q \<^bold>\ r)\ shows \A \ (p \<^bold>\ r)\ proof - have \tautology ((p \<^bold>\ q) \<^bold>\ ( (q \<^bold>\ r) \<^bold>\ (p \<^bold>\ r)))\ by fastforce then show ?thesis by (meson A1 R1 assms(1) assms(2)) qed lemma K_imp_trans': assumes \A \ (q \<^bold>\ r)\ shows \A \ ((p \<^bold>\ q) \<^bold>\ (p \<^bold>\ r))\ proof - have \tautology ((q \<^bold>\ r) \<^bold>\ ((p \<^bold>\ q) \<^bold>\ (p \<^bold>\ r)))\ by fastforce then show ?thesis using A1 R1 assms by blast qed lemma K_imply_multi: assumes \A \ (a \<^bold>\ b)\ and \A \ (a \<^bold>\ c)\ shows \A \ (a \<^bold>\ (b\<^bold>\c))\ proof - have \tautology ((a\<^bold>\b)\<^bold>\(a\<^bold>\c)\<^bold>\(a\<^bold>\(b\<^bold>\c)))\ by force then have \A \ ((a\<^bold>\b)\<^bold>\(a\<^bold>\c)\<^bold>\(a\<^bold>\(b\<^bold>\c)))\ using A1 by blast then have \A \ ((a\<^bold>\c)\<^bold>\(a\<^bold>\(b\<^bold>\c)))\ using assms(1) R1 by blast then show ?thesis using assms(2) R1 by blast qed lemma K_multi_imply: assumes \A \ (a \<^bold>\ b \<^bold>\ c)\ shows \A \ ((a\<^bold>\b) \<^bold>\ c)\ proof - have \tautology ((a \<^bold>\ b \<^bold>\ c) \<^bold>\ ((a\<^bold>\b) \<^bold>\ c))\ by force then have \A \ ((a \<^bold>\ b \<^bold>\ c) \<^bold>\ ((a\<^bold>\b) \<^bold>\ c))\ using A1 by blast then show ?thesis using assms R1 by blast qed lemma K_thm: \A \ ((K i p) \<^bold>\ (L i q) \<^bold>\ L i (p \<^bold>\ q))\ proof - have \tautology (p \<^bold>\ (\<^bold>\(p \<^bold>\ q)) \<^bold>\ \<^bold>\ q)\ by force then have \A \ (p \<^bold>\ (\<^bold>\(p \<^bold>\ q)) \<^bold>\ \<^bold>\ q)\ by (simp add: A1) then have \A \ ((K i p) \<^bold>\ K i ((\<^bold>\(p \<^bold>\ q)) \<^bold>\ \<^bold>\ q))\ and \A \ (K i ((\<^bold>\(p \<^bold>\ q)) \<^bold>\ \<^bold>\ q) \<^bold>\ K i (\<^bold>\(p \<^bold>\ q)) \<^bold>\ K i (\<^bold>\ q))\ apply (simp add: K_map) by (meson K_A2') then have \A \ ((K i p) \<^bold>\ K i (\<^bold>\(p \<^bold>\ q)) \<^bold>\ K i (\<^bold>\ q))\ using K_imp_trans by blast then have \A \ ((K i p) \<^bold>\ L i (q) \<^bold>\ L i (p \<^bold>\ q))\ by (metis AK.simps K_imp_trans duality_taut) then show ?thesis by (simp add: K_multi_imply) qed primrec conjunct :: \'i fm list \ 'i fm\ where \conjunct [] = \<^bold>\\ | \conjunct (p#ps) = (p \<^bold>\ conjunct ps)\ lemma imply_conjunct: \tautology ((imply G p) \<^bold>\ ((conjunct G) \<^bold>\ p))\ apply(induction G) apply simp by force lemma conjunct_imply: \tautology (((conjunct G) \<^bold>\ p) \<^bold>\(imply G p))\ by (induct G) simp_all lemma K_imply_conjunct: assumes \A \ imply G p\ shows \A \ ((conjunct G) \<^bold>\ p)\ using A1 R1 assms imply_conjunct by blast lemma K_conjunct_imply: assumes \A \ ((conjunct G) \<^bold>\ p)\ shows \A \ imply G p\ using A1 R1 assms conjunct_imply by blast lemma K_conj_imply_factor: fixes A :: \('i fm \ bool)\ shows \A \ ((((K i p) \<^bold>\ (K i q)) \<^bold>\ r) \<^bold>\((K i (p \<^bold>\ q)) \<^bold>\ r))\ proof - have *: \A \ ((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ (K i q)))\ proof (rule ccontr) assume \\ A \ ((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ (K i q)))\ then have \consistent A {\<^bold>\((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ (K i q)))}\ by (metis imply.simps(1) inconsistent_imply insert_is_Un list.set(1)) let ?V = \Extend A {\<^bold>\((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ (K i q)))}\ let ?M = \\\ = mcss A, \ = reach A, \ = pi\\ have \?V \ \ ?M \ ?M, ?V \ \<^bold>\((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ (K i q)))\ using canonical_model \consistent A {\<^bold>\ (K i (p \<^bold>\ q) \<^bold>\ K i p \<^bold>\ K i q)}\ insert_iff mem_Collect_eq by fastforce then have o: \?M, ?V \ ((K i (p \<^bold>\ q)) \<^bold>\ \<^bold>\((K i p) \<^bold>\ (K i q)))\ by auto then have \?M, ?V \ (K i (p \<^bold>\ q))\ \(?M, ?V \ \<^bold>\(K i p)) \ (?M, ?V \ \<^bold>\(K i q))\ by auto then have \\ U \ \ ?M \ \ ?M i ?V. ?M, U \(p \<^bold>\ q)\ \\ U \ \ ?M \ \ ?M i ?V. ?M, U \ ((\<^bold>\p) \<^bold>\ (\<^bold>\q))\ using o by auto then show False by simp qed then have \A \ (((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ (K i q))) \<^bold>\ ((((K i p) \<^bold>\ (K i q)) \<^bold>\ r) \<^bold>\ ((K i (p \<^bold>\ q)) \<^bold>\ r)))\ by (simp add: A1) then show ?thesis using "*" R1 by blast qed lemma K_conjunction_in: \A \ (K i (p \<^bold>\ q) \<^bold>\ ((K i p) \<^bold>\ K i q))\ proof - have o1: \A \ ((p\<^bold>\q) \<^bold>\ p)\ and o2: \A \ ((p\<^bold>\q) \<^bold>\ q)\ apply(simp add: A1) by (simp add: A1) then have c1: \A \ (K i (p \<^bold>\ q) \<^bold>\ K i q)\ and c2: \A \ (K i (p \<^bold>\ q) \<^bold>\ K i p)\ apply (simp add: K_map o2) by (simp add: K_map o1) then show ?thesis by (simp add: K_imply_multi) qed lemma K_conjunction_in_mult: \A \ ((K i (conjunct G)) \<^bold>\ conjunct (map (K i) G))\ proof (induct G) case Nil then show ?case by (simp add: A1) case (Cons a G) then have \A \ ((K i (conjunct (a#G))) \<^bold>\ (K i (a \<^bold>\ conjunct G)))\ and \A \ ((K i (a \<^bold>\ conjunct G)) \<^bold>\ ((K i a) \<^bold>\ K i (conjunct G)))\ apply (simp add: A1) by (metis K_conjunction_in) then have o1: \A \ ((K i (conjunct (a#G))) \<^bold>\ ((K i a) \<^bold>\ K i (conjunct G)))\ using K_imp_trans by blast then have \A \ (((K i a) \<^bold>\ K i (conjunct G)) \<^bold>\ K i a)\ and o2: \A \ (((K i a) \<^bold>\ K i (conjunct G)) \<^bold>\ conjunct (map (K i) G))\ apply (simp add: A1) by (metis Cons.hyps K_imply_Cons K_multi_imply imply.simps(1) imply.simps(2)) then have \A \ (((K i a) \<^bold>\ K i (conjunct G)) \<^bold>\ (K i a) \<^bold>\ conjunct (map (K i) G))\ using K_imply_multi by blast then show ?case using K_imp_trans o1 by auto qed lemma K_conjunction_out: \A \ ((K i p) \<^bold>\ (K i q) \<^bold>\ K i (p \<^bold>\ q))\ proof - have \A \ (p \<^bold>\ q \<^bold>\ p\<^bold>\q)\ by (simp add: A1) then have \A \ K i (p \<^bold>\ q \<^bold>\ p\<^bold>\q)\ by (simp add: R2) then have \A \ ((K i p) \<^bold>\ K i (q \<^bold>\ p\<^bold>\q))\ by (simp add: K_map \A \ (p \<^bold>\ q \<^bold>\ p \<^bold>\ q)\) then have \A \ ((K i p) \<^bold>\ (K i q) \<^bold>\ K i (p\<^bold>\q))\ by (meson K_A2' K_imp_trans) then show ?thesis using K_multi_imply by blast qed lemma K_conjunction_out_mult: \A \ (conjunct (map (K i) G) \<^bold>\ (K i (conjunct G)))\ proof (induct G) case Nil then show ?case by (metis A1 K_imply_conjunct Nil_is_map_conv R2 conjunct.simps(1) eval.simps(5) imply.simps(1)) case (Cons a G) then have \A \ ((conjunct (map (K i) (a#G))) \<^bold>\ ((K i a) \<^bold>\ conjunct (map (K i) G)))\ by (simp add: A1) then have *: \A \ (((K i a) \<^bold>\ conjunct (map (K i) G))\<^bold>\ (K i a) \<^bold>\ K i (conjunct G))\ by (metis Cons.hyps K_imply_Cons K_imply_head K_imply_multi K_multi_imply imply.simps(1) imply.simps(2)) then have \A \ (((K i a) \<^bold>\ K i (conjunct G))\<^bold>\ K i (conjunct (a#G)))\ by (simp add: K_conjunction_out) then show ?case using "*" K_imp_trans by auto qed subsection \More on mcs's properties\ lemma mcs_conjunction: assumes \consistent A V\ and \maximal A V\ shows \p \ V \ q \ V \ (p \<^bold>\ q) \ V\ proof - have \tautology (p \<^bold>\ q \<^bold>\ (p\<^bold>\q))\ by force then have \(p \<^bold>\ q \<^bold>\ (p\<^bold>\q)) \ V\ using A1 assms(1) assms(2) deriv_in_maximal by blast then have \p \ V \ (q \<^bold>\ (p\<^bold>\q)) \ V\ by (meson assms(1) assms(2) consequent_in_maximal) then show ?thesis using assms(1) assms(2) consequent_in_maximal by blast qed lemma mcs_conjunction_mult: assumes \consistent A V\ and \maximal A V\ shows \(set (S :: ('i fm list)) \ V \ finite (set S)) \ (conjunct S) \ V\ proof(induct S) case Nil then show ?case by (metis K_Boole assms(1) assms(2) conjunct.simps(1) consistent_def inconsistent_subset maximal_def) case (Cons a S) then have \set S \ set (a#S)\ by (meson set_subset_Cons) then have c1: \ set (a # S) \ V \ finite (set (a # S)) \ conjunct (S) \ V \ a \ V\ using Cons by fastforce then have \ conjunct (S) \ V \ a \ V \ (conjunct (a#S)) \ V \ using assms(1) assms(2) mcs_conjunction by auto then show ?case using c1 by fastforce qed lemma reach_dualK: assumes \consistent A V\ \maximal A V\ and \consistent A W\ \maximal A W\ \W \ reach A i V\ shows \\p. p \ W \ (L i p) \ V\ proof (rule ccontr) assume \\ (\p. p \ W \ (L i p) \ V)\ then obtain p' where *: \p' \ W \ (L i p') \ V\ by presburger then have \(\<^bold>\ L i p') \ V\ using assms(1) assms(2) assms(3) assms(4) assms(5) exactly_one_in_maximal by blast then have \K i (\<^bold>\ p') \ V\ using assms(1) assms(2) exactly_one_in_maximal by blast then have \(\<^bold>\ p') \ W\ using assms(5) by blast then show False by (meson "*" assms(3) assms(4) exactly_one_in_maximal) qed lemma dual_reach: assumes \consistent A V\ \maximal A V\ shows \(L i p) \ V \ (\ W. W \ reach A i V \ p \ W)\ proof - have \(\W. W \ {W. known V i \ W} \ p \ W) \ (\W. W \ {W. known V i \ W} \ (\<^bold>\p) \ W)\ by blast then have \(\W. W \ {W. known V i \ W} \ (\<^bold>\p) \ W) \ (\ W. W \ reach A i V \ (\<^bold>\p) \ W)\ by fastforce then have \(\ W. W \ reach A i V \ (\<^bold>\p) \ W) \ ((K i (\<^bold>\p)) \ V)\ by blast then have \((K i (\<^bold>\p)) \ V) \ (\((L i p) \ V))\ using assms(1) assms(2) exactly_one_in_maximal by blast then have \(\W. W \ {W. known V i \ W} \ p \ W) \ \((L i p) \ V)\ by blast then show ?thesis by blast qed section \Ax.2\ definition weakly_directed :: \('i, 's) kripke \ bool\ where \weakly_directed M \ \i. \s \ \ M. \t \ \ M. \r \ \ M. (r \ \ M i s \ t \ \ M i s)\(\ u \ \ M. (u \ \ M i r \ u \ \ M i t))\ inductive Ax_2 :: \'i fm \ bool\ where \Ax_2 (\<^bold>\ K i (\<^bold>\ K i p) \<^bold>\ K i (\<^bold>\ K i (\<^bold>\ p)))\ subsection \Soundness\ theorem weakly_directed: assumes \weakly_directed M\ \w \ \ M\ shows \M, w \ (L i (K i p) \<^bold>\ K i (L i p))\ proof assume \M, w \ (L i (K i p))\ then have \\v \ \ M \ \ M i w. M, v \ K i p\ by simp then have \\v \ \ M \ \ M i w. \u \ \ M \ \ M i v. M, u \ p\ using \weakly_directed M\ \w \ \ M\ unfolding weakly_directed_def by (metis IntE IntI semantics.simps(6)) then have \\v \ \ M \ \ M i w. M, v \ L i p\ by simp then show \M, w \ K i (L i p)\ by simp qed lemma soundness_Ax_2: \Ax_2 p \ weakly_directed M \ w \ \ M \ M, w \ p\ by (induct p rule: Ax_2.induct) (meson weakly_directed) subsection \Imply completeness\ lemma Ax_2_weakly_directed: fixes A :: \'i fm \ bool\ assumes \\p. Ax_2 p \ A p\ \consistent A V\ \maximal A V\ and \consistent A W\ \maximal A W\ \consistent A U\ \maximal A U\ and \W \ reach A i V\ \U \ reach A i V\ shows \\X. (consistent A X) \ (maximal A X) \ X \ (reach A i W) \ (reach A i U)\ proof (rule ccontr) assume \\ ?thesis\ let ?S = \(known W i) \ (known U i)\ have \\ consistent A ?S\ by (smt (verit, best) Int_Collect \\X. consistent A X \ maximal A X \ X \ {Wa. known W i \ Wa} \ {W. known U i \ W}\ maximal_extension mem_Collect_eq sup.bounded_iff) then obtain S' where *: \A \ imply S' \<^bold>\\ \set S' \ ?S\ \finite (set S')\ unfolding consistent_def by blast let ?U = \filter (\p. p \ (known U i)) S'\ let ?W = \filter (\p. p \ (known W i)) S'\ let ?p = \conjunct ?U\ and ?q = \conjunct ?W\ have \(set ?U) \ (set ?W) = (set S')\ using * by auto then have \A \ imply ?U (imply ?W \<^bold>\)\ using K_imply_weaken imply_append by (metis (mono_tags, lifting) "*"(1) set_append subset_refl) then have \A \ (?p \<^bold>\ (imply ?W \<^bold>\))\ using K_imply_conjunct by blast then have \tautology ((imply ?W \<^bold>\) \<^bold>\ (?q \<^bold>\ \<^bold>\))\ using imply_conjunct by blast then have \A \ ((imply ?W \<^bold>\) \<^bold>\ (?q \<^bold>\ \<^bold>\))\ using A1 by blast then have \A \ (?p \<^bold>\ (?q \<^bold>\ \<^bold>\))\ using K_imp_trans \A \ (conjunct (filter (\p. p \ known U i) S') \<^bold>\ imply (filter (\p. p \ known W i) S') \<^bold>\)\ by blast then have o1: \A \ ((?p \<^bold>\ ?q) \<^bold>\ \<^bold>\)\ by (meson K_multi_imply) moreover have \set ?U \ (known U i)\ and \set ?W \ (known W i)\ and \\ p. p \ set ?U \ (K i p) \ U\ and \\ p. p \ set ?W \ (K i p) \ W\ by auto then have \set (map (K i) ?U) \ U\ and c1: \set (map (K i) ?W) \ W\ apply (metis (mono_tags, lifting) imageE set_map subsetI) by auto then have c2: \conjunct (map (K i) ?U) \ U\ and c2': \conjunct (map (K i) ?W) \ W\ using assms(6) assms(7) mcs_conjunction_mult apply blast using assms(4) assms(5) c1 mcs_conjunction_mult by blast then have \((conjunct (map (K i) ?U)) \<^bold>\ (K i ?p)) \ U\ and c3: \((conjunct (map (K i) ?W)) \<^bold>\ (K i ?q)) \ W\ apply (meson K_conjunction_out_mult assms(6) assms(7) deriv_in_maximal) by (meson K_conjunction_out_mult assms(4) assms(5) deriv_in_maximal) then have c4: \(K i ?p) \ U\ and c4': \(K i ?q) \ W\ using assms(6) assms(7) c2 consequent_in_maximal apply blast using assms(4) assms(5) c2' c3 consequent_in_maximal by blast then have \(L i (K i ?p)) \ V\ and c5: \(L i (K i ?q)) \ V\ using assms(2) assms(3) assms(6) assms(7) assms(9) exactly_one_in_maximal apply blast using assms(2) assms(3) assms(4) assms(5) assms(8) c4' exactly_one_in_maximal by blast then have \(K i (L i ?p)) \ V\ by (meson Ax_2.intros assms(1) assms(2) assms(3) ax_in_maximal consequent_in_maximal) then have \((K i (L i ?p)) \<^bold>\ (L i (K i ?q))) \ V\ using assms(2) assms(3) c5 mcs_conjunction by blast then have \(L i ((L i ?p) \<^bold>\ K i ?q)) \ V\ by (meson K_thm assms(2) assms(3) consequent_in_maximal deriv_in_maximal) then have \(L i ((K i ?q) \<^bold>\ L i ?p)) \ V\ - by (smt (z3) \K i (L i (conjunct (filter (\p. p \ known U i) S'))) \ V\ assms(2) assms(3) assms(4) assms(5) assms(8) c4' exactly_one_in_maximal mcs_conjunction mem_Collect_eq subset_iff) + by (smt (verit) \K i (L i (conjunct (filter (\p. p \ known U i) S'))) \ V\ assms(2) assms(3) assms(4) assms(5) assms(8) c4' exactly_one_in_maximal mcs_conjunction mem_Collect_eq subset_iff) then obtain Z' where z1:\(consistent A Z') \ (maximal A Z')\ and z2:\Z' \ (reach A i V)\ and z3: \((K i ?q) \<^bold>\ L i ?p) \ Z'\ using \K i (L i (conjunct (filter (\p. p \ known U i) S'))) \ V\ assms(4) assms(5) assms(8) c4' mcs_conjunction by blast then have z4: \(L i (?q \<^bold>\ ?p)) \ Z'\ by (metis K_thm consequent_in_maximal deriv_in_maximal) then have o2:\(L i (L i (?q \<^bold>\ ?p))) \ V\ using assms(2) assms(3) mcs_properties(2) z1 z2 by blast then have \A \ K i (K i (((?p \<^bold>\ ?q) \<^bold>\ \<^bold>\)))\ by (metis R2 o1) then have o3:\K i (K i (((?p \<^bold>\ ?q) \<^bold>\ \<^bold>\))) \ V\ using assms(2) assms(3) deriv_in_maximal by blast then obtain X1 where x1:\(consistent A X1) \ (maximal A X1)\ and x2:\X1 \ (reach A i V)\ and x3: \(L i (?q \<^bold>\ ?p)) \ X1\ using z1 z2 z4 by blast then have x4:\(K i (((?p \<^bold>\ ?q) \<^bold>\ \<^bold>\))) \ X1\ using o3 by blast then have t:\\x. \y. tautology (((x\<^bold>\y)\<^bold>\\<^bold>\)\<^bold>\\<^bold>\(y\<^bold>\x))\ by (metis eval.simps(4) eval.simps(5)) then have \(((?p\<^bold>\?q)\<^bold>\\<^bold>\)\<^bold>\\<^bold>\(?q\<^bold>\?p)) \ X1\ using A1 deriv_in_maximal x1 by blast then have \K i (((?p\<^bold>\?q)\<^bold>\\<^bold>\)\<^bold>\\<^bold>\(?q\<^bold>\?p)) \ X1\ by (meson A1 R2 deriv_in_maximal t x1) then have \(K i ((?p\<^bold>\?q)\<^bold>\\<^bold>\)\<^bold>\ K i (\<^bold>\(?q\<^bold>\?p))) \ X1\ by (meson K_A2' consequent_in_maximal deriv_in_maximal x1) then have \(K i ((?p\<^bold>\?q)\<^bold>\\<^bold>\)\<^bold>\ (\<^bold>\ L i(?q\<^bold>\?p))) \ X1\ using consequent_in_maximal exactly_one_in_maximal x1 x3 x4 by blast then have \(\<^bold>\ L i(?q\<^bold>\?p)) \ X1 \ (L i(?q\<^bold>\?p)) \ X1\ using consequent_in_maximal x1 x4 x3 by blast then show False using exactly_one_in_maximal x1 by blast qed lemma mcs\<^sub>_\<^sub>2_weakly_directed: fixes A :: \'i fm \ bool\ assumes \\p. Ax_2 p \ A p\ shows \weakly_directed \\ = mcss A, \ = reach A, \ = pi\\ unfolding weakly_directed_def proof (intro allI ballI, auto) fix i V U W assume \consistent A V\ \maximal A V\ \consistent A U\ \maximal A U\ \consistent A W\ \maximal A W\ \known V i \ U\ \known V i \ W\ then have \\X. (consistent A X) \ (maximal A X) \ X \ (reach A i W) \ (reach A i U)\ using Ax_2_weakly_directed [where A=A and V=V and W=W and U=U] assms IntD2 by simp then have \\X. (consistent A X) \ (maximal A X) \ X \ (reach A i W) \ X \ (reach A i U)\ by simp then show \\X. (consistent A X) \ (maximal A X) \ known W i \ X \ known U i \ X\ by auto qed lemma imply_completeness_K_2: assumes valid: \\(M :: ('i, 'i fm set) kripke). \w \ \ M. weakly_directed M \ (\q \ G. M, w \ q) \ M, w \ p\ shows \\qs. set qs \ G \ (Ax_2 \ imply qs p)\ proof (rule ccontr) assume \\qs. set qs \ G \ Ax_2 \ imply qs p\ then have *: \\qs. set qs \ G \ \ Ax_2 \ imply ((\<^bold>\ p) # qs) \<^bold>\\ using K_Boole by blast let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend Ax_2 ?S\ let ?M = \\\ = mcss Ax_2, \ = reach Ax_2, \ = pi\\ have \consistent Ax_2 ?S\ using * by (metis K_imply_Cons consistent_def inconsistent_subset) then have \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent Ax_2 ?V\ \maximal Ax_2 ?V\ using canonical_model unfolding list_all_def by fastforce+ moreover have \weakly_directed ?M\ using mcs\<^sub>_\<^sub>2_weakly_directed [where A=Ax_2] by fast ultimately have \?M, ?V \ p\ using valid by auto then show False using \?M, ?V \ (\<^bold>\ p)\ by simp qed section \System S4.2\ abbreviation SystemS4_2 :: \'i fm \ bool\ ("\\<^sub>S\<^sub>4\<^sub>2 _" [50] 50) where \\\<^sub>S\<^sub>4\<^sub>2 p \ AxT \ Ax4 \ Ax_2 \ p\ abbreviation AxS4_2 :: \'i fm \ bool\ where \AxS4_2 \ AxT \ Ax4 \ Ax_2\ subsection \Soundness\ abbreviation w_directed_preorder :: \('i, 'w) kripke \ bool\ where \w_directed_preorder M \ reflexive M \ transitive M \ weakly_directed M\ lemma soundness_AxS4_2: \AxS4_2 p \ w_directed_preorder M \ w \ \ M \ M, w \ p\ using soundness_AxT soundness_Ax4 soundness_Ax_2 by metis lemma soundness\<^sub>S\<^sub>4\<^sub>2: \\\<^sub>S\<^sub>4\<^sub>2 p \ w_directed_preorder M \ w \ \ M \ M, w \ p\ using soundness soundness_AxS4_2 . subsection \Completeness\ lemma imply_completeness_S4_2: assumes valid: \\(M :: ('i, 'i fm set) kripke). \w \ \ M. w_directed_preorder M \ (\q \ G. M, w \ q) \ M, w \ p\ shows \\qs. set qs \ G \ (AxS4_2 \ imply qs p)\ proof (rule ccontr) assume \\qs. set qs \ G \ AxS4_2 \ imply qs p\ then have *: \\qs. set qs \ G \ \ AxS4_2 \ imply ((\<^bold>\ p) # qs) \<^bold>\\ using K_Boole by blast let ?S = \{\<^bold>\ p} \ G\ let ?V = \Extend AxS4_2 ?S\ let ?M = \\\ = mcss AxS4_2, \ = reach AxS4_2, \ = pi\\ have \consistent AxS4_2 ?S\ using * by (metis (no_types, lifting) K_imply_Cons consistent_def inconsistent_subset) then have \?M, ?V \ (\<^bold>\ p)\ \\q \ G. ?M, ?V \ q\ \consistent AxS4_2 ?V\ \maximal AxS4_2 ?V\ using canonical_model unfolding list_all_def by fastforce+ moreover have \w_directed_preorder ?M\ using reflexive\<^sub>T[of AxS4_2] mcs\<^sub>_\<^sub>2_weakly_directed[of AxS4_2] transitive\<^sub>K\<^sub>4[of AxS4_2] by auto ultimately have \?M, ?V \ p\ using valid by auto then show False using \?M, ?V \ (\<^bold>\ p)\ by simp qed lemma completeness\<^sub>S\<^sub>4\<^sub>2: assumes \\(M :: ('i, 'i fm set) kripke). \w \ \ M. w_directed_preorder M \ M, w \ p\ shows \\\<^sub>S\<^sub>4\<^sub>2 p\ using assms imply_completeness_S4_2[where G=\{}\] by auto abbreviation \valid\<^sub>S\<^sub>4\<^sub>2 p \ \(M :: (nat, nat fm set) kripke). \w \ \ M. w_directed_preorder M \ M, w \ p\ theorem main\<^sub>S\<^sub>4\<^sub>2: \valid\<^sub>S\<^sub>4\<^sub>2 p \ \\<^sub>S\<^sub>4\<^sub>2 p\ using soundness\<^sub>S\<^sub>4\<^sub>2 completeness\<^sub>S\<^sub>4\<^sub>2 by fast corollary assumes \w_directed_preorder M\ \w \ \ M\ shows \valid\<^sub>S\<^sub>4\<^sub>2 p \ M, w \ p\ using assms soundness\<^sub>S\<^sub>4\<^sub>2 completeness\<^sub>S\<^sub>4\<^sub>2 by fast section \Topological S4 axioms\ abbreviation DoubleImp (infixr "\<^bold>\" 25) where \(p\<^bold>\q) \ ((p \<^bold>\ q) \<^bold>\ (q \<^bold>\ p))\ inductive System_topoS4 :: \'i fm \ bool\ ("\\<^sub>T\<^sub>o\<^sub>p _" [50] 50) where A1': \tautology p \ \\<^sub>T\<^sub>o\<^sub>p p\ | AR: \\\<^sub>T\<^sub>o\<^sub>p ((K i (p \<^bold>\ q)) \<^bold>\ ((K i p) \<^bold>\ K i q))\ | AT': \\\<^sub>T\<^sub>o\<^sub>p (K i p \<^bold>\ p)\ | A4': \\\<^sub>T\<^sub>o\<^sub>p (K i p \<^bold>\ K i (K i p))\ | AN: \\\<^sub>T\<^sub>o\<^sub>p K i \<^bold>\\ | R1': \\\<^sub>T\<^sub>o\<^sub>p p \ \\<^sub>T\<^sub>o\<^sub>p (p \<^bold>\ q) \ \\<^sub>T\<^sub>o\<^sub>p q\ | RM: \\\<^sub>T\<^sub>o\<^sub>p (p \<^bold>\ q) \ \\<^sub>T\<^sub>o\<^sub>p ((K i p) \<^bold>\ K i q)\ lemma topoS4_trans: \\\<^sub>T\<^sub>o\<^sub>p ((p \<^bold>\ q) \<^bold>\ (q \<^bold>\ r) \<^bold>\ p \<^bold>\ r)\ by (simp add: A1') lemma topoS4_conjElim: \\\<^sub>T\<^sub>o\<^sub>p (p \<^bold>\ q \<^bold>\ q)\ by (simp add: A1') lemma topoS4_AxK: \\\<^sub>T\<^sub>o\<^sub>p (K i p \<^bold>\ K i (p \<^bold>\ q) \<^bold>\ K i q)\ proof - have \\\<^sub>T\<^sub>o\<^sub>p ((p \<^bold>\ (p \<^bold>\ q)) \<^bold>\ q)\ using A1' by force then have *: \\\<^sub>T\<^sub>o\<^sub>p (K i (p \<^bold>\ (p \<^bold>\ q)) \<^bold>\ K i q)\ using RM by fastforce then have \\\<^sub>T\<^sub>o\<^sub>p (K i p \<^bold>\ K i (p \<^bold>\ q) \<^bold>\ K i (p \<^bold>\ (p \<^bold>\ q)))\ using AR topoS4_conjElim System_topoS4.simps by fast then show ?thesis by (meson "*" System_topoS4.R1' topoS4_trans) qed lemma topoS4_NecR: assumes \\\<^sub>T\<^sub>o\<^sub>p p\ shows\\\<^sub>T\<^sub>o\<^sub>p K i p\ proof - have \\\<^sub>T\<^sub>o\<^sub>p (\<^bold>\ \<^bold>\ p)\ using assms by (metis System_topoS4.A1' System_topoS4.R1' conjunct.simps(1) imply.simps(1) imply_conjunct) then have \\\<^sub>T\<^sub>o\<^sub>p (K i \<^bold>\ \<^bold>\ K i p)\ using RM by force then show ?thesis by (meson AN System_topoS4.R1') qed lemma empty_S4: "{} \\<^sub>S\<^sub>4 p \ AxT \ Ax4 \ p" by simp lemma S4_topoS4: \{} \\<^sub>S\<^sub>4 p \ \\<^sub>T\<^sub>o\<^sub>p p\ unfolding empty_S4 proof (induct p rule: AK.induct) case (A2 i p q) then show ?case using topoS4_AxK . next case (Ax p) then show ?case using AT' A4' by (metis AxT.cases Ax4.cases) next case (R2 p) then show ?case by (simp add: topoS4_NecR) qed (meson System_topoS4.intros)+ lemma topoS4_S4: fixes p :: \'i fm\ shows \\\<^sub>T\<^sub>o\<^sub>p p \ {} \\<^sub>S\<^sub>4 p\ unfolding empty_S4 proof (induct p rule: System_topoS4.induct) case (AT' i p) then show ?case by (simp add: Ax AxT.intros) next case (A4' i p) then show ?case by (simp add: Ax Ax4.intros) next case (AR i p q) then show ?case by (meson K_conj_imply_factor K_conjunction_in K_conjunction_out K_imp_trans' K_imply_multi R1) next case (AN i) then have *: \AxT \ Ax4 \ \<^bold>\\ by (simp add: A1) then show ?case by (simp add: * R2) next case (RM p q i) then have \AxT \ Ax4 \ K i (p \<^bold>\ q)\ by (simp add: R2) then show ?case by (simp add: K_map RM.hyps(2)) qed (meson AK.intros)+ theorem main\<^sub>S\<^sub>4': \{} \\<^sub>S\<^sub>4 p \ (\\<^sub>T\<^sub>o\<^sub>p p)\ using main\<^sub>S\<^sub>4[of "{}"] S4_topoS4 topoS4_S4 by fast end diff --git a/thys/Sturm_Tarski/Pseudo_Remainder_Sequence.thy b/thys/Sturm_Tarski/Pseudo_Remainder_Sequence.thy --- a/thys/Sturm_Tarski/Pseudo_Remainder_Sequence.thy +++ b/thys/Sturm_Tarski/Pseudo_Remainder_Sequence.thy @@ -1,566 +1,566 @@ (* File: Pseudo_Remainder_Sequence.thy Author: Wenda Li *) section \An implementation for calculating pseudo remainder sequences\ theory Pseudo_Remainder_Sequence imports Sturm_Tarski "HOL-Computational_Algebra.Computational_Algebra" (*TODO: we should consider to move this to the standard library*) "Polynomial_Interpolation.Ring_Hom_Poly" begin subsection \Misc\ function spmods :: "'a::idom poly \ 'a poly \ ('a poly) list" where "spmods p q= (if p=0 then [] else let m=(if even(degree p+1-degree q) then -1 else -lead_coeff q) in Cons p (spmods q (smult m (pseudo_mod p q))))" by auto termination apply (relation "measure (\(p,q).if p=0 then 0 else if q=0 then 1 else 2+degree q)") by (simp_all add: degree_pseudo_mod_less) declare spmods.simps[simp del] lemma spmods_0[simp]: "spmods 0 q = []" "spmods p 0 = (if p=0 then [] else [p])" by (auto simp:spmods.simps) lemma spmods_nil_eq:"spmods p q = [] \ (p=0)" by (metis list.distinct(1) spmods.elims) lemma changes_poly_at_alternative: "changes_poly_at ps a= changes (map (\p. sign(poly p a)) ps)" "changes_poly_at ps a= changes (map (\p. sgn(poly p a)) ps)" unfolding changes_poly_at_def subgoal by (subst changes_map_sign_eq) (auto simp add:comp_def) subgoal by (subst changes_map_sgn_eq) (auto simp add:comp_def) done lemma smods_smult_length: assumes "a\0" "b\0" shows "length (smods p q) = length (smods (smult a p) (smult b q))" using assms proof (induct "smods p q" arbitrary:p q a b ) case Nil thus ?case by (simp split:if_splits) next case (Cons x xs) hence "p\0" by auto define r where "r\- (p mod q)" have "smods q r = xs" using Cons.hyps(2) `p\0` unfolding r_def by auto hence "length (smods q r) = length (smods (smult b q) (smult a r))" using Cons.hyps(1)[of q r b a] Cons by auto moreover have "smult a p\0" using `a\0` `p\0` by auto moreover have "-((smult a p) mod (smult b q)) = (smult a r)" by (simp add: Cons.prems(2) mod_smult_left mod_smult_right r_def) ultimately show ?case unfolding r_def by auto qed lemma smods_smult_nth[rule_format]: fixes p q::"real poly" assumes "a\0" "b\0" defines "xs\smods p q" and "ys\smods (smult a p) (smult b q)" shows "\n0" by auto define r where "r\- (p mod q)" have xs:"xs=smods q r" "p#xs=smods p q" using Cons.hyps(2) `p\0` unfolding r_def by auto define ys where "ys\smods (smult b q) (smult a r)" have "- ((smult a p) mod (smult b q)) = smult a r" by (simp add: Cons.hyps(4) mod_smult_left mod_smult_right r_def) hence ys:"smult a p # ys = smods (smult a p) (smult b q)" using `p\0` `a\0` unfolding ys_def r_def by auto have hyps:"\n. n ys ! n = (if even n then smult b (xs ! n) else smult a (xs ! n))" using Cons.hyps(1)[of q r b a,folded xs ys_def] `a\0` `b\0` by auto thus ?case apply (fold xs ys) apply auto by (case_tac n,auto)+ qed lemma smods_smult_sgn_map_eq: fixes x::real assumes "m>0" defines "f\\p. sgn(poly p x)" shows "map f (smods p (smult m q)) = map f (smods p q)" "map sgn_pos_inf (smods p (smult m q)) = map sgn_pos_inf (smods p q)" "map sgn_neg_inf (smods p (smult m q)) = map sgn_neg_inf (smods p q)" proof - define xs ys where "xs\smods p q" and "ys\smods p (smult m q)" have "m\0" using `m>0` by simp have len_eq:"length xs =length ys" using smods_smult_length[of 1 m] `m>0` unfolding xs_def ys_def by auto moreover have "(map f xs) ! i = (map f ys) ! i" "(map sgn_pos_inf xs) ! i = (map sgn_pos_inf ys) ! i" "(map sgn_neg_inf xs) ! i = (map sgn_neg_inf ys) ! i" when "i0`,of _ p q,unfolded smult_1_left, folded xs_def ys_def,OF `i0` by (auto simp add:sgn_mult sgn_pos_inf_def sgn_neg_inf_def lead_coeff_smult) qed ultimately show "map f (smods p (smult m q)) = map f (smods p q)" "map sgn_pos_inf (smods p (smult m q)) = map sgn_pos_inf (smods p q)" "map sgn_neg_inf (smods p (smult m q)) = map sgn_neg_inf (smods p q)" apply (fold xs_def ys_def) by (auto intro: nth_equalityI) qed lemma changes_poly_at_smods_smult: assumes "m>0" shows "changes_poly_at (smods p (smult m q)) x =changes_poly_at (smods p q) x " using smods_smult_sgn_map_eq[OF `m>0`] by (metis changes_poly_at_alternative(2)) lemma spmods_smods_sgn_map_eq: fixes p q::"real poly" and x::real defines "f\\p. sgn (poly p x)" shows "map f (smods p q) = map f (spmods p q)" "map sgn_pos_inf (smods p q) = map sgn_pos_inf (spmods p q)" "map sgn_neg_inf (smods p q) = map sgn_neg_inf (spmods p q)" proof (induct "spmods p q" arbitrary:p q) case Nil hence "p=0" using spmods_nil_eq by metis thus "map f (smods p q) = map f (spmods p q)" "map sgn_pos_inf (smods p q) = map sgn_pos_inf (spmods p q)" "map sgn_neg_inf (smods p q) = map sgn_neg_inf (spmods p q)" by auto next case (Cons p' xs) hence "p\0" by auto define r where "r\- (p mod q)" define exp where " exp\degree p +1 - degree q" define m where "m\(if even exp then 1 else lead_coeff q) * (lead_coeff q ^ exp)" have xs1:"p#xs=spmods p q" by (metis (no_types) Cons.hyps(4) list.distinct(1) list.inject spmods.simps) have xs2:"xs=spmods q (smult m r)" when "q\0" proof - define m' where "m'\if even exp then - 1 else - lead_coeff q" have "smult m' (pseudo_mod p q) = smult m r" unfolding m_def m'_def r_def apply (subst pseudo_mod_mod[symmetric]) using that exp_def by auto thus ?thesis using `p\0` xs1 unfolding r_def by (simp add:spmods.simps[of p q,folded exp_def, folded m'_def] del:spmods.simps) qed define ys where "ys\smods q r" have ys:"p#ys=smods p q" using `p\0` unfolding ys_def r_def by auto have qm:"q\0 \ m>0" using `p\0` unfolding m_def apply auto subgoal by (simp add: zero_less_power_eq) subgoal using zero_less_power_eq by fastforce done show "map f (smods p q) = map f (spmods p q)" proof (cases "q\0") case True then have "map f (spmods q (smult m r)) = map f (smods q r)" using smods_smult_sgn_map_eq(1)[of m x q r,folded f_def] qm Cons.hyps(1)[OF xs2,folded f_def] by simp thus ?thesis apply (fold xs1 xs2[OF True] ys ys_def) by auto next case False thus ?thesis by auto qed show "map sgn_pos_inf (smods p q) = map sgn_pos_inf (spmods p q)" proof (cases "q\0") case True then have "map sgn_pos_inf (spmods q (smult m r)) = map sgn_pos_inf (smods q r)" using Cons.hyps(2)[OF xs2,folded f_def] qm[OF True] smods_smult_sgn_map_eq(2)[of m q r,folded f_def] by auto thus ?thesis apply (fold xs1 xs2[OF True] ys ys_def) by (simp add:f_def) next case False thus ?thesis by auto qed show "map sgn_neg_inf (smods p q) = map sgn_neg_inf (spmods p q)" proof (cases "q\0") case True then have "map sgn_neg_inf (spmods q (smult m r)) = map sgn_neg_inf (smods q r)" using Cons.hyps(3)[OF xs2,folded f_def] qm[OF True] smods_smult_sgn_map_eq(3)[of m q r,folded f_def] by auto thus ?thesis apply (fold xs1 xs2[OF True] ys ys_def) by (simp add:f_def) next case False thus ?thesis by auto qed qed subsection \Converting @{typ "rat poly"} to @{typ "int poly"} by clearing the denominators\ definition int_of_rat::"rat \ int" where "int_of_rat = inv of_int" lemma of_rat_inj[simp]: "inj of_rat" by (simp add: linorder_injI) lemma (in ring_char_0) of_int_inj[simp]: "inj of_int" by (simp add: inj_on_def) lemma int_of_rat_id: "int_of_rat o of_int = id" unfolding int_of_rat_def by auto lemma int_of_rat_0[simp]:"int_of_rat 0 = 0" by (metis id_apply int_of_rat_id o_def of_int_0) lemma int_of_rat_inv:"r\\ \ of_int (int_of_rat r) = r" unfolding int_of_rat_def by (simp add: Ints_def f_inv_into_f) lemma int_of_rat_0_iff:"x\\ \ int_of_rat x = 0 \ x = 0" using int_of_rat_inv by force lemma [code]:"int_of_rat r = (let (a,b) = quotient_of r in if b=1 then a else Code.abort (STR ''Failed to convert rat to int'') (\_. int_of_rat r))" apply (auto simp add:split_beta int_of_rat_def) by (metis Fract_of_int_quotient inv_f_eq of_int_inj of_int_rat quotient_of_div surjective_pairing) definition de_lcm::"rat poly \ int" where "de_lcm p = Lcm(set(map (\x. snd (quotient_of x)) (coeffs p)))" lemma de_lcm_pCons:"de_lcm (pCons a p) = lcm (snd (quotient_of a)) (de_lcm p)" unfolding de_lcm_def by (cases "a=0\p=0",auto) lemma de_lcm_0[simp]:"de_lcm 0 = 1" unfolding de_lcm_def by auto lemma de_lcm_pos[simp]:"de_lcm p > 0" apply (induct p) apply (auto simp add:de_lcm_pCons) by (metis lcm_pos_int less_numeral_extra(3) quotient_of_denom_pos')+ lemma de_lcm_ints: fixes x::rat shows "x\set (coeffs p) \ rat_of_int (de_lcm p) * x \ \" proof (induct p) case 0 then show ?case by auto next case (pCons a p) define a1 a2 where "a1\fst (quotient_of a)" and "a2\snd (quotient_of a)" have a:"a=(rat_of_int a1)/(rat_of_int a2)" and "a2>0" using quotient_of_denom_pos'[of a] unfolding a1_def a2_def by (auto simp add: quotient_of_div) define mp1 where "mp1\a2 div gcd (de_lcm p) a2" define mp2 where "mp2\de_lcm p div gcd a2 (de_lcm p) " have lcm_times1:"lcm a2 (de_lcm p) = de_lcm p * mp1" using lcm_altdef_int[of "de_lcm p" a2,folded mp1_def] `a2>0` unfolding mp1_def apply (subst div_mult_swap) by (auto simp add: abs_of_pos gcd.commute lcm_altdef_int mult.commute) have lcm_times2:"lcm a2 (de_lcm p) = a2 * mp2" using lcm_altdef_int[of a2 "de_lcm p",folded mp1_def] `a2>0` unfolding mp2_def by (subst div_mult_swap, auto simp add:abs_of_pos) show ?case proof (cases "x \ set (coeffs p)") case True show ?thesis using pCons(2)[OF True] - by (smt Ints_mult Ints_of_int a2_def de_lcm_pCons lcm_times1 + by (smt (verit) Ints_mult Ints_of_int a2_def de_lcm_pCons lcm_times1 mult.assoc mult.commute of_int_mult) next case False then have "x=a" using pCons cCons_not_0_eq coeffs_pCons_eq_cCons insert_iff list.set(2) not_0_cCons_eq by fastforce show ?thesis unfolding `x=a` de_lcm_pCons apply (fold a2_def,unfold a) by (simp add: de_lcm_pCons lcm_times2 of_rat_divide) qed qed definition clear_de::"rat poly \ int poly" where "clear_de p = (SOME q. (map_poly of_int q) = smult (of_int (de_lcm p)) p)" lemma clear_de:"of_int_poly(clear_de p) = smult (of_int (de_lcm p)) p" proof - have "\q. (of_int_poly q) = smult (of_int (de_lcm p)) p" proof (induct p) case 0 show ?case by (metis map_poly_0 smult_0_right) next case (pCons a p) then obtain q1::"int poly" where q1:"of_int_poly q1 = smult (rat_of_int (de_lcm p)) p" by auto define a1 a2 where "a1\fst (quotient_of a)" and "a2\snd (quotient_of a)" have a:"a=(rat_of_int a1)/ (rat_of_int a2)" and "a2>0" using quotient_of_denom_pos' quotient_of_div unfolding a1_def a2_def by auto define mp1 where "mp1\a2 div gcd (de_lcm p) a2" define mp2 where "mp2\de_lcm p div gcd a2 (de_lcm p) " have lcm_times1:"lcm a2 (de_lcm p) = de_lcm p * mp1" using lcm_altdef_int[of "de_lcm p" a2,folded mp1_def] `a2>0` unfolding mp1_def by (subst div_mult_swap, auto simp add: abs_of_pos gcd.commute lcm_altdef_int mult.commute) have lcm_times2:"lcm a2 (de_lcm p) = a2 * mp2" using lcm_altdef_int[of a2 "de_lcm p",folded mp1_def] `a2>0` unfolding mp2_def by (subst div_mult_swap, auto simp add:abs_of_pos) define q2 where "q2\pCons (mp2 * a1) (smult mp1 q1)" have "of_int_poly q2 = smult (rat_of_int (de_lcm (pCons a p))) (pCons a p)" using `a2>0` apply (simp add:de_lcm_pCons ) apply (fold a2_def) apply (unfold a) apply (subst lcm_times2,subst lcm_times1) by (simp add: Polynomial.map_poly_pCons mult.commute of_int_hom.map_poly_hom_smult q1 q2_def) then show ?case by auto qed then show ?thesis unfolding clear_de_def by (meson someI_ex) qed lemma clear_de_0[simp]:"clear_de 0 = 0" using clear_de[of 0] by auto lemma [code abstract]: "coeffs (clear_de p) = (let lcm = de_lcm p in map (\x. int_of_rat (of_int lcm * x)) (coeffs p))" proof - define mul where "mul\rat_of_int (de_lcm p)" have "map_poly int_of_rat (of_int_poly q) = q" for q apply (subst map_poly_map_poly) by (auto simp add:int_of_rat_id) then have clear_eq:"clear_de p = map_poly int_of_rat (smult (of_int (de_lcm p)) p)" using arg_cong[where f="map_poly int_of_rat",OF clear_de] by auto show ?thesis proof (cases "p=0") case True then show ?thesis by auto next case False define g where "g\(\x. int_of_rat (rat_of_int (de_lcm p) * x))" have "de_lcm p \ 0" using de_lcm_pos by (metis less_irrefl) moreover have "last (coeffs p) \0" by (simp add: False last_coeffs_eq_coeff_degree) have False when asm:"last (map g (coeffs p)) =0" proof - have "coeffs p \[]" using False by auto hence "g (last (coeffs p)) = 0" using asm last_map[of "coeffs p" g] by auto hence "last (coeffs p) = 0" unfolding g_def using `coeffs p \[]` `de_lcm p \ 0` apply (subst (asm) int_of_rat_0_iff) by (auto intro!: de_lcm_ints ) thus False using `last (coeffs p) \0` by simp qed ultimately show ?thesis apply (auto simp add: coeffs_smult clear_eq comp_def smult_conv_map_poly map_poly_map_poly coeffs_map_poly) apply (fold g_def) by (metis False Ring_Hom_Poly.coeffs_map_poly coeffs_eq_Nil last_coeffs_eq_coeff_degree last_map) qed qed subsection \Sign variations for pseudo-remainder sequences\ locale order_hom = fixes hom :: "'a :: ord \ 'b :: ord" assumes hom_less: "x < y \ hom x < hom y" and hom_less_eq: "x \ y \ hom x \ hom y" locale linordered_idom_hom = order_hom hom + inj_idom_hom hom for hom :: "'a :: linordered_idom \ 'b :: linordered_idom" begin lemma sgn_sign:"sgn (hom x) = of_int (sign x)" by (simp add: sign_def hom_less sgn_if) end locale hom_pseudo_smods= comm_semiring_hom hom + r1:linordered_idom_hom R\<^sub>1 + r2:linordered_idom_hom R\<^sub>2 for hom::"'a::linordered_idom \ 'b::{comm_semiring_1,linordered_idom}" and R\<^sub>1::"'a \ real" and R\<^sub>2::"'b \ real" + assumes R_hom:"R\<^sub>1 x = R\<^sub>2 (hom x)" begin (* hom R2 'a \ 'b \ real rat/float int p x *) lemma map_poly_R_hom_commute: "poly (map_poly R\<^sub>1 p) (R\<^sub>2 x) = R\<^sub>2 (poly (map_poly hom p) x)" apply (induct p) using r2.hom_add r2.hom_mult R_hom by auto definition changes_hpoly_at::"'a poly list \ 'b \ int" where "changes_hpoly_at ps a= changes (map (\p. eval_poly hom p a) ps)" lemma changes_hpoly_at_Nil[simp]: "changes_hpoly_at [] a = 0" unfolding changes_hpoly_at_def by simp definition changes_itv_spmods:: "'b \ 'b \ 'a poly \ 'a poly \ int" where "changes_itv_spmods a b p q= (let ps = spmods p q in changes_hpoly_at ps a - changes_hpoly_at ps b)" definition changes_gt_spmods:: "'b \ 'a poly \ 'a poly \ int" where "changes_gt_spmods a p q= (let ps = spmods p q in changes_hpoly_at ps a - changes_poly_pos_inf ps)" definition changes_le_spmods:: "'b \ 'a poly \ 'a poly \ int" where "changes_le_spmods b p q= (let ps = spmods p q in changes_poly_neg_inf ps - changes_hpoly_at ps b)" definition changes_R_spmods:: "'a poly \ 'a poly \ int" where "changes_R_spmods p q= (let ps= spmods p q in changes_poly_neg_inf ps - changes_poly_pos_inf ps)" lemma changes_spmods_smods: shows "changes_itv_spmods a b p q = changes_itv_smods (R\<^sub>2 a) (R\<^sub>2 b) (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" and "changes_R_spmods p q = changes_R_smods (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" and "changes_gt_spmods a p q = changes_gt_smods (R\<^sub>2 a) (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" and "changes_le_spmods b p q = changes_le_smods (R\<^sub>2 b) (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" proof - define pp qq where "pp = map_poly R\<^sub>1 p" and "qq = map_poly R\<^sub>1 q" have spmods_eq:"spmods (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q) = map (map_poly R\<^sub>1) (spmods p q)" proof (induct "spmods p q" arbitrary:p q ) case Nil thus ?case by (metis list.simps(8) map_poly_0 spmods_nil_eq) next case (Cons p' xs) hence "p\0" by auto define m where "m\(if even (degree p + 1 - degree q) then - 1 else - lead_coeff q)" define r where "r\smult m (pseudo_mod p q)" have xs1:"p#xs=spmods p q" by (metis (no_types) Cons.hyps(2) list.distinct(1) list.inject spmods.simps) have xs2:"xs=spmods q r" using xs1 \p\0\ r_def by (auto simp add:spmods.simps[of p q,folded exp_def,folded m_def]) define ys where "ys\spmods (map_poly R\<^sub>1 q) (map_poly R\<^sub>1 r)" have ys:"(map_poly R\<^sub>1 p)#ys=spmods (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" using \p\0\ unfolding ys_def r_def apply (subst (2) spmods.simps) unfolding m_def by (auto simp:r1.pseudo_mod_hom hom_distribs) show ?case using Cons.hyps(1)[OF xs2] apply (fold xs1 xs2 ys ys_def) by auto qed have changes_eq_at:"changes_poly_at (smods pp qq) (R\<^sub>2 x) = changes_hpoly_at (spmods p q) x" (is "?L=?R") for x proof - define ff where "ff = (\p. sgn (poly p (R\<^sub>2 x)))" have "?L = changes (map ff (smods pp qq))" using changes_poly_at_alternative unfolding ff_def by blast also have "... = changes (map ff (spmods pp qq))" unfolding ff_def using spmods_smods_sgn_map_eq by simp also have "... = changes (map ff (map (map_poly R\<^sub>1) (spmods p q)))" unfolding pp_def qq_def using spmods_eq by simp also have "... = ?R" proof - have "ff \ map_poly R\<^sub>1 = sign \ (\p. eval_poly hom p x)" unfolding ff_def comp_def by (simp add: map_poly_R_hom_commute poly_map_poly_eval_poly r2.sgn_sign) then show ?thesis unfolding changes_hpoly_at_def apply (subst (2) changes_map_sign_of_int_eq) by (simp add:comp_def) qed finally show ?thesis . qed have changes_eq_neg_inf: "changes_poly_neg_inf (smods pp qq) = changes_poly_neg_inf (spmods p q)" (is "?L=?R") proof - have "?L = changes (map sgn_neg_inf (map (map_poly R\<^sub>1) (spmods p q)))" unfolding changes_poly_neg_inf_def spmods_smods_sgn_map_eq by (simp add: spmods_eq[folded pp_def qq_def]) also have "... = changes (map (sgn_neg_inf \ (map_poly R\<^sub>1)) (spmods p q))" using map_map by simp also have "... = changes (map ((sign:: _ \ real) \ sgn_neg_inf) (spmods p q))" proof - have "(sgn_neg_inf \ (map_poly R\<^sub>1)) = of_int o sign \ sgn_neg_inf" unfolding sgn_neg_inf_def comp_def by (auto simp:r1.sgn_sign) then show ?thesis by (simp add:comp_def) qed also have "... = changes (map sgn_neg_inf (spmods p q))" apply (subst (2) changes_map_sign_of_int_eq) by (simp add:comp_def) also have "... = ?R" unfolding changes_poly_neg_inf_def by simp finally show ?thesis . qed have changes_eq_pos_inf: "changes_poly_pos_inf (smods pp qq) = changes_poly_pos_inf (spmods p q)" (is "?L=?R") proof - have "?L = changes (map sgn_pos_inf (map (map_poly R\<^sub>1) (spmods p q)))" unfolding changes_poly_pos_inf_def spmods_smods_sgn_map_eq by (simp add: spmods_eq[folded pp_def qq_def]) also have "... = changes (map (sgn_pos_inf \ (map_poly R\<^sub>1)) (spmods p q))" using map_map by simp also have "... = changes (map ((sign:: _ \ real) \ sgn_pos_inf) (spmods p q))" proof - have "(sgn_pos_inf \ (map_poly R\<^sub>1)) = of_int o sign \ sgn_pos_inf" unfolding sgn_pos_inf_def comp_def by (auto simp:r1.sgn_sign) then show ?thesis by (auto simp:comp_def) qed also have "... = changes (map sgn_pos_inf (spmods p q))" apply (subst (2) changes_map_sign_of_int_eq) by (simp add:comp_def) also have "... = ?R" unfolding changes_poly_pos_inf_def by simp finally show ?thesis . qed show "changes_itv_spmods a b p q = changes_itv_smods (R\<^sub>2 a) (R\<^sub>2 b) (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" unfolding changes_itv_spmods_def changes_itv_smods_def using changes_eq_at by (simp add: Let_def pp_def qq_def) show "changes_R_spmods p q = changes_R_smods (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" unfolding changes_R_spmods_def changes_R_smods_def Let_def using changes_eq_neg_inf changes_eq_pos_inf by (simp add: pp_def qq_def) show "changes_gt_spmods a p q = changes_gt_smods (R\<^sub>2 a) (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" unfolding changes_gt_spmods_def changes_gt_smods_def Let_def using changes_eq_at changes_eq_pos_inf by (simp add: pp_def qq_def) show "changes_le_spmods b p q = changes_le_smods (R\<^sub>2 b) (map_poly R\<^sub>1 p) (map_poly R\<^sub>1 q)" unfolding changes_le_spmods_def changes_le_smods_def Let_def using changes_eq_at changes_eq_neg_inf by (simp add: pp_def qq_def) qed end end \ No newline at end of file diff --git a/thys/Van_der_Waerden/Van_der_Waerden.thy b/thys/Van_der_Waerden/Van_der_Waerden.thy --- a/thys/Van_der_Waerden/Van_der_Waerden.thy +++ b/thys/Van_der_Waerden/Van_der_Waerden.thy @@ -1,943 +1,943 @@ theory Van_der_Waerden imports Main "HOL-Library.FuncSet" Digits begin section \Van der Waerden's Theorem\ text \In combinatorics, Van der Waerden's Theorem is about arithmetic progressions of a certain length of the same colour in a colouring of an interval. In order to state the theorem and to prove it, we need to formally introduce arithmetic progressions. We will express $k$-colourings as functions mapping an integer interval to the set $\{0,\dots , k-1 \}$ of colours.\ subsection \Arithmetic progressions\ text \A sequence of integer numbers with the same step size is called an arithmetic progression. We say an $m$-fold arithmetic progression is an arithmetic progression with multiple step lengths.\ text \ Arithmetic progressions are defined in the following using the variables: \begin{tabular}{lcp{8cm}} $start$:& \int\& starting value\\ $step$:& \nat\& positive integer for step length\\ $i$:& \nat\& $i$-th value in the arithmetic progression \\ \end{tabular}\ definition arith_prog :: "int \ nat \ nat \ int" where "arith_prog start step i = start + int (i * step)" text \ An $m$-fold arithmetic progression (which we will also call a multi-arithmetic progression) is defined in the following using the variables: \begin{tabular}{lcp{8cm}} $dims$:& \nat\& number of dimensions/step directions of $m$-fold arithmetic progression\\ $start$:& \int\& starting value\\ $steps$:& \nat \ nat\& function of steps, returns step in $i$-th dimension for $i\in[0..nat \ nat\& function of coefficients, returns coefficient in $i$-th dimension for $i\in[0.. definition multi_arith_prog :: "nat \ int \ (nat \ nat) \ (nat \ nat) \ int" where "multi_arith_prog dims start steps c = start + int (\iAn $m$-fold arithmetic progression of dimension $1$ is also an arithmetic progression and vice versa. This is shown in the following lemmas.\ lemma multi_to_arith_prog: "multi_arith_prog 1 start steps c = arith_prog start (steps 0) (c 0)" unfolding multi_arith_prog_def arith_prog_def by auto lemma arith_prog_to_multi: "arith_prog start step c = multi_arith_prog 1 start (\_. step) (\_. c)" unfolding multi_arith_prog_def arith_prog_def by auto text \To show that an arithmetic progression is well-defined, we introduce the following predicate. It assures that \arith_prog start step ` [0.. is contained in the integer interval $[a..b]$.\ definition is_arith_prog_on :: "nat \ int \ nat \ int \ int \ bool" where "is_arith_prog_on l start step a b \ (start \ a \ arith_prog start step (l-1) \ b)" text \Furthermore, we have monotonicity for arithmetic progressions.\ lemma arith_prog_mono: assumes "c \ c'" shows "arith_prog start step c \ arith_prog start step c'" using assms unfolding arith_prog_def by (auto intro: mult_mono) text \Now, we state the well-definedness of an arithmetic progression of length $l$ in an integer interval $[a..b]$. Indeed, \is_arith_prog_on\ guarantees that every element of \arith_prog start step\ of length $l$ lies in $[a..b]$.\ lemma is_arith_prog_onD: assumes "is_arith_prog_on l start step a b" assumes "c \ {0.. {a..b}" proof - have "arith_prog start step 0 \ arith_prog start step c" by (rule arith_prog_mono) auto hence "arith_prog start step c \ a" using assms by (simp add: arith_prog_def is_arith_prog_on_def add_increasing2) moreover have "arith_prog start step (l-1) \ arith_prog start step c" by (rule arith_prog_mono) (use assms(2) in auto) hence "arith_prog start step c \ b" using assms unfolding arith_prog_def is_arith_prog_on_def by linarith ultimately show ?thesis by auto qed text \We also need a predicate for an $m$-fold arithmetic progression to be well-defined. It assures that \multi_arith_prog start step ` [0.. is contained in $[a..b]$.\ definition is_multi_arith_prog_on :: "nat \ nat \ int \ (nat \ nat) \ int \ int \ bool" where "is_multi_arith_prog_on l m start steps a b \ (start \ a \ multi_arith_prog m start steps (\_. l-1) \ b)" text \Moreover, we have monotonicity for $m$-fold arithmetic progressions as well.\ lemma multi_arith_prog_mono: assumes "\i. i < m \ c i \ c' i" shows "multi_arith_prog m start steps c \ multi_arith_prog m start steps c'" using assms unfolding multi_arith_prog_def by (auto intro!: sum_mono intro: mult_right_mono) text \Finally, we get the well-definedness for $m$-fold arithmetic progressions of length $l$. Here, \is_multi_arith_prog_on\ guarantees that every element of \multi_arith_prog start step\ of length $l$ lies in $[a..b]$.\ lemma is_multi_arith_prog_onD: assumes "is_multi_arith_prog_on l m start steps a b" assumes "c \ {0.. {0.. {a..b}" proof - have "multi_arith_prog m start steps (\_. 0) \ multi_arith_prog m start steps c" by (rule multi_arith_prog_mono) auto hence "multi_arith_prog m start steps c \ a" using assms by (simp add: multi_arith_prog_def is_multi_arith_prog_on_def) moreover have "multi_arith_prog m start steps (\_. l-1) \ multi_arith_prog m start steps c" by (rule multi_arith_prog_mono) (use assms in force) hence "multi_arith_prog m start steps c \ b" using assms by (simp add: multi_arith_prog_def is_multi_arith_prog_on_def) ultimately show ?thesis by auto qed subsection \Van der Waerden's Theorem\ text \The property for a number $n$ to fulfill Van der Waerden's theorem is the following:\\ For a $k$-colouring col of $[a..b]$ there exist \begin{itemize} \item $start$: starting value of an arithmetic progression \item $step$: step length of an arithmetic progression \item $j$: colour \end{itemize} such that \arith_prog start step\ is a valid arithmetic progression of length $l$ lying in $[a..b]$ of the same colour $j$. The following variables will be used:\\ \begin{tabular}{lcp{8cm}} $k$:& \nat\& number of colours in segment colouring on $[a..b]$\\ $l$:& \nat\& length of arithmetic progression\\ $n$:& \nat\& number fulfilling Van der Waerden's Theorem\\ \end{tabular} \ definition vdw :: "nat \ nat \ nat \ bool" where "vdw k l n \ (\a b col. b + 1 \ a + int n \ col \ {a..b} \ {.. (\j start step. j < k \ step > 0 \ is_arith_prog_on l start step a b \ arith_prog start step ` {.. col -` {j} \ {a..b}))" text \To better work with the property of Van der Waerden's theorem, we introduce an elimination rule.\ lemma vdwE: assumes "vdw k l n" "b + 1 \ a + int n" "col \ {a..b} \ {.. 0" "is_arith_prog_on l start step a b" "arith_prog start step ` {.. col -` {j} \ {a..b}" using assms that unfolding vdw_def by metis text \Van der Waerden's theorem implies that the number fulfilling it is positive. This is show in the following lemma.\ lemma vdw_imp_pos: assumes "vdw k l n" "l > 0" shows "n > 0" proof (rule Nat.gr0I) assume [simp]: "n = 0" show False using assms by (elim vdwE[where a = 1 and b = 0 and col = "\_. 0"]) (auto simp: lessThan_empty_iff) qed text \Van der Waerden's Theorem is trivial for a non-existent colouring. It also makes no sense for arithmetic progressions of length 0.\ lemma vdw_0_left [simp, intro]: "n>0 \ vdw 0 l n" by (auto simp: vdw_def) text \In the case of $k=1$, Van der Waerden's Theorem holds. Then every number has the same colour, hence also the arithmetic progression. A possible choice for the number fulfilling Van der Waerden Theorem is $l$.\ lemma vdw_1_left: assumes "l>0" shows "vdw 1 l l" unfolding vdw_def proof (safe, goal_cases) case (1 a b col) have "arith_prog a 1 ` {.. {a..b}" using 1(1) by (auto simp: arith_prog_def) also have "{a..b} = col -` {0} \ {a..b}" using 1(2) by auto finally have "arith_prog a 1 ` {.. col -` {0} \ {a..b}" by auto moreover have "is_arith_prog_on l a 1 a b" unfolding is_arith_prog_on_def arith_prog_def using 1 assms by auto ultimately show "\j start step. j < 1 \ 0 < step \ is_arith_prog_on l start step a b \ arith_prog start step ` {.. col -` {j} \ {a..b}" by auto qed text \In the case $l=1$, Van der Waerden's Theorem holds. As the length of the arithmetic progression is $1$, it consists of just one element. Thus every nonempty integer interval fulfills the Van der Waerden property. We can prove $N_{k,1}$ to be $1$.\ lemma vdw_1_right: "vdw k 1 1" unfolding vdw_def proof safe fix a b :: int and col :: "int \ nat" assume *: "a + int 1 \ b + 1" "col \ {a..b} \ {.. col -` {col a} \ {a..b}" using * by auto finally have "arith_prog a 1 ` {..<1} \ col -` {col a} \ {a..b}" by auto moreover have "is_arith_prog_on 1 a 1 a b" unfolding is_arith_prog_on_def arith_prog_def using * by auto ultimately show "\j start step. j < k \ 0 < step \ is_arith_prog_on 1 start step a b \ arith_prog start step ` {..<1} \ col -` {j} \ {a..b}" using \col a by blast qed text \In the case $l=2$, Van der Waerden's Theorem holds as well. Here, any two distinct numbers form an arithmetic progression of length $2$. Thus we only have to find two numbers with the same colour. Using the pigeonhole principle on $k+1$ values, we can find two integers with the same colour.\ lemma vdw_2_right: "vdw k 2 (k+1)" unfolding vdw_def proof safe fix a b :: int and col :: "int \ nat" assume *: "a + int (k + 1) \ b + 1" "col \ {a..b} \ {.. {.. card {a..b}" using *(1) by auto ultimately have "card (col ` {a..b}) < card {a..b}" using * by (metis card_lessThan card_mono finite_lessThan le_less_trans less_add_one not_le) then have "\ inj_on col {a..b}" using pigeonhole[of col "{a..b}"] by auto then obtain start start_step where pigeon: "col start = col start_step" "start < start_step" "start \ {a..b}" "start_step \ {a..b}" using inj_onI[of "{a..b}" col] by (metis not_less_iff_gr_or_eq) define step where "step = nat (start_step - start)" define j where "j = col start" have "j < k" unfolding j_def using *(2) pigeon(3) by auto moreover have "0 < step" unfolding step_def using pigeon(2) by auto moreover have "is_arith_prog_on 2 start step a b" unfolding is_arith_prog_on_def arith_prog_def step_def using pigeon by auto moreover { have "arith_prog start step i \ {start, start_step}" if "i<2" for i using that arith_prog_def step_def by (auto simp: less_2_cases_iff) also have "\ \ col -` {j} \ {a..b}" using pigeon unfolding j_def by auto finally have "arith_prog start step ` {..<2} \ col -` {j} \ {a..b}" by auto } ultimately show "\j start step. j < k \ 0 < step \ is_arith_prog_on 2 start step a b \ arith_prog start step ` {..<2} \ col -` {j} \ {a..b}" by blast qed text \In order to prove Van der Waerden's Theorem, we first prove a slightly different lemma. The statement goes as follows:\\ For a $k$-colouring $col$ on $[a..b]$ there exist \begin{itemize} \item $start$: starting value of an arithmetic progression \item $steps$: step length of an arithmetic progression \end{itemize} such that \f = multi_arith_prog m start step\ is a valid $m$-fold arithmetic progression of length $l$ lying in $[a..b]$ such that for every $snat\& number of colours in segment colouring of $[a..b]$\\ $m$:& \nat\& dimension of $m$-fold arithmetic progression\\ $l$:& \nat\& $l+1$ is length of $m$-fold arithmetic progression\\ $n$:& \nat\& number fulfilling \vdw_lemma\\\ \end{tabular} \ definition vdw_lemma :: "nat \ nat \ nat \ nat \ bool" where "vdw_lemma k m l n \ (\a b col. b + 1 \ a + int n \ col \ {a..b} \ {.. (\start steps. (\i 0) \ is_multi_arith_prog_on (l+1) m start steps a b \ ( let f = multi_arith_prog m start steps in (\c \ {0.. {0..l}. \s j \ s. c j < l) \ col (f c) = col (f (\i. if i \ s then 0 else c i))))))" text \To better work with this property, we introduce an elimination rule for \vdw_lemma\.\ lemma vdw_lemmaE: fixes a b :: int assumes "vdw_lemma k m l n" "b + 1 \ a + int n" "col \ {a..b} \ {..i. i < m \ steps i > 0" "is_multi_arith_prog_on (l+1) m start steps a b" "let f = multi_arith_prog m start steps in \c \ {0.. {0..l}. \s j \ s. c j < l) \ col (f c) = col (f (\i. if i \ s then 0 else c i))" using assms that unfolding vdw_lemma_def by blast text \To simplify the following proof, we show the following formula.\ lemma sum_mod_poly: assumes "(k::nat)>0" shows "(k - 1) * (\ n\{..nn = int k ^ q - 1" by (induction q) (auto simp: algebra_simps) also have "\ < int (k ^ q)" by simp finally show ?thesis by linarith qed text \The proof of Van der Waerden's Theorem now proceeds in three steps:\\ \begin{itemize} \item Firstly, we show that the \vdw\ property for all $k$ proves the \vdw_lemma\ for fixed $l$ but arbitrary $k$ and $m$. This is done by induction over $m$. \item Secondly, we show that \vdw_lemma\ implies the induction step of \vdw\ using the pigeonhole principle. \item Lastly, we combine the previous steps in an induction over $l$ to show Van der Waerden's Theorem in the general setting. \end{itemize}\ text \Firstly, we need to show that \vdw\ for arbitrary $k$ implies \vdw_lemma\ for fixed $l$. As mentioned earlier, we use induction over $m$.\ lemma vdw_imp_vdw_lemma: fixes l assumes vdw_assms: "\k'. k'>0 \ \n_k'. vdw k' l n_k'" and "l \ 2" and "m > 0" and "k > 0" shows "\N. vdw_lemma k m l N" using \m>0\ \k>0\ proof (induction m rule: less_induct) case (less m) consider "m=1" | "m>1" using less.prems by linarith then show ?case proof cases text \ Case $m=1$: Show \vdw_lemma\ for arithmetic progression, Induction start. \ assume "m = 1" obtain n where vdw: "vdw k l n" using vdw_assms \k>0\ by blast define N where "N = 2*n" have "l>0" and "l>1" using \l\2\ by auto have "vdw_lemma k m l N" unfolding vdw_lemma_def proof (safe, goal_cases) case (1 a b col) text \ Divide $[a..b]$ in two intervals $I_1$, $I_2$ of same length and obtain arithmetic progression of length $l$ in $I_1$. \ have col_restr: "col \ {a..a + int n - 1} \ {.. 0" "is_arith_prog_on l start step a (a + int n -1)" "arith_prog start step ` {.. col -` {j} \ {a..a + int n - 1}" using vdw 1 unfolding N_def by (elim vdwE)(auto simp:is_arith_prog_on_def) have range_prog_lessThan_l: "arith_prog start step i \ {a..a + int n -1}" if "i < l" for i using that prog by auto have "{a..a + int n-1}\{a..b}" using N_def "1"(1) by auto then have "a + 2* int n - 1 \ b" using 1(1) unfolding N_def by auto text \ Show that \arith_prog start step\ is an arithmetic progression of length $l+1$ in $[a..b]$. \ have prog_in_ivl: "arith_prog start step i \ {a..b}" if "i \ l" for i proof (cases "i=l") case False have "i{a..a + int n-1}\{a..b}\ by force next case True text \ Show $\step\\leq |I_1|$ then have \arith_prog start step (l+1)\[a..b]\ as \arith_prog start step (l+1) = arith_prog start step l + step\ \ have "start \ {a..a + int n -1}" using range_prog_lessThan_l[of 0] unfolding arith_prog_def by (simp add: \0 < l\) moreover have "start + int step \ {a..a + int n -1}" using range_prog_lessThan_l[of 1] unfolding arith_prog_def by (metis \1 < l\ mult.left_neutral) ultimately have "step \ n" by auto have "arith_prog start step (l-1) \ {a..a + int n -1}" using range_prog_lessThan_l[of "l-1"] unfolding arith_prog_def using \0 < l\ diff_less less_numeral_extra(1) by blast moreover have "arith_prog start step l = arith_prog start step (l-1) + int step" unfolding arith_prog_def using \0 < l\ mult_eq_if by force ultimately have "arith_prog start step l \ {a..b}" using \step\n\ N_def \a + 2* int n -1 \ b\ by auto then show ?thesis using range_prog_lessThan_l using True by force qed have col_prog_eq: "col (arith_prog start step k) = j" if "k < l" for k using prog that by blast define steps :: "nat \ nat" where steps_def: "steps = (\i. step)" define f where "f = multi_arith_prog 1 start steps" have rel_prop_1: "col (f c) = col (f (\i. if i < s then 0 else c i))" if "c \ {0..<1} \ {0..l}" "s<1" "\j\s. c j < l" for c s using that by auto have arith_prog_on: "is_multi_arith_prog_on (l+1) m start steps a b" using prog(3) unfolding is_arith_prog_on_def is_multi_arith_prog_on_def using \m=1\ arith_prog_to_multi steps_def prog_in_ivl by auto show ?case by (rule exI[of _ start], rule exI[of _ steps]) (use rel_prop_1 \step > 0\ \m = 1\ arith_prog_on col_prog_eq multi_to_arith_prog in \auto simp: f_def Let_def steps_def\) qed then show ?case .. next text \ Case $m>1$: Show \vdw_lemma\ for $m$-fold arithmetic progression, Induction step $(m-1) \longrightarrow m$. \ assume "m>1" obtain q where vdw_lemma_IH:"vdw_lemma k (m-1) l q" using \1 < m\ less by force have "k^q>0" using \k>0\ by auto obtain n_kq where vdw: "vdw (k^q) l n_kq" using vdw_assms \k^q>0\ by blast define N where "N = q + 2 * n_kq" text \Idea: $[a..b] = I_1 \cup I_2$ where $|I_1| = 2*n_{k,q}$ and $|I_2| = q$. Divide $I_1$ into blocks of length $q$ and define a new colouring on the set of $q$-blocks where the colour of the block is the $k$-basis representation where the $i$-th digit corresponds to the colour of the $i$-th element in the block. Get an arithmetic progression of $q$-blocks of length $l+1$ in $I_1$, such that the first $l$ $q$-blocks have the same colour. The step of the block-arithmetic progression is going to be the additional step in the induction over $m$. \ have "vdw_lemma k m l N" unfolding vdw_lemma_def proof (safe, goal_cases) case (1 a b col) have "n_kq>0" using vdw_imp_pos vdw \l\2\ by auto then have "N>0" by (simp add:N_def) then have "a\b" using 1 by auto then have "k>0" using 1 by (intro Nat.gr0I) force have "l>0" and "l>1" using \l\2\ by auto interpret digits k by (simp add: \0 < k\ digits_def) define col1 where "col1 = (\ x. from_digits q (\y. col (x + y)))" have range_col1: "col1\{a..a + int n_kq - 1} \ {..{a..a + int n_kq - 1}" then have col_xn:"col (x + int n)\{.. k - 1" if "nk>0\ by (auto) have "(\n (\n = (k-1) * (\n < k^q" using sum_mod_poly \k>0\ by auto finally show "col1 x 0" "is_arith_prog_on l start step a (a + int n_kq - 1)" "arith_prog start step ` {.. col1 -` {j} \ {a..a + int n_kq -1}" using vdw range_col1 by (elim vdwE) (auto simp: \k>0\) have range_prog_lessThan_l: "arith_prog start step i \ {a..a + int n_kq -1}" if "i < l" for i using that prog by auto have prog_in_ivl: "arith_prog start step i \ {a..a + 2 * int n_kq -1}" if "i \ l" for i proof (cases "i=l") case False then have "i {a..a + int n_kq -1}" using range_prog_lessThan_l[of 0] unfolding arith_prog_def by (simp add: \0 < l\) moreover have "start + step \ {a..a + int n_kq -1}" using range_prog_lessThan_l[of 1] unfolding arith_prog_def by (metis \1 < l\ mult.left_neutral) ultimately have "step \ n_kq" by auto have "arith_prog start step (l-1) \ {a..a + int n_kq -1}" using range_prog_lessThan_l[of "l-1"] unfolding arith_prog_def using \0 < l\ diff_less less_numeral_extra(1) by blast moreover have "arith_prog start step l = arith_prog start step (l-1) + step" unfolding arith_prog_def using \0 < l\ mult_eq_if by force ultimately have "arith_prog start step l \ {a..a + 2 * int n_kq - 1}" using \step\n_kq\ by auto then show ?thesis using range_prog_lessThan_l using True by force qed have col_prog_eq: "col1 (arith_prog start step k) = j" if "k < l" for k using prog that by blast have digit_col1:"digit (col1 x) y = col (x+int y)" if "x\{a..{..j'. j' x+j'\{a..b}" using "1"(1) N_def that(1) by force then have "\j'. j' (\y. col (x+int y)) j' < k" using 1 that by auto then show "digit (from_digits q (\xa. col (x + int xa))) y = col (x + int y)" using digit_from_digits that 1 by auto qed text \ Impact on the colour when taking the block-step. \ have one_step_more: "col (arith_prog start' step i) = digit j (nat (start'-start))" if "start'\{start..{.. start'" using that by simp have shift_arith_prog: "arith_prog start step i + (start' - start) = arith_prog start' step i" unfolding arith_prog_def by simp define diff where "diff = nat (start'-start)" have "diff \{..{a..a + 2 * int n_kq-1}" using prog(4) that by auto ultimately show ?thesis using digit_col1[where x = "arith_prog start step i" and y = "diff"] prog 1 \diff \{.. by auto qed then show ?thesis unfolding diff_def 1 by (auto simp: \start\start'\ shift_arith_prog) qed have one_step_more': "col (arith_prog start' step i) = col (arith_prog start' step 0)" if "start'\{start..{.. start + int q - 1 + 1" by linarith have "{start..start + int q-1} \ {a..b}" using prog N_def 1(1) by (force simp: arith_prog_def is_arith_prog_on_def) then have col': "col \ {start..start + int q-1} \ {.. Obtain an $(m-1)$-fold arithmetic progression in the starting $q$-bolck of the block arithmetic progression. \ obtain start_m steps_m where step_m_pos: "\i. i < m - 1 \ 0 < steps_m i" and is_multi_arith_prog: "is_multi_arith_prog_on (l+1) (m - 1) start_m steps_m start (start + int q - 1)" and g_aux: "let g = multi_arith_prog (m - 1) start_m steps_m in \c\{0.. {0..l}. \sj\s. c j < l) \ col (g c) = col (g (\i. if i \ s then 0 else c i))" by (rule vdw_lemmaE[OF vdw_lemma_IH start_q col']) blast define g where "g = multi_arith_prog (m-1) start_m steps_m" have g: "col (g c) = col (g (\i. if i \ s then 0 else c i))" if "c \ {0..<(m-1)} \ {0..l}" "s < m - 1" "\j \ s. c j < l" for c s using g_aux that unfolding g_def Let_def by blast have range_g: "g c \ {start..start + int q - 1}" if "c \ {0.. {0..<(l+1)}" for c using is_multi_arith_prog_onD[OF is_multi_arith_prog that] by (auto simp: g_def) text \Obtain an $m$-fold arithmetic progression by adding the block-step.\ define steps :: "nat \ nat" where steps_def: "steps = (\i. (if i=0 then step else steps_m (i-1)))" define f where "f = multi_arith_prog m start_m steps" have f_step_g: "f c = int (c 0*step) + g (c \ Suc)" for c proof - have "f c = start_m + int (\i = start_m + int (c 0 * steps 0) + int (\i = start_m + int (c 0 * step) + int (\i Show that this $m$-fold arithmetic progression fulfills all needed properties. \ have steps_gr_0: "\i start_m" using is_multi_arith_prog unfolding is_multi_arith_prog_on_def using is_arith_prog_on_def prog(3) by force moreover { have "f (\_. l) = arith_prog (g ((\_. l) \ Suc)) step l" using f_step_g unfolding arith_prog_def by auto also have "g ((\_. l) \ Suc) \ start + q" using range_g[of "(\_. l) \ Suc"] by auto then have "arith_prog (g ((\_. l) \ Suc)) step l \ arith_prog start step l + q" unfolding arith_prog_def by auto also have "\\ b" using prog_in_ivl[of l] using is_multi_arith_prog unfolding is_multi_arith_prog_on_def using "1"(1) N_def by auto finally have "f (\_. l) \ b" by auto } ultimately show ?thesis unfolding is_multi_arith_prog_on_def f_def by auto qed text \ Show the relational property for all $s$. \ have rel_prop_1: "col (f c) = col (f (\i. if i \ s then 0 else c i))" if "c \ {0.. {0..l}" "sj\s. c j < l" for c s proof (cases "s = 0") case True have "c 0 < l" using that(3) True by auto have range_c_Suc: "c \ Suc \ {0.. {0..l}" using that(1) by auto have "f c = arith_prog (g (c \ Suc)) step (c 0)" using f_step_g unfolding arith_prog_def by auto then have "col (f c) = col (arith_prog (g (c \ Suc)) step 0)" using one_step_more'[of "g (c \ Suc)" "c 0"] \c 0 < l\ range_g[of "c \ Suc"] range_c_Suc atLeastLessThanSuc_atLeastAtMost by auto also { have "(\xx=1..x. x-1)" "Suc"]) (auto simp: steps_def split:if_splits) also have "\ = (\x Suc)) step 0 = f (\i. if i \ s then 0 else c i)" unfolding f_def g_def multi_arith_prog_def arith_prog_def using True by auto } finally show ?thesis by auto next case False hence s_greater_0: "s > 0" by auto have range_c_Suc: "c \ Suc \ {0.. {0..l}" using that(1) by auto have "c 0 < l" using \s>0\ that by auto have g_IH: "col (g c') = col (g (\i. if i \ s' then 0 else c' i))" if "c' \ {0.. {0..l}" "s'j\s'. c' j < l" for c' s' using g_aux that unfolding multi_arith_prog_def g_def by (auto simp: Let_def) have g_shift_IH: "col (g (c \ Suc)) = col (g ((\i. if i\{1..t} then 0 else c i) \ Suc))" if "c \ {1.. {0..l}" "t\{1..j\{1..t}. c j < l" for c t proof - have "(\i. (if i \ t - 1 then 0 else (c \ Suc) i)) = (\i. (if i \ {1..t} then 0 else c i)) \ Suc" using that by (auto split: if_splits simp:fun_eq_iff) then have right: "g (\i. if i \ (t-1) then 0 else (c \ Suc) i) = g ((\i. if i\{1..t} then 0 else c i) \ Suc)" by auto have "(c \ Suc)\ {0.. {0..l}" using that(1) by auto moreover have "t-1j\t-1. (c \ Suc) j < l" using that by auto ultimately have "col (g (c \ Suc)) = col (g (\i. (if i \ t-1 then 0 else (c \ Suc) i)))" using g_IH[of "(c \ Suc)" "t-1"] by auto with right show ?thesis by auto qed have "col (f c) = col (int (c 0 * step) + g (c \ Suc))" using f_step_g by simp also have "int (c 0 * step) + g (c \ Suc) = arith_prog (g (c \ Suc)) step (c 0)" by (simp add: arith_prog_def) also have "col \ = col (arith_prog (g (c \ Suc)) step 0)" using one_step_more'[of "g (c \ Suc)" "c 0"] \c 0 < l\ range_g[of "c \ Suc"] range_c_Suc atLeastLessThanSuc_atLeastAtMost by auto also have "\ = col (g (c \ Suc))" unfolding arith_prog_def by auto also have "\ = col (g ((\i. if i\{1..s} then 0 else c i) \ Suc))" using g_shift_IH[of "c" s] \s>0\ that by force also have "\ = col ((\c. int (c 0 * step) + g (c \ Suc))(\i. if i\s then 0 else c i))" by (auto simp: g_def multi_arith_prog_def) also have "\ = col (f (\i. if i \ s then 0 else c i))" unfolding f_step_g by auto finally show ?thesis by simp qed show ?case by (rule exI[of _ start_m], rule exI[of _ steps]) (use steps_gr_0 is_multi_on_f rel_prop_1 in \auto simp: f_def Let_def steps_def\) qed then show ?case .. qed qed text \ Secondly, we show that \vdw_lemma\ implies the induction step of Van der Waerden's Theorem using the pigeonhole principle. \ lemma vdw_lemma_imp_vdw: assumes "vdw_lemma k k l N" shows "vdw k (Suc l) N" unfolding vdw_def proof (safe, goal_cases) text \Idea: Proof uses pigeonhole principle to guarantee the existence of an arithmetic progression of length $l+1$ with the same colour. \ case (1 a b col) obtain start steps where prog: "\i. i < k \ steps i > 0" "is_multi_arith_prog_on (l+1) k start steps a b" "let f = multi_arith_prog k start steps in \c \ {0.. {0..l}. \s j \ s. c j < l) \ col (f c) = col (f (\i. if i \ s then 0 else c i))" using assms 1 by (elim vdw_lemmaE[where a=a and b=b and col=col and m=k and k=k and l=l and n=N]) auto text \ Obtain a $k$-fold arithmetic progression $f$ of length $l$ from assumptions. \ define f where "f = multi_arith_prog k start steps" have rel_propE: "col (f c) = col (f (\i. if i \ s then 0 else c i))" if "c \ {0.. {0..l}" "s j \ s. c j < l" for c s using prog(3) that unfolding f_def Let_def by auto text \There are $k+1$ values $a_r = f(0,\dots,0,l,\dots,l)$ with $0\leq r\leq k$ zeros.\ define a_r where "a_r = (\r. f (\i. (if i {a..b}" unfolding a_r_def f_def by (intro is_multi_arith_prog_onD[OF prog(2)]) auto thus ?thesis using 1 by blast qed then have "(col \ a_r) ` {.. {.. a_r) ` {.. card {.. inj_on (col \ a_r) {.. a_r" "{..Using the pigeonhole principle get $r_1$ and $r_2$ where $a_{r_1}$ and $a_{r_2}$ have the same colour.\ then obtain r1 r2 where pigeon_cols: "r1\{..{.. a_r) r1 = (col \ a_r) r2" by (metis (mono_tags, lifting) linear linorder_inj_onI) text \ Show that the following function $h$ is an arithmetic progression which fulfills all properties for Van der Waerden's Theorem. \ define h where "h = (\x. f (\i. (if ir1 by (intro arg_cong[where f = f]) auto moreover have "h l = a_r r1" unfolding h_def a_r_def using \r1 by (metis le_eq_less_or_eq less_le_trans) ultimately have "col (h 0) = col (h l)" using pigeon_cols(4) by auto have h_col: "col (h 0) = col (h i)" if "i\{..col (h 0) = col (h l)\ by auto next case False then have "i{0.. {0..l}" using that by auto moreover have "(\j\r2-1. ?c j < l)" using \i pigeon_cols(3) by force ultimately have "col (f ?c) = col (f (\i. if i \ r2-1 then 0 else ?c i))" using rel_propE[of ?c "r2-1"] pigeon_cols by simp then show ?thesis unfolding h_def f_def - by (smt (z3) Nat.lessE One_nat_def add_diff_cancel_left' + by (smt (verit) Nat.lessE One_nat_def add_diff_cancel_left' le_less less_Suc_eq_le multi_arith_prog_mono plus_1_eq_Suc) qed define h_start where "h_start = start + l*(\i\{r2..i\{r1..xx = r2..x = r1..xx = (\xx=r1..x. int (if x < r1 then 0 else y) * int (steps x))"] \r1 by auto also have "\ = (\x=r1.. = int y * (\x=r1..x. int (if x < r1 then 0 else if x < r2 then y else l) * int (steps x))" have "(\xx=r1.. {r2..r1 < r2\ \r2 < k\ by auto also have "(\x\\. aux_left x) = (\x=r1..x=r2..x=r1..x=r1..x=r2..x=r2..r1 < r2\ by (intro sum.cong) (auto simp: aux_left_def) finally show ?thesis by (simp add: aux_left_def sum_distrib_left) qed then show ?thesis unfolding arith_prog_def h_start_def h_step_def h_def f_def multi_arith_prog_def by (auto split:if_splits) qed define j where "j = col (h 0)" have case_j: "jcol (h 0) = col (h l)\ \h l = a_r r1\ j_def pigeon_cols(1) by auto have case_step: "h_step > 0" unfolding h_step_def using pigeon_cols by (intro sum_pos prog(1)) auto have range_h: "h i \ {a..b}" if "i < l + 1" for i unfolding h_def f_def by (rule is_multi_arith_prog_onD[OF prog(2)]) (use that in auto) have case_on: "is_arith_prog_on (l+1) h_start h_step a b" unfolding is_arith_prog_on_def h_arith_prog using range_h[of 0] range_h[of l] by (auto simp: Max_ge[of "{a..b}"] Min_le[of "{a..b}"] h_arith_prog arith_prog_def) have case_col: "h ` {.. col -` {j} \ {a..b}" using h_col range_h unfolding j_def by auto show ?case using case_j case_step case_on case_col by (auto simp: h_arith_prog) qed text \ Lastly, we assemble all lemmas to finally prove Van der Waerden's Theorem by induction on $l$. The cases $l=1$ and the induction start $l=2$ are treated separately and have been shown earlier.\ theorem van_der_Waerden: assumes "l>0" "k>0" shows "\n. vdw k l n" using assms proof (induction l arbitrary: k rule: less_induct) case (less l) consider "l=1" | "l=2" | "l>2" using less.prems by linarith then show ?case proof (cases) assume "l=1" then show ?thesis using vdw_1_right by auto next assume "l=2" then show ?thesis using vdw_2_right by auto next assume "l > 2" then have "2\l-1" by auto from less.IH[of "l-1"] \l>2\ have "\k'. k'>0 \ \n. vdw k' (l-1) n" by auto with vdw_imp_vdw_lemma[of "l-1" k k] \l-1\2\ \k>0\ obtain N where "vdw_lemma k k (l-1) N" by auto then have "vdw k l N" using vdw_lemma_imp_vdw[of k "l-1" N] by (simp add: less.prems(1)) then show ?thesis by auto qed qed end