diff --git a/thys/Adaptive_State_Counting/ASC/ASC_Sufficiency.thy b/thys/Adaptive_State_Counting/ASC/ASC_Sufficiency.thy --- a/thys/Adaptive_State_Counting/ASC/ASC_Sufficiency.thy +++ b/thys/Adaptive_State_Counting/ASC/ASC_Sufficiency.thy @@ -1,874 +1,877 @@ theory ASC_Sufficiency imports ASC_Suite begin section \ Sufficiency of the test suite to test for reduction \ text \ This section provides a proof that the test suite generated by the adaptive state counting algorithm is sufficient to test for reduction. \ subsection \ Properties of minimal sequences to failures extending the deterministic state cover \ text \ The following two lemmata show that minimal sequences to failures extending the deterministic state cover do not with their extending suffix visit any state twice or visit a state also reached by a sequence in the chosen permutation of reactions to the deterministic state cover. \ lemma minimal_sequence_to_failure_extending_implies_Rep_Pre : assumes "minimal_sequence_to_failure_extending V M1 M2 vs xs" and "OFSM M1" and "OFSM M2" and "test_tools M2 M1 FAIL PM V \" and "V'' \ N (vs@xs') M1 V" and "prefix xs' xs" shows "\ Rep_Pre M2 M1 vs xs'" proof assume "Rep_Pre M2 M1 vs xs'" then obtain xs1 xs2 s1 s2 where "prefix xs1 xs2" "prefix xs2 xs'" "xs1 \ xs2" "io_targets M2 (initial M2) (vs @ xs1) = {s2}" "io_targets M2 (initial M2) (vs @ xs2) = {s2}" "io_targets M1 (initial M1) (vs @ xs1) = {s1}" "io_targets M1 (initial M1) (vs @ xs2) = {s1}" by auto then have "s2 \ io_targets M2 (initial M2) (vs @ xs1)" "s2 \ io_targets M2 (initial M2) (vs @ xs2)" "s1 \ io_targets M1 (initial M1) (vs @ xs1)" "s1 \ io_targets M1 (initial M1) (vs @ xs2)" by auto have "vs@xs1 \ L M1" using io_target_implies_L[OF \s1 \ io_targets M1 (initial M1) (vs @ xs1)\] by assumption have "vs@xs2 \ L M1" using io_target_implies_L[OF \s1 \ io_targets M1 (initial M1) (vs @ xs2)\] by assumption have "vs@xs1 \ L M2" using io_target_implies_L[OF \s2 \ io_targets M2 (initial M2) (vs @ xs1)\] by assumption have "vs@xs2 \ L M2" using io_target_implies_L[OF \s2 \ io_targets M2 (initial M2) (vs @ xs2)\] by assumption obtain tr1_1 where "path M1 (vs@xs1 || tr1_1) (initial M1)" "length tr1_1 = length (vs@xs1)" "target (vs@xs1 || tr1_1) (initial M1) = s1" using \s1 \ io_targets M1 (initial M1) (vs @ xs1)\ by auto obtain tr1_2 where "path M1 (vs@xs2 || tr1_2) (initial M1)" "length tr1_2 = length (vs@xs2)" "target (vs@xs2 || tr1_2) (initial M1) = s1" using \s1 \ io_targets M1 (initial M1) (vs @ xs2)\ by auto obtain tr2_1 where "path M2 (vs@xs1 || tr2_1) (initial M2)" "length tr2_1 = length (vs@xs1)" "target (vs@xs1 || tr2_1) (initial M2) = s2" using \s2 \ io_targets M2 (initial M2) (vs @ xs1)\ by auto obtain tr2_2 where "path M2 (vs@xs2 || tr2_2) (initial M2)" "length tr2_2 = length (vs@xs2)" "target (vs@xs2 || tr2_2) (initial M2) = s2" using \s2 \ io_targets M2 (initial M2) (vs @ xs2)\ by auto have "productF M2 M1 FAIL PM" using assms(4) by auto have "well_formed M1" using assms(2) by auto have "well_formed M2" using assms(3) by auto have "observable PM" by (meson assms(2) assms(3) assms(4) observable_productF) have "length (vs@xs1) = length tr2_1" using \length tr2_1 = length (vs @ xs1)\ by presburger then have "length tr2_1 = length tr1_1" using \length tr1_1 = length (vs@xs1)\ by presburger have "vs@xs1 \ L PM" using productF_path_inclusion[OF \length (vs@xs1) = length tr2_1\ \length tr2_1 = length tr1_1\ \productF M2 M1 FAIL PM\ \well_formed M2\ \well_formed M1\] by (meson Int_iff \productF M2 M1 FAIL PM\ \vs @ xs1 \ L M1\ \vs @ xs1 \ L M2\ \well_formed M1\ \well_formed M2\ productF_language) have "length (vs@xs2) = length tr2_2" using \length tr2_2 = length (vs @ xs2)\ by presburger then have "length tr2_2 = length tr1_2" using \length tr1_2 = length (vs@xs2)\ by presburger have "vs@xs2 \ L PM" using productF_path_inclusion[OF \length (vs@xs2) = length tr2_2\ \length tr2_2 = length tr1_2\ \productF M2 M1 FAIL PM\ \well_formed M2\ \well_formed M1\] by (meson Int_iff \productF M2 M1 FAIL PM\ \vs @ xs2 \ L M1\ \vs @ xs2 \ L M2\ \well_formed M1\ \well_formed M2\ productF_language) have "io_targets PM (initial M2, initial M1) (vs @ xs1) = {(s2, s1)}" using productF_path_io_targets_reverse [OF \productF M2 M1 FAIL PM\ \s2 \ io_targets M2 (initial M2) (vs @ xs1)\ \s1 \ io_targets M1 (initial M1) (vs @ xs1)\ \vs @ xs1 \ L M2\ \vs @ xs1 \ L M1\ ] proof - have "\c f. c \ initial (f::('a, 'b, 'c) FSM) \ c \ nodes f" by blast then show ?thesis by (metis (no_types) \\observable M2; observable M1; well_formed M2; well_formed M1; initial M2 \ nodes M2; initial M1 \ nodes M1\ \ io_targets PM (initial M2, initial M1) (vs @ xs1) = {(s2, s1)}\ assms(2) assms(3)) qed have "io_targets PM (initial M2, initial M1) (vs @ xs2) = {(s2, s1)}" using productF_path_io_targets_reverse [OF \productF M2 M1 FAIL PM\ \s2 \ io_targets M2 (initial M2) (vs @ xs2)\ \s1 \ io_targets M1 (initial M1) (vs @ xs2)\ \vs @ xs2 \ L M2\ \vs @ xs2 \ L M1\ ] proof - have "\c f. c \ initial (f::('a, 'b, 'c) FSM) \ c \ nodes f" by blast then show ?thesis by (metis (no_types) \\observable M2; observable M1; well_formed M2; well_formed M1; initial M2 \ nodes M2; initial M1 \ nodes M1\ \ io_targets PM (initial M2, initial M1) (vs @ xs2) = {(s2, s1)}\ assms(2) assms(3)) qed have "prefix (vs @ xs1) (vs @ xs2)" using \prefix xs1 xs2\ by auto have "sequence_to_failure M1 M2 (vs@xs)" using assms(1) by auto have "prefix (vs@xs1) (vs@xs')" using \prefix xs1 xs2\ \prefix xs2 xs'\ prefix_order.dual_order.trans same_prefix_prefix by blast have "prefix (vs@xs2) (vs@xs')" using \prefix xs2 xs'\ prefix_order.dual_order.trans same_prefix_prefix by blast have "io_targets PM (initial PM) (vs @ xs1) = {(s2,s1)}" using \io_targets PM (initial M2, initial M1) (vs @ xs1) = {(s2, s1)}\ assms(4) by auto have "io_targets PM (initial PM) (vs @ xs2) = {(s2,s1)}" using \io_targets PM (initial M2, initial M1) (vs @ xs2) = {(s2, s1)}\ assms(4) by auto have "(vs @ xs2) @ (drop (length xs2) xs) = vs@xs" by (metis \prefix xs2 xs'\ append_eq_appendI append_eq_conv_conj assms(6) prefixE) moreover have "io_targets PM (initial PM) (vs@xs) = {FAIL}" using sequence_to_failure_reaches_FAIL_ob[OF \sequence_to_failure M1 M2 (vs@xs)\ assms(2,3) \productF M2 M1 FAIL PM\] by assumption ultimately have "io_targets PM (initial PM) ((vs @ xs2) @ (drop (length xs2) xs)) = {FAIL}" by auto have "io_targets PM (s2,s1) (drop (length xs2) xs) = {FAIL}" using observable_io_targets_split [OF \observable PM\ \io_targets PM (initial PM) ((vs @ xs2) @ (drop (length xs2) xs)) = {FAIL}\ \io_targets PM (initial PM) (vs @ xs2) = {(s2, s1)}\] by assumption have "io_targets PM (initial PM) (vs@xs1@(drop (length xs2) xs)) = {FAIL}" using observable_io_targets_append [OF \observable PM\ \io_targets PM (initial PM) (vs @ xs1) = {(s2,s1)}\ \io_targets PM (s2,s1) (drop (length xs2) xs) = {FAIL}\] by simp have "sequence_to_failure M1 M2 (vs@xs1@(drop (length xs2) xs))" using sequence_to_failure_alt_def [OF \io_targets PM (initial PM) (vs@xs1@(drop (length xs2) xs)) = {FAIL}\ assms(2,3)] assms(4) by blast have "length xs1 < length xs2" - using \prefix xs1 xs2\ \xs1 \ xs2\ prefix_length_prefix by fastforce - have "xs = (xs1 @ (drop (length xs1) xs))" - by (metis (no_types) \(vs @ xs2) @ drop (length xs2) xs = vs @ xs\ \prefix xs1 xs2\ - append_assoc append_eq_conv_conj prefixE) - have "length xs1 < length xs" - using \prefix xs1 xs2\ \prefix xs2 xs'\ \xs = xs1 @ drop (length xs1) xs\ \xs1 \ xs2\ - assms(6) leI - by fastforce + using \prefix xs1 xs2\ \xs1 \ xs2\ prefix_length_prefix by fastforce + + have prefix_drop: "ys = ys1 @ (drop (length ys1)) ys" if "prefix ys1 ys" + for ys ys1 :: "('a \ 'b) list" + using that by (induction ys1) (auto elim: prefixE) + then have "xs = (xs1 @ (drop (length xs1) xs))" + using \prefix xs1 xs2\ \prefix xs2 xs'\ \prefix xs' xs\ by simp + then have "length xs1 < length xs" + using prefix_drop[OF \prefix xs2 xs'\] \prefix xs2 xs'\ \prefix xs' xs\ + using \length xs1 < length xs2\ + by (auto dest!: prefix_length_le) have "length (xs1@(drop (length xs2) xs)) < length xs" using \length xs1 < length xs2\ \length xs1 < length xs\ by auto have "vs \ L\<^sub>i\<^sub>n M1 V \ sequence_to_failure M1 M2 (vs @ xs1@(drop (length xs2) xs)) \ length (xs1@(drop (length xs2) xs)) < length xs" using \length (xs1 @ drop (length xs2) xs) < length xs\ \sequence_to_failure M1 M2 (vs @ xs1 @ drop (length xs2) xs)\ assms(1) minimal_sequence_to_failure_extending.simps by blast then have "\ minimal_sequence_to_failure_extending V M1 M2 vs xs" by (meson minimal_sequence_to_failure_extending.elims(2)) then show "False" using assms(1) by linarith qed lemma minimal_sequence_to_failure_extending_implies_Rep_Cov : assumes "minimal_sequence_to_failure_extending V M1 M2 vs xs" and "OFSM M1" and "OFSM M2" and "test_tools M2 M1 FAIL PM V \" and "V'' \ N (vs@xsR) M1 V" and "prefix xsR xs" shows "\ Rep_Cov M2 M1 V'' vs xsR" proof assume "Rep_Cov M2 M1 V'' vs xsR" then obtain xs' vs' s2 s1 where "xs' \ []" "prefix xs' xsR" "vs' \ V''" "io_targets M2 (initial M2) (vs @ xs') = {s2}" "io_targets M2 (initial M2) (vs') = {s2}" "io_targets M1 (initial M1) (vs @ xs') = {s1}" "io_targets M1 (initial M1) (vs') = {s1}" by auto then have "s2 \ io_targets M2 (initial M2) (vs @ xs')" "s2 \ io_targets M2 (initial M2) (vs')" "s1 \ io_targets M1 (initial M1) (vs @ xs')" "s1 \ io_targets M1 (initial M1) (vs')" by auto have "vs@xs' \ L M1" using io_target_implies_L[OF \s1 \ io_targets M1 (initial M1) (vs @ xs')\] by assumption have "vs' \ L M1" using io_target_implies_L[OF \s1 \ io_targets M1 (initial M1) (vs')\] by assumption have "vs@xs' \ L M2" using io_target_implies_L[OF \s2 \ io_targets M2 (initial M2) (vs @ xs')\] by assumption have "vs' \ L M2" using io_target_implies_L[OF \s2 \ io_targets M2 (initial M2) (vs')\] by assumption obtain tr1_1 where "path M1 (vs@xs' || tr1_1) (initial M1)" "length tr1_1 = length (vs@xs')" "target (vs@xs' || tr1_1) (initial M1) = s1" using \s1 \ io_targets M1 (initial M1) (vs @ xs')\ by auto obtain tr1_2 where "path M1 (vs' || tr1_2) (initial M1)" "length tr1_2 = length (vs')" "target (vs' || tr1_2) (initial M1) = s1" using \s1 \ io_targets M1 (initial M1) (vs')\ by auto obtain tr2_1 where "path M2 (vs@xs' || tr2_1) (initial M2)" "length tr2_1 = length (vs@xs')" "target (vs@xs' || tr2_1) (initial M2) = s2" using \s2 \ io_targets M2 (initial M2) (vs @ xs')\ by auto obtain tr2_2 where "path M2 (vs' || tr2_2) (initial M2)" "length tr2_2 = length (vs')" "target (vs' || tr2_2) (initial M2) = s2" using \s2 \ io_targets M2 (initial M2) (vs')\ by auto have "productF M2 M1 FAIL PM" using assms(4) by auto have "well_formed M1" using assms(2) by auto have "well_formed M2" using assms(3) by auto have "observable PM" by (meson assms(2) assms(3) assms(4) observable_productF) have "length (vs@xs') = length tr2_1" using \length tr2_1 = length (vs @ xs')\ by presburger then have "length tr2_1 = length tr1_1" using \length tr1_1 = length (vs@xs')\ by presburger have "vs@xs' \ L PM" using productF_path_inclusion[OF \length (vs@xs') = length tr2_1\ \length tr2_1 = length tr1_1\ \productF M2 M1 FAIL PM\ \well_formed M2\ \well_formed M1\] by (meson Int_iff \productF M2 M1 FAIL PM\ \vs @ xs' \ L M1\ \vs @ xs' \ L M2\ \well_formed M1\ \well_formed M2\ productF_language) have "length (vs') = length tr2_2" using \length tr2_2 = length (vs')\ by presburger then have "length tr2_2 = length tr1_2" using \length tr1_2 = length (vs')\ by presburger have "vs' \ L PM" using productF_path_inclusion[OF \length (vs') = length tr2_2\ \length tr2_2 = length tr1_2\ \productF M2 M1 FAIL PM\ \well_formed M2\ \well_formed M1\] by (meson Int_iff \productF M2 M1 FAIL PM\ \vs' \ L M1\ \vs' \ L M2\ \well_formed M1\ \well_formed M2\ productF_language) have "io_targets PM (initial M2, initial M1) (vs @ xs') = {(s2, s1)}" using productF_path_io_targets_reverse [OF \productF M2 M1 FAIL PM\ \s2 \ io_targets M2 (initial M2) (vs @ xs')\ \s1 \ io_targets M1 (initial M1) (vs @ xs')\ \vs @ xs' \ L M2\ \vs @ xs' \ L M1\ ] proof - have "\c f. c \ initial (f::('a, 'b, 'c) FSM) \ c \ nodes f" by blast then show ?thesis by (metis (no_types) \\observable M2; observable M1; well_formed M2; well_formed M1; initial M2 \ nodes M2; initial M1 \ nodes M1\ \ io_targets PM (initial M2, initial M1) (vs @ xs') = {(s2, s1)}\ assms(2) assms(3)) qed have "io_targets PM (initial M2, initial M1) (vs') = {(s2, s1)}" using productF_path_io_targets_reverse [OF \productF M2 M1 FAIL PM\ \s2 \ io_targets M2 (initial M2) (vs')\ \s1 \ io_targets M1 (initial M1) (vs')\ \vs' \ L M2\ \vs' \ L M1\ ] proof - have "\c f. c \ initial (f::('a, 'b, 'c) FSM) \ c \ nodes f" by blast then show ?thesis by (metis (no_types) \\observable M2; observable M1; well_formed M2; well_formed M1; initial M2 \ nodes M2; initial M1 \ nodes M1\ \ io_targets PM (initial M2, initial M1) (vs') = {(s2, s1)}\ assms(2) assms(3)) qed have "io_targets PM (initial PM) (vs') = {(s2, s1)}" by (metis (no_types) \io_targets PM (initial M2, initial M1) vs' = {(s2, s1)}\ \productF M2 M1 FAIL PM\ productF_simps(4)) have "sequence_to_failure M1 M2 (vs@xs)" using assms(1) by auto have "xs = xs' @ (drop (length xs') xs)" by (metis \prefix xs' xsR\ append_assoc append_eq_conv_conj assms(6) prefixE) then have "io_targets PM (initial M2, initial M1) (vs @ xs' @ (drop (length xs') xs)) = {FAIL}" by (metis \productF M2 M1 FAIL PM\ \sequence_to_failure M1 M2 (vs @ xs)\ assms(2) assms(3) productF_simps(4) sequence_to_failure_reaches_FAIL_ob) then have "io_targets PM (initial M2, initial M1) ((vs @ xs') @ (drop (length xs') xs)) = {FAIL}" by auto have "io_targets PM (s2, s1) (drop (length xs') xs) = {FAIL}" using observable_io_targets_split [OF \observable PM\ \io_targets PM (initial M2,initial M1) ((vs @ xs') @ (drop (length xs') xs)) = {FAIL}\ \io_targets PM (initial M2, initial M1) (vs @ xs') = {(s2, s1)}\] by assumption have "io_targets PM (initial PM) (vs' @ (drop (length xs') xs)) = {FAIL}" using observable_io_targets_append [OF \observable PM\ \io_targets PM (initial PM) (vs') = {(s2, s1)}\ \io_targets PM (s2, s1) (drop (length xs') xs) = {FAIL}\] by assumption have "sequence_to_failure M1 M2 (vs' @ (drop (length xs') xs))" using sequence_to_failure_alt_def [OF \io_targets PM (initial PM) (vs' @ (drop (length xs') xs)) = {FAIL}\ assms(2,3)] assms(4) by blast have "length (drop (length xs') xs) < length xs" by (metis (no_types) \xs = xs' @ drop (length xs') xs\ \xs' \ []\ length_append length_greater_0_conv less_add_same_cancel2) have "vs' \ L\<^sub>i\<^sub>n M1 V" proof - have "V'' \ Perm V M1" using assms(5) unfolding N.simps by blast then obtain f where f_def : "V'' = image f V \ (\ v \ V . f v \ language_state_for_input M1 (initial M1) v)" unfolding Perm.simps by blast then obtain v where "v \ V" "vs' = f v" using \vs' \ V''\ by auto then have "vs' \ language_state_for_input M1 (initial M1) v" using f_def by auto have "language_state_for_input M1 (initial M1) v = L\<^sub>i\<^sub>n M1 {v}" by auto moreover have "{v} \ V" using \v \ V\ by blast ultimately have "language_state_for_input M1 (initial M1) v \ L\<^sub>i\<^sub>n M1 V" unfolding language_state_for_inputs.simps language_state_for_input.simps by blast then show ?thesis using\vs' \ language_state_for_input M1 (initial M1) v\ by blast qed have "\ minimal_sequence_to_failure_extending V M1 M2 vs xs" using \vs' \ L\<^sub>i\<^sub>n M1 V\ \sequence_to_failure M1 M2 (vs' @ (drop (length xs') xs))\ \length (drop (length xs') xs) < length xs\ using minimal_sequence_to_failure_extending.elims(2) by blast then show "False" using assms(1) by linarith qed lemma mstfe_no_repetition : assumes "minimal_sequence_to_failure_extending V M1 M2 vs xs" and "OFSM M1" and "OFSM M2" and "test_tools M2 M1 FAIL PM V \" and "V'' \ N (vs@xs') M1 V" and "prefix xs' xs" shows "\ Rep_Pre M2 M1 vs xs'" and "\ Rep_Cov M2 M1 V'' vs xs'" using minimal_sequence_to_failure_extending_implies_Rep_Pre[OF assms] minimal_sequence_to_failure_extending_implies_Rep_Cov[OF assms] by linarith+ subsection \ Sufficiency of the test suite to test for reduction \ text \ The following lemma proves that set of input sequences generated in the final iteration of the @{verbatim TS} function constitutes a test suite sufficient to test for reduction the FSMs it has been generated for. This proof is performed by contradiction: If the test suite is not sufficient, then some minimal sequence to a failure extending the deterministic state cover must exist. Due to the test suite being assumed insufficient, this sequence cannot be contained in it and hence a prefix of it must have been contained in one of the sets calculated by the @{verbatim R} function. This is only possible if the prefix is not a minimal sequence to a failure extending the deterministic state cover or if the test suite observes a failure, both of which violates the assumptions. \ lemma asc_sufficiency : assumes "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "final_iteration M2 M1 \ V m i" shows "M1 \\(TS M2 M1 \ V m i) . \\ M2 \ M1 \ M2" proof assume "atc_io_reduction_on_sets M1 (TS M2 M1 \ V m i) \ M2" show "M1 \ M2" proof (rule ccontr) let ?TS = "\ n . TS M2 M1 \ V m n" let ?C = "\ n . C M2 M1 \ V m n" let ?RM = "\ n . RM M2 M1 \ V m n" assume "\ M1 \ M2" obtain vs xs where "minimal_sequence_to_failure_extending V M1 M2 vs xs" using assms(1) assms(2) assms(4) minimal_sequence_to_failure_extending_det_state_cover_ob[OF _ _ _ _ \\ M1 \ M2\, of V] by blast then have "vs \ L\<^sub>i\<^sub>n M1 V" "sequence_to_failure M1 M2 (vs @ xs)" "\ (\ io' . \ w' \ L\<^sub>i\<^sub>n M1 V . sequence_to_failure M1 M2 (w' @ io') \ length io' < length xs)" by auto then have "vs@xs \ L M1 - L M2" by auto have "vs@xs \ L\<^sub>i\<^sub>n M1 {map fst (vs@xs)}" by (metis (full_types) Diff_iff \vs @ xs \ L M1 - L M2\ insertI1 language_state_for_inputs_map_fst) have "vs@xs \ L\<^sub>i\<^sub>n M2 {map fst (vs@xs)}" by (meson Diff_iff \vs @ xs \ L M1 - L M2\ language_state_for_inputs_in_language_state subsetCE) have "finite V" using det_state_cover_finite assms(4,2) by auto then have "finite (?TS i)" using TS_finite[of V M2] assms(2) by auto then have "io_reduction_on M1 (?TS i) M2" using io_reduction_from_atc_io_reduction [OF \atc_io_reduction_on_sets M1 (TS M2 M1 \ V m i) \ M2\] by auto have "map fst (vs@xs) \ ?TS i" proof - have f1: "\ps P Pa. (ps::('a \ 'b) list) \ P - Pa \ ps \ P \ ps \ Pa" by blast have "\P Pa ps. \ P \ Pa \ (ps::('a \ 'b) list) \ Pa \ ps \ P" by blast then show ?thesis using f1 by (metis (no_types) \vs @ xs \ L M1 - L M2\ \io_reduction_on M1 (?TS i) M2\ language_state_for_inputs_in_language_state language_state_for_inputs_map_fst) qed have "map fst vs \ V" using \vs \ L\<^sub>i\<^sub>n M1 V\ by auto let ?stf = "map fst (vs@xs)" let ?stfV = "map fst vs" let ?stfX = "map fst xs" have "?stf = ?stfV @ ?stfX" by simp then have "?stfV @ ?stfX \ ?TS i" using \?stf \ ?TS i\ by auto have "mcp (?stfV @ ?stfX) V ?stfV" by (metis \map fst (vs @ xs) = map fst vs @ map fst xs\ \minimal_sequence_to_failure_extending V M1 M2 vs xs\ assms(1) assms(2) assms(4) minimal_sequence_to_failure_extending_mcp) have "set ?stf \ inputs M1" by (meson DiffD1 \vs @ xs \ L M1 - L M2\ assms(1) language_state_inputs) then have "set ?stf \ inputs M2" using assms(3) by blast moreover have "set ?stf = set ?stfV \ set ?stfX" by simp ultimately have "set ?stfX \ inputs M2" by blast obtain xr j where "xr \ ?stfX" "prefix xr ?stfX" "Suc j \ i" "?stfV@xr \ RM M2 M1 \ V m (Suc j)" using TS_non_containment_causes_final_suc[OF \?stfV @ ?stfX \ ?TS i\ \mcp (?stfV @ ?stfX) V ?stfV\ \set ?stfX \ inputs M2\ assms(5,2)] by blast let ?yr = "take (length xr) (map snd xs)" have "length ?yr = length xr" using \prefix xr (map fst xs)\ prefix_length_le by fastforce have "(xr || ?yr) = take (length xr) xs" by (metis (no_types, hide_lams) \prefix xr (map fst xs)\ append_eq_conv_conj prefixE take_zip zip_map_fst_snd) have "prefix (vs@(xr || ?yr)) (vs@xs)" by (simp add: \xr || take (length xr) (map snd xs) = take (length xr) xs\ take_is_prefix) have "xr = take (length xr) (map fst xs)" by (metis \length (take (length xr) (map snd xs)) = length xr\ \xr || take (length xr) (map snd xs) = take (length xr) xs\ map_fst_zip take_map) have "vs@(xr || ?yr) \ L M1" by (metis DiffD1 \prefix (vs @ (xr || take (length xr) (map snd xs))) (vs @ xs)\ \vs @ xs \ L M1 - L M2\ language_state_prefix prefixE) then have "vs@(xr || ?yr) \ L\<^sub>i\<^sub>n M1 {?stfV @ xr}" by (metis \length (take (length xr) (map snd xs)) = length xr\ insertI1 language_state_for_inputs_map_fst map_append map_fst_zip) have "length xr < length xs" by (metis \xr = take (length xr) (map fst xs)\ \xr \ map fst xs\ not_le_imp_less take_all take_map) from \?stfV@xr \ RM M2 M1 \ V m (Suc j)\ have "?stfV@xr \ {xs' \ C M2 M1 \ V m (Suc j) . (\ (L\<^sub>i\<^sub>n M1 {xs'} \ L\<^sub>i\<^sub>n M2 {xs'})) \ (\ io \ L\<^sub>i\<^sub>n M1 {xs'} . (\ V'' \ N io M1 V . (\ S1 . (\ vs xs . io = (vs@xs) \ mcp (vs@xs) V'' vs \ S1 \ nodes M2 \ (\ s1 \ S1 . \ s2 \ S1 . s1 \ s2 \ (\ io1 \ RP M2 s1 vs xs V'' . \ io2 \ RP M2 s2 vs xs V'' . B M1 io1 \ \ B M1 io2 \ )) \ m < LB M2 M1 vs xs (TS M2 M1 \ V m j \ V) S1 \ V'' ))))}" unfolding RM.simps by blast moreover have "\ xs' \ ?C (Suc j) . L\<^sub>i\<^sub>n M1 {xs'} \ L\<^sub>i\<^sub>n M2 {xs'}" proof fix xs' assume "xs' \ ?C (Suc j)" from \Suc j \ i\ have "?C (Suc j) \ ?TS i" using C_subset TS_subset by blast then have "{xs'} \ ?TS i" using \xs' \ ?C (Suc j)\ by blast show "L\<^sub>i\<^sub>n M1 {xs'} \ L\<^sub>i\<^sub>n M2 {xs'}" using io_reduction_on_subset[OF \io_reduction_on M1 (?TS i) M2\ \{xs'} \ ?TS i\] by assumption qed ultimately have "(\ io \ L\<^sub>i\<^sub>n M1 {?stfV@xr} . (\ V'' \ N io M1 V . (\ S1 . (\ vs xs . io = (vs@xs) \ mcp (vs@xs) V'' vs \ S1 \ nodes M2 \ (\ s1 \ S1 . \ s2 \ S1 . s1 \ s2 \ (\ io1 \ RP M2 s1 vs xs V'' . \ io2 \ RP M2 s2 vs xs V'' . B M1 io1 \ \ B M1 io2 \ )) \ m < LB M2 M1 vs xs (TS M2 M1 \ V m j \ V) S1 \ V'' ))))" by blast then have " (\ V'' \ N (vs@(xr || ?yr)) M1 V . (\ S1 . (\ vs' xs' . vs@(xr || ?yr) = (vs'@xs') \ mcp (vs'@xs') V'' vs' \ S1 \ nodes M2 \ (\ s1 \ S1 . \ s2 \ S1 . s1 \ s2 \ (\ io1 \ RP M2 s1 vs' xs' V'' . \ io2 \ RP M2 s2 vs' xs' V'' . B M1 io1 \ \ B M1 io2 \ )) \ m < LB M2 M1 vs' xs' (TS M2 M1 \ V m j \ V) S1 \ V'' )))" using \vs@(xr || ?yr) \ L\<^sub>i\<^sub>n M1 {?stfV @ xr}\ by blast then obtain V'' S1 vs' xs' where RM_impl : "V'' \ N (vs@(xr || ?yr)) M1 V" "vs@(xr || ?yr) = (vs'@xs')" "mcp (vs'@xs') V'' vs'" "S1 \ nodes M2" "(\ s1 \ S1 . \ s2 \ S1 . s1 \ s2 \ (\ io1 \ RP M2 s1 vs' xs' V'' . \ io2 \ RP M2 s2 vs' xs' V'' . B M1 io1 \ \ B M1 io2 \ ))" " m < LB M2 M1 vs' xs' (TS M2 M1 \ V m j \ V) S1 \ V''" by blast have "?stfV = mcp' (map fst (vs @ (xr || take (length xr) (map snd xs)))) V" by (metis (full_types) \length (take (length xr) (map snd xs)) = length xr\ \mcp (map fst vs @ map fst xs) V (map fst vs)\ \prefix xr (map fst xs)\ map_append map_fst_zip mcp'_intro mcp_prefix_of_suffix) have "is_det_state_cover M2 V" using assms(4) by blast moreover have "well_formed M2" using assms(2) by auto moreover have "finite V" using det_state_cover_finite assms(4,2) by auto ultimately have "vs \ V''" "vs = mcp' (vs @ (xr || take (length xr) (map snd xs))) V''" using N_mcp_prefix[OF \?stfV = mcp' (map fst (vs @ (xr || take (length xr) (map snd xs)))) V\ \V'' \ N (vs@(xr || ?yr)) M1 V\, of M2] by simp+ have "vs' = vs" by (metis (no_types) \mcp (vs' @ xs') V'' vs'\ \vs = mcp' (vs @ (xr || take (length xr) (map snd xs))) V''\ \vs @ (xr || take (length xr) (map snd xs)) = vs' @ xs'\ mcp'_intro) then have "xs' = (xr || ?yr)" using \vs @ (xr || take (length xr) (map snd xs)) = vs' @ xs'\ by blast have "V \ ?TS i" proof - have "1 \ i" using \Suc j \ i\ by linarith then have "?TS 1 \ ?TS i" using TS_subset by blast then show ?thesis by auto qed have "?stfV@xr \ ?C (Suc j)" using \?stfV@xr \ RM M2 M1 \ V m (Suc j)\ unfolding RM.simps by blast \ \show that the prerequisites (@{verbatim Prereq}) for @{verbatim LB} are met by construction\ have "(\vs'a\V''. prefix vs'a (vs' @ xs') \ length vs'a \ length vs')" using \mcp (vs' @ xs') V'' vs'\ by auto moreover have "atc_io_reduction_on_sets M1 (?TS j \ V) \ M2" proof - have "j < i" using \Suc j \ i\ by auto then have "?TS j \ ?TS i" by (simp add: TS_subset) then show ?thesis using atc_io_reduction_on_subset [OF \atc_io_reduction_on_sets M1 (TS M2 M1 \ V m i) \ M2\, of "?TS j"] by (meson Un_subset_iff \V \ ?TS i\ \atc_io_reduction_on_sets M1 (TS M2 M1 \ V m i) \ M2\ atc_io_reduction_on_subset) qed moreover have "finite (?TS j \ V)" proof - have "finite (?TS j)" using TS_finite[OF \finite V\, of M2 M1 \ m j] assms(2) by auto then show ?thesis using \finite V\ by blast qed moreover have "V \ ?TS j \ V" by blast moreover have "(\ p . (prefix p xs' \ p \ xs') \ map fst (vs' @ p) \ ?TS j \ V)" proof fix p show "prefix p xs' \ p \ xs' \ map fst (vs' @ p) \ TS M2 M1 \ V m j \ V" proof assume "prefix p xs' \ p \ xs'" have "prefix (map fst (vs' @ p)) (map fst (vs' @ xs'))" by (simp add: \prefix p xs' \ p \ xs'\ map_mono_prefix) have "prefix (map fst (vs' @ p)) (?stfV @ xr)" using \length (take (length xr) (map snd xs)) = length xr\ \prefix (map fst (vs' @ p)) (map fst (vs' @ xs'))\ \vs' = vs\ \xs' = xr || take (length xr) (map snd xs)\ by auto then have "prefix (map fst vs' @ map fst p) (?stfV @ xr)" by simp then have "prefix (map fst p) xr" by (simp add: \vs' = vs\) have "?stfV @ xr \ ?TS (Suc j)" proof (cases j) case 0 then show ?thesis using \map fst vs @ xr \ C M2 M1 \ V m (Suc j)\ by auto next case (Suc nat) then show ?thesis using TS.simps(3) \map fst vs @ xr \ C M2 M1 \ V m (Suc j)\ by blast qed have "mcp (map fst vs @ xr) V (map fst vs)" using \mcp (map fst vs @ map fst xs) V (map fst vs)\ \prefix xr (map fst xs)\ mcp_prefix_of_suffix by blast have "map fst vs @ map fst p \ TS M2 M1 \ V m (Suc j)" using TS_prefix_containment[OF \?stfV @ xr \ ?TS (Suc j)\ \mcp (map fst vs @ xr) V (map fst vs)\ \prefix (map fst p) xr\] by assumption have "Suc (length xr) = (Suc j)" using C_index[OF \?stfV@xr \ ?C (Suc j)\ \mcp (map fst vs @ xr) V (map fst vs)\] by assumption have"Suc (length p) < (Suc j)" proof - have "map fst xs' = xr" by (metis \xr = take (length xr) (map fst xs)\ \xr || take (length xr) (map snd xs) = take (length xr) xs\ \xs' = xr || take (length xr) (map snd xs)\ take_map) then show ?thesis by (metis (no_types) Suc_less_eq \Suc (length xr) = Suc j\ \prefix p xs' \ p \ xs'\ append_eq_conv_conj length_map nat_less_le prefixE prefix_length_le take_all) qed have "mcp (map fst vs @ map fst p) V (map fst vs)" using \mcp (map fst vs @ xr) V (map fst vs)\ \prefix (map fst p) xr\ mcp_prefix_of_suffix by blast then have "map fst vs @ map fst p \ ?C (Suc (length (map fst p)))" using TS_index(2)[OF \map fst vs @ map fst p \ TS M2 M1 \ V m (Suc j)\] by auto have "map fst vs @ map fst p \ ?TS j" using TS_union[of M2 M1 \ V m j] proof - have "Suc (length p) \ {0..Suc (length p) < Suc j\ by force then show ?thesis by (metis UN_I \TS M2 M1 \ V m j = (\j\set [0.. V m j)\ \map fst vs @ map fst p \ C M2 M1 \ V m (Suc (length (map fst p)))\ length_map set_upt) qed then show "map fst (vs' @ p) \ TS M2 M1 \ V m j \ V" by (simp add: \vs' = vs\) qed qed moreover have "vs' @ xs' \ L M2 \ L M1" by (metis (no_types, lifting) IntI RM_impl(2) \\xs'\C M2 M1 \ V m (Suc j). L\<^sub>i\<^sub>n M1 {xs'} \ L\<^sub>i\<^sub>n M2 {xs'}\ \map fst vs @ xr \ C M2 M1 \ V m (Suc j)\ \vs @ (xr || take (length xr) (map snd xs)) \ L\<^sub>i\<^sub>n M1 {map fst vs @ xr}\ language_state_for_inputs_in_language_state subsetCE) ultimately have "Prereq M2 M1 vs' xs' (?TS j \ V) S1 \ V''" using RM_impl(4,5) unfolding Prereq.simps by blast have "V'' \ Perm V M1" using \V'' \ N (vs@(xr || ?yr)) M1 V\ unfolding N.simps by blast have \prefix (xr || ?yr) xs\ by (simp add: \xr || take (length xr) (map snd xs) = take (length xr) xs\ take_is_prefix) \ \ show that furthermore neither @{verbatim Rep_Pre} nor @{verbatim Rep_Cov} holds \ have "\ Rep_Pre M2 M1 vs (xr || ?yr)" using minimal_sequence_to_failure_extending_implies_Rep_Pre [OF \minimal_sequence_to_failure_extending V M1 M2 vs xs\ assms(1,2) \test_tools M2 M1 FAIL PM V \\ RM_impl(1) \prefix (xr || take (length xr) (map snd xs)) xs\] by assumption then have "\ Rep_Pre M2 M1 vs' xs'" using \vs' = vs\ \xs' = xr || ?yr\ by blast have "\ Rep_Cov M2 M1 V'' vs (xr || ?yr)" using minimal_sequence_to_failure_extending_implies_Rep_Cov [OF \minimal_sequence_to_failure_extending V M1 M2 vs xs\ assms(1,2) \test_tools M2 M1 FAIL PM V \\ RM_impl(1) \prefix (xr || take (length xr) (map snd xs)) xs\] by assumption then have "\ Rep_Cov M2 M1 V'' vs' xs'" using \vs' = vs\ \xs' = xr || ?yr\ by blast have "vs'@xs' \ L M1" using \vs @ (xr || take (length xr) (map snd xs)) \ L M1\ \vs' = vs\ \xs' = xr || take (length xr) (map snd xs)\ by blast \ \ therefore it is impossible to remove the prefix of the minimal sequence to a failure, as this would require @{verbatim M1} to have more than m states \ have "LB M2 M1 vs' xs' (?TS j \ V) S1 \ V'' \ card (nodes M1)" using LB_count[OF \vs'@xs' \ L M1\ assms(1,2,3) \test_tools M2 M1 FAIL PM V \\ \V'' \ Perm V M1\ \Prereq M2 M1 vs' xs' (?TS j \ V) S1 \ V''\ \\ Rep_Pre M2 M1 vs' xs'\ \ \ Rep_Cov M2 M1 V'' vs' xs'\] by assumption then have "LB M2 M1 vs' xs' (?TS j \ V) S1 \ V'' \ m" using assms(3) by linarith then show "False" using \m < LB M2 M1 vs' xs' (?TS j \ V) S1 \ V''\ by linarith qed qed subsection \ Main result \ text \ The following lemmata add to the previous result to show that some FSM @{verbatim M1} is a reduction of FSM @{verbatim M2} if and only if it is a reduction on the test suite generated by the adaptive state counting algorithm for these FSMs. \ lemma asc_soundness : assumes "OFSM M1" and "OFSM M2" shows "M1 \ M2 \ atc_io_reduction_on_sets M1 T \ M2" using atc_io_reduction_on_sets_reduction assms by blast lemma asc_main_theorem : assumes "OFSM M1" and "OFSM M2" and "asc_fault_domain M2 M1 m" and "test_tools M2 M1 FAIL PM V \" and "final_iteration M2 M1 \ V m i" shows "M1 \ M2 \ atc_io_reduction_on_sets M1 (TS M2 M1 \ V m i) \ M2" by (metis asc_sufficiency assms(1-5) atc_io_reduction_on_sets_reduction) end \ No newline at end of file diff --git a/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy b/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy --- a/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy +++ b/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy @@ -1,939 +1,925 @@ section \CCW for Arbitrary Points in the Plane\ theory Counterclockwise_2D_Arbitrary imports Counterclockwise_2D_Strict begin subsection \Interpretation of Knuth's axioms in the plane\ definition lex::"point \ point \ bool" where "lex p q \ (fst p < fst q \ fst p = fst q \ snd p < snd q \ p = q)" definition psi::"point \ point \ point \ bool" where "psi p q r \ (lex p q \ lex q r)" definition ccw::"point \ point \ point \ bool" where "ccw p q r \ ccw' p q r \ (det3 p q r = 0 \ (psi p q r \ psi q r p \ psi r p q))" interpretation ccw: linorder_list0 "ccw x" for x . lemma ccw'_imp_ccw: "ccw' a b c \ ccw a b c" by (simp add: ccw_def) lemma ccw_ncoll_imp_ccw: "ccw a b c \ \coll a b c \ ccw' a b c" by (simp add: ccw_def) lemma ccw_translate: "ccw p (p + q) (p + r) = ccw 0 q r" by (auto simp: ccw_def psi_def lex_def) lemma ccw_translate_origin: "NO_MATCH 0 p \ ccw p q r = ccw 0 (q - p) (r - p)" using ccw_translate[of p "q - p" "r - p"] by simp lemma psi_scale: "psi (r *\<^sub>R a) (r *\<^sub>R b) 0 = (if r > 0 then psi a b 0 else if r < 0 then psi 0 b a else True)" "psi (r *\<^sub>R a) 0 (r *\<^sub>R b) = (if r > 0 then psi a 0 b else if r < 0 then psi b 0 a else True)" "psi 0 (r *\<^sub>R a) (r *\<^sub>R b) = (if r > 0 then psi 0 a b else if r < 0 then psi b a 0 else True)" by (auto simp: psi_def lex_def det3_def' not_less algebra_split_simps) lemma ccw_scale23: "ccw 0 a b \ r > 0 \ ccw 0 (r *\<^sub>R a) (r *\<^sub>R b)" by (auto simp: ccw_def psi_scale) lemma psi_notI: "distinct3 p q r \ psi p q r \ \ psi q p r" by (auto simp: algebra_simps psi_def lex_def) -lemma not_lex_eq: "\ lex a b \ lex b a \ a \ b" +lemma not_lex_eq: "\ lex a b \ lex b a \ b \ a" by (auto simp: algebra_simps lex_def prod_eq_iff) lemma lex_trans: "lex a b \ lex b c \ lex a c" by (auto simp: lex_def) lemma lex_sym_eqI: "lex a b \ lex b a \ a = b" and lex_sym_eq_iff: "lex a b \ lex b a \ a = b" by (auto simp: lex_def) lemma lex_refl[simp]: "lex p p" by (metis not_lex_eq) lemma psi_disjuncts: "distinct3 p q r \ psi p q r \ psi p r q \ psi q r p \ psi q p r \ psi r p q \ psi r q p" by (auto simp: psi_def not_lex_eq) lemma nlex_ccw_left: "lex x 0 \ ccw 0 (0, 1) x" by (auto simp: ccw_def lex_def psi_def ccw'_def det3_def') interpretation ccw_system123 ccw apply unfold_locales subgoal by (force simp: ccw_def ccw'_def det3_def' algebra_simps) subgoal by (force simp: ccw_def ccw'_def det3_def' psi_def algebra_simps lex_sym_eq_iff) subgoal by (drule psi_disjuncts) (force simp: ccw_def ccw'_def det3_def' algebra_simps) done lemma lex_scaleR_nonneg: "lex a b \ r \ 0 \ lex a (a + r *\<^sub>R (b - a))" by (auto simp: lex_def) lemma lex_scale1_zero: "lex (v *\<^sub>R u) 0 = (if v > 0 then lex u 0 else if v < 0 then lex 0 u else True)" and lex_scale2_zero: "lex 0 (v *\<^sub>R u) = (if v > 0 then lex 0 u else if v < 0 then lex u 0 else True)" by (auto simp: lex_def prod_eq_iff less_eq_prod_def algebra_split_simps) lemma nlex_add: assumes "lex a 0" "lex b 0" shows "lex (a + b) 0" using assms by (auto simp: lex_def) lemma nlex_sum: assumes "finite X" assumes "\x. x \ X \ lex (f x) 0" shows "lex (sum f X) 0" using assms by induction (auto intro!: nlex_add) lemma abs_add_nlex: assumes "coll 0 a b" assumes "lex a 0" assumes "lex b 0" shows "abs (a + b) = abs a + abs b" proof (rule antisym[OF abs_triangle_ineq]) have "fst (\a\ + \b\) \ fst \a + b\" using assms by (auto simp add: det3_def' abs_prod_def lex_def) moreover { assume H: "fst a < 0" "fst b < 0" hence "snd b \ 0 \ snd a \ 0" using assms by (auto simp: lex_def det3_def' mult.commute) (metis mult_le_cancel_left_neg mult_zero_right)+ hence "\snd a\ + \snd b\ \ \snd a + snd b\" using H by auto } hence "snd (\a\ + \b\) \ snd \a + b\" using assms by (auto simp add: det3_def' abs_prod_def lex_def) ultimately show "\a\ + \b\ \ \a + b\" unfolding less_eq_prod_def .. qed lemma lex_sum_list: "(\x. x \ set xs \ lex x 0) \ lex (sum_list xs) 0" by (induct xs) (auto simp: nlex_add) lemma abs_sum_list_coll: assumes coll: "list_all (coll 0 x) xs" assumes "x \ 0" assumes up: "list_all (\x. lex x 0) xs" shows "abs (sum_list xs) = sum_list (map abs xs)" using assms proof (induct xs) case (Cons y ys) hence "coll 0 x y" "coll 0 x (sum_list ys)" by (auto simp: list_all_iff intro!: coll_sum_list) hence "coll 0 y (sum_list ys)" using \x \ 0\ by (rule coll_trans) hence "\y + sum_list ys\ = abs y + abs (sum_list ys)" using Cons by (subst abs_add_nlex) (auto simp: list_all_iff lex_sum_list) thus ?case using Cons by simp qed simp lemma lex_diff1: "lex (a - b) c = lex a (c + b)" and lex_diff2: "lex c (a - b) = lex (c + b) a" by (auto simp: lex_def) lemma sum_list_eq_0_iff_nonpos: fixes xs::"'a::ordered_ab_group_add list" shows "list_all (\x. x \ 0) xs \ sum_list xs = 0 \ (\n\set xs. n = 0)" by (auto simp: list_all_iff sum_list_sum_nth sum_nonpos_eq_0_iff) (auto simp add: in_set_conv_nth) lemma sum_list_nlex_eq_zeroI: assumes nlex: "list_all (\x. lex x 0) xs" assumes "sum_list xs = 0" assumes "x \ set xs" shows "x = 0" proof - from assms(2) have z1: "sum_list (map fst xs) = 0" and z2: "sum_list (map snd xs) = 0" by (auto simp: prod_eq_iff fst_sum_list snd_sum_list) from nlex have "list_all (\x. x \ 0) (map fst xs)" by (auto simp: lex_def list_all_iff) from sum_list_eq_0_iff_nonpos[OF this] z1 nlex have z1': "list_all (\x. x = 0) (map fst xs)" and "list_all (\x. x \ 0) (map snd xs)" by (auto simp: list_all_iff lex_def) from sum_list_eq_0_iff_nonpos[OF this(2)] z2 have "list_all (\x. x = 0) (map snd xs)" by (simp add: list_all_iff) with z1' show "x = 0" by (auto simp: list_all_iff zero_prod_def assms prod_eq_iff) qed lemma sum_list_eq0I: "(\x\set xs. x = 0) \ sum_list xs = 0" by (induct xs) auto lemma sum_list_nlex_eq_zero_iff: assumes nlex: "list_all (\x. lex x 0) xs" shows "sum_list xs = 0 \ list_all ((=) 0) xs" using assms by (auto intro: sum_list_nlex_eq_zeroI sum_list_eq0I simp: list_all_iff) lemma assumes "lex p q" "lex q r" "0 \ a" "0 \ b" "0 \ c" "a + b + c = 1" assumes comb_def: "comb = a *\<^sub>R p + b *\<^sub>R q + c *\<^sub>R r" shows lex_convex3: "lex p comb" "lex comb r" proof - from convex3_alt[OF assms(3-6), of p q r] obtain u v where uv: "a *\<^sub>R p + b *\<^sub>R q + c *\<^sub>R r = p + u *\<^sub>R (q - p) + v *\<^sub>R (r - p)" "0 \ u" "0 \ v" "u + v \ 1" . have "lex p r" using assms by (metis lex_trans) hence "lex (v *\<^sub>R (p - r)) 0" using uv by (simp add: lex_scale1_zero lex_diff1) also have "lex 0 (u *\<^sub>R (q - p))" using \lex p q\ uv by (simp add: lex_scale2_zero lex_diff2) finally (lex_trans) show "lex p comb" unfolding comb_def uv by (simp add: lex_def prod_eq_iff algebra_simps) from comb_def have comb_def': "comb = c *\<^sub>R r + b *\<^sub>R q + a *\<^sub>R p" by simp from assms have "c + b + a = 1" by simp from convex3_alt[OF assms(5,4,3) this, of r q p] obtain u v where uv: "c *\<^sub>R r + b *\<^sub>R q + a *\<^sub>R p = r + u *\<^sub>R (q - r) + v *\<^sub>R (p - r)" "0 \ u" "0 \ v" "u + v \ 1" by auto have "lex (u *\<^sub>R (q - r)) 0" using uv \lex q r\ by (simp add: lex_scale1_zero lex_diff1) also have "lex 0 (v *\<^sub>R (r - p))" using uv \lex p r\ by (simp add: lex_scale2_zero lex_diff2) finally (lex_trans) show "lex comb r" unfolding comb_def' uv by (simp add: lex_def prod_eq_iff algebra_simps) qed lemma lex_convex_self2: assumes "lex p q" "0 \ a" "a \ 1" defines "r \ a *\<^sub>R p + (1 - a) *\<^sub>R q" shows "lex p r" (is ?th1) and "lex r q" (is ?th2) using lex_convex3[OF \lex p q\, of q a "1 - a" 0 r] assms by (simp_all add: r_def) lemma lex_uminus0[simp]: "lex (-a) 0 = lex 0 a" by (auto simp: lex_def) lemma lex_fst_zero_imp: "fst x = 0 \ lex x 0 \ lex y 0 \ \coll 0 x y \ ccw' 0 y x" by (auto simp: ccw'_def det3_def' lex_def algebra_split_simps) lemma lex_ccw_left: "lex x y \ r > 0 \ ccw y (y + (0, r)) x" by (auto simp: ccw_def ccw'_def det3_def' algebra_simps lex_def psi_def) lemma lex_translate_origin: "NO_MATCH 0 a \ lex a b = lex 0 (b - a)" by (auto simp: lex_def) subsection \Order prover setup\ definition "lexs p q \ (lex p q \ p \ q)" lemma lexs_irrefl: "\ lexs p p" and lexs_imp_lex: "lexs x y \ lex x y" and not_lexs: "(\ lexs x y) = (lex y x)" and not_lex: "(\ lex x y) = (lexs y x)" and eq_lex_refl: "x = y \ lex x y" by (auto simp: lexs_def lex_def prod_eq_iff) lemma lexs_trans: "lexs x y \ lexs y z \ lexs x z" and lexs_lex_trans: "lexs x y \ lex y z \ lexs x z" and lex_lexs_trans: "lex x y \ lexs y z \ lexs x z" and lex_neq_trans: "lex a b \ a \ b \ lexs a b" and neq_lex_trans: "a \ b \ lex a b \ lexs a b" and lexs_imp_neq: "lexs a b \ a \ b" by (auto simp: lexs_def lex_def prod_eq_iff) -declare - lexs_irrefl[THEN notE, order add less_reflE: linorder "(=) :: point => point => bool" lex lexs] -declare lex_refl[order add le_refl: linorder "(=) :: point => point => bool" lex lexs] -declare lexs_imp_lex[order add less_imp_le: linorder "(=) :: point => point => bool" lex lexs] -declare - not_lexs[THEN iffD2, order add not_lessI: linorder "(=) :: point => point => bool" lex lexs] -declare not_lex[THEN iffD2, order add not_leI: linorder "(=) :: point => point => bool" lex lexs] -declare - not_lexs[THEN iffD1, order add not_lessD: linorder "(=) :: point => point => bool" lex lexs] -declare not_lex[THEN iffD1, order add not_leD: linorder "(=) :: point => point => bool" lex lexs] -declare lex_sym_eqI[order add eqI: linorder "(=) :: point => point => bool" lex lexs] -declare eq_lex_refl[order add eqD1: linorder "(=) :: point => point => bool" lex lexs] -declare sym[THEN eq_lex_refl, order add eqD2: linorder "(=) :: point => point => bool" lex lexs] -declare lexs_trans[order add less_trans: linorder "(=) :: point => point => bool" lex lexs] -declare lexs_lex_trans[order add less_le_trans: linorder "(=) :: point => point => bool" lex lexs] -declare lex_lexs_trans[order add le_less_trans: linorder "(=) :: point => point => bool" lex lexs] -declare lex_trans[order add le_trans: linorder "(=) :: point => point => bool" lex lexs] -declare lex_neq_trans[order add le_neq_trans: linorder "(=) :: point => point => bool" lex lexs] -declare neq_lex_trans[order add neq_le_trans: linorder "(=) :: point => point => bool" lex lexs] -declare lexs_imp_neq[order add less_imp_neq: linorder "(=) :: point => point => bool" lex lexs] -declare - eq_neq_eq_imp_neq[order add eq_neq_eq_imp_neq: linorder "(=) :: point => point => bool" lex lexs] -declare not_sym[order add not_sym: linorder "(=) :: point => point => bool" lex lexs] - +local_setup \ + HOL_Order_Tac.declare_linorder { + ops = {eq = @{term \(=) :: point \ point \ bool\}, le = @{term \lex\}, lt = @{term \lexs\}}, + thms = {trans = @{thm lex_trans}, refl = @{thm lex_refl}, eqD1 = @{thm eq_lex_refl}, + eqD2 = @{thm eq_lex_refl[OF sym]}, antisym = @{thm lex_sym_eqI}, contr = @{thm notE}}, + conv_thms = {less_le = @{thm eq_reflection[OF lexs_def]}, + nless_le = @{thm eq_reflection[OF not_lexs]}, + nle_le = @{thm eq_reflection[OF not_lex_eq]}} + } +\ subsection \Contradictions\ lemma assumes d: "distinct4 s p q r" shows contra1: "\(lex p q \ lex q r \ lex r s \ indelta s p q r)" (is ?th1) and contra2: "\(lex s p \ lex p q \ lex q r \ indelta s p q r)" (is ?th2) and contra3: "\(lex p r \ lex p s \ lex q r \ lex q s \ insquare p r q s)" (is ?th3) proof - { assume "det3 s p q = 0" "det3 s q r = 0" "det3 s r p = 0" "det3 p q r = 0" hence ?th1 ?th2 ?th3 using d by (auto simp add: det3_def' ccw'_def ccw_def psi_def algebra_simps) } moreover { assume *: "\(det3 s p q = 0 \ det3 s q r = 0 \ det3 s r p = 0 \ det3 p q r = 0)" { assume d0: "det3 p q r = 0" with d have "?th1 \ ?th2" by (force simp add: det3_def' ccw'_def ccw_def psi_def algebra_simps) } moreover { assume dp: "det3 p q r \ 0" have "?th1 \ ?th2" unfolding de_Morgan_disj[symmetric] proof (rule notI, goal_cases) case prems: 1 hence **: "indelta s p q r" by auto hence nonnegs: "det3 p q r \ 0" "0 \ det3 s q r" "0 \ det3 p s r" "0 \ det3 p q s" by (auto simp: ccw_def ccw'_def det3_def' algebra_simps) hence det_pos: "det3 p q r > 0" using dp by simp have det_eq: "det3 s q r + det3 p s r + det3 p q s = det3 p q r" by (auto simp: ccw_def det3_def' algebra_simps) hence det_div_eq: "det3 s q r / det3 p q r + det3 p s r / det3 p q r + det3 p q s / det3 p q r = 1" using det_pos by (auto simp: field_simps) from lex_convex3[OF _ _ _ _ _ det_div_eq convex_comb_dets[OF det_pos, of s]] have "lex p s" "lex s r" using prems by (auto simp: nonnegs) with prems d show False by (simp add: lex_sym_eq_iff) qed } moreover have ?th3 proof (safe, goal_cases) case prems: 1 have nonnegs: "det3 p r q \ 0" "det3 r q s \ 0" "det3 s p r \ 0" "det3 q s p \ 0" using prems by (auto simp add: ccw_def ccw'_def less_eq_real_def) have dets_eq: "det3 p r q + det3 q s p = det3 r q s + det3 s p r" by (auto simp: det3_def') hence **: "det3 p r q = 0 \ det3 q s p = 0 \ det3 r q s = 0 \ det3 s p r = 0" using prems by (auto simp: ccw_def ccw'_def) moreover { fix p r q s assume det_pos: "det3 p r q > 0" assume dets_eq: "det3 p r q + det3 q s p = det3 r q s + det3 s p r" assume nonnegs:"det3 r q s \ 0" "det3 s p r \ 0" "det3 q s p \ 0" assume g14: "lex p r" "lex p s" "lex q r" "lex q s" assume d: "distinct4 s p q r" let ?sum = "(det3 p r q + det3 q s p) / det3 p r q" have eqs: "det3 s p r = det3 p r s" "det3 r q s = det3 s r q" "det3 q s p = - det3 p s q" by (auto simp: det3_def' algebra_simps) from convex_comb_dets[OF det_pos, of s] have "((det3 p r q / det3 p r q) *\<^sub>R s + (det3 q s p / det3 p r q) *\<^sub>R r) /\<^sub>R ?sum = ((det3 r q s / det3 p r q) *\<^sub>R p + (det3 s p r / det3 p r q) *\<^sub>R q) /\<^sub>R ?sum" unfolding eqs by (simp add: algebra_simps prod_eq_iff) hence srpq: "(det3 p r q / det3 p r q / ?sum) *\<^sub>R s + (det3 q s p / det3 p r q / ?sum) *\<^sub>R r = (det3 r q s / det3 p r q / ?sum) *\<^sub>R p + (det3 s p r / det3 p r q / ?sum) *\<^sub>R q" (is "?s *\<^sub>R s + ?r *\<^sub>R r = ?p *\<^sub>R p + ?q *\<^sub>R q") using det_pos by (simp add: algebra_simps inverse_eq_divide) have eqs: "?s + ?r = 1" "?p + ?q = 1" and s: "?s \ 0" "?s \ 1" and r: "?r \ 0" "?r \ 1" and p: "?p \ 0" "?p \ 1" and q: "?q \ 0" "?q \ 1" unfolding add_divide_distrib[symmetric] using det_pos nonnegs dets_eq by (auto) from eqs have eqs': "1 - ?s = ?r" "1 - ?r = ?s" "1 - ?p = ?q" "1 - ?q = ?p" by auto have comm: "?r *\<^sub>R r + ?s *\<^sub>R s = ?s *\<^sub>R s + ?r *\<^sub>R r" "?q *\<^sub>R q + ?p *\<^sub>R p = ?p *\<^sub>R p + ?q *\<^sub>R q" by simp_all define K where "K = (det3 r q s / det3 p r q / ?sum) *\<^sub>R p + (det3 s p r / det3 p r q / ?sum) *\<^sub>R q" note rewrs = eqs' comm srpq K_def[symmetric] from lex_convex_self2[OF _ s, of s r, unfolded rewrs] lex_convex_self2[OF _ r, of r s, unfolded rewrs] lex_convex_self2[OF _ p, of p q, unfolded rewrs] lex_convex_self2[OF _ q, of q p, unfolded rewrs] have False using g14 d det_pos by (metis lex_trans not_lex_eq) } note wlog = this from dets_eq have 1: "det3 q s p + det3 p r q = det3 s p r + det3 r q s" by simp from d have d': "distinct4 r q p s" by auto note wlog[of q s p r, OF _ 1 nonnegs(3,2,1) prems(4,3,2,1) d'] wlog[of p r q s, OF _ dets_eq nonnegs(2,3,4) prems(1-4) d] ultimately show False using nonnegs d * by (auto simp: less_eq_real_def det3_def' algebra_simps) qed ultimately have ?th1 ?th2 ?th3 by blast+ } ultimately show ?th1 ?th2 ?th3 by force+ qed lemma ccw'_subst_psi_disj: assumes "det3 t r s = 0" assumes "psi t r s \ psi t s r \ psi s r t" assumes "s \ t" assumes "ccw' t r p" shows "ccw' t s p" proof cases assume "r \ s" from assms have "r \ t" by (auto simp: det3_def' ccw'_def algebra_simps) from assms have "det3 r s t = 0" by (auto simp: algebra_simps det3_def') from coll_ex_scaling[OF assms(3) this] obtain x where s: "r = s + x *\<^sub>R (t - s)" by auto from assms(4)[simplified s] have "0 < det3 0 (s + x *\<^sub>R (t - s) - t) (p - t)" by (auto simp: algebra_simps det3_def' ccw'_def) also have "s + x *\<^sub>R (t - s) - t = (1 - x) *\<^sub>R (s - t)" by (simp add: algebra_simps) finally have ccw': "ccw' 0 ((1 - x) *\<^sub>R (s - t)) (p - t)" by (simp add: ccw'_def) hence neq: "x \ 1" by (auto simp add: det3_def' ccw'_def) have tr: "fst s < fst r \ fst t = fst s \ snd t \ snd r" by (simp add: s) from s have "fst (r - s) = fst (x *\<^sub>R (t - s))" "snd (r - s) = snd (x *\<^sub>R (t - s))" by (auto simp: ) hence "x = (if fst (t - s) = 0 then snd (r - s) / snd (t - s) else fst (r - s) / fst (t - s))" using \s \ t\ by (auto simp add: field_simps prod_eq_iff) also have "\ \ 1" using assms by (auto simp: lex_def psi_def tr) finally have "x < 1" using neq by simp thus ?thesis using ccw' by (auto simp: ccw'.translate_origin) qed (insert assms, simp) lemma lex_contr: assumes "distinct4 t s q r" assumes "lex t s" "lex s r" assumes "det3 t s r = 0" assumes "ccw' t s q" assumes "ccw' t q r" shows "False" using ccw'_subst_psi_disj[of t s r q] assms by (cases "r = t") (auto simp: det3_def' algebra_simps psi_def ccw'_def) lemma contra4: assumes "distinct4 s r q p" assumes lex: "lex q p" "lex p r" "lex r s" assumes ccw: "ccw r q s" "ccw r s p" "ccw r q p" shows False proof cases assume c: "ccw s q p" from c have *: "indelta s r q p" using assms by simp with contra1[OF assms(1)] have "\ (lex r q \ lex q p \ lex p s)" by blast hence "\ lex q p" using \ccw s q p\ contra1 cyclic assms nondegenerate by blast thus False using assms by simp next assume "\ ccw s q p" with ccw have "ccw q s p \ ccw s p r \ ccw p r q \ ccw r q s" by (metis assms(1) ccw'.cyclic ccw_def not_ccw'_eq psi_disjuncts) moreover from lex have "lex q r" "lex q s" "lex p r" "lex p s" by order+ ultimately show False using contra3[of r q p s] \distinct4 s r q p\ by blast qed lemma not_coll_ordered_lexI: assumes "l \ x0" and "lex x1 r" and "lex x1 l" and "lex r x0" and "lex l x0" and "ccw' x0 l x1" and "ccw' x0 x1 r" shows "det3 x0 l r \ 0" proof assume "coll x0 l r" from \coll x0 l r\ have 1: "coll 0 (l - x0) (r - x0)" by (simp add: det3_def' algebra_simps) from \lex r x0\ have 2: "lex (r - x0) 0" by (auto simp add: lex_def) from \lex l x0\ have 3: "lex (l - x0) 0" by (auto simp add: lex_def) from \ccw' x0 l x1\ have 4: "ccw' 0 (l - x0) (x1 - x0)" by (simp add: det3_def' ccw'_def algebra_simps) from \ccw' x0 x1 r\ have 5: "ccw' 0 (x1 - x0) (r - x0)" by (simp add: det3_def' ccw'_def algebra_simps) from \lex x1 r\ have 6: "lex 0 (r - x0 - (x1 - x0))" by (auto simp: lex_def) from \lex x1 l\ have 7: "lex 0 (l - x0 - (x1 - x0))" by (auto simp: lex_def) define r' where "r' = r - x0" define l' where "l' = l - x0" define x0' where "x0' = x1 - x0" from 1 2 3 4 5 6 7 have rs: "coll 0 l' r'" "lex r' 0" "lex l' 0" "ccw' 0 l' x0'" "ccw' 0 x0' r'" "lex 0 (r' - x0')" "lex 0 (l' - x0')" unfolding r'_def[symmetric] l'_def[symmetric] x0'_def[symmetric] by auto from assms have "l' \ 0" by (auto simp: l'_def) from coll_scale[OF \coll 0 l' _\ this] obtain y where y: "r' = y *\<^sub>R l'" by auto { assume "y > 0" with rs have False by (auto simp: det3_def' algebra_simps y ccw'_def) } moreover { assume "y < 0" with rs have False by (auto simp: lex_def not_less algebra_simps algebra_split_simps y ccw'_def) } moreover { assume "y = 0" from this rs have False by (simp add: ccw'_def y) } ultimately show False by arith qed interpretation ccw_system4 ccw proof unfold_locales fix p q r t assume ccw: "ccw t q r" "ccw p t r" "ccw p q t" show "ccw p q r" proof (cases "det3 t q r = 0 \ det3 p t r = 0 \ det3 p q t = 0") case True { assume "psi t q r \ psi q r t \ psi r t q" "psi p t r \ psi t r p \ psi r p t" "psi p q t \ psi q t p \ psi t p q" hence "psi p q r \ psi q r p \ psi r p q" using lex_sym_eq_iff psi_def by blast } with True ccw show ?thesis by (simp add: det3_def' algebra_simps ccw_def ccw'_def) next case False hence "0 \ det3 t q r" "0 \ det3 p t r" "0 \ det3 p q t" using ccw by (auto simp: less_eq_real_def ccw_def ccw'_def) with False show ?thesis by (auto simp: ccw_def det3_def' algebra_simps ccw'_def intro!: disjI1) qed qed lemma lex_total: "lex t q \ t \ q \ lex q t \ t \ q \ t = q" by auto lemma ccw_two_up_contra: assumes c: "ccw' t p q" "ccw' t q r" assumes ccws: "ccw t s p" "ccw t s q" "ccw t s r" "ccw t p q" "ccw t q r" "ccw t r p" assumes distinct: "distinct5 t s p q r" shows False proof - from ccws have nn: "det3 t s p \ 0" "det3 t s q \ 0" "det3 t s r \ 0" "det3 t r p \ 0" by (auto simp add: less_eq_real_def ccw_def ccw'_def) with c det_identity[of t p q s r] have tsr: "coll t s r" and tsp: "coll t s p" by (auto simp: add_nonneg_eq_0_iff ccw'_def) moreover have trp: "coll t r p" by (metis ccw'_subst_collinear distinct not_ccw'_eq tsr tsp) ultimately have tpr: "coll t p r" by (auto simp: det3_def' algebra_simps) moreover have psi: "psi t p r \ psi t r p \ psi r p t" unfolding psi_def proof - have ntsr: "\ ccw' t s r" "\ ccw' t r s" using tsr by (auto simp: not_ccw'_eq det3_def' algebra_simps) have f8: "\ ccw' t r s" using tsr not_ccw'_eq by blast have f9: "\ ccw' t r p" using tpr by (simp add: not_ccw'_eq) have f10: "(lex t r \ lex r p \ lex r p \ lex p t \ lex p t \ lex t r)" using ccw_def ccws(6) psi_def f9 by auto have "\ ccw' t r q" using c(2) not_ccw'_eq by blast moreover have "\coll t q s" using ntsr ccw'_subst_collinear distinct c(2) by blast hence "ccw' t s q" by (meson ccw_def ccws(2) not_ccw'_eq) moreover from tsr tsp \coll t r p\ have "coll t p s" "coll t p r" "coll t r s" by (auto simp add: det3_def' algebra_simps) ultimately show "lex t p \ lex p r \ lex t r \ lex r p \ lex r p \ lex p t" by (metis ccw'_subst_psi_disj distinct ccw_def ccws(3) contra4 tsp ntsr(1) f10 lex_total psi_def trp) qed moreover from distinct have "r \ t" by auto ultimately have "ccw' t r q" using c(1) by (rule ccw'_subst_psi_disj) thus False using c(2) by (simp add: ccw'_contra) qed lemma ccw_transitive_contr: fixes t s p q r assumes ccws: "ccw t s p" "ccw t s q" "ccw t s r" "ccw t p q" "ccw t q r" "ccw t r p" assumes distinct: "distinct5 t s p q r" shows False proof - from ccws distinct have *: "ccw p t r" "ccw p q t" by (metis cyclic)+ with distinct have "ccw r p q" using interior[OF _ _ ccws(5) *, of UNIV] by (auto intro: cyclic) from ccws have nonnegs: "det3 t s p \ 0" "det3 t s q \ 0" "det3 t s r \ 0" "det3 t p q \ 0" "det3 t q r \ 0" "det3 t r p \ 0" by (auto simp add: less_eq_real_def ccw_def ccw'_def) { assume "ccw' t p q" "ccw' t q r" "ccw' t r p" hence False using ccw_two_up_contra ccws distinct by blast } moreover { assume c: "coll t q r" "coll t r p" with distinct four_points_aligned(1)[OF c, of s] have "coll t p q" by auto hence "(psi t p q \ psi p q t \ psi q t p)" "psi t q r \ psi q r t \ psi r t q" "psi t r p \ psi r p t \ psi p t r" using ccws(4,5,6) c by (simp_all add: ccw_def ccw'_def) hence False using distinct by (auto simp: psi_def ccw'_def) } moreover { assume c: "det3 t p q = 0" "det3 t q r > 0" "det3 t r p = 0" have "\x. det3 t q r = 0 \ t = x \ r = q \ q = x \ r = p \ p = x \ r = x" by (meson c(1) c(3) distinct four_points_aligned(1)) hence False by (metis (full_types) c(2) distinct less_irrefl) } moreover { assume c: "det3 t p q = 0" "det3 t q r = 0" "det3 t r p > 0" have "\x. det3 t r p = 0 \ t = x \ r = x \ q = x \ p = x" by (meson c(1) c(2) distinct four_points_aligned(1)) hence False by (metis (no_types) c(3) distinct less_numeral_extra(3)) } moreover { assume c: "ccw' t p q" "ccw' t q r" from ccw_two_up_contra[OF this ccws distinct] have False . } moreover { assume c: "ccw' t p q" "ccw' t r p" from ccw_two_up_contra[OF this(2,1), of s] ccws distinct have False by auto } moreover { assume c: "ccw' t q r" "ccw' t r p" from ccw_two_up_contra[OF this, of s] ccws distinct have False by auto } ultimately show "False" using \0 \ det3 t p q\ \0 \ det3 t q r\\0 \ det3 t r p\ by (auto simp: less_eq_real_def ccw'_def) qed interpretation ccw: ccw_system ccw by unfold_locales (metis ccw_transitive_contr nondegenerate) lemma ccw_scaleR1: "det3 0 xr P \ 0 \ 0 < e \ ccw 0 xr P \ ccw 0 (e*\<^sub>Rxr) P" by (simp add: ccw_def) lemma ccw_scaleR2: "det3 0 xr P \ 0 \ 0 < e \ ccw 0 xr P \ ccw 0 xr (e*\<^sub>RP)" by (simp add: ccw_def) lemma ccw_translate3_aux: assumes "\coll 0 a b" assumes "x < 1" assumes "ccw 0 (a - x*\<^sub>Ra) (b - x *\<^sub>R a)" shows "ccw 0 a b" proof - from assms have "\ coll 0 (a - x*\<^sub>Ra) (b - x *\<^sub>R a)" by simp with assms have "ccw' 0 ((1 - x) *\<^sub>R a) (b - x *\<^sub>R a)" by (simp add: algebra_simps ccw_def) thus "ccw 0 a b" using \x < 1\ by (simp add: ccw_def) qed lemma ccw_translate3_minus: "det3 0 a b \ 0 \ x < 1 \ ccw 0 a (b - x *\<^sub>R a) \ ccw 0 a b" using ccw_translate3_aux[of a b x] ccw_scaleR1[of a "b - x *\<^sub>R a" "1-x" ] by (auto simp add: algebra_simps) lemma ccw_translate3: "det3 0 a b \ 0 \ x < 1 \ ccw 0 a b \ ccw 0 a (x *\<^sub>R a + b)" by (rule ccw_translate3_minus) (auto simp add: algebra_simps) lemma ccw_switch23: "det3 0 P Q \ 0 \ (\ ccw 0 Q P \ ccw 0 P Q)" by (auto simp: ccw_def algebra_simps not_ccw'_eq ccw'_not_coll) lemma ccw0_upward: "fst a > 0 \ snd a = 0 \ snd b > snd a \ ccw 0 a b" by (auto simp: ccw_def det3_def' algebra_simps ccw'_def) lemma ccw_uminus3[simp]: "det3 a b c \ 0 \ ccw (-a) (-b) (-c) = ccw a b c" by (auto simp: ccw_def ccw'_def algebra_simps det3_def') lemma coll_minus_eq: "coll (x - a) (x - b) (x - c) = coll a b c" by (auto simp: det3_def' algebra_simps) lemma ccw_minus3: "\ coll a b c \ ccw (x - a) (x - b) (x - c) \ ccw a b c" by (simp add: ccw_def coll_minus_eq) lemma ccw0_uminus[simp]: "\ coll 0 a b \ ccw 0 (-a) (-b) \ ccw 0 a b" using ccw_uminus3[of 0 a b] by simp lemma lex_convex2: assumes "lex p q" "lex p r" "0 \ u" "u \ 1" shows "lex p (u *\<^sub>R q + (1 - u) *\<^sub>R r)" proof cases note \lex p q\ also assume "lex q r" hence "lex q (u *\<^sub>R q + (1 - u) *\<^sub>R r)" using \0 \ u\ \u \ 1\ by (rule lex_convex_self2) finally (lex_trans) show ?thesis . next note \lex p r\ also assume "\ lex q r" hence "lex r q" by simp hence "lex r ((1 - u) *\<^sub>R r + (1 - (1 - u)) *\<^sub>R q)" using \0 \ u\ \u \ 1\ by (intro lex_convex_self2) simp_all finally (lex_trans) show ?thesis by (simp add: ac_simps) qed lemma lex_convex2': assumes "lex q p" "lex r p" "0 \ u" "u \ 1" shows "lex (u *\<^sub>R q + (1 - u) *\<^sub>R r) p" proof - have "lex (- p) (u *\<^sub>R (-q) + (1 - u) *\<^sub>R (-r))" using assms by (intro lex_convex2) (auto simp: lex_def) thus ?thesis by (auto simp: lex_def algebra_simps) qed lemma psi_convex1: assumes "psi c a b" assumes "psi d a b" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi (u *\<^sub>R c + v *\<^sub>R d) a b" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2' lex_convex2) qed lemma psi_convex2: assumes "psi a c b" assumes "psi a d b" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi a (u *\<^sub>R c + v *\<^sub>R d) b" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2' lex_convex2) qed lemma psi_convex3: assumes "psi a b c" assumes "psi a b d" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2) qed lemma coll_convex: assumes "coll a b c" "coll a b d" assumes "0 \ u" "0 \ v" "u + v = 1" shows "coll a b (u *\<^sub>R c + v *\<^sub>R d)" proof cases assume "a \ b" with assms(1, 2) obtain x y where xy: "c - a = x *\<^sub>R (b - a)" "d - a = y *\<^sub>R (b - a)" by (auto simp: det3_translate_origin dest!: coll_scale) from assms have "(u + v) *\<^sub>R a = a" by simp hence "u *\<^sub>R c + v *\<^sub>R d - a = u *\<^sub>R (c - a) + v *\<^sub>R (d - a)" by (simp add: algebra_simps) also have "\ = u *\<^sub>R x *\<^sub>R (b - a) + v *\<^sub>R y *\<^sub>R (b - a)" by (simp add: xy) also have "\ = (u * x + v * y) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "coll 0 (b - a) \" by (simp add: coll_scaleR_right_eq) finally show ?thesis by (auto simp: det3_translate_origin) qed simp lemma (in ccw_vector_space) convex3: assumes "u \ 0" "v \ 0" "u + v = 1" "ccw a b d" "ccw a b c" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - have "v = 1 - u" using assms by simp hence "ccw 0 (b - a) (u *\<^sub>R (c - a) + v *\<^sub>R (d - a))" using assms by (cases "u = 0" "v = 0" rule: bool.exhaust[case_product bool.exhaust]) (auto simp add: translate_origin intro!: add3) also have "(u + v) *\<^sub>R a = a" by (simp add: assms) hence "u *\<^sub>R (c - a) + v *\<^sub>R (d - a) = u *\<^sub>R c + v *\<^sub>R d - a" by (auto simp: algebra_simps) finally show ?thesis by (simp add: translate_origin) qed lemma ccw_self[simp]: "ccw a a b" "ccw b a a" by (auto simp: ccw_def psi_def intro: cyclic) lemma ccw_sefl'[simp]: "ccw a b a" by (rule cyclic) simp lemma ccw_convex': assumes uv: "u \ 0" "v \ 0" "u + v = 1" assumes "ccw a b c" and 1: "coll a b c" assumes "ccw a b d" and 2: "\ coll a b d" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have u: "0 \ u" "u \ 1" and v: "v = 1 - u" by (auto simp: algebra_simps) let ?c = "u *\<^sub>R c + v *\<^sub>R d" from 1 have abd: "ccw' a b d" using assms by (auto simp: ccw_def) { assume 2: "\ coll a b c" from 2 have "ccw' a b c" using assms by (auto simp: ccw_def) with abd have "ccw' a b ?c" using assms by (auto intro!: ccw'.convex3) hence ?thesis by (simp add: ccw_def) } moreover { assume 2: "coll a b c" { assume "a = b" hence ?thesis by simp } moreover { assume "v = 0" hence ?thesis by (auto simp: v assms) } moreover { assume "v \ 0" "a \ b" have "coll c a b" using 2 by (auto simp: det3_def' algebra_simps) from coll_ex_scaling[OF \a \ b\ this] obtain r where c: "c = a + r *\<^sub>R (b - a)" by auto have *: "u *\<^sub>R (a + r *\<^sub>R (b - a)) + v *\<^sub>R d - a = (u * r) *\<^sub>R (b - a) + (1 - u) *\<^sub>R (d - a)" by (auto simp: algebra_simps v) have "ccw' a b ?c" using \v \ 0\ uv abd by (simp add: ccw'.translate_origin c *) hence ?thesis by (simp add: ccw_def) } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma ccw_convex: assumes uv: "u \ 0" "v \ 0" "u + v = 1" assumes "ccw a b c" assumes "ccw a b d" assumes lex: "coll a b c \ coll a b d \ lex b a" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have u: "0 \ u" "u \ 1" and v: "v = 1 - u" by (auto simp: algebra_simps) let ?c = "u *\<^sub>R c + v *\<^sub>R d" { assume coll: "coll a b c \ coll a b d" hence "coll a b ?c" by (auto intro!: coll_convex assms) moreover from coll have "psi a b c \ psi b c a \ psi c a b" "psi a b d \ psi b d a \ psi d a b" using assms by (auto simp add: ccw_def ccw'_not_coll) hence "psi a b ?c \ psi b ?c a \ psi ?c a b" using coll uv lex by (auto simp: psi_def ccw_def not_lex lexs_def v intro: lex_convex2 lex_convex2') ultimately have ?thesis by (simp add: ccw_def) } moreover { assume 1: "\ coll a b d" and 2: "\ coll a b c" from 1 have abd: "ccw' a b d" using assms by (auto simp: ccw_def) from 2 have "ccw' a b c" using assms by (auto simp: ccw_def) with abd have "ccw' a b ?c" using assms by (auto intro!: ccw'.convex3) hence ?thesis by (simp add: ccw_def) } moreover { assume "\ coll a b d" "coll a b c" have ?thesis by (rule ccw_convex') fact+ } moreover { assume 1: "coll a b d" and 2: "\ coll a b c" have "0 \ 1 - u" using assms by (auto ) from ccw_convex'[OF this \0 \ u\ _ \ccw a b d\ 1 \ccw a b c\ 2] have ?thesis by (simp add: algebra_simps v) } ultimately show ?thesis by blast qed interpretation ccw: ccw_convex ccw S "\a b. lex b a" for S by unfold_locales (rule ccw_convex) lemma ccw_sorted_scaleR: "ccw.sortedP 0 xs \ r > 0 \ ccw.sortedP 0 (map ((*\<^sub>R) r) xs)" by (induct xs) (auto intro!: ccw.sortedP.Cons ccw_scale23 elim!: ccw.sortedP_Cons simp del: scaleR_Pair) lemma ccw_sorted_implies_ccw'_sortedP: assumes nonaligned: "\y z. y \ set Ps \ z \ set Ps \ y \ z \ \ coll 0 y z" assumes sorted: "linorder_list0.sortedP (ccw 0) Ps" assumes "distinct Ps" shows "linorder_list0.sortedP (ccw' 0 ) Ps" using assms proof (induction Ps) case (Cons P Ps) { fix p assume p: "p \ set Ps" moreover from p Cons.prems have "ccw 0 P p" by (auto elim!: linorder_list0.sortedP_Cons intro: Cons) ultimately have "ccw' 0 P p" using \distinct (P#Ps)\ by (intro ccw_ncoll_imp_ccw Cons) auto } moreover have "linorder_list0.sortedP (ccw' 0) Ps" using Cons.prems by (intro Cons) (auto elim!: linorder_list0.sortedP_Cons intro: Cons) ultimately show ?case by (auto intro!: linorder_list0.Cons ) qed (auto intro: linorder_list0.Nil) end diff --git a/thys/Allen_Calculus/nest.thy b/thys/Allen_Calculus/nest.thy --- a/thys/Allen_Calculus/nest.thy +++ b/thys/Allen_Calculus/nest.thy @@ -1,810 +1,809 @@ (* Title: Allen's qualitative temporal calculus Author: Fadoua Ghourabi (fadouaghourabi@gmail.com) Affiliation: Ochanomizu University, Japan *) theory nest imports Main jointly_exhaustive examples "HOL-Eisbach.Eisbach_Tools" begin section \Nests\ text\Nests are sets of intervals that share a meeting point. We define relation before between nests that give the ordering properties of points.\ subsection \Definitions\ type_synonym 'a nest = "'a set" definition (in arelations) BEGIN :: "'a \ 'a nest" where "BEGIN i = {j | j. (j,i) \ ov \ s \ m \ f^-1 \ d^-1 \ e \ s^-1}" definition (in arelations) END :: "'a \ 'a nest" where "END i = {j | j. (j,i) \ e \ ov^-1 \ s^-1 \ d^-1 \ f \ f^-1 \ m^-1}" definition (in arelations) NEST ::"'a nest \ bool" where "NEST S \ \i. \ i \ (S = BEGIN i \ S = END i)" definition (in arelations) before :: "'a nest \ 'a nest \ bool" (infix "\" 100) where "before N M \ NEST N \ NEST M \ (\n m. \<^cancel>\\ m \ \ n \\ n \ N \ m \ M \ (n,m) \ b)" subsection \Properties of Nests\ lemma intv1: assumes "\ i" shows "i \ BEGIN i" unfolding BEGIN_def by (simp add:e assms) lemma intv2: assumes "\ i" shows "i \ END i" unfolding END_def by (simp add: e assms) lemma NEST_nonempty: assumes "NEST S" shows "S \ {}" using assms unfolding NEST_def by (insert intv1 intv2, auto) lemma NEST_BEGIN: assumes "\ i" shows "NEST (BEGIN i)" using NEST_def assms by auto lemma NEST_END: assumes "\ i" shows "NEST (END i)" using NEST_def assms by auto lemma before: assumes a:"\ i" shows "BEGIN i \ END i" proof - obtain p where pi:"(p,i) \ m" using a M3 m by blast then have p:"p \ BEGIN i" using BEGIN_def by auto obtain q where qi:"(q,i) \ m^-1" using a M3 m by blast then have q:"q \ END i" using END_def by auto from pi qi have c1:"(p,q) \ b" using b m by blast with c1 p q assms show ?thesis by (auto simp:NEST_def before_def) qed lemma meets: fixes i j assumes "\ i" and "\ j" shows "(i,j) \ m = ((END i) = (BEGIN j))" proof assume ij:"(i,j) \ m" then have ibj:"i \ (BEGIN j)" unfolding BEGIN_def by auto from ij have ji:"(j,i) \ m^-1" by simp then have jeo:"j \ (END i)" unfolding END_def by simp show "((END i) = (BEGIN j))" proof {fix x::"'a" assume a:"x \ (END i)" then have asimp:"(x,i) \ e \ ov\ \ s\ \ d\ \ f \ m\ \ f^-1" unfolding END_def by auto then have "x \ (BEGIN j)" using ij eovisidifmifiOm by (auto simp:BEGIN_def) } thus conc1:"END i \ BEGIN j" by auto next {fix x assume b:"x \ (BEGIN j)" then have bsimp:"(x,j) \ ov \ s\ m \ f^-1 \ d^-1 \ e \ s^-1" unfolding BEGIN_def by auto then have "x \ (END i)" using ij ovsmfidiesiOmi by (auto simp:END_def) }thus conc2:"BEGIN j \ END i" by auto qed next assume a0:"END (i::'a) = BEGIN (j::'a)" show "(i,j) \ m" proof (rule ccontr) assume a:"(i,j) \ m" then have "\i\j" using m by auto from a have "(i,j) \ b \ ov \ s \ d \ f^-1 \ e \ f \ s^-1 \ d^-1 \ ov^-1 \ m^-1 \ b^-1" using assms JE by auto thus False proof (auto) {assume ij:"(i,j) \ e" obtain p where ip:"i\p" using M3 assms(1) by auto then have pi:"(p,i)\ m^-1" using m by auto then have "p \ END i" using END_def by auto with a0 have pj:"p \ (BEGIN j) " by auto from ij pi have "(p,j) \ m^-1" by (simp add: e) with pj show ?thesis apply (auto simp:BEGIN_def) using m_rules by auto[7] } next {assume ij: "(j,i) \ m" obtain p where ip:"i\p" using M3 assms(1) by auto then have pi:"(p,i)\ m^-1" using m by auto then have "p \ END i" using END_def by auto with a0 have pj:"p \ (BEGIN j) " by auto from ij have "(i,j) \ m^-1" by simp with pi have "(p,j) \ b^-1" using cmimi by auto with pj show ?thesis apply (auto simp:BEGIN_def) using b_rules by auto } next {assume ij:"(i,j)\ b" have ii:"(i,i)\e" and "i \ END i" using assms intv2 e by auto with a0 have j:"i \ BEGIN j" by simp with ij show ?thesis apply (auto simp:BEGIN_def) using b_rules by auto } next { assume ji:"(j,i) \ b" then have ij:"(i,j) \ b^-1" by simp have ii:"(i,i)\e" and "i \ END i" using assms intv2 e by auto with a0 have j:"i \ BEGIN j" by simp with ij show ?thesis apply (auto simp:BEGIN_def) using b_rules by auto} next {assume ij:"(i,j)\ov" then obtain u v::"'a" where iu:"i\u" and uv:"u\v" and uv:"u\v" using ov by blast from iu have "u \ END i" using m END_def by auto with a0 have u:"u \ BEGIN j" by simp from iu have "(u,i) \ m^-1" using m by auto with ij have uj:"(u,j) \ ov^-1 \ d \ f" using covim by auto show ?thesis using u uj apply (auto simp:BEGIN_def) using ov_rules eovi apply auto[9] using s_rules apply auto[2] using d_rules apply auto[5] using f_rules by auto[5] } next {assume "(j,i) \ ov" then have ij:"(i,j)\ov^-1" by simp let ?p = i from ij have pi:"(?p, i) \ e" by (simp add:e) from ij have pj:"(?p,j) \ ov^-1" by simp from pi have "?p \ END i" using END_def by auto with a0 have "?p \ (BEGIN j) " by auto with pj show ?thesis apply (auto simp:BEGIN_def) using ov_rules by auto } next {assume ij:"(i,j) \ s" then obtain p q t where ip:"i\p" and pq:"p\q" and jq:"j\q" and ti:"t\i" and tj:"t\j" using s by blast from ip have "(p,i) \ m^-1" using m by auto then have "p \ END i" using END_def by auto with a0 have p:"p \ BEGIN j" by simp from ti ip pq tj jq have "(p,j) \ f" using f by blast with p show ?thesis apply (auto simp:BEGIN_def) using f_rules by auto } next {assume "(j,i) \ s" then have ij:"(i,j) \ s^-1" by simp then obtain u v where ju:"j\u" and uv:"u\v" and iv:"i\v" using s by blast from iv have "(v,i) \ m^-1" using m by blast then have "v \ END i" using END_def by auto with a0 have v:"v \ BEGIN j" by simp from ju uv have "(v,j) \ b^-1" using b by auto with v show ?thesis apply (auto simp:BEGIN_def) using b_rules by auto} next {assume ij:"(i,j) \ f" have "(i,i) \ e" and "i \ END i" by (simp add: e) (auto simp: assms intv2 ) with a0 have "i \ BEGIN j" by simp with ij show ?thesis apply (auto simp:BEGIN_def) using f_rules by auto } next {assume "(j,i) \ f" then have "(i,j)\f^-1" by simp then obtain u where ju:"j\u" and iu:"i\u" using f by auto then have ui:"(u,i) \ m^-1" and "u \ END i" apply (simp add: converse.intros m) using END_def iu m by auto with a0 have ubj:"u \ BEGIN j" by simp from ju have "(u,j) \ m^-1" by (simp add: converse.intros m) with ubj show ?thesis apply (auto simp:BEGIN_def) using m_rules by auto } next {assume ij:"(i,j) \ d" then have "(i,i) \ e" and "i \ END i" using assms e by (blast, simp add: intv2) with a0 have "i \ BEGIN j" by simp with ij show ?thesis apply (auto simp:BEGIN_def) using d_rules by auto} next {assume ji:"(j,i) \ d" then have "(i,j) \ d^-1" using d by simp then obtain u v where ju:"j\u" and uv:"u\v" and iv:"i\v" using d using ji by blast then have "(v,i) \ m^-1" and "v \ END i" using m END_def by auto with a0 ju uv have vj:"(v,j) \ b^-1" and "v \ BEGIN j" using b by auto with vj show ?thesis apply (auto simp:BEGIN_def) using b_rules by auto} qed qed qed lemma starts: fixes i j assumes "\ i" and "\ j" shows "((i,j) \ s \ s^-1 \ e) = (BEGIN i = BEGIN j)" proof assume a3:"(i,j) \ s \ s^-1 \ e" show "BEGIN i = BEGIN j" proof - { fix x assume "x \ BEGIN i" then have "(x,i) \ ov \ s \ m \ f\ \ d\ \ e \ s\" unfolding BEGIN_def by auto hence "x \ BEGIN j" using a3 ovsmfidiesiOssie by (auto simp:BEGIN_def) } note c1 = this { fix x assume "x \ BEGIN j" then have xj:"(x,j) \ ov \ s \ m \ f\ \ d\ \ e \ s\" unfolding BEGIN_def by auto then have "x \ BEGIN i" apply (insert converseI[OF a3] xj) apply (subst (asm) converse_Un)+ apply (subst (asm) converse_converse) using ovsmfidiesiOssie by (auto simp:BEGIN_def) } note c2 = this from c1 have "BEGIN i \ BEGIN j" by auto moreover with c2 have "BEGIN j \ BEGIN i" by auto ultimately show ?thesis by auto qed next assume a4:"BEGIN i = BEGIN j" with assms have "i \ BEGIN j" and "j \ BEGIN i" using intv1 by auto then have ij:"(i,j) \ ov \ s \ m \ f^-1 \ d^-1 \ e \ s^-1" and ji:"(j,i) \ ov \ s \ m \ f^-1 \ d^-1 \ e \ s^-1" unfolding BEGIN_def by auto then have ijov:"(i,j) \ ov" apply auto using ov_rules by auto from ij ji have ijm:"(i,j) \ m" apply (simp_all, elim disjE, simp_all) using ov_rules apply auto[13] using s_rules apply auto[11] using m_rules apply auto[9] using f_rules apply auto[7] using d_rules apply auto[5] using m_rules by auto[4] from ij ji have ijfi:"(i,j) \ f^-1" apply (simp_all, elim disjE, simp_all) using ov_rules apply auto[13] using s_rules apply auto[11] using m_rules apply auto[9] using f_rules apply auto[7] using d_rules apply auto[5] using f_rules by auto[4] from ij ji have ijdi:"(i,j) \ d^-1" apply (simp_all, elim disjE, simp_all) using ov_rules apply auto[13] using s_rules apply auto[11] using m_rules apply auto[9] using f_rules apply auto[7] using d_rules apply auto[5] using d_rules by auto[4] from ij ijm ijov ijfi ijdi show "(i, j) \ s \ s\ \ e" by auto qed lemma xj_set:"x \ {a |a. (a, j) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\} = ((x,j) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\)" by blast lemma ends: fixes i j assumes "\ i" and "\ j" shows "((i,j) \ f \ f^-1 \ e) = (END i = END j)" proof assume a3:"(i,j) \ f \ f^-1 \ e" show "END i = END j" proof - { fix x assume "x \ END i" then have "(x,i) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto then have "x \ END j" using a3 unfolding END_def apply (subst xj_set) using ceovisidiffimi_ffie_simp by simp } note c1 =this { fix x assume "x \ END j" then have "(x,j) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto then have "x \ END i" using a3 unfolding END_def by (metis Un_iff ceovisidiffimi_ffie_simp converse_iff eei mem_Collect_eq) } note c2 = this from c1 have "END i \ END j" by auto moreover with c2 have "END j \ END i" by auto ultimately show ?thesis by auto qed (*} note conc = this *) next assume a4:"END i = END j" with assms have "i \ END j" and "j \ END i" using intv2 by auto then have ij:"(i,j) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" and ji:"(j,i) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto then have ijov:"(i,j) \ ov^-1" apply (simp_all, elim disjE, simp_all) using eov es ed efi ef em eov apply auto[13] using ov_rules apply auto[11] using s_rules apply auto[9] using d_rules apply auto[7] using f_rules apply auto[8] using movi by auto from ij ji have ijm:"(i,j) \ m^-1" apply (simp_all, elim disjE, simp_all) using m_rules by auto from ij ji have ijfi:"(i,j) \ s^-1" apply (simp_all, elim disjE, simp_all) using s_rules by auto from ij ji have ijdi:"(i,j) \ d^-1" apply (simp_all, elim disjE, simp_all) using d_rules by auto from ij ijm ijov ijfi ijdi show "(i, j) \ f \ f\ \ e" by auto qed lemma before_irrefl: fixes a shows "\ a \ a" proof (rule ccontr, auto) assume a0:"a \ a" then have "NEST a" unfolding before_def by auto then obtain i where i:"a = BEGIN i \ a = END i" unfolding NEST_def by auto from i show False proof assume "a = BEGIN i" with a0 have "BEGIN i \ BEGIN i" by simp then obtain p q where "p\ BEGIN i" and "q \ BEGIN i" and b:"(p,q) \ b" unfolding before_def by auto then have a1:"(p,i) \ ov \ s \ m \ f\ \ d\ \ e \ s\" and a2:"(i,q) \ ov^-1 \ s^-1 \ m^-1 \ f \ d \ e \ s" unfolding BEGIN_def apply auto using eei apply fastforce by (simp add: e)+ with b show False using piiq[of p i q] - apply auto - using b_rules using disjoint_iff_not_equal by auto + using b_rules by safe fast+ next assume "a = END i" with a0 have "END i \ END i" by simp then obtain p q where "p\ END i" and "q \ END i" and b:"(p,q) \ b" unfolding before_def by auto then have a1:"(p,i) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" and a2:"(q,i) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto with b show False apply (subst (asm) converse_iff[THEN sym]) using cbi_alpha1ialpha4mi neq_bi_alpha1ialpha4mi relcomp.relcompI subsetCE by blast qed qed lemma BEGIN_before: fixes i j assumes "\ i" and "\ j" shows "BEGIN i \ BEGIN j = ((i,j) \ b \ m \ ov \ f\ \ d\)" proof assume a3:"BEGIN i \ BEGIN j" from a3 obtain p q where pa:"p \ BEGIN i" and qc:"q \ BEGIN j" and pq:"(p,q) \ b" unfolding before_def by auto then obtain r where "p\r" and "r\q" using b by auto then have pr:"(p,r) \ m" and rq:"(r,q) \ m" using m by auto from pa have pi:"(p,i) \ ov \ s \ m \ f\ \ d\ \ e \ s\" unfolding BEGIN_def by auto moreover with pr have "(r,p) \ m^-1" by simp ultimately have "(r,i) \ d \ f \ ov^-1 \ e \ f^-1 \ m^-1 \ b^-1 \ s \ s^-1" using cmiov cmis cmim cmifi cmidi cmisi apply ( simp_all,elim disjE, auto) by (simp add: e) then have ir:"(i,r) \ d^-1 \ f^-1 \ ov \ e \ f \ m \ b \ s^-1 \ s" by (metis (mono_tags, lifting) converseD converse_Un converse_converse eei) from qc have "(q,j) \ ov \ s \ m \ f\ \ d\ \ e \ s\" unfolding BEGIN_def by auto with rq have rj:"(r,j) \ b \ s \ m " using m_ovsmfidiesi using contra_subsetD relcomp.relcompI by blast with ir have c1:"(i,j) \ b \ m \ ov \ f\ \ d\ \ d \ e \ s \ s\" using difibs by blast {assume "(i,j) \ s\ (i,j)\s^-1 \ (i,j) \ e" then have "BEGIN i = BEGIN j" using starts Un_iff assms(1) assms(2) by blast with a3 have False by (simp add: before_irrefl)} from c1 have c1':"(i,j) \ b \ m \ ov \ f\ \ d\ \ d " using \(i, j) \ s \ (i, j) \ s\ \ (i, j) \ e \ False\ by blast {assume "(i,j) \ d" with pi have "(p,j) \ e \ s \ d \ ov \ ov^-1 \ s^-1 \ f \ f^-1 \ d^-1" using ovsmfidiesi_d using relcomp.relcompI subsetCE by blast with pq have "(q,j) \ b^-1 \ d \ f \ ov^-1 \ m^-1" apply (subst (asm) converse_iff[THEN sym]) using cbi_esdovovisiffidi by blast with qc have False unfolding BEGIN_def apply (subgoal_tac "(q, j) \ ov \ s \ m \ f\ \ d\ \ e \ s\") prefer 2 apply simp using neq_beta2i_alpha2alpha5m by auto } with c1' show "((i, j) \ b \ m \ ov \ f\ \ d\)" by auto next assume "(i, j) \ b \ m \ ov \ f\ \ d\" then show "BEGIN i \ BEGIN j" proof ( simp_all,elim disjE, simp_all) assume "(i,j) \ b" thus ?thesis using intv1 using before_def NEST_BEGIN assms by metis next assume iu:"(i,j) \ m" obtain l where li:"(l,i) \ m" using M3 m meets_wd assms by blast with iu have "(l,j) \ b" using cmm by auto moreover from li have "l \ (BEGIN i)" using BEGIN_def by auto ultimately show ?thesis using intv1 before_def NEST_BEGIN assms by blast next assume iu:"(i,j) \ ov" obtain l where li:"(l,i) \ m" using M3 m meets_wd assms by blast with iu have "(l,j) \ b" using cmov by auto moreover from li have "l \ (BEGIN i)" using BEGIN_def by auto ultimately show ?thesis using intv1 before_def NEST_BEGIN assms by blast next assume iu:"(j,i) \ f" obtain l where li:"(l,i) \ m" using M3 m meets_wd assms by blast with iu have "(l,j) \ b" using cmfi by auto moreover from li have "l \ (BEGIN i)" using BEGIN_def by auto ultimately show ?thesis using intv1 before_def NEST_BEGIN assms by blast next assume iu:"(j,i) \ d" obtain l where li:"(l,i) \ m" using M3 m meets_wd assms by blast with iu have "(l,j) \ b" using cmdi by auto moreover from li have "l \ (BEGIN i)" using BEGIN_def by auto ultimately show ?thesis using intv1 before_def NEST_BEGIN assms by blast qed qed lemma BEGIN_END_before: fixes i j assumes "\ i" and "\ j" shows "BEGIN i \ END j = ((i,j) \ e \ b \ m \ ov \ ov^-1 \ s \ s^-1 \ f \ f\ \ d \ d\) " proof assume a3:"BEGIN i \ END j" then obtain p q where pa:"p \ BEGIN i" and qc:"q \ END j" and pq:"(p,q) \ b" unfolding before_def by auto then obtain r where "p\r" and "r\q" using b by auto then have pr:"(p,r) \ m" and rq:"(r,q) \ m" using m by auto from pa have pi:"(p,i) \ ov \ s \ m \ f\ \ d\ \ e \ s\" unfolding BEGIN_def by auto moreover with pr have "(r,p) \ m^-1" by simp ultimately have "(r,i) \ d \ f \ ov^-1 \ e \ f^-1 \ m^-1 \ b^-1 \ s \ s^-1" using cmiov cmis cmim cmifi cmidi e cmisi by ( simp_all,elim disjE, auto simp:e) then have ir:"(i,r) \ d^-1 \ f^-1 \ ov \ e \ f \ m \ b \ s^-1 \ s" by (metis (mono_tags, lifting) converseD converse_Un converse_converse eei) from qc have "(q,j) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto with rq have rj:"(r,j) \ m \ ov \ s \ d \ b \ f^-1 \ f \ e" using cm_alpha1ialpha4mi by blast with ir show c1:"(i,j) \ e \ b \ m \ ov \ ov^-1 \ s \ s^-1 \ f \ f\ \ d \ d\" using difimov by blast next assume a4:"(i, j) \ e \ b \ m \ ov \ ov\ \ s \ s\ \ f \ f\ \ d \ d\" then show "BEGIN i \ END j" proof ( simp_all,elim disjE, simp_all) assume "(i,j) \ e" obtain l k where l:"l\i" and "i\k" using M3 meets_wd assms by blast with \(i,j) \ e\ have k:"j\k" by (simp add: e) from l k have "(l,i) \ m" and "(k,j) \ m^-1" using m by auto then have "l \ BEGIN i" and "k \ END j" using BEGIN_def END_def by auto moreover from l \i\k\ have "(l,k) \ b" using b by auto ultimately show ?thesis using before_def assms NEST_BEGIN NEST_END by blast next assume "(i,j) \ b" then show ?thesis using before_def assms NEST_BEGIN NEST_END intv1[of i] intv2[of j] by auto next assume "(i,j) \ m" obtain l where "l\i" using M3 assms by blast then have "l\BEGIN i" using m BEGIN_def by auto moreover from \(i,j)\m\ \l\i\ have "(l,j) \ b" using b m by blast ultimately show ?thesis using intv2[of j] assms NEST_BEGIN NEST_END before_def by blast next assume "(i,j) \ ov" then obtain l k where li:"l\i" and lk:"l\k" and ku:"k\j" using ov by blast from li have "l \ BEGIN i" using m BEGIN_def by auto moreover from lk ku have "(l,j) \ b" using b by auto ultimately show ?thesis using intv2[of j] assms NEST_BEGIN NEST_END before_def by blast next assume "(j,i) \ ov" then obtain l k v where uv:"j\v" and lk:"l\k" and kv:"k\v" and li:"l\i" using ov by blast from li have "l \ BEGIN i" using m BEGIN_def by auto moreover from uv have "v \ END j" using m END_def by auto moreover from lk kv have "(l,v) \ b" using b by auto ultimately show ?thesis using assms NEST_BEGIN NEST_END before_def by blast next assume "(i,j) \ s" then obtain v v' where iv:"i\v" and vvp:"v\v'" and "j\v'" using s by blast then have "v' \ END j" using END_def m by auto moreover from iv vvp have "(i,v') \ b" using b by auto ultimately show ?thesis using intv1[of i] assms NEST_BEGIN NEST_END before_def by blast next assume "(j,i) \ s" then obtain l v where li:"l\i" and lu:"l\j" and "j\v" using s by blast then have "v \ END j" using m END_def by auto moreover from li have "l \ BEGIN i" using m BEGIN_def by auto moreover from lu \j\v\ have "(l,v) \ b" using b by auto ultimately show ?thesis using assms NEST_BEGIN NEST_END before_def by blast next assume "(i,j) : f" then obtain l v where li:"l\i" and iv:"i\v" and "j\v" using f by blast then have "v \ END j" using m END_def by auto moreover from li have "l \ BEGIN i" using m BEGIN_def by auto moreover from iv li have "(l,v) \ b" using b by auto ultimately show ?thesis using assms NEST_BEGIN NEST_END before_def by blast next assume "(j,i) \ f" then obtain l v where li:"l\i" and iv:"i\v" and "j\v" using f by blast then have "v \ END j" using m END_def by auto moreover from li have "l \ BEGIN i" using m BEGIN_def by auto moreover from iv li have "(l,v) \ b" using b by auto ultimately show ?thesis using assms NEST_BEGIN NEST_END before_def by blast next assume "(i,j) : d" then obtain k v where ik:"i\k" and kv:"k\v" and "j\v" using d by blast then have "v \ END j" using END_def m by auto moreover from ik kv have "(i,v) \ b" using b by auto ultimately show ?thesis using intv1[of i] assms NEST_BEGIN NEST_END before_def by blast next assume "(j,i) \ d" then obtain l k where "l\i" and lk:"l\k" and ku:"k\j" using d by blast then have "l \ BEGIN i" using BEGIN_def m by auto moreover from lk ku have "(l,j) \ b" using b by auto ultimately show ?thesis using intv2[of j] assms NEST_BEGIN NEST_END before_def by blast qed qed lemma END_BEGIN_before: fixes i j assumes "\ i" and "\ j" shows "END i \ BEGIN j = ((i,j) \ b) " proof assume a3:"END i \ BEGIN j" from a3 obtain p q where pa:"p \ END i" and qc:"q \ BEGIN j" and pq:"(p,q) \ b" unfolding before_def by auto then obtain r where "p\r" and "r\q" using b by auto then have pr:"(p,r) \ m" and rq:"(r,q) \ m" using m by auto from pa have pi:"(p,i) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto moreover with pr have "(r,p) \ m^-1" by simp ultimately have "(r,i) \ m^-1 \ b^-1" using e cmiovi cmisi cmidi cmif cmifi cmimi by ( simp_all,elim disjE, auto simp:e) then have ir:"(i,r) \ m \ b" by simp from qc have "(q,j) \ ov \ s \ m \ f\ \ d\ \ e \ s\" unfolding BEGIN_def by auto with rq have rj:"(r,j) \ b \ m " using cmov cms cmm cmfi cmdi e cmsi by (simp_all, elim disjE, auto simp:e) with ir show "(i,j) \ b" using cmb cmm cbm cbb by auto next assume "(i,j) \ b" thus "END i \ BEGIN j" using intv1[of j] intv2[of i] assms before_def NEST_END NEST_BEGIN by auto qed lemma END_END_before: fixes i j assumes "\ i" and "\ j" shows "END i \ END j = ((i,j) \ b \ m \ ov \ s \ d) " proof assume a3:"END i \ END j" from a3 obtain p q where pa:"p \ END i" and qc:"q \ END j" and pq:"(p,q) \ b" unfolding before_def by auto then obtain r where "p\r" and "r\q" using b by auto then have pr:"(p,r) \ m" and rq:"(r,q) \ m" using m by auto from pa have pi:"(p,i) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto moreover with pr have "(r,p) \ m^-1" by simp ultimately have "(r,i) \ m^-1 \ b^-1" using e cmiovi cmisi cmidi cmif cmifi cmimi by ( simp_all,elim disjE, auto simp:e) then have ir:"(i,r) \ m \ b" by simp from qc have "(q,j) \ e \ ov\ \ s\ \ d\ \ f \ f\ \ m\" unfolding END_def by auto with rq have rj:"(r,j) \ m \ ov \ s \ d \ b \ f^-1 \ e \ f " using e cmovi cmsi cmdi cmf cmfi cmmi by (simp_all, elim disjE, auto simp:e) with ir show "(i,j) \ b \ m \ ov \ s \ d" using cmm cmov cms cmd cmb cmfi e cmf cbm cbov cbs cbd cbb cbfi cbf by (simp_all, elim disjE, auto simp:e) next assume "(i, j) \ b \ m \ ov \ s \ d" then show "END i \ END j" proof ( simp_all,elim disjE, simp_all) assume "(i,j) \ b" thus ?thesis using intv2[of i] intv2[of j] assms NEST_END before_def by blast next assume "(i,j) \ m" obtain v where "j\v" using M3 assms by blast with \(i,j) \ m\ have "(i,v) \b" using b m by blast moreover from \j\v\ have "v \ END j" using m END_def by auto ultimately show ?thesis using intv2[of i] assms NEST_END before_def by blast next assume "(i,j) : ov" then obtain v v' where iv:"i\v" and vvp:"v\v'" and "j\v'" using ov by blast then have "v' \ END j" using m END_def by auto moreover from iv vvp have "(i,v') \ b" using b by auto ultimately show ?thesis using intv2[of i] assms NEST_END before_def by blast next assume "(i,j) \ s" then obtain v v' where iv:"i\v" and vvp:"v\v'" and "j\v'" using s by blast then have "v' \ END j" using m END_def by auto moreover from iv vvp have "(i,v') \ b" using b by auto ultimately show ?thesis using intv2[of i] assms NEST_END before_def by blast next assume "(i,j) \ d" then obtain v v' where iv:"i\v" and vvp:"v\v'" and "j\v'" using d by blast then have "v' \ END j" using m END_def by auto moreover from iv vvp have "(i,v') \ b" using b by auto ultimately show ?thesis using intv2[of i] assms NEST_END before_def by blast qed qed lemma overlaps: assumes "\ i" and "\ j" shows "(i,j) \ ov = ((BEGIN i \ BEGIN j) \ (BEGIN j \ END i) \ (END i \ END j))" proof assume a:"(i,j) \ ov" then obtain n t q u v where nt:"n\t" and tj:"t\j" and tq:"t\q" and qu:"q\u" and iu:"i\u" and uv:"u\v" and jv:"j\v" and "n\i" using ov by blast then have ni:"(n,i) \ m" using m by blast then have n:"n \ BEGIN i" unfolding BEGIN_def by auto from nt tj have nj:"(n,j) \ b" using b by auto have "j \ BEGIN j" using assms(2) by (simp add: intv1) with assms n nj have c1:"BEGIN i \ BEGIN j" unfolding before_def using NEST_BEGIN by blast from tj have a1:"(t,j) \ m" and a2:"t \ BEGIN j" using m BEGIN_def by auto from iu have "(u,i) \ m^-1" and "u \ END i" using m END_def by auto with assms tq qu a2 have c2:"BEGIN j \ END i" unfolding before_def using b NEST_BEGIN NEST_END by blast have "i \ END i" by (simp add: assms intv2) moreover with jv have "v \ END j" using m END_def by auto moreover with iu uv have "(i,v) \ b" using b by auto ultimately have c3:"END i \ END j" using assms NEST_END before_def by blast show "((BEGIN i \ BEGIN j) \ (BEGIN j \ END i) \ (END i \ END j))" using c1 c2 c3 by simp next assume a0:"((BEGIN i \ BEGIN j) \ (BEGIN j \ END i) \ (END i \ END j))" then have "(i,j) \ b \ m \ ov \ f\ \ d\ \ (i,j) \ e \ b^-1 \ m^-1 \ ov^-1 \ ov \ s^-1 \ s \ f^-1 \ f \ d^-1 \ d \ (i,j) \ b \ m \ ov \ s \ d" using BEGIN_before BEGIN_END_before END_END_before assms by (metis (no_types, lifting) converseD converse_Un converse_converse eei) then have "(i,j) \ (b \ m \ ov \ f\ \ d\) \ (e \ b^-1 \ m^-1 \ ov^-1 \ ov \ s^-1 \ s \ f^-1 \ f \ d^-1 \ d) \ (b \ m \ ov \ s \ d)" by (auto) then show "(i,j) \ ov" using inter_ov by blast qed subsection \Ordering of nests\ class strict_order = fixes ls::"'a nest \ 'a nest \ bool" assumes irrefl:"\ ls a a" and trans:"ls a c \ ls c g \ ls a g" and asym:"ls a c \ \ ls c a" class total_strict_order = strict_order + assumes trichotomy: "a = c \ (\ (ls a c) \ \ (ls c a))" interpretation nest:total_strict_order "(\) " proof { fix a::"'a nest" show "\ a \ a" by (simp add: before_irrefl) } note irrefl_nest = this {fix a c::"'a nest" assume "a = c" show "\ a \ c \ \ c \ a" by (simp add: \a = c\ irrefl_nest)} note trichotomy_nest = this {fix a c g::"'a nest" assume a:"a \ c" and c:" c \ g" show " a \ g" proof - from a c have na:"NEST a" and nc:"NEST c" and ng:"NEST g" unfolding before_def by auto from na obtain i where i:"a = BEGIN i \ a = END i" and wdi:"\ i" unfolding NEST_def by auto from nc obtain j where j:"c = BEGIN j \ c = END j" and wdj:"\ j" unfolding NEST_def by auto from ng obtain u where u:"g = BEGIN u \ g = END u" and wdu:"\ u" unfolding NEST_def by auto from i j u show ?thesis proof (elim disjE, auto) assume abi:"a = BEGIN i" and cbj:"c = BEGIN j" and gbu:"g = BEGIN u" from abi cbj a wdi wdj have "(i,j) \ b \ m \ ov \ f\ \ d\ " using BEGIN_before by auto moreover from cbj gbu c wdj wdu have "(j,u) \ b \ m \ ov \ f\ \ d\" using BEGIN_before by auto ultimately have c1:"(i,u) \ b \ m \ ov \ f^-1 \ d^-1" using cbeta2_beta2 by blast then have "a \ g" by (simp add: BEGIN_before abi gbu wdi wdu) thus "BEGIN i \ BEGIN u" using abi gbu by auto next assume abi:"a = BEGIN i" and cbj:"c = BEGIN j" and geu:"g = END u" from abi cbj a wdi wdj have "(i,j) \ b \ m \ ov \ f\ \ d\ " using BEGIN_before by auto moreover from cbj geu c wdj wdu have "(j,u) : e \ b \ m \ ov \ ov\ \ s \ s\ \ f \ f\ \ d \ d\" using BEGIN_END_before by auto ultimately have "(i,u) \ e \ b \ m \ ov \ ov\ \ s \ s\ \ f \ f\ \ d \ d\" using cbeta2_gammabm by blast then have "a \ g" by (simp add: BEGIN_END_before abi geu wdi wdj wdu) thus "BEGIN i \ END u" using abi geu by auto next assume abi:"a = BEGIN i" and cej:"c = END j" and gbu:"g = BEGIN u" from abi cej a wdi wdj have ij:"(i,j) : e \ b \ m \ ov \ ov\ \ s \ s\ \ f \ f\ \ d \ d\" using BEGIN_END_before by auto from cej gbu c wdj wdu have "(j,u) \ b " using END_BEGIN_before by auto with ij have "(i,u) \ b \ m \ ov \ f^-1 \ d^-1" using ebmovovissifsiddib by ( auto) thus "BEGIN i \ BEGIN u" by (simp add: BEGIN_before abi gbu wdi wdu) next assume abi:"a = BEGIN i" and cej:"c = END j" and geu:"g = END u" with a have "(i,j) \ e \ b \ m \ ov \ ov\ \ s \ s\ \ f \ f\ \ d \ d\" using BEGIN_END_before wdi wdj by auto moreover from cej geu c wdj wdu have "(j,u) \ b \ m \ ov \ s \ d" using END_END_before by auto ultimately have "(i,u) \ b \ m \ ov \ s \ d \ f^-1 \ d^-1 \ ov^-1 \ s\ \ f \ e" using ebmovovissiffiddibmovsd by blast thus "BEGIN i \ END u" using BEGIN_END_before wdi wdu by auto next assume aei:"a = END i" and cbj:"c = BEGIN j" and gbu:"g = BEGIN u" from a aei cbj wdi wdj have "(i,j) \ b" using END_BEGIN_before by auto moreover from c cbj gbu wdj wdu have "(j,u) \ b \ m \ ov \ f\ \ d\" using BEGIN_before by auto ultimately have "(i,u) : b" using cbb cbm cbov cbfi cbdi by (simp_all, elim disjE, auto) thus "END i \ BEGIN u" using END_BEGIN_before wdi wdu by auto next assume aei:"a = END i" and cbj:"c = BEGIN j" and geu:"g = END u" from a aei cbj wdi wdj have "(i,j) \ b" using END_BEGIN_before by auto moreover from c cbj geu wdj wdu have "(j,u) \ e \ b \ m \ ov \ ov\ \ s \ s\ \ f \ f\ \ d \ d\" using BEGIN_END_before by auto ultimately have "(i,u) \ b \ m \ ov \ s \ d" using bebmovovissiffiddi by blast thus "END i \ END u" using END_END_before wdi wdu by auto next assume aei:"a = END i" and cej:"c = END j" and gbu:"g = BEGIN u" from aei cej wdi wdj have "(i,j) \ b \ m \ ov \ s \ d" using END_END_before a by auto moreover from cej gbu c wdj wdu have "(j,u) \ b" using END_BEGIN_before by auto ultimately have "(i,u) \ b" using cbb cmb covb csb cdb by (simp_all, elim disjE, auto) thus "END i \ BEGIN u" using END_BEGIN_before wdi wdu by auto next assume aei:"a = END i" and cej:"c = END j" and geu:"g = END u" from aei cej wdi wdj have "(i,j) \ b \ m \ ov \ s \ d" using END_END_before a by auto moreover from cej geu c wdj wdu have "(j,u) \ b \ m \ ov \ s \ d" using END_END_before by auto ultimately have "(i,u) \ b \ m \ ov \ s \ d" using calpha1_alpha1 by auto thus "END i \ END u" using END_END_before wdi wdu by auto qed qed} note trans_nest = this { fix a c::"'a nest" assume a:"a \ c" show "\ c \ a" apply (rule ccontr, auto) using a irrefl_nest trans_nest by blast} qed end diff --git a/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs.thy b/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs.thy --- a/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs.thy +++ b/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs.thy @@ -1,513 +1,512 @@ section \Boolean Programs\ theory BoolProgs imports CAVA_Base.CAVA_Base begin subsection \Syntax and Semantics\ datatype bexp = TT | FF | V nat | Not bexp | And bexp bexp | Or bexp bexp type_synonym state = "bitset" fun bval :: "bexp \ state \ bool" where "bval TT s = True" | "bval FF s = False" | "bval (V n) s = bs_mem n s" | "bval (Not b) s = (\ bval b s)" | "bval (And b\<^sub>1 b\<^sub>2) s = (bval b\<^sub>1 s & bval b\<^sub>2 s)" | "bval (Or b\<^sub>1 b\<^sub>2) s = (bval b\<^sub>1 s | bval b\<^sub>2 s)" datatype instr = AssI "nat list" "bexp list" | TestI bexp int | ChoiceI "(bexp * int) list" | GotoI int type_synonym config = "nat * state" type_synonym bprog = "instr array" text \ Semantics Notice: To be equivalent in semantics with SPIN, there is no such thing as a finite run: \begin{itemize} \item Deadlocks (i.e. empty Choice) are self-loops \item program termination is self-loop \end{itemize} \ fun exec :: "instr \ config \ config list" where "exec instr (pc,s) = (case instr of AssI ns bs \ let bvs = zip ns (map (\b. bval b s) bs) in [(pc + 1, foldl (\s (n,bv). set_bit s n bv) s bvs)] | TestI b d \ [if bval b s then (pc+1, s) else (nat(int(pc+1)+d), s)] | ChoiceI bis \ let succs = [(nat(int(pc+1)+i), s) . (b,i) <- bis, bval b s] in if succs = [] then [(pc,s)] else succs | GotoI d \ [(nat(int(pc+1)+d),s)])" (* exec' is merely a stuttering optimization. It summarizes chains of choices or jumps, as long as they go forward in the program. The forward-condition is to ensure termination. TODO: Make this explicit, and prove stuttering equivalence! *) function exec' :: "bprog \ state \ nat \ nat list" where "exec' ins s pc = ( if pc < array_length ins then ( case (array_get ins pc) of AssI ns bs \ [pc] | TestI b d \ ( if bval b s then exec' ins s (pc+1) else let pc'=(nat(int(pc+1)+d)) in if pc'>pc then exec' ins s pc' else [pc'] ) | ChoiceI bis \ let succs = [(nat(int(pc+1)+i)) . (b,i) <- bis, bval b s] in if succs = [] then [pc] else concat (map (\pc'. if pc'>pc then exec' ins s pc' else [pc']) succs) | GotoI d \ let pc' = nat(int(pc+1)+d) in (if pc'>pc then exec' ins s pc' else [pc']) ) else [pc] )" by pat_completeness auto termination apply (relation "measure (%(ins,s,pc). array_length ins - pc)") apply auto done fun nexts1 :: "bprog \ config \ config list" where "nexts1 ins (pc,s) = ( if pc < array_length ins then exec (array_get ins pc) (pc,s) else [(pc,s)])" fun nexts :: "bprog \ config \ config list" where "nexts ins (pc,s) = concat ( map (\(pc,s). map (\pc. (pc,s)) (exec' ins s pc)) (nexts1 ins (pc,s)) )" declare nexts.simps [simp del] datatype com = SKIP | Assign "nat list" "bexp list" | Seq com com | GC "(bexp * com)list" | IfTE bexp com com | While bexp com locale BoolProg_Syntax begin notation Assign ("_ ::= _" [999, 61] 61) and Seq ("_;/ _" [60, 61] 60) and GC ("IF _ FI") and IfTE ("(IF _/ THEN _/ ELSE _)" [0, 61, 61] 61) and While ("(WHILE _/ DO _)" [0, 61] 61) end context begin interpretation BoolProg_Syntax . fun comp' :: "com \ instr list" where "comp' SKIP = []" | "comp' (Assign n b) = [AssI n b]" | "comp' (c1;c2) = comp' c1 @ comp' c2" | "comp' (IF gcs FI) = (let cgcs = map (\(b,c). (b,comp' c)) gcs in let addbc = (\(b,cc) (bis,ins). let cc' = cc @ (if ins = [] then [] else [GotoI (int(length ins))]) in let bis' = map (\(b,i). (b, i + int(length cc'))) bis in ((b,0)#bis', cc' @ ins)) in let (bis,ins) = foldr addbc cgcs ([],[]) in ChoiceI bis # ins)" | "comp' (IF b THEN c1 ELSE c2) = (let ins1 = comp' c1 in let ins2 = comp' c2 in let i1 = int(length ins1 + 1) in let i2 = int(length ins2) in TestI b i1 # ins1 @ GotoI i2 # ins2)" | "comp' (WHILE b DO c) = (let ins = comp' c in let i = int(length ins + 1) in TestI b i # ins @ [GotoI (-(i+1))])" (* a test: *) value "comp' (IF [(V 0, [1,0] ::= [TT, FF]), (V 1, [0] ::= [TT])] FI)" end definition comp :: "com \ bprog" where "comp = array_of_list \ comp'" (* Optimization - Gotos: Resolve Goto-Chains: If a Goto referres to another Goto -- add the second offset to the first one. If a Goto has a negative offset ignore it, to avoid problems with loops. *) fun opt' where "opt' (GotoI d) ys = (let next = \i. (case i of GotoI d \ d + 1 | _ \ 0) in if d < 0 \ nat d \ length ys then (GotoI d)#ys else let d' = d + next (ys ! nat d) in (GotoI d' # ys))" | "opt' x ys = x#ys" definition opt :: "instr list \ instr list" where "opt instr = foldr opt' instr []" definition optcomp :: "com \ bprog" where "optcomp \ array_of_list \ opt \ comp'" subsection \Finiteness of reachable configurations\ inductive_set reachable_configs for bp :: bprog and c\<^sub>s :: config \ \start configuration\ where "c\<^sub>s \ reachable_configs bp c\<^sub>s" | "c \ reachable_configs bp c\<^sub>s \ x \ set (nexts bp c) \ x \ reachable_configs bp c\<^sub>s" lemmas reachable_configs_induct = reachable_configs.induct[split_format(complete),case_names 0 1] fun offsets :: "instr \ int set" where "offsets (AssI _ _) = {0}" | "offsets (TestI _ i) = {0,i}" | "offsets (ChoiceI bis) = set(map snd bis) \ {0}" | "offsets (GotoI i) = {i}" definition offsets_is :: "instr list \ int set" where "offsets_is ins = (UN instr : set ins. offsets instr)" definition max_next_pcs :: "instr list \ nat set" where "max_next_pcs ins = {nat(int(length ins + 1) + i) |i. i : offsets_is ins}" lemma finite_max_next_pcs: "finite(max_next_pcs bp)" proof- { fix instr have "finite (offsets instr)" by(cases instr) auto } moreover { fix ins have "max_next_pcs ins = (UN i : offsets_is ins. {nat(int(length ins + 1) + i)})" by(auto simp add: max_next_pcs_def) } ultimately show ?thesis by(auto simp add: offsets_is_def) qed (* TODO: Move *) lemma (in linorder) le_Max_insertI1: "\ finite A; x \ b \ \ x \ Max (insert b A)" by (metis Max_ge finite.insertI insert_iff order_trans) lemma (in linorder) le_Max_insertI2: "\ finite A; A \ {}; x \ Max A \ \ x \ Max (insert b A)" by(auto simp add: max_def not_le simp del: Max_less_iff) lemma max_next_pcs_not_empty: "pc x : set (exec (bp!pc) (pc,s)) \ max_next_pcs bp \ {}" apply(drule nth_mem) apply(fastforce simp: max_next_pcs_def offsets_is_def split: instr.splits) done lemma Max_lem2: assumes "pc < length bp" and "(pc', s') \ set (exec (bp!pc) (pc, s))" shows "pc' \ Max (max_next_pcs bp)" using assms proof (cases "bp ! pc") case (ChoiceI l) show ?thesis proof (cases "pc' = pc") case True with assms ChoiceI show ?thesis by (auto simp: Max_ge_iff max_next_pcs_not_empty finite_max_next_pcs) (force simp add: max_next_pcs_def offsets_is_def dest: nth_mem) next case False with ChoiceI assms obtain b i where bi: "bval b s" "(b,i) \ set l" "pc' = nat(int(pc+1)+i)" by (auto split: if_split_asm) with ChoiceI assms have "i \ \(offsets ` (set bp))" by (force dest: nth_mem) with bi assms have "\a. (a \ max_next_pcs bp \ pc' \ a)" unfolding max_next_pcs_def offsets_is_def by force thus ?thesis by (auto simp: Max_ge_iff max_next_pcs_not_empty[OF assms] finite_max_next_pcs) qed qed (auto simp: Max_ge_iff max_next_pcs_not_empty finite_max_next_pcs, (force simp add: max_next_pcs_def offsets_is_def dest: nth_mem split: if_split_asm)+) lemma Max_lem1: "\ pc < length bp; (pc', s') \ set (exec (bp ! pc) (pc, s))\ \ pc' \ Max (insert x (max_next_pcs bp))" apply(rule le_Max_insertI2) apply (simp add: finite_max_next_pcs) apply(simp add: max_next_pcs_not_empty) apply(auto intro!: Max_lem2 simp del:exec.simps) done definition "pc_bound bp \ max (Max (max_next_pcs (list_of_array bp)) + 1) (array_length bp + 1)" declare exec'.simps[simp del] lemma [simp]: "length (list_of_array a) = array_length a" by (cases a) auto lemma aux2: assumes A: "pc < array_length ins" assumes B: "ofs \ offsets_is (list_of_array ins)" shows "nat (1 + int pc + ofs) < pc_bound ins" proof - have "nat (int (1 + array_length ins) + ofs) \ max_next_pcs (list_of_array ins)" using B unfolding max_next_pcs_def by auto with A show ?thesis unfolding pc_bound_def apply - apply (rule max.strict_coboundedI1) apply auto apply (drule Max_ge[OF finite_max_next_pcs]) apply simp done qed lemma array_idx_in_set: "\ pc < array_length ins; array_get ins pc = x \ \ x \ set (list_of_array ins)" by (induct ins) auto lemma rcs_aux: assumes "pc < pc_bound bp" assumes "pc'\set (exec' bp s pc)" shows "pc' < pc_bound bp" using assms proof (induction bp s pc arbitrary: pc' rule: exec'.induct[case_names C]) case (C ins s pc pc') from C.prems show ?case apply (subst (asm) exec'.simps) apply (split if_split_asm instr.split_asm)+ apply (simp add: pc_bound_def) apply (simp split: if_split_asm add: Let_def) apply (frule (2) C.IH(1), auto) [] apply (auto simp: pc_bound_def) [] apply (frule (2) C.IH(2), auto) [] apply (rename_tac bexp int) apply (subgoal_tac "int \ offsets_is (list_of_array ins)") apply (blast intro: aux2) apply (auto simp: offsets_is_def) [] apply (rule_tac x="TestI bexp int" in bexI, auto simp: array_idx_in_set) [] apply (rename_tac list) - apply (clarsimp split: if_split_asm simp add: Let_def) - apply (elim disjE conjE, auto) [] + apply (auto split: if_split_asm)[] apply (frule (1) C.IH(3), auto) [] apply (force) apply (force) apply (subgoal_tac "ba \ offsets_is (list_of_array ins)") apply (blast intro: aux2) apply (auto simp: offsets_is_def) [] apply (rule_tac x="ChoiceI list" in bexI, auto simp: array_idx_in_set) [] apply (rename_tac int) apply (simp split: if_split_asm add: Let_def) apply (frule (1) C.IH(4), auto) [] apply (subgoal_tac "int \ offsets_is (list_of_array ins)") apply (blast intro: aux2) apply (auto simp: offsets_is_def) [] apply (rule_tac x="GotoI int" in bexI, auto simp: array_idx_in_set) [] apply simp done qed primrec bexp_vars :: "bexp \ nat set" where "bexp_vars TT = {}" | "bexp_vars FF = {}" | "bexp_vars (V n) = {n}" | "bexp_vars (Not b) = bexp_vars b" | "bexp_vars (And b1 b2) = bexp_vars b1 \ bexp_vars b2" | "bexp_vars (Or b1 b2) = bexp_vars b1 \ bexp_vars b2" primrec instr_vars :: "instr \ nat set" where "instr_vars (AssI xs bs) = set xs \ \(bexp_vars`set bs)" | "instr_vars (TestI b _) = bexp_vars b" | "instr_vars (ChoiceI cs) = \(bexp_vars`fst`set cs)" | "instr_vars (GotoI _) = {}" find_consts "'a array \ 'a list" definition bprog_vars :: "bprog \ nat set" where "bprog_vars bp = \(instr_vars`set (list_of_array bp))" definition "state_bound bp s0 \ {s. bs_\ s - bprog_vars bp = bs_\ s0 - bprog_vars bp}" abbreviation "config_bound bp s0 \ {0..< pc_bound bp} \ state_bound bp s0" lemma exec_bound: assumes PCB: "pc < array_length bp" assumes SB: "s \ state_bound bp s0" shows "set (exec (array_get bp pc) (pc,s)) \ config_bound bp s0" proof (clarsimp simp del: exec.simps, intro conjI) obtain instrs where BP_eq[simp]: "bp = Array instrs" by (cases bp) from PCB have PCB'[simp]: "pc < length instrs" by simp fix pc' s' assume STEP: "(pc',s') \ set (exec (array_get bp pc) (pc,s))" hence STEP': "(pc',s') \ set (exec (instrs!pc) (pc,s))" by simp show "pc' < pc_bound bp" using Max_lem2[OF PCB' STEP'] unfolding pc_bound_def by simp show "s' \ state_bound bp s0" using STEP' SB proof (cases "instrs!pc") case (AssI xs vs) have "set xs \ instr_vars (instrs!pc)" by (simp add: AssI) also have "\ \ bprog_vars bp" apply (simp add: bprog_vars_def) by (metis PCB' UN_upper nth_mem) finally have XSB: "set xs \ bprog_vars bp" . { fix x s v assume A: "x \ bprog_vars bp" "s \ state_bound bp s0" have SB_CNV: "bs_\ (set_bit s x v) = (if v then (insert x (bs_\ s)) else (bs_\ s - {x}))" apply (cases v) apply simp_all apply (fold bs_insert_def bs_delete_def) apply (simp_all add: bs_correct) done from A have "set_bit s x v \ state_bound bp s0" unfolding state_bound_def by (auto simp: SB_CNV) } note aux=this { fix vs have "foldl (\s (x, y). set_bit s x y) s (zip xs vs) \ state_bound (Array instrs) s0" using SB XSB apply (induct xs arbitrary: vs s) apply simp apply (case_tac vs) apply simp using aux apply (auto) done } note aux2=this thus ?thesis using STEP' by (simp add: AssI) qed (auto split: if_split_asm) qed lemma in_bound_step: notes [simp del] = exec.simps assumes BOUND: "c \ config_bound bp s0" assumes STEP: "c'\set (nexts bp c)" shows "c' \ config_bound bp s0" using BOUND STEP apply (cases c) apply (auto simp add: nexts.simps split: if_split_asm) apply (frule (2) exec_bound[THEN subsetD]) apply clarsimp apply (frule (1) rcs_aux) apply simp apply (frule (2) exec_bound[THEN subsetD]) apply clarsimp apply (frule (1) rcs_aux) apply simp done lemma reachable_configs_in_bound: "c \ config_bound bp s0 \ reachable_configs bp c \ config_bound bp s0" proof fix c' assume "c' \ reachable_configs bp c" "c \ config_bound bp s0" thus "c' \ config_bound bp s0" apply induction apply simp by (rule in_bound_step) qed lemma reachable_configs_out_of_bound: "(pc',s')\reachable_configs bp (pc,s) \ \ pc < pc_bound bp \ (pc',s') = (pc,s)" proof (induct rule: reachable_configs_induct) case (1 pc' s' pc'' s'') hence [simp]: "pc'=pc" "s'=s" by auto from 1(4) have "\ pc < array_length bp" unfolding pc_bound_def by auto with 1(3) show ?case by (auto simp add: nexts.simps exec'.simps) qed auto lemma finite_bexp_vars[simp, intro!]: "finite (bexp_vars be)" by (induction be) auto lemma finite_instr_vars[simp, intro!]: "finite (instr_vars ins)" by (cases ins) auto lemma finite_bprog_vars[simp, intro!]: "finite (bprog_vars bp)" unfolding bprog_vars_def by simp lemma finite_state_bound[simp, intro!]: "finite (state_bound bp s0)" unfolding state_bound_def apply (rule finite_imageD[where f=bs_\]) apply (rule finite_subset[where B = "{s. s - bprog_vars bp = bs_\ s0 - bprog_vars bp}"]) apply auto [] apply (rule finite_if_eq_beyond_finite) apply simp apply (rule inj_onI) apply (fold bs_eq_def) apply (auto simp: bs_eq_correct) done lemma finite_config_bound[simp, intro!]: "finite (config_bound bp s0)" by blast lemma reachable_configs_finite[simp, intro!]: "finite (reachable_configs bp c)" proof (cases c, clarsimp) fix pc s show "finite (reachable_configs bp (pc, s))" proof (cases "pc < pc_bound bp") case False from reachable_configs_out_of_bound[OF _ False, where s=s] have "reachable_configs bp (pc, s) \ {(pc,s)}" by auto thus ?thesis by (rule finite_subset) auto next case True hence "(pc,s) \ config_bound bp s" by (simp add: state_bound_def) thus ?thesis by (rule finite_subset[OF reachable_configs_in_bound]) simp qed qed definition "bpc_is_run bpc r \ let (bp,c)=bpc in r 0 = c \ (\i. r (Suc i) \ set (BoolProgs.nexts bp (r i)))" definition "bpc_props c \ bs_\ (snd c)" definition "bpc_lang bpc \ {bpc_props o r | r. bpc_is_run bpc r}" (* Printing *) (*definition print_list where "print_list f sep s = (let f' = (\str s. if str = [] then f s else str @ sep @ f s) in ''['' @ (foldl f' [] s) @ '']'')" fun bool_list_to_prop_list where "bool_list_to_prop_list _ [] props = props" | "bool_list_to_prop_list n (x#xs) props = (let props' = if x then n#props else props in bool_list_to_prop_list (Suc n) xs props')" *) fun print_config :: "(nat \ string) \ (bitset \ string) \ config \ string" where "print_config f fx (p,s) = f p @ '' '' @ fx s" end diff --git a/thys/CAVA_LTL_Modelchecker/SM/Analysis/SM_Variables.thy b/thys/CAVA_LTL_Modelchecker/SM/Analysis/SM_Variables.thy --- a/thys/CAVA_LTL_Modelchecker/SM/Analysis/SM_Variables.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/Analysis/SM_Variables.thy @@ -1,239 +1,240 @@ section \Accessed Variables\ theory SM_Variables imports "../Refine/SM_Pid" begin subsection "Expressions" primrec udvars_of_exp :: "exp \ ident set" where "udvars_of_exp (e_var x) = {x}" | "udvars_of_exp (e_localvar x) = {}" | "udvars_of_exp (e_globalvar x) = {}" | "udvars_of_exp (e_const c) = {}" | "udvars_of_exp (e_bin bop e1 e2) = (udvars_of_exp e1) \ (udvars_of_exp e2)" | "udvars_of_exp (e_un uop e) = (udvars_of_exp e)" primrec lvars_of_exp :: "exp \ ident set" where "lvars_of_exp (e_var x) = {}" | "lvars_of_exp (e_localvar x) = {x}" | "lvars_of_exp (e_globalvar x) = {}" | "lvars_of_exp (e_const c) = {}" | "lvars_of_exp (e_bin bop e1 e2) = (lvars_of_exp e1) \ (lvars_of_exp e2)" | "lvars_of_exp (e_un uop e) = (lvars_of_exp e)" primrec gvars_of_exp :: "exp \ ident set" where "gvars_of_exp (e_var x) = {}" | "gvars_of_exp (e_localvar x) = {}" | "gvars_of_exp (e_globalvar x) = {x}" | "gvars_of_exp (e_const c) = {}" | "gvars_of_exp (e_bin bop e1 e2) = (gvars_of_exp e1) \ (gvars_of_exp e2)" | "gvars_of_exp (e_un uop e) = (gvars_of_exp e)" subsection "Commands" primrec join_cmd_exp :: "'a \ ('a \ 'a \ 'a) \ (exp \ 'a) \ cmd \ 'a" where "join_cmd_exp z join f (Assign c x e) = join (f c) (f e)" | "join_cmd_exp z join f (Assign_local c x e) = join (f c) (f e)" | "join_cmd_exp z join f (Assign_global c x e) = join (f c) (f e)" | "join_cmd_exp z join f (Test e) = f e" | "join_cmd_exp z join f (Skip) = z" | "join_cmd_exp z join f (Seq c1 c2) = join (join_cmd_exp z join f c1) (join_cmd_exp z join f c2)" | "join_cmd_exp z join f (Or c1 c2) = join (join_cmd_exp z join f c1) (join_cmd_exp z join f c2)" | "join_cmd_exp z join f (Break) = z" | "join_cmd_exp z join f (Continue) = z" | "join_cmd_exp z join f (Iterate c1 c2) = join (join_cmd_exp z join f c1) (join_cmd_exp z join f c2)" | "join_cmd_exp z join f (Invalid) = z" abbreviation "union_cmd_exp \ join_cmd_exp {} (\)" abbreviation "all_cmd_exp \ join_cmd_exp True (\)" abbreviation "ex_cmd_exp \ join_cmd_exp False (\)" abbreviation "ud_rvars_of_cmd == union_cmd_exp udvars_of_exp" abbreviation "g_rvars_of_cmd == union_cmd_exp gvars_of_exp" abbreviation "l_rvars_of_cmd == union_cmd_exp lvars_of_exp" primrec ud_wvars_of_cmd :: "cmd \ ident set" where "ud_wvars_of_cmd (Assign c x e) = {x}" | "ud_wvars_of_cmd (Assign_local c x e) = {}" | "ud_wvars_of_cmd (Assign_global c x e) = {}" | "ud_wvars_of_cmd (Test e) = {}" | "ud_wvars_of_cmd (Skip) = {}" | "ud_wvars_of_cmd (Seq c1 c2) = (ud_wvars_of_cmd c1 \ ud_wvars_of_cmd c2)" | "ud_wvars_of_cmd (Or c1 c2) = (ud_wvars_of_cmd c1 \ ud_wvars_of_cmd c2)" | "ud_wvars_of_cmd (Break) = {}" | "ud_wvars_of_cmd (Continue) = {}" | "ud_wvars_of_cmd (Iterate c1 c2) = (ud_wvars_of_cmd c1 \ ud_wvars_of_cmd c2)" | "ud_wvars_of_cmd (Invalid) = {}" primrec l_wvars_of_cmd :: "cmd \ ident set" where "l_wvars_of_cmd (Assign c x e) = {}" | "l_wvars_of_cmd (Assign_local c x e) = {x}" | "l_wvars_of_cmd (Assign_global c x e) = {}" | "l_wvars_of_cmd (Test e) = {}" | "l_wvars_of_cmd (Skip) = {}" | "l_wvars_of_cmd (Seq c1 c2) = (l_wvars_of_cmd c1 \ l_wvars_of_cmd c2)" | "l_wvars_of_cmd (Or c1 c2) = (l_wvars_of_cmd c1 \ l_wvars_of_cmd c2)" | "l_wvars_of_cmd (Break) = {}" | "l_wvars_of_cmd (Continue) = {}" | "l_wvars_of_cmd (Iterate c1 c2) = (l_wvars_of_cmd c1 \ l_wvars_of_cmd c2)" | "l_wvars_of_cmd (Invalid) = {}" primrec g_wvars_of_cmd :: "cmd \ ident set" where "g_wvars_of_cmd (Assign c x e) = {}" | "g_wvars_of_cmd (Assign_local c x e) = {}" | "g_wvars_of_cmd (Assign_global c x e) = {x}" | "g_wvars_of_cmd (Test e) = {}" | "g_wvars_of_cmd (Skip) = {}" | "g_wvars_of_cmd (Seq c1 c2) = (g_wvars_of_cmd c1 \ g_wvars_of_cmd c2)" | "g_wvars_of_cmd (Or c1 c2) = (g_wvars_of_cmd c1 \ g_wvars_of_cmd c2)" | "g_wvars_of_cmd (Break) = {}" | "g_wvars_of_cmd (Continue) = {}" | "g_wvars_of_cmd (Iterate c1 c2) = (g_wvars_of_cmd c1 \ g_wvars_of_cmd c2)" | "g_wvars_of_cmd (Invalid) = {}" definition "udvars_of_cmd c \ ud_rvars_of_cmd c \ ud_wvars_of_cmd c" definition "lvars_of_cmd c \ l_rvars_of_cmd c \ l_wvars_of_cmd c" definition "gvars_of_cmd c \ g_rvars_of_cmd c \ g_wvars_of_cmd c" fun read_globals :: "action \ ident set" where "read_globals (AAssign c _ e) = gvars_of_exp c \ udvars_of_exp c \ gvars_of_exp e \ udvars_of_exp e" | "read_globals (AAssign_local c _ e) = gvars_of_exp c \ udvars_of_exp c \ gvars_of_exp e \ udvars_of_exp e" | "read_globals (AAssign_global c _ e) = gvars_of_exp c \ udvars_of_exp c \ gvars_of_exp e \ udvars_of_exp e" | "read_globals (ATest e) = gvars_of_exp e \ udvars_of_exp e" | "read_globals (ASkip) = {}" fun write_globals :: "action \ ident set" where "write_globals (AAssign _ x _) = {x}" | "write_globals (AAssign_local _ _ _) = {}" | "write_globals (AAssign_global _ x _) = {x}" | "write_globals (ATest e) = {}" | "write_globals (ASkip) = {}" abbreviation "rw_globals a \ read_globals a \ write_globals a" definition eq_on :: "ident set \ valuation \ valuation \ bool" where "eq_on X s1 s2 \ \x\X. s1 x = s2 x" lemma eq_onD: "\eq_on X s1 s2; x\X\ \ s1 x = s2 x" unfolding eq_on_def by auto lemma eq_on_refl[simp]: "eq_on X s s" and eq_on_sym: "eq_on X s1 s2 \ eq_on X s2 s1" and eq_on_sym_eq: "eq_on X s1 s2 \ eq_on X s2 s1" and eq_on_trans[trans]: "\eq_on Xa s1 s2; eq_on Xb s2 s3\ \ eq_on (Xa\Xb) s1 s3" and eq_on_UNIV_eq[simp]: "eq_on UNIV s1 s2 \ s1=s2" unfolding eq_on_def by auto lemma eq_on_join: "\eq_on X s s'; eq_on Y s s'\ \ eq_on (X\Y) s s'" unfolding eq_on_def by auto abbreviation ls_eq_on :: "ident set \ local_state \ local_state \ bool" where "ls_eq_on X ls1 ls2 \ eq_on X (local_state.variables ls1) (local_state.variables ls2)" abbreviation gs_eq_on :: "ident set \ global_state \ global_state \ bool" where "gs_eq_on X ls1 ls2 \ eq_on X (global_state.variables ls1) (global_state.variables ls2)" lemma eval_dep_vars: assumes "gs_eq_on X gs gs'" assumes "udvars_of_exp e \ X" "gvars_of_exp e \ X" shows "eval_exp e (ls,gs) = eval_exp e (ls,gs')" using assms apply (induction e) apply (auto split: option.splits simp: eq_on_def) done lemma en_dep_read: assumes "gs_eq_on X gs gs'" assumes "read_globals a \ X" shows "la_en' (ls,gs) a = la_en' (ls,gs') a" using assms apply (cases a) apply (auto simp: la_en'_def eval_dep_vars split: Option.bind_splits) done lemma ex_mod_limit: assumes "la_ex' (ls,gs) a = (ls',gs')" shows "gs_eq_on (-write_globals a) gs gs'" using assms apply (cases a) apply (auto simp: la_ex'_def eval_dep_vars eq_on_def - split: option.splits Option.bind_splits if_split_asm) + split: option.splits Option.bind_splits) + apply (auto split: if_split_asm) done lemma ex_dep_pres: assumes "gs_eq_on X gs1 gs2" assumes "read_globals a \ X" assumes "la_ex' (ls,gs1) a = (ls1',gs1')" assumes "la_ex' (ls,gs2) a = (ls2',gs2')" shows "ls1'=ls2' \ gs_eq_on X gs1' gs2'" using assms apply (cases a) apply (auto simp: la_ex'_def eval_dep_vars eq_on_def split: Option.bind_splits option.splits bool.splits if_split_asm) done lemma ex_swap_global: assumes DJ: "write_globals a1 \ read_globals a2 = {}" "write_globals a2 \ read_globals a1 = {}" "write_globals a1 \ write_globals a2 = {}" assumes EXA: "la_ex' (ls1,gs) a1 = (ls1a, gsa)" "la_ex' (ls2, gsa) a2 = (ls2a, gsa')" assumes EXB: "la_ex' (ls2,gs) a2 = (ls2b, gsb)" "la_ex' (ls1, gsb) a1 = (ls1b, gsb')" shows "ls1a=ls1b" "ls2a=ls2b" "gsa'=gsb'" proof - from ex_dep_pres[OF ex_mod_limit[OF EXB(1)] _ EXA(1) EXB(2)] DJ have "ls1a = ls1b" and E1: "gs_eq_on (- write_globals a2) gsa gsb'" by auto thus [simp]: "ls1a=ls1b" by simp from ex_dep_pres[OF ex_mod_limit[OF EXA(1), THEN eq_on_sym] _ EXA(2) EXB(1) ] DJ have "ls2a = ls2b" and E2: "gs_eq_on (- write_globals a1) gsa' gsb" by auto thus [simp]: "ls2a=ls2b" by simp from DJ have [simp]: "- write_globals a1 \ - write_globals a2 = UNIV" by auto from eq_on_trans[OF E2 ex_mod_limit[OF EXB(2)]] have "gs_eq_on (- write_globals a1) gsa' gsb'" by simp also from eq_on_trans[OF ex_mod_limit[OF EXA(2), THEN eq_on_sym] E1] have "gs_eq_on (- write_globals a2) gsa' gsb'" by simp finally (eq_on_join) have "gs_eq_on UNIV gsa' gsb'" by simp thus "gsa' = gsb'" apply (cases gsa', cases gsb') apply simp done qed lemma ex_pres_en: assumes EX: "la_ex' (ls1,gs) a1 = (ls1',gs')" assumes DJ: "write_globals a1 \ read_globals a2 = {}" shows "la_en' (ls2,gs') a2 = la_en' (ls2,gs) a2" using en_dep_read[OF ex_mod_limit[OF EX], of a2] DJ by auto end diff --git a/thys/Circus/Denotational_Semantics.thy b/thys/Circus/Denotational_Semantics.thy --- a/thys/Circus/Denotational_Semantics.thy +++ b/thys/Circus/Denotational_Semantics.thy @@ -1,1055 +1,1039 @@ section \Denotational semantics of Circus actions\ theory Denotational_Semantics imports Circus_Actions Var_list begin text \In this section, we introduce the definitions of Circus actions denotational semantics. We provide the proof of well-formedness of every action. We also provide proofs concerning the monotonicity of operators over actions.\ subsection \Skip\ definition Skip :: "('\::ev_eq,'\) action" where "Skip \ action_of (R (true \ \(A, A'). tr A' = tr A \ \wait A' \ more A = more A'))" lemma Skip_is_action: "(R (true \ \(A, A'). tr A' = tr A \ \wait A' \ more A = more A')) \ {p. is_CSP_process p}" apply (simp) apply (rule rd_is_CSP) by auto lemmas Skip_is_CSP = Skip_is_action[simplified] lemma relation_of_Skip: "relation_of Skip = (R (true \ \(A, A'). tr A' = tr A \ \wait A' \ more A = more A'))" by (simp add: Skip_def action_of_inverse Skip_is_CSP) definition CSP3::"(('\::ev_eq,'\) alphabet_rp) Healthiness_condition" where "CSP3 (P) \ relation_of Skip ;; P" definition CSP4::"(('\::ev_eq,'\) alphabet_rp) Healthiness_condition" where "CSP4 (P) \ P ;; relation_of Skip" lemma Skip_is_CSP3: "(relation_of Skip) is CSP3 healthy" apply (auto simp: relation_of_Skip rp_defs design_defs fun_eq_iff CSP3_def) apply (split cond_splits, simp_all)+ apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ prefer 3 apply (split cond_splits, simp_all)+ apply (auto simp add: prefix_def) done lemma Skip_is_CSP4: "(relation_of Skip) is CSP4 healthy" apply (auto simp: relation_of_Skip rp_defs design_defs fun_eq_iff CSP4_def) apply (split cond_splits, simp_all)+ apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ prefer 3 apply (split cond_splits, simp_all)+ apply (auto simp add: prefix_def) done lemma Skip_comp_absorb: "(relation_of Skip ;; relation_of Skip) = relation_of Skip" apply (auto simp: relation_of_Skip fun_eq_iff rp_defs true_def design_defs) apply (clarsimp split: cond_splits)+ apply (case_tac "ok aa", simp_all) apply (erule disjE)+ apply (clarsimp simp: prefix_def) apply (clarsimp simp: prefix_def) apply (erule disjE)+ apply (clarsimp simp: prefix_def) apply (clarsimp simp: prefix_def) apply (erule disjE)+ apply (clarsimp simp: prefix_def) apply (clarsimp simp: prefix_def) apply (case_tac "ok aa", simp_all) apply (clarsimp simp: prefix_def) apply (clarsimp split: cond_splits)+ apply (rule_tac b=a in comp_intro) apply (clarsimp split: cond_splits)+ apply (rule_tac b=a in comp_intro) apply (clarsimp split: cond_splits)+ done subsection \Stop\ definition Stop :: "('\::ev_eq,'\) action" where "Stop \ action_of (R (true \ \(A, A'). tr A' = tr A \ wait A'))" lemma Stop_is_action: "(R (true \ \(A, A'). tr A' = tr A \ wait A')) \ {p. is_CSP_process p}" apply (simp) apply (rule rd_is_CSP) by auto lemmas Stop_is_CSP = Stop_is_action[simplified] lemma relation_of_Stop: "relation_of Stop = (R (true \ \(A, A'). tr A' = tr A \ wait A'))" by (simp add: Stop_def action_of_inverse Stop_is_CSP) lemma Stop_is_CSP3: "(relation_of Stop) is CSP3 healthy" apply (auto simp: relation_of_Stop relation_of_Skip rp_defs design_defs fun_eq_iff CSP3_def) apply (rule_tac b=a in comp_intro) apply (split cond_splits, auto) apply (split cond_splits)+ apply (simp_all) apply (case_tac "ok aa", simp_all) apply (case_tac "tr aa \ tr ba", simp_all) apply (case_tac "ok ba", simp_all) apply (case_tac "tr ba \ tr c", simp_all) apply (rule disjI1) apply (simp add: prefix_def) apply (erule exE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (rule disjI1) apply (simp add: prefix_def) apply (erule exE | erule conjE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (split cond_splits)+ apply (simp_all add: true_def) apply (erule disjE) apply (simp add: prefix_def) apply (erule exE | erule conjE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (auto simp add: prefix_def) done lemma Stop_is_CSP4: "(relation_of Stop) is CSP4 healthy" apply (auto simp: relation_of_Stop relation_of_Skip rp_defs design_defs fun_eq_iff CSP4_def) apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ apply (case_tac "ok aa", simp_all) apply (case_tac "tr aa \ tr ba", simp_all) apply (case_tac "ok ba", simp_all) apply (case_tac "tr ba \ tr c", simp_all) apply (rule disjI1) apply (simp add: prefix_def) apply (erule exE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (rule disjI1) apply (simp add: prefix_def) apply (erule exE | erule conjE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (split cond_splits)+ apply (simp_all add: true_def) apply (erule disjE) apply (simp add: prefix_def) apply (erule exE | erule conjE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (auto simp add: prefix_def) done subsection \Chaos\ definition Chaos :: "('\::ev_eq,'\) action" where "Chaos \ action_of (R(false \ true))" lemma Chaos_is_action: "(R(false \ true)) \ {p. is_CSP_process p}" apply (simp) apply (rule rd_is_CSP) by auto lemmas Chaos_is_CSP = Chaos_is_action[simplified] lemma relation_of_Chaos: "relation_of Chaos = (R(false \ true))" by (simp add: Chaos_def action_of_inverse Chaos_is_CSP) subsection \State update actions\ definition Pre ::"'\ relation \ '\ predicate" where "Pre sc \ \A. \ A'. sc (A, A')" definition state_update_before :: "'\ relation \ ('\::ev_eq,'\) action \ ('\,'\) action" where "state_update_before sc Ac = action_of(R ((\(A, A'). (Pre sc) (more A)) \ (\(A, A'). sc (more A, more A') & \wait A' & tr A = tr A')) ;; relation_of Ac)" lemma state_update_before_is_action: "(R ((\(A, A'). (Pre sc) (more A)) \ (\(A, A').sc (more A, more A') & \wait A' & tr A = tr A')) ;; relation_of Ac) \ {p. is_CSP_process p}" apply (simp) apply (rule seq_CSP) apply (rule rd_is_CSP1) apply (auto simp: R_idem2 Healthy_def relation_of_CSP) done lemmas state_update_before_is_CSP = state_update_before_is_action[simplified] lemma relation_of_state_update_before: "relation_of (state_update_before sc Ac) = (R ((\(A, A'). (Pre sc) (more A)) \ (\(A, A'). sc (more A, more A') & \wait A' & tr A = tr A')) ;; relation_of Ac)" by (simp add: state_update_before_def action_of_inverse state_update_before_is_CSP) lemma mono_state_update_before: "mono (state_update_before sc)" by (auto simp: mono_def less_eq_action ref_def relation_of_state_update_before design_defs rp_defs fun_eq_iff split: cond_splits dest: relation_of_spec_f_f[simplified] relation_of_spec_t_f[simplified]) lemma state_update_before_is_CSP3: "relation_of (state_update_before sc Ac) is CSP3 healthy" apply (auto simp: relation_of_state_update_before relation_of_Skip rp_defs design_defs fun_eq_iff CSP3_def) apply (rule_tac b=aa in comp_intro) apply (split cond_splits, auto) apply (split cond_splits, simp_all)+ apply (rule_tac b=bb in comp_intro) apply (split cond_splits, simp_all)+ apply (case_tac "ok aa", simp_all) apply (case_tac "tr aa \ tr ab", simp_all) apply (case_tac "ok ab", simp_all) apply (case_tac "tr ab \ tr bb", simp_all) apply (rule disjI1) apply (simp add: prefix_def) apply (erule exE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (rule_tac b=bb in comp_intro) apply (split cond_splits, simp_all)+ apply (rule disjI1) apply (simp add: prefix_def) apply (erule exE | erule conjE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (rule_tac b=bb in comp_intro) apply (split cond_splits, simp_all)+ apply (simp_all add: true_def) apply (erule disjE) apply (simp add: prefix_def) apply (erule exE | erule conjE)+ apply (rule_tac x="zs@zsa" in exI, simp) apply (auto simp add: prefix_def) done lemma state_update_before_is_CSP4: assumes A : "relation_of Ac is CSP4 healthy" shows "relation_of (state_update_before sc Ac) is CSP4 healthy" apply (auto simp: relation_of_state_update_before relation_of_Skip rp_defs design_defs fun_eq_iff CSP4_def) apply (rule_tac b=c in comp_intro) apply (rule_tac b=ba in comp_intro, simp_all) apply (split cond_splits, simp_all) apply (rule_tac b=bb in comp_intro, simp_all) apply (subst A[simplified design_defs rp_defs CSP4_def relation_of_Skip]) apply (auto simp: rp_defs) done definition state_update_after :: "'\ relation \ ('\::ev_eq,'\) action \ ('\,'\) action" where "state_update_after sc Ac = action_of(relation_of Ac ;; R (true \ (\(A, A'). sc (more A, more A') & \wait A' & tr A = tr A')))" lemma state_update_after_is_action: "(relation_of Ac ;; R (true \ (\(A, A'). sc (more A, more A') & \wait A' & tr A = tr A'))) \ {p. is_CSP_process p}" apply (simp) apply (rule seq_CSP) apply (auto simp: relation_of_CSP[simplified is_CSP_process_def]) apply (rule rd_is_CSP, auto) done lemmas state_update_after_is_CSP = state_update_after_is_action[simplified] lemma relation_of_state_update_after: "relation_of (state_update_after sc Ac) = (relation_of Ac ;; R (true \ (\(A, A'). sc (more A, more A') & \wait A' & tr A = tr A')))" by (simp add: state_update_after_def action_of_inverse state_update_after_is_CSP) lemma mono_state_update_after: "mono (state_update_after sc)" by (auto simp: mono_def less_eq_action ref_def relation_of_state_update_after design_defs rp_defs fun_eq_iff split: cond_splits dest: relation_of_spec_f_f[simplified] relation_of_spec_t_f[simplified]) lemma state_update_after_is_CSP3: assumes A : "relation_of Ac is CSP3 healthy" shows "relation_of (state_update_after sc Ac) is CSP3 healthy" apply (auto simp: relation_of_state_update_after relation_of_Skip rp_defs design_defs fun_eq_iff CSP3_def) apply (rule_tac b=aa in comp_intro) apply (split cond_splits, auto) apply (rule_tac b=bb in comp_intro, simp_all) apply (subst A[simplified design_defs rp_defs CSP3_def relation_of_Skip]) apply (auto simp: rp_defs) done lemma state_update_after_is_CSP4: "relation_of (state_update_after sc Ac) is CSP4 healthy" apply (auto simp: relation_of_state_update_after relation_of_Skip rp_defs design_defs fun_eq_iff CSP4_def) apply (rule_tac b=c in comp_intro) apply (rule_tac b=ba in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (case_tac "ok bb", simp_all) apply (case_tac "tr bb \ tr c", simp_all) apply (case_tac "ok ca", simp_all) apply (case_tac "tr ca \ tr c", simp_all) -apply (simp add: prefix_def) -apply (erule exE)+ -apply (erule_tac x="zs@zsa" in allE, simp) -apply (rule_tac b=bb in comp_intro, simp_all) -apply (split cond_splits, simp_all add: true_def)+ -apply (case_tac "ok ca", simp_all) -apply (case_tac "tr ca \ tr c", simp_all) -apply (simp add: prefix_def) -apply (erule exE | erule conjE)+ -apply (rule_tac x="zsa@zs" in exI, simp) -apply (rule_tac b=bb in comp_intro, simp_all) -apply (split cond_splits, simp_all)+ -apply (case_tac "tr bb \ tr c", simp_all) -apply (simp add: prefix_def) -apply (erule exE | erule conjE)+ -apply (erule_tac x="zsa@zs" in allE, simp) -apply (auto simp add: prefix_def) +apply (auto simp add: prefix_def comp_def true_def split: cond_splits) done subsection \Sequential composition\ definition Seq::"('\::ev_eq,'\) action \ ('\,'\) action \ ('\,'\) action" (infixl "`;`" 24) where "P `;` Q \ action_of (relation_of P ;; relation_of Q)" lemma Seq_is_action: "(relation_of P ;; relation_of Q) \ {p. is_CSP_process p}" apply (simp) apply (rule seq_CSP[OF relation_of_CSP[THEN CSP_is_CSP1] relation_of_CSP[THEN CSP_is_R] relation_of_CSP]) done lemmas Seq_is_CSP = Seq_is_action[simplified] lemma relation_of_Seq: "relation_of (P `;` Q) = (relation_of P ;; relation_of Q)" by (simp add: Seq_def action_of_inverse Seq_is_CSP) lemma mono_Seq: "mono ((`;`) P)" by (auto simp: mono_def less_eq_action ref_def relation_of_Seq) lemma CSP3_imp_left_Skip: assumes A: "relation_of P is CSP3 healthy" shows "(Skip `;` P) = P" apply (subst relation_of_inject[symmetric]) apply (simp add: relation_of_Seq A[simplified design_defs CSP3_def, symmetric]) done lemma CSP4_imp_right_Skip: assumes A: "relation_of P is CSP4 healthy" shows "(P `;` Skip) = P" apply (subst relation_of_inject[symmetric]) apply (simp add: relation_of_Seq A[simplified design_defs CSP4_def, symmetric]) done lemma Seq_assoc: "(A `;` (B `;` C)) = ((A `;` B) `;` C)" by (auto simp: relation_of_inject[symmetric] fun_eq_iff relation_of_Seq rp_defs design_defs) lemma Skip_absorb: "(Skip `;` Skip) = Skip" by (auto simp: Skip_comp_absorb relation_of_inject[symmetric] relation_of_Seq) subsection \Internal choice\ definition Ndet::"('\::ev_eq,'\) action \ ('\,'\) action \ ('\,'\) action" (infixl "\" 18) where "P \ Q \ action_of ((relation_of P) \ (relation_of Q))" lemma Ndet_is_action: "((relation_of P) \ (relation_of Q)) \ {p. is_CSP_process p}" apply (simp) apply (rule disj_CSP) apply (simp_all add: relation_of_CSP) done lemmas Ndet_is_CSP = Ndet_is_action[simplified] lemma relation_of_Ndet: "relation_of (P \ Q) = ((relation_of P) \ (relation_of Q))" by (simp add: Ndet_def action_of_inverse Ndet_is_CSP) lemma mono_Ndet: "mono ((\) P)" by (auto simp: mono_def less_eq_action ref_def relation_of_Ndet) subsection \External choice\ definition Det::"('\::ev_eq,'\) action \ ('\,'\) action \ ('\,'\) action" (infixl "\" 18) where "P \ Q \ action_of(R((\((relation_of P)\<^sup>f\<^sub>f) \ \((relation_of Q)\<^sup>f\<^sub>f)) \ (((relation_of P)\<^sup>t\<^sub>f \ ((relation_of Q)\<^sup>t\<^sub>f)) \ \(A, A'). tr A = tr A' \ wait A' \ ((relation_of P)\<^sup>t\<^sub>f \ ((relation_of Q)\<^sup>t\<^sub>f)))))" lemma Det_is_action: "(R((\((relation_of P)\<^sup>f\<^sub>f) \ \((relation_of Q)\<^sup>f\<^sub>f)) \ (((relation_of P)\<^sup>t\<^sub>f \ ((relation_of Q)\<^sup>t\<^sub>f)) \ \(A, A'). tr A = tr A' \ wait A' \ ((relation_of P)\<^sup>t\<^sub>f \ ((relation_of Q)\<^sup>t\<^sub>f))))) \ {p. is_CSP_process p}" apply (simp add: spec_def) apply (rule rd_is_CSP) apply (auto) done lemmas Det_is_CSP = Det_is_action[simplified] lemma relation_of_Det: "relation_of (P \ Q) = (R((\((relation_of P)\<^sup>f\<^sub>f) \ \((relation_of Q)\<^sup>f\<^sub>f)) \ (((relation_of P)\<^sup>t\<^sub>f \ ((relation_of Q)\<^sup>t\<^sub>f)) \ \(A, A'). tr A = tr A' \ wait A' \ ((relation_of P)\<^sup>t\<^sub>f \ ((relation_of Q)\<^sup>t\<^sub>f)))))" apply (unfold Det_def) apply (rule action_of_inverse) apply (rule Det_is_action) done lemma mono_Det: "mono ((\) P)" by (auto simp: mono_def less_eq_action ref_def relation_of_Det design_defs rp_defs fun_eq_iff split: cond_splits dest: relation_of_spec_f_f[simplified] relation_of_spec_t_f[simplified]) subsection \Reactive design assignment\ definition "rd_assign s = action_of (R (true \ \(A, A'). ref A' = ref A \ tr A' = tr A \ \wait A' \ more A' = s))" lemma rd_assign_is_action: "(R (true \ \(A, A'). ref A' = ref A \ tr A' = tr A \ \wait A' \ more A' = s)) \ {p. is_CSP_process p}" apply (auto simp:) apply (rule rd_is_CSP) by auto lemmas rd_assign_is_CSP = rd_assign_is_action[simplified] lemma relation_of_rd_assign: "relation_of (rd_assign s) = (R (true \ \(A, A'). ref A' = ref A \ tr A' = tr A \ \wait A' \ more A' = s))" by (simp add: rd_assign_def action_of_inverse rd_assign_is_CSP) subsection \Local state external choice\ definition Loc::"'\ \ ('\::ev_eq,'\) action \ '\ \ ('\,'\) action \ ('\,'\) action" ("'(()loc _ \ _ ') \ '(()loc _ \ _ ')") where "(loc s1 \ P) \ (loc s2 \ Q) \ ((rd_assign s1)`;`P) \ ((rd_assign s2)`;` Q)" subsection \Schema expression\ definition Schema :: "'\ relation \ ('\::ev_eq,'\) action" where "Schema sc \ action_of(R ((\(A, A'). (Pre sc) (more A)) \ (\(A, A'). sc (more A, more A') \ \wait A' \ tr A = tr A')))" lemma Schema_is_action: "(R ((\(A, A'). (Pre sc) (more A)) \ (\(A, A'). sc (more A, more A') & \wait A' & tr A = tr A'))) \ {p. is_CSP_process p}" apply (simp) apply (rule rd_is_CSP) apply (auto) done lemmas Schema_is_CSP = Schema_is_action[simplified] lemma relation_of_Schema: "relation_of (Schema sc) = (R ((\(A, A'). (Pre sc) (more A)) \ (\(A, A'). sc (more A, more A') \ \wait A' \ tr A = tr A')))" by (simp add: Schema_def action_of_inverse Schema_is_CSP) lemma Schema_is_state_update_before: "Schema u = state_update_before u Skip" apply (subst relation_of_inject[symmetric]) apply (auto simp: relation_of_Schema relation_of_state_update_before relation_of_Skip rp_defs fun_eq_iff design_defs) apply (split cond_splits, simp_all) apply (rule comp_intro) apply (split cond_splits, simp_all)+ apply (rule comp_intro) apply (split cond_splits, simp_all)+ prefer 3 apply (split cond_splits, simp_all)+ apply (auto simp: prefix_def) done subsection \Parallel composition\ type_synonym '\ local_state = "('\ \ ('\ \ '\ \ '\))" fun MergeSt :: "'\ local_state \ '\ local_state \ ('\,'\) relation_rp" where "MergeSt (s1,s1') (s2,s2') = ((\(S, S'). (s1' s1) (more S) = more S');; (\(S::('\,'\) alphabet_rp, S'). (s2' s2) (more S) = more S'))" definition listCons ::"'\ \ '\ list list \ '\ list list" ("_ ## _") where "a ## l = ((map (Cons a)) l)" fun ParMergel :: "'\::ev_eq list \ '\ list \ '\ set \ '\ list list" where "ParMergel [] [] cs = [[]]" | "ParMergel [] (b#tr2) cs = (if (filter_chan_set b cs) then [[]] else (b ## (ParMergel [] tr2 cs)))" | "ParMergel (a#tr1) [] cs = (if (filter_chan_set a cs) then [[]] else (a ## (ParMergel tr1 [] cs)))" | "ParMergel (a#tr1) (b#tr2) cs = (if (filter_chan_set a cs) then (if (ev_eq a b) then (a ## (ParMergel tr1 tr2 cs)) else (if (filter_chan_set b cs) then [[]] else (b ## (ParMergel (a#tr1) tr2 cs)))) else (if (filter_chan_set b cs) then (a ## (ParMergel tr1 (b#tr2) cs)) else (a ## (ParMergel tr1 (b#tr2) cs)) @ (b ## (ParMergel (a#tr1) tr2 cs))))" definition ParMerge::"'\::ev_eq list \ '\ list \ '\ set \ '\ list set" where "ParMerge tr1 tr2 cs = set (ParMergel tr1 tr2 cs)" lemma set_Cons1: "tr1 \ set l \ a # tr1 \ (#) a ` set l" by (auto) lemma tr_in_set_eq: "(tr1 \ (#) b ` set l) = (tr1 \ [] \ hd tr1 = b \ tl tr1 \ set l)" by (induct l) auto definition M_par::"(('\::ev_eq), '\) alpha_rp_scheme \ ('\ \ '\ \ '\) \ ('\, '\) alpha_rp_scheme \ ('\ \ '\ \ '\) \ ('\ set) \ ('\, '\) relation_rp" where "M_par s1 x1 s2 x2 cs = ((\(S, S'). ((diff_tr S' S) \ ParMerge (diff_tr s1 S) (diff_tr s2 S) cs & ev_eq (tr_filter (tr s1) cs) (tr_filter (tr s2) cs))) \ ((\(S, S'). (wait s1 \ wait s2) \ ref S' \ ((((ref s1)\(ref s2))\cs)\(((ref s1)\(ref s2))-cs))) \ wait o snd \ ((\(S, S'). (\wait s1 \ \wait s2)) \ MergeSt ((more s1), x1) ((more s2), x2))))" definition Par::"('\::ev_eq,'\) action \ ('\ \ '\ \ '\) \ '\ set \ ('\ \ '\ \ '\) \ ('\,'\) action \ ('\,'\) action" ("_ \ _ | _ | _ \ _") where "A1 \ ns1 | cs | ns2 \ A2 \ (action_of (R ((\ (S, S'). \ (\ tr1 tr2. ((relation_of A1)\<^sup>f\<^sub>f ;; (\ (S, S'). tr1 = (tr S))) (S, S') \ (spec False (wait S) (relation_of A2) ;; (\ (S, _). tr2 = (tr S))) (S, S') \ ((tr_filter tr1 cs) = (tr_filter tr2 cs))) \ \ (\ tr1 tr2. (spec False (wait S) (relation_of A1);;(\(S, _). tr1 = tr S)) (S, S') \ ((relation_of A2)\<^sup>f\<^sub>f ;; (\(S, S'). tr2 = (tr S))) (S, S') \ ((tr_filter tr1 cs) = (tr_filter tr2 cs)))) \ (\ (S, S'). (\ s1 s2. ((\ (A, A'). (relation_of A1)\<^sup>t\<^sub>f (A, s1) \ ((relation_of A2)\<^sup>t\<^sub>f (A, s2)));; M_par s1 ns1 s2 ns2 cs) (S, S'))))))" lemma Par_is_action: "(R ((\ (S, S'). \ (\ tr1 tr2. ((relation_of A1)\<^sup>f\<^sub>f ;; (\ (S, S'). tr1 = (tr S))) (S, S') \ (spec False (wait S) (relation_of A2) ;; (\ (S, S'). tr2 = (tr S))) (S, S') \ ((tr_filter tr1 cs) = (tr_filter tr2 cs))) \ \ (\ tr1 tr2. (spec False (wait S) (relation_of A1);;(\(S, _). tr1 = tr S)) (S, S') \ ((relation_of A2)\<^sup>f\<^sub>f ;; (\ (S, S'). tr2 = (tr S))) (S, S') \ ((tr_filter tr1 cs) = (tr_filter tr2 cs)))) \ (\ (S, S'). (\ s1 s2. ((\ (A, A'). (relation_of A1)\<^sup>t\<^sub>f (A, s1) \ ((relation_of A2)\<^sup>t\<^sub>f (A, s2)));; M_par s1 ns1 s2 ns2 cs) (S, S'))))) \ {p. is_CSP_process p}" apply (simp) apply (rule rd_is_CSP) apply (blast) done lemmas Par_is_CSP = Par_is_action[simplified] lemma relation_of_Par: "relation_of (A1 \ ns1 | cs | ns2 \ A2) = (R ((\ (S, S'). \ (\ tr1 tr2. ((relation_of A1)\<^sup>f\<^sub>f ;; (\ (S, S'). tr1 = (tr S))) (S, S') \ (spec False (wait S) (relation_of A2) ;; (\ (S, S'). tr2 = (tr S))) (S, S') \ ((tr_filter tr1 cs) = (tr_filter tr2 cs))) \ \ (\ tr1 tr2. (spec False (wait S) (relation_of A1);;(\(S, _). tr1 = tr S)) (S, S') \ ((relation_of A2)\<^sup>f\<^sub>f ;; (\ (S, S'). tr2 = (tr S))) (S, S') \ ((tr_filter tr1 cs) = (tr_filter tr2 cs)))) \ (\ (S, S'). (\ s1 s2. ((\ (A, A'). (relation_of A1)\<^sup>t\<^sub>f (A, s1) \ ((relation_of A2)\<^sup>t\<^sub>f (A, s2)));; M_par s1 ns1 s2 ns2 cs) (S, S')))))" apply (unfold Par_def) apply (rule action_of_inverse) apply (rule Par_is_action) done lemma mono_Par: "mono (\Q. P \ ns1 | cs | ns2 \ Q)" apply (auto simp: mono_def less_eq_action ref_def relation_of_Par design_defs fun_eq_iff rp_defs split: cond_splits) apply (auto simp: rp_defs dest: relation_of_spec_f_f[simplified] relation_of_spec_t_f[simplified]) apply (erule_tac x="tr ba" in allE, auto) apply (erule notE) apply (auto dest: relation_of_spec_f_f relation_of_spec_t_f) done subsection \Local parallel block\ definition ParLoc::"'\ \ ('\ \ '\ \ '\) \ ('\::ev_eq, '\) action \ '\ set \ '\ \ ('\ \ '\ \ '\) \ ('\,'\) action \ ('\,'\) action" ("'(()par _ | _ \ _ ') \ _ \ '(()par _ | _ \ _ ')") where "(par s1 | ns1 \ P) \ cs \ (par s2 | ns2 \ Q) \ ((rd_assign s1)`;`P) \ ns1 | cs | ns2 \ ((rd_assign s2)`;` Q)" subsection \Assignment\ definition ASSIGN::"('v, '\) var_list \ ('\ \ 'v) \ ('\::ev_eq, '\) action" where "ASSIGN x e \ action_of (R (true \ (\ (S, S'). tr S' = tr S \ \wait S' \ (more S' = (update x (\_. (e (more S)))) (more S)))))" syntax "_assign"::"id \ ('\ \ 'v) \ ('\, '\) action" ("_ `:=` _") translations "y `:=` vv" => "CONST ASSIGN (VAR y) vv" lemma Assign_is_action: "(R (true \ (\ (S, S'). tr S' = tr S \ \wait S' \ (more S' = (update x (\_. (e (more S)))) (more S))))) \ {p. is_CSP_process p}" apply (simp) apply (rule rd_is_CSP) apply (blast) done lemmas Assign_is_CSP = Assign_is_action[simplified] lemma relation_of_Assign: "relation_of (ASSIGN x e) = (R (true \ (\ (S, S'). tr S' = tr S \ \wait S' \ (more S' = (update x (\_. (e (more S)))) (more S)))))" by (simp add: ASSIGN_def action_of_inverse Assign_is_CSP) lemma Assign_is_state_update_before: "ASSIGN x e = state_update_before (\ (s, s') . s' = (update x (\_. (e s))) s) Skip" apply (subst relation_of_inject[symmetric]) apply (auto simp: relation_of_Assign relation_of_state_update_before relation_of_Skip rp_defs fun_eq_iff Pre_def update_def design_defs) apply (split cond_splits, simp_all)+ apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=b in comp_intro) apply (split cond_splits, simp_all)+ defer apply (split cond_splits, simp_all)+ prefer 3 apply (split cond_splits, simp_all)+ apply (auto simp add: prefix_def) done subsection \Variable scope\ definition Var::"('v, '\) var_list \('\, '\) action \ ('\::ev_eq,'\) action" where "Var v A \ action_of( (R(true \ (\ (A, A'). \ a. tr A' = tr A \ \wait A' \ more A' = (increase v a (more A)))));; (relation_of A;; (R(true \ (\ (A, A'). tr A' = tr A \ \wait A' \ more A' = (decrease v (more A)))))))" syntax "_var"::"idt \ ('\, '\) action \ ('\, '\) action" ("var _ \ _" [1000] 999) translations "var y \ Act" => "CONST Var (VAR_LIST y) Act" lemma Var_is_action: "((R(true \ (\ (A, A'). \ a. tr A' = tr A \ \wait A' \ more A' = (increase v a (more A)))));; (relation_of A;; (R(true \ (\ (A, A'). tr A' = tr A \ \wait A' \ more A' = (decrease v (more A))))))) \ {p. is_CSP_process p}" apply (simp) apply (rule seq_CSP) prefer 3 apply (rule seq_CSP) apply (auto simp: relation_of_CSP1 relation_of_R) apply (rule rd_is_CSP) apply (auto simp: csp_defs rp_defs design_defs fun_eq_iff prefix_def increase_def decrease_def split: cond_splits) done lemmas Var_is_CSP = Var_is_action[simplified] lemma relation_of_Var: "relation_of (Var v A) = ((R(true \ (\ (A, A'). \ a. tr A' = tr A \ \wait A' \ more A' = (increase v a (more A)))));; (relation_of A;; (R(true \ (\ (A, A'). tr A' = tr A \ \wait A' \ more A' = (decrease v (more A)))))))" apply (simp only: Var_def) apply (rule action_of_inverse) apply (rule Var_is_action) done lemma mono_Var : "mono (Var x)" by (auto simp: mono_def less_eq_action ref_def relation_of_Var) definition Let::"('v, '\) var_list \('\, '\) action \ ('\::ev_eq,'\) action" where "Let v A \ action_of((relation_of A;; (R(true \ (\ (A, A'). tr A' = tr A \ \wait A' \ more A' = (decrease v (more A)))))))" syntax "_let"::"idt \ ('\, '\) action \ ('\, '\) action" ("let _ \ _" [1000] 999) translations "let y \ Act" => "CONST Let (VAR_LIST y) Act" lemma Let_is_action: "(relation_of A;; (R(true \ (\ (A, A'). tr A' = tr A \ \wait A' \ more A' = (decrease v (more A)))))) \ {p. is_CSP_process p}" apply (simp) apply (rule seq_CSP) apply (auto simp: relation_of_CSP1 relation_of_R) apply (rule rd_is_CSP) apply (auto) done lemmas Let_is_CSP = Let_is_action[simplified] lemma relation_of_Let: "relation_of (Let v A) = (relation_of A;; (R(true \ (\ (A, A'). tr A' = tr A \ \wait A' \ more A' = (decrease v (more A))))))" by (simp add: Let_def action_of_inverse Let_is_CSP) lemma mono_Let : "mono (Let x)" by (auto simp: mono_def less_eq_action ref_def relation_of_Let) lemma Var_is_state_update_before: "Var v A = state_update_before (\ (s, s'). \ a. s' = increase v a s) (Let v A)" apply (subst relation_of_inject[symmetric]) apply (auto simp: relation_of_Var relation_of_Let relation_of_state_update_before relation_of_Skip fun_eq_iff) apply (auto simp: rp_defs fun_eq_iff Pre_def design_defs) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ defer apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ defer apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ apply (case_tac "\A' a. A' = increase v a (alpha_rp.more aa)", simp_all add: true_def) apply (erule_tac x="increase v a (alpha_rp.more aa)" in allE) apply (erule_tac x="a" in allE, simp) apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ apply (rule_tac b=ab in comp_intro) apply (split cond_splits, simp_all)+ apply (case_tac "\A' a. A' = increase v a (alpha_rp.more aa)", simp_all add: true_def) apply (erule_tac x="increase v a (alpha_rp.more aa)" in allE) apply (erule_tac x="a" in allE, simp) apply (rule_tac b=bb in comp_intro, simp_all) apply (split cond_splits, simp_all)+ done lemma Let_is_state_update_after: "Let v A = state_update_after (\ (s, s'). s' = decrease v s) A" apply (subst relation_of_inject[symmetric]) apply (auto simp: relation_of_Var relation_of_Let relation_of_state_update_after relation_of_Skip fun_eq_iff) apply (auto simp: rp_defs fun_eq_iff Pre_def design_defs) apply (auto split: cond_splits) done subsection \Guarded action\ definition Guard::"'\ predicate \ ('\::ev_eq, '\) action \ ('\, '\) action" ("_ `&` _") where "g `&` P \ action_of(R (((g o more o fst) \ \ ((relation_of P)\<^sup>f\<^sub>f)) \ (((g o more o fst) \ ((relation_of P)\<^sup>t\<^sub>f)) \ ((\(g o more o fst)) \ (\ (A, A'). tr A' = tr A \ wait A')))))" lemma Guard_is_action: "(R ( ((g o more o fst) \ \ ((relation_of P)\<^sup>f\<^sub>f)) \ (((g o more o fst) \ ((relation_of P)\<^sup>t\<^sub>f)) \ ((\(g o more o fst)) \ (\ (A, A'). tr A' = tr A \ wait A'))))) \ {p. is_CSP_process p}" by (auto simp add: spec_def intro: rd_is_CSP) lemmas Guard_is_CSP = Guard_is_action[simplified] lemma relation_of_Guard: "relation_of (g `&` P) = (R (((g o more o fst) \ \ ((relation_of P)\<^sup>f\<^sub>f)) \ (((g o more o fst) \ ((relation_of P)\<^sup>t\<^sub>f)) \ ((\(g o more o fst)) \ (\ (A, A'). tr A' = tr A \ wait A')))))" apply (unfold Guard_def) apply (subst action_of_inverse) apply (simp_all only: Guard_is_action) done lemma mono_Guard : "mono (Guard g)" apply (auto simp: mono_def less_eq_action ref_def rp_defs design_defs relation_of_Guard split: cond_splits) apply (auto dest: relation_of_spec_f_f relation_of_spec_t_f) done lemma false_Guard: "false `&` P = Stop" apply (subst relation_of_inject[symmetric]) apply (subst relation_of_Stop) apply (subst relation_of_Guard) apply (simp add: fun_eq_iff utp_defs csp_defs design_defs rp_defs) done lemma false_Guard1: "\ a b. g (alpha_rp.more a) = False \ (relation_of (g `&` P)) (a, b) = (relation_of Stop) (a, b)" apply (subst relation_of_Guard) apply (subst relation_of_Stop) apply (auto simp: fun_eq_iff csp_defs design_defs rp_defs split: cond_splits) done lemma true_Guard: "true `&` P = P" apply (subst relation_of_inject[symmetric]) apply (subst relation_of_Guard) apply (subst CSP_is_rd[OF relation_of_CSP]) back back apply (simp add: fun_eq_iff utp_defs csp_defs design_defs rp_defs) done lemma true_Guard1: "\ a b. g (alpha_rp.more a) = True \ (relation_of (g `&` P)) (a, b) = (relation_of P) (a, b)" apply (subst relation_of_Guard) apply (subst CSP_is_rd[OF relation_of_CSP]) back back apply (auto simp: fun_eq_iff csp_defs design_defs rp_defs split: cond_splits) done lemma Guard_is_state_update_before: "g `&` P = state_update_before (\ (s, s') . g s) P" apply (subst relation_of_inject[symmetric]) apply (auto simp: relation_of_Guard relation_of_state_update_before relation_of_Skip rp_defs fun_eq_iff Pre_def update_def design_defs) apply (rule_tac b=a in comp_intro) apply (split cond_splits, simp_all)+ apply (subst CSP_is_rd) apply (simp_all add: relation_of_CSP rp_defs design_defs fun_eq_iff) apply (split cond_splits, simp_all)+ apply (auto) apply (subst (asm) CSP_is_rd) apply (simp_all add: relation_of_CSP rp_defs design_defs fun_eq_iff) apply (split cond_splits, simp_all)+ apply (subst (asm) CSP_is_rd) apply (simp_all add: relation_of_CSP rp_defs design_defs fun_eq_iff) apply (split cond_splits, simp_all)+ apply (subst CSP_is_rd) apply (simp_all add: relation_of_CSP rp_defs design_defs fun_eq_iff) apply (split cond_splits, simp_all)+ apply (subst CSP_is_rd) apply (simp_all add: relation_of_CSP rp_defs design_defs fun_eq_iff) apply (split cond_splits, simp_all)+ apply (auto) defer apply (split cond_splits, simp_all)+ apply (subst (asm) CSP_is_rd) apply (simp_all add: relation_of_CSP rp_defs design_defs fun_eq_iff) apply (split cond_splits, simp_all)+ defer apply (rule disjI1) defer apply (case_tac "g (alpha_rp.more aa)", simp_all) apply (rule)+ apply (simp add: impl_def) defer oops subsection \Prefixed action\ definition do where "do e \ (\(A, A'). tr A = tr A' \ (e (more A)) \ (ref A')) \ wait o snd \ (\(A, A'). tr A' = (tr A @[(e (more A))]))" definition do_I::"('\ \'\) \ '\ set \ ('\, '\) relation_rp" where "do_I c S \ ((\(A, A'). tr A = tr A' & S \ (ref A') = {}) \ wait o snd \ (\(A, A'). hd (tr A' - tr A) \ S & (c (more A) = (last (tr A')))))" (* definition do_I::"('v \ '\) \ ('v, '\) var_list \ 'v set \ ('\, '\) relation_rp" where "do_I c x P \ ((\(A, A'). tr A = tr A' \ (c`P) \ (ref A') = {}) \ wait o fst \ (\(A, A'). hd (tr A' - tr A) \ (c`P) \ (c (select x (more A)) = (last (tr A')))))" *) definition iPrefix::"('\ \'\::ev_eq) \ ('\ relation) \ (('\, '\) action \ ('\, '\) action) \ ('\ \ '\ set) \ ('\, '\) action \ ('\, '\) action" where "iPrefix c i j S P \ action_of(R(true \ (\ (A, A'). (do_I c (S (more A))) (A, A') & more A' = more A)))`;` P" definition oPrefix::"('\ \'\) \ ('\::ev_eq, '\) action \ ('\, '\) action" where "oPrefix c P \ action_of(R(true \ (do c) \ (\ (A, A'). more A' = more A)))`;` P" definition Prefix0::"'\ \ ('\::ev_eq, '\) action \ ('\, '\) action" where "Prefix0 c P \ action_of(R(true \ (do (\ _. c)) \ (\ (A, A'). more A' = more A)))`;` P" definition read::"('v \ '\) \ ('v, '\) var_list \ ('\::ev_eq, '\) action \ ('\, '\) action" where "read c x P \ iPrefix (\ A. c (select x A)) (\ (s, s'). \ a. s' = increase x a s) (Let x) (\_. range c) P" definition read1::"('v \ '\) \ ('v, '\) var_list \ ('\ \ 'v set) \ ('\::ev_eq, '\) action \ ('\, '\) action" where "read1 c x S P \ iPrefix (\ A. c (select x A)) (\ (s, s'). \ a. a\(S s) & s' = increase x a s) (Let x) (\s. c`(S s)) P" definition write1::"('v \ '\) \ ('\ \ 'v) \ ('\::ev_eq, '\) action \ ('\, '\) action" where "write1 c a P \ oPrefix (\ A. c (a A)) P" definition write0::"'\ \ ('\::ev_eq, '\) action \ ('\, '\) action" where "write0 c P \ Prefix0 c P" syntax "_read" ::"[id, pttrn, ('\, '\) action] => ('\, '\) action" ("(_`?`_ /\ _)") "_readS" ::"[id, pttrn, '\ set,('\, '\) action] => ('\, '\) action" ("(_`?`_`:`_ /\ _)") "_readSS" ::"[id, pttrn, '\ => '\ set,('\, '\) action] => ('\, '\) action" ("(_`?`_`\`_ /\ _)") "_write" ::"[id, '\, ('\, '\) action] => ('\, '\) action" ("(_`!`_ /\ _)") "_writeS"::"['\, ('\, '\) action] => ('\, '\) action" ("(_ /\ _)") translations "_read c p P" == "CONST read c (VAR_LIST p) P" "_readS c p b P" == "CONST read1 c (VAR_LIST p) (\_. b) P" "_readSS c p b P" == "CONST read1 c (VAR_LIST p) b P" "_write c p P" == "CONST write1 c p P" "_writeS a P" == "CONST write0 a P" lemma Prefix_is_action: "(R(true \ (do c) \ (\ (A, A'). more A' = more A))) \ {p. is_CSP_process p}" by (auto intro: rd_is_CSP) lemma Prefix1_is_action: "(R(true \ \(A, A'). do_I c (S (alpha_rp.more A)) (A, A') \ alpha_rp.more A' = alpha_rp.more A)) \ {p. is_CSP_process p}" by (auto intro: rd_is_CSP) lemma Prefix0_is_action: "(R(true \ (do c) \ (\ (A, A'). more A' = more A))) \ {p. is_CSP_process p}" by (auto intro: rd_is_CSP) lemmas Prefix_is_CSP = Prefix_is_action[simplified] lemmas Prefix1_is_CSP = Prefix1_is_action[simplified] lemmas Prefix0_is_CSP = Prefix0_is_action[simplified] lemma relation_of_iPrefix: "relation_of (iPrefix c i j S P) = ((R(true \ (\ (A, A'). (do_I c (S (more A))) (A, A') & more A' = more A)));; relation_of P)" by (simp add: iPrefix_def relation_of_Seq action_of_inverse Prefix1_is_CSP) lemma relation_of_oPrefix: "relation_of (oPrefix c P) = ((R(true \ (do c) \ (\ (A, A'). more A' = more A)));; relation_of P)" by (simp add: oPrefix_def relation_of_Seq action_of_inverse Prefix_is_CSP) lemma relation_of_Prefix0: "relation_of (Prefix0 c P) = ((R(true \ (do (\ _. c)) \ (\ (A, A'). more A' = more A)));; relation_of P)" by (simp add: Prefix0_def relation_of_Seq action_of_inverse Prefix0_is_CSP) lemma mono_iPrefix : "mono (iPrefix c i j s)" by (auto simp: mono_def less_eq_action ref_def relation_of_iPrefix) lemma mono_oPrefix : "mono (oPrefix c)" by (auto simp: mono_def less_eq_action ref_def relation_of_oPrefix) lemma mono_Prefix0 : "mono(Prefix0 c)" by (auto simp: mono_def less_eq_action ref_def relation_of_Prefix0) subsection \Hiding\ definition Hide::"('\::ev_eq, '\) action \ '\ set \ ('\, '\) action" (infixl "\\" 18) where "P \\ cs \ action_of(R(\(S, S'). \ s. (diff_tr S' S) = (tr_filter (s - (tr S)) cs) & (relation_of P)(S, S'\tr := s, ref := (ref S') \ cs \));; (relation_of Skip))" definition "hid P cs == (R(\(S, S'). \ s. (diff_tr S' S) = (tr_filter (s - (tr S)) cs) & (relation_of P)(S, S'\tr := s, ref := (ref S') \ cs \)) ;; (relation_of Skip))" lemma hid_is_R: "hid P cs is R healthy" apply (simp add: hid_def) apply (rule seq_R) apply (simp add: Healthy_def R_idem2) apply (rule CSP_is_R) apply (rule relation_of_CSP) done lemma hid_Skip: "hid P cs = (hid P cs ;; relation_of Skip)" by (simp add: hid_def comp_assoc[symmetric] Skip_comp_absorb) lemma hid_is_CSP1: "hid P cs is CSP1 healthy" apply (auto simp: design_defs CSP1_def hid_def rp_defs fun_eq_iff) apply (rule_tac b="a" in comp_intro) apply (clarsimp split: cond_splits) apply (subst CSP_is_rd, auto simp: rp_defs relation_of_CSP design_defs fun_eq_iff split: cond_splits) apply (auto simp: diff_tr_def relation_of_Skip rp_defs design_defs true_def split: cond_splits) apply (rule_tac x="[]" in exI, auto) done lemma hid_is_CSP2: "hid P cs is CSP2 healthy" apply (simp add: hid_def) apply (rule seq_CSP2) apply (rule CSP_is_CSP2) apply (rule relation_of_CSP) done lemma hid_is_CSP: "is_CSP_process (hid P cs)" by (auto simp: csp_defs hid_is_CSP1 hid_is_R hid_is_CSP2) lemma Hide_is_action: "(R(\(S, S'). \ s. (diff_tr S' S) = (tr_filter (s - (tr S)) cs) & (relation_of P)(S, S'\tr := s, ref := (ref S') \ cs \));; (relation_of Skip)) \ {p. is_CSP_process p}" by (simp add: hid_is_CSP[simplified hid_def]) lemmas Hide_is_CSP = Hide_is_action[simplified] lemma relation_of_Hide: "relation_of (P \\ cs) = (R(\(S, S'). \ s. (diff_tr S' S) = (tr_filter (s - (tr S)) cs) & (relation_of P)(S, S'\tr :=s, ref := (ref S') \ cs \));; (relation_of Skip))" by (simp add: Hide_def action_of_inverse Hide_is_CSP) lemma mono_Hide : "mono(\ P. P \\ cs)" by (auto simp: mono_def less_eq_action ref_def prefix_def utp_defs relation_of_Hide rp_defs) subsection \Recursion\ text \To represent the recursion operator "\\\" over actions, we use the universal least fix-point operator "@{const lfp}" defined in the HOL library for lattices. The operator "@{const lfp}" is inherited from the "Complete Lattice class" under some conditions. All theorems defined over this operator can be reused.\ text \In the @{theory Circus.Circus_Actions} theory, we presented the proof that Circus actions form a complete lattice. The Knaster-Tarski Theorem (in its simplest formulation) states that any monotone function on a complete lattice has a least fixed-point. This is a consequence of the basic boundary properties of the complete lattice operations. Instantiating the complete lattice class allows one to inherit these properties with the definition of the least fixed-point for monotonic functions over Circus actions. \ syntax "_MU"::"[idt, idt \ ('\, '\) action] \ ('\, '\) action" ("\ _ \ _") translations "_MU X P" == "CONST lfp (\ X. P)" (*<*) text\Instead fo the following:\ lemma is_action_REP_Mu: shows "is_CSP_process (relation_of (lfp P))" oops text\... we refer to the proof of @{thm Sup_is_action} and its analogue who capture the essence of this proof at the level of the type instantiation.\ text\Monotonicity: STATUS: probably critical. Does not seem to be necessary for parameterless Circus.\ lemma mono_Mu: assumes A : "mono P" and B : "\ X. mono (P X)" shows "mono (lfp P)" apply (subst lfp_unfold) apply (auto simp: A B) done term Nat.Suc (*>*) end diff --git a/thys/Farkas/Farkas.thy b/thys/Farkas/Farkas.thy --- a/thys/Farkas/Farkas.thy +++ b/thys/Farkas/Farkas.thy @@ -1,1633 +1,1633 @@ (* Authors: R. Bottesch, M. W. Haslbeck, R. Thiemann *) section \Farkas Coefficients via the Simplex Algorithm of Duterte and de~Moura\ text \Let $c_1,\dots,c_n$ be a finite list of linear inequalities. Let $C$ be a list of pairs $(r_i,c_i)$ where $r_i$ is a rational number. We say that $C$ is a list of \emph{Farkas coefficients} if the sum of all products $r_i \cdot c_i$ results in an inequality that is trivially unsatisfiable. Farkas' Lemma states that a finite set of non-strict linear inequalities is unsatisfiable if and only if Farkas coefficients exist. We will prove this lemma with the help of the simplex algorithm of Dutertre and de~Moura's. Note that the simplex implementation works on four layers, and we will formulate and prove a variant of Farkas' Lemma for each of these layers.\ theory Farkas imports Simplex.Simplex begin subsection \Linear Inequalities\ text \Both Farkas' Lemma and Motzkin's Transposition Theorem require linear combinations of inequalities. To this end we define one type that permits strict and non-strict inequalities which are always of the form ``polynomial R constant'' where R is either $\leq$ or $<$. On this type we can then define a commutative monoid.\ text \A type for the two relations: less-or-equal and less-than.\ datatype le_rel = Leq_Rel | Lt_Rel primrec rel_of :: "le_rel \ 'a :: lrv \ 'a \ bool" where "rel_of Leq_Rel = (\)" | "rel_of Lt_Rel = (<)" instantiation le_rel :: comm_monoid_add begin definition "zero_le_rel = Leq_Rel" fun plus_le_rel where "plus_le_rel Leq_Rel Leq_Rel = Leq_Rel" | "plus_le_rel _ _ = Lt_Rel" instance proof fix a b c :: le_rel show "a + b + c = a + (b + c)" by (cases a; cases b; cases c, auto) show "a + b = b + a" by (cases a; cases b, auto) show "0 + a = a" unfolding zero_le_rel_def by (cases a, auto) qed end lemma Leq_Rel_0: "Leq_Rel = 0" unfolding zero_le_rel_def by simp datatype 'a le_constraint = Le_Constraint (lec_rel: le_rel) (lec_poly: linear_poly) (lec_const: 'a) abbreviation (input) "Leqc \ Le_Constraint Leq_Rel" instantiation le_constraint :: (lrv) comm_monoid_add begin fun plus_le_constraint :: "'a le_constraint \ 'a le_constraint \ 'a le_constraint" where "plus_le_constraint (Le_Constraint r1 p1 c1) (Le_Constraint r2 p2 c2) = (Le_Constraint (r1 + r2) (p1 + p2) (c1 + c2))" definition zero_le_constraint :: "'a le_constraint" where "zero_le_constraint = Leqc 0 0" instance proof fix a b c :: "'a le_constraint" show "0 + a = a" by (cases a, auto simp: zero_le_constraint_def Leq_Rel_0) show "a + b = b + a" by (cases a; cases b, auto simp: ac_simps) show "a + b + c = a + (b + c)" by (cases a; cases b; cases c, auto simp: ac_simps) qed end primrec satisfiable_le_constraint :: "'a::lrv valuation \ 'a le_constraint \ bool" (infixl "\\<^sub>l\<^sub>e" 100) where "(v \\<^sub>l\<^sub>e (Le_Constraint rel l r)) \ (rel_of rel (l\v\) r)" lemma satisfies_zero_le_constraint: "v \\<^sub>l\<^sub>e 0" by (simp add: valuate_zero zero_le_constraint_def) lemma satisfies_sum_le_constraints: assumes "v \\<^sub>l\<^sub>e c" "v \\<^sub>l\<^sub>e d" shows "v \\<^sub>l\<^sub>e (c + d)" proof - obtain lc rc ld rd rel1 rel2 where cd: "c = Le_Constraint rel1 lc rc" "d = Le_Constraint rel2 ld rd" by (cases c; cases d, auto) have 1: "rel_of rel1 (lc\v\) rc" using assms cd by auto have 2: "rel_of rel2 (ld\v\) rd" using assms cd by auto from 1 have le1: "lc\v\ \ rc" by (cases rel1, auto) from 2 have le2: "ld\v\ \ rd" by (cases rel2, auto) from 1 2 le1 le2 have "rel_of (rel1 + rel2) ((lc\v\) + (ld\v\)) (rc + rd)" apply (cases rel1; cases rel2; simp add: add_mono) by (metis add.commute le_less_trans order.strict_iff_order plus_less)+ thus ?thesis by (auto simp: cd valuate_add) qed lemma satisfies_sumlist_le_constraints: assumes "\ c. c \ set (cs :: 'a :: lrv le_constraint list) \ v \\<^sub>l\<^sub>e c" shows "v \\<^sub>l\<^sub>e sum_list cs" using assms by (induct cs, auto intro: satisfies_zero_le_constraint satisfies_sum_le_constraints) lemma sum_list_lec: "sum_list ls = Le_Constraint (sum_list (map lec_rel ls)) (sum_list (map lec_poly ls)) (sum_list (map lec_const ls))" proof (induct ls) case Nil show ?case by (auto simp: zero_le_constraint_def Leq_Rel_0) next case (Cons l ls) show ?case by (cases l, auto simp: Cons) qed lemma sum_list_Leq_Rel: "((\x\C. lec_rel (f x)) = Leq_Rel) \ (\ x \ set C. lec_rel (f x) = Leq_Rel)" proof (induct C) case (Cons c C) show ?case proof (cases "lec_rel (f c)") case Leq_Rel show ?thesis using Cons by (simp add: Leq_Rel Leq_Rel_0) qed simp qed (simp add: Leq_Rel_0) subsection \Farkas' Lemma on Layer 4\ text \On layer 4 the algorithm works on a state containing a tableau, atoms (or bounds), an assignment and a satisfiability flag. Only non-strict inequalities appear at this level. In order to even state a variant of Farkas' Lemma on layer 4, we need conversions from atoms to non-strict constraints and then further to linear inequalities of type @{type le_constraint}. The latter conversion is a partial operation, since non-strict constraints of type @{type ns_constraint} permit greater-or-equal constraints, whereas @{type le_constraint} allows only less-or-equal.\ text \The advantage of first going via @{type ns_constraint} is that this type permits a multiplication with arbitrary rational numbers (the direction of the inequality must be flipped when multiplying by a negative number, which is not possible with @{type le_constraint}).\ instantiation ns_constraint :: (scaleRat) scaleRat begin fun scaleRat_ns_constraint :: "rat \ 'a ns_constraint \ 'a ns_constraint" where "scaleRat_ns_constraint r (LEQ_ns p c) = (if (r < 0) then GEQ_ns (r *R p) (r *R c) else LEQ_ns (r *R p) (r *R c))" | "scaleRat_ns_constraint r (GEQ_ns p c) = (if (r > 0) then GEQ_ns (r *R p) (r *R c) else LEQ_ns (r *R p) (r *R c))" instance .. end lemma sat_scale_rat_ns: assumes "v \\<^sub>n\<^sub>s ns" shows "v \\<^sub>n\<^sub>s (f *R ns)" proof - have "f < 0 | f = 0 | f > 0" by auto then show ?thesis using assms by (cases ns, auto simp: valuate_scaleRat scaleRat_leq1 scaleRat_leq2) qed lemma scaleRat_scaleRat_ns_constraint: assumes "a \ 0 \ b \ 0" shows "a *R (b *R (c :: 'a :: lrv ns_constraint)) = (a * b) *R c" proof - have "b > 0 \ b < 0 \ b = 0" by linarith moreover have "a > 0 \ a < 0 \ a = 0" by linarith ultimately show ?thesis using assms by (elim disjE; cases c, auto simp add: not_le not_less mult_neg_pos mult_neg_neg mult_nonpos_nonneg mult_nonpos_nonpos mult_nonneg_nonpos mult_pos_neg) qed fun lec_of_nsc where "lec_of_nsc (LEQ_ns p c) = (Leqc p c)" fun is_leq_ns where "is_leq_ns (LEQ_ns p c) = True" | "is_leq_ns (GEQ_ns p c) = False" lemma lec_of_nsc: assumes "is_leq_ns c" shows "(v \\<^sub>l\<^sub>e lec_of_nsc c) \ (v \\<^sub>n\<^sub>s c)" using assms by (cases c, auto) fun nsc_of_atom where "nsc_of_atom (Leq var b) = LEQ_ns (lp_monom 1 var) b" | "nsc_of_atom (Geq var b) = GEQ_ns (lp_monom 1 var) b" lemma nsc_of_atom: "v \\<^sub>n\<^sub>s nsc_of_atom a \ v \\<^sub>a a" by (cases a, auto) text \We say that $C$ is a list of Farkas coefficients \emph{for a given tableau $t$ and atom set $as$}, if it is a list of pairs $(r,a)$ such that $a \in as$, $r$ is non-zero, $r \cdot a$ is a `less-than-or-equal'-constraint, and the linear combination of inequalities must result in an inequality of the form $p \leq c$, where $c < 0$ and $t \models p = 0$.\ definition farkas_coefficients_atoms_tableau where "farkas_coefficients_atoms_tableau (as :: 'a :: lrv atom set) t C = (\ p c. (\(r,a) \ set C. a \ as \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0) \ (\(r,a) \ C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c \ c < 0 \ (\ v :: 'a valuation. v \\<^sub>t t \(p\v\ = 0)))" text \We first prove that if the check-function detects a conflict, then Farkas coefficients do exist for the tableau and atom set for which the conflict is detected.\ definition bound_atoms :: "('i, 'a) state \ 'a atom set" ("\\<^sub>A") where "bound_atoms s = (\(v,x). Geq v x) ` (set_of_map (\\<^sub>l s)) \ (\(v,x). Leq v x) ` (set_of_map (\\<^sub>u s))" context PivotUpdateMinVars begin lemma farkas_check: assumes check: "check s' = s" and U: "\ s" "\ \ s'" and inv: "\ s'" "\ (\ s')" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" "\ s'" and index: "index_valid as s'" shows "\ C. farkas_coefficients_atoms_tableau (snd ` as) (\ s') C" proof - let ?Q = "\ s f p c C. set C \ \\<^sub>A s \ distinct C \ (\a \ set C. is_leq_ns (f (atom_var a) *R nsc_of_atom a) \ f (atom_var a) \ 0) \ (\a \ C. lec_of_nsc (f (atom_var a) *R nsc_of_atom a)) = Leqc p c \ c < 0 \ (\ v :: 'a valuation. v \\<^sub>t \ s \(p\v\ = 0))" let ?P = "\ s. \ s \ (\ f p c C. ?Q s f p c C)" have "?P (check s')" proof (induct rule: check_induct''[OF inv, of ?P]) case (3 s x\<^sub>i dir I) have dir: "dir = Positive \ dir = Negative" by fact let ?eq = "(eq_for_lvar (\ s) x\<^sub>i)" define X\<^sub>j where "X\<^sub>j = rvars_eq ?eq" define XL\<^sub>j where "XL\<^sub>j = Abstract_Linear_Poly.vars_list (rhs ?eq)" have [simp]: "set XL\<^sub>j = X\<^sub>j" unfolding XL\<^sub>j_def X\<^sub>j_def using set_vars_list by blast have XL\<^sub>j_distinct: "distinct XL\<^sub>j" unfolding XL\<^sub>j_def using distinct_vars_list by simp define A where "A = coeff (rhs ?eq)" have bounds_id: "\\<^sub>A (set_unsat I s) = \\<^sub>A s" "\\<^sub>u (set_unsat I s) = \\<^sub>u s" "\\<^sub>l (set_unsat I s) = \\<^sub>l s" by (auto simp: boundsl_def boundsu_def bound_atoms_def) have t_id: "\ (set_unsat I s) = \ s" by simp have u_id: "\ (set_unsat I s) = True" by simp let ?p = "rhs ?eq - lp_monom 1 x\<^sub>i" have p_eval_zero: "?p \ v \ = 0" if "v \\<^sub>t \ s" for v :: "'a valuation" proof - have eqT: "?eq \ set (\ s)" by (simp add: 3(7) eq_for_lvar local.min_lvar_not_in_bounds_lvars) have "v \\<^sub>e ?eq" using that eqT satisfies_tableau_def by blast also have "?eq = (lhs ?eq, rhs ?eq)" by (cases ?eq, auto) also have "lhs ?eq = x\<^sub>i" by (simp add: 3(7) eq_for_lvar local.min_lvar_not_in_bounds_lvars) finally have "v \\<^sub>e (x\<^sub>i, rhs ?eq)" . then show ?thesis by (auto simp: satisfies_eq_iff valuate_minus) qed have Xj_rvars: "X\<^sub>j \ rvars (\ s)" unfolding X\<^sub>j_def using 3 min_lvar_not_in_bounds_lvars rvars_of_lvar_rvars by blast have xi_lvars: "x\<^sub>i \ lvars (\ s)" using 3 min_lvar_not_in_bounds_lvars rvars_of_lvar_rvars by blast have "lvars (\ s) \ rvars (\ s) = {}" using 3 normalized_tableau_def by auto with xi_lvars Xj_rvars have xi_Xj: "x\<^sub>i \ X\<^sub>j" by blast have rhs_eval_xi: "(rhs (eq_for_lvar (\ s) x\<^sub>i)) \\\ s\\ = \\ s\ x\<^sub>i" proof - have *: "(rhs eq) \ v \ = v (lhs eq)" if "v \\<^sub>e eq" for v :: "'a valuation" and eq using satisfies_eq_def that by metis moreover have "\\ s\ \\<^sub>e eq_for_lvar (\ s) x\<^sub>i" using 3 satisfies_tableau_def eq_for_lvar curr_val_satisfies_no_lhs_def xi_lvars by blast ultimately show ?thesis using eq_for_lvar xi_lvars by simp qed let ?\\<^sub>l = "Direction.LB dir" let ?\\<^sub>u = "Direction.UB dir" let ?lt = "Direction.lt dir" let ?le = "Simplex.le ?lt" let ?Geq = "Direction.GE dir" let ?Leq = "Direction.LE dir" have 0: "(if A x < 0 then ?\\<^sub>l s x = Some (\\ s\ x) else ?\\<^sub>u s x = Some (\\ s\ x)) \ A x \ 0" if x: "x \ X\<^sub>j" for x proof - have "Some (\\ s\ x) = (?\\<^sub>l s x)" if "A x < 0" proof - have cmp: "\ \\<^sub>l\<^sub>b ?lt (\\ s\ x) (?\\<^sub>l s x)" using x that dir min_rvar_incdec_eq_None[OF 3(9)] unfolding X\<^sub>j_def A_def by auto then obtain c where c: "?\\<^sub>l s x = Some c" by (cases "?\\<^sub>l s x", auto simp: bound_compare_defs) also have "c = \\ s\ x" proof - have "x \ rvars (\ s)" using that x Xj_rvars by blast then have "x \ (- lvars (\ s))" using 3 unfolding normalized_tableau_def by auto moreover have "\x\(- lvars (\ s)). in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" using 3 unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_bounds_set.simps) ultimately have "in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" by blast moreover have "?le (\\ s\ x) c" using cmp c dir unfolding bound_compare_defs by auto ultimately show ?thesis - using c dir by auto + using c dir by (auto simp del: Simplex.bounds_lg) qed then show ?thesis using c by simp qed moreover have "Some (\\ s\ x) = (?\\<^sub>u s x)" if "0 < A x" proof - have cmp: "\ \\<^sub>u\<^sub>b ?lt (\\ s\ x) (?\\<^sub>u s x)" using x that min_rvar_incdec_eq_None[OF 3(9)] unfolding X\<^sub>j_def A_def by auto then obtain c where c: "?\\<^sub>u s x = Some c" by (cases "?\\<^sub>u s x", auto simp: bound_compare_defs) also have "c = \\ s\ x" proof - have "x \ rvars (\ s)" using that x Xj_rvars by blast then have "x \ (- lvars (\ s))" using 3 unfolding normalized_tableau_def by auto moreover have "\x\(- lvars (\ s)). in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" using 3 unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_bounds_set.simps) ultimately have "in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" by blast moreover have "?le c (\\ s\ x)" using cmp c dir unfolding bound_compare_defs by auto ultimately show ?thesis - using c dir by auto + using c dir by (auto simp del: Simplex.bounds_lg) qed then show ?thesis using c by simp qed moreover have "A x \ 0" using that coeff_zero unfolding A_def X\<^sub>j_def by auto ultimately show ?thesis using that by auto qed have l_Ba: "l \ \\<^sub>A s" if "l \ {?Geq x\<^sub>i (the (?\\<^sub>l s x\<^sub>i))}" for l proof - from that have l: "l = ?Geq x\<^sub>i (the (?\\<^sub>l s x\<^sub>i))" by simp from 3(8) obtain c where bl': "?\\<^sub>l s x\<^sub>i = Some c" by (cases "?\\<^sub>l s x\<^sub>i", auto simp: bound_compare_defs) hence bl: "(x\<^sub>i, c) \ set_of_map (?\\<^sub>l s)" unfolding set_of_map_def by auto show "l \ \\<^sub>A s" unfolding l bound_atoms_def using bl bl' dir by auto qed let ?negA = "filter (\ x. A x < 0) XL\<^sub>j" let ?posA = "filter (\ x. \ A x < 0) XL\<^sub>j" define neg where "neg = (if dir = Positive then (\ x :: rat. x) else uminus)" define negP where "negP = (if dir = Positive then (\ x :: linear_poly. x) else uminus)" define nega where "nega = (if dir = Positive then (\ x :: 'a. x) else uminus)" from dir have dirn: "dir = Positive \ neg = (\ x. x) \ negP = (\ x. x) \ nega = (\ x. x) \ dir = Negative \ neg = uminus \ negP = uminus \ nega = uminus" unfolding neg_def negP_def nega_def by auto define C where "C = map (\x. ?Geq x (the (?\\<^sub>l s x))) ?negA @ map (\ x. ?Leq x (the (?\\<^sub>u s x))) ?posA @ [?Geq x\<^sub>i (the (?\\<^sub>l s x\<^sub>i))]" define f where "f = (\x. if x = x\<^sub>i then neg (-1) else neg (A x))" define c where "c = (\x\C. lec_const (lec_of_nsc (f (atom_var x) *R nsc_of_atom x)))" let ?q = "negP ?p" show ?case unfolding bounds_id t_id u_id proof (intro exI impI conjI allI) show "v \\<^sub>t \ s \ ?q \ v \ = 0" for v :: "'a valuation" using dirn p_eval_zero[of v] by (auto simp: valuate_minus) show "set C \ \\<^sub>A s" unfolding C_def set_append set_map set_filter list.simps using 0 l_Ba dir by (intro Un_least subsetI) (force simp: bound_atoms_def set_of_map_def)+ show is_leq: "\a\set C. is_leq_ns (f (atom_var a) *R nsc_of_atom a) \ f (atom_var a) \ 0" using dirn xi_Xj 0 unfolding C_def f_def by (elim disjE, auto) show "(\a \ C. lec_of_nsc (f (atom_var a) *R nsc_of_atom a)) = Leqc ?q c" unfolding sum_list_lec le_constraint.simps map_map o_def proof (intro conjI) define scale_poly :: "'a atom \ linear_poly" where "scale_poly = (\x. lec_poly (lec_of_nsc (f (atom_var x) *R nsc_of_atom x)))" have "(\x\C. scale_poly x) = (\x\?negA. scale_poly (?Geq x (the (?\\<^sub>l s x)))) + (\x\?posA. scale_poly (?Leq x (the (?\\<^sub>u s x)))) - negP (lp_monom 1 x\<^sub>i)" unfolding C_def using dirn by (auto simp add: comp_def scale_poly_def f_def) also have "(\x\?negA. scale_poly (?Geq x (the (?\\<^sub>l s x)))) = (\x\ ?negA. negP (A x *R lp_monom 1 x))" unfolding scale_poly_def f_def using dirn xi_Xj by (subst map_cong) auto also have "(\x\?posA. scale_poly (?Leq x (the (?\\<^sub>u s x)))) = (\x\ ?posA. negP (A x *R lp_monom 1 x))" unfolding scale_poly_def f_def using dirn xi_Xj by (subst map_cong) auto also have "(\x\ ?negA. negP (A x *R lp_monom 1 x)) + (\x\ ?posA. negP (A x *R lp_monom 1 x)) = negP (rhs (eq_for_lvar (\ s) x\<^sub>i))" using dirn XL\<^sub>j_distinct coeff_zero by (elim disjE; intro poly_eqI, auto intro!: poly_eqI simp add: coeff_sum_list A_def X\<^sub>j_def uminus_sum_list_map[unfolded o_def, symmetric]) finally show "(\x\C. lec_poly (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))) = ?q" unfolding scale_poly_def using dirn by auto show "(\x\C. lec_rel (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))) = Leq_Rel" unfolding sum_list_Leq_Rel proof fix c assume c: "c \ set C" show "lec_rel (lec_of_nsc (f (atom_var c) *R nsc_of_atom c)) = Leq_Rel" using is_leq[rule_format, OF c] by (cases "f (atom_var c) *R nsc_of_atom c", auto) qed qed (simp add: c_def) show "c < 0" proof - define scale_const_f :: "'a atom \ 'a" where "scale_const_f x = lec_const (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))" for x obtain d where bl': "?\\<^sub>l s x\<^sub>i = Some d" using 3 by (cases "?\\<^sub>l s x\<^sub>i", auto simp: bound_compare_defs) have "c = (\x\map (\x. ?Geq x (the (?\\<^sub>l s x))) ?negA. scale_const_f x) + (\x\ map (\x. ?Leq x (the (?\\<^sub>u s x))) ?posA. scale_const_f x) - nega d" unfolding c_def C_def f_def scale_const_f_def using dirn rhs_eval_xi bl' by auto also have "(\x\map (\x. ?Geq x (the (?\\<^sub>l s x))) ?negA. scale_const_f x) = (\x\ ?negA. nega (A x *R the (?\\<^sub>l s x)))" using xi_Xj dirn by (subst map_cong) (auto simp add: f_def scale_const_f_def) also have "\ = (\x\?negA. nega (A x *R \\ s\ x))" using 0 by (subst map_cong) auto also have "(\x\map (\x. ?Leq x (the (?\\<^sub>u s x))) ?posA. scale_const_f x) = (\x\ ?posA. nega (A x *R the (?\\<^sub>u s x)))" using xi_Xj dirn by (subst map_cong) (auto simp add: f_def scale_const_f_def) also have "\ = (\x\ ?posA. nega (A x *R \\ s\ x))" using 0 by (subst map_cong) auto also have "(\x\?negA. nega (A x *R \\ s\ x)) + (\x\?posA. nega (A x *R \\ s\ x)) = (\x\?negA @ ?posA. nega (A x *R \\ s\ x))" by auto also have "\ = (\x\ X\<^sub>j. nega (A x *R \\ s\ x))" using XL\<^sub>j_distinct by (subst sum_list_distinct_conv_sum_set) (auto intro!: sum.cong) also have "\ = nega (\x\ X\<^sub>j. (A x *R \\ s\ x))" using dirn by (auto simp: sum_negf) also have "(\x\ X\<^sub>j. (A x *R \\ s\ x)) = ((rhs ?eq) \\\ s\\)" unfolding A_def X\<^sub>j_def by (subst linear_poly_sum) (auto simp add: sum_negf) also have "\ = \\ s\ x\<^sub>i" using rhs_eval_xi by blast also have "nega (\\ s\ x\<^sub>i) - nega d < 0" proof - have "?lt (\\ s\ x\<^sub>i) d" using dirn 3(2-) bl' by (elim disjE, auto simp: bound_compare_defs) thus ?thesis using dirn unfolding minus_lt[symmetric] by auto qed finally show ?thesis . qed show "distinct C" unfolding C_def using XL\<^sub>j_distinct xi_Xj dirn by (auto simp add: inj_on_def distinct_map) qed qed (insert U, blast+) then obtain f p c C where Qs: "?Q s f p c C" using U unfolding check by blast from index[folded check_tableau_index_valid[OF U(2) inv(3,4,2,1)]] check have index: "index_valid as s" by auto from check_tableau_equiv[OF U(2) inv(3,4,2,1), unfolded check] have id: "v \\<^sub>t \ s = v \\<^sub>t \ s'" for v :: "'a valuation" by auto let ?C = "map (\ a. (f (atom_var a), a)) C" have "set C \ \\<^sub>A s" using Qs by blast also have "\ \ snd ` as" using index unfolding bound_atoms_def index_valid_def set_of_map_def boundsl_def boundsu_def o_def by force finally have sub: "snd ` set ?C \ snd ` as" by force show ?thesis unfolding farkas_coefficients_atoms_tableau_def by (intro exI[of _ p] exI[of _ c] exI[of _ ?C] conjI, insert Qs[unfolded id] sub, (force simp: o_def)+) qed end text \Next, we show that a conflict found by the assert-bound function also gives rise to Farkas coefficients.\ context Update begin lemma farkas_assert_bound: assumes inv: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" "\ s" and index: "index_valid as s" and U: "\ (assert_bound ia s)" shows "\ C. farkas_coefficients_atoms_tableau (snd ` (insert ia as)) (\ s) C" proof - obtain i a where ia[simp]: "ia = (i,a)" by force let ?A = "snd ` insert ia as" have "\ x c d. Leq x c \ ?A \ Geq x d \ ?A \ c < d" proof (cases a) case (Geq x d) let ?s = "update\\ (Direction.UBI_upd (Direction (\x y. y < x) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>l_update Geq Leq (\))) i x d s" have id: "\ ?s = \ s" by auto have norm: "\ (\ ?s)" using inv by auto have val: "\ ?s" using inv(4) unfolding tableau_valuated_def by simp have idd: "x \ lvars (\ ?s) \ \ (update x d ?s) = \ ?s" by (rule update_unsat_id[OF norm val]) from U[unfolded ia Geq] inv(1) id idd have "\\<^sub>l\<^sub>b (\x y. y < x) d (\\<^sub>u s x)" by (auto split: if_splits simp: Let_def) then obtain c where Bu: "\\<^sub>u s x = Some c" and lt: "c < d" by (cases "\\<^sub>u s x", auto simp: bound_compare_defs) from Bu obtain j where "Mapping.lookup (\\<^sub>i\<^sub>u s) x = Some (j,c)" unfolding boundsu_def by auto with index[unfolded index_valid_def] have "(j, Leq x c) \ as" by auto hence xc: "Leq x c \ ?A" by force have xd: "Geq x d \ ?A" unfolding ia Geq by force from xc xd lt show ?thesis by auto next case (Leq x c) let ?s = "update\\ (Direction.UBI_upd (Direction (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>u_update Leq Geq (\))) i x c s" have id: "\ ?s = \ s" by auto have norm: "\ (\ ?s)" using inv by auto have val: "\ ?s" using inv(4) unfolding tableau_valuated_def by simp have idd: "x \ lvars (\ ?s) \ \ (update x c ?s) = \ ?s" by (rule update_unsat_id[OF norm val]) from U[unfolded ia Leq] inv(1) id idd have "\\<^sub>l\<^sub>b (<) c (\\<^sub>l s x)" by (auto split: if_splits simp: Let_def) then obtain d where Bl: "\\<^sub>l s x = Some d" and lt: "c < d" by (cases "\\<^sub>l s x", auto simp: bound_compare_defs) from Bl obtain j where "Mapping.lookup (\\<^sub>i\<^sub>l s) x = Some (j,d)" unfolding boundsl_def by auto with index[unfolded index_valid_def] have "(j, Geq x d) \ as" by auto hence xd: "Geq x d \ ?A" by force have xc: "Leq x c \ ?A" unfolding ia Leq by force from xc xd lt show ?thesis by auto qed then obtain x c d where c: "Leq x c \ ?A" and d: "Geq x d \ ?A" and cd: "c < d" by blast show ?thesis unfolding farkas_coefficients_atoms_tableau_def proof (intro exI conjI allI) let ?C = "[(-1, Geq x d), (1,Leq x c)]" show "\(r,a)\set ?C. a \ ?A \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0" using c d by auto show "c - d < 0" using cd using minus_lt by auto qed (auto simp: valuate_zero) qed end text \Moreover, we prove that all other steps of the simplex algorithm on layer~4, such as pivoting, asserting bounds without conflict, etc., preserve Farkas coefficients.\ lemma farkas_coefficients_atoms_tableau_mono: assumes "as \ bs" shows "farkas_coefficients_atoms_tableau as t C \ farkas_coefficients_atoms_tableau bs t C" using assms unfolding farkas_coefficients_atoms_tableau_def by force locale AssertAllState''' = AssertAllState'' init ass_bnd chk + Update update + PivotUpdateMinVars eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update for init and ass_bnd :: "'i \ 'a :: lrv atom \ _" and chk :: "('i, 'a) state \ ('i, 'a) state" and update :: "nat \ 'a :: lrv \ ('i, 'a) state \ ('i, 'a) state" and eq_idx_for_lvar :: "tableau \ var \ nat" and min_lvar_not_in_bounds :: "('i,'a::lrv) state \ var option" and min_rvar_incdec_eq :: "('i,'a) Direction \ ('i,'a) state \ eq \ 'i list + var" and pivot_and_update :: "var \ var \ 'a \ ('i,'a) state \ ('i,'a) state" + assumes ass_bnd: "ass_bnd = Update.assert_bound update" and chk: "chk = PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update" context AssertAllState''' begin lemma farkas_assert_bound_loop: assumes "\ (assert_bound_loop as (init t))" and norm: "\ t" shows "\ C. farkas_coefficients_atoms_tableau (snd ` set as) t C" proof - let ?P = "\ as s. \ s \ (\ C. farkas_coefficients_atoms_tableau (snd ` as) (\ s) C)" let ?s = "assert_bound_loop as (init t)" have "\ \ (init t)" by (rule init_unsat_flag) have "\ (assert_bound_loop as (init t)) = t \ (\ (assert_bound_loop as (init t)) \ (\ C. farkas_coefficients_atoms_tableau (snd ` set as) (\ (init t)) C))" proof (rule AssertAllState''Induct[OF norm], unfold ass_bnd, goal_cases) case 1 have "\ \ (init t)" by (rule init_unsat_flag) moreover have "\ (init t) = t" by (rule init_tableau_id) ultimately show ?case by auto next case (2 as bs s) hence "snd ` as \ snd ` bs" by auto from farkas_coefficients_atoms_tableau_mono[OF this] 2(2) show ?case by auto next case (3 s a ats) let ?s = "assert_bound a s" have tab: "\ ?s = \ s" unfolding ass_bnd by (rule assert_bound_nolhs_tableau_id, insert 3, auto) have t: "t = \ s" using 3 by simp show ?case unfolding t tab proof (intro conjI impI refl) assume "\ ?s" from farkas_assert_bound[OF 3(1,3-6,8) this] show "\ C. farkas_coefficients_atoms_tableau (snd ` insert a (set ats)) (\ (init (\ s))) C" unfolding t[symmetric] init_tableau_id . qed qed thus ?thesis unfolding init_tableau_id using assms by blast qed text \Now we get to the main result for layer~4: If the main algorithm returns unsat, then there are Farkas coefficients for the tableau and atom set that were given as input for this layer.\ lemma farkas_assert_all_state: assumes U: "\ (assert_all_state t as)" and norm: "\ t" shows "\ C. farkas_coefficients_atoms_tableau (snd ` set as) t C" proof - let ?s = "assert_bound_loop as (init t)" show ?thesis proof (cases "\ (assert_bound_loop as (init t))") case True from farkas_assert_bound_loop[OF this norm] show ?thesis by auto next case False from AssertAllState''_tableau_id[OF norm] have T: "\ ?s = t" unfolding init_tableau_id . from U have U: "\ (check ?s)" unfolding chk[symmetric] by simp show ?thesis proof (rule farkas_check[OF refl U False, unfolded T, OF _ norm]) from AssertAllState''_precond[OF norm, unfolded Let_def] False show "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s" "\ ?s" "\ ?s" by blast+ from AssertAllState''_index_valid[OF norm] show "index_valid (set as) ?s" . qed qed qed subsection \Farkas' Lemma on Layer 3\ text \There is only a small difference between layers 3 and 4, namely that there is no simplex algorithm (@{const assert_all_state}) on layer 3, but just a tableau and atoms.\ text \Hence, one task is to link the unsatisfiability flag on layer 4 with unsatisfiability of the original tableau and atoms (layer 3). This can be done via the existing soundness results of the simplex algorithm. Moreover, we give an easy proof that the existence of Farkas coefficients for a tableau and set of atoms implies unsatisfiability.\ end lemma farkas_coefficients_atoms_tableau_unsat: assumes "farkas_coefficients_atoms_tableau as t C" shows "\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" proof assume "\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" then obtain v where *: "v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" by auto then obtain p c where isleq: "(\(r,a) \ set C. a \ as \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0)" and leq: "(\(r,a) \ C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c" and cltz: "c < 0" and p0: "p\v\ = 0" using assms farkas_coefficients_atoms_tableau_def by blast have fa: "\(r,a) \ set C. v \\<^sub>a a" using * isleq leq satisfies_atom_set_def by force { fix r a assume a: "(r,a) \ set C" from a fa have va: "v \\<^sub>a a" unfolding satisfies_atom_set_def by auto hence v: "v \\<^sub>n\<^sub>s (r *R nsc_of_atom a)" by (auto simp: nsc_of_atom sat_scale_rat_ns) from a isleq have "is_leq_ns (r *R nsc_of_atom a)" by auto from lec_of_nsc[OF this] v have "v \\<^sub>l\<^sub>e lec_of_nsc (r *R nsc_of_atom a)" by blast } note v = this have "v \\<^sub>l\<^sub>e Leqc p c" unfolding leq[symmetric] by (rule satisfies_sumlist_le_constraints, insert v, auto) then have "0 \ c" using p0 by auto then show False using cltz by auto qed text \Next is the main result for layer~3: a tableau and a finite set of atoms are unsatisfiable if and only if there is a list of Farkas coefficients for the set of atoms and the tableau.\ lemma farkas_coefficients_atoms_tableau: assumes norm: "\ t" and fin: "finite as" shows "(\ C. farkas_coefficients_atoms_tableau as t C) \ (\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as)" proof from finite_list[OF fin] obtain bs where as: "as = set bs" by auto assume unsat: "\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" let ?as = "map (\ x. ((),x)) bs" interpret AssertAllState''' init_state assert_bound_code check_code update_code eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code by (unfold_locales, auto simp: assert_bound_code_def check_code_def) let ?call = "assert_all t ?as" have id: "snd ` set ?as = as" unfolding as by force from assert_all_sat[OF norm, of ?as, unfolded id] unsat obtain I where "?call = Inl I" by (cases ?call, auto) from this[unfolded assert_all_def Let_def] have "\ (assert_all_state_code t ?as)" by (auto split: if_splits simp: assert_all_state_code_def) from farkas_assert_all_state[OF this[unfolded assert_all_state_code_def] norm, unfolded id] show "\ C. farkas_coefficients_atoms_tableau as t C" . qed (insert farkas_coefficients_atoms_tableau_unsat, auto) subsection \Farkas' Lemma on Layer 2\ text \The main difference between layers 2 and 3 is the introduction of slack-variables in layer 3 via the preprocess-function. Our task here is to show that Farkas coefficients at layer 3 (where slack-variables are used) can be converted into Farkas coefficients for layer 2 (before the preprocessing).\ text \We also need to adapt the previos notion of Farkas coefficients, which was used in @{const farkas_coefficients_atoms_tableau}, for layer~2. At layer 3, Farkas coefficients are the coefficients in a linear combination of atoms that evaluates to an inequality of the form $p \leq c$, where $p$ is a linear polynomial, $c < 0$, and $t \models p = 0$ holds. At layer 2, the atoms are replaced by non-strict constraints where the left-hand side is a polynomial in the original variables, but the corresponding linear combination (with Farkas coefficients) evaluates directly to the inequality $0 \leq c$, with $c < 0$. The implication $t \models p = 0$ is no longer possible in this layer, since there is no tableau $t$, nor is it needed, since $p$ is $0$. Thus, the statement defining Farkas coefficients must be changed accordingly.\ definition farkas_coefficients_ns where "farkas_coefficients_ns ns C = (\ c. (\(r, n) \ set C. n \ ns \ is_leq_ns (r *R n) \ r \ 0) \ (\(r, n) \ C. lec_of_nsc (r *R n)) = Leqc 0 c \ c < 0)" text \The easy part is to prove that Farkas coefficients imply unsatisfiability.\ lemma farkas_coefficients_ns_unsat: assumes "farkas_coefficients_ns ns C" shows "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" proof assume "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" then obtain v where *: "v \\<^sub>n\<^sub>s\<^sub>s ns" by auto obtain c where isleq: "(\(a,n) \ set C. n \ ns \ is_leq_ns (a *R n) \ a \ 0)" and leq: "(\(a,n) \ C. lec_of_nsc (a *R n)) = Leqc 0 c" and cltz: "c < 0" using assms farkas_coefficients_ns_def by blast { fix a n assume n: "(a,n) \ set C" from n * isleq have "v \\<^sub>n\<^sub>s n" by auto hence v: "v \\<^sub>n\<^sub>s (a *R n)" by (rule sat_scale_rat_ns) from n isleq have "is_leq_ns (a *R n)" by auto from lec_of_nsc[OF this] v have "v \\<^sub>l\<^sub>e lec_of_nsc (a *R n)" by blast } note v = this have "v \\<^sub>l\<^sub>e Leqc 0 c" unfolding leq[symmetric] by (rule satisfies_sumlist_le_constraints, insert v, auto) then show False using cltz by (metis leD satisfiable_le_constraint.simps valuate_zero rel_of.simps(1)) qed text \In order to eliminate the need for a tableau, we require the notion of an arbitrary substitution on polynomials, where all variables can be replaced at once. The existing simplex formalization provides only a function to replace one variable at a time.\ definition subst_poly :: "(var \ linear_poly) \ linear_poly \ linear_poly" where "subst_poly \ p = (\ x \ vars p. coeff p x *R \ x)" lemma subst_poly_0[simp]: "subst_poly \ 0 = 0" unfolding subst_poly_def by simp lemma valuate_subst_poly: "(subst_poly \ p) \ v \ = (p \ (\ x. ((\ x) \ v \)) \)" by (subst (2) linear_poly_sum, unfold subst_poly_def valuate_sum valuate_scaleRat, simp) lemma subst_poly_add: "subst_poly \ (p + q) = subst_poly \ p + subst_poly \ q" by (rule linear_poly_eqI, unfold valuate_add valuate_subst_poly, simp) fun subst_poly_lec :: "(var \ linear_poly) \ 'a le_constraint \ 'a le_constraint" where "subst_poly_lec \ (Le_Constraint rel p c) = Le_Constraint rel (subst_poly \ p) c" lemma subst_poly_lec_0[simp]: "subst_poly_lec \ 0 = 0" unfolding zero_le_constraint_def by simp lemma subst_poly_lec_add: "subst_poly_lec \ (c1 + c2) = subst_poly_lec \ c1 + subst_poly_lec \ c2" by (cases c1; cases c2, auto simp: subst_poly_add) lemma subst_poly_lec_sum_list: "subst_poly_lec \ (sum_list ps) = sum_list (map (subst_poly_lec \) ps)" by (induct ps, auto simp: subst_poly_lec_add) lemma subst_poly_lp_monom[simp]: "subst_poly \ (lp_monom r x) = r *R \ x" unfolding subst_poly_def by (simp add: vars_lp_monom) lemma subst_poly_scaleRat: "subst_poly \ (r *R p) = r *R (subst_poly \ p)" by (rule linear_poly_eqI, unfold valuate_scaleRat valuate_subst_poly, simp) text \We need several auxiliary properties of the preprocess-function which are not present in the simplex formalization.\ lemma Tableau_is_monom_preprocess': assumes "(x, p) \ set (Tableau (preprocess' cs start))" shows "\ is_monom p" using assms by(induction cs start rule: preprocess'.induct) (auto simp add: Let_def split: if_splits option.splits) lemma preprocess'_atoms_to_constraints': assumes "preprocess' cs start = S" shows "set (Atoms S) \ {(i,qdelta_constraint_to_atom c v) | i c v. (i,c) \ set cs \ (\ is_monom (poly c) \ Poly_Mapping S (poly c) = Some v)}" unfolding assms(1)[symmetric] by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits, force+) lemma monom_of_atom_coeff: assumes "is_monom (poly ns)" "a = qdelta_constraint_to_atom ns v" shows "(monom_coeff (poly ns)) *R nsc_of_atom a = ns" using assms is_monom_monom_coeff_not_zero by(cases a; cases ns) (auto split: atom.split ns_constraint.split simp add: monom_poly_assemble field_simps) text \The next lemma provides the functionality that is required to convert an atom back to a non-strict constraint, i.e., it is a kind of inverse of the preprocess-function.\ lemma preprocess'_atoms_to_constraints: assumes S: "preprocess' cs start = S" and start: "start = start_fresh_variable cs" and ns: "ns = (case a of Leq v c \ LEQ_ns q c | Geq v c \ GEQ_ns q c)" and "a \ snd ` set (Atoms S)" shows "(atom_var a \ fst ` set (Tableau S) \ (\ r. r \ 0 \ r *R nsc_of_atom a \ snd ` set cs)) \ ((atom_var a, q) \ set (Tableau S) \ ns \ snd ` set cs)" proof - let ?S = "preprocess' cs start" from assms(4) obtain i where ia: "(i,a) \ set (Atoms S)" by auto with preprocess'_atoms_to_constraints'[OF assms(1)] obtain c v where a: "a = qdelta_constraint_to_atom c v" and c: "(i,c) \ set cs" and nmonom: "\ is_monom (poly c) \ Poly_Mapping S (poly c) = Some v" by blast hence c': "c \ snd ` set cs" by force let ?p = "poly c" show ?thesis proof (cases "is_monom ?p") case True hence av: "atom_var a = monom_var ?p" unfolding a by (cases c, auto) from Tableau_is_monom_preprocess'[of _ ?p cs start] True have "(x, ?p) \ set (Tableau ?S)" for x by blast { assume "(atom_var a, q) \ set (Tableau S)" hence "(monom_var ?p, q) \ set (Tableau S)" unfolding av by simp hence "monom_var ?p \ lvars (Tableau S)" unfolding lvars_def by force from lvars_tableau_ge_start[rule_format, OF this[folded S]] have "monom_var ?p \ start" unfolding S . moreover have "monom_var ?p \ vars_constraints (map snd cs)" using True c by (auto intro!: bexI[of _ "(i,c)"] simp: monom_var_in_vars) ultimately have False using start_fresh_variable_fresh[of cs, folded start] by force } moreover from monom_of_atom_coeff[OF True a] True have "\r. r \ 0 \ r *R nsc_of_atom a = c" by (intro exI[of _ "monom_coeff ?p"], auto, cases a, auto) ultimately show ?thesis using c' by auto next case False hence av: "atom_var a = v" unfolding a by (cases c, auto) from nmonom[OF False] have "Poly_Mapping S ?p = Some v" . from preprocess'_Tableau_Poly_Mapping_Some[OF this[folded S]] have tab: "(atom_var a, ?p) \ set (Tableau (preprocess' cs start))" unfolding av by simp hence "atom_var a \ fst ` set (Tableau S)" unfolding S by force moreover { assume "(atom_var a, q) \ set (Tableau S)" from tab this have qp: "q = ?p" unfolding S using lvars_distinct[of cs start, unfolded S lhs_def] by (simp add: case_prod_beta' eq_key_imp_eq_value) have "ns = c" unfolding ns qp using av a False by (cases c, auto) hence "ns \ snd ` set cs" using c' by blast } ultimately show ?thesis by blast qed qed text \Next follows the major technical lemma of this part, namely that Farkas coefficients on layer~3 for preprocessed constraints can be converted into Farkas coefficients on layer~2.\ lemma farkas_coefficients_preprocess': assumes pp: "preprocess' cs (start_fresh_variable cs) = S" and ft: "farkas_coefficients_atoms_tableau (snd ` set (Atoms S)) (Tableau S) C" shows "\ C. farkas_coefficients_ns (snd ` set cs) C" proof - note ft[unfolded farkas_coefficients_atoms_tableau_def] obtain p c where 0: "\ (r,a) \ set C. a \ snd ` set (Atoms S) \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0" "(\(r,a)\C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c" "c < 0" "\v :: QDelta valuation. v \\<^sub>t Tableau S \ p \ v \ = 0" using ft unfolding farkas_coefficients_atoms_tableau_def by blast note 0 = 0(1)[rule_format, of "(a, b)" for a b, unfolded split] 0(2-) let ?T = "Tableau S" define \ :: "var \ linear_poly" where "\ = (\ x. case map_of ?T x of Some p \ p | None \ lp_monom 1 x)" let ?P = "(\r a s ns. ns \ (snd ` set cs) \ is_leq_ns (s *R ns) \ s \ 0 \ subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (s *R ns))" have "\s ns. ?P r a s ns" if ra: "(r,a) \ set C" for r a proof - have a: "a \ snd ` set (Atoms S)" using ra 0 by force from 0 ra have is_leq: "is_leq_ns (r *R nsc_of_atom a)" and r0: "r \ 0" by auto let ?x = "atom_var a" show ?thesis proof (cases "map_of ?T ?x") case (Some q) hence \: "\ ?x = q" unfolding \_def by auto from Some have xqT: "(?x, q) \ set ?T" by (rule map_of_SomeD) define ns where "ns = (case a of Leq v c \ LEQ_ns q c | Geq v c \ GEQ_ns q c)" from preprocess'_atoms_to_constraints[OF pp refl ns_def a] xqT have ns_mem: "ns \ snd ` set cs" by blast have id: "subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (r *R ns)" using is_leq \ by (cases a, auto simp: ns_def subst_poly_scaleRat) from id is_leq \ have is_leq: "is_leq_ns (r *R ns)" by (cases a, auto simp: ns_def) show ?thesis by (intro exI[of _ r] exI[of _ ns] conjI ns_mem id is_leq conjI r0) next case None hence "?x \ fst ` set ?T" by (meson map_of_eq_None_iff) from preprocess'_atoms_to_constraints[OF pp refl refl a] this obtain rr where rr: "rr *R nsc_of_atom a \ (snd ` set cs)" and rr0: "rr \ 0" by blast from None have \: "\ ?x = lp_monom 1 ?x" unfolding \_def by simp define ns where "ns = rr *R nsc_of_atom a" define s where "s = r / rr" from rr0 r0 have s0: "s \ 0" unfolding s_def by auto from is_leq \ have "subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (r *R nsc_of_atom a)" by (cases a, auto simp: subst_poly_scaleRat) moreover have "r *R nsc_of_atom a = s *R ns" unfolding ns_def s_def scaleRat_scaleRat_ns_constraint[OF rr0] using rr0 by simp ultimately have "subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (s *R ns)" "is_leq_ns (s *R ns)" using is_leq by auto then show ?thesis unfolding ns_def using rr s0 by blast qed qed hence "\ ra. \ s ns. (fst ra, snd ra) \ set C \ ?P (fst ra) (snd ra) s ns" by blast from choice[OF this] obtain s where s: "\ ra. \ ns. (fst ra, snd ra) \ set C \ ?P (fst ra) (snd ra) (s ra) ns" by blast from choice[OF this] obtain ns where ns: "\ r a. (r,a) \ set C \ ?P r a (s (r,a)) (ns (r,a))" by force define NC where "NC = map (\(r,a). (s (r,a), ns (r,a))) C" have "(\(s, ns)\map (\(r,a). (s (r,a), ns (r,a))) C'. lec_of_nsc (s *R ns)) = (\(r, a)\C'. subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)))" if "set C' \ set C" for C' using that proof (induction C') case Nil then show ?case by simp next case (Cons a C') have "(\x\a # C'. lec_of_nsc (s x *R ns x)) = lec_of_nsc (s a *R ns a) + (\x\C'. lec_of_nsc (s x *R ns x))" by simp also have "(\x\C'. lec_of_nsc (s x *R ns x)) = (\(r, a)\C'. subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)))" using Cons by (auto simp add: case_prod_beta' comp_def) also have "lec_of_nsc (s a *R ns a) = subst_poly_lec \ (lec_of_nsc (fst a *R nsc_of_atom (snd a)))" proof - have "a \ set C" using Cons by simp then show ?thesis using ns by auto qed finally show ?case by (auto simp add: case_prod_beta' comp_def) qed also have "(\(r, a)\C. subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a))) = subst_poly_lec \ (\(r, a)\C. (lec_of_nsc (r *R nsc_of_atom a)))" by (auto simp add: subst_poly_lec_sum_list case_prod_beta' comp_def) also have "(\(r, a)\C. (lec_of_nsc (r *R nsc_of_atom a))) = Leqc p c" using 0 by blast also have "subst_poly_lec \ (Leqc p c) = Leqc (subst_poly \ p) c" by simp also have "subst_poly \ p = 0" proof (rule all_valuate_zero) fix v :: "QDelta valuation" have "(subst_poly \ p) \ v \ = (p \ \x. ((\ x) \ v \) \)" by (rule valuate_subst_poly) also have "\ = 0" proof (rule 0(4)) have "(\ a) \ v \ = (q \ \x. ((\ x) \ v \) \)" if "(a, q) \ set (Tableau S)" for a q proof - have "distinct (map fst ?T)" using normalized_tableau_preprocess' assms unfolding normalized_tableau_def lhs_def by (auto simp add: case_prod_beta') then have 0: "\ a = q" unfolding \_def using that by auto have "q \ v \ = (q \ \x. ((\ x) \ v \) \)" proof - have "vars q \ rvars ?T" unfolding rvars_def using that by force moreover have "(\ x) \ v \ = v x" if "x \ rvars ?T" for x proof - have "x \ lvars (Tableau S)" using that normalized_tableau_preprocess' assms unfolding normalized_tableau_def by blast then have "x \ fst ` set (Tableau S)" unfolding lvars_def by force then have "map_of ?T x = None" using map_of_eq_None_iff by metis then have "\ x = lp_monom 1 x" unfolding \_def by auto also have "(lp_monom 1 x) \ v \ = v x" by auto finally show ?thesis . qed ultimately show ?thesis by (auto intro!: valuate_depend) qed then show ?thesis using 0 by blast qed then show "(\x. ((\ x) \ v \)) \\<^sub>t ?T" using 0 by (auto simp add: satisfies_tableau_def satisfies_eq_def) qed finally show "(subst_poly \ p) \ v \ = 0" . qed finally have "(\(s, n)\NC. lec_of_nsc (s *R n)) = Le_Constraint Leq_Rel 0 c" unfolding NC_def by blast moreover have "ns (r,a) \ snd ` set cs" "is_leq_ns (s (r, a) *R ns (r, a))" "s (r, a) \ 0" if "(r, a) \ set C" for r a using ns that by force+ ultimately have "farkas_coefficients_ns (snd ` set cs) NC" unfolding farkas_coefficients_ns_def NC_def using 0 by force then show ?thesis by blast qed lemma preprocess'_unsat_indexD: "i \ set (UnsatIndices (preprocess' ns j)) \ \ c. poly c = 0 \ \ zero_satisfies c \ (i,c) \ set ns" by (induct ns j rule: preprocess'.induct, auto simp: Let_def split: if_splits option.splits) lemma preprocess'_unsat_index_farkas_coefficients_ns: assumes "i \ set (UnsatIndices (preprocess' ns j))" shows "\ C. farkas_coefficients_ns (snd ` set ns) C" proof - from preprocess'_unsat_indexD[OF assms] obtain c where contr: "poly c = 0" "\ zero_satisfies c" and mem: "(i,c) \ set ns" by auto from mem have mem: "c \ snd ` set ns" by force let ?c = "ns_constraint_const c" define r where "r = (case c of LEQ_ns _ _ \ 1 | _ \ (-1 :: rat))" define d where "d = (case c of LEQ_ns _ _ \ ?c | _ \ - ?c)" have [simp]: "(- x < 0) = (0 < x)" for x :: QDelta using uminus_less_lrv[of _ 0] by simp show ?thesis unfolding farkas_coefficients_ns_def by (intro exI[of _ "[(r,c)]"] exI[of _ d], insert mem contr, cases "c", auto simp: r_def d_def) qed text \The combination of the previous results easily provides the main result of this section: a finite set of non-strict constraints on layer~2 is unsatisfiable if and only if there are Farkas coefficients. Again, here we use results from the simplex formalization, namely soundness of the preprocess-function.\ lemma farkas_coefficients_ns: assumes "finite (ns :: QDelta ns_constraint set)" shows "(\ C. farkas_coefficients_ns ns C) \ (\ v. v \\<^sub>n\<^sub>s\<^sub>s ns)" proof assume "\ C. farkas_coefficients_ns ns C" from farkas_coefficients_ns_unsat this show "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" by blast next assume unsat: "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" from finite_list[OF assms] obtain nsl where ns: "ns = set nsl" by auto let ?cs = "map (Pair ()) nsl" obtain I t ias where part1: "preprocess_part_1 ?cs = (t,ias,I)" by (cases "preprocess_part_1 ?cs", auto) let ?as = "snd ` set ias" let ?s = "start_fresh_variable ?cs" have fin: "finite ?as" by auto have id: "ias = Atoms (preprocess' ?cs ?s)" "t = Tableau (preprocess' ?cs ?s)" "I = UnsatIndices (preprocess' ?cs ?s)" using part1 unfolding preprocess_part_1_def Let_def by auto have norm: "\ t" using normalized_tableau_preprocess'[of ?cs] unfolding id . { fix v assume "v \\<^sub>a\<^sub>s ?as" "v \\<^sub>t t" from preprocess'_sat[OF this[unfolded id], folded id] unsat[unfolded ns] have "set I \ {}" by auto then obtain i where "i \ set I" using all_not_in_conv by blast from preprocess'_unsat_index_farkas_coefficients_ns[OF this[unfolded id]] have "\C. farkas_coefficients_ns (snd ` set ?cs) C" by simp } with farkas_coefficients_atoms_tableau[OF norm fin] obtain C where "farkas_coefficients_atoms_tableau ?as t C \ (\C. farkas_coefficients_ns (snd ` set ?cs) C)" by blast from farkas_coefficients_preprocess'[of ?cs, OF refl] this have "\ C. farkas_coefficients_ns (snd ` set ?cs) C" using part1 unfolding preprocess_part_1_def Let_def by auto also have "snd ` set ?cs = ns" unfolding ns by force finally show "\ C. farkas_coefficients_ns ns C" . qed subsection \Farkas' Lemma on Layer 1\ text \The main difference of layers 1 and 2 is the restriction to non-strict constraints via delta-rationals. Since we now work with another constraint type, @{type constraint}, we again need translations into linear inequalities of type @{type le_constraint}. Moreover, we also need to define scaling of constraints where flipping the comparison sign may be required.\ fun is_le :: "constraint \ bool" where "is_le (LT _ _) = True" | "is_le (LEQ _ _) = True" | "is_le _ = False" fun lec_of_constraint where "lec_of_constraint (LEQ p c) = (Le_Constraint Leq_Rel p c)" | "lec_of_constraint (LT p c) = (Le_Constraint Lt_Rel p c)" lemma lec_of_constraint: assumes "is_le c" shows "(v \\<^sub>l\<^sub>e (lec_of_constraint c)) \ (v \\<^sub>c c)" using assms by (cases c, auto) instantiation constraint :: scaleRat begin fun scaleRat_constraint :: "rat \ constraint \ constraint" where "scaleRat_constraint r cc = (if r = 0 then LEQ 0 0 else (case cc of LEQ p c \ (if (r < 0) then GEQ (r *R p) (r *R c) else LEQ (r *R p) (r *R c)) | LT p c \ (if (r < 0) then GT (r *R p) (r *R c) else LT (r *R p) (r *R c)) | GEQ p c \ (if (r > 0) then GEQ (r *R p) (r *R c) else LEQ (r *R p) (r *R c)) | GT p c \ (if (r > 0) then GT (r *R p) (r *R c) else LT (r *R p) (r *R c)) | LTPP p q \ (if (r < 0) then GT (r *R (p - q)) 0 else LT (r *R (p - q)) 0) | LEQPP p q \ (if (r < 0) then GEQ (r *R (p - q)) 0 else LEQ (r *R (p - q)) 0) | GTPP p q \ (if (r > 0) then GT (r *R (p - q)) 0 else LT (r *R (p - q)) 0) | GEQPP p q \ (if (r > 0) then GEQ (r *R (p - q)) 0 else LEQ (r *R (p - q)) 0) | EQPP p q \ LEQ (r *R (p - q)) 0 \ \We do not keep equality, since the aim is to convert the scaled constraints into inequalities, which will then be summed up.\ | EQ p c \ LEQ (r *R p) (r *R c) ))" instance .. end lemma sat_scale_rat: assumes "(v :: rat valuation) \\<^sub>c c" shows "v \\<^sub>c (r *R c)" proof - have "r < 0 \ r = 0 \ r > 0" by auto then show ?thesis using assms by (cases c, auto simp: right_diff_distrib valuate_minus valuate_scaleRat scaleRat_leq1 scaleRat_leq2 valuate_zero) qed text \In the following definition of Farkas coefficients (for layer 1), the main difference to @{const farkas_coefficients_ns} is that the linear combination evaluates either to a strict inequality where the constant must be non-positive, or to a non-strict inequality where the constant must be negative.\ definition farkas_coefficients where "farkas_coefficients cs C = (\ d rel. (\ (r,c) \ set C. c \ cs \ is_le (r *R c) \ r \ 0) \ (\ (r,c) \ C. lec_of_constraint (r *R c)) = Le_Constraint rel 0 d \ (rel = Leq_Rel \ d < 0 \ rel = Lt_Rel \ d \ 0))" text \Again, the existence Farkas coefficients immediately implies unsatisfiability.\ lemma farkas_coefficients_unsat: assumes "farkas_coefficients cs C" shows "\ v. v \\<^sub>c\<^sub>s cs" proof assume "\ v. v \\<^sub>c\<^sub>s cs" then obtain v where *: "v \\<^sub>c\<^sub>s cs" by auto obtain d rel where isleq: "(\(r,c) \ set C. c \ cs \ is_le (r *R c) \ r \ 0)" and leq: "(\ (r,c) \ C. lec_of_constraint (r *R c)) = Le_Constraint rel 0 d" and choice: "rel = Lt_Rel \ d \ 0 \ rel = Leq_Rel \ d < 0" using assms farkas_coefficients_def by blast { fix r c assume c: "(r,c) \ set C" from c * isleq have "v \\<^sub>c c" by auto hence v: "v \\<^sub>c (r *R c)" by (rule sat_scale_rat) from c isleq have "is_le (r *R c)" by auto from lec_of_constraint[OF this] v have "v \\<^sub>l\<^sub>e lec_of_constraint (r *R c)" by blast } note v = this have "v \\<^sub>l\<^sub>e Le_Constraint rel 0 d" unfolding leq[symmetric] by (rule satisfies_sumlist_le_constraints, insert v, auto) then show False using choice by (cases rel, auto simp: valuate_zero) qed text \Now follows the difficult implication. The major part is proving that the translation @{const constraint_to_qdelta_constraint} preserves the existence of Farkas coefficients via pointwise compatibility of the sum. Here, compatibility links a strict or non-strict inequality from the input constraint to a translated non-strict inequality over delta-rationals.\ fun compatible_cs where "compatible_cs (Le_Constraint Leq_Rel p c) (Le_Constraint Leq_Rel q d) = (q = p \ d = QDelta c 0)" | "compatible_cs (Le_Constraint Lt_Rel p c) (Le_Constraint Leq_Rel q d) = (q = p \ qdfst d = c)" | "compatible_cs _ _ = False" lemma compatible_cs_0_0: "compatible_cs 0 0" by code_simp lemma compatible_cs_plus: "compatible_cs c1 d1 \ compatible_cs c2 d2 \ compatible_cs (c1 + c2) (d1 + d2)" by (cases c1; cases d1; cases c2; cases d2; cases "lec_rel c1"; cases "lec_rel d1"; cases "lec_rel c2"; cases "lec_rel d2"; auto simp: plus_QDelta_def) lemma unsat_farkas_coefficients: assumes "\ v. v \\<^sub>c\<^sub>s cs" and fin: "finite cs" shows "\ C. farkas_coefficients cs C" proof - from finite_list[OF fin] obtain csl where cs: "cs = set csl" by blast let ?csl = "map (Pair ()) csl" let ?ns = "(snd ` set (to_ns ?csl))" let ?nsl = "concat (map constraint_to_qdelta_constraint csl)" have id: "snd ` set ?csl = cs" unfolding cs by force have id2: "?ns = set ?nsl" unfolding to_ns_def set_concat by force from SolveExec'Default.to_ns_sat[of ?csl, unfolded id] assms have unsat: "\ v. \v\ \\<^sub>n\<^sub>s\<^sub>s ?ns" by metis have fin: "finite ?ns" by auto have "\ v. v \\<^sub>n\<^sub>s\<^sub>s ?ns" proof assume "\ v. v \\<^sub>n\<^sub>s\<^sub>s ?ns" then obtain v where model: "v \\<^sub>n\<^sub>s\<^sub>s ?ns" by blast let ?v = "Mapping.Mapping (\ x. Some (v x))" have "v = \?v\" by (intro ext, auto simp: map2fun_def Mapping.lookup.abs_eq) from model this unsat show False by metis qed from farkas_coefficients_ns[OF fin] this id2 obtain N where farkas: "farkas_coefficients_ns (set ?nsl) N" by metis from this[unfolded farkas_coefficients_ns_def] obtain d where is_leq: "\ a n. (a,n) \ set N \ n \ set ?nsl \ is_leq_ns (a *R n) \ a \ 0" and sum: "(\(a,n)\N. lec_of_nsc (a *R n)) = Le_Constraint Leq_Rel 0 d" and d0: "d < 0" by blast let ?prop = "\ NN C. (\ (a,c) \ set C. c \ cs \ is_le (a *R c) \ a \ 0) \ compatible_cs (\ (a,c) \ C. lec_of_constraint (a *R c)) (\(a,n)\NN. lec_of_nsc (a *R n))" have "set NN \ set N \ \ C. ?prop NN C" for NN proof (induct NN) case Nil have "?prop Nil Nil" by (simp add: compatible_cs_0_0) thus ?case by blast next case (Cons an NN) obtain a n where an: "an = (a,n)" by force from Cons an obtain C where IH: "?prop NN C" and n: "(a,n) \ set N" by auto have compat_CN: "compatible_cs (\(f, c)\C. lec_of_constraint (f *R c)) (\(a,n)\NN. lec_of_nsc (a *R n))" using IH by blast from n is_leq obtain c where c: "c \ cs" and nc: "n \ set (constraint_to_qdelta_constraint c)" unfolding cs by force from is_leq[OF n] have is_leq: "is_leq_ns (a *R n) \ a \ 0" by blast have is_less: "is_le (a *R c)" and a0: "a \ 0" and compat_cn: "compatible_cs (lec_of_constraint (a *R c)) (lec_of_nsc (a *R n))" by (atomize(full), cases c, insert is_leq nc, auto simp: QDelta_0_0 scaleRat_QDelta_def qdsnd_0 qdfst_0) let ?C = "Cons (a, c) C" let ?N = "Cons (a, n) NN" have "?prop ?N ?C" unfolding an proof (intro conjI) show "\ (a,c) \ set ?C. c \ cs \ is_le (a *R c) \ a \ 0" using IH is_less a0 c by auto show "compatible_cs (\(a, c)\?C. lec_of_constraint (a *R c)) (\(a,n)\?N. lec_of_nsc (a *R n))" using compatible_cs_plus[OF compat_cn compat_CN] by simp qed thus ?case unfolding an by blast qed from this[OF subset_refl, unfolded sum] obtain C where is_less: "(\(a, c)\set C. c \ cs \ is_le (a *R c) \ a \ 0)" and compat: "compatible_cs (\(f, c)\C. lec_of_constraint (f *R c)) (Le_Constraint Leq_Rel 0 d)" (is "compatible_cs ?sum _") by blast obtain rel p e where "?sum = Le_Constraint rel p e" by (cases ?sum) with compat have sum: "?sum = Le_Constraint rel 0 e" by (cases rel, auto) have e: "(rel = Leq_Rel \ e < 0 \ rel = Lt_Rel \ e \ 0)" using compat[unfolded sum] d0 by (cases rel, auto simp: less_QDelta_def qdfst_0 qdsnd_0) show ?thesis unfolding farkas_coefficients_def by (intro exI conjI, rule is_less, rule sum, insert e, auto) qed text \Finally we can prove on layer 1 that a finite set of constraints is unsatisfiable if and only if there are Farkas coefficients.\ lemma farkas_coefficients: assumes "finite cs" shows "(\ C. farkas_coefficients cs C) \ (\ v. v \\<^sub>c\<^sub>s cs)" using farkas_coefficients_unsat unsat_farkas_coefficients[OF _ assms] by blast section \Corollaries from the Literature\ text \In this section, we convert the previous variations of Farkas' Lemma into more well-known forms of this result. Moreover, instead of referring to the various constraint types of the simplex formalization, we now speak solely about constraints of type @{type le_constraint}.\ subsection \Farkas' Lemma on Delta-Rationals\ text \We start with Lemma~2 of \cite{Bromberger2017}, a variant of Farkas' Lemma for delta-rationals. To be more precise, it states that a set of non-strict inequalities over delta-rationals is unsatisfiable if and only if there is a linear combination of the inequalities that results in a trivial unsatisfiable constraint $0 < const$ for some negative constant $const$. We can easily prove this statement via the lemma @{thm [source] farkas_coefficients_ns} and some conversions between the different constraint types.\ lemma Farkas'_Lemma_Delta_Rationals: fixes cs :: "QDelta le_constraint set" assumes only_non_strict: "lec_rel ` cs \ {Leq_Rel}" and fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const. (\ (r, c) \ set C. r > 0 \ c \ cs) \ (\ (r,c) \ C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const \ const < 0)" (is "?lhs = ?rhs") proof - { fix c assume "c \ cs" with only_non_strict have "lec_rel c = Leq_Rel" by auto then have "\ p const. c = Leqc p const" by (cases c, auto) } note leqc = this let ?to_ns = "\ c. LEQ_ns (lec_poly c) (lec_const c)" let ?ns = "?to_ns ` cs" from fin have fin: "finite ?ns" by auto have "v \\<^sub>n\<^sub>s\<^sub>s ?ns \ (\ c \ cs. v \\<^sub>l\<^sub>e c)" for v using leqc by fastforce hence "?lhs = (\ v. v \\<^sub>n\<^sub>s\<^sub>s ?ns)" by simp also have "\ = (\C. farkas_coefficients_ns ?ns C)" unfolding farkas_coefficients_ns[OF fin] .. also have "\ = ?rhs" proof assume "\ C. farkas_coefficients_ns ?ns C" then obtain C const where is_leq: "\ (s, n) \ set C. n \ ?ns \ is_leq_ns (s *R n) \ s \ 0" and sum: "(\(s, n)\C. lec_of_nsc (s *R n)) = Leqc 0 const" and c0: "const < 0" unfolding farkas_coefficients_ns_def by blast let ?C = "map (\ (s,n). (s,lec_of_nsc n)) C" show ?rhs proof (intro exI[of _ ?C] exI[of _ const] conjI c0, unfold sum[symmetric] map_map o_def set_map, intro ballI, clarify) { fix s n assume sn: "(s, n) \ set C" with is_leq have n_ns: "n \ ?ns" and is_leq: "is_leq_ns (s *R n)" "s \ 0" by force+ from n_ns obtain c where c: "c \ cs" and n: "n = LEQ_ns (lec_poly c) (lec_const c)" by auto with leqc[OF c] obtain p d where cs: "Leqc p d \ cs" and n: "n = LEQ_ns p d" by auto from is_leq[unfolded n] have s0: "s > 0" by (auto split: if_splits) let ?n = "lec_of_nsc n" from cs n have mem: "?n \ cs" by auto show "0 < s \ ?n \ cs" using s0 mem by blast have "Leqc (s *R lec_poly ?n) (s *R lec_const ?n) = lec_of_nsc (s *R n)" unfolding n using s0 by simp } note id = this show "(\x\C. case case x of (s, n) \ (s, lec_of_nsc n) of (r, c) \ Leqc (r *R lec_poly c) (r *R lec_const c)) = (\(s, n)\C. lec_of_nsc (s *R n))" (is "sum_list (map ?f C) = sum_list (map ?g C)") proof (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl]) fix pair assume mem: "pair \ set C" then obtain s n where pair: "pair = (s,n)" by force show "?f pair = ?g pair" unfolding pair split using id[OF mem[unfolded pair]] . qed qed next assume ?rhs then obtain C const where C: "\ r c. (r,c) \ set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" and const: "const < 0" by blast let ?C = "map (\ (r,c). (r, ?to_ns c)) C" show "\ C. farkas_coefficients_ns ?ns C" unfolding farkas_coefficients_ns_def proof (intro exI[of _ ?C] exI[of _ const] conjI const, unfold sum[symmetric]) show "\(s, n)\set ?C. n \ ?ns \ is_leq_ns (s *R n) \ s \ 0" using C by fastforce show "(\(s, n)\?C. lec_of_nsc (s *R n)) = (\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c))" unfolding map_map o_def by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, force) qed qed finally show ?thesis . qed subsection \Motzkin's Transposition Theorem or the Kuhn-Fourier Theorem\ text \Next, we prove a generalization of Farkas' Lemma that permits arbitrary combinations of strict and non-strict inequalities: Motzkin's Transposition Theorem which is also known as the Kuhn--Fourier Theorem. The proof is mainly based on the lemma @{thm [source] farkas_coefficients}, again requiring conversions between constraint types.\ theorem Motzkin's_transposition_theorem: fixes cs :: "rat le_constraint set" assumes fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const rel. (\ (r, c) \ set C. r > 0 \ c \ cs) \ (\ (r,c) \ C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const \ (rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0))" (is "?lhs = ?rhs") proof - let ?to_cs = "\ c. (case lec_rel c of Leq_Rel \ LEQ | _ \ LT) (lec_poly c) (lec_const c)" have to_cs: "v \\<^sub>c ?to_cs c \ v \\<^sub>l\<^sub>e c" for v c by (cases c, cases "lec_rel c", auto) let ?cs = "?to_cs ` cs" from fin have fin: "finite ?cs" by auto have "v \\<^sub>c\<^sub>s ?cs \ (\ c \ cs. v \\<^sub>l\<^sub>e c)" for v using to_cs by auto hence "?lhs = (\ v. v \\<^sub>c\<^sub>s ?cs)" by simp also have "\ = (\C. farkas_coefficients ?cs C)" unfolding farkas_coefficients[OF fin] .. also have "\ = ?rhs" proof assume "\ C. farkas_coefficients ?cs C" then obtain C const rel where is_leq: "\ (s, n) \ set C. n \ ?cs \ is_le (s *R n) \ s \ 0" and sum: "(\(s, n)\C. lec_of_constraint (s *R n)) = Le_Constraint rel 0 const" and c0: "(rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0)" unfolding farkas_coefficients_def by blast let ?C = "map (\ (s,n). (s,lec_of_constraint n)) C" show ?rhs proof (intro exI[of _ ?C] exI[of _ const] exI[of _ rel] conjI c0, unfold map_map o_def set_map sum[symmetric], intro ballI, clarify) { fix s n assume sn: "(s, n) \ set C" with is_leq have n_ns: "n \ ?cs" and is_leq: "is_le (s *R n)" and s0: "s \ 0" by force+ from n_ns obtain c where c: "c \ cs" and n: "n = ?to_cs c" by auto from is_leq[unfolded n] have "s \ 0" by (cases "lec_rel c", auto split: if_splits) with s0 have s0: "s > 0" by auto let ?c = "lec_of_constraint n" from c n have mem: "?c \ cs" by (cases c, cases "lec_rel c", auto) show "0 < s \ ?c \ cs" using s0 mem by blast have "lec_of_constraint (s *R n) = Le_Constraint (lec_rel ?c) (s *R lec_poly ?c) (s *R lec_const ?c)" unfolding n using s0 by (cases c, cases "lec_rel c", auto) } note id = this show "(\x\C. case case x of (s, n) \ (s, lec_of_constraint n) of (r, c) \ Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = (\(s, n)\C. lec_of_constraint (s *R n))" (is "sum_list (map ?f C) = sum_list (map ?g C)") proof (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl]) fix pair assume mem: "pair \ set C" obtain r c where pair: "pair = (r,c)" by force show "?f pair = ?g pair" unfolding pair split id[OF mem[unfolded pair]] .. qed qed next assume ?rhs then obtain C const rel where C: "\ r c. (r,c) \ set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const" and const: "rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0" by blast let ?C = "map (\ (r,c). (r, ?to_cs c)) C" show "\ C. farkas_coefficients ?cs C" unfolding farkas_coefficients_def proof (intro exI[of _ ?C] exI[of _ const] exI[of _ rel] conjI const, unfold sum[symmetric]) show "\(s, n)\set ?C. n \ ?cs \ is_le (s *R n) \ s \ 0" using C by (fastforce split: le_rel.splits) show "(\(s, n)\?C. lec_of_constraint (s *R n)) = (\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))" unfolding map_map o_def by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, fastforce split: le_rel.splits) qed qed finally show ?thesis . qed subsection \Farkas' Lemma\ text \Finally we derive the commonly used form of Farkas' Lemma, which easily follows from @{thm [source] Motzkin's_transposition_theorem}. It only permits non-strict inequalities and, as a result, the sum of inequalities will always be non-strict.\ lemma Farkas'_Lemma: fixes cs :: "rat le_constraint set" assumes only_non_strict: "lec_rel ` cs \ {Leq_Rel}" and fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const. (\ (r, c) \ set C. r > 0 \ c \ cs) \ (\ (r,c) \ C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const \ const < 0)" (is "_ = ?rhs") proof - { fix c assume "c \ cs" with only_non_strict have "lec_rel c = Leq_Rel" by auto then have "\ p const. c = Leqc p const" by (cases c, auto) } note leqc = this let ?lhs = "\C const rel. (\(r, c)\set C. 0 < r \ c \ cs) \ (\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const \ (rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0)" show ?thesis unfolding Motzkin's_transposition_theorem[OF fin] proof assume ?rhs then obtain C const where C: "\ r c. (r, c)\set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" and const: "const < 0" by blast show ?lhs proof (intro exI[of _ C] exI[of _ const] exI[of _ Leq_Rel] conjI) show "\ (r,c) \ set C. 0 < r \ c \ cs" using C by force show "(\(r, c)\ C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" unfolding sum[symmetric] by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, force dest!: leqc) qed (insert const, auto) next assume ?lhs then obtain C const rel where C: "\ r c. (r, c)\set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const" and const: "rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0" by blast have id: "(\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = (\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c))" (is "_ = ?sum") by (rule arg_cong[of _ _ sum_list], rule map_cong, auto dest!: C leqc) have "lec_rel ?sum = Leq_Rel" unfolding sum_list_lec by (auto simp add: sum_list_Leq_Rel o_def) with sum[unfolded id] have rel: "rel = Leq_Rel" by auto with const have const: "const < 0" by auto show ?rhs by (intro exI[of _ C] exI[of _ const] conjI const, insert sum id C rel, force+) qed qed text \We also present slightly modified versions\ lemma sum_list_map_filter_sum: fixes f :: "'a \ 'b :: comm_monoid_add" shows "sum_list (map f (filter g xs)) + sum_list (map f (filter (Not o g) xs)) = sum_list (map f xs)" by (induct xs, auto simp: ac_simps) text \A version where every constraint obtains exactly one coefficient and where 0 coefficients are allowed.\ lemma Farkas'_Lemma_set_sum: fixes cs :: "rat le_constraint set" assumes only_non_strict: "lec_rel ` cs \ {Leq_Rel}" and fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const. (\ c \ cs. C c \ 0) \ (\ c \ cs. Leqc ((C c) *R lec_poly c) ((C c) *R lec_const c)) = Leqc 0 const \ const < 0)" unfolding Farkas'_Lemma[OF only_non_strict fin] proof ((standard; elim exE conjE), goal_cases) case (2 C const) from finite_distinct_list[OF fin] obtain csl where csl: "set csl = cs" and dist: "distinct csl" by auto let ?list = "filter (\ c. C c \ 0) csl" let ?C = "map (\ c. (C c, c)) ?list" show ?case proof (intro exI[of _ ?C] exI[of _ const] conjI) have "(\(r, c)\?C. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = (\(r, c)\map (\c. (C c, c)) csl. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))" unfolding map_map by (rule sum_list_map_filter, auto simp: zero_le_constraint_def) also have "\ = Le_Constraint Leq_Rel 0 const" unfolding 2(2)[symmetric] csl[symmetric] unfolding sum.distinct_set_conv_list[OF dist] map_map o_def split .. finally show "(\(r, c)\?C. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint Leq_Rel 0 const" by auto show "const < 0" by fact show "\(r, c)\set ?C. 0 < r \ c \ cs" using 2(1) unfolding set_map set_filter csl by auto qed next case (1 C const) define CC where "CC = (\ c. sum_list (map fst (filter (\ rc. snd rc = c) C)))" show "(\ C const. (\ c \ cs. C c \ 0) \ (\ c \ cs. Leqc ((C c) *R lec_poly c) ((C c) *R lec_const c)) = Leqc 0 const \ const < 0)" proof (intro exI[of _ CC] exI[of _ const] conjI) show "\c\cs. 0 \ CC c" unfolding CC_def using 1(1) by (force intro!: sum_list_nonneg) show "const < 0" by fact from 1 have snd: "snd ` set C \ cs" by auto show "(\c\cs. Le_Constraint Leq_Rel (CC c *R lec_poly c) (CC c *R lec_const c)) = Le_Constraint Leq_Rel 0 const" unfolding 1(2)[symmetric] using fin snd unfolding CC_def proof (induct cs arbitrary: C rule: finite_induct) case empty hence C: "C = []" by auto thus ?case by simp next case *: (insert c cs C) let ?D = "filter (Not \ (\rc. snd rc = c)) C" from * have "snd ` set ?D \ cs" by auto note IH = *(3)[OF this] have id: "(\a\ ?D. case a of (r, c) \ Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = (\(r, c)\?D. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))" by (induct C, force+) show ?case unfolding sum.insert[OF *(1,2)] unfolding sum_list_map_filter_sum[of _ "\ rc. snd rc = c" C, symmetric] proof (rule arg_cong2[of _ _ _ _ "(+)"], goal_cases) case 2 show ?case unfolding IH[symmetric] by (rule sum.cong, insert *(2,1), auto intro!: arg_cong[of _ _ "\ xs. sum_list (map _ xs)"], (induct C, auto)+) next case 1 show ?case proof (rule sym, induct C) case (Cons rc C) thus ?case by (cases "rc", cases "snd rc = c", auto simp: field_simps scaleRat_left_distrib) qed (auto simp: zero_le_constraint_def) qed qed qed qed text \A version with indexed constraints, i.e., in particular where constraints may occur several times.\ lemma Farkas'_Lemma_indexed: fixes c :: "nat \ rat le_constraint" assumes only_non_strict: "lec_rel ` c ` Is \ {Leq_Rel}" and fin: "finite Is" shows "(\ v. \ i \ Is. v \\<^sub>l\<^sub>e c i) \ (\ C const. (\ i \ Is. C i \ 0) \ (\ i \ Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const \ const < 0)" proof - let ?C = "c ` Is" have fin: "finite ?C" using fin by auto have "(\ v. \ i \ Is. v \\<^sub>l\<^sub>e c i) = (\ v. \ cc \ ?C. v \\<^sub>l\<^sub>e cc)" by force also have "\ = (\ C const. (\ i \ Is. C i \ 0) \ (\ i \ Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const \ const < 0)" (is "?l = ?r") proof assume ?r then obtain C const where r: "(\ i \ Is. C i \ 0)" and eq: "(\ i \ Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const" and "const < 0" by auto from finite_distinct_list[OF \finite Is\] obtain Isl where isl: "set Isl = Is" and dist: "distinct Isl" by auto let ?CC = "filter (\ rc. fst rc \ 0) (map (\ i. (C i, c i)) Isl)" show ?l unfolding Farkas'_Lemma[OF only_non_strict fin] proof (intro exI[of _ ?CC] exI[of _ const] conjI) show "const < 0" by fact show "\ (r, ca) \ set ?CC. 0 < r \ ca \ ?C" using r(1) isl by auto show "(\(r, c)\?CC. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint Leq_Rel 0 const" unfolding eq[symmetric] by (subst sum_list_map_filter, force simp: zero_le_constraint_def, unfold map_map o_def, subst sum_list_distinct_conv_sum_set[OF dist], rule sum.cong, auto simp: isl) qed next assume ?l from this[unfolded Farkas'_Lemma_set_sum[OF only_non_strict fin]] obtain C const where nonneg: "(\c\ ?C. 0 \ C c)" and sum: "(\c\ ?C. Le_Constraint Leq_Rel (C c *R lec_poly c) (C c *R lec_const c)) = Le_Constraint Leq_Rel 0 const" and const: "const < 0" by blast define I where "I = (\ i. (C (c i) / rat_of_nat (card (Is \ { j. c i = c j}))))" show ?r proof (intro exI[of _ I] exI[of _ const] conjI const) show "\i \ Is. 0 \ I i" using nonneg unfolding I_def by auto show "(\ i \ Is. Le_Constraint Leq_Rel (I i *R lec_poly (c i)) (I i *R lec_const (c i))) = Le_Constraint Leq_Rel 0 const" unfolding sum[symmetric] unfolding sum.image_gen[OF \finite Is\, of _ c] proof (rule sum.cong[OF refl], goal_cases) case (1 cc) define II where "II = (Is \ {j. cc = c j})" from 1 have "II \ {}" unfolding II_def by auto moreover have finII: "finite II" using \finite Is\ unfolding II_def by auto ultimately have card: "card II \ 0" by auto let ?C = "\ II. rat_of_nat (card II)" define ii where "ii = C cc / rat_of_nat (card II)" have "(\i\{x \ Is. c x = cc}. Le_Constraint Leq_Rel (I i *R lec_poly (c i)) (I i *R lec_const (c i))) = (\ i\ II. Le_Constraint Leq_Rel (ii *R lec_poly cc) (ii *R lec_const cc))" unfolding I_def ii_def II_def by (rule sum.cong, auto) also have "\ = Le_Constraint Leq_Rel ((?C II * ii) *R lec_poly cc) ((?C II * ii) *R lec_const cc)" using finII by (induct II rule: finite_induct, auto simp: zero_le_constraint_def field_simps scaleRat_left_distrib) also have "?C II * ii = C cc" unfolding ii_def using card by auto finally show ?case . qed qed qed finally show ?thesis . qed -end \ No newline at end of file +end diff --git a/thys/Gabow_SCC/Gabow_GBG.thy b/thys/Gabow_SCC/Gabow_GBG.thy --- a/thys/Gabow_SCC/Gabow_GBG.thy +++ b/thys/Gabow_SCC/Gabow_GBG.thy @@ -1,2101 +1,2102 @@ section \Lasso Finding Algorithm for Generalized B\"uchi Graphs \label{sec:gbg}\ theory Gabow_GBG imports Gabow_Skeleton CAVA_Automata.Lasso Find_Path begin (* TODO: convenience locale, consider merging this with invariants *) locale igb_fr_graph = igb_graph G + fr_graph G for G :: "('Q,'more) igb_graph_rec_scheme" lemma igb_fr_graphI: assumes "igb_graph G" assumes "finite ((g_E G)\<^sup>* `` g_V0 G)" shows "igb_fr_graph G" proof - interpret igb_graph G by fact show ?thesis using assms(2) by unfold_locales qed text \ We implement an algorithm that computes witnesses for the non-emptiness of Generalized B\"uchi Graphs (GBG). \ section \Specification\ context igb_graph begin definition ce_correct \ \Specifies a correct counter-example\ where "ce_correct Vr Vl \ (\pr pl. Vr \ E\<^sup>*``V0 \ Vl \ E\<^sup>*``V0 \ \Only reachable nodes are covered\ \ set pr\Vr \ set pl\Vl \ \The paths are inside the specified sets\ \ Vl\Vl \ (E \ Vl\Vl)\<^sup>* \ \\Vl\ is mutually connected\ \ Vl\Vl \ E \ {} \ \\Vl\ is non-trivial\ \ is_lasso_prpl (pr,pl)) \ \Paths form a lasso\ " definition find_ce_spec :: "('Q set \ 'Q set) option nres" where "find_ce_spec \ SPEC (\r. case r of None \ (\prpl. \is_lasso_prpl prpl) | Some (Vr,Vl) \ ce_correct Vr Vl )" definition find_lasso_spec :: "('Q list \ 'Q list) option nres" where "find_lasso_spec \ SPEC (\r. case r of None \ (\prpl. \is_lasso_prpl prpl) | Some prpl \ is_lasso_prpl prpl )" end section \Invariant Extension\ text \Extension of the outer invariant:\ context igb_fr_graph begin definition no_acc_over \ \Specifies that there is no accepting cycle touching a set of nodes\ where "no_acc_over D \ \(\v\D. \pl. pl\[] \ path E v pl v \ (\iq\set pl. i\acc q))" definition "fgl_outer_invar_ext \ \it (brk,D). case brk of None \ no_acc_over D | Some (Vr,Vl) \ ce_correct Vr Vl" definition "fgl_outer_invar \ \it (brk,D). case brk of None \ outer_invar it D \ no_acc_over D | Some (Vr,Vl) \ ce_correct Vr Vl" end text \Extension of the inner invariant:\ locale fgl_invar_loc = invar_loc G v0 D0 p D pE + igb_graph G for G :: "('Q, 'more) igb_graph_rec_scheme" and v0 D0 and brk :: "('Q set \ 'Q set) option" and p D pE + assumes no_acc: "brk=None \ \(\v pl. pl\[] \ path lvE v pl v \ (\iq\set pl. i\acc q))" \ \No accepting cycle over visited edges\ assumes acc: "brk=Some (Vr,Vl) \ ce_correct Vr Vl" begin lemma locale_this: "fgl_invar_loc G v0 D0 brk p D pE" by unfold_locales lemma invar_loc_this: "invar_loc G v0 D0 p D pE" by unfold_locales lemma eas_gba_graph_this: "igb_graph G" by unfold_locales end definition (in igb_graph) "fgl_invar v0 D0 \ \(brk, p, D, pE). fgl_invar_loc G v0 D0 brk p D pE" section \Definition of the Lasso-Finding Algorithm\ context igb_fr_graph begin definition find_ce :: "('Q set \ 'Q set) option nres" where "find_ce \ do { let D = {}; (brk,_)\FOREACHci fgl_outer_invar V0 (\(brk,_). brk=None) (\v0 (brk,D0). do { if v0\D0 then do { let s = (None,initial v0 D0); (brk,p,D,pE) \ WHILEIT (fgl_invar v0 D0) (\(brk,p,D,pE). brk=None \ p \ []) (\(_,p,D,pE). do { \ \Select edge from end of path\ (vo,(p,D,pE)) \ select_edge (p,D,pE); ASSERT (p\[]); case vo of Some v \ do { if v \ \(set p) then do { \ \Collapse\ let (p,D,pE) = collapse v (p,D,pE); ASSERT (p\[]); if \iq\last p. i\acc q then RETURN (Some (\(set (butlast p)),last p),p,D,pE) else RETURN (None,p,D,pE) } else if v\D then do { \ \Edge to new node. Append to path\ RETURN (None,push v (p,D,pE)) } else RETURN (None,p,D,pE) } | None \ do { \ \No more outgoing edges from current node on path\ ASSERT (pE \ last p \ UNIV = {}); RETURN (None,pop (p,D,pE)) } }) s; ASSERT (brk=None \ (p=[] \ pE={})); RETURN (brk,D) } else RETURN (brk,D0) }) (None,D); RETURN brk }" end section \Invariant Preservation\ context igb_fr_graph begin definition "fgl_invar_part \ \(brk, p, D, pE). fgl_invar_loc_axioms G brk p D pE" lemma fgl_outer_invarI[intro?]: "\ brk=None \ outer_invar it D; \brk=None \ outer_invar it D\ \ fgl_outer_invar_ext it (brk,D)\ \ fgl_outer_invar it (brk,D)" unfolding outer_invar_def fgl_outer_invar_ext_def fgl_outer_invar_def apply (auto split: prod.splits option.splits) done lemma fgl_invarI[intro?]: "\ invar v0 D0 PDPE; invar v0 D0 PDPE \ fgl_invar_part (B,PDPE)\ \ fgl_invar v0 D0 (B,PDPE)" unfolding invar_def fgl_invar_part_def fgl_invar_def apply (simp split: prod.split_asm) apply intro_locales apply (simp add: invar_loc_def) apply assumption done lemma fgl_invar_initial: assumes OINV: "fgl_outer_invar it (None,D0)" assumes A: "v0\it" "v0\D0" shows "fgl_invar_part (None, initial v0 D0)" proof - from OINV interpret outer_invar_loc G it D0 by (simp add: fgl_outer_invar_def outer_invar_def) from OINV have no_acc: "no_acc_over D0" by (simp add: fgl_outer_invar_def fgl_outer_invar_ext_def) { fix v pl assume "pl \ []" and P: "path (vE [{v0}] D0 (E \ {v0} \ UNIV)) v pl v" hence 1: "v\D0" by (cases pl) (auto simp: path_cons_conv vE_def touched_def) have 2: "path E v pl v" using path_mono[OF vE_ss_E P] . note 1 2 } note AUX1=this show ?thesis unfolding fgl_invar_part_def apply (simp split: prod.splits add: initial_def) apply unfold_locales using \v0\D0\ using AUX1 no_acc unfolding no_acc_over_def apply blast by simp qed lemma fgl_invar_pop: assumes INV: "fgl_invar v0 D0 (None,p,D,pE)" assumes INV': "invar v0 D0 (pop (p,D,pE))" assumes NE[simp]: "p\[]" assumes NO': "pE \ last p \ UNIV = {}" shows "fgl_invar_part (None, pop (p,D,pE))" proof - from INV interpret fgl_invar_loc G v0 D0 None p D pE by (simp add: fgl_invar_def) show ?thesis apply (unfold fgl_invar_part_def pop_def) apply (simp split: prod.splits) apply unfold_locales unfolding vE_pop[OF NE] using no_acc apply auto [] apply simp done qed lemma fgl_invar_collapse_ce_aux: assumes INV: "invar v0 D0 (p, D, pE)" assumes NE[simp]: "p\[]" assumes NONTRIV: "vE p D pE \ (last p \ last p) \ {}" assumes ACC: "\iq\last p. i\acc q" shows "fgl_invar_part (Some (\(set (butlast p)), last p), p, D, pE)" proof - from INV interpret invar_loc G v0 D0 p D pE by (simp add: invar_def) txt \The last collapsed node on the path contains states from all accepting sets. As it is strongly connected and reachable, we get a counter-example. Here, we explicitely construct the lasso.\ let ?Er = "E \ (\(set (butlast p)) \ UNIV)" txt \We choose a node in the last Cnode, that is reachable only using former Cnodes.\ obtain w where "(v0,w)\?Er\<^sup>*" "w\last p" proof cases assume "length p = 1" hence "v0\last p" using root_v0 by (cases p) auto thus thesis by (auto intro: that) next assume "length p\1" hence "length p > 1" by (cases p) auto hence "Suc (length p - 2) < length p" by auto from p_connected'[OF this] obtain u v where UIP: "u\p!(length p - 2)" and VIP: "v\p!(length p - 1)" and "(u,v)\lvE" using \length p > 1\ by auto from root_v0 have V0IP: "v0\p!0" by (cases p) auto from VIP have "v\last p" by (cases p rule: rev_cases) auto from pathI[OF V0IP UIP] \length p > 1\ have "(v0,u)\(lvE \ \(set (butlast p)) \ \(set (butlast p)))\<^sup>*" (is "_ \ \\<^sup>*") by (simp add: path_seg_butlast) also have "\ \ ?Er" using lvE_ss_E by auto finally (rtrancl_mono_mp[rotated]) have "(v0,u)\?Er\<^sup>*" . also note \(u,v)\lvE\ UIP hence "(u,v)\?Er" using lvE_ss_E \length p > 1\ apply (auto simp: Bex_def in_set_conv_nth) by (metis One_nat_def Suc_lessE \Suc (length p - 2) < length p\ diff_Suc_1 length_butlast nth_butlast) finally show ?thesis by (rule that) fact qed then obtain "pr" where P_REACH: "path E v0 pr w" and R_SS: "set pr \ \(set (butlast p))" apply - apply (erule rtrancl_is_path) apply (frule path_nodes_edges) apply (auto dest!: order_trans[OF _ image_Int_subset] dest: path_mono[of _ E, rotated]) done have [simp]: "last p = p!(length p - 1)" by (cases p rule: rev_cases) auto txt \From that node, we construct a lasso by inductively appending a path for each accepting set\ { fix na assume na_def: "na = num_acc" have "\pl. pl\[] \ path (lvE \ last p\last p) w pl w \ (\iq\set pl. i\acc q)" using ACC unfolding na_def[symmetric] proof (induction na) case 0 from NONTRIV obtain u v where "(u,v)\lvE \ last p \ last p" "u\last p" "v\last p" by auto from cnode_connectedI \w\last p\ \u\last p\ have "(w,u)\(lvE \ last p \ last p)\<^sup>*" by auto also note \(u,v)\lvE \ last p \ last p\ also (rtrancl_into_trancl1) from cnode_connectedI \v\last p\ \w\last p\ have "(v,w)\(lvE \ last p \ last p)\<^sup>*" by auto finally obtain pl where "pl\[]" "path (lvE \ last p \ last p) w pl w" by (rule trancl_is_path) thus ?case by auto next case (Suc n) from Suc.prems have "\iq\last p. i\acc q" by auto with Suc.IH obtain pl where IH: "pl\[]" "path (lvE \ last p \ last p) w pl w" "\iq\set pl. i\acc q" by blast from Suc.prems obtain v where "v\last p" and "n\acc v" by auto from cnode_connectedI \w\last p\ \v\last p\ have "(w,v)\(lvE \ last p \ last p)\<^sup>*" by auto then obtain pl1 where P1: "path (lvE \ last p \ last p) w pl1 v" by (rule rtrancl_is_path) also from cnode_connectedI \w\last p\ \v\last p\ have "(v,w)\(lvE \ last p \ last p)\<^sup>*" by auto then obtain pl2 where P2: "path (lvE \ last p \ last p) v pl2 w" by (rule rtrancl_is_path) also (path_conc) note IH(2) finally (path_conc) have P: "path (lvE \ last p \ last p) w (pl1@pl2@pl) w" by simp moreover from IH(1) have "pl1@pl2@pl \ []" by simp moreover have "\i'q\set (pl1@pl2@pl). i'\acc q" using IH(3) by auto moreover have "v\set (pl1@pl2@pl)" using P1 P2 P IH(1) apply (cases pl2, simp_all add: path_cons_conv path_conc_conv) apply (cases pl, simp_all add: path_cons_conv) apply (cases pl1, simp_all add: path_cons_conv) done with \n\acc v\ have "\q\set (pl1@pl2@pl). n\acc q" by auto ultimately show ?case apply (intro exI conjI) apply assumption+ apply (auto elim: less_SucE) done qed } then obtain pl where pl: "pl\[]" "path (lvE \ last p\last p) w pl w" "\iq\set pl. i\acc q" by blast hence "path E w pl w" and L_SS: "set pl \ last p" apply - apply (frule path_mono[of _ E, rotated]) using lvE_ss_E apply auto [2] apply (drule path_nodes_edges) apply (drule order_trans[OF _ image_Int_subset]) apply auto [] done have LASSO: "is_lasso_prpl (pr,pl)" unfolding is_lasso_prpl_def is_lasso_prpl_pre_def using \path E w pl w\ P_REACH pl by auto from p_sc have "last p \ last p \ (lvE \ last p \ last p)\<^sup>*" by auto with lvE_ss_E have VL_CLOSED: "last p \ last p \ (E \ last p \ last p)\<^sup>*" apply (erule_tac order_trans) apply (rule rtrancl_mono) by blast have NONTRIV': "last p \ last p \ E \ {}" by (metis Int_commute NONTRIV disjoint_mono lvE_ss_E subset_refl) from order_trans[OF path_touched touched_reachable] have LP_REACH: "last p \ E\<^sup>*``V0" and BLP_REACH: "\(set (butlast p)) \ E\<^sup>*``V0" apply - apply (cases p rule: rev_cases) apply simp apply auto [] apply (cases p rule: rev_cases) apply simp apply auto [] done show ?thesis apply (simp add: fgl_invar_part_def) apply unfold_locales apply simp using LASSO R_SS L_SS VL_CLOSED NONTRIV' LP_REACH BLP_REACH unfolding ce_correct_def apply simp apply blast done qed lemma fgl_invar_collapse_ce: fixes u v assumes INV: "fgl_invar v0 D0 (None,p,D,pE)" defines "pE' \ pE - {(u,v)}" assumes CFMT: "(p',D',pE'') = collapse v (p,D,pE')" assumes INV': "invar v0 D0 (p',D',pE'')" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" and "u\last p" assumes BACK: "v\\(set p)" assumes ACC: "\iq\last p'. i\acc q" defines i_def: "i \ idx_of p v" shows "fgl_invar_part ( Some (\(set (butlast p')), last p'), collapse v (p,D,pE'))" proof - from CFMT have p'_def: "p' = collapse_aux p i" and [simp]: "D'=D" "pE''=pE'" by (simp_all add: collapse_def i_def) from INV interpret fgl_invar_loc G v0 D0 None p D pE by (simp add: fgl_invar_def) from idx_of_props[OF BACK] have "ip!i" by (simp_all add: i_def) have "u\last p'" using \u\last p\ \i unfolding p'_def collapse_aux_def apply (simp add: last_drop last_snoc) by (metis Misc.last_in_set drop_eq_Nil last_drop not_le) moreover have "v\last p'" using \v\p!i\ \i unfolding p'_def collapse_aux_def by (metis UnionI append_Nil Cons_nth_drop_Suc in_set_conv_decomp last_snoc) ultimately have "vE p' D pE' \ last p' \ last p' \ {}" unfolding p'_def pE'_def by (auto simp: E) have "p'\[]" by (simp add: p'_def collapse_aux_def) have [simp]: "collapse v (p,D,pE') = (p',D,pE')" unfolding collapse_def p'_def i_def by simp show ?thesis apply simp apply (rule fgl_invar_collapse_ce_aux) using INV' apply simp apply fact+ done qed lemma fgl_invar_collapse_nce: fixes u v assumes INV: "fgl_invar v0 D0 (None,p,D,pE)" defines "pE' \ pE - {(u,v)}" assumes CFMT: "(p',D',pE'') = collapse v (p,D,pE')" assumes INV': "invar v0 D0 (p',D',pE'')" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" and "u\last p" assumes BACK: "v\\(set p)" assumes NACC: "jq\last p'. j\acc q" defines "i \ idx_of p v" shows "fgl_invar_part (None, collapse v (p,D,pE'))" proof - from CFMT have p'_def: "p' = collapse_aux p i" and [simp]: "D'=D" "pE''=pE'" by (simp_all add: collapse_def i_def) have [simp]: "collapse v (p,D,pE') = (p',D,pE')" by (simp add: collapse_def p'_def i_def) from INV interpret fgl_invar_loc G v0 D0 None p D pE by (simp add: fgl_invar_def) from INV' interpret inv': invar_loc G v0 D0 p' D pE' by (simp add: invar_def) define vE' where "vE' = vE p' D pE'" have vE'_alt: "vE' = insert (u,v) lvE" by (simp add: vE'_def p'_def pE'_def E) from idx_of_props[OF BACK] have "ip!i" by (simp_all add: i_def) have "u\last p'" using \u\last p\ \i unfolding p'_def collapse_aux_def apply (simp add: last_drop last_snoc) by (metis Misc.last_in_set drop_eq_Nil last_drop leD) moreover have "v\last p'" using \v\p!i\ \i unfolding p'_def collapse_aux_def by (metis UnionI append_Nil Cons_nth_drop_Suc in_set_conv_decomp last_snoc) ultimately have "vE' \ last p' \ last p' \ {}" unfolding vE'_alt by (auto) have "p'\[]" by (simp add: p'_def collapse_aux_def) { txt \ We show that no visited strongly connected component contains states from all acceptance sets.\ fix w pl txt \For this, we chose a non-trivial loop inside the visited edges\ assume P: "path vE' w pl w" and NT: "pl\[]" txt \And show that there is one acceptance set disjoint with the nodes of the loop\ have "\iq\set pl. i\acc q" proof cases assume "set pl \ last p' = {}" \ \Case: The loop is outside the last Cnode\ with path_restrict[OF P] \u\last p'\ \v\last p'\ have "path lvE w pl w" apply - apply (drule path_mono[of _ lvE, rotated]) unfolding vE'_alt by auto with no_acc NT show ?thesis by auto next assume "set pl \ last p' \ {}" \ \Case: The loop touches the last Cnode\ txt \Then, the loop must be completely inside the last CNode\ from inv'.loop_in_lastnode[folded vE'_def, OF P \p'\[]\ this] have "w\last p'" "set pl \ last p'" . with NACC show ?thesis by blast qed } note AUX_no_acc = this show ?thesis apply (simp add: fgl_invar_part_def) apply unfold_locales using AUX_no_acc[unfolded vE'_def] apply auto [] apply simp done qed lemma collapse_ne: "([],D',pE') \ collapse v (p,D,pE)" by (simp add: collapse_def collapse_aux_def Let_def) lemma fgl_invar_push: assumes INV: "fgl_invar v0 D0 (None,p,D,pE)" assumes BRK[simp]: "brk=None" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" and UIL: "u\last p" assumes VNE: "v\\(set p)" "v\D" assumes INV': "invar v0 D0 (push v (p,D,pE - {(u,v)}))" shows "fgl_invar_part (None, push v (p,D,pE - {(u,v)}))" proof - from INV interpret fgl_invar_loc G v0 D0 None p D pE by (simp add: fgl_invar_def) define pE' where "pE' = (pE - {(u,v)} \ E\{v}\UNIV)" have [simp]: "push v (p,D,pE - {(u,v)}) = (p@[{v}],D,pE')" by (simp add: push_def pE'_def) from INV' interpret inv': invar_loc G v0 D0 "(p@[{v}])" D "pE'" by (simp add: invar_def) note defs_fold = vE_push[OF E UIL VNE, folded pE'_def] { txt \We show that there still is no loop that contains all accepting nodes. For this, we choose some loop.\ fix w pl assume P: "path (insert (u,v) lvE) w pl w" and [simp]: "pl\[]" have "\iq\set pl. i\acc q" proof cases assume "v\set pl" \ \Case: The newly pushed last cnode is on the loop\ txt \Then the loop is entirely on the last cnode\ with inv'.loop_in_lastnode[unfolded defs_fold, OF P] have [simp]: "w=v" and SPL: "set pl = {v}" by auto txt \However, we then either have that the last cnode is contained in the last but one cnode, or that there is a visited edge inside the last cnode.\ from P SPL have "u=v \ (v,v)\lvE" apply (cases pl) apply (auto simp: path_cons_conv) apply (case_tac list) apply (auto simp: path_cons_conv) done txt \Both leads to a contradiction\ hence False proof assume "u=v" \ \This is impossible, as @{text "u"} was on the original path, but @{text "v"} was not\ with UIL VNE show False by auto next assume "(v,v)\lvE" \ \This is impossible, as all visited edges are from touched nodes, but @{text "v"} was untouched\ with vE_touched VNE show False unfolding touched_def by auto qed thus ?thesis .. next assume A: "v\set pl" \ \Case: The newly pushed last cnode is not on the loop\ txt \Then, the path lays inside the old visited edges\ have "path lvE w pl w" proof - have "w\set pl" using P by (cases pl) (auto simp: path_cons_conv) with A show ?thesis using path_restrict[OF P] apply - apply (drule path_mono[of _ lvE, rotated]) apply (cases pl, auto) [] apply assumption done qed txt \And thus, the proposition follows from the invariant on the old state\ with no_acc show ?thesis apply simp using \pl\[]\ by blast qed } note AUX_no_acc = this show ?thesis unfolding fgl_invar_part_def apply simp apply unfold_locales unfolding defs_fold using AUX_no_acc apply auto [] apply simp done qed lemma fgl_invar_skip: assumes INV: "fgl_invar v0 D0 (None,p,D,pE)" assumes BRK[simp]: "brk=None" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" and UIL: "u\last p" assumes VID: "v\D" assumes INV': "invar v0 D0 (p, D, (pE - {(u,v)}))" shows "fgl_invar_part (None, p, D, (pE - {(u,v)}))" proof - from INV interpret fgl_invar_loc G v0 D0 None p D pE by (simp add: fgl_invar_def) from INV' interpret inv': invar_loc G v0 D0 p D "(pE - {(u,v)})" by (simp add: invar_def) { txt \We show that there still is no loop that contains all accepting nodes. For this, we choose some loop.\ fix w pl assume P: "path (insert (u,v) lvE) w pl w" and [simp]: "pl\[]" from P have "\iq\set pl. i\acc q" proof (cases rule: path_edge_rev_cases) case no_use \ \Case: The loop does not use the new edge\ txt \The proposition follows from the invariant for the old state\ with no_acc show ?thesis apply simp using \pl\[]\ by blast next case (split p1 p2) \ \Case: The loop uses the new edge\ txt \As done is closed under transitions, the nodes of the edge have already been visited\ from split(2) D_closed_vE_rtrancl have WID: "w\D" using VID by (auto dest!: path_is_rtrancl) from split(1) WID D_closed_vE_rtrancl have "u\D" apply (cases rule: path_edge_cases) apply (auto dest!: path_is_rtrancl) done txt \Which is a contradition to the assumptions\ with UIL p_not_D have False by (cases p rule: rev_cases) auto thus ?thesis .. qed } note AUX_no_acc = this show ?thesis apply (simp add: fgl_invar_part_def) apply unfold_locales unfolding vE_remove[OF NE E] using AUX_no_acc apply auto [] apply simp done qed lemma fgl_outer_invar_initial: "outer_invar V0 {} \ fgl_outer_invar_ext V0 (None, {})" unfolding fgl_outer_invar_ext_def apply (simp add: no_acc_over_def) done lemma fgl_outer_invar_brk: assumes INV: "fgl_invar v0 D0 (Some (Vr,Vl),p,D,pE)" shows "fgl_outer_invar_ext anyIt (Some (Vr,Vl), anyD)" proof - from INV interpret fgl_invar_loc G v0 D0 "Some (Vr,Vl)" p D pE by (simp add: fgl_invar_def) from acc show ?thesis by (simp add: fgl_outer_invar_ext_def) qed lemma fgl_outer_invar_newnode_nobrk: assumes A: "v0\D0" "v0\it" assumes OINV: "fgl_outer_invar it (None,D0)" assumes INV: "fgl_invar v0 D0 (None,[],D',pE)" shows "fgl_outer_invar_ext (it-{v0}) (None,D')" proof - from OINV interpret outer_invar_loc G it D0 unfolding fgl_outer_invar_def outer_invar_def by simp from INV interpret inv: fgl_invar_loc G v0 D0 None "[]" D' pE unfolding fgl_invar_def by simp from inv.pE_fin have [simp]: "pE = {}" by simp { fix v pl assume A: "v\D'" "path E v pl v" have "path (E \ D' \ UNIV) v pl v" apply (rule path_mono[OF _ path_restrict_closed[OF inv.D_closed A]]) by auto } note AUX1=this show ?thesis unfolding fgl_outer_invar_ext_def apply simp using inv.no_acc AUX1 apply (auto simp add: vE_def touched_def no_acc_over_def) [] done qed lemma fgl_outer_invar_newnode: assumes A: "v0\D0" "v0\it" assumes OINV: "fgl_outer_invar it (None,D0)" assumes INV: "fgl_invar v0 D0 (brk,p,D',pE)" assumes CASES: "(\Vr Vl. brk = Some (Vr, Vl)) \ p = []" shows "fgl_outer_invar_ext (it-{v0}) (brk,D')" using CASES apply (elim disjE1) using fgl_outer_invar_brk[of v0 D0 _ _ p D' pE] INV apply - apply (auto, assumption) [] using fgl_outer_invar_newnode_nobrk[OF A] OINV INV apply auto [] done lemma fgl_outer_invar_Dnode: assumes "fgl_outer_invar it (None, D)" "v\D" shows "fgl_outer_invar_ext (it - {v}) (None, D)" using assms by (auto simp: fgl_outer_invar_def fgl_outer_invar_ext_def) lemma fgl_fin_no_lasso: assumes A: "fgl_outer_invar {} (None, D)" assumes B: "is_lasso_prpl prpl" shows "False" proof - obtain "pr" pl where [simp]: "prpl = (pr,pl)" by (cases prpl) from A have NA: "no_acc_over D" by (simp add: fgl_outer_invar_def fgl_outer_invar_ext_def) from A have "outer_invar {} D" by (simp add: fgl_outer_invar_def) with fin_outer_D_is_reachable have [simp]: "D=E\<^sup>*``V0" by simp from NA B show False apply (simp add: no_acc_over_def is_lasso_prpl_def is_lasso_prpl_pre_def) apply clarsimp apply (blast dest: path_is_rtrancl) done qed lemma fgl_fin_lasso: assumes A: "fgl_outer_invar it (Some (Vr,Vl), D)" shows "ce_correct Vr Vl" using A by (simp add: fgl_outer_invar_def fgl_outer_invar_ext_def) lemmas fgl_invar_preserve = fgl_invar_initial fgl_invar_push fgl_invar_pop fgl_invar_collapse_ce fgl_invar_collapse_nce fgl_invar_skip fgl_outer_invar_newnode fgl_outer_invar_Dnode invar_initial outer_invar_initial fgl_invar_initial fgl_outer_invar_initial fgl_fin_no_lasso fgl_fin_lasso end section \Main Correctness Proof\ context igb_fr_graph begin lemma outer_invar_from_fgl_invarI: "fgl_outer_invar it (None,D) \ outer_invar it D" unfolding fgl_outer_invar_def outer_invar_def by (simp split: prod.splits) lemma invar_from_fgl_invarI: "fgl_invar v0 D0 (B,PDPE) \ invar v0 D0 PDPE" unfolding fgl_invar_def invar_def apply (simp split: prod.splits) unfolding fgl_invar_loc_def by simp theorem find_ce_correct: "find_ce \ find_ce_spec" proof - note [simp del] = Union_iff show ?thesis unfolding find_ce_def find_ce_spec_def select_edge_def select_def apply (refine_rcg WHILEIT_rule[where R="inv_image (abs_wf_rel v0) snd" for v0] refine_vcg ) using [[goals_limit = 5]] apply (vc_solve rec: fgl_invarI fgl_outer_invarI intro: invar_from_fgl_invarI outer_invar_from_fgl_invarI dest!: sym[of "collapse a b" for a b] simp: collapse_ne simp: pE_fin'[OF invar_from_fgl_invarI] finite_V0 solve: invar_preserve solve: asm_rl[of "_ \ _ = {}"] solve: fgl_invar_preserve) done qed end section "Emptiness Check" text \Using the lasso-finding algorithm, we can define an emptiness check\ context igb_fr_graph begin definition "abs_is_empty \ do { ce \ find_ce; RETURN (ce = None) }" theorem abs_is_empty_correct: "abs_is_empty \ SPEC (\res. res \ (\r. \is_acc_run r))" unfolding abs_is_empty_def apply (refine_rcg refine_vcg order_trans[OF find_ce_correct, unfolded find_ce_spec_def]) unfolding ce_correct_def using lasso_accepted accepted_lasso apply (clarsimp split: option.splits) apply (metis is_lasso_prpl_of_lasso surj_pair) by (metis is_lasso_prpl_conv) definition "abs_is_empty_ce \ do { ce \ find_ce; case ce of None \ RETURN None | Some (Vr,Vl) \ do { ASSERT (\pr pl. set pr \ Vr \ set pl \ Vl \ Vl \ Vl \ (E \ Vl\Vl)\<^sup>* \ is_lasso_prpl (pr,pl)); (pr,pl) \ SPEC (\(pr,pl). set pr \ Vr \ set pl \ Vl \ Vl \ Vl \ (E \ Vl\Vl)\<^sup>* \ is_lasso_prpl (pr,pl)); RETURN (Some (pr,pl)) } }" theorem abs_is_empty_ce_correct: "abs_is_empty_ce \ SPEC (\res. case res of None \ (\r. \is_acc_run r) | Some (pr,pl) \ is_acc_run (pr\pl\<^sup>\) )" unfolding abs_is_empty_ce_def apply (refine_rcg refine_vcg order_trans[OF find_ce_correct, unfolded find_ce_spec_def]) apply (clarsimp_all simp: ce_correct_def) using accepted_lasso finite_reachableE_V0 apply (metis is_lasso_prpl_of_lasso surj_pair) apply blast apply (simp add: lasso_prpl_acc_run) done end section \Refinement\ text \ In this section, we refine the lasso finding algorithm to use efficient data structures. First, we explicitely keep track of the set of acceptance classes for every c-node on the path. Second, we use Gabow's data structure to represent the path. \ subsection \Addition of Explicit Accepting Sets\ text \In a first step, we explicitely keep track of the current set of acceptance classes for every c-node on the path.\ type_synonym 'a abs_gstate = "nat set list \ 'a abs_state" type_synonym 'a ce = "('a set \ 'a set) option" type_synonym 'a abs_gostate = "'a ce \ 'a set" context igb_fr_graph begin definition gstate_invar :: "'Q abs_gstate \ bool" where "gstate_invar \ \(a,p,D,pE). a = map (\V. \(acc`V)) p" definition "gstate_rel \ br snd gstate_invar" lemma gstate_rel_sv[relator_props,simp,intro!]: "single_valued gstate_rel" by (simp add: gstate_rel_def) definition (in -) gcollapse_aux :: "nat set list \ 'a set list \ nat \ nat set list \ 'a set list" where "gcollapse_aux a p i \ (take i a @ [\(set (drop i a))],take i p @ [\(set (drop i p))])" definition (in -) gcollapse :: "'a \ 'a abs_gstate \ 'a abs_gstate" where "gcollapse v APDPE \ let (a,p,D,pE)=APDPE; i=idx_of p v; (a,p) = gcollapse_aux a p i in (a,p,D,pE)" definition "gpush v s \ let (a,s) = s in (a@[acc v],push v s)" definition "gpop s \ let (a,s) = s in (butlast a,pop s)" definition ginitial :: "'Q \ 'Q abs_gostate \ 'Q abs_gstate" where "ginitial v0 s0 \ ([acc v0], initial v0 (snd s0))" definition goinitial :: "'Q abs_gostate" where "goinitial \ (None,{})" definition go_is_no_brk :: "'Q abs_gostate \ bool" where "go_is_no_brk s \ fst s = None" definition goD :: "'Q abs_gostate \ 'Q set" where "goD s \ snd s" definition goBrk :: "'Q abs_gostate \ 'Q ce" where "goBrk s \ fst s" definition gto_outer :: "'Q ce \ 'Q abs_gstate \ 'Q abs_gostate" where "gto_outer brk s \ let (A,p,D,pE)=s in (brk,D)" definition "gselect_edge s \ do { let (a,s)=s; (r,s)\select_edge s; RETURN (r,a,s) }" definition gfind_ce :: "('Q set \ 'Q set) option nres" where "gfind_ce \ do { let os = goinitial; os\FOREACHci fgl_outer_invar V0 (go_is_no_brk) (\v0 s0. do { if v0\goD s0 then do { let s = (None,ginitial v0 s0); (brk,a,p,D,pE) \ WHILEIT (\(brk,a,s). fgl_invar v0 (goD s0) (brk,s)) (\(brk,a,p,D,pE). brk=None \ p \ []) (\(_,a,p,D,pE). do { \ \Select edge from end of path\ (vo,(a,p,D,pE)) \ gselect_edge (a,p,D,pE); ASSERT (p\[]); case vo of Some v \ do { if v \ \(set p) then do { \ \Collapse\ let (a,p,D,pE) = gcollapse v (a,p,D,pE); ASSERT (p\[]); ASSERT (a\[]); if last a = {0..(set (butlast p)),last p),a,p,D,pE) else RETURN (None,a,p,D,pE) } else if v\D then do { \ \Edge to new node. Append to path\ RETURN (None,gpush v (a,p,D,pE)) } else RETURN (None,a,p,D,pE) } | None \ do { \ \No more outgoing edges from current node on path\ ASSERT (pE \ last p \ UNIV = {}); RETURN (None,gpop (a,p,D,pE)) } }) s; ASSERT (brk=None \ (p=[] \ pE={})); RETURN (gto_outer brk (a,p,D,pE)) } else RETURN s0 }) os; RETURN (goBrk os) }" lemma gcollapse_refine: "\(v',v)\Id; (s',s)\gstate_rel\ \ (gcollapse v' s',collapse v s)\gstate_rel" unfolding gcollapse_def collapse_def collapse_aux_def gcollapse_aux_def apply (simp add: gstate_rel_def br_def Let_def) unfolding gstate_invar_def[abs_def] apply (auto split: prod.splits simp: take_map drop_map) done lemma gpush_refine: "\(v',v)\Id; (s',s)\gstate_rel\ \ (gpush v' s',push v s)\gstate_rel" unfolding gpush_def push_def apply (simp add: gstate_rel_def br_def) unfolding gstate_invar_def[abs_def] apply (auto split: prod.splits) done lemma gpop_refine: "\(s',s)\gstate_rel\ \ (gpop s',pop s)\gstate_rel" unfolding gpop_def pop_def apply (simp add: gstate_rel_def br_def) unfolding gstate_invar_def[abs_def] apply (auto split: prod.splits simp: map_butlast) done lemma ginitial_refine: "(ginitial x (None, b), initial x b) \ gstate_rel" unfolding ginitial_def gstate_rel_def br_def gstate_invar_def initial_def by auto lemma oinitial_b_refine: "((None,{}),(None,{}))\Id\\<^sub>rId" by simp lemma gselect_edge_refine: "\(s',s)\gstate_rel\ \ gselect_edge s' \\(\Id\option_rel \\<^sub>r gstate_rel) (select_edge s)" unfolding gselect_edge_def select_edge_def apply (simp add: pw_le_iff refine_pw_simps prod_rel_sv split: prod.splits option.splits) apply (auto simp: gstate_rel_def br_def gstate_invar_def) done lemma last_acc_impl: assumes "p\[]" assumes "((a',p',D',pE'),(p,D,pE))\gstate_rel" shows "(last a' = {0..iq\last p. i\acc q)" using assms acc_bound unfolding gstate_rel_def br_def gstate_invar_def by (auto simp: last_map) lemma fglr_aux1: assumes V: "(v',v)\Id" and S: "(s',s)\gstate_rel" and P: "\a' p' D' pE' p D pE. ((a',p',D',pE'),(p,D,pE))\gstate_rel \ f' a' p' D' pE' \\R (f p D pE)" shows "(let (a',p',D',pE') = gcollapse v' s' in f' a' p' D' pE') \ \R (let (p,D,pE) = collapse v s in f p D pE)" apply (auto split: prod.splits) apply (rule P) using gcollapse_refine[OF V S] apply simp done lemma gstate_invar_empty: "gstate_invar (a,[],D,pE) \ a=[]" "gstate_invar ([],p,D,pE) \ p=[]" by (auto simp add: gstate_invar_def) lemma find_ce_refine: "gfind_ce \\Id find_ce" unfolding gfind_ce_def find_ce_def unfolding goinitial_def go_is_no_brk_def[abs_def] goD_def goBrk_def gto_outer_def using [[goals_limit = 1]] apply (refine_rcg gselect_edge_refine prod_relI[OF IdI gpop_refine] prod_relI[OF IdI gpush_refine] fglr_aux1 last_acc_impl oinitial_b_refine inj_on_id ) apply refine_dref_type apply (simp_all add: ginitial_refine) apply (vc_solve (nopre) solve: asm_rl simp: gstate_rel_def br_def gstate_invar_empty) done end subsection \Refinement to Gabow's Data Structure\ subsubsection \Preliminaries\ definition Un_set_drop_impl :: "nat \ 'a set list \ 'a set nres" \ \Executable version of @{text "\set (drop i A)"}, using indexing to access @{text "A"}\ where "Un_set_drop_impl i A \ do { (_,res) \ WHILET (\(i,res). i < length A) (\(i,res). do { ASSERT (i res; let i = i + 1; RETURN (i,res) }) (i,{}); RETURN res }" lemma Un_set_drop_impl_correct: "Un_set_drop_impl i A \ SPEC (\r. r=\(set (drop i A)))" unfolding Un_set_drop_impl_def apply (refine_rcg WHILET_rule[where I="\(i',res). res=\(set ((drop i (take i' A)))) \ i\i'" and R="measure (\(i',_). length A - i')"] refine_vcg) apply (auto simp: take_Suc_conv_app_nth) done schematic_goal Un_set_drop_code_aux: assumes [autoref_rules]: "(es_impl,{})\\R\Rs" assumes [autoref_rules]: "(un_impl,(\))\\R\Rs\\R\Rs\\R\Rs" shows "(?c,Un_set_drop_impl)\nat_rel \ \\R\Rs\as_rel \ \\R\Rs\nres_rel" unfolding Un_set_drop_impl_def[abs_def] apply (autoref (trace, keep_goal)) done concrete_definition Un_set_drop_code uses Un_set_drop_code_aux schematic_goal Un_set_drop_tr_aux: "RETURN ?c \ Un_set_drop_code es_impl un_impl i A" unfolding Un_set_drop_code_def by refine_transfer concrete_definition Un_set_drop_tr for es_impl un_impl i A uses Un_set_drop_tr_aux lemma Un_set_drop_autoref[autoref_rules]: assumes "GEN_OP es_impl {} (\R\Rs)" assumes "GEN_OP un_impl (\) (\R\Rs\\R\Rs\\R\Rs)" shows "(\i A. RETURN (Un_set_drop_tr es_impl un_impl i A),Un_set_drop_impl) \nat_rel \ \\R\Rs\as_rel \ \\R\Rs\nres_rel" apply (intro fun_relI nres_relI) apply (rule order_trans[OF Un_set_drop_tr.refine]) using Un_set_drop_code.refine[of es_impl Rs R un_impl, param_fo, THEN nres_relD] using assms by simp subsubsection \Actual Refinement\ type_synonym 'Q gGS = "nat set list \ 'Q GS" type_synonym 'Q goGS = "'Q ce \ 'Q oGS" context igb_graph begin definition gGS_invar :: "'Q gGS \ bool" where "gGS_invar s \ let (a,S,B,I,P) = s in GS_invar (S,B,I,P) \ length a = length B \ \(set a) \ {0.. :: "'Q gGS \ 'Q abs_gstate" where "gGS_\ s \ let (a,s)=s in (a,GS.\ s)" definition "gGS_rel \ br gGS_\ gGS_invar" lemma gGS_rel_sv[relator_props,intro!,simp]: "single_valued gGS_rel" unfolding gGS_rel_def by auto definition goGS_invar :: "'Q goGS \ bool" where "goGS_invar s \ let (brk,ogs)=s in brk=None \ oGS_invar ogs" definition "goGS_\ s \ let (brk,ogs)=s in (brk,oGS_\ ogs)" definition "goGS_rel \ br goGS_\ goGS_invar" lemma goGS_rel_sv[relator_props,intro!,simp]: "single_valued goGS_rel" unfolding goGS_rel_def by auto end context igb_fr_graph begin lemma gGS_relE: assumes "(s',(a,p,D,pE))\gGS_rel" obtains S' B' I' P' where "s'=(a,S',B',I',P')" and "((S',B',I',P'),(p,D,pE))\GS_rel" and "length a = length B'" and "\(set a) \ {0.._def GS.\_def) apply (rule that) apply (simp only:) apply (auto simp: GS_rel_def br_def gGS_invar_def GS.\_def) done definition goinitial_impl :: "'Q goGS" where "goinitial_impl \ (None,Map.empty)" lemma goinitial_impl_refine: "(goinitial_impl,goinitial)\goGS_rel" by (auto simp: goinitial_impl_def goinitial_def goGS_rel_def br_def simp: goGS_\_def goGS_invar_def oGS_\_def oGS_invar_def) definition gto_outer_impl :: "'Q ce \ 'Q gGS \ 'Q goGS" where "gto_outer_impl brk s \ let (A,S,B,I,P)=s in (brk,I)" lemma gto_outer_refine: assumes A: "brk = None \ (p=[] \ pE={})" assumes B: "(s, (A,p, D, pE)) \ gGS_rel" assumes C: "(brk',brk)\Id" shows "(gto_outer_impl brk' s,gto_outer brk (A,p,D,pE))\goGS_rel" proof (cases s) fix A S B I P assume [simp]: "s=(A,S,B,I,P)" show ?thesis using C apply (cases brk) using assms I_to_outer[of S B I P D] apply (auto simp: goGS_rel_def br_def goGS_\_def gto_outer_def gto_outer_impl_def goGS_invar_def simp: gGS_rel_def oGS_rel_def GS_rel_def gGS_\_def gGS_invar_def GS.\_def) [] using B apply (auto simp: gto_outer_def gto_outer_impl_def simp: br_def goGS_rel_def goGS_invar_def goGS_\_def oGS_\_def simp: gGS_rel_def gGS_\_def GS.\_def GS.D_\_def ) done qed definition "gpush_impl v s \ let (a,s)=s in (a@[acc v], push_impl v s)" lemma gpush_impl_refine: assumes B: "(s',(a,p,D,pE))\gGS_rel" assumes A: "(v',v)\Id" assumes PRE: "v' \ \(set p)" "v' \ D" shows "(gpush_impl v' s', gpush v (a,p,D,pE))\gGS_rel" proof - from B obtain S' B' I' P' where [simp]: "s'=(a,S',B',I',P')" and OSR: "((S',B',I',P'),(p,D,pE))\GS_rel" and L: "length a = length B'" and R: "\(set a) \ {0.._def GS_rel_def br_def apply (auto dest: AUX1) done qed definition gpop_impl :: "'Q gGS \ 'Q gGS nres" where "gpop_impl s \ do { let (a,s)=s; s\pop_impl s; ASSERT (a\[]); let a = butlast a; RETURN (a,s) }" lemma gpop_impl_refine: assumes A: "(s',(a,p,D,pE))\gGS_rel" assumes PRE: "p \ []" "pE \ last p \ UNIV = {}" shows "gpop_impl s' \ \gGS_rel (RETURN (gpop (a,p,D,pE)))" proof - from A obtain S' B' I' P' where [simp]: "s'=(a,S',B',I',P')" and OSR: "((S',B',I',P'),(p,D,pE))\GS_rel" and L: "length a = length B'" and R: "\(set a) \ {0..[]" using L by (auto simp add: GS_rel_def br_def GS.\_def GS.p_\_def) { fix S B I P S' B' I' P' assume "nofail (pop_impl ((S, B, I, P)::'a GS))" "inres (pop_impl ((S, B, I, P)::'a GS)) (S', B', I', P')" hence "length B' = length B - Suc 0" apply (simp add: pop_impl_def GS.pop_impl_def Let_def refine_pw_simps) apply auto done } note AUX1=this from A L show ?thesis unfolding gpop_impl_def gpop_def gGS_rel_def gGS_\_def br_def apply (simp add: Let_def) using pop_refine[OF OSR PRE] apply (simp add: pw_le_iff refine_pw_simps split: prod.splits) unfolding gGS_rel_def gGS_invar_def gGS_\_def GS_rel_def GS.\_def br_def apply (auto dest!: AUX1 in_set_butlastD iff: Sup_le_iff) done qed definition gselect_edge_impl :: "'Q gGS \ ('Q option \ 'Q gGS) nres" where "gselect_edge_impl s \ do { let (a,s)=s; (vo,s)\select_edge_impl s; RETURN (vo,a,s) }" thm select_edge_refine lemma gselect_edge_impl_refine: assumes A: "(s', a, p, D, pE) \ gGS_rel" assumes PRE: "p \ []" shows "gselect_edge_impl s' \ \(Id \\<^sub>r gGS_rel) (gselect_edge (a, p, D, pE))" proof - from A obtain S' B' I' P' where [simp]: "s'=(a,S',B',I',P')" and OSR: "((S',B',I',P'),(p,D,pE))\GS_rel" and L: "length a = length B'" and R: "\(set a) \ {0.._def gGS_invar_def GS_rel_def GS.\_def apply (simp split: prod.splits) apply clarsimp using R apply (auto simp: L dest: AUX1) done qed term GS.idx_of_impl thm GS_invar.idx_of_correct definition gcollapse_impl_aux :: "'Q \ 'Q gGS \ 'Q gGS nres" where "gcollapse_impl_aux v s \ do { let (A,s)=s; \<^cancel>\ASSERT (v\\set (GS.p_\ s));\ i \ GS.idx_of_impl s v; s \ collapse_impl v s; ASSERT (i < length A); us \ Un_set_drop_impl i A; let A = take i A @ [us]; RETURN (A,s) }" term collapse lemma gcollapse_alt: "gcollapse v APDPE = ( let (a,p,D,pE)=APDPE; i=idx_of p v; s=collapse v (p,D,pE); us=\(set (drop i a)); a = take i a @ [us] in (a,s))" unfolding gcollapse_def gcollapse_aux_def collapse_def collapse_aux_def by auto thm collapse_refine lemma gcollapse_impl_aux_refine: assumes A: "(s', a, p, D, pE) \ gGS_rel" assumes B: "(v',v)\Id" assumes PRE: "v\\(set p)" shows "gcollapse_impl_aux v' s' \ \ gGS_rel (RETURN (gcollapse v (a, p, D, pE)))" proof - note [simp] = Let_def from A obtain S' B' I' P' where [simp]: "s'=(a,S',B',I',P')" and OSR: "((S',B',I',P'),(p,D,pE))\GS_rel" and L: "length a = length B'" and R: "\(set a) \ {0.. (S',B',I',P') = p" by (simp add: GS_rel_def br_def GS.\_def) from OSR PRE have PRE': "v \ \(set (GS.p_\ (S',B',I',P')))" by (simp add: GS_rel_def br_def GS.\_def) from OSR have GS_invar: "GS_invar (S',B',I',P')" by (simp add: GS_rel_def br_def) term GS.B { fix s assume "collapse v (p, D, pE) = (GS.p_\ s, GS.D_\ s, GS.pE_\ s)" hence "length (GS.B s) = Suc (idx_of p v)" unfolding collapse_def collapse_aux_def Let_def apply (cases s) apply (auto simp: GS.p_\_def) apply (drule arg_cong[where f=length]) using GS_invar.p_\_disjoint_sym[OF GS_invar] and PRE \GS.p_\ (S', B', I', P') = p\ idx_of_props(1)[of p v] by simp } note AUX1 = this show ?thesis unfolding gcollapse_alt gcollapse_impl_aux_def apply simp apply (rule RETURN_as_SPEC_refine) apply (refine_rcg order_trans[OF GS_invar.idx_of_correct[OF GS_invar PRE']] order_trans[OF collapse_refine[OF OSR B PRE, simplified]] refine_vcg ) using PRE' apply simp apply (simp add: L) using Un_set_drop_impl_correct acc_bound R apply (simp add: refine_pw_simps pw_le_iff) unfolding gGS_rel_def GS_rel_def GS.\_def br_def gGS_\_def gGS_invar_def apply (clarsimp simp: L dest!: AUX1) apply (auto dest!: AUX1 simp: L) apply (force dest!: in_set_dropD) [] apply (force dest!: in_set_takeD) [] done qed definition gcollapse_impl :: "'Q \ 'Q gGS \ 'Q gGS nres" where "gcollapse_impl v s \ do { let (A,S,B,I,P)=s; i \ GS.idx_of_impl (S,B,I,P) v; ASSERT (i+1 \ length B); let B = take (i+1) B; ASSERT (i < length A); us\Un_set_drop_impl i A; let A = take i A @ [us]; RETURN (A,S,B,I,P) }" lemma gcollapse_impl_aux_opt_refine: "gcollapse_impl v s \ gcollapse_impl_aux v s" unfolding gcollapse_impl_def gcollapse_impl_aux_def collapse_impl_def GS.collapse_impl_def apply (simp add: refine_pw_simps pw_le_iff split: prod.splits) apply blast done lemma gcollapse_impl_refine: assumes A: "(s', a, p, D, pE) \ gGS_rel" assumes B: "(v',v)\Id" assumes PRE: "v\\(set p)" shows "gcollapse_impl v' s' \ \ gGS_rel (RETURN (gcollapse v (a, p, D, pE)))" using order_trans[OF gcollapse_impl_aux_opt_refine gcollapse_impl_aux_refine[OF assms]] . definition ginitial_impl :: "'Q \ 'Q goGS \ 'Q gGS" where "ginitial_impl v0 s0 \ ([acc v0],initial_impl v0 (snd s0))" lemma ginitial_impl_refine: assumes A: "v0\goD s0" "go_is_no_brk s0" assumes REL: "(s0i,s0)\goGS_rel" "(v0i,v0)\Id" shows "(ginitial_impl v0i s0i,ginitial v0 s0)\gGS_rel" unfolding ginitial_impl_def ginitial_def using REL initial_refine[OF A(1) _ REL(2), of "snd s0i"] A(2) apply (auto simp: gGS_rel_def br_def gGS_\_def gGS_invar_def goGS_rel_def goGS_\_def simp: go_is_no_brk_def goD_def oGS_rel_def GS_rel_def goGS_invar_def split: prod.splits ) using acc_bound apply (fastforce simp: initial_impl_def GS_initial_impl_def)+ done definition gpath_is_empty_impl :: "'Q gGS \ bool" where "gpath_is_empty_impl s = path_is_empty_impl (snd s)" lemma gpath_is_empty_refine: "(s,(a,p,D,pE))\gGS_rel \ gpath_is_empty_impl s \ p=[]" unfolding gpath_is_empty_impl_def using path_is_empty_refine by (fastforce simp: gGS_rel_def br_def gGS_invar_def gGS_\_def GS.\_def) definition gis_on_stack_impl :: "'Q \ 'Q gGS \ bool" where "gis_on_stack_impl v s = is_on_stack_impl v (snd s)" lemma gis_on_stack_refine: "\(s,(a,p,D,pE))\gGS_rel\ \ gis_on_stack_impl v s \ v\\(set p)" unfolding gis_on_stack_impl_def using is_on_stack_refine by (fastforce simp: gGS_rel_def br_def gGS_invar_def gGS_\_def GS.\_def) definition gis_done_impl :: "'Q \ 'Q gGS \ bool" where "gis_done_impl v s \ is_done_impl v (snd s)" thm is_done_refine lemma gis_done_refine: "(s,(a,p,D,pE))\gGS_rel \ gis_done_impl v s \ (v \ D)" using is_done_refine[of "(snd s)" v] by (auto simp: gGS_rel_def br_def gGS_\_def gGS_invar_def GS.\_def gis_done_impl_def) definition (in -) "on_stack_less I u v \ case I v of Some (STACK j) \ j False" definition (in -) "on_stack_ge I l v \ case I v of Some (STACK j) \ l\j | _ \ False" lemma (in GS_invar) set_butlast_p_refine: assumes PRE: "p_\\[]" shows "Collect (on_stack_less I (last B)) = \(set (butlast p_\))" (is "?L=?R") proof (intro equalityI subsetI) from PRE have [simp]: "B\[]" by (auto simp: p_\_def) have [simp]: "S\[]" by (simp add: empty_eq) { fix v assume "v\?L" then obtain j where [simp]: "I v = Some (STACK j)" and "jj have "jj] find_seg_correct[OF \j] have "v\seg (find_seg j)" "find_seg j < length B" by auto moreover with \j have "find_seg j < length B - 1" (* What follows is an unreadable, auto-generated structured proof that replaces the following smt-call: by (smt GS.seg_start_def `seg_start (find_seg j) \ j`)*) proof - have f1: "\x\<^sub>1 x. \ (x\<^sub>1::nat) < x\<^sub>1 - x" using less_imp_diff_less by blast have "j \ last B" by (metis \j < last B\ less_le) hence f2: "\x\<^sub>1. \ last B < x\<^sub>1 \ \ x\<^sub>1 \ j" using f1 by (metis diff_diff_cancel le_trans) have "\x\<^sub>1. seg_end x\<^sub>1 \ j \ \ x\<^sub>1 < find_seg j" by (metis \seg_start (find_seg j) \ j\ calculation(2) le_trans seg_end_less_start) thus "find_seg j < length B - 1" using f1 f2 by (metis GS.seg_start_def \B \ []\ \j < B ! (length B - 1)\ \seg_start (find_seg j) \ j\ calculation(2) diff_diff_cancel last_conv_nth nat_neq_iff seg_start_less_end) qed ultimately show "v\?R" by (auto simp: p_\_def map_butlast[symmetric] butlast_upt) } { fix v assume "v\?R" then obtain i where "iseg i" by (auto simp: p_\_def map_butlast[symmetric] butlast_upt) then obtain j where "j < seg_end i" and "v=S!j" by (auto simp: seg_def) hence "j length B - 1" using \i by (auto simp: seg_end_def last_conv_nth split: if_split_asm) with sorted_nth_mono[OF B_sorted \i+1 \ length B - 1\] have "jj < seg_end i\ have "ji + 1 \ length B - 1\ add_lessD1 less_imp_diff_less less_le_not_le nat_neq_iff seg_end_bound) (*by (smt `i < length B - 1` seg_end_bound)*) with I_consistent \v=S!j\ have "I v = Some (STACK j)" by auto ultimately show "v\?L" by (auto simp: on_stack_less_def) } qed lemma (in GS_invar) set_last_p_refine: assumes PRE: "p_\\[]" shows "Collect (on_stack_ge I (last B)) = last p_\" (is "?L=?R") proof (intro equalityI subsetI) from PRE have [simp]: "B\[]" by (auto simp: p_\_def) have [simp]: "S\[]" by (simp add: empty_eq) { fix v assume "v\?L" then obtain j where [simp]: "I v = Some (STACK j)" and "j\last B" by (auto simp: on_stack_ge_def split: option.splits node_state.splits) from I_consistent[of v j] have [simp]: "jseg (length B - 1)" using \j\last B\ by (auto simp: seg_def last_conv_nth seg_start_def seg_end_def) thus "v\last p_\" by (auto simp: p_\_def last_map) } { fix v assume "v\?R" hence "v\seg (length B - 1)" by (auto simp: p_\_def last_map) then obtain j where "v=S!j" "j\last B" "jj\last B\ show "v\?L" by (auto simp: on_stack_ge_def) } qed definition ce_impl :: "'Q gGS \ (('Q set \ 'Q set) option \ 'Q gGS) nres" where "ce_impl s \ do { let (a,S,B,I,P) = s; ASSERT (B\[]); let bls = Collect (on_stack_less I (last B)); let ls = Collect (on_stack_ge I (last B)); RETURN (Some (bls, ls),a,S,B,I,P) }" lemma ce_impl_refine: assumes A: "(s,(a,p,D,pE))\gGS_rel" assumes PRE: "p\[]" shows "ce_impl s \ \(Id\\<^sub>rgGS_rel) (RETURN (Some (\(set (butlast p)),last p),a,p,D,pE))" proof - from A obtain S' B' I' P' where [simp]: "s=(a,S',B',I',P')" and OSR: "((S',B',I',P'),(p,D,pE))\GS_rel" and L: "length a = length B'" by (rule gGS_relE) from OSR have [simp]: "GS.p_\ (S',B',I',P') = p" by (simp add: GS_rel_def br_def GS.\_def) from PRE have NE': "GS.p_\ (S', B', I', P') \ []" by simp hence BNE[simp]: "B'\[]" by (simp add: GS.p_\_def) from OSR have GS_invar: "GS_invar (S',B',I',P')" by (simp add: GS_rel_def br_def) show ?thesis using GS_invar.set_butlast_p_refine[OF GS_invar NE'] using GS_invar.set_last_p_refine[OF GS_invar NE'] unfolding ce_impl_def using A by auto qed definition "last_is_acc_impl s \ do { let (a,_)=s; ASSERT (a\[]); RETURN (\ilast a) }" lemma last_is_acc_impl_refine: assumes A: "(s,(a,p,D,pE))\gGS_rel" assumes PRE: "a\[]" shows "last_is_acc_impl s \ RETURN (last a = {0.. {0.._def by auto hence C: "(\ilast a) \ (last a = {0.._def br_def split: prod.splits) show ?thesis unfolding last_is_acc_impl_def by (auto simp: gGS_rel_def br_def gGS_\_def C PRE split: prod.splits) qed definition go_is_no_brk_impl :: "'Q goGS \ bool" where "go_is_no_brk_impl s \ fst s = None" lemma go_is_no_brk_refine: "(s,s')\goGS_rel \ go_is_no_brk_impl s \ go_is_no_brk s'" unfolding go_is_no_brk_def go_is_no_brk_impl_def by (auto simp: goGS_rel_def br_def goGS_\_def split: prod.splits) definition goD_impl :: "'Q goGS \ 'Q oGS" where "goD_impl s \ snd s" lemma goD_refine: "go_is_no_brk s' \ (s,s')\goGS_rel \ (goD_impl s, goD s')\oGS_rel" unfolding goD_impl_def goD_def by (auto simp: goGS_rel_def br_def goGS_\_def goGS_invar_def oGS_rel_def go_is_no_brk_def) definition go_is_done_impl :: "'Q \ 'Q goGS \ bool" where "go_is_done_impl v s \ is_done_oimpl v (snd s)" thm is_done_orefine lemma go_is_done_impl_refine: "\go_is_no_brk s'; (s,s')\goGS_rel; (v,v')\Id\ \ go_is_done_impl v s \ (v'\goD s')" using is_done_orefine unfolding go_is_done_impl_def goD_def go_is_no_brk_def apply (fastforce simp: goGS_rel_def br_def goGS_invar_def goGS_\_def) done definition goBrk_impl :: "'Q goGS \ 'Q ce" where "goBrk_impl \ fst" lemma goBrk_refine: "(s,s')\goGS_rel \ (goBrk_impl s, goBrk s')\Id" unfolding goBrk_impl_def goBrk_def by (auto simp: goGS_rel_def br_def goGS_\_def split: prod.splits) definition find_ce_impl :: "('Q set \ 'Q set) option nres" where "find_ce_impl \ do { stat_start_nres; let os=goinitial_impl; os\FOREACHci (\it os. fgl_outer_invar it (goGS_\ os)) V0 (go_is_no_brk_impl) (\v0 s0. do { if \go_is_done_impl v0 s0 then do { let s = (None,ginitial_impl v0 s0); (brk,s)\WHILEIT (\(brk,s). fgl_invar v0 (oGS_\ (goD_impl s0)) (brk,snd (gGS_\ s))) (\(brk,s). brk=None \ \gpath_is_empty_impl s) (\(l,s). do { \ \Select edge from end of path\ (vo,s) \ gselect_edge_impl s; case vo of Some v \ do { if gis_on_stack_impl v s then do { s\gcollapse_impl v s; b\last_is_acc_impl s; if b then ce_impl s else RETURN (None,s) } else if \gis_done_impl v s then do { \ \Edge to new node. Append to path\ RETURN (None,gpush_impl v s) } else do { \ \Edge to done node. Skip\ RETURN (None,s) } } | None \ do { \ \No more outgoing edges from current node on path\ s\gpop_impl s; RETURN (None,s) } }) (s); RETURN (gto_outer_impl brk s) } else RETURN s0 }) os; stat_stop_nres; RETURN (goBrk_impl os) }" lemma find_ce_impl_refine: "find_ce_impl \ \Id gfind_ce" proof - note [refine2] = prod_relI[OF IdI[of None] ginitial_impl_refine] have [refine]: "\s a p D pE. \ (s,(a,p,D,pE))\gGS_rel; p \ []; pE \ last p \ UNIV = {} \ \ gpop_impl s \ (\s. RETURN (None, s)) \ SPEC (\c. (c, None, gpop (a,p,D,pE)) \ Id \\<^sub>r gGS_rel)" apply (drule (2) gpop_impl_refine) apply (fastforce simp add: pw_le_iff refine_pw_simps) done note [[goals_limit = 1]] note FOREACHci_refine_rcg'[refine del] show ?thesis unfolding find_ce_impl_def gfind_ce_def apply (refine_rcg bind_refine' prod_relI IdI inj_on_id gselect_edge_impl_refine gpush_impl_refine oinitial_refine ginitial_impl_refine bind_Let_refine2[OF gcollapse_impl_refine] if_bind_cond_refine[OF last_is_acc_impl_refine] ce_impl_refine goinitial_impl_refine gto_outer_refine goBrk_refine FOREACHci_refine_rcg'[where R=goGS_rel, OF inj_on_id] ) apply refine_dref_type apply (simp_all add: go_is_no_brk_refine go_is_done_impl_refine) apply (auto simp: goGS_rel_def br_def) [] apply (auto simp: goGS_rel_def br_def goGS_\_def gGS_\_def gGS_rel_def goD_def goD_impl_def) [] apply (auto dest: gpath_is_empty_refine ) [] apply (auto dest: gis_on_stack_refine) [] apply (auto dest: gis_done_refine) [] done qed end section \Constructing a Lasso from Counterexample\ subsection \Lassos in GBAs\ context igb_fr_graph begin definition reconstruct_reach :: "'Q set \ 'Q set \ ('Q list \ 'Q) nres" \ \Reconstruct the reaching path of a lasso\ where "reconstruct_reach Vr Vl \ do { res \ find_path (E\Vr\UNIV) V0 (\v. v\Vl); ASSERT (res \ None); RETURN (the res) }" lemma reconstruct_reach_correct: assumes CEC: "ce_correct Vr Vl" shows "reconstruct_reach Vr Vl \ SPEC (\(pr,va). \v0\V0. path E v0 pr va \ va\Vl)" proof - have FIN_aux: "finite ((E \ Vr \ UNIV)\<^sup>* `` V0)" by (metis finite_reachableE_V0 finite_subset inf_sup_ord(1) inf_sup_ord(2) inf_top.left_neutral reachable_mono) { fix u p v assume P: "path E u p v" and SS: "set p \ Vr" have "path (E \ Vr\UNIV) u p v" apply (rule path_mono[OF _ path_restrict[OF P]]) using SS by auto } note P_CONV=this from CEC obtain v0 "pr" va where "v0\V0" "set pr \ Vr" "va\Vl" "path (E \ Vr\UNIV) v0 pr va" unfolding ce_correct_def is_lasso_prpl_def is_lasso_prpl_pre_def by (force simp: neq_Nil_conv path_simps dest: P_CONV) hence 1: "va \ (E \ Vr \ UNIV)\<^sup>* `` V0" by (auto dest: path_is_rtrancl) show ?thesis using assms unfolding reconstruct_reach_def apply (refine_rcg refine_vcg order_trans[OF find_path_ex_rule]) apply (clarsimp_all simp: FIN_aux finite_V0) using \va\Vl\ 1 apply auto [] apply (auto dest: path_mono[of "E \ Vr \ UNIV" E, simplified]) [] done qed definition "rec_loop_invar Vl va s \ let (v,p,cS) = s in va \ E\<^sup>*``V0 \ path E va p v \ cS = acc v \ (\(acc`set p)) \ va \ Vl \ v \ Vl \ set p \ Vl" definition reconstruct_lasso :: "'Q set \ 'Q set \ ('Q list \ 'Q list) nres" \ \Reconstruct lasso\ where "reconstruct_lasso Vr Vl \ do { (pr,va) \ reconstruct_reach Vr Vl; let cS_full = {0.. UNIV\Vl; (vd,p,_) \ WHILEIT (rec_loop_invar Vl va) (\(_,_,cS). cS \ cS_full) (\(v,p,cS). do { ASSERT (\v'. (v,v')\E\<^sup>* \ \ (acc v' \ cS)); sr \ find_path E {v} (\v. \ (acc v \ cS)); ASSERT (sr \ None); let (p_seg,v) = the sr; RETURN (v,p@p_seg,cS \ acc v) }) (va,[],acc va); p_close_r \ (if p=[] then find_path1 E vd ((=) va) else find_path E {vd} ((=) va)); ASSERT (p_close_r \ None); let (p_close,_) = the p_close_r; RETURN (pr, p@p_close) }" lemma (in igb_fr_graph) reconstruct_lasso_correct: assumes CEC: "ce_correct Vr Vl" shows "reconstruct_lasso Vr Vl \ SPEC (is_lasso_prpl)" proof - let ?E = "E \ UNIV \ Vl" have E_SS: "E \ Vl \ Vl \ ?E" by auto from CEC have REACH: "Vl \ E\<^sup>*``V0" and CONN: "Vl\Vl \ (E \ Vl\Vl)\<^sup>*" and NONTRIV: "Vl\Vl \ E \ {}" and NES[simp]: "Vl\{}" and ALL: "\(acc`Vl) = {0..(_::'Q,_::'Q list,cS). cS))" hence WF: "wf term_rel" by simp { fix va assume "va \ Vl" hence "rec_loop_invar Vl va (va, [], acc va)" unfolding rec_loop_invar_def using REACH by auto } note INVAR_INITIAL = this { fix v p cS va assume "rec_loop_invar Vl va (v, p, cS)" hence "finite ((?E)\<^sup>* `` {v})" apply - apply (rule finite_subset[where B="E\<^sup>*``V0"]) unfolding rec_loop_invar_def using REACH apply (clarsimp_all dest!: path_is_rtrancl) apply (drule rtrancl_mono_mp[where U="?E" and V=E, rotated], (auto) []) by (metis rev_ImageI rtrancl_trans) } note FIN1 = this { fix va v :: 'Q and p cS assume INV: "rec_loop_invar Vl va (v,p,cS)" and NC: "cS \ {0..cS" unfolding rec_loop_invar_def by auto blast with ALL obtain v' where "v'\Vl" "\ acc v' \ cS" by simp (smt UN_iff atLeastLessThan_iff le0 subsetCE) moreover with CONN INV have "(v,v')\(E \ Vl \ Vl)\<^sup>*" unfolding rec_loop_invar_def by auto hence "(v,v')\?E\<^sup>*" using rtrancl_mono_mp[OF E_SS] by blast ultimately have "\v'. (v,v')\(?E)\<^sup>* \ \ acc v' \ cS" by auto } note ASSERT1 = this { fix va v p cS v' p' assume "rec_loop_invar Vl va (v, p, cS)" and "path (?E) v p' v'" and "\ (acc v' \ cS)" and "\v\set p'. acc v \ cS" hence "rec_loop_invar Vl va (v', p@p', cS \ acc v')" unfolding rec_loop_invar_def apply simp apply (intro conjI) apply (auto simp: path_simps dest: path_mono[of "?E" E, simplified]) [] apply (cases p') apply (auto simp: path_simps) [2] apply (cases p' rule: rev_cases) apply (auto simp: path_simps) [2] apply (erule path_set_induct) apply auto [2] done } note INV_PRES = this { fix va v p cS v' p' assume "rec_loop_invar Vl va (v, p, cS)" and "path ?E v p' v'" and "\ (acc v' \ cS)" and "\v\set p'. acc v \ cS" hence "((v', p@p', cS \ acc v'), (v,p,cS)) \ term_rel" unfolding term_rel_def rec_loop_invar_def by (auto simp: finite_psupset_def) } note VAR = this have CONN1: "Vl \ Vl \ (?E)\<^sup>+" proof clarify fix a b assume "a\Vl" "b\Vl" from NONTRIV obtain u v where E: "(u,v)\(E \ Vl\Vl)" by auto from CONN \a\Vl\ E have "(a,u)\(E\Vl\Vl)\<^sup>*" by auto also note E also (rtrancl_into_trancl1) from CONN \b\Vl\ E have "(v,b)\(E\Vl\Vl)\<^sup>*" by auto finally show "(a,b)\(?E)\<^sup>+" using trancl_mono[OF _ E_SS] by auto qed { fix va v p cS assume "rec_loop_invar Vl va (v, p, cS)" hence "(v,va) \ (?E)\<^sup>+" unfolding rec_loop_invar_def using CONN1 by auto } note CLOSE1 = this { fix va v p cS assume "rec_loop_invar Vl va (v, p, cS)" hence "(v,va) \ (?E)\<^sup>*" unfolding rec_loop_invar_def using CONN rtrancl_mono[OF E_SS] by auto } note CLOSE2 = this { fix "pr" vd pl va v0 assume "rec_loop_invar Vl va (vd, [], {0.. Vl" "v0 \ V0" "path E v0 pr va" "pl \ []" "path ?E vd pl va" hence "is_lasso_prpl (pr, pl)" unfolding is_lasso_prpl_def is_lasso_prpl_pre_def rec_loop_invar_def by (auto simp: neq_Nil_conv path_simps dest!: path_mono[of "?E" E, simplified]) [] } note INV_POST1 = this { fix va v p p' "pr" v0 assume INV: "rec_loop_invar Vl va (v,p,{0..[]" "path ?E v p' va" and PR: "v0\V0" "path E v0 pr va" from INV have "\iq\insert v (set p). i \ acc q" unfolding rec_loop_invar_def by auto moreover from INV 1 have "insert v (set p) \ set p \ set p'" unfolding rec_loop_invar_def - apply (cases p, simp) + apply (cases p) + apply blast apply (cases p') apply (auto simp: path_simps) done ultimately have ACC: "\iq\set p \ set p'. i \ acc q" by blast from INV 1 have PL: "path E va (p @ p') va" by (auto simp: rec_loop_invar_def path_simps dest!: path_mono[of "?E" E, simplified]) [] have "is_lasso_prpl (pr,p@p')" unfolding is_lasso_prpl_def is_lasso_prpl_pre_def apply (clarsimp simp: ACC) using PR PL \p\[]\ by auto } note INV_POST2 = this show ?thesis unfolding reconstruct_lasso_def apply (refine_rcg WF order_trans[OF reconstruct_reach_correct] order_trans[OF find_path_ex_rule] order_trans[OF find_path1_ex_rule] refine_vcg ) apply (vc_solve del: subsetI solve: ASSERT1 INV_PRES asm_rl VAR CLOSE1 CLOSE2 INV_POST1 INV_POST2 simp: INVAR_INITIAL FIN1 CEC) done qed definition find_lasso where "find_lasso \ do { ce \ find_ce_spec; case ce of None \ RETURN None | Some (Vr,Vl) \ do { l \ reconstruct_lasso Vr Vl; RETURN (Some l) } }" lemma (in igb_fr_graph) find_lasso_correct: "find_lasso \ find_lasso_spec" unfolding find_lasso_spec_def find_lasso_def find_ce_spec_def apply (refine_rcg refine_vcg order_trans[OF reconstruct_lasso_correct]) apply auto done end end diff --git a/thys/JinjaThreads/BV/BVSpecTypeSafe.thy b/thys/JinjaThreads/BV/BVSpecTypeSafe.thy --- a/thys/JinjaThreads/BV/BVSpecTypeSafe.thy +++ b/thys/JinjaThreads/BV/BVSpecTypeSafe.thy @@ -1,1841 +1,1844 @@ (* Title: JinjaThreads/BV/BVSpecTypeSafe.thy Author: Cornelia Pusch, Gerwin Klein, Andreas Lochbihler *) section \BV Type Safety Proof \label{sec:BVSpecTypeSafe}\ theory BVSpecTypeSafe imports BVConform "../Common/ExternalCallWF" begin declare listE_length [simp del] text \ This theory contains proof that the specification of the bytecode verifier only admits type safe programs. \ subsection \Preliminaries\ text \ Simp and intro setup for the type safety proof: \ context JVM_heap_conf_base begin lemmas widen_rules [intro] = conf_widen confT_widen confs_widens confTs_widen end subsection \Exception Handling\ text \ For the \Invoke\ instruction the BV has checked all handlers that guard the current \pc\. \ lemma Invoke_handlers: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set (relevant_entries P (Invoke n M) pc xt). (case D of None \ True | Some D' \ P \ C \\<^sup>* D') \ pc \ {f.. pc' = h \ d' = d" by (induct xt) (auto simp add: relevant_entries_def matches_ex_entry_def is_relevant_entry_def split: if_split_asm) lemma match_is_relevant: assumes rv: "\D'. P \ D \\<^sup>* D' \ is_relevant_class (ins ! i) P D'" assumes match: "match_ex_table P D pc xt = Some (pc',d')" shows "\(f,t,D',h,d) \ set (relevant_entries P (ins ! i) pc xt). (case D' of None \ True | Some D'' \ P \ D \\<^sup>* D'') \ pc \ {f.. pc' = h \ d' = d" using rv match by(fastforce simp add: relevant_entries_def is_relevant_entry_def matches_ex_entry_def dest: match_ex_table_SomeD) context JVM_heap_conf_base begin lemma exception_step_conform: fixes \' :: "('addr, 'heap) jvm_state" assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes correct: "\ \ t:(\xcp\, h, fr # frs) \" shows "\ \ t:exception_step P xcp h fr frs \" proof - obtain stk loc C M pc where fr: "fr = (stk, loc, C, M, pc)" by(cases fr) from correct obtain Ts T mxs mxl\<^sub>0 ins xt where meth: "P \ C sees M:Ts \ T = \(mxs,mxl\<^sub>0,ins,xt)\ in C" by (simp add: correct_state_def fr) blast from correct meth fr obtain D where hxcp: "typeof_addr h xcp = \Class_type D\" and DsubThrowable: "P \ D \\<^sup>* Throwable" and rv: "\D'. P \ D \\<^sup>* D' \ is_relevant_class (instrs_of P C M ! pc) P D'" by(fastforce simp add: correct_state_def dest: sees_method_fun) from meth have [simp]: "ex_table_of P C M = xt" by simp from correct have tconf: "P,h \ t \t" by(simp add: correct_state_def) show ?thesis proof(cases "match_ex_table P D pc xt") case None with correct fr meth hxcp show ?thesis by(fastforce simp add: correct_state_def cname_of_def split: list.split) next case (Some pc_d) then obtain pc' d' where pcd: "pc_d = (pc', d')" and match: "match_ex_table P D pc xt = Some (pc',d')" by (cases pc_d) auto from match_is_relevant[OF rv match] meth obtain f t D' where rv: "(f, t, D', pc', d') \ set (relevant_entries P (ins ! pc) pc xt)" and DsubD': "(case D' of None \ True | Some D'' \ P \ D \\<^sup>* D'')" and pc: "pc \ {f.._pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" unfolding correct_state_def fr by(auto dest: sees_method_fun) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" by (unfold conf_f_def) auto from stk have [simp]: "size stk = size ST" .. from wtp meth correct fr have wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (auto simp add: correct_state_def conf_f_def dest: sees_method_fun elim!: wt_jvm_prog_impl_wt_instr) from wt \_pc have eff: "\(pc', s')\set (xcpt_eff (ins!pc) P pc (ST,LT) xt). pc' < size ins \ P \ s' \' \ C M!pc'" by (auto simp add: defs1) let ?stk' = "Addr xcp # drop (length stk - d') stk" let ?f = "(?stk', loc, C, M, pc')" have conf: "P,h \ Addr xcp :\ Class (case D' of None \ Throwable | Some D'' \ D'')" using DsubD' hxcp DsubThrowable by(auto simp add: conf_def) obtain ST' LT' where \_pc': "\ C M ! pc' = Some (ST', LT')" and pc': "pc' < size ins" and less: "P \ (Class D # drop (size ST - d') ST, LT) \\<^sub>i (ST', LT')" proof(cases D') case Some thus ?thesis using eff rv DsubD' conf that by(fastforce simp add: xcpt_eff_def sup_state_opt_any_Some intro: widen_trans[OF widen_subcls]) next case None with that eff rv conf DsubThrowable show ?thesis by(fastforce simp add: xcpt_eff_def sup_state_opt_any_Some intro: widen_trans[OF widen_subcls]) qed with conf loc stk hxcp have "conf_f P h (ST',LT') ins ?f" by (auto simp add: defs1 conf_def intro: list_all2_dropI) with meth h_ok frames \_pc' fr match hxcp tconf preh show ?thesis unfolding correct_state_def by(fastforce dest: sees_method_fun simp add: cname_of_def) qed qed end subsection \Single Instructions\ text \ In this subsection we prove for each single (welltyped) instruction that the state after execution of the instruction still conforms. Since we have already handled raised exceptions above, we can now assume that no exception has been raised in this step. \ context JVM_conf_read begin declare defs1 [simp] lemma Invoke_correct: fixes \' :: "('addr, 'heap) jvm_state" assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth_C: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins ! pc = Invoke M' n" assumes wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes approx: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes exec: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\ \" proof - note split_paired_Ex [simp del] from wtprog obtain wfmb where wfprog: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from ins meth_C approx obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (fastforce dest: sees_method_fun) from ins wti \_pc have n: "n < size ST" by simp show ?thesis proof(cases "stk!n = Null") case True with ins heap_ok \_pc frame frames exec meth_C tconf preh show ?thesis by(fastforce elim: wf_preallocatedE[OF wfprog, where C=NullPointer]) next case False note Null = this have NT: "ST!n \ NT" proof assume "ST!n = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with n have "P,h \ stk!n :\ ST!n" by (simp add: list_all2_conv_all_nth) ultimately have "stk!n = Null" by simp with Null show False by contradiction qed from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" by simp from NT ins wti \_pc have pc': "pc+1 < size ins" by simp from NT ins wti \_pc obtain ST' LT' where pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (ST', LT')" and LT': "P \ LT [\\<^sub>\] LT'" by(auto simp add: neq_Nil_conv sup_state_opt_any_Some split: if_split_asm) with NT ins wti \_pc obtain D D' TTs TT m where D: "class_type_of' (ST!n) = \D\" and m_D: "P \ D sees M': TTs\TT = m in D'" and Ts: "P \ rev (take n ST) [\] TTs" and ST': "P \ (TT # drop (n+1) ST) [\] ST'" by(auto) from n stk D have "P,h \ stk!n :\ ST ! n" by (auto simp add: list_all2_conv_all_nth) from \P,h \ stk!n :\ ST ! n\ Null D obtain U a where Addr: "stk!n = Addr a" and obj: "typeof_addr h a = Some U" and UsubSTn: "P \ ty_of_htype U \ ST ! n" by(cases "stk ! n")(auto simp add: conf_def widen_Class) from D UsubSTn obtain C' where C': "class_type_of' (ty_of_htype U) = \C'\" and C'subD: "P \ C' \\<^sup>* D" by(rule widen_is_class_type_of) simp with wfprog m_D obtain Ts' T' D'' meth' where m_C': "P \ C' sees M': Ts'\T' = meth' in D''" and T': "P \ T' \ TT" and Ts': "P \ TTs [\] Ts'" by (auto dest: sees_method_mono) from Ts n have [simp]: "size TTs = n" by (auto dest: list_all2_lengthD simp: min_def) with Ts' have [simp]: "size Ts' = n" by (auto dest: list_all2_lengthD) from m_C' wfprog obtain mD'': "P \ D'' sees M':Ts'\T'=meth' in D''" by (fast dest: sees_method_idemp) { fix mxs' mxl' ins' xt' assume [simp]: "meth' = \(mxs', mxl', ins', xt')\" let ?loc' = "Addr a # rev (take n stk) @ replicate mxl' undefined_value" let ?f' = "([], ?loc', D'', M', 0)" let ?f = "(stk, loc, C, M, pc)" from Addr obj m_C' ins meth_C exec C' False have s': "\ = (None, h, ?f' # ?f # frs)" by(auto split: if_split_asm) moreover from wtprog mD'' obtain start: "wt_start P D'' Ts' mxl' (\ D'' M')" and ins': "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT\<^sub>0 where LT\<^sub>0: "\ D'' M' ! 0 = Some ([], LT\<^sub>0)" by (clarsimp simp add: wt_start_def defs1 sup_state_opt_any_Some) moreover have "conf_f P h ([], LT\<^sub>0) ins' ?f'" proof - let ?LT = "OK (Class D'') # (map OK Ts') @ (replicate mxl' Err)" from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" by simp also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\\<^sub>\] map OK Ts'" by simp also have "P,h \ replicate mxl' undefined_value [:\\<^sub>\] replicate mxl' Err" by simp also from m_C' have "P \ C' \\<^sup>* D''" by (rule sees_method_decl_above) from obj heap_ok have "is_htype P U" by (rule typeof_addr_is_type) with C' have "P \ ty_of_htype U \ Class C'" by(cases U)(simp_all add: widen_array_object) with \P \ C' \\<^sup>* D''\ obj C' have "P,h \ Addr a :\ Class D''" by (auto simp add: conf_def intro: widen_trans) ultimately have "P,h \ ?loc' [:\\<^sub>\] ?LT" by simp also from start LT\<^sub>0 have "P \ \ [\\<^sub>\] LT\<^sub>0" by (simp add: wt_start_def) finally have "P,h \ ?loc' [:\\<^sub>\] LT\<^sub>0" . thus ?thesis using ins' by simp qed ultimately have ?thesis using s' \_pc approx meth_C m_D T' ins D tconf C' mD'' by (fastforce dest: sees_method_fun [of _ C]) } moreover { assume [simp]: "meth' = Native" with wfprog m_C' have "D''\M'(Ts') :: T'" by(simp add: sees_wf_native) with C' m_C' have nec: "is_native P U M'" by(auto intro: is_native.intros) from ins n Addr obj exec m_C' C' obtain va h' tas' where va: "(tas', va, h') \ red_external_aggr P t a M' (rev (take n stk)) h" and \: "\ = extRet2JVM n h' stk loc C M pc frs va" by(auto) from va nec obj have hext: "h \ h'" by(auto intro: red_external_aggr_hext) with frames have frames': "conf_fs P h' \ M (length Ts) T frs" by(rule conf_fs_hext) from preh hext have preh': "preallocated h'" by(rule preallocated_hext) from va nec obj tconf have tconf': "P,h' \ t \t" by(auto dest: red_external_aggr_preserves_tconf) from hext obj have obj': "typeof_addr h' a = \U\" by(rule typeof_addr_hext_mono) from stk have "P,h \ take n stk [:\] take n ST" by(rule list_all2_takeI) then obtain Us where "map typeof\<^bsub>h\<^esub> (take n stk) = map Some Us" "P \ Us [\] take n ST" by(auto simp add: confs_conv_map) hence Us: "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some (rev Us)" "P \ rev Us [\] rev (take n ST)" by- (simp only: rev_map[symmetric], simp) from \P \ rev Us [\] rev (take n ST)\ Ts Ts' have "P \ rev Us [\] Ts'" by(blast intro: widens_trans) with obj \map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some (rev Us)\ C' m_C' have wtext': "P,h \ a\M'(rev (take n stk)) : T'" by(simp add: external_WT'.intros) from va have va': "P,t \ \a\M'(rev (take n stk)),h\ -tas'\ext \va,h'\" by(unfold WT_red_external_list_conv[OF wfprog wtext' tconf]) with heap_ok wtext' tconf wfprog have heap_ok': "hconf h'" by(auto dest: external_call_hconf) have ?thesis proof(cases va) case (RetExc a') from frame hext have "conf_f P h' (ST, LT) ins (stk, loc, C, M, pc)" by(rule conf_f_hext) with \ tconf' heap_ok' meth_C \_pc frames' RetExc red_external_conf_extRet[OF wfprog va' wtext' heap_ok preh tconf] ins preh' show ?thesis by(fastforce simp add: conf_def widen_Class) next case RetStaySame from frame hext have "conf_f P h' (ST, LT) ins (stk, loc, C, M, pc)" by(rule conf_f_hext) with \ heap_ok' meth_C \_pc RetStaySame frames' tconf' preh' show ?thesis by fastforce next case (RetVal v) with \ have \: "\ = (None, h', (v # drop (n+1) stk, loc, C, M, pc+1) # frs)" by simp from heap_ok wtext' va' RetVal preh tconf have "P,h' \ v :\ T'" by(auto dest: red_external_conf_extRet[OF wfprog]) from stk have "P,h \ drop (n + 1) stk [:\] drop (n+1) ST" by(rule list_all2_dropI) hence "P,h' \ drop (n + 1) stk [:\] drop (n+1) ST" using hext by(rule confs_hext) with \P,h' \ v :\ T'\ have "P,h' \ v # drop (n + 1) stk [:\] T' # drop (n+1) ST" by(auto simp add: conf_def intro: widen_trans) also with NT ins wti \_pc \' nec False D m_D T' have "P \ (T' # drop (n + 1) ST) [\] ST'" by(auto dest: sees_method_fun intro: widen_trans) also from loc hext have "P,h' \ loc [:\\<^sub>\] LT" by(rule confTs_hext) hence "P,h' \ loc [:\\<^sub>\] LT'" using LT' by(rule confTs_widen) ultimately show ?thesis using \hconf h'\ \ meth_C \' pc' frames' tconf' preh' by fastforce qed } ultimately show ?thesis by(cases meth') auto qed qed declare list_all2_Cons2 [iff] lemma Return_correct: assumes wt_prog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins ! pc = Return" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes correct: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes s': "(tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs)" shows "\ \ t:\'\" proof - from wt_prog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from meth ins s' correct have "frs = [] \ ?thesis" by (simp add: correct_state_def) moreover { fix f frs' assume frs': "frs = f#frs'" moreover obtain stk' loc' C' M' pc' where f: "f = (stk',loc',C',M',pc')" by (cases f) moreover note meth ins s' ultimately have \': "\' = (None,h,(hd stk#(drop (1+size Ts) stk'),loc',C',M',pc'+1)#frs')" (is "\' = (None,h,?f'#frs')") by simp from correct meth obtain ST LT where h_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h (ST, LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from \_pc ins wt obtain U ST\<^sub>0 where "ST = U # ST\<^sub>0" "P \ U \ T" by (simp add: wt_instr_def app_def) blast with wf frame have hd_stk: "P,h \ hd stk :\ T" by (auto simp add: conf_f_def) from f frs' frames obtain ST' LT' Ts'' T'' mxs' mxl\<^sub>0' ins' xt' Ts' T' where \': "\ C' M' ! pc' = Some (ST', LT')" and meth_C': "P \ C' sees M':Ts''\T''=\(mxs',mxl\<^sub>0',ins',xt')\ in C'" and ins': "ins' ! pc' = Invoke M (size Ts)" and D: "\D m D'. class_type_of' (ST' ! (size Ts)) = Some D \ P \ D sees M: Ts'\T' = m in D'" and T': "P \ T \ T'" and frame': "conf_f P h (ST',LT') ins' f" and conf_fs: "conf_fs P h \ M' (size Ts'') T'' frs'" by clarsimp blast from f frame' obtain stk': "P,h \ stk' [:\] ST'" and loc': "P,h \ loc' [:\\<^sub>\] LT'" and pc': "pc' < size ins'" by (simp add: conf_f_def) from wt_prog meth_C' pc' have wti: "P,T'',mxs',size ins',xt' \ ins'!pc',pc' :: \ C' M'" by (rule wt_jvm_prog_impl_wt_instr) obtain aTs ST'' LT'' where \_suc: "\ C' M' ! Suc pc' = Some (ST'', LT'')" and less: "P \ (T' # drop (size Ts+1) ST', LT') \\<^sub>i (ST'', LT'')" and suc_pc': "Suc pc' < size ins'" using ins' \' D T' wti by(fastforce simp add: sup_state_opt_any_Some split: if_split_asm) from hd_stk T' have hd_stk': "P,h \ hd stk :\ T'" .. have frame'': "conf_f P h (ST'',LT'') ins' ?f'" proof - from stk' have "P,h \ drop (1+size Ts) stk' [:\] drop (1+size Ts) ST'" .. moreover with hd_stk' less have "P,h \ hd stk # drop (1+size Ts) stk' [:\] ST''" by auto moreover from wf loc' less have "P,h \ loc' [:\\<^sub>\] LT''" by auto moreover note suc_pc' ultimately show ?thesis by (simp add: conf_f_def) qed with \' frs' f meth h_ok hd_stk \_suc frames meth_C' \' tconf preh have ?thesis by (fastforce dest: sees_method_fun [of _ C']) } ultimately show ?thesis by (cases frs) blast+ qed declare sup_state_opt_any_Some [iff] declare not_Err_eq [iff] lemma Load_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins!pc = Load idx; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\ ; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\' \" by (fastforce dest: sees_method_fun [of _ C] elim!: confTs_confT_sup) declare [[simproc del: list_to_set_comprehension]] lemma Store_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins!pc = Store idx; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply clarsimp apply (blast intro!: list_all2_update_cong) done lemma Push_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins!pc = Push v; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply clarsimp apply (blast dest: typeof_lit_conf) done declare [[simproc add: list_to_set_comprehension]] lemma Checkcast_correct: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins!pc = Checkcast D; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs \ \ \ \ t:\ \" using wf_preallocatedD[of "\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)" P h ClassCast] apply (clarsimp simp add: wf_jvm_prog_phi_def split: if_split_asm) apply(drule (1) sees_method_fun) apply(fastforce simp add: conf_def intro: widen_trans) apply (drule (1) sees_method_fun) apply(fastforce simp add: conf_def intro: widen_trans) done lemma Instanceof_correct: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins!pc = Instanceof Ty; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs \ \ \ \ t:\ \" apply (clarsimp simp add: wf_jvm_prog_phi_def split: if_split_asm) apply (drule (1) sees_method_fun) apply fastforce done declare split_paired_All [simp del] end lemma widens_Cons [iff]: "P \ (T # Ts) [\] Us = (\z zs. Us = z # zs \ P \ T \ z \ P \ Ts [\] zs)" by(rule list_all2_Cons1) context heap_conf_base begin end context JVM_conf_read begin lemma Getfield_correct: assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes i: "ins!pc = Getfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes cf: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "(tas, \') \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\'\" proof - from mC cf obtain ST LT where "h\": "hconf h" and tconf: "P,h \ t \t" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (fastforce dest: sees_method_fun) from i \ wt obtain oT ST'' vT ST' LT' vT' fm where oT: "P \ oT \ Class D" and ST: "ST = oT # ST''" and F: "P \ D sees F:vT (fm) in D" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (vT'#ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and vT': "P \ vT \ vT'" by fastforce from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto show ?thesis proof(cases "ref = Null") case True with tconf "h\" i xc stk' mC fs \ ST'' ref ST loc pc' wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False from ref oT have "P,h \ ref :\ Class D" .. with False obtain a U' D' where a: "ref = Addr a" and h: "typeof_addr h a = Some U'" and U': "D' = class_type_of U'" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD2) { fix v assume read: "heap_read h a (CField D F) v" from D' F have has_field: "P \ D' has F:vT (fm) in D" by (blast intro: has_field_mono has_visible_field) with h have "P,h \ a@CField D F : vT" unfolding U' .. with read have v: "P,h \ v :\ vT" using "h\" by(rule heap_read_conf) from ST'' ST' have "P,h \ stk' [:\] ST'" .. moreover from v vT' have "P,h \ v :\ vT'" by blast moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover note "h\" mC \' pc' v fs tconf preh ultimately have "\ \ t:(None, h, (v#stk',loc,C,M,pc+1)#frs) \" by fastforce } with a h i mC stk' xc show ?thesis by auto qed qed lemma Putfield_correct: assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes i: "ins!pc = Putfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes cf: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "(tas, \') \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\' \" proof - from mC cf obtain ST LT where "h\": "hconf h" and tconf: "P,h \ t \t" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (fastforce dest: sees_method_fun) from i \ wt obtain vT vT' oT ST'' ST' LT' fm where ST: "ST = vT # oT # ST''" and field: "P \ D sees F:vT' (fm) in D" and oT: "P \ oT \ Class D" and vT: "P \ vT \ vT'" and pc': "pc+1 < size ins" and \': "\ C M!(pc+1) = Some (ST',LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" by clarsimp from stk ST obtain v ref stk' where stk': "stk = v#ref#stk'" and v: "P,h \ v :\ vT" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto show ?thesis proof(cases "ref = Null") case True with tconf "h\" i xc stk' mC fs \ ST'' ref ST loc pc' v wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False from ref oT have "P,h \ ref :\ Class D" .. with False obtain a U' D' where a: "ref = Addr a" and h: "typeof_addr h a = Some U'" and U': "D' = class_type_of U'" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD2) from v vT have vT': "P,h \ v :\ vT'" .. from field D' have has_field: "P \ D' has F:vT' (fm) in D" by (blast intro: has_field_mono has_visible_field) with h have al: "P,h \ a@CField D F : vT'" unfolding U' .. let ?f' = "(stk',loc,C,M,pc+1)" { fix h' assume "write": "heap_write h a (CField D F) v h'" hence hext: "h \ h'" by(rule hext_heap_write) with preh have "preallocated h'" by(rule preallocated_hext) moreover from "write" "h\" al vT' have "hconf h'" by(rule hconf_heap_write_mono) moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. from this hext have "P,h' \ stk' [:\] ST'" by (rule confs_hext) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. from this hext have "P,h' \ loc [:\\<^sub>\] LT'" by (rule confTs_hext) moreover from fs hext have "conf_fs P h' \ M (size Ts) T frs" by (rule conf_fs_hext) moreover note mC \' pc' moreover from tconf hext have "P,h' \ t \t" by(rule tconf_hext_mono) ultimately have "\ \ t:(None, h', ?f'#frs) \" by fastforce } with a h i mC stk' xc show ?thesis by(auto simp del: correct_state_def) qed qed lemma CAS_correct: assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes i: "ins!pc = CAS F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes cf: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "(tas, \') \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\' \" proof - from mC cf obtain ST LT where "h\": "hconf h" and tconf: "P,h \ t \t" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (fastforce dest: sees_method_fun) from i \ wt obtain T1 T2 T3 T' ST'' ST' LT' fm where ST: "ST = T3 # T2 # T1 # ST''" and field: "P \ D sees F:T' (fm) in D" and oT: "P \ T1 \ Class D" and T2: "P \ T2 \ T'" and T3: "P \ T3 \ T'" and pc': "pc+1 < size ins" and \': "\ C M!(pc+1) = Some (Boolean # ST',LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" by clarsimp from stk ST obtain v'' v' v stk' where stk': "stk = v''#v'#v#stk'" and v: "P,h \ v :\ T1" and v': "P,h \ v' :\ T2" and v'': "P,h \ v'' :\ T3" and ST'': "P,h \ stk' [:\] ST''" by auto show ?thesis proof(cases "v = Null") case True with tconf "h\" i xc stk' mC fs \ ST'' v ST loc pc' v' v'' wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False from v oT have "P,h \ v :\ Class D" .. with False obtain a U' D' where a: "v = Addr a" and h: "typeof_addr h a = Some U'" and U': "D' = class_type_of U'" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD2) from v' T2 have vT': "P,h \ v' :\ T'" .. from v'' T3 have vT'': "P,h \ v'' :\ T'" .. from field D' have has_field: "P \ D' has F:T' (fm) in D" by (blast intro: has_field_mono has_visible_field) with h have al: "P,h \ a@CField D F : T'" unfolding U' .. from ST'' ST' have stk'': "P,h \ stk' [:\] ST'" .. from loc LT' have loc': "P,h \ loc [:\\<^sub>\] LT'" .. { fix h' assume "write": "heap_write h a (CField D F) v'' h'" hence hext: "h \ h'" by(rule hext_heap_write) with preh have "preallocated h'" by(rule preallocated_hext) moreover from "write" "h\" al vT'' have "hconf h'" by(rule hconf_heap_write_mono) moreover from stk'' hext have "P,h' \ stk' [:\] ST'" by (rule confs_hext) moreover from loc' hext have "P,h' \ loc [:\\<^sub>\] LT'" by (rule confTs_hext) moreover from fs hext have "conf_fs P h' \ M (size Ts) T frs" by (rule conf_fs_hext) moreover note mC \' pc' moreover let ?f' = "(Bool True # stk',loc,C,M,pc+1)" from tconf hext have "P,h' \ t \t" by(rule tconf_hext_mono) ultimately have "\ \ t:(None, h', ?f'#frs) \" by fastforce } moreover { let ?f' = "(Bool False # stk',loc,C,M,pc+1)" have "\ \ t:(None, h, ?f'#frs) \" using tconf "h\" preh mC \' stk'' loc' pc' fs by fastforce } ultimately show ?thesis using a h i mC stk' xc by(auto simp del: correct_state_def) qed qed lemma New_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = New X" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from \_pc ins wt obtain ST' LT' where is_class_X: "is_class P X" and mxs: "size ST < mxs" and suc_pc: "pc+1 < size ins" and \_suc: "\ C M!(pc+1) = Some (ST', LT')" and less: "P \ (Class X # ST, LT) \\<^sub>i (ST', LT')" by auto show ?thesis proof(cases "allocate h (Class_type X) = {}") case True with frame frames tconf suc_pc no_x ins meth \_pc wf_preallocatedD[OF wf, of h OutOfMemory] preh is_class_X heap_ok show ?thesis by(fastforce intro: tconf_hext_mono confs_hext confTs_hext conf_fs_hext) next case False with ins meth no_x obtain h' oref where new: "(h', oref) \ allocate h (Class_type X)" and \': "\ = (None, h', (Addr oref#stk,loc,C,M,pc+1)#frs)" (is "\ = (None, h', ?f # frs)") by auto from new have hext: "h \ h'" by(rule hext_allocate) with preh have preh': "preallocated h'" by(rule preallocated_hext) from new heap_ok is_class_X have heap_ok': "hconf h'" by(auto intro: hconf_allocate_mono) with new is_class_X have h': "typeof_addr h' oref = \Class_type X\" by(auto dest: allocate_SomeD) note heap_ok' \' moreover from frame less suc_pc wf h' hext have "conf_f P h' (ST', LT') ins ?f" apply (clarsimp simp add: fun_upd_apply conf_def split_beta) apply (auto intro: confs_hext confTs_hext) done moreover from frames hext have "conf_fs P h' \ M (size Ts) T frs" by (rule conf_fs_hext) moreover from tconf hext have "P,h' \ t \t" by(rule tconf_hext_mono) ultimately show ?thesis using meth \_suc preh' by fastforce qed qed lemma Goto_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = Goto branch; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\' \" apply clarsimp apply (drule (1) sees_method_fun) apply fastforce done declare [[simproc del: list_to_set_comprehension]] lemma IfFalse_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = IfFalse branch; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply fastforce done declare [[simproc add: list_to_set_comprehension]] lemma BinOp_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = BinOpInstr bop; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply(clarsimp simp add: conf_def) apply(drule (2) WTrt_binop_widen_mono) apply clarsimp apply(frule (2) binop_progress) apply(clarsimp split: sum.split_asm) apply(frule (5) binop_type) apply(fastforce intro: widen_trans simp add: conf_def) apply(frule (5) binop_type) apply(clarsimp simp add: conf_def) apply(clarsimp simp add: widen_Class) apply(fastforce intro: widen_trans dest: binop_relevant_class simp add: cname_of_def conf_def) done lemma Pop_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = Pop; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply fastforce done lemma Dup_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = Dup; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply fastforce done lemma Swap_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = Swap; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs) \ \ \ \ t:\'\" apply clarsimp apply (drule (1) sees_method_fun) apply fastforce done declare [[simproc del: list_to_set_comprehension]] lemma Throw_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; ins ! pc = ThrowExc; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; \ \ t:(None, h, (stk,loc,C,M,pc)#frs)\; (tas, \') \ exec_instr (ins!pc) P t h stk loc C M pc frs \ \ \ \ t:\'\" using wf_preallocatedD[of wt P h NullPointer] apply(clarsimp) apply(drule (1) sees_method_fun) apply(auto) apply fastforce apply fastforce apply(drule (1) non_npD) apply fastforce+ done declare [[simproc add: list_to_set_comprehension]] lemma NewArray_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = NewArray X" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t:(None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from ins \_pc wt obtain ST'' X' ST' LT' where ST: "ST = Integer # ST''" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (X'#ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and XX': "P \ X\\ \ X'" and suc_pc: "pc+1 < size ins" and is_type_X: "is_type P (X\\)" by(fastforce dest: Array_widen) from stk ST obtain si stk' where si: "stk = Intg si # stk'" by(auto simp add: conf_def) show ?thesis proof(cases "si allocate h (Array_type X (nat (sint si))) = {}") case True with frame frames tconf heap_ok suc_pc no_x ins meth \_pc si preh wf_preallocatedD[OF wf, of h OutOfMemory] wf_preallocatedD[OF wf, of h NegativeArraySize] show ?thesis by(fastforce intro: tconf_hext_mono confs_hext confTs_hext conf_fs_hext split: if_split_asm)+ next case False with ins meth si no_x obtain h' oref where new: "(h', oref) \ allocate h (Array_type X (nat (sint si)))" and \': "\ = (None, h', (Addr oref#tl stk,loc,C,M,pc+1)#frs)" (is "\ = (None, h', ?f # frs)") by(auto split: if_split_asm) from new have hext: "h \ h'" by(rule hext_allocate) with preh have preh': "preallocated h'" by(rule preallocated_hext) from new heap_ok is_type_X have heap_ok': "hconf h'" by(auto intro: hconf_allocate_mono) from False have si': "0 <=s si" by auto with new is_type_X have h': "typeof_addr h' oref = \Array_type X (nat (sint si))\" by(auto dest: allocate_SomeD) note \' heap_ok' moreover from frame ST' ST LT' suc_pc wf XX' h' hext have "conf_f P h' (X' # ST', LT') ins ?f" by(clarsimp simp add: fun_upd_apply conf_def split_beta)(auto intro: confs_hext confTs_hext) moreover from frames hext have "conf_fs P h' \ M (size Ts) T frs" by (rule conf_fs_hext) moreover from tconf hext have "P,h' \ t \t" by(rule tconf_hext_mono) ultimately show ?thesis using meth \' preh' by fastforce qed qed lemma ALoad_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = ALoad" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t: (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t:\ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from ins wt \_pc have lST: "length ST > 1" by(auto) show ?thesis proof(cases "hd (tl stk) = Null") case True with ins no_x heap_ok tconf \_pc stk loc frame frames meth wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False note stkNN = this have STNN: "hd (tl ST) \ NT" proof assume "hd (tl ST) = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with lST have "P,h \ hd (tl stk) :\ hd (tl ST)" by (cases ST, auto, case_tac list, auto) ultimately have "hd (tl stk) = Null" by simp with stkNN show False by contradiction qed with stkNN ins \_pc wt obtain ST'' X X' ST' LT' where ST: "ST = Integer # X\\ # ST''" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (X'#ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and XX': "P \ X \ X'" and suc_pc: "pc+1 < size ins" by(fastforce) from stk ST obtain ref idx stk' where stk': "stk = idx#ref#stk'" and idx: "P,h \ idx :\ Integer" and ref: "P,h \ ref :\ X\\" and ST'': "P,h \ stk' [:\] ST''" by auto from stkNN stk' have "ref \ Null" by(simp) with ref obtain a Xel n where a: "ref = Addr a" and ha: "typeof_addr h a = \Array_type Xel n\" and Xel: "P \ Xel \ X" by(cases ref)(fastforce simp add: conf_def widen_Array)+ from idx obtain idxI where idxI: "idx = Intg idxI" by(auto simp add: conf_def) show ?thesis proof(cases "0 <=s idxI \ sint idxI < int n") case True hence si': "0 <=s idxI" "sint idxI < int n" by auto hence "nat (sint idxI) < n" by (simp add: word_sle_eq nat_less_iff) with ha have al: "P,h \ a@ACell (nat (sint idxI)) : Xel" .. { fix v assume read: "heap_read h a (ACell (nat (sint idxI))) v" hence v: "P,h \ v :\ Xel" using al heap_ok by(rule heap_read_conf) let ?f = "(v # stk', loc, C, M, pc + 1)" from frame ST' ST LT' suc_pc wf XX' Xel idxI si' v ST'' have "conf_f P h (X' # ST', LT') ins ?f" by(auto intro: widen_trans simp add: conf_def) hence "\ \ t:(None, h, ?f # frs) \" using meth \' heap_ok \_pc frames tconf preh by fastforce } with ins meth si' stk' a ha no_x idxI idx show ?thesis by(auto simp del: correct_state_def split: if_split_asm) next case False with stk' idxI ins no_x heap_ok tconf meth a ha Xel \_pc frame frames - wf_preallocatedD[OF wf, of h ArrayIndexOutOfBounds] preh - show ?thesis by(fastforce split: if_split_asm) + wf_preallocatedD[OF wf, of h ArrayIndexOutOfBounds] + show ?thesis + by (fastforce simp: preh split: if_split_asm simp del: Listn.lesub_list_impl_same_size) qed qed qed lemma AStore_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = AStore" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t: (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t: \ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from ins wt \_pc have lST: "length ST > 2" by(auto) show ?thesis proof(cases "hd (tl (tl stk)) = Null") case True with ins no_x heap_ok tconf \_pc stk loc frame frames meth wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False note stkNN = this have STNN: "hd (tl (tl ST)) \ NT" proof assume "hd (tl (tl ST)) = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with lST have "P,h \ hd (tl (tl stk)) :\ hd (tl (tl ST))" by (cases ST, auto, case_tac list, auto, case_tac lista, auto) ultimately have "hd (tl (tl stk)) = Null" by simp with stkNN show False by contradiction qed with ins stkNN \_pc wt obtain ST'' Y X ST' LT' where ST: "ST = Y # Integer # X\\ # ST''" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and suc_pc: "pc+1 < size ins" by(fastforce) from stk ST obtain ref e idx stk' where stk': "stk = e#idx#ref#stk'" and idx: "P,h \ idx :\ Integer" and ref: "P,h \ ref :\ X\\" and e: "P,h \ e :\ Y" and ST'': "P,h \ stk' [:\] ST''" by auto from stkNN stk' have "ref \ Null" by(simp) with ref obtain a Xel n where a: "ref = Addr a" and ha: "typeof_addr h a = \Array_type Xel n\" and Xel: "P \ Xel \ X" by(cases ref)(fastforce simp add: conf_def widen_Array)+ from idx obtain idxI where idxI: "idx = Intg idxI" by(auto simp add: conf_def) show ?thesis proof(cases "0 <=s idxI \ sint idxI < int n") case True hence si': "0 <=s idxI" "sint idxI < int n" by simp_all from e obtain Te where Te: "typeof\<^bsub>h\<^esub> e = \Te\" "P \ Te \ Y" by(auto simp add: conf_def) show ?thesis proof(cases "P \ Te \ Xel") case True with Te have eXel: "P,h \ e :\ Xel" by(auto simp add: conf_def intro: widen_trans) { fix h' assume "write": "heap_write h a (ACell (nat (sint idxI))) e h'" hence hext: "h \ h'" by(rule hext_heap_write) with preh have preh': "preallocated h'" by(rule preallocated_hext) let ?f = "(stk', loc, C, M, pc + 1)" from si' have "nat (sint idxI) < n" by (simp add: word_sle_eq nat_less_iff) with ha have "P,h \ a@ACell (nat (sint idxI)) : Xel" .. with "write" heap_ok have heap_ok': "hconf h'" using eXel by(rule hconf_heap_write_mono) moreover from ST stk stk' ST' have "P,h \ stk' [:\] ST'" by auto with hext have stk'': "P,h' \ stk' [:\] ST'" by- (rule confs_hext) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. with hext have "P,h' \ loc [:\\<^sub>\] LT'" by - (rule confTs_hext) moreover with frame ST' ST LT' suc_pc wf Xel idxI si' stk'' have "conf_f P h' (ST', LT') ins ?f" by(clarsimp) with frames hext have "conf_fs P h' \ M (size Ts) T frs" by- (rule conf_fs_hext) moreover from tconf hext have "P,h' \ t \t" by(rule tconf_hext_mono) ultimately have "\ \ t:(None, h', ?f # frs) \" using meth \' \_pc suc_pc preh' by(fastforce) } with True si' ins meth stk' a ha no_x idxI idx Te show ?thesis by(auto split: if_split_asm simp del: correct_state_def intro: widen_trans) next case False - with stk' idxI ins no_x heap_ok tconf meth a ha Xel Te \_pc frame frames si' preh + with stk' idxI ins no_x heap_ok tconf meth a ha Xel Te \_pc frame frames si' wf_preallocatedD[OF wf, of h ArrayStore] - show ?thesis by(fastforce split: if_split_asm) + show ?thesis + by (fastforce split: if_splits list.splits simp: preh + simp del: Listn.lesub_list_impl_same_size) qed next case False with stk' idxI ins no_x heap_ok tconf meth a ha Xel \_pc frame frames preh wf_preallocatedD[OF wf, of h ArrayIndexOutOfBounds] show ?thesis by(fastforce split: if_split_asm) qed qed qed lemma ALength_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = ALength" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t: (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t: \ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from ins wt \_pc have lST: "length ST > 0" by(auto) show ?thesis proof(cases "hd stk = Null") case True with ins no_x heap_ok tconf \_pc stk loc frame frames meth wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False note stkNN = this have STNN: "hd ST \ NT" proof assume "hd ST = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with lST have "P,h \ hd stk :\ hd ST" by (cases ST, auto) ultimately have "hd stk = Null" by simp with stkNN show False by contradiction qed with stkNN ins \_pc wt obtain ST'' X ST' LT' where ST: "ST = (X\\) # ST''" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (ST', LT')" and ST': "P \ (Integer # ST'') [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and suc_pc: "pc+1 < size ins" by(fastforce) from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ X\\" and ST'': "P,h \ stk' [:\] ST''" by auto from stkNN stk' have "ref \ Null" by(simp) with ref obtain a Xel n where a: "ref = Addr a" and ha: "typeof_addr h a = \Array_type Xel n\" and Xel: "P \ Xel \ X" by(cases ref)(fastforce simp add: conf_def widen_Array)+ from ins meth stk' a ha no_x have \': "\ = (None, h, (Intg (word_of_int (int n)) # stk', loc, C, M, pc + 1) # frs)" (is "\ = (None, h, ?f # frs)") by(auto) moreover from ST stk stk' ST' have "P,h \ Intg si # stk' [:\] ST'" by(auto) with frame ST' ST LT' suc_pc wf have "conf_f P h (ST', LT') ins ?f" by(fastforce intro: widen_trans) ultimately show ?thesis using meth \' heap_ok \_pc frames tconf preh by fastforce qed qed lemma MEnter_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = MEnter" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t: (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t: \ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from ins wt \_pc have lST: "length ST > 0" by(auto) show ?thesis proof(cases "hd stk = Null") case True with ins no_x heap_ok tconf \_pc stk loc frame frames meth wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False note stkNN = this have STNN: "hd ST \ NT" proof assume "hd ST = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with lST have "P,h \ hd stk :\ hd ST" by (cases ST, auto) ultimately have "hd stk = Null" by simp with stkNN show False by contradiction qed with stkNN ins \_pc wt obtain ST'' X ST' LT' where ST: "ST = X # ST''" and refT: "is_refT X" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and suc_pc: "pc+1 < size ins" by(fastforce) from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ X" by auto from stkNN stk' have "ref \ Null" by(simp) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover from ST stk stk' ST' have "P,h \ stk' [:\] ST'" by(auto) ultimately show ?thesis using meth \' heap_ok \_pc suc_pc frames loc LT' no_x ins stk' ST' tconf preh by(fastforce) qed qed lemma MExit_correct: assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C" assumes ins: "ins!pc = MExit" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes conf: "\ \ t: (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "(tas, \) \ exec_instr (ins!pc) P t h stk loc C M pc frs" shows "\ \ t: \ \" proof - from ins conf meth obtain ST LT where heap_ok: "hconf h" and tconf: "P,h \ t \t" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" and preh: "preallocated h" by (auto dest: sees_method_fun) from ins wt \_pc have lST: "length ST > 0" by(auto) show ?thesis proof(cases "hd stk = Null") case True with ins no_x heap_ok tconf \_pc stk loc frame frames meth wf_preallocatedD[OF wf, of h NullPointer] preh show ?thesis by(fastforce) next case False note stkNN = this have STNN: "hd ST \ NT" proof assume "hd ST = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with lST have "P,h \ hd stk :\ hd ST" by (cases ST, auto) ultimately have "hd stk = Null" by simp with stkNN show False by contradiction qed with stkNN ins \_pc wt obtain ST'' X ST' LT' where ST: "ST = X # ST''" and refT: "is_refT X" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and suc_pc: "pc+1 < size ins" by(fastforce) from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ X" by auto from stkNN stk' have "ref \ Null" by(simp) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover from ST stk stk' ST' have "P,h \ stk' [:\] ST'" by(auto) ultimately show ?thesis using meth \' heap_ok \_pc suc_pc frames loc LT' no_x ins stk' ST' tconf frame preh wf_preallocatedD[OF wf, of h IllegalMonitorState] by(fastforce) qed qed text \ The next theorem collects the results of the sections above, i.e.~exception handling and the execution step for each instruction. It states type safety for single step execution: in welltyped programs, a conforming state is transformed into another conforming state when one instruction is executed. \ theorem instr_correct: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:Ts\T=\(mxs,mxl\<^sub>0,ins,xt)\ in C; (tas, \') \ exec P t (None, h, (stk,loc,C,M,pc)#frs); \ \ t: (None, h, (stk,loc,C,M,pc)#frs)\ \ \ \ \ t: \'\" apply (subgoal_tac "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M") prefer 2 apply (erule wt_jvm_prog_impl_wt_instr, assumption) apply clarsimp apply (drule (1) sees_method_fun) apply simp apply(unfold exec.simps Let_def set_map) apply (frule wt_jvm_progD, erule exE) apply (cases "ins ! pc") apply (rule Load_correct, assumption+, fastforce) apply (rule Store_correct, assumption+, fastforce) apply (rule Push_correct, assumption+, fastforce) apply (rule New_correct, assumption+, fastforce) apply (rule NewArray_correct, assumption+, fastforce) apply (rule ALoad_correct, assumption+, fastforce) apply (rule AStore_correct, assumption+, fastforce) apply (rule ALength_correct, assumption+, fastforce) apply (rule Getfield_correct, assumption+, fastforce) apply (rule Putfield_correct, assumption+, fastforce) apply (rule CAS_correct, assumption+, fastforce) apply (rule Checkcast_correct, assumption+, fastforce) apply (rule Instanceof_correct, assumption+, fastforce) apply (rule Invoke_correct, assumption+, fastforce) apply (rule Return_correct, assumption+, fastforce simp add: split_beta) apply (rule Pop_correct, assumption+, fastforce) apply (rule Dup_correct, assumption+, fastforce) apply (rule Swap_correct, assumption+, fastforce) apply (rule BinOp_correct, assumption+, fastforce) apply (rule Goto_correct, assumption+, fastforce) apply (rule IfFalse_correct, assumption+, fastforce) apply (rule Throw_correct, assumption+, fastforce) apply (rule MEnter_correct, assumption+, fastforce) apply (rule MExit_correct, assumption+, fastforce) done declare defs1 [simp del] end subsection \Main\ lemma (in JVM_conf_read) BV_correct_1 [rule_format]: "\\. \ wf_jvm_prog\<^bsub>\\<^esub> P; \ \ t: \\\ \ P,t \ \ -tas-jvm\ \' \ \ \ t: \'\" apply (simp only: split_tupled_all exec_1_iff) apply (rename_tac xp h frs) apply (case_tac xp) apply (case_tac frs) apply simp apply (simp only: split_tupled_all) apply hypsubst apply (frule correct_state_impl_Some_method) apply clarify apply (rule instr_correct) apply assumption+ apply clarify apply(case_tac frs) apply simp apply(clarsimp simp only: exec.simps set_simps) apply(erule (1) exception_step_conform) done theorem (in JVM_progress) progress: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" and cs: "\ \ t: (xcp, h, f # frs)\" shows "\ta \'. P,t \ (xcp, h, f # frs) -ta-jvm\ \'" proof - obtain stk loc C M pc where f: "f = (stk, loc, C, M, pc)" by(cases f) with cs obtain Ts T mxs mxl\<^sub>0 "is" xt ST LT where hconf: "hconf h" and sees: "P \ C sees M: Ts\T = \(mxs, mxl\<^sub>0, is, xt)\ in C" and \_pc: "\ C M ! pc = \(ST, LT)\" and ST: "P,h \ stk [:\] ST" and LT: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < length is" by(auto simp add: defs1) show ?thesis proof(cases xcp) case Some thus ?thesis unfolding f exec_1_iff by auto next case [simp]: None note [simp del] = split_paired_Ex note [simp] = defs1 list_all2_Cons2 from wt obtain wf_md where wf: "wf_prog wf_md P" by(auto dest: wt_jvm_progD) from wt sees pc have wt: "P,T,mxs,size is,xt \ is!pc,pc :: \ C M" by(rule wt_jvm_prog_impl_wt_instr) have "\ta \'. (ta, \') \ exec_instr (is ! pc) P t h stk loc C M pc frs" proof(cases "is ! pc") case [simp]: ALoad with wt \_pc have lST: "length ST > 1" by(auto) show ?thesis proof(cases "hd (tl stk) = Null") case True thus ?thesis by simp next case False have STNN: "hd (tl ST) \ NT" proof assume "hd (tl ST) = NT" moreover from ST lST have "P,h \ hd (tl stk) :\ hd (tl ST)" by (cases ST)(auto, case_tac list, auto) ultimately have "hd (tl stk) = Null" by simp with False show False by contradiction qed with False \_pc wt obtain ST'' X where "ST = Integer # X\\ # ST''" by auto with ST obtain ref idx stk' where stk': "stk = idx#ref#stk'" and idx: "P,h \ idx :\ Integer" and ref: "P,h \ ref :\ X\\" by(auto) from False stk' have "ref \ Null" by(simp) with ref obtain a Xel n where a: "ref = Addr a" and ha: "typeof_addr h a = \Array_type Xel n\" and Xel: "P \ Xel \ X" by(cases ref)(fastforce simp add: conf_def widen_Array)+ from idx obtain idxI where idxI: "idx = Intg idxI" by(auto simp add: conf_def) show ?thesis proof(cases "0 <=s idxI \ sint idxI < int n") case True hence si': "0 <=s idxI" "sint idxI < int n" by auto hence "nat (sint idxI) < n" by (simp add: word_sle_eq nat_less_iff) with ha have al: "P,h \ a@ACell (nat (sint idxI)) : Xel" .. from heap_read_total[OF hconf this] True False ha stk' idxI a show ?thesis by auto next case False with ha stk' idxI a show ?thesis by auto qed qed next case [simp]: AStore from wt \_pc have lST: "length ST > 2" by(auto) show ?thesis proof(cases "hd (tl (tl stk)) = Null") case True thus ?thesis by(fastforce) next case False note stkNN = this have STNN: "hd (tl (tl ST)) \ NT" proof assume "hd (tl (tl ST)) = NT" moreover from ST lST have "P,h \ hd (tl (tl stk)) :\ hd (tl (tl ST))" by (cases ST, auto, case_tac list, auto, case_tac lista, auto) ultimately have "hd (tl (tl stk)) = Null" by simp with stkNN show False by contradiction qed with stkNN \_pc wt obtain ST'' Y X where "ST = Y # Integer # X\\ # ST''" by(fastforce) with ST obtain ref e idx stk' where stk': "stk = e#idx#ref#stk'" and idx: "P,h \ idx :\ Integer" and ref: "P,h \ ref :\ X\\" and e: "P,h \ e :\ Y" by auto from stkNN stk' have "ref \ Null" by(simp) with ref obtain a Xel n where a: "ref = Addr a" and ha: "typeof_addr h a = \Array_type Xel n\" and Xel: "P \ Xel \ X" by(cases ref)(fastforce simp add: conf_def widen_Array)+ from idx obtain idxI where idxI: "idx = Intg idxI" by(auto simp add: conf_def) show ?thesis proof(cases "0 <=s idxI \ sint idxI < int n") case True hence si': "0 <=s idxI" "sint idxI < int n" by simp_all hence "nat (sint idxI) < n" by (simp add: word_sle_eq nat_less_iff) with ha have adal: "P,h \ a@ACell (nat (sint idxI)) : Xel" .. show ?thesis proof(cases "P \ the (typeof\<^bsub>h\<^esub> e) \ Xel") case False with ha stk' idxI a show ?thesis by auto next case True hence "P,h \ e :\ Xel" using e by(auto simp add: conf_def) from heap_write_total[OF hconf adal this] ha stk' idxI a show ?thesis by auto qed next case False with ha stk' idxI a show ?thesis by auto qed qed next case [simp]: (Getfield F D) from \_pc wt obtain oT ST'' vT fm where oT: "P \ oT \ Class D" and "ST = oT # ST''" and F: "P \ D sees F:vT (fm) in D" by fastforce with ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ oT" by auto show ?thesis proof(cases "ref = Null") case True thus ?thesis using stk' by auto next case False from ref oT have "P,h \ ref :\ Class D" .. with False obtain a U' D' where a: "ref = Addr a" and h: "typeof_addr h a = Some U'" and U': "D' = class_type_of U'" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD2) from D' F have has_field: "P \ D' has F:vT (fm) in D" by (blast intro: has_field_mono has_visible_field) with h have "P,h \ a@CField D F : vT" unfolding U' .. from heap_read_total[OF hconf this] show ?thesis using stk' a by auto qed next case [simp]: (Putfield F D) from \_pc wt obtain vT vT' oT ST'' fm where "ST = vT # oT # ST''" and field: "P \ D sees F:vT' (fm) in D" and oT: "P \ oT \ Class D" and vT': "P \ vT \ vT'" by fastforce with ST obtain v ref stk' where stk': "stk = v#ref#stk'" and ref: "P,h \ ref :\ oT" and v: "P,h \ v :\ vT" by auto show ?thesis proof(cases "ref = Null") case True with stk' show ?thesis by auto next case False from ref oT have "P,h \ ref :\ Class D" .. with False obtain a U' D' where a: "ref = Addr a" and h: "typeof_addr h a = Some U'" and U': "D' = class_type_of U'" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD2) from field D' have has_field: "P \ D' has F:vT' (fm) in D" by (blast intro: has_field_mono has_visible_field) with h have al: "P,h \ a@CField D F : vT'" unfolding U' .. from v vT' have "P,h \ v :\ vT'" by auto from heap_write_total[OF hconf al this] v a stk' h show ?thesis by auto qed next case [simp]: (CAS F D) from \_pc wt obtain T' T1 T2 T3 ST'' fm where "ST = T3 # T2 # T1 # ST''" and field: "P \ D sees F:T' (fm) in D" and oT: "P \ T1 \ Class D" and vT': "P \ T2 \ T'" "P \ T3 \ T'" by fastforce with ST obtain v v' v'' stk' where stk': "stk = v''#v'#v#stk'" and v: "P,h \ v :\ T1" and v': "P,h \ v' :\ T2" and v'': "P,h \ v'' :\ T3" by auto show ?thesis proof(cases "v= Null") case True with stk' show ?thesis by auto next case False from v oT have "P,h \ v :\ Class D" .. with False obtain a U' D' where a: "v = Addr a" and h: "typeof_addr h a = Some U'" and U': "D' = class_type_of U'" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD2) from field D' have has_field: "P \ D' has F:T' (fm) in D" by (blast intro: has_field_mono has_visible_field) with h have al: "P,h \ a@CField D F : T'" unfolding U' .. from v' vT' have "P,h \ v' :\ T'" by auto from heap_read_total[OF hconf al] obtain v''' where v''': "heap_read h a (CField D F) v'''" by blast show ?thesis proof(cases "v''' = v'") case True from v'' vT' have "P,h \ v'' :\ T'" by auto from heap_write_total[OF hconf al this] v a stk' h v''' True show ?thesis by auto next case False from v''' v a stk' h False show ?thesis by auto qed qed next case [simp]: (Invoke M' n) from wt \_pc have n: "n < size ST" by simp show ?thesis proof(cases "stk!n = Null") case True thus ?thesis by simp next case False note Null = this have NT: "ST!n \ NT" proof assume "ST!n = NT" moreover from ST n have "P,h \ stk!n :\ ST!n" by (simp add: list_all2_conv_all_nth) ultimately have "stk!n = Null" by simp with Null show False by contradiction qed from NT wt \_pc obtain D D' Ts T m where D: "class_type_of' (ST!n) = Some D" and m_D: "P \ D sees M': Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" by auto from n ST D have "P,h \ stk!n :\ ST!n" by (auto simp add: list_all2_conv_all_nth) from \P,h \ stk!n :\ ST!n\ Null D obtain a T' where Addr: "stk!n = Addr a" and obj: "typeof_addr h a = Some T'" and T'subSTn: "P \ ty_of_htype T' \ ST ! n" by(cases "stk ! n")(auto simp add: conf_def widen_Class) from D T'subSTn obtain C' where C': "class_type_of' (ty_of_htype T') = \C'\" and C'subD: "P \ C' \\<^sup>* D" by(rule widen_is_class_type_of) simp from Call_lemma[OF m_D C'subD wf] obtain D' Ts' T' m' where Call': "P \ C' sees M': Ts'\T' = m' in D'" "P \ Ts [\] Ts'" "P \ T' \ T" "P \ C' \\<^sup>* D'" "is_type P T'" "\T\set Ts'. is_type P T" by blast show ?thesis proof(cases m') case Some with Call' C' obj Addr C' C'subD show ?thesis by(auto) next case [simp]: None from ST have "P,h \ take n stk [:\] take n ST" by(rule list_all2_takeI) then obtain Us where "map typeof\<^bsub>h\<^esub> (take n stk) = map Some Us" "P \ Us [\] take n ST" by(auto simp add: confs_conv_map) hence Us: "map typeof\<^bsub>h\<^esub> (rev (take n stk)) = map Some (rev Us)" "P \ rev Us [\] rev (take n ST)" by- (simp only: rev_map[symmetric], simp) with Ts \P \ Ts [\] Ts'\ have "P \ rev Us [\] Ts'" by(blast intro: widens_trans) with obj Us Call' C' have "P,h \ a\M'(rev (take n stk)) : T'" by(auto intro!: external_WT'.intros) from external_call_progress[OF wf this hconf, of t] obj Addr Call' C' show ?thesis by(auto dest!: red_external_imp_red_external_aggr) qed qed qed(auto 4 4 simp add: split_beta split: if_split_asm) thus ?thesis using sees None unfolding f exec_1_iff by(simp del: split_paired_Ex) qed qed lemma (in JVM_heap_conf) BV_correct_initial: shows "\ wf_jvm_prog\<^bsub>\\<^esub> P; start_heap_ok; P \ C sees M:Ts\T = \m\ in D; P,start_heap \ vs [:\] Ts \ \ \ \ start_tid:JVM_start_state' P C M vs \" apply (cases m) apply (unfold JVM_start_state'_def) apply (unfold correct_state_def) apply (clarsimp) apply (frule wt_jvm_progD) apply (erule exE) apply (frule wf_prog_wf_syscls) apply (rule conjI) apply(erule (1) tconf_start_heap_start_tid) apply(rule conjI) apply (simp add: wf_jvm_prog_phi_def hconf_start_heap) apply(frule sees_method_idemp) apply (frule wt_jvm_prog_impl_wt_start, assumption+) apply (unfold conf_f_def wt_start_def) apply(auto simp add: sup_state_opt_any_Some) apply(erule preallocated_start_heap) apply(rule exI conjI|assumption)+ apply(auto simp add: list_all2_append1) apply(auto dest: list_all2_lengthD intro!: exI) done end diff --git a/thys/JinjaThreads/MM/DRF_JVM.thy b/thys/JinjaThreads/MM/DRF_JVM.thy --- a/thys/JinjaThreads/MM/DRF_JVM.thy +++ b/thys/JinjaThreads/MM/DRF_JVM.thy @@ -1,1167 +1,1167 @@ (* Title: JinjaThreads/MM/DRF_JVM.thy Author: Andreas Lochbihler *) section \JMM Instantiation for bytecode\ theory DRF_JVM imports JMM_Common JMM_JVM "../BV/BVProgressThreaded" SC_Legal begin subsection \DRF guarantee for the JVM\ abbreviation (input) ka_xcp :: "'addr option \ 'addr set" where "ka_xcp \ set_option" primrec jvm_ka :: "'addr jvm_thread_state \ 'addr set" where "jvm_ka (xcp, frs) = ka_xcp xcp \ (\(stk, loc, C, M, pc) \ set frs. (\v \ set stk. ka_Val v) \ (\v \ set loc. ka_Val v))" context heap begin lemma red_external_aggr_read_mem_typeable: "\ (ta, va, h') \ red_external_aggr P t a M vs h; ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ \T'. P,h \ ad@al : T'" by(auto simp add: red_external_aggr_def split_beta split: if_split_asm dest: heap_clone_read_typeable) end context JVM_heap_base begin definition jvm_known_addrs :: "'thread_id \ 'addr jvm_thread_state \ 'addr set" where "jvm_known_addrs t xcpfrs = {thread_id2addr t} \ jvm_ka xcpfrs \ set start_addrs" end context JVM_heap begin lemma exec_instr_known_addrs: assumes ok: "start_heap_ok" and exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" shows "jvm_known_addrs t (xcp', frs') \ jvm_known_addrs t (None, (stk, loc, C, M, pc) # frs) \ new_obs_addrs \ta\\<^bsub>o\<^esub>" proof - note [simp] = jvm_known_addrs_def new_obs_addrs_def addr_of_sys_xcpt_start_addr[OF ok] subset_Un1 subset_Un2 subset_insert ka_Val_subset_new_obs_Addr_ReadMem SUP_subset_mono split_beta neq_Nil_conv tl_conv_drop set_drop_subset is_Ref_def from exec check show ?thesis proof(cases "i") case Load with exec check show ?thesis by auto next case (Store V) with exec check show ?thesis using set_update_subset_insert[of loc V] by(clarsimp simp del: set_update_subsetI) blast next case (Push v) with check have "ka_Val v = {}" by(cases v) simp_all with Push exec check show ?thesis by(simp) next case (CAS F D) then show ?thesis using exec check by(clarsimp split: if_split_asm)(fastforce dest!: in_set_dropD)+ next case (Invoke M' n) show ?thesis proof(cases "stk ! n = Null") case True with exec check Invoke show ?thesis by(simp) next case [simp]: False with check Invoke obtain a where stkn: "stk ! n = Addr a" "n < length stk" by auto hence a: "a \ (\v \ set stk. ka_Val v)" by(fastforce dest: nth_mem) show ?thesis proof(cases "snd (snd (snd (method P (class_type_of (the (typeof_addr h (the_Addr (stk ! n))))) M'))) = Native") case True with exec check Invoke a stkn show ?thesis apply clarsimp apply(drule red_external_aggr_known_addrs_mono[OF ok], simp) apply(auto dest!: in_set_takeD dest: bspec subsetD split: extCallRet.split_asm simp add: has_method_def is_native.simps) done next case False with exec check Invoke a stkn show ?thesis by(auto simp add: set_replicate_conv_if dest!: in_set_takeD) qed qed next case Swap with exec check show ?thesis by(cases stk)(simp, case_tac list, auto) next case (BinOpInstr bop) with exec check show ?thesis using binop_known_addrs[OF ok, of bop "hd (drop (Suc 0) stk)" "hd stk"] apply(cases stk) apply(simp, case_tac list, simp) apply clarsimp apply(drule (2) binop_progress) apply(auto 6 2 split: sum.split_asm) done next case MExit with exec check show ?thesis by(auto split: if_split_asm) qed(clarsimp split: if_split_asm)+ qed lemma exec_d_known_addrs_mono: assumes ok: "start_heap_ok" and exec: "mexecd P t (xcpfrs, h) ta (xcpfrs', h')" shows "jvm_known_addrs t xcpfrs' \ jvm_known_addrs t xcpfrs \ new_obs_addrs \ta\\<^bsub>o\<^esub>" using exec apply(cases xcpfrs) apply(cases xcpfrs') apply(simp add: split_beta) apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(fastforce simp add: check_def split_beta del: subsetI dest!: exec_instr_known_addrs[OF ok]) apply(fastforce simp add: jvm_known_addrs_def split_beta dest!: in_set_dropD) done lemma exec_instr_known_addrs_ReadMem: assumes exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" and read: "ReadMem ad al v \ set \ta\\<^bsub>o\<^esub>" shows "ad \ jvm_known_addrs t (None, (stk, loc, C, M, pc) # frs)" using assms proof(cases i) case ALoad thus ?thesis using assms by(cases stk)(case_tac [2] list, auto simp add: split_beta is_Ref_def jvm_known_addrs_def split: if_split_asm) next case (Invoke M n) with check have "stk ! n \ Null \ the_Addr (stk ! n) \ ka_Val (stk ! n)" "stk ! n \ set stk" by(auto simp add: is_Ref_def) with assms Invoke show ?thesis by(auto simp add: split_beta is_Ref_def simp del: ka_Val.simps nth_mem split: if_split_asm dest!: red_external_aggr_known_addrs_ReadMem in_set_takeD del: is_AddrE)(auto simp add: jvm_known_addrs_def simp del: ka_Val.simps nth_mem del: is_AddrE) next case Getfield thus ?thesis using assms by(auto simp add: jvm_known_addrs_def neq_Nil_conv is_Ref_def split: if_split_asm) next case CAS thus ?thesis using assms apply(cases stk; simp) subgoal for v stk apply(cases stk; simp) subgoal for v stk by(cases stk)(auto split: if_split_asm simp add: jvm_known_addrs_def is_Ref_def) done done qed(auto simp add: split_beta is_Ref_def neq_Nil_conv split: if_split_asm) lemma mexecd_known_addrs_ReadMem: "\ mexecd P t (xcpfrs, h) ta (xcpfrs', h'); ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ ad \ jvm_known_addrs t xcpfrs" apply(cases xcpfrs) apply(cases xcpfrs') apply simp apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(auto simp add: check_def dest: exec_instr_known_addrs_ReadMem) done lemma exec_instr_known_addrs_WriteMem: assumes exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" and "write": "\ta\\<^bsub>o\<^esub> ! n = WriteMem ad al (Addr a)" "n < length \ta\\<^bsub>o\<^esub>" shows "a \ jvm_known_addrs t (None, (stk, loc, C, M, pc) # frs) \ a \ new_obs_addrs (take n \ta\\<^bsub>o\<^esub>)" using assms proof(cases i) case (Invoke M n) with check have "stk ! n \ Null \ the_Addr (stk ! n) \ ka_Val (stk ! n)" "stk ! n \ set stk" by(auto simp add: is_Ref_def) thus ?thesis using assms Invoke by(auto simp add: is_Ref_def split_beta split: if_split_asm simp del: ka_Val.simps nth_mem dest!: red_external_aggr_known_addrs_WriteMem in_set_takeD del: is_AddrE)(auto simp add: jvm_known_addrs_def del: is_AddrE) next case AStore with assms show ?thesis by(cases stk)(auto simp add: jvm_known_addrs_def split: if_split_asm) next case Putfield with assms show ?thesis by(cases stk)(auto simp add: jvm_known_addrs_def split: if_split_asm) next case CAS with assms show ?thesis apply(cases stk; simp) subgoal for v stk apply(cases stk; simp) subgoal for v stk by(cases stk)(auto split: if_split_asm simp add: take_Cons' jvm_known_addrs_def) done done qed(auto simp add: split_beta split: if_split_asm) lemma mexecd_known_addrs_WriteMem: "\ mexecd P t (xcpfrs, h) ta (xcpfrs', h'); \ta\\<^bsub>o\<^esub> ! n = WriteMem ad al (Addr a); n < length \ta\\<^bsub>o\<^esub> \ \ a \ jvm_known_addrs t xcpfrs \ a \ new_obs_addrs (take n \ta\\<^bsub>o\<^esub>)" apply(cases xcpfrs) apply(cases xcpfrs') apply simp apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(auto simp add: check_def dest: exec_instr_known_addrs_WriteMem) done lemma exec_instr_known_addrs_new_thread: assumes exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" and new: "NewThread t' x' h'' \ set \ta\\<^bsub>t\<^esub>" shows "jvm_known_addrs t' x' \ jvm_known_addrs t (None, (stk, loc, C, M, pc) # frs)" using assms proof(cases i) case (Invoke M n) with assms have "stk ! n \ Null \ the_Addr (stk ! n) \ ka_Val (stk ! n) \ thread_id2addr (addr2thread_id (the_Addr (stk ! n))) = the_Addr (stk ! n)" "stk ! n \ set stk" apply(auto simp add: is_Ref_def split: if_split_asm) apply(frule red_external_aggr_NewThread_idD, simp, simp) apply(drule red_external_aggr_new_thread_sub_thread) apply(auto intro: addr2thread_id_inverse) done with assms Invoke show ?thesis apply(auto simp add: is_Ref_def split_beta split: if_split_asm simp del: nth_mem del: is_AddrE) apply(drule red_external_aggr_NewThread_idD) apply(auto simp add: extNTA2JVM_def jvm_known_addrs_def split_beta simp del: nth_mem del: is_AddrE) done qed(auto simp add: split_beta split: if_split_asm) lemma mexecd_known_addrs_new_thread: "\ mexecd P t (xcpfrs, h) ta (xcpfrs', h'); NewThread t' x' h'' \ set \ta\\<^bsub>t\<^esub> \ \ jvm_known_addrs t' x' \ jvm_known_addrs t xcpfrs" apply(cases xcpfrs) apply(cases xcpfrs') apply simp apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(auto 4 3 simp add: check_def dest: exec_instr_known_addrs_new_thread) done lemma exec_instr_New_same_addr_same: "\ (ta, xcp', h', frs') \ exec_instr ins P t h stk loc C M pc frs; \ta\\<^bsub>o\<^esub> ! i = NewHeapElem a x; i < length \ta\\<^bsub>o\<^esub>; \ta\\<^bsub>o\<^esub> ! j = NewHeapElem a x'; j < length \ta\\<^bsub>o\<^esub> \ \ i = j" apply(cases ins) apply(auto simp add: nth_Cons' split: prod.split_asm if_split_asm) apply(auto split: extCallRet.split_asm dest: red_external_aggr_New_same_addr_same) done lemma exec_New_same_addr_same: "\ (ta, xcp', h', frs') \ exec P t (xcp, h, frs); \ta\\<^bsub>o\<^esub> ! i = NewHeapElem a x; i < length \ta\\<^bsub>o\<^esub>; \ta\\<^bsub>o\<^esub> ! j = NewHeapElem a x'; j < length \ta\\<^bsub>o\<^esub> \ \ i = j" apply(cases "(P, t, xcp, h, frs)" rule: exec.cases) apply(auto dest: exec_instr_New_same_addr_same) done lemma exec_1_d_New_same_addr_same: "\ P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', h', frs'); \ta\\<^bsub>o\<^esub> ! i = NewHeapElem a x; i < length \ta\\<^bsub>o\<^esub>; \ta\\<^bsub>o\<^esub> ! j = NewHeapElem a x'; j < length \ta\\<^bsub>o\<^esub> \ \ i = j" by(erule jvmd_NormalE)(rule exec_New_same_addr_same) end locale JVM_allocated_heap = allocated_heap + constrains addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and allocated :: "'heap \ 'addr set" and P :: "'addr jvm_prog" sublocale JVM_allocated_heap < JVM_heap by(unfold_locales) context JVM_allocated_heap begin lemma exec_instr_allocated_mono: "\ (ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs; check_instr i P h stk loc C M pc frs \ \ allocated h \ allocated h'" apply(cases i) apply(auto 4 4 simp add: split_beta has_method_def is_native.simps split: if_split_asm sum.split_asm intro: allocate_allocated_mono dest: heap_write_allocated_same dest!: red_external_aggr_allocated_mono del: subsetI) done lemma mexecd_allocated_mono: "mexecd P t (xcpfrs, h) ta (xcpfrs', h') \ allocated h \ allocated h'" apply(cases xcpfrs) apply(cases xcpfrs') apply(simp) apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(auto del: subsetI simp add: check_def dest: exec_instr_allocated_mono) done lemma exec_instr_allocatedD: "\ (ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs; check_instr i P h stk loc C M pc frs; NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub> \ \ ad \ allocated h' \ ad \ allocated h" apply(cases i) apply(auto 4 4 split: if_split_asm prod.split_asm dest: allocate_allocatedD dest!: red_external_aggr_allocatedD simp add: has_method_def is_native.simps) done lemma mexecd_allocatedD: "\ mexecd P t (xcpfrs, h) ta (xcpfrs', h'); NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub> \ \ ad \ allocated h' \ ad \ allocated h" apply(cases xcpfrs) apply(cases xcpfrs') apply(simp) apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(auto del: subsetI dest: exec_instr_allocatedD simp add: check_def) done lemma exec_instr_NewHeapElemD: "\ (ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs; check_instr i P h stk loc C M pc frs; ad \ allocated h'; ad \ allocated h \ \ \CTn. NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>" apply(cases i) apply(auto 4 3 split: if_split_asm prod.split_asm sum.split_asm dest: allocate_allocatedD heap_write_allocated_same dest!: red_external_aggr_NewHeapElemD simp add: is_native.simps has_method_def) done lemma mexecd_NewHeapElemD: "\ mexecd P t (xcpfrs, h) ta (xcpfrs', h'); ad \ allocated h'; ad \ allocated h \ \ \CTn. NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>" apply(cases xcpfrs) apply(cases xcpfrs') apply(simp) apply(erule jvmd_NormalE) apply(cases "fst xcpfrs") apply(auto dest: exec_instr_NewHeapElemD simp add: check_def) done lemma mexecd_allocated_multithreaded: "allocated_multithreaded addr2thread_id thread_id2addr empty_heap allocate typeof_addr heap_write allocated JVM_final (mexecd P) P" proof fix t x m ta x' m' assume "mexecd P t (x, m) ta (x', m')" thus "allocated m \ allocated m'" by(rule mexecd_allocated_mono) next fix x t m ta x' m' ad CTn assume "mexecd P t (x, m) ta (x', m')" and "NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>" thus "ad \ allocated m' \ ad \ allocated m" by(rule mexecd_allocatedD) next fix t x m ta x' m' ad assume "mexecd P t (x, m) ta (x', m')" and "ad \ allocated m'" "ad \ allocated m" thus "\CTn. NewHeapElem ad CTn \ set \ta\\<^bsub>o\<^esub>" by(rule mexecd_NewHeapElemD) next fix t x m ta x' m' i a CTn j CTn' assume "mexecd P t (x, m) ta (x', m')" and "\ta\\<^bsub>o\<^esub> ! i = NewHeapElem a CTn" "i < length \ta\\<^bsub>o\<^esub>" and "\ta\\<^bsub>o\<^esub> ! j = NewHeapElem a CTn'" "j < length \ta\\<^bsub>o\<^esub>" thus "i = j" by(auto dest: exec_1_d_New_same_addr_same simp add: split_beta) qed end sublocale JVM_allocated_heap < execd_mthr: allocated_multithreaded addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated JVM_final "mexecd P" P by(rule mexecd_allocated_multithreaded) context JVM_allocated_heap begin lemma mexecd_known_addrs: assumes wf: "wf_prog wfmd P" and ok: "start_heap_ok" shows "known_addrs addr2thread_id thread_id2addr empty_heap allocate typeof_addr heap_write allocated jvm_known_addrs JVM_final (mexecd P) P" proof fix t x m ta x' m' assume "mexecd P t (x, m) ta (x', m')" thus "jvm_known_addrs t x' \ jvm_known_addrs t x \ new_obs_addrs \ta\\<^bsub>o\<^esub>" by(rule exec_d_known_addrs_mono[OF ok]) next fix t x m ta x' m' t' x'' m'' assume "mexecd P t (x, m) ta (x', m')" and "NewThread t' x'' m'' \ set \ta\\<^bsub>t\<^esub>" thus "jvm_known_addrs t' x'' \ jvm_known_addrs t x" by(rule mexecd_known_addrs_new_thread) next fix t x m ta x' m' ad al v assume "mexecd P t (x, m) ta (x', m')" and "ReadMem ad al v \ set \ta\\<^bsub>o\<^esub>" thus "ad \ jvm_known_addrs t x" by(rule mexecd_known_addrs_ReadMem) next fix t x m ta x' m' n ad al ad' assume "mexecd P t (x, m) ta (x', m')" and "\ta\\<^bsub>o\<^esub> ! n = WriteMem ad al (Addr ad')" "n < length \ta\\<^bsub>o\<^esub>" thus "ad' \ jvm_known_addrs t x \ ad' \ new_obs_addrs (take n \ta\\<^bsub>o\<^esub>)" by(rule mexecd_known_addrs_WriteMem) qed end context JVM_heap begin lemma exec_instr_read_typeable: assumes exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" and read: "ReadMem ad al v \ set \ta\\<^bsub>o\<^esub>" shows "\T'. P,h \ ad@al : T'" using exec check read proof(cases i) case ALoad with assms show ?thesis by(fastforce simp add: split_beta is_Ref_def nat_less_iff word_sless_alt intro: addr_loc_type.intros split: if_split_asm) next case (Getfield F D) with assms show ?thesis by(clarsimp simp add: split_beta is_Ref_def split: if_split_asm)(blast intro: addr_loc_type.intros dest: has_visible_field has_field_mono) next case (Invoke M n) with exec check read obtain a vs ta' va T where "(ta', va, h') \ red_external_aggr P t a M vs h" and "ReadMem ad al v \ set \ta'\\<^bsub>o\<^esub>" by(auto split: if_split_asm simp add: is_Ref_def) thus ?thesis by(rule red_external_aggr_read_mem_typeable) next case (CAS F D) with assms show ?thesis by(clarsimp simp add: split_beta is_Ref_def conf_def split: if_split_asm) (force intro: addr_loc_type.intros dest: has_visible_field[THEN has_field_mono]) qed(auto simp add: split_beta is_Ref_def split: if_split_asm) lemma exec_1_d_read_typeable: "\ P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', h', frs'); ReadMem ad al v \ set \ta\\<^bsub>o\<^esub> \ \ \T'. P,h \ ad@al : T'" apply(erule jvmd_NormalE) apply(cases "(P, t, xcp, h, frs)" rule: exec.cases) apply(auto intro: exec_instr_read_typeable simp add: check_def) done end sublocale JVM_heap_base < execd_mthr: if_multithreaded JVM_final "mexecd P" convert_RA for P by(unfold_locales) context JVM_heap_conf begin lemma JVM_conf_read_heap_read_typed: "JVM_conf_read addr2thread_id thread_id2addr empty_heap allocate typeof_addr (heap_read_typed P) heap_write hconf P" proof - interpret conf: heap_conf_read addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr "heap_read_typed P" heap_write hconf P by(rule heap_conf_read_heap_read_typed) show ?thesis by(unfold_locales) qed lemma exec_instr_New_typeof_addrD: "\ (ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs; check_instr i P h stk loc C M pc frs; hconf h; NewHeapElem a x \ set \ta\\<^bsub>o\<^esub> \ \ typeof_addr h' a = Some x" apply(cases i) apply(auto dest: allocate_SomeD split: prod.split_asm if_split_asm) apply(auto 4 4 split: extCallRet.split_asm dest!: red_external_aggr_New_typeof_addrD simp add: has_method_def is_native.simps) done lemma exec_1_d_New_typeof_addrD: "\ P,t \ Normal (xcp, h, frs) -ta-jvmd\ Normal (xcp', h', frs'); NewHeapElem a x \ set \ta\\<^bsub>o\<^esub>; hconf h \ \ typeof_addr h' a = Some x" apply(erule jvmd_NormalE) apply(cases "xcp") apply(auto dest: exec_instr_New_typeof_addrD simp add: check_def) done lemma exec_instr_non_speculative_typeable: assumes exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" and sc: "non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>))" and vs_conf: "vs_conf P h vs" and hconf: "hconf h" shows "(ta, xcp', h', frs') \ JVM_heap_base.exec_instr addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr (heap_read_typed P) heap_write i P t h stk loc C M pc frs" proof - note [simp] = JVM_heap_base.exec_instr.simps and [split] = if_split_asm prod.split_asm sum.split_asm and [split del] = if_split from assms show "?thesis" proof(cases i) case ALoad with assms show ?thesis by(auto 4 3 intro!: heap_read_typedI dest: vs_confD addr_loc_type_fun) next case Getfield with assms show ?thesis by(auto 4 3 intro!: heap_read_typedI dest: vs_confD addr_loc_type_fun) next case CAS with assms show ?thesis by(auto 4 3 intro!: heap_read_typedI dest: vs_confD addr_loc_type_fun) next case Invoke with assms show ?thesis by(fastforce dest: red_external_aggr_non_speculative_typeable simp add: has_method_def is_native.simps) qed(auto) qed lemma exec_instr_non_speculative_vs_conf: assumes exec: "(ta, xcp', h', frs') \ exec_instr i P t h stk loc C M pc frs" and check: "check_instr i P h stk loc C M pc frs" and sc: "non_speculative P vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" and vs_conf: "vs_conf P h vs" and hconf: "hconf h" shows "vs_conf P h' (w_values P vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" proof - note [simp] = JVM_heap_base.exec_instr.simps take_Cons' and [split] = if_split_asm prod.split_asm sum.split_asm and [split del] = if_split from assms show ?thesis proof(cases i) case New with assms show ?thesis by(auto 4 4 dest: hext_allocate vs_conf_allocate intro: vs_conf_hext) next case NewArray with assms show ?thesis by(auto 4 4 dest: hext_allocate vs_conf_allocate intro: vs_conf_hext cong: if_cong) next case Invoke with assms show ?thesis by(fastforce dest: red_external_aggr_non_speculative_vs_conf simp add: has_method_def is_native.simps) next case AStore { assume "hd (tl (tl stk)) \ Null" and "\ the_Intg (hd (tl stk)) int (alen_of_htype (the (typeof_addr h (the_Addr (hd (tl (tl stk))))))) \ sint (the_Intg (hd (tl stk)))" and "P \ the (typeof\<^bsub>h\<^esub> (hd stk)) \ the_Array (ty_of_htype (the (typeof_addr h (the_Addr (hd (tl (tl stk)))))))" moreover hence "nat (sint (the_Intg (hd (tl stk)))) < alen_of_htype (the (typeof_addr h (the_Addr (hd (tl (tl stk))))))" by(auto simp add: not_le nat_less_iff word_sle_eq word_sless_eq not_less) with assms AStore have "nat (sint (the_Intg (hd (tl stk)))) < alen_of_htype (the (typeof_addr h' (the_Addr (hd (tl (tl stk))))))" by(auto dest!: hext_arrD hext_heap_write) ultimately have "\T. P,h' \ the_Addr (hd (tl (tl stk)))@ACell (nat (sint (the_Intg (hd (tl stk))))) : T \ P,h' \ hd stk :\ T" using assms AStore by(auto 4 4 simp add: is_Ref_def conf_def dest!: hext_heap_write dest: hext_arrD intro!: addr_loc_type.intros intro: typeof_addr_hext_mono type_of_hext_type_of) } thus ?thesis using assms AStore by(auto intro!: vs_confI)(blast intro: addr_loc_type_hext_mono conf_hext dest: hext_heap_write vs_confD)+ next case Putfield show ?thesis using assms Putfield by(auto intro!: vs_confI dest!: hext_heap_write)(blast intro: addr_loc_type.intros addr_loc_type_hext_mono typeof_addr_hext_mono has_field_mono[OF has_visible_field] conf_hext dest: vs_confD)+ next case CAS show ?thesis using assms CAS by(auto intro!: vs_confI dest!: hext_heap_write)(blast intro: addr_loc_type.intros addr_loc_type_hext_mono typeof_addr_hext_mono has_field_mono[OF has_visible_field] conf_hext dest: vs_confD)+ qed(auto) qed lemma mexecd_non_speculative_typeable: "\ P,t \ Normal (xcp, h, stk) -ta-jvmd\ Normal (xcp', h', frs'); non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>)); vs_conf P h vs; hconf h \ \ JVM_heap_base.exec_1_d addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr (heap_read_typed P) heap_write P t (Normal (xcp, h, stk)) ta (Normal (xcp', h', frs'))" apply(erule jvmd_NormalE) apply(cases xcp) apply(auto intro!: JVM_heap_base.exec_1_d.intros simp add: JVM_heap_base.exec_d_def check_def JVM_heap_base.exec.simps intro: exec_instr_non_speculative_typeable) done lemma mexecd_non_speculative_vs_conf: "\ P,t \ Normal (xcp, h, stk) -ta-jvmd\ Normal (xcp', h', frs'); non_speculative P vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>))); vs_conf P h vs; hconf h \ \ vs_conf P h' (w_values P vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" apply(erule jvmd_NormalE) apply(cases xcp) apply(auto intro!: JVM_heap_base.exec_1_d.intros simp add: JVM_heap_base.exec_d_def check_def JVM_heap_base.exec.simps intro: exec_instr_non_speculative_vs_conf) done end locale JVM_allocated_heap_conf = JVM_heap_conf addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_allocated_heap addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated P for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and allocated :: "'heap \ 'addr set" and P :: "'addr jvm_prog" begin lemma mexecd_known_addrs_typing: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and ok: "start_heap_ok" shows "known_addrs_typing addr2thread_id thread_id2addr empty_heap allocate typeof_addr heap_write allocated jvm_known_addrs JVM_final (mexecd P) (\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \) P" proof - from wf obtain wf_md where "wf_prog wf_md P" by(blast dest: wt_jvm_progD) then interpret known_addrs addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated jvm_known_addrs JVM_final "mexecd P" P using ok by(rule mexecd_known_addrs) show ?thesis proof fix t x m ta x' m' assume "mexecd P t (x, m) ta (x', m')" thus "m \ m'" by(auto simp add: split_beta intro: exec_1_d_hext) next fix t x m ta x' m' vs assume exec: "mexecd P t (x, m) ta (x', m')" and ts_ok: "(\(xcp, frstls) h. \ \ t:(xcp, h, frstls) \) x m" and vs: "vs_conf P m vs" and ns: "non_speculative P vs (llist_of (map NormalAction \ta\\<^bsub>o\<^esub>))" let ?mexecd = "JVM_heap_base.mexecd addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr (heap_read_typed P) heap_write P" have lift: "lifting_wf JVM_final ?mexecd (\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \)" by(intro JVM_conf_read.lifting_wf_correct_state_d JVM_conf_read_heap_read_typed wf) from exec ns vs ts_ok have exec': "?mexecd t (x, m) ta (x', m')" by(auto simp add: split_beta correct_state_def dest: mexecd_non_speculative_typeable) thus "(\(xcp, frstls) h. \ \ t:(xcp, h, frstls) \) x' m'" using ts_ok by(rule lifting_wf.preserves_red[OF lift]) { fix t'' x'' m'' assume New: "NewThread t'' x'' m'' \ set \ta\\<^bsub>t\<^esub>" with exec have "m'' = snd (x', m')" by(rule execd_mthr.new_thread_memory) thus "(\(xcp, frstls) h. \ \ t'':(xcp, h, frstls) \) x'' m''" using lifting_wf.preserves_NewThread[where ?r="?mexecd", OF lift exec' ts_ok] New by auto } { fix t'' x'' assume "(\(xcp, frstls) h. \ \ t'':(xcp, h, frstls) \) x'' m" with lift exec' ts_ok show "(\(xcp, frstls) h. \ \ t'':(xcp, h, frstls) \) x'' m'" by(rule lifting_wf.preserves_other) } next fix t x m ta x' m' vs n assume exec: "mexecd P t (x, m) ta (x', m')" and ts_ok: "(\(xcp, frstls) h. \ \ t:(xcp, h, frstls) \) x m" and vs: "vs_conf P m vs" and ns: "non_speculative P vs (llist_of (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" thus "vs_conf P m' (w_values P vs (take n (map NormalAction \ta\\<^bsub>o\<^esub>)))" by(auto simp add: correct_state_def dest: mexecd_non_speculative_vs_conf) next fix t x m ta x' m' ad al v assume "mexecd P t (x, m) ta (x', m')" and "(\(xcp, frstls) h. \ \ t:(xcp, h, frstls) \) x m" and "ReadMem ad al v \ set \ta\\<^bsub>o\<^esub>" thus "\T. P,m \ ad@al : T" by(auto simp add: correct_state_def split_beta dest: exec_1_d_read_typeable) next fix t x m ta x' m' ad hT assume "mexecd P t (x, m) ta (x', m')" and "(\(xcp, frstls) h. \ \ t:(xcp, h, frstls) \) x m" and "NewHeapElem ad hT \ set \ta\\<^bsub>o\<^esub>" thus "typeof_addr m' ad = \hT\" by(auto dest: exec_1_d_New_typeof_addrD[where x="hT"] simp add: split_beta correct_state_def) qed qed lemma executions_sc: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and wf_start: "wf_start_state P C M vs" and vs2: "\(ka_Val ` set vs) \ set start_addrs" shows "executions_sc_hb (JVMd_\ P C M vs status) P" (is "executions_sc_hb ?E P") proof - from wf_start obtain Ts T meth D where ok: "start_heap_ok" and sees: "P \ C sees M:Ts\T=\meth\ in D" and vs1: "P,start_heap \ vs [:\] Ts" by cases interpret known_addrs_typing addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated jvm_known_addrs JVM_final "mexecd P" "\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \" P using wf ok by(rule mexecd_known_addrs_typing) from wf obtain wf_md where wf': "wf_prog wf_md P" by(blast dest: wt_jvm_progD) hence "wf_syscls P" by(rule wf_prog_wf_syscls) thus ?thesis proof(rule executions_sc_hb) from correct_jvm_state_initial[OF wf wf_start] show "correct_state_ts \ (thr (JVM_start_state P C M vs)) start_heap" by(simp add: correct_jvm_state_def start_state_def split_beta) next show "jvm_known_addrs start_tid ((\(mxs, mxl0, b) vs. (None, [([], Null # vs @ replicate mxl0 undefined_value, fst (method P C M), M, 0)])) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" using vs2 by(auto simp add: split_beta start_addrs_allocated jvm_known_addrs_def intro: start_tid_start_addrs[OF \wf_syscls P\ ok]) qed qed end declare split_paired_Ex [simp del] declare eq_upto_seq_inconsist_simps [simp] context JVM_progress begin abbreviation (input) jvm_non_speculative_read_bound :: nat where "jvm_non_speculative_read_bound \ 2" lemma exec_instr_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and vs: "vs_conf P (shr s) vs" and hconf: "hconf (shr s)" and exec_i: "(ta, xcp', h', frs') \ exec_instr i P t (shr s) stk loc C M pc frs" and check: "check_instr i P (shr s) stk loc C M pc frs" and aok: "execd_mthr.mthr.if.actions_ok s t ta" and i: "I < length \ta\\<^bsub>o\<^esub>" and read: "\ta\\<^bsub>o\<^esub> ! I = ReadMem a'' al'' v" and v': "v' \ w_values P vs (map NormalAction (take I \ta\\<^bsub>o\<^esub>)) (a'', al'')" and ns: "non_speculative P vs (llist_of (map NormalAction (take I \ta\\<^bsub>o\<^esub>)))" shows "\ta' xcp'' h'' frs''. (ta', xcp'', h'', frs'') \ exec_instr i P t (shr s) stk loc C M pc frs \ execd_mthr.mthr.if.actions_ok s t ta' \ I < length \ta'\\<^bsub>o\<^esub> \ take I \ta'\\<^bsub>o\<^esub> = take I \ta\\<^bsub>o\<^esub> \ \ta'\\<^bsub>o\<^esub> ! I = ReadMem a'' al'' v' \ length \ta'\\<^bsub>o\<^esub> \ max jvm_non_speculative_read_bound (length \ta\\<^bsub>o\<^esub>)" using exec_i i read proof(cases i) case [simp]: ALoad let ?a = "the_Addr (hd (tl stk))" let ?i = "the_Intg (hd stk)" from exec_i i read have Null: "hd (tl stk) \ Null" and bounds: "0 <=s ?i" "sint ?i < int (alen_of_htype (the (typeof_addr (shr s) ?a)))" and [simp]: "I = 0" "a'' = ?a" "al'' = ACell (nat (sint ?i))" by(auto split: if_split_asm) from Null check obtain a T n where a: "length stk > 1" "hd (tl stk) = Addr a" and type: "typeof_addr (shr s) ?a = \Array_type T n\" by(fastforce simp add: is_Ref_def) from bounds type have "nat (sint ?i) < n" by (simp add: word_sle_eq nat_less_iff) with type have adal: "P,shr s \ ?a@ACell (nat (sint ?i)) : T" by(rule addr_loc_type.intros) from v' vs adal have "P,shr s \ v' :\ T" by(auto dest!: vs_confD dest: addr_loc_type_fun) with hrt adal have "heap_read (shr s) ?a (ACell (nat (sint ?i))) v'" using hconf by(rule heap_read_typeableD) - with type bounds Null aok exec_i show ?thesis by(fastforce) + with type Null aok exec_i show ?thesis apply auto using bounds by fastforce+ next case [simp]: (Getfield F D) let ?a = "the_Addr (hd stk)" from exec_i i read have Null: "hd stk \ Null" and [simp]: "I = 0" "a'' = ?a" "al'' = CField D F" by(auto split: if_split_asm) with check obtain U T fm C' a where sees: "P \ D sees F:T (fm) in D" and type: "typeof_addr (shr s) ?a = \U\" and sub: "P \ class_type_of U \\<^sup>* D" and a: "hd stk = Addr a" "length stk > 0" by(auto simp add: is_Ref_def) from has_visible_field[OF sees] sub have "P \ class_type_of U has F:T (fm) in D" by(rule has_field_mono) with type have adal: "P,shr s \ ?a@CField D F : T" by(rule addr_loc_type.intros) from v' vs adal have "P,shr s \ v' :\ T" by(auto dest!: vs_confD dest: addr_loc_type_fun) with hrt adal have "heap_read (shr s) ?a (CField D F) v'" using hconf by(rule heap_read_typeableD) with type Null aok exec_i show ?thesis by(fastforce) next case [simp]: (CAS F D) let ?a = "the_Addr (hd (tl (tl stk)))" from exec_i i read have Null: "hd (tl (tl stk)) \ Null" and [simp]: "I = 0" "a'' = ?a" "al'' = CField D F" by(auto split: if_split_asm simp add: nth_Cons') with check obtain U T fm C' a where sees: "P \ D sees F:T (fm) in D" and type: "typeof_addr (shr s) ?a = \U\" and sub: "P \ class_type_of U \\<^sup>* D" and a: "hd (tl (tl stk)) = Addr a" "length stk > 2" and v: "P,shr s \ hd stk :\ T" by(auto simp add: is_Ref_def) from has_visible_field[OF sees] sub have "P \ class_type_of U has F:T (fm) in D" by(rule has_field_mono) with type have adal: "P,shr s \ ?a@CField D F : T" by(rule addr_loc_type.intros) from v' vs adal have "P,shr s \ v' :\ T" by(auto dest!: vs_confD dest: addr_loc_type_fun) with hrt adal have read: "heap_read (shr s) ?a (CField D F) v'" using hconf by(rule heap_read_typeableD) show ?thesis proof(cases "v' = hd (tl stk)") case True from heap_write_total[OF hconf adal v] a obtain h' where "heap_write (shr s) a (CField D F) (hd stk) h'" by auto then show ?thesis using read a True aok exec_i by fastforce next case False then show ?thesis using read a aok exec_i by(fastforce intro!: disjI2) qed next case [simp]: (Invoke M n) let ?a = "the_Addr (stk ! n)" let ?vs = "rev (take n stk)" from exec_i i read have Null: "stk ! n \ Null" and iec: "snd (snd (snd (method P (class_type_of (the (typeof_addr (shr s) ?a))) M))) = Native" by(auto split: if_split_asm) with check obtain a T Ts Tr D where a: "stk ! n = Addr a" "n < length stk" and type: "typeof_addr (shr s) ?a = \T\" and extwt: "P \ class_type_of T sees M:Ts\Tr = Native in D" "D\M(Ts) :: Tr" by(auto simp add: is_Ref_def has_method_def) from extwt have native: "is_native P T M" by(auto simp add: is_native.simps) from Null iec type exec_i obtain ta' va where red: "(ta', va, h') \ red_external_aggr P t ?a M ?vs (shr s)" and ta: "ta = extTA2JVM P ta'" by(fastforce) from aok ta have aok': "execd_mthr.mthr.if.actions_ok s t ta'" by simp from red_external_aggr_non_speculative_read[OF hrt vs red[unfolded a the_Addr.simps] _ aok' hconf, of I a'' al'' v v'] native type i read v' ns a ta obtain ta'' va'' h'' where "(ta'', va'', h'') \ red_external_aggr P t a M (rev (take n stk)) (shr s)" and "execd_mthr.mthr.if.actions_ok s t ta''" and "I < length \ta''\\<^bsub>o\<^esub>" "take I \ta''\\<^bsub>o\<^esub> = take I \ta'\\<^bsub>o\<^esub>" and "\ta''\\<^bsub>o\<^esub> ! I = ReadMem a'' al'' v'" "length \ta''\\<^bsub>o\<^esub> \ length \ta'\\<^bsub>o\<^esub>" by auto thus ?thesis using Null iec ta extwt a type by(cases va'') force+ qed(auto simp add: split_beta split: if_split_asm) lemma exec_1_d_non_speculative_read: assumes hrt: "heap_read_typeable hconf P" and vs: "vs_conf P (shr s) vs" and exec: "P,t \ Normal (xcp, shr s, frs) -ta-jvmd\ Normal (xcp', h', frs')" and aok: "execd_mthr.mthr.if.actions_ok s t ta" and hconf: "hconf (shr s)" and i: "I < length \ta\\<^bsub>o\<^esub>" and read: "\ta\\<^bsub>o\<^esub> ! I = ReadMem a'' al'' v" and v': "v' \ w_values P vs (map NormalAction (take I \ta\\<^bsub>o\<^esub>)) (a'', al'')" and ns: "non_speculative P vs (llist_of (map NormalAction (take I \ta\\<^bsub>o\<^esub>)))" shows "\ta' xcp'' h'' frs''. P,t \ Normal (xcp, shr s, frs) -ta'-jvmd\ Normal (xcp'', h'', frs'') \ execd_mthr.mthr.if.actions_ok s t ta' \ I < length \ta'\\<^bsub>o\<^esub> \ take I \ta'\\<^bsub>o\<^esub> = take I \ta\\<^bsub>o\<^esub> \ \ta'\\<^bsub>o\<^esub> ! I = ReadMem a'' al'' v' \ length \ta'\\<^bsub>o\<^esub> \ max jvm_non_speculative_read_bound (length \ta\\<^bsub>o\<^esub>)" using assms apply - apply(erule jvmd_NormalE) apply(cases "(P, t, xcp, shr s, frs)" rule: exec.cases) apply simp defer apply simp apply clarsimp apply(drule (3) exec_instr_non_speculative_read) apply(clarsimp simp add: check_def has_method_def) apply simp apply(rule i) apply(rule read) apply(rule v') apply(rule ns) apply(clarsimp simp add: exec_1_d.simps exec_d_def) done end declare split_paired_Ex [simp] declare eq_upto_seq_inconsist_simps [simp del] locale JVM_allocated_progress = JVM_progress addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_allocated_heap_conf addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf allocated P for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and allocated :: "'heap \ 'addr set" and P :: "'addr jvm_prog" begin lemma non_speculative_read: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "execd_mthr.if.non_speculative_read jvm_non_speculative_read_bound (init_fin_lift_state status (JVM_start_state P C M vs)) (w_values P (\_. {}) (map snd (lift_start_obs start_tid start_heap_obs)))" (is "execd_mthr.if.non_speculative_read _ ?start_state ?start_vs") proof(rule execd_mthr.if.non_speculative_readI) fix ttas s' t x ta x' m' i ad al v v' assume \Red: "execd_mthr.mthr.if.RedT P ?start_state ttas s'" and sc: "non_speculative P ?start_vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and ts't: "thr s' t = \(x, no_wait_locks)\" and red: "execd_mthr.init_fin P t (x, shr s') ta (x', m')" and aok: "execd_mthr.mthr.if.actions_ok s' t ta" and i: "i < length \ta\\<^bsub>o\<^esub>" and ns': "non_speculative P (w_values P ?start_vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (take i \ta\\<^bsub>o\<^esub>))" and read: "\ta\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v)" and v': "v' \ w_values P ?start_vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas) @ take i \ta\\<^bsub>o\<^esub>) (ad, al)" from wf_start obtain Ts T meth D where ok: "start_heap_ok" and sees: "P \ C sees M:Ts\T = \meth\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases let ?conv = "\ttas. concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)" let ?vs' = "w_values P ?start_vs (?conv ttas)" let ?wt_ok = "init_fin_lift (\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \)" let ?start_obs = "map snd (lift_start_obs start_tid start_heap_obs)" from wf obtain wf_md where wf': "wf_prog wf_md P" by(blast dest: wt_jvm_progD) interpret known_addrs_typing addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated jvm_known_addrs JVM_final "mexecd P" "\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \" using wf ok by(rule mexecd_known_addrs_typing) from conf have len2: "length vs = length Ts" by(rule list_all2_lengthD) from correct_jvm_state_initial[OF wf wf_start] have "correct_state_ts \ (thr (JVM_start_state P C M vs)) start_heap" by(simp add: correct_jvm_state_def start_state_def split_beta) hence ts_ok_start: "ts_ok ?wt_ok (thr ?start_state) (shr ?start_state)" unfolding ts_ok_init_fin_lift_init_fin_lift_state by(simp add: start_state_def split_beta) have sc': "non_speculative P ?start_vs (lmap snd (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) (llist_of ttas))))" using sc by(simp add: lmap_lconcat llist.map_comp o_def split_def lconcat_llist_of[symmetric]) from start_state_vs_conf[OF wf_prog_wf_syscls[OF wf']] have vs_conf_start: "vs_conf P (shr ?start_state) ?start_vs" by(simp add:init_fin_lift_state_conv_simps start_state_def split_beta) with \Red ts_ok_start sc have wt': "ts_ok ?wt_ok (thr s') (shr s')" and vs': "vs_conf P (shr s') ?vs'" by(rule if_RedT_non_speculative_invar)+ from red i read obtain xcp frs xcp' frs' ta' where x: "x = (Running, xcp, frs)" and x': "x' = (Running, xcp', frs')" and ta: "ta = convert_TA_initial (convert_obs_initial ta')" and red': "P,t \ Normal (xcp, shr s', frs) -ta'-jvmd\ Normal (xcp', m', frs')" by cases fastforce+ from ts't wt' x have hconf: "hconf (shr s')" by(auto dest!: ts_okD simp add: correct_state_def) have aok': "execd_mthr.mthr.if.actions_ok s' t ta'" using aok unfolding ta by simp from i read v' ns' ta have "i < length \ta'\\<^bsub>o\<^esub>" and "\ta'\\<^bsub>o\<^esub> ! i = ReadMem ad al v" and "v' \ w_values P ?vs' (map NormalAction (take i \ta'\\<^bsub>o\<^esub>)) (ad, al)" and "non_speculative P ?vs' (llist_of (map NormalAction (take i \ta'\\<^bsub>o\<^esub>)))" by(simp_all add: take_map) from exec_1_d_non_speculative_read[OF hrt vs' red' aok' hconf this] obtain ta'' xcp'' frs'' h'' where red'': "P,t \ Normal (xcp, shr s', frs) -ta''-jvmd\ Normal (xcp'', h'', frs'')" and aok'': "execd_mthr.mthr.if.actions_ok s' t ta''" and i'': " i < length \ta''\\<^bsub>o\<^esub>" and eq'': "take i \ta''\\<^bsub>o\<^esub> = take i \ta'\\<^bsub>o\<^esub>" and read'': "\ta''\\<^bsub>o\<^esub> ! i = ReadMem ad al v'" and len'': "length \ta''\\<^bsub>o\<^esub> \ max jvm_non_speculative_read_bound (length \ta'\\<^bsub>o\<^esub>)" by blast let ?x' = "(Running, xcp'', frs'')" let ?ta' = "convert_TA_initial (convert_obs_initial ta'')" from red'' have "execd_mthr.init_fin P t (x, shr s') ?ta' (?x', h'')" unfolding x by -(rule execd_mthr.init_fin.NormalAction, simp) moreover from aok'' have "execd_mthr.mthr.if.actions_ok s' t ?ta'" by simp moreover from i'' have "i < length \?ta'\\<^bsub>o\<^esub>" by simp moreover from eq'' have "take i \?ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub>" unfolding ta by(simp add: take_map) moreover from read'' i'' have "\?ta'\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v')" by(simp add: nth_map) moreover from len'' have "length \?ta'\\<^bsub>o\<^esub> \ max jvm_non_speculative_read_bound (length \ta\\<^bsub>o\<^esub>)" unfolding ta by simp ultimately show "\ta' x'' m''. execd_mthr.init_fin P t (x, shr s') ta' (x'', m'') \ execd_mthr.mthr.if.actions_ok s' t ta' \ i < length \ta'\\<^bsub>o\<^esub> \ take i \ta'\\<^bsub>o\<^esub> = take i \ta\\<^bsub>o\<^esub> \ \ta'\\<^bsub>o\<^esub> ! i = NormalAction (ReadMem ad al v') \ length \ta'\\<^bsub>o\<^esub> \ max jvm_non_speculative_read_bound (length \ta\\<^bsub>o\<^esub>)" by blast qed lemma JVM_cut_and_update: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "execd_mthr.if.cut_and_update (init_fin_lift_state status (JVM_start_state P C M vs)) (mrw_values P Map.empty (map snd (lift_start_obs start_tid start_heap_obs)))" proof - from wf_start obtain Ts T meth D where ok: "start_heap_ok" and sees: "P \ C sees M:Ts\T = \meth\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases interpret known_addrs_typing addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated jvm_known_addrs JVM_final "mexecd P" "\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \" using wf ok by(rule mexecd_known_addrs_typing) let ?start_vs = "w_values P (\_. {}) (map snd (lift_start_obs start_tid start_heap_obs))" let ?wt_ok = "init_fin_lift (\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \)" let ?start_obs = "map snd (lift_start_obs start_tid start_heap_obs)" let ?start_state = "init_fin_lift_state status (JVM_start_state P C M vs)" from wf obtain wf_md where wf': "wf_prog wf_md P" by(blast dest: wt_jvm_progD) hence "wf_syscls P" by(rule wf_prog_wf_syscls) moreover note non_speculative_read[OF wf hrt wf_start ka] moreover have "ts_ok ?wt_ok (thr ?start_state) (shr ?start_state)" using correct_jvm_state_initial[OF wf wf_start] by(simp add: correct_jvm_state_def start_state_def split_beta) moreover have ka: "jvm_known_addrs start_tid ((\(mxs, mxl0, b) vs. (None, [([], Null # vs @ replicate mxl0 undefined_value, fst (method P C M), M, 0)])) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" using ka by(auto simp add: split_beta start_addrs_allocated jvm_known_addrs_def intro: start_tid_start_addrs[OF \wf_syscls P\ ok]) ultimately show ?thesis by(rule non_speculative_read_into_cut_and_update) qed lemma JVM_drf: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "drf (JVMd_\ P C M vs status) P" proof - from wf_start obtain Ts T meth D where ok: "start_heap_ok" and sees: "P \ C sees M:Ts\T = \meth\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases from wf obtain wf_md where wf': "wf_prog wf_md P" by(blast dest: wt_jvm_progD) hence "wf_syscls P" by(rule wf_prog_wf_syscls) with JVM_cut_and_update[OF assms] show ?thesis proof(rule known_addrs_typing.drf[OF mexecd_known_addrs_typing[OF wf ok]]) from correct_jvm_state_initial[OF wf wf_start] show "correct_state_ts \ (thr (JVM_start_state P C M vs)) start_heap" by(simp add: correct_jvm_state_def start_state_def split_beta) next show "jvm_known_addrs start_tid ((\(mxs, mxl0, b) vs. (None, [([], Null # vs @ replicate mxl0 undefined_value, fst (method P C M), M, 0)])) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" using ka by(auto simp add: split_beta start_addrs_allocated jvm_known_addrs_def intro: start_tid_start_addrs[OF \wf_syscls P\ ok]) qed qed lemma JVM_sc_legal: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "sc_legal (JVMd_\ P C M vs status) P" proof - from wf_start obtain Ts T meth D where ok: "start_heap_ok" and sees: "P \ C sees M:Ts\T = \meth\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases interpret known_addrs_typing addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write allocated jvm_known_addrs JVM_final "mexecd P" "\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \" using wf ok by(rule mexecd_known_addrs_typing) from wf obtain wf_md where wf': "wf_prog wf_md P" by(blast dest: wt_jvm_progD) hence "wf_syscls P" by(rule wf_prog_wf_syscls) let ?start_vs = "w_values P (\_. {}) (map snd (lift_start_obs start_tid start_heap_obs))" let ?wt_ok = "init_fin_lift (\t (xcp, frstls) h. \ \ t: (xcp, h, frstls) \)" let ?start_obs = "map snd (lift_start_obs start_tid start_heap_obs)" let ?start_state = "init_fin_lift_state status (JVM_start_state P C M vs)" note \wf_syscls P\ non_speculative_read[OF wf hrt wf_start ka] moreover have "ts_ok ?wt_ok (thr ?start_state) (shr ?start_state)" using correct_jvm_state_initial[OF wf wf_start] by(simp add: correct_jvm_state_def start_state_def split_beta) moreover have ka_allocated: "jvm_known_addrs start_tid ((\(mxs, mxl0, b) vs. (None, [([], Null # vs @ replicate mxl0 undefined_value, fst (method P C M), M, 0)])) (the (snd (snd (snd (method P C M))))) vs) \ allocated start_heap" using ka by(auto simp add: split_beta start_addrs_allocated jvm_known_addrs_def intro: start_tid_start_addrs[OF \wf_syscls P\ ok]) ultimately have "execd_mthr.if.hb_completion ?start_state (lift_start_obs start_tid start_heap_obs)" by(rule non_speculative_read_into_hb_completion) thus ?thesis using \wf_syscls P\ proof(rule sc_legal) from correct_jvm_state_initial[OF wf wf_start] show "correct_state_ts \ (thr (JVM_start_state P C M vs)) start_heap" by(simp add: correct_jvm_state_def start_state_def split_beta) qed(rule ka_allocated) qed lemma JVM_jmm_consistent: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "jmm_consistent (JVMd_\ P C M vs status) P" (is "jmm_consistent ?\ P") proof - interpret drf "?\" P using assms by(rule JVM_drf) interpret sc_legal "?\" P using assms by(rule JVM_sc_legal) show ?thesis by unfold_locales qed lemma JVM_ex_sc_exec: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "\E ws. E \ JVMd_\ P C M vs status \ P \ (E, ws) \ \ sequentially_consistent P (E, ws)" (is "\E ws. _ \ ?\ \ _") proof - interpret jmm: executions_sc_hb ?\ P using assms by -(rule executions_sc) let ?start_state = "init_fin_lift_state status (JVM_start_state P C M vs)" let ?start_mrw = "mrw_values P Map.empty (map snd (lift_start_obs start_tid start_heap_obs))" from execd_mthr.if.sequential_completion_Runs[OF execd_mthr.if.cut_and_update_imp_sc_completion[OF JVM_cut_and_update[OF assms]] ta_seq_consist_convert_RA] obtain ttas where Red: "execd_mthr.mthr.if.mthr.Runs P ?start_state ttas" and sc: "ta_seq_consist P ?start_mrw (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) ttas))" by blast let ?E = "lappend (llist_of (lift_start_obs start_tid start_heap_obs)) (lconcat (lmap (\(t, ta). llist_of (map (Pair t) \ta\\<^bsub>o\<^esub>)) ttas))" from Red have "?E \ ?\" by(blast intro: execd_mthr.mthr.if.\.intros) moreover from Red have tsa: "thread_start_actions_ok ?E" by(blast intro: execd_mthr.thread_start_actions_ok_init_fin execd_mthr.mthr.if.\.intros) from sc have "ta_seq_consist P Map.empty (lmap snd ?E)" unfolding lmap_lappend_distrib lmap_lconcat llist.map_comp split_def o_def lmap_llist_of map_map snd_conv by(simp add: ta_seq_consist_lappend ta_seq_consist_start_heap_obs) from ta_seq_consist_imp_sequentially_consistent[OF tsa jmm.\_new_actions_for_fun[OF \?E \ ?\\] this] obtain ws where "sequentially_consistent P (?E, ws)" "P \ (?E, ws) \" by iprover ultimately show ?thesis by blast qed theorem JVM_consistent: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and hrt: "heap_read_typeable hconf P" and wf_start: "wf_start_state P C M vs" and ka: "\(ka_Val ` set vs) \ set start_addrs" shows "\E ws. legal_execution P (JVMd_\ P C M vs status) (E, ws)" proof - let ?\ = "JVMd_\ P C M vs status" interpret sc_legal "?\" P using assms by(rule JVM_sc_legal) from JVM_ex_sc_exec[OF assms] obtain E ws where "E \ ?\" "P \ (E, ws) \" "sequentially_consistent P (E, ws)" by blast hence "legal_execution P ?\ (E, ws)" by(rule SC_is_legal) thus ?thesis by blast qed end text \ One could now also prove that the aggressive JVM satisfies @{term drf}. The key would be that \welltyped_commute\ also holds for @{term "non_speculative"} prefixes from start. \ end diff --git a/thys/Median_Of_Medians_Selection/Median_Of_Medians_Selection.thy b/thys/Median_Of_Medians_Selection/Median_Of_Medians_Selection.thy --- a/thys/Median_Of_Medians_Selection/Median_Of_Medians_Selection.thy +++ b/thys/Median_Of_Medians_Selection/Median_Of_Medians_Selection.thy @@ -1,974 +1,975 @@ (* File: Median_Of_Medians_Selection.thy Author: Manuel Eberl, TU München The median-of-medians selection algorithm, which runs deterministically in linear time. *) section \The Median-of-Medians Selection Algorithm\ theory Median_Of_Medians_Selection imports Complex_Main "HOL-Library.Multiset" begin subsection \Some facts about lists and multisets\ lemma mset_concat: "mset (concat xss) = sum_list (map mset xss)" by (induction xss) simp_all lemma set_mset_sum_list [simp]: "set_mset (sum_list xs) = (\x\set xs. set_mset x)" by (induction xs) auto lemma filter_mset_image_mset: "filter_mset P (image_mset f A) = image_mset f (filter_mset (\x. P (f x)) A)" by (induction A) auto lemma filter_mset_sum_list: "filter_mset P (sum_list xs) = sum_list (map (filter_mset P) xs)" by (induction xs) simp_all lemma sum_mset_mset_mono: assumes "(\x. x \# A \ f x \# g x)" shows "(\x\#A. f x) \# (\x\#A. g x)" using assms by (induction A) (auto intro!: subset_mset.add_mono) lemma mset_filter_mono: assumes "A \# B" "\x. x \# A \ P x \ Q x" shows "filter_mset P A \# filter_mset Q B" by (rule mset_subset_eqI) (insert assms, auto simp: mset_subset_eq_count count_eq_zero_iff) lemma size_mset_sum_mset_distrib: "size (sum_mset A :: 'a multiset) = sum_mset (image_mset size A)" by (induction A) auto lemma sum_mset_mono: assumes "\x. x \# A \ f x \ (g x :: 'a :: {ordered_ab_semigroup_add,comm_monoid_add})" shows "(\x\#A. f x) \ (\x\#A. g x)" using assms by (induction A) (auto intro!: add_mono) lemma filter_mset_is_empty_iff: "filter_mset P A = {#} \ (\x. x \# A \ \P x)" by (auto simp: multiset_eq_iff count_eq_zero_iff) lemma sorted_filter_less_subset_take: assumes "sorted xs" "i < length xs" shows "{# x \# mset xs. x < xs ! i #} \# mset (take i xs)" using assms proof (induction xs arbitrary: i rule: list.induct) case (Cons x xs i) show ?case proof (cases i) case 0 thus ?thesis using Cons.prems by (auto simp: filter_mset_is_empty_iff) next case (Suc i') have "{#y \# mset (x # xs). y < (x # xs) ! i#} \# add_mset x {#y \# mset xs. y < xs ! i'#}" using Suc Cons.prems by (auto) also have "\ \# add_mset x (mset (take i' xs))" unfolding mset_subset_eq_add_mset_cancel using Cons.prems Suc by (intro Cons.IH) (auto) also have "\ = mset (take i (x # xs))" by (simp add: Suc) finally show ?thesis . qed qed auto lemma sorted_filter_greater_subset_drop: assumes "sorted xs" "i < length xs" shows "{# x \# mset xs. x > xs ! i #} \# mset (drop (Suc i) xs)" using assms proof (induction xs arbitrary: i rule: list.induct) case (Cons x xs i) show ?case proof (cases i) case 0 thus ?thesis by (auto simp: sorted_append filter_mset_is_empty_iff) next case (Suc i') have "{#y \# mset (x # xs). y > (x # xs) ! i#} \# {#y \# mset xs. y > xs ! i'#}" using Suc Cons.prems by (auto simp: set_conv_nth) also have "\ \# mset (drop (Suc i') xs)" using Cons.prems Suc by (intro Cons.IH) (auto) also have "\ = mset (drop (Suc i) (x # xs))" by (simp add: Suc) finally show ?thesis . qed qed auto subsection \The dual order type\ text \ The following type is a copy of a given ordered base type, but with the ordering reversed. This will be useful later because we can do some of our reasoning simply by symmetry. \ typedef 'a dual_ord = "UNIV :: 'a set" morphisms of_dual_ord to_dual_ord by auto setup_lifting type_definition_dual_ord instantiation dual_ord :: (ord) ord begin lift_definition less_eq_dual_ord :: "'a dual_ord \ 'a dual_ord \ bool" is "\a b :: 'a. a \ b" . lift_definition less_dual_ord :: "'a dual_ord \ 'a dual_ord \ bool" is "\a b :: 'a. a > b" . instance .. end instance dual_ord :: (preorder) preorder by standard (transfer; force simp: less_le_not_le intro: order_trans)+ instance dual_ord :: (linorder) linorder by standard (transfer; force simp: not_le)+ subsection \Chopping a list into equal-sized sublists\ (* TODO: Move to library? *) function chop :: "nat \ 'a list \ 'a list list" where "chop n [] = []" | "chop 0 xs = []" | "n > 0 \ xs \ [] \ chop n xs = take n xs # chop n (drop n xs)" by force+ termination by lexicographic_order context includes lifting_syntax begin lemma chop_transfer [transfer_rule]: "((=) ===> list_all2 R ===> list_all2 (list_all2 R)) chop chop" proof (intro rel_funI) fix m n ::nat and xs :: "'a list" and ys :: "'b list" assume "m = n" "list_all2 R xs ys" from this(2) have "list_all2 (list_all2 R) (chop n xs) (chop n ys)" proof (induction n xs arbitrary: ys rule: chop.induct) case (3 n xs ys) hence "ys \ []" by auto with 3 show ?case by auto qed auto with \m = n\ show "list_all2 (list_all2 R) (chop m xs) (chop n ys)" by simp qed end lemma chop_reduce: "chop n xs = (if n = 0 \ xs = [] then [] else take n xs # chop n (drop n xs))" by (cases "n = 0"; cases "xs = []") auto lemma concat_chop [simp]: "n > 0 \ concat (chop n xs) = xs" by (induction n xs rule: chop.induct) auto lemma chop_elem_not_Nil [simp,dest]: "ys \ set (chop n xs) \ ys \ []" by (induction n xs rule: chop.induct) (auto simp: eq_commute[of "[]"]) lemma chop_eq_Nil_iff [simp]: "chop n xs = [] \ n = 0 \ xs = []" by (induction n xs rule: chop.induct) auto lemma chop_ge_length_eq: "n > 0 \ xs \ [] \ n \ length xs \ chop n xs = [xs]" by simp lemma length_chop_part_le: "ys \ set (chop n xs) \ length ys \ n" by (induction n xs rule: chop.induct) auto lemma length_nth_chop: assumes "i < length (chop n xs)" shows "length (chop n xs ! i) = (if i = length (chop n xs) - 1 \ \n dvd length xs then length xs mod n else n)" proof (cases "n = 0") case False thus ?thesis using assms proof (induction n xs arbitrary: i rule: chop.induct) case (3 n xs i) show ?case proof (cases i) case 0 thus ?thesis using "3.prems" by (cases "length xs < n") (auto simp: le_Suc_eq dest: dvd_imp_le) next case [simp]: (Suc i') with "3.prems" have [simp]: "xs \ []" by auto with "3.prems" have *: "length xs > n" by (cases "length xs \ n") simp_all with "3.prems" have "chop n xs ! i = chop n (drop n xs) ! i'" by simp also have "length \ = (if i = length (chop n xs) - 1 \ \ n dvd (length xs - n) then (length xs - n) mod n else n)" by (subst "3.IH") (use Suc "3.prems" in auto) also have "n dvd (length xs - n) \ n dvd length xs" using * by (subst dvd_minus_self) auto also have "(length xs - n) mod n = length xs mod n" using * by (subst le_mod_geq [symmetric]) auto finally show ?thesis . qed qed auto qed (insert assms, auto) lemma length_chop: assumes "n > 0" shows "length (chop n xs) = nat \length xs / n\" using assms proof (induction n xs rule: chop.induct) case (3 n xs) show ?case proof (cases "length xs \ n") case False hence "\real (length xs) / real n\ = 1" using "3.hyps" by (intro ceiling_unique) auto with False show ?thesis using "3.prems" "3.hyps" by (auto simp: chop_ge_length_eq not_le) next case True hence "real (length xs) = real n + real (length (drop n xs))" by simp also have "\ / real n = real (length (drop n xs)) / real n + 1" using \n > 0\ by (simp add: divide_simps) also have "ceiling \ = ceiling (real (length (drop n xs)) / real n) + 1" by simp also have "nat \ = nat (ceiling (real (length (drop n xs)) / real n)) + nat 1" by (intro nat_add_distrib[OF order.trans[OF _ ceiling_mono[of 0]]]) auto also have "\ = length (chop n xs)" using \n > 0\ "3.hyps" by (subst "3.IH" [symmetric]) auto finally show ?thesis .. qed qed auto lemma sum_msets_chop: "n > 0 \ (\ys\chop n xs. mset ys) = mset xs" by (subst mset_concat [symmetric]) simp_all lemma UN_sets_chop: "n > 0 \ (\ys\set (chop n xs). set ys) = set xs" by (simp only: set_concat [symmetric] concat_chop) lemma in_set_chopD [dest]: assumes "x \ set ys" "ys \ set (chop d xs)" shows "x \ set xs" proof (cases "d > 0") case True thus ?thesis by (subst UN_sets_chop [symmetric]) (use assms in auto) qed (use assms in auto) subsection \$k$-th order statistics and medians\ text \ This returns the $k$-th smallest element of a list. This is also known as the $k$-th order statistic. \ definition select :: "nat \ 'a list \ ('a :: linorder)" where "select k xs = sort xs ! k" text \ The median of a list, where, for lists of even lengths, the smaller one is favoured: \ definition median where "median xs = select ((length xs - 1) div 2) xs" lemma select_in_set [intro,simp]: assumes "k < length xs" shows "select k xs \ set xs" proof - from assms have "sort xs ! k \ set (sort xs)" by (intro nth_mem) auto also have "set (sort xs) = set xs" by simp finally show ?thesis by (simp add: select_def) qed lemma median_in_set [intro, simp]: assumes "xs \ []" shows "median xs \ set xs" proof - from assms have "length xs > 0" by auto hence "(length xs - 1) div 2 < length xs" by linarith thus ?thesis by (simp add: median_def) qed text \ We show that selection and medians does not depend on the order of the elements: \ lemma sort_cong: "mset xs = mset ys \ sort xs = sort ys" by (rule properties_for_sort) simp_all lemma select_cong: "k = k' \ mset xs = mset xs' \ select k xs = select k' xs'" by (auto simp: select_def dest: sort_cong) lemma median_cong: "mset xs = mset xs' \ median xs = median xs'" unfolding median_def by (intro select_cong) (auto dest: mset_eq_length) text \ Selection distributes over appending lists under certain conditions: \ lemma sort_append: assumes "\x y. x \ set xs \ y \ set ys \ x \ y" shows "sort (xs @ ys) = sort xs @ sort ys" using assms by (intro properties_for_sort) (auto simp: sorted_append) lemma select_append: assumes "\y z. y \ set ys \ z \ set zs \ y \ z" shows "k < length ys \ select k (ys @ zs) = select k ys" "k \ {length ys.. select k (ys @ zs) = select (k - length ys) zs" using assms by (simp_all add: select_def sort_append nth_append) lemma select_append': assumes "\y z. y \ set ys \ z \ set zs \ y \ z" "k < length ys + length zs" shows "select k (ys @ zs) = (if k < length ys then select k ys else select (k - length ys) zs)" using assms by (auto intro!: select_append) text \ We can find simple upper bounds for the number of elements that are strictly less than (resp. greater than) the median of a list. \ lemma size_less_than_median: "size {#y \# mset xs. y < median xs#} \ (length xs - 1) div 2" proof (cases "xs = []") case False hence "length xs > 0" by simp hence "(length xs - 1) div 2 < length xs" by linarith hence "size {#y \# mset (sort xs). y < median xs#} \ size (mset (take ((length xs - 1) div 2) (sort xs)))" unfolding median_def select_def using False by (intro size_mset_mono sorted_filter_less_subset_take) auto thus ?thesis using False by simp qed auto lemma size_greater_than_median: "size {#y \# mset xs. y > median xs#} \ length xs div 2" proof (cases "xs = []") case False hence "length xs > 0" by simp hence "(length xs - 1) div 2 < length xs" by linarith hence "size {#y \# mset (sort xs). y > median xs#} \ size (mset (drop (Suc ((length xs - 1) div 2)) (sort xs)))" unfolding median_def select_def using False by (intro size_mset_mono sorted_filter_greater_subset_drop) auto hence "size (filter_mset (\y. y > median xs) (mset xs)) \ length xs - Suc ((length xs - 1) div 2)" by simp also have "\ = length xs div 2" by linarith finally show ?thesis . qed auto subsection \A more liberal notion of medians\ text \ We now define a more relaxed version of being ``a median'' as opposed to being ``\emph{the} median''. A value is a median if at most half the values in the list are strictly smaller than it and at most half are strictly greater. Note that, by this definition, the median does not even have to be in the list itself. \ definition is_median :: "'a :: linorder \ 'a list \ bool" where "is_median x xs \ length (filter (\y. y < x) xs) \ length xs div 2 \ length (filter (\y. y > x) xs) \ length xs div 2" text \ We set up some transfer rules for @{const is_median}. In particular, we have a rule that shows that something is a median for a list iff it is a median on that list w.\,r.\,t.\ the dual order, which will later allow us to argue by symmetry. \ context includes lifting_syntax begin lemma transfer_is_median [transfer_rule]: assumes [transfer_rule]: "(r ===> r ===> (=)) (<) (<)" shows "(r ===> list_all2 r ===> (=)) is_median is_median" unfolding is_median_def by transfer_prover lemma list_all2_eq_fun_conv_map: "list_all2 (\x y. x = f y) xs ys \ xs = map f ys" proof assume "list_all2 (\x y. x = f y) xs ys" thus "xs = map f ys" by induction auto next assume "xs = map f ys" moreover have "list_all2 (\x y. x = f y) (map f ys) ys" by (induction ys) auto ultimately show "list_all2 (\x y. x = f y) xs ys" by simp qed lemma transfer_is_median_dual_ord [transfer_rule]: "(pcr_dual_ord (=) ===> list_all2 (pcr_dual_ord (=)) ===> (=)) is_median is_median" by (auto simp: pcr_dual_ord_def cr_dual_ord_def OO_def rel_fun_def is_median_def list_all2_eq_fun_conv_map o_def less_dual_ord.rep_eq) end lemma is_median_to_dual_ord_iff [simp]: "is_median (to_dual_ord x) (map to_dual_ord xs) \ is_median x xs" unfolding is_median_def by transfer auto text \ The following is an obviously equivalent definition of @{const is_median} in terms of multisets that is occasionally nicer to use. \ lemma is_median_altdef: "is_median x xs \ size (filter_mset (\y. y < x) (mset xs)) \ length xs div 2 \ size (filter_mset (\y. y > x) (mset xs)) \ length xs div 2" proof - have *: "length (filter P xs) = size (filter_mset P (mset xs))" for P and xs :: "'a list" by (simp flip: mset_filter) show ?thesis by (simp only: is_median_def *) qed lemma is_median_cong: assumes "x = y" "mset xs = mset ys" shows "is_median x xs \ is_median y ys" unfolding is_median_altdef by (simp only: assms mset_eq_length[OF assms(2)]) text \ If an element is the median of a list of odd length, we can add any element to the list and the element is still a median. Conversely, if we want to compute a median of a list with even length $n$, we can simply drop one element and reduce the problem to a median of a list of size $n - 1$. \ lemma is_median_Cons_odd: assumes "is_median x xs" and "odd (length xs)" shows "is_median x (y # xs)" using assms by (auto simp: is_median_def) text \ And, of course, \emph{the} median is a median. \ lemma is_median_median [simp,intro]: "is_median (median xs) xs" using size_less_than_median[of xs] size_greater_than_median[of xs] unfolding is_median_def size_mset [symmetric] mset_filter by linarith+ subsection \Properties of a median-of-medians\ text \ We can now bound the number of list elements that can be strictly smaller than a median-of-medians of a chopped-up list (where each part has length $d$ except for the last one, which can also be shorter). The core argument is that at least roughly half of the medians of the sublists are greater or equal to the median-of-medians, and about $\frac{d}{2}$ elements in each such sublist are greater than or equal to their median and thereby also than the median-of-medians. \ lemma size_less_than_median_of_medians_strong: fixes xs :: "'a :: linorder list" and d :: nat assumes d: "d > 0" assumes median: "\xs. xs \ [] \ length xs \ d \ is_median (med xs) xs" assumes median': "is_median x (map med (chop d xs))" defines "m \ length (chop d xs)" shows "size {#y \# mset xs. y < x#} \ m * (d div 2) + m div 2 * ((d + 1) div 2)" proof - define n where [simp]: "n = length xs" \ \The medians of the sublists\ define M where "M = mset (map med (chop d xs))" define YS where "YS = mset (chop d xs)" \ \The sublists with a smaller median than the median-of-medians @{term x} and the rest.\ define YS1 where "YS1 = filter_mset (\ys. med ys < x) (mset (chop d xs))" define YS2 where "YS2 = filter_mset (\ys. \(med ys < x)) (mset (chop d xs))" \ \At most roughly half of the lists have a median that is smaller than @{term M}\ have "size YS1 = size (image_mset med YS1)" by simp also have "image_mset med YS1 = {#y \# mset (map med (chop d xs)). y < x#}" unfolding YS1_def by (subst filter_mset_image_mset [symmetric]) simp_all also have "size \ \ (length (map med (chop d xs))) div 2" using median' unfolding is_median_altdef by simp also have "\ = m div 2" by (simp add: m_def) finally have size_YS1: "size YS1 \ m div 2" . have "m = size (mset (chop d xs))" by (simp add: m_def) also have "mset (chop d xs) = YS1 + YS2" unfolding YS1_def YS2_def by (rule multiset_partition) finally have m_eq: "m = size YS1 + size YS2" by simp \ \We estimate the number of elements less than @{term x} by grouping them into elements coming from @{term YS1} and elements coming from @{term YS2}. In the first case, we just note that no more than @{term d} elements can come from each sublist, whereas in the second case, we make the analysis more precise and note that only elements that are less than the median of their sublist can be less than @{term x}.\ have "{# y \# mset xs. y < x#} = {# y \# (\ys\chop d xs. mset ys). y < x#}" using d by (subst sum_msets_chop) simp_all also have "\ = (\ys\chop d xs. {#y \# mset ys. y < x#})" by (subst filter_mset_sum_list) (simp add: o_def) also have "\ = (\ys\#YS. {#y \# mset ys. y < x#})" unfolding YS_def by (subst sum_mset_sum_list [symmetric]) simp_all also have "YS = YS1 + YS2" by (simp add: YS_def YS1_def YS2_def not_le) also have "(\ys\#\. {#y \# mset ys. y < x#}) = (\ys\#YS1. {#y \# mset ys. y < x#}) + (\ys\#YS2. {#y \# mset ys. y < x#})" by simp also have "\ \# (\ys\#YS1. mset ys) + (\ys\#YS2. {#y \# mset ys. y < med ys#})" by (intro subset_mset.add_mono sum_mset_mset_mono mset_filter_mono) (auto simp: YS2_def) finally have "{# y \# mset xs. y < x #} \# \" . hence "size {# y \# mset xs. y < x #} \ size \" by (rule size_mset_mono) \ \We do some further straightforward estimations and arrive at our goal.\ also have "\ = (\ys\#YS1. length ys) + (\x\#YS2. size {#y \# mset x. y < med x#})" by (simp add: size_mset_sum_mset_distrib multiset.map_comp o_def) also have "(\ys\#YS1. length ys) \ (\ys\#YS1. d)" by (intro sum_mset_mono) (auto simp: YS1_def length_chop_part_le) also have "\ = size YS1 * d" by simp also have d: "d = (d div 2) + ((d + 1) div 2)" using d by linarith have "size YS1 * d = size YS1 * (d div 2) + size YS1 * ((d + 1) div 2)" by (subst d) (simp add: algebra_simps) also have "(\ys\#YS2. size {#y \# mset ys. y < med ys#}) \ (\ys\#YS2. length ys div 2)" proof (intro sum_mset_mono size_less_than_median, goal_cases) case (1 ys) hence "ys \ []" "length ys \ d" by (auto simp: YS2_def length_chop_part_le) from median[OF this] show ?case by (auto simp: is_median_altdef) qed also have "\ \ (\ys\#YS2. d div 2)" by (intro sum_mset_mono div_le_mono diff_le_mono) (auto simp: YS2_def dest: length_chop_part_le) also have "\ = size YS2 * (d div 2)" by simp also have "size YS1 * (d div 2) + size YS1 * ((d + 1) div 2) + \ = m * (d div 2) + size YS1 * ((d + 1) div 2)" by (simp add: m_eq algebra_simps) also have "size YS1 * ((d + 1) div 2) \ (m div 2) * ((d + 1) div 2)" by (intro mult_right_mono size_YS1) auto finally show "size {#y \# mset xs. y < x#} \ m * (d div 2) + m div 2 * ((d + 1) div 2)" by simp_all qed text \ We now focus on the case of an odd chopping size and make some further estimations to simplify the above result a little bit. \ theorem size_less_than_median_of_medians: fixes xs :: "'a :: linorder list" and d :: nat assumes median: "\xs. xs \ [] \ length xs \ Suc (2 * d) \ is_median (med xs) xs" assumes median': "is_median x (map med (chop (Suc (2*d)) xs))" defines "n \ length xs" defines "c \ (3 * real d + 1) / (2 * (2 * d + 1))" shows "size {#y \# mset xs. y < x#} \ nat \c * n\ + (5 * d) div 2 + 1" proof (cases "xs = []") case False define m where "m = length (chop (Suc (2*d)) xs)" have "real (m div 2) \ real (nat \real n / (1 + 2 * real d)\) / 2" by (simp add: m_def length_chop n_def) also have "real (nat \real n / (1 + 2 * real d)\) = of_int \real n / (1 + 2 * real d)\" by (intro of_nat_nat) (auto simp: divide_simps) also have "\ / 2 \ (real n / (1 + 2 * real d) + 1) / 2" by (intro divide_right_mono) linarith+ also have "\ = n / (2 * (2 * real d + 1)) + 1 / 2" by (simp add: field_simps) finally have m: "real (m div 2) \ \" . have "size {#y \# mset xs. y < x#} \ d * m + Suc d * (m div 2)" using size_less_than_median_of_medians_strong[of "Suc (2 * d)" med x xs] assms unfolding m_def by (simp add: algebra_simps) also have "\ \ d * (2 * (m div 2) + 1) + Suc d * (m div 2)" by (intro add_mono mult_left_mono) linarith+ also have "\ = (3 * d + 1) * (m div 2) + d" by (simp add: algebra_simps) finally have "real (size {#y \# mset xs. y < x#}) \ real \" by (subst of_nat_le_iff) also have "\ \ (3 * real d + 1) * (n / (2 * (2 * d + 1)) + 1/2) + real d" unfolding of_nat_add of_nat_mult of_nat_1 of_nat_numeral by (intro add_mono mult_mono order.refl m) (auto simp: m_def length_chop n_def add_ac) also have "\ = c * real n + (5 * real d + 1) / 2" by (simp add: field_simps c_def) also have "\ \ real (nat \c * n\ + ((5 * d) div 2 + 1))" unfolding of_nat_add by (intro add_mono) (linarith, simp add: field_simps) finally show ?thesis by (subst (asm) of_nat_le_iff) (simp_all add: add_ac) qed auto text \ We get the analogous result for the number of elements that are greater than a median-of-medians by looking at the dual order and using the \emph{transfer} method. \ theorem size_greater_than_median_of_medians: fixes xs :: "'a :: linorder list" and d :: nat assumes median: "\xs. xs \ [] \ length xs \ Suc (2 * d) \ is_median (med xs) xs" assumes median': "is_median x (map med (chop (Suc (2*d)) xs))" defines "n \ length xs" defines "c \ (3 * real d + 1) / (2 * (2 * d + 1))" shows "size {#y \# mset xs. y > x#} \ nat \c * n\ + (5 * d) div 2 + 1" proof - include lifting_syntax define med' where "med' = (\xs. to_dual_ord (med (map of_dual_ord xs)))" have "xs = map of_dual_ord ys" if "list_all2 cr_dual_ord xs ys" for xs :: "'a list" and ys using that by induction (auto simp: cr_dual_ord_def) hence [transfer_rule]: "(list_all2 (pcr_dual_ord (=)) ===> pcr_dual_ord (=)) med med'" by (auto simp: rel_fun_def pcr_dual_ord_def OO_def med'_def cr_dual_ord_def dual_ord.to_dual_ord_inverse) have "size {#y \# mset xs. y > x#} = length (filter (\y. y > x) xs)" by (subst size_mset [symmetric]) (simp only: mset_filter) also have "\ = length (map to_dual_ord (filter (\y. y > x) xs))" by simp also have "(\y. y > x) = (\y. to_dual_ord y < to_dual_ord x)" by transfer simp_all hence "length (map to_dual_ord (filter (\y. y > x) xs)) = length (map to_dual_ord (filter \ xs))" by simp also have "\ = length (filter (\y. y < to_dual_ord x) (map to_dual_ord xs))" unfolding filter_map o_def by simp also have "\ = size {#y \# mset (map to_dual_ord xs). y < to_dual_ord x#}" by (subst size_mset [symmetric]) (simp only: mset_filter) also have "\ \ nat \(3 * real d + 1) / real (2 * (2 * d + 1)) * length (map to_dual_ord xs)\ + 5 * d div 2 + 1" proof (intro size_less_than_median_of_medians) fix xs :: "'a dual_ord list" assume xs: "xs \ []" "length xs \ Suc (2 * d)" from xs show "is_median (med' xs) xs" by (transfer fixing: d) (rule median) next show "is_median (to_dual_ord x) (map med' (chop (Suc (2 * d)) (map to_dual_ord xs)))" by (transfer fixing: d x xs) (use median' in simp_all) qed finally show ?thesis by (simp add: n_def c_def) qed text \ The most important case is that of chopping size 5, since that is the most practical one for the median-of-medians selection algorithm. For it, we obtain the following nice and simple bounds: \ corollary size_less_greater_median_of_medians_5: fixes xs :: "'a :: linorder list" assumes "\xs. xs \ [] \ length xs \ 5 \ is_median (med xs) xs" assumes "is_median x (map med (chop 5 xs))" shows "length (filter (\y. y < x) xs) \ nat \0.7 * length xs\ + 6" and "length (filter (\y. y > x) xs) \ nat \0.7 * length xs\ + 6" using size_less_than_median_of_medians[of 2 med x xs] size_greater_than_median_of_medians[of 2 med x xs] assms by (simp_all add: size_mset [symmetric] mset_filter mult_ac add_ac del: size_mset) subsection \The recursive step\ text \ We now turn to the actual selection algorithm itself. The following simple reduction lemma illustrates the idea of the algorithm quite well already, but it has the disadvantage that, if one were to use it as a recursive algorithm, it would only work for lists with distinct elements. If the list contains repeated elements, this may not even terminate. The basic idea is that we choose some pivot element, partition the list into elements that are bigger than the pivot and those that are not, and then recurse into one of these (hopefully smaller) lists. \ theorem select_rec_partition: assumes "d > 0" "k < length xs" shows "select k xs = ( let (ys, zs) = partition (\y. y \ x) xs in if k < length ys then select k ys else select (k - length ys) zs )" (is "_ = ?rhs") proof - define ys zs where "ys = filter (\y. y \ x) xs" and "zs = filter (\y. \(y \ x)) xs" have "select k xs = select k (ys @ zs)" by (intro select_cong) (simp_all add: ys_def zs_def) also have "\ = (if k < length ys then select k ys else select (k - length ys) zs)" using assms(2) by (intro select_append') (auto simp: ys_def zs_def sum_length_filter_compl) finally show ?thesis by (simp add: ys_def zs_def Let_def o_def) qed text \ The following variant uses a three-way partitioning function instead. This way, the size of the list in the final recursive call decreases by a factor of at least $\frac{3d'+1}{2(2d'+1)}$ by the previous estimates, given that the chopping size is $d = 2d'+1$. For a chopping size of 5, we get a factor of $0.7$. \ definition threeway_partition :: "'a \ 'a :: linorder list \ 'a list \ 'a list \ 'a list" where "threeway_partition x xs = (filter (\y. y < x) xs, filter (\y. y = x) xs, filter (\y. y > x) xs)" lemma threeway_partition_code [code]: "threeway_partition x [] = ([], [], [])" "threeway_partition x (y # ys) = (case threeway_partition x ys of (ls, es, gs) \ if y < x then (y # ls, es, gs) else if x = y then (ls, y # es, gs) else (ls, es, y # gs))" by (auto simp: threeway_partition_def) theorem select_rec_threeway_partition: assumes "d > 0" "k < length xs" shows "select k xs = ( let (ls, es, gs) = threeway_partition x xs; nl = length ls; ne = length es in if k < nl then select k ls else if k < nl + ne then x else select (k - nl - ne) gs )" (is "_ = ?rhs") proof - define ls es gs where "ls = filter (\y. y < x) xs" and "es = filter (\y. y = x) xs" and "gs = filter (\y. y > x) xs" define nl ne where [simp]: "nl = length ls" "ne = length es" have mset_eq: "mset xs = mset ls + mset es + mset gs" unfolding ls_def es_def gs_def by (induction xs) auto have length_eq: "length xs = length ls + length es + length gs" unfolding ls_def es_def gs_def - by (induction xs) auto + by (induction xs) (auto simp del: filter_True) + have [simp]: "select i es = x" if "i < length es" for i proof - have "select i es \ set (sort es)" unfolding select_def using that by (intro nth_mem) auto hence "select i es \ set es" using that by (auto simp: select_def) also have "set es \ {x}" unfolding es_def by (induction es) auto finally show ?thesis by simp qed have "select k xs = select k (ls @ (es @ gs))" by (intro select_cong) (simp_all add: mset_eq) also have "\ = (if k < nl then select k ls else select (k - nl) (es @ gs))" unfolding nl_ne_def using assms by (intro select_append') (auto simp: ls_def es_def gs_def length_eq) also have "\ = (if k < nl then select k ls else if k < nl + ne then x else select (k - nl - ne) gs)" (is "?lhs' = ?rhs'") proof (cases "k < nl") case False hence "?lhs' = select (k - nl) (es @ gs)" by simp also have "\ = (if k - nl < ne then select (k - nl) es else select (k - nl - ne) gs)" unfolding nl_ne_def using assms False by (intro select_append') (auto simp: ls_def es_def gs_def length_eq) also have "\ = (if k - nl < ne then x else select (k - nl - ne) gs)" by simp also from False have "\ = ?rhs'" by auto finally show ?thesis . qed simp_all also have "\ = ?rhs" by (simp add: threeway_partition_def Let_def ls_def es_def gs_def) finally show ?thesis . qed text \ By the above results, it can be seen quite easily that, in each recursive step, the algorithm takes a list of length $n$, does $O(n)$ work for the chopping, computing the medians of the sublists, and partitioning, and it calls itself recursively with lists of size at most $\lceil 0.2n\rceil$ and $\lceil 0.7n\rceil + 6$, respectively. This means that the runtime of the algorithm is bounded above by the Akra--Bazzi-style recurrence \[T(n) = T(\lceil 0.2n\rceil) + T(\lceil 0.7n\rceil + 6) + O(n)\] which, by the Akra--Bazzi theorem, can be shown to fulfil $T\in \Theta(n)$. However, a proper analysis of this would require an actual execution model and some way of measuring the runtime of the algorithm, which is not what we aim to do here. Additionally, the entire algorithm can be performed in-place in an imperative way, but this because quite tedious. Instead of this, we will now focus on developing the above recursion into an executable functional algorithm. \ subsection \Medians of lists of length at most 5\ text \ We now show some basic results about how to efficiently find a median of a list of size at most 5. For length 1 or 2, this is trivial, since we can just pick any element. For length 3 and 4, we need at most three comparisons. For length 5, we need at most six comparisons. This allows us to save some comparisons compared with the naive method of performing insertion sort and then returning the element in the middle. \ definition median_3 :: "'a :: linorder \ _" where "median_3 a b c = (if a \ b then if b \ c then b else max a c else if c \ b then b else min a c)" lemma median_3: "median_3 a b c = median [a, b, c]" by (auto simp: median_3_def median_def select_def min_def max_def) definition median_5_aux :: "'a :: linorder \ _" where "median_5_aux x1 x2 x3 x4 x5 = ( if x2 \ x3 then if x2 \ x4 then min x3 x4 else min x2 x5 else if x4 \ x3 then min x3 x5 else min x2 x4)" lemma median_5_aux: assumes "x1 \ x2" "x4 \ x5" "x1 \ x4" shows "median_5_aux x1 x2 x3 x4 x5 = median [x1,x2,x3,x4,x5]" using assms by (auto simp: median_5_aux_def median_def select_def min_def) definition median_5 :: "'a :: linorder \ _" where "median_5 a b c d e = ( let (x1, x2) = (if a \ b then (a, b) else (b, a)); (x4, x5) = (if d \ e then (d, e) else (e, d)) in if x1 \ x4 then median_5_aux x1 x2 c x4 x5 else median_5_aux x4 x5 c x1 x2)" lemma median_5: "median_5 a b c d e = median [a, b, c, d, e]" by (auto simp: median_5_def Let_def median_5_aux intro: median_cong) fun median_le_5 where "median_le_5 [a] = a" | "median_le_5 [a,b] = a" | "median_le_5 [a,b,c] = median_3 a b c" | "median_le_5 [a,b,c,d] = median_3 a b c" | "median_le_5 [a,b,c,d,e] = median_5 a b c d e" | "median_le_5 _ = undefined" lemma median_5_in_set: "median_5 a b c d e \ {a, b, c, d, e}" proof - have "median_5 a b c d e \ set [a, b, c, d, e]" unfolding median_5 by (rule median_in_set) auto thus ?thesis by simp qed lemma median_le_5_in_set: assumes "xs \ []" "length xs \ 5" shows "median_le_5 xs \ set xs" proof (cases xs rule: median_le_5.cases) case (5 a b c d e) with median_5_in_set[of a b c d e] show ?thesis by simp qed (insert assms, auto simp: median_3_def min_def max_def) lemma median_le_5: assumes "xs \ []" "length xs \ 5" shows "is_median (median_le_5 xs) xs" proof (cases xs rule: median_le_5.cases) case (3 a b c) have "is_median (median xs) xs" by simp also have "median xs = median_3 a b c" by (simp add: median_3 3) finally show ?thesis using 3 by simp next case (4 a b c d) have "is_median (median [a,b,c]) [a,b,c]" by simp also have "median [a,b,c] = median_3 a b c" by (simp add: median_3 4) finally have "is_median (median_3 a b c) (d # [a,b,c])" by (rule is_median_Cons_odd) auto also have "?this \ is_median (median_3 a b c) [a,b,c,d]" by (intro is_median_cong) auto finally show ?thesis using 4 by simp next case (5 a b c d e) have "is_median (median xs) xs" by simp also have "median xs = median_5 a b c d e" by (simp add: median_5 5) finally show ?thesis using 5 by simp qed (insert assms, auto simp: is_median_def) subsection \Median-of-medians selection algorithm\ text \ The fast selection function now simply computes the median-of-medians of the chopped-up list as a pivot, partitions the list into with respect to that pivot, and recurses into one of the resulting sublists. \ function fast_select where "fast_select k xs = ( if length xs \ 20 then sort xs ! k else let x = fast_select (((length xs + 4) div 5 - 1) div 2) (map median_le_5 (chop 5 xs)); (ls, es, gs) = threeway_partition x xs in if k < length ls then fast_select k ls else if k < length ls + length es then x else fast_select (k - length ls - length es) gs )" by auto text \ The correctness of this is obvious from the above theorems, but the proof is still somewhat complicated by the fact that termination depends on the correctness of the function. \ lemma fast_select_correct_aux: assumes "fast_select_dom (k, xs)" "k < length xs" shows "fast_select k xs = select k xs" using assms proof induction case (1 k xs) show ?case proof (cases "length xs \ 20") case True thus ?thesis using "1.prems" "1.hyps" by (subst fast_select.psimps) (auto simp: select_def) next case False define x where "x = fast_select (((length xs + 4) div 5 - Suc 0) div 2) (map median_le_5 (chop 5 xs))" define ls where "ls = filter (\y. y < x) xs" define es where "es = filter (\y. y = x) xs" define gs where "gs = filter (\y. y > x) xs" define nl ne where "nl = length ls" and "ne = length es" note defs = nl_def ne_def x_def ls_def es_def gs_def have tw: "(ls, es, gs) = threeway_partition (fast_select (((length xs + 4) div 5 - 1) div 2) (map median_le_5 (chop 5 xs))) xs" unfolding threeway_partition_def defs One_nat_def .. have tw': "(ls, es, gs) = threeway_partition x xs" by (simp add: tw x_def) have "fast_select k xs = (if k < nl then fast_select k ls else if k < nl + ne then x else fast_select (k - nl - ne) gs)" using "1.hyps" False by (subst fast_select.psimps) (simp_all add: threeway_partition_def defs [symmetric]) also have "\ = (if k < nl then select k ls else if k < nl + ne then x else select (k - nl - ne) gs)" proof (intro if_cong refl) assume *: "k < nl" show "fast_select k ls = select k ls" by (rule 1; (rule refl tw)?) (insert *, auto simp: False threeway_partition_def ls_def x_def nl_def)+ next assume *: "\k < nl" "\k < nl + ne" have **: "length xs = length ls + length es + length gs" - unfolding ls_def es_def gs_def by (induction xs) auto + unfolding ls_def es_def gs_def by (induction xs) (auto simp del: filter_True) show "fast_select (k - nl - ne) gs = select (k - nl - ne) gs" unfolding nl_def ne_def by (rule 1; (rule refl tw)?) (insert False * ** \k < length xs\, auto simp: nl_def ne_def) qed also have "\ = select k xs" using \k < length xs\ by (subst (3) select_rec_threeway_partition[of "5::nat" _ _ x]) (unfold Let_def nl_def ne_def ls_def gs_def es_def x_def threeway_partition_def, simp_all) finally show ?thesis . qed qed text \ Termination of the algorithm is reasonably obvious because the lists that are recursed into never contain the pivot (the median-of-medians), while the original list clearly does. The proof is still somewhat technical though. \ lemma fast_select_termination: "All fast_select_dom" proof (relation "measure (length \ snd)"; (safe)?, goal_cases) case (1 k xs) thus ?case by (auto simp: length_chop nat_less_iff ceiling_less_iff) next fix k :: nat and xs ls es gs :: "'a list" define x where "x = fast_select (((length xs + 4) div 5 - 1) div 2) (map median_le_5 (chop 5 xs))" assume A: "\ length xs \ 20" "(ls, es, gs) = threeway_partition x xs" "fast_select_dom (((length xs + 4) div 5 - 1) div 2, map median_le_5 (chop 5 xs))" from A have eq: "ls = filter (\y. y < x) xs" "gs = filter (\y. y > x) xs" by (simp_all add: x_def threeway_partition_def) have len: "(length xs + 4) div 5 = nat \length xs / 5\" by linarith have less: "(nat \real (length xs) / 5\ - Suc 0) div 2 < nat \real (length xs) / 5\" using A(1) by linarith have "x = select (((length xs + 4) div 5 - 1) div 2) (map median_le_5 (chop 5 xs))" using less unfolding x_def by (intro fast_select_correct_aux A) (auto simp: length_chop len) also have "\ = median (map median_le_5 (chop 5 xs))" by (simp add: median_def len length_chop) finally have x: "x = \" . moreover { have "x \ set (map median_le_5 (chop 5 xs))" using A(1) unfolding x by (intro median_in_set) auto also have "\ \ (\ys\set (chop 5 xs). {median_le_5 ys})" by auto also have "\ \ (\ys\set (chop 5 xs). set ys)" using A(1) by (intro UN_mono) (auto simp: median_le_5_in_set length_chop_part_le) also have "\ = set xs" by (subst UN_sets_chop) auto finally have "x \ set xs" . } ultimately show "((k, ls), k, xs) \ measure (length \ snd)" and "((k - length ls - length es, gs), k, xs) \ measure (length \ snd)" using A(1) by (auto simp: eq intro!: length_filter_less[of x]) qed text \ We now have all the ingredients to show that @{const fast_select} terminates and does, indeed, compute the $k$-th order statistic. \ termination fast_select by (rule fast_select_termination) theorem fast_select_correct: "k < length xs \ fast_select k xs = select k xs" using fast_select_termination by (intro fast_select_correct_aux) auto text \ The following version is then suitable for code export. \ lemma fast_select_code [code]: "fast_select k xs = ( if length xs \ 20 then fold insort xs [] ! k else let x = fast_select (((length xs + 4) div 5 - 1) div 2) (map median_le_5 (chop 5 xs)); (ls, es, gs) = threeway_partition x xs; nl = length ls; ne = nl + length es in if k < nl then fast_select k ls else if k < ne then x else fast_select (k - ne) gs )" by (subst fast_select.simps) (simp_all only: Let_def algebra_simps sort_conv_fold) lemma select_code [code]: "select k xs = (if k < length xs then fast_select k xs else Code.abort (STR ''Selection index out of bounds.'') (\_. select k xs))" proof (cases "k < length xs") case True thus ?thesis by (simp only: if_True fast_select_correct) qed (simp_all only: Code.abort_def if_False) end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Concrete_Reachability_Analysis_C1.thy b/thys/Ordinary_Differential_Equations/Numerics/Concrete_Reachability_Analysis_C1.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Concrete_Reachability_Analysis_C1.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Concrete_Reachability_Analysis_C1.thy @@ -1,1716 +1,1714 @@ theory Concrete_Reachability_Analysis_C1 imports Concrete_Reachability_Analysis Abstract_Reachability_Analysis_C1 begin definition "op_card_vec TYPE('a) = CARD('a)" lemma op_card_vec_pat_def[autoref_op_pat_def]: "CARD('a) \ OP (op_card_vec TYPE('a))" by (auto simp: op_card_vec_def) lemma op_card_vec_impl[autoref_rules]: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::enum rvec) E" shows "(E, op_card_vec TYPE('a)) \ nat_rel" using assms by (auto simp: op_card_vec_def) context approximate_sets begin sublocale approximate_sets_ode' where ode_ops = ode_ops for ode_ops ..\ \parametrized by \ode_ops\\ end context approximate_sets begin lemma nonneg_reals_autoref[autoref_rules]: "(None, nonneg_reals) \ \Id\phantom_rel" and pos_reals_autoref[autoref_rules]: "(None, pos_reals) \ \Id\phantom_rel" by (auto simp: phantom_rel_def) lemma appr1_relI: assumes "c1_info_invar DIM('n::executable_euclidean_space) X0i" shows "(X0i, (c1_info_of_appr X0i::'n c1_info set)) \ appr1_rel" using assms apply (cases "snd X0i") subgoal apply (simp add: c1_info_of_appr_def c1_info_invar_def) unfolding appr1_rel_internal apply (rule UnI1) apply auto apply (rule exI[where x="fst X0i"]) apply safe subgoal by (auto simp: prod_eq_iff) subgoal apply (rule exI[where x="eucl_of_list ` set_of_appr (fst X0i)"]) apply (auto simp: appr_rel_def) by (auto simp: appr_rell_internal lv_rel_def set_rel_br br_chain length_set_of_appr intro!: brI) done subgoal for D apply (simp add: c1_info_of_appr_def c1_info_invar_def) unfolding appr1_rel_internal apply (rule UnI2) apply (auto simp: set_rel_br) apply (rule exI[where x="fst X0i"]) apply (rule exI[where x=D]) apply safe subgoal by (auto simp: prod_eq_iff) subgoal by (auto simp: appr_rell_internal lv_rel_def set_rel_br br_chain length_set_of_appr intro!: brI) (auto simp: power2_eq_square) done done lemma appr1_rel_br: "appr1_rel = br (c1_info_of_appr::_\('n c1_info)set) (c1_info_invar DIM('n::executable_euclidean_space))" apply (auto simp: dest!: brD intro!: appr1_relI) apply (rule brI) subgoal by (auto simp: appr1_rel_internal c1_info_of_appr_def appr_rel_br set_rel_br dest!: brD) subgoal by (auto simp: c1_info_invar_def appr1_rel_internal appr_rel_br power2_eq_square dest!: brD) done lemma appr1_rel_aux: "{((xs, Some ys), X) |xs ys X. (xs @ ys, X) \ appr_rel \ length ys = (length xs)\<^sup>2} O \br flow1_of_vec1 top\set_rel = {((xs, Some ys), X::'n eucl1 set) |xs ys X. X = (\xs. flow1_of_vec1 (eucl_of_list xs)) ` set_of_appr (xs @ ys) \ length xs = DIM((real, 'n::enum) vec) \ length ys = DIM((real, 'n) vec) * DIM((real, 'n) vec)}" apply (auto simp: set_rel_br appr_rel_br power2_eq_square dest!: brD) apply (rule relcompI) apply simp apply (rule brI) apply (rule refl) apply simp apply (rule brI) defer apply simp apply auto done lemma flow1_of_list_def': shows "flow1_of_list xs = flow1_of_vec1 (eucl_of_list xs)" by (auto simp: flow1_of_list_def flow1_of_vec1_def eucl_of_list_prod blinfun_of_list_eq_blinfun_of_vmatrix) lemma appr1_rel_def: "appr1_rel = {((xs, None ), X \ UNIV)| xs X. (xs, X) \ appr_rel} \ {((xs, Some ys), X)| xs ys X. (xs @ ys, X) \ appr_rel \ length ys = (length xs)\<^sup>2} O \br flow1_of_vec1 top\set_rel" unfolding appr1_rel_internal flow1_of_list_def'[abs_def] appr1_rel_aux .. lemmas [autoref_rel_intf] = REL_INTFI[of appr1_rel i_appr1] lemma [autoref_op_pat]: "(`) flow1_of_vec1 \ OP op_image_flow1_of_vec1" by auto lemma op_image_flow1_of_vec1[autoref_rules]: assumes "DIM_precond TYPE('a rvec) E" shows "(\xs. (take E xs, Some (drop E xs)), op_image_flow1_of_vec1::('a::enum) vec1 set\_) \ appr_rel \ appr1_rel" using assms apply (auto simp: appr1_rel_def set_rel_br flow1_of_vec1_def[abs_def] intro!: brI elim!: notE split: option.splits list.splits) apply (rule relcompI[OF _ brI[OF refl]]) apply (auto simp: power2_eq_square min_def appr_rel_br br_def) done lemma index_autoref[autoref_rules]: "(index, index) \ \lv_rel\list_rel \ lv_rel \ nat_rel" unfolding index_def[abs_def] find_index_def apply parametricity apply (auto simp: lv_rel_def br_def list_rel_def) using list_of_eucl_eucl_of_list by force lemma [autoref_op_pat]: "(`) fst \ OP op_image_fst" by auto lemma op_image_fst_flow1[autoref_rules]: shows "(\x. fst x, op_image_fst::_\'n::executable_euclidean_space set) \ appr1_rel \ appr_rel" apply (auto simp: appr1_rel_internal flow1_of_list_def set_rel_br image_image power2_eq_square dest!: brD) apply (auto simp: br_def appr_rel_br length_set_of_appr image_image eucl_of_list_prod dest!: set_of_appr_takeD) subgoal for xs ys a apply (rule image_eqI[where x="take DIM('n) a"]) by (auto intro!: take_set_of_apprI dest: length_set_of_appr) subgoal for xs ys a apply (frule set_of_appr_ex_append2[where b=ys]) apply auto subgoal for r apply (rule image_eqI[where x="a @ r"]) apply (auto simp: length_set_of_appr ) apply (rule eucl_of_list_eqI) by (auto dest!: length_set_of_appr) done done lemma op_image_fste_impl[autoref_rules]: "((\(_, x, _). x), op_image_fste) \ appr1e_rel \ appr_rel" by (auto simp: image_image split_beta' scaleR2_rel_def dest!: op_image_fst_flow1[param_fo] brD) lemma DIM_precond_vec1I[autoref_rules_raw]: assumes "DIM_precond TYPE('n::enum rvec) E" shows "DIM_precond TYPE('n::enum vec1) (E + E*E)" using assms by (auto simp: ) lemma vec1rep_impl[autoref_rules]: "(\(a, bs). RETURN (map_option ((@) a) bs), vec1rep) \ appr1_rel \ \\appr_rel\option_rel\nres_rel" apply (auto simp: vec1rep_def appr1_rel_def set_rel_br appr_rel_def power2_eq_square nres_rel_def dest!: brD intro!: RETURN_SPEC_refine) subgoal for xs ys a b apply (rule exI[where x="Some (eucl_of_list ` set_of_appr (xs @ ys))"]) apply (auto simp: appr_rell_internal image_image lv_rel_def set_rel_br length_set_of_appr intro!: brI dest!: brD) done done lemma [autoref_op_pat]: "X \ UNIV \ OP op_times_UNIV $ X" by simp lemma op_times_UNIV_impl[autoref_rules]: "(\x. (x, None), op_times_UNIV) \ appr_rel \ appr1_rel" by (auto simp: appr1_rel_internal) schematic_goal solve_poincare_plane_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(ni, n) \ lv_rel" and CX[autoref_rules]: "(CXi, CX) \ appr1_rel" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" shows "(nres_of (?R), solve_poincare_plane odo n (CX::'n eucl1 set)) \ \appr1_rel\nres_rel" unfolding autoref_tag_defs unfolding solve_poincare_plane_def including art by autoref_monadic concrete_definition solve_poincare_plane_impl for ni CXi uses solve_poincare_plane_impl lemmas solve_poincare_plane_impl_refine[autoref_rules] = solve_poincare_plane_impl.refine[autoref_higher_order_rule (1)] sublocale autoref_op_pat_def solve_poincare_plane . lemma [autoref_rules_raw]: assumes "DIM_precond TYPE((real, 'n::enum) vec) K" shows "DIM_precond TYPE(((real, 'n) vec, 'n) vec) (K * K)" using assms by auto lemma embed1_impl[autoref_rules]: assumes "DIM_precond TYPE((real, 'n::enum) vec) E" shows "((\x. x @ replicate (E * E) 0), embed1::'n rvec\_) \ lv_rel \ lv_rel" using assms by (auto simp: lv_rel_def br_def eucl_of_list_prod) definition "var_ode_ops_impl = (\(ode_ops, _, _, vode_slps). (var_ode_ops ode_ops, (case vode_slps of None => let _ = print_msg_impl ''ODE solver not properly initialized: vode_slps missing!'' in ode_slps_of (approximate_sets_ode'.var_ode_ops ode_ops) | Some (vode_slps) => vode_slps), None, None))" lemma var_ode_ops[autoref_rules]: "(var_ode_ops_impl, var_ode_ops) \ ode_ops_rel \ ode_ops_rel" by (auto simp: var_ode_ops_impl_def ode_ops_rel_def init_ode_ops_def split: option.splits) schematic_goal choose_step1_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" "ncc_precond TYPE('n vec1)" "ncc_precond TYPE('n rvec)" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ appr1_rel" "(hi, h) \ rnv_rel" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" notes [autoref_post_simps] = fst_conv shows "(nres_of ?f, choose_step1 odo X h) \ \rnv_rel \\<^sub>r appr1_rel \\<^sub>r appr1_rel \\<^sub>r appr1_rel\nres_rel" unfolding choose_step1_def including art by autoref_monadic concrete_definition choose_step1_impl for Xi hi uses choose_step1_impl lemmas [autoref_rules] = choose_step1_impl.refine[autoref_higher_order_rule (1 2 3)] sublocale autoref_op_pat_def choose_step1 . lemma op_image_zerofst_impl[autoref_rules]: "(\(_, x). (appr_of_ivl ops (replicate E 0) (replicate E 0), x), op_image_zerofst::'n c1_info set \ _) \ appr1_rel \ appr1_rel" if "DIM_precond (TYPE('n::executable_euclidean_space)) E" using that apply (auto simp: appr1_rel_br dest!: brD intro!: brI) subgoal by (force simp: c1_info_of_appr_def image_image flow1_of_list_def set_of_appr_of_ivl_point_append eucl_of_list_prod c1_info_invar_def length_set_of_appr split: option.splits elim!: mem_set_of_appr_appendE cong del: image_cong_simp) subgoal for a b c d apply (auto simp: c1_info_of_appr_def split: option.splits) subgoal using set_of_appr_nonempty[of a] by (force simp del: set_of_appr_nonempty) subgoal supply [simp del] = eucl_of_list_take_DIM apply (auto simp: image_image set_of_appr_of_ivl_point_append flow1_of_list_def) apply (frule set_of_appr_ex_append1[where b=a]) apply auto apply (rule image_eqI) prefer 2 apply assumption by (auto simp: eucl_of_list_prod c1_info_invar_def dest!: length_set_of_appr) done subgoal by (auto simp: c1_info_of_appr_def flow1_of_vec1_def image_image set_of_appr_of_ivl_point_append eucl_of_list_prod c1_info_invar_def length_set_of_appr split: option.splits elim!: mem_set_of_appr_appendE) done sublocale autoref_op_pat_def op_image_zerofst . lemma op_image_zerofst_vec_impl[autoref_rules]: "(\x. (appr_of_ivl ops (replicate E 0) (replicate E 0) @ drop E x), op_image_zerofst_vec::'n vec1 set \ _) \ appr_rel \ appr_rel" if "DIM_precond (TYPE('n::enum rvec)) E" using that apply (auto simp: appr_rel_br set_of_appr_of_ivl_point_append image_image eucl_of_list_prod dest!: brD intro!: brI dest: drop_set_of_apprD[where n="CARD('n)"] cong del: image_cong_simp) subgoal for a b apply (drule set_of_appr_dropD) apply safe apply (rule image_eqI) defer apply assumption apply (auto simp: eucl_of_list_prod) apply (rule eucl_of_list_eq_takeI) apply simp done done sublocale autoref_op_pat_def op_image_zerofst_vec . lemma [autoref_op_pat_def]: "embed1 ` X \ OP op_image_embed1 $ X" by auto lemma op_image_embed1_impl[autoref_rules]: assumes "DIM_precond TYPE((real, 'n::enum) vec) E" shows "(\x. x@appr_of_ivl ops (replicate (E*E) 0) (replicate (E*E) 0), op_image_embed1::'n rvec set \ _) \ appr_rel \ appr_rel" using assms by (force simp: appr_rel_br br_def set_of_appr_of_ivl_point_append set_of_appr_of_ivl_append_point image_image eucl_of_list_prod length_set_of_appr) sublocale autoref_op_pat_def op_image_embed1 . lemma sv_appr1_rel[relator_props]: "single_valued appr1_rel" apply (auto simp: appr1_rel_internal appr_rel_def intro!: relator_props single_valued_union) apply (auto simp: single_valued_def) apply (auto simp: lv_rel_def set_rel_br) apply (auto simp: br_def) apply (rule imageI) apply (metis single_valued_def sv_appr_rell) by (metis imageI single_valued_def sv_appr_rell) schematic_goal inter_sctn1_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(Xi, (X::'n eucl1 set)) \ appr1_rel" "(hi, h) \ \lv_rel\sctn_rel" shows "(nres_of ?f, inter_sctn1_spec X h) \ \appr1_rel \\<^sub>r appr1_rel\nres_rel" unfolding autoref_tag_defs unfolding inter_sctn1_spec_def including art by autoref_monadic concrete_definition inter_sctn1_impl for Xi hi uses inter_sctn1_impl lemmas [autoref_rules] = inter_sctn1_impl.refine[autoref_higher_order_rule (1 2)] sublocale autoref_op_pat_def inter_sctn1_spec . schematic_goal op_image_fst_coll_nres_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::executable_euclidean_space) E" assumes [autoref_rules]: "(XSi, (XS::'n c1_info set)) \ clw_rel appr1_rel" shows "(RETURN ?r, op_image_fst_coll_nres XS) \ \clw_rel appr_rel\nres_rel" unfolding autoref_tag_defs unfolding op_image_fst_coll_nres_def including art by (autoref_monadic (plain)) concrete_definition op_image_fst_coll_nres_impl for XSi uses op_image_fst_coll_nres_impl lemmas [autoref_rules] = op_image_fst_coll_nres_impl.refine[autoref_higher_order_rule (1)] sublocale autoref_op_pat_def op_image_fst_coll_nres . lemma [autoref_op_pat]: "(`) fst \ OP op_image_fst_coll" by auto lemma op_image_fst_coll_impl[autoref_rules]: assumes "DIM_precond TYPE('n::executable_euclidean_space) E" shows "(op_image_fst_coll_nres_impl, op_image_fst_coll::_\'n set) \ clw_rel appr1_rel \ clw_rel appr_rel" apply rule subgoal premises prems for x using nres_rel_trans2[OF op_image_fst_coll_nres_spec[OF order_refl] op_image_fst_coll_nres_impl.refine[OF assms, simplified, OF prems]] by (auto simp: nres_rel_def RETURN_RES_refine_iff) done schematic_goal fst_safe_coll_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::executable_euclidean_space) E" assumes [autoref_rules]: "(XSi, (XS::'n c1_info set)) \ clw_rel appr1_rel" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" shows "(nres_of ?r, fst_safe_coll odo XS) \ \clw_rel appr_rel\nres_rel" unfolding autoref_tag_defs unfolding fst_safe_coll_def including art by autoref_monadic concrete_definition fst_safe_coll_impl for XSi uses fst_safe_coll_impl lemmas [autoref_rules] = fst_safe_coll_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def fst_safe_coll_impl . lemma [autoref_op_pat]: "(`) flow1_of_vec1 \ OP op_image_flow1_of_vec1_coll" by auto lemma op_image_flow1_of_vec1_coll[autoref_rules]: "(map (\x. (take E x, Some (drop E x))), op_image_flow1_of_vec1_coll::_\'n eucl1 set) \ clw_rel appr_rel \ clw_rel appr1_rel" if "DIM_precond TYPE('n::enum rvec) E" apply (rule lift_clw_rel_map) apply (rule relator_props) apply (rule relator_props) unfolding op_image_flow1_of_vec1_coll_def op_image_flow1_of_vec1_def[symmetric] apply (rule op_image_flow1_of_vec1) using that by auto sublocale autoref_op_pat_def op_image_flow1_of_vec1_coll . schematic_goal vec1reps_impl: assumes [autoref_rules]: "(Xi, X) \ clw_rel appr1_rel" shows "(RETURN ?r, vec1reps X) \ \\clw_rel appr_rel\option_rel\nres_rel" unfolding vec1reps_def including art by (autoref_monadic (plain)) concrete_definition vec1reps_impl for Xi uses vec1reps_impl lemma vec1reps_impl_refine[autoref_rules]: "(\x. RETURN (vec1reps_impl x), vec1reps) \ clw_rel appr1_rel \ \\clw_rel appr_rel\option_rel\nres_rel" using vec1reps_impl.refine by force sublocale autoref_op_pat_def vec1reps . abbreviation "intersection_STATE_rel \ (appr1_rel \\<^sub>r \Id\phantom_rel \\<^sub>r clw_rel appr1_rel \\<^sub>r clw_rel appr1_rel \\<^sub>r clw_rel (\appr_rel, \lv_rel\sbelows_rel\inter_rel) \\<^sub>r bool_rel \\<^sub>r bool_rel)" lemma print_set_impl1[autoref_rules]: shows "(\a s. printing_fun optns a (list_of_appr1 s), print_set1) \ bool_rel \ A \ Id" by auto sublocale autoref_op_pat_def print_set1 . lemma trace_set1_impl1[autoref_rules]: shows "(\s a. tracing_fun optns s (map_option list_of_appr1 a), trace_set1) \ string_rel \ \A\option_rel \ Id" by auto sublocale autoref_op_pat_def trace_set1 . lemma print_set_impl1e[autoref_rules]: shows "(\a s. printing_fun optns a (list_of_appr1e s), print_set1e) \ bool_rel \ A \ Id" by auto sublocale autoref_op_pat_def print_set1e . lemma trace_set1_impl1e[autoref_rules]: shows "(\s a. tracing_fun optns s (map_option (list_of_appr1e) a), trace_set1e) \ string_rel \ \A\option_rel \ Id" by auto sublocale autoref_op_pat_def trace_set1e . schematic_goal split_spec_param1_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::enum rvec) E" assumes [autoref_rules]: "(Xi, X) \ appr1_rel" notes [autoref_post_simps] = case_prod_eta shows "(nres_of (?f), split_spec_param1 (X::'a eucl1 set)) \ \appr1_rel \\<^sub>r appr1_rel\nres_rel" unfolding autoref_tag_defs unfolding split_spec_param1_def including art by autoref_monadic concrete_definition split_spec_param1_impl for Xi uses split_spec_param1_impl lemmas split_spec_param1_refine[autoref_rules] = split_spec_param1_impl.refine[autoref_higher_order_rule (1)] sublocale autoref_op_pat_def split_spec_param1 . lemma [autoref_op_pat del]: "{} \ OP op_empty_default" "{} \ OP op_empty_coll" and [autoref_op_pat_def del]: "get_inter p \ OP (get_inter p)" by simp_all lemma fst_image_c1_info_of_appr: "c1_info_invar (DIM('a)) X \ (fst ` c1_info_of_appr X::'a::executable_euclidean_space set) = eucl_of_list ` (set_of_appr (fst X))" apply (auto simp: c1_info_invar_def power2_eq_square image_image flow1_of_list_def c1_info_of_appr_def flow1_of_vec1_def eucl_of_list_prod split: option.splits) subgoal for a b by (rule image_eqI[where x="take DIM('a) b"]) (auto intro!: take_set_of_apprI simp: length_set_of_appr) subgoal for a b apply (frule set_of_appr_ex_append2[where b=a]) apply auto subgoal for r by (rule image_eqI[where x="b@r"]) (auto intro!: eucl_of_list_eqI dest!: length_set_of_appr) done done lemma op_image_fst_colle_impl[autoref_rules]: "(map (\(_, x, _). x), op_image_fst_colle) \ clw_rel appr1e_rel \ clw_rel appr_rel" apply (rule fun_relI) unfolding appr_rel_br apply (rule map_mem_clw_rel_br) unfolding appr1_rel_br unfolding scaleR2_rel_br unfolding clw_rel_br apply (auto dest!: brD simp: image_Union split_beta') apply (drule bspec, assumption) apply auto apply (drule bspec, assumption) apply (auto simp: fst_image_c1_info_of_appr) apply (rule bexI) prefer 2 apply assumption apply (auto simp: scaleR2_rel_br scaleR2_def image_def c1_info_of_appr_def split: option.splits) subgoal for a b c d e f g h i apply (rule bexI[where x="take DIM('a) i"]) by (auto intro!: take_set_of_apprI simp: flow1_of_list_def eucl_of_list_prod c1_info_invar_def length_set_of_appr) subgoal by (auto intro!: take_set_of_apprI simp: flow1_of_vec1_def eucl_of_list_prod length_set_of_appr c1_info_invar_def) done sublocale autoref_op_pat_def op_image_fst_colle . lemma is_empty_appr1_rel[autoref_rules]: "(\_. False, is_empty) \ appr1_rel \ bool_rel" by (auto simp: appr1_rel_internal set_rel_br) (auto simp: appr_rel_br br_def) sublocale autoref_op_pat_def is_empty . schematic_goal split_spec_param1e_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::enum rvec) E" assumes [autoref_rules]: "(Xi, X) \ \appr1_rel\scaleR2_rel" notes [autoref_post_simps] = case_prod_eta shows "(nres_of (?f), split_spec_param1e (X::'a eucl1 set)) \ \\appr1_rel\scaleR2_rel \\<^sub>r \appr1_rel\scaleR2_rel\nres_rel" unfolding autoref_tag_defs unfolding split_spec_param1e_def including art by autoref_monadic concrete_definition split_spec_param1e_impl for Xi uses split_spec_param1e_impl lemmas split_spec_param1e_refine[autoref_rules] = split_spec_param1e_impl.refine[autoref_higher_order_rule (1)] sublocale autoref_op_pat_def split_spec_param1e . schematic_goal reduce_spec1_impl: "(nres_of ?r, reduce_spec1 C X) \ \appr1_rel\nres_rel" if [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" and [autoref_rules]: "(Xi, X::'n eucl1 set) \ appr1_rel" "(Ci, C) \ reduce_argument_rel TYPE('b)" unfolding reduce_spec1_def including art by autoref_monadic concrete_definition reduce_spec1_impl for Ci Xi uses reduce_spec1_impl lemmas reduce_spec1_impl_refine[autoref_rules] = reduce_spec1_impl.refine[autoref_higher_order_rule (1)] sublocale autoref_op_pat_def reduce_spec1 . schematic_goal reduce_spec1e_impl: "(nres_of ?r, reduce_spec1e C X) \ \\appr1_rel\scaleR2_rel\nres_rel" if [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" and [autoref_rules]: "(Xi, X::'n eucl1 set) \ \appr1_rel\scaleR2_rel" "(Ci, C) \ reduce_argument_rel TYPE('b)" unfolding reduce_spec1e_def including art by autoref_monadic concrete_definition reduce_spec1e_impl for Ci Xi uses reduce_spec1e_impl lemmas reduce_spec1e_impl_refine[autoref_rules] = reduce_spec1e_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def reduce_spec1e . lemma eq_spec_impl[autoref_rules]: "(\a b. RETURN (a = b), eq_spec) \ A \ A \ \bool_rel\nres_rel" if "PREFER single_valued A" using that by (auto simp: nres_rel_def single_valued_def) schematic_goal select_with_inter_impl: assumes [relator_props]: "single_valued A" "single_valued P" assumes [autoref_rules]: "(ci, c) \ clw_rel (\A, P\inter_rel)" "(ai, a) \ clw_rel A" shows "(RETURN ?r, select_with_inter $ c $ a) \ \clw_rel (\A, P\inter_rel)\nres_rel" unfolding select_with_inter_def including art by (autoref_monadic (plain)) concrete_definition select_with_inter_impl for ci ai uses select_with_inter_impl lemmas [autoref_rules] = select_with_inter_impl.refine[OF PREFER_sv_D PREFER_sv_D] sublocale autoref_op_pat_def select_with_inter . schematic_goal choose_step1e_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" "ncc_precond TYPE('n vec1)" "ncc_precond TYPE('n rvec)" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ appr1e_rel" "(hi, h) \ rnv_rel" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" shows "(nres_of ?r, choose_step1e odo X h) \ \rnv_rel \\<^sub>r appr1_rel \\<^sub>r appr_rel \\<^sub>r appr1e_rel\nres_rel" unfolding choose_step1e_def including art by autoref_monadic concrete_definition choose_step1e_impl for Xi hi uses choose_step1e_impl lemmas [autoref_rules] = choose_step1e_impl.refine[autoref_higher_order_rule (1 2 3)] sublocale autoref_op_pat_def choose_step1e . lemma pre_split_reduce_impl[autoref_rules]: "(\ro. RETURN (pre_split_reduce ro), pre_split_reduce_spec) \ reach_optns_rel \ \reduce_argument_rel TYPE('b)\nres_rel" by (auto simp: pre_split_reduce_spec_def nres_rel_def reduce_argument_rel_def RETURN_RES_refine) schematic_goal step_split_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ appr1e_rel" and [autoref_rules]: "(odoi, odo) \ ode_ops_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" shows "(nres_of (?f), step_split odo ro X)\\clw_rel appr1e_rel\nres_rel" using assms unfolding step_split_def[abs_def] including art by autoref_monadic concrete_definition step_split_impl for odoi Xi uses step_split_impl lemmas [autoref_rules] = step_split_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def step_split . schematic_goal width_spec_appr1_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ appr1_rel" shows "(?r, width_spec_appr1 X) \ \rnv_rel\nres_rel" unfolding width_spec_appr1_def by autoref_monadic concrete_definition width_spec_appr1_impl for Xi uses width_spec_appr1_impl lemmas width_spec_appr1_impl_refine[autoref_rules] = width_spec_appr1_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def width_spec_appr1 . schematic_goal split_under_threshold_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" assumes [autoref_rules]: "(thi, th) \ rnv_rel" "(Xi, X) \ clw_rel (\appr1_rel\scaleR2_rel)" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" shows "(nres_of ?x, split_under_threshold ro th (X::'n eucl1 set)) \ \clw_rel (\appr1_rel\scaleR2_rel)\nres_rel" unfolding autoref_tag_defs unfolding split_under_threshold_def by autoref_monadic concrete_definition split_under_threshold_impl for thi Xi uses split_under_threshold_impl lemmas [autoref_rules] = split_under_threshold_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def split_under_threshold . schematic_goal pre_intersection_step_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ appr1e_rel" "(hi, (h::real)) \ rnv_rel" and [autoref_rules]: "(roptnsi, roptns) \ reach_optns_rel" "(odoi, odo) \ ode_ops_rel" shows "(nres_of ?r, pre_intersection_step odo roptns X h) \ \clw_rel (iinfo_rel appr1e_rel) \\<^sub>r clw_rel appr_rel \\<^sub>r clw_rel (iinfo_rel appr1e_rel)\nres_rel" unfolding pre_intersection_step_def including art by autoref_monadic concrete_definition pre_intersection_step_impl for roptnsi Xi hi uses pre_intersection_step_impl lemmas [autoref_rules] = pre_intersection_step_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def pre_intersection_step . schematic_goal subset_spec_plane_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a) E" assumes [autoref_rules]: "(Xi, X::'a::executable_euclidean_space set) \ lvivl_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" shows "(nres_of ?R, subset_spec_plane X sctn) \ \bool_rel\nres_rel" unfolding subset_spec_plane_def by autoref_monadic concrete_definition subset_spec_plane_impl for Xi sctni uses subset_spec_plane_impl lemmas [autoref_rules] = subset_spec_plane_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def subset_spec_plane . schematic_goal op_eventually_within_sctn_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::executable_euclidean_space) E" assumes [autoref_rules]: "(Xi, X::'a set) \ appr_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" "(Si, S) \ lvivl_rel" shows "(nres_of ?R, op_eventually_within_sctn X sctn S) \ \bool_rel\nres_rel" unfolding op_eventually_within_sctn_def including art by autoref_monadic concrete_definition op_eventually_within_sctn_impl for Xi sctni Si uses op_eventually_within_sctn_impl lemmas [autoref_rules] = op_eventually_within_sctn_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def op_eventually_within_sctn . schematic_goal nonzero_component_within_impl: "(nres_of ?r, nonzero_component_within odo ivl sctn (PDP::'n eucl1 set)) \ \bool_rel\nres_rel" if [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" and [autoref_rules]: "(ivli, ivl) \ lvivl_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" "(PDPi, PDP) \ appr1_rel" "(odoi, odo) \ ode_ops_rel" unfolding nonzero_component_within_def including art by autoref_monadic concrete_definition nonzero_component_within_impl uses nonzero_component_within_impl lemmas [autoref_rules] = nonzero_component_within_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def nonzero_component_within . schematic_goal disjoints_spec_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(Xi, (X::'n::enum rvec set)) \ clw_rel appr_rel" "(Yi, (Y::'n rvec set)) \ clw_rel lvivl_rel" shows "(nres_of ?f, disjoints_spec X Y) \ \bool_rel\nres_rel" unfolding autoref_tag_defs unfolding disjoints_spec_def op_coll_is_empty_def[symmetric] including art by autoref_monadic concrete_definition disjoints_spec_impl for Xi Yi uses disjoints_spec_impl lemmas [autoref_rules] = disjoints_spec_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def disjoints_spec . schematic_goal do_intersection_body_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(hi, h) \ rnv_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel lvivl_rel" and civl[autoref_rules]: "(ivli, ivl::'n rvec set) \ lvivl_rel" and csctns[autoref_rules]: "(sctni, sctn::'n rvec sctn) \ \lv_rel\sctn_rel" and [autoref_rules]: "(STATEi, STATE) \ intersection_STATE_rel" notes [intro, simp] = list_set_rel_finiteD closed_ivl_rel[OF civl] shows "(nres_of ?f, do_intersection_body odo guards ivl sctn h STATE) \ \intersection_STATE_rel\nres_rel" unfolding do_intersection_body_def by autoref_monadic concrete_definition do_intersection_body_impl for odoi guardsi ivli sctni hi STATEi uses do_intersection_body_impl lemmas do_intersection_body_impl_refine[autoref_rules] = do_intersection_body_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def do_intersection_body . schematic_goal do_intersection_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(Xi, X) \ appr1_rel" "(hi, h) \ rnv_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (\lvivl_rel, \lv_rel\plane_rel\inter_rel)" and civl[autoref_rules]: "(ivli, ivl::'n rvec set) \ lvivl_rel" and csctns[autoref_rules]: "(sctni, sctn::'n rvec sctn) \ \lv_rel\sctn_rel" notes [intro, simp] = list_set_rel_finiteD closed_ivl_rel[OF civl] shows "(nres_of ?f, do_intersection odo guards ivl sctn (X::'n eucl1 set) h)\ \bool_rel \\<^sub>r clw_rel appr1_rel \\<^sub>r clw_rel appr1_rel \\<^sub>r clw_rel (\appr_rel, \lv_rel\sbelows_rel\inter_rel)\nres_rel" unfolding autoref_tag_defs unfolding do_intersection_def including art by autoref_monadic concrete_definition do_intersection_impl for odoi guardsi ivli sctni Xi hi uses do_intersection_impl lemmas do_intersection_impl_refine[autoref_rules] = do_intersection_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def do_intersection . schematic_goal tolerate_error1_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) dd" assumes [autoref_rules]: "(Yi, Y::'n eucl1 set) \ appr1e_rel" assumes [autoref_rules]: "(Ei, E) \ appr1_rel" shows "(nres_of ?r, tolerate_error1 Y E) \ \bool_rel \\<^sub>r rnv_rel\nres_rel" unfolding tolerate_error1_def including art by autoref_monadic concrete_definition tolerate_error1_impl for dd Yi Ei uses tolerate_error1_impl lemmas tolerate_error1_refine[autoref_rules] = tolerate_error1_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def tolerate_error1 . lemma lower_impl[autoref_rules]: "(lower, lower) \ Id \ Id" and upper_impl[autoref_rules]: "(lower, lower) \ Id \ Id" by auto schematic_goal step_adapt_time_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(hi, h) \ rnv_rel" "(Xi, X::'n eucl1 set) \ (appr1e_rel)" shows "(nres_of ?f, step_adapt_time odo X h)\\rnv_rel \\<^sub>r appr_rel \\<^sub>r appr1e_rel \\<^sub>r rnv_rel\nres_rel" unfolding step_adapt_time_def[abs_def] including art by autoref_monadic concrete_definition step_adapt_time_impl for Xi hi uses step_adapt_time_impl lemmas [autoref_rules] = step_adapt_time_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def step_adapt_time . schematic_goal resolve_step_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(hi, h) \ rnv_rel" "(Xi, X::'n eucl1 set) \ (appr1e_rel)" "(roptnsi, roptns) \ reach_optns_rel" shows "(nres_of ?f, resolve_step odo roptns X h)\\rnv_rel \\<^sub>r clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r rnv_rel\nres_rel" unfolding resolve_step_def[abs_def] including art by autoref_monadic concrete_definition resolve_step_impl for roptnsi Xi hi uses resolve_step_impl lemmas [autoref_rules] = resolve_step_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def resolve_step . sublocale autoref_op_pat_def fst_safe_coll . schematic_goal reach_cont_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS) \ clw_rel appr1e_rel" "(guardsi, guards::'n rvec set) \ clw_rel (iplane_rel lvivl_rel)" and [autoref_rules]: "(roptnsi, roptns) \ reach_optns_rel" notes [relator_props, autoref_rules_raw] = sv_appr1_rel shows "(nres_of (?f::?'f dres), reach_cont odo roptns guards XS)\?R" unfolding autoref_tag_defs unfolding reach_cont_def including art by autoref_monadic concrete_definition reach_cont_impl for guardsi XSi uses reach_cont_impl lemmas reach_cont_ho[autoref_rules] = reach_cont_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def reach_cont . schematic_goal reach_cont_par_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS) \ clw_rel appr1e_rel" "(guardsi, guards::'n rvec set) \ clw_rel (iplane_rel lvivl_rel)" and [autoref_rules]: "(roptnsi, roptns) \ reach_optns_rel" shows "(nres_of (?f::?'f dres), reach_cont_par odo roptns guards XS)\?R" unfolding autoref_tag_defs unfolding reach_cont_par_def including art by autoref_monadic concrete_definition reach_cont_par_impl for roptnsi guardsi XSi uses reach_cont_par_impl lemmas reach_cont_par_impl_refine[autoref_rules] = reach_cont_par_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def reach_cont_par . schematic_goal subset_iplane_coll_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::executable_euclidean_space) E" assumes [autoref_rules]: "(xi, x::'a set) \ iplane_rel lvivl_rel" assumes [autoref_rules]: "(icsi, ics) \ clw_rel (iplane_rel lvivl_rel)" shows "(nres_of ?r, subset_iplane_coll x ics) \ \bool_rel\nres_rel" unfolding subset_iplane_coll_def including art by autoref_monadic concrete_definition subset_iplane_coll_impl uses subset_iplane_coll_impl lemmas subset_iplane_coll_impl_refine[autoref_rules] = subset_iplane_coll_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def subset_iplane_coll . schematic_goal subsets_iplane_coll_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::executable_euclidean_space) E" assumes [autoref_rules]: "(xi, x::'a set set) \ \iplane_rel lvivl_rel\list_wset_rel" assumes [autoref_rules]: "(icsi, ics) \ clw_rel (iplane_rel lvivl_rel)" shows "(nres_of ?r, subsets_iplane_coll x ics) \ \bool_rel\nres_rel" unfolding subsets_iplane_coll_def including art by autoref_monadic concrete_definition subsets_iplane_coll_impl uses subsets_iplane_coll_impl lemmas [autoref_rules] = subsets_iplane_coll_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def subsets_iplane_coll . schematic_goal symstart_coll_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS::'n eucl1 set) \ clw_rel appr1e_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" shows "(nres_of ?r, symstart_coll $ odo $ symstart $ XS) \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" unfolding symstart_coll_def including art by autoref_monadic concrete_definition symstart_coll_impl for symstartd XSi uses symstart_coll_impl lemmas [autoref_rules] = symstart_coll_impl.refine sublocale autoref_op_pat_def symstart_coll . schematic_goal reach_cont_symstart_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS::'n eucl1 set) \ clw_rel appr1e_rel" "(guardsi, guards::'n rvec set) \ clw_rel (iplane_rel lvivl_rel)" "(roptnsi, roptns) \ reach_optns_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" shows "(nres_of (?r), reach_cont_symstart $ odo $ roptns $ symstart $ guards $ XS) \ \clw_rel appr_rel \\<^sub>r clw_rel (\iplane_rel lvivl_rel::(_ \ 'n rvec set)set, iinfo_rel appr1e_rel\info_rel)\nres_rel" unfolding autoref_tag_defs unfolding reach_cont_symstart_def Let_def including art by autoref_monadic concrete_definition reach_cont_symstart_impl for roptnsi symstartd XSi uses reach_cont_symstart_impl lemmas [autoref_rules] = reach_cont_symstart_impl.refine sublocale autoref_op_pat_def reach_cont_symstart . lemma sv_reach_conts_impl_aux: "single_valued (clw_rel (iinfo_rel appr1e_rel))" by (auto intro!: relator_props) schematic_goal reach_conts_impl: assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS) \ clw_rel appr1e_rel" "(guardsi, guards::'n rvec set) \ clw_rel (iplane_rel lvivl_rel)" and [autoref_rules]: "(roptnsi, roptns) \ reach_optns_rel" notes [simp] = list_wset_rel_finite[OF sv_reach_conts_impl_aux] assumes "(trapi, trap) \ ghost_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" shows "(nres_of (?f::?'f dres), reach_conts $ odo $ roptns $ symstart $ trap $ guards $ XS)\?R" unfolding autoref_tag_defs unfolding reach_conts_def including art by autoref_monadic concrete_definition reach_conts_impl for odoi guardsi XSi uses reach_conts_impl lemmas [autoref_rules] = reach_conts_impl.refine sublocale autoref_op_pat_def reach_conts . lemma get_sctns_autoref[autoref_rules]: "(\x. RETURN x, get_sctns) \ \R\halfspaces_rel \ \\\R\sctn_rel\list_set_rel\nres_rel" by (auto simp: get_sctns_def nres_rel_def halfspaces_rel_def br_def intro!: RETURN_SPEC_refine) sublocale autoref_op_pat_def get_sctns . schematic_goal leaves_halfspace_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes nccp[autoref_rules_raw]: "ncc_precond TYPE('n vec1)" notes [simp] = ncc_precondD[OF nccp] assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(Si, S) \ \lv_rel\halfspaces_rel" assumes [autoref_rules]: "(Xi, X::'n rvec set) \ clw_rel appr_rel" shows "(nres_of ?r, leaves_halfspace $ odo $ S $ X) \ \\\lv_rel\sctn_rel\option_rel\nres_rel" unfolding leaves_halfspace_def including art by autoref_monadic concrete_definition leaves_halfspace_impl for Si Xi uses leaves_halfspace_impl lemmas [autoref_rules] = leaves_halfspace_impl.refine sublocale autoref_op_pat_def leaves_halfspace . schematic_goal poincare_start_on_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes ncc2[autoref_rules_raw]: "ncc_precond TYPE('n::enum rvec)" assumes ncc2[autoref_rules_raw]: "ncc_precond TYPE('n::enum vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(sctni, sctn) \ \lv_rel\sctn_rel" "(guardsi, guards) \ clw_rel lvivl_rel" "(X0i, X0::'n eucl1 set) \ clw_rel (appr1e_rel)" shows "(nres_of (?f), poincare_start_on $ odo $ guards $ sctn $ X0) \ \clw_rel appr1e_rel \\<^sub>r clw_rel appr_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_start_on_def including art by autoref_monadic concrete_definition poincare_start_on_impl for guardsi sctni X0i uses poincare_start_on_impl lemmas [autoref_rules] = poincare_start_on_impl.refine sublocale autoref_op_pat_def poincare_start_on . lemma isets_of_iivls[autoref_rules]: assumes "PREFER single_valued A" assumes le[THEN GEN_OP_D, param_fo]: "GEN_OP le (\) ((lv_rel::(_ \ 'a::executable_euclidean_space)set) \ lv_rel \ bool_rel)" shows "(\xs. map (\((i, s), x). (appr_of_ivl ops i s, x)) [((i,s), x) \ xs. le i s], isets_of_iivls::_\'a set) \ clw_rel (\lvivl_rel, A\inter_rel) \ clw_rel (\appr_rel, A\inter_rel)" apply (rule fun_relI) using assms apply (auto elim!: single_valued_as_brE) unfolding appr_rel_br ivl_rel_br clw_rel_br lvivl_rel_br inter_rel_br apply (auto simp: br_def set_of_ivl_def) subgoal for a b c d e f g apply (rule exI[where x=e]) apply (rule exI[where x=f]) apply (rule exI[where x=g]) apply (rule conjI) apply (assumption) apply (rule conjI) subgoal using transfer_operations1[where 'a='a, of "eucl_of_list e" "eucl_of_list f" e f] le[of e _ f _, OF lv_relI lv_relI] by (auto simp: appr_rel_br br_def lvivl_rel_br set_of_ivl_def lv_rel_def) subgoal apply (drule bspec, assumption) using transfer_operations1[where 'a='a, of "eucl_of_list e" "eucl_of_list f" e f] le[of e _ f _, OF lv_relI lv_relI] apply (auto simp: appr_rel_br br_def lvivl_rel_br set_of_ivl_def lv_rel_def) using atLeastAtMost_iff apply blast - apply (drule order_trans) - apply assumption apply simp done done subgoal for a b c d e f g apply (drule bspec, assumption) using transfer_operations1[where 'a='a, of "eucl_of_list d" "eucl_of_list e" d e] le[of d _ e _, OF lv_relI lv_relI] by (auto simp: appr_rel_br br_def lvivl_rel_br set_of_ivl_def lv_rel_def intro!: bexI) subgoal for a b c d e f apply (drule bspec, assumption) using transfer_operations1[where 'a='a, of "eucl_of_list d" "eucl_of_list e" d e] le[of d _ e _, OF lv_relI lv_relI] by (auto simp: appr_rel_br br_def lvivl_rel_br set_of_ivl_def lv_rel_def intro!: bexI) done sublocale autoref_op_pat_def isets_of_iivls . lemma [autoref_op_pat]: "X \ UNIV \ OP op_times_UNIV_coll $ X" by simp lemma op_times_UNIV_coll_impl[autoref_rules]: "(map (\x. (x, None)), op_times_UNIV_coll) \ clw_rel appr_rel \ clw_rel appr1_rel" apply (rule lift_clw_rel_map) apply (rule relator_props) apply (rule relator_props) unfolding op_times_UNIV_coll_def op_times_UNIV_def[symmetric] apply (rule op_times_UNIV_impl) by auto sublocale autoref_op_pat_def op_times_UNIV_coll . schematic_goal do_intersection_core_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ iinfo_rel appr1e_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (iplane_rel lvivl_rel)" and csctns[autoref_rules]: "(sctni, sctn) \ \lv_rel\sctn_rel" and csctns[autoref_rules]: "(ivli, ivl) \ lvivl_rel" notes [simp] = list_set_rel_finiteD shows "(nres_of ?f, do_intersection_core odo guards ivl sctn X) \ \clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr1e_rel\nres_rel" unfolding do_intersection_core_def[abs_def] including art by autoref_monadic concrete_definition do_intersection_core_impl for guardsi ivli sctni Xi uses do_intersection_core_impl sublocale autoref_op_pat_def do_intersection_core . lemmas do_intersection_core_impl_refine[autoref_rules] = do_intersection_core_impl.refine[autoref_higher_order_rule(1 2 3)] lemma finite_ra1eicacacslsbicae1lw: "(xc, x'c) \ \\rnv_rel, appr1e_rel\info_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel (\appr_rel, \lv_rel\sbelows_rel\inter_rel) \\<^sub>r clw_rel appr1e_rel\list_wset_rel \ finite x'c" for x'c::"('n::enum eucl1 set * 'n eucl1 set * 'n eucl1 set * 'n rvec set * 'n eucl1 set) set" apply (rule list_wset_rel_finite) by (auto intro!: relator_props) schematic_goal do_intersection_coll_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE((real, 'n::enum) vec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(Xi, X::'n eucl1 set) \ clw_rel (iinfo_rel appr1e_rel)" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (iplane_rel lvivl_rel)" and csctns[autoref_rules]: "(sctni, sctn) \ \lv_rel\sctn_rel" and csctns[autoref_rules]: "(ivli, ivl) \ lvivl_rel" notes [simp] = finite_ra1eicacacslsbicae1lw[where 'n='n] shows "(nres_of ?f, do_intersection_coll odo guards ivl sctn X) \ \clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr1e_rel\nres_rel" unfolding do_intersection_coll_def[abs_def] including art by autoref_monadic concrete_definition do_intersection_coll_impl for guardsi ivli sctni Xi uses do_intersection_coll_impl lemmas [autoref_rules] = do_intersection_coll_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def do_intersection_coll . schematic_goal op_enlarge_ivl_sctn_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::executable_euclidean_space) E" assumes [autoref_rules]: "(ivli, ivl::'a set) \ lvivl_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" "(di, d) \ rnv_rel" shows "(nres_of ?R, op_enlarge_ivl_sctn $ ivl $ sctn $ d) \ \lvivl_rel\nres_rel" unfolding op_enlarge_ivl_sctn_def including art by autoref_monadic concrete_definition op_enlarge_ivl_sctn_impl for ivli sctni di uses op_enlarge_ivl_sctn_impl lemmas [autoref_rules] = op_enlarge_ivl_sctn_impl.refine sublocale autoref_op_pat_def op_enlarge_ivl_sctn . schematic_goal resolve_ivlplanes_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS::('n rvec set \ 'n eucl1 set) set) \ \iplane_rel lvivl_rel \\<^sub>r clw_rel (iinfo_rel appr1e_rel)\list_wset_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (iplane_rel lvivl_rel)" and osctns[autoref_rules]: "(ivlplanesi, ivlplanes) \ clw_rel (iplane_rel lvivl_rel)" notes [intro, simp] = list_set_rel_finiteD shows "(nres_of ?r, resolve_ivlplanes odo guards ivlplanes XS) \ \\clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel)\list_wset_rel\nres_rel" unfolding autoref_tag_defs unfolding resolve_ivlplanes_def including art by autoref_monadic concrete_definition resolve_ivlplanes_impl for guardsi ivlplanesi XSi uses resolve_ivlplanes_impl lemmas [autoref_rules] = resolve_ivlplanes_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def resolve_ivlplanes . schematic_goal poincare_onto_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS::'n eucl1 set) \ clw_rel appr1e_rel" assumes [autoref_rules]: "(CXSi, CXS::'n rvec set) \ clw_rel appr_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (iplane_rel lvivl_rel)" and osctns[autoref_rules]: "(ivlplanesi, ivlplanes) \ clw_rel (iplane_rel lvivl_rel)" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" assumes [autoref_rules]: "((), trap) \ ghost_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" notes [intro, simp] = list_set_rel_finiteD shows "(nres_of ?r, poincare_onto $ odo $ ro $ symstart $ trap $ guards $ ivlplanes $ XS $ CXS) \ \\clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto_def including art by autoref_monadic concrete_definition poincare_onto_impl for odoi roi symstarti guardsi ivlplanesi XSi uses poincare_onto_impl lemmas [autoref_rules] = poincare_onto_impl.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def poincare_onto . schematic_goal empty_remainders_impl: assumes [autoref_rules]: "(PSi, PS) \ \clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel" shows "(nres_of ?r, empty_remainders PS) \ \bool_rel\nres_rel" unfolding empty_remainders_def including art by autoref_monadic concrete_definition empty_remainders_impl uses empty_remainders_impl lemmas empty_remainders_impl_refine[autoref_rules] = empty_remainders_impl.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def empty_remainders . lemma empty_trap_impl[autoref_rules]: "((), empty_trap) \ ghost_rel" by (auto intro!: ghost_relI) sublocale autoref_op_pat_def empty_trap . lemma empty_symstart_impl:\ \why this? \ "((\x. nres_of (dRETURN ([], [x]))), empty_symstart) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" unfolding empty_symstart_def using mk_coll[unfolded autoref_tag_defs, OF sv_appr1e_rel[OF sv_appr1_rel], param_fo] by (auto intro!: nres_relI simp:) lemma empty_symstart_nres_rel[autoref_rules]: "((\x. RETURN ([], [x])), empty_symstart::'n::enum eucl1 set\_) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" using mk_coll[OF PREFER_I[of single_valued, OF sv_appr1e_rel[OF sv_appr1_rel]], param_fo, of x y for x and y::"'n eucl1 set"] by (auto simp: mk_coll_def[abs_def] nres_rel_def) sublocale autoref_op_pat_def empty_symstart . lemma empty_symstart_dres_nres_rel: "((\x. dRETURN ([], [x])), empty_symstart::'n::enum eucl1 set\_) \ (appr1e_rel) \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\dres_nres_rel" using mk_coll[OF PREFER_I[of single_valued, OF sv_appr1e_rel[OF sv_appr1_rel]], param_fo, of x y for x and y::"'n eucl1 set"] by (auto simp: mk_coll_def[abs_def] dres_nres_rel_def) schematic_goal poincare_onto_empty_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS::'n eucl1 set) \ clw_rel appr1e_rel" assumes [autoref_rules]: "(CXSi, CXS::'n rvec set) \ clw_rel appr_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (iplane_rel lvivl_rel)" and osctns[autoref_rules]: "(ivlplanesi, ivlplanes) \ clw_rel (iplane_rel lvivl_rel)" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" notes [intro, simp] = list_set_rel_finiteD shows "(nres_of (?r), poincare_onto_empty odo ro guards ivlplanes XS CXS) \ \\clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto_empty_def including art apply (rule autoref_monadicI) apply (autoref phases: id_op rel_inf fix_rel)\ \TODO: what is going wrong here?\ apply (simp only: autoref_tag_defs) apply (rule poincare_onto_impl.refine[unfolded autoref_tag_defs]) apply fact+ apply (rule ghost_relI) apply (rule empty_symstart_impl) apply refine_transfer apply refine_transfer done concrete_definition poincare_onto_empty_impl for guardsi XSi CXSi uses poincare_onto_empty_impl lemmas [autoref_rules] = poincare_onto_empty_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def poincare_onto_empty . lemma sv_thingy: "single_valued (clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r \lv_rel\ivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (\appr_rel, \lv_rel\sbelows_rel\inter_rel) \\<^sub>r clw_rel appr_rel)" by (intro relator_props) schematic_goal poincare_onto2_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" assumes [autoref_rules]: "(XSi, XS::'n eucl1 set) \ clw_rel appr1e_rel" and osctns[autoref_rules]: "(guardsi, guards) \ clw_rel (iplane_rel lvivl_rel)" and osctns[autoref_rules]: "(ivlplanesi, ivlplanes) \ clw_rel (iplane_rel lvivl_rel)" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" assumes [autoref_rules]: "((), trap) \ ghost_rel" notes [intro, simp] = list_set_rel_finiteD list_wset_rel_finite[OF sv_thingy] shows "(nres_of (?r), poincare_onto2 $ odo $ ro $ symstart $ trap $ guards $ ivlplanes $ XS) \ \\bool_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto2_def including art by autoref_monadic concrete_definition poincare_onto2_impl for odoi guardsi XSi uses poincare_onto2_impl lemmas [autoref_rules] = poincare_onto2_impl.refine sublocale autoref_op_pat_def poincare_onto2 . schematic_goal width_spec_ivl_impl: assumes [autoref_rules]: "(Mi, M) \ nat_rel" "(xi, x) \ lvivl_rel" shows "(RETURN ?r, width_spec_ivl M x) \ \rnv_rel\nres_rel" unfolding width_spec_ivl_def including art by (autoref_monadic (plain)) concrete_definition width_spec_ivl_impl for Mi xi uses width_spec_ivl_impl lemmas width_spec_ivl_impl_refine[autoref_rules] = width_spec_ivl_impl.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def width_spec_ivl . schematic_goal partition_ivl_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('a::executable_euclidean_space) E" assumes [autoref_rules]: "(xsi, xs::'a set)\ clw_rel lvivl_rel" "(roi, ro) \ reach_optns_rel" shows "(nres_of (?f), partition_ivl ro xs)\\clw_rel lvivl_rel\nres_rel" unfolding partition_ivl_def[abs_def] including art by autoref_monadic concrete_definition partition_ivl_impl for roi xsi uses partition_ivl_impl lemmas [autoref_rules] = partition_ivl_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def partition_ivl . schematic_goal vec1repse_impl: assumes [autoref_rules]: "(Xi, X) \ clw_rel appr1e_rel" shows "(nres_of ?r, vec1repse X) \ \\clw_rel appre_rel\option_rel\nres_rel" unfolding vec1repse_def including art by autoref_monadic concrete_definition vec1repse_impl for Xi uses vec1repse_impl lemmas vec1repse_impl_refine[autoref_rules] = vec1repse_impl.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def vec1repse . schematic_goal scaleR2_rep1_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(Yi, Y::'n vec1 set) \ appre_rel" shows "(nres_of ?r, scaleR2_rep1 Y) \ \elvivl_rel\nres_rel" unfolding scaleR2_rep1_def including art by autoref_monadic concrete_definition scaleR2_rep1_impl uses scaleR2_rep1_impl lemmas [autoref_rules] = scaleR2_rep1_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def scaleR2_rep1 . schematic_goal ivlse_of_setse_impl: "(nres_of ?r, ivlse_of_setse X) \ \clw_rel elvivl_rel\nres_rel" if [autoref_rules_raw]:"DIM_precond TYPE('n::enum rvec) E" and [autoref_rules]: "(Xi, X::'n vec1 set) \ clw_rel (appre_rel)" unfolding ivlse_of_setse_def including art by autoref_monadic concrete_definition ivlse_of_setse_impl uses ivlse_of_setse_impl lemmas [autoref_rules] = ivlse_of_setse_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def ivlse_of_setse . schematic_goal setse_of_ivlse_impl: "(nres_of ?r, setse_of_ivlse X) \ \clw_rel (appre_rel)\nres_rel" if [autoref_rules]: "(Xi, X) \ clw_rel elvivl_rel" unfolding setse_of_ivlse_def including art by autoref_monadic concrete_definition setse_of_ivlse_impl uses setse_of_ivlse_impl lemmas setse_of_ivlse_impl_refine[autoref_rules] = setse_of_ivlse_impl.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def setse_of_ivlse . lemma op_image_flow1_of_vec1_colle[autoref_rules]: "(map (\(lu, x). (lu, (take E x, Some (drop E x)))), op_image_flow1_of_vec1_colle::_\'n eucl1 set) \ clw_rel appre_rel \ clw_rel appr1e_rel" if "DIM_precond TYPE('n::enum rvec) E" apply (rule lift_clw_rel_map) apply (rule relator_props) apply (rule relator_props) apply (rule relator_props) apply (rule relator_props) apply (rule lift_scaleR2) unfolding op_image_flow1_of_vec1_colle_def op_image_flow1_of_vec1_coll_def op_image_flow1_of_vec1_def[symmetric] apply (rule op_image_flow1_of_vec1) using that subgoal by force subgoal for l u x unfolding op_image_flow1_of_vec1_def flow1_of_vec1_def scaleR2_def apply (auto simp: image_def vimage_def) subgoal for a b c d e apply (rule exI[where x="c *\<^sub>R e"]) apply (auto simp: blinfun_of_vmatrix_scaleR) apply (rule exI[where x="c"]) apply auto apply (rule bexI) prefer 2 apply assumption apply auto done subgoal for a b c d e apply (rule exI[where x="c"]) apply (auto simp: blinfun_of_vmatrix_scaleR) apply (rule exI[where x="blinfun_of_vmatrix e"]) apply auto apply (rule bexI) prefer 2 apply assumption apply auto done done subgoal by auto done sublocale autoref_op_pat_def op_image_flow1_of_vec1_colle . schematic_goal okay_granularity_impl: assumes [autoref_rules]: "(ivli,ivl::'n::enum vec1 set)\ lvivl_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" shows "(nres_of ?f, okay_granularity ro ivl) \ \bool_rel\nres_rel" unfolding okay_granularity_def[abs_def] including art by autoref_monadic concrete_definition okay_granularity_impl for roi ivli uses okay_granularity_impl lemmas [autoref_rules] = okay_granularity_impl.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def okay_granularity . lemma le_post_inter_granularity_op[autoref_rules]: "(\roi (ls, us). RETURN (width_appr ops optns (appr_of_ivl ops ls us) \ post_inter_granularity roi), (le_post_inter_granularity_op::_\'a::executable_euclidean_space set\_)) \ (reach_optns_rel \ lvivl_rel \ \bool_rel\nres_rel)" by (auto simp: nres_rel_def le_post_inter_granularity_op_def) sublocale autoref_op_pat_def le_post_inter_granularity_op . schematic_goal partition_set_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(xsi,xs::'n eucl1 set)\ clw_rel appr1e_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" shows "(nres_of (?f), partition_set ro xs) \ \clw_rel appr1e_rel\nres_rel" unfolding partition_set_def including art by autoref_monadic concrete_definition partition_set_impl for roi xsi uses partition_set_impl lemmas [autoref_rules] = partition_set_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def partition_set . schematic_goal partition_sets_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(xsi,xs::(bool \ 'n eucl1 set \ _)set)\ \bool_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" shows "(nres_of (?f), partition_sets ro xs)\\clw_rel appr1e_rel\nres_rel" unfolding partition_sets_def[abs_def] including art by autoref_monadic concrete_definition partition_sets_impl for roi xsi uses partition_sets_impl lemmas [autoref_rules] = partition_sets_impl.refine[autoref_higher_order_rule(1)] sublocale autoref_op_pat_def partition_sets . lemma [autoref_rules]: assumes "PREFER single_valued A" shows "(\xs. case xs of [x] \ RETURN x | _ \ SUCCEED, singleton_spec) \ \A\list_wset_rel \ \A\nres_rel" using assms by (auto simp: nres_rel_def singleton_spec_def list_wset_rel_def set_rel_br split: list.splits elim!: single_valued_as_brE dest!: brD intro!: RETURN_SPEC_refine brI) sublocale autoref_op_pat_def singleton_spec . lemma closed_ivl_prod8_list_rel: assumes "(xl, x'l) \ \bool_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r \lv_rel\ivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (\appr_rel, \lv_rel\sbelows_rel\inter_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel" shows "(\(b, X, PS1, PS2, R, ivl, sctn, CX, CXS)\x'l. closed ivl)" using assms unfolding list_wset_rel_def set_rel_sv[OF single_valued_Id_arbitrary_interface] apply (subst (asm) set_rel_sv) subgoal by (auto simp: Id_arbitrary_interface_def intro!: relator_props intro: single_valuedI) by (auto simp: Id_arbitrary_interface_def) lemma poincare_onto_series_rec_list_eq:\ \TODO: here is a problem if interrupt gets uncurried, too\ "poincare_onto_series odo interrupt trap guards XS ivl sctn ro = rec_list (\(((((trap), XS0), ivl), sctn), ro). let guard0 = mk_coll (mk_inter ivl (plane_of sctn)) in ASSUME (closed guard0) \ (\_. (poincare_onto2 odo (ro ::: reach_optns_rel) (interrupt:::appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel) trap (op_empty_coll ::: clw_rel (\\lv_rel\ivl_rel, \lv_rel\plane_rel\inter_rel)) guard0 XS0 ::: \\bool_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel\nres_rel) \ (\(XS1). singleton_spec XS1 \ (\(b, X, PS1, PS2, R, ivl', sctn', CX, CXS). CHECKs ''poincare_onto_series: last return!'' (ivl' = ivl \ sctn' = sctn) \ (\_. RETURN PS2))))) (\x xs rr (((((trap), XS0), ivl), sctn), ro0). case x of (guard, ro) \ ASSUME (closed ivl) \ (\_. let guard0 = mk_coll (mk_inter ivl (plane_of sctn)) in ASSUME (closed guard0) \ (\_. ASSUME (\(guard, ro)\set (x # xs). closed guard) \ (\_. let guardset = \(guard, ro)\set ((guard0, ro0) # xs). guard in (poincare_onto2 odo ro (interrupt:::appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel) trap (guardset ::: clw_rel (\\lv_rel\ivl_rel, \lv_rel\plane_rel\inter_rel)) guard XS0 ::: \\bool_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r clw_rel appr1e_rel \\<^sub>r lvivl_rel \\<^sub>r \lv_rel\sctn_rel \\<^sub>r clw_rel (isbelows_rel appr_rel) \\<^sub>r clw_rel appr_rel\list_wset_rel\nres_rel) \ (\(XS1). ASSUME (\(b, X, PS, PS2, R, ivl, sctn, CX, CXS)\XS1. closed ivl) \ (\_. partition_sets ro XS1 \ (\XS2. fst_safe_colle odo XS2 \ (\_. rr (((((trap), XS2), ivl), sctn), ro0 ::: reach_optns_rel) \ RETURN)))))))) guards (((((trap), XS), ivl), sctn), ro)" by (induction guards arbitrary: XS ivl sctn ro) (auto simp: split_beta' split: prod.splits) schematic_goal poincare_onto_series_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS::'n eucl1 set) \ clw_rel appr1e_rel" and osctns[autoref_rules]: "(guardsi, guards) \ \clw_rel (iplane_rel lvivl_rel)\\<^sub>rreach_optns_rel\list_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" "(ivli, ivl) \ lvivl_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" and [autoref_rules]: "((), trap) \ ghost_rel" notes [intro, simp] = list_set_rel_finiteD closed_ivl_prod3_list_rel closed_clw_rel_iplane_rel closed_ivl_prod8_list_rel notes [autoref_rules_raw] = ghost_relI[of x for x::"'n eucl1 set"] shows "(nres_of ?r, poincare_onto_series $ odo $ symstart $ trap $ guards $ XS $ ivl $ sctn $ ro) \ \clw_rel appr1e_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto_series_rec_list_eq including art apply autoref_monadic apply (rule ghost_relI) apply (autoref phases: trans) apply (autoref phases: trans) apply (rule ghost_relI) apply (autoref phases: trans) apply (autoref phases: trans) apply simp apply (autoref phases: trans) apply (autoref phases: trans) apply simp apply (refine_transfer) done concrete_definition poincare_onto_series_impl for symstartd guardsi XSi ivli sctni roi uses poincare_onto_series_impl lemmas [autoref_rules] = poincare_onto_series_impl.refine sublocale autoref_op_pat_def poincare_onto_series . schematic_goal poincare_onto_from_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS) \ clw_rel appr1e_rel" and [autoref_rules]: "(Si, S) \ \lv_rel\halfspaces_rel" and osctns[autoref_rules]: "(guardsi, guards) \ \clw_rel (iplane_rel lvivl_rel)\\<^sub>rreach_optns_rel\list_rel" and civl[autoref_rules]: "(ivli, ivl::'n rvec set) \ lvivl_rel" and csctns[autoref_rules]: "(sctni, sctn::'n rvec sctn) \ \lv_rel\sctn_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" and [autoref_rules]: "((), trap) \ ghost_rel" notes [intro, simp] = list_set_rel_finiteD closed_ivl_rel[OF civl] closed_ivl_prod3_list_rel shows "(nres_of ?r, poincare_onto_from $ odo $ symstart $ trap $ S $ guards $ ivl $ sctn $ ro $ XS) \ \clw_rel appr1e_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto_from_def including art by autoref_monadic concrete_definition poincare_onto_from_impl for symstartd Si guardsi ivli sctni roi XSi uses poincare_onto_from_impl lemmas [autoref_rules] = poincare_onto_from_impl.refine sublocale autoref_op_pat_def poincare_onto_from . schematic_goal subset_spec1_impl: "(nres_of ?r, subset_spec1 R P dP) \ \bool_rel\nres_rel" if [autoref_rules]: "(Ri, R) \ appr1_rel" "(Pimpl, P) \ lvivl_rel" "(dPi, dP) \ \lvivl_rel\(default_rel UNIV)" unfolding subset_spec1_def including art by autoref_monadic lemmas [autoref_rules] = subset_spec1_impl[autoref_higher_order_rule] sublocale autoref_op_pat_def subset_spec1 . schematic_goal subset_spec1_collc: "(nres_of (?f), subset_spec1_coll R P dP) \ \bool_rel\nres_rel" if [autoref_rules]: "(Ri, R) \ clw_rel appr1_rel" "(Pimpl, P) \ lvivl_rel" "(dPi, dP) \ \lvivl_rel\(default_rel UNIV)" unfolding subset_spec1_coll_def including art by autoref_monadic concrete_definition subset_spec1_collc for Ri Pimpl dPi uses subset_spec1_collc lemmas subset_spec1_collc_refine[autoref_rules] = subset_spec1_collc.refine[autoref_higher_order_rule] sublocale autoref_op_pat_def subset_spec1_coll . schematic_goal one_step_until_time_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(X0i, X0::'n eucl1 set) \ appr1e_rel" assumes [autoref_rules]: "(phi, ph) \ bool_rel" assumes [autoref_rules]: "(t1i, t1) \ rnv_rel" notes [autoref_tyrel] = ty_REL[where 'a="real" and R="Id"] shows "(nres_of ?f, one_step_until_time odo X0 ph t1)\\appr1e_rel \\<^sub>r \clw_rel appr_rel\phantom_rel\nres_rel" unfolding one_step_until_time_def[abs_def] including art by autoref_monadic concrete_definition one_step_until_time_impl for odoi X0i phi t1i uses one_step_until_time_impl lemmas one_step_until_time_impl_refine[autoref_rules] = one_step_until_time_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def one_step_until_time . schematic_goal ivl_of_appr1_coll_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules]: "(Xi, X::'n::enum rvec set) \ clw_rel appr_rel" shows "(nres_of ?r, ivl_of_eucl_coll X) \ \appr1_rel\nres_rel" unfolding ivl_of_eucl_coll_def by autoref_monadic concrete_definition ivl_of_appr1_coll_impl uses ivl_of_appr1_coll_impl sublocale autoref_op_pat_def ivl_of_eucl_coll . lemmas ivl_of_appr1_coll_impl_refine[autoref_rules] = ivl_of_appr1_coll_impl.refine[autoref_higher_order_rule(1)] schematic_goal one_step_until_time_ivl_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(X0i, X0::'n eucl1 set) \ appr1e_rel" assumes [autoref_rules]: "(phi, ph) \ bool_rel" assumes [autoref_rules]: "(t1i, t1) \ rnv_rel" assumes [autoref_rules]: "(t2i, t2) \ rnv_rel" shows "(nres_of ?r, one_step_until_time_ivl odo X0 ph t1 t2) \ \appr1e_rel \\<^sub>r \clw_rel appr_rel\phantom_rel\nres_rel" unfolding one_step_until_time_ivl_def including art by autoref_monadic concrete_definition one_step_until_time_ivl_impl for X0i phi t1i t2i uses one_step_until_time_ivl_impl lemmas [autoref_rules] = one_step_until_time_ivl_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def one_step_until_time_ivl . schematic_goal poincare_onto_from_in_ivl_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS) \ clw_rel appr1e_rel" and [autoref_rules]: "(Si, S) \ \lv_rel\halfspaces_rel" and osctns[autoref_rules]: "(guardsi, guards) \ \clw_rel (iplane_rel lvivl_rel)\\<^sub>rreach_optns_rel\list_rel" and civl[autoref_rules]: "(ivli, ivl::'n rvec set) \ lvivl_rel" and csctns[autoref_rules]: "(sctni, sctn::'n rvec sctn) \ \lv_rel\sctn_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" assumes [autoref_rules]: "(symstarti, symstart::'n eucl1 set \ ('n rvec set \ 'n eucl1 set)nres) \ appr1e_rel \ \clw_rel appr_rel \\<^sub>r clw_rel appr1e_rel\nres_rel" assumes [unfolded autoref_tag_defs, refine_transfer]: "\X. TRANSFER (nres_of (symstartd X) \ symstarti X)" and [autoref_rules]: "((), trap) \ ghost_rel" "(Pimpl, P) \ lvivl_rel" "(dPi, dP) \ \lvivl_rel\(default_rel UNIV)" notes [intro, simp] = list_set_rel_finiteD closed_ivl_rel[OF civl] closed_ivl_prod3_list_rel shows "(nres_of ?r, poincare_onto_from_in_ivl $ odo $ symstart $ trap $ S $ guards $ ivl $ sctn $ ro $ XS $ P $ dP) \ \bool_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto_from_in_ivl_def including art by autoref_monadic concrete_definition poincare_onto_from_in_ivl_impl for E odoi symstartd Si guardsi ivli sctni roi XSi Pimpl dPi uses poincare_onto_from_in_ivl_impl lemmas [autoref_rules] = poincare_onto_from_in_ivl_impl.refine sublocale autoref_op_pat_def poincare_onto_from_in_ivl . lemma TRANSFER_I: "x \ TRANSFER x" by simp lemma dres_nres_rel_nres_relD: "(symstartd, symstart) \ A \ \B\dres_nres_rel \ (\x. nres_of (symstartd x), symstart) \ A \ \B\nres_rel" by (auto simp: dres_nres_rel_def nres_rel_def dest!: fun_relD) lemma c1_info_of_apprsI: assumes "(b, a) \ clw_rel appr1_rel" assumes "x \ a" shows "x \ c1_info_of_apprs b" using assms by (auto simp: appr1_rel_br clw_rel_br c1_info_of_apprs_def dest!: brD) lemma clw_rel_appr1_relI: assumes "\X. X \ set XS \ c1_info_invar CARD('n::enum) X" shows "(XS, c1_info_of_apprs XS::('n rvec\_)set) \ clw_rel appr1_rel" by (auto simp: appr1_rel_br clw_rel_br c1_info_of_apprs_def intro!: brI assms) lemma c1_info_of_appr'I: assumes "(b, a) \ \clw_rel appr1_rel\phantom_rel" assumes "x \ a" shows "x \ c1_info_of_appr' b" using assms by (auto simp add: c1_info_of_appr'_def intro!: c1_info_of_apprsI split: option.splits) lemma appr1e_relI: assumes "c1_info_invare CARD('n::enum) X0i" shows "(X0i, c1_info_of_appre X0i::'n eucl1 set) \ appr1e_rel" using assms apply (cases X0i) apply (auto simp: scaleR2_rel_def c1_info_of_appre_def c1_info_invare_def) apply (rule relcompI) apply (rule prod_relI) apply (rule IdI) apply (rule appr1_relI) apply (auto simp: vimage_def intro!: brI) apply (metis ereal_dense2 less_imp_le) apply (rule relcompI) apply (rule prod_relI) apply (rule IdI) apply (rule appr1_relI) apply (auto simp: vimage_def intro!: brI) by (metis basic_trans_rules(23) ereal_cases ereal_less_eq(1) ereal_top order_eq_refl) lemma c1_info_of_apprI: assumes "(b, a) \ appr1_rel" assumes "x \ a" shows "x \ c1_info_of_appr b" using assms apply (auto simp add: c1_info_of_appr_def c1_info_invar_def appr1_rel_internal appr_rel_def lv_rel_def set_rel_br dest!: brD split: option.splits) apply (auto simp add: appr_rell_internal dest!: brD) done lemma c1_info_of_appreI: assumes "(lub, a) \ appr1e_rel" assumes "x \ a" shows "x \ c1_info_of_appre lub" using assms apply (auto simp add: scaleR2_def c1_info_of_appre_def image_def vimage_def scaleR2_rel_def dest!: brD intro!: c1_info_of_apprsI split: option.splits) subgoal for a b c d e f g h i apply (rule exI[where x=g]) apply (rule conjI, assumption)+ apply (rule bexI) prefer 2 apply (rule c1_info_of_apprI) apply assumption apply assumption apply simp done done lemma c1_info_of_apprseI: assumes "(b, a) \ clw_rel appr1e_rel" assumes "x \ a" shows "x \ c1_info_of_apprse b" using assms by (force simp: appr1_rel_br scaleR2_rel_br clw_rel_br c1_info_of_appre_def c1_info_of_apprse_def dest!: brD) lemma clw_rel_appr1e_relI: assumes "\X. X \ set XS \ c1_info_invare CARD('n::enum) X" shows "(XS, c1_info_of_apprse XS::('n rvec\_)set) \ clw_rel appr1e_rel" using assms apply (auto simp: c1_info_of_apprse_def c1_info_of_appre_def c1_info_invare_def) unfolding appr1_rel_br scaleR2_rel_br clw_rel_br apply (rule brI) apply (auto simp: c1_info_invar_def vimage_def) subgoal premises prems for a b c d using prems(1)[OF prems(2)] by (cases a; cases b) auto done schematic_goal one_step_until_time_ivl_in_ivl_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n::enum rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(X0i, X0::'n eucl1 set) \ appr1e_rel" assumes [autoref_rules]: "(t1i, t1) \ rnv_rel" assumes [autoref_rules]: "(t2i, t2) \ rnv_rel" "(Ri, R) \ lvivl_rel" "(dRi, dR) \ \lvivl_rel\(default_rel UNIV)" shows "(nres_of ?r, one_step_until_time_ivl_in_ivl odo X0 t1 t2 R dR) \ \bool_rel\nres_rel" unfolding one_step_until_time_ivl_in_ivl_def including art by autoref_monadic concrete_definition one_step_until_time_ivl_in_ivl_impl for odoi X0i t1i t2i Ri dRi uses one_step_until_time_ivl_in_ivl_impl lemmas one_step_until_time_ivl_in_ivl_impl_refine[autoref_rules] = one_step_until_time_ivl_in_ivl_impl.refine[autoref_higher_order_rule(1 2 3)] sublocale autoref_op_pat_def one_step_until_time_ivl_in_ivl . schematic_goal poincare_onto_in_ivl_impl: assumes [autoref_rules_raw]: "DIM_precond TYPE('n::enum rvec) E" assumes [autoref_rules_raw]: "ncc_precond TYPE('n::enum rvec)" assumes [autoref_rules_raw]: "ncc_precond TYPE('n vec1)" assumes [autoref_rules]: "(odoi, odo) \ ode_ops_rel" assumes [autoref_rules]: "(XSi, XS) \ clw_rel appr1e_rel" and osctns[autoref_rules]: "(guardsi, guards) \ \clw_rel (iplane_rel lvivl_rel)\\<^sub>rreach_optns_rel\list_rel" and civl[autoref_rules]: "(ivli, ivl::'n rvec set) \ lvivl_rel" and csctns[autoref_rules]: "(sctni, sctn::'n rvec sctn) \ \lv_rel\sctn_rel" and [autoref_rules]: "(roi, ro) \ reach_optns_rel" "(Pimpl, P::'n rvec set) \ lvivl_rel" "(dPi, dP:: ((real, 'n) vec, 'n) vec set) \ \lvivl_rel\(default_rel UNIV)" notes [intro, simp] = list_set_rel_finiteD closed_ivl_rel[OF civl] closed_ivl_prod3_list_rel shows "(nres_of ?r, poincare_onto_in_ivl odo guards ivl sctn ro XS P dP) \ \bool_rel\nres_rel" unfolding autoref_tag_defs unfolding poincare_onto_in_ivl_def including art apply (rule autoref_monadicI) apply (autoref phases: id_op rel_inf fix_rel) apply (autoref_trans_step) apply (autoref_trans_step) apply (autoref_trans_step) apply (simp only: autoref_tag_defs) apply (rule poincare_onto_series_impl.refine[unfolded autoref_tag_defs])\ \TODO: why?\ apply fact+ apply (rule empty_symstart_impl) apply refine_transfer apply (rule ghost_relI) apply (autoref phases: trans) unfolding autoref_tag_defs by refine_transfer concrete_definition poincare_onto_in_ivl_impl for E odoi guardsi ivli sctni roi XSi Pimpl dPi uses poincare_onto_in_ivl_impl lemmas [autoref_rules] = poincare_onto_in_ivl_impl.refine[autoref_higher_order_rule(1 2 3)] subsection \Main (executable) interfaces to the ODE solver, with initialization\ definition "carries_c1 = Not o Option.is_none o (snd o snd)" definition "solves_poincare_map odo symstart S guards ivli sctni roi XS P dP \ poincare_onto_from_in_ivl_impl (D odo) (init_ode_ops True (carries_c1 (hd XS)) odo) symstart S guards ivli sctni roi XS P dP = dRETURN True" definition "solves_poincare_map' odo S = solves_poincare_map odo (\x. dRETURN ([], [x])) [S]" definition "one_step_until_time_ivl_in_ivl_check odo X t0 t1 Ri dRi \ one_step_until_time_ivl_in_ivl_impl (D odo) (init_ode_ops True (carries_c1 X) odo) X t0 t1 Ri dRi = dRETURN True" definition "solves_poincare_map_onto odo guards ivli sctni roi XS P dP \ poincare_onto_in_ivl_impl (D odo) (init_ode_ops True (carries_c1 (hd XS)) odo) guards ivli sctni roi XS P dP = dRETURN True" end context approximate_sets begin lemma c1_info_of_appre_c0_I: "(x, d) \ c1_info_of_appre ((1, 1), X0, None)" if "list_of_eucl x \ set_of_appr X0" using that by (force simp: c1_info_of_appre_def c1_info_of_appr_def) lemma lvivl'_invar_None[simp]: "lvivl'_invar n None" by (auto simp: lvivl'_invar_def) lemma c1_info_invar_None: "c1_info_invar n (u, None) \ length u = n" by (auto simp: c1_info_invar_def) lemma c1_info_invare_None: "c1_info_invare n ((l, u), x, None) \((l < u \ -\ < l \ l \ u \ u < \) \ length x = n)" by (auto simp: c1_info_invare_def Let_def c1_info_invar_None) end end \ No newline at end of file diff --git a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy @@ -1,2896 +1,2892 @@ theory Refine_Reachability_Analysis_C1 imports Abstract_Reachability_Analysis_C1 Refine_Reachability_Analysis begin lemma fst_flow1_of_vec1[simp]: "fst (flow1_of_vec1 x) = fst x" by (auto simp: flow1_of_vec1_def) lemma fst_vec1_of_flow[simp]: "fst (vec1_of_flow1 x) = fst x" by (auto simp: vec1_of_flow1_def) context approximate_sets_ode' begin lemma poincare_mapsto_scaleR2I: "poincare_mapsto P (scaleR2 x1 x2 baa) UNIV x1b (scaleR2 x1 x2 aca)" if "poincare_mapsto P (baa) UNIV x1b (aca)" using that apply (auto simp: poincare_mapsto_def scaleR2_def image_def vimage_def) apply (drule bspec, assumption) apply auto apply (rule exI, rule conjI, assumption) apply (rule exI, rule conjI, assumption, rule conjI, assumption) apply (rule bexI) prefer 2 apply assumption apply (auto simp: scaleR_blinfun_compose_right) done context includes ode_ops.lifting begin lemma var_safe_form_eq[simp]: "var.safe_form = safe_form" unfolding var.safe_form_def by transfer (auto simp: var_ode_ops_def safe_form_def) lemma var_ode_e: "var.ode_e = ode_e'" unfolding var.ode_e_def by transfer (auto simp: var_ode_ops_def) end lemma wd_imp_var_wd[refine_vcg, intro]: "wd (TYPE('n rvec)) \ var.wd (TYPE('n::enum vec1))" unfolding var.wd_def by (auto simp: wd_def length_concat o_def sum_list_distinct_conv_sum_set concat_map_map_index var_ode_e D_def ode_e'_def intro!: max_Var_floatariths_mmult_fa[le] max_Var_floatariths_mapI max_Var_floatarith_FDERIV_floatarith[le] max_Var_floatariths_fold_const_fa[le] max_Var_floatarith_le_max_Var_floatariths_nthI max_Var_floatariths_list_updateI max_Var_floatariths_replicateI) lemma safe_eq: assumes "wd TYPE('n::enum rvec)" shows "var.Csafe = ((Csafe \ UNIV)::'n vec1 set)" using assms var.wdD[OF wd_imp_var_wd[OF assms]] wdD[OF assms] unfolding var.safe_def safe_def var.wd_def wd_def var.Csafe_def Csafe_def unfolding ode_e'_def var_ode_e apply (auto simp: D_def) subgoal apply (subst interpret_form_max_Var_cong) prefer 2 apply assumption by (auto simp: nth_Basis_list_prod) subgoal for a b apply (drule isFDERIV_appendD1) apply simp apply simp apply (auto intro!: max_Var_floatariths_fold_const_fa[le])[] apply (rule isFDERIV_max_Var_congI, assumption) by (auto simp: nth_Basis_list_prod) subgoal apply (subst interpret_form_max_Var_cong) prefer 2 apply assumption by (auto simp: nth_Basis_list_prod) subgoal for a b apply (rule isFDERIV_appendI1) apply (rule isFDERIV_max_Var_congI, assumption) apply (auto simp: nth_Basis_list_prod) apply (auto simp: isFDERIV_def FDERIV_floatariths_def in_set_conv_nth isDERIV_inner_iff length_concat o_def sum_list_distinct_conv_sum_set concat_map_map_index intro!: isDERIV_FDERIV_floatarith isDERIV_mmult_fa_nth) apply (rule isDERIV_max_Var_floatarithI[where ys="list_of_eucl a"]) subgoal for i j k apply (cases "i < CARD('n)") subgoal by auto subgoal apply (rule isDERIV_max_VarI) apply (rule max_Var_floatarith_le_max_Var_floatariths_nthI) apply force apply auto done done subgoal for i j k l by (auto dest!: max_Var_floatariths_lessI simp: nth_Basis_list_prod) subgoal by (auto simp: nth_list_update) done done lemma var_ode_eq: fixes x::"'n::enum vec1" assumes "wd TYPE('n rvec)" and [simp]: "(fst x) \ Csafe" shows "var.ode x = (ode (fst x), matrix (ode_d1 (fst x)) ** snd x)" proof - have "interpret_floatariths ode_e (list_of_eucl x) = interpret_floatariths ode_e (list_of_eucl (fst x))" apply (rule interpret_floatariths_max_Var_cong) using wdD[OF \wd _\] by (auto simp: list_of_eucl_nth_if nth_Basis_list_prod inner_prod_def) moreover have "eucl_of_list (interpret_floatariths (mmult_fa D D D (concat (map (\j. map (\i. FDERIV_floatarith (ode_e ! j) [0.. index_Basis_list_nth length_list_of_eucl) done qed done ultimately show ?thesis unfolding var.ode_def ode_def unfolding ode_e'_def var_ode_e by (auto simp: wdD[OF \wd _\] ode_d1_def intro!: euclidean_eqI[where 'a="'n vec1"]) qed lemma var_existence_ivl_imp_existence_ivl: fixes x::"'n::enum vec1" assumes wd: "wd TYPE('n rvec)" assumes t: "t \ var.existence_ivl0 x" shows "t \ existence_ivl0 (fst x)" proof (rule existence_ivl_maximal_segment) from var.flow_solves_ode[OF UNIV_I var.mem_existence_ivl_iv_defined(2), OF t] have D: "(var.flow0 x solves_ode (\_. var.ode)) {0--t} (var.Csafe)" apply (rule solves_ode_on_subset) apply (rule var.closed_segment_subset_existence_ivl) apply (rule t) apply simp done show "((\t. fst (var.flow0 x t)) solves_ode (\_. ode)) {0--t} (Csafe)" using var.closed_segment_subset_existence_ivl[OF t] apply (auto simp: has_vderiv_on_def has_vector_derivative_def subset_iff intro!: solves_odeI derivative_eq_intros) apply (rule refl) apply (rule refl) apply (rule refl) apply (auto simp: var.flowderiv_def ) apply (subst var_ode_eq[OF wd(1)]) apply (auto simp: blinfun.bilinear_simps) subgoal for s using solves_odeD(2)[OF D, of s] by (subst(asm) (3) safe_eq[OF wd]) (auto ) subgoal for s using solves_odeD(2)[OF D, of s] by (subst(asm) (3) safe_eq[OF wd]) (auto ) done next show "fst (var.flow0 x 0) = fst x" apply (subst var.flow_initial_time) apply simp apply (rule var.mem_existence_ivl_iv_defined[OF t]) apply auto done qed simp lemma existence_ivl_imp_var_existence_ivl: fixes x::"'n::enum rvec" assumes wd: "wd TYPE('n rvec)" assumes t: "t \ existence_ivl0 x" shows "t \ var.existence_ivl0 ((x, W)::'n vec1)" proof (rule var.existence_ivl_maximal_segment) from flow_solves_ode[OF UNIV_I mem_existence_ivl_iv_defined(2), OF t] have D: "(flow0 x solves_ode (\_. ode)) {0--t} (Csafe)" apply (rule solves_ode_on_subset) apply (rule closed_segment_subset_existence_ivl) apply (rule t) apply simp done show "((\t. (flow0 x t, matrix (Dflow x t) ** W)) solves_ode (\_. var.ode)) {0--t} (var.Csafe)" using closed_segment_subset_existence_ivl[OF t] apply (auto simp: has_vderiv_on_def has_vector_derivative_def subset_iff intro!: solves_odeI derivative_eq_intros) apply (rule refl) apply (rule refl) apply (rule refl) apply (rule has_derivative_at_withinI) apply (rule Dflow_has_derivative) apply force apply (rule refl) apply (auto simp: flowderiv_def ) apply (subst var_ode_eq) apply (auto simp: blinfun.bilinear_simps matrix_blinfun_compose wd intro!: ext) subgoal for s h unfolding matrix_scaleR matrix_blinfun_compose matrix_mul_assoc matrix_scaleR_right .. subgoal for s using solves_odeD(2)[OF D, of s] safe_eq[OF wd] by auto done next have "x \ Csafe" by rule fact then show "(flow0 x 0, matrix (blinfun_apply (Dflow x 0)) ** W) = (x, W)" apply (auto ) apply (vector matrix_def matrix_matrix_mult_def axis_def) by (auto simp: if_distrib if_distribR cong: if_cong) qed auto theorem var_existence_ivl0_eq_existence_ivl0: fixes x::"'n::enum vec1" assumes wd: "wd TYPE('n rvec)" shows "var.existence_ivl0 (x::'n vec1) = existence_ivl0 (fst x)" apply safe subgoal by (rule var_existence_ivl_imp_existence_ivl[OF wd, of _ "x", simplified], simp) subgoal by (rule existence_ivl_imp_var_existence_ivl[OF wd, of _ "fst x" "snd x", unfolded prod.collapse]) done theorem var_flow_eq_flow_Dflow: fixes x::"'n::enum vec1" assumes wd: "wd TYPE('n rvec)" assumes t: "t \ var.existence_ivl0 x" shows "var.flow0 x t = vec1_of_flow1 (flow0 (fst x) t, Dflow (fst x) t o\<^sub>L blinfun_of_vmatrix (snd x)) " proof - have x: "x \ var.Csafe" by (rule var.mem_existence_ivl_iv_defined[OF t]) then have "fst x \ Csafe" by (subst (asm) safe_eq[OF wd]) auto then have sx[simp]: "(fst x) \ Csafe" by simp show ?thesis proof (rule var.flow_unique_on[OF t]) show "vec1_of_flow1 (flow0 (fst x) 0, Dflow (fst x) 0 o\<^sub>L blinfun_of_vmatrix (snd x)) = x" by (auto simp: vec1_of_flow1_def x) show "((\a. vec1_of_flow1 (flow0 (fst x) a, Dflow (fst x) a o\<^sub>L blinfun_of_vmatrix (snd x))) has_vderiv_on (\t. var.ode (vec1_of_flow1 (flow0 (fst x) t, Dflow (fst x) t o\<^sub>L blinfun_of_vmatrix (snd x))))) (var.existence_ivl0 x)" apply (auto simp: has_vderiv_on_def has_vector_derivative_def vec1_of_flow1_def at_within_open[OF _ var.open_existence_ivl] flowderiv_def intro!: derivative_eq_intros var_existence_ivl_imp_existence_ivl[OF wd] Dflow_has_derivative ext) apply (subst var_ode_eq[OF wd(1)]) apply (auto simp: blinfun.bilinear_simps) subgoal for t using flow_in_domain[of t "fst x"] by (simp add: var_existence_ivl_imp_existence_ivl[OF wd]) subgoal for t h by (simp add: matrix_blinfun_compose matrix_scaleR matrix_mul_assoc matrix_scaleR_right) done fix t assume "t \ var.existence_ivl0 x" then show "vec1_of_flow1 (flow0 (fst x) t, Dflow (fst x) t o\<^sub>L blinfun_of_vmatrix (snd x)) \ var.Csafe" by (subst safe_eq[OF wd]) (auto simp: vec1_of_flow1_def dest!: var_existence_ivl_imp_existence_ivl[OF wd] flow_in_domain) qed qed theorem flow_Dflow_eq_var_flow: fixes x::"'n::enum rvec" assumes wd: "wd TYPE('n rvec)" assumes t: "t \ existence_ivl0 x" shows "(flow0 x t, Dflow x t o\<^sub>L W) = flow1_of_vec1 (var.flow0 (x, matrix W) t::'n vec1)" using var_flow_eq_flow_Dflow[OF wd existence_ivl_imp_var_existence_ivl[OF wd t]] unfolding var_flow_eq_flow_Dflow[OF wd existence_ivl_imp_var_existence_ivl[OF wd t]] by (auto simp: flow1_of_vec1_def vec1_of_flow1_def) context includes blinfun.lifting begin lemma flow1_of_vec1_vec1_of_flow1[simp]: "flow1_of_vec1 (vec1_of_flow1 X) = X" unfolding vec1_of_flow1_def flow1_of_vec1_def by (transfer) auto end lemma var_flowpipe0_flowpipe: assumes wd: "wd TYPE('n::enum rvec)" assumes "var.flowpipe0 X0 hl hu (CX) X1" assumes "fst ` X0 \ Csafe" assumes "fst ` CX \ Csafe" assumes "fst ` X1 \ Csafe" shows "flowpipe (flow1_of_vec1 ` X0) hl hu (flow1_of_vec1 ` (CX::'n vec1 set)) (flow1_of_vec1 ` X1)" using assms unfolding flowpipe_def var.flowpipe0_def apply safe subgoal by (auto simp add: flow1_of_vec1_def vec1_of_flow1_def safe_eq[OF wd]) subgoal by (auto simp add: flow1_of_vec1_def vec1_of_flow1_def safe_eq[OF wd]) subgoal by (auto simp add: flow1_of_vec1_def vec1_of_flow1_def safe_eq[OF wd]) subgoal for x W y V h apply (drule bspec[where x="(y, V)"], force) apply (drule bspec, assumption) by (simp add: var_existence_ivl0_eq_existence_ivl0[OF wd] flow1_of_vec1_def) subgoal for x W y V h apply (drule bspec[where x="(y, V)"], force) apply (drule bspec, assumption) apply (subst flow_Dflow_eq_var_flow[OF wd], force simp: var_existence_ivl0_eq_existence_ivl0[OF wd] flow1_of_vec1_def) apply (rule imageI) by (simp add: vec1_of_flow1_def flow1_of_vec1_def) subgoal for x W y V h h' apply (drule bspec[where x="vec1_of_flow1 (x, W)"], force) apply (drule bspec, assumption) apply (subst flow_Dflow_eq_var_flow[OF wd]) apply (subst (asm) var_existence_ivl0_eq_existence_ivl0[OF wd]) apply (simp add: flow1_of_vec1_def) subgoal by (meson local.existence_ivl_initial_time local.mem_existence_ivl_iv_defined(1) local.mem_existence_ivl_iv_defined(2) mem_is_interval_1_I mvar.interval) subgoal apply (rule imageI) by (simp add: vec1_of_flow1_def flow1_of_vec1_def) done done theorem einterpret_solve_poincare_fas: assumes wd: "wd TYPE('n rvec)" assumes "length CXs = D + D*D" "n < D" assumes nz: "ode (fst (eucl_of_list CXs::'n vec1)) \ Basis_list ! n \ 0" shows "flow1_of_vec1 (einterpret (solve_poincare_fas n) CXs::'n::enum vec1) = (let (x, d) = flow1_of_vec1 (eucl_of_list CXs::'n vec1) in (x, d - (blinfun_scaleR_left (ode (x)) o\<^sub>L (blinfun_scaleR_left (inverse (ode x \ Basis_list ! n)) o\<^sub>L (blinfun_inner_left (Basis_list ! n) o\<^sub>L d)))))" using assms apply (auto intro!: simp: flow1_of_vec1_def solve_poincare_fas_def) subgoal apply (auto intro!: euclidean_eqI[where 'a="'n rvec"]) apply (subst eucl_of_list_prod) by (auto simp: eucl_of_list_prod length_concat o_def sum_list_distinct_conv_sum_set D_def Let_def wdD[OF wd] take_eq_map_nth) subgoal premises prems proof - have ode_e_eq: "interpret_floatarith (ode_e ! i) (map ((!) CXs) [0.. Csafe" shows "var.choose_step (X0::'n vec1 set) h \ SPEC (\(h', _, RES_ivl, RES::'n vec1 set). 0 < h' \ h' \ h \ flowpipe (flow1_of_vec1 ` X0) h' h' (flow1_of_vec1 ` RES_ivl) (flow1_of_vec1 ` RES))" apply refine_vcg apply (auto simp: ) apply (frule var.flowpipe0_safeD) apply (drule var_flowpipe0_flowpipe[rotated]) by (auto simp: safe_eq wd) lemma max_Var_floatariths_solve_poincare_fas[le]: assumes wd: "wd (TYPE('n::enum rvec))" shows "i < D \ max_Var_floatariths (solve_poincare_fas i) \ D + D * D" by (auto simp: solve_poincare_fas_def concat_map_map_index Let_def intro!: max_Var_floatariths_leI Suc_leI) (auto intro!: max_Var_floatarith_le_max_Var_floatariths_nthI max_Var_floatariths_ode_e_wd[OF wd] simp: wdD[OF wd]) lemma length_solve_poincare_fas[simp]: "length (solve_poincare_fas n) = D + D * D" by (auto simp: solve_poincare_fas_def length_concat o_def sum_list_distinct_conv_sum_set D_def Let_def) theorem interpret_floatariths_solve_poincare_fas: assumes wd: "wd TYPE('n::enum rvec)" assumes "length CXs = D + D*D" "n < D" assumes nz: "ode (fst (eucl_of_list CXs::'n vec1)) \ Basis_list ! n \ 0" shows "interpret_floatariths (solve_poincare_fas n) CXs = list_of_eucl (vec1_of_flow1 (let (x, d) = flow1_of_vec1 (eucl_of_list CXs::'n vec1) in (x, d - (blinfun_scaleR_left (ode (x)) o\<^sub>L (blinfun_scaleR_left (inverse (ode x \ Basis_list ! n)) o\<^sub>L (blinfun_inner_left (Basis_list ! n) o\<^sub>L d))))))" using arg_cong[where f="list_of_eucl::'n vec1 \ _", OF arg_cong[where f=vec1_of_flow1, OF einterpret_solve_poincare_fas[OF assms]]] apply (auto simp: ) apply (subst (asm) list_of_eucl_eucl_of_list) apply (auto simp: ) apply (auto simp: wdD[OF wd]) done lemma length_solve_poincare_slp[simp]: "length solve_poincare_slp = D" by (auto simp: solve_poincare_slp_def) lemma ne_zero_lemma: assumes "ode ` fst ` CX \ FC" "\b\FC. b \ n \ 0" "(a, b) \ CX" "ode a \ n = 0" shows "False" proof - have "(a, b) \ CX" by fact then have "ode (fst (a, b)) \ ode ` fst ` CX" by blast also have "\ \ FC" by fact finally have "ode a \ FC" by simp with assms show False by auto qed lemma ne_zero_lemma2: assumes "ode ` fst ` flow1_of_vec1 ` env \ F" "\x\F. x \ n \ 0" "(a, b) \ env" "flow1_of_vec1 (a, b) = (a', b')" "ode a' \ n = 0" shows False proof - have "(a', b') \ flow1_of_vec1 ` env" apply (rule image_eqI) using assms by auto then have "ode (fst (a', b')) \ ode ` fst ` \" by blast also from assms have "\ \ F" by simp finally have "ode a' \ F" by simp with assms have "ode a' \ n \ 0" by auto with assms show False by simp qed lemma solve_poincare_plane[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" assumes "n \ Basis" shows "solve_poincare_plane (n::'n::enum rvec) CX \ SPEC (\PDP. fst ` PDP \ Csafe \ (\(x, d) \ CX. (x, d - (blinfun_scaleR_left (ode x) o\<^sub>L (blinfun_scaleR_left (inverse (ode x \ n)) o\<^sub>L (blinfun_inner_left n o\<^sub>L d)))) \ PDP) \ (\(x, d) \ PDP. ode x \ n \ 0))" unfolding solve_poincare_plane_def apply (refine_vcg) subgoal using assms by auto subgoal using assms by auto subgoal using assms by auto subgoal using assms by (auto simp: solve_poincare_slp_def) subgoal using assms by auto subgoal for C1 FC _ CX' CX'' P P1 FP _ apply auto apply (drule bspec, assumption) apply (rule image_eqI) prefer 2 apply assumption apply (subst einterpret_solve_poincare_fas) subgoal using wd by auto subgoal using wd by auto subgoal using wd by auto subgoal using wd assms by (auto elim!: ne_zero_lemma) subgoal using wd assms by (auto simp: ) done subgoal by (auto elim!: ne_zero_lemma2) done lemma choose_step1_flowpipe[le, refine_vcg]: assumes wd[refine_vcg]: "wd TYPE('n::enum rvec)" shows "choose_step1 (X0::'n eucl1 set) h \ SPEC (\(h', _, RES_ivl, RES::'n eucl1 set). 0 < h' \ h' \ h \ flowpipe X0 h' h' RES_ivl RES)" using assms unfolding choose_step1_def by (refine_vcg choose_step'_flowpipe[le] wd) (auto simp: image_image, auto simp: safe_eq vec1_of_flow1_def flowpipe0_imp_flowpipe env_len_def) lemma image_flow1_of_vec1I: "vec1_of_flow1 x \ X \ x \ flow1_of_vec1 ` X" by (rule image_eqI) (rule flow1_of_vec1_vec1_of_flow1[symmetric]) lemma inter_sctn1_spec[le, refine_vcg]: "inter_sctn1_spec X sctn \ SPEC (\(R, S). X \ plane_of sctn \ UNIV \ R \ fst ` R \ plane_of sctn \ X \ plane_of sctn \ UNIV \ S \ fst ` S \ plane_of sctn)" unfolding inter_sctn1_spec_def apply (refine_vcg, auto) subgoal by (rule image_flow1_of_vec1I) (auto simp: plane_of_def inner_prod_def) subgoal by (auto simp: plane_of_def inner_prod_def) subgoal by (rule image_flow1_of_vec1I) (force simp: set_plus_def plane_of_def inner_prod_def vec1_of_flow1_def) subgoal by (force simp: set_plus_def) done lemma fst_safe_coll[le, refine_vcg]: "wd TYPE('a) \ fst_safe_coll (X::('a::executable_euclidean_space*'c) set) \ SPEC (\R. R = fst ` X \ fst ` X \ Csafe)" unfolding fst_safe_coll_def by refine_vcg lemma vec1reps[THEN order_trans, refine_vcg]: "vec1reps CX \ SPEC (\R. case R of None \ True | Some X \ X = vec1_of_flow1 ` CX)" unfolding vec1reps_def apply (refine_vcg FORWEAK_mono_rule[where I="\XS R. case R of None \ True | Some R \ vec1_of_flow1 ` (\XS) \ R \ R \ vec1_of_flow1 ` CX"]) by (auto simp: split: option.splits) force+ lemma nonzero_component_within[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "nonzero_component_within ivl sctn (PDP::'n eucl1 set) \ SPEC (\b. (b \ (\x\PDP. fst x \ ivl \ (\\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl))) \ fst ` PDP \ Csafe \ (\x\PDP. ode (fst x) \ normal sctn \ 0))" unfolding nonzero_component_within_def by refine_vcg auto lemma do_intersection_invar_inside: "do_intersection_invar guards b ivl sctn X (e, f, m, n, p, q, True) \ fst ` e \ sabove_halfspace sctn \ fst ` mn \ ivl \ mn = m \ mn = n \ do_intersection_spec UNIV guards ivl sctn X (mn, p)" subgoal premises prems proof - from prems have e: "e \ sbelow_halfspace sctn \ UNIV = {}" by (auto simp: halfspace_simps plane_of_def) with prems(1) have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} X UNIV p m" "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} X UNIV p n" "e \ sbelow_halfspace sctn \ UNIV = {}" "fst ` X \ b = {}" "fst ` X \ sbelow_halfspace sctn" "ivl \ plane (normal sctn) (pstn sctn)" "fst ` X \ p" "fst ` m \ Csafe" "fst ` n \ Csafe" "p \ Csafe" "fst ` e \ Csafe" "f \ {0..}" "p \ sbelow_halfspace sctn - guards" "e \ (- guards) \ UNIV" "fst ` (m \ n) \ guards = {}" "0 \ (\x. ode x \ normal sctn) ` fst ` (m \ n)" "\x\m \ n. \\<^sub>F x in at (fst x) within plane (normal sctn) (pstn sctn). x \ ivl" by (auto simp: do_intersection_invar_def do_intersection_spec_def plane_of_def) then show ?thesis using prems(2-) by (auto simp: do_intersection_spec_def plane_of_def halfspace_simps) qed done lemma do_intersection_body_lemma: assumes "flowsto A T (i \ UNIV) (X' \ sbelow_halfspace sctn \ UNIV)" "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} B UNIV i PS " "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} B UNIV i PS2" "T \ {0..}" "i \ sbelow_halfspace sctn - guards" "fst ` (A \ B) \ sbelow_halfspace sctn" "fst ` PS \ Csafe " "fst ` PS2 \ Csafe " \X = A \ B\ assumes ivl: "closed ivl" "ivl \ plane_of sctn" assumes normal_Basis: "\normal sctn\ \ Basis" and inter_empties: "fst ` Y \ GUARDS = {}" "fst ` CX' \ GUARDS = {}" "fst ` PDP' \ GUARDS = {}" "fst ` PDP'' \ GUARDS = {}" and h': "0 < h'" "h' \ h" and safe: "fst ` PDP \ Csafe" "fst ` CX' \ Csafe" "fst ` PDP' \ Csafe" "fst ` PDP'' \ Csafe" and PDP: "\(x,d)\CX'. (x, d - (blinfun_scaleR_left (ode x) o\<^sub>L (blinfun_scaleR_left (inverse (ode x \ \normal sctn\)) o\<^sub>L (blinfun_inner_left \normal sctn\ o\<^sub>L d)))) \ PDP" and PDP': "PDP \ plane_of sctn \ UNIV \ PDP'" and PDP'': "PDP \ plane_of sctn \ UNIV \ PDP''" and evin: "\x\PDP'. fst x \ ivl \ (\\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl)" "\x\PDP''. fst x \ ivl \ (\\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl)" and through: "\(x, d)\PDP. ode x \ \normal sctn\ \ 0" "\x\PDP'. ode (fst x) \ normal sctn \ 0" "\x\PDP''. ode (fst x) \ normal sctn \ 0" and plane: "fst ` PDP' \ plane_of sctn" "fst ` PDP'' \ plane_of sctn" and flowpipe: "flowpipe X' h' h' CX' Y" shows "\A B. X = A \ B \ flowsto A {0<..} ((fst ` CX' \ sbelow_halfspace sctn \ i) \ UNIV) (Y \ sbelow_halfspace sctn \ UNIV) \ poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} B UNIV (fst ` CX' \ sbelow_halfspace sctn \ i) (PDP' \ PS) \ poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} B UNIV (fst ` CX' \ sbelow_halfspace sctn \ i) (PDP'' \ PS2)" proof - from flowpipe have 1: "flowpipe (X' \ (sbelow_halfspace sctn) \ UNIV) h' h' CX' Y" by (rule flowpipe_subset) (use flowpipe in \auto dest!: flowpipe_safeD\) have 2: "fst ` (X' \ (sbelow_halfspace sctn) \ UNIV) \ {x. pstn sctn \ x \ normal sctn} = {}" by (auto simp: halfspace_simps plane_of_def) from normal_Basis have 3: "normal sctn \ 0" by (auto simp: ) note 4 = \closed ivl\ from \ivl \ plane_of sctn\ have 5: "ivl \ plane (normal sctn) (pstn sctn)" by (auto simp: plane_of_def) have 6: "(x, d) \ CX' \ x \ plane (normal sctn) (pstn sctn) \ (x, d - (blinfun_scaleR_left (ode x) o\<^sub>L (blinfun_scaleR_left (inverse (ode x \ normal sctn)) o\<^sub>L (blinfun_inner_left (normal sctn) o\<^sub>L d)))) \ PDP' \ PDP''" for x d unfolding PDP_abs_lemma[OF normal_Basis] apply (drule PDP[rule_format, of "(x, d)", unfolded split_beta' fst_conv snd_conv]) using PDP' PDP'' by (auto simp: plane_of_def) from normal_Basis through have 7: "(x, d) \ PDP' \ ode x \ normal sctn \ 0" for x d by (auto elim!: abs_in_BasisE) have 8: "(x, d) \ PDP' \ x \ ivl" for x d using evin by auto have 9: "(x, d) \ PDP' \ \\<^sub>F x in at x within plane (normal sctn) (pstn sctn). x \ ivl" for x d using evin by (auto simp add: plane_of_def) obtain X1 X2 where X1X2: "X' \ sbelow_halfspace sctn \ UNIV = X1 \ X2" and X1: "flowsto X1 {0<..h'} (CX' \ {x. x \ normal sctn < pstn sctn} \ UNIV) (CX' \ {x \ ivl. x \ normal sctn = pstn sctn} \ UNIV)" and X2: "flowsto X2 {h'..h'} (CX' \ {x. x \ normal sctn < pstn sctn} \ UNIV) (Y \ {x. x \ normal sctn < pstn sctn} \ UNIV)" and P: "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} X1 UNIV (fst ` CX' \ {x. x \ normal sctn < pstn sctn}) (PDP' \ PDP'')" by (rule flowpipe_split_at_above_halfspace[OF 1 2 3 4 5 6 7 8 9]) (auto simp: Ball_def) from \flowsto A _ _ _\[unfolded X1X2] obtain p1 p2 where p1p2: "A = p1 \ p2" and p1: "flowsto p1 T (i \ UNIV) X1" and p2: "flowsto p2 T (i \ UNIV) X2" by (rule flowsto_unionE) have "A \ B = p2 \ (p1 \ B)" using \A = p1 \ p2\ by auto moreover from flowsto_trans[OF p2 X2] have "flowsto p2 {0<..} ((fst ` CX' \ (sbelow_halfspace sctn) \ i) \ UNIV) (Y \ (sbelow_halfspace sctn) \ UNIV)" apply (rule flowsto_subset) subgoal by (auto simp: halfspace_simps) subgoal using h' \T \ _\ by (auto simp: halfspace_simps intro!: add_nonneg_pos) subgoal using flowpipe_source_subset[OF 1, unfolded X1X2] X1X2 apply auto by (auto simp: halfspace_simps) subgoal by (auto simp: halfspace_simps) done moreover have cls: "closed {x \ ivl. x \ normal sctn = pstn sctn}" by (rule closed_levelset_within continuous_intros \closed ivl\)+ from flowsto_trans[OF p1 X1] have ftt: "flowsto p1 ({s + t |s t. s \ T \ t \ {0<..h'}}) (i \ UNIV \ CX' \ {x. x \ normal sctn < pstn sctn} \ UNIV \ X1 \ X1) (X1 - X1 \ CX' \ {x \ ivl. x \ normal sctn = pstn sctn} \ UNIV)" by auto from X1X2 have X1_sb: "X1 \ sbelow_halfspace sctn \ UNIV" by auto have "{x \ ivl. x \ normal sctn = pstn sctn} \ UNIV \ (i \ UNIV \ CX' \ {x. x \ normal sctn < pstn sctn} \ UNIV \ X1) = {}" apply (intro Int_Un_eq_emptyI) subgoal using \i \ sbelow_halfspace sctn - guards\ by (auto simp: halfspace_simps) subgoal by (auto simp: halfspace_simps) subgoal using X1_sb by (auto simp: halfspace_simps) done then have inter_empty: "{x \ ivl. x \ normal sctn = pstn sctn} \ UNIV \ (i \ UNIV \ CX' \ {x. x \ normal sctn < pstn sctn} \ UNIV \ X1 \ X1) = {}" by auto have p1ret: "returns_to {x \ ivl. x \ normal sctn = pstn sctn} x" and p1pm: "poincare_map {x \ ivl. x \ normal sctn = pstn sctn} x \ fst ` (PDP' \ PDP'')" if "(x, d) \ p1" for x d apply (rule flowsto_poincareD[OF ftt _ inter_empty _ _ _ order_refl]) subgoal by auto subgoal by fact subgoal using \T \ _\ by auto subgoal using that by auto subgoal apply (rule flowsto_poincareD[OF ftt _ inter_empty]) subgoal by auto subgoal by fact subgoal using \T \ _\ by auto subgoal using that by auto subgoal using 6 by force done done have crt: "isCont (return_time {x \ ivl. x \ normal sctn - pstn sctn = 0}) x" if "(x, d) \ p1" for x d apply (rule return_time_isCont_outside[where Ds="\_. blinfun_inner_left (normal sctn)"]) subgoal by (simp add: p1ret[OF that]) subgoal by fact subgoal by (auto intro!: derivative_eq_intros) subgoal by simp subgoal apply simp using p1pm[OF that] by (auto dest!: 7) subgoal using p1pm[OF that] by (auto dest!: 9 simp: eventually_at_filter) subgoal using \fst ` (A \ B) \ sbelow_halfspace sctn\ that p1p2 by (auto simp: halfspace_simps) done have pmij: "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} p1 UNIV (fst ` (i \ UNIV \ X1) \ fst ` CX' \ {x. x \ normal sctn < pstn sctn}) (PDP' \ PDP'')" apply (rule flowsto_poincare_trans[OF \flowsto _ _ _ X1\ P]) subgoal using \T \ {0..}\ by auto subgoal by auto subgoal using \i \ sbelow_halfspace sctn - guards\ X1X2 by (force simp: halfspace_simps) subgoal by fact subgoal for x d using crt by simp subgoal by auto done from pmij have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} p1 UNIV (fst ` (i \ UNIV \ X1) \ fst ` CX' \ {x. x \ normal sctn < pstn sctn}) PDP'" apply (rule poincare_mapsto_subset) using \fst ` PDP' \ Csafe\ by auto from this \poincare_mapsto _ _ _ i PS\ have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} (p1 \ B) UNIV ((fst ` (i \ UNIV \ X1) \ fst ` CX' \ {x. x \ normal sctn < pstn sctn}) \ i) (PDP' \ PS)" by (intro poincare_mapsto_unionI) (auto simp: plane_of_def) then have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} (p1 \ B) UNIV (fst ` CX' \ sbelow_halfspace sctn \ i) (PDP' \ PS)" apply (rule poincare_mapsto_subset) subgoal by auto subgoal by auto subgoal using flowpipe_source_subset[OF 1, unfolded X1X2] X1X2 apply (auto simp: halfspace_simps subset_iff) done subgoal using safe \fst ` PS \ Csafe\ by auto done moreover from pmij have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} p1 UNIV (fst ` (i \ UNIV \ X1) \ fst ` CX' \ {x. x \ normal sctn < pstn sctn}) PDP''" apply (rule poincare_mapsto_subset) using \fst ` PDP'' \ Csafe\ by auto from this \poincare_mapsto _ _ _ i PS2\ have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} (p1 \ B) UNIV ((fst ` (i \ UNIV \ X1) \ fst ` CX' \ {x. x \ normal sctn < pstn sctn}) \ i) (PDP'' \ PS2)" by (intro poincare_mapsto_unionI) (auto simp: plane_of_def) then have "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} (p1 \ B) UNIV (fst ` CX' \ sbelow_halfspace sctn \ i) (PDP'' \ PS2)" apply (rule poincare_mapsto_subset) subgoal by auto subgoal by auto subgoal using flowpipe_source_subset[OF 1, unfolded X1X2] X1X2 apply (auto simp: halfspace_simps subset_iff) done subgoal using safe \fst ` PS2 \ Csafe\ by auto done ultimately show ?thesis unfolding \X = A \ B\ by blast qed lemma do_intersection_body_spec: fixes guards::"'n::enum rvec set" assumes invar: "do_intersection_invar guards GUARDS ivl sctn X (X', T, PS, PS2, i, True, True)" and wdp[refine_vcg]: "wd TYPE('n rvec)" and X: "fst ` X \ Csafe" and ivl: "closed ivl" and GUARDS: "guards \ GUARDS" shows "do_intersection_body GUARDS ivl sctn h (X', T, PS, PS2, i, True, True) \ SPEC (do_intersection_invar guards GUARDS ivl sctn X)" proof - from invar obtain A B where AB: "fst ` (A \ B) \ GUARDS = {} " "fst ` (A \ B) \ sbelow_halfspace sctn " "ivl \ plane_of sctn " "fst ` (A \ B) \ i " "fst ` PS \ Csafe " "fst ` PS2 \ Csafe " "i \ Csafe " "fst ` X' \ Csafe " "T \ {0..}" "i \ sbelow_halfspace sctn - guards " "X' \ (- guards) \ UNIV " "fst ` (PS \ PS2) \ guards = {} " "0 \ (\x. ode x \ normal sctn) ` fst ` (PS \ PS2) " "\x\PS \ PS2. \\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl " "X = A \ B " "flowsto A T (i \ UNIV) (X' \ sbelow_halfspace sctn \ UNIV)" "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} B UNIV i PS " "poincare_mapsto {x \ ivl. x \ normal sctn = pstn sctn} B UNIV i PS2" by (auto simp: do_intersection_invar_def) have ev_in_ivl: "\\<^sub>F x in at p within plane_of sctn. x \ ivl" if \\x\d. fst x \ ivl \ (\\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl)\ \\x\e. fst x \ ivl \ (\\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl)\ \(p, q) \ d \ (p, q) \ PS \ (p, q) \ e \ (p, q) \ PS2\ for p q d e using \\x\PS \ PS2. \\<^sub>F x in at (fst x) within plane_of sctn. x \ ivl\ using that by (auto dest!: bspec[where x="(p, q)"]) show ?thesis unfolding do_intersection_body_def do_intersection_invar_def apply simp apply (refine_vcg, clarsimp_all) subgoal using AB by (auto simp: ) subgoal using AB by (auto simp: ) subgoal using AB by (auto simp: ) subgoal apply (rule conjI) subgoal using AB by auto\ \unnecessarily slow\ subgoal using AB by fastforce done subgoal using AB by (auto simp: ) subgoal using AB by (auto simp: ) subgoal using AB by (auto simp: ) subgoal by (auto dest!: flowpipe_safeD) subgoal apply safe subgoal using AB GUARDS by auto subgoal using AB by auto subgoal using AB by auto subgoal using AB GUARDS by auto subgoal using AB by auto subgoal using AB by auto done subgoal using AB GUARDS by auto subgoal using AB GUARDS by auto\ \unnecessarily slow\ subgoal using AB GUARDS by auto subgoal using AB assms by (auto intro: ev_in_ivl) subgoal using AB assms apply - by (rule do_intersection_body_lemma) done qed lemma do_intersection_spec[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "do_intersection guards ivl sctn (X::'n eucl1 set) h \ SPEC (\(inside, P, P2, CX). (inside \ (do_intersection_spec UNIV guards ivl sctn X (P, CX) \ do_intersection_spec UNIV guards ivl sctn X (P2, CX)) \ fst ` X \ CX))" using assms unfolding do_intersection_def autoref_tag_defs apply (refine_vcg, clarsimp_all) subgoal unfolding do_intersection_invar_def apply clarsimp apply (intro conjI) apply force apply force apply force apply (rule exI[where x=X]) apply (rule exI[where x="{}"]) by (auto intro!: flowsto_self) subgoal by (rule do_intersection_body_spec) subgoal by (rule do_intersection_invar_inside, assumption) auto subgoal by (rule do_intersection_invar_inside, assumption) auto subgoal by (auto simp: plane_of_def halfspace_simps do_intersection_invar_def) done lemma mem_flow1_of_vec1_image_iff[simp]: "(c, d) \ flow1_of_vec1 ` a \ vec1_of_flow1 (c, d) \ a" by force lemma mem_vec1_of_flow1_image_iff[simp]: "(c, d) \ vec1_of_flow1 ` a \ flow1_of_vec1 (c, d) \ a" by force lemma split_spec_param1[le, refine_vcg]: "split_spec_param1 X \ SPEC (\(A, B). X \ A \ B)" unfolding split_spec_param1_def apply (refine_vcg) apply (auto simp add: subset_iff split: option.splits) by (metis flow1_of_vec1_vec1_of_flow1 surjective_pairing) lemma do_intersection_spec_empty: "X = {} \ Y = {} \ do_intersection_spec S sctns ivl sctn X ({}, Y)" by (auto simp: do_intersection_spec_def halfspaces_union) lemma do_intersection_spec_subset: "do_intersection_spec S osctns ivl csctns Y (a, b) \ X \ Y \ do_intersection_spec S osctns ivl csctns X (a, b)" by (auto simp: do_intersection_spec_def halfspaces_union intro: flowsto_subset poincare_mapsto_subset) lemma do_intersection_spec_union: "do_intersection_spec S osctns ivl csctns a (b, c) \ do_intersection_spec S osctns ivl csctns f (g, h) \ do_intersection_spec S osctns ivl csctns (a \ f) (b \ g, c \ h)" by (auto simp: do_intersection_spec_def intro!: poincare_mapsto_unionI) lemma scaleR2_rep_of_coll[le, refine_vcg]: "scaleR2_rep_coll X \ SPEC (\((l, u), Y). X \ scaleR2 l u Y)" unfolding scaleR2_rep_coll_def apply (refine_vcg FORWEAK_mono_rule[where I="\Xs ((l, u), Y). \Xs \ scaleR2 l u Y"]) subgoal by (auto intro: scaleR2_subset) subgoal apply clarsimp apply safe subgoal by (auto elim: scaleR2_subset) subgoal apply (rule set_rev_mp, assumption) apply (rule order_trans) apply (rule Union_upper, assumption) apply (rule order_trans, assumption) apply (rule subsetI) apply (erule scaleR2_subset) by (auto ) done done lemma split_spec_param1e[le, refine_vcg]: "split_spec_param1e X \ SPEC (\(A, B). X \ A \ B)" unfolding split_spec_param1e_def apply (refine_vcg) apply clarsimp apply (thin_tac "_ \ {}") apply (auto simp: scaleR2_def vimage_def image_def) apply (rule exI, rule conjI, assumption, rule conjI, assumption) apply (auto simp: split_beta') apply (drule_tac x = x in spec) apply auto by (metis (no_types, lifting) UnE prod.sel(1) prod.sel(2) subset_eq) lemma reduce_spec1[le, refine_vcg]: "reduce_spec1 ro X \ SPEC (\R. X \ R)" unfolding reduce_spec1_def by refine_vcg auto lemma reduce_spec1e[le, refine_vcg]: "reduce_spec1e ro X \ SPEC (\R. X \ R)" unfolding reduce_spec1e_def by refine_vcg (auto simp: scaleR2_def image_def vimage_def, force) lemma split_under_threshold[le, refine_vcg]: "split_under_threshold ro th X \ SPEC (\R. X \ R)" unfolding split_under_threshold_def autoref_tag_defs by (refine_vcg) auto lemma step_split[le, refine_vcg]: "wd TYPE((real, 'n::enum) vec) \ step_split ro (X::'n eucl1 set) \ SPEC (\Y. X \ Y \ fst ` Y \ Csafe)" unfolding step_split_def by (refine_vcg refine_vcg) auto lemma tolerate_error_SPEC[THEN order_trans, refine_vcg]: "tolerate_error Y E \ SPEC (\b. True)" unfolding tolerate_error_def by refine_vcg lemma flowpipe_scaleR2I: "flowpipe (scaleR2 x1 x2 bc) x1a x1a (fst ` aca \ UNIV) (scaleR2 x1 x2 bca)" if "flowpipe (bc) x1a x1a (fst ` aca \ UNIV) (bca)" using that apply (auto simp: flowpipe_def scaleR2_def) apply (drule bspec, assumption) apply (auto simp: image_def vimage_def ) apply (rule exI, rule conjI, assumption, rule conjI, assumption) apply (rule bexI) prefer 2 apply assumption by (auto simp: scaleR_blinfun_compose_right) lemma choose_step1e_flowpipe[le, refine_vcg]: assumes vwd[refine_vcg]: "wd TYPE('n::enum rvec)" shows "choose_step1e (X0::'n eucl1 set) h \ SPEC (\(h', _, RES_ivl, RES::'n eucl1 set). 0 < h' \ h' \ h \ flowpipe X0 h' h' (RES_ivl \ UNIV) RES)" unfolding choose_step1e_def apply (refine_vcg) apply (auto intro: flowpipe_scaleR2I) apply (erule contrapos_np) apply (auto intro!: flowpipe_scaleR2I) apply (rule flowpipe_subset) apply assumption apply (auto dest!: flowpipe_safeD) done lemma width_spec_appr1[THEN order_trans, refine_vcg]: "width_spec_appr1 X \ SPEC (\_. True)" unfolding width_spec_appr1_def by refine_vcg lemma tolerate_error1_SPEC[THEN order_trans, refine_vcg]: "tolerate_error1 Y E \ SPEC (\b. True)" unfolding tolerate_error1_def by refine_vcg lemma step_adapt_time[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "step_adapt_time (X::'n eucl1 set) h \ SPEC (\(t, CX, X1, h). flowpipe X t t (CX \ UNIV) X1)" unfolding step_adapt_time_def autoref_tag_defs apply (refine_vcg refine_vcg, clarsimp) apply (auto simp: flowpipe_def) apply force done lemma resolve_step[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "resolve_step roptns (X::'n::enum eucl1 set) h \ SPEC (\(_, CX, X1, _). flowsto X {0..} (CX \ UNIV) X1 \ X \ X1 \ CX \ UNIV \ X1 \ CX \ UNIV \ Csafe \ UNIV)" unfolding resolve_step_def autoref_tag_defs apply (refine_vcg refine_vcg) subgoal by (rule flowsto_self) auto subgoal by auto subgoal by auto subgoal apply clarsimp apply (frule flowpipe_imp_flowsto_nonneg) apply (rule flowsto_subset, assumption) by auto subgoal by (auto dest: flowpipe_source_subset) subgoal by (auto dest!: flowpipe_safeD) done lemma pre_intersection_step[THEN order_trans, refine_vcg]: "pre_intersection_step ro (X::'n eucl1 set) h \ SPEC (\(X', CX, G). X \ X' \ G \ X \ X' \ G \ CX \ UNIV)" if [refine_vcg]: "wd TYPE('n::enum rvec)" unfolding pre_intersection_step_def autoref_tag_defs by (refine_vcg) auto lemma [THEN order_trans, refine_vcg]: "select_with_inter ci a \ SPEC (\_. True)" unfolding select_with_inter_def by (refine_vcg FORWEAK_mono_rule[where I="\_ _. True"]) lemmas [refine_vcg del] = scaleR2_rep_of_coll lemma fst_scaleR2_image[simp]: "ad \ ereal r \ ereal r \ bd \ fst ` scaleR2 ad bd be = fst ` be" by (cases ad; cases bd; force simp: scaleR2_def image_image split_beta' vimage_def) lemma scaleR2_rep_of_coll2[le, refine_vcg]: "scaleR2_rep_coll X \ SPEC (\((l, u), Y). X \ scaleR2 l u Y \ fst ` X = fst ` Y)" unfolding scaleR2_rep_coll_def supply [simp del] = mem_scaleR2_union apply (refine_vcg FORWEAK_mono_rule[where I="\Xs ((l, u), Y). \Xs \ scaleR2 l u Y \ fst ` \Xs \ fst ` Y \ fst ` Y \ fst ` X"]) apply (auto intro: scaleR2_subset) subgoal by (auto simp: scaleR2_def) subgoal by (auto simp: scaleR2_def image_def vimage_def, fastforce) subgoal apply (rule scaleR2_subset) apply (rule subsetD) apply assumption apply auto done subgoal by force subgoal for a b c d e f g h i j k l apply (rule scaleR2_subset) apply (rule subsetD) apply assumption by auto subgoal by (auto simp: scaleR2_def) subgoal by (auto simp: scaleR2_def) subgoal by (auto simp: scaleR2_def image_def vimage_def, fastforce) done lemma reach_cont[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "reach_cont roptns guards (X::'n eucl1 set) \ SPEC (\(CX, G). G \ (CX \ UNIV) \ (Csafe - guards) \ UNIV \ X \ G \ CX \ UNIV \ flowsto X {0..} (CX \ UNIV) G)" using [[simproc del: defined_all]] unfolding reach_cont_def autoref_tag_defs apply (refine_vcg, clarsimp_all simp add: cancel_times_UNIV_subset) subgoal by (rule flowsto_self) (auto simp: ) subgoal by (force simp: scaleR2_def) subgoal by (fastforce simp: scaleR2_def vimage_def image_def) subgoal premises prems for _ _ _ _ _ _ _ g using \flowsto X _ _ (g \ _ \ _)\ \flowsto g _ _ _\ apply (rule flowsto_stepI) using prems by auto subgoal apply safe subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto done subgoal by auto subgoal by (rule flowsto_subset, assumption) auto subgoal apply safe subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by fastforce subgoal by auto subgoal by auto subgoal by (metis (mono_tags, lifting) Diff_eq_empty_iff Diff_iff IntI) done subgoal apply safe subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto subgoal by auto done subgoal by auto done lemma reach_cont_par[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "reach_cont_par roptns guards (X::'n eucl1 set) \ SPEC (\(CX, G). G \ (CX \ UNIV) \ (Csafe - guards) \ UNIV \ X \ G \ CX \ UNIV \ flowsto X {0..} (CX \ UNIV) G)" unfolding reach_cont_par_def apply refine_vcg apply auto apply force apply force apply force apply force subgoal apply (rule bexI) prefer 2 apply assumption by auto subgoal apply (rule bexI) prefer 2 apply assumption by auto subgoal for R apply (rule flowsto_source_Union) apply (drule bspec, assumption) apply auto apply (rule flowsto_subset, assumption) apply auto done done lemma subset_iplane_coll[THEN order_trans, refine_vcg]: "subset_iplane_coll x ics \ SPEC (\b. b \ x \ ics)" unfolding subset_iplane_coll_def apply refine_vcg subgoal for X icss by (refine_vcg FORWEAK_mono_rule[where I="\ic b. b \ X \ \(icss)"]) auto done lemma subsets_iplane_coll[THEN order_trans, refine_vcg]: "subsets_iplane_coll x ics \ SPEC (\b. b \ \x \ ics)" unfolding subsets_iplane_coll_def by (refine_vcg FORWEAK_mono_rule[where I="\x b. (b \ \x \ ics)"]) auto lemma symstart_coll[THEN order_trans, refine_vcg]: assumes [refine_vcg]: "wd (TYPE('n::enum rvec))" assumes [le, refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) (X))" shows "symstart_coll symstart X0 \ SPEC (\(CX, X). flowsto ((X0::'n eucl1 set) - trap \ UNIV) {0..} (CX \ UNIV) X)" unfolding symstart_coll_def autoref_tag_defs apply (refine_vcg FORWEAK_mono_rule[where I="\X (CY, Y). flowsto (\X - trap \ UNIV) {0..} (CY \ UNIV) Y"], clarsimp_all) subgoal by force subgoal for a b c d e by (rule flowsto_subset, assumption) auto subgoal by force subgoal for a b c d e f g unfolding Un_Diff apply (rule flowsto_source_unionI) subgoal by (rule flowsto_subset, assumption) auto subgoal by (rule flowsto_subset, assumption) auto done done lemma reach_cont_symstart[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" assumes [le, refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) (X))" shows "reach_cont_symstart roptns symstart guards (X::'n eucl1 set) \ SPEC (\(CX, G). G \ (CX \ UNIV) \ (Csafe - guards) \ UNIV \ X \ CX \ UNIV \ G \ CX \ UNIV \ flowsto (X - trap \ UNIV) {0..} (CX \ UNIV) (G))" unfolding reach_cont_symstart_def autoref_tag_defs apply (refine_vcg, clarsimp_all) subgoal by (auto simp: times_subset_iff) subgoal by auto subgoal by auto subgoal for a b c d e f g apply (rule flowsto_stepI[OF _ _ order_refl]) apply assumption by assumption auto done lemma reach_conts[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) X)" shows "reach_conts roptns symstart trap guards (X::'n eucl1 set) \ SPEC (\(CX, IGs, X0). \(snd ` IGs) \ (CX \ UNIV) \ (Csafe - guards) \ UNIV \ X \ CX \ UNIV \ \(snd ` IGs) \ CX \ UNIV \ \(fst ` IGs) \ guards \ X = \(X0 ` (snd ` IGs)) \ (\(I, G) \ IGs. flowsto (X0 G - trap \ UNIV) {0..} (CX \ UNIV) G))" unfolding reach_conts_def autoref_tag_defs apply (refine_vcg, clarsimp_all) subgoal for a b apply (erule flowsto_Diff_to_Union_funE) apply (force simp: split_beta') subgoal for f apply (rule exI[where x=f]) by (auto simp: split_beta') done subgoal by (auto) subgoal by (auto) subgoal by (auto) done lemma leaves_halfspace[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "leaves_halfspace S (X::'n::enum rvec set) \ SPEC (\b. case b of None \ S = UNIV | Some sctn \ (S = below_halfspace sctn \ X \ plane_of sctn \ (\x \ X. ode x \ normal sctn < 0)))" unfolding leaves_halfspace_def autoref_tag_defs op_set_to_list_def apply (refine_vcg, clarsimp_all) subgoal by (force simp add: halfspace_simps plane_of_def) done lemma poincare_start_on[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "poincare_start_on guards sctn (X0::'n eucl1 set) \ SPEC (\(X1S, CX1S). fst ` (X1S \ (CX1S \ UNIV)) \ Csafe \ fst ` X1S \ sbelow_halfspace sctn \ fst ` (X1S \ (CX1S \ UNIV)) \ guards = {} \ (X0 \ (CX1S \ UNIV)) \ (\(x, d) \ CX1S \ UNIV. ode x \ normal sctn < 0) \ flowsto X0 pos_reals ((CX1S \ UNIV) \ (sbelow_halfspace sctn \ UNIV)) X1S)" unfolding poincare_start_on_def autoref_tag_defs apply refine_vcg apply (rule FORWEAK_mono_rule[where I="\X0S (X1S, CX1S). flowsto (\X0S) pos_reals ((CX1S \ UNIV) \ sbelow_halfspace sctn \ UNIV) X1S \ fst ` (X1S \ (CX1S \ UNIV)) \ Csafe \ (\X0S) \ X0 \ (\X0S) \ (CX1S \ UNIV) \ fst ` (X1S \ (CX1S \ UNIV)) \ guards = {} \ (\(x, d) \ (CX1S \ UNIV). ode x \ normal sctn < 0) \ fst ` X1S \ sbelow_halfspace sctn"]) subgoal by (refine_vcg) subgoal for A B apply (refine_vcg) subgoal apply (auto simp: dest!: flowpipe_imp_flowsto) apply (rule flowsto_subset) apply (rule flowsto_stays_sbelow[where sctn=sctn]) apply (rule flowsto_subset) apply assumption apply (rule order_refl) apply force apply (rule order_refl) apply (rule order_refl) apply (auto simp: halfspace_simps) apply (rule le_less_trans) prefer 2 apply assumption apply (drule bspec) apply (rule subsetD, assumption) prefer 2 apply assumption apply auto done subgoal by auto subgoal by force subgoal by (auto simp: dest!: flowpipe_source_subset) subgoal by auto subgoal apply (auto simp: halfspace_simps subset_iff) apply (rule le_less_trans[rotated], assumption) by fastforce done subgoal by (auto intro: flowsto_subset) force subgoal for a b c d using assms apply (refine_vcg, clarsimp_all) subgoal for e f g h i j k l m n apply (rule flowsto_source_unionI) subgoal apply (drule flowpipe_imp_flowsto, assumption) apply (rule flowsto_subset[OF flowsto_stays_sbelow[where sctn=sctn] order_refl]) apply (rule flowsto_subset[OF _ order_refl], assumption) apply force apply (rule order_refl) apply (rule order_refl) apply (auto simp: halfspace_simps) apply (rule le_less_trans) prefer 2 apply assumption apply (drule bspec) apply (rule subsetD, assumption) prefer 2 apply assumption apply auto done by (auto intro!: flowsto_source_unionI dest!: flowpipe_imp_flowsto intro: flowsto_subset[OF _ order_refl]) subgoal apply (auto simp: subset_iff) apply (auto simp: image_Un) done subgoal by auto subgoal by (auto dest!: flowpipe_source_subset) subgoal by auto subgoal apply (auto simp: halfspace_simps subset_iff) apply (rule le_less_trans[rotated], assumption) by fastforce subgoal by auto done subgoal by auto done lemma op_inter_fst_coll[le, refine_vcg]: "op_inter_fst_coll X Y \ SPEC (\R. R = X \ Y \ UNIV)" unfolding op_inter_fst_coll_def by (refine_vcg FORWEAK_mono_rule[where I="\Xs R. \Xs \ Y \ UNIV \ R \ R \ X \ Y \ UNIV"]) auto lemma scaleRe_ivl_coll_spec[le, refine_vcg]: "scaleRe_ivl_coll_spec l u X \ SPEC (\Y. Y = scaleR2 l u X)" unfolding scaleRe_ivl_coll_spec_def apply (refine_vcg FORWEAK_mono_rule[where I="\Xs R. scaleR2 l u (\Xs) \ R \ R \ scaleR2 l u X"]) apply (auto simp: intro: scaleR2_subset) subgoal by (force simp: intro: scaleR2_subset) done lemma do_intersection_spec_scaleR2I: "do_intersection_spec UNIV sctns ivl sctn (scaleR2 x1 x2 baa) (scaleR2 x1 x2 aca, x1b)" if "do_intersection_spec UNIV sctns ivl sctn (baa) (aca, x1b)" using that by (auto simp: do_intersection_spec_def intro!: poincare_mapsto_scaleR2I) (auto simp: scaleR2_def image_def vimage_def) lemma do_intersection_core[refine_vcg, le]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "do_intersection_core sctns ivl sctn (X::'n eucl1 set) \ SPEC (\(P1, P2, CX, X0s). do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P1, CX) \ do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P2, CX) \ fst ` (X - X0s) \ CX \ X0s \ X)" unfolding do_intersection_core_def autoref_tag_defs apply (refine_vcg assms, clarsimp_all) subgoal by (rule do_intersection_spec_scaleR2I) (auto simp: do_intersection_spec_def intro!: ) subgoal by (rule do_intersection_spec_scaleR2I) (auto simp: do_intersection_spec_def intro!: ) subgoal by (fastforce simp: scaleR2_def) subgoal by (auto simp: do_intersection_spec_def) subgoal by (auto simp: do_intersection_spec_def) done lemma do_intersection_spec_Union: "do_intersection_spec S sctns ivl sctn (\X) A" if "\x. x \ X \ do_intersection_spec S sctns ivl sctn x A" "X \ {}" using that(2) unfolding do_intersection_spec_def apply clarsimp apply safe subgoal by (rule poincare_mapsto_Union) (auto simp: do_intersection_spec_def dest!: that(1)) subgoal by (auto simp: do_intersection_spec_def dest!: that(1)) subgoal by (auto simp: do_intersection_spec_def dest!: that(1)) subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1)) subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1)) subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1)) subgoal by (force simp: do_intersection_spec_def dest!: that(1)) subgoal by (auto simp: do_intersection_spec_def dest!: that(1)) subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1)) subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1)) done lemma do_intersection_spec_subset2: "do_intersection_spec S p ivl sctn X1 (ab, CY) \ CY \ CX \ CX \ Csafe \ CX \ p = {} \ CX \ ivl \ plane_of sctn = {} \ X0 \ X1 \ do_intersection_spec S p ivl sctn X0 (ab, CX)" by (auto simp: do_intersection_spec_def intro: poincare_mapsto_subset) lemma do_intersection_spec_Union3: "do_intersection_spec S osctns ivl csctns (\x\X. a x) ((\x\X. b x), (\x\X. c x))" if "finite X" "X \ {}" "\x. x \ X \ do_intersection_spec S osctns ivl csctns (a x) (b x, c x)" using that proof induction case empty then show ?case by (auto simp: ) next case (insert x F) show ?case apply (cases "F = {}") subgoal using insert by simp subgoal apply simp apply (rule do_intersection_spec_union) apply (rule insert.prems) apply simp apply (rule insert.IH) apply (assumption) apply (rule insert.prems) apply simp done done qed lemma do_intersection_coll[le]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "do_intersection_coll sctns ivl sctn (X::'n eucl1 set) \ SPEC (\(P1, P2, CX, X0s). do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P1, CX) \ do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P2, CX) \ fst ` (X - X0s) \ CX \ X0s \ X)" unfolding do_intersection_coll_def autoref_tag_defs apply (refine_vcg, clarsimp_all) subgoal apply (rule do_intersection_spec_subset[OF _ diff_subset]) apply (rule do_intersection_spec_Union3) subgoal by auto subgoal by auto subgoal by auto done subgoal apply (rule do_intersection_spec_subset[OF _ diff_subset]) apply (rule do_intersection_spec_Union3) subgoal by auto subgoal by auto subgoal by auto done subgoal by fastforce subgoal by fastforce done lemma do_intersection_flowsto_trans_outside: assumes "flowsto XS0 {0..} (CX \ UNIV) X1" assumes "do_intersection_spec UNIV guards ivl sctn X1 (P, CP)" assumes "fst ` X1 \ CP" assumes "{x \ ivl. x \ plane_of sctn} \ CX = {}" assumes "guards \ (CX \ CP) = {}" assumes "XS0 \ CX \ UNIV" assumes "closed ivl" assumes "CX \ Csafe" shows "do_intersection_spec UNIV guards ivl sctn XS0 (P, CX \ CP)" using assms apply (auto simp: do_intersection_spec_def) subgoal apply (rule flowsto_poincare_trans, assumption, assumption) subgoal by simp subgoal by auto subgoal using assms(3) by auto subgoal by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def) subgoal premises prems for x d proof - have [intro, simp]: "closed {x \ ivl. x \ plane_of sctn} " "closed {x \ ivl. x \ normal sctn = pstn sctn}" by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def assms) from flowsto_poincare_mapsto_trans_flowsto[OF \flowsto _ _ _ _\ \poincare_mapsto _ _ _ _ _\ _ _ order_refl] have ft: "flowsto XS0 {0<..} (X1 \ CX \ UNIV \ CP \ UNIV) (fst ` P \ UNIV)" by (auto simp: ) then have ret: "returns_to {x \ ivl. x \ normal sctn - pstn sctn = 0} x" apply (rule returns_to_flowstoI[OF _ _ _ _ _ _ order_refl]) using prems by (auto simp: plane_of_def) have pm: "poincare_map {x \ ivl. x \ normal sctn = pstn sctn} x \ fst ` P" apply (rule poincare_map_mem_flowstoI[OF ft]) using prems by (auto simp: plane_of_def) from pm prems have "\\<^sub>F x in at (poincare_map {x \ ivl. x \ normal sctn = pstn sctn} x) within plane_of sctn. x \ ivl" by auto from ret have "isCont (return_time {x \ ivl. x \ normal sctn - pstn sctn = 0}) x" apply (rule return_time_isCont_outside) using prems pm by (auto simp: eventually_at_filter plane_of_def intro!: assms derivative_eq_intros) then show "isCont (return_time {x \ ivl. x \ plane_of sctn}) x" by (simp add: plane_of_def) qed subgoal by simp done done lemma do_intersection_coll_flowsto[le]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" assumes ft: "flowsto X0 {0..} (CX0 \ UNIV) X" assumes X_subset: "X \ CX0 \ UNIV" assumes X0_subset: "X0 \ CX0 \ UNIV" and CX0_safe: "CX0 \ Csafe" assumes ci: "closed ivl" assumes disj: "ivl \ plane_of sctn \ CX0 = {}" "sctns \ CX0 = {}" shows "do_intersection_coll sctns ivl sctn (X::'n eucl1 set) \ SPEC (\(P1, P2, CX, X0s). \A. do_intersection_spec UNIV sctns ivl sctn A (P1, CX0 \ CX) \ do_intersection_spec UNIV sctns ivl sctn A (P2, CX0 \ CX) \ flowsto (X0 - A) {0..} (CX0 \ UNIV) X0s \ A \ X0 \ P1 \ X0s = {} \ P2 \ X0s = {})" apply (rule do_intersection_coll) apply (rule wd) proof (clarsimp, goal_cases) case (1 P1 P2 CX R) from ft have "flowsto X0 {0..} (CX0 \ UNIV) (X - R \ R)" by (rule flowsto_subset) auto from flowsto_union_DiffE[OF this] obtain A where AB: "A \ X0" and A: "flowsto A {0..} (CX0 \ UNIV) (X - R)" and B: "flowsto (X0 - A) {0..} (CX0 \ UNIV) (R)" by auto have di: "do_intersection_spec UNIV sctns ivl sctn A (P1, CX0 \ CX)" apply (rule do_intersection_flowsto_trans_outside[OF A 1(1)]) subgoal using 1 by simp subgoal using disj by auto subgoal using 1 disj by (auto simp: do_intersection_spec_def) subgoal using X0_subset AB by (auto simp: do_intersection_spec_def) subgoal using ci by simp subgoal using CX0_safe . done then have "P1 \ (ivl \ plane_of sctn) \ UNIV" by (auto simp: do_intersection_spec_def) then have disjoint: "P1 \ R = {}" using \R \ X\ disj X_subset apply (auto simp: subset_iff) by (metis (no_types, lifting) Int_iff disjoint_iff_not_equal) have di2: "do_intersection_spec UNIV sctns ivl sctn A (P2, CX0 \ CX)" apply (rule do_intersection_flowsto_trans_outside[OF A 1(2)]) subgoal using 1 by simp subgoal using disj by auto subgoal using 1 disj by (auto simp: do_intersection_spec_def) subgoal using X0_subset AB by (auto simp: do_intersection_spec_def) subgoal using ci by simp subgoal using CX0_safe . done then have "P2 \ (ivl \ plane_of sctn) \ UNIV" by (auto simp: do_intersection_spec_def) then have "P2 \ R = {}" using \R \ X\ disj X_subset apply (auto simp: subset_iff) by (metis (no_types, lifting) Int_iff disjoint_iff_not_equal) from AB this disjoint di di2 B show ?case by (auto simp:) qed lemma op_enlarge_ivl_sctn[le, refine_vcg]: "op_enlarge_ivl_sctn ivl sctn d \ SPEC (\ivl'. ivl \ ivl')" unfolding op_enlarge_ivl_sctn_def apply refine_vcg unfolding plane_of_def apply (safe intro!: eventually_in_planerectI) apply (auto intro!: simp: eucl_le[where 'a='a] inner_sum_left inner_Basis if_distrib algebra_simps cong: if_cong) done lemma resolve_ivlplanes[le]: assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)" assumes "\x\Xg. case x of (I, G) \ flowsto (XSf G) {0..} (CXS \ UNIV) G" "(\x\Xg. snd x) \ (Csafe - (ivlplanes \ guards)) \ UNIV" "CXS \ UNIV \ (Csafe - (ivlplanes \ guards)) \ UNIV" "(\a\Xg. XSf (snd a)) \ (CXS::'a rvec set) \ UNIV" "(\x\Xg. snd x) \ CXS \ UNIV" "(\x\Xg. fst x) \ ivlplanes \ guards" shows "resolve_ivlplanes guards ivlplanes Xg \ SPEC (\PS. CXS \ (guards \ ivlplanes) = {} \ CXS \ Csafe \ (\R0 P0. (\x\PS. P0 x) \ (\x\PS. R0 x) = (\a\Xg. XSf (snd a))\ (\x\PS. case x of (X, P1, P2, R, ivl, sctn, CX) \ ivl \ plane_of sctn \ ivlplanes \ closed ivl \ P0 (X, P1, P2, R, ivl, sctn, CX) \ R0 (X, P1, P2, R, ivl, sctn, CX) = {} \ R0 (X, P1, P2, R, ivl, sctn, CX) \ (CXS \ UNIV) \ flowsto (R0 (X, P1, P2, R, ivl, sctn, CX)) {0..} (CXS \ UNIV) R \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P2, CXS \ CX))))" using assms unfolding resolve_ivlplanes_def apply clarsimp_all apply (refine_vcg FORWEAK_mono_rule[where I="\Xgs PS. (\R0 P0. snd ` Xgs \ fst ` PS \ fst ` PS \ snd ` Xg \ (\(X, P1, P2, R, ivl, sctn, CX) \ PS. P0 (X, P1, P2, R, ivl, sctn, CX) \ R0 (X, P1, P2, R, ivl, sctn, CX) = XSf X \ ivl \ plane_of sctn \ ivlplanes \ closed ivl \ P0 (X, P1, P2, R, ivl, sctn, CX) \ R0 (X, P1, P2, R, ivl, sctn, CX) = {} \ R0 (X, P1, P2, R, ivl, sctn, CX) \ (CXS \ UNIV) \ flowsto (R0 (X, P1, P2, R, ivl, sctn, CX)) {0..} (CXS \ UNIV) R \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P2, CXS \ CX)))"], clarsimp_all) using [[goals_limit=1]] subgoal by auto subgoal by auto subgoal for a b c apply (frule bspec, assumption, clarsimp) apply (rule do_intersection_coll_flowsto) apply (rule wd) apply assumption apply force apply force apply blast apply assumption subgoal premises prems proof - have "(b \ plane_of c, a) \ Xg" using prems by simp with \(\x\Xg. fst x) \ ivlplanes \ guards\ have "b \ plane_of c \ ivlplanes \ guards" by (force simp: subset_iff) then show ?thesis using \CXS \ UNIV \ (Csafe - (ivlplanes \ guards)) \ UNIV\ by auto qed subgoal by (auto simp: subset_iff) subgoal apply (refine_vcg, clarsimp_all) apply force apply (intro exI conjI)defer defer defer apply assumption+ apply simp apply force apply force apply force done done subgoal by (auto simp: subset_iff) blast subgoal for a b c d e f R0 P0 apply (frule bspec, assumption, clarsimp) apply (rule do_intersection_coll_flowsto) apply (rule wd) apply assumption subgoal apply (rule order_trans[where y="(\x\Xg. snd x)"]) by auto subgoal apply (rule order_trans) defer apply assumption by auto subgoal by blast subgoal by simp subgoal premises prems proof - have "(d \ plane_of e, c) \ Xg" using prems by simp with \(\x\Xg. fst x) \ ivlplanes \ guards\ have "d \ plane_of e \ ivlplanes \ guards" by (force simp: subset_iff) then show ?thesis using \CXS \ UNIV \ (Csafe - (ivlplanes \ guards)) \ UNIV\ by auto qed subgoal by (auto simp: subset_iff) subgoal apply (refine_vcg, clarsimp_all) subgoal by (auto simp: subset_iff) subgoal by (auto simp: ) subgoal for x1 x1' x2 x3 A apply (rule exI[where x="R0((c, x1, x1', x3, d, e, x2):=(XSf c - A))"]) apply (rule exI[where x="P0((c, x1, x1', x3, d, e, x2):=A)"]) apply clarsimp apply (rule conjI) subgoal by auto apply (rule conjI) subgoal premises prems using prems apply (auto simp: subset_iff) by fastforce apply clarsimp subgoal apply (drule bspec, assumption) apply (drule bspec, assumption) by force done done done subgoal by (auto simp: subset_iff) subgoal by (auto simp: subset_iff) subgoal for a R0 P0 apply (rule exI[where x=R0]) apply (rule exI[where x=P0]) apply (rule conjI) subgoal premises prems proof - note prems show ?thesis using prems(9,8) by fastforce qed by auto done lemma poincare_onto[le, refine_vcg]: assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) X)" assumes CXS0: "CXS0 \ (guards \ ivlplanes) = {}" shows "poincare_onto ro symstart trap guards ivlplanes (XS0::'a eucl1 set) CXS0 \ SPEC (\PS. (\R0 P0. \(P0 ` PS \ R0 ` PS) = XS0 - trap \ UNIV \ (\(X, P1, P2, R, ivl, sctn, CX, CXS) \ PS. ivl \ plane_of sctn \ ivlplanes \ closed ivl \ XS0 \ CXS \ UNIV \ CXS0 \ CXS \ CXS \ (guards \ ivlplanes) = {} \ P0 (X, P1, P2, R, ivl, sctn, CX, CXS) \ R0 (X, P1, P2, R, ivl, sctn, CX, CXS) = {} \ R0 (X, P1, P2, R, ivl, sctn, CX, CXS) \ CXS \ UNIV \ flowsto (R0 (X, P1, P2, R, ivl, sctn, CX, CXS)) {0..} (CXS \ UNIV) R \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \ CX)) ))" unfolding poincare_onto_def autoref_tag_defs using [[goals_limit=1]] apply (refine_vcg, clarsimp_all) apply (refine_vcg resolve_ivlplanes[OF wd]) subgoal by force apply clarsimp subgoal for a b c d R0 P0 apply (rule exI[where x="\(X, P1, P2, R, ivl, sctn, CX, CXS). R0 (X, P1, P2, R, ivl, sctn, CX)"]) apply (rule exI[where x="\(X, P1, P2, R, ivl, sctn, CX, CXS). P0 (X, P1, P2, R, ivl, sctn, CX)"]) apply (rule conjI) subgoal premises prems using \(\x\d. P0 x) \ (\x\d. R0 x) = (\x\b. c (snd x)) - trap \ UNIV\ by auto subgoal apply clarsimp apply (drule bspec, assumption)+ apply (rule conjI, force) apply (rule conjI, force) apply (rule conjI, force) apply (rule conjI) subgoal using CXS0 by (auto simp: ) apply (rule conjI, force) apply (rule conjI, force) apply (rule conjI) subgoal by (auto intro: flowsto_subset) subgoal apply clarsimp apply (rule conjI) subgoal apply (rule do_intersection_spec_subset2, assumption) subgoal by force subgoal by (force simp: do_intersection_spec_def) subgoal using CXS0 by (auto simp: do_intersection_spec_def) subgoal using CXS0 by (auto simp: do_intersection_spec_def) subgoal by auto done subgoal apply (rule do_intersection_spec_subset2, assumption) subgoal by force subgoal by (force simp: do_intersection_spec_def) subgoal using CXS0 by (auto simp: do_intersection_spec_def) subgoal using CXS0 by (auto simp: do_intersection_spec_def) subgoal by auto done done done done done lemma empty_remainders[le, refine_vcg]: "empty_remainders PS \ SPEC (\b. b \ (\(X, P1, P2, R, ivl, sctn, CX) \ PS. R = {}))" unfolding empty_remainders_def by (refine_vcg FORWEAK_mono_rule[where I="\Xs b. b \ (\(X, P1, P2, R, ivl, sctn, CX) \ Xs. R = {})"]) auto lemma poincare_onto_empty[le, refine_vcg]: assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)" assumes CXS0: "CXS0 \ (guards \ ivlplanes) = {}" shows "poincare_onto_empty ro guards ivlplanes (XS0::'a eucl1 set) CXS0 \ SPEC (\(PS). (\R0 P0. \(P0 ` PS \ R0 ` PS) = XS0 \ (\(X, P1, P2, R, ivl, sctn, CX, CXS) \ PS. ivl \ plane_of sctn \ ivlplanes \ closed ivl \ XS0 \ CXS \ UNIV \ CXS0 \ CXS \ CXS \ (guards \ ivlplanes) = {} \ P0 (X, P1, P2, R, ivl, sctn, CX, CXS) \ R0 (X, P1, P2, R, ivl, sctn, CX, CXS) = {} \ R0 (X, P1, P2, R, ivl, sctn, CX, CXS) \ CXS \ UNIV \ flowsto (R0 (X, P1, P2, R, ivl, sctn, CX, CXS)) {0..} (CXS \ UNIV) R \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \ CX)) ))" using CXS0 unfolding poincare_onto_empty_def autoref_tag_defs by (refine_vcg) (auto intro!: flowsto_self) lemma do_intersection_spec_union2: assumes "do_intersection_spec S osctns ivl csctns a (b, c)" "do_intersection_spec S osctns ivl csctns f (b, c)" shows "do_intersection_spec S osctns ivl csctns (a \ f) (b, c)" using do_intersection_spec_union[OF assms] by auto lemma poincare_onto2[le, refine_vcg]: assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) X)" notes [refine_vcg_def] = op_set_ndelete_spec shows "poincare_onto2 ro symstart trap guards ivlplanes (XS0::'a eucl1 set) \ SPEC (\(PS). (\P0. \(P0 ` PS) = XS0 - trap \ UNIV \ (\(s, X, P1, P2, R, ivl, sctn, CX, CXS) \ PS. XS0 \ CXS \ UNIV \ do_intersection_spec UNIV guards ivl sctn (P0 (s, X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (P0 (s, X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \ CX))))" unfolding poincare_onto2_def autoref_tag_defs apply (refine_vcg, clarsimp_all) subgoal for PS R0 P0 apply (rule FORWEAK_mono_rule_empty[where I="\PS1 PS2. (\X0. \(R0 ` PS1) \ \(X0 ` PS2) \ (\(X, P1, P2, R, ivl, sctn, CX, CXS) \ PS2. XS0 \ CXS \ UNIV \ do_intersection_spec UNIV guards ivl sctn (X0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (X0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \ CX)))"]) subgoal by refine_vcg subgoal by auto subgoal by auto subgoal apply clarsimp subgoal for c apply (rule exI[where x=c]) apply (rule conjI) apply (rule order_trans) prefer 2 apply assumption apply (rule UN_mono) apply assumption apply (rule order_refl) apply assumption done done subgoal for \ apply (clarsimp) subgoal for X0 apply (rule exI[where x="\(b, x). (if b then X0 x else P0 x) \ XS0 - trap \ UNIV "]) apply (rule conjI) subgoal premises prems using \(\x\PS. P0 x) \ (\x\PS. R0 x) = XS0 - trap \ UNIV\ \(\x\PS. R0 x) \ (\x\\. X0 x)\ by auto subgoal by (auto intro: do_intersection_spec_subset) done done apply clarsimp subgoal for a b b' c d e f g h i j apply (cases "c = {}") subgoal by (auto intro!: exI[where x="j"]) subgoal using [[goals_limit=1]] apply clarsimp apply refine_vcg subgoal premises prems for k l proof - note prems then show ?thesis apply - apply (drule bspec, assumption)+ apply clarsimp subgoal premises prems using \g \ (guards \ \k) = {}\ \l = k - {d \ plane_of e} \ l = k\ \d \ plane_of e \ \k\ by auto done qed apply simp apply (drule bspec, assumption) apply simp apply (erule exE conjE)+ subgoal for k l m n p q apply (subgoal_tac "\x. x \ m \ p x = {}") defer subgoal for x proof goal_cases case 1 from 1(10,15,24) show ?case by (auto dest!: bspec[where x=x]) qed apply simp subgoal premises prems proof - note prems from prems have "finite (q ` m)" "flowsto (R0 (a, b, b', c, d, e, f, g)) {0..} (g \ UNIV) (\(q ` m))" by auto from flowsto_Union_funE[OF this] obtain XGs where XGs: "\G. G \ q ` m \ flowsto (XGs G) {0..} (g \ UNIV) G" "R0 (a, b, b', c, d, e, f, g) = \(XGs ` (q ` m))" by metis define q0 where "q0 = XGs o q" have "case x of (X, P1, P2, R, ivl, sctn, CX, CXS) \ do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \ CX) \ do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \ CX)" if "x \ m" for x proof (clarsimp, goal_cases) case (1 X P1 P2 R ivl sctn CX CXS) with prems(10)[rule_format, OF \x \ m\] prems(15)[rule_format, OF \x \ m\] \_ = c\ have *: "R = {}" "x = (X, P1, P2, {}, ivl, sctn, CX, CXS)" "ivl \ plane_of sctn \ \l" "closed ivl" "c \ CXS \ UNIV" "g \ CXS" "\(q ` m) \ CXS \ UNIV" "CXS \ (guards \ \l) = {}" "p (X, P1, P2, {}, ivl, sctn, CX, CXS) = {}" "p (X, P1, P2, R, ivl, sctn, CX, CXS) \ CXS \ UNIV" "do_intersection_spec UNIV guards ivl sctn (q (X, P1, P2, {}, ivl, sctn, CX, CXS)) (P1, CXS \ CX)" "do_intersection_spec UNIV guards ivl sctn (q (X, P1, P2, {}, ivl, sctn, CX, CXS)) (P2, CXS \ CX)" by auto have "do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, (CXS \ CX) \ (CXS \ CX))" apply (rule do_intersection_flowsto_trans_outside) apply (simp add: q0_def) apply (rule flowsto_subset) apply (rule XGs) using \x \ m\ apply (rule imageI) using 1 apply force apply force using * apply force apply (rule order_refl) using * apply (auto intro!: *)[] subgoal using * \x \ m\ by (auto simp add: ) subgoal using * by (auto simp: do_intersection_spec_def) subgoal using * by (auto simp: do_intersection_spec_def) subgoal proof - have "q0 (X, P1, P2, R, ivl, sctn, CX, CXS) \ XGs (q x)" by (auto simp: q0_def 1) also have "\ \ R0 (a, b, b', c, d, e, f, g)" using \x \m\ XGs by auto also have "\ \ (CXS \ CX) \ UNIV" using prems(20) \g \ CXS\ by auto finally show ?thesis by simp qed subgoal by fact subgoal using * by (auto simp: do_intersection_spec_def) done moreover have "do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, (CXS \ CX) \ (CXS \ CX))" apply (rule do_intersection_flowsto_trans_outside) apply (simp add: q0_def) apply (rule flowsto_subset) apply (rule XGs) using \x \ m\ apply (rule imageI) using 1 apply force apply force using * apply force apply (rule order_refl) using * apply (auto intro!: *)[] subgoal using * \x \ m\ by (auto simp add: ) subgoal using * by (auto simp: do_intersection_spec_def) subgoal using * by (auto simp: do_intersection_spec_def) subgoal proof - have "q0 (X, P1, P2, R, ivl, sctn, CX, CXS) \ XGs (q x)" by (auto simp: q0_def 1) also have "\ \ R0 (a, b, b', c, d, e, f, g)" using \x \m\ XGs by auto also have "\ \ (CXS \ CX) \ UNIV" using prems(20) \g \ CXS\ by auto finally show ?thesis by simp qed subgoal by fact subgoal using * by (auto simp: do_intersection_spec_def) done ultimately show ?case by (simp add: ) qed note q0 = this have q0': "(a, aa, aa', ab, ac, ad, ae, b) \ m \ XS0 \ b \ UNIV" for a aa aa' ab ac ad ae b apply (drule prems(15)[rule_format]) using \XS0 \ g \ UNIV\ by auto from prems show ?thesis apply (intro exI[where x="\x. if x \ i \ m then j x \ q0 x else if x \ i then j x else q0 x"] conjI) subgoal 1 premises prems unfolding XGs apply simp by (auto simp: q0_def) subgoal premises _ by (rule order_trans[OF \(\x\h. R0 x) \ (\x\i. j x)\]) auto subgoal premises _ using prems(6)[rule_format] q0 apply auto subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2) subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2) subgoal by (auto intro!: do_intersection_spec_union2) subgoal by (auto dest!: prems(6)[rule_format] q0' intro!: do_intersection_spec_union2) subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2) subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2) subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2) subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2) done done qed done done done done done lemma width_spec_ivl[THEN order_trans, refine_vcg]: "width_spec_ivl M X \ SPEC (\x. True)" unfolding width_spec_ivl_def by (refine_vcg) lemma partition_ivl_spec[le, refine_vcg]: shows "partition_ivl cg XS \ SPEC (\YS. XS \ YS)" using [[simproc del: defined_all]] unfolding partition_ivl_def autoref_tag_defs apply (refine_vcg, clarsimp_all) subgoal by fastforce subgoal by fastforce subgoal by fastforce subgoal by fastforce subgoal premises prems for a b c d e f ws g h i j k l m n proof - note prems have disj: "\A Aa. n \ A \ \ XS \ A \ Aa \ n \ Aa" using prems by blast then have "n \ g" using prems by (metis (no_types) Un_iff atLeastAtMost_iff subset_iff) then show ?thesis using disj prems by (meson atLeastAtMost_iff) qed done lemma op_inter_fst_ivl_scaleR2[le,refine_vcg]: "op_inter_fst_ivl_scaleR2 X Y \ SPEC (\R. X \ (Y \ UNIV) = R)" unfolding op_inter_fst_ivl_scaleR2_def apply refine_vcg apply (auto simp: scaleR2_def) subgoal for a b c d e f g h i j k by (rule image_eqI[where x="(i, (j, k))"]; fastforce) subgoal for a b c d e f g h i j k by (rule image_eqI[where x="(i, (j, k))"]; fastforce) done lemma op_inter_fst_ivl_coll_scaleR2[le,refine_vcg]: "op_inter_fst_ivl_coll_scaleR2 X Y \ SPEC (\R. X \ (Y \ UNIV) = R)" unfolding op_inter_fst_ivl_coll_scaleR2_def by (refine_vcg FORWEAK_mono_rule[where I="\Xs R. (\Xs) \ (Y \ UNIV) \ R \ R \ X \ (Y \ UNIV)"]) auto lemma op_inter_ivl_co[le, refine_vcg]: "op_ivl_of_ivl_coll X \ SPEC (\R. X \ R)" unfolding op_ivl_of_ivl_coll_def apply (refine_vcg FORWEAK_mono_rule[where I="\R (l, u). \R \ {l .. u}"]) apply auto apply (metis Set.basic_monos(7) Sup_le_iff atLeastAtMost_iff inf.coboundedI2 inf_sup_aci(1)) by (meson Set.basic_monos(7) UnionI atLeastAtMost_iff le_supI1) lemma op_inter_ivl_coll_scaleR2[le,refine_vcg]: "op_inter_ivl_coll_scaleR2 X Y \ SPEC (\R. X \ (Y \ UNIV) \ R)" unfolding op_inter_ivl_coll_scaleR2_def apply refine_vcg subgoal for _ _ _ A l u by (auto, rule scaleR2_subset[where i'=l and j'=u and k'=A], auto) done lemma [le, refine_vcg]: "op_image_fst_ivl_coll X \ SPEC (\R. R = fst ` X)" unfolding op_image_fst_ivl_coll_def apply (refine_vcg FORWEAK_mono_rule[where I="\Xs R. fst ` (\Xs) \ R \ R \ fst ` X"]) apply auto apply force+ done lemma op_single_inter_ivl[le, refine_vcg]: "op_single_inter_ivl a fxs \ SPEC (\R. a \ fxs \ R)" unfolding op_single_inter_ivl_def by refine_vcg auto lemma partition_ivle_spec[le, refine_vcg]: shows "partition_ivle cg XS \ SPEC (\YS. XS \ YS)" unfolding partition_ivle_def autoref_tag_defs supply [refine_vcg del] = scaleR2_rep_of_coll2 and [refine_vcg] = scaleR2_rep_of_coll apply (refine_vcg) subgoal by (fastforce simp: scaleR2_def) subgoal by auto apply clarsimp subgoal by (fastforce simp: scaleR2_def) done lemma vec1repse[THEN order_trans, refine_vcg]: "vec1repse CX \ SPEC (\R. case R of None \ True | Some X \ X = vec1_of_flow1 ` CX)" unfolding vec1repse_def apply (refine_vcg FORWEAK_mono_rule[where I="\XS R. case R of None \ True | Some R \ vec1_of_flow1 ` (\XS) \ R \ R \ vec1_of_flow1 ` CX"]) apply (auto simp: scaleR2_def split: option.splits) subgoal for a b c d e f g h i j apply (auto simp: vimage_def image_def) apply (rule exI[where x="h"]) apply auto apply (rule exI[where x=f]) apply (rule exI[where x="matrix j"]) apply auto apply (rule bexI) by (auto simp: vec1_of_flow1_def matrix_scaleR) subgoal for a b c d e f g h i j apply (rule bexI) defer apply assumption apply (rule image_eqI[where x="(f, g, j)"]) by (auto simp: flow1_of_vec1_def vec1_of_flow1_def matrix_scaleR[symmetric]) subgoal by fastforce subgoal for a b c d e f g h i j k l apply (auto simp: vimage_def image_def) apply (rule exI[where x="j"]) apply auto apply (rule exI[where x=h]) apply (rule exI[where x="matrix l"]) apply auto apply (rule bexI) by (auto simp: vec1_of_flow1_def matrix_scaleR) subgoal by fastforce subgoal for a b c d e f g h i j k l apply (rule bexI) defer apply assumption apply (rule image_eqI[where x="(h, i, l)"]) by (auto simp: flow1_of_vec1_def vec1_of_flow1_def matrix_scaleR[symmetric]) done lemma scaleR2_rep1[le, refine_vcg]: "scaleR2_rep1 Y \ SPEC (\R. Y \ R)" unfolding scaleR2_rep1_def apply refine_vcg subgoal by (auto simp: norm2_slp_def) subgoal for a b c d e y z f g h i j prec k l m n p q r s apply (auto simp: scaleR2_def image_def vimage_def) subgoal premises prems for B C D E proof - define ij where "ij = (i + j) / 2" from prems have "ij > 0" by (auto simp: ij_def) show ?thesis unfolding ij_def[symmetric] apply (rule exI[where x="1 / ij * B"]) apply (intro conjI) prefer 3 apply (rule bexI[where x="(D, ij *\<^sub>R E)"]) subgoal using \ij > 0\ by auto subgoal using prems using \(D, E) \ c\ \c \ {(n, p)..(q, r)}\ \ij > 0\ by (auto simp: ij_def[symmetric] intro!: scaleR_left_mono) subgoal using \d \ ereal B\ \0 < ij\ \0 < d\ apply (cases d) apply (simp only: times_ereal.simps ereal_less_eq) apply (rule mult_mono) apply (rule real_divl) by auto subgoal using \0 < d\ \d \ ereal B\ \ereal B \ e\ \0 < ij\ \0 < e\ \0 < real_divr prec 1 ((i + j) / 2)\ unfolding ij_def[symmetric] apply (cases e; cases d) apply (simp only: times_ereal.simps ereal_less_eq) apply (rule mult_mono) apply (rule real_divr) by auto done qed done done lemma reduce_ivl[le, refine_vcg]: "reduce_ivl Y b \ SPEC (\R. Y \ R)" unfolding reduce_ivl_def apply refine_vcg apply (auto simp add: scaleR2_def image_def vimage_def plane_of_def ) - prefer 2 - subgoal using basic_trans_rules(23) by blast - prefer 3 - subgoal using basic_trans_rules(23) by blast proof goal_cases case (1 i0 i1 s0 s1 y0 y1) from 1 have le: "1 \ (y1 \ b) / (i1 \ b)" by (auto simp: min_def dest!: inner_Basis_mono[OF _ \b \ Basis\]) show ?case apply (rule exI[where x="(y1 \ b) / (i1 \ b)"]) apply (rule conjI) apply fact apply (rule bexI[where x="(y0, ((i1 \ b) / (y1 \ b)) *\<^sub>R y1)"]) subgoal using 1 le by simp subgoal using 1 le apply simp apply (rule conjI) subgoal apply (auto simp: eucl_le[where 'a="'c"]) apply (auto simp: divide_simps) apply (subst mult.commute) subgoal for i apply (cases " y1 \ b \ i1 \ b") apply (rule order_trans) apply (rule mult_left_mono[where b="y1 \ i"]) apply (auto simp: mult_le_cancel_right) apply (cases "i1 \ i \ 0") apply (rule order_trans) apply (rule mult_right_mono_neg[where b="i1 \ b"]) apply auto by (auto simp: not_le inner_Basis split: if_splits dest!: bspec[where x=i]) done subgoal apply (auto simp: eucl_le[where 'a="'c"]) subgoal for i apply (cases "i = b") apply (auto simp: divide_simps) subgoal by (auto simp: divide_simps algebra_simps) subgoal apply (auto simp: divide_simps algebra_simps inner_Basis) apply (subst mult.commute) apply (rule order_trans) apply (rule mult_right_mono[where b="s1 \ i"]) apply simp apply simp apply (rule mult_left_mono) by auto done done done done next case (2 i0 i1 s0 s1 y0 y1) from 2 have le: "1 \ (y1 \ b) / (s1 \ b)" by (auto simp: min_def abs_real_def divide_simps dest!: inner_Basis_mono[OF _ \b \ Basis\]) show ?case apply (rule exI[where x="(y1 \ b) / (s1 \ b)"]) apply (rule conjI) apply fact apply (rule bexI[where x="(y0, ((s1 \ b) / (y1 \ b)) *\<^sub>R y1)"]) subgoal using 2 le by simp subgoal using 2 le apply simp apply (rule conjI) subgoal apply (auto simp: eucl_le[where 'a="'c"]) subgoal for i apply (cases "i = b") apply (auto simp: divide_simps) subgoal by (auto simp: divide_simps algebra_simps) subgoal apply (auto simp: divide_simps algebra_simps inner_Basis) apply (subst mult.commute) apply (cases "y1 \ i \ 0") apply (rule order_trans) apply (rule mult_left_mono_neg[where b="y1 \ b"]) apply (auto simp: mult_le_cancel_right not_le) apply (rule order_trans) apply (rule mult_right_mono_neg[where b="i1 \ i"]) apply (auto intro!: mult_left_mono_neg) done done done subgoal apply (auto simp: eucl_le[where 'a="'c"]) subgoal for i apply (cases "i = b") subgoal by (auto simp: divide_simps algebra_simps) subgoal apply (auto simp: divide_simps algebra_simps inner_Basis) apply (subst mult.commute) apply (cases "y1 \ i \ 0") apply (rule order_trans) apply (rule mult_left_mono_neg[where b="y1 \ i"]) apply simp apply simp apply (rule mult_right_mono) apply force apply force proof - assume a1: "\i\Basis. s1 \ b * (if b = i then 1 else 0) \ s1 \ i" assume a2: "i \ Basis" assume a3: "i \ b" assume a4: "y1 \ b < 0" assume a5: "s1 \ b < 0" assume a6: "\ 0 \ y1 \ i" have "s1 \ b * (if b = i then 1 else 0) \ s1 \ i" using a2 a1 by metis then have f7: "0 \ s1 \ i" using a3 by (metis (full_types) mult_zero_right) have f8: "y1 \ b \ 0" using a4 by (metis eucl_less_le_not_le) have "s1 \ b \ 0" using a5 by (metis eucl_less_le_not_le) then show "y1 \ b * (s1 \ i) \ s1 \ b * (y1 \ i)" using f8 f7 a6 by (metis mult_right_mono_le mult_zero_left zero_le_mult_iff zero_le_square) qed done done done done qed lemma reduce_ivle[le, refine_vcg]: "reduce_ivle Y b \ SPEC (\R. Y \ R)" using [[simproc del: defined_all]] unfolding reduce_ivle_def apply refine_vcg apply (auto simp: scaleR2_def image_def vimage_def) subgoal for a b c d e f g h i j k apply (drule subsetD, assumption) apply auto subgoal for l m apply (rule exI[where x="l * g"]) apply (intro conjI) subgoal unfolding times_ereal.simps[symmetric] apply (rule ereal_mult_mono) subgoal by (cases e) auto subgoal by (cases b) auto subgoal by (cases b) auto subgoal by (cases e) auto done subgoal unfolding times_ereal.simps[symmetric] apply (rule ereal_mult_mono) subgoal by (cases b) auto subgoal by (cases b) auto subgoal by (cases b) auto subgoal by (cases e) auto done subgoal by force done done done lemma reduces_ivle[le, refine_vcg]: "reduces_ivle X \ SPEC (\R. X \ R)" unfolding reduces_ivle_def by refine_vcg auto lemma ivlse_of_setse[le, refine_vcg]: "ivlse_of_setse X \ SPEC (\R. X \ R)" unfolding ivlse_of_setse_def by (refine_vcg FORWEAK_mono_rule[where I="\Xs R. \Xs \ R"]) (auto simp: scaleR2_def image_def vimage_def) lemma setse_of_ivlse[le, refine_vcg]: "setse_of_ivlse X \ SPEC (\R. R = X)" unfolding setse_of_ivlse_def apply (refine_vcg FORWEAK_mono_rule[where I="\Xs R. \Xs \ R \ R \ X"]) apply clarsimp_all subgoal by (rule bexI) subgoal by auto subgoal by auto subgoal by auto done lemma partition_set_spec[le, refine_vcg]: shows "partition_set ro XS \ SPEC (\YS. XS \ YS)" unfolding partition_set_def autoref_tag_defs apply (refine_vcg) subgoal by (fastforce simp: scaleR2_def vimage_def image_def) subgoal by fastforce done lemma partition_sets_spec[le, refine_vcg]: shows "partition_sets ro XS \ SPEC (\YS. (\(_, _, PS, _, _, _, _, _) \ XS. PS) \ YS)" unfolding partition_sets_def autoref_tag_defs by (refine_vcg FORWEAK_mono_rule[where I="\X Y. (\(_, _, PS, _, _, _, _, _) \ X. PS) \ Y"]) auto lemma do_intersection_poincare_mapstos_trans: assumes pm: "\i. i \ I \ poincare_mapsto (p i) (X0 i) UNIV (CX i) (X1 i)" assumes di: "do_intersection_spec UNIV guards ivl sctn (\i\I. X1 i) (P, CP)" assumes "\i. i \ I \ fst ` (X1 i) \ CP" assumes "\i. i \ I \ {x \ ivl. x \ plane_of sctn} \ CX i = {}" assumes "\i. i \ I \ guards \ (CX i \ CP) = {}" assumes "\i. i \ I \ X0 i \ CX i \ UNIV" assumes "\i. i \ I \ closed (p i)" assumes "closed ivl" assumes "\i. i \ I \ CX i \ Csafe" shows "do_intersection_spec UNIV guards ivl sctn (\i\I. X0 i) (P, (\i\I. CX i) \ CP)" apply (auto simp: do_intersection_spec_def) subgoal apply (simp del: UN_simps add: UN_extend_simps) apply (rule impI) apply (thin_tac "I \ {}") subgoal proof - from di have pmi: "poincare_mapsto {x \ ivl. x \ plane_of sctn} (X1 i) UNIV CP P" if "i \ I" for i by (auto simp: do_intersection_spec_def intro: poincare_mapsto_subset that) show ?thesis apply (rule poincare_mapsto_UnionI) apply (rule poincare_mapsto_trans[OF pm pmi]) apply clarsimp_all subgoal s1 using assms by (auto simp: do_intersection_spec_def) subgoal using assms apply (auto simp: do_intersection_spec_def) apply blast by (metis (mono_tags, lifting) s1 mem_Collect_eq mem_simps(2) mem_simps(4)) subgoal using assms by auto subgoal using assms by auto subgoal premises prems for i x d proof - note prems have [intro, simp]: "closed {x \ ivl. x \ plane_of sctn} " "closed {x \ ivl. x \ normal sctn = pstn sctn}" by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def assms) have set_eq: "(CX i \ CP) \ UNIV = (fst ` X1 i \ UNIV \ CX i \ UNIV \ CP \ UNIV)" using assms prems by auto have empty_inter: "{x \ ivl. x \ normal sctn - pstn sctn = 0} \ UNIV \ (CX i \ CP) \ UNIV = {}" apply safe subgoal using assms(4)[of i] \i \ I\ by (auto simp: plane_of_def ) subgoal using assms(4)[of i] using prems assms by (auto simp: plane_of_def do_intersection_spec_def) done have ft: "flowsto (X0 i) {0<..} ((CX i \ CP) \ UNIV) (fst ` P \ UNIV)" unfolding set_eq apply (rule flowsto_poincare_mapsto_trans_flowsto[OF poincare_mapsto_imp_flowsto[OF pm[OF \i \ I\]] pmi[OF \i \ I\] _ _ order_refl]) using assms prems by (auto) then have ret: "returns_to {x \ ivl. x \ normal sctn - pstn sctn = 0} x" apply (rule returns_to_flowstoI[OF _ _ _ _ _ _ order_refl]) subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal by (rule empty_inter) subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def) done have pm: "poincare_map {x \ ivl. x \ normal sctn = pstn sctn} x \ fst ` P" apply (rule poincare_map_mem_flowstoI[OF ft]) subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using empty_inter by simp subgoal by auto subgoal by auto subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal by auto done from ret have "isCont (return_time {x \ ivl. x \ normal sctn - pstn sctn = 0}) x" apply (rule return_time_isCont_outside) subgoal by fact apply (force intro!: derivative_eq_intros) subgoal by (auto intro!: continuous_intros) subgoal using prems pm assms by (auto simp: do_intersection_spec_def) subgoal using prems pm assms by (auto simp: eventually_at_filter plane_of_def do_intersection_spec_def) subgoal proof - have "x \ CX i" using \_ \ I \ X0 _ \ CX _ \ UNIV\[OF \i \ I\] \(x, _) \ _\ by auto with assms(4)[OF \i \ I\] show ?thesis by (auto simp: plane_of_def) qed done then show "isCont (return_time {x \ ivl. x \ plane_of sctn}) x" by (simp add: plane_of_def) qed done qed done subgoal using assms by (fastforce simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (fastforce simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms(9) by (fastforce simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def) done lemma flow_in_stable_setD: "flow0 x0 t \ stable_set trap \ t \ existence_ivl0 x0 \ x0 \ stable_set trap" apply (auto simp: stable_set_def) proof goal_cases case (1 s) then show ?case apply (cases "s \ t") apply (meson atLeastAtMost_iff contra_subsetD local.ivl_subset_existence_ivl) using contra_subsetD local.existence_ivl_reverse local.existence_ivl_trans' local.flows_reverse by fastforce next case (2) have "((\s. flow0 x0 (t + s)) \ trap) (at_top)" proof (rule Lim_transform_eventually) have "\\<^sub>F x in at_top. x > max t 0" by (simp add: max_def) then show "\\<^sub>F x in at_top. flow0 (flow0 x0 t) x = flow0 x0 (t + x)" apply eventually_elim apply (subst flow_trans) using 2 by auto qed (use 2 in auto) then show ?case by (simp add: tendsto_at_top_translate_iff ac_simps) qed lemma poincare_mapsto_avoid_trap: assumes "poincare_mapsto p (X0 - trap \ UNIV) S CX P" assumes "closed p" assumes trapprop[THEN stable_onD]: "stable_on (CX \ fst ` P) trap" shows "poincare_mapsto p (X0 - trap \ UNIV) S CX (P - trap \ UNIV)" using assms(1,2) apply (auto simp: poincare_mapsto_def) apply (drule bspec, force) apply auto subgoal for x0 d0 D apply (rule exI[where x=D]) apply (auto dest!: trapprop simp: poincare_map_def intro!: return_time_exivl assms(1,2) return_time_pos) subgoal for s by (cases "s = return_time p x0") (auto simp: ) done done lemma poincare_onto_series[le, refine_vcg]: assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) (X))" assumes trapprop: "stable_on (Csafe - (ivl \ plane_of sctn)) trap" shows "poincare_onto_series symstart trap guards (X0::'a eucl1 set) ivl sctn ro \ SPEC (\XS. do_intersection_spec UNIV {} ivl sctn (X0 - trap \ UNIV) (XS, Csafe - (ivl \ plane_of sctn)) \ fst ` X0 - trap \ Csafe - (ivl \ plane_of sctn))" proof (induction guards arbitrary: X0) case Nil then show ?case apply (simp add:) apply refine_vcg apply (clarsimp simp add: ivlsctn_to_set_def) apply (rule do_intersection_spec_subset2, assumption) subgoal by (auto simp: do_intersection_spec_def) subgoal by (auto simp: do_intersection_spec_def) subgoal by (auto simp: do_intersection_spec_def) subgoal by (auto simp: do_intersection_spec_def) subgoal by (auto simp: do_intersection_spec_def) subgoal by (auto simp: do_intersection_spec_def) done next case (Cons a guards) note Cons.IH[simplified, le, refine_vcg] show ?case supply [[simproc del: defined_all]] apply auto apply refine_vcg apply clarsimp_all defer subgoal premises prems for b c d e f g h proof - from prems have "(f, g) \ (\x\c. h x)" by auto then obtain x where "x \ c" "(f, g) \ (h x)" by auto then show ?thesis using prems(14)[rule_format, OF \x \ c\] prems(5-7) by (cases x) (auto simp: do_intersection_spec_def) qed subgoal premises prems for c ro d e f proof - let ?s = "trap \ UNIV" note prems from \do_intersection_spec _ _ _ _ _ _ \ have disro: "do_intersection_spec UNIV {} ivl sctn ((\i\ro. case i of (_, _, PS, _, _, _, _, _, _) \ PS - ?s)) (e, Csafe - ivl \ plane_of sctn)" apply (rule do_intersection_spec_subset) using prems by auto have subset: "(Csafe - ivl \ plane (normal sctn) (pstn sctn)) \ (snd (snd (snd (snd (snd (snd (snd (snd i))))))) \ fst (snd (snd (snd (snd (snd (snd (snd i))))))) \ fst ` fst (snd (snd i)))" if "i \ ro" for i using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that] apply (clarsimp ) subgoal for s X P1 P2 R ivla sctna CX CXS apply (rule conjI) subgoal by (auto simp: plane_of_def) subgoal by (auto simp: plane_of_def) done done have pmro: "poincare_mapsto (case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ {x \ ivla. x \ plane_of sctna}) (f i - ?s) UNIV (case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ CXS \ CX) (case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ P1)" if "i \ ro" for i using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that] by (auto intro: poincare_mapsto_subset) then have pmro: "poincare_mapsto (case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ {x \ ivla. x \ plane_of sctna}) (f i - ?s) UNIV (case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ CXS \ CX) (case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ P1 - ?s)" if "i \ ro" for i unfolding split_beta' apply (rule poincare_mapsto_avoid_trap) using that prems assms by (auto intro!: closed_levelset_within continuous_intros stable_on_mono[OF _ subset] simp: plane_of_def) have "do_intersection_spec UNIV {} ivl sctn (\i\ro. f i - ?s) (e, (\i\ro. case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \ CXS \ CX) \ (Csafe - ivl \ plane_of sctn))" apply (rule do_intersection_poincare_mapstos_trans[OF pmro disro]) subgoal by auto subgoal premises that for i using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that] using [[simproc del: defined_all]] by (auto simp: do_intersection_spec_def) subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def) subgoal by auto subgoal premises that for i using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that] prems(11) that by (auto simp: do_intersection_spec_def) subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def) subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def) subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def) done then show ?thesis unfolding \(\x\ro. f x) = X0 - trap \ UNIV\ apply (rule do_intersection_spec_subset2) subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def) using prems by (auto simp: do_intersection_spec_def intro: poincare_mapsto_subset) qed done qed lemma do_intersection_flowsto_trans_return: assumes "flowsto XS0 {0<..} (CX \ UNIV) X1" assumes "do_intersection_spec UNIV guards ivl sctn X1 (P, CP)" assumes "fst ` X1 \ CP" assumes "{x \ ivl. x \ plane_of sctn} \ CX = {}" assumes "guards \ (CX \ CP) = {}" assumes "closed ivl" assumes "CX \ sbelow_halfspace sctn \ Csafe" assumes subset_plane: "fst ` XS0 \ plane_of sctn \ ivl" assumes down: "\x d. (x, d) \ XS0 \ ode x \ normal sctn < 0" "\x. x \ CX \ ode x \ normal sctn < 0" shows "do_intersection_spec (below_halfspace sctn) guards ivl sctn XS0 (P, CX \ CP)" using assms apply (auto simp: do_intersection_spec_def) subgoal apply (rule flowsto_poincare_trans, assumption, assumption) subgoal by simp subgoal by auto subgoal using assms(3) by auto subgoal by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def) prefer 2 subgoal by (auto simp add: plane_of_def halfspace_simps) subgoal premises prems for x d proof - have [intro, simp]: "closed {x \ ivl. x \ plane_of sctn} " "closed {x \ ivl. x \ normal sctn = pstn sctn}" by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def assms) from subset_plane have "fst ` XS0 \ below_halfspace sctn" by (auto simp: ) from flowsto_stays_sbelow[OF \flowsto _ _ _ _\ this down(2)] have ft_below: "flowsto XS0 pos_reals (CX \ UNIV \ sbelow_halfspace sctn \ UNIV) X1" by auto from flowsto_poincare_mapsto_trans_flowsto[OF ft_below \poincare_mapsto _ _ _ _ _\ _ _ order_refl] have ft: "flowsto XS0 {0<..} (X1 \ CX \ UNIV \ sbelow_halfspace sctn \ UNIV \ CP \ UNIV) (fst ` P \ UNIV)" by (auto simp: ) have ret: "returns_to {x \ ivl. x \ normal sctn - pstn sctn = 0} x" apply (rule returns_to_flowstoI[OF ft]) using prems by (auto simp: plane_of_def halfspace_simps) have pm: "poincare_map {x \ ivl. x \ normal sctn = pstn sctn} x \ fst ` P" apply (rule poincare_map_mem_flowstoI[OF ft]) using prems by (auto simp: plane_of_def halfspace_simps) from pm prems have evmem: "\\<^sub>F x in at (poincare_map {x \ ivl. x \ normal sctn = pstn sctn} x) within plane_of sctn. x \ ivl" by auto from ret have "continuous (at x within {x. x \ normal sctn - pstn sctn \ 0}) (return_time {x \ ivl. x \ normal sctn - pstn sctn = 0})" apply (rule return_time_continuous_below) apply (rule derivative_eq_intros refl)+ apply force subgoal using \closed ivl\ by auto subgoal using prems pm by (auto simp: plane_of_def eventually_at_filter) subgoal by (auto intro!: ) subgoal using prems pm by auto subgoal using prems by auto subgoal using prems pm by (auto intro!: assms simp: plane_of_def) subgoal using prems pm by auto done then show "continuous (at x within below_halfspace sctn) (return_time {x \ ivl. x \ plane_of sctn})" by (simp add: plane_of_def halfspace_simps) qed done done lemma do_intersection_spec_sctn_cong: assumes "sctn = sctn' \ (normal sctn = - normal sctn' \ pstn sctn = - pstn sctn')" shows "do_intersection_spec a b c sctn d e = do_intersection_spec a b c sctn' d e" using assms by (auto simp: do_intersection_spec_def plane_of_def set_eq_iff intro!: ) lemma poincare_onto_from[le, refine_vcg]: assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) (X))" assumes trapprop: "stable_on (Csafe - (ivl \ plane_of sctn)) trap" shows "poincare_onto_from symstart trap S guards ivl sctn ro (XS0::'a eucl1 set) \ SPEC (poincare_mapsto (ivl \ plane_of sctn) (XS0 - trap \ UNIV) S (Csafe - ivl \ plane_of sctn))" unfolding poincare_onto_from_def autoref_tag_defs apply (refine_vcg, clarsimp_all simp: trapprop) subgoal by (auto simp: do_intersection_spec_def Int_def intro: poincare_mapsto_subset) subgoal premises prems for a b c d e f proof - note prems from trapprop have stable: "stable_on (fst ` (e \ UNIV \ sbelow_halfspace a \ UNIV \ d)) trap" apply (rule stable_on_mono) using \fst ` (d \ e \ UNIV) \ Csafe\ \a = sctn \ normal a = - normal sctn \ pstn a = - pstn sctn\ \fst ` d \ sbelow_halfspace a\ by (auto simp: halfspace_simps plane_of_def image_Un) from prems(16) have "flowsto (XS0 - trap \ UNIV) {0<..} (e \ UNIV \ sbelow_halfspace a \ UNIV) d" by (rule flowsto_subset) auto then have ft: "flowsto (XS0 - trap \ UNIV) {0<..} ((e \ sbelow_halfspace a) \ UNIV) (d - trap \ UNIV)" by (auto intro!: flowsto_mapsto_avoid_trap stable simp: Times_Int_distrib1) from prems(8) have di: "do_intersection_spec UNIV {} ivl a (d - trap \ UNIV) (f, Csafe - ivl \ plane_of sctn)" apply (subst do_intersection_spec_sctn_cong) defer apply assumption using prems(2) by auto have "do_intersection_spec (below_halfspace a) {} ivl a (XS0 - trap \ UNIV) (f, e \ sbelow_halfspace a \ (Csafe - ivl \ plane_of sctn))" apply (rule do_intersection_flowsto_trans_return[OF ft di]) subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def) subgoal by (auto simp: halfspace_simps plane_of_def) subgoal using prems by (auto simp: halfspace_simps plane_of_def) subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def) subgoal using prems by (auto simp: image_Un) subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def) subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def) subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def) done moreover have "plane_of a = plane_of sctn" using prems(2) by (auto simp: plane_of_def) ultimately show ?thesis apply (auto simp add: do_intersection_spec_def Int_def) apply (rule poincare_mapsto_subset, assumption) by auto qed done lemma subset_spec1[refine_vcg]: "subset_spec1 R P dP \ SPEC (\b. b \ R \ flow1_of_vec1 ` (P \ dP))" unfolding subset_spec1_def by refine_vcg (auto simp: vec1_of_flow1_def) lemma subset_spec1_coll[le, refine_vcg]: "subset_spec1_coll R P dP \ subset_spec R (flow1_of_vec1 ` (P \ dP))" unfolding autoref_tag_defs subset_spec_def subset_spec1_coll_def by (refine_vcg) (auto simp: subset_iff set_of_ivl_def) lemma one_step_until_time_spec[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "one_step_until_time (X0::'n eucl1 set) CX t1 \ SPEC (\(R, CX). (\(x0, d0) \ X0. t1 \ existence_ivl0 x0 \ (flow0 x0 t1, Dflow x0 t1 o\<^sub>L d0) \ R \ (\t \ {0 .. t1}. flow0 x0 t \ CX)) \ fst ` R \ CX \ Csafe)" using [[simproc del: defined_all]] unfolding one_step_until_time_def autoref_tag_defs apply (refine_vcg WHILE_rule[where I="\(t, h, X, CX). fst ` X \ Csafe \ CX \ Csafe \ 0 \ h \ 0 \ t \ t \ t1 \ (\(x0, d0) \ X0. t \ existence_ivl0 x0 \ (flow0 x0 t, Dflow x0 t o\<^sub>L d0) \ X \ (\s \ {0 .. t}. flow0 x0 s \ CX))"]) subgoal by auto subgoal by (force simp: flowpipe_def existence_ivl_trans flow_trans) subgoal by (auto simp: flowpipe_def existence_ivl_trans flow_trans) apply clarsimp subgoal for startstep rk2_param a b c d e f g h i j apply (safe) subgoal by (auto simp: flowpipe_def intro!: existence_ivl_trans flow_trans) subgoal apply (subst flow_trans, force) subgoal by (auto simp: flowpipe_def intro!: existence_ivl_trans flow_trans) apply (subst Dflow_trans, force) subgoal by (auto simp: flowpipe_def intro!: existence_ivl_trans flow_trans) by (auto simp: blinfun_compose_assoc flowpipe_def) subgoal for s apply (drule bspec[where x="(i, j)"], assumption) apply auto apply (cases "s \ a") subgoal by auto subgoal apply (auto simp: blinfun_compose_assoc flowpipe_def) apply (drule bspec, assumption) apply auto proof goal_cases case 1 have a: "a \ existence_ivl0 i" using 1 by auto have sa: "s - a \ existence_ivl0 (flow0 i a)" using "1"(15) "1"(19) "1"(20) local.ivl_subset_existence_ivl by fastforce have "flow0 i s = flow0 (flow0 i a) (s - a)" by (auto simp: a sa flow_trans[symmetric]) also have "\ \ f" using 1 by auto finally show ?case using 1 by simp qed done done subgoal by auto done text \solve ODE until the time interval \{t1 .. t2}\\ lemma ivl_of_eucl1_coll[THEN order_trans, refine_vcg]: "ivl_of_eucl_coll X \ SPEC (\R. X \ UNIV \ R)" unfolding ivl_of_eucl_coll_def by refine_vcg auto lemma one_step_until_time_ivl_spec[le, refine_vcg]: assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))" shows "one_step_until_time_ivl (X0::'n eucl1 set) CX t1 t2 \ SPEC (\(R, CX). (\(x0, d0) \ X0. {t1 .. t2} \ existence_ivl0 x0 \ (\t \ {t1 .. t2}. (flow0 x0 t, Dflow x0 t o\<^sub>L d0) \ R) \ (\t \ {0 .. t1}. (flow0 x0 t) \ CX)) \ fst ` R \ CX \ Csafe)" unfolding one_step_until_time_ivl_def apply (refine_vcg, clarsimp_all) subgoal for X CX Y CY CY' x0 d0 apply (drule bspec, assumption, clarsimp) apply (drule bspec, assumption, clarsimp simp add: nonneg_interval_mem_existence_ivlI) apply (rule subsetD, assumption) subgoal for t apply (drule bspec[where x=0], force) apply (drule bspec[where x="t - t1"], force) using interval_subset_existence_ivl[of t1 x0 t2] by (auto simp: flow_trans') done done lemma empty_symstart_flowsto: "X0 \ Csafe \ UNIV \ RETURN ({}, X0) \ SPEC (\(CX, X). flowsto (X0 - {} \ UNIV) {0..} (CX \ UNIV) X)" by (auto intro!: flowsto_self) subsection \Poincare map returning to\ lemma poincare_onto_from_ivla[le, refine_vcg]: assumes [refine_vcg]: "wd TYPE('n::enum rvec)" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) (X))" assumes trapprop[refine_vcg]: "stable_on (Csafe - (ivl \ plane_of sctn)) trap" shows "poincare_onto_from symstart trap S guards ivl sctn ro (XS0::'n eucl1 set) \ SPEC (\P. wd TYPE((real, 'n) vec) \ poincare_mapsto (ivl \ plane_of sctn) (XS0 - trap \ UNIV) S (Csafe - ivl \ plane_of sctn) P)" by (refine_vcg) subsection \Poincare map onto (from outside of target)\ subsection \One step method (reachability in time)\ lemma c0_info_of_apprsI: assumes "(b, a) \ clw_rel appr_rel" assumes "x \ a" shows "x \ c0_info_of_apprs b" using assms by (auto simp: appr_rel_br clw_rel_br c0_info_of_apprs_def c0_info_of_appr_def dest!: brD) lemma c0_info_of_appr'I: assumes "(b, a) \ \clw_rel appr_rel\phantom_rel" assumes "x \ a" shows "x \ c0_info_of_appr' b" using assms by (auto simp add: c0_info_of_appr'_def intro!: c0_info_of_apprsI split: option.splits) lemma poincare_onto_from_in_ivl[le, refine_vcg]: assumes [refine_vcg]: "wd TYPE('n::enum rvec)" assumes [refine_vcg]: "\X0. X0 \ Csafe \ UNIV \ symstart X0 \ SPEC (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) (X))" assumes trapprop: "stable_on (Csafe - (ivl \ plane_of sctn)) trap" shows "poincare_onto_from_in_ivl symstart trap S guards ivl sctn ro (XS0::'n::enum eucl1 set) P dP \ SPEC (\b. b \ poincare_mapsto (ivl \ plane_of sctn) (XS0 - trap \ UNIV) S (Csafe - ivl \ plane_of sctn) (flow1_of_vec1 ` (P \ dP)))" unfolding poincare_onto_from_in_ivl_def apply (refine_vcg, clarsimp_all) apply (rule trapprop) apply (rule poincare_mapsto_subset) apply assumption by (auto simp: ) lemma lvivl_default_relI: "(dRi, set_of_lvivl' dRi::'e::executable_euclidean_space set) \ \lvivl_rel\default_rel UNIV" if "lvivl'_invar DIM('e) dRi" using that by (auto simp: set_of_lvivl'_def set_of_lvivl_def set_of_ivl_def lvivl'_invar_def intro!: mem_default_relI lvivl_relI) lemma stable_on_empty[simp]: "stable_on A {}" by (auto simp: stable_on_def) lemma poincare_onto_in_ivl[le, refine_vcg]: assumes [simp]: "length (ode_e) = CARD('n::enum)" shows "poincare_onto_in_ivl guards ivl sctn ro (XS0::'n::enum eucl1 set) P dP \ SPEC (\b. b \ poincare_mapsto (ivl \ plane_of sctn) (XS0) UNIV (Csafe - ivl \ plane_of sctn) (flow1_of_vec1 ` (P \ dP)))" proof - have wd[refine_vcg]: "wd TYPE((real, 'n) vec)" by (simp add: wd_def) show ?thesis unfolding poincare_onto_in_ivl_def apply (refine_vcg) subgoal by (auto intro!: flowsto_self) subgoal apply (clarsimp simp add: do_intersection_spec_def Int_def[symmetric]) apply (rule poincare_mapsto_subset) apply assumption by auto done qed end end \ No newline at end of file diff --git a/thys/Simplex/Simplex.thy b/thys/Simplex/Simplex.thy --- a/thys/Simplex/Simplex.thy +++ b/thys/Simplex/Simplex.thy @@ -1,8324 +1,8326 @@ (* Authors: F. Maric, M. Spasic, R. Thiemann *) section \The Simplex Algorithm\ theory Simplex imports Linear_Poly_Maps QDelta Rel_Chain Simplex_Algebra "HOL-Library.Multiset" "HOL-Library.RBT_Mapping" "HOL-Library.Code_Target_Numeral" begin text\Linear constraints are of the form \p \ c\ or \p\<^sub>1 \ p\<^sub>2\, where \p\, \p\<^sub>1\, and \p\<^sub>2\, are linear polynomials, \c\ is a rational constant and \\ \ {<, >, \, \, =}\. Their abstract syntax is given by the \constraint\ type, and semantics is given by the relation \\\<^sub>c\, defined straightforwardly by primitive recursion over the \constraint\ type. A set of constraints is satisfied, denoted by \\\<^sub>c\<^sub>s\, if all constraints are. There is also an indexed version \\\<^sub>i\<^sub>c\<^sub>s\ which takes an explicit set of indices and then only demands that these constraints are satisfied.\ datatype constraint = LT linear_poly rat | GT linear_poly rat | LEQ linear_poly rat | GEQ linear_poly rat | EQ linear_poly rat | LTPP linear_poly linear_poly | GTPP linear_poly linear_poly | LEQPP linear_poly linear_poly | GEQPP linear_poly linear_poly | EQPP linear_poly linear_poly text \Indexed constraints are just pairs of indices and constraints. Indices will be used to identify constraints, e.g., to easily specify an unsatisfiable core by a list of indices.\ type_synonym 'i i_constraint = "'i \ constraint" abbreviation (input) restrict_to :: "'i set \ ('i \ 'a) set \ 'a set" where "restrict_to I xs \ snd ` (xs \ (I \ UNIV))" text \The operation @{const restrict_to} is used to select constraints for a given index set.\ abbreviation (input) flat :: "('i \ 'a) set \ 'a set" where "flat xs \ snd ` xs" text \The operation @{const flat} is used to drop indices from a set of indexed constraints.\ abbreviation (input) flat_list :: "('i \ 'a) list \ 'a list" where "flat_list xs \ map snd xs" primrec satisfies_constraint :: "'a :: lrv valuation \ constraint \ bool" (infixl "\\<^sub>c" 100) where "v \\<^sub>c (LT l r) \ (l\v\) < r *R 1" | "v \\<^sub>c GT l r \ (l\v\) > r *R 1" | "v \\<^sub>c LEQ l r \ (l\v\) \ r *R 1" | "v \\<^sub>c GEQ l r \ (l\v\) \ r *R 1" | "v \\<^sub>c EQ l r \ (l\v\) = r *R 1" | "v \\<^sub>c LTPP l1 l2 \ (l1\v\) < (l2\v\)" | "v \\<^sub>c GTPP l1 l2 \ (l1\v\) > (l2\v\)" | "v \\<^sub>c LEQPP l1 l2 \ (l1\v\) \ (l2\v\)" | "v \\<^sub>c GEQPP l1 l2 \ (l1\v\) \ (l2\v\)" | "v \\<^sub>c EQPP l1 l2 \ (l1\v\) = (l2\v\)" abbreviation satisfies_constraints :: "rat valuation \ constraint set \ bool" (infixl "\\<^sub>c\<^sub>s" 100) where "v \\<^sub>c\<^sub>s cs \ \ c \ cs. v \\<^sub>c c" lemma unsat_mono: assumes "\ (\ v. v \\<^sub>c\<^sub>s cs)" and "cs \ ds" shows "\ (\ v. v \\<^sub>c\<^sub>s ds)" using assms by auto fun i_satisfies_cs (infixl "\\<^sub>i\<^sub>c\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>c\<^sub>s cs \ v \\<^sub>c\<^sub>s restrict_to I cs" definition distinct_indices :: "('i \ 'c) list \ bool" where "distinct_indices as = (distinct (map fst as))" lemma distinct_indicesD: "distinct_indices as \ (i,x) \ set as \ (i,y) \ set as \ x = y" unfolding distinct_indices_def by (rule eq_key_imp_eq_value) text \For the unsat-core predicate we only demand minimality in case that the indices are distinct. Otherwise, minimality does in general not hold. For instance, consider the input constraints $c_1: x < 0$, $c_2: x > 2$ and $c_2: x < 1$ where the index $c_2$ occurs twice. If the simplex-method first encounters constraint $c_1$, then it will detect that there is a conflict between $c_1$ and the first $c_2$-constraint. Consequently, the index-set $\{c_1,c_2\}$ will be returned, but this set is not minimal since $\{c_2\}$ is already unsatisfiable.\ definition minimal_unsat_core :: "'i set \ 'i i_constraint list \ bool" where "minimal_unsat_core I ics = ((I \ fst ` set ics) \ (\ (\ v. (I,v) \\<^sub>i\<^sub>c\<^sub>s set ics)) \ (distinct_indices ics \ (\ J. J \ I \ (\ v. (J,v) \\<^sub>i\<^sub>c\<^sub>s set ics))))" subsection \Procedure Specification\ abbreviation (input) Unsat where "Unsat \ Inl" abbreviation (input) Sat where "Sat \ Inr" text\The specification for the satisfiability check procedure is given by:\ locale Solve = \ \Decide if the given list of constraints is satisfiable. Return either an unsat core, or a satisfying valuation.\ fixes solve :: "'i i_constraint list \ 'i list + rat valuation" \ \If the status @{const Sat} is returned, then returned valuation satisfies all constraints.\ assumes simplex_sat: "solve cs = Sat v \ v \\<^sub>c\<^sub>s flat (set cs)" \ \If the status @{const Unsat} is returned, then constraints are unsatisfiable, i.e., an unsatisfiable core is returned.\ assumes simplex_unsat: "solve cs = Unsat I \ minimal_unsat_core (set I) cs" abbreviation (input) look where "look \ Mapping.lookup" abbreviation (input) upd where "upd \ Mapping.update" lemma look_upd: "look (upd k v m) = (look m)(k \ v)" by (transfer, auto) lemmas look_upd_simps[simp] = look_upd Mapping.lookup_empty definition map2fun:: "(var, 'a :: zero) mapping \ var \ 'a" where "map2fun v \ \x. case look v x of None \ 0 | Some y \ y" syntax "_map2fun" :: "(var, 'a) mapping \ var \ 'a" ("\_\") translations "\v\" == "CONST map2fun v" lemma map2fun_def': "\v\ x \ case Mapping.lookup v x of None \ 0 | Some y \ y" by (auto simp add: map2fun_def) text\Note that the above specification requires returning a valuation (defined as a HOL function), which is not efficiently executable. In order to enable more efficient data structures for representing valuations, a refinement of this specification is needed and the function \solve\ is replaced by the function \solve_exec\ returning optional \(var, rat) mapping\ instead of \var \ rat\ function. This way, efficient data structures for representing mappings can be easily plugged-in during code generation \cite{florian-refinement}. A conversion from the \mapping\ datatype to HOL function is denoted by \\_\\ and given by: @{thm map2fun_def'[no_vars]}.\ locale SolveExec = fixes solve_exec :: "'i i_constraint list \ 'i list + (var, rat) mapping" assumes simplex_sat0: "solve_exec cs = Sat v \ \v\ \\<^sub>c\<^sub>s flat (set cs)" assumes simplex_unsat0: "solve_exec cs = Unsat I \ minimal_unsat_core (set I) cs" begin definition solve where "solve cs \ case solve_exec cs of Sat v \ Sat \v\ | Unsat c \ Unsat c" end sublocale SolveExec < Solve solve by (unfold_locales, insert simplex_sat0 simplex_unsat0, auto simp: solve_def split: sum.splits) subsection \Handling Strict Inequalities\ text\The first step of the procedure is removing all equalities and strict inequalities. Equalities can be easily rewritten to non-strict inequalities. Removing strict inequalities can be done by replacing the list of constraints by a new one, formulated over an extension \\'\ of the space of rationals \\\. \\'\ must have a structure of a linearly ordered vector space over \\\ (represented by the type class \lrv\) and must guarantee that if some non-strict constraints are satisfied in \\'\, then there is a satisfying valuation for the original constraints in \\\. Our final implementation uses the \\\<^sub>\\ space, defined in \cite{simplex-rad} (basic idea is to replace \p < c\ by \p \ c - \\ and \p > c\ by \p \ c + \\ for a symbolic parameter \\\). So, all constraints are reduced to the form \p \ b\, where \p\ is a linear polynomial (still over \\\), \b\ is constant from \\'\ and \\ \ {\, \}\. The non-strict constraints are represented by the type \'a ns_constraint\, and their semantics is denoted by \\\<^sub>n\<^sub>s\ and \\\<^sub>n\<^sub>s\<^sub>s\. The indexed variant is \\\<^sub>i\<^sub>n\<^sub>s\<^sub>s\.\ datatype 'a ns_constraint = LEQ_ns linear_poly 'a | GEQ_ns linear_poly 'a type_synonym ('i,'a) i_ns_constraint = "'i \ 'a ns_constraint" primrec satisfiable_ns_constraint :: "'a::lrv valuation \ 'a ns_constraint \ bool" (infixl "\\<^sub>n\<^sub>s" 100) where "v \\<^sub>n\<^sub>s LEQ_ns l r \ l\v\ \ r" | "v \\<^sub>n\<^sub>s GEQ_ns l r \ l\v\ \ r" abbreviation satisfies_ns_constraints :: "'a::lrv valuation \ 'a ns_constraint set \ bool" (infixl "\\<^sub>n\<^sub>s\<^sub>s " 100) where "v \\<^sub>n\<^sub>s\<^sub>s cs \ \ c \ cs. v \\<^sub>n\<^sub>s c" fun i_satisfies_ns_constraints :: "'i set \ 'a::lrv valuation \ ('i,'a) i_ns_constraint set \ bool" (infixl "\\<^sub>i\<^sub>n\<^sub>s\<^sub>s " 100) where "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs \ v \\<^sub>n\<^sub>s\<^sub>s restrict_to I cs" lemma i_satisfies_ns_constraints_mono: "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs \ J \ I \ (J,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs" by auto primrec poly :: "'a ns_constraint \ linear_poly" where "poly (LEQ_ns p a) = p" | "poly (GEQ_ns p a) = p" primrec ns_constraint_const :: "'a ns_constraint \ 'a" where "ns_constraint_const (LEQ_ns p a) = a" | "ns_constraint_const (GEQ_ns p a) = a" definition distinct_indices_ns :: "('i,'a :: lrv) i_ns_constraint set \ bool" where "distinct_indices_ns ns = ((\ n1 n2 i. (i,n1) \ ns \ (i,n2) \ ns \ poly n1 = poly n2 \ ns_constraint_const n1 = ns_constraint_const n2))" definition minimal_unsat_core_ns :: "'i set \ ('i,'a :: lrv) i_ns_constraint set \ bool" where "minimal_unsat_core_ns I cs = ((I \ fst ` cs) \ (\ (\ v. (I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs)) \ (distinct_indices_ns cs \ (\ J \ I. \ v. (J,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs)))" text\Specification of reduction of constraints to non-strict form is given by:\ locale To_ns = \ \Convert a constraint to an equisatisfiable non-strict constraint list. The conversion must work for arbitrary subsets of constraints -- selected by some index set I -- in order to carry over unsat-cores and in order to support incremental simplex solving.\ fixes to_ns :: "'i i_constraint list \ ('i,'a::lrv) i_ns_constraint list" \ \Convert the valuation that satisfies all non-strict constraints to the valuation that satisfies all initial constraints.\ fixes from_ns :: "(var, 'a) mapping \ 'a ns_constraint list \ (var, rat) mapping" assumes to_ns_unsat: "minimal_unsat_core_ns I (set (to_ns cs)) \ minimal_unsat_core I cs" assumes i_to_ns_sat: "(I,\v'\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs) \ (I,\from_ns v' (flat_list (to_ns cs))\) \\<^sub>i\<^sub>c\<^sub>s set cs" assumes to_ns_indices: "fst ` set (to_ns cs) = fst ` set cs" assumes distinct_cond: "distinct_indices cs \ distinct_indices_ns (set (to_ns cs))" begin lemma to_ns_sat: "\v'\ \\<^sub>n\<^sub>s\<^sub>s flat (set (to_ns cs)) \ \from_ns v' (flat_list (to_ns cs))\ \\<^sub>c\<^sub>s flat (set cs)" using i_to_ns_sat[of UNIV v' cs] by auto end locale Solve_exec_ns = fixes solve_exec_ns :: "('i,'a::lrv) i_ns_constraint list \ 'i list + (var, 'a) mapping" assumes simplex_ns_sat: "solve_exec_ns cs = Sat v \ \v\ \\<^sub>n\<^sub>s\<^sub>s flat (set cs)" assumes simplex_ns_unsat: "solve_exec_ns cs = Unsat I \ minimal_unsat_core_ns (set I) (set cs)" text\After the transformation, the procedure is reduced to solving only the non-strict constraints, implemented in the \solve_exec_ns\ function having an analogous specification to the \solve\ function. If \to_ns\, \from_ns\ and \solve_exec_ns\ are available, the \solve_exec\ function can be easily defined and it can be easily shown that this definition satisfies its specification (also analogous to \solve\). \ locale SolveExec' = To_ns to_ns from_ns + Solve_exec_ns solve_exec_ns for to_ns:: "'i i_constraint list \ ('i,'a::lrv) i_ns_constraint list" and from_ns :: "(var, 'a) mapping \ 'a ns_constraint list \ (var, rat) mapping" and solve_exec_ns :: "('i,'a) i_ns_constraint list \ 'i list + (var, 'a) mapping" begin definition solve_exec where "solve_exec cs \ let cs' = to_ns cs in case solve_exec_ns cs' of Sat v \ Sat (from_ns v (flat_list cs')) | Unsat is \ Unsat is" end sublocale SolveExec' < SolveExec solve_exec by (unfold_locales, insert simplex_ns_sat simplex_ns_unsat to_ns_sat to_ns_unsat, (force simp: solve_exec_def Let_def split: sum.splits)+) subsection \Preprocessing\ text\The next step in the procedure rewrites a list of non-strict constraints into an equisatisfiable form consisting of a list of linear equations (called the \emph{tableau}) and of a list of \emph{atoms} of the form \x\<^sub>i \ b\<^sub>i\ where \x\<^sub>i\ is a variable and \b\<^sub>i\ is a constant (from the extension field). The transformation is straightforward and introduces auxiliary variables for linear polynomials occurring in the initial formula. For example, \[x\<^sub>1 + x\<^sub>2 \ b\<^sub>1, x\<^sub>1 + x\<^sub>2 \ b\<^sub>2, x\<^sub>2 \ b\<^sub>3]\ can be transformed to the tableau \[x\<^sub>3 = x\<^sub>1 + x\<^sub>2]\ and atoms \[x\<^sub>3 \ b\<^sub>1, x\<^sub>3 \ b\<^sub>2, x\<^sub>2 \ b\<^sub>3]\.\ type_synonym eq = "var \ linear_poly" primrec lhs :: "eq \ var" where "lhs (l, r) = l" primrec rhs :: "eq \ linear_poly" where "rhs (l, r) = r" abbreviation rvars_eq :: "eq \ var set" where "rvars_eq eq \ vars (rhs eq)" definition satisfies_eq :: "'a::rational_vector valuation \ eq \ bool" (infixl "\\<^sub>e" 100) where "v \\<^sub>e eq \ v (lhs eq) = (rhs eq)\v\" lemma satisfies_eq_iff: "v \\<^sub>e (x, p) \ v x = p\v\" by (simp add: satisfies_eq_def) type_synonym tableau ="eq list" definition satisfies_tableau ::"'a::rational_vector valuation \ tableau \ bool" (infixl "\\<^sub>t" 100) where "v \\<^sub>t t \ \ e \ set t. v \\<^sub>e e" definition lvars :: "tableau \ var set" where "lvars t = set (map lhs t)" definition rvars :: "tableau \ var set" where "rvars t = \ (set (map rvars_eq t))" abbreviation tvars where "tvars t \ lvars t \ rvars t" text \The condition that the rhss are non-zero is required to obtain minimal unsatisfiable cores. To observe the problem with 0 as rhs, consider the tableau $x = 0$ in combination with atom $(A: x \leq 0)$ where then $(B: x \geq 1)$ is asserted. In this case, the unsat core would be computed as $\{A,B\}$, although already $\{B\}$ is unsatisfiable.\ definition normalized_tableau :: "tableau \ bool" ("\") where "normalized_tableau t \ distinct (map lhs t) \ lvars t \ rvars t = {} \ 0 \ rhs ` set t" text\Equations are of the form \x = p\, where \x\ is a variable and \p\ is a polynomial, and are represented by the type \eq = var \ linear_poly\. Semantics of equations is given by @{thm satisfies_eq_iff[no_vars]}. Tableau is represented as a list of equations, by the type \tableau = eq list\. Semantics for a tableau is given by @{thm satisfies_tableau_def[no_vars]}. Functions \lvars\ and \rvars\ return sets of variables appearing on the left hand side (lhs) and the right hand side (rhs) of a tableau. Lhs variables are called \emph{basic} while rhs variables are called \emph{non-basic} variables. A tableau \t\ is \emph{normalized} (denoted by @{term "\ t"}) iff no variable occurs on the lhs of two equations in a tableau and if sets of lhs and rhs variables are distinct.\ lemma normalized_tableau_unique_eq_for_lvar: assumes "\ t" shows "\ x \ lvars t. \! p. (x, p) \ set t" proof (safe) fix x assume "x \ lvars t" then show "\p. (x, p) \ set t" unfolding lvars_def by auto next fix x p1 p2 assume *: "(x, p1) \ set t" "(x, p2) \ set t" then show "p1 = p2" using \\ t\ unfolding normalized_tableau_def by (force simp add: distinct_map inj_on_def) qed lemma recalc_tableau_lvars: assumes "\ t" shows "\ v. \ v'. (\ x \ rvars t. v x = v' x) \ v' \\<^sub>t t" proof fix v let ?v' = "\ x. if x \ lvars t then (THE p. (x, p) \ set t) \ v \ else v x" show "\ v'. (\ x \ rvars t. v x = v' x) \ v' \\<^sub>t t" proof (rule_tac x="?v'" in exI, rule conjI) show "\x\rvars t. v x = ?v' x" using \\ t\ unfolding normalized_tableau_def by auto show "?v' \\<^sub>t t" unfolding satisfies_tableau_def satisfies_eq_def proof fix e assume "e \ set t" obtain l r where e: "e = (l,r)" by force show "?v' (lhs e) = rhs e \ ?v' \" proof- have "(lhs e, rhs e) \ set t" using \e \ set t\ e by auto have "\!p. (lhs e, p) \ set t" using \\ t\ normalized_tableau_unique_eq_for_lvar[of t] using \e \ set t\ unfolding lvars_def by auto let ?p = "THE p. (lhs e, p) \ set t" have "(lhs e, ?p) \ set t" apply (rule theI') using \\!p. (lhs e, p) \ set t\ by auto then have "?p = rhs e" using \(lhs e, rhs e) \ set t\ using \\!p. (lhs e, p) \ set t\ by auto moreover have "?v' (lhs e) = ?p \ v \" using \e \ set t\ unfolding lvars_def by simp moreover have "rhs e \ ?v' \ = rhs e \ v \" apply (rule valuate_depend) using \\ t\ \e \ set t\ unfolding normalized_tableau_def by (auto simp add: lvars_def rvars_def) ultimately show ?thesis by auto qed qed qed qed lemma tableau_perm: assumes "lvars t1 = lvars t2" "rvars t1 = rvars t2" "\ t1" "\ t2" "\ v::'a::lrv valuation. v \\<^sub>t t1 \ v \\<^sub>t t2" shows "mset t1 = mset t2" proof- { fix t1 t2 assume "lvars t1 = lvars t2" "rvars t1 = rvars t2" "\ t1" "\ v::'a::lrv valuation. v \\<^sub>t t1 \ v \\<^sub>t t2" have "set t1 \ set t2" proof (safe) fix a b assume "(a, b) \ set t1" then have "a \ lvars t1" unfolding lvars_def by force then have "a \ lvars t2" using \lvars t1 = lvars t2\ by simp then obtain b' where "(a, b') \ set t2" unfolding lvars_def by force have "\v::'a valuation. \v'. (\x\vars (b - b'). v' x = v x) \ (b - b') \ v' \ = 0" proof fix v::"'a valuation" obtain v' where "v' \\<^sub>t t1" "\ x \ rvars t1. v x = v' x" using recalc_tableau_lvars[of t1] \\ t1\ by auto have "v' \\<^sub>t t2" using \v' \\<^sub>t t1\ \\ v. v \\<^sub>t t1 \ v \\<^sub>t t2\ by simp have "b \v'\ = b' \v'\" using \(a, b) \ set t1\ \v' \\<^sub>t t1\ using \(a, b') \ set t2\ \v' \\<^sub>t t2\ unfolding satisfies_tableau_def satisfies_eq_def by force then have "(b - b') \v'\ = 0" using valuate_minus[of b b' v'] by auto moreover have "vars b \ rvars t1" "vars b' \ rvars t1" using \(a, b) \ set t1\ \(a, b') \ set t2\ \rvars t1 = rvars t2\ unfolding rvars_def by force+ then have "vars (b - b') \ rvars t1" using vars_minus[of b b'] by blast then have "\x\vars (b - b'). v' x = v x" using \\ x \ rvars t1. v x = v' x\ by auto ultimately show "\v'. (\x\vars (b - b'). v' x = v x) \ (b - b') \ v' \ = 0" by auto qed then have "b = b'" using all_val[of "b - b'"] by simp then show "(a, b) \ set t2" using \(a, b') \ set t2\ by simp qed } note * = this have "set t1 = set t2" using *[of t1 t2] *[of t2 t1] using assms by auto moreover have "distinct t1" "distinct t2" using \\ t1\ \\ t2\ unfolding normalized_tableau_def by (auto simp add: distinct_map) ultimately show ?thesis by (auto simp add: set_eq_iff_mset_eq_distinct) qed text\Elementary atoms are represented by the type \'a atom\ and semantics for atoms and sets of atoms is denoted by \\\<^sub>a\ and \\\<^sub>a\<^sub>s\ and given by: \ datatype 'a atom = Leq var 'a | Geq var 'a primrec atom_var::"'a atom \ var" where "atom_var (Leq var a) = var" | "atom_var (Geq var a) = var" primrec atom_const::"'a atom \ 'a" where "atom_const (Leq var a) = a" | "atom_const (Geq var a) = a" primrec satisfies_atom :: "'a::linorder valuation \ 'a atom \ bool" (infixl "\\<^sub>a" 100) where "v \\<^sub>a Leq x c \ v x \ c" | "v \\<^sub>a Geq x c \ v x \ c" definition satisfies_atom_set :: "'a::linorder valuation \ 'a atom set \ bool" (infixl "\\<^sub>a\<^sub>s" 100) where "v \\<^sub>a\<^sub>s as \ \ a \ as. v \\<^sub>a a" definition satisfies_atom' :: "'a::linorder valuation \ 'a atom \ bool" (infixl "\\<^sub>a\<^sub>e" 100) where "v \\<^sub>a\<^sub>e a \ v (atom_var a) = atom_const a" lemma satisfies_atom'_stronger: "v \\<^sub>a\<^sub>e a \ v \\<^sub>a a" by (cases a, auto simp: satisfies_atom'_def) abbreviation satisfies_atom_set' :: "'a::linorder valuation \ 'a atom set \ bool" (infixl "\\<^sub>a\<^sub>e\<^sub>s" 100) where "v \\<^sub>a\<^sub>e\<^sub>s as \ \ a \ as. v \\<^sub>a\<^sub>e a" lemma satisfies_atom_set'_stronger: "v \\<^sub>a\<^sub>e\<^sub>s as \ v \\<^sub>a\<^sub>s as" using satisfies_atom'_stronger unfolding satisfies_atom_set_def by auto text \There is also the indexed variant of an atom\ type_synonym ('i,'a) i_atom = "'i \ 'a atom" fun i_satisfies_atom_set :: "'i set \ 'a::linorder valuation \ ('i,'a) i_atom set \ bool" (infixl "\\<^sub>i\<^sub>a\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>a\<^sub>s as \ v \\<^sub>a\<^sub>s restrict_to I as" fun i_satisfies_atom_set' :: "'i set \ 'a::linorder valuation \ ('i,'a) i_atom set \ bool" (infixl "\\<^sub>i\<^sub>a\<^sub>e\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as \ v \\<^sub>a\<^sub>e\<^sub>s restrict_to I as" lemma i_satisfies_atom_set'_stronger: "Iv \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as \ Iv \\<^sub>i\<^sub>a\<^sub>s as" using satisfies_atom_set'_stronger[of _ "snd Iv"] by (cases Iv, auto) lemma satisifies_atom_restrict_to_Cons: "v \\<^sub>a\<^sub>s restrict_to I (set as) \ (i \ I \ v \\<^sub>a a) \ v \\<^sub>a\<^sub>s restrict_to I (set ((i,a) # as))" unfolding satisfies_atom_set_def by auto lemma satisfies_tableau_Cons: "v \\<^sub>t t \ v \\<^sub>e e \ v \\<^sub>t (e # t)" unfolding satisfies_tableau_def by auto definition distinct_indices_atoms :: "('i,'a) i_atom set \ bool" where "distinct_indices_atoms as = (\ i a b. (i,a) \ as \ (i,b) \ as \ atom_var a = atom_var b \ atom_const a = atom_const b)" text\ The specification of the preprocessing function is given by:\ locale Preprocess = fixes preprocess::"('i,'a::lrv) i_ns_constraint list \ tableau\ ('i,'a) i_atom list \ ((var,'a) mapping \ (var,'a) mapping) \ 'i list" assumes \ \The returned tableau is always normalized.\ preprocess_tableau_normalized: "preprocess cs = (t,as,trans_v,U) \ \ t" and \ \Tableau and atoms are equisatisfiable with starting non-strict constraints.\ i_preprocess_sat: "\ v. preprocess cs = (t,as,trans_v,U) \ I \ set U = {} \ (I,\v\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t \ (I,\trans_v v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" and preprocess_unsat: "preprocess cs = (t, as,trans_v,U) \ (I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs \ \ v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>t t" and \ \distinct indices on ns-constraints ensures distinct indices in atoms\ preprocess_distinct: "preprocess cs = (t, as,trans_v, U) \ distinct_indices_ns (set cs) \ distinct_indices_atoms (set as)" and \ \unsat indices\ preprocess_unsat_indices: "preprocess cs = (t, as,trans_v, U) \ i \ set U \ \ (\ v. ({i},v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs)" and \ \preprocessing cannot introduce new indices\ preprocess_index: "preprocess cs = (t,as,trans_v, U) \ fst ` set as \ set U \ fst ` set cs" begin lemma preprocess_sat: "preprocess cs = (t,as,trans_v,U) \ U = [] \ \v\ \\<^sub>a\<^sub>s flat (set as) \ \v\ \\<^sub>t t \ \trans_v v\ \\<^sub>n\<^sub>s\<^sub>s flat (set cs)" using i_preprocess_sat[of cs t as trans_v U UNIV v] by auto end definition minimal_unsat_core_tabl_atoms :: "'i set \ tableau \ ('i,'a::lrv) i_atom set \ bool" where "minimal_unsat_core_tabl_atoms I t as = ( I \ fst ` as \ (\ (\ v. v \\<^sub>t t \ (I,v) \\<^sub>i\<^sub>a\<^sub>s as)) \ (distinct_indices_atoms as \ (\ J \ I. \ v. v \\<^sub>t t \ (J,v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as)))" lemma minimal_unsat_core_tabl_atomsD: assumes "minimal_unsat_core_tabl_atoms I t as" shows "I \ fst ` as" "\ (\ v. v \\<^sub>t t \ (I,v) \\<^sub>i\<^sub>a\<^sub>s as)" "distinct_indices_atoms as \ J \ I \ \ v. v \\<^sub>t t \ (J,v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as" using assms unfolding minimal_unsat_core_tabl_atoms_def by auto locale AssertAll = fixes assert_all :: "tableau \ ('i,'a::lrv) i_atom list \ 'i list + (var, 'a)mapping" assumes assert_all_sat: "\ t \ assert_all t as = Sat v \ \v\ \\<^sub>t t \ \v\ \\<^sub>a\<^sub>s flat (set as)" assumes assert_all_unsat: "\ t \ assert_all t as = Unsat I \ minimal_unsat_core_tabl_atoms (set I) t (set as)" text\Once the preprocessing is done and tableau and atoms are obtained, their satisfiability is checked by the \assert_all\ function. Its precondition is that the starting tableau is normalized, and its specification is analogue to the one for the \solve\ function. If \preprocess\ and \assert_all\ are available, the \solve_exec_ns\ can be defined, and it can easily be shown that this definition satisfies the specification.\ locale Solve_exec_ns' = Preprocess preprocess + AssertAll assert_all for preprocess:: "('i,'a::lrv) i_ns_constraint list \ tableau \ ('i,'a) i_atom list \ ((var,'a)mapping \ (var,'a)mapping) \ 'i list" and assert_all :: "tableau \ ('i,'a::lrv) i_atom list \ 'i list + (var, 'a) mapping" begin definition solve_exec_ns where "solve_exec_ns s \ case preprocess s of (t,as,trans_v,ui) \ (case ui of i # _ \ Inl [i] | _ \ (case assert_all t as of Inl I \ Inl I | Inr v \ Inr (trans_v v))) " end context Preprocess begin lemma preprocess_unsat_index: assumes prep: "preprocess cs = (t,as,trans_v,ui)" and i: "i \ set ui" shows "minimal_unsat_core_ns {i} (set cs)" proof - from preprocess_index[OF prep] have "set ui \ fst ` set cs" by auto with i have i': "i \ fst ` set cs" by auto from preprocess_unsat_indices[OF prep i] show ?thesis unfolding minimal_unsat_core_ns_def using i' by auto qed lemma preprocess_minimal_unsat_core: assumes prep: "preprocess cs = (t,as,trans_v,ui)" and unsat: "minimal_unsat_core_tabl_atoms I t (set as)" and inter: "I \ set ui = {}" shows "minimal_unsat_core_ns I (set cs)" proof - from preprocess_tableau_normalized[OF prep] have t: "\ t" . from preprocess_index[OF prep] have index: "fst ` set as \ set ui \ fst ` set cs" by auto from minimal_unsat_core_tabl_atomsD(1,2)[OF unsat] preprocess_unsat[OF prep, of I] have 1: "I \ fst ` set as" "\ (\ v. (I, v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs)" by force+ show "minimal_unsat_core_ns I (set cs)" unfolding minimal_unsat_core_ns_def proof (intro conjI impI allI 1(2)) show "I \ fst ` set cs" using 1 index by auto fix J assume "distinct_indices_ns (set cs)" "J \ I" with preprocess_distinct[OF prep] have "distinct_indices_atoms (set as)" "J \ I" by auto from minimal_unsat_core_tabl_atomsD(3)[OF unsat this] obtain v where model: "v \\<^sub>t t" "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s set as" by auto from i_satisfies_atom_set'_stronger[OF model(2)] have model': "(J, v) \\<^sub>i\<^sub>a\<^sub>s set as" . define w where "w = Mapping.Mapping (\ x. Some (v x))" have "v = \w\" unfolding w_def map2fun_def by (intro ext, transfer, auto) with model model' have "\w\ \\<^sub>t t" "(J, \w\) \\<^sub>i\<^sub>a\<^sub>s set as" by auto from i_preprocess_sat[OF prep _ this(2,1)] \J \ I\ inter have "(J, \trans_v w\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" by auto then show "\ w. (J, w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" by auto qed qed end sublocale Solve_exec_ns' < Solve_exec_ns solve_exec_ns proof fix cs obtain t as trans_v ui where prep: "preprocess cs = (t,as,trans_v,ui)" by (cases "preprocess cs") from preprocess_tableau_normalized[OF prep] have t: "\ t" . from preprocess_index[OF prep] have index: "fst ` set as \ set ui \ fst ` set cs" by auto note solve = solve_exec_ns_def[of cs, unfolded prep split] { fix v assume "solve_exec_ns cs = Sat v" from this[unfolded solve] t assert_all_sat[OF t] preprocess_sat[OF prep] show " \v\ \\<^sub>n\<^sub>s\<^sub>s flat (set cs)" by (auto split: sum.splits list.splits) } { fix I assume res: "solve_exec_ns cs = Unsat I" show "minimal_unsat_core_ns (set I) (set cs)" proof (cases ui) case (Cons i uis) hence I: "I = [i]" using res[unfolded solve] by auto from preprocess_unsat_index[OF prep, of i] I Cons index show ?thesis by auto next case Nil from res[unfolded solve Nil] have assert: "assert_all t as = Unsat I" by (auto split: sum.splits) from assert_all_unsat[OF t assert] have "minimal_unsat_core_tabl_atoms (set I) t (set as)" . from preprocess_minimal_unsat_core[OF prep this] Nil show "minimal_unsat_core_ns (set I) (set cs)" by simp qed } qed subsection\Incrementally Asserting Atoms\ text\The function @{term assert_all} can be implemented by iteratively asserting one by one atom from the given list of atoms. \ type_synonym 'a bounds = "var \ 'a" text\Asserted atoms will be stored in a form of \emph{bounds} for a given variable. Bounds are of the form \l\<^sub>i \ x\<^sub>i \ u\<^sub>i\, where \l\<^sub>i\ and \u\<^sub>i\ and are either scalars or $\pm \infty$. Each time a new atom is asserted, a bound for the corresponding variable is updated (checking for conflict with the previous bounds). Since bounds for a variable can be either finite or $\pm \infty$, they are represented by (partial) maps from variables to values (\'a bounds = var \ 'a\). Upper and lower bounds are represented separately. Infinite bounds map to @{term None} and this is reflected in the semantics: \begin{small} \c \\<^sub>u\<^sub>b b \ case b of None \ False | Some b' \ c \ b'\ \c \\<^sub>u\<^sub>b b \ case b of None \ True | Some b' \ c \ b'\ \end{small} \noindent Strict comparisons, and comparisons with lower bounds are performed similarly. \ abbreviation (input) le where "le lt x y \ lt x y \ x = y" definition geub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ False | Some b' \ le lt b' c" definition gtub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ False | Some b' \ lt b' c" definition leub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ True | Some b' \ le lt c b'" definition ltub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ True | Some b' \ lt c b'" definition lelb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ False | Some b' \ le lt c b'" definition ltlb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ False | Some b' \ lt c b'" definition gelb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ True | Some b' \ le lt b' c" definition gtlb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ True | Some b' \ lt b' c" definition ge_ubound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>u\<^sub>b" 100) where "c \\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition gt_ubound :: "'a::linorder \ 'a option \ bool" (infixl ">\<^sub>u\<^sub>b" 100) where "c >\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition le_ubound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>u\<^sub>b" 100) where "c \\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition lt_ubound :: "'a::linorder \ 'a option \ bool" (infixl "<\<^sub>u\<^sub>b" 100) where "c <\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition le_lbound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>l\<^sub>b" 100) where "c \\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" definition lt_lbound :: "'a::linorder \ 'a option \ bool" (infixl "<\<^sub>l\<^sub>b" 100) where "c <\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" definition ge_lbound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>l\<^sub>b" 100) where "c \\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" definition gt_lbound :: "'a::linorder \ 'a option \ bool" (infixl ">\<^sub>l\<^sub>b" 100) where "c >\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" lemmas bound_compare'_defs = geub_def gtub_def leub_def ltub_def gelb_def gtlb_def lelb_def ltlb_def lemmas bound_compare''_defs = ge_ubound_def gt_ubound_def le_ubound_def lt_ubound_def le_lbound_def lt_lbound_def ge_lbound_def gt_lbound_def lemmas bound_compare_defs = bound_compare'_defs bound_compare''_defs lemma opposite_dir [simp]: "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" by (case_tac[!] b) (auto simp add: bound_compare'_defs) (* Auxiliary lemmas about bound comparison *) lemma [simp]: "\ c \\<^sub>u\<^sub>b None " "\ c \\<^sub>l\<^sub>b None" by (auto simp add: bound_compare_defs) lemma neg_bounds_compare: "(\ (c \\<^sub>u\<^sub>b b)) \ c <\<^sub>u\<^sub>b b" "(\ (c \\<^sub>u\<^sub>b b)) \ c >\<^sub>u\<^sub>b b" "(\ (c >\<^sub>u\<^sub>b b)) \ c \\<^sub>u\<^sub>b b" "(\ (c <\<^sub>u\<^sub>b b)) \ c \\<^sub>u\<^sub>b b" "(\ (c \\<^sub>l\<^sub>b b)) \ c >\<^sub>l\<^sub>b b" "(\ (c \\<^sub>l\<^sub>b b)) \ c <\<^sub>l\<^sub>b b" "(\ (c <\<^sub>l\<^sub>b b)) \ c \\<^sub>l\<^sub>b b" "(\ (c >\<^sub>l\<^sub>b b)) \ c \\<^sub>l\<^sub>b b" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma bounds_compare_contradictory [simp]: "\c \\<^sub>u\<^sub>b b; c <\<^sub>u\<^sub>b b\ \ False" "\c \\<^sub>u\<^sub>b b; c >\<^sub>u\<^sub>b b\ \ False" "\c >\<^sub>u\<^sub>b b; c \\<^sub>u\<^sub>b b\ \ False" "\c <\<^sub>u\<^sub>b b; c \\<^sub>u\<^sub>b b\ \ False" "\c \\<^sub>l\<^sub>b b; c >\<^sub>l\<^sub>b b\ \ False" "\c \\<^sub>l\<^sub>b b; c <\<^sub>l\<^sub>b b\ \ False" "\c <\<^sub>l\<^sub>b b; c \\<^sub>l\<^sub>b b\ \ False" "\c >\<^sub>l\<^sub>b b; c \\<^sub>l\<^sub>b b\ \ False" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma compare_strict_nonstrict: "x <\<^sub>u\<^sub>b b \ x \\<^sub>u\<^sub>b b" "x >\<^sub>u\<^sub>b b \ x \\<^sub>u\<^sub>b b" "x <\<^sub>l\<^sub>b b \ x \\<^sub>l\<^sub>b b" "x >\<^sub>l\<^sub>b b \ x \\<^sub>l\<^sub>b b" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma [simp]: "\ x \ c; c <\<^sub>u\<^sub>b b \ \ x <\<^sub>u\<^sub>b b" "\ x < c; c \\<^sub>u\<^sub>b b \ \ x <\<^sub>u\<^sub>b b" "\ x \ c; c \\<^sub>u\<^sub>b b \ \ x \\<^sub>u\<^sub>b b" "\ x \ c; c >\<^sub>l\<^sub>b b \ \ x >\<^sub>l\<^sub>b b" "\ x > c; c \\<^sub>l\<^sub>b b \ \ x >\<^sub>l\<^sub>b b" "\ x \ c; c \\<^sub>l\<^sub>b b \ \ x \\<^sub>l\<^sub>b b" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma bounds_lg [simp]: "\ c >\<^sub>u\<^sub>b b; x \\<^sub>u\<^sub>b b\ \ x < c" "\ c \\<^sub>u\<^sub>b b; x <\<^sub>u\<^sub>b b\ \ x < c" "\ c \\<^sub>u\<^sub>b b; x \\<^sub>u\<^sub>b b\ \ x \ c" "\ c <\<^sub>l\<^sub>b b; x \\<^sub>l\<^sub>b b\ \ x > c" "\ c \\<^sub>l\<^sub>b b; x >\<^sub>l\<^sub>b b\ \ x > c" "\ c \\<^sub>l\<^sub>b b; x \\<^sub>l\<^sub>b b\ \ x \ c" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma bounds_compare_Some [simp]: "x \\<^sub>u\<^sub>b Some c \ x \ c" "x \\<^sub>u\<^sub>b Some c \ x \ c" "x <\<^sub>u\<^sub>b Some c \ x < c" "x >\<^sub>u\<^sub>b Some c \ x > c" "x \\<^sub>l\<^sub>b Some c \ x \ c" "x \\<^sub>l\<^sub>b Some c \ x \ c" "x >\<^sub>l\<^sub>b Some c \ x > c" "x <\<^sub>l\<^sub>b Some c \ x < c" by (auto simp add: bound_compare_defs) fun in_bounds where "in_bounds x v (lb, ub) = (v x \\<^sub>l\<^sub>b lb x \ v x \\<^sub>u\<^sub>b ub x)" fun satisfies_bounds :: "'a::linorder valuation \ 'a bounds \ 'a bounds \ bool" (infixl "\\<^sub>b" 100) where "v \\<^sub>b b \ (\ x. in_bounds x v b)" declare satisfies_bounds.simps [simp del] lemma satisfies_bounds_iff: "v \\<^sub>b (lb, ub) \ (\ x. v x \\<^sub>l\<^sub>b lb x \ v x \\<^sub>u\<^sub>b ub x)" by (auto simp add: satisfies_bounds.simps) lemma not_in_bounds: "\ (in_bounds x v (lb, ub)) = (v x <\<^sub>l\<^sub>b lb x \ v x >\<^sub>u\<^sub>b ub x)" using bounds_compare_contradictory(7) using bounds_compare_contradictory(2) using neg_bounds_compare(7)[of "v x" "lb x"] using neg_bounds_compare(2)[of "v x" "ub x"] by auto fun atoms_equiv_bounds :: "'a::linorder atom set \ 'a bounds \ 'a bounds \ bool" (infixl "\" 100) where "as \ (lb, ub) \ (\ v. v \\<^sub>a\<^sub>s as \ v \\<^sub>b (lb, ub))" declare atoms_equiv_bounds.simps [simp del] lemma atoms_equiv_bounds_simps: "as \ (lb, ub) \ \ v. v \\<^sub>a\<^sub>s as \ v \\<^sub>b (lb, ub)" by (simp add: atoms_equiv_bounds.simps) text\A valuation satisfies bounds iff the value of each variable respects both its lower and upper bound, i.e, @{thm satisfies_bounds_iff[no_vars]}. Asserted atoms are precisely encoded by the current bounds in a state (denoted by \\\) if every valuation satisfies them iff it satisfies the bounds, i.e., @{thm atoms_equiv_bounds_simps[no_vars, iff]}.\ text\The procedure also keeps track of a valuation that is a candidate solution. Whenever a new atom is asserted, it is checked whether the valuation is still satisfying. If not, the procedure tries to fix that by changing it and changing the tableau if necessary (but so that it remains equivalent to the initial tableau).\ text\Therefore, the state of the procedure stores the tableau (denoted by \\\), lower and upper bounds (denoted by \\\<^sub>l\ and \\\<^sub>u\, and ordered pair of lower and upper bounds denoted by \\\), candidate solution (denoted by \\\) and a flag (denoted by \\\) indicating if unsatisfiability has been detected so far:\ text\Since we also need to now about the indices of atoms, actually, the bounds are also indexed, and in addition to the flag for unsatisfiability, we also store an optional unsat core.\ type_synonym 'i bound_index = "var \ 'i" type_synonym ('i,'a) bounds_index = "(var, ('i \ 'a))mapping" datatype ('i,'a) state = State (\: "tableau") (\\<^sub>i\<^sub>l: "('i,'a) bounds_index") (\\<^sub>i\<^sub>u: "('i,'a) bounds_index") (\: "(var, 'a) mapping") (\: bool) (\\<^sub>c: "'i list option") definition indexl :: "('i,'a) state \ 'i bound_index" ("\\<^sub>l") where "\\<^sub>l s = (fst o the) o look (\\<^sub>i\<^sub>l s)" definition boundsl :: "('i,'a) state \ 'a bounds" ("\\<^sub>l") where "\\<^sub>l s = map_option snd o look (\\<^sub>i\<^sub>l s)" definition indexu :: "('i,'a) state \ 'i bound_index" ("\\<^sub>u") where "\\<^sub>u s = (fst o the) o look (\\<^sub>i\<^sub>u s)" definition boundsu :: "('i,'a) state \ 'a bounds" ("\\<^sub>u") where "\\<^sub>u s = map_option snd o look (\\<^sub>i\<^sub>u s)" abbreviation BoundsIndicesMap ("\\<^sub>i") where "\\<^sub>i s \ (\\<^sub>i\<^sub>l s, \\<^sub>i\<^sub>u s)" abbreviation Bounds :: "('i,'a) state \ 'a bounds \ 'a bounds" ("\") where "\ s \ (\\<^sub>l s, \\<^sub>u s)" abbreviation Indices :: "('i,'a) state \ 'i bound_index \ 'i bound_index" ("\") where "\ s \ (\\<^sub>l s, \\<^sub>u s)" abbreviation BoundsIndices :: "('i,'a) state \ ('a bounds \ 'a bounds) \ ('i bound_index \ 'i bound_index)" ("\\") where "\\ s \ (\ s, \ s)" fun satisfies_bounds_index :: "'i set \ 'a::lrv valuation \ ('a bounds \ 'a bounds) \ ('i bound_index \ 'i bound_index) \ bool" (infixl "\\<^sub>i\<^sub>b" 100) where "(I,v) \\<^sub>i\<^sub>b ((BL,BU),(IL,IU)) \ ( (\ x c. BL x = Some c \ IL x \ I \ v x \ c) \ (\ x c. BU x = Some c \ IU x \ I \ v x \ c))" declare satisfies_bounds_index.simps[simp del] fun satisfies_bounds_index' :: "'i set \ 'a::lrv valuation \ ('a bounds \ 'a bounds) \ ('i bound_index \ 'i bound_index) \ bool" (infixl "\\<^sub>i\<^sub>b\<^sub>e" 100) where "(I,v) \\<^sub>i\<^sub>b\<^sub>e ((BL,BU),(IL,IU)) \ ( (\ x c. BL x = Some c \ IL x \ I \ v x = c) \ (\ x c. BU x = Some c \ IU x \ I \ v x = c))" declare satisfies_bounds_index'.simps[simp del] fun atoms_imply_bounds_index :: "('i,'a::lrv) i_atom set \ ('a bounds \ 'a bounds) \ ('i bound_index \ 'i bound_index) \ bool" (infixl "\\<^sub>i" 100) where "as \\<^sub>i bi \ (\ I v. (I,v) \\<^sub>i\<^sub>a\<^sub>s as \ (I,v) \\<^sub>i\<^sub>b bi)" declare atoms_imply_bounds_index.simps[simp del] lemma i_satisfies_atom_set_mono: "as \ as' \ v \\<^sub>i\<^sub>a\<^sub>s as' \ v \\<^sub>i\<^sub>a\<^sub>s as" by (cases v, auto simp: satisfies_atom_set_def) lemma atoms_imply_bounds_index_mono: "as \ as' \ as \\<^sub>i bi \ as' \\<^sub>i bi" unfolding atoms_imply_bounds_index.simps using i_satisfies_atom_set_mono by blast definition satisfies_state :: "'a::lrv valuation \ ('i,'a) state \ bool" (infixl "\\<^sub>s" 100) where "v \\<^sub>s s \ v \\<^sub>b \ s \ v \\<^sub>t \ s" definition curr_val_satisfies_state :: "('i,'a::lrv) state \ bool" ("\") where "\ s \ \\ s\ \\<^sub>s s" fun satisfies_state_index :: "'i set \ 'a::lrv valuation \ ('i,'a) state \ bool" (infixl "\\<^sub>i\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>s s \ (v \\<^sub>t \ s \ (I,v) \\<^sub>i\<^sub>b \\ s)" declare satisfies_state_index.simps[simp del] fun satisfies_state_index' :: "'i set \ 'a::lrv valuation \ ('i,'a) state \ bool" (infixl "\\<^sub>i\<^sub>s\<^sub>e" 100) where "(I,v) \\<^sub>i\<^sub>s\<^sub>e s \ (v \\<^sub>t \ s \ (I,v) \\<^sub>i\<^sub>b\<^sub>e \\ s)" declare satisfies_state_index'.simps[simp del] definition indices_state :: "('i,'a)state \ 'i set" where "indices_state s = { i. \ x b. look (\\<^sub>i\<^sub>l s) x = Some (i,b) \ look (\\<^sub>i\<^sub>u s) x = Some (i,b)}" text \distinctness requires that for each index $i$, there is at most one variable $x$ and bound $b$ such that $x \leq b$ or $x \geq b$ or both are enforced.\ definition distinct_indices_state :: "('i,'a)state \ bool" where "distinct_indices_state s = (\ i x b x' b'. ((look (\\<^sub>i\<^sub>l s) x = Some (i,b) \ look (\\<^sub>i\<^sub>u s) x = Some (i,b)) \ (look (\\<^sub>i\<^sub>l s) x' = Some (i,b') \ look (\\<^sub>i\<^sub>u s) x' = Some (i,b')) \ (x = x' \ b = b')))" lemma distinct_indices_stateD: assumes "distinct_indices_state s" shows "look (\\<^sub>i\<^sub>l s) x = Some (i,b) \ look (\\<^sub>i\<^sub>u s) x = Some (i,b) \ look (\\<^sub>i\<^sub>l s) x' = Some (i,b') \ look (\\<^sub>i\<^sub>u s) x' = Some (i,b') \ x = x' \ b = b'" using assms unfolding distinct_indices_state_def by blast+ definition unsat_state_core :: "('i,'a::lrv) state \ bool" where "unsat_state_core s = (set (the (\\<^sub>c s)) \ indices_state s \ (\ (\ v. (set (the (\\<^sub>c s)),v) \\<^sub>i\<^sub>s s)))" definition subsets_sat_core :: "('i,'a::lrv) state \ bool" where "subsets_sat_core s = ((\ I. I \ set (the (\\<^sub>c s)) \ (\ v. (I,v) \\<^sub>i\<^sub>s\<^sub>e s)))" definition minimal_unsat_state_core :: "('i,'a::lrv) state \ bool" where "minimal_unsat_state_core s = (unsat_state_core s \ (distinct_indices_state s \ subsets_sat_core s))" lemma minimal_unsat_core_tabl_atoms_mono: assumes sub: "as \ bs" and unsat: "minimal_unsat_core_tabl_atoms I t as" shows "minimal_unsat_core_tabl_atoms I t bs" unfolding minimal_unsat_core_tabl_atoms_def proof (intro conjI impI allI) note min = unsat[unfolded minimal_unsat_core_tabl_atoms_def] from min have I: "I \ fst ` as" by blast with sub show "I \ fst ` bs" by blast from min have "(\v. v \\<^sub>t t \ (I, v) \\<^sub>i\<^sub>a\<^sub>s as)" by blast with i_satisfies_atom_set_mono[OF sub] show "(\v. v \\<^sub>t t \ (I, v) \\<^sub>i\<^sub>a\<^sub>s bs)" by blast fix J assume J: "J \ I" and dist_bs: "distinct_indices_atoms bs" hence dist: "distinct_indices_atoms as" using sub unfolding distinct_indices_atoms_def by blast from min dist J obtain v where v: "v \\<^sub>t t" "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as" by blast have "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s bs" unfolding i_satisfies_atom_set'.simps proof (intro ballI) fix a assume "a \ snd ` (bs \ J \ UNIV)" then obtain i where ia: "(i,a) \ bs" and i: "i \ J" by force with J have "i \ I" by auto with I obtain b where ib: "(i,b) \ as" by force with sub have ib': "(i,b) \ bs" by auto from dist_bs[unfolded distinct_indices_atoms_def, rule_format, OF ia ib'] have id: "atom_var a = atom_var b" "atom_const a = atom_const b" by auto from v(2)[unfolded i_satisfies_atom_set'.simps] i ib have "v \\<^sub>a\<^sub>e b" by force thus "v \\<^sub>a\<^sub>e a" using id unfolding satisfies_atom'_def by auto qed with v show "\v. v \\<^sub>t t \ (J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s bs" by blast qed lemma state_satisfies_index: assumes "v \\<^sub>s s" shows "(I,v) \\<^sub>i\<^sub>s s" unfolding satisfies_state_index.simps satisfies_bounds_index.simps proof (intro conjI impI allI) fix x c from assms[unfolded satisfies_state_def satisfies_bounds.simps, simplified] have "v \\<^sub>t \ s" and bnd: "v x \\<^sub>l\<^sub>b \\<^sub>l s x" "v x \\<^sub>u\<^sub>b \\<^sub>u s x" by auto show "v \\<^sub>t \ s" by fact show "\\<^sub>l s x = Some c \ \\<^sub>l s x \ I \ c \ v x" using bnd(1) by auto show "\\<^sub>u s x = Some c \ \\<^sub>u s x \ I \ v x \ c" using bnd(2) by auto qed lemma unsat_state_core_unsat: "unsat_state_core s \ \ (\ v. v \\<^sub>s s)" unfolding unsat_state_core_def using state_satisfies_index by blast definition tableau_valuated ("\") where "\ s \ \ x \ tvars (\ s). Mapping.lookup (\ s) x \ None" definition index_valid where "index_valid as (s :: ('i,'a) state) = (\ x b i. (look (\\<^sub>i\<^sub>l s) x = Some (i,b) \ ((i, Geq x b) \ as)) \ (look (\\<^sub>i\<^sub>u s) x = Some (i,b) \ ((i, Leq x b) \ as)))" lemma index_valid_indices_state: "index_valid as s \ indices_state s \ fst ` as" unfolding index_valid_def indices_state_def by force lemma index_valid_mono: "as \ bs \ index_valid as s \ index_valid bs s" unfolding index_valid_def by blast lemma index_valid_distinct_indices: assumes "index_valid as s" and "distinct_indices_atoms as" shows "distinct_indices_state s" unfolding distinct_indices_state_def proof (intro allI impI, goal_cases) case (1 i x b y c) note valid = assms(1)[unfolded index_valid_def, rule_format] from 1(1) valid[of x i b] have "(i, Geq x b) \ as \ (i, Leq x b) \ as" by auto then obtain a where a: "(i,a) \ as" "atom_var a = x" "atom_const a = b" by auto from 1(2) valid[of y i c] have y: "(i, Geq y c) \ as \ (i, Leq y c) \ as" by auto then obtain a' where a': "(i,a') \ as" "atom_var a' = y" "atom_const a' = c" by auto from assms(2)[unfolded distinct_indices_atoms_def, rule_format, OF a(1) a'(1)] show ?case using a a' by auto qed text\To be a solution of the initial problem, a valuation should satisfy the initial tableau and list of atoms. Since tableau is changed only by equivalency preserving transformations and asserted atoms are encoded in the bounds, a valuation is a solution if it satisfies both the tableau and the bounds in the final state (when all atoms have been asserted). So, a valuation \v\ satisfies a state \s\ (denoted by \\\<^sub>s\) if it satisfies the tableau and the bounds, i.e., @{thm satisfies_state_def [no_vars]}. Since \\\ should be a candidate solution, it should satisfy the state (unless the \\\ flag is raised). This is denoted by \\ s\ and defined by @{thm curr_val_satisfies_state_def[no_vars]}. \\ s\ will denote that all variables of \\ s\ are explicitly valuated in \\ s\.\ definition update\\ where [simp]: "update\\ field_update i x c s = field_update (upd x (i,c)) s" fun \\<^sub>i\<^sub>u_update where "\\<^sub>i\<^sub>u_update up (State T BIL BIU V U UC) = State T BIL (up BIU) V U UC" fun \\<^sub>i\<^sub>l_update where "\\<^sub>i\<^sub>l_update up (State T BIL BIU V U UC) = State T (up BIL) BIU V U UC" fun \_update where "\_update V (State T BIL BIU V_old U UC) = State T BIL BIU V U UC" fun \_update where "\_update T (State T_old BIL BIU V U UC) = State T BIL BIU V U UC" lemma update_simps[simp]: "\\<^sub>i\<^sub>u (\\<^sub>i\<^sub>u_update up s) = up (\\<^sub>i\<^sub>u s)" "\\<^sub>i\<^sub>l (\\<^sub>i\<^sub>u_update up s) = \\<^sub>i\<^sub>l s" "\ (\\<^sub>i\<^sub>u_update up s) = \ s" "\ (\\<^sub>i\<^sub>u_update up s) = \ s" "\ (\\<^sub>i\<^sub>u_update up s) = \ s" "\\<^sub>c (\\<^sub>i\<^sub>u_update up s) = \\<^sub>c s" "\\<^sub>i\<^sub>l (\\<^sub>i\<^sub>l_update up s) = up (\\<^sub>i\<^sub>l s)" "\\<^sub>i\<^sub>u (\\<^sub>i\<^sub>l_update up s) = \\<^sub>i\<^sub>u s" "\ (\\<^sub>i\<^sub>l_update up s) = \ s" "\ (\\<^sub>i\<^sub>l_update up s) = \ s" "\ (\\<^sub>i\<^sub>l_update up s) = \ s" "\\<^sub>c (\\<^sub>i\<^sub>l_update up s) = \\<^sub>c s" "\ (\_update V s) = V" "\\<^sub>i\<^sub>l (\_update V s) = \\<^sub>i\<^sub>l s" "\\<^sub>i\<^sub>u (\_update V s) = \\<^sub>i\<^sub>u s" "\ (\_update V s) = \ s" "\ (\_update V s) = \ s" "\\<^sub>c (\_update V s) = \\<^sub>c s" "\ (\_update T s) = T" "\\<^sub>i\<^sub>l (\_update T s) = \\<^sub>i\<^sub>l s" "\\<^sub>i\<^sub>u (\_update T s) = \\<^sub>i\<^sub>u s" "\ (\_update T s) = \ s" "\ (\_update T s) = \ s" "\\<^sub>c (\_update T s) = \\<^sub>c s" by (atomize(full), cases s, auto) declare \\<^sub>i\<^sub>u_update.simps[simp del] \\<^sub>i\<^sub>l_update.simps[simp del] fun set_unsat :: "'i list \ ('i,'a) state \ ('i,'a) state" where "set_unsat I (State T BIL BIU V U UC) = State T BIL BIU V True (Some (remdups I))" lemma set_unsat_simps[simp]: "\\<^sub>i\<^sub>l (set_unsat I s) = \\<^sub>i\<^sub>l s" "\\<^sub>i\<^sub>u (set_unsat I s) = \\<^sub>i\<^sub>u s" "\ (set_unsat I s) = \ s" "\ (set_unsat I s) = \ s" "\ (set_unsat I s) = True" "\\<^sub>c (set_unsat I s) = Some (remdups I)" by (atomize(full), cases s, auto) datatype ('i,'a) Direction = Direction (lt: "'a::linorder \ 'a \ bool") (LBI: "('i,'a) state \ ('i,'a) bounds_index") (UBI: "('i,'a) state \ ('i,'a) bounds_index") (LB: "('i,'a) state \ 'a bounds") (UB: "('i,'a) state \ 'a bounds") (LI: "('i,'a) state \ 'i bound_index") (UI: "('i,'a) state \ 'i bound_index") (UBI_upd: "(('i,'a) bounds_index \ ('i,'a) bounds_index) \ ('i,'a) state \ ('i,'a) state") (LE: "var \ 'a \ 'a atom") (GE: "var \ 'a \ 'a atom") (le_rat: "rat \ rat \ bool") definition Positive where [simp]: "Positive \ Direction (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>u_update Leq Geq (\)" definition Negative where [simp]: "Negative \ Direction (>) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>l_update Geq Leq (\)" text\Assuming that the \\\ flag and the current valuation \\\ in the final state determine the solution of a problem, the \assert_all\ function can be reduced to the \assert_all_state\ function that operates on the states: @{text[display] "assert_all t as \ let s = assert_all_state t as in if (\ s) then (False, None) else (True, Some (\ s))" } \ text\Specification for the \assert_all_state\ can be directly obtained from the specification of \assert_all\, and it describes the connection between the valuation in the final state and the initial tableau and atoms. However, we will make an additional refinement step and give stronger assumptions about the \assert_all_state\ function that describes the connection between the initial tableau and atoms with the tableau and bounds in the final state.\ locale AssertAllState = fixes assert_all_state::"tableau \ ('i,'a::lrv) i_atom list \ ('i,'a) state" assumes \ \The final and the initial tableau are equivalent.\ assert_all_state_tableau_equiv: "\ t \ assert_all_state t as = s' \ (v::'a valuation) \\<^sub>t t \ v \\<^sub>t \ s'" and \ \If @{term \} is not raised, then the valuation in the final state satisfies its tableau and its bounds (that are, in this case, equivalent to the set of all asserted bounds).\ assert_all_state_sat: "\ t \ assert_all_state t as = s' \ \ \ s' \ \ s'" and assert_all_state_sat_atoms_equiv_bounds: "\ t \ assert_all_state t as = s' \ \ \ s' \ flat (set as) \ \ s'" and \ \If @{term \} is raised, then there is no valuation satisfying the tableau and the bounds in the final state (that are, in this case, equivalent to a subset of asserted atoms).\ assert_all_state_unsat: "\ t \ assert_all_state t as = s' \ \ s' \ minimal_unsat_state_core s'" and assert_all_state_unsat_atoms_equiv_bounds: "\ t \ assert_all_state t as = s' \ \ s' \ set as \\<^sub>i \\ s'" and \ \The set of indices is taken from the constraints\ assert_all_state_indices: "\ t \ assert_all_state t as = s \ indices_state s \ fst ` set as" and assert_all_index_valid: "\ t \ assert_all_state t as = s \ index_valid (set as) s" begin definition assert_all where "assert_all t as \ let s = assert_all_state t as in if (\ s) then Unsat (the (\\<^sub>c s)) else Sat (\ s)" end text\The \assert_all_state\ function can be implemented by first applying the \init\ function that creates an initial state based on the starting tableau, and then by iteratively applying the \assert\ function for each atom in the starting atoms list.\ text\ \begin{small} \assert_loop as s \ foldl (\ s' a. if (\ s') then s' else assert a s') s as\ \assert_all_state t as \ assert_loop ats (init t)\ \end{small} \ locale Init' = fixes init :: "tableau \ ('i,'a::lrv) state" assumes init'_tableau_normalized: "\ t \ \ (\ (init t))" assumes init'_tableau_equiv: "\ t \ (v::'a valuation) \\<^sub>t t = v \\<^sub>t \ (init t)" assumes init'_sat: "\ t \ \ \ (init t) \ \ (init t)" assumes init'_unsat: "\ t \ \ (init t) \ minimal_unsat_state_core (init t)" assumes init'_atoms_equiv_bounds: "\ t \ {} \ \ (init t)" assumes init'_atoms_imply_bounds_index: "\ t \ {} \\<^sub>i \\ (init t)" text\Specification for \init\ can be obtained from the specification of \asser_all_state\ since all its assumptions must also hold for \init\ (when the list of atoms is empty). Also, since \init\ is the first step in the \assert_all_state\ implementation, the precondition for \init\ the same as for the \assert_all_state\. However, unsatisfiability is never going to be detected during initialization and @{term \} flag is never going to be raised. Also, the tableau in the initial state can just be initialized with the starting tableau. The condition @{term "{} \ \ (init t)"} is equivalent to asking that initial bounds are empty. Therefore, specification for \init\ can be refined to:\ locale Init = fixes init::"tableau \ ('i,'a::lrv) state" assumes \ \Tableau in the initial state for @{text t} is @{text t}:\ init_tableau_id: "\ (init t) = t" and \ \Since unsatisfiability is not detected, @{text \} flag must not be set:\ init_unsat_flag: "\ \ (init t)" and \ \The current valuation must satisfy the tableau:\ init_satisfies_tableau: "\\ (init t)\ \\<^sub>t t" and \ \In an inital state no atoms are yet asserted so the bounds must be empty:\ init_bounds: "\\<^sub>i\<^sub>l (init t) = Mapping.empty" "\\<^sub>i\<^sub>u (init t) = Mapping.empty" and \ \All tableau vars are valuated:\ init_tableau_valuated: "\ (init t)" begin lemma init_satisfies_bounds: "\\ (init t)\ \\<^sub>b \ (init t)" using init_bounds unfolding satisfies_bounds.simps in_bounds.simps bound_compare_defs by (auto simp: boundsl_def boundsu_def) lemma init_satisfies: "\ (init t)" using init_satisfies_tableau init_satisfies_bounds init_tableau_id unfolding curr_val_satisfies_state_def satisfies_state_def by simp lemma init_atoms_equiv_bounds: "{} \ \ (init t)" using init_bounds unfolding atoms_equiv_bounds.simps satisfies_bounds.simps in_bounds.simps satisfies_atom_set_def unfolding bound_compare_defs by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) lemma init_atoms_imply_bounds_index: "{} \\<^sub>i \\ (init t)" using init_bounds unfolding atoms_imply_bounds_index.simps satisfies_bounds_index.simps in_bounds.simps i_satisfies_atom_set.simps satisfies_atom_set_def unfolding bound_compare_defs by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) lemma init_tableau_normalized: "\ t \ \ (\ (init t))" using init_tableau_id by simp lemma init_index_valid: "index_valid as (init t)" using init_bounds unfolding index_valid_def by auto lemma init_indices: "indices_state (init t) = {}" unfolding indices_state_def init_bounds by auto end sublocale Init < Init' init using init_tableau_id init_satisfies init_unsat_flag init_atoms_equiv_bounds init_atoms_imply_bounds_index by unfold_locales auto abbreviation vars_list where "vars_list t \ remdups (map lhs t @ (concat (map (Abstract_Linear_Poly.vars_list \ rhs) t)))" lemma "tvars t = set (vars_list t)" by (auto simp add: set_vars_list lvars_def rvars_def) text\\smallskip The \assert\ function asserts a single atom. Since the \init\ function does not raise the \\\ flag, from the definition of \assert_loop\, it is clear that the flag is not raised when the \assert\ function is called. Moreover, the assumptions about the \assert_all_state\ imply that the loop invariant must be that if the \\\ flag is not raised, then the current valuation must satisfy the state (i.e., \\ s\). The \assert\ function will be more easily implemented if it is always applied to a state with a normalized and valuated tableau, so we make this another loop invariant. Therefore, the precondition for the \assert a s\ function call is that \\ \ s\, \\ s\, \\ (\ s)\ and \\ s\ hold. The specification for \assert\ directly follows from the specification of \assert_all_state\ (except that it is additionally required that bounds reflect asserted atoms also when unsatisfiability is detected, and that it is required that \assert\ keeps the tableau normalized and valuated).\ locale Assert = fixes assert::"('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" assumes \ \Tableau remains equivalent to the previous one and normalized and valuated.\ assert_tableau: "\\ \ s; \ s; \ (\ s); \ s\ \ let s' = assert a s in ((v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ s') \ \ (\ s') \ \ s'" and \ \If the @{term \} flag is not raised, then the current valuation is updated so that it satisfies the current tableau and the current bounds.\ assert_sat: "\\ \ s; \ s; \ (\ s); \ s\ \ \ \ (assert a s) \ \ (assert a s)" and \ \The set of asserted atoms remains equivalent to the bounds in the state.\ assert_atoms_equiv_bounds: "\\ \ s; \ s; \ (\ s); \ s\ \ flat ats \ \ s \ flat (ats \ {a}) \ \ (assert a s)" and \ \There is a subset of asserted atoms which remains index-equivalent to the bounds in the state.\ assert_atoms_imply_bounds_index: "\\ \ s; \ s; \ (\ s); \ s\ \ ats \\<^sub>i \\ s \ insert a ats \\<^sub>i \\ (assert a s)" and \ \If the @{term \} flag is raised, then there is no valuation that satisfies both the current tableau and the current bounds.\ assert_unsat: "\\ \ s; \ s; \ (\ s); \ s; index_valid ats s\ \ \ (assert a s) \ minimal_unsat_state_core (assert a s)" and assert_index_valid: "\\ \ s; \ s; \ (\ s); \ s\ \ index_valid ats s \ index_valid (insert a ats) (assert a s)" begin lemma assert_tableau_equiv: "\\ \ s; \ s; \ (\ s); \ s\ \ (v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ (assert a s)" using assert_tableau by (simp add: Let_def) lemma assert_tableau_normalized: "\\ \ s; \ s; \ (\ s); \ s\ \ \ (\ (assert a s))" using assert_tableau by (simp add: Let_def) lemma assert_tableau_valuated: "\\ \ s; \ s; \ (\ s); \ s\ \ \ (assert a s)" using assert_tableau by (simp add: Let_def) end locale AssertAllState' = Init init + Assert assert for init :: "tableau \ ('i,'a::lrv) state" and assert :: "('i,'a) i_atom \ ('i,'a) state \ ('i,'a) state" begin definition assert_loop where "assert_loop as s \ foldl (\ s' a. if (\ s') then s' else assert a s') s as" definition assert_all_state where [simp]: "assert_all_state t as \ assert_loop as (init t)" lemma AssertAllState'_precond: "\ t \ \ (\ (assert_all_state t as)) \ (\ (assert_all_state t as)) \ (\ \ (assert_all_state t as) \ \ (assert_all_state t as))" unfolding assert_all_state_def assert_loop_def using init_satisfies init_tableau_normalized init_index_valid using assert_sat assert_tableau_normalized init_tableau_valuated assert_tableau_valuated by (induct as rule: rev_induct) auto lemma AssertAllState'Induct: assumes "\ t" "P {} (init t)" "\ as bs t. as \ bs \ P as t \ P bs t" "\ s a as. \\ \ s; \ s; \ (\ s); \ s; P as s; index_valid as s\ \ P (insert a as) (assert a s)" shows "P (set as) (assert_all_state t as)" proof - have "P (set as) (assert_all_state t as) \ index_valid (set as) (assert_all_state t as)" proof (induct as rule: rev_induct) case Nil then show ?case unfolding assert_all_state_def assert_loop_def using assms(2) init_index_valid by auto next case (snoc a as') let ?f = "\s' a. if \ s' then s' else assert a s'" let ?s = "foldl ?f (init t) as'" show ?case proof (cases "\ ?s") case True from snoc index_valid_mono[of _ "set (a # as')" "(assert_all_state t as')"] have index: "index_valid (set (a # as')) (assert_all_state t as')" by auto from snoc assms(3)[of "set as'" "set (a # as')"] have P: "P (set (a # as')) (assert_all_state t as')" by auto show ?thesis using True P index unfolding assert_all_state_def assert_loop_def by simp next case False then show ?thesis using snoc using assms(1) assms(4) using AssertAllState'_precond assert_index_valid unfolding assert_all_state_def assert_loop_def by auto qed qed then show ?thesis .. qed lemma AssertAllState_index_valid: "\ t \ index_valid (set as) (assert_all_state t as)" by (rule AssertAllState'Induct, auto intro: assert_index_valid init_index_valid index_valid_mono) lemma AssertAllState'_sat_atoms_equiv_bounds: "\ t \ \ \ (assert_all_state t as) \ flat (set as) \ \ (assert_all_state t as)" using AssertAllState'_precond using init_atoms_equiv_bounds assert_atoms_equiv_bounds unfolding assert_all_state_def assert_loop_def by (induct as rule: rev_induct) auto lemma AssertAllState'_unsat_atoms_implies_bounds: assumes "\ t" shows "set as \\<^sub>i \\ (assert_all_state t as)" proof (induct as rule: rev_induct) case Nil then show ?case using assms init_atoms_imply_bounds_index unfolding assert_all_state_def assert_loop_def by simp next case (snoc a as') let ?s = "assert_all_state t as'" show ?case proof (cases "\ ?s") case True then show ?thesis using snoc atoms_imply_bounds_index_mono[of "set as'" "set (as' @ [a])"] unfolding assert_all_state_def assert_loop_def by auto next case False then have id: "assert_all_state t (as' @ [a]) = assert a ?s" unfolding assert_all_state_def assert_loop_def by simp from snoc have as': "set as' \\<^sub>i \\ ?s" by auto from AssertAllState'_precond[of t as'] assms False have "\ ?s" "\ (\ ?s)" "\ ?s" by auto from assert_atoms_imply_bounds_index[OF False this as', of a] show ?thesis unfolding id by auto qed qed end text\Under these assumptions, it can easily be shown (mainly by induction) that the previously shown implementation of \assert_all_state\ satisfies its specification.\ sublocale AssertAllState' < AssertAllState assert_all_state proof fix v::"'a valuation" and t as s' assume *: "\ t" and id: "assert_all_state t as = s'" note idsym = id[symmetric] show "v \\<^sub>t t = v \\<^sub>t \ s'" unfolding idsym using init_tableau_id[of t] assert_tableau_equiv[of _ v] by (induct rule: AssertAllState'Induct) (auto simp add: * ) show "\ \ s' \ \ s'" unfolding idsym using AssertAllState'_precond by (simp add: * ) show "\ \ s' \ flat (set as) \ \ s'" unfolding idsym using * by (rule AssertAllState'_sat_atoms_equiv_bounds) show "\ s' \ set as \\<^sub>i \\ s'" using * unfolding idsym by (rule AssertAllState'_unsat_atoms_implies_bounds) show "\ s' \ minimal_unsat_state_core s'" using init_unsat_flag assert_unsat assert_index_valid unfolding idsym by (induct rule: AssertAllState'Induct) (auto simp add: * ) show "indices_state s' \ fst ` set as" unfolding idsym using * by (intro index_valid_indices_state, induct rule: AssertAllState'Induct, auto simp: init_index_valid index_valid_mono assert_index_valid) show "index_valid (set as) s'" using "*" AssertAllState_index_valid idsym by blast qed subsection\Asserting Single Atoms\ text\The @{term assert} function is split in two phases. First, @{term assert_bound} updates the bounds and checks only for conflicts cheap to detect. Next, \check\ performs the full simplex algorithm. The \assert\ function can be implemented as \assert a s = check (assert_bound a s)\. Note that it is also possible to do the first phase for several asserted atoms, and only then to let the expensive second phase work. \medskip Asserting an atom \x \ b\ begins with the function \assert_bound\. If the atom is subsumed by the current bounds, then no changes are performed. Otherwise, bounds for \x\ are changed to incorporate the atom. If the atom is inconsistent with the previous bounds for \x\, the @{term \} flag is raised. If \x\ is not a lhs variable in the current tableau and if the value for \x\ in the current valuation violates the new bound \b\, the value for \x\ can be updated and set to \b\, meanwhile updating the values for lhs variables of the tableau so that it remains satisfied. Otherwise, no changes to the current valuation are performed.\ fun satisfies_bounds_set :: "'a::linorder valuation \ 'a bounds \ 'a bounds \ var set \ bool" where "satisfies_bounds_set v (lb, ub) S \ (\ x \ S. in_bounds x v (lb, ub))" declare satisfies_bounds_set.simps [simp del] syntax "_satisfies_bounds_set" :: "(var \ 'a::linorder) \ 'a bounds \ 'a bounds \ var set \ bool" ("_ \\<^sub>b _ \/ _") translations "v \\<^sub>b b \ S" == "CONST satisfies_bounds_set v b S" lemma satisfies_bounds_set_iff: "v \\<^sub>b (lb, ub) \ S \ (\ x \ S. v x \\<^sub>l\<^sub>b lb x \ v x \\<^sub>u\<^sub>b ub x)" by (simp add: satisfies_bounds_set.simps) definition curr_val_satisfies_no_lhs ("\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s") where "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \\ s\ \\<^sub>t (\ s) \ (\\ s\ \\<^sub>b (\ s) \ (- lvars (\ s)))" lemma satisfies_satisfies_no_lhs: "\ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" by (simp add: curr_val_satisfies_state_def satisfies_state_def curr_val_satisfies_no_lhs_def satisfies_bounds.simps satisfies_bounds_set.simps) definition bounds_consistent :: "('i,'a::linorder) state \ bool" ("\") where "\ s \ \ x. if \\<^sub>l s x = None \ \\<^sub>u s x = None then True else the (\\<^sub>l s x) \ the (\\<^sub>u s x)" text\So, the \assert_bound\ function must ensure that the given atom is included in the bounds, that the tableau remains satisfied by the valuation and that all variables except the lhs variables in the tableau are within their bounds. To formalize this, we introduce the notation \v \\<^sub>b (lb, ub) \ S\, and define @{thm satisfies_bounds_set_iff[no_vars]}, and @{thm curr_val_satisfies_no_lhs_def[no_vars]}. The \assert_bound\ function raises the \\\ flag if and only if lower and upper bounds overlap. This is formalized as @{thm bounds_consistent_def[no_vars]}.\ lemma satisfies_bounds_consistent: "(v::'a::linorder valuation) \\<^sub>b \ s \ \ s" unfolding satisfies_bounds.simps in_bounds.simps bounds_consistent_def bound_compare_defs by (auto split: option.split) force lemma satisfies_consistent: "\ s \ \ s" by (auto simp add: curr_val_satisfies_state_def satisfies_state_def satisfies_bounds_consistent) lemma bounds_consistent_geq_lb: "\\ s; \\<^sub>u s x\<^sub>i = Some c\ \ c \\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i" unfolding bounds_consistent_def by (cases "\\<^sub>l s x\<^sub>i", auto simp add: bound_compare_defs split: if_splits) (erule_tac x="x\<^sub>i" in allE, auto) lemma bounds_consistent_leq_ub: "\\ s; \\<^sub>l s x\<^sub>i = Some c\ \ c \\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i" unfolding bounds_consistent_def by (cases "\\<^sub>u s x\<^sub>i", auto simp add: bound_compare_defs split: if_splits) (erule_tac x="x\<^sub>i" in allE, auto) lemma bounds_consistent_gt_ub: "\\ s; c <\<^sub>l\<^sub>b \\<^sub>l s x \ \ \ c >\<^sub>u\<^sub>b \\<^sub>u s x" unfolding bounds_consistent_def by (case_tac[!] "\\<^sub>l s x", case_tac[!] "\\<^sub>u s x") (auto simp add: bound_compare_defs, erule_tac x="x" in allE, simp) lemma bounds_consistent_lt_lb: "\\ s; c >\<^sub>u\<^sub>b \\<^sub>u s x \ \ \ c <\<^sub>l\<^sub>b \\<^sub>l s x" unfolding bounds_consistent_def by (case_tac[!] "\\<^sub>l s x", case_tac[!] "\\<^sub>u s x") (auto simp add: bound_compare_defs, erule_tac x="x" in allE, simp) text\Since the \assert_bound\ is the first step in the \assert\ function implementation, the preconditions for \assert_bound\ are the same as preconditions for the \assert\ function. The specifiction for the \assert_bound\ is:\ locale AssertBound = fixes assert_bound::"('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" assumes \ \The tableau remains unchanged and valuated.\ assert_bound_tableau: "\\ \ s; \ s; \ (\ s); \ s\ \ assert_bound a s = s' \ \ s' = \ s \ \ s'" and \ \If the @{term \} flag is not set, all but the lhs variables in the tableau remain within their bounds, the new valuation satisfies the tableau, and bounds do not overlap.\ assert_bound_sat: "\\ \ s; \ s; \ (\ s); \ s\ \ assert_bound a s = s' \ \ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ s'" and \ \The set of asserted atoms remains equivalent to the bounds in the state.\ assert_bound_atoms_equiv_bounds: "\\ \ s; \ s; \ (\ s); \ s\ \ flat ats \ \ s \ flat (ats \ {a}) \ \ (assert_bound a s)" and assert_bound_atoms_imply_bounds_index: "\\ \ s; \ s; \ (\ s); \ s\ \ ats \\<^sub>i \\ s \ insert a ats \\<^sub>i \\ (assert_bound a s)" and \ \@{term \} flag is raised, only if the bounds became inconsistent:\ assert_bound_unsat: "\\ \ s; \ s; \ (\ s); \ s\ \ index_valid as s \ assert_bound a s = s' \ \ s' \ minimal_unsat_state_core s'" and assert_bound_index_valid: "\\ \ s; \ s; \ (\ s); \ s\ \ index_valid as s \ index_valid (insert a as) (assert_bound a s)" begin lemma assert_bound_tableau_id: "\\ \ s; \ s; \ (\ s); \ s\ \ \ (assert_bound a s) = \ s" using assert_bound_tableau by blast lemma assert_bound_tableau_valuated: "\\ \ s; \ s; \ (\ s); \ s\ \ \ (assert_bound a s)" using assert_bound_tableau by blast end locale AssertBoundNoLhs = fixes assert_bound :: "('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" assumes assert_bound_nolhs_tableau_id: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ \ (assert_bound a s) = \ s" assumes assert_bound_nolhs_sat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ \ \ (assert_bound a s) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (assert_bound a s) \ \ (assert_bound a s)" assumes assert_bound_nolhs_atoms_equiv_bounds: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ flat ats \ \ s \ flat (ats \ {a}) \ \ (assert_bound a s)" assumes assert_bound_nolhs_atoms_imply_bounds_index: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ ats \\<^sub>i \\ s \ insert a ats \\<^sub>i \\ (assert_bound a s)" assumes assert_bound_nolhs_unsat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ index_valid as s \ \ (assert_bound a s) \ minimal_unsat_state_core (assert_bound a s)" assumes assert_bound_nolhs_tableau_valuated: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ \ (assert_bound a s)" assumes assert_bound_nolhs_index_valid: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ index_valid as s \ index_valid (insert a as) (assert_bound a s)" sublocale AssertBoundNoLhs < AssertBound by unfold_locales ((metis satisfies_satisfies_no_lhs satisfies_consistent assert_bound_nolhs_tableau_id assert_bound_nolhs_sat assert_bound_nolhs_atoms_equiv_bounds assert_bound_nolhs_index_valid assert_bound_nolhs_atoms_imply_bounds_index assert_bound_nolhs_unsat assert_bound_nolhs_tableau_valuated)+) text\The second phase of \assert\, the \check\ function, is the heart of the Simplex algorithm. It is always called after \assert_bound\, but in two different situations. In the first case \assert_bound\ raised the \\\ flag and then \check\ should retain the flag and should not perform any changes. In the second case \assert_bound\ did not raise the \\\ flag, so \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\, \\ s\, \\ (\ s)\, and \\ s\ hold.\ locale Check = fixes check::"('i,'a::lrv) state \ ('i,'a) state" assumes \ \If @{text check} is called from an inconsistent state, the state is unchanged.\ check_unsat_id: "\ s \ check s = s" and \ \The tableau remains equivalent to the previous one, normalized and valuated.\ check_tableau: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ let s' = check s in ((v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ s') \ \ (\ s') \ \ s'" and \ \The bounds remain unchanged.\ check_bounds_id: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \\<^sub>i (check s) = \\<^sub>i s" and \ \If @{term \} flag is not raised, the current valuation @{text "\"} satisfies both the tableau and the bounds and if it is raised, there is no valuation that satisfies them.\ check_sat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ \ (check s) \ \ (check s)" and check_unsat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ (check s) \ minimal_unsat_state_core (check s)" begin lemma check_tableau_equiv: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ (v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ (check s)" using check_tableau by (simp add: Let_def) lemma check_tableau_index_valid: assumes "\ \ s" " \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" "\ (\ s)" "\ s" shows "index_valid as (check s) = index_valid as s" unfolding index_valid_def using check_bounds_id[OF assms] by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) lemma check_tableau_normalized: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ (\ (check s))" using check_tableau by (simp add: Let_def) lemma check_tableau_valuated: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ (check s)" using check_tableau by (simp add: Let_def) lemma check_indices_state: assumes "\ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ \ s \ \ s" "\ \ s \ \ (\ s)" "\ \ s \ \ s" shows "indices_state (check s) = indices_state s" using check_bounds_id[OF _ assms] check_unsat_id[of s] unfolding indices_state_def by (cases "\ s", auto) lemma check_distinct_indices_state: assumes "\ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ \ s \ \ s" "\ \ s \ \ (\ s)" "\ \ s \ \ s" shows "distinct_indices_state (check s) = distinct_indices_state s" using check_bounds_id[OF _ assms] check_unsat_id[of s] unfolding distinct_indices_state_def by (cases "\ s", auto) end locale Assert' = AssertBound assert_bound + Check check for assert_bound :: "('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" and check :: "('i,'a::lrv) state \ ('i,'a) state" begin definition assert :: "('i,'a) i_atom \ ('i,'a) state \ ('i,'a) state" where "assert a s \ check (assert_bound a s)" lemma Assert'Precond: assumes "\ \ s" "\ s" "\ (\ s)" "\ s" shows "\ (\ (assert_bound a s))" "\ \ (assert_bound a s) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (assert_bound a s) \ \ (assert_bound a s)" "\ (assert_bound a s)" using assms using assert_bound_tableau_id assert_bound_sat assert_bound_tableau_valuated by (auto simp add: satisfies_bounds_consistent Let_def) end sublocale Assert' < Assert assert proof fix s::"('i,'a) state" and v::"'a valuation" and a::"('i,'a) i_atom" assume *: "\ \ s" "\ s" "\ (\ s)" "\ s" have "\ (\ (assert a s))" using check_tableau_normalized[of "assert_bound a s"] check_unsat_id[of "assert_bound a s"] * using assert_bound_sat[of s a] Assert'Precond[of s a] by (force simp add: assert_def) moreover have "v \\<^sub>t \ s = v \\<^sub>t \ (assert a s)" using check_tableau_equiv[of "assert_bound a s" v] * using check_unsat_id[of "assert_bound a s"] by (force simp add: assert_def Assert'Precond assert_bound_sat assert_bound_tableau_id) moreover have "\ (assert a s)" using assert_bound_tableau_valuated[of s a] * using check_tableau_valuated[of "assert_bound a s"] using check_unsat_id[of "assert_bound a s"] by (cases "\ (assert_bound a s)") (auto simp add: Assert'Precond assert_def) ultimately show "let s' = assert a s in (v \\<^sub>t \ s = v \\<^sub>t \ s') \ \ (\ s') \ \ s'" by (simp add: Let_def) next fix s::"('i,'a) state" and a::"('i,'a) i_atom" assume "\ \ s" "\ s" "\ (\ s)" "\ s" then show "\ \ (assert a s) \ \ (assert a s)" unfolding assert_def using check_unsat_id[of "assert_bound a s"] using check_sat[of "assert_bound a s"] by (force simp add: Assert'Precond) next fix s::"('i,'a) state" and a::"('i,'a) i_atom" and ats::"('i,'a) i_atom set" assume "\ \ s" "\ s" "\ (\ s)" "\ s" then show "flat ats \ \ s \ flat (ats \ {a}) \ \ (assert a s)" using assert_bound_atoms_equiv_bounds using check_unsat_id[of "assert_bound a s"] check_bounds_id by (cases "\ (assert_bound a s)") (auto simp add: Assert'Precond assert_def simp: indexl_def indexu_def boundsl_def boundsu_def) next fix s::"('i,'a) state" and a::"('i,'a) i_atom" and ats assume *: "\ \ s" "\ s" "\ (\ s)" "\ s" "\ (assert a s)" "index_valid ats s" show "minimal_unsat_state_core (assert a s)" proof (cases "\ (assert_bound a s)") case True then show ?thesis unfolding assert_def using * assert_bound_unsat check_tableau_equiv[of "assert_bound a s"] check_bounds_id using check_unsat_id[of "assert_bound a s"] by (auto simp add: Assert'Precond satisfies_state_def Let_def) next case False then show ?thesis unfolding assert_def using * assert_bound_sat[of s a] check_unsat Assert'Precond by (metis assert_def) qed next fix ats fix s::"('i,'a) state" and a::"('i,'a) i_atom" assume *: "index_valid ats s" and **: "\ \ s" "\ s" "\ (\ s)" "\ s" have *: "index_valid (insert a ats) (assert_bound a s)" using assert_bound_index_valid[OF ** *] . show "index_valid (insert a ats) (assert a s)" proof (cases "\ (assert_bound a s)") case True show ?thesis unfolding assert_def check_unsat_id[OF True] using * . next case False show ?thesis unfolding assert_def using Assert'Precond[OF **, of a] False * by (subst check_tableau_index_valid[OF False], auto) qed next fix s ats a let ?s = "assert_bound a s" assume *: "\ \ s" "\ s" "\ (\ s)" "\ s" "ats \\<^sub>i \\ s" from assert_bound_atoms_imply_bounds_index[OF this, of a] have as: "insert a ats \\<^sub>i \\ (assert_bound a s)" by auto show "insert a ats \\<^sub>i \\ (assert a s)" proof (cases "\ ?s") case True from check_unsat_id[OF True] as show ?thesis unfolding assert_def by auto next case False from assert_bound_tableau_id[OF *(1-4)] * have t: "\ (\ ?s)" by simp from assert_bound_tableau_valuated[OF *(1-4)] have v: "\ ?s" . from assert_bound_sat[OF *(1-4) refl False] have "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s" "\ ?s" by auto from check_bounds_id[OF False this t v] as show ?thesis unfolding assert_def by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) qed qed text\Under these assumptions for \assert_bound\ and \check\, it can be easily shown that the implementation of \assert\ (previously given) satisfies its specification.\ locale AssertAllState'' = Init init + AssertBoundNoLhs assert_bound + Check check for init :: "tableau \ ('i,'a::lrv) state" and assert_bound :: "('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" and check :: "('i,'a::lrv) state \ ('i,'a) state" begin definition assert_bound_loop where "assert_bound_loop ats s \ foldl (\ s' a. if (\ s') then s' else assert_bound a s') s ats" definition assert_all_state where [simp]: "assert_all_state t ats \ check (assert_bound_loop ats (init t))" text\However, for efficiency reasons, we want to allow implementations that delay the \check\ function call and call it after several \assert_bound\ calls. For example: \smallskip \begin{small} @{thm assert_bound_loop_def[no_vars]} @{thm assert_all_state_def[no_vars]} \end{small} \smallskip Then, the loop consists only of \assert_bound\ calls, so \assert_bound\ postcondition must imply its precondition. This is not the case, since variables on the lhs may be out of their bounds. Therefore, we make a refinement and specify weaker preconditions (replace \\ s\, by \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ and \\ s\) for \assert_bound\, and show that these preconditions are still good enough to prove the correctnes of this alternative \assert_all_state\ definition.\ lemma AssertAllState''_precond': assumes "\ (\ s)" "\ s" "\ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \ s" shows "let s' = assert_bound_loop ats s in \ (\ s') \ \ s' \ (\ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ s')" using assms using assert_bound_nolhs_tableau_id assert_bound_nolhs_sat assert_bound_nolhs_tableau_valuated by (induct ats rule: rev_induct) (auto simp add: assert_bound_loop_def Let_def) lemma AssertAllState''_precond: assumes "\ t" shows "let s' = assert_bound_loop ats (init t) in \ (\ s') \ \ s' \ (\ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ s')" using assms using AssertAllState''_precond'[of "init t" ats] by (simp add: Let_def init_tableau_id init_unsat_flag init_satisfies satisfies_consistent satisfies_satisfies_no_lhs init_tableau_valuated) lemma AssertAllState''Induct: assumes "\ t" "P {} (init t)" "\ as bs t. as \ bs \ P as t \ P bs t" "\ s a ats. \\ \ s; \\ s\ \\<^sub>t \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s; P (set ats) s; index_valid (set ats) s\ \ P (insert a (set ats)) (assert_bound a s)" shows "P (set ats) (assert_bound_loop ats (init t))" proof - have "P (set ats) (assert_bound_loop ats (init t)) \ index_valid (set ats) (assert_bound_loop ats (init t))" proof (induct ats rule: rev_induct) case Nil then show ?case unfolding assert_bound_loop_def using assms(2) init_index_valid by simp next case (snoc a as') let ?s = "assert_bound_loop as' (init t)" from snoc index_valid_mono[of _ "set (a # as')" "assert_bound_loop as' (init t)"] have index: "index_valid (set (a # as')) (assert_bound_loop as' (init t))" by auto from snoc assms(3)[of "set as'" "set (a # as')"] have P: "P (set (a # as')) (assert_bound_loop as' (init t))" by auto show ?case proof (cases "\ ?s") case True then show ?thesis using P index unfolding assert_bound_loop_def by simp next case False have id': "set (as' @ [a]) = insert a (set as')" by simp have id: "assert_bound_loop (as' @ [a]) (init t) = assert_bound a (assert_bound_loop as' (init t))" using False unfolding assert_bound_loop_def by auto from snoc have index: "index_valid (set as') ?s" by simp show ?thesis unfolding id unfolding id' using False snoc AssertAllState''_precond[OF assms(1)] by (intro conjI assert_bound_nolhs_index_valid index assms(4); (force simp: Let_def curr_val_satisfies_no_lhs_def)?) qed qed then show ?thesis .. qed lemma AssertAllState''_tableau_id: "\ t \ \ (assert_bound_loop ats (init t)) = \ (init t)" by (rule AssertAllState''Induct) (auto simp add: init_tableau_id assert_bound_nolhs_tableau_id) lemma AssertAllState''_sat: "\ t \ \ \ (assert_bound_loop ats (init t)) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (assert_bound_loop ats (init t)) \ \ (assert_bound_loop ats (init t))" by (rule AssertAllState''Induct) (auto simp add: init_unsat_flag init_satisfies satisfies_consistent satisfies_satisfies_no_lhs assert_bound_nolhs_sat) lemma AssertAllState''_unsat: "\ t \ \ (assert_bound_loop ats (init t)) \ minimal_unsat_state_core (assert_bound_loop ats (init t))" by (rule AssertAllState''Induct) (auto simp add: init_tableau_id assert_bound_nolhs_unsat init_unsat_flag) lemma AssertAllState''_sat_atoms_equiv_bounds: "\ t \ \ \ (assert_bound_loop ats (init t)) \ flat (set ats) \ \ (assert_bound_loop ats (init t))" using AssertAllState''_precond using assert_bound_nolhs_atoms_equiv_bounds init_atoms_equiv_bounds by (induct ats rule: rev_induct) (auto simp add: Let_def assert_bound_loop_def) lemma AssertAllState''_atoms_imply_bounds_index: assumes "\ t" shows "set ats \\<^sub>i \\ (assert_bound_loop ats (init t))" proof (induct ats rule: rev_induct) case Nil then show ?case unfolding assert_bound_loop_def using init_atoms_imply_bounds_index assms by simp next case (snoc a ats') let ?s = "assert_bound_loop ats' (init t)" show ?case proof (cases "\ ?s") case True then show ?thesis using snoc atoms_imply_bounds_index_mono[of "set ats'" "set (ats' @ [a])"] unfolding assert_bound_loop_def by auto next case False then have id: "assert_bound_loop (ats' @ [a]) (init t) = assert_bound a ?s" unfolding assert_bound_loop_def by auto from snoc have ats: "set ats' \\<^sub>i \\ ?s" by auto from AssertAllState''_precond[of t ats', OF assms, unfolded Let_def] False have *: "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s" "\ (\ ?s)" "\ ?s" "\ ?s" by auto show ?thesis unfolding id using assert_bound_nolhs_atoms_imply_bounds_index[OF False * ats, of a] by auto qed qed lemma AssertAllState''_index_valid: "\ t \ index_valid (set ats) (assert_bound_loop ats (init t))" by (rule AssertAllState''Induct, auto simp: init_index_valid index_valid_mono assert_bound_nolhs_index_valid) end sublocale AssertAllState'' < AssertAllState assert_all_state proof fix v::"'a valuation" and t ats s' assume *: "\ t" and "assert_all_state t ats = s'" then have s': "s' = assert_all_state t ats" by simp let ?s' = "assert_bound_loop ats (init t)" show "v \\<^sub>t t = v \\<^sub>t \ s'" unfolding assert_all_state_def s' using * check_tableau_equiv[of ?s' v] AssertAllState''_tableau_id[of t ats] init_tableau_id[of t] using AssertAllState''_sat[of t ats] check_unsat_id[of ?s'] using AssertAllState''_precond[of t ats] by force show "\ \ s' \ \ s'" unfolding assert_all_state_def s' using * AssertAllState''_precond[of t ats] using check_sat check_unsat_id by (force simp add: Let_def) show "\ s' \ minimal_unsat_state_core s'" using * check_unsat check_unsat_id[of ?s'] check_bounds_id using AssertAllState''_unsat[of t ats] AssertAllState''_precond[of t ats] s' by (force simp add: Let_def satisfies_state_def) show "\ \ s' \ flat (set ats) \ \ s'" unfolding assert_all_state_def s' using * AssertAllState''_precond[of t ats] using check_bounds_id[of ?s'] check_unsat_id[of ?s'] using AssertAllState''_sat_atoms_equiv_bounds[of t ats] by (force simp add: Let_def simp: indexl_def indexu_def boundsl_def boundsu_def) show "\ s' \ set ats \\<^sub>i \\ s'" unfolding assert_all_state_def s' using * AssertAllState''_precond[of t ats] unfolding Let_def using check_bounds_id[of ?s'] using AssertAllState''_atoms_imply_bounds_index[of t ats] using check_unsat_id[of ?s'] by (cases "\ ?s'") (auto simp add: Let_def simp: indexl_def indexu_def boundsl_def boundsu_def) show "index_valid (set ats) s'" unfolding assert_all_state_def s' using * AssertAllState''_precond[of t ats] AssertAllState''_index_valid[OF *, of ats] unfolding Let_def using check_tableau_index_valid[of ?s'] using check_unsat_id[of ?s'] by (cases "\ ?s'", auto) show "indices_state s' \ fst ` set ats" by (intro index_valid_indices_state, fact) qed subsection\Update and Pivot\ text\Both \assert_bound\ and \check\ need to update the valuation so that the tableau remains satisfied. If the value for a variable not on the lhs of the tableau is changed, this can be done rather easily (once the value of that variable is changed, one should recalculate and change the values for all lhs variables of the tableau). The \update\ function does this, and it is specified by:\ locale Update = fixes update::"var \ 'a::lrv \ ('i,'a) state \ ('i,'a) state" assumes \ \Tableau, bounds, and the unsatisfiability flag are preserved.\ update_id: "\\ (\ s); \ s; x \ lvars (\ s)\ \ let s' = update x c s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>c s" and \ \Tableau remains valuated.\ update_tableau_valuated: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \ (update x v s)" and \ \The given variable @{text "x"} in the updated valuation is set to the given value @{text "v"} while all other variables (except those on the lhs of the tableau) are unchanged.\ update_valuation_nonlhs: "\\ (\ s); \ s; x \ lvars (\ s)\ \ x' \ lvars (\ s) \ look (\ (update x v s)) x' = (if x = x' then Some v else look (\ s) x')" and \ \Updated valuation continues to satisfy the tableau.\ update_satisfies_tableau: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \\ s\ \\<^sub>t \ s \ \\ (update x c s)\ \\<^sub>t \ s" begin lemma update_bounds_id: assumes "\ (\ s)" "\ s" "x \ lvars (\ s)" shows "\\<^sub>i (update x c s) = \\<^sub>i s" "\\ (update x c s) = \\ s" "\\<^sub>l (update x c s) = \\<^sub>l s" "\\<^sub>u (update x c s) = \\<^sub>u s" using update_id assms by (auto simp add: Let_def simp: indexl_def indexu_def boundsl_def boundsu_def) lemma update_indices_state_id: assumes "\ (\ s)" "\ s" "x \ lvars (\ s)" shows "indices_state (update x c s) = indices_state s" using update_bounds_id[OF assms] unfolding indices_state_def by auto lemma update_tableau_id: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \ (update x c s) = \ s" using update_id by (auto simp add: Let_def) lemma update_unsat_id: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \ (update x c s) = \ s" using update_id by (auto simp add: Let_def) lemma update_unsat_core_id: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \\<^sub>c (update x c s) = \\<^sub>c s" using update_id by (auto simp add: Let_def) definition assert_bound' where [simp]: "assert_bound' dir i x c s \ (if (\\<^sub>u\<^sub>b (lt dir)) c (UB dir s x) then s else let s' = update\\ (UBI_upd dir) i x c s in if (\\<^sub>l\<^sub>b (lt dir)) c ((LB dir) s x) then set_unsat [i, ((LI dir) s x)] s' else if x \ lvars (\ s') \ (lt dir) c (\\ s\ x) then update x c s' else s')" fun assert_bound :: "('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" where "assert_bound (i,Leq x c) s = assert_bound' Positive i x c s" | "assert_bound (i,Geq x c) s = assert_bound' Negative i x c s" lemma assert_bound'_cases: assumes "\\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \ P s" assumes "\\ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)\ \ P (set_unsat [i, ((LI dir) s x)] (update\\ (UBI_upd dir) i x c s))" assumes "\x \ lvars (\ s); (lt dir) c (\\ s\ x); \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x))\ \ P (update x c (update\\ (UBI_upd dir) i x c s))" assumes "\\ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)); x \ lvars (\ s)\ \ P (update\\ (UBI_upd dir) i x c s)" assumes "\\ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)); \ ((lt dir) c (\\ s\ x))\ \ P (update\\ (UBI_upd dir) i x c s)" assumes "dir = Positive \ dir = Negative" shows "P (assert_bound' dir i x c s)" proof (cases "\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)") case True then show ?thesis using assms(1) by simp next case False show ?thesis proof (cases "\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)") case True then show ?thesis using \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ using assms(2) by simp next case False let ?s = "update\\ (UBI_upd dir) i x c s" show ?thesis proof (cases "x \ lvars (\ ?s) \ (lt dir) c (\\ s\ x)") case True then show ?thesis using \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \\ \\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)\ using assms(3) assms(6) by auto next case False then have "x \ lvars (\ ?s) \ \ ((lt dir) c (\\ s\ x))" by simp then show ?thesis proof assume "x \ lvars (\ ?s)" then show ?thesis using \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \\ \\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)\ using assms(4) assms(6) by auto next assume "\ (lt dir) c (\\ s\ x)" then show ?thesis using \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \\ \\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)\ using assms(5) assms(6) by simp qed qed qed qed lemma assert_bound_cases: assumes "\ c x dir. \ dir = Positive \ dir = Negative; a = LE dir x c; \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x) \ \ P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) s" assumes "\ c x dir. \dir = Positive \ dir = Negative; a = LE dir x c; \ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x); \\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x) \ \ P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) (set_unsat [i, ((LI dir) s x)] (update\\ (UBI_upd dir) i x c s))" assumes "\ c x dir. \ dir = Positive \ dir = Negative; a = LE dir x c; x \ lvars (\ s); (lt dir) c (\\ s\ x); \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)) \ \ P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) (update x c ((update\\ (UBI_upd dir) i x c s)))" assumes "\ c x dir. \ dir = Positive \ dir = Negative; a = LE dir x c; x \ lvars (\ s); \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)) \ \ P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) ((update\\ (UBI_upd dir) i x c s))" assumes "\ c x dir. \ dir = Positive \ dir = Negative; a = LE dir x c; \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)); \ ((lt dir) c (\\ s\ x)) \ \ P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) ((update\\ (UBI_upd dir) i x c s))" assumes "\ s. P s = P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s" assumes "\ s. P s = P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s" shows "P (assert_bound (i,a) s)" proof (cases a) case (Leq x c) then show ?thesis apply (simp del: assert_bound'_def) apply (rule assert_bound'_cases, simp_all) using assms(1)[of Positive x c] using assms(2)[of Positive x c] using assms(3)[of Positive x c] using assms(4)[of Positive x c] using assms(5)[of Positive x c] using assms(7) by auto next case (Geq x c) then show ?thesis apply (simp del: assert_bound'_def) apply (rule assert_bound'_cases) using assms(1)[of Negative x c] using assms(2)[of Negative x c] using assms(3)[of Negative x c] using assms(4)[of Negative x c] using assms(5)[of Negative x c] using assms(6) by auto qed end lemma set_unsat_bounds_id: "\ (set_unsat I s) = \ s" unfolding boundsl_def boundsu_def by auto lemma decrease_ub_satisfied_inverse: assumes lt: "\\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" and dir: "dir = Positive \ dir = Negative" assumes v: "v \\<^sub>b \ (update\\ (UBI_upd dir) i x c s)" shows "v \\<^sub>b \ s" unfolding satisfies_bounds.simps proof fix x' show "in_bounds x' v (\ s)" proof (cases "x = x'") case False then show ?thesis using v dir unfolding satisfies_bounds.simps by (auto split: if_splits simp: boundsl_def boundsu_def) next case True then show ?thesis using v dir unfolding satisfies_bounds.simps using lt by (erule_tac x="x'" in allE) (auto simp add: lt_ubound_def[THEN sym] gt_lbound_def[THEN sym] compare_strict_nonstrict boundsl_def boundsu_def) qed qed lemma atoms_equiv_bounds_extend: fixes x c dir assumes "dir = Positive \ dir = Negative" "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "ats \ \ s" shows "(ats \ {LE dir x c}) \ \ (update\\ (UBI_upd dir) i x c s)" unfolding atoms_equiv_bounds.simps proof fix v let ?s = "update\\ (UBI_upd dir) i x c s" show "v \\<^sub>a\<^sub>s (ats \ {LE dir x c}) = v \\<^sub>b \ ?s" proof assume "v \\<^sub>a\<^sub>s (ats \ {LE dir x c})" then have "v \\<^sub>a\<^sub>s ats" "le (lt dir) (v x) c" using \dir = Positive \ dir = Negative\ unfolding satisfies_atom_set_def by auto show "v \\<^sub>b \ ?s" unfolding satisfies_bounds.simps proof fix x' show "in_bounds x' v (\ ?s)" using \v \\<^sub>a\<^sub>s ats\ \le (lt dir) (v x) c\ \ats \ \ s\ using \dir = Positive \ dir = Negative\ unfolding atoms_equiv_bounds.simps satisfies_bounds.simps by (cases "x = x'") (auto simp: boundsl_def boundsu_def) qed next assume "v \\<^sub>b \ ?s" then have "v \\<^sub>b \ s" using \\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)\ using \dir = Positive \ dir = Negative\ using decrease_ub_satisfied_inverse[of dir c s x v] using neg_bounds_compare(1)[of c "\\<^sub>u s x"] using neg_bounds_compare(5)[of c "\\<^sub>l s x"] by (auto simp add: lt_ubound_def[THEN sym] ge_ubound_def[THEN sym] le_lbound_def[THEN sym] gt_lbound_def[THEN sym]) show "v \\<^sub>a\<^sub>s (ats \ {LE dir x c})" unfolding satisfies_atom_set_def proof fix a assume "a \ ats \ {LE dir x c}" then show "v \\<^sub>a a" proof assume "a \ {LE dir x c}" then show ?thesis using \v \\<^sub>b \ ?s\ using \dir = Positive \ dir = Negative\ unfolding satisfies_bounds.simps by (auto split: if_splits simp: boundsl_def boundsu_def) next assume "a \ ats" then show ?thesis using \ats \ \ s\ using \v \\<^sub>b \ s\ unfolding atoms_equiv_bounds.simps satisfies_atom_set_def by auto qed qed qed qed lemma bounds_updates: "\\<^sub>l (\\<^sub>i\<^sub>u_update u s) = \\<^sub>l s" "\\<^sub>u (\\<^sub>i\<^sub>l_update u s) = \\<^sub>u s" "\\<^sub>u (\\<^sub>i\<^sub>u_update (upd x (i,c)) s) = (\\<^sub>u s) (x \ c)" "\\<^sub>l (\\<^sub>i\<^sub>l_update (upd x (i,c)) s) = (\\<^sub>l s) (x \ c)" by (auto simp: boundsl_def boundsu_def) locale EqForLVar = fixes eq_idx_for_lvar :: "tableau \ var \ nat" assumes eq_idx_for_lvar: "\x \ lvars t\ \ eq_idx_for_lvar t x < length t \ lhs (t ! eq_idx_for_lvar t x) = x" begin definition eq_for_lvar :: "tableau \ var \ eq" where "eq_for_lvar t v \ t ! (eq_idx_for_lvar t v)" lemma eq_for_lvar: "\x \ lvars t\ \ eq_for_lvar t x \ set t \ lhs (eq_for_lvar t x) = x" unfolding eq_for_lvar_def using eq_idx_for_lvar by auto abbreviation rvars_of_lvar where "rvars_of_lvar t x \ rvars_eq (eq_for_lvar t x)" lemma rvars_of_lvar_rvars: assumes "x \ lvars t" shows "rvars_of_lvar t x \ rvars t" using assms eq_for_lvar[of x t] unfolding rvars_def by auto end text \Updating changes the value of \x\ and then updates values of all lhs variables so that the tableau remains satisfied. This can be based on a function that recalculates rhs polynomial values in the changed valuation:\ locale RhsEqVal = fixes rhs_eq_val::"(var, 'a::lrv) mapping \ var \ 'a \ eq \ 'a" \ \@{text rhs_eq_val} computes the value of the rhs of @{text e} in @{text "\v\(x := c)"}.\ assumes rhs_eq_val: "\v\ \\<^sub>e e \ rhs_eq_val v x c e = rhs e \ \v\ (x := c) \" begin text\\noindent Then, the next implementation of \update\ satisfies its specification:\ abbreviation update_eq where "update_eq v x c v' e \ upd (lhs e) (rhs_eq_val v x c e) v'" definition update :: "var \ 'a \ ('i,'a) state \ ('i,'a) state" where "update x c s \ \_update (upd x c (foldl (update_eq (\ s) x c) (\ s) (\ s))) s" lemma update_no_set_none: shows "look (\ s) y \ None \ look (foldl (update_eq (\ s) x v) (\ s) t) y \ None" by (induct t rule: rev_induct, auto simp: lookup_update') lemma update_no_left: assumes "y \ lvars t" shows "look (\ s) y = look (foldl (update_eq (\ s) x v) (\ s) t) y" using assms by (induct t rule: rev_induct) (auto simp add: lvars_def lookup_update') lemma update_left: assumes "y \ lvars t" shows "\ eq \ set t. lhs eq = y \ look (foldl (update_eq (\ s) x v) (\ s) t) y = Some (rhs_eq_val (\ s) x v eq)" using assms by (induct t rule: rev_induct) (auto simp add: lvars_def lookup_update') lemma update_valuate_rhs: assumes "e \ set (\ s)" "\ (\ s)" shows "rhs e \ \\ (update x c s)\ \ = rhs e \ \\ s\ (x := c) \" proof (rule valuate_depend, safe) fix y assume "y \ rvars_eq e" then have "y \ lvars (\ s)" using \\ (\ s)\ \e \ set (\ s)\ by (auto simp add: normalized_tableau_def rvars_def) then show "\\ (update x c s)\ y = (\\ s\(x := c)) y" using update_no_left[of y "\ s" s x c] by (auto simp add: update_def map2fun_def lookup_update') qed end sublocale RhsEqVal < Update update proof fix s::"('i,'a) state" and x c show "let s' = update x c s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>c s" by (simp add: Let_def update_def add: boundsl_def boundsu_def indexl_def indexu_def) next fix s::"('i,'a) state" and x c assume "\ (\ s)" "\ s" "x \ lvars (\ s)" then show "\ (update x c s)" using update_no_set_none[of s] by (simp add: Let_def update_def tableau_valuated_def lookup_update') next fix s::"('i,'a) state" and x x' c assume "\ (\ s)" "\ s" "x \ lvars (\ s)" show "x' \ lvars (\ s) \ look (\ (update x c s)) x' = (if x = x' then Some c else look (\ s) x')" using update_no_left[of x' "\ s" s x c] unfolding update_def lvars_def Let_def by (auto simp: lookup_update') next fix s::"('i,'a) state" and x c assume "\ (\ s)" "\ s" "x \ lvars (\ s)" have "\\ s\ \\<^sub>t \ s \ \e \ set (\ s). \\ (update x c s)\ \\<^sub>e e" proof fix e assume "e \ set (\ s)" "\\ s\ \\<^sub>t \ s" then have "\\ s\ \\<^sub>e e" by (simp add: satisfies_tableau_def) have "x \ lhs e" using \x \ lvars (\ s)\ \e \ set (\ s)\ by (auto simp add: lvars_def) then have "\\ (update x c s)\ (lhs e) = rhs_eq_val (\ s) x c e" using update_left[of "lhs e" "\ s" s x c] \e \ set (\ s)\ \\ (\ s)\ by (auto simp add: lvars_def lookup_update' update_def Let_def map2fun_def normalized_tableau_def distinct_map inj_on_def) then show "\\ (update x c s)\ \\<^sub>e e" using \\\ s\ \\<^sub>e e\ \e \ set (\ s)\ \x \ lvars (\ s)\ \\ (\ s)\ using rhs_eq_val by (simp add: satisfies_eq_def update_valuate_rhs) qed then show "\\ s\ \\<^sub>t \ s \ \\ (update x c s)\ \\<^sub>t \ s" by(simp add: satisfies_tableau_def update_def) qed text\To update the valuation for a variable that is on the lhs of the tableau it should first be swapped with some rhs variable of its equation, in an operation called \emph{pivoting}. Pivoting has the precondition that the tableau is normalized and that it is always called for a lhs variable of the tableau, and a rhs variable in the equation with that lhs variable. The set of rhs variables for the given lhs variable is found using the \rvars_of_lvar\ function (specified in a very simple locale \EqForLVar\, that we do not print).\ locale Pivot = EqForLVar + fixes pivot::"var \ var \ ('i,'a::lrv) state \ ('i,'a) state" assumes \ \Valuation, bounds, and the unsatisfiability flag are not changed.\ pivot_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ let s' = pivot x\<^sub>i x\<^sub>j s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>c s" and \ \The tableau remains equivalent to the previous one and normalized.\ pivot_tableau: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ let s' = pivot x\<^sub>i x\<^sub>j s in ((v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ s') \ \ (\ s') " and \ \@{text "x\<^sub>i"} and @{text "x\<^sub>j"} are swapped, while the other variables do not change sides.\ pivot_vars': "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ let s' = pivot x\<^sub>i x\<^sub>j s in rvars(\ s') = rvars(\ s)-{x\<^sub>j}\{x\<^sub>i} \ lvars(\ s') = lvars(\ s)-{x\<^sub>i}\{x\<^sub>j}" begin lemma pivot_bounds_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>i (pivot x\<^sub>i x\<^sub>j s) = \\<^sub>i s" using pivot_id by (simp add: Let_def) lemma pivot_bounds_id': assumes "\ (\ s)" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\\ (pivot x\<^sub>i x\<^sub>j s) = \\ s" "\ (pivot x\<^sub>i x\<^sub>j s) = \ s" "\ (pivot x\<^sub>i x\<^sub>j s) = \ s" using pivot_bounds_id[OF assms] by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) lemma pivot_valuation_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot x\<^sub>i x\<^sub>j s) = \ s" using pivot_id by (simp add: Let_def) lemma pivot_unsat_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot x\<^sub>i x\<^sub>j s) = \ s" using pivot_id by (simp add: Let_def) lemma pivot_unsat_core_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>c (pivot x\<^sub>i x\<^sub>j s) = \\<^sub>c s" using pivot_id by (simp add: Let_def) lemma pivot_tableau_equiv: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ (v::'a valuation) \\<^sub>t \ s = v \\<^sub>t \ (pivot x\<^sub>i x\<^sub>j s)" using pivot_tableau by (simp add: Let_def) lemma pivot_tableau_normalized: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (\ (pivot x\<^sub>i x\<^sub>j s))" using pivot_tableau by (simp add: Let_def) lemma pivot_rvars: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ rvars (\ (pivot x\<^sub>i x\<^sub>j s)) = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" using pivot_vars' by (simp add: Let_def) lemma pivot_lvars: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ lvars (\ (pivot x\<^sub>i x\<^sub>j s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" using pivot_vars' by (simp add: Let_def) lemma pivot_vars: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ tvars (\ (pivot x\<^sub>i x\<^sub>j s)) = tvars (\ s) " using pivot_lvars[of s x\<^sub>i x\<^sub>j] pivot_rvars[of s x\<^sub>i x\<^sub>j] using rvars_of_lvar_rvars[of x\<^sub>i "\ s"] by auto lemma pivot_tableau_valuated: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i; \ s\ \ \ (pivot x\<^sub>i x\<^sub>j s)" using pivot_valuation_id pivot_vars by (auto simp add: tableau_valuated_def) end text\Functions \pivot\ and \update\ can be used to implement the \check\ function. In its context, \pivot\ and \update\ functions are always called together, so the following definition can be used: @{prop "pivot_and_update x\<^sub>i x\<^sub>j c s = update x\<^sub>i c (pivot x\<^sub>i x\<^sub>j s)"}. It is possible to make a more efficient implementation of \pivot_and_update\ that does not use separate implementations of \pivot\ and \update\. To allow this, a separate specification for \pivot_and_update\ can be given. It can be easily shown that the \pivot_and_update\ definition above satisfies this specification.\ locale PivotAndUpdate = EqForLVar + fixes pivot_and_update :: "var \ var \ 'a::lrv \ ('i,'a) state \ ('i,'a) state" assumes pivotandupdate_unsat_id: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" assumes pivotandupdate_unsat_core_id: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>c (pivot_and_update x\<^sub>i x\<^sub>j c s) = \\<^sub>c s" assumes pivotandupdate_bounds_id: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>i (pivot_and_update x\<^sub>i x\<^sub>j c s) = \\<^sub>i s" assumes pivotandupdate_tableau_normalized: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (\ (pivot_and_update x\<^sub>i x\<^sub>j c s))" assumes pivotandupdate_tableau_equiv: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ (v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ (pivot_and_update x\<^sub>i x\<^sub>j c s)" assumes pivotandupdate_satisfies_tableau: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\ s\ \\<^sub>t \ s \ \\ (pivot_and_update x\<^sub>i x\<^sub>j c s)\ \\<^sub>t \ s" assumes pivotandupdate_rvars: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ rvars (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" assumes pivotandupdate_lvars: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ lvars (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" assumes pivotandupdate_valuation_nonlhs: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ x \ lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j} \ look (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) x = (if x = x\<^sub>i then Some c else look (\ s) x)" assumes pivotandupdate_tableau_valuated: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot_and_update x\<^sub>i x\<^sub>j c s)" begin lemma pivotandupdate_bounds_id': assumes "\ (\ s)" "\ s" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \\ s" "\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" "\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" using pivotandupdate_bounds_id[OF assms] by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) lemma pivotandupdate_valuation_xi: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ look (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) x\<^sub>i = Some c" using pivotandupdate_valuation_nonlhs[of s x\<^sub>i x\<^sub>j x\<^sub>i c] using rvars_of_lvar_rvars by (auto simp add: normalized_tableau_def) lemma pivotandupdate_valuation_other_nolhs: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i; x \ lvars (\ s); x \ x\<^sub>j\ \ look (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) x = look (\ s) x" using pivotandupdate_valuation_nonlhs[of s x\<^sub>i x\<^sub>j x c] by auto lemma pivotandupdate_nolhs: "\ \ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \\<^sub>l s x\<^sub>i = Some c \ \\<^sub>u s x\<^sub>i = Some c\ \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (pivot_and_update x\<^sub>i x\<^sub>j c s)" using pivotandupdate_satisfies_tableau[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_tableau_equiv[of s x\<^sub>i x\<^sub>j _ c] using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>j _ c] using pivotandupdate_lvars[of s x\<^sub>i x\<^sub>j c] by (auto simp add: curr_val_satisfies_no_lhs_def satisfies_bounds.simps satisfies_bounds_set.simps bounds_consistent_geq_lb bounds_consistent_leq_ub map2fun_def pivotandupdate_bounds_id') lemma pivotandupdate_bounds_consistent: assumes "\ (\ s)" "\ s" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" using assms pivotandupdate_bounds_id'[of s x\<^sub>i x\<^sub>j c] by (simp add: bounds_consistent_def) end locale PivotUpdate = Pivot eq_idx_for_lvar pivot + Update update for eq_idx_for_lvar :: "tableau \ var \ nat" and pivot :: "var \ var \ ('i,'a::lrv) state \ ('i,'a) state" and update :: "var \ 'a \ ('i,'a) state \ ('i,'a) state" begin definition pivot_and_update :: "var \ var \ 'a \ ('i,'a) state \ ('i,'a) state" where [simp]: "pivot_and_update x\<^sub>i x\<^sub>j c s \ update x\<^sub>i c (pivot x\<^sub>i x\<^sub>j s)" lemma pivot_update_precond: assumes "\ (\ s)" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\ (\ (pivot x\<^sub>i x\<^sub>j s))" "x\<^sub>i \ lvars (\ (pivot x\<^sub>i x\<^sub>j s))" proof- from assms have "x\<^sub>i \ x\<^sub>j" using rvars_of_lvar_rvars[of x\<^sub>i "\ s"] by (auto simp add: normalized_tableau_def) then show "\ (\ (pivot x\<^sub>i x\<^sub>j s))" "x\<^sub>i \ lvars (\ (pivot x\<^sub>i x\<^sub>j s))" using assms using pivot_tableau_normalized[of s x\<^sub>i x\<^sub>j] using pivot_lvars[of s x\<^sub>i x\<^sub>j] by auto qed end sublocale PivotUpdate < PivotAndUpdate eq_idx_for_lvar pivot_and_update using pivot_update_precond using update_unsat_id pivot_unsat_id pivot_unsat_core_id update_bounds_id pivot_bounds_id update_tableau_id pivot_tableau_normalized pivot_tableau_equiv update_satisfies_tableau pivot_valuation_id pivot_lvars pivot_rvars update_valuation_nonlhs update_valuation_nonlhs pivot_tableau_valuated update_tableau_valuated update_unsat_core_id by (unfold_locales, auto) text\Given the @{term update} function, \assert_bound\ can be implemented as follows. \vspace{-2mm} @{text[display] "assert_bound (Leq x c) s \ if c \\<^sub>u\<^sub>b \\<^sub>u s x then s else let s' = s \ \\<^sub>u := (\\<^sub>u s) (x := Some c) \ in if c <\<^sub>l\<^sub>b \\<^sub>l s x then s' \ \ := True \ else if x \ lvars (\ s') \ c < \\ s\ x then update x c s' else s'" } \vspace{-2mm} \noindent The case of \Geq x c\ atoms is analogous (a systematic way to avoid symmetries is discussed in Section \ref{sec:symmetries}). This implementation satisfies both its specifications. \ lemma indices_state_set_unsat: "indices_state (set_unsat I s) = indices_state s" by (cases s, auto simp: indices_state_def) lemma \\_set_unsat: "\\ (set_unsat I s) = \\ s" by (cases s, auto simp: boundsl_def boundsu_def indexl_def indexu_def) lemma satisfies_tableau_cong: assumes "\ x. x \ tvars t \ v x = w x" shows "(v \\<^sub>t t) = (w \\<^sub>t t)" unfolding satisfies_tableau_def satisfies_eq_def by (intro ball_cong[OF refl] arg_cong2[of _ _ _ _ "(=)"] valuate_depend, insert assms, auto simp: lvars_def rvars_def) lemma satisfying_state_valuation_to_atom_tabl: assumes J: "J \ indices_state s" and model: "(J, v) \\<^sub>i\<^sub>s\<^sub>e s" and ivalid: "index_valid as s" and dist: "distinct_indices_atoms as" shows "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as" "v \\<^sub>t \ s" unfolding i_satisfies_atom_set'.simps proof (intro ballI) from model[unfolded satisfies_state_index'.simps] have model: "v \\<^sub>t \ s" "(J, v) \\<^sub>i\<^sub>b\<^sub>e \\ s" by auto show "v \\<^sub>t \ s" by fact fix a assume "a \ restrict_to J as" then obtain i where iJ: "i \ J" and mem: "(i,a) \ as" by auto with J have "i \ indices_state s" by auto from this[unfolded indices_state_def] obtain x c where look: "look (\\<^sub>i\<^sub>l s) x = Some (i, c) \ look (\\<^sub>i\<^sub>u s) x = Some (i, c)" by auto with ivalid[unfolded index_valid_def] obtain b where "(i,b) \ as" "atom_var b = x" "atom_const b = c" by force with dist[unfolded distinct_indices_atoms_def, rule_format, OF this(1) mem] have a: "atom_var a = x" "atom_const a = c" by auto from model(2)[unfolded satisfies_bounds_index'.simps] look iJ have "v x = c" by (auto simp: boundsu_def boundsl_def indexu_def indexl_def) thus "v \\<^sub>a\<^sub>e a" unfolding satisfies_atom'_def a . qed text \Note that in order to ensure minimality of the unsat cores, pivoting is required.\ sublocale AssertAllState < AssertAll assert_all proof fix t as v I assume D: "\ t" from D show "assert_all t as = Sat v \ \v\ \\<^sub>t t \ \v\ \\<^sub>a\<^sub>s flat (set as)" unfolding Let_def assert_all_def using assert_all_state_tableau_equiv[OF D refl] using assert_all_state_sat[OF D refl] using assert_all_state_sat_atoms_equiv_bounds[OF D refl, of as] unfolding atoms_equiv_bounds.simps curr_val_satisfies_state_def satisfies_state_def satisfies_atom_set_def by (auto simp: Let_def split: if_splits) let ?s = "assert_all_state t as" assume "assert_all t as = Unsat I" then have i: "I = the (\\<^sub>c ?s)" and U: "\ ?s" unfolding assert_all_def Let_def by (auto split: if_splits) from assert_all_index_valid[OF D refl, of as] have ivalid: "index_valid (set as) ?s" . note unsat = assert_all_state_unsat[OF D refl U, unfolded minimal_unsat_state_core_def unsat_state_core_def i[symmetric]] from unsat have "set I \ indices_state ?s" by auto also have "\ \ fst ` set as" using assert_all_state_indices[OF D refl] . finally have indices: "set I \ fst ` set as" . show "minimal_unsat_core_tabl_atoms (set I) t (set as)" unfolding minimal_unsat_core_tabl_atoms_def proof (intro conjI impI allI indices, clarify) fix v assume model: "v \\<^sub>t t" "(set I, v) \\<^sub>i\<^sub>a\<^sub>s set as" from unsat have no_model: "\ ((set I, v) \\<^sub>i\<^sub>s ?s)" by auto from assert_all_state_unsat_atoms_equiv_bounds[OF D refl U] have equiv: "set as \\<^sub>i \\ ?s" by auto from assert_all_state_tableau_equiv[OF D refl, of v] model have model_t: "v \\<^sub>t \ ?s" by auto have model_as': "(set I, v) \\<^sub>i\<^sub>a\<^sub>s set as" using model(2) by (auto simp: satisfies_atom_set_def) with equiv model_t have "(set I, v) \\<^sub>i\<^sub>s ?s" unfolding satisfies_state_index.simps atoms_imply_bounds_index.simps by simp with no_model show False by simp next fix J assume dist: "distinct_indices_atoms (set as)" and J: "J \ set I" from J unsat[unfolded subsets_sat_core_def, folded i] have J': "J \ indices_state ?s" by auto from index_valid_distinct_indices[OF ivalid dist] J unsat[unfolded subsets_sat_core_def, folded i] obtain v where model: "(J, v) \\<^sub>i\<^sub>s\<^sub>e ?s" by blast have "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s set as" "v \\<^sub>t t" using satisfying_state_valuation_to_atom_tabl[OF J' model ivalid dist] assert_all_state_tableau_equiv[OF D refl] by auto then show "\ v. v \\<^sub>t t \ (J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s set as" by blast qed qed lemma (in Update) update_to_assert_bound_no_lhs: assumes pivot: "Pivot eqlvar (pivot :: var \ var \ ('i,'a) state \ ('i,'a) state)" shows "AssertBoundNoLhs assert_bound" proof fix s::"('i,'a) state" and a assume "\ \ s" "\ (\ s)" "\ s" then show "\ (assert_bound a s) = \ s" by (cases a, cases "snd a") (auto simp add: Let_def update_tableau_id tableau_valuated_def) next fix s::"('i,'a) state" and ia and as assume *: "\ \ s" "\ (\ s)" "\ s" and **: "\ (assert_bound ia s)" and index: "index_valid as s" and consistent: "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" obtain i a where ia: "ia = (i,a)" by force let ?modelU = "\ lt UB UI s v x c i. UB s x = Some c \ UI s x = i \ i \ set (the (\\<^sub>c s)) \ (lt (v x) c \ v x = c)" let ?modelL = "\ lt LB LI s v x c i. LB s x = Some c \ LI s x = i \ i \ set (the (\\<^sub>c s)) \ (lt c (v x) \ c = v x)" let ?modelIU = "\ I lt UB UI s v x c i. UB s x = Some c \ UI s x = i \ i \ I \ (v x = c)" let ?modelIL = "\ I lt LB LI s v x c i. LB s x = Some c \ LI s x = i \ i \ I \ (v x = c)" let ?P' = "\ lt UBI LBI UB LB UBI_upd UI LI LE GE s. \ s \ (set (the (\\<^sub>c s)) \ indices_state s \ \ (\v. (v \\<^sub>t \ s \ (\ x c i. ?modelU lt UB UI s v x c i) \ (\ x c i. ?modelL lt LB LI s v x c i)))) \ (distinct_indices_state s \ (\ I. I \ set (the (\\<^sub>c s)) \ (\ v. v \\<^sub>t \ s \ (\ x c i. ?modelIU I lt UB UI s v x c i) \ (\ x c i. ?modelIL I lt LB LI s v x c i))))" have "\ (assert_bound ia s) \ (unsat_state_core (assert_bound ia s) \ (distinct_indices_state (assert_bound ia s) \ subsets_sat_core (assert_bound ia s)))" (is "?P (assert_bound ia s)") unfolding ia proof (rule assert_bound_cases[of _ _ ?P']) fix s' :: "('i,'a) state" have id: "((x :: 'a) < y \ x = y) \ x \ y" "((x :: 'a) > y \ x = y) \ x \ y" for x y by auto have id': "?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u undefined \\<^sub>l \\<^sub>u Geq Leq s' = ?P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l undefined \\<^sub>u \\<^sub>l Leq Geq s'" by (intro arg_cong[of _ _ "\ y. _ \ y"] arg_cong[of _ _ "\ x. _ \ x"], intro arg_cong2[of _ _ _ _ "(\)"] arg_cong[of _ _ "\ y. _ \ y"] arg_cong[of _ _ "\ y. \ x \ set (the (\\<^sub>c s')). y x"] ext arg_cong[of _ _ Not], unfold id, auto) show "?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u undefined \\<^sub>l \\<^sub>u Geq Leq s'" unfolding satisfies_state_def satisfies_bounds_index.simps satisfies_bounds.simps in_bounds.simps unsat_state_core_def satisfies_state_index.simps subsets_sat_core_def satisfies_state_index'.simps satisfies_bounds_index'.simps unfolding bound_compare''_defs id by ((intro arg_cong[of _ _ "\ x. _ \ x"] arg_cong[of _ _ "\ x. _ \ x"], intro arg_cong2[of _ _ _ _ "(\)"] refl arg_cong[of _ _ "\ x. _ \ x"] arg_cong[of _ _ Not] arg_cong[of _ _ "\ y. \ x \ set (the (\\<^sub>c s')). y x"] ext; intro arg_cong[of _ _ Ex] ext), auto) then show "?P s' = ?P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l undefined \\<^sub>u \\<^sub>l Leq Geq s'" unfolding id' . next fix c::'a and x::nat and dir assume "\\<^sub>l\<^sub>b (lt dir) c (LB dir s x)" and dir: "dir = Positive \ dir = Negative" then obtain d where some: "LB dir s x = Some d" and lt: "lt dir c d" by (auto simp: bound_compare'_defs split: option.splits) from index[unfolded index_valid_def, rule_format, of x _ d] some dir obtain j where ind: "LI dir s x = j" "look (LBI dir s) x = Some (j,d)" and ge: "(j, GE dir x d) \ as" by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) let ?s = "set_unsat [i, ((LI dir) s x)] (update\\ (UBI_upd dir) i x c s)" let ?ss = "update\\ (UBI_upd dir) i x c s" show "?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) ?s" proof (intro conjI impI allI, goal_cases) case 1 thus ?case using dir ind ge lt some by (force simp: indices_state_def split: if_splits) next case 2 { fix v assume vU: "\ x c i. ?modelU (lt dir) (UB dir) (UI dir) ?s v x c i" assume vL: "\ x c i. ?modelL (lt dir) (LB dir) (LI dir) ?s v x c i" from dir have "UB dir ?s x = Some c" "UI dir ?s x = i" by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) from vU[rule_format, OF this] have vx_le_c: "lt dir (v x) c \ v x = c" by auto from dir ind some have *: "LB dir ?s x = Some d" "LI dir ?s x = j" by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) have d_le_vx: "lt dir d (v x) \ d = v x" by (intro vL[rule_format, OF *], insert some ind, auto) from dir d_le_vx vx_le_c lt - have False by auto + have False by (auto simp del: Simplex.bounds_lg) } thus ?case by blast next case (3 I) then obtain j where I: "I \ {j}" by (auto split: if_splits) from 3 have dist: "distinct_indices_state ?ss" unfolding distinct_indices_state_def by auto have id1: "UB dir ?s y = UB dir ?ss y" "LB dir ?s y = LB dir ?ss y" "UI dir ?s y = UI dir ?ss y" "LI dir ?s y = LI dir ?ss y" "\ ?s = \ s" "set (the (\\<^sub>c ?s)) = {i,LI dir s x}" for y using dir by (auto simp: boundsu_def boundsl_def indexu_def indexl_def) from I have id: "(\ k. P1 k \ P2 k \ k \ I \ Q k) \ (I = {} \ (P1 j \ P2 j \ Q j))" for P1 P2 Q by auto have id2: "(UB dir s xa = Some ca \ UI dir s xa = j \ P) = (look (UBI dir s) xa = Some (j,ca) \ P)" "(LB dir s xa = Some ca \ LI dir s xa = j \ P) = (look (LBI dir s) xa = Some (j,ca) \ P)" for xa ca P s using dir by (auto simp: boundsu_def indexu_def boundsl_def indexl_def) have "\v. v \\<^sub>t \ s \ (\xa ca ia. UB dir ?ss xa = Some ca \ UI dir ?ss xa = ia \ ia \ I \ v xa = ca) \ (\xa ca ia. LB dir ?ss xa = Some ca \ LI dir ?ss xa = ia \ ia \ I \ v xa = ca)" proof (cases "\ xa ca. look (UBI dir ?ss) xa = Some (j,ca) \ look (LBI dir ?ss) xa = Some (j,ca)") case False thus ?thesis unfolding id id2 using consistent unfolding curr_val_satisfies_no_lhs_def by (intro exI[of _ "\\ s\"], auto) next case True from consistent have val: " \\ s\ \\<^sub>t \ s" unfolding curr_val_satisfies_no_lhs_def by auto define ss where ss: "ss = ?ss" from True obtain y b where "look (UBI dir ?ss) y = Some (j,b) \ look (LBI dir ?ss) y = Some (j,b)" by force then have id3: "(look (LBI dir ss) yy = Some (j,bb) \ look (UBI dir ss) yy = Some (j,bb)) \ (yy = y \ bb = b)" for yy bb using distinct_indices_stateD(1)[OF dist, of y j b yy bb] using dir unfolding ss[symmetric] by (auto simp: boundsu_def boundsl_def indexu_def indexl_def) have "\v. v \\<^sub>t \ s \ v y = b" proof (cases "y \ lvars (\ s)") case False let ?v = "\\ (update y b s)\" show ?thesis proof (intro exI[of _ ?v] conjI) from update_satisfies_tableau[OF *(2,3) False] val show "?v \\<^sub>t \ s" by simp from update_valuation_nonlhs[OF *(2,3) False, of y b] False show "?v y = b" by (simp add: map2fun_def') qed next case True from *(2)[unfolded normalized_tableau_def] have zero: "0 \ rhs ` set (\ s)" by auto interpret Pivot eqlvar pivot by fact interpret PivotUpdate eqlvar pivot update .. let ?eq = "eq_for_lvar (\ s) y" from eq_for_lvar[OF True] have "?eq \ set (\ s)" "lhs ?eq = y" by auto with zero have rhs: "rhs ?eq \ 0" by force hence "rvars_eq ?eq \ {}" by (simp add: vars_empty_zero) then obtain z where z: "z \ rvars_eq ?eq" by auto let ?v = "\ (pivot_and_update y z b s)" let ?vv = "\?v\" from pivotandupdate_valuation_xi[OF *(2,3) True z] have "look ?v y = Some b" . hence vv: "?vv y = b" unfolding map2fun_def' by auto show ?thesis proof (intro exI[of _ ?vv] conjI vv) show "?vv \\<^sub>t \ s" using pivotandupdate_satisfies_tableau[OF *(2,3) True z] val by auto qed qed thus ?thesis unfolding id id2 ss[symmetric] using id3 by metis qed thus ?case unfolding id1 . qed next fix c::'a and x::nat and dir assume **: "dir = Positive \ dir = Negative" "a = LE dir x c" "x \ lvars (\ s)" "lt dir c (\\ s\ x)" "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "\ \\<^sub>l\<^sub>b (lt dir) c (LB dir s x)" let ?s = "update\\ (UBI_upd dir) i x c s" show "?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) (update x c ?s)" using * ** by (auto simp add: update_unsat_id tableau_valuated_def) qed (auto simp add: * update_unsat_id tableau_valuated_def) with ** show "minimal_unsat_state_core (assert_bound ia s)" by (auto simp: minimal_unsat_state_core_def) next fix s::"('i,'a) state" and ia assume *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" "\ (\ s)" "\ s" and **: "\ \ (assert_bound ia s)" (is ?lhs) obtain i a where ia: "ia = (i,a)" by force have "\\ (assert_bound ia s)\ \\<^sub>t \ (assert_bound ia s)" proof- let ?P = "\ lt UBI LBI UB LB UBI_upd UI LI LE GE s. \\ s\ \\<^sub>t \ s" show ?thesis unfolding ia proof (rule assert_bound_cases[of _ _ ?P]) fix c x and dir :: "('i,'a) Direction" let ?s' = "update\\ (UBI_upd dir) i x c s" assume "x \ lvars (\ s)" "(lt dir) c (\\ s\ x)" "dir = Positive \ dir = Negative" then show "\\ (update x c ?s')\ \\<^sub>t \ (update x c ?s')" using * using update_satisfies_tableau[of ?s' x c] update_tableau_id by (auto simp add: curr_val_satisfies_no_lhs_def tableau_valuated_def) qed (insert *, auto simp add: curr_val_satisfies_no_lhs_def) qed moreover have "\ \ (assert_bound ia s) \ \\ (assert_bound ia s)\ \\<^sub>b \ (assert_bound ia s) \ - lvars (\ (assert_bound ia s))" (is "?P (assert_bound ia s)") proof- let ?P' = "\ lt UBI LBI UB LB UB_upd UI LI LE GE s. \ \ s \ (\x\- lvars (\ s). \\<^sub>l\<^sub>b lt (\\ s\ x) (LB s x) \ \\<^sub>u\<^sub>b lt (\\ s\ x) (UB s x))" let ?P'' = "\ dir. ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)" have x: "\ s'. ?P s' = ?P' (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s'" unfolding satisfies_bounds_set.simps in_bounds.simps bound_compare_defs by (auto split: option.split) show ?thesis unfolding ia proof (rule assert_bound_cases[of _ _ ?P']) fix dir :: "('i,'a) Direction" assume "dir = Positive \ dir = Negative" then show "?P'' dir s" using x[of s] xx[of s] \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ by (auto simp add: curr_val_satisfies_no_lhs_def) next fix x c and dir :: "('i,'a) Direction" let ?s' = "update\\ (UBI_upd dir) i x c s" assume "x \ lvars (\ s)" "dir = Positive \ dir = Negative" then have "?P ?s'" using \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ by (auto simp add: satisfies_bounds_set.simps curr_val_satisfies_no_lhs_def boundsl_def boundsu_def indexl_def indexu_def) then show "?P'' dir ?s'" using x[of ?s'] xx[of ?s'] \dir = Positive \ dir = Negative\ by auto next fix c x and dir :: "('i,'a) Direction" let ?s' = "update\\ (UBI_upd dir) i x c s" assume "\ lt dir c (\\ s\ x)" "dir = Positive \ dir = Negative" then show "?P'' dir ?s'" using \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ by (auto simp add: satisfies_bounds_set.simps curr_val_satisfies_no_lhs_def simp: boundsl_def boundsu_def indexl_def indexu_def) (auto simp add: bound_compare_defs) next fix c x and dir :: "('i,'a) Direction" let ?s' = "update x c (update\\ (UBI_upd dir) i x c s)" assume "x \ lvars (\ s)" "\ \\<^sub>l\<^sub>b (lt dir) c (LB dir s x)" "dir = Positive \ dir = Negative" show "?P'' dir ?s'" proof (rule impI, rule ballI) fix y assume "\ \ ?s'" "y \ - lvars (\ ?s')" show "\\<^sub>l\<^sub>b (lt dir) (\\ ?s'\ y) (LB dir ?s' y) \ \\<^sub>u\<^sub>b (lt dir) (\\ ?s'\ y) (UB dir ?s' y)" proof (cases "x = y") case True then show ?thesis using \x \ lvars (\ s)\ using \y \ - lvars (\ ?s')\ using \\ \\<^sub>l\<^sub>b (lt dir) c (LB dir s x)\ using \dir = Positive \ dir = Negative\ using neg_bounds_compare(7) neg_bounds_compare(3) using * by (auto simp add: update_valuation_nonlhs update_tableau_id update_bounds_id bound_compare''_defs map2fun_def tableau_valuated_def bounds_updates) (force simp add: bound_compare'_defs)+ next case False then show ?thesis using \x \ lvars (\ s)\ \y \ - lvars (\ ?s')\ using \dir = Positive \ dir = Negative\ * by (auto simp add: update_valuation_nonlhs update_tableau_id update_bounds_id bound_compare''_defs satisfies_bounds_set.simps curr_val_satisfies_no_lhs_def map2fun_def tableau_valuated_def bounds_updates) qed qed qed (auto simp add: x xx) qed moreover have "\ \ (assert_bound ia s) \ \ (assert_bound ia s)" (is "?P (assert_bound ia s)") proof- let ?P' = "\ lt UBI LBI UB LB UBI_upd UI LI LE GE s. \ \ s \ (\x. if LB s x = None \ UB s x = None then True else lt (the (LB s x)) (the (UB s x)) \ (the (LB s x) = the (UB s x)))" let ?P'' = "\ dir. ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)" have x: "\ s'. ?P s' = ?P' (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s'" unfolding bounds_consistent_def by auto show ?thesis unfolding ia proof (rule assert_bound_cases[of _ _ ?P']) fix dir :: "('i,'a) Direction" assume "dir = Positive \ dir = Negative" then show "?P'' dir s" using \\ s\ by (auto simp add: bounds_consistent_def) (erule_tac x=x in allE, auto)+ next fix x c and dir :: "('i,'a) Direction" let ?s' = "update x c (update\\ (UBI_upd dir) i x c s)" assume "dir = Positive \ dir = Negative" "x \ lvars (\ s)" "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "\ \\<^sub>l\<^sub>b (lt dir) c (LB dir s x)" then show "?P'' dir ?s'" using \\ s\ * unfolding bounds_consistent_def by (auto simp add: update_bounds_id tableau_valuated_def bounds_updates split: if_splits) (force simp add: bound_compare'_defs, erule_tac x=xa in allE, simp)+ next fix x c and dir :: "('i,'a) Direction" let ?s' = "update\\ (UBI_upd dir) i x c s" assume "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "\ \\<^sub>l\<^sub>b (lt dir) c (LB dir s x)" "dir = Positive \ dir = Negative" then have "?P'' dir ?s'" using \\ s\ unfolding bounds_consistent_def by (auto split: if_splits simp: bounds_updates) (force simp add: bound_compare'_defs, erule_tac x=xa in allE, simp)+ then show "?P'' dir ?s'" "?P'' dir ?s'" by simp_all qed (auto simp add: x xx) qed ultimately show "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (assert_bound ia s) \ \ (assert_bound ia s)" using \?lhs\ unfolding curr_val_satisfies_no_lhs_def by simp next fix s :: "('i,'a) state" and ats and ia :: "('i,'a) i_atom" assume "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" obtain i a where ia: "ia = (i,a)" by force { fix ats let ?P' = "\ lt UBI LBI UB LB UB_upd UI LI LE GE s'. ats \ \ s \ (ats \ {a}) \ \ s'" let ?P'' = "\ dir. ?P' (lt dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)" have "ats \ \ s \ (ats \ {a}) \ \ (assert_bound ia s)" (is "?P (assert_bound ia s)") unfolding ia proof (rule assert_bound_cases[of _ _ ?P']) fix x c and dir :: "('i,'a) Direction" assume "dir = Positive \ dir = Negative" "a = LE dir x c" "\\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" then show "?P s" unfolding atoms_equiv_bounds.simps satisfies_atom_set_def satisfies_bounds.simps by auto (erule_tac x=x in allE, force simp add: bound_compare_defs)+ next fix x c and dir :: "('i,'a) Direction" let ?s' = "set_unsat [i, ((LI dir) s x)] (update\\ (UBI_upd dir) i x c s)" assume "dir = Positive \ dir = Negative" "a = LE dir x c" "\ (\\<^sub>u\<^sub>b (lt dir) c (UB dir s x))" then show "?P ?s'" unfolding set_unsat_bounds_id using atoms_equiv_bounds_extend[of dir c s x ats i] by auto next fix x c and dir :: "('i,'a) Direction" let ?s' = "update\\ (UBI_upd dir) i x c s" assume "dir = Positive \ dir = Negative" "a = LE dir x c" "\ (\\<^sub>u\<^sub>b (lt dir) c (UB dir s x))" then have "?P ?s'" using atoms_equiv_bounds_extend[of dir c s x ats i] by auto then show "?P ?s'" "?P ?s'" by simp_all next fix x c and dir :: "('i,'a) Direction" let ?s = "update\\ (UBI_upd dir) i x c s" let ?s' = "update x c ?s" assume *: "dir = Positive \ dir = Negative" "a = LE dir x c" "\ (\\<^sub>u\<^sub>b (lt dir) c (UB dir s x))" "x \ lvars (\ s)" then have "\ (\ ?s)" "\ ?s" "x \ lvars (\ ?s)" using \\ (\ s)\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ \\ s\ by (auto simp: tableau_valuated_def) from update_bounds_id[OF this, of c] have "\\<^sub>i ?s' = \\<^sub>i ?s" by blast then have id: "\ ?s' = \ ?s" unfolding boundsl_def boundsu_def by auto show "?P ?s'" unfolding id \a = LE dir x c\ by (intro impI atoms_equiv_bounds_extend[rule_format] *(1,3)) qed simp_all } then show "flat ats \ \ s \ flat (ats \ {ia}) \ \ (assert_bound ia s)" unfolding ia by auto next fix s :: "('i,'a) state" and ats and ia :: "('i,'a) i_atom" obtain i a where ia: "ia = (i,a)" by force assume "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" have *: "\ dir x c s. dir = Positive \ dir = Negative \ \ (update\\ (UBI_upd dir) i x c s) = \ s" "\ s y I . \ (set_unsat I s) = \ s" by (auto simp add: tableau_valuated_def) show "\ (assert_bound ia s)" (is "?P (assert_bound ia s)") proof- let ?P' = "\ lt UBI LBI UB LB UB_upd UI LI LE GE s'. \ s'" let ?P'' = "\ dir. ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)" show ?thesis unfolding ia proof (rule assert_bound_cases[of _ _ ?P']) fix x c and dir :: "('i,'a) Direction" let ?s' = "update\\ (UBI_upd dir) i x c s" assume "dir = Positive \ dir = Negative" then have "\ ?s'" using *(1)[of dir x c s] \\ s\ by simp then show "\ (set_unsat [i, ((LI dir) s x)] ?s')" using *(2) by auto next fix x c and dir :: "('i,'a) Direction" assume *: "x \ lvars (\ s)" "dir = Positive \ dir = Negative" let ?s = "update\\ (UBI_upd dir) i x c s" let ?s' = "update x c ?s" from * show "\ ?s'" using \\ (\ s)\ \\ s\ using update_tableau_valuated[of ?s x c] by (auto simp add: tableau_valuated_def) qed (insert \\ s\ *(1), auto) qed next fix s :: "('i,'a) state" and as and ia :: "('i,'a) i_atom" obtain i a where ia: "ia = (i,a)" by force assume *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" and valid: "index_valid as s" have id: "\ dir x c s. dir = Positive \ dir = Negative \ \ (update\\ (UBI_upd dir) i x c s) = \ s" "\ s y I. \ (set_unsat I s) = \ s" by (auto simp add: tableau_valuated_def) let ?I = "insert (i,a) as" define I where "I = ?I" from index_valid_mono[OF _ valid] have valid: "index_valid I s" unfolding I_def by auto have I: "(i,a) \ I" unfolding I_def by auto let ?P = "\ s. index_valid I s" let ?P' = "\ (lt :: 'a \ 'a \ bool) (UBI :: ('i,'a) state \ ('i,'a) bounds_index) (LBI :: ('i,'a) state \ ('i,'a) bounds_index) (UB :: ('i,'a) state \ 'a bounds) (LB :: ('i,'a) state \ 'a bounds) (UBI_upd :: (('i,'a) bounds_index \ ('i,'a) bounds_index) \ ('i,'a) state \ ('i,'a) state) (UI :: ('i,'a) state \ 'i bound_index) (LI :: ('i,'a) state \ 'i bound_index) LE GE s'. (\ x c i. look (UBI s') x = Some (i,c) \ (i,LE (x :: var) c) \ I) \ (\ x c i. look (LBI s') x = Some (i,c) \ (i,GE (x :: nat) c) \ I)" define P where "P = ?P'" let ?P'' = "\ (dir :: ('i,'a) Direction). P (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)" have x: "\ s'. ?P s' = P (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = P (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s'" unfolding satisfies_bounds_set.simps in_bounds.simps bound_compare_defs index_valid_def P_def by (auto split: option.split simp: indexl_def indexu_def boundsl_def boundsu_def) from valid have P'': "dir = Positive \ dir = Negative \ ?P'' dir s" for dir using x[of s] xx[of s] by auto have UTrue: "dir = Positive \ dir = Negative \ ?P'' dir s \ ?P'' dir (set_unsat I s)" for dir s I unfolding P_def by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) have updateIB: "a = LE dir x c \ dir = Positive \ dir = Negative \ ?P'' dir s \ ?P'' dir (update\\ (UBI_upd dir) i x c s)" for dir x c s unfolding P_def using I by (auto split: if_splits simp: simp: boundsl_def boundsu_def indexl_def indexu_def) show "index_valid (insert ia as) (assert_bound ia s)" unfolding ia I_def[symmetric] proof ((rule assert_bound_cases[of _ _ P]; (intro UTrue x xx updateIB P'')?)) fix x c and dir :: "('i,'a) Direction" assume **: "dir = Positive \ dir = Negative" "a = LE dir x c" "x \ lvars (\ s)" let ?s = "(update\\ (UBI_upd dir) i x c s)" define s' where "s' = ?s" have 1: "\ (\ ?s)" using * ** by auto have 2: "\ ?s" using id(1) ** * \\ s\ by auto have 3: "x \ lvars (\ ?s)" using id(1) ** * \\ s\ by auto have "?P'' dir ?s" using ** by (intro updateIB P'') auto with update_id[of ?s x c, OF 1 2 3, unfolded Let_def] **(1) show "P (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) (update x c (update\\ (UBI_upd dir) i x c s))" unfolding P_def s'_def[symmetric] by auto qed auto next fix s and ia :: "('i,'a) i_atom" and ats :: "('i,'a) i_atom set" assume *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" "\ s" and ats: "ats \\<^sub>i \\ s" obtain i a where ia: "ia = (i,a)" by force have id: "\ dir x c s. dir = Positive \ dir = Negative \ \ (update\\ (UBI_upd dir) i x c s) = \ s" "\ s I. \ (set_unsat I s) = \ s" by (auto simp add: tableau_valuated_def) have idlt: "(c < (a :: 'a) \ c = a) = (c \ a)" "(a < c \ c = a) = (c \ a)" for a c by auto define A where "A = insert (i,a) ats" let ?P = "\ (s :: ('i,'a) state). A \\<^sub>i \\ s" let ?Q = "\ bs (lt :: 'a \ 'a \ bool) (UBI :: ('i,'a) state \ ('i,'a) bounds_index) (LBI :: ('i,'a) state \ ('i,'a) bounds_index) (UB :: ('i,'a) state \ 'a bounds) (LB :: ('i,'a) state \ 'a bounds) (UBI_upd :: (('i,'a) bounds_index \ ('i,'a) bounds_index) \ ('i,'a) state \ ('i,'a) state) UI LI (LE :: nat \ 'a \ 'a atom) (GE :: nat \ 'a \ 'a atom) s'. (\ I v. (I :: 'i set,v) \\<^sub>i\<^sub>a\<^sub>s bs \ ((\ x c. LB s' x = Some c \ LI s' x \ I \ lt c (v x) \ c = v x) \ (\ x c. UB s' x = Some c \ UI s' x \ I \ lt (v x) c \ v x = c)))" define Q where "Q = ?Q" let ?P' = "Q A" have equiv: "bs \\<^sub>i \\ s' \ Q bs (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" "bs \\<^sub>i \\ s' \ Q bs (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s'" for bs s' unfolding satisfies_bounds_set.simps in_bounds.simps bound_compare_defs index_valid_def Q_def atoms_imply_bounds_index.simps by (atomize(full), (intro conjI iff_exI allI arg_cong2[of _ _ _ _ "(\)"] refl iff_allI arg_cong2[of _ _ _ _ "(=)"]; unfold satisfies_bounds_index.simps idlt), auto) have x: "\ s'. ?P s' = ?P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s'" using equiv by blast+ from ats equiv[of ats s] have Q_ats: "Q ats (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s" "Q ats (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s" by auto let ?P'' = "\ (dir :: ('i,'a) Direction). ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)" have P_upd: "dir = Positive \ dir = Negative \ ?P'' dir (set_unsat I s) = ?P'' dir s" for s I dir unfolding Q_def by (intro iff_exI arg_cong2[of _ _ _ _ "(\)"] refl iff_allI arg_cong2[of _ _ _ _ "(=)"] arg_cong2[of _ _ _ _ "(\)"], auto simp: boundsl_def boundsu_def indexl_def indexu_def) have P_upd: "dir = Positive \ dir = Negative \ ?P'' dir s \ ?P'' dir (set_unsat I s)" for s I dir using P_upd[of dir] by blast have ats_sub: "ats \ A" unfolding A_def by auto { fix x c and dir :: "('i,'a) Direction" assume dir: "dir = Positive \ dir = Negative" and a: "a = LE dir x c" from Q_ats dir have Q_ats: "Q ats (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) s" by auto have "?P'' dir (update\\ (UBI_upd dir) i x c s)" unfolding Q_def proof (intro allI impI conjI) fix I v y d assume IvA: "(I, v) \\<^sub>i\<^sub>a\<^sub>s A" from i_satisfies_atom_set_mono[OF ats_sub this] have "(I, v) \\<^sub>i\<^sub>a\<^sub>s ats" by auto from Q_ats[unfolded Q_def, rule_format, OF this] have s_bnds: "LB dir s x = Some c \ LI dir s x \ I \ lt dir c (v x) \ c = v x" "UB dir s x = Some c \ UI dir s x \ I \ lt dir (v x) c \ v x = c" for x c by auto from IvA[unfolded A_def, unfolded i_satisfies_atom_set.simps satisfies_atom_set_def, simplified] have va: "i \ I \ v \\<^sub>a a" by auto with a dir have vc: "i \ I \ lt dir (v x) c \ v x = c" by auto let ?s = "(update\\ (UBI_upd dir) i x c s)" show "LB dir ?s y = Some d \ LI dir ?s y \ I \ lt dir d (v y) \ d = v y" "UB dir ?s y = Some d \ UI dir ?s y \ I \ lt dir (v y) d \ v y = d" proof (atomize(full), goal_cases) case 1 consider (main) "y = x" "UI dir ?s x = i" | (easy1) "x \ y" | (easy2) "x = y" "UI dir ?s y \ i" by blast then show ?case proof cases case easy1 then show ?thesis using s_bnds[of y d] dir by (fastforce simp: boundsl_def boundsu_def indexl_def indexu_def) next case easy2 then show ?thesis using s_bnds[of y d] dir by (fastforce simp: boundsl_def boundsu_def indexl_def indexu_def) next case main note s_bnds = s_bnds[of x] show ?thesis unfolding main using s_bnds dir vc by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) qed qed qed } note main = this have Ps: "dir = Positive \ dir = Negative \ ?P'' dir s" for dir using Q_ats unfolding Q_def using i_satisfies_atom_set_mono[OF ats_sub] by fastforce have "?P (assert_bound (i,a) s)" proof ((rule assert_bound_cases[of _ _ ?P']; (intro x xx P_upd main Ps)?)) fix c x and dir :: "('i,'a) Direction" assume **: "dir = Positive \ dir = Negative" "a = LE dir x c" "x \ lvars (\ s)" let ?s = "update\\ (UBI_upd dir) i x c s" define s' where "s' = ?s" from main[OF **(1-2)] have P: "?P'' dir s'" unfolding s'_def . have 1: "\ (\ ?s)" using * ** by auto have 2: "\ ?s" using id(1) ** * \\ s\ by auto have 3: "x \ lvars (\ ?s)" using id(1) ** * \\ s\ by auto have "\ (\ s')" "\ s'" "x \ lvars (\ s')" using 1 2 3 unfolding s'_def by auto from update_bounds_id[OF this, of c] **(1) have "?P'' dir (update x c s') = ?P'' dir s'" unfolding Q_def by (intro iff_allI arg_cong2[of _ _ _ _ "(\)"] arg_cong2[of _ _ _ _ "(\)"] refl, auto) with P show "?P'' dir (update x c ?s)" unfolding s'_def by blast qed auto then show "insert ia ats \\<^sub>i \\ (assert_bound ia s)" unfolding ia A_def by blast qed text \Pivoting the tableau can be reduced to pivoting single equations, and substituting variable by polynomials. These operations are specified by:\ locale PivotEq = fixes pivot_eq::"eq \ var \ eq" assumes \ \Lhs var of @{text eq} and @{text x\<^sub>j} are swapped, while the other variables do not change sides.\ vars_pivot_eq:" \x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ let eq' = pivot_eq eq x\<^sub>j in lhs eq' = x\<^sub>j \ rvars_eq eq' = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" and \ \Pivoting keeps the equation equisatisfiable.\ equiv_pivot_eq: "\x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ (v::'a::lrv valuation) \\<^sub>e pivot_eq eq x\<^sub>j \ v \\<^sub>e eq" begin lemma lhs_pivot_eq: "\x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ lhs (pivot_eq eq x\<^sub>j) = x\<^sub>j" using vars_pivot_eq by (simp add: Let_def) lemma rvars_pivot_eq: "\x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ rvars_eq (pivot_eq eq x\<^sub>j) = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" using vars_pivot_eq by (simp add: Let_def) end abbreviation doublesub (" _ \s _ \s _" [50,51,51] 50) where "doublesub a b c \ a \ b \ b \ c" locale SubstVar = fixes subst_var::"var \ linear_poly \ linear_poly \ linear_poly" assumes \ \Effect of @{text "subst_var x\<^sub>j lp' lp"} on @{text lp} variables.\ vars_subst_var': "(vars lp - {x\<^sub>j}) - vars lp' \s vars (subst_var x\<^sub>j lp' lp) \s (vars lp - {x\<^sub>j}) \ vars lp'"and subst_no_effect: "x\<^sub>j \ vars lp \ subst_var x\<^sub>j lp' lp = lp" and subst_with_effect: "x\<^sub>j \ vars lp \ x \ vars lp' - vars lp \ x \ vars (subst_var x\<^sub>j lp' lp)" and \ \Effect of @{text "subst_var x\<^sub>j lp' lp"} on @{text lp} value.\ equiv_subst_var: "(v::'a :: lrv valuation) x\<^sub>j = lp' \v\ \ lp \v\ = (subst_var x\<^sub>j lp' lp) \v\" begin lemma vars_subst_var: "vars (subst_var x\<^sub>j lp' lp) \ (vars lp - {x\<^sub>j}) \ vars lp'" using vars_subst_var' by simp lemma vars_subst_var_supset: "vars (subst_var x\<^sub>j lp' lp) \ (vars lp - {x\<^sub>j}) - vars lp'" using vars_subst_var' by simp definition subst_var_eq :: "var \ linear_poly \ eq \ eq" where "subst_var_eq v lp' eq \ (lhs eq, subst_var v lp' (rhs eq))" lemma rvars_eq_subst_var_eq: shows "rvars_eq (subst_var_eq x\<^sub>j lp eq) \ (rvars_eq eq - {x\<^sub>j}) \ vars lp" unfolding subst_var_eq_def by (auto simp add: vars_subst_var) lemma rvars_eq_subst_var_eq_supset: "rvars_eq (subst_var_eq x\<^sub>j lp eq) \ (rvars_eq eq) - {x\<^sub>j} - (vars lp)" unfolding subst_var_eq_def by (simp add: vars_subst_var_supset) lemma equiv_subst_var_eq: assumes "(v::'a valuation) \\<^sub>e (x\<^sub>j, lp')" shows "v \\<^sub>e eq \ v \\<^sub>e subst_var_eq x\<^sub>j lp' eq" using assms unfolding subst_var_eq_def unfolding satisfies_eq_def using equiv_subst_var[of v x\<^sub>j lp' "rhs eq"] by auto end locale Pivot' = EqForLVar + PivotEq + SubstVar begin definition pivot_tableau' :: "var \ var \ tableau \ tableau" where "pivot_tableau' x\<^sub>i x\<^sub>j t \ let x\<^sub>i_idx = eq_idx_for_lvar t x\<^sub>i; eq = t ! x\<^sub>i_idx; eq' = pivot_eq eq x\<^sub>j in map (\ idx. if idx = x\<^sub>i_idx then eq' else subst_var_eq x\<^sub>j (rhs eq') (t ! idx) ) [0.. var \ ('i,'a::lrv) state \ ('i,'a) state" where "pivot' x\<^sub>i x\<^sub>j s \ \_update (pivot_tableau' x\<^sub>i x\<^sub>j (\ s)) s" text\\noindent Then, the next implementation of \pivot\ satisfies its specification:\ definition pivot_tableau :: "var \ var \ tableau \ tableau" where "pivot_tableau x\<^sub>i x\<^sub>j t \ let eq = eq_for_lvar t x\<^sub>i; eq' = pivot_eq eq x\<^sub>j in map (\ e. if lhs e = lhs eq then eq' else subst_var_eq x\<^sub>j (rhs eq') e) t" definition pivot :: "var \ var \ ('i,'a::lrv) state \ ('i,'a) state" where "pivot x\<^sub>i x\<^sub>j s \ \_update (pivot_tableau x\<^sub>i x\<^sub>j (\ s)) s" lemma pivot_tableau'pivot_tableau: assumes "\ t" "x\<^sub>i \ lvars t" shows "pivot_tableau' x\<^sub>i x\<^sub>j t = pivot_tableau x\<^sub>i x\<^sub>j t" proof- let ?f = "\idx. if idx = eq_idx_for_lvar t x\<^sub>i then pivot_eq (t ! eq_idx_for_lvar t x\<^sub>i) x\<^sub>j else subst_var_eq x\<^sub>j (rhs (pivot_eq (t ! eq_idx_for_lvar t x\<^sub>i) x\<^sub>j)) (t ! idx)" let ?f' = "\e. if lhs e = lhs (eq_for_lvar t x\<^sub>i) then pivot_eq (eq_for_lvar t x\<^sub>i) x\<^sub>j else subst_var_eq x\<^sub>j (rhs (pivot_eq (eq_for_lvar t x\<^sub>i) x\<^sub>j)) e" have "\ i < length t. ?f' (t ! i) = ?f i" proof(safe) fix i assume "i < length t" then have "t ! i \ set t" "i < length t" by auto moreover have "t ! eq_idx_for_lvar t x\<^sub>i \ set t" "eq_idx_for_lvar t x\<^sub>i < length t" using eq_for_lvar[of x\<^sub>i t] \x\<^sub>i \ lvars t\ eq_idx_for_lvar[of x\<^sub>i t] by (auto simp add: eq_for_lvar_def) ultimately have "lhs (t ! i) = lhs (t ! eq_idx_for_lvar t x\<^sub>i) \ t ! i = t ! (eq_idx_for_lvar t x\<^sub>i)" "distinct t" using \\ t\ unfolding normalized_tableau_def by (auto simp add: distinct_map inj_on_def) then have "lhs (t ! i) = lhs (t ! eq_idx_for_lvar t x\<^sub>i) \ i = eq_idx_for_lvar t x\<^sub>i" using \i < length t\ \eq_idx_for_lvar t x\<^sub>i < length t\ by (auto simp add: distinct_conv_nth) then show "?f' (t ! i) = ?f i" by (auto simp add: eq_for_lvar_def) qed then show "pivot_tableau' x\<^sub>i x\<^sub>j t = pivot_tableau x\<^sub>i x\<^sub>j t" unfolding pivot_tableau'_def pivot_tableau_def unfolding Let_def by (auto simp add: map_reindex) qed lemma pivot'pivot: fixes s :: "('i,'a::lrv)state" assumes "\ (\ s)" "x\<^sub>i \ lvars (\ s)" shows "pivot' x\<^sub>i x\<^sub>j s = pivot x\<^sub>i x\<^sub>j s" using pivot_tableau'pivot_tableau[OF assms] unfolding pivot_def pivot'_def by auto end sublocale Pivot' < Pivot eq_idx_for_lvar pivot proof fix s::"('i,'a) state" and x\<^sub>i x\<^sub>j and v::"'a valuation" assume "\ (\ s)" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" show "let s' = pivot x\<^sub>i x\<^sub>j s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>c s" unfolding pivot_def by (auto simp add: Let_def simp: boundsl_def boundsu_def indexl_def indexu_def) let ?t = "\ s" let ?idx = "eq_idx_for_lvar ?t x\<^sub>i" let ?eq = "?t ! ?idx" let ?eq' = "pivot_eq ?eq x\<^sub>j" have "?idx < length ?t" "lhs (?t ! ?idx) = x\<^sub>i" using \x\<^sub>i \ lvars ?t\ using eq_idx_for_lvar by auto have "distinct (map lhs ?t)" using \\ ?t\ unfolding normalized_tableau_def by simp have "x\<^sub>j \ rvars_eq ?eq" using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ unfolding eq_for_lvar_def by simp then have "x\<^sub>j \ rvars ?t" using \?idx < length ?t\ using in_set_conv_nth[of ?eq ?t] by (auto simp add: rvars_def) then have "x\<^sub>j \ lvars ?t" using \\ ?t\ unfolding normalized_tableau_def by auto have "x\<^sub>i \ rvars ?t" using \x\<^sub>i \ lvars ?t\ \\ ?t\ unfolding normalized_tableau_def rvars_def by auto then have "x\<^sub>i \ rvars_eq ?eq" unfolding rvars_def using \?idx < length ?t\ using in_set_conv_nth[of ?eq ?t] by auto have "x\<^sub>i \ x\<^sub>j" using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ rvars_eq ?eq\ by auto have "?eq' = (x\<^sub>j, rhs ?eq')" using lhs_pivot_eq[of x\<^sub>j ?eq] using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ \lhs (?t ! ?idx) = x\<^sub>i\ \x\<^sub>i \ rvars_eq ?eq\ by (auto simp add: eq_for_lvar_def) (cases "?eq'", simp)+ let ?I1 = "[0..?idx < length ?t\ by (rule interval_3split) then have map_lhs_pivot: "map lhs (\ (pivot' x\<^sub>i x\<^sub>j s)) = map (\idx. lhs (?t ! idx)) ?I1 @ [x\<^sub>j] @ map (\idx. lhs (?t ! idx)) ?I2" using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ \lhs (?t ! ?idx) = x\<^sub>i\ \x\<^sub>i \ rvars_eq ?eq\ by (auto simp add: Let_def subst_var_eq_def eq_for_lvar_def lhs_pivot_eq pivot'_def pivot_tableau'_def) have lvars_pivot: "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" proof- have "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = {x\<^sub>j} \ (\idx. lhs (?t ! idx)) ` ({0..?idx < length ?t\ \?eq' = (x\<^sub>j, rhs ?eq')\ by (cases ?eq', auto simp add: Let_def pivot'_def pivot_tableau'_def lvars_def subst_var_eq_def)+ also have "... = {x\<^sub>j} \ (((\idx. lhs (?t ! idx)) ` {0..?idx < length ?t\ \distinct (map lhs ?t)\ by (auto simp add: distinct_conv_nth) also have "... = {x\<^sub>j} \ (set (map lhs ?t) - {x\<^sub>i})" using \lhs (?t ! ?idx) = x\<^sub>i\ by (auto simp add: in_set_conv_nth rev_image_eqI) (auto simp add: image_def) finally show "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" by (simp add: lvars_def) qed moreover have rvars_pivot: "rvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" proof- have "rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>j})" using rvars_pivot_eq[of x\<^sub>j ?eq] using \lhs (?t ! ?idx) = x\<^sub>i\ using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ rvars_eq ?eq\ by simp let ?S1 = "rvars_eq ?eq'" let ?S2 = "\idx\({0..j (rhs ?eq') (?t ! idx))" have "rvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = ?S1 \ ?S2" unfolding pivot'_def pivot_tableau'_def rvars_def using \?idx < length ?t\ by (auto simp add: Let_def split: if_splits) also have "... = {x\<^sub>i} \ (rvars ?t - {x\<^sub>j})" (is "?S1 \ ?S2 = ?rhs") proof show "?S1 \ ?S2 \ ?rhs" proof- have "?S1 \ ?rhs" using \?idx < length ?t\ unfolding rvars_def using \rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>j})\ by (force simp add: in_set_conv_nth) moreover have "?S2 \ ?rhs" proof- have "?S2 \ (\idx\{0..j}) \ rvars_eq ?eq')" apply (rule UN_mono) using rvars_eq_subst_var_eq by auto also have "... \ rvars_eq ?eq' \ (\idx\{0..j})" by auto also have "... = rvars_eq ?eq' \ (rvars ?t - {x\<^sub>j})" unfolding rvars_def by (force simp add: in_set_conv_nth) finally show ?thesis using \rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>j})\ unfolding rvars_def using \?idx < length ?t\ by auto qed ultimately show ?thesis by simp qed next show "?rhs \ ?S1 \ ?S2" proof fix x assume "x \ ?rhs" show "x \ ?S1 \ ?S2" proof (cases "x \ rvars_eq ?eq'") case True then show ?thesis by auto next case False let ?S2' = "\idx\({0..j}) - rvars_eq ?eq'" have "x \ ?S2'" using False \x \ ?rhs\ using \rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>j})\ unfolding rvars_def by (force simp add: in_set_conv_nth) moreover have "?S2 \ ?S2'" apply (rule UN_mono) using rvars_eq_subst_var_eq_supset[of _ x\<^sub>j "rhs ?eq'" ] by auto ultimately show ?thesis by auto qed qed qed ultimately show ?thesis by simp qed ultimately show "let s' = pivot x\<^sub>i x\<^sub>j s in rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i} \ lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" using pivot'pivot[where ?'i = 'i] using \\ (\ s)\ \x\<^sub>i \ lvars (\ s)\ by (simp add: Let_def) have "\ (\ (pivot' x\<^sub>i x\<^sub>j s))" unfolding normalized_tableau_def proof have "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) \ rvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = {}" (is ?g1) using \\ (\ s)\ unfolding normalized_tableau_def using lvars_pivot rvars_pivot using \x\<^sub>i \ x\<^sub>j\ by auto moreover have "0 \ rhs ` set (\ (pivot' x\<^sub>i x\<^sub>j s))" (is ?g2) proof let ?eq = "eq_for_lvar (\ s) x\<^sub>i" from eq_for_lvar[OF \x\<^sub>i \ lvars (\ s)\] have "?eq \ set (\ s)" and var: "lhs ?eq = x\<^sub>i" by auto have "lhs ?eq \ rvars_eq ?eq" using \\ (\ s)\ \?eq \ set (\ s)\ using \x\<^sub>i \ rvars_eq (\ s ! eq_idx_for_lvar (\ s) x\<^sub>i)\ eq_for_lvar_def var by auto from vars_pivot_eq[OF \x\<^sub>j \ rvars_eq ?eq\ this] have vars_pivot: "lhs (pivot_eq ?eq x\<^sub>j) = x\<^sub>j" "rvars_eq (pivot_eq ?eq x\<^sub>j) = {lhs (eq_for_lvar (\ s) x\<^sub>i)} \ (rvars_eq (eq_for_lvar (\ s) x\<^sub>i) - {x\<^sub>j})" unfolding Let_def by auto from vars_pivot(2) have rhs_pivot0: "rhs (pivot_eq ?eq x\<^sub>j) \ 0" using vars_zero by auto assume "0 \ rhs ` set (\ (pivot' x\<^sub>i x\<^sub>j s))" from this[unfolded pivot'pivot[OF \\ (\ s)\ \x\<^sub>i \ lvars (\ s)\] pivot_def] have "0 \ rhs ` set (pivot_tableau x\<^sub>i x\<^sub>j (\ s))" by simp from this[unfolded pivot_tableau_def Let_def var, unfolded var] rhs_pivot0 obtain e where "e \ set (\ s)" "lhs e \ x\<^sub>i" and rvars_eq: "rvars_eq (subst_var_eq x\<^sub>j (rhs (pivot_eq ?eq x\<^sub>j)) e) = {}" by (auto simp: vars_zero) from rvars_eq[unfolded subst_var_eq_def] have empty: "vars (subst_var x\<^sub>j (rhs (pivot_eq ?eq x\<^sub>j)) (rhs e)) = {}" by auto show False proof (cases "x\<^sub>j \ vars (rhs e)") case False from empty[unfolded subst_no_effect[OF False]] have "rvars_eq e = {}" by auto hence "rhs e = 0" using zero_coeff_zero coeff_zero by auto with \e \ set (\ s)\ \\ (\ s)\ show False unfolding normalized_tableau_def by auto next case True from \e \ set (\ s)\ have "rvars_eq e \ rvars (\ s)" unfolding rvars_def by auto hence "x\<^sub>i \ vars (rhs (pivot_eq ?eq x\<^sub>j)) - rvars_eq e" unfolding vars_pivot(2) var using \\ (\ s)\[unfolded normalized_tableau_def] \x\<^sub>i \ lvars (\ s)\ by auto from subst_with_effect[OF True this] rvars_eq show ?thesis by (simp add: subst_var_eq_def) qed qed ultimately show "?g1 \ ?g2" .. show "distinct (map lhs (\ (pivot' x\<^sub>i x\<^sub>j s)))" using map_parametrize_idx[of lhs ?t] using map_lhs_pivot using \distinct (map lhs ?t)\ using interval_3split[of ?idx "length ?t"] \?idx < length ?t\ using \x\<^sub>j \ lvars ?t\ unfolding lvars_def by auto qed moreover have "v \\<^sub>t ?t = v \\<^sub>t \ (pivot' x\<^sub>i x\<^sub>j s)" unfolding satisfies_tableau_def proof assume "\e\set (?t). v \\<^sub>e e" show "\e\set (\ (pivot' x\<^sub>i x\<^sub>j s)). v \\<^sub>e e" proof- have "v \\<^sub>e ?eq'" using \x\<^sub>i \ rvars_eq ?eq\ using \?idx < length ?t\ \\e\set (?t). v \\<^sub>e e\ using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ lvars ?t\ by (simp add: equiv_pivot_eq eq_idx_for_lvar) moreover { fix idx assume "idx < length ?t" "idx \ ?idx" have "v \\<^sub>e subst_var_eq x\<^sub>j (rhs ?eq') (?t ! idx)" using \?eq' = (x\<^sub>j, rhs ?eq')\ using \v \\<^sub>e ?eq'\ \idx < length ?t\ \\e\set (?t). v \\<^sub>e e\ using equiv_subst_var_eq[of v x\<^sub>j "rhs ?eq'" "?t ! idx"] by auto } ultimately show ?thesis by (auto simp add: pivot'_def pivot_tableau'_def Let_def) qed next assume "\e\set (\ (pivot' x\<^sub>i x\<^sub>j s)). v \\<^sub>e e" then have "v \\<^sub>e ?eq'" "\ idx. \idx < length ?t; idx \ ?idx \ \ v \\<^sub>e subst_var_eq x\<^sub>j (rhs ?eq') (?t ! idx)" using \?idx < length ?t\ unfolding pivot'_def pivot_tableau'_def by (auto simp add: Let_def) show "\e\set (\ s). v \\<^sub>e e" proof- { fix idx assume "idx < length ?t" have "v \\<^sub>e (?t ! idx)" proof (cases "idx = ?idx") case True then show ?thesis using \v \\<^sub>e ?eq'\ using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ lvars ?t\ \x\<^sub>i \ rvars_eq ?eq\ by (simp add: eq_idx_for_lvar equiv_pivot_eq) next case False then show ?thesis using \idx < length ?t\ using \\idx < length ?t; idx \ ?idx \ \ v \\<^sub>e subst_var_eq x\<^sub>j (rhs ?eq') (?t ! idx)\ using \v \\<^sub>e ?eq'\ \?eq' = (x\<^sub>j, rhs ?eq')\ using equiv_subst_var_eq[of v x\<^sub>j "rhs ?eq'" "?t ! idx"] by auto qed } then show ?thesis by (force simp add: in_set_conv_nth) qed qed ultimately show "let s' = pivot x\<^sub>i x\<^sub>j s in v \\<^sub>t \ s = v \\<^sub>t \ s' \ \ (\ s')" using pivot'pivot[where ?'i = 'i] using \\ (\ s)\ \x\<^sub>i \ lvars (\ s)\ by (simp add: Let_def) qed subsection\Check implementation\ text\The \check\ function is called when all rhs variables are in bounds, and it checks if there is a lhs variable that is not. If there is no such variable, then satisfiability is detected and \check\ succeeds. If there is a lhs variable \x\<^sub>i\ out of its bounds, a rhs variable \x\<^sub>j\ is sought which allows pivoting with \x\<^sub>i\ and updating \x\<^sub>i\ to its violated bound. If \x\<^sub>i\ is under its lower bound it must be increased, and if \x\<^sub>j\ has a positive coefficient it must be increased so it must be under its upper bound and if it has a negative coefficient it must be decreased so it must be above its lower bound. The case when \x\<^sub>i\ is above its upper bound is symmetric (avoiding symmetries is discussed in Section \ref{sec:symmetries}). If there is no such \x\<^sub>j\, unsatisfiability is detected and \check\ fails. The procedure is recursively repeated, until it either succeeds or fails. To ensure termination, variables \x\<^sub>i\ and \x\<^sub>j\ must be chosen with respect to a fixed variable ordering. For choosing these variables auxiliary functions \min_lvar_not_in_bounds\, \min_rvar_inc\ and \min_rvar_dec\ are specified (each in its own locale). For, example: \ locale MinLVarNotInBounds = fixes min_lvar_not_in_bounds::"('i,'a::lrv) state \ var option" assumes min_lvar_not_in_bounds_None: "min_lvar_not_in_bounds s = None \ (\x\lvars (\ s). in_bounds x \\ s\ (\ s))" and min_lvar_not_in_bounds_Some': "min_lvar_not_in_bounds s = Some x\<^sub>i \ x\<^sub>i\lvars (\ s) \ \in_bounds x\<^sub>i \\ s\ (\ s) \ (\x\lvars (\ s). x < x\<^sub>i \ in_bounds x \\ s\ (\ s))" begin lemma min_lvar_not_in_bounds_None': "min_lvar_not_in_bounds s = None \ (\\ s\ \\<^sub>b \ s \ lvars (\ s))" unfolding satisfies_bounds_set.simps by (rule min_lvar_not_in_bounds_None) lemma min_lvar_not_in_bounds_lvars:"min_lvar_not_in_bounds s = Some x\<^sub>i \ x\<^sub>i \ lvars (\ s)" using min_lvar_not_in_bounds_Some' by simp lemma min_lvar_not_in_bounds_Some: "min_lvar_not_in_bounds s = Some x\<^sub>i \ \ in_bounds x\<^sub>i \\ s\ (\ s)" using min_lvar_not_in_bounds_Some' by simp lemma min_lvar_not_in_bounds_Some_min: "min_lvar_not_in_bounds s = Some x\<^sub>i \ (\ x \ lvars (\ s). x < x\<^sub>i \ in_bounds x \\ s\ (\ s))" using min_lvar_not_in_bounds_Some' by simp end abbreviation reasable_var where "reasable_var dir x eq s \ (coeff (rhs eq) x > 0 \ \\<^sub>u\<^sub>b (lt dir) (\\ s\ x) (UB dir s x)) \ (coeff (rhs eq) x < 0 \ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x) (LB dir s x))" locale MinRVarsEq = fixes min_rvar_incdec_eq :: "('i,'a) Direction \ ('i,'a::lrv) state \ eq \ 'i list + var" assumes min_rvar_incdec_eq_None: "min_rvar_incdec_eq dir s eq = Inl is \ (\ x \ rvars_eq eq. \ reasable_var dir x eq s) \ (set is = {LI dir s (lhs eq)} \ {LI dir s x | x. x \ rvars_eq eq \ coeff (rhs eq) x < 0} \ {UI dir s x | x. x \ rvars_eq eq \ coeff (rhs eq) x > 0}) \ ((dir = Positive \ dir = Negative) \ LI dir s (lhs eq) \ indices_state s \ set is \ indices_state s)" assumes min_rvar_incdec_eq_Some_rvars: "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j \ x\<^sub>j \ rvars_eq eq" assumes min_rvar_incdec_eq_Some_incdec: "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j \ reasable_var dir x\<^sub>j eq s" assumes min_rvar_incdec_eq_Some_min: "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j \ (\ x \ rvars_eq eq. x < x\<^sub>j \ \ reasable_var dir x eq s)" begin lemma min_rvar_incdec_eq_None': assumes *: "dir = Positive \ dir = Negative" and min: "min_rvar_incdec_eq dir s eq = Inl is" and sub: "I = set is" and Iv: "(I,v) \\<^sub>i\<^sub>b \\ s" shows "le (lt dir) ((rhs eq) \v\) ((rhs eq) \\\ s\\)" proof - have "\ x \ rvars_eq eq. \ reasable_var dir x eq s" using min using min_rvar_incdec_eq_None by simp have "\ x \ rvars_eq eq. (0 < coeff (rhs eq) x \ le (lt dir) 0 (\\ s\ x - v x)) \ (coeff (rhs eq) x < 0 \ le (lt dir) (\\ s\ x - v x) 0)" proof (safe) fix x assume x: "x \ rvars_eq eq" "0 < coeff (rhs eq) x" "0 \ \\ s\ x - v x" then have "\ (\\<^sub>u\<^sub>b (lt dir) (\\ s\ x) (UB dir s x))" using \\ x \ rvars_eq eq. \ reasable_var dir x eq s\ by auto then have "\\<^sub>u\<^sub>b (lt dir) (\\ s\ x) (UB dir s x)" using * by (cases "UB dir s x") (auto simp add: bound_compare_defs) moreover from min_rvar_incdec_eq_None[OF min] x sub have "UI dir s x \ I" by auto from Iv * this have "\\<^sub>u\<^sub>b (lt dir) (v x) (UB dir s x)" unfolding satisfies_bounds_index.simps by (cases "UB dir s x", auto simp: indexl_def indexu_def boundsl_def boundsu_def bound_compare'_defs) (fastforce)+ ultimately have "le (lt dir) (v x) (\\ s\ x)" using * by (cases "UB dir s x") (auto simp add: bound_compare_defs) then show "lt dir 0 (\\ s\ x - v x)" using \0 \ \\ s\ x - v x\ * using minus_gt[of "v x" "\\ s\ x"] minus_lt[of "\\ s\ x" "v x"] - by auto + by (auto simp del: Simplex.bounds_lg) next fix x assume x: "x \ rvars_eq eq" "0 > coeff (rhs eq) x" "\\ s\ x - v x \ 0" then have "\ (\\<^sub>l\<^sub>b (lt dir) (\\ s\ x) (LB dir s x))" using \\ x \ rvars_eq eq. \ reasable_var dir x eq s\ by auto then have "\\<^sub>l\<^sub>b (lt dir) (\\ s\ x) (LB dir s x)" using * by (cases "LB dir s x") (auto simp add: bound_compare_defs) moreover from min_rvar_incdec_eq_None[OF min] x sub have "LI dir s x \ I" by auto from Iv * this have "\\<^sub>l\<^sub>b (lt dir) (v x) (LB dir s x)" unfolding satisfies_bounds_index.simps by (cases "LB dir s x", auto simp: indexl_def indexu_def boundsl_def boundsu_def bound_compare'_defs) (fastforce)+ ultimately have "le (lt dir) (\\ s\ x) (v x)" using * by (cases "LB dir s x") (auto simp add: bound_compare_defs) then show "lt dir (\\ s\ x - v x) 0" using \\\ s\ x - v x \ 0\ * using minus_lt[of "\\ s\ x" "v x"] minus_gt[of "v x" "\\ s\ x"] - by auto + by (auto simp del: Simplex.bounds_lg) qed then have "le (lt dir) 0 (rhs eq \ \ x. \\ s\ x - v x\)" using * apply auto using valuate_nonneg[of "rhs eq" "\x. \\ s\ x - v x"] - apply force + apply (force simp del: Simplex.bounds_lg) using valuate_nonpos[of "rhs eq" "\x. \\ s\ x - v x"] - apply force + apply (force simp del: Simplex.bounds_lg) done then show "le (lt dir) rhs eq \ v \ rhs eq \ \\ s\ \" using \dir = Positive \ dir = Negative\ using minus_gt[of "rhs eq \ v \" "rhs eq \ \\ s\ \"] - by (auto simp add: valuate_diff[THEN sym]) + by (auto simp add: valuate_diff[THEN sym] simp del: Simplex.bounds_lg) qed end locale MinRVars = EqForLVar + MinRVarsEq min_rvar_incdec_eq for min_rvar_incdec_eq :: "('i, 'a :: lrv) Direction \ _" begin abbreviation min_rvar_incdec :: "('i,'a) Direction \ ('i,'a) state \ var \ 'i list + var" where "min_rvar_incdec dir s x\<^sub>i \ min_rvar_incdec_eq dir s (eq_for_lvar (\ s) x\<^sub>i)" end locale MinVars = MinLVarNotInBounds min_lvar_not_in_bounds + MinRVars eq_idx_for_lvar min_rvar_incdec_eq for min_lvar_not_in_bounds :: "('i,'a::lrv) state \ _" and eq_idx_for_lvar and min_rvar_incdec_eq :: "('i, 'a :: lrv) Direction \ _" locale PivotUpdateMinVars = PivotAndUpdate eq_idx_for_lvar pivot_and_update + MinVars min_lvar_not_in_bounds eq_idx_for_lvar min_rvar_incdec_eq for eq_idx_for_lvar :: "tableau \ var \ nat" and min_lvar_not_in_bounds :: "('i,'a::lrv) state \ var option" and min_rvar_incdec_eq :: "('i,'a) Direction \ ('i,'a) state \ eq \ 'i list + var" and pivot_and_update :: "var \ var \ 'a \ ('i,'a) state \ ('i,'a) state" begin definition check' where "check' dir x\<^sub>i s \ let l\<^sub>i = the (LB dir s x\<^sub>i); x\<^sub>j' = min_rvar_incdec dir s x\<^sub>i in case x\<^sub>j' of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s" lemma check'_cases: assumes "\ I. \min_rvar_incdec dir s x\<^sub>i = Inl I; check' dir x\<^sub>i s = set_unsat I s\ \ P (set_unsat I s)" assumes "\ x\<^sub>j l\<^sub>i. \min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j; l\<^sub>i = the (LB dir s x\<^sub>i); check' dir x\<^sub>i s = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s\ \ P (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)" shows "P (check' dir x\<^sub>i s)" using assms unfolding check'_def by (cases "min_rvar_incdec dir s x\<^sub>i", auto) partial_function (tailrec) check where "check s = (if \ s then s else let x\<^sub>i' = min_lvar_not_in_bounds s in case x\<^sub>i' of None \ s | Some x\<^sub>i \ let dir = if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative in check (check' dir x\<^sub>i s))" declare check.simps[code] inductive check_dom where step: "\\x\<^sub>i. \\ \ s; Some x\<^sub>i = min_lvar_not_in_bounds s; \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ \ check_dom (check' Positive x\<^sub>i s); \x\<^sub>i. \\ \ s; Some x\<^sub>i = min_lvar_not_in_bounds s; \ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ \ check_dom (check' Negative x\<^sub>i s)\ \ check_dom s" text\ The definition of \check\ can be given by: @{text[display] "check s \ if \ s then s else let x\<^sub>i' = min_lvar_not_in_bounds s in case x\<^sub>i' of None \ s | Some x\<^sub>i \ if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then check (check_inc x\<^sub>i s) else check (check_dec x\<^sub>i s)" } @{text[display] "check_inc x\<^sub>i s \ let l\<^sub>i = the (\\<^sub>l s x\<^sub>i); x\<^sub>j' = min_rvar_inc s x\<^sub>i in case x\<^sub>j' of None \ s \ \ := True \ | Some x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s" } The definition of \check_dec\ is analogous. It is shown (mainly by induction) that this definition satisfies the \check\ specification. Note that this definition uses general recursion, so its termination is non-trivial. It has been shown that it terminates for all states satisfying the check preconditions. The proof is based on the proof outline given in \cite{simplex-rad}. It is very technically involved, but conceptually uninteresting so we do not discuss it in more details.\ lemma pivotandupdate_check_precond: assumes "dir = (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative)" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j" "l\<^sub>i = the (LB dir s x\<^sub>i)" "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" " \ s" shows "\ (\ (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s) \ \ (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s) \ \ (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)" proof- have "\\<^sub>l s x\<^sub>i = Some l\<^sub>i \ \\<^sub>u s x\<^sub>i = Some l\<^sub>i" using \l\<^sub>i = the (LB dir s x\<^sub>i)\ \dir = (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative)\ using \min_lvar_not_in_bounds s = Some x\<^sub>i\ min_lvar_not_in_bounds_Some[of s x\<^sub>i] using \\ s\ by (case_tac[!] "\\<^sub>l s x\<^sub>i", case_tac[!] "\\<^sub>u s x\<^sub>i") (auto simp add: bounds_consistent_def bound_compare_defs) then show ?thesis using assms using pivotandupdate_tableau_normalized[of s x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_nolhs[of s x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_bounds_consistent[of s x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_tableau_valuated[of s x\<^sub>i x\<^sub>j l\<^sub>i] by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars) qed (* -------------------------------------------------------------------------- *) (* Termination *) (* -------------------------------------------------------------------------- *) abbreviation gt_state' where "gt_state' dir s s' x\<^sub>i x\<^sub>j l\<^sub>i \ min_lvar_not_in_bounds s = Some x\<^sub>i \ l\<^sub>i = the (LB dir s x\<^sub>i) \ min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j \ s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s" definition gt_state :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\\<^sub>x" 100) where "s \\<^sub>x s' \ \ x\<^sub>i x\<^sub>j l\<^sub>i. let dir = if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative in gt_state' dir s s' x\<^sub>i x\<^sub>j l\<^sub>i" abbreviation succ :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\" 100) where "s \ s' \ \ (\ s) \ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \ s \ s \\<^sub>x s' \ \\<^sub>i s' = \\<^sub>i s \ \\<^sub>c s' = \\<^sub>c s" abbreviation succ_rel :: "('i,'a) state rel" where "succ_rel \ {(s, s'). s \ s'}" abbreviation succ_rel_trancl :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\\<^sup>+" 100) where "s \\<^sup>+ s' \ (s, s') \ succ_rel\<^sup>+" abbreviation succ_rel_rtrancl :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\\<^sup>*" 100) where "s \\<^sup>* s' \ (s, s') \ succ_rel\<^sup>*" lemma succ_vars: assumes "s \ s'" obtains x\<^sub>i x\<^sub>j where "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" "x\<^sub>j \ rvars (\ s)" "lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" proof- from assms obtain x\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" "lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j] using pivotandupdate_lvars[of s x\<^sub>i x\<^sub>j] by auto moreover have "x\<^sub>j \ rvars (\ s)" using \x\<^sub>i \ lvars (\ s)\ using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ using eq_for_lvar[of x\<^sub>i "\ s"] unfolding rvars_def by auto ultimately have "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" "x\<^sub>j \ rvars (\ s)" "lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" by auto then show thesis .. qed lemma succ_vars_id: assumes "s \ s'" shows "lvars (\ s) \ rvars (\ s) = lvars (\ s') \ rvars (\ s')" using assms by (rule succ_vars) auto lemma succ_inv: assumes "s \ s'" shows "\ (\ s')" "\ s'" "\ s'" "\\<^sub>i s = \\<^sub>i s'" "(v::'a valuation) \\<^sub>t (\ s) \ v \\<^sub>t (\ s')" proof- from assms obtain x\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then show "\ (\ s')" "\ s'" "\ s'" "\\<^sub>i s = \\<^sub>i s'" "(v::'a valuation) \\<^sub>t (\ s) \ v \\<^sub>t (\ s')" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using pivotandupdate_tableau_normalized[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_bounds_consistent[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_bounds_id[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_tableau_equiv using pivotandupdate_tableau_valuated by auto qed lemma succ_rvar_valuation_id: assumes "s \ s'" "x \ rvars (\ s)" "x \ rvars (\ s')" shows "\\ s\ x = \\ s'\ x" proof- from assms obtain x\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then show ?thesis using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using \x \ rvars (\ s)\ \x \ rvars (\ s')\ using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>j x c] by (force simp add: normalized_tableau_def map2fun_def) qed lemma succ_min_lvar_not_in_bounds: assumes "s \ s'" "xr \ lvars (\ s)" "xr \ rvars (\ s')" shows "\ in_bounds xr (\\ s\) (\ s)" "\ x \ lvars (\ s). x < xr \ in_bounds x (\\ s\) (\ s)" proof- from assms obtain x\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then have "x\<^sub>i = xr" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using \xr \ lvars (\ s)\ \xr \ rvars (\ s')\ using pivotandupdate_rvars by (auto simp add: normalized_tableau_def) then show "\ in_bounds xr (\\ s\) (\ s)" "\ x \ lvars (\ s). x < xr \ in_bounds x (\\ s\) (\ s)" using \min_lvar_not_in_bounds s = Some x\<^sub>i\ using min_lvar_not_in_bounds_Some min_lvar_not_in_bounds_Some_min by simp_all qed lemma succ_min_rvar: assumes "s \ s'" "xs \ lvars (\ s)" "xs \ rvars (\ s')" "xr \ rvars (\ s)" "xr \ lvars (\ s')" "eq = eq_for_lvar (\ s) xs" and dir: "dir = Positive \ dir = Negative" shows "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ xs) (LB dir s xs) \ reasable_var dir xr eq s \ (\ x \ rvars_eq eq. x < xr \ \ reasable_var dir x eq s)" proof- from assms(1) obtain x\<^sub>i x\<^sub>j c where"\ (\ s) \ \ s \ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "gt_state' (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative) s s' x\<^sub>i x\<^sub>j c" by (auto simp add: gt_state_def Let_def) then have "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" and *: "(\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i \ min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j) \ (\ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j)" by (auto split: if_splits) then have "xr = x\<^sub>j" "xs = x\<^sub>i" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using \xr \ rvars (\ s)\ \xr \ lvars (\ s')\ using \xs \ lvars (\ s)\ \xs \ rvars (\ s')\ using pivotandupdate_lvars pivotandupdate_rvars by (auto simp add: normalized_tableau_def) show "\ (\\<^sub>l\<^sub>b (lt dir) (\\ s\ xs) (LB dir s xs)) \ reasable_var dir xr eq s \ (\ x \ rvars_eq eq. x < xr \ \ reasable_var dir x eq s)" proof assume "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ xs) (LB dir s xs)" then have "\\<^sub>l\<^sub>b (lt dir) (\\ s\ xs) (LB dir s xs)" using dir by (cases "LB dir s xs") (auto simp add: bound_compare_defs) moreover then have "\ (\\<^sub>u\<^sub>b (lt dir) (\\ s\ xs) (UB dir s xs))" using \\ s\ dir using bounds_consistent_gt_ub bounds_consistent_lt_lb by (force simp add: bound_compare''_defs) ultimately have "min_rvar_incdec dir s xs = Inr xr" using * \xr = x\<^sub>j\ \xs = x\<^sub>i\ dir by (auto simp add: bound_compare''_defs) then show "reasable_var dir xr eq s \ (\ x \ rvars_eq eq. x < xr \ \ reasable_var dir x eq s)" using \eq = eq_for_lvar (\ s) xs\ using min_rvar_incdec_eq_Some_min[of dir s eq xr] using min_rvar_incdec_eq_Some_incdec[of dir s eq xr] by simp qed qed lemma succ_set_on_bound: assumes "s \ s'" "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ s')" and dir: "dir = Positive \ dir = Negative" shows "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (LB dir s x\<^sub>i)" "\\ s'\ x\<^sub>i = the (\\<^sub>l s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (\\<^sub>u s x\<^sub>i)" proof- from assms(1) obtain x\<^sub>i' x\<^sub>j c where"\ (\ s) \ \ s \ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "gt_state' (if \\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' then Positive else Negative) s s' x\<^sub>i' x\<^sub>j c" by (auto simp add: gt_state_def Let_def) then have "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i'" "s' = pivot_and_update x\<^sub>i' x\<^sub>j c s" and *: "(\\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' \ c = the (\\<^sub>l s x\<^sub>i') \ min_rvar_incdec Positive s x\<^sub>i' = Inr x\<^sub>j) \ (\ \\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' \ c = the (\\<^sub>u s x\<^sub>i') \ min_rvar_incdec Negative s x\<^sub>i' = Inr x\<^sub>j)" by (auto split: if_splits) then have "x\<^sub>i = x\<^sub>i'" "x\<^sub>i' \ lvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i')" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i'] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i'" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i'" x\<^sub>j] using \x\<^sub>i \ lvars (\ s)\ \x\<^sub>i \ rvars (\ s')\ using pivotandupdate_rvars by (auto simp add: normalized_tableau_def) show "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (LB dir s x\<^sub>i)" proof assume "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)" then have "\\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)" using dir by (cases "LB dir s x\<^sub>i") (auto simp add: bound_compare_defs) moreover then have "\ \\<^sub>u\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (UB dir s x\<^sub>i)" using \\ s\ dir using bounds_consistent_gt_ub bounds_consistent_lt_lb by (force simp add: bound_compare''_defs) ultimately show "\\ s'\ x\<^sub>i = the (LB dir s x\<^sub>i)" using * \x\<^sub>i = x\<^sub>i'\ \s' = pivot_and_update x\<^sub>i' x\<^sub>j c s\ using \\ (\ s)\ \\ s\ \x\<^sub>i' \ lvars (\ s)\ \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i')\ using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] dir by (case_tac[!] "\\<^sub>l s x\<^sub>i'", case_tac[!] "\\<^sub>u s x\<^sub>i'") (auto simp add: bound_compare_defs map2fun_def) qed have "\ \\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' \ \\ s\ x\<^sub>i' >\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i'" using \min_lvar_not_in_bounds s = Some x\<^sub>i'\ using min_lvar_not_in_bounds_Some[of s x\<^sub>i'] using not_in_bounds[of x\<^sub>i' "\\ s\" "\\<^sub>l s" "\\<^sub>u s"] by auto then show "\\ s'\ x\<^sub>i = the (\\<^sub>l s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (\\<^sub>u s x\<^sub>i)" using \\ (\ s)\ \\ s\ \x\<^sub>i' \ lvars (\ s)\ \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i')\ using \s' = pivot_and_update x\<^sub>i' x\<^sub>j c s\ \x\<^sub>i = x\<^sub>i'\ using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] using * by (case_tac[!] "\\<^sub>l s x\<^sub>i'", case_tac[!] "\\<^sub>u s x\<^sub>i'") (auto simp add: map2fun_def bound_compare_defs) qed lemma succ_rvar_valuation: assumes "s \ s'" "x \ rvars (\ s')" shows "\\ s'\ x = \\ s\ x \ \\ s'\ x = the (\\<^sub>l s x) \ \\ s'\ x = the (\\<^sub>u s x)" proof- from assms obtain x\<^sub>i x\<^sub>j b where "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)" "s' = pivot_and_update x\<^sub>i x\<^sub>j b s" unfolding gt_state_def by (auto simp add: Let_def split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" "x\<^sub>j \ rvars (\ s)" "x\<^sub>j \ lvars (\ s)" "x\<^sub>i \ x\<^sub>j" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using rvars_of_lvar_rvars \\ (\ s)\ by (auto simp add: normalized_tableau_def) then have "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" "x \ rvars (\ s) \ x = x\<^sub>i" "x \ x\<^sub>j" "x \ x\<^sub>i \ x \ lvars (\ s)" using \x \ rvars (\ s')\ using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j] using \\ (\ s)\ \\ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ by (auto simp add: normalized_tableau_def) then show ?thesis using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j b] using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>j x b] using \x\<^sub>i \ lvars (\ s)\ \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ using \\ (\ s)\ \\ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ \b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)\ by (auto simp add: map2fun_def) qed lemma succ_no_vars_valuation: assumes "s \ s'" "x \ tvars (\ s')" shows "look (\ s') x = look (\ s) x" proof- from assms obtain x\<^sub>i x\<^sub>j b where "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)" "s' = pivot_and_update x\<^sub>i x\<^sub>j b s" unfolding gt_state_def by (auto simp add: Let_def split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" "x\<^sub>j \ rvars (\ s)" "x\<^sub>j \ lvars (\ s)" "x\<^sub>i \ x\<^sub>j" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using rvars_of_lvar_rvars \\ (\ s)\ by (auto simp add: normalized_tableau_def) then show ?thesis using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>j x b] using \\ (\ s)\ \\ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ using \x \ tvars (\ s')\ using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j] using pivotandupdate_lvars[of s x\<^sub>i x\<^sub>j] by (auto simp add: map2fun_def) qed lemma succ_valuation_satisfies: assumes "s \ s'" "\\ s\ \\<^sub>t \ s" shows "\\ s'\ \\<^sub>t \ s'" proof- from \s \ s'\ obtain x\<^sub>i x\<^sub>j b where "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)" "s' = pivot_and_update x\<^sub>i x\<^sub>j b s" unfolding gt_state_def by (auto simp add: Let_def split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] \\ (\ s)\ by (auto simp add: normalized_tableau_def) then show ?thesis using pivotandupdate_satisfies_tableau[of s x\<^sub>i x\<^sub>j b] using pivotandupdate_tableau_equiv[of s x\<^sub>i x\<^sub>j ] using \\ (\ s)\ \\ s\ \\\ s\ \\<^sub>t \ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ by auto qed lemma succ_tableau_valuated: assumes "s \ s'" "\ s" shows "\ s'" using succ_inv(2) assms by blast (* -------------------------------------------------------------------------- *) abbreviation succ_chain where "succ_chain l \ rel_chain l succ_rel" lemma succ_chain_induct: assumes *: "succ_chain l" "i \ j" "j < length l" assumes base: "\ i. P i i" assumes step: "\ i. l ! i \ (l ! (i + 1)) \ P i (i + 1)" assumes trans: "\ i j k. \P i j; P j k; i < j; j \ k\ \ P i k" shows "P i j" using * proof (induct "j - i" arbitrary: i) case 0 then show ?case by (simp add: base) next case (Suc k) have "P (i + 1) j" using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5) by auto moreover have "P i (i + 1)" proof (rule step) show "l ! i \ (l ! (i + 1))" using Suc(2) Suc(3) Suc(5) unfolding rel_chain_def by auto qed ultimately show ?case using trans[of i "i + 1" j] Suc(2) by simp qed lemma succ_chain_bounds_id: assumes "succ_chain l" "i \ j" "j < length l" shows "\\<^sub>i (l ! i) = \\<^sub>i (l ! j)" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "\\<^sub>i (l ! i) = \\<^sub>i (l ! (i + 1))" by (rule succ_inv(4)) qed simp_all lemma succ_chain_vars_id': assumes "succ_chain l" "i \ j" "j < length l" shows "lvars (\ (l ! i)) \ rvars (\ (l ! i)) = lvars (\ (l ! j)) \ rvars (\ (l ! j))" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "tvars (\ (l ! i)) = tvars (\ (l ! (i + 1)))" by (rule succ_vars_id) qed simp_all lemma succ_chain_vars_id: assumes "succ_chain l" "i < length l" "j < length l" shows "lvars (\ (l ! i)) \ rvars (\ (l ! i)) = lvars (\ (l ! j)) \ rvars (\ (l ! j))" proof (cases "i \ j") case True then show ?thesis using assms succ_chain_vars_id'[of l i j] by simp next case False then have "j \ i" by simp then show ?thesis using assms succ_chain_vars_id'[of l j i] by simp qed lemma succ_chain_tableau_equiv': assumes "succ_chain l" "i \ j" "j < length l" shows "(v::'a valuation) \\<^sub>t \ (l ! i) \ v \\<^sub>t \ (l ! j)" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "v \\<^sub>t \ (l ! i) = v \\<^sub>t \ (l ! (i + 1))" by (rule succ_inv(5)) qed simp_all lemma succ_chain_tableau_equiv: assumes "succ_chain l" "i < length l" "j < length l" shows "(v::'a valuation) \\<^sub>t \ (l ! i) \ v \\<^sub>t \ (l ! j)" proof (cases "i \ j") case True then show ?thesis using assms succ_chain_tableau_equiv'[of l i j v] by simp next case False then have "j \ i" by auto then show ?thesis using assms succ_chain_tableau_equiv'[of l j i v] by simp qed lemma succ_chain_no_vars_valuation: assumes "succ_chain l" "i \ j" "j < length l" shows "\ x. x \ tvars (\ (l ! i)) \ look (\ (l ! i)) x = look (\ (l ! j)) x" (is "?P i j") using assms proof (induct "j - i" arbitrary: i) case 0 then show ?case by simp next case (Suc k) have "?P (i + 1) j" using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5) by auto moreover have "?P (i + 1) i" proof (rule+, rule succ_no_vars_valuation) show "l ! i \ (l ! (i + 1))" using Suc(2) Suc(3) Suc(5) unfolding rel_chain_def by auto qed moreover have "tvars (\ (l ! i)) = tvars (\ (l ! (i + 1)))" proof (rule succ_vars_id) show "l ! i \ (l ! (i + 1))" using Suc(2) Suc(3) Suc(5) unfolding rel_chain_def by simp qed ultimately show ?case by simp qed lemma succ_chain_rvar_valuation: assumes "succ_chain l" "i \ j" "j < length l" shows "\x\rvars (\ (l ! j)). \\ (l ! j)\ x = \\ (l ! i)\ x \ \\ (l ! j)\ x = the (\\<^sub>l (l ! i) x) \ \\ (l ! j)\ x = the (\\<^sub>u (l ! i) x)" (is "?P i j") using assms proof (induct "j - i" arbitrary: j) case 0 then show ?case by simp next case (Suc k) have "k = j - 1 - i" "succ_chain l" "i \ j - 1" "j - 1 < length l" "j > 0" using Suc(2) Suc(3) Suc(4) Suc(5) by auto then have ji: "?P i (j - 1)" using Suc(1) by simp have "l ! (j - 1) \ (l ! j)" using Suc(3) \j < length l\ \j > 0\ unfolding rel_chain_def by (erule_tac x="j - 1" in allE) simp then have jj: "?P (j - 1) j" using succ_rvar_valuation by auto obtain x\<^sub>i x\<^sub>j where vars: "x\<^sub>i \ lvars (\ (l ! (j - 1)))" "x\<^sub>j \ rvars (\ (l ! (j - 1)))" "rvars (\ (l ! j)) = rvars (\ (l ! (j - 1))) - {x\<^sub>j} \ {x\<^sub>i}" using \l ! (j - 1) \ (l ! j)\ by (rule succ_vars) simp then have bounds: "\\<^sub>l (l ! (j - 1)) = \\<^sub>l (l ! i)" "\\<^sub>l (l ! j) = \\<^sub>l (l ! i)" "\\<^sub>u (l ! (j - 1)) = \\<^sub>u (l ! i)" "\\<^sub>u (l ! j) = \\<^sub>u (l ! i)" using \succ_chain l\ using succ_chain_bounds_id[of l i "j - 1", THEN sym] \j - 1 < length l\ \i \ j - 1\ using succ_chain_bounds_id[of l "j - 1" j, THEN sym] \j < length l\ by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) show ?case proof fix x assume "x \ rvars (\ (l ! j))" then have "x \ x\<^sub>j \ x \ rvars (\ (l ! (j - 1))) \ x = x\<^sub>i" using vars by auto then show "\\ (l ! j)\ x = \\ (l ! i)\ x \ \\ (l ! j)\ x = the (\\<^sub>l (l ! i) x) \ \\ (l ! j)\ x = the (\\<^sub>u (l ! i) x)" proof assume "x \ x\<^sub>j \ x \ rvars (\ (l ! (j - 1)))" then show ?thesis using jj \x \ rvars (\ (l ! j))\ ji using bounds by force next assume "x = x\<^sub>i" then show ?thesis using succ_set_on_bound(2)[of "l ! (j - 1)" "l ! j" x\<^sub>i] \l ! (j - 1) \ (l ! j)\ using vars bounds by auto qed qed qed lemma succ_chain_valuation_satisfies: assumes "succ_chain l" "i \ j" "j < length l" shows "\\ (l ! i)\ \\<^sub>t \ (l ! i) \ \\ (l ! j)\ \\<^sub>t \ (l ! j)" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "\\ (l ! i)\ \\<^sub>t \ (l ! i) \ \\ (l ! (i + 1))\ \\<^sub>t \ (l ! (i + 1))" using succ_valuation_satisfies by auto qed simp_all lemma succ_chain_tableau_valuated: assumes "succ_chain l" "i \ j" "j < length l" shows "\ (l ! i) \ \ (l ! j)" using assms proof(rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "\ (l ! i) \ \ (l ! (i + 1))" using succ_tableau_valuated by auto qed simp_all abbreviation swap_lr where "swap_lr l i x \ i + 1 < length l \ x \ lvars (\ (l ! i)) \ x \ rvars (\ (l ! (i + 1)))" abbreviation swap_rl where "swap_rl l i x \ i + 1 < length l \ x \ rvars (\ (l ! i)) \ x \ lvars (\ (l ! (i + 1)))" abbreviation always_r where "always_r l i j x \ \ k. i \ k \ k \ j \ x \ rvars (\ (l ! k))" lemma succ_chain_always_r_valuation_id: assumes "succ_chain l" "i \ j" "j < length l" shows "always_r l i j x \ \\ (l ! i)\ x = \\ (l ! j)\ x" (is "?P i j") using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "?P i (i + 1)" using succ_rvar_valuation_id by simp qed simp_all lemma succ_chain_swap_rl_exists: assumes "succ_chain l" "i < j" "j < length l" "x \ rvars (\ (l ! i))" "x \ lvars (\ (l ! j))" shows "\ k. i \ k \ k < j \ swap_rl l k x" using assms proof (induct "j - i" arbitrary: i) case 0 then show ?case by simp next case (Suc k) have "l ! i \ (l ! (i + 1))" using Suc(3) Suc(4) Suc(5) unfolding rel_chain_def by auto then have "\ (\ (l ! (i + 1)))" by (rule succ_inv) show ?case proof (cases "x \ rvars (\ (l ! (i + 1)))") case True then have "j \ i + 1" using Suc(7) \\ (\ (l ! (i + 1)))\ by (auto simp add: normalized_tableau_def) have "k = j - Suc i" using Suc(2) by simp then obtain k where "k \ i + 1" "k < j" "swap_rl l k x" using \x \ rvars (\ (l ! (i + 1)))\ \j \ i + 1\ using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5) Suc(6) Suc(7) by auto then show ?thesis by (rule_tac x="k" in exI) simp next case False then have "x \ lvars (\ (l ! (i + 1)))" using Suc(6) using \l ! i \ (l ! (i + 1))\ succ_vars_id by auto then show ?thesis using Suc(4) Suc(5) Suc(6) by force qed qed lemma succ_chain_swap_lr_exists: assumes "succ_chain l" "i < j" "j < length l" "x \ lvars (\ (l ! i))" "x \ rvars (\ (l ! j))" shows "\ k. i \ k \ k < j \ swap_lr l k x" using assms proof (induct "j - i" arbitrary: i) case 0 then show ?case by simp next case (Suc k) have "l ! i \ (l ! (i + 1))" using Suc(3) Suc(4) Suc(5) unfolding rel_chain_def by auto then have "\ (\ (l ! (i + 1)))" by (rule succ_inv) show ?case proof (cases "x \ lvars (\ (l ! (i + 1)))") case True then have "j \ i + 1" using Suc(7) \\ (\ (l ! (i + 1)))\ by (auto simp add: normalized_tableau_def) have "k = j - Suc i" using Suc(2) by simp then obtain k where "k \ i + 1" "k < j" "swap_lr l k x" using \x \ lvars (\ (l ! (i + 1)))\ \j \ i + 1\ using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5) Suc(6) Suc(7) by auto then show ?thesis by (rule_tac x="k" in exI) simp next case False then have "x \ rvars (\ (l ! (i + 1)))" using Suc(6) using \l ! i \ (l ! (i + 1))\ succ_vars_id by auto then show ?thesis using Suc(4) Suc(5) Suc(6) by force qed qed (* -------------------------------------------------------------------------- *) lemma finite_tableaus_aux: shows "finite {t. lvars t = L \ rvars t = V - L \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t t0)}" (is "finite (?Al L)") proof (cases "?Al L = {}") case True show ?thesis by (subst True) simp next case False then have "\ t. t \ ?Al L" by auto let ?t = "SOME t. t \ ?Al L" have "?t \ ?Al L" using \\ t. t \ ?Al L\ by (rule someI_ex) have "?Al L \ {t. mset t = mset ?t}" proof fix x assume "x \ ?Al L" have "mset x = mset ?t" apply (rule tableau_perm) using \?t \ ?Al L\ \x \ ?Al L\ by auto then show "x \ {t. mset t = mset ?t}" by simp qed moreover have "finite {t. mset t = mset ?t}" by (fact mset_eq_finite) ultimately show ?thesis by (rule finite_subset) qed lemma finite_tableaus: assumes "finite V" shows "finite {t. tvars t = V \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t t0)}" (is "finite ?A") proof- let ?Al = "\ L. {t. lvars t = L \ rvars t = V - L \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t t0)}" have "?A = \ (?Al ` {L. L \ V})" by (auto simp add: normalized_tableau_def) then show ?thesis using \finite V\ using finite_tableaus_aux by auto qed lemma finite_accessible_tableaus: shows "finite (\ ` {s'. s \\<^sup>* s'})" proof- have "{s'. s \\<^sup>* s'} = {s'. s \\<^sup>+ s'} \ {s}" by (auto simp add: rtrancl_eq_or_trancl) moreover have "finite (\ ` {s'. s \\<^sup>+ s'})" (is "finite ?A") proof- let ?T = "{t. tvars t = tvars (\ s) \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t(\ s))}" have "?A \ ?T" proof fix t assume "t \ ?A" then obtain s' where "s \\<^sup>+ s'" "t = \ s'" by auto then obtain l where *: "l \ []" "1 < length l" "hd l = s" "last l = s'" "succ_chain l" using trancl_rel_chain[of s s' succ_rel] by auto show "t \ ?T" proof- have "tvars (\ s') = tvars (\ s)" using succ_chain_vars_id[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] by simp moreover have "\ (\ s')" using \s \\<^sup>+ s'\ using succ_inv(1)[of _ s'] by (auto dest: tranclD2) moreover have "\v::'a valuation. v \\<^sub>t \ s' = v \\<^sub>t \ s" using succ_chain_tableau_equiv[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] by auto ultimately show ?thesis using \t = \ s'\ by simp qed qed moreover have "finite (tvars (\ s))" by (auto simp add: lvars_def rvars_def finite_vars) ultimately show ?thesis using finite_tableaus[of "tvars (\ s)" "\ s"] by (auto simp add: finite_subset) qed ultimately show ?thesis by simp qed abbreviation check_valuation where "check_valuation (v::'a valuation) v0 bl0 bu0 t0 V \ \ t. tvars t = V \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t t0) \ v \\<^sub>t t \ (\ x \ rvars t. v x = v0 x \ v x = bl0 x \ v x = bu0 x) \ (\ x. x \ V \ v x = v0 x)" lemma finite_valuations: assumes "finite V" shows "finite {v::'a valuation. check_valuation v v0 bl0 bu0 t0 V}" (is "finite ?A") proof- let ?Al = "\ L. {t. lvars t = L \ rvars t = V - L \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t t0)}" let ?Vt = "\ t. {v::'a valuation. v \\<^sub>t t \ (\ x \ rvars t. v x = v0 x \ v x = bl0 x \ v x = bu0 x) \ (\ x. x \ V \ v x = v0 x)}" have "finite {L. L \ V}" using \finite V\ by auto have "\ L. L \ V \ finite (?Al L)" using finite_tableaus_aux by auto have "\ L t. L \ V \ t \ ?Al L \ finite (?Vt t)" proof (safe) fix L t assume "lvars t \ V" "rvars t = V - lvars t" "\ t" "\v. v \\<^sub>t t = v \\<^sub>t t0" then have "rvars t \ lvars t = V" by auto let ?f = "\ v x. if x \ rvars t then v x else 0" have "inj_on ?f (?Vt t)" unfolding inj_on_def proof (safe, rule ext) fix v1 v2 x assume "(\x. if x \ rvars t then v1 x else (0 :: 'a)) = (\x. if x \ rvars t then v2 x else (0 :: 'a))" (is "?f1 = ?f2") have "\x\rvars t. v1 x = v2 x" proof fix x assume "x \ rvars t" then show "v1 x = v2 x" using \?f1 = ?f2\ fun_cong[of ?f1 ?f2 x] by auto qed assume *: "v1 \\<^sub>t t" "v2 \\<^sub>t t" "\x. x \ V \ v1 x = v0 x" "\x. x \ V \ v2 x = v0 x" show "v1 x = v2 x" proof (cases "x \ lvars t") case False then show ?thesis using * \\x\rvars t. v1 x = v2 x\ \rvars t \ lvars t = V\ by auto next case True let ?eq = "eq_for_lvar t x" have "?eq \ set t \ lhs ?eq = x" using eq_for_lvar \x \ lvars t\ by simp then have "v1 x = rhs ?eq \ v1 \" "v2 x = rhs ?eq \ v2 \" using \v1 \\<^sub>t t\ \v2 \\<^sub>t t\ unfolding satisfies_tableau_def satisfies_eq_def by auto moreover have "rhs ?eq \ v1 \ = rhs ?eq \ v2 \" apply (rule valuate_depend) using \\x\rvars t. v1 x = v2 x\ \?eq \ set t \ lhs ?eq = x\ unfolding rvars_def by auto ultimately show ?thesis by simp qed qed let ?R = "{v. \ x. if x \ rvars t then v x = v0 x \ v x = bl0 x \ v x = bu0 x else v x = 0 }" have "?f ` (?Vt t) \ ?R" by auto moreover have "finite ?R" proof- have "finite (rvars t)" using \finite V\ \rvars t \ lvars t = V\ using finite_subset[of "rvars t" V] by auto moreover let ?R' = "{v. \ x. if x \ rvars t then v x \ {v0 x, bl0 x, bu0 x} else v x = 0}" have "?R = ?R'" by auto ultimately show ?thesis using finite_fun_args[of "rvars t" "\ x. {v0 x, bl0 x, bu0 x}" "\ x. 0"] by auto qed ultimately have "finite (?f ` (?Vt t))" by (simp add: finite_subset) then show "finite (?Vt t)" using \inj_on ?f (?Vt t)\ by (auto dest: finite_imageD) qed have "?A = \ (\ (((`) ?Vt) ` (?Al ` {L. L \ V})))" (is "?A = ?A'") by (auto simp add: normalized_tableau_def cong del: image_cong_simp) moreover have "finite ?A'" proof (rule finite_Union) show "finite (\ (((`) ?Vt) ` (?Al ` {L. L \ V})))" using \finite {L. L \ V}\ \\ L. L \ V \ finite (?Al L)\ by auto next fix M assume "M \ \ (((`) ?Vt) ` (?Al ` {L. L \ V}))" then obtain L t where "L \ V" "t \ ?Al L" "M = ?Vt t" by blast then show "finite M" using \\ L t. L \ V \ t \ ?Al L \ finite (?Vt t)\ by blast qed ultimately show ?thesis by simp qed lemma finite_accessible_valuations: shows "finite (\ ` {s'. s \\<^sup>* s'})" proof- have "{s'. s \\<^sup>* s'} = {s'. s \\<^sup>+ s'} \ {s}" by (auto simp add: rtrancl_eq_or_trancl) moreover have "finite (\ ` {s'. s \\<^sup>+ s'})" (is "finite ?A") proof- let ?P = "\ v. check_valuation v (\\ s\) (\ x. the (\\<^sub>l s x)) (\ x. the (\\<^sub>u s x)) (\ s) (tvars (\ s))" let ?P' = "\ v::(var, 'a) mapping. \ t. tvars t = tvars (\ s) \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t \ s) \ \v\ \\<^sub>t t \ (\ x \ rvars t. \v\ x = \\ s\ x \ \v\ x = the (\\<^sub>l s x) \ \v\ x = the (\\<^sub>u s x)) \ (\ x. x \ tvars (\ s) \ look v x = look (\ s) x) \ (\ x. x \ tvars (\ s) \ look v x \ None)" have "finite (tvars (\ s))" by (auto simp add: lvars_def rvars_def finite_vars) then have "finite {v. ?P v}" using finite_valuations[of "tvars (\ s)" "\ s" "\\ s\" "\ x. the (\\<^sub>l s x)" "\ x. the (\\<^sub>u s x)"] by auto moreover have "map2fun ` {v. ?P' v} \ {v. ?P v}" by (auto simp add: map2fun_def) ultimately have "finite (map2fun ` {v. ?P' v})" by (auto simp add: finite_subset) moreover have "inj_on map2fun {v. ?P' v}" unfolding inj_on_def proof (safe) fix x y assume "\x\ = \y\" and *: "\x. x \ Simplex.tvars (\ s) \ look y x = look (\ s) x" "\xa. xa \ Simplex.tvars (\ s) \ look x xa = look (\ s) xa" "\x. x \ Simplex.tvars (\ s) \ look y x \ None" "\xa. xa \ Simplex.tvars (\ s) \ look x xa \ None" show "x = y" proof (rule mapping_eqI) fix k have "\x\ k = \y\ k" using \\x\ = \y\\ by simp then show "look x k = look y k" using * by (cases "k \ tvars (\ s)") (auto simp add: map2fun_def split: option.split) qed qed ultimately have "finite {v. ?P' v}" by (rule finite_imageD) moreover have "?A \ {v. ?P' v}" proof (safe) fix s' assume "s \\<^sup>+ s'" then obtain l where *: "l \ []" "1 < length l" "hd l = s" "last l = s'" "succ_chain l" using trancl_rel_chain[of s s' succ_rel] by auto show "?P' (\ s')" proof- have "\ s" "\ (\ s)" "\\ s\ \\<^sub>t \ s" using \s \\<^sup>+ s'\ using tranclD[of s s' succ_rel] by (auto simp add: curr_val_satisfies_no_lhs_def) have "tvars (\ s') = tvars (\ s)" using succ_chain_vars_id[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] by simp moreover have "\(\ s')" using \s \\<^sup>+ s'\ using succ_inv(1)[of _ s'] by (auto dest: tranclD2) moreover have "\v::'a valuation. v \\<^sub>t \ s' = v \\<^sub>t \ s" using succ_chain_tableau_equiv[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] by auto moreover have "\\ s'\ \\<^sub>t \ s'" using succ_chain_valuation_satisfies[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] \\\ s\ \\<^sub>t \ s\ by simp moreover have "\x\rvars (\ s'). \\ s'\ x = \\ s\ x \ \\ s'\ x = the (\\<^sub>l s x) \ \\ s'\ x = the (\\<^sub>u s x)" using succ_chain_rvar_valuation[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] by auto moreover have "\x. x \ tvars (\ s) \ look (\ s') x = look (\ s) x" using succ_chain_no_vars_valuation[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] by auto moreover have "\x. x \ Simplex.tvars (\ s') \ look (\ s') x \ None" using succ_chain_tableau_valuated[of l 0 "length l - 1"] using * hd_conv_nth[of l] last_conv_nth[of l] using \tvars (\ s') = tvars (\ s)\ \\ s\ by (auto simp add: tableau_valuated_def) ultimately show ?thesis by (rule_tac x="\ s'" in exI) auto qed qed ultimately show ?thesis by (auto simp add: finite_subset) qed ultimately show ?thesis by simp qed lemma accessible_bounds: shows "\\<^sub>i ` {s'. s \\<^sup>* s'} = {\\<^sub>i s}" proof - have "s \\<^sup>* s' \ \\<^sub>i s' = \\<^sub>i s" for s' by (induct s s' rule: rtrancl.induct, auto) then show ?thesis by blast qed lemma accessible_unsat_core: shows "\\<^sub>c ` {s'. s \\<^sup>* s'} = {\\<^sub>c s}" proof - have "s \\<^sup>* s' \ \\<^sub>c s' = \\<^sub>c s" for s' by (induct s s' rule: rtrancl.induct, auto) then show ?thesis by blast qed lemma state_eqI: "\\<^sub>i\<^sub>l s = \\<^sub>i\<^sub>l s' \ \\<^sub>i\<^sub>u s = \\<^sub>i\<^sub>u s' \ \ s = \ s' \ \ s = \ s' \ \ s = \ s' \ \\<^sub>c s = \\<^sub>c s' \ s = s'" by (cases s, cases s', auto) lemma finite_accessible_states: shows "finite {s'. s \\<^sup>* s'}" (is "finite ?A") proof- let ?V = "\ ` ?A" let ?T = "\ ` ?A" let ?P = "?V \ ?T \ {\\<^sub>i s} \ {True, False} \ {\\<^sub>c s}" have "finite ?P" using finite_accessible_valuations finite_accessible_tableaus by auto moreover let ?f = "\ s. (\ s, \ s, \\<^sub>i s, \ s, \\<^sub>c s)" have "?f ` ?A \ ?P" using accessible_bounds[of s] accessible_unsat_core[of s] by auto moreover have "inj_on ?f ?A" unfolding inj_on_def by (auto intro: state_eqI) ultimately show ?thesis using finite_imageD [of ?f ?A] using finite_subset by auto qed (* -------------------------------------------------------------------------- *) lemma acyclic_suc_rel: "acyclic succ_rel" proof (rule acyclicI, rule allI) fix s show "(s, s) \ succ_rel\<^sup>+" proof assume "s \\<^sup>+ s" then obtain l where "l \ []" "length l > 1" "hd l = s" "last l = s" "succ_chain l" using trancl_rel_chain[of s s succ_rel] by auto have "l ! 0 = s" using \l \ []\ \hd l = s\ by (simp add: hd_conv_nth) then have "s \ (l ! 1)" using \succ_chain l\ unfolding rel_chain_def using \length l > 1\ by auto then have "\ (\ s)" by simp let ?enter_rvars = "{x. \ sl. swap_lr l sl x}" have "finite ?enter_rvars" proof- let ?all_vars = "\ (set (map (\ t. lvars t \ rvars t) (map \ l)))" have "finite ?all_vars" by (auto simp add: lvars_def rvars_def finite_vars) moreover have "?enter_rvars \ ?all_vars" by force ultimately show ?thesis by (simp add: finite_subset) qed let ?xr = "Max ?enter_rvars" have "?xr \ ?enter_rvars" proof (rule Max_in) show "?enter_rvars \ {}" proof- from \s \ (l ! 1)\ obtain x\<^sub>i x\<^sub>j :: var where "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ (l ! 1))" by (rule succ_vars) auto then have "x\<^sub>i \ ?enter_rvars" using \hd l = s\ \l \ []\ \length l > 1\ by (auto simp add: hd_conv_nth) then show ?thesis by auto qed next show "finite ?enter_rvars" using \finite ?enter_rvars\ . qed then obtain xr sl where "xr = ?xr" "swap_lr l sl xr" by auto then have "sl + 1 < length l" by simp have "(l ! sl) \ (l ! (sl + 1))" using \sl + 1 < length l\ \succ_chain l\ unfolding rel_chain_def by auto have "length l > 2" proof (rule ccontr) assume "\ ?thesis" with \length l > 1\ have "length l = 2" by auto then have "last l = l ! 1" by (cases l) (auto simp add: last_conv_nth nth_Cons split: nat.split) then have "xr \ lvars (\ s)" "xr \ rvars (\ s)" using \length l = 2\ using \swap_lr l sl xr\ using \hd l = s\ \last l = s\ \l \ []\ by (auto simp add: hd_conv_nth) then show False using \\ (\ s)\ unfolding normalized_tableau_def by auto qed obtain l' where "hd l' = l ! (sl + 1)" "last l' = l ! sl" "length l' = length l - 1" "succ_chain l'" and l'_l: "\ i. i + 1 < length l' \ (\ j. j + 1 < length l \ l' ! i = l ! j \ l' ! (i + 1) = l ! (j + 1))" using \length l > 2\ \sl + 1 < length l\ \hd l = s\ \last l = s\ \succ_chain l\ using reorder_cyclic_list[of l s sl] by blast then have "xr \ rvars (\ (hd l'))" "xr \ lvars (\ (last l'))" "length l' > 1" "l' \ []" using \swap_lr l sl xr\ \length l > 2\ by auto then have "\ sp. swap_rl l' sp xr" using \succ_chain l'\ using succ_chain_swap_rl_exists[of l' 0 "length l' - 1" xr] by (auto simp add: hd_conv_nth last_conv_nth) then have "\ sp. swap_rl l' sp xr \ (\ sp'. sp' < sp \ \ swap_rl l' sp' xr)" by (rule min_element) then obtain sp where "swap_rl l' sp xr" "\ sp'. sp' < sp \ \ swap_rl l' sp' xr" by blast then have "sp + 1 < length l'" by simp have "\\ (l' ! 0)\ xr = \\ (l' ! sp)\ xr" proof- have "always_r l' 0 sp xr" using \xr \ rvars (\ (hd l'))\ \sp + 1 < length l'\ \\ sp'. sp' < sp \ \ swap_rl l' sp' xr\ proof (induct sp) case 0 then have "l' \ []" by auto then show ?case using 0(1) by (auto simp add: hd_conv_nth) next case (Suc sp') show ?case proof (safe) fix k assume "k \ Suc sp'" show "xr \ rvars (\ (l' ! k))" proof (cases "k = sp' + 1") case False then show ?thesis using Suc \k \ Suc sp'\ by auto next case True then have "xr \ rvars (\ (l' ! (k - 1)))" using Suc by auto moreover then have "xr \ lvars (\ (l' ! k))" using True Suc(3) Suc(4) by auto moreover have "(l' ! (k - 1)) \ (l' ! k)" using \succ_chain l'\ using Suc(3) True by (simp add: rel_chain_def) ultimately show ?thesis using succ_vars_id[of "l' ! (k - 1)" "l' ! k"] by auto qed qed qed then show ?thesis using \sp + 1 < length l'\ using \succ_chain l'\ using succ_chain_always_r_valuation_id by simp qed have "(l' ! sp) \ (l' ! (sp+1))" using \sp + 1 < length l'\ \succ_chain l'\ unfolding rel_chain_def by simp then obtain xs xr' :: var where "xs \ lvars (\ (l' ! sp))" "xr \ rvars (\ (l' ! sp))" "swap_lr l' sp xs" apply (rule succ_vars) using \swap_rl l' sp xr\ \sp + 1 < length l'\ by auto then have "xs \ xr" using \(l' ! sp) \ (l' ! (sp+1))\ by (auto simp add: normalized_tableau_def) obtain sp' where "l' ! sp = l ! sp'" "l' ! (sp + 1) = l ! (sp' + 1)" "sp' + 1 < length l" using \sp + 1 < length l'\ l'_l by auto have "xs \ ?enter_rvars" using \swap_lr l' sp xs\ l'_l by force have "xs < xr" proof- have "xs \ ?xr" using \finite ?enter_rvars\ \xs \ ?enter_rvars\ by (rule Max_ge) then show ?thesis using \xr = ?xr\ \xs \ xr\ by simp qed let ?sl = "l ! sl" let ?sp = "l' ! sp" let ?eq = "eq_for_lvar (\ ?sp) xs" let ?bl = "\ ?sl" let ?bp = "\ ?sp" have "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sl" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sp" using \l ! sl \ (l ! (sl + 1))\ using \l' ! sp \ (l' ! (sp+ 1))\ by simp_all have "\\<^sub>i ?sp = \\<^sub>i ?sl" proof- have "\\<^sub>i (l' ! sp) = \\<^sub>i (l' ! (length l' - 1))" using \sp + 1 < length l'\ \succ_chain l'\ using succ_chain_bounds_id by auto then have "\\<^sub>i (last l') = \\<^sub>i (l' ! sp)" using \l' \ []\ by (simp add: last_conv_nth) then show ?thesis using \last l' = l ! sl\ by simp qed have diff_satified: "\?bl\ xs - \?bp\ xs = ((rhs ?eq) \ \?bl\ \) - ((rhs ?eq) \ \?bp\ \)" proof- have "\?bp\ \\<^sub>e ?eq" using \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sp\ using eq_for_lvar[of xs "\ ?sp"] using \xs \ lvars (\ (l' ! sp))\ unfolding curr_val_satisfies_no_lhs_def satisfies_tableau_def by auto moreover have "\?bl\ \\<^sub>e ?eq" proof- have "\\ (l ! sl)\ \\<^sub>t \ (l' ! sp)" using \l' ! sp = l ! sp'\ \sp' + 1 < length l\ \sl + 1 < length l\ using \succ_chain l\ using succ_chain_tableau_equiv[of l sl sp'] using \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sl\ unfolding curr_val_satisfies_no_lhs_def by simp then show ?thesis unfolding satisfies_tableau_def using eq_for_lvar using \xs \ lvars (\ (l' ! sp))\ by simp qed moreover have "lhs ?eq = xs" using \xs \ lvars (\ (l' ! sp))\ using eq_for_lvar by simp ultimately show ?thesis unfolding satisfies_eq_def by auto qed have "\ in_bounds xr \?bl\ (\ ?sl)" using \l ! sl \ (l ! (sl + 1))\ \swap_lr l sl xr\ using succ_min_lvar_not_in_bounds(1)[of ?sl "l ! (sl + 1)" xr] by simp have "\ x. x < xr \ in_bounds x \?bl\ (\ ?sl)" proof (safe) fix x assume "x < xr" show "in_bounds x \?bl\ (\ ?sl)" proof (cases "x \ lvars (\ ?sl)") case True then show ?thesis using succ_min_lvar_not_in_bounds(2)[of ?sl "l ! (sl + 1)" xr] using \l ! sl \ (l ! (sl + 1))\ \swap_lr l sl xr\ \x < xr\ by simp next case False then show ?thesis using \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sl\ unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_bounds_set.simps) qed qed then have "in_bounds xs \?bl\ (\ ?sl)" using \xs < xr\ by simp have "\ in_bounds xs \?bp\ (\ ?sp)" using \l' ! sp \ (l' ! (sp + 1))\ \swap_lr l' sp xs\ using succ_min_lvar_not_in_bounds(1)[of ?sp "l' ! (sp + 1)" xs] by simp have "\ x \ rvars_eq ?eq. x > xr \ \?bp\ x = \?bl\ x" proof (safe) fix x assume "x \ rvars_eq ?eq" "x > xr" then have "always_r l' 0 (length l' - 1) x" proof (safe) fix k assume "x \ rvars_eq ?eq" "x > xr" "0 \ k" "k \ length l' - 1" obtain k' where "l ! k' = l' ! k" "k' < length l" using l'_l \k \ length l' - 1\ \length l' > 1\ apply (cases "k > 0") apply (erule_tac x="k - 1" in allE) apply (drule mp) by auto let ?eq' = "eq_for_lvar (\ (l ! sp')) xs" have "\ x \ rvars_eq ?eq'. x > xr \ always_r l 0 (length l - 1) x" proof (safe) fix x k assume "x \ rvars_eq ?eq'" "xr < x" "0 \ k" "k \ length l - 1" then have "x \ rvars (\ (l ! sp'))" using eq_for_lvar[of xs "\ (l ! sp')"] using \swap_lr l' sp xs\ \l' ! sp = l ! sp'\ by (auto simp add: rvars_def) have *: "\ i. i < sp' \ x \ rvars (\ (l ! i))" proof (safe, rule ccontr) fix i assume "i < sp'" "x \ rvars (\ (l ! i))" then have "x \ lvars (\ (l ! i))" using \x \ rvars (\ (l ! sp'))\ using \sp' + 1 < length l\ using \succ_chain l\ using succ_chain_vars_id[of l i sp'] by auto obtain i' where "swap_lr l i' x" using \x \ lvars (\ (l ! i))\ using \x \ rvars (\ (l ! sp'))\ using \i < sp'\ \sp' + 1 < length l\ using \succ_chain l\ using succ_chain_swap_lr_exists[of l i sp' x] by auto then have "x \ ?enter_rvars" by auto then have "x \ ?xr" using \finite ?enter_rvars\ using Max_ge[of ?enter_rvars x] by simp then show False using \x > xr\ using \xr = ?xr\ by simp qed then have "x \ rvars (\ (last l))" using \hd l = s\ \last l = s\ \l \ []\ using \x \ rvars (\ (l ! sp'))\ by (auto simp add: hd_conv_nth) show "x \ rvars (\ (l ! k))" proof (cases "k = length l - 1") case True then show ?thesis using \x \ rvars (\ (last l))\ using \l \ []\ by (simp add: last_conv_nth) next case False then have "k < length l - 1" using \k \ length l - 1\ by simp then have "k < length l" using \length l > 1\ by auto show ?thesis proof (rule ccontr) assume "\ ?thesis" then have "x \ lvars (\ (l ! k))" using \x \ rvars (\ (l ! sp'))\ using \sp' + 1 < length l\ \k < length l\ using succ_chain_vars_id[of l k sp'] using \succ_chain l\ \l \ []\ by auto obtain i' where "swap_lr l i' x" using \succ_chain l\ using \x \ lvars (\ (l ! k))\ using \x \ rvars (\ (last l))\ using \k < length l - 1\ \l \ []\ using succ_chain_swap_lr_exists[of l k "length l - 1" x] by (auto simp add: last_conv_nth) then have "x \ ?enter_rvars" by auto then have "x \ ?xr" using \finite ?enter_rvars\ using Max_ge[of ?enter_rvars x] by simp then show False using \x > xr\ using \xr = ?xr\ by simp qed qed qed then have "x \ rvars (\ (l ! k'))" using \x \ rvars_eq ?eq\ \x > xr\ \k' < length l\ using \l' ! sp = l ! sp'\ by simp then show "x \ rvars (\ (l' ! k))" using \l ! k' = l' ! k\ by simp qed then have "\?bp\ x = \\ (l' ! (length l' - 1))\ x" using \succ_chain l'\ \sp + 1 < length l'\ by (auto intro!: succ_chain_always_r_valuation_id[rule_format]) then have "\?bp\ x = \\ (last l')\ x" using \l' \ []\ by (simp add: last_conv_nth) then show "\?bp\ x = \?bl\ x" using \last l' = l ! sl\ by simp qed have "\?bp\ xr = \\ (l ! (sl + 1))\ xr" using \\\ (l' ! 0)\ xr = \\ (l' ! sp)\ xr\ using \hd l' = l ! (sl + 1)\ \l' \ []\ by (simp add: hd_conv_nth) { fix dir1 dir2 :: "('i,'a) Direction" assume dir1: "dir1 = (if \?bl\ xr <\<^sub>l\<^sub>b \\<^sub>l ?sl xr then Positive else Negative)" then have "\\<^sub>l\<^sub>b (lt dir1) (\?bl\ xr) (LB dir1 ?sl xr)" using \\ in_bounds xr \?bl\ (\ ?sl)\ using neg_bounds_compare(7) neg_bounds_compare(3) by (auto simp add: bound_compare''_defs) then have "\ \\<^sub>l\<^sub>b (lt dir1) (\?bl\ xr) (LB dir1 ?sl xr)" using bounds_compare_contradictory(7) bounds_compare_contradictory(3) neg_bounds_compare(6) dir1 unfolding bound_compare''_defs by auto force have "LB dir1 ?sl xr \ None" using \\\<^sub>l\<^sub>b (lt dir1) (\?bl\ xr) (LB dir1 ?sl xr)\ by (cases "LB dir1 ?sl xr") (auto simp add: bound_compare_defs) assume dir2: "dir2 = (if \?bp\ xs <\<^sub>l\<^sub>b \\<^sub>l ?sp xs then Positive else Negative)" then have "\\<^sub>l\<^sub>b (lt dir2) (\?bp\ xs) (LB dir2 ?sp xs)" using \\ in_bounds xs \?bp\ (\ ?sp)\ using neg_bounds_compare(2) neg_bounds_compare(6) by (auto simp add: bound_compare''_defs) then have "\ \\<^sub>l\<^sub>b (lt dir2) (\?bp\ xs) (LB dir2 ?sp xs)" using bounds_compare_contradictory(3) bounds_compare_contradictory(7) neg_bounds_compare(6) dir2 unfolding bound_compare''_defs by auto force then have "\ x \ rvars_eq ?eq. x < xr \ \ reasable_var dir2 x ?eq ?sp" using succ_min_rvar[of ?sp "l' ! (sp + 1)" xs xr ?eq] using \l' ! sp \ (l' ! (sp + 1))\ using \swap_lr l' sp xs\ \swap_rl l' sp xr\ dir2 unfolding bound_compare''_defs by auto have "LB dir2 ?sp xs \ None" using \\\<^sub>l\<^sub>b (lt dir2) (\?bp\ xs) (LB dir2 ?sp xs)\ by (cases "LB dir2 ?sp xs") (auto simp add: bound_compare_defs) have *: "\ x \ rvars_eq ?eq. x < xr \ ((coeff (rhs ?eq) x > 0 \ \\<^sub>u\<^sub>b (lt dir2) (\?bp\ x) (UB dir2 ?sp x)) \ (coeff (rhs ?eq) x < 0 \ \\<^sub>l\<^sub>b (lt dir2) (\?bp\ x) (LB dir2 ?sp x)))" proof (safe) fix x assume "x \ rvars_eq ?eq" "x < xr" "coeff (rhs ?eq) x > 0" then have "\ \\<^sub>u\<^sub>b (lt dir2) (\?bp\ x) (UB dir2 ?sp x)" using \\ x \ rvars_eq ?eq. x < xr \ \ reasable_var dir2 x ?eq ?sp\ by simp then show "\\<^sub>u\<^sub>b (lt dir2) (\?bp\ x) (UB dir2 ?sp x)" using dir2 neg_bounds_compare(4) neg_bounds_compare(8) unfolding bound_compare''_defs by force next fix x assume "x \ rvars_eq ?eq" "x < xr" "coeff (rhs ?eq) x < 0" then have "\ \\<^sub>l\<^sub>b (lt dir2) (\?bp\ x) (LB dir2 ?sp x)" using \\ x \ rvars_eq ?eq. x < xr \ \ reasable_var dir2 x ?eq ?sp\ by simp then show "\\<^sub>l\<^sub>b (lt dir2) (\?bp\ x) (LB dir2 ?sp x)" using dir2 neg_bounds_compare(4) neg_bounds_compare(8) dir2 unfolding bound_compare''_defs by force qed have "(lt dir2) (\?bp\ xs) (\?bl\ xs)" using \\\<^sub>l\<^sub>b (lt dir2) (\?bp\ xs) (LB dir2 ?sp xs)\ using \\\<^sub>i ?sp = \\<^sub>i ?sl\ dir2 using \in_bounds xs \?bl\ (\ ?sl)\ by (auto simp add: bound_compare''_defs simp: indexl_def indexu_def boundsl_def boundsu_def) then have "(lt dir2) 0 (\?bl\ xs - \?bp\ xs)" using dir2 by (auto simp add: minus_gt[THEN sym] minus_lt[THEN sym]) moreover have "le (lt dir2) ((rhs ?eq) \ \?bl\ \ - (rhs ?eq) \ \?bp\ \) 0" proof- have "\ x \ rvars_eq ?eq. (0 < coeff (rhs ?eq) x \ le (lt dir2) 0 (\?bp\ x - \?bl\ x)) \ (coeff (rhs ?eq) x < 0 \ le (lt dir2) (\?bp\ x - \?bl\ x) 0)" proof fix x assume "x \ rvars_eq ?eq" show "(0 < coeff (rhs ?eq) x \ le (lt dir2) 0 (\?bp\ x - \?bl\ x)) \ (coeff (rhs ?eq) x < 0 \ le (lt dir2) (\?bp\ x - \?bl\ x) 0)" proof (cases "x < xr") case True then have "in_bounds x \?bl\ (\ ?sl)" using \\ x. x < xr \ in_bounds x \?bl\ (\ ?sl)\ by simp show ?thesis proof (safe) assume "coeff (rhs ?eq) x > 0" "0 \ \?bp\ x - \?bl\ x" then have "\\<^sub>u\<^sub>b (lt dir2) (\\ (l' ! sp)\ x) (UB dir2 (l' ! sp) x)" using * \x < xr\ \x \ rvars_eq ?eq\ by simp then have "le (lt dir2) (\?bl\ x) (\?bp\ x)" using \in_bounds x \?bl\ (\ ?sl)\ \\\<^sub>i ?sp = \\<^sub>i ?sl\ dir2 apply (auto simp add: bound_compare''_defs) using bounds_lg(3)[of "\?bp\ x" "\\<^sub>u (l ! sl) x" "\?bl\ x"] using bounds_lg(6)[of "\?bp\ x" "\\<^sub>l (l ! sl) x" "\?bl\ x"] unfolding bound_compare''_defs by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) then show "lt dir2 0 (\?bp\ x - \?bl\ x)" using \0 \ \?bp\ x - \?bl\ x\ using minus_gt[of "\?bl\ x" "\?bp\ x"] minus_lt[of "\?bp\ x" "\?bl\ x"] dir2 - by auto + by (auto simp del: Simplex.bounds_lg) next assume "coeff (rhs ?eq) x < 0" "\?bp\ x - \?bl\ x \ 0" then have "\\<^sub>l\<^sub>b (lt dir2) (\\ (l' ! sp)\ x) (LB dir2 (l' ! sp) x)" using * \x < xr\ \x \ rvars_eq ?eq\ by simp then have "le (lt dir2) (\?bp\ x) (\?bl\ x)" using \in_bounds x \?bl\ (\ ?sl)\ \\\<^sub>i ?sp = \\<^sub>i ?sl\ dir2 apply (auto simp add: bound_compare''_defs) using bounds_lg(3)[of "\?bp\ x" "\\<^sub>u (l ! sl) x" "\?bl\ x"] using bounds_lg(6)[of "\?bp\ x" "\\<^sub>l (l ! sl) x" "\?bl\ x"] unfolding bound_compare''_defs by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) then show "lt dir2 (\?bp\ x - \?bl\ x) 0" using \\?bp\ x - \?bl\ x \ 0\ using minus_gt[of "\?bl\ x" "\?bp\ x"] minus_lt[of "\?bp\ x" "\?bl\ x"] dir2 - by auto + by (auto simp del: Simplex.bounds_lg) qed next case False show ?thesis proof (cases "x = xr") case True have "\\ (l ! (sl + 1))\ xr = the (LB dir1 ?sl xr)" using \l ! sl \ (l ! (sl + 1))\ using \swap_lr l sl xr\ using succ_set_on_bound(1)[of "l ! sl" "l ! (sl + 1)" xr] using \\ \\<^sub>l\<^sub>b (lt dir1) (\?bl\ xr) (LB dir1 ?sl xr)\ dir1 unfolding bound_compare''_defs by auto then have "\?bp\ xr = the (LB dir1 ?sl xr)" using \\?bp\ xr = \\ (l ! (sl + 1))\ xr\ by simp then have "lt dir1 (\?bl\ xr) (\?bp\ xr)" using \LB dir1 ?sl xr \ None\ using \\\<^sub>l\<^sub>b (lt dir1) (\?bl\ xr) (LB dir1 ?sl xr)\ dir1 by (auto simp add: bound_compare_defs) moreover have "reasable_var dir2 xr ?eq ?sp" using \\ \\<^sub>l\<^sub>b (lt dir2) (\?bp\ xs) (LB dir2 ?sp xs)\ using \l' ! sp \ (l' ! (sp + 1))\ using \swap_lr l' sp xs\ \swap_rl l' sp xr\ using succ_min_rvar[of "l' ! sp" "l' ! (sp + 1)"xs xr ?eq] dir2 unfolding bound_compare''_defs by auto then have "if dir1 = dir2 then coeff (rhs ?eq) xr > 0 else coeff (rhs ?eq) xr < 0" using \\?bp\ xr = the (LB dir1 ?sl xr)\ using \\\<^sub>i ?sp = \\<^sub>i ?sl\[THEN sym] dir1 using \LB dir1 ?sl xr \ None\ dir1 dir2 by (auto split: if_splits simp add: bound_compare_defs indexl_def indexu_def boundsl_def boundsu_def) moreover have "dir1 = Positive \ dir1 = Negative" "dir2 = Positive \ dir2 = Negative" using dir1 dir2 by auto ultimately show ?thesis using \x = xr\ using minus_lt[of "\?bp\ xr" "\?bl\ xr"] minus_gt[of "\?bl\ xr" "\?bp\ xr"] - by (auto split: if_splits) + by (auto split: if_splits simp del: Simplex.bounds_lg) next case False then have "x > xr" using \\ x < xr\ by simp then have "\?bp\ x = \?bl\ x" using \\ x \ rvars_eq ?eq. x > xr \ \?bp\ x = \?bl\ x\ using \x \ rvars_eq ?eq\ by simp then show ?thesis by simp qed qed qed then have "le (lt dir2) 0 (rhs ?eq \ \ x. \?bp\ x - \?bl\ x \)" using dir2 apply auto using valuate_nonneg[of "rhs ?eq" "\ x. \?bp\ x - \?bl\ x"] - apply force + apply (force simp del: Simplex.bounds_lg) using valuate_nonpos[of "rhs ?eq" "\ x. \?bp\ x - \?bl\ x"] - apply force + apply (force simp del: Simplex.bounds_lg) done then have "le (lt dir2) 0 ((rhs ?eq) \ \?bp\ \ - (rhs ?eq) \ \?bl\ \)" by (subst valuate_diff)+ simp then have "le (lt dir2) ((rhs ?eq) \ \?bl\ \) ((rhs ?eq) \ \?bp\ \)" using minus_lt[of "(rhs ?eq) \ \?bp\ \" "(rhs ?eq) \ \?bl\ \"] dir2 - by auto + by (auto simp del: Simplex.bounds_lg) then show ?thesis using dir2 using minus_lt[of "(rhs ?eq) \ \?bl\ \" "(rhs ?eq) \ \?bp\ \"] using minus_gt[of "(rhs ?eq) \ \?bp\ \" "(rhs ?eq) \ \?bl\ \"] - by auto + by (auto simp del: Simplex.bounds_lg) qed ultimately have False using diff_satified dir2 - by (auto split: if_splits) + by (auto split: if_splits simp del: Simplex.bounds_lg) } then show False by auto qed qed (* -------------------------------------------------------------------------- *) lemma check_unsat_terminates: assumes "\ s" shows "check_dom s" by (rule check_dom.intros) (auto simp add: assms) lemma check_sat_terminates'_aux: assumes dir: "dir = (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative)" and *: "\ s'. \s \ s'; \ s'; \ (\ s'); \ s'; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ check_dom s'" and "\ s" "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ \ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "\\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)" shows "check_dom (case min_rvar_incdec dir s x\<^sub>i of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j (the (LB dir s x\<^sub>i)) s)" proof (cases "min_rvar_incdec dir s x\<^sub>i") case Inl then show ?thesis using check_unsat_terminates by simp next case (Inr x\<^sub>j) then have xj: "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" using min_rvar_incdec_eq_Some_rvars[of _ s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using dir by simp let ?s' = "pivot_and_update x\<^sub>i x\<^sub>j (the (LB dir s x\<^sub>i)) s" have "check_dom ?s'" proof (rule * ) show **: "\ ?s'" "\ (\ ?s')" "\ ?s'" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s'" using \min_lvar_not_in_bounds s = Some x\<^sub>i\ Inr using \\ s\ \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ dir using pivotandupdate_check_precond by auto have xi: "x\<^sub>i \ lvars (\ s)" using assms(8) min_lvar_not_in_bounds_lvars by blast show "s \ ?s'" unfolding gt_state_def using \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ \\ s\ using \min_lvar_not_in_bounds s = Some x\<^sub>i\ \\\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)\ Inr dir by (intro conjI pivotandupdate_bounds_id pivotandupdate_unsat_core_id, auto intro!: xj xi) qed then show ?thesis using Inr by simp qed lemma check_sat_terminates': assumes "\ s" "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "s\<^sub>0 \\<^sup>* s" shows "check_dom s" using assms proof (induct s rule: wf_induct[of "{(y, x). s\<^sub>0 \\<^sup>* x \ x \ y}"]) show "wf {(y, x). s\<^sub>0 \\<^sup>* x \ x \ y}" proof (rule finite_acyclic_wf) let ?A = "{(s', s). s\<^sub>0 \\<^sup>* s \ s \ s'}" let ?B = "{s. s\<^sub>0 \\<^sup>* s}" have "?A \ ?B \ ?B" proof fix p assume "p \ ?A" then have "fst p \ ?B" "snd p \ ?B" using rtrancl_into_trancl1[of s\<^sub>0 "snd p" succ_rel "fst p"] by auto then show "p \ ?B \ ?B" using mem_Sigma_iff[of "fst p" "snd p"] by auto qed then show "finite ?A" using finite_accessible_states[of s\<^sub>0] using finite_subset[of ?A "?B \ ?B"] by simp show "acyclic ?A" proof- have "?A \ succ_rel\" by auto then show ?thesis using acyclic_converse acyclic_subset using acyclic_suc_rel by auto qed qed next fix s assume "\ s'. (s', s) \ {(y, x). s\<^sub>0 \\<^sup>* x \ x \ y} \ \ s' \ \ (\ s') \ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ s\<^sub>0 \\<^sup>* s' \ check_dom s'" "\ s" "\ (\ s)" "\ s" " \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "s\<^sub>0 \\<^sup>* s" then have *: "\ s'. \s \ s'; \ s'; \ (\ s'); \ s'; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ check_dom s'" using rtrancl_into_trancl1[of s\<^sub>0 s succ_rel] using trancl_into_rtrancl[of s\<^sub>0 _ succ_rel] by auto show "check_dom s" proof (rule check_dom.intros, simp_all add: check'_def, unfold Positive_def[symmetric], unfold Negative_def[symmetric]) fix x\<^sub>i assume "\ \ s" "Some x\<^sub>i = min_lvar_not_in_bounds s" "\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i" have "\\<^sub>l s x\<^sub>i = LB Positive s x\<^sub>i" by simp show "check_dom (case min_rvar_incdec Positive s x\<^sub>i of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j (the (\\<^sub>l s x\<^sub>i)) s)" apply (subst \\\<^sub>l s x\<^sub>i = LB Positive s x\<^sub>i\) apply (rule check_sat_terminates'_aux[of Positive s x\<^sub>i]) using \\ s\ \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ * using \\ \ s\ \Some x\<^sub>i = min_lvar_not_in_bounds s\ \\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ by (simp_all add: bound_compare''_defs) next fix x\<^sub>i assume "\ \ s" "Some x\<^sub>i = min_lvar_not_in_bounds s" "\ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i" then have "\\ s\ x\<^sub>i >\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i" using min_lvar_not_in_bounds_Some[of s x\<^sub>i] using neg_bounds_compare(7) neg_bounds_compare(2) by auto have "\\<^sub>u s x\<^sub>i = LB Negative s x\<^sub>i" by simp show "check_dom (case min_rvar_incdec Negative s x\<^sub>i of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j (the (\\<^sub>u s x\<^sub>i)) s)" apply (subst \\\<^sub>u s x\<^sub>i = LB Negative s x\<^sub>i\) apply (rule check_sat_terminates'_aux) using \\ s\ \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ * using \\ \ s\ \Some x\<^sub>i = min_lvar_not_in_bounds s\ \\\ s\ x\<^sub>i >\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i\ \\ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ by (simp_all add: bound_compare''_defs) qed qed lemma check_sat_terminates: assumes "\ s" "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" shows "check_dom s" using assms using check_sat_terminates'[of s s] by simp lemma check_cases: assumes "\ s \ P s" assumes "\\ \ s; min_lvar_not_in_bounds s = None\ \ P s" assumes "\ x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I\ \ P (set_unsat I s)" assumes "\ x\<^sub>i x\<^sub>j l\<^sub>i dir. \dir = (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative); \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j; l\<^sub>i = the (LB dir s x\<^sub>i); check' dir x\<^sub>i s = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s\ \ P (check (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s))" assumes "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" shows "P (check s)" proof (cases "\ s") case True then show ?thesis using assms(1) using check.simps[of s] by simp next case False show ?thesis proof (cases "min_lvar_not_in_bounds s") case None then show ?thesis using \\ \ s\ using assms(2) \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ using check.simps[of s] by simp next case (Some x\<^sub>i) let ?dir = "if (\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i) then (Positive :: ('i,'a)Direction) else Negative" let ?s' = "check' ?dir x\<^sub>i s" have "\\<^sub>l\<^sub>b (lt ?dir) (\\ s\ x\<^sub>i) (LB ?dir s x\<^sub>i)" using \min_lvar_not_in_bounds s = Some x\<^sub>i\ using min_lvar_not_in_bounds_Some[of s x\<^sub>i] using not_in_bounds[of x\<^sub>i "\\ s\" "\\<^sub>l s" "\\<^sub>u s"] by (auto split: if_splits simp add: bound_compare''_defs) have "P (check ?s')" apply (rule check'_cases) using \\ \ s\ \min_lvar_not_in_bounds s = Some x\<^sub>i\ \\\<^sub>l\<^sub>b (lt ?dir) (\\ s\ x\<^sub>i) (LB ?dir s x\<^sub>i)\ using assms(3)[of ?dir x\<^sub>i] using assms(4)[of ?dir x\<^sub>i] using check.simps[of "set_unsat (_ :: 'i list) s"] using \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ by (auto simp add: bounds_consistent_def curr_val_satisfies_no_lhs_def) then show ?thesis using \\ \ s\ \min_lvar_not_in_bounds s = Some x\<^sub>i\ using check.simps[of s] using \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ by auto qed qed lemma check_induct: fixes s :: "('i,'a) state" assumes *: "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" assumes **: "\ s. \ s \ P s s" "\ s. \\ \ s; min_lvar_not_in_bounds s = None\ \ P s s" "\ s x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I\ \ P s (set_unsat I s)" assumes step': "\ s x\<^sub>i x\<^sub>j l\<^sub>i. \\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ \ P s (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)" assumes trans': "\ si sj sk. \P si sj; P sj sk\ \ P si sk" shows "P s (check s)" proof- have "check_dom s" using * by (simp add: check_sat_terminates) then show ?thesis using * proof (induct s rule: check_dom.induct) case (step s') show ?case proof (rule check_cases) fix x\<^sub>i x\<^sub>j l\<^sub>i dir let ?dir = "if \\ s'\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s' x\<^sub>i then Positive else Negative" let ?s' = "check' dir x\<^sub>i s'" assume "\ \ s'" "min_lvar_not_in_bounds s' = Some x\<^sub>i" "min_rvar_incdec dir s' x\<^sub>i = Inr x\<^sub>j" "l\<^sub>i = the (LB dir s' x\<^sub>i)" "?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'" "dir = ?dir" moreover then have "\ ?s'" "\ (\ ?s')" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s'" "\ ?s'" using \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ using pivotandupdate_check_precond[of dir s' x\<^sub>i x\<^sub>j l\<^sub>i] by auto ultimately have "P (check' dir x\<^sub>i s') (check (check' dir x\<^sub>i s'))" using step(2)[of x\<^sub>i] step(4)[of x\<^sub>i] \\ (\ s')\ \\ s'\ by auto then show "P s' (check (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'))" using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ \\ (\ s')\ \\ s'\ using \min_lvar_not_in_bounds s' = Some x\<^sub>i\ \min_rvar_incdec dir s' x\<^sub>i = Inr x\<^sub>j\ using step'[of s' x\<^sub>i x\<^sub>j l\<^sub>i] using trans'[of s' ?s' "check ?s'"] by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars) qed (simp_all add: \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ **) qed qed lemma check_induct': fixes s :: "('i,'a) state" assumes "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" assumes "\ s x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I; P s\ \ P (set_unsat I s)" assumes "\ s x\<^sub>i x\<^sub>j l\<^sub>i. \\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i); P s\ \ P (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)" assumes "P s" shows "P (check s)" proof- have "P s \ P (check s)" by (rule check_induct) (simp_all add: assms) then show ?thesis using \P s\ by simp qed lemma check_induct'': fixes s :: "('i,'a) state" assumes *: "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" assumes **: "\ s \ P s" "\ s. \\ s; \ (\ s); \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ \ s; min_lvar_not_in_bounds s = None\ \ P s" "\ s x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ s; \ (\ s); \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I\ \ P (set_unsat I s)" shows "P (check s)" proof (cases "\ s") case True then show ?thesis using \\ s \ P s\ by (simp add: check.simps) next case False have "check_dom s" using * by (simp add: check_sat_terminates) then show ?thesis using * False proof (induct s rule: check_dom.induct) case (step s') show ?case proof (rule check_cases) fix x\<^sub>i x\<^sub>j l\<^sub>i dir let ?dir = "if \\ s'\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s' x\<^sub>i then Positive else Negative" let ?s' = "check' dir x\<^sub>i s'" assume "\ \ s'" "min_lvar_not_in_bounds s' = Some x\<^sub>i" "min_rvar_incdec dir s' x\<^sub>i = Inr x\<^sub>j" "l\<^sub>i = the (LB dir s' x\<^sub>i)" "?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'" "dir = ?dir" moreover then have "\ ?s'" "\ (\ ?s')" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s'" "\ ?s'" "\ \ ?s'" using \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ using pivotandupdate_check_precond[of dir s' x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_unsat_id[of s' x\<^sub>i x\<^sub>j l\<^sub>i] by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars) ultimately have "P (check (check' dir x\<^sub>i s'))" using step(2)[of x\<^sub>i] step(4)[of x\<^sub>i] \\ (\ s')\ \\ s'\ by auto then show "P (check (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'))" using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ by simp qed (simp_all add: \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ \\ \ s'\ ** ) qed qed end lemma poly_eval_update: "(p \ v ( x := c :: 'a :: lrv) \) = (p \ v \) + coeff p x *R (c - v x)" proof (transfer, simp, goal_cases) case (1 p v x c) hence fin: "finite {v. p v \ 0}" by simp have "(\y\{v. p v \ 0}. p y *R (if y = x then c else v y)) = (\y\{v. p v \ 0} \ {x}. p y *R (if y = x then c else v y)) + (\y\{v. p v \ 0} \ (UNIV - {x}). p y *R (if y = x then c else v y))" (is "?l = ?a + ?b") by (subst sum.union_disjoint[symmetric], auto intro: sum.cong fin) also have "?a = (if p x = 0 then 0 else p x *R c)" by auto also have "\ = p x *R c" by auto also have "?b = (\y\{v. p v \ 0} \ (UNIV - {x}). p y *R v y)" (is "_ = ?c") by (rule sum.cong, auto) finally have l: "?l = p x *R c + ?c" . define r where "r = (\y\{v. p v \ 0}. p y *R v y) + p x *R (c - v x)" have "r = (\y\{v. p v \ 0}. p y *R v y) + p x *R (c - v x)" by (simp add: r_def) also have "(\y\{v. p v \ 0}. p y *R v y) = (\y\{v. p v \ 0} \ {x}. p y *R v y) + ?c" (is "_ = ?d + _") by (subst sum.union_disjoint[symmetric], auto intro: sum.cong fin) also have "?d = (if p x = 0 then 0 else p x *R v x)" by auto also have "\ = p x *R v x" by auto finally have "(p x *R (c - v x) + p x *R v x) + ?c = r" by simp also have "(p x *R (c - v x) + p x *R v x) = p x *R c" unfolding scaleRat_right_distrib[symmetric] by simp finally have r: "p x *R c + ?c = r" . show ?case unfolding l r r_def .. qed context PivotUpdateMinVars begin context fixes rhs_eq_val :: "(var, 'a::lrv) mapping \ var \ 'a \ eq \ 'a" assumes "RhsEqVal rhs_eq_val" begin lemma check_minimal_unsat_state_core: assumes *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" "\ (\ s)" "\ s" shows "\ (check s) \ minimal_unsat_state_core (check s)" (is "?P (check s)") proof (rule check_induct'') fix s' :: "('i,'a) state" and x\<^sub>i dir I assume nolhs: "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" and min_rvar: "min_rvar_incdec dir s' x\<^sub>i = Inl I" and sat: "\ \ s'" and min_lvar: "min_lvar_not_in_bounds s' = Some x\<^sub>i" and dir: "dir = Positive \ dir = Negative" and lt: "\\<^sub>l\<^sub>b (lt dir) (\\ s'\ x\<^sub>i) (LB dir s' x\<^sub>i)" and norm: "\ (\ s')" and valuated: "\ s'" let ?eq = "eq_for_lvar (\ s') x\<^sub>i" have unsat_core: "set (the (\\<^sub>c (set_unsat I s'))) = set I" by auto obtain l\<^sub>i where LB_Some: "LB dir s' x\<^sub>i = Some l\<^sub>i" and lt: "lt dir (\\ s'\ x\<^sub>i) l\<^sub>i" using lt by (cases "LB dir s' x\<^sub>i") (auto simp add: bound_compare_defs) from LB_Some dir obtain i where LBI: "look (LBI dir s') x\<^sub>i = Some (i,l\<^sub>i)" and LI: "LI dir s' x\<^sub>i = i" by (auto simp: simp: indexl_def indexu_def boundsl_def boundsu_def) from min_rvar_incdec_eq_None[OF min_rvar] dir have Is': "LI dir s' (lhs (eq_for_lvar (\ s') x\<^sub>i)) \ indices_state s' \ set I \ indices_state s'" and reasable: "\ x. x \ rvars_eq ?eq \ \ reasable_var dir x ?eq s'" and setI: "set I = {LI dir s' (lhs ?eq)} \ {LI dir s' x |x. x \ rvars_eq ?eq \ coeff (rhs ?eq) x < 0} \ {UI dir s' x |x. x \ rvars_eq ?eq \ 0 < coeff (rhs ?eq) x}" (is "_ = ?L \ ?R1 \ ?R2") by auto note setI also have id: "lhs ?eq = x\<^sub>i" by (simp add: EqForLVar.eq_for_lvar EqForLVar_axioms min_lvar min_lvar_not_in_bounds_lvars) finally have iI: "i \ set I" unfolding LI by auto note setI = setI[unfolded id] have "LI dir s' x\<^sub>i \ indices_state s'" using LBI LI unfolding indices_state_def using dir by force from Is'[unfolded id, OF this] have Is': "set I \ indices_state s'" . have "x\<^sub>i \ lvars (\ s')" using min_lvar by (simp add: min_lvar_not_in_bounds_lvars) then have **: "?eq \ set (\ s')" "lhs ?eq = x\<^sub>i" by (auto simp add: eq_for_lvar) have Is': "set I \ indices_state (set_unsat I s')" using Is' * unfolding indices_state_def by auto have "\\ s'\ \\<^sub>t \ s'" and b: "\\ s'\ \\<^sub>b \ s' \ - lvars (\ s')" using nolhs[unfolded curr_val_satisfies_no_lhs_def] by auto from norm[unfolded normalized_tableau_def] have lvars_rvars: "lvars (\ s') \ rvars (\ s') = {}" by auto hence in_bnds: "x \ rvars (\ s') \ in_bounds x \\ s'\ (\ s')" for x by (intro b[unfolded satisfies_bounds_set.simps, rule_format, of x], auto) { assume dist: "distinct_indices_state (set_unsat I s')" hence "distinct_indices_state s'" unfolding distinct_indices_state_def by auto note dist = this[unfolded distinct_indices_state_def, rule_format] { fix x c i y assume c: "look (\\<^sub>i\<^sub>l s') x = Some (i,c) \ look (\\<^sub>i\<^sub>u s') x = Some (i,c)" and y: "y \ rvars_eq ?eq" and coeff: "coeff (rhs ?eq) y < 0 \ i = LI dir s' y \ coeff (rhs ?eq) y > 0 \ i = UI dir s' y" { assume coeff: "coeff (rhs ?eq) y < 0" and i: "i = LI dir s' y" from reasable[OF y] coeff have not_gt: "\ (\\<^sub>l\<^sub>b (lt dir) (\\ s'\ y) (LB dir s' y))" by auto then obtain d where LB: "LB dir s' y = Some d" using dir by (cases "LB dir s' y", auto simp: bound_compare_defs) with not_gt have le: "le (lt dir) (\\ s'\ y) d" using dir by (auto simp: bound_compare_defs) from LB have "look (LBI dir s') y = Some (i, d)" unfolding i using dir by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) with c dist[of x i c y d] dir have yx: "y = x" "d = c" by auto from y[unfolded yx] have "x \ rvars (\ s')" using **(1) unfolding rvars_def by force - from in_bnds[OF this] le LB not_gt i have "\\ s'\ x = c" unfolding yx using dir by auto + from in_bnds[OF this] le LB not_gt i have "\\ s'\ x = c" unfolding yx using dir + by (auto simp del: Simplex.bounds_lg) note yx(1) this } moreover { assume coeff: "coeff (rhs ?eq) y > 0" and i: "i = UI dir s' y" from reasable[OF y] coeff have not_gt: "\ (\\<^sub>u\<^sub>b (lt dir) (\\ s'\ y) (UB dir s' y))" by auto then obtain d where UB: "UB dir s' y = Some d" using dir by (cases "UB dir s' y", auto simp: bound_compare_defs) with not_gt have le: "le (lt dir) d (\\ s'\ y)" using dir by (auto simp: bound_compare_defs) from UB have "look (UBI dir s') y = Some (i, d)" unfolding i using dir by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) with c dist[of x i c y d] dir have yx: "y = x" "d = c" by auto from y[unfolded yx] have "x \ rvars (\ s')" using **(1) unfolding rvars_def by force - from in_bnds[OF this] le UB not_gt i have "\\ s'\ x = c" unfolding yx using dir by auto + from in_bnds[OF this] le UB not_gt i have "\\ s'\ x = c" unfolding yx using dir + by (auto simp del: Simplex.bounds_lg) note yx(1) this } ultimately have "y = x" "\\ s'\ x = c" using coeff by blast+ } note x_vars_main = this { fix x c i assume c: "look (\\<^sub>i\<^sub>l s') x = Some (i,c) \ look (\\<^sub>i\<^sub>u s') x = Some (i,c)" and i: "i \ ?R1 \ ?R2" from i obtain y where y: "y \ rvars_eq ?eq" and coeff: "coeff (rhs ?eq) y < 0 \ i = LI dir s' y \ coeff (rhs ?eq) y > 0 \ i = UI dir s' y" by auto from x_vars_main[OF c y coeff] have "y = x" "\\ s'\ x = c" using coeff by blast+ with y have "x \ rvars_eq ?eq" "x \ rvars (\ s')" "\\ s'\ x = c" using **(1) unfolding rvars_def by force+ } note x_rvars = this have R1R2: "(?R1 \ ?R2, \\ s'\) \\<^sub>i\<^sub>s\<^sub>e s'" unfolding satisfies_state_index'.simps proof (intro conjI) show "\\ s'\ \\<^sub>t \ s'" by fact show "(?R1 \ ?R2, \\ s'\) \\<^sub>i\<^sub>b\<^sub>e \\ s'" unfolding satisfies_bounds_index'.simps proof (intro conjI impI allI) fix x c assume c: "\\<^sub>l s' x = Some c" and i: "\\<^sub>l s' x \ ?R1 \ ?R2" from c have ci: "look (\\<^sub>i\<^sub>l s') x = Some (\\<^sub>l s' x, c)" unfolding boundsl_def indexl_def by auto from x_rvars[OF _ i] ci show "\\ s'\ x = c" by auto next fix x c assume c: "\\<^sub>u s' x = Some c" and i: "\\<^sub>u s' x \ ?R1 \ ?R2" from c have ci: "look (\\<^sub>i\<^sub>u s') x = Some (\\<^sub>u s' x, c)" unfolding boundsu_def indexu_def by auto from x_rvars[OF _ i] ci show "\\ s'\ x = c" by auto qed qed have id1: "set (the (\\<^sub>c (set_unsat I s'))) = set I" "\ x. x \\<^sub>i\<^sub>s\<^sub>e set_unsat I s' \ x \\<^sub>i\<^sub>s\<^sub>e s'" by (auto simp: satisfies_state_index'.simps boundsl_def boundsu_def indexl_def indexu_def) have "subsets_sat_core (set_unsat I s')" unfolding subsets_sat_core_def id1 proof (intro allI impI) fix J assume sub: "J \ set I" show "\v. (J, v) \\<^sub>i\<^sub>s\<^sub>e s'" proof (cases "J \ ?R1 \ ?R2") case True with R1R2 have "(J, \\ s'\) \\<^sub>i\<^sub>s\<^sub>e s'" unfolding satisfies_state_index'.simps satisfies_bounds_index'.simps by blast thus ?thesis by blast next case False with sub obtain k where k: "k \ ?R1 \ ?R2" "k \ J" "k \ set I" unfolding setI by auto from k(1) obtain y where y: "y \ rvars_eq ?eq" and coeff: "coeff (rhs ?eq) y < 0 \ k = LI dir s' y \ coeff (rhs ?eq) y > 0 \ k = UI dir s' y" by auto hence cy0: "coeff (rhs ?eq) y \ 0" by auto from y **(1) have ry: "y \ rvars (\ s')" unfolding rvars_def by force hence yl: "y \ lvars (\ s')" using lvars_rvars by blast interpret rev: RhsEqVal rhs_eq_val by fact note update = rev.update_valuation_nonlhs[THEN mp, OF norm valuated yl] define diff where "diff = l\<^sub>i - \\ s'\ x\<^sub>i" have "\\ s'\ x\<^sub>i < l\<^sub>i \ 0 < l\<^sub>i - \\ s'\ x\<^sub>i" "l\<^sub>i < \\ s'\ x\<^sub>i \ l\<^sub>i - \\ s'\ x\<^sub>i < 0" using minus_gt by (blast, insert minus_lt, blast) - with lt dir have diff: "lt dir 0 diff" by (auto simp: diff_def) + with lt dir have diff: "lt dir 0 diff" by (auto simp: diff_def simp del: Simplex.bounds_lg) define up where "up = inverse (coeff (rhs ?eq) y) *R diff" define v where "v = \\ (rev.update y (\\ s'\ y + up) s')\" show ?thesis unfolding satisfies_state_index'.simps proof (intro exI[of _ v] conjI) show "v \\<^sub>t \ s'" unfolding v_def using rev.update_satisfies_tableau[OF norm valuated yl] \\\ s'\ \\<^sub>t \ s'\ by auto with **(1) have "v \\<^sub>e ?eq" unfolding satisfies_tableau_def by auto from this[unfolded satisfies_eq_def id] have v_xi: "v x\<^sub>i = (rhs ?eq \ v \)" . from \\\ s'\ \\<^sub>t \ s'\ **(1) have "\\ s'\ \\<^sub>e ?eq" unfolding satisfies_tableau_def by auto hence V_xi: "\\ s'\ x\<^sub>i = (rhs ?eq \ \\ s'\ \)" unfolding satisfies_eq_def id . have "v x\<^sub>i = \\ s'\ x\<^sub>i + coeff (rhs ?eq) y *R up" unfolding v_xi unfolding v_def rev.update_valuate_rhs[OF **(1) norm] poly_eval_update V_xi by simp also have "\ = l\<^sub>i" unfolding up_def diff_def scaleRat_scaleRat using cy0 by simp finally have v_xi_l: "v x\<^sub>i = l\<^sub>i" . { assume both: "\\<^sub>u s' y \ ?R1 \ ?R2" "\\<^sub>u s' y \ None" "\\<^sub>l s' y \ ?R1 \ ?R2" "\\<^sub>l s' y \ None" and diff: "\\<^sub>l s' y \ \\<^sub>u s' y" from both(1) dir obtain xu cu where looku: "look (\\<^sub>i\<^sub>l s') xu = Some (\\<^sub>u s' y, cu) \ look (\\<^sub>i\<^sub>u s') xu = Some (\\<^sub>u s' y,cu)" by (smt Is' indices_state_def le_sup_iff mem_Collect_eq setI set_unsat_simps subsetCE) from both(1) obtain xu' where "xu' \ rvars_eq ?eq" "coeff (rhs ?eq) xu' < 0 \ \\<^sub>u s' y = LI dir s' xu' \ coeff (rhs ?eq) xu' > 0 \ \\<^sub>u s' y = UI dir s' xu'" by blast with x_vars_main(1)[OF looku this] have xu: "xu \ rvars_eq ?eq" "coeff (rhs ?eq) xu < 0 \ \\<^sub>u s' y = LI dir s' xu \ coeff (rhs ?eq) xu > 0 \ \\<^sub>u s' y = UI dir s' xu" by auto { assume "xu \ y" with dist[OF looku, of y] have "look (\\<^sub>i\<^sub>u s') y = None" by (cases "look (\\<^sub>i\<^sub>u s') y", auto simp: boundsu_def indexu_def, blast) with both(2) have False by (simp add: boundsu_def) } hence xu_y: "xu = y" by blast from both(3) dir obtain xl cl where lookl: "look (\\<^sub>i\<^sub>l s') xl = Some (\\<^sub>l s' y, cl) \ look (\\<^sub>i\<^sub>u s') xl = Some (\\<^sub>l s' y,cl)" by (smt Is' indices_state_def le_sup_iff mem_Collect_eq setI set_unsat_simps subsetCE) from both(3) obtain xl' where "xl' \ rvars_eq ?eq" "coeff (rhs ?eq) xl' < 0 \ \\<^sub>l s' y = LI dir s' xl' \ coeff (rhs ?eq) xl' > 0 \ \\<^sub>l s' y = UI dir s' xl'" by blast with x_vars_main(1)[OF lookl this] have xl: "xl \ rvars_eq ?eq" "coeff (rhs ?eq) xl < 0 \ \\<^sub>l s' y = LI dir s' xl \ coeff (rhs ?eq) xl > 0 \ \\<^sub>l s' y = UI dir s' xl" by auto { assume "xl \ y" with dist[OF lookl, of y] have "look (\\<^sub>i\<^sub>l s') y = None" by (cases "look (\\<^sub>i\<^sub>l s') y", auto simp: boundsl_def indexl_def, blast) with both(4) have False by (simp add: boundsl_def) } hence xl_y: "xl = y" by blast from xu(2) xl(2) diff have diff: "xu \ xl" by auto with xu_y xl_y have False by simp } note both_y_False = this show "(J, v) \\<^sub>i\<^sub>b\<^sub>e \\ s'" unfolding satisfies_bounds_index'.simps proof (intro conjI allI impI) fix x c assume x: "\\<^sub>l s' x = Some c" "\\<^sub>l s' x \ J" with k have not_k: "\\<^sub>l s' x \ k" by auto from x have ci: "look (\\<^sub>i\<^sub>l s') x = Some (\\<^sub>l s' x, c)" unfolding boundsl_def indexl_def by auto show "v x = c" proof (cases "\\<^sub>l s' x = i") case False hence iR12: "\\<^sub>l s' x \ ?R1 \ ?R2" using sub x unfolding setI LI by blast from x_rvars(2-3)[OF _ iR12] ci have xr: "x \ rvars (\ s')" and val: "\\ s'\ x = c" by auto with lvars_rvars have xl: "x \ lvars (\ s')" by auto show ?thesis proof (cases "x = y") case False thus ?thesis using val unfolding v_def map2fun_def' update[OF xl] using val by auto next case True note coeff = coeff[folded True] from coeff not_k dir ci have Iu: "\\<^sub>u s' x = k" by auto with ci Iu x(2) k sub False True have both: "\\<^sub>u s' y \ ?R1 \ ?R2" "\\<^sub>l s' y \ ?R1 \ ?R2" and diff: "\\<^sub>l s' y \ \\<^sub>u s' y" unfolding setI LI by auto have "\\<^sub>l s' y \ None" using x True by simp from both_y_False[OF both(1) _ both(2) this diff] have "\\<^sub>u s' y = None" by metis with reasable[OF y] dir coeff True have "dir = Negative \ 0 < coeff (rhs ?eq) y" "dir = Positive \ 0 > coeff (rhs ?eq) y" by (auto simp: bound_compare_defs) with dir coeff[unfolded True] have "k = \\<^sub>l s' y" by auto with diff Iu False True have False by auto thus ?thesis .. qed next case True from LBI ci[unfolded True] dir dist[unfolded distinct_indices_state_def, rule_format, of x i c x\<^sub>i l\<^sub>i] have xxi: "x = x\<^sub>i" and c: "c = l\<^sub>i" by auto have vxi: "v x = l\<^sub>i" unfolding xxi v_xi_l .. thus ?thesis unfolding c by simp qed next fix x c assume x: "\\<^sub>u s' x = Some c" "\\<^sub>u s' x \ J" with k have not_k: "\\<^sub>u s' x \ k" by auto from x have ci: "look (\\<^sub>i\<^sub>u s') x = Some (\\<^sub>u s' x, c)" unfolding boundsu_def indexu_def by auto show "v x = c" proof (cases "\\<^sub>u s' x = i") case False hence iR12: "\\<^sub>u s' x \ ?R1 \ ?R2" using sub x unfolding setI LI by blast from x_rvars(2-3)[OF _ iR12] ci have xr: "x \ rvars (\ s')" and val: "\\ s'\ x = c" by auto with lvars_rvars have xl: "x \ lvars (\ s')" by auto show ?thesis proof (cases "x = y") case False thus ?thesis using val unfolding v_def map2fun_def' update[OF xl] using val by auto next case True note coeff = coeff[folded True] from coeff not_k dir ci have Iu: "\\<^sub>l s' x = k" by auto with ci Iu x(2) k sub False True have both: "\\<^sub>u s' y \ ?R1 \ ?R2" "\\<^sub>l s' y \ ?R1 \ ?R2" and diff: "\\<^sub>l s' y \ \\<^sub>u s' y" unfolding setI LI by auto have "\\<^sub>u s' y \ None" using x True by simp from both_y_False[OF both(1) this both(2) _ diff] have "\\<^sub>l s' y = None" by metis with reasable[OF y] dir coeff True have "dir = Negative \ 0 > coeff (rhs ?eq) y" "dir = Positive \ 0 < coeff (rhs ?eq) y" by (auto simp: bound_compare_defs) with dir coeff[unfolded True] have "k = \\<^sub>u s' y" by auto with diff Iu False True have False by auto thus ?thesis .. qed next case True from LBI ci[unfolded True] dir dist[unfolded distinct_indices_state_def, rule_format, of x i c x\<^sub>i l\<^sub>i] have xxi: "x = x\<^sub>i" and c: "c = l\<^sub>i" by auto have vxi: "v x = l\<^sub>i" unfolding xxi v_xi_l .. thus ?thesis unfolding c by simp qed qed qed qed qed } note minimal_core = this have unsat_core: "unsat_state_core (set_unsat I s')" unfolding unsat_state_core_def unsat_core proof (intro impI conjI Is', clarify) fix v assume "(set I, v) \\<^sub>i\<^sub>s set_unsat I s'" then have Iv: "(set I, v) \\<^sub>i\<^sub>s s'" unfolding satisfies_state_index.simps by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) from Iv have vt: "v \\<^sub>t \ s'" and Iv: "(set I, v) \\<^sub>i\<^sub>b \\ s'" unfolding satisfies_state_index.simps by auto have lt_le_eq: "\ x y :: 'a. (x < y) \ (x \ y \ x \ y)" by auto from Iv dir have lb: "\ x i c l. look (LBI dir s') x = Some (i,l) \ i \ set I \ le (lt dir) l (v x)" unfolding satisfies_bounds_index.simps by (auto simp: lt_le_eq indexl_def indexu_def boundsl_def boundsu_def) from lb[OF LBI iI] have li_x: "le (lt dir) l\<^sub>i (v x\<^sub>i)" . have "\\ s'\ \\<^sub>e ?eq" using nolhs \?eq \ set (\ s')\ unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_tableau_def) then have "\\ s'\ x\<^sub>i = (rhs ?eq) \ \\ s'\ \" using \lhs ?eq = x\<^sub>i\ by (simp add: satisfies_eq_def) moreover have "v \\<^sub>e ?eq" using vt \?eq \ set (\ s')\ by (simp add: satisfies_state_def satisfies_tableau_def) then have "v x\<^sub>i = (rhs ?eq) \ v \" using \lhs ?eq = x\<^sub>i\ by (simp add: satisfies_eq_def) moreover have "\\<^sub>l\<^sub>b (lt dir) (v x\<^sub>i) (LB dir s' x\<^sub>i)" using li_x dir unfolding LB_Some by (auto simp: bound_compare'_defs) moreover from min_rvar_incdec_eq_None'[rule_format, OF dir min_rvar refl Iv] have "le (lt dir) (rhs (?eq) \v\) (rhs (?eq) \ \\ s'\ \)" . ultimately show False using dir lt LB_Some by (auto simp add: bound_compare_defs) qed thus "\ (set_unsat I s') \ minimal_unsat_state_core (set_unsat I s')" using minimal_core by (auto simp: minimal_unsat_state_core_def) qed (simp_all add: *) lemma Check_check: "Check check" proof fix s :: "('i,'a) state" assume "\ s" then show "check s = s" by (simp add: check.simps) next fix s :: "('i,'a) state" and v :: "'a valuation" assume *: "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" then have "v \\<^sub>t \ s = v \\<^sub>t \ (check s)" by (rule check_induct, simp_all add: pivotandupdate_tableau_equiv) moreover have "\ (\ (check s))" by (rule check_induct', simp_all add: * pivotandupdate_tableau_normalized) moreover have "\ (check s)" proof (rule check_induct', simp_all add: * pivotandupdate_tableau_valuated) fix s I show "\ s \ \ (set_unsat I s)" by (simp add: tableau_valuated_def) qed ultimately show "let s' = check s in v \\<^sub>t \ s = v \\<^sub>t \ s' \ \ (\ s') \ \ s'" by (simp add: Let_def) next fix s :: "('i,'a) state" assume *: "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" from * show "\\<^sub>i (check s) = \\<^sub>i s" by (rule check_induct, simp_all add: pivotandupdate_bounds_id) next fix s :: "('i,'a) state" assume *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" "\ (\ s)" "\ s" have "\ \ (check s) \ \ (check s)" proof (rule check_induct'', simp_all add: *) fix s assume "min_lvar_not_in_bounds s = None" "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" then show " \ s" using min_lvar_not_in_bounds_None[of s] unfolding curr_val_satisfies_state_def satisfies_state_def unfolding curr_val_satisfies_no_lhs_def by (auto simp add: satisfies_bounds_set.simps satisfies_bounds.simps) qed then show "\ \ (check s) \ \ (check s)" by blast next fix s :: "('i,'a) state" assume *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" "\ (\ s)" "\ s" have "\ (check s) \ minimal_unsat_state_core (check s)" by (rule check_minimal_unsat_state_core[OF *]) then show "\ (check s) \ minimal_unsat_state_core (check s)" by blast qed end end subsection\Symmetries\ text\\label{sec:symmetries} Simplex algorithm exhibits many symmetric cases. For example, \assert_bound\ treats atoms \Leq x c\ and \Geq x c\ in a symmetric manner, \check_inc\ and \check_dec\ are symmetric, etc. These symmetric cases differ only in several aspects: order relations between numbers (\<\ vs \>\ and \\\ vs \\\), the role of lower and upper bounds (\\\<^sub>l\ vs \\\<^sub>u\) and their updating functions, comparisons with bounds (e.g., \\\<^sub>u\<^sub>b\ vs \\\<^sub>l\<^sub>b\ or \<\<^sub>l\<^sub>b\ vs \>\<^sub>u\<^sub>b\), and atom constructors (\Leq\ and \Geq\). These can be attributed to two different orientations (positive and negative) of rational axis. To avoid duplicating definitions and proofs, \assert_bound\ definition cases for \Leq\ and \Geq\ are replaced by a call to a newly introduced function parametrized by a \Direction\ --- a record containing minimal set of aspects listed above that differ in two definition cases such that other aspects can be derived from them (e.g., only \<\ need to be stored while \\\ can be derived from it). Two constants of the type \Direction\ are defined: \Positive\ (with \<\, \\\ orders, \\\<^sub>l\ for lower and \\\<^sub>u\ for upper bounds and their corresponding updating functions, and \Leq\ constructor) and \Negative\ (completely opposite from the previous one). Similarly, \check_inc\ and \check_dec\ are replaced by a new function \check_incdec\ parametrized by a \Direction\. All lemmas, previously repeated for each symmetric instance, were replaced by a more abstract one, again parametrized by a \Direction\ parameter. \vspace{-3mm} \ (*-------------------------------------------------------------------------- *) subsection\Concrete implementation\ (*-------------------------------------------------------------------------- *) (* -------------------------------------------------------------------------- *) (* Init init_state *) (* -------------------------------------------------------------------------- *) text\It is easy to give a concrete implementation of the initial state constructor, which satisfies the specification of the @{term Init} locale. For example:\ definition init_state :: "_ \ ('i,'a :: zero)state" where "init_state t = State t Mapping.empty Mapping.empty (Mapping.tabulate (vars_list t) (\ v. 0)) False None" interpretation Init "init_state :: _ \ ('i,'a :: lrv)state" proof fix t let ?init = "init_state t :: ('i,'a)state" show "\\ ?init\ \\<^sub>t t" unfolding satisfies_tableau_def satisfies_eq_def proof (safe) fix l r assume "(l, r) \ set t" then have "l \ set (vars_list t)" "vars r \ set (vars_list t)" by (auto simp: set_vars_list) (transfer, force) then have *: "vars r \ lhs ` set t \ (\x\set t. rvars_eq x)" by (auto simp: set_vars_list) have "\\ ?init\ l = (0::'a)" using \l \ set (vars_list t)\ unfolding init_state_def by (auto simp: map2fun_def lookup_tabulate) moreover have "r \ \\ ?init\ \ = (0::'a)" using * proof (transfer fixing: t, goal_cases) case (1 r) { fix x assume "x\{v. r v \ 0}" then have "r x *R \\ ?init\ x = (0::'a)" using 1 unfolding init_state_def by (auto simp add: map2fun_def lookup_tabulate comp_def restrict_map_def set_vars_list Abstract_Linear_Poly.vars_def) } then show ?case by auto qed ultimately show "\\ ?init\ (lhs (l, r)) = rhs (l, r) \ \\ ?init\ \" by auto qed next fix t show "\ (init_state t)" unfolding init_state_def by (auto simp add: lookup_tabulate tableau_valuated_def comp_def restrict_map_def set_vars_list lvars_def rvars_def) qed (simp_all add: init_state_def add: boundsl_def boundsu_def indexl_def indexu_def) (* -------------------------------------------------------------------------- *) (* MinVars min_lvar_not_in_bounds min_rvar_incdec_eq *) (* -------------------------------------------------------------------------- *) definition min_lvar_not_in_bounds :: "('i,'a::{linorder,zero}) state \ var option" where "min_lvar_not_in_bounds s \ min_satisfying (\ x. \ in_bounds x (\\ s\) (\ s)) (map lhs (\ s))" interpretation MinLVarNotInBounds "min_lvar_not_in_bounds :: ('i,'a::lrv) state \ _" proof fix s::"('i,'a) state" show "min_lvar_not_in_bounds s = None \ (\x\lvars (\ s). in_bounds x (\\ s\) (\ s))" unfolding min_lvar_not_in_bounds_def lvars_def using min_satisfying_None by blast next fix s x\<^sub>i show "min_lvar_not_in_bounds s = Some x\<^sub>i \ x\<^sub>i \ lvars (\ s) \ \ in_bounds x\<^sub>i \\ s\ (\ s) \ (\x\lvars (\ s). x < x\<^sub>i \ in_bounds x \\ s\ (\ s))" unfolding min_lvar_not_in_bounds_def lvars_def using min_satisfying_Some by blast+ qed \ \all variables in vs have either a positive or a negative coefficient, so no equal-zero test required.\ definition unsat_indices :: "('i,'a :: linorder) Direction \ ('i,'a) state \ var list \ eq \ 'i list" where "unsat_indices dir s vs eq = (let r = rhs eq; li = LI dir s; ui = UI dir s in remdups (li (lhs eq) # map (\ x. if coeff r x < 0 then li x else ui x) vs))" definition min_rvar_incdec_eq :: "('i,'a) Direction \ ('i,'a::lrv) state \ eq \ 'i list + var" where "min_rvar_incdec_eq dir s eq = (let rvars = Abstract_Linear_Poly.vars_list (rhs eq) in case min_satisfying (\ x. reasable_var dir x eq s) rvars of None \ Inl (unsat_indices dir s rvars eq) | Some x\<^sub>j \ Inr x\<^sub>j)" interpretation MinRVarsEq "min_rvar_incdec_eq :: ('i,'a :: lrv) Direction \ _" proof fix s eq "is" and dir :: "('i,'a) Direction" let ?min = "min_satisfying (\ x. reasable_var dir x eq s) (Abstract_Linear_Poly.vars_list (rhs eq))" let ?vars = "Abstract_Linear_Poly.vars_list (rhs eq)" { assume "min_rvar_incdec_eq dir s eq = Inl is" from this[unfolded min_rvar_incdec_eq_def Let_def, simplified] have "?min = None" and I: "set is = set (unsat_indices dir s ?vars eq)" by (cases ?min, auto)+ from this min_satisfying_None set_vars_list have 1: "\ x. x \ rvars_eq eq \ \ reasable_var dir x eq s" by blast { fix i assume "i \ set is" and dir: "dir = Positive \ dir = Negative" and lhs_eq: "LI dir s (lhs eq) \ indices_state s" from this[unfolded I unsat_indices_def Let_def] consider (lhs) "i = LI dir s (lhs eq)" | (LI_rhs) x where "i = LI dir s x" "x \ rvars_eq eq" "coeff (rhs eq) x < 0" | (UI_rhs) x where "i = UI dir s x" "x \ rvars_eq eq" "coeff (rhs eq) x \ 0" by (auto split: if_splits simp: set_vars_list) then have "i \ indices_state s" proof cases case lhs show ?thesis unfolding lhs using lhs_eq by auto next case LI_rhs from 1[OF LI_rhs(2)] LI_rhs(3) have "\ (\\<^sub>l\<^sub>b (lt dir) (\\ s\ x) (LB dir s x))" by auto then show ?thesis unfolding LI_rhs(1) unfolding indices_state_def using dir by (auto simp: bound_compare'_defs boundsl_def boundsu_def indexl_def indexu_def split: option.splits intro!: exI[of _ x]) auto next case UI_rhs from UI_rhs(2) have "coeff (rhs eq) x \ 0" by (simp add: coeff_zero) with UI_rhs(3) have "0 < coeff (rhs eq) x" by auto from 1[OF UI_rhs(2)] this have "\ (\\<^sub>u\<^sub>b (lt dir) (\\ s\ x) (UB dir s x))" by auto then show ?thesis unfolding UI_rhs(1) unfolding indices_state_def using dir by (auto simp: bound_compare'_defs boundsl_def boundsu_def indexl_def indexu_def split: option.splits intro!: exI[of _ x]) auto qed } then have 2: "dir = Positive \ dir = Negative \ LI dir s (lhs eq) \ indices_state s \ set is \ indices_state s" by auto show " (\ x \ rvars_eq eq. \ reasable_var dir x eq s) \ set is = {LI dir s (lhs eq)} \ {LI dir s x |x. x \ rvars_eq eq \ coeff (rhs eq) x < 0} \ {UI dir s x |x. x \ rvars_eq eq \ 0 < coeff (rhs eq) x} \ (dir = Positive \ dir = Negative \ LI dir s (lhs eq) \ indices_state s \ set is \ indices_state s)" proof (intro conjI impI 2, goal_cases) case 2 have "set is = {LI dir s (lhs eq)} \ LI dir s ` (rvars_eq eq \ {x. coeff (rhs eq) x < 0}) \ UI dir s ` (rvars_eq eq \ {x. \ coeff (rhs eq) x < 0})" unfolding I unsat_indices_def Let_def by (auto simp add: set_vars_list) also have "\ = {LI dir s (lhs eq)} \ LI dir s ` {x. x \ rvars_eq eq \ coeff (rhs eq) x < 0} \ UI dir s ` { x. x \ rvars_eq eq \ 0 < coeff (rhs eq) x}" proof (intro arg_cong2[of _ _ _ _ "(\)"] arg_cong[of _ _ "\ x. _ ` x"] refl, goal_cases) case 2 { fix x assume "x \ rvars_eq eq" hence "coeff (rhs eq) x \ 0" by (simp add: coeff_zero) hence or: "coeff (rhs eq) x < 0 \ coeff (rhs eq) x > 0" by auto assume "\ coeff (rhs eq) x < 0" hence "coeff (rhs eq) x > 0" using or by simp } note [dest] = this show ?case by auto qed auto finally show "set is = {LI dir s (lhs eq)} \ {LI dir s x |x. x \ rvars_eq eq \ coeff (rhs eq) x < 0} \ {UI dir s x |x. x \ rvars_eq eq \ 0 < coeff (rhs eq) x}" by auto qed (insert 1, auto) } fix x\<^sub>j assume "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j" from this[unfolded min_rvar_incdec_eq_def Let_def] have "?min = Some x\<^sub>j" by (cases ?min, auto) then show "x\<^sub>j \ rvars_eq eq" "reasable_var dir x\<^sub>j eq s" "(\ x' \ rvars_eq eq. x' < x\<^sub>j \ \ reasable_var dir x' eq s)" using min_satisfying_Some set_vars_list by blast+ qed (* -------------------------------------------------------------------------- *) (* EqForLVar eq_idx_for_lvar *) (* -------------------------------------------------------------------------- *) primrec eq_idx_for_lvar_aux :: "tableau \ var \ nat \ nat" where "eq_idx_for_lvar_aux [] x i = i" | "eq_idx_for_lvar_aux (eq # t) x i = (if lhs eq = x then i else eq_idx_for_lvar_aux t x (i+1))" definition eq_idx_for_lvar where "eq_idx_for_lvar t x \ eq_idx_for_lvar_aux t x 0" lemma eq_idx_for_lvar_aux: assumes "x \ lvars t" shows "let idx = eq_idx_for_lvar_aux t x i in i \ idx \ idx < i + length t \ lhs (t ! (idx - i)) = x" using assms proof (induct t arbitrary: i) case Nil then show ?case by (simp add: lvars_def) next case (Cons eq t) show ?case using Cons(1)[of "i+1"] Cons(2) by (cases "x = lhs eq") (auto simp add: Let_def lvars_def nth_Cons') qed global_interpretation EqForLVarDefault: EqForLVar eq_idx_for_lvar defines eq_for_lvar_code = EqForLVarDefault.eq_for_lvar proof (unfold_locales) fix x t assume "x \ lvars t" then show "eq_idx_for_lvar t x < length t \ lhs (t ! eq_idx_for_lvar t x) = x" using eq_idx_for_lvar_aux[of x t 0] by (simp add: Let_def eq_idx_for_lvar_def) qed (* -------------------------------------------------------------------------- *) (* PivotEq pivot_eq *) (* -------------------------------------------------------------------------- *) definition pivot_eq :: "eq \ var \ eq" where "pivot_eq e y \ let cy = coeff (rhs e) y in (y, (-1/cy) *R ((rhs e) - cy *R (Var y)) + (1/cy) *R (Var (lhs e)))" lemma pivot_eq_satisfies_eq: assumes "y \ rvars_eq e" shows "v \\<^sub>e e = v \\<^sub>e pivot_eq e y" using assms using scaleRat_right_distrib[of "1 / Rep_linear_poly (rhs e) y" "- (rhs e \ v \)" "v (lhs e)"] using Groups.group_add_class.minus_unique[of "- ((rhs e) \ v \)" "v (lhs e)"] unfolding coeff_def vars_def by (simp add: coeff_def vars_def Let_def pivot_eq_def satisfies_eq_def) (auto simp add: rational_vector.scale_right_diff_distrib valuate_add valuate_minus valuate_uminus valuate_scaleRat valuate_Var) lemma pivot_eq_rvars: assumes "x \ vars (rhs (pivot_eq e v))" "x \ lhs e" "coeff (rhs e) v \ 0" "v \ lhs e" shows "x \ vars (rhs e)" proof- have "v \ vars ((1 / coeff (rhs e) v) *R (rhs e - coeff (rhs e) v *R Var v))" using coeff_zero by force then have "x \ v" using assms(1) assms(3) assms(4) using vars_plus[of "(-1 / coeff (rhs e) v) *R (rhs e - coeff (rhs e) v *R Var v)" "(1 / coeff (rhs e) v) *R Var (lhs e)"] by (auto simp add: Let_def vars_scaleRat pivot_eq_def) then show ?thesis using assms using vars_plus[of "(-1 / coeff (rhs e) v) *R (rhs e - coeff (rhs e) v *R Var v)" "(1 / coeff (rhs e) v) *R Var (lhs e)"] using vars_minus[of "rhs e" "coeff (rhs e) v *R Var v"] by (auto simp add: vars_scaleRat Let_def pivot_eq_def) qed interpretation PivotEq pivot_eq proof fix eq x\<^sub>j assume "x\<^sub>j \ rvars_eq eq" "lhs eq \ rvars_eq eq" have "lhs (pivot_eq eq x\<^sub>j) = x\<^sub>j" unfolding pivot_eq_def by (simp add: Let_def) moreover have "rvars_eq (pivot_eq eq x\<^sub>j) = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" proof show "rvars_eq (pivot_eq eq x\<^sub>j) \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" proof fix x assume "x \ rvars_eq (pivot_eq eq x\<^sub>j)" have *: "coeff (rhs (pivot_eq eq x\<^sub>j)) x\<^sub>j = 0" using \x\<^sub>j \ rvars_eq eq\ \lhs eq \ rvars_eq eq\ using coeff_Var2[of "lhs eq" x\<^sub>j] by (auto simp add: Let_def pivot_eq_def) have "coeff (rhs eq) x\<^sub>j \ 0" using \x\<^sub>j \ rvars_eq eq\ using coeff_zero by (cases eq) (auto simp add:) then show "x \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" using pivot_eq_rvars[of x eq x\<^sub>j] using \x \ rvars_eq (pivot_eq eq x\<^sub>j)\ \x\<^sub>j \ rvars_eq eq\ \lhs eq \ rvars_eq eq\ using coeff_zero * by auto qed show "{lhs eq} \ (rvars_eq eq - {x\<^sub>j}) \ rvars_eq (pivot_eq eq x\<^sub>j)" proof fix x assume "x \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" have *: "coeff (rhs eq) (lhs eq) = 0" using coeff_zero using \lhs eq \ rvars_eq eq\ by auto have **: "coeff (rhs eq) x\<^sub>j \ 0" using \x\<^sub>j \ rvars_eq eq\ by (simp add: coeff_zero) have ***: "x \ rvars_eq eq \ coeff (Var (lhs eq)) x = 0" using \lhs eq \ rvars_eq eq\ using coeff_Var2[of "lhs eq" x] by auto have "coeff (Var x\<^sub>j) (lhs eq) = 0" using \x\<^sub>j \ rvars_eq eq\ \lhs eq \ rvars_eq eq\ using coeff_Var2[of x\<^sub>j "lhs eq"] by auto then have "coeff (rhs (pivot_eq eq x\<^sub>j)) x \ 0" using \x \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})\ * ** *** using coeff_zero[of "rhs eq" x] by (auto simp add: Let_def coeff_Var2 pivot_eq_def) then show "x \ rvars_eq (pivot_eq eq x\<^sub>j)" by (simp add: coeff_zero) qed qed ultimately show "let eq' = pivot_eq eq x\<^sub>j in lhs eq' = x\<^sub>j \ rvars_eq eq' = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" by (simp add: Let_def) next fix v eq x\<^sub>j assume "x\<^sub>j \ rvars_eq eq" then show "v \\<^sub>e pivot_eq eq x\<^sub>j = v \\<^sub>e eq" using pivot_eq_satisfies_eq by blast qed (* -------------------------------------------------------------------------- *) (* SubstVar subst_var *) (* -------------------------------------------------------------------------- *) definition subst_var:: "var \ linear_poly \ linear_poly \ linear_poly" where "subst_var v lp' lp \ lp + (coeff lp v) *R lp' - (coeff lp v) *R (Var v)" definition "subst_var_eq_code = SubstVar.subst_var_eq subst_var" global_interpretation SubstVar subst_var rewrites "SubstVar.subst_var_eq subst_var = subst_var_eq_code" proof (unfold_locales) fix x\<^sub>j lp' lp have *: "\x. \x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j); x \ vars lp'\ \ x \ vars lp" proof- fix x assume "x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j)" then have "coeff (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j) x \ 0" using coeff_zero by force assume "x \ vars lp'" then have "coeff lp' x = 0" using coeff_zero by auto show "x \ vars lp" proof(rule ccontr) assume "x \ vars lp" then have "coeff lp x = 0" using coeff_zero by auto then show False using \coeff (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j) x \ 0\ using \coeff lp' x = 0\ by (cases "x = x\<^sub>j") (auto simp add: coeff_Var2) qed qed have "vars (subst_var x\<^sub>j lp' lp) \ (vars lp - {x\<^sub>j}) \ vars lp'" unfolding subst_var_def using coeff_zero[of "lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j" x\<^sub>j] using coeff_zero[of lp' x\<^sub>j] using * by auto moreover have "\x. \x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j); x \ vars lp; x \ vars lp'\ \ x = x\<^sub>j" proof- fix x assume "x \ vars lp" "x \ vars lp'" then have "coeff lp x \ 0" "coeff lp' x = 0" using coeff_zero by auto assume "x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j)" then have "coeff (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j) x = 0" using coeff_zero by force then show "x = x\<^sub>j" using \coeff lp x \ 0\ \coeff lp' x = 0\ by (cases "x = x\<^sub>j") (auto simp add: coeff_Var2) qed then have "vars lp - {x\<^sub>j} - vars lp' \ vars (subst_var x\<^sub>j lp' lp)" by (auto simp add: subst_var_def) ultimately show "vars lp - {x\<^sub>j} - vars lp' \s vars (subst_var x\<^sub>j lp' lp) \s vars lp - {x\<^sub>j} \ vars lp'" by simp next fix v x\<^sub>j lp' lp show "v x\<^sub>j = lp' \ v \ \ lp \ v \ = (subst_var x\<^sub>j lp' lp) \ v \" unfolding subst_var_def using valuate_minus[of "lp + coeff lp x\<^sub>j *R lp'" "coeff lp x\<^sub>j *R Var x\<^sub>j" v] using valuate_add[of lp "coeff lp x\<^sub>j *R lp'" v] using valuate_scaleRat[of "coeff lp x\<^sub>j" lp' v] valuate_scaleRat[of "coeff lp x\<^sub>j" "Var x\<^sub>j" v] using valuate_Var[of x\<^sub>j v] by auto next fix x\<^sub>j lp lp' assume "x\<^sub>j \ vars lp" hence 0: "coeff lp x\<^sub>j = 0" using coeff_zero by blast show "subst_var x\<^sub>j lp' lp = lp" unfolding subst_var_def 0 by simp next fix x\<^sub>j lp x lp' assume "x\<^sub>j \ vars lp" "x \ vars lp' - vars lp" hence x: "x \ x\<^sub>j" and 0: "coeff lp x = 0" and no0: "coeff lp x\<^sub>j \ 0" "coeff lp' x \ 0" using coeff_zero by blast+ from x have 00: "coeff (Var x\<^sub>j) x = 0" using coeff_Var2 by auto show "x \ vars (subst_var x\<^sub>j lp' lp)" unfolding subst_var_def coeff_zero[symmetric] by (simp add: 0 00 no0) qed (simp_all add: subst_var_eq_code_def) (* -------------------------------------------------------------------------- *) (* Update update *) (* -------------------------------------------------------------------------- *) definition rhs_eq_val where "rhs_eq_val v x\<^sub>i c e \ let x\<^sub>j = lhs e; a\<^sub>i\<^sub>j = coeff (rhs e) x\<^sub>i in \v\ x\<^sub>j + a\<^sub>i\<^sub>j *R (c - \v\ x\<^sub>i)" definition "update_code = RhsEqVal.update rhs_eq_val" definition "assert_bound'_code = Update.assert_bound' update_code" definition "assert_bound_code = Update.assert_bound update_code" global_interpretation RhsEqValDefault': RhsEqVal rhs_eq_val rewrites "RhsEqVal.update rhs_eq_val = update_code" and "Update.assert_bound update_code = assert_bound_code" and "Update.assert_bound' update_code = assert_bound'_code" proof unfold_locales fix v x c e assume "\v\ \\<^sub>e e" then show "rhs_eq_val v x c e = rhs e \ \v\(x := c) \" unfolding rhs_eq_val_def Let_def using valuate_update_x[of "rhs e" x "\v\" "\v\(x := c)"] by (auto simp add: satisfies_eq_def) qed (auto simp: update_code_def assert_bound'_code_def assert_bound_code_def) sublocale PivotUpdateMinVars < Check check proof (rule Check_check) show "RhsEqVal rhs_eq_val" .. qed definition "pivot_code = Pivot'.pivot eq_idx_for_lvar pivot_eq subst_var" definition "pivot_tableau_code = Pivot'.pivot_tableau eq_idx_for_lvar pivot_eq subst_var" global_interpretation Pivot'Default: Pivot' eq_idx_for_lvar pivot_eq subst_var rewrites "Pivot'.pivot eq_idx_for_lvar pivot_eq subst_var = pivot_code" and "Pivot'.pivot_tableau eq_idx_for_lvar pivot_eq subst_var = pivot_tableau_code" and "SubstVar.subst_var_eq subst_var = subst_var_eq_code" by (unfold_locales, auto simp: pivot_tableau_code_def pivot_code_def subst_var_eq_code_def) definition "pivot_and_update_code = PivotUpdate.pivot_and_update pivot_code update_code" global_interpretation PivotUpdateDefault: PivotUpdate eq_idx_for_lvar pivot_code update_code rewrites "PivotUpdate.pivot_and_update pivot_code update_code = pivot_and_update_code" by (unfold_locales, auto simp: pivot_and_update_code_def) sublocale Update < AssertBoundNoLhs assert_bound proof (rule update_to_assert_bound_no_lhs) show "Pivot eq_idx_for_lvar pivot_code" .. qed definition "check_code = PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code" definition "check'_code = PivotUpdateMinVars.check' eq_idx_for_lvar min_rvar_incdec_eq pivot_and_update_code" global_interpretation PivotUpdateMinVarsDefault: PivotUpdateMinVars eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code rewrites "PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code = check_code" and "PivotUpdateMinVars.check' eq_idx_for_lvar min_rvar_incdec_eq pivot_and_update_code = check'_code" by (unfold_locales) (simp_all add: check_code_def check'_code_def) definition "assert_code = Assert'.assert assert_bound_code check_code" global_interpretation Assert'Default: Assert' assert_bound_code check_code rewrites "Assert'.assert assert_bound_code check_code = assert_code" by (unfold_locales, auto simp: assert_code_def) definition "assert_bound_loop_code = AssertAllState''.assert_bound_loop assert_bound_code" definition "assert_all_state_code = AssertAllState''.assert_all_state init_state assert_bound_code check_code" definition "assert_all_code = AssertAllState.assert_all assert_all_state_code" global_interpretation AssertAllStateDefault: AssertAllState'' init_state assert_bound_code check_code rewrites "AssertAllState''.assert_bound_loop assert_bound_code = assert_bound_loop_code" and "AssertAllState''.assert_all_state init_state assert_bound_code check_code = assert_all_state_code" and "AssertAllState.assert_all assert_all_state_code = assert_all_code" by unfold_locales (simp_all add: assert_bound_loop_code_def assert_all_state_code_def assert_all_code_def) (* -------------------------------------------------------------------------- *) (* Preprocess preprocess *) (* -------------------------------------------------------------------------- *) primrec monom_to_atom:: "QDelta ns_constraint \ QDelta atom" where "monom_to_atom (LEQ_ns l r) = (if (monom_coeff l < 0) then (Geq (monom_var l) (r /R monom_coeff l)) else (Leq (monom_var l) (r /R monom_coeff l)))" | "monom_to_atom (GEQ_ns l r) = (if (monom_coeff l < 0) then (Leq (monom_var l) (r /R monom_coeff l)) else (Geq (monom_var l) (r /R monom_coeff l)))" primrec qdelta_constraint_to_atom:: "QDelta ns_constraint \ var \ QDelta atom" where "qdelta_constraint_to_atom (LEQ_ns l r) v = (if (is_monom l) then (monom_to_atom (LEQ_ns l r)) else (Leq v r))" | "qdelta_constraint_to_atom (GEQ_ns l r) v = (if (is_monom l) then (monom_to_atom (GEQ_ns l r)) else (Geq v r))" primrec qdelta_constraint_to_atom':: "QDelta ns_constraint \ var \ QDelta atom" where "qdelta_constraint_to_atom' (LEQ_ns l r) v = (Leq v r)" | "qdelta_constraint_to_atom' (GEQ_ns l r) v = (Geq v r)" fun linear_poly_to_eq:: "linear_poly \ var \ eq" where "linear_poly_to_eq p v = (v, p)" datatype 'i istate = IState (FirstFreshVariable: var) (Tableau: tableau) (Atoms: "('i,QDelta) i_atom list") (Poly_Mapping: "linear_poly \ var") (UnsatIndices: "'i list") primrec zero_satisfies :: "'a :: lrv ns_constraint \ bool" where "zero_satisfies (LEQ_ns l r) \ 0 \ r" | "zero_satisfies (GEQ_ns l r) \ 0 \ r" lemma zero_satisfies: "poly c = 0 \ zero_satisfies c \ v \\<^sub>n\<^sub>s c" by (cases c, auto simp: valuate_zero) lemma not_zero_satisfies: "poly c = 0 \ \ zero_satisfies c \ \ v \\<^sub>n\<^sub>s c" by (cases c, auto simp: valuate_zero) fun preprocess' :: "('i,QDelta) i_ns_constraint list \ var \ 'i istate" where "preprocess' [] v = IState v [] [] (\ p. None) []" | "preprocess' ((i,h) # t) v = (let s' = preprocess' t v; p = poly h; is_monom_h = is_monom p; v' = FirstFreshVariable s'; t' = Tableau s'; a' = Atoms s'; m' = Poly_Mapping s'; u' = UnsatIndices s' in if is_monom_h then IState v' t' ((i,qdelta_constraint_to_atom h v') # a') m' u' else if p = 0 then if zero_satisfies h then s' else IState v' t' a' m' (i # u') else (case m' p of Some v \ IState v' t' ((i,qdelta_constraint_to_atom h v) # a') m' u' | None \ IState (v' + 1) (linear_poly_to_eq p v' # t') ((i,qdelta_constraint_to_atom h v') # a') (m' (p \ v')) u') )" lemma preprocess'_simps: "preprocess' ((i,h) # t) v = (let s' = preprocess' t v; p = poly h; is_monom_h = is_monom p; v' = FirstFreshVariable s'; t' = Tableau s'; a' = Atoms s'; m' = Poly_Mapping s'; u' = UnsatIndices s' in if is_monom_h then IState v' t' ((i,monom_to_atom h) # a') m' u' else if p = 0 then if zero_satisfies h then s' else IState v' t' a' m' (i # u') else (case m' p of Some v \ IState v' t' ((i,qdelta_constraint_to_atom' h v) # a') m' u' | None \ IState (v' + 1) (linear_poly_to_eq p v' # t') ((i,qdelta_constraint_to_atom' h v') # a') (m' (p \ v')) u') )" by (cases h, auto simp add: Let_def split: option.splits) lemmas preprocess'_code = preprocess'.simps(1) preprocess'_simps declare preprocess'_code[code] text \Normalization of constraints helps to identify same polynomials, e.g., the constraints $x + y \leq 5$ and $-2x-2y \leq -12$ will be normalized to $x + y \leq 5$ and $x + y \geq 6$, so that only one slack-variable will be introduced for the polynomial $x+y$, and not another one for $-2x-2y$. Normalization will take care that the max-var of the polynomial in the constraint will have coefficient 1 (if the polynomial is non-zero)\ fun normalize_ns_constraint :: "'a :: lrv ns_constraint \ 'a ns_constraint" where "normalize_ns_constraint (LEQ_ns l r) = (let v = max_var l; c = coeff l v in if c = 0 then LEQ_ns l r else let ic = inverse c in if c < 0 then GEQ_ns (ic *R l) (scaleRat ic r) else LEQ_ns (ic *R l) (scaleRat ic r))" | "normalize_ns_constraint (GEQ_ns l r) = (let v = max_var l; c = coeff l v in if c = 0 then GEQ_ns l r else let ic = inverse c in if c < 0 then LEQ_ns (ic *R l) (scaleRat ic r) else GEQ_ns (ic *R l) (scaleRat ic r))" lemma normalize_ns_constraint[simp]: "v \\<^sub>n\<^sub>s (normalize_ns_constraint c) \ v \\<^sub>n\<^sub>s (c :: 'a :: lrv ns_constraint)" proof - let ?c = "coeff (poly c) (max_var (poly c))" consider (0) "?c = 0" | (pos) "?c > 0" | (neg) "?c < 0" by linarith thus ?thesis proof cases case 0 thus ?thesis by (cases c, auto) next case pos from pos have id: "a /R ?c \ b /R ?c \ (a :: 'a) \ b" for a b using scaleRat_leq1 by fastforce show ?thesis using pos id by (cases c, auto simp: Let_def valuate_scaleRat id) next case neg from neg have id: "a /R ?c \ b /R ?c \ (a :: 'a) \ b" for a b using scaleRat_leq2 by fastforce show ?thesis using neg id by (cases c, auto simp: Let_def valuate_scaleRat id) qed qed declare normalize_ns_constraint.simps[simp del] lemma i_satisfies_normalize_ns_constraint[simp]: "Iv \\<^sub>i\<^sub>n\<^sub>s\<^sub>s (map_prod id normalize_ns_constraint ` cs) \ Iv \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs" by (cases Iv, force) abbreviation max_var:: "QDelta ns_constraint \ var" where "max_var C \ Abstract_Linear_Poly.max_var (poly C)" fun start_fresh_variable :: "('i,QDelta) i_ns_constraint list \ var" where "start_fresh_variable [] = 0" | "start_fresh_variable ((i,h)#t) = max (max_var h + 1) (start_fresh_variable t)" definition preprocess_part_1 :: "('i,QDelta) i_ns_constraint list \ tableau \ (('i,QDelta) i_atom list) \ 'i list" where "preprocess_part_1 l \ let start = start_fresh_variable l; is = preprocess' l start in (Tableau is, Atoms is, UnsatIndices is)" lemma lhs_linear_poly_to_eq [simp]: "lhs (linear_poly_to_eq h v) = v" by (cases h) auto lemma rvars_eq_linear_poly_to_eq [simp]: "rvars_eq (linear_poly_to_eq h v) = vars h" by simp lemma fresh_var_monoinc: "FirstFreshVariable (preprocess' cs start) \ start" by (induct cs) (auto simp add: Let_def split: option.splits) abbreviation vars_constraints where "vars_constraints cs \ \ (set (map vars (map poly cs)))" lemma start_fresh_variable_fresh: "\ var \ vars_constraints (flat_list cs). var < start_fresh_variable cs" using max_var_max by (induct cs, auto simp add: max_def) force+ lemma vars_tableau_vars_constraints: "rvars (Tableau (preprocess' cs start)) \ vars_constraints (flat_list cs)" by (induct cs start rule: preprocess'.induct) (auto simp add: rvars_def Let_def split: option.splits) lemma lvars_tableau_ge_start: "\ var \ lvars (Tableau (preprocess' cs start)). var \ start" by (induct cs start rule: preprocess'.induct) (auto simp add: Let_def lvars_def fresh_var_monoinc split: option.splits) lemma rhs_no_zero_tableau_start: "0 \ rhs ` set (Tableau (preprocess' cs start))" by (induct cs start rule: preprocess'.induct, auto simp add: Let_def rvars_def fresh_var_monoinc split: option.splits) lemma first_fresh_variable_not_in_lvars: "\var \ lvars (Tableau (preprocess' cs start)). FirstFreshVariable (preprocess' cs start) > var" by (induct cs start rule: preprocess'.induct) (auto simp add: Let_def lvars_def split: option.splits) lemma sat_atom_sat_eq_sat_constraint_non_monom: assumes "v \\<^sub>a qdelta_constraint_to_atom h var" "v \\<^sub>e linear_poly_to_eq (poly h) var" "\ is_monom (poly h)" shows "v \\<^sub>n\<^sub>s h" using assms by (cases h) (auto simp add: satisfies_eq_def split: if_splits) lemma qdelta_constraint_to_atom_monom: assumes "is_monom (poly h)" shows "v \\<^sub>a qdelta_constraint_to_atom h var \ v \\<^sub>n\<^sub>s h" proof (cases h) case (LEQ_ns l a) then show ?thesis using assms using monom_valuate[of _ v] apply auto using scaleRat_leq2[of "a /R monom_coeff l" "v (monom_var l)" "monom_coeff l"] using divide_leq1[of "monom_coeff l" "v (monom_var l)" a] apply (force, simp add: divide_rat_def) using scaleRat_leq1[of "v (monom_var l)" "a /R monom_coeff l" "monom_coeff l"] using is_monom_monom_coeff_not_zero[of l] using divide_leq[of "monom_coeff l" "v (monom_var l)" a] using is_monom_monom_coeff_not_zero[of l] by (simp_all add: divide_rat_def) next case (GEQ_ns l a) then show ?thesis using assms using monom_valuate[of _ v] apply auto using scaleRat_leq2[of "v (monom_var l)" "a /R monom_coeff l" "monom_coeff l"] using divide_geq1[of a "monom_coeff l" "v (monom_var l)"] apply (force, simp add: divide_rat_def) using scaleRat_leq1[of "a /R monom_coeff l" "v (monom_var l)" "monom_coeff l"] using is_monom_monom_coeff_not_zero[of l] using divide_geq[of a "monom_coeff l" "v (monom_var l)"] using is_monom_monom_coeff_not_zero[of l] by (simp_all add: divide_rat_def) qed lemma preprocess'_Tableau_Poly_Mapping_None: "(Poly_Mapping (preprocess' cs start)) p = None \ linear_poly_to_eq p v \ set (Tableau (preprocess' cs start))" by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits if_splits) lemma preprocess'_Tableau_Poly_Mapping_Some: "(Poly_Mapping (preprocess' cs start)) p = Some v \ linear_poly_to_eq p v \ set (Tableau (preprocess' cs start))" by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits if_splits) lemma preprocess'_Tableau_Poly_Mapping_Some': "(Poly_Mapping (preprocess' cs start)) p = Some v \ \ h. poly h = p \ \ is_monom (poly h) \ qdelta_constraint_to_atom h v \ flat (set (Atoms (preprocess' cs start)))" by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits if_splits) lemma not_one_le_zero_qdelta: "\ (1 \ (0 :: QDelta))" by code_simp lemma one_zero_contra[dest,consumes 2]: "1 \ x \ (x :: QDelta) \ 0 \ False" using order.trans[of 1 x 0] not_one_le_zero_qdelta by simp lemma i_preprocess'_sat: assumes "(I,v) \\<^sub>i\<^sub>a\<^sub>s set (Atoms (preprocess' s start))" "v \\<^sub>t Tableau (preprocess' s start)" "I \ set (UnsatIndices (preprocess' s start)) = {}" shows "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set s" using assms by (induct s start rule: preprocess'.induct) (auto simp add: Let_def satisfies_atom_set_def satisfies_tableau_def qdelta_constraint_to_atom_monom sat_atom_sat_eq_sat_constraint_non_monom split: if_splits option.splits dest!: preprocess'_Tableau_Poly_Mapping_Some zero_satisfies) lemma preprocess'_sat: assumes "v \\<^sub>a\<^sub>s flat (set (Atoms (preprocess' s start)))" "v \\<^sub>t Tableau (preprocess' s start)" "set (UnsatIndices (preprocess' s start)) = {}" shows "v \\<^sub>n\<^sub>s\<^sub>s flat (set s)" using i_preprocess'_sat[of UNIV v s start] assms by simp lemma sat_constraint_valuation: assumes "\ var \ vars (poly c). v1 var = v2 var" shows "v1 \\<^sub>n\<^sub>s c \ v2 \\<^sub>n\<^sub>s c" using assms using valuate_depend by (cases c) (force)+ lemma atom_var_first: assumes "a \ flat (set (Atoms (preprocess' cs start)))" "\ var \ vars_constraints (flat_list cs). var < start" shows "atom_var a < FirstFreshVariable (preprocess' cs start)" using assms proof(induct cs arbitrary: a) case (Cons hh t a) obtain i h where hh: "hh = (i,h)" by force let ?s = "preprocess' t start" show ?case proof(cases "a \ flat (set (Atoms ?s))") case True then show ?thesis using Cons(1)[of a] Cons(3) hh by (auto simp add: Let_def split: option.splits) next case False consider (monom) "is_monom (poly h)" | (normal) "\ is_monom (poly h)" "poly h \ 0" "(Poly_Mapping ?s) (poly h) = None" | (old) var where "\ is_monom (poly h)" "poly h \ 0" "(Poly_Mapping ?s) (poly h) = Some var" | (zero) "\ is_monom (poly h)" "poly h = 0" by auto then show ?thesis proof cases case monom from Cons(3) monom_var_in_vars hh monom have "monom_var (poly h) < start" by auto moreover from False have "a = qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))" using Cons(2) hh monom by (auto simp: Let_def) ultimately show ?thesis using fresh_var_monoinc[of start t] hh monom by (cases a; cases h) (auto simp add: Let_def ) next case normal have "a = qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))" using False normal Cons(2) hh by (auto simp: Let_def) then show ?thesis using hh normal by (cases a; cases h) (auto simp add: Let_def ) next case (old var) from preprocess'_Tableau_Poly_Mapping_Some'[OF old(3)] obtain h' where "poly h' = poly h" "qdelta_constraint_to_atom h' var \ flat (set (Atoms ?s))" by blast from Cons(1)[OF this(2)] Cons(3) this(1) old(1) have var: "var < FirstFreshVariable ?s" by (cases h', auto) have "a = qdelta_constraint_to_atom h var" using False old Cons(2) hh by (auto simp: Let_def) then have a: "atom_var a = var" using old by (cases a; cases h; auto simp: Let_def) show ?thesis unfolding a hh by (simp add: old Let_def var) next case zero from False show ?thesis using Cons(2) hh zero by (auto simp: Let_def split: if_splits) qed qed qed simp lemma satisfies_tableau_satisfies_tableau: assumes "v1 \\<^sub>t t" "\ var \ tvars t. v1 var = v2 var" shows "v2 \\<^sub>t t" using assms using valuate_depend[of _ v1 v2] by (force simp add: lvars_def rvars_def satisfies_eq_def satisfies_tableau_def) lemma preprocess'_unsat_indices: assumes "i \ set (UnsatIndices (preprocess' s start))" shows "\ ({i},v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set s" using assms proof (induct s start rule: preprocess'.induct) case (2 j h t v) then show ?case by (auto simp: Let_def not_zero_satisfies split: if_splits option.splits) qed simp lemma preprocess'_unsat: assumes "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set s" "vars_constraints (flat_list s) \ V" "\var \ V. var < start" shows "\v'. (\var \ V. v var = v' var) \ v' \\<^sub>a\<^sub>s restrict_to I (set (Atoms (preprocess' s start))) \ v' \\<^sub>t Tableau (preprocess' s start)" using assms proof(induct s) case Nil show ?case by (auto simp add: satisfies_atom_set_def satisfies_tableau_def) next case (Cons hh t) obtain i h where hh: "hh = (i,h)" by force from Cons hh obtain v' where var: "(\var\V. v var = v' var)" and v'_as: "v' \\<^sub>a\<^sub>s restrict_to I (set (Atoms (preprocess' t start)))" and v'_t: "v' \\<^sub>t Tableau (preprocess' t start)" and vars_h: "vars_constraints [h] \ V" by auto from Cons(2)[unfolded hh] have i: "i \ I \ v \\<^sub>n\<^sub>s h" by auto have "\ var \ vars (poly h). v var = v' var" using \(\var\V. v var = v' var)\ Cons(3) hh by auto then have vh_v'h: "v \\<^sub>n\<^sub>s h \ v' \\<^sub>n\<^sub>s h" by (rule sat_constraint_valuation) show ?case proof(cases "is_monom (poly h)") case True then have id: "is_monom (poly h) = True" by simp show ?thesis unfolding hh preprocess'.simps Let_def id if_True istate.simps istate.sel proof (intro exI[of _ v'] conjI v'_t var satisifies_atom_restrict_to_Cons[OF v'_as]) assume "i \ I" from i[OF this] var vh_v'h show "v' \\<^sub>a qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))" unfolding qdelta_constraint_to_atom_monom[OF True] by auto qed next case False then have id: "is_monom (poly h) = False" by simp let ?s = "preprocess' t start" let ?x = "FirstFreshVariable ?s" show ?thesis proof (cases "poly h = 0") case zero: False hence id': "(poly h = 0) = False" by simp let ?look = "(Poly_Mapping ?s) (poly h)" show ?thesis proof (cases ?look) case None let ?y = "poly h \ v'\" let ?v' = "v'(?x:=?y)" show ?thesis unfolding preprocess'.simps hh Let_def id id' if_False istate.simps istate.sel None option.simps proof (rule exI[of _ ?v'], intro conjI satisifies_atom_restrict_to_Cons satisfies_tableau_Cons) show vars': "(\var\V. v var = ?v' var)" using \(\var\V. v var = v' var)\ using fresh_var_monoinc[of start t] using Cons(4) by auto { assume "i \ I" from vh_v'h i[OF this] False show "?v' \\<^sub>a qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))" by (cases h, auto) } let ?atoms = "restrict_to I (set (Atoms (preprocess' t start)))" show "?v' \\<^sub>a\<^sub>s ?atoms" unfolding satisfies_atom_set_def proof fix a assume "a \ ?atoms" then have "v' \\<^sub>a a" using \v' \\<^sub>a\<^sub>s ?atoms\ hh by (force simp add: satisfies_atom_set_def) then show "?v' \\<^sub>a a" using \a \ ?atoms\ atom_var_first[of a t start] using Cons(3) Cons(4) by (cases a) auto qed show "?v' \\<^sub>e linear_poly_to_eq (poly h) (FirstFreshVariable (preprocess' t start))" using Cons(3) Cons(4) using valuate_depend[of "poly h" v' "v'(FirstFreshVariable (preprocess' t start) := (poly h) \ v' \)"] using fresh_var_monoinc[of start t] hh by (cases h) (force simp add: satisfies_eq_def)+ have "FirstFreshVariable (preprocess' t start) \ tvars (Tableau (preprocess' t start))" using first_fresh_variable_not_in_lvars[of t start] using Cons(3) Cons(4) using vars_tableau_vars_constraints[of t start] using fresh_var_monoinc[of start t] by force then show "?v' \\<^sub>t Tableau (preprocess' t start)" using \v' \\<^sub>t Tableau (preprocess' t start)\ using satisfies_tableau_satisfies_tableau[of v' "Tableau (preprocess' t start)" ?v'] by auto qed next case (Some var) from preprocess'_Tableau_Poly_Mapping_Some[OF Some] have "linear_poly_to_eq (poly h) var \ set (Tableau ?s)" by auto with v'_t[unfolded satisfies_tableau_def] have v'_h_var: "v' \\<^sub>e linear_poly_to_eq (poly h) var" by auto show ?thesis unfolding preprocess'.simps hh Let_def id id' if_False istate.simps istate.sel Some option.simps proof (intro exI[of _ v'] conjI var v'_t satisifies_atom_restrict_to_Cons satisfies_tableau_Cons v'_as) assume "i \ I" from vh_v'h i[OF this] False v'_h_var show "v' \\<^sub>a qdelta_constraint_to_atom h var" by (cases h, auto simp: satisfies_eq_iff) qed qed next case zero: True hence id': "(poly h = 0) = True" by simp show ?thesis proof (cases "zero_satisfies h") case True hence id'': "zero_satisfies h = True" by simp show ?thesis unfolding hh preprocess'.simps Let_def id id' id'' if_True if_False istate.simps istate.sel by (intro exI[of _ v'] conjI v'_t var v'_as) next case False hence id'': "zero_satisfies h = False" by simp { assume "i \ I" from i[OF this] not_zero_satisfies[OF zero False] have False by simp } note no_I = this show ?thesis unfolding hh preprocess'.simps Let_def id id' id'' if_True if_False istate.simps istate.sel proof (rule Cons(1)[OF _ _ Cons(4)]) show "(I, v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set t" using Cons(2) by auto show "vars_constraints (map snd t) \ V" using Cons(3) by force qed qed qed qed qed lemma lvars_distinct: "distinct (map lhs (Tableau (preprocess' cs start)))" using first_fresh_variable_not_in_lvars[where ?'a = 'a] by (induct cs, auto simp add: Let_def lvars_def) (force split: option.splits) lemma normalized_tableau_preprocess': "\ (Tableau (preprocess' cs (start_fresh_variable cs)))" proof - let ?s = "start_fresh_variable cs" show ?thesis using lvars_distinct[of cs ?s] using lvars_tableau_ge_start[of cs ?s] using vars_tableau_vars_constraints[of cs ?s] using start_fresh_variable_fresh[of cs] unfolding normalized_tableau_def Let_def by (smt disjoint_iff_not_equal inf.absorb_iff2 inf.strict_order_iff rhs_no_zero_tableau_start subsetD) qed text \Improved preprocessing: Deletion. An equation x = p can be deleted from the tableau, if x does not occur in the atoms.\ lemma delete_lhs_var: assumes norm: "\ t" and t: "t = t1 @ (x,p) # t2" and t': "t' = t1 @ t2" and tv: "tv = (\ v. upd x (p \ \v\ \) v)" and x_as: "x \ atom_var ` snd ` set as" shows "\ t'" \ \new tableau is normalized\ "\w\ \\<^sub>t t' \ \tv w\ \\<^sub>t t" \ \solution of new tableau is translated to solution of old tableau\ "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" \ \solution translation also works for bounds\ "v \\<^sub>t t \ v \\<^sub>t t'" \ \solution of old tableau is also solution for new tableau\ proof - have tv: "\tv w\ = \w\ (x := p \ \w\ \)" unfolding tv map2fun_def' by auto from norm show "\ t'" unfolding t t' normalized_tableau_def by (auto simp: lvars_def rvars_def) show "v \\<^sub>t t \ v \\<^sub>t t'" unfolding t t' satisfies_tableau_def by auto from norm have dist: "distinct (map lhs t)" "lvars t \ rvars t = {}" unfolding normalized_tableau_def by auto from x_as have x_as: "x \ atom_var ` snd ` (set as \ I \ UNIV)" by auto have "(I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as" unfolding i_satisfies_atom_set.simps satisfies_atom_set_def proof (intro ball_cong[OF refl]) fix a assume "a \ snd ` (set as \ I \ UNIV)" with x_as have "x \ atom_var a" by auto then show "\tv w\ \\<^sub>a a = \w\ \\<^sub>a a" unfolding tv by (cases a, auto) qed then show "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" by blast assume w: "\w\ \\<^sub>t t'" from dist(2)[unfolded t] have xp: "x \ vars p" unfolding lvars_def rvars_def by auto { fix eq assume mem: "eq \ set t1 \ set t2" then have "eq \ set t'" unfolding t' by auto with w have w: "\w\ \\<^sub>e eq" unfolding satisfies_tableau_def by auto obtain y q where eq: "eq = (y,q)" by force from mem[unfolded eq] dist(1)[unfolded t] have xy: "x \ y" by force from mem[unfolded eq] dist(2)[unfolded t] have xq: "x \ vars q" unfolding lvars_def rvars_def by auto from w have "\tv w\ \\<^sub>e eq" unfolding tv eq satisfies_eq_iff using xy xq by (auto intro!: valuate_depend) } moreover have "\tv w\ \\<^sub>e (x,p)" unfolding satisfies_eq_iff tv using xp by (auto intro!: valuate_depend) ultimately show "\tv w\ \\<^sub>t t" unfolding t satisfies_tableau_def by auto qed definition pivot_tableau_eq :: "tableau \ eq \ tableau \ var \ tableau \ eq \ tableau" where "pivot_tableau_eq t1 eq t2 x \ let eq' = pivot_eq eq x; m = map (\ e. subst_var_eq x (rhs eq') e) in (m t1, eq', m t2)" lemma pivot_tableau_eq: assumes t: "t = t1 @ eq # t2" "t' = t1' @ eq' # t2'" and x: "x \ rvars_eq eq" and norm: "\ t" and pte: "pivot_tableau_eq t1 eq t2 x = (t1',eq',t2')" shows "\ t'" "lhs eq' = x" "(v :: 'a :: lrv valuation) \\<^sub>t t' \ v \\<^sub>t t" proof - let ?s = "\ t. State t undefined undefined undefined undefined undefined" let ?y = "lhs eq" have yl: "?y \ lvars t" unfolding t lvars_def by auto from norm have eq_t12: "?y \ lhs ` (set t1 \ set t2)" unfolding normalized_tableau_def t lvars_def by auto have eq: "eq_for_lvar_code t ?y = eq" by (metis (mono_tags, lifting) EqForLVarDefault.eq_for_lvar Un_insert_right eq_t12 image_iff insert_iff list.set(2) set_append t(1) yl) have *: "(?y, b) \ set t1 \ ?y \ lhs ` (set t1)" for b t1 by (metis image_eqI lhs.simps) have pivot: "pivot_tableau_code ?y x t = t'" unfolding Pivot'Default.pivot_tableau_def Let_def eq using pte[symmetric] unfolding t pivot_tableau_eq_def Let_def using eq_t12 by (auto dest!: *) note thms = Pivot'Default.pivot_vars' Pivot'Default.pivot_tableau note thms = thms[unfolded Pivot'Default.pivot_def, of "?s t", simplified, OF norm yl, unfolded eq, OF x, unfolded pivot] from thms(1) thms(2)[of v] show "\ t'" "v \\<^sub>t t' \ v \\<^sub>t t" by auto show "lhs eq' = x" using pte[symmetric] unfolding t pivot_tableau_eq_def Let_def pivot_eq_def by auto qed function preprocess_opt :: "var set \ tableau \ tableau \ tableau \ ((var,'a :: lrv)mapping \ (var,'a)mapping)" where "preprocess_opt X t1 [] = (t1,id)" | "preprocess_opt X t1 ((x,p) # t2) = (if x \ X then case preprocess_opt X t1 t2 of (t,tv) \ (t, (\ v. upd x (p \ \v\ \) v) o tv) else case find (\ x. x \ X) (Abstract_Linear_Poly.vars_list p) of None \ preprocess_opt X ((x,p) # t1) t2 | Some y \ case pivot_tableau_eq t1 (x,p) t2 y of (tt1,(z,q),tt2) \ case preprocess_opt X tt1 tt2 of (t,tv) \ (t, (\ v. upd z (q \ \v\ \) v) o tv))" by pat_completeness auto termination by (relation "measure (\ (X,t1,t2). length t2)", auto simp: pivot_tableau_eq_def Let_def) lemma preprocess_opt: assumes "X = atom_var ` snd ` set as" "preprocess_opt X t1 t2 = (t',tv)" "\ t" "t = rev t1 @ t2" shows "\ t'" "(\w\ :: 'a :: lrv valuation) \\<^sub>t t' \ \tv w\ \\<^sub>t t" "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" "v \\<^sub>t t \ (v :: 'a valuation) \\<^sub>t t'" using assms proof (atomize(full), induct X t1 t2 arbitrary: t tv w rule: preprocess_opt.induct) case (1 X t1 t tv) then show ?case by (auto simp: normalized_tableau_def lvars_def rvars_def satisfies_tableau_def simp flip: rev_map) next case (2 X t1 x p t2 t tv w) note IH = 2(1-3) note X = 2(4) note res = 2(5) have norm: "\ t" by fact have t: "t = rev t1 @ (x, p) # t2" by fact show ?case proof (cases "x \ X") case False with res obtain tv' where res: "preprocess_opt X t1 t2 = (t', tv')" and tv: "tv = (\v. Mapping.update x (p \ \v\ \) v) o tv'" by (auto split: prod.splits) note delete = delete_lhs_var[OF norm t refl refl False[unfolded X]] note IH = IH(1)[OF False X res delete(1) refl] from delete(2)[of "tv' w"] delete(3)[of I "tv' w"] delete(4)[of v] IH[of w] show ?thesis unfolding tv o_def by auto next case True then have "\ x \ X" by simp note IH = IH(2-3)[OF this] show ?thesis proof (cases "find (\x. x \ X) (Abstract_Linear_Poly.vars_list p)") case None with res True have pre: "preprocess_opt X ((x, p) # t1) t2 = (t', tv)" by auto from t have t: "t = rev ((x, p) # t1) @ t2" by simp from IH(1)[OF None X pre norm t] show ?thesis . next case (Some z) from Some[unfolded find_Some_iff] have zX: "z \ X" and "z \ set (Abstract_Linear_Poly.vars_list p)" unfolding set_conv_nth by auto then have z: "z \ rvars_eq (x, p)" by (simp add: set_vars_list) obtain tt1 z' q tt2 where pte: "pivot_tableau_eq t1 (x, p) t2 z = (tt1,(z',q),tt2)" by (cases "pivot_tableau_eq t1 (x, p) t2 z", auto) then have pte_rev: "pivot_tableau_eq (rev t1) (x, p) t2 z = (rev tt1,(z',q),tt2)" unfolding pivot_tableau_eq_def Let_def by (auto simp: rev_map) note eq = pivot_tableau_eq[OF t refl z norm pte_rev] then have z': "z' = z" by auto note eq = eq(1,3)[unfolded z'] note pte = pte[unfolded z'] note pte_rev = pte_rev[unfolded z'] note delete = delete_lhs_var[OF eq(1) refl refl refl zX[unfolded X]] from res[unfolded preprocess_opt.simps Some option.simps pte] True obtain tv' where res: "preprocess_opt X tt1 tt2 = (t', tv')" and tv: "tv = (\v. Mapping.update z (q \ \v\ \) v) o tv'" by (auto split: prod.splits) note IH = IH(2)[OF Some, unfolded pte, OF refl refl refl X res delete(1) refl] from IH[of w] delete(2)[of "tv' w"] delete(3)[of I "tv' w"] delete(4)[of v] show ?thesis unfolding tv o_def eq(2) by auto qed qed qed definition "preprocess_part_2 as t = preprocess_opt (atom_var ` snd ` set as) [] t" lemma preprocess_part_2: assumes "preprocess_part_2 as t = (t',tv)" "\ t" shows "\ t'" "(\w\ :: 'a :: lrv valuation) \\<^sub>t t' \ \tv w\ \\<^sub>t t" "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" "v \\<^sub>t t \ (v :: 'a valuation) \\<^sub>t t'" using preprocess_opt[OF refl assms(1)[unfolded preprocess_part_2_def] assms(2)] by auto definition preprocess :: "('i,QDelta) i_ns_constraint list \ _ \ _ \ (_ \ (var,QDelta)mapping) \ 'i list" where "preprocess l = (case preprocess_part_1 (map (map_prod id normalize_ns_constraint) l) of (t,as,ui) \ case preprocess_part_2 as t of (t,tv) \ (t,as,tv,ui))" lemma preprocess: assumes id: "preprocess cs = (t, as, trans_v, ui)" shows "\ t" "fst ` set as \ set ui \ fst ` set cs" "distinct_indices_ns (set cs) \ distinct_indices_atoms (set as)" "I \ set ui = {} \ (I, \v\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t \ (I, \trans_v v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" "i \ set ui \ \v. ({i}, v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" "\ v. (I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs \ \v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>t t" proof - define ncs where "ncs = map (map_prod id normalize_ns_constraint) cs" have ncs: "fst ` set ncs = fst ` set cs" "\ Iv. Iv \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set ncs \ Iv \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" unfolding ncs_def by force auto from id obtain t1 where part1: "preprocess_part_1 ncs = (t1,as,ui)" unfolding preprocess_def by (auto simp: ncs_def split: prod.splits) from id[unfolded preprocess_def part1 split ncs_def[symmetric]] have part_2: "preprocess_part_2 as t1 = (t,trans_v)" by (auto split: prod.splits) have norm: "\ t1" using normalized_tableau_preprocess' part1 by (auto simp: preprocess_part_1_def Let_def) note part_2 = preprocess_part_2[OF part_2 norm] show "\ t" by fact have unsat: "(I,\v\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t1 \ I \ set ui = {} \ (I,\v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set ncs" for v using part1[unfolded preprocess_part_1_def Let_def, simplified] i_preprocess'_sat[of I] by blast with part_2(2,3) show "I \ set ui = {} \ (I,\v\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t \ (I,\trans_v v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" by (auto simp: ncs) from part1[unfolded preprocess_part_1_def Let_def] obtain var where as: "as = Atoms (preprocess' ncs var)" and ui: "ui = UnsatIndices (preprocess' ncs var)" by auto note min_defs = distinct_indices_atoms_def distinct_indices_ns_def have min1: "(distinct_indices_ns (set ncs) \ (\ k a. (k,a) \ set as \ (\ v p. a = qdelta_constraint_to_atom p v \ (k,p) \ set ncs \ (\ is_monom (poly p) \ Poly_Mapping (preprocess' ncs var) (poly p) = Some v) ))) \ fst ` set as \ set ui \ fst ` set ncs" unfolding as ui proof (induct ncs var rule: preprocess'.induct) case (2 i h t v) hence sub: "fst ` set (Atoms (preprocess' t v)) \ set (UnsatIndices (preprocess' t v)) \ fst ` set t" by auto show ?case proof (intro conjI impI allI, goal_cases) show "fst ` set (Atoms (preprocess' ((i, h) # t) v)) \ set (UnsatIndices (preprocess' ((i,h) #t) v)) \ fst ` set ((i, h) # t)" using sub by (auto simp: Let_def split: option.splits) next case (1 k a) hence min': "distinct_indices_ns (set t)" unfolding min_defs list.simps by blast note IH = 2[THEN conjunct1, rule_format, OF min'] show ?case proof (cases "(k,a) \ set (Atoms (preprocess' t v))") case True from IH[OF this] show ?thesis by (force simp: Let_def split: option.splits if_split) next case new: False with 1(2) have ki: "k = i" by (auto simp: Let_def split: if_splits option.splits) show ?thesis proof (cases "is_monom (poly h)") case True thus ?thesis using new 1(2) by (auto simp: Let_def True intro!: exI) next case no_monom: False thus ?thesis using new 1(2) by (auto simp: Let_def no_monom split: option.splits if_splits intro!: exI) qed qed qed qed (auto simp: min_defs) then show "fst ` set as \ set ui \ fst ` set cs" by (auto simp: ncs) { assume mini: "distinct_indices_ns (set cs)" have mini: "distinct_indices_ns (set ncs)" unfolding distinct_indices_ns_def proof (intro impI allI, goal_cases) case (1 n1 n2 i) from 1(1) obtain c1 where c1: "(i,c1) \ set cs" and n1: "n1 = normalize_ns_constraint c1" unfolding ncs_def by auto from 1(2) obtain c2 where c2: "(i,c2) \ set cs" and n2: "n2 = normalize_ns_constraint c2" unfolding ncs_def by auto from mini[unfolded distinct_indices_ns_def, rule_format, OF c1 c2] show ?case unfolding n1 n2 by (cases c1; cases c2; auto simp: normalize_ns_constraint.simps Let_def) qed note min = min1[THEN conjunct1, rule_format, OF this] show "distinct_indices_atoms (set as)" unfolding distinct_indices_atoms_def proof (intro allI impI) fix i a b assume a: "(i,a) \ set as" and b: "(i,b) \ set as" from min[OF a] obtain v p where aa: "a = qdelta_constraint_to_atom p v" "(i, p) \ set ncs" "\ is_monom (poly p) \ Poly_Mapping (preprocess' ncs var) (poly p) = Some v" by auto from min[OF b] obtain w q where bb: "b = qdelta_constraint_to_atom q w" "(i, q) \ set ncs" "\ is_monom (poly q) \ Poly_Mapping (preprocess' ncs var) (poly q) = Some w" by auto from mini[unfolded distinct_indices_ns_def, rule_format, OF aa(2) bb(2)] have *: "poly p = poly q" "ns_constraint_const p = ns_constraint_const q" by auto show "atom_var a = atom_var b \ atom_const a = atom_const b" proof (cases "is_monom (poly q)") case True thus ?thesis unfolding aa(1) bb(1) using * by (cases p; cases q, auto) next case False thus ?thesis unfolding aa(1) bb(1) using * aa(3) bb(3) by (cases p; cases q, auto) qed qed } show "i \ set ui \ \v. ({i}, v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" using preprocess'_unsat_indices[of i ncs] part1 unfolding preprocess_part_1_def Let_def by (auto simp: ncs) assume "\ w. (I,w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" then obtain w where "(I,w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" by blast hence "(I,w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set ncs" unfolding ncs . from preprocess'_unsat[OF this _ start_fresh_variable_fresh, of ncs] have "\v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>t t1" using part1 unfolding preprocess_part_1_def Let_def by auto then show "\v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>t t" using part_2(4) by auto qed interpretation PreprocessDefault: Preprocess preprocess by (unfold_locales; rule preprocess, auto) global_interpretation Solve_exec_ns'Default: Solve_exec_ns' preprocess assert_all_code defines solve_exec_ns_code = Solve_exec_ns'Default.solve_exec_ns by unfold_locales (* -------------------------------------------------------------------------- *) (* To_ns to_ns from_ns *) (* -------------------------------------------------------------------------- *) primrec constraint_to_qdelta_constraint:: "constraint \ QDelta ns_constraint list" where "constraint_to_qdelta_constraint (LT l r) = [LEQ_ns l (QDelta.QDelta r (-1))]" | "constraint_to_qdelta_constraint (GT l r) = [GEQ_ns l (QDelta.QDelta r 1)]" | "constraint_to_qdelta_constraint (LEQ l r) = [LEQ_ns l (QDelta.QDelta r 0)]" | "constraint_to_qdelta_constraint (GEQ l r) = [GEQ_ns l (QDelta.QDelta r 0)]" | "constraint_to_qdelta_constraint (EQ l r) = [LEQ_ns l (QDelta.QDelta r 0), GEQ_ns l (QDelta.QDelta r 0)]" | "constraint_to_qdelta_constraint (LTPP l1 l2) = [LEQ_ns (l1 - l2) (QDelta.QDelta 0 (-1))]" | "constraint_to_qdelta_constraint (GTPP l1 l2) = [GEQ_ns (l1 - l2) (QDelta.QDelta 0 1)]" | "constraint_to_qdelta_constraint (LEQPP l1 l2) = [LEQ_ns (l1 - l2) 0]" | "constraint_to_qdelta_constraint (GEQPP l1 l2) = [GEQ_ns (l1 - l2) 0]" | "constraint_to_qdelta_constraint (EQPP l1 l2) = [LEQ_ns (l1 - l2) 0, GEQ_ns (l1 - l2) 0]" primrec i_constraint_to_qdelta_constraint:: "'i i_constraint \ ('i,QDelta) i_ns_constraint list" where "i_constraint_to_qdelta_constraint (i,c) = map (Pair i) (constraint_to_qdelta_constraint c)" definition to_ns :: "'i i_constraint list \ ('i,QDelta) i_ns_constraint list" where "to_ns l \ concat (map i_constraint_to_qdelta_constraint l)" primrec \0_val :: "QDelta ns_constraint \ QDelta valuation \ rat" where "\0_val (LEQ_ns lll rrr) vl = \0 lll\vl\ rrr" | "\0_val (GEQ_ns lll rrr) vl = \0 rrr lll\vl\" primrec \0_val_min :: "QDelta ns_constraint list \ QDelta valuation \ rat" where "\0_val_min [] vl = 1" | "\0_val_min (h#t) vl = min (\0_val_min t vl) (\0_val h vl)" abbreviation vars_list_constraints where "vars_list_constraints cs \ remdups (concat (map Abstract_Linear_Poly.vars_list (map poly cs)))" definition from_ns ::"(var, QDelta) mapping \ QDelta ns_constraint list \ (var, rat) mapping" where "from_ns vl cs \ let \ = \0_val_min cs \vl\ in Mapping.tabulate (vars_list_constraints cs) (\ var. val (\vl\ var) \)" global_interpretation SolveExec'Default: SolveExec' to_ns from_ns solve_exec_ns_code defines solve_exec_code = SolveExec'Default.solve_exec and solve_code = SolveExec'Default.solve proof unfold_locales { fix ics :: "'i i_constraint list" and v' and I let ?to_ns = "to_ns ics" let ?flat = "set ?to_ns" assume sat: "(I,\v'\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s ?flat" define cs where "cs = map snd (filter (\ ic. fst ic \ I) ics)" define to_ns' where to_ns: "to_ns' = (\ l. concat (map constraint_to_qdelta_constraint l))" show "(I,\from_ns v' (flat_list ?to_ns)\) \\<^sub>i\<^sub>c\<^sub>s set ics" unfolding i_satisfies_cs.simps proof let ?listf = "map (\C. case C of (LEQ_ns l r) \ (l\\v'\\, r) | (GEQ_ns l r) \ (r, l\\v'\\) )" let ?to_ns = "\ ics. to_ns' (map snd (filter (\ic. fst ic \ I) ics))" let ?list = "?listf (to_ns' cs)" (* index-filtered list *) let ?f_list = "flat_list (to_ns ics)" let ?flist = "?listf ?f_list" (* full list *) obtain i_list where i_list: "?list = i_list" by force obtain f_list where f_list: "?flist = f_list" by force have if_list: "set i_list \ set f_list" unfolding i_list[symmetric] f_list[symmetric] to_ns_def to_ns set_map set_concat cs_def by (intro image_mono, force) have "\ qd1 qd2. (qd1, qd2) \ set ?list \ qd1 \ qd2" proof- fix qd1 qd2 assume "(qd1, qd2) \ set ?list" then show "qd1 \ qd2" using sat unfolding cs_def proof(induct ics) case Nil then show ?case by (simp add: to_ns) next case (Cons h t) obtain i c where h: "h = (i,c)" by force from Cons(2) consider (ic) "(qd1,qd2) \ set (?listf (?to_ns [(i,c)]))" | (t) "(qd1,qd2) \ set (?listf (?to_ns t))" unfolding to_ns h set_map set_concat by fastforce then show ?case proof cases case t from Cons(1)[OF this] Cons(3) show ?thesis unfolding to_ns_def by auto next case ic note ic = ic[unfolded to_ns, simplified] from ic have i: "(i \ I) = True" by (cases "i \ I", auto) note ic = ic[unfolded i if_True, simplified] from Cons(3)[unfolded h] i have "\v'\ \\<^sub>n\<^sub>s\<^sub>s set (to_ns' [c])" unfolding i_satisfies_ns_constraints.simps unfolding to_ns to_ns_def by force with ic show ?thesis by (induct c) (auto simp add: to_ns) qed qed qed then have l1: "\ > 0 \ \ \ (\_min ?list) \ \qd1 qd2. (qd1, qd2) \ set ?list \ val qd1 \ \ val qd2 \" for \ unfolding i_list by (simp add: delta_gt_zero delta_min[of i_list]) have "\_min ?flist \ \_min ?list" unfolding f_list i_list by (rule delta_min_mono[OF if_list]) from l1[OF delta_gt_zero this] have l1: "\qd1 qd2. (qd1, qd2) \ set ?list \ val qd1 (\_min f_list) \ val qd2 (\_min f_list)" unfolding f_list . have "\0_val_min (flat_list (to_ns ics)) \v'\ = \_min f_list" unfolding f_list[symmetric] proof(induct ics) case Nil show ?case by (simp add: to_ns_def) next case (Cons h t) then show ?case by (cases h; cases "snd h") (auto simp add: to_ns_def) qed then have l2: "from_ns v' ?f_list = Mapping.tabulate (vars_list_constraints ?f_list) (\ var. val (\v'\ var) (\_min f_list))" by (auto simp add: from_ns_def) fix c assume "c \ restrict_to I (set ics)" then obtain i where i: "i \ I" and mem: "(i,c) \ set ics" by auto from mem show "\from_ns v' ?f_list\ \\<^sub>c c" proof (induct c) case (LT lll rrr) then have "(lll\\v'\\, (QDelta.QDelta rrr (-1))) \ set ?list" using i unfolding cs_def by (force simp add: to_ns) then have "val (lll\\v'\\) (\_min f_list) \ val (QDelta.QDelta rrr (-1)) (\_min f_list)" using l1 by simp moreover have "lll\(\x. val (\v'\ x) (\_min f_list))\ = lll\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars lll" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using LT by (auto simp add: comp_def lookup_tabulate restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "lll\\from_ns v' ?f_list\\ \ (val (QDelta.QDelta rrr (-1)) (\_min f_list))" by (auto simp add: valuate_rat_valuate) then show ?case using delta_gt_zero[of f_list] by (simp add: val_def) next case (GT lll rrr) then have "((QDelta.QDelta rrr 1), lll\\v'\\) \ set ?list" using i unfolding cs_def by (force simp add: to_ns) then have "val (lll\\v'\\) (\_min f_list) \ val (QDelta.QDelta rrr 1) (\_min f_list)" using l1 by simp moreover have "lll\(\x. val (\v'\ x) (\_min f_list))\ = lll\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars lll" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using GT by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "lll\\from_ns v' ?f_list\\ \ val (QDelta.QDelta rrr 1) (\_min f_list)" using l2 by (simp add: valuate_rat_valuate) then show ?case using delta_gt_zero[of f_list] by (simp add: val_def) next case (LEQ lll rrr) then have "(lll\\v'\\, (QDelta.QDelta rrr 0) ) \ set ?list" using i unfolding cs_def by (force simp add: to_ns) then have "val (lll\\v'\\) (\_min f_list) \ val (QDelta.QDelta rrr 0) (\_min f_list)" using l1 by simp moreover have "lll\(\x. val (\v'\ x) (\_min f_list))\ = lll\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars lll" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using LEQ by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "lll\\from_ns v' ?f_list\\ \ val (QDelta.QDelta rrr 0) (\_min f_list)" using l2 by (simp add: valuate_rat_valuate) then show ?case by (simp add: val_def) next case (GEQ lll rrr) then have "((QDelta.QDelta rrr 0), lll\\v'\\) \ set ?list" using i unfolding cs_def by (force simp add: to_ns) then have "val (lll\\v'\\) (\_min f_list) \ val (QDelta.QDelta rrr 0) (\_min f_list)" using l1 by simp moreover have "lll\(\x. val (\v'\ x) (\_min f_list))\ = lll\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars lll" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using GEQ by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "lll\\from_ns v' ?f_list\\ \ val (QDelta.QDelta rrr 0) (\_min f_list)" using l2 by (simp add: valuate_rat_valuate) then show ?case by (simp add: val_def) next case (EQ lll rrr) then have "((QDelta.QDelta rrr 0), lll\\v'\\) \ set ?list" and "(lll\\v'\\, (QDelta.QDelta rrr 0) ) \ set ?list" using i unfolding cs_def by (force simp add: to_ns)+ then have "val (lll\\v'\\) (\_min f_list) \ val (QDelta.QDelta rrr 0) (\_min f_list)" and "val (lll\\v'\\) (\_min f_list) \ val (QDelta.QDelta rrr 0) (\_min f_list)" using l1 by simp_all moreover have "lll\(\x. val (\v'\ x) (\_min f_list))\ = lll\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars lll" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using EQ by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "lll\\from_ns v' ?f_list\\ \ val (QDelta.QDelta rrr 0) (\_min f_list)" and "lll\\from_ns v' ?f_list\\ \ val (QDelta.QDelta rrr 0) (\_min f_list)" using l1 by (auto simp add: valuate_rat_valuate) then show ?case by (simp add: val_def) next case (LTPP ll1 ll2) then have "((ll1-ll2)\\v'\\, (QDelta.QDelta 0 (-1)) ) \ set ?list" using i unfolding cs_def by (force simp add: to_ns) then have "val ((ll1-ll2)\\v'\\) (\_min f_list) \ val (QDelta.QDelta 0 (-1)) (\_min f_list)" using l1 by simp moreover have "(ll1-ll2)\(\x. val (\v'\ x) (\_min f_list))\ = (ll1-ll2)\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars (ll1 - ll2)" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using LTPP by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "(ll1-ll2)\\from_ns v' ?f_list\\ \ val (QDelta.QDelta 0 (-1)) (\_min f_list)" using l1 by (simp add: valuate_rat_valuate) then show ?case using delta_gt_zero[of f_list] by (simp add: val_def valuate_minus) next case (GTPP ll1 ll2) then have "((QDelta.QDelta 0 1), (ll1-ll2)\\v'\\) \ set ?list" using i unfolding cs_def by (force simp add: to_ns) then have "val ((ll1-ll2)\\v'\\) (\_min f_list) \ val (QDelta.QDelta 0 1) (\_min f_list)" using l1 by simp moreover have "(ll1-ll2)\(\x. val (\v'\ x) (\_min f_list))\ = (ll1-ll2)\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars (ll1 - ll2)" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using GTPP by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "(ll1-ll2)\\from_ns v' ?f_list\\ \ val (QDelta.QDelta 0 1) (\_min f_list)" using l1 by (simp add: valuate_rat_valuate) then show ?case using delta_gt_zero[of f_list] by (simp add: val_def valuate_minus) next case (LEQPP ll1 ll2) then have "((ll1-ll2)\\v'\\, (QDelta.QDelta 0 0) ) \ set ?list" using i unfolding cs_def by (force simp add: to_ns zero_QDelta_def) then have "val ((ll1-ll2)\\v'\\) (\_min f_list) \ val (QDelta.QDelta 0 0) (\_min f_list)" using l1 by simp moreover have "(ll1-ll2)\(\x. val (\v'\ x) (\_min f_list))\ = (ll1-ll2)\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars (ll1 - ll2)" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using LEQPP by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "(ll1-ll2)\\from_ns v' ?f_list\\ \ val (QDelta.QDelta 0 0) (\_min f_list)" using l1 by (simp add: valuate_rat_valuate) then show ?case using delta_gt_zero[of f_list] by (simp add: val_def valuate_minus) next case (GEQPP ll1 ll2) then have "((QDelta.QDelta 0 0), (ll1-ll2)\\v'\\) \ set ?list" using i unfolding cs_def by (force simp add: to_ns zero_QDelta_def) then have "val ((ll1-ll2)\\v'\\) (\_min f_list) \ val (QDelta.QDelta 0 0) (\_min f_list)" using l1 by simp moreover have "(ll1-ll2)\(\x. val (\v'\ x) (\_min f_list))\ = (ll1-ll2)\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars (ll1 - ll2)" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using GEQPP by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "(ll1-ll2)\\from_ns v' ?f_list\\ \ val (QDelta.QDelta 0 0) (\_min f_list)" using l1 by (simp add: valuate_rat_valuate) then show ?case using delta_gt_zero[of f_list] by (simp add: val_def valuate_minus) next case (EQPP ll1 ll2) then have "((ll1-ll2)\\v'\\, (QDelta.QDelta 0 0) ) \ set ?list" and "((QDelta.QDelta 0 0), (ll1-ll2)\\v'\\) \ set ?list" using i unfolding cs_def by (force simp add: to_ns zero_QDelta_def)+ then have "val ((ll1-ll2)\\v'\\) (\_min f_list) \ val (QDelta.QDelta 0 0) (\_min f_list)" and "val ((ll1-ll2)\\v'\\) (\_min f_list) \ val (QDelta.QDelta 0 0) (\_min f_list)" using l1 by simp_all moreover have "(ll1-ll2)\(\x. val (\v'\ x) (\_min f_list))\ = (ll1-ll2)\\from_ns v' ?f_list\\" proof (rule valuate_depend, rule) fix x assume "x \ vars (ll1 - ll2)" then show "val (\v'\ x) (\_min f_list) = \from_ns v' ?f_list\ x" using l2 using EQPP by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def') qed ultimately have "(ll1-ll2)\\from_ns v' ?f_list\\ \ val (QDelta.QDelta 0 0) (\_min f_list)" and "(ll1-ll2)\\from_ns v' ?f_list\\ \ val (QDelta.QDelta 0 0) (\_min f_list)" using l1 by (auto simp add: valuate_rat_valuate) then show ?case by (simp add: val_def valuate_minus) qed qed } note sat = this fix cs :: "('i \ constraint) list" have set_to_ns: "set (to_ns cs) = { (i,n) | i n c. (i,c) \ set cs \ n \ set (constraint_to_qdelta_constraint c)}" unfolding to_ns_def by auto show indices: "fst ` set (to_ns cs) = fst ` set cs" proof show "fst ` set (to_ns cs) \ fst ` set cs" unfolding set_to_ns by force { fix i assume "i \ fst ` set cs" then obtain c where "(i,c) \ set cs" by force hence "i \ fst ` set (to_ns cs)" unfolding set_to_ns by (cases c; force) } thus "fst ` set cs \ fst ` set (to_ns cs)" by blast qed { assume dist: "distinct_indices cs" show "distinct_indices_ns (set (to_ns cs))" unfolding distinct_indices_ns_def proof (intro allI impI conjI notI) fix n1 n2 i assume "(i,n1) \ set (to_ns cs)" "(i,n2) \ set (to_ns cs)" then obtain c1 c2 where i: "(i,c1) \ set cs" "(i,c2) \ set cs" and n: "n1 \ set (constraint_to_qdelta_constraint c1)" "n2 \ set (constraint_to_qdelta_constraint c2)" unfolding set_to_ns by auto from dist have "distinct (map fst cs)" unfolding distinct_indices_def by auto with i have c12: "c1 = c2" by (metis eq_key_imp_eq_value) note n = n[unfolded c12] show "poly n1 = poly n2" using n by (cases c2, auto) show "ns_constraint_const n1 = ns_constraint_const n2" using n by (cases c2, auto) qed } note mini = this fix I mode assume unsat: "minimal_unsat_core_ns I (set (to_ns cs))" note unsat = unsat[unfolded minimal_unsat_core_ns_def indices] hence indices: "I \ fst ` set cs" by auto show "minimal_unsat_core I cs" unfolding minimal_unsat_core_def proof (intro conjI indices impI allI, clarify) fix v assume v: "(I,v) \\<^sub>i\<^sub>c\<^sub>s set cs" let ?v = "\var. QDelta.QDelta (v var) 0" have "(I,?v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s (set (to_ns cs))" using v proof(induct cs) case (Cons ic cs) obtain i c where ic: "ic = (i,c)" by force from Cons(2-) ic have rec: "(I,v) \\<^sub>i\<^sub>c\<^sub>s set cs" and c: "i \ I \ v \\<^sub>c c" by auto { fix jn assume i: "i \ I" and "jn \ set (to_ns [(i,c)])" then have "jn \ set (i_constraint_to_qdelta_constraint (i,c))" unfolding to_ns_def by auto then obtain n where n: "n \ set (constraint_to_qdelta_constraint c)" and jn: "jn = (i,n)" by force from c[OF i] have c: "v \\<^sub>c c" by force from c n jn have "?v \\<^sub>n\<^sub>s snd jn" by (cases c) (auto simp add: less_eq_QDelta_def to_ns_def valuate_valuate_rat valuate_minus zero_QDelta_def) } note main = this from Cons(1)[OF rec] have IH: "(I,?v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs)" . show ?case unfolding i_satisfies_ns_constraints.simps proof (intro ballI) fix x assume "x \ snd ` (set (to_ns (ic # cs)) \ I \ UNIV)" then consider (1) "x \ snd ` (set (to_ns cs) \ I \ UNIV)" | (2) "x \ snd ` (set (to_ns [(i,c)]) \ I \ UNIV)" unfolding ic to_ns_def by auto then show "?v \\<^sub>n\<^sub>s x" proof cases case 1 then show ?thesis using IH by auto next case 2 then obtain jn where x: "snd jn = x" and "jn \ set (to_ns [(i,c)]) \ I \ UNIV" by auto with main[of jn] show ?thesis unfolding to_ns_def by auto qed qed qed (simp add: to_ns_def) with unsat show False unfolding minimal_unsat_core_ns_def by simp blast next fix J assume *: "distinct_indices cs" "J \ I" hence "distinct_indices_ns (set (to_ns cs))" using mini by auto with * unsat obtain v where model: "(J, v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs)" by blast define w where "w = Mapping.Mapping (\ x. Some (v x))" have "v = \w\" unfolding w_def map2fun_def by (intro ext, transfer, auto) with model have model: "(J, \w\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs)" by auto from sat[OF this] show " \v. (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs" by blast qed qed (* cleanup *) hide_const (open) le lt LE GE LB UB LI UI LBI UBI UBI_upd le_rat inv zero Var add flat flat_list restrict_to look upd (* -------------------------------------------------------------------------- *) (* Main soundness lemma and executability *) (* -------------------------------------------------------------------------- *) text \Simplex version with indexed constraints as input\ definition simplex_index :: "'i i_constraint list \ 'i list + (var, rat) mapping" where "simplex_index = solve_exec_code" lemma simplex_index: "simplex_index cs = Unsat I \ set I \ fst ` set cs \ \ (\ v. (set I, v) \\<^sub>i\<^sub>c\<^sub>s set cs) \ (distinct_indices cs \ (\ J \ set I. (\ v. (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs)))" \ \minimal unsat core\ "simplex_index cs = Sat v \ \v\ \\<^sub>c\<^sub>s (snd ` set cs)" \ \satisfying assingment\ proof (unfold simplex_index_def) assume "solve_exec_code cs = Unsat I" from SolveExec'Default.simplex_unsat0[OF this] have core: "minimal_unsat_core (set I) cs" by auto then show "set I \ fst ` set cs \ \ (\ v. (set I, v) \\<^sub>i\<^sub>c\<^sub>s set cs) \ (distinct_indices cs \ (\J\set I. \v. (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs))" unfolding minimal_unsat_core_def by auto next assume "solve_exec_code cs = Sat v" from SolveExec'Default.simplex_sat0[OF this] show "\v\ \\<^sub>c\<^sub>s (snd ` set cs)" . qed text \Simplex version where indices will be created\ definition simplex where "simplex cs = simplex_index (zip [0.. \ (\ v. v \\<^sub>c\<^sub>s set cs)" \ \unsat of original constraints\ "simplex cs = Unsat I \ set I \ {0.. \ (\ v. v \\<^sub>c\<^sub>s {cs ! i | i. i \ set I}) \ (\J\set I. \v. v \\<^sub>c\<^sub>s {cs ! i |i. i \ J})" \ \minimal unsat core\ "simplex cs = Sat v \ \v\ \\<^sub>c\<^sub>s set cs" \ \satisfying assignment\ proof (unfold simplex_def) let ?cs = "zip [0.. {0 ..< length cs}" and core: "\v. v \\<^sub>c\<^sub>s (snd ` (set ?cs \ set I \ UNIV))" "(distinct_indices (zip [0.. (\ J \ set I. \v. v \\<^sub>c\<^sub>s (snd ` (set ?cs \ J \ UNIV))))" by (auto simp flip: set_map) note core(2) also have "distinct_indices (zip [0.. J \ set I. \v. v \\<^sub>c\<^sub>s (snd ` (set ?cs \ J \ UNIV))) = (\ J \ set I. \v. v \\<^sub>c\<^sub>s { cs ! i | i. i \ J})" using index by (intro all_cong1 imp_cong ex_cong1 arg_cong[of _ _ "\ x. _ \\<^sub>c\<^sub>s x"] refl, force simp: set_zip) finally have core': "(\J\set I. \v. v \\<^sub>c\<^sub>s {cs ! i |i. i \ J}) " . note unsat = unsat_mono[OF core(1)] show "\ (\ v. v \\<^sub>c\<^sub>s set cs)" by (rule unsat, auto simp: set_zip) show "set I \ {0.. \ (\ v. v \\<^sub>c\<^sub>s {cs ! i | i. i \ set I}) \ (\J\set I. \v. v \\<^sub>c\<^sub>s {cs ! i |i. i \ J})" by (intro conjI index core', rule unsat, auto simp: set_zip) next assume "simplex_index (zip [0..v\ \\<^sub>c\<^sub>s set cs" by (auto simp flip: set_map) qed text \check executability\ lemma "case simplex [LTPP (lp_monom 2 1) (lp_monom 3 2 - lp_monom 3 0), GT (lp_monom 1 1) 5] of Sat _ \ True | Unsat _ \ False" by eval text \check unsat core\ lemma "case simplex_index [ (1 :: int, LT (lp_monom 1 1) 4), (2, GTPP (lp_monom 2 1) (lp_monom 1 2)), (3, EQPP (lp_monom 1 1) (lp_monom 2 2)), (4, GT (lp_monom 2 2) 5), (5, GT (lp_monom 3 0) 7)] of Sat _ \ False | Unsat I \ set I = {1,3,4}" \ \Constraints 1,3,4 are unsat core\ by eval end \ No newline at end of file diff --git a/thys/Types_Tableaus_and_Goedels_God/AndersonProof.thy b/thys/Types_Tableaus_and_Goedels_God/AndersonProof.thy --- a/thys/Types_Tableaus_and_Goedels_God/AndersonProof.thy +++ b/thys/Types_Tableaus_and_Goedels_God/AndersonProof.thy @@ -1,243 +1,243 @@ (*<*) theory AndersonProof imports IHOML begin -nitpick_params[user_axioms=true, show_all, expect=genuine, format = 4, atoms e = a b c d] +nitpick_params[user_axioms=true, show_all, expect=genuine, format = 4, atoms e = a b c d, timeout=60] sledgehammer_params[verbose=true] (*>*) section \Anderson's Alternative\ text\ In this final section, we verify Anderson's emendation of G\"odel's argument, as it is presented in the last part of the textbook by Fitting (pp. 169-171). \ subsection \General Definitions\ abbreviation existencePredicate::"\\\\" ("E!") where "E! x \ \w. (\<^bold>\\<^sup>Ey. y\<^bold>\x) w" consts positiveProperty::"\\\\\\\" ("\

") abbreviation God::"\\\\" ("G\<^sup>A") where "G\<^sup>A \ \x. \<^bold>\Y. (\

Y) \<^bold>\ \<^bold>\(Y x)" abbreviation Entailment::"\\\\\\,\\\\\" (infix "\" 60) where "X \ Y \ \<^bold>\(\<^bold>\\<^sup>Ez. X z \<^bold>\ Y z)" subsection \Part I - God's Existence is Possible\ axiomatization where A1a:"\\<^bold>\X. \

(\<^bold>\X) \<^bold>\ \<^bold>\(\

X) \" and \ \Axiom 11.3A\ A2: "\\<^bold>\X Y. (\

X \<^bold>\ (X \ Y)) \<^bold>\ \

Y\" and \ \Axiom 11.5\ T2: "\\

G\<^sup>A\" \ \Proposition 11.16\ lemma True nitpick[satisfy] oops \ \model found: axioms are consistent\ theorem T1: "\\<^bold>\X. \

X \<^bold>\ \<^bold>\\<^bold>\\<^sup>E X\" using A1a A2 by blast \ \positive properties are possibly instantiated\ theorem T3: "\\<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" using T1 T2 by simp \ \God exists possibly\ subsection \Part II - God's Existence is Necessary if Possible\ text\ \\

\ now satisfies only one of the stability conditions. But since the argument uses an \emph{S5} logic, the other stability condition is implied. Therefore \\

\ becomes rigid (see p. 124). \ axiomatization where A4a: "\\<^bold>\X. \

X \<^bold>\ \<^bold>\(\

X)\" \ \axiom 11.11\ text\ We again postulate our \emph{S5} axioms: \ axiomatization where refl: "reflexive aRel" and tran: "transitive aRel" and symm: "symmetric aRel" lemma True nitpick[satisfy] oops \ \model found: so far all axioms consistent\ abbreviation rigidPred::"('t\io)\io" where "rigidPred \ \ (\\. \<^bold>\((\z. \ \<^bold>\ z) \<^bold>\\)) \<^bold>\\" lemma A4b: "\\<^bold>\X. \<^bold>\(\

X) \<^bold>\ \<^bold>\\<^bold>\(\

X)\" using A4a symm by auto \ \symmetry is needed (which corresponds to \emph{B} axiom)\ lemma "\rigidPred \

\" using A4a A4b by blast \ \@{text "\

"} is therefore rigid in a \emph{B} logic\ text\ Essence, Anderson Version (Definition 11.34) \ abbreviation essenceOf::"\\\\\\,\\" ("\\<^sup>A") where "\\<^sup>A Y x \ (\<^bold>\Z. \<^bold>\(Z x) \<^bold>\ Y \ Z)" text\ Necessary Existence, Anderson Version (Definition 11.35) \ abbreviation necessaryExistencePred::"\\\\" ("NE\<^sup>A") where "NE\<^sup>A x \ (\w. (\<^bold>\Y. \\<^sup>A Y x \<^bold>\ \<^bold>\\<^bold>\\<^sup>E Y) w)" text\ Theorem 11.36 - If g is God-like, then the property of being God-like is the essence of g. \ text\ As shown before, this theorem's proof could be completely automatized for G\"odel's and Fitting's variants. For Anderson's version however, we had to provide Isabelle with some help based on the corresponding natural-language proof given by Anderson (see @{cite "anderson90:_some_emend_of_goedel_ontol_proof"} Theorem 2*, p. 296) \ (*Anderson's Proof: Suppose that g is God-like* and necessarily has a property Q. Then by definition (of "God-like*"), that property is positive. But necessarily, if Q is positive, then if anything is God-like*, then it has Q -again by the definition of "God-like* ," together with the fact that if something has a property necessarily, then it has the property. But if a property is positive, then it is necessarily positive (Axiom 4). Hence, if Q is positive, then it is entailed by being God-like* (by modal logic-as in the original Theorem 2). But Q is positive and hence is entailed by being God-like*. Thus we have proved that if an entity is God-like* and has a property essentially, then that property is entailed by the property of being God-like*. Suppose a property Q is entailed by the property of being God-like*. Then Q is positive by Axioms 2 and 3* and therefore, since g is God-like*, g has Q necessarily (by the definition of "God-like*"). Hence, if something is God-like*, it has a property essentially if and only if that property is entailed by being God-like-i.e., God-likeness* is an essence* of that thing. Q.E.D.*) theorem GodIsEssential: "\\<^bold>\x. G\<^sup>A x \<^bold>\ (\\<^sup>A G\<^sup>A x)\" proof - { fix w { fix g { assume "G\<^sup>A g w" hence 1: "\Y. (\

Y w) \ (\<^bold>\(Y g)) w" by simp { fix Q from 1 have 2: "(\

Q w) \ (\<^bold>\(Q g)) w" by (rule allE) have "(\<^bold>\(Q g)) w \ (G\<^sup>A \ Q) w" \ \we need to prove @{text "\"} and @{text "\"}\ proof assume "(\<^bold>\(Q g)) w" \ \suppose g is God-like and necessarily has Q\ hence 3: "(\

Q w)" using 2 by simp \ \then Q is positive\ { fix u have "(\

Q u) \ (\x. G\<^sup>A x u \ (\<^bold>\(Q x)) u)" by auto \ \using the definition of God-like\ have "(\

Q u) \ (\x. G\<^sup>A x u \ ((Q x)) u)" using refl by auto \ \and using @{text "\(\ x) \ \ x"}\ } hence "\z. (\

Q z) \ (\x. G\<^sup>A x z \ Q x z)" by (rule allI) hence "\\

Q \<^bold>\ (\<^bold>\x. G\<^sup>A x \<^bold>\ Q x)\" by auto \ \if Q is positive, then whatever is God-like has Q\ hence "\\<^bold>\(\

Q \<^bold>\ (\<^bold>\x. G\<^sup>A x \<^bold>\ Q x))\" by (rule NEC) hence "\(\<^bold>\(\

Q)) \<^bold>\ \<^bold>\(\<^bold>\x. G\<^sup>A x \<^bold>\ Q x)\" using K by auto hence "\(\<^bold>\(\

Q)) \<^bold>\ G\<^sup>A \ Q\" by simp hence "((\<^bold>\(\

Q)) \<^bold>\ G\<^sup>A \ Q) w" by (rule allE) hence 4: "(\<^bold>\(\

Q)) w \ (G\<^sup>A \ Q) w" by simp (*if a property is necessarily positive, then it is entailed by being God-like*) have "\\<^bold>\X. \

X \<^bold>\ \<^bold>\(\

X)\" by (rule A4a) \ \using axiom 4\ hence "(\<^bold>\X. \

X \<^bold>\ (\<^bold>\(\

X))) w" by (rule allE) hence "\

Q w \ (\<^bold>\(\

Q)) w" by (rule allE) hence "\

Q w \ (G\<^sup>A \ Q) w" using 4 by simp (*if Q is positive, then it is entailed by being God-like*) thus "(G\<^sup>A \ Q) w" using 3 by (rule mp) \ \@{text "\"} direction\ next assume 5: "(G\<^sup>A \ Q) w" \ \suppose Q is entailed by being God-like\ have "\\<^bold>\X Y. (\

X \<^bold>\ (X \ Y)) \<^bold>\ \

Y\" by (rule A2) hence "(\<^bold>\X Y. (\

X \<^bold>\ (X \ Y)) \<^bold>\ \

Y) w" by (rule allE) hence "\X Y. (\

X w \ (X \ Y) w) \ \

Y w" by simp hence "\Y. (\

G\<^sup>A w \ (G\<^sup>A \ Y) w) \ \

Y w" by (rule allE) hence 6: "(\

G\<^sup>A w \ (G\<^sup>A \ Q) w) \ \

Q w" by (rule allE) have "\\

G\<^sup>A\" by (rule T2) hence "\

G\<^sup>A w" by (rule allE) hence "\

G\<^sup>A w \ (G\<^sup>A \ Q) w" using 5 by (rule conjI) from 6 this have "\

Q w" by (rule mp) \ \Q is positive by A2 and T2\ thus "(\<^bold>\(Q g)) w" using 2 by simp (*@{text "\"} direction *) qed } hence "\Z. (\<^bold>\(Z g)) w \ (G\<^sup>A \ Z) w" by (rule allI) hence "(\<^bold>\Z. \<^bold>\(Z g) \<^bold>\ G\<^sup>A \ Z) w" by simp hence "\\<^sup>A G\<^sup>A g w" by simp } hence "G\<^sup>A g w \ \\<^sup>A G\<^sup>A g w " by (rule impI) } hence "\x. G\<^sup>A x w \ \\<^sup>A G\<^sup>A x w " by (rule allI) } thus ?thesis by (rule allI) qed text\ Axiom 11.37 (Anderson's version of 11.25) \ axiomatization where A5: "\\

NE\<^sup>A\" lemma True nitpick[satisfy] oops \ \model found: so far all axioms consistent\ text\ Theorem 11.38 - Possibilist existence of God implies necessary actualist existence: \ theorem GodExistenceImpliesNecExistence: "\\<^bold>\ G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" proof - { fix w { assume "\x. G\<^sup>A x w" then obtain g where 1: "G\<^sup>A g w" .. hence "NE\<^sup>A g w" using A5 by blast \ \axiom 11.25\ hence "\Y. (\\<^sup>A Y g w) \ (\<^bold>\\<^bold>\\<^sup>E Y) w" by simp hence 2: "(\\<^sup>A G\<^sup>A g w) \ (\<^bold>\\<^bold>\\<^sup>E G\<^sup>A) w" by (rule allE) have "(\<^bold>\x. G\<^sup>A x \<^bold>\ (\\<^sup>A G\<^sup>A x)) w" using GodIsEssential by (rule allE) \ \GodIsEssential follows from Axioms 11.11 and 11.3B\ hence "(G\<^sup>A g \<^bold>\ (\\<^sup>A G\<^sup>A g)) w" by (rule allE) hence "G\<^sup>A g w \ \\<^sup>A G\<^sup>A g w" by blast from this 1 have 3: "\\<^sup>A G\<^sup>A g w" by (rule mp) from 2 3 have "(\<^bold>\\<^bold>\\<^sup>E G\<^sup>A) w" by (rule mp) } hence "(\x. G\<^sup>A x w) \ (\<^bold>\\<^bold>\\<^sup>E G\<^sup>A) w" by (rule impI) hence "((\<^bold>\x. G\<^sup>A x) \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A) w" by simp } thus ?thesis by (rule allI) qed text\ Some useful rules: \ lemma modal_distr: "\\<^bold>\(\ \<^bold>\ \)\ \ \(\<^bold>\\ \<^bold>\ \<^bold>\\)\" by blast lemma modal_trans: "(\\ \<^bold>\ \\ \ \\ \<^bold>\ \\) \ \\ \<^bold>\ \\" by simp text\ Anderson's version of Theorem 11.27 \ theorem possExistenceImpliesNecEx: "\\<^bold>\\<^bold>\ G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" \ \local consequence\ proof - have "\\<^bold>\ G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" using GodExistenceImpliesNecExistence by simp \ \follows from Axioms 11.11, 11.25 and 11.3B\ hence "\\<^bold>\(\<^bold>\ G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A)\" using NEC by simp hence 1: "\\<^bold>\\<^bold>\ G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" by (rule modal_distr) have 2: "\\<^bold>\\<^bold>\\<^bold>\\<^sup>E G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" using symm tran by metis from 1 2 have "\\<^bold>\\<^bold>\ G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^bold>\\<^sup>E G\<^sup>A\ \ \\<^bold>\\<^bold>\\<^bold>\\<^sup>E G\<^sup>A \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" by simp thus ?thesis by (rule modal_trans) qed lemma T4: "\\<^bold>\\<^bold>\ G\<^sup>A\ \ \\<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" using possExistenceImpliesNecEx by (rule localImpGlobalCons) \ \global consequence\ text\ Conclusion - Necessary (actualist) existence of God: \ lemma GodNecExists: "\\<^bold>\\<^bold>\\<^sup>E G\<^sup>A\" using T3 T4 by metis subsection \Modal Collapse\ text\ Modal collapse is countersatisfiable \ lemma "\\<^bold>\\.(\ \<^bold>\ (\<^bold>\ \))\" nitpick oops text\ \pagebreak \ section \Conclusion\ text\ We presented a shallow semantical embedding in Isabelle/HOL for an intensional higher-order modal logic (a successor of Montague/Gallin intensional logics) as introduced by M. Fitting in his textbook \emph{Types, Tableaus and G\"odel's God} @{cite "Fitting"}. We subsequently employed this logic to formalise and verify all results (theorems, examples and exercises) relevant to the discussion of G\"odel's ontological argument in the last part of Fitting's book. Three different versions of the ontological argument have been considered: the first one by G\"odel himself (respectively, Scott), the second one by Fitting and the last one by Anderson. \ text\ By employing an interactive theorem-prover like Isabelle, we were not only able to verify Fitting's results, but also to guarantee consistency. We could prove even stronger versions of many of the theorems and find better countermodels (i.e. with smaller cardinality) than the ones presented in the book. Another interesting aspect was the possibility to explore the implications of alternative formalisations for definitions and theorems which shed light on interesting philosophical issues concerning entailment, essentialism and free will, which are currently the subject of some follow-up analysis. \ text\ The latest developments in \emph{automated theorem proving} allow us to engage in much more experimentation during the formalisation and assessment of arguments than ever before. The potential reduction (of several orders of magnitude) in the time needed for proving or disproving theorems (compared to pen-and-paper proofs), results in almost real-time feedback about the suitability of our speculations. The practical benefits of computer-supported argumentation go beyond mere quantitative (easier, faster and more reliable proofs). The advantages are also qualitative, since it fosters a different approach to argumentation: We can now work iteratively (by `trial-and-error') on an argument by making gradual adjustments to its definitions, axioms and theorems. This allows us to continuously expose and revise the assumptions we indirectly commit ourselves everytime we opt for some particular formalisation. \pagebreak \ (*<*) end (*>*) diff --git a/thys/Types_Tableaus_and_Goedels_God/FittingProof.thy b/thys/Types_Tableaus_and_Goedels_God/FittingProof.thy --- a/thys/Types_Tableaus_and_Goedels_God/FittingProof.thy +++ b/thys/Types_Tableaus_and_Goedels_God/FittingProof.thy @@ -1,184 +1,184 @@ (*<*) theory FittingProof imports IHOML begin -nitpick_params[user_axioms=true, show_all, expect=genuine, format = 4, atoms e = a b c d] +nitpick_params[user_axioms=true, show_all, expect=genuine, format = 4, atoms e = a b c d, timeout=60] sledgehammer_params[verbose=true] (*>*) section \Fitting's Solution\ text\ In this section we consider Fitting's solution to the objections raised in his discussion of G\"odel's Argument pp. 164-9, especially the problem of \emph{modal collapse}, which has been metaphysically interpreted as implying a rejection of free will. Since we are generally commited to the existence of free will (in a pre-theoretical sense), such a result is philosophically unappealing and rather seen as a problem in the argument's formalisation. \ text\ This part of the book still leaves several details unspecified and the reader is thus compelled to fill in the gaps. As a result, we came across some premises and theorems allowing for different formalisations and therefore leading to disparate implications. Only some of those cases are shown here for illustrative purposes. The options we have chosen here are such that they indeed validate the argument (and we assume that they correspond to Fitting's intention. \ subsection \General Definitions\ text\ The following is an existence predicate for our object-language. (We have previously shown it is equivalent to its meta-logical counterpart.) \ abbreviation existencePredicate::"\\\\" ("E!") where "E! x \ (\w. (\<^bold>\\<^sup>Ey. y\<^bold>\x) w)" text\ Reminder: The `\\_\\' parenthesis are used to convert an extensional object into its `rigid' intensional counterpart (e.g. \\\\ \ \w. \\). \ consts positiveProperty::"\\\\\\" ("\

") abbreviation God::"\\\\" ("G") where "G \ (\x. \<^bold>\Y. \

Y \<^bold>\ \Y x\)" abbreviation God_star::"\\\\" ("G*") where "G* \ (\x. \<^bold>\Y. \

Y \<^bold>\ \Y x\)" abbreviation Entailment::"\\\\\,\\\\" (infix "\" 60) where "X \ Y \ \<^bold>\(\<^bold>\\<^sup>Ez. \X z\ \<^bold>\ \Y z\)" subsection \Part I - God's Existence is Possible\ axiomatization where A1a:"\\<^bold>\X. \

(\X) \<^bold>\ \<^bold>\(\

X) \" and \ \axiom 11.3A\ A1b:"\\<^bold>\X. \<^bold>\(\

X) \<^bold>\ \

(\X)\" and \ \axiom 11.3B\ A2: "\\<^bold>\X Y. (\

X \<^bold>\ (X \ Y)) \<^bold>\ \

Y\" and \ \axiom 11.5\ T2: "\\

\G\" \ \proposition 11.16 (modified)\ lemma True nitpick[satisfy] oops \ \model found: axioms are consistent\ lemma "\D\" using A1a A1b A2 by blast \ \axioms already imply \emph{D} axiom\ lemma GodDefsAreEquivalent: "\\<^bold>\x. G x \<^bold>\ G* x\" using A1b by fastforce text\ \bigbreak \ text\ \emph{T1} (Positive properties are possibly instantiated) can be formalised in two different ways: \ theorem T1a: "\\<^bold>\X::\\\. \

X \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. \X z\)\" using A1a A2 by blast \ \this is the one used in the book\ theorem T1b: "\\<^bold>\X::\\\\. \

\X \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. X z)\" nitpick oops \ \this one is also possible but not valid so we won't use it\ text\ Some interesting (non-)equivalences: \ lemma "\\<^bold>\\<^bold>\\<^sup>E (Q::\\\\) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E \<^bold>\Q)\" by simp lemma "\\<^bold>\\<^bold>\\<^sup>E (Q::\\\\) \<^bold>\ ((\X. \<^bold>\\<^bold>\\<^sup>E X) Q)\" by simp lemma "\\<^bold>\\<^bold>\\<^sup>E (Q::\\\\) \<^bold>\ ((\X. \<^bold>\\<^bold>\\<^sup>E \<^bold>\X) Q)\" by simp lemma "\\<^bold>\\<^bold>\\<^sup>E (Q::\\\\) \<^bold>\ ((\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\Q)\" nitpick oops \ \not equivalent!\ text\ \emph{T3} (God exists possibly) can be formalised in two different ways, using a \emph{de re} or a \emph{de dicto} reading. \ theorem T3_deRe: "\(\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G\" using T1a T2 by simp theorem T3_deDicto: "\\<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" nitpick oops \ \countersatisfiable\ text\ From the last two theorems, we think \T3_deRe\ should be the version originally implied in the book, since \T3_deDicto\ is not valid (\emph{T1b} were valid but it isn't) \ lemma assumes T1b: "\\<^bold>\X. \

\X \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. X z)\" shows T3_deDicto: "\\<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" using assms T2 by simp subsection \Part II - God's Existence is Necessary if Possible\ text\ In this variant \\

\ also designates rigidly, as shown in the last section. \ axiomatization where A4a: "\\<^bold>\X. \

X \<^bold>\ \<^bold>\(\

X)\" \ \axiom 11.11\ lemma A4b: "\\<^bold>\X. \<^bold>\(\

X) \<^bold>\ \<^bold>\\<^bold>\(\

X)\" using A1a A1b A4a by blast lemma True nitpick[satisfy] oops \ \model found: so far all axioms consistent\ abbreviation essenceOf::"\\\\\,\\" ("\") where "\ Y x \ \Y x\ \<^bold>\ (\<^bold>\Z::\\\. \Z x\ \<^bold>\ Y \ Z)" text\ Theorem 11.20 - Informal Proposition 5 \ theorem GodIsEssential: "\\<^bold>\x. G x \<^bold>\ ((\ \\<^sub>1G) x)\" using A1b by metis text\ Theorem 11.21 \ theorem God_starIsEssential: "\\<^bold>\x. G* x \<^bold>\ ((\ \\<^sub>1G*) x)\" by meson abbreviation necExistencePred:: "\\\\" ("NE") where "NE x \ \w. (\<^bold>\Y. \ Y x \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. \Y z\)) w" text\ \bigbreak \ text\ Informal Axiom 5 \ axiomatization where A5: "\\

\NE\" lemma True nitpick[satisfy] oops \ \model found: so far all axioms consistent\ text\ Reminder: We use \\G\ instead of \G\ because it is more explicit. See (non-)equivalences above. \ lemma "\\<^bold>\ G \<^bold>\ \<^bold>\ \<^bold>\G\" by simp lemma "\\<^bold>\\<^sup>E G \<^bold>\ \<^bold>\\<^sup>E \<^bold>\G\" by simp lemma "\\<^bold>\\<^bold>\\<^sup>E G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" by simp text\ Theorem 11.26 (Informal Proposition 7) - (possibilist) existence of God implies necessary (actualist) existence. \ text\ There are two different ways of formalising this theorem. Both of them are proven valid: \ text\ First version: \ theorem GodExImpliesNecEx_v1: "\\<^bold>\ \<^bold>\G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" proof - { fix w { assume "\x. G x w" then obtain g where 1: "G g w" .. hence "NE g w" using A5 by auto hence "\Y. (\ Y g w) \ (\<^bold>\(\<^bold>\\<^sup>Ez. \Y z\)) w" by simp hence "(\ (\x. G x w) g w) \ (\<^bold>\(\<^bold>\\<^sup>Ez. \(\x. G x w) z\)) w" by (rule allE) hence 2: "((\ \\<^sub>1G) g w) \ (\<^bold>\(\<^bold>\\<^sup>E G)) w" using A4b by meson have "(\<^bold>\x. G x \<^bold>\ ((\ \\<^sub>1G) x)) w" using GodIsEssential by (rule allE) hence "(G g \<^bold>\ ((\ \\<^sub>1G) g)) w" by (rule allE) hence "G g w \ (\ \\<^sub>1G) g w" by simp from this 1 have 3: "(\ \\<^sub>1G) g w" by (rule mp) from 2 3 have "(\<^bold>\\<^bold>\\<^sup>E G) w" by (rule mp) } hence "(\x. G x w) \ (\<^bold>\\<^bold>\\<^sup>E G) w" by (rule impI) hence "((\<^bold>\x. G x) \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G) w" by simp } thus ?thesis by (rule allI) qed text\ Second version (which can be proven directly by automated tools using the previous version): \ theorem GodExImpliesNecEx_v2: "\\<^bold>\ \<^bold>\G \<^bold>\ ((\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G)\" using A4a GodExImpliesNecEx_v1 by metis text\ In contrast to G\"odel's argument (as presented by Fitting), the following theorems can be proven in logic \emph{K} (the \emph{S5} axioms are no longer needed): \ text\ Theorem 11.27 - Informal Proposition 8 \ theorem possExImpliesNecEx_v1: "\\<^bold>\\<^bold>\ \<^bold>\G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" using GodExImpliesNecEx_v1 T3_deRe by metis theorem possExImpliesNecEx_v2: "\(\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G \<^bold>\ ((\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G)\" using GodExImpliesNecEx_v2 by blast text\ Corollaries: \ lemma T4_v1: "\\<^bold>\\<^bold>\ \<^bold>\G\ \ \\<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" using possExImpliesNecEx_v1 by simp lemma T4_v2: "\(\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G\ \ \(\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G\" using possExImpliesNecEx_v2 by simp subsection \Conclusion (\emph{De Re} and \emph{De Dicto} Reading)\ text\ Version I - Necessary Existence of God (\emph{de dicto}): \ lemma GodNecExists_v1: "\\<^bold>\\<^bold>\\<^sup>E \<^bold>\G\" using GodExImpliesNecEx_v1 T3_deRe by fastforce \ \corollary 11.28\ lemma God_starNecExists_v1: "\\<^bold>\\<^bold>\\<^sup>E \<^bold>\G*\" using GodNecExists_v1 GodDefsAreEquivalent by simp lemma "\\<^bold>\(\X. \<^bold>\\<^sup>E X) \<^bold>\G*\" using God_starNecExists_v1 by simp \ \\emph{de dicto} shown here explicitly\ text\ Version II - Necessary Existence of God (\emph{de re}) \ lemma GodNecExists_v2: "\(\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G\" using T3_deRe T4_v2 by blast lemma God_starNecExists_v2: "\(\X. \<^bold>\\<^bold>\\<^sup>E X) \<^bold>\G*\" using GodNecExists_v2 GodDefsAreEquivalent by simp subsection \Modal Collapse\ text\ Modal collapse is countersatisfiable even in \emph{S5}. Note that countermodels with a cardinality of one for the domain of individuals are found by \emph{Nitpick} (the countermodel shown in the book has cardinality of two). \ lemma "\\<^bold>\\.(\ \<^bold>\ (\<^bold>\ \))\" nitpick[card 't=1, card i=2] oops \ \countermodel found in \emph{K}\ axiomatization where S5: "equivalence aRel" \ \assume \emph{S5} logic\ lemma "\\<^bold>\\.(\ \<^bold>\ (\<^bold>\ \))\" nitpick[card 't=1, card i=2] oops \ \countermodel also found in \emph{S5}\ text\ \pagebreak \ (*<*) end (*>*) diff --git a/thys/Types_Tableaus_and_Goedels_God/GoedelProof_P2.thy b/thys/Types_Tableaus_and_Goedels_God/GoedelProof_P2.thy --- a/thys/Types_Tableaus_and_Goedels_God/GoedelProof_P2.thy +++ b/thys/Types_Tableaus_and_Goedels_God/GoedelProof_P2.thy @@ -1,298 +1,298 @@ (*<*) theory GoedelProof_P2 imports IHOML begin -nitpick_params[user_axioms=true, show_all, expect=genuine, format = 3, atoms e = a b c d] +nitpick_params[user_axioms=true, show_all, expect=genuine, format = 3, atoms e = a b c d, timeout=60] sledgehammer_params[verbose=true] (*>*) subsection \Part II - God's Existence is Necessary if Possible\ text\ We show here that God's necessary existence follows from its possible existence by adding some additional (potentially controversial) assumptions including an \emph{essentialist} premise and the \emph{S5} axioms. Further results like monotheism and the rejection of free will (\emph{modal collapse}) are also proved. \ subsubsection \General Definitions\ abbreviation existencePredicate::"\\\\" ("E!") where "E! x \ (\w. (\<^bold>\\<^sup>Ey. y\<^bold>\x) w)" consts positiveProperty::"\\\\\\\" ("\

") abbreviation God::"\\\\" ("G") where "G \ (\x. \<^bold>\Y. \

Y \<^bold>\ Y x)" abbreviation God_star::"\\\\" ("G*") where "G* \ (\x. \<^bold>\Y. \

Y \<^bold>\ Y x)" abbreviation Entailment::"\\\\\\,\\\\\" (infix "\" 60) where "X \ Y \ \<^bold>\(\<^bold>\\<^sup>Ez. X z \<^bold>\ Y z)" subsubsection \Results from Part I\ text\ Note that the only use G\"odel makes of axiom A3 is to show that being Godlike is a positive property (\emph{T2}). We follow therefore Scott's proposal and take (\emph{T2}) directly as an axiom: \ axiomatization where A1a:"\\<^bold>\X. \

(\<^bold>\X) \<^bold>\ \<^bold>\(\

X) \" and \ \axiom 11.3A\ A1b:"\\<^bold>\X. \<^bold>\(\

X) \<^bold>\ \

(\<^bold>\X)\" and \ \axiom 11.3B\ A2: "\\<^bold>\X Y. (\

X \<^bold>\ (X \ Y)) \<^bold>\ \

Y\" and \ \axiom 11.5\ T2: "\\

G\" \ \proposition 11.16\ lemma True nitpick[satisfy] oops \ \model found: axioms are consistent\ lemma "\D\" using A1a A1b A2 by blast \ \axioms already imply \emph{D} axiom\ lemma GodDefsAreEquivalent: "\\<^bold>\x. G x \<^bold>\ G* x\" using A1b by fastforce theorem T1: "\\<^bold>\X. \

X \<^bold>\ \<^bold>\\<^bold>\\<^sup>E X\" using A1a A2 by blast \ \positive properties are possibly instantiated\ theorem T3: "\\<^bold>\\<^bold>\\<^sup>E G\" using T1 T2 by simp \ \God exists possibly\ subsubsection \Axioms\ text\ \\

\ satisfies the so-called stability conditions (see @{cite "Fitting"}, p. 124), which means it designates rigidly (note that this makes for an \emph{essentialist} assumption). \ axiomatization where A4a: "\\<^bold>\X. \

X \<^bold>\ \<^bold>\(\

X)\" \ \axiom 11.11\ lemma A4b: "\\<^bold>\X. \<^bold>\(\

X) \<^bold>\ \<^bold>\\<^bold>\(\

X)\" using A1a A1b A4a by blast abbreviation rigidPred::"('t\io)\io" where "rigidPred \ \ (\\. \<^bold>\((\z. \ \<^bold>\ z) \<^bold>\\)) \<^bold>\\" lemma "\rigidPred \

\" using A4a A4b by blast \ \@{term "\

"} is therefore rigid\ lemma True nitpick[satisfy] oops \ \model found: so far all axioms A1-4 consistent\ text\ \bigbreak \ subsubsection \Theorems\ text\ Remark: Essence is defined here (and in Fitting's variant) in the version of Scott; G\"odel's original version leads to the inconsistency reported in @{cite C55 and C60} \ abbreviation essenceOf::"\\\\\\,\\" ("\") where "\ Y x \ (Y x) \<^bold>\ (\<^bold>\Z. Z x \<^bold>\ Y \ Z)" abbreviation beingIdenticalTo::"\\\\\\" ("id") where "id x \ (\y. y\<^bold>\x)" \ \note that \emph{id} is a rigid predicate\ text\ Theorem 11.20 - Informal Proposition 5 \ theorem GodIsEssential: "\\<^bold>\x. G x \<^bold>\ (\ G x)\" using A1b A4a by metis text\ Theorem 11.21 \ theorem "\\<^bold>\x. G* x \<^bold>\ (\ G* x)\" using A4a by meson text\ Theorem 11.22 - Something can have only one essence: \ theorem "\\<^bold>\X Y z. (\ X z \<^bold>\ \ Y z) \<^bold>\ (X \ Y)\" by meson text\ Theorem 11.23 - An essence is a complete characterization of an individual: \ theorem EssencesCharacterizeCompletely: "\\<^bold>\X y. \ X y \<^bold>\ (X \ (id y))\" proof (rule ccontr) assume "\ \\<^bold>\X y. \ X y \<^bold>\ (X \ (id y))\" hence "\w. \(( \<^bold>\X y. \ X y \<^bold>\ X \ id y) w)" by simp then obtain w where "\(( \<^bold>\X y. \ X y \<^bold>\ X \ id y) w)" .. hence "(\<^bold>\X y. \ X y \<^bold>\ \<^bold>\(X \ id y)) w" by simp hence "\X y. \ X y w \ (\<^bold>\(X \ id y)) w" by simp then obtain P where "\y. \ P y w \ (\<^bold>\(P \ id y)) w" .. then obtain a where 1: "\ P a w \ (\<^bold>\(P \ id a)) w" .. hence 2: "\ P a w" by (rule conjunct1) from 1 have "(\<^bold>\(P \ id a)) w" by (rule conjunct2) hence "\x. \z. w r x \ existsAt z x \ P z x \ \(a = z)" by blast then obtain w1 where "\z. w r w1 \ existsAt z w1 \ P z w1 \ \(a = z)" .. then obtain b where 3: "w r w1 \ existsAt b w1 \ P b w1 \ \(a = b)" .. hence "w r w1" by simp from 3 have "existsAt b w1" by simp from 3 have "P b w1" by simp from 3 have 4: " \(a = b)" by simp from 2 have "P a w" by simp from 2 have "\Y. Y a w \ ((P \ Y) w)" by auto hence "(\<^bold>\(id b)) a w \ (P \ (\<^bold>\(id b))) w" by (rule allE) hence "\(\<^bold>\(id b)) a w \ ((P \ (\<^bold>\(id b))) w)" by blast then show False proof assume "\(\<^bold>\(id b)) a w" hence "a = b" by simp thus False using 4 by auto next assume "((P \ (\<^bold>\(id b))) w)" hence "\x. \z. (w r x \ existsAt z x \ P z x) \ (\<^bold>\(id b)) z x" by blast hence "\z. (w r w1 \ existsAt z w1 \ P z w1) \ (\<^bold>\(id b)) z w1" by (rule allE) hence "(w r w1 \ existsAt b w1 \ P b w1) \ (\<^bold>\(id b)) b w1" by (rule allE) hence "\(w r w1 \ existsAt b w1 \ P b w1) \ (\<^bold>\(id b)) b w1" by simp hence "(\<^bold>\(id b)) b w" using 3 by simp hence "\(b=b)" by simp thus False by simp qed qed text\ Definition 11.24 - Necessary Existence (Informal Definition 6): \ abbreviation necessaryExistencePred::"\\\\" ("NE") where "NE x \ (\w. (\<^bold>\Y. \ Y x \<^bold>\ \<^bold>\\<^bold>\\<^sup>E Y) w)" text\ Axiom 11.25 (Informal Axiom 5) \ axiomatization where A5: "\\

NE\" lemma True nitpick[satisfy] oops \ \model found: so far all axioms consistent\ text\ Theorem 11.26 (Informal Proposition 7) - Possibilist existence of God implies necessary actualist existence: \ theorem GodExistenceImpliesNecExistence: "\\<^bold>\ G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\" proof - { fix w { assume "\x. G x w" then obtain g where 1: "G g w" .. hence "NE g w" using A5 by auto \ \axiom 11.25\ hence "\Y. (\ Y g w) \ (\<^bold>\\<^bold>\\<^sup>E Y) w" by simp hence 2: "(\ G g w) \ (\<^bold>\\<^bold>\\<^sup>E G) w" by (rule allE) have "(\<^bold>\x. G x \<^bold>\ (\ G x)) w" using GodIsEssential by (rule allE) \ \GodIsEssential follows from Axioms 11.11 and 11.3B\ hence "(G g \<^bold>\ (\ G g)) w" by (rule allE) hence "G g w \ \ G g w" by simp from this 1 have 3: "\ G g w" by (rule mp) from 2 3 have "(\<^bold>\\<^bold>\\<^sup>E G) w" by (rule mp) } hence "(\x. G x w) \ (\<^bold>\\<^bold>\\<^sup>E G) w" by (rule impI) hence "((\<^bold>\x. G x) \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G) w" by simp } thus ?thesis by (rule allI) qed text\ \emph{Modal collapse} is countersatisfiable (unless we introduce S5 axioms): \ lemma "\\<^bold>\\.(\ \<^bold>\ (\<^bold>\ \))\" nitpick oops text\ We postulate semantic frame conditions for some modal logics. Taken together, reflexivity, transitivity and symmetry make for an equivalence relation and therefore an \emph{S5} logic (via \emph{Sahlqvist correspondence}). We prefer to postulate them individually here in order to get more detailed information about their relevance in the proofs presented below. \ axiomatization where refl: "reflexive aRel" and tran: "transitive aRel" and symm: "symmetric aRel" lemma True nitpick[satisfy] oops \ \model found: axioms still consistent\ text\ Using an \emph{S5} logic, \emph{modal collapse} (\\\<^bold>\\.(\ \<^bold>\ (\<^bold>\ \))\\) is actually valid (see `More Objections' some pages below) \ text\ We prove some useful inference rules: \ lemma modal_distr: "\\<^bold>\(\ \<^bold>\ \)\ \ \(\<^bold>\\ \<^bold>\ \<^bold>\\)\" by blast lemma modal_trans: "(\\ \<^bold>\ \\ \ \\ \<^bold>\ \\) \ \\ \<^bold>\ \\" by simp text\ Theorem 11.27 - Informal Proposition 8. Note that only symmetry and transitivity for the accessibility relation are used. \ theorem possExistenceImpliesNecEx: "\\<^bold>\\<^bold>\ G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\" \ \local consequence\ proof - have "\\<^bold>\ G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\" using GodExistenceImpliesNecExistence by simp \ \follows from Axioms 11.11, 11.25 and 11.3B\ hence "\\<^bold>\(\<^bold>\ G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G)\" using NEC by simp hence 1: "\\<^bold>\\<^bold>\ G \<^bold>\ \<^bold>\\<^bold>\\<^bold>\\<^sup>E G\" by (rule modal_distr) have 2: "\\<^bold>\\<^bold>\\<^bold>\\<^sup>E G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\" using symm tran by metis \ \frame conditions\ from 1 2 have "\\<^bold>\\<^bold>\ G \<^bold>\ \<^bold>\\<^bold>\\<^bold>\\<^sup>E G\ \ \\<^bold>\\<^bold>\\<^bold>\\<^sup>E G \<^bold>\ \<^bold>\\<^bold>\\<^sup>E G\" by simp thus ?thesis by (rule modal_trans) qed lemma T4: "\\<^bold>\\<^bold>\ G\ \ \\<^bold>\\<^bold>\\<^sup>E G\" using possExistenceImpliesNecEx by (rule localImpGlobalCons) \ \global consequence\ text\ Corollary 11.28 - Necessary (actualist) existence of God (for both definitions); reflexivity is still not used: \ lemma GodNecExists: "\\<^bold>\\<^bold>\\<^sup>E G\" using T3 T4 by metis lemma God_starNecExists: "\\<^bold>\\<^bold>\\<^sup>E G*\" using GodNecExists GodDefsAreEquivalent by simp subsubsection \Monotheism\ text\ Monotheism for non-normal models (with Leibniz equality) follows directly from God having all and only positive properties: \ theorem Monotheism_LeibnizEq: "\\<^bold>\x. G x \<^bold>\ (\<^bold>\y. G y \<^bold>\ (x \<^bold>\\<^sup>L y))\" using GodDefsAreEquivalent by simp text\ Monotheism for normal models is trickier. We need to consider some previous results (p. 162): \ lemma GodExistenceIsValid: "\\<^bold>\\<^sup>E G\" using GodNecExists refl by auto \ \reflexivity is now required by the solver\ text\ Proposition 11.29: \ theorem Monotheism_normalModel: "\\<^bold>\x.\<^bold>\y. G y \<^bold>\ x \<^bold>\ y\" proof - { fix w have "\\<^bold>\\<^sup>E G\" using GodExistenceIsValid by simp \ \follows from corollary 11.28\ hence "(\<^bold>\\<^sup>E G) w" by (rule allE) then obtain g where 1: "existsAt g w \ G g w" .. hence 2: "\ G g w" using GodIsEssential by blast \ \follows from ax. 11.11/11.3B\ { fix y have "G y w \ (g \<^bold>\ y) w" proof assume "G y w" hence 3: "\ G y w" using GodIsEssential by blast have "(\ G y \<^bold>\ (G \ id y)) w" using EssencesCharacterizeCompletely by simp \ \follows from theorem 11.23\ hence "\ G y w \ ((G \ id y) w)" by simp from this 3 have "(G \ id y) w" by (rule mp) hence "(\<^bold>\(\<^bold>\\<^sup>Ez. G z \<^bold>\ z \<^bold>\ y)) w" by simp hence "\x. w r x \ ((\z. (existsAt z x \ G z x) \ z = y))" by auto hence "w r w \ ((\z. (existsAt z w \ G z w) \ z = y))" by (rule allE) hence "\z. (w r w \ existsAt z w \ G z w) \ z = y" by auto hence 4: "(w r w \ existsAt g w \ G g w) \ g = y" by (rule allE) have "w r w" using refl by simp \ \using frame reflexivity (Axiom M)\ hence "w r w \ (existsAt g w \ G g w)" using 1 by (rule conjI) from 4 this have "g = y" by (rule mp) thus "(g \<^bold>\ y) w" by simp next assume "(g \<^bold>\ y) w" from this 2 have "\ G y w" by simp thus "G y w " by (rule conjunct1) qed } hence "\y. G y w \ (g \<^bold>\ y) w" by (rule allI) hence "\x. (\y. G y w \ (x \<^bold>\ y) w)" by (rule exI) hence "(\<^bold>\x. (\<^bold>\y. G y \<^bold>\ (x \<^bold>\ y))) w" by simp } thus ?thesis by (rule allI) qed text\ \bigbreak \ text\ Corollary 11.30: \ lemma GodImpliesExistence: "\\<^bold>\x. G x \<^bold>\ E! x\" using GodExistenceIsValid Monotheism_normalModel by metis subsubsection \Positive Properties are Necessarily Instantiated\ lemma PosPropertiesNecExist: "\\<^bold>\Y. \

Y \<^bold>\ \<^bold>\\<^bold>\\<^sup>E Y\" using GodNecExists A4a by meson \ \proposition 11.31: follows from corollary 11.28 and axiom A4a\ subsubsection \More Objections\ text\ Fitting discusses the objection raised by Sobel @{cite "sobel2004logic"}, who argues that G\"odel's axiom system is too strong: it implies that whatever is the case is so necessarily, i.e. the modal system collapses (\\ \ \\\). The \emph{modal collapse} has been philosophically interpreted as implying the absence of free will. \ text\ We start by proving an useful FOL lemma: \ lemma useful: "(\x. \ x \ \) \ ((\x. \ x) \ \)" by simp text\ In the context of our S5 axioms, the \emph{modal collapse} becomes valid (pp. 163-4): \ lemma ModalCollapse: "\\<^bold>\\.(\ \<^bold>\ (\<^bold>\ \))\" proof - { fix w { fix Q have "(\<^bold>\x. G x \<^bold>\ (\ G x)) w" using GodIsEssential by (rule allE) \ \follows from Axioms 11.11 and 11.3B\ hence "\x. G x w \ \ G x w" by simp hence "\x. G x w \ (\<^bold>\Z. Z x \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. G z \<^bold>\ Z z)) w" by force hence "\x. G x w \ ((\y. Q) x \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. G z \<^bold>\ (\y. Q) z)) w" by force hence "\x. G x w \ (Q \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. G z \<^bold>\ Q)) w" by simp hence 1: "(\x. G x w) \ ((Q \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. G z \<^bold>\ Q)) w)" by (rule useful) have "\x. G x w" using GodExistenceIsValid by auto from 1 this have "(Q \<^bold>\ \<^bold>\(\<^bold>\\<^sup>Ez. G z \<^bold>\ Q)) w" by (rule mp) hence "(Q \<^bold>\ \<^bold>\((\<^bold>\\<^sup>Ez. G z) \<^bold>\ Q)) w" using useful by blast hence "(Q \<^bold>\ (\<^bold>\(\<^bold>\\<^sup>Ez. G z) \<^bold>\ \<^bold>\Q)) w" by simp hence "(Q \<^bold>\ \<^bold>\Q) w" using GodNecExists by simp } hence "(\<^bold>\\. \ \<^bold>\ \<^bold>\ \) w" by (rule allI) } thus ?thesis by (rule allI) qed text\ \pagebreak \ (*<*) end (*>*)