diff --git a/thys/Blue_Eyes/Blue_Eyes.thy b/thys/Blue_Eyes/Blue_Eyes.thy --- a/thys/Blue_Eyes/Blue_Eyes.thy +++ b/thys/Blue_Eyes/Blue_Eyes.thy @@ -1,424 +1,424 @@ (*<*) theory Blue_Eyes imports Main begin (*>*) section \Introduction\ text \The original problem statement @{cite xkcd} explains the puzzle well: \begin{quotation} A group of people with assorted eye colors live on an island. They are all perfect logicians -- if a conclusion can be logically deduced, they will do it instantly. No one knows the color of their eyes. Every night at midnight, a ferry stops at the island. Any islanders who have figured out the color of their own eyes then leave the island, and the rest stay. Everyone can see everyone else at all times and keeps a count of the number of people they see with each eye color (excluding themselves), but they cannot otherwise communicate. Everyone on the island knows all the rules in this paragraph. On this island there are 100 blue-eyed people, 100 brown-eyed people, and the Guru (she happens to have green eyes). So any given blue-eyed person can see 100 people with brown eyes and 99 people with blue eyes (and one with green), but that does not tell him his own eye color; as far as he knows the totals could be 101 brown and 99 blue. Or 100 brown, 99 blue, and he could have red eyes. The Guru is allowed to speak once (let's say at noon), on one day in all their endless years on the island. Standing before the islanders, she says the following: ``I can see someone who has blue eyes.'' Who leaves the island, and on what night? \end{quotation} It might seem weird that the Guru's declaration gives anyone any new information. For an informal discussion, see \cite[Section~1.1]{fagin1995}.\ section \Modeling the world \label{sec:world}\ text \We begin by fixing two type variables: @{typ "'color"} and @{typ "'person"}. The puzzle doesn't specify how many eye colors are possible, but four are mentioned. Crucially, we must assume they are distinct. We specify the existence of colors other than blue and brown, even though we don't mention them later, because when blue and brown are the only possible colors, the puzzle has a different solution — the brown-eyed logicians may leave one day after the blue-eyed ones. We refrain from specifying the exact population of the island, choosing to only assume it is finite and denote a specific person as the Guru. We could also model the Guru as an outside entity instead of a participant. This doesn't change the answer and results in a slightly simpler proof, but is less faithful to the problem statement.\ context fixes blue brown green red :: 'color assumes colors_distinct: "distinct [blue, brown, green, red]" fixes guru :: 'person assumes "finite (UNIV :: 'person set)" begin text \It's slightly tricky to formalize the behavior of perfect logicians. The representation we use is centered around the type of a @{emph \world\}, which describes the entire state of the environment. In our case, it's a function @{typ "'person => 'color"} that assigns an eye color to everyone.@{footnote \We would introduce a type synonym, but at the time of writing Isabelle doesn't support including type variables fixed by a locale in a type synonym.\} The only condition known to everyone and not dependent on the observer is Guru's declaration:\ definition valid :: "('person \ 'color) \ bool" where "valid w \ (\p. p \ guru \ w p = blue)" text \We then define the function @{term "possible n p w w'"}, which returns @{term True} if on day \n\ the potential world \w'\ is plausible from the perspective of person \p\, based on the observations they made in the actual world \w\. Then, @{term "leaves n p w"} is @{term True} if \p\ is able to unambiguously deduce the color of their own eyes, i.e. if it is the same in all possible worlds. Note that if \p\ actually left many moons ago, this function still returns @{term True}.\ fun leaves :: "nat \ 'person \ ('person \ 'color) \ bool" and possible :: "nat \ 'person \ ('person \ 'color) \ ('person \ 'color) \ bool" where "leaves n p w = (\w'. possible n p w w' \ w' p = w p)" | "possible n p w w' \ valid w \ valid w' \ (\p' \ p. w p' = w' p') \ (\n' < n. \p'. leaves n' p' w = leaves n' p' w')" text \Naturally, the act of someone leaving can be observed by others, thus the two definitions are mutually recursive. As such, we need to instruct the simplifier to not unfold these definitions endlessly.\ declare possible.simps[simp del] leaves.simps[simp del] text \A world is possible if \<^enum> The Guru's declaration holds. \<^enum> The eye color of everyone but the observer matches. \<^enum> The same people left on each of the previous days. Moreover, we require that the actual world \w\ is \valid\, so that the relation is symmetric:\ lemma possible_sym: "possible n p w w' = possible n p w' w" by (auto simp: possible.simps) text \In fact, \possible n p\ is an equivalence relation:\ lemma possible_refl: "valid w \ possible n p w w" by (auto simp: possible.simps) lemma possible_trans: "possible n p w1 w2 \ possible n p w2 w3 \ possible n p w1 w3" by (auto simp: possible.simps) section \Eye colors other than blue\ text \Since there is no way to distinguish between the colors other than blue, only the blue-eyed people will ever leave. To formalize this notion, we define a function that takes a world and replaces the eye color of a specified person. The original color is specified too, so that the transformation composes nicely with the recursive hypothetical worlds of @{const possible}.\ definition try_swap :: "'person \ 'color \ 'color \ ('person \ 'color) \ ('person \ 'color)" where "try_swap p c\<^sub>1 c\<^sub>2 w x = (if c\<^sub>1 = blue \ c\<^sub>2 = blue \ x \ p then w x else Fun.swap c\<^sub>1 c\<^sub>2 id (w x))" lemma try_swap_valid[simp]: "valid (try_swap p c\<^sub>1 c\<^sub>2 w) = valid w" - by (auto simp add: try_swap_def valid_def swap_def) + by (auto simp add: try_swap_def valid_def swap_id_eq) lemma try_swap_eq[simp]: "try_swap p c\<^sub>1 c\<^sub>2 w x = try_swap p c\<^sub>1 c\<^sub>2 w' x \ w x = w' x" - by (auto simp add: try_swap_def swap_def) + by (auto simp add: try_swap_def swap_id_eq) lemma try_swap_inv[simp]: "try_swap p c\<^sub>1 c\<^sub>2 (try_swap p c\<^sub>1 c\<^sub>2 w) = w" - by (rule ext) (auto simp add: try_swap_def swap_def) + by (rule ext) (auto simp add: try_swap_def swap_id_eq) lemma leaves_try_swap[simp]: assumes "valid w" shows "leaves n p (try_swap p' c\<^sub>1 c\<^sub>2 w) = leaves n p w" using assms proof (induction n arbitrary: p w rule: less_induct) case (less n) have "leaves n p w" if "leaves n p (try_swap p' c\<^sub>1 c\<^sub>2 w)" for w proof (unfold leaves.simps; rule+) fix w' assume "possible n p w w'" then have "possible n p (try_swap p' c\<^sub>1 c\<^sub>2 w) (try_swap p' c\<^sub>1 c\<^sub>2 w')" by (fastforce simp: possible.simps less.IH) with `leaves n p (try_swap p' c\<^sub>1 c\<^sub>2 w)` have "try_swap p' c\<^sub>1 c\<^sub>2 w' p = try_swap p' c\<^sub>1 c\<^sub>2 w p" unfolding leaves.simps by simp thus "w' p = w p" by simp qed with try_swap_inv show ?case by auto qed text \This lets us prove that only blue-eyed people will ever leave the island.\ proposition only_blue_eyes_leave: assumes "leaves n p w" and "valid w" shows "w p = blue" proof (rule ccontr) assume "w p \ blue" then obtain c where c: "w p \ c" "c \ blue" using colors_distinct by (metis distinct_length_2_or_more) let ?w' = "try_swap p (w p) c w" have "possible n p w ?w'" using `valid w` apply (simp add: possible.simps) by (auto simp: try_swap_def) moreover have "?w' p \ w p" using c `w p \ blue` by (auto simp: try_swap_def) ultimately have "\ leaves n p w" by (auto simp: leaves.simps) with assms show False by simp qed section "The blue-eyed logicians" text \We will now consider the behavior of the logicians with blue eyes. First, some simple lemmas. Reasoning about set cardinalities often requires considering infinite sets separately. Usefully, all sets of people are finite by assumption.\ lemma people_finite[simp]: "finite (S::'person set)" proof (rule finite_subset) show "S \ UNIV" by auto show "finite (UNIV::'person set)" by fact qed text \Secondly, we prove a destruction rule for @{const possible}. It is strictly weaker than the definition, but thanks to the simpler form, it's easier to guide the automation with it.\ lemma possibleD_colors: assumes "possible n p w w'" and "p' \ p" shows "w' p' = w p'" using assms unfolding possible.simps by simp text \A central concept in the reasoning is the set of blue-eyed people someone can see.\ definition blues_seen :: "('person \ 'color) \ 'person \ 'person set" where "blues_seen w p = {p'. w p' = blue} - {p}" lemma blues_seen_others: assumes "w p' = blue" and "p \ p'" shows "w p = blue \ card (blues_seen w p) = card (blues_seen w p')" and "w p \ blue \ card (blues_seen w p) = Suc (card (blues_seen w p'))" proof - assume "w p = blue" then have "blues_seen w p' = blues_seen w p \ {p} - {p'}" by (auto simp add: blues_seen_def) moreover have "p \ blues_seen w p" unfolding blues_seen_def by auto moreover have "p' \ blues_seen w p \ {p}" unfolding blues_seen_def using `p \ p'` `w p' = blue` by auto ultimately show "card (blues_seen w p) = card (blues_seen w p')" by simp next assume "w p \ blue" then have "blues_seen w p' = blues_seen w p - {p'}" by (auto simp add: blues_seen_def) moreover have "p' \ blues_seen w p" unfolding blues_seen_def using `p \ p'` `w p' = blue` by auto ultimately show "card (blues_seen w p) = Suc (card (blues_seen w p'))" by (simp only: card_Suc_Diff1 people_finite) qed lemma blues_seen_same[simp]: assumes "possible n p w w'" shows "blues_seen w' p = blues_seen w p" using assms by (auto simp: blues_seen_def possible.simps) lemma possible_blues_seen: assumes "possible n p w w'" assumes "w p' = blue" and "p \ p'" shows "w' p = blue \ card (blues_seen w p) = card (blues_seen w' p')" and "w' p \ blue \ card (blues_seen w p) = Suc (card (blues_seen w' p'))" using possibleD_colors[OF `possible n p w w'`] and blues_seen_others assms by (auto simp flip: blues_seen_same) text \Finally, the crux of the solution. We proceed by strong induction.\ lemma blue_leaves: assumes "w p = blue" and "valid w" and guru: "w guru \ blue" shows "leaves n p w \ n \ card (blues_seen w p)" using assms proof (induction n arbitrary: p w rule: less_induct) case (less n) show ?case proof \ \First, we show that day \n\ is sufficient to deduce that the eyes are blue.\ assume "n \ card (blues_seen w p)" have "w' p = blue" if "possible n p w w'" for w' proof (cases "card (blues_seen w' p)") case 0 moreover from `possible n p w w'` have "valid w'" by (simp add: possible.simps) ultimately show "w' p = blue" unfolding valid_def blues_seen_def by auto next case (Suc k) \ \We consider the behavior of somebody else, who also has blue eyes.\ then have "blues_seen w' p \ {}" by auto then obtain p' where "w' p' = blue" and "p \ p'" unfolding blues_seen_def by auto then have "w p' = blue" using possibleD_colors[OF `possible n p w w'`] by simp have "p \ guru" using `w p = blue` and `w guru \ blue` by auto hence "w' guru \ blue" using `w guru \ blue` and possibleD_colors[OF `possible n p w w'`] by simp have "valid w'" using `possible n p w w'` unfolding possible.simps by simp show "w' p = blue" proof (rule ccontr) assume "w' p \ blue" \ \If our eyes weren't blue, then \p'\ would see one blue-eyed person less than us.\ with possible_blues_seen[OF `possible n p w w'` `w p' = blue` `p \ p'`] have *: "card (blues_seen w p) = Suc (card (blues_seen w' p'))" by simp \ \By induction, they would've left on day \k = blues_seen w' p'\.\ let ?k = "card (blues_seen w' p')" have "?k < n" using `n \ card (blues_seen w p)` and * by simp hence "leaves ?k p' w'" using `valid w'` `w' p' = blue` `w' guru \ blue` by (intro less.IH[THEN iffD2]; auto) \ \However, we know that actually, \p'\ didn't leave that day yet.\ moreover have "\ leaves ?k p' w" proof assume "leaves ?k p' w" then have "?k \ card (blues_seen w p')" using `?k < n` `w p' = blue` `valid w` `w guru \ blue` by (intro less.IH[THEN iffD1]; auto) have "card (blues_seen w p) = card (blues_seen w p')" by (intro blues_seen_others; fact) with * have "?k < card (blues_seen w p')" by simp with `?k \ card (blues_seen w p')` show False by simp qed moreover have "leaves ?k p' w' = leaves ?k p' w" using `possible n p w w'` `?k < n` unfolding possible.simps by simp ultimately show False by simp qed qed thus "leaves n p w" unfolding leaves.simps using `w p = blue` by simp next \ \Then, we show that it's not possible to deduce the eye color any earlier.\ { assume "n < card (blues_seen w p)" \ \Consider a hypothetical world where \p\ has brown eyes instead. We will prove that this world is \possible\.\ let ?w' = "w(p := brown)" have "?w' guru \ blue" using `w guru \ blue` `w p = blue` by auto have "valid ?w'" proof - from `n < card (blues_seen w p)` have "card (blues_seen w p) \ 0" by auto hence "blues_seen w p \ {}" by auto then obtain p' where "p' \ blues_seen w p" by auto hence "p \ p'" and "w p' = blue" by (auto simp: blues_seen_def) hence "?w' p' = blue" by auto with `?w' guru \ blue` show "valid ?w'" unfolding valid_def by auto qed moreover have "leaves n' p' w = leaves n' p' ?w'" if "n' < n" for n' p' proof - have not_leavesI: "\leaves n' p' w'" if "valid w'" "w' guru \ blue" and P: "w' p' = blue \ n' < card (blues_seen w' p')" for w' proof (cases "w' p' = blue") case True then have "leaves n' p' w' \ n' \ card (blues_seen w' p')" using less.IH `n' < n` `valid w'` `w' guru \ blue` by simp with P[OF `w' p' = blue`] show "\leaves n' p' w'" by simp next case False then show "\ leaves n' p' w'" using only_blue_eyes_leave `valid w'` by auto qed have "\leaves n' p' w" proof (intro not_leavesI) assume "w p' = blue" with `w p = blue` have "card (blues_seen w p) = card (blues_seen w p')" apply (cases "p = p'", simp) by (intro blues_seen_others; auto) with `n' < n` and `n < card (blues_seen w p)` show "n' < card (blues_seen w p')" by simp qed fact+ moreover have "\ leaves n' p' ?w'" proof (intro not_leavesI) assume "?w' p' = blue" with colors_distinct have "p \ p'" and "?w' p \ blue" by auto hence "card (blues_seen ?w' p) = Suc (card (blues_seen ?w' p'))" using `?w' p' = blue` by (intro blues_seen_others; auto) moreover have "blues_seen w p = blues_seen ?w' p" unfolding blues_seen_def by auto ultimately show "n' < card (blues_seen ?w' p')" using `n' < n` and `n < card (blues_seen w p)` by auto qed fact+ ultimately show "leaves n' p' w = leaves n' p' ?w'" by simp qed ultimately have "possible n p w ?w'" using `valid w` by (auto simp: possible.simps) moreover have "?w' p \ blue" using colors_distinct by auto ultimately have "\ leaves n p w" unfolding leaves.simps using `w p = blue` by blast } then show "leaves n p w \ n \ card (blues_seen w p)" by fastforce qed qed text \This can be combined into a theorem that describes the behavior of the logicians based on the objective count of blue-eyed people, and not the count by a specific person. The xkcd puzzle is the instance where \n = 99\.\ theorem blue_eyes: assumes "card {p. w p = blue} = Suc n" and "valid w" and "w guru \ blue" shows "leaves k p w \ w p = blue \ k \ n" proof (cases "w p = blue") case True with assms have "card (blues_seen w p) = n" unfolding blues_seen_def by simp then show ?thesis using `w p = blue` `valid w` `w guru \ blue` blue_leaves by simp next case False then show ?thesis using only_blue_eyes_leave `valid w` by auto qed end (*<*) end (*>*) section \Future work\ text \After completing this formalization, I have been made aware of epistemic logic. The @{emph \possible worlds\} model in \cref{sec:world} turns out to be quite similar to the usual semantics of this logic. It might be interesting to solve this puzzle within the axiom system of epistemic logic, without explicit reasoning about possible worlds.\ \ No newline at end of file diff --git a/thys/Cayley_Hamilton/Square_Matrix.thy b/thys/Cayley_Hamilton/Square_Matrix.thy --- a/thys/Cayley_Hamilton/Square_Matrix.thy +++ b/thys/Cayley_Hamilton/Square_Matrix.thy @@ -1,751 +1,751 @@ (* Title: Cayley_Hamilton/Square_Matrix.thy Author: Johannes Hölzl, TU München *) theory Square_Matrix imports "HOL-Analysis.Determinants" "HOL-Analysis.Cartesian_Euclidean_Space" begin lemma smult_axis: "x *s axis i y = axis i (x * y::_::mult_zero)" by (simp add: axis_def vec_eq_iff) typedef ('a, 'n) sq_matrix = "UNIV :: ('n \ 'n \ 'a) set" morphisms to_fun of_fun by (rule UNIV_witness) syntax "_sq_matrix" :: "type \ type \ type" ("(_ ^^/ _)" [15, 16] 15) parse_translation \ let fun vec t u = Syntax.const @{type_syntax sq_matrix} $ t $ u; fun sq_matrix_tr [t, u] = (case Term_Position.strip_positions u of v as Free (x, _) => if Lexicon.is_tid x then vec t (Syntax.const @{syntax_const "_ofsort"} $ v $ Syntax.const @{class_syntax finite}) else vec t u | _ => vec t u) in [(@{syntax_const "_sq_matrix"}, K sq_matrix_tr)] end \ setup_lifting type_definition_sq_matrix lift_definition map_sq_matrix :: "('a \ 'c) \ 'a^^'b \ 'c^^'b" is "\f M i j. f (M i j)" . lift_definition from_vec :: "'a^'n^'n \ 'a^^'n" is "\M i j. M $ i $ j" . lift_definition to_vec :: "'a^^'n \ 'a^'n^'n" is "\M. \ i j. M i j" . lemma from_vec_eq_iff: "from_vec M = from_vec N \ M = N" by transfer (auto simp: vec_eq_iff fun_eq_iff) lemma to_vec_from_vec[simp]: "to_vec (from_vec M) = M" by transfer (simp add: vec_eq_iff) lemma from_vec_to_vec[simp]: "from_vec (to_vec M) = M" by transfer (simp add: vec_eq_iff fun_eq_iff) lemma map_sq_matrix_compose[simp]: "map_sq_matrix f (map_sq_matrix g M) = map_sq_matrix (\x. f (g x)) M" by transfer simp lemma map_sq_matrix_ident[simp]: "map_sq_matrix (\x. x) M = M" by transfer (simp add: vec_eq_iff) lemma map_sq_matrix_cong: "M = N \ (\i j. f (to_fun N i j) = g (to_fun N i j)) \ map_sq_matrix f M = map_sq_matrix g N" by transfer simp lift_definition diag :: "'a::zero \ 'a^^'n" is "\k i j. if i = j then k else 0" . lemma diag_eq_iff: "diag x = diag y \ x = y" by transfer (simp add: fun_eq_iff) lemma map_sq_matrix_diag[simp]: "f 0 = 0 \ map_sq_matrix f (diag c) = diag (f c)" by transfer (simp add: fun_eq_iff) lift_definition smult_sq_matrix :: "'a::times \ 'a^^'n \ 'a^^'n" (infixr "*\<^sub>S" 75) is "\c M i j. c * M i j" . lemma smult_map_sq_matrix: "(\y. f (x * y) = z * f y) \ map_sq_matrix f (x *\<^sub>S A) = z *\<^sub>S map_sq_matrix f A" by transfer simp lemma map_sq_matrix_smult: "c *\<^sub>S map_sq_matrix f A = map_sq_matrix (\x. c * f x) A" by transfer simp lemma one_smult[simp]: "(1::_::monoid_mult) *\<^sub>S x = x" by transfer simp lemma smult_diag: "x *\<^sub>S diag y = diag (x * y::_::mult_zero)" by transfer (simp add: fun_eq_iff) instantiation sq_matrix :: (semigroup_add, finite) semigroup_add begin lift_definition plus_sq_matrix :: "'a^^'b \ 'a^^'b \ 'a^^'b" is "\A B i j. A i j + B i j" . instance by standard (transfer, simp add: field_simps) end lemma map_sq_matrix_add: "(\a b. f (a + b) = f a + f b) \ map_sq_matrix f (A + B) = map_sq_matrix f A + map_sq_matrix f B" by transfer simp lemma add_map_sq_matrix: "map_sq_matrix f A + map_sq_matrix g A = map_sq_matrix (\x. f x + g x) A" by transfer simp instantiation sq_matrix :: (monoid_add, finite) monoid_add begin lift_definition zero_sq_matrix :: "'a^^'b" is "\i j. 0" . instance by standard (transfer, simp)+ end lemma diag_0: "diag 0 = 0" by transfer simp lemma diag_0_eq: "diag x = 0 \ x = 0" by transfer (simp add: fun_eq_iff) lemma zero_map_sq_matrix: "f 0 = 0 \ map_sq_matrix f 0 = 0" by transfer simp lemma map_sq_matrix_0[simp]: "map_sq_matrix (\x. 0) A = 0" by transfer simp instance sq_matrix :: (ab_semigroup_add, finite) ab_semigroup_add by standard (transfer, simp add: field_simps)+ instantiation sq_matrix :: (minus, finite) minus begin lift_definition minus_sq_matrix :: "'a^^'b \ 'a^^'b \ 'a^^'b" is "\A B i j. A i j - B i j" . instance .. end instantiation sq_matrix :: (group_add, finite) group_add begin lift_definition uminus_sq_matrix :: "'a^^'b \ 'a^^'b" is "uminus" . instance by standard (transfer, simp)+ end lemma map_sq_matrix_diff: "(\a b. f (a - b) = f a - f b) \ map_sq_matrix f (A - B) = map_sq_matrix f A - map_sq_matrix f B" by transfer (simp add: vec_eq_iff) lemma smult_diff: fixes a :: "'a::comm_ring_1" shows "a *\<^sub>S (A - B) = a *\<^sub>S A - a *\<^sub>S B" by transfer (simp add: field_simps) instance sq_matrix :: (cancel_semigroup_add, finite) cancel_semigroup_add by (standard; transfer, simp add: field_simps fun_eq_iff) instance sq_matrix :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add by (standard; transfer, simp add: field_simps) instance sq_matrix :: (comm_monoid_add, finite) comm_monoid_add by (standard; transfer, simp add: field_simps) lemma map_sq_matrix_sum: "f 0 = 0 \ (\a b. f (a + b) = f a + f b) \ map_sq_matrix f (\i\I. A i) = (\i\I. map_sq_matrix f (A i))" by (induction I rule: infinite_finite_induct) (auto simp: zero_map_sq_matrix map_sq_matrix_add) lemma sum_map_sq_matrix: "(\i\I. map_sq_matrix (f i) A) = map_sq_matrix (\x. \i\I. f i x) A" by (induction I rule: infinite_finite_induct) (simp_all add: add_map_sq_matrix) lemma smult_zero[simp]: fixes a :: "'a::ring_1" shows "a *\<^sub>S 0 = 0" by transfer (simp add: vec_eq_iff) lemma smult_right_add: fixes a :: "'a::ring_1" shows "a *\<^sub>S (x + y) = a *\<^sub>S x + a *\<^sub>S y" by transfer (simp add: vec_eq_iff field_simps) lemma smult_sum: fixes a :: "'a::ring_1" shows "(\i\I. a *\<^sub>S f i) = a *\<^sub>S (sum f I)" by (induction I rule: infinite_finite_induct) (simp_all add: smult_right_add vec_eq_iff) instance sq_matrix :: (ab_group_add, finite) ab_group_add by standard (transfer, simp add: field_simps)+ instantiation sq_matrix :: ("semiring_0", finite) semiring_0 begin lift_definition times_sq_matrix :: "'a^^'b \ 'a^^'b \ 'a^^'b" is "\M N i j. \k\UNIV. M i k * N k j" . instance proof fix a b c :: "'a^^'b" show "a * b * c = a * (b * c)" by transfer (auto simp: fun_eq_iff sum_distrib_left sum_distrib_right field_simps intro: sum.swap) qed (transfer, simp add: vec_eq_iff sum.distrib field_simps)+ end lemma diag_mult: "diag x * A = x *\<^sub>S A" by transfer (simp add: if_distrib[where f="\x. x * a" for a] sum.If_cases) lemma mult_diag: fixes x :: "'a::comm_ring_1" shows "A * diag x = x *\<^sub>S A" by transfer (simp add: if_distrib[where f="\x. a * x" for a] mult.commute sum.If_cases) lemma smult_mult1: fixes a :: "'a::comm_ring_1" shows "a *\<^sub>S (A * B) = (a *\<^sub>S A) * B" by transfer (simp add: sum_distrib_left field_simps) lemma smult_mult2: fixes a :: "'a::comm_ring_1" shows "a *\<^sub>S (A * B) = A * (a *\<^sub>S B)" by transfer (simp add: sum_distrib_left field_simps) lemma map_sq_matrix_mult: fixes f :: "'a::semiring_1 \ 'b::semiring_1" assumes f: "\a b. f (a + b) = f a + f b" "\a b. f (a * b) = f a * f b" "f 0 = 0" shows "map_sq_matrix f (A * B) = map_sq_matrix f A * map_sq_matrix f B" proof (transfer fixing: f) fix A B :: "'c \ 'c \ 'a" { fix I i j have "f (\k\I. A i k * B k j) = (\k\I. f (A i k) * f (B k j))" by (induction I rule: infinite_finite_induct) (auto simp add: f) } then show "(\i j. f (\k\UNIV. A i k * B k j)) = (\i j. \k\UNIV. f (A i k) * f (B k j))" by simp qed lemma from_vec_mult[simp]: "from_vec (M ** N) = from_vec M * from_vec N" by transfer (simp add: matrix_matrix_mult_def fun_eq_iff vec_eq_iff) instantiation sq_matrix :: ("semiring_1", finite) semiring_1 begin lift_definition one_sq_matrix :: "'a^^'b" is "\i j. if i = j then 1 else 0" . instance by standard (transfer, simp add: fun_eq_iff sum.If_cases if_distrib[where f="\x. x * b" for b] if_distrib[where f="\x. b * x" for b])+ end instance sq_matrix :: ("semiring_1", finite) numeral .. lemma diag_1: "diag 1 = 1" by transfer simp lemma diag_1_eq: "diag x = 1 \ x = 1" by transfer (simp add: fun_eq_iff) instance sq_matrix :: ("ring_1", finite) ring_1 by standard simp_all interpretation sq_matrix: vector_space smult_sq_matrix by standard (transfer, simp add: vec_eq_iff field_simps)+ instantiation sq_matrix :: (real_vector, finite) real_vector begin lift_definition scaleR_sq_matrix :: "real \ 'a^^'b \ 'a^^'b" is "\r A i j. r *\<^sub>R A i j" . instance by standard (transfer, simp add: scaleR_add_right scaleR_add_left)+ end instance sq_matrix :: ("semiring_1", finite) Rings.dvd .. lift_definition transpose :: "'a^^'n \ 'a^^'n" is "\M i j. M j i" . lemma transpose_transpose[simp]: "transpose (transpose A) = A" by transfer simp lemma transpose_diag[simp]: "transpose (diag c) = diag c" by transfer (simp add: fun_eq_iff) lemma transpose_zero[simp]: "transpose 0 = 0" by transfer simp lemma transpose_one[simp]: "transpose 1 = 1" by transfer (simp add: fun_eq_iff) lemma transpose_add[simp]: "transpose (A + B) = transpose A + transpose B" by transfer simp lemma transpose_minus[simp]: "transpose (A - B) = transpose A - transpose B" by transfer simp lemma transpose_uminus[simp]: "transpose (- A) = - transpose A" by transfer (simp add: fun_eq_iff) lemma transpose_mult[simp]: "transpose (A * B :: 'a::comm_semiring_0^^'n) = transpose B * transpose A" by transfer (simp add: field_simps) lift_definition trace :: "'a::comm_monoid_add^^'n \ 'a" is "\M. \i\UNIV. M i i" . lemma trace_diag[simp]: "trace (diag c::'a::semiring_1^^'n) = of_nat CARD('n) * c" by transfer simp lemma trace_0[simp]: "trace 0 = 0" by transfer simp lemma trace_1[simp]: "trace (1::'a::semiring_1^^'n) = of_nat CARD('n)" by transfer simp lemma trace_plus[simp]: "trace (A + B) = trace A + trace B" by transfer (simp add: sum.distrib) lemma trace_minus[simp]: "trace (A - B) = (trace A - trace B::_::ab_group_add)" by transfer (simp add: sum_subtractf) lemma trace_uminus[simp]: "trace (- A) = - (trace A::_::ab_group_add)" by transfer (simp add: sum_negf) lemma trace_smult[simp]: "trace (s *\<^sub>S A) = (s * trace A::_::semiring_0)" by transfer (simp add: sum_distrib_left) lemma trace_transpose[simp]: "trace (transpose A) = trace A" by transfer simp lemma trace_mult_symm: fixes A B :: "'a::comm_semiring_0^^'n" shows "trace (A * B) = trace (B * A)" by transfer (auto intro: sum.swap simp: mult.commute) lift_definition det :: "'a::comm_ring_1^^'n \ 'a" is "\A. (\p|p permutes UNIV. of_int (sign p) * (\i\UNIV. A i (p i)))" . lemma det_eq: "det A = (\p|p permutes UNIV. of_int (sign p) * (\i\UNIV. to_fun A i (p i)))" by transfer rule lemma permutes_UNIV_permutation: "permutation p \ p permutes (UNIV::_::finite)" by (auto simp: permutation_permutes permutes_def) lemma det_0[simp]: "det 0 = 0" by transfer (simp add: zero_power) lemma det_transpose: "det (transpose A) = det A" apply transfer apply (subst sum_permutations_inverse) apply (rule sum.cong[OF refl]) apply (simp add: sign_inverse permutes_UNIV_permutation) apply (subst prod.reindex_bij_betw[symmetric]) apply (rule permutes_imp_bij) apply assumption apply (simp add: permutes_inverses) done lemma det_diagonal: fixes A :: "'a::comm_ring_1^^'n" shows "(\i j. i \ j \ to_fun A i j = 0) \ det A = (\i\UNIV. to_fun A i i)" proof transfer fix A :: "'n \ 'n \ 'a" assume neq: "\i j. i \ j \ A i j = 0" let ?pp = "\p. of_int (sign p) * (\i\UNIV. A i (p i))" { fix p :: "'n \ 'n" assume p: "p permutes UNIV" "p \ id" then obtain i where i: "i \ p i" unfolding id_def by metis with neq[OF i] have "(\i\UNIV. A i (p i)) = 0" by (intro prod_zero) auto } then have "(\p | p permutes UNIV. ?pp p) = (\p\{id}. ?pp p)" by (intro sum.mono_neutral_cong_right) (auto intro: permutes_id) then show "(\p | p permutes UNIV. ?pp p) = (\i\UNIV. A i i)" by (simp add: sign_id) qed lemma det_1[simp]: "det (1::'a::comm_ring_1^^'n) = 1" by (subst det_diagonal) (transfer, simp)+ lemma det_lowerdiagonal: fixes A :: "'a::comm_ring_1^^'n::{finite,wellorder}" shows "(\i j. i < j \ to_fun A i j = 0) \ det A = (\i\UNIV. to_fun A i i)" proof transfer fix A :: "'n \ 'n \ 'a" assume ld: "\i j. i < j \ A i j = 0" let ?pp = "\p. of_int (sign p) * (\i\UNIV. A i (p i))" { fix p :: "'n \ 'n" assume p: "p permutes UNIV" "p \ id" with permutes_natset_le[OF p(1)] obtain i where i: "p i > i" by (metis not_le) with ld[OF i] have "(\i\UNIV. A i (p i)) = 0" by (intro prod_zero) auto } then have "(\p | p permutes UNIV. ?pp p) = (\p\{id}. ?pp p)" by (intro sum.mono_neutral_cong_right) (auto intro: permutes_id) then show "(\p | p permutes UNIV. ?pp p) = (\i\UNIV. A i i)" by (simp add: sign_id) qed lemma det_upperdiagonal: fixes A :: "'a::comm_ring_1^^'n::{finite, wellorder}" shows "(\i j. j < i \ to_fun A i j = 0) \ det A = (\i\UNIV. to_fun A i i)" using det_lowerdiagonal[of "transpose A"] unfolding det_transpose transpose.rep_eq . lift_definition perm_rows :: "'a^^'b \ ('b \ 'b) \ 'a^^'b" is "\M p i j. M (p i) j" . lift_definition perm_cols :: "'a^^'b \ ('b \ 'b) \ 'a^^'b" is "\M p i j. M i (p j)" . lift_definition upd_rows :: "'a^^'b \ 'b set \ ('b \ 'a^'b) \ 'a^^'b" is "\M S v i j. if i \ S then v i $ j else M i j" . lift_definition upd_cols :: "'a^^'b \ 'b set \ ('b \ 'a^'b) \ 'a^^'b" is "\M S v i j. if j \ S then v j $ i else M i j" . lift_definition upd_row :: "'a^^'b \ 'b \ 'a^'b \ 'a^^'b" is "\M i' v i j. if i = i' then v $ j else M i j" . lift_definition upd_col :: "'a^^'b \ 'b \ 'a^'b \ 'a^^'b" is "\M j' v i j. if j = j' then v $ i else M i j" . lift_definition row :: "'a^^'b \ 'b \ 'a^'b" is "\M i. \ j. M i j" . lift_definition col :: "'a^^'b \ 'b \ 'a^'b" is "\M j. \ i. M i j" . lemma perm_rows_transpose: "perm_rows (transpose M) p = transpose (perm_cols M p)" by transfer simp lemma perm_cols_transpose: "perm_cols (transpose M) p = transpose (perm_rows M p)" by transfer simp lemma upd_row_transpose: "upd_row (transpose M) i p = transpose (upd_col M i p)" by transfer simp lemma upd_col_transpose: "upd_col (transpose M) i p = transpose (upd_row M i p)" by transfer simp lemma upd_rows_transpose: "upd_rows (transpose M) i p = transpose (upd_cols M i p)" by transfer simp lemma upd_cols_transpose: "upd_cols (transpose M) i p = transpose (upd_rows M i p)" by transfer simp lemma upd_rows_empty[simp]: "upd_rows M {} f = M" by transfer simp lemma upd_cols_empty[simp]: "upd_cols M {} f = M" by transfer simp lemma upd_rows_single[simp]: "upd_rows M {i} f = upd_row M i (f i)" by transfer (simp add: fun_eq_iff) lemma upd_cols_single[simp]: "upd_cols M {i} f = upd_col M i (f i)" by transfer (simp add: fun_eq_iff) lemma upd_rows_insert: "upd_rows M (insert i I) f = upd_row (upd_rows M I f) i (f i)" by transfer (auto simp: fun_eq_iff) lemma upd_rows_insert_rev: "upd_rows M (insert i I) f = upd_rows (upd_row M i (f i)) I f" by transfer (auto simp: fun_eq_iff) lemma upd_rows_upd_row_swap: "i \ I \ upd_rows (upd_row M i x) I f = upd_row (upd_rows M I f) i x" by transfer (simp add: fun_eq_iff) lemma upd_cols_insert: "upd_cols M (insert i I) f = upd_col (upd_cols M I f) i (f i)" by transfer (auto simp: fun_eq_iff) lemma upd_cols_insert_rev: "upd_cols M (insert i I) f = upd_cols (upd_col M i (f i)) I f" by transfer (auto simp: fun_eq_iff) lemma upd_cols_upd_col_swap: "i \ I \ upd_cols (upd_col M i x) I f = upd_col (upd_cols M I f) i x" by transfer (simp add: fun_eq_iff) lemma upd_rows_cong[cong]: "M = N \ T = S \ (\s. s \ S =simp=> f s = g s) \ upd_rows M T f = upd_rows N S g" unfolding simp_implies_def by transfer (auto simp: fun_eq_iff) lemma upd_cols_cong[cong]: "M = N \ T = S \ (\s. s \ S =simp=> f s = g s) \ upd_cols M T f = upd_cols N S g" unfolding simp_implies_def by transfer (auto simp: fun_eq_iff) lemma row_upd_row_If: "row (upd_row M i x) j = (if i = j then x else row M j)" by transfer (simp add: vec_eq_iff fun_eq_iff) lemma row_upd_row[simp]: "row (upd_row M i x) i = x" by (simp add: row_upd_row_If) lemma col_upd_col_If: "col (upd_col M i x) j = (if i = j then x else col M j)" by transfer (simp add: vec_eq_iff) lemma col_upd_col[simp]: "col (upd_col M i x) i = x" by (simp add: col_upd_col_If) lemma upd_row_row[simp]: "upd_row M i (row M i) = M" by transfer (simp add: fun_eq_iff) lemma upd_row_upd_row_cancel[simp]: "upd_row (upd_row M i x) i y = upd_row M i y" by transfer (simp add: fun_eq_iff) lemma upd_col_upd_col_cancel[simp]: "upd_col (upd_col M i x) i y = upd_col M i y" by transfer (simp add: fun_eq_iff) lemma upd_col_col[simp]: "upd_col M i (col M i) = M" by transfer (simp add: fun_eq_iff) lemma row_transpose: "row (transpose M) i = col M i" by transfer simp lemma col_transpose: "col (transpose M) i = row M i" by transfer simp lemma det_perm_cols: fixes A :: "'a::comm_ring_1^^'n" assumes p: "p permutes UNIV" shows "det (perm_cols A p) = of_int (sign p) * det A" proof (transfer fixing: p) fix A :: "'n \ 'n \ 'a" from p have "(\q | q permutes UNIV. of_int (sign q) * (\i\UNIV. A i (p (q i)))) = (\q | q permutes UNIV. of_int (sign (inv p \ q)) * (\i\UNIV. A i (q i)))" by (intro sum.reindex_bij_witness[where j="\q. p \ q" and i="\q. inv p \ q"]) (auto simp: comp_assoc[symmetric] permutes_inv_o permutes_compose permutes_inv) with p show "(\q | q permutes UNIV. of_int (sign q) * (\i\UNIV. A i (p (q i)))) = of_int (sign p) * (\p | p permutes UNIV. of_int (sign p) * (\i\UNIV. A i (p i)))" by (auto intro!: sum.cong simp: sum_distrib_left sign_compose permutes_inv sign_inverse permutes_UNIV_permutation) qed lemma det_perm_rows: fixes A :: "'a::comm_ring_1^^'n" assumes p: "p permutes UNIV" shows "det (perm_rows A p) = of_int (sign p) * det A" using det_perm_cols[OF p, of "transpose A"] by (simp add: det_transpose perm_cols_transpose) lemma det_row_add: "det (upd_row M i (a + b)) = det (upd_row M i a) + det (upd_row M i b)" by transfer (simp add: prod.If_cases sum.distrib[symmetric] field_simps) lemma det_row_mul: "det (upd_row M i (c *s a)) = c * det (upd_row M i a)" by transfer (simp add: prod.If_cases sum_distrib_left field_simps) lemma det_row_uminus: "det (upd_row M i (- a)) = - det (upd_row M i a)" by (simp add: vector_sneg_minus1 det_row_mul) lemma det_row_minus: "det (upd_row M i (a - b)) = det (upd_row M i a) - det (upd_row M i b)" unfolding diff_conv_add_uminus det_row_add det_row_uminus .. lemma det_row_0: "det (upd_row M i 0) = 0" using det_row_mul[of M i 0] by simp lemma det_row_sum: "det (upd_row M i (\s\S. a s)) = (\s\S. det (upd_row M i (a s)))" by (induction S rule: infinite_finite_induct) (simp_all add: det_row_0 det_row_add) lemma det_col_add: "det (upd_col M i (a + b)) = det (upd_col M i a) + det (upd_col M i b)" using det_row_add[of "transpose M" i a b] by (simp add: upd_row_transpose det_transpose) lemma det_col_mul: "det (upd_col M i (c *s a)) = c * det (upd_col M i a)" using det_row_mul[of "transpose M" i c a] by (simp add: upd_row_transpose det_transpose) lemma det_col_uminus: "det (upd_col M i (- a)) = - det (upd_col M i a)" by (simp add: vector_sneg_minus1 det_col_mul) lemma det_col_minus: "det (upd_col M i (a - b)) = det (upd_col M i a) - det (upd_col M i b)" unfolding diff_conv_add_uminus det_col_add det_col_uminus .. lemma det_col_0: "det (upd_col M i 0) = 0" using det_col_mul[of M i 0] by simp lemma det_col_sum: "det (upd_col M i (\s\S. a s)) = (\s\S. det (upd_col M i (a s)))" by (induction S rule: infinite_finite_induct) (simp_all add: det_col_0 det_col_add) (* Proof by Jose Divason *) lemma det_identical_cols: assumes "i \ i'" shows "col A i = col A i' \ det A = 0" proof (transfer fixing: i i') fix A :: "'a \ 'a \ 'b" assume "(\ j. A j i) = (\ i. A i i')" then have [simp]: "\j q. A j (Fun.swap i i' id (q j)) = A j (q j)" - by (auto simp: vec_eq_iff swap_def) + by (auto simp: vec_eq_iff swap_id_eq) let ?p = "\p. of_int (sign p) * (\i\UNIV. A i (p i))" let ?s = "\q. Fun.swap i i' id \ q" let ?E = "{p. p permutes UNIV \ evenperm p}" have [simp]: "inj_on ?s ?E" - by (auto simp: inj_on_def fun_eq_iff swap_def) + by (auto simp: inj_on_def fun_eq_iff swap_id_eq) note p = permutes_UNIV_permutation evenperm_comp permutes_swap_id evenperm_swap permutes_compose sign_compose sign_swap_id from \i \ i'\ have *: "evenperm q" if "q \ ?s`?E" "q permutes UNIV" for q using that by (auto simp add: comp_assoc[symmetric] image_iff p elim!: allE[of _ "?s q"]) have "(\p | p permutes UNIV. ?p p) = (\p \ ?E \ ?s`?E. ?p p)" by (auto simp add: permutes_compose permutes_swap_id intro: * sum.cong) also have "\ = (\p\?E. ?p p) + (\p\?s`?E. ?p p)" by (intro sum.union_disjoint) (auto simp: p \i \ i'\) also have "(\p\?s`?E. ?p p) = (\p\?E. - ?p p)" using \i \ i'\ by (subst sum.reindex) (auto intro!: sum.cong simp: p) finally show "(\p | p permutes UNIV. ?p p) = 0" by (simp add: sum_negf) qed lemma det_identical_rows: "i \ i' \ row A i = row A i' \ det A = 0" using det_identical_cols[of i i' "transpose A"] by (simp add: det_transpose col_transpose) lemma det_cols_sum: "det (upd_cols M T (\i. \s\S. a i s)) = (\f\T \\<^sub>E S. det (upd_cols M T (\i. a i (f i))))" proof (induct T arbitrary: M rule: infinite_finite_induct) case (insert i T) have "(\f\insert i T \\<^sub>E S. det (upd_cols M (insert i T) (\i. a i (f i)))) = (\s\S. \f\T \\<^sub>E S. det (upd_cols (upd_col M i (a i s)) T (\i. a i (f i))))" unfolding sum.cartesian_product PiE_insert_eq using \i \ T\ by (subst sum.reindex[OF inj_combinator[OF \i \ T\]]) (auto intro!: sum.cong arg_cong[where f=det] upd_cols_cong simp: upd_cols_insert_rev simp_implies_def) also have "\ = det (upd_col (upd_cols M T (\i. sum (a i) S)) i (\s\S. a i s))" unfolding insert(3)[symmetric] by (simp add: upd_cols_upd_col_swap[OF \i \ T\] det_col_sum) finally show ?case by (simp add: upd_cols_insert) qed auto lemma det_rows_sum: "det (upd_rows M T (\i. \s\S. a i s)) = (\f\T \\<^sub>E S. det (upd_rows M T (\i. a i (f i))))" using det_cols_sum[of "transpose M" T a S] by (simp add: upd_cols_transpose det_transpose) lemma det_rows_mult: "det (upd_rows M T (\i. c i *s a i)) = (\i\T. c i) * det (upd_rows M T a)" by transfer (simp add: prod.If_cases sum_distrib_left field_simps prod.distrib) lemma det_cols_mult: "det (upd_cols M T (\i. c i *s a i)) = (\i\T. c i) * det (upd_cols M T a)" using det_rows_mult[of "transpose M" T c a] by (simp add: det_transpose upd_rows_transpose) lemma det_perm_rows_If: "det (perm_rows B f) = (if f permutes UNIV then of_int (sign f) * det B else 0)" proof cases assume "\ f permutes UNIV" moreover with bij_imp_permutes[of f UNIV] have "\ inj f" using finite_UNIV_inj_surj[of f] by (auto simp: bij_betw_def) then obtain i j where "f i = f j" "i \ j" by (auto simp: inj_on_def) moreover then have "row (perm_rows B f) i = row (perm_rows B f) j" by transfer (auto simp: vec_eq_iff) ultimately show ?thesis by (simp add: det_identical_rows) qed (simp add: det_perm_rows) lemma det_mult: "det (A * B) = det A * det B" proof - have "A * B = upd_rows 0 UNIV (\i. \j\UNIV. to_fun A i j *s row B j)" by transfer simp moreover have "\f. upd_rows 0 UNIV (\i. Square_Matrix.row B (f i)) = perm_rows B f" by transfer simp moreover have "det A = (\p | p permutes UNIV. of_int (sign p) * (\i\UNIV. to_fun A i (p i)))" by transfer rule ultimately show ?thesis by (auto simp add: det_rows_sum det_rows_mult sum_distrib_right det_perm_rows_If split: if_split_asm intro!: sum.mono_neutral_cong_right) qed lift_definition minor :: "'a^^'b \ 'b \ 'b \ 'a::semiring_1^^'b" is "\A i j k l. if k = i \ l = j then 1 else if k = i \ l = j then 0 else A k l" . lemma minor_transpose: "minor (transpose A) i j = transpose (minor A j i)" by transfer (auto simp: fun_eq_iff) lemma minor_eq_row_col: "minor M i j = upd_row (upd_col M j (axis i 1)) i (axis j 1)" by transfer (simp add: fun_eq_iff axis_def) lemma minor_eq_col_row: "minor M i j = upd_col (upd_row M i (axis j 1)) j (axis i 1)" by transfer (simp add: fun_eq_iff axis_def) lemma row_minor: "row (minor M i j) i = axis j 1" by (simp add: minor_eq_row_col) lemma col_minor: "col (minor M i j) j = axis i 1" by (simp add: minor_eq_col_row) lemma det_minor_row': "row B i = axis j 1 \ det (minor B i j) = det B" proof (induction "{k. to_fun B k j \ 0} - {i}" arbitrary: B rule: infinite_finite_induct) case empty then have "\k. k \ i \ to_fun B k j = 0" by (auto simp add: card_eq_0_iff) with empty.prems have "axis i 1 = col B j" by transfer (auto simp: vec_eq_iff axis_def) with empty.prems[symmetric] show ?case by (simp add: minor_eq_row_col) next case (insert r NZ) then have r: "r \ i" "to_fun B r j \ 0" by auto let ?B' = "upd_row B r (row B r - (to_fun B r j) *s row B i)" have "det (minor ?B' i j) = det ?B'" proof (rule insert.hyps) show "NZ = {k. to_fun ?B' k j \ 0} - {i}" using insert.hyps(2,4) insert.prems by transfer (auto simp add: axis_def set_eq_iff) show "row ?B' i = axis j 1" using r insert by (simp add: row_upd_row_If) qed also have "minor ?B' i j = minor B i j" using r insert.prems by transfer (simp add: fun_eq_iff axis_def) also have "det ?B' = det B" using \r \ i\ by (simp add: det_row_minus det_row_mul det_identical_rows[OF \r \ i\] row_upd_row_If) finally show ?case . qed simp lemma det_minor_row: "det (minor B i j) = det (upd_row B i (axis j 1))" proof - have "det (minor (upd_row B i (axis j 1)) i j) = det (upd_row B i (axis j 1))" by (rule det_minor_row') simp then show ?thesis by (simp add: minor_eq_col_row) qed lemma det_minor_col: "det (minor B i j) = det (upd_col B j (axis i 1))" using det_minor_row[of "transpose B" j i] by (simp add: minor_transpose det_transpose upd_row_transpose) lift_definition cofactor :: "'a^^'b \ 'a::comm_ring_1^^'b" is "\A i j. det (minor A i j)" . lemma cofactor_transpose: "cofactor (transpose A) = transpose (cofactor A)" by (simp add: cofactor_def minor_transpose det_transpose transpose.rep_eq to_fun_inject[symmetric] of_fun_inverse) definition "adjugate A = transpose (cofactor A)" lemma adjugate_transpose: "adjugate (transpose A) = transpose (adjugate A)" by (simp add: adjugate_def cofactor_transpose) theorem adjugate_mult_det: "adjugate A * A = diag (det A)" proof (intro to_fun_inject[THEN iffD1] fun_eq_iff[THEN iffD2] allI) fix i k have "to_fun (adjugate A * A) i k = (\j\UNIV. to_fun A j k * det (minor A j i))" by (simp add: adjugate_def times_sq_matrix.rep_eq transpose.rep_eq cofactor_def mult.commute of_fun_inverse) also have "\ = det (upd_col A i (\j\UNIV. to_fun A j k *s axis j 1))" by (simp add: det_minor_col det_col_mul det_col_sum) also have "(\j\UNIV. to_fun A j k *s axis j 1) = col A k" by transfer (simp add: smult_axis vec_eq_iff, simp add: axis_def sum.If_cases) also have "det (upd_col A i (col A k)) = (if i = k then det A else 0)" by (auto simp: col_upd_col_If det_identical_cols[of i k]) also have "\ = to_fun (diag (det A)) i k" by (simp add: diag.rep_eq) finally show "to_fun (adjugate A * A) i k = to_fun (diag (det A)) i k" . qed lemma mult_adjugate_det: "A * adjugate A = diag (det A)" proof- have "transpose (transpose (A * adjugate A)) = transpose (diag (det A))" unfolding transpose_mult adjugate_transpose[symmetric] adjugate_mult_det det_transpose .. then show ?thesis by simp qed end diff --git a/thys/Derangements/Derangements.thy b/thys/Derangements/Derangements.thy --- a/thys/Derangements/Derangements.thy +++ b/thys/Derangements/Derangements.thy @@ -1,515 +1,518 @@ (* Author: Lukas Bulwahn *) section \Derangements\ theory Derangements imports Complex_Main "HOL-Library.Permutations" begin subsection \Preliminaries\ subsubsection \Additions to @{theory HOL.Finite_Set} Theory\ lemma card_product_dependent: assumes "finite S" "\x \ S. finite (T x)" shows "card {(x, y). x \ S \ y \ T x} = (\x \ S. card (T x))" using card_SigmaI[OF assms, symmetric] by (auto intro!: arg_cong[where f=card] simp add: Sigma_def) subsubsection \Additions to @{theory "HOL-Library.Permutations"} Theory\ lemma permutes_imp_bij': assumes "p permutes S" shows "bij p" using assms by (meson bij_def permutes_inj permutes_surj) lemma permutesE: assumes "p permutes S" obtains "bij p" "\x. x \ S \ p x = x" using assms by (simp add: permutes_def permutes_imp_bij') lemma bij_imp_permutes': assumes "bij p" "\x. x \ A \ p x = x" shows "p permutes A" using assms bij_imp_permutes permutes_superset by force lemma permutes_swap: assumes "p permutes S" shows "Fun.swap x y p permutes (insert x (insert y S))" proof - from assms have "p permutes (insert x (insert y S))" by (meson permutes_subset subset_insertI) moreover have "Fun.swap x y id permutes (insert x (insert y S))" by (simp add: permutes_swap_id) ultimately show "Fun.swap x y p permutes (insert x (insert y S))" by (metis comp_id comp_swap permutes_compose) qed lemma bij_extends: "bij p \ p x = x \ bij (p(x := y, inv p y := x))" unfolding bij_def proof (rule conjI; erule conjE) assume a: "inj p" "p x = x" show "inj (p(x := y, inv p y := x))" proof (intro injI) fix z z' assume "(p(x := y, inv p y := x)) z = (p(x := y, inv p y := x)) z'" from this a show "z = z'" by (auto split: if_split_asm simp add: inv_f_eq inj_eq) qed next assume a: "inj p" "surj p" "p x = x" { fix x' from a have "(p(x := y, inv p y := x)) (((inv p)(y := x, x := inv p y)) x') = x'" by (auto split: if_split_asm) (metis surj_f_inv_f)+ } from this show "surj (p(x := y, inv p y := x))" by (metis surjI) qed lemma permutes_add_one: assumes "p permutes S" "x \ S" "y \ S" shows "p(x := y, inv p y := x) permutes (insert x S)" proof (rule bij_imp_permutes') from assms show "bij (p(x := y, inv p y := x))" by (meson bij_extends permutes_def permutes_imp_bij') from assms show "\z. z \ insert x S \ (p(x := y, inv p y := x)) z = z" by (metis fun_upd_apply insertCI permutes_def permutes_inverses(1)) qed lemma permutations_skip_one: assumes "p permutes S" "x : S" shows "p(x := x, inv p x := p x) permutes (S - {x})" proof (rule bij_imp_permutes') from assms show "\z. z \ S - {x} \ (p(x := x, inv p x := p x)) z = z" by (auto elim: permutesE simp add: bij_inv_eq_iff) (simp add: assms(1) permutes_in_image permutes_inv) from assms have "inj (p(x := x, inv p x := p x))" by (intro injI) (auto split: if_split_asm; metis permutes_inverses(2))+ - from assms this show "bij (p(x := x, inv p x := p x))" - by (metis UNIV_I bij_betw_imageI bij_betw_swap_iff permutes_inj permutes_surj surj_f_inv_f swap_def) + moreover have "surj (p(x := x, inv p x := p x))" + using assms UNIV_I bij_betw_swap_iff permutes_inj permutes_surj surj_f_inv_f + by (metis (no_types, hide_lams) Fun.swap_def bij_betw_def) + ultimately show "bij (p(x := x, inv p x := p x))" + by (rule bijI) qed lemma permutes_drop_cycle_size_two: assumes "p permutes S" "p (p x) = x" shows "Fun.swap x (p x) p permutes (S - {x, p x})" using assms by (auto intro!: bij_imp_permutes' elim!: permutesE) (metis swap_apply(1,3)) subsection \Fixpoint-Free Permutations\ definition derangements :: "nat set \ (nat \ nat) set" where "derangements S = {p. p permutes S \ (\x \ S. p x \ x)}" lemma derangementsI: assumes "p permutes S" "\x. x \ S \ p x \ x" shows "p \ derangements S" using assms unfolding derangements_def by auto lemma derangementsE: assumes "d : derangements S" obtains "d permutes S" "\x\S. d x \ x" using assms unfolding derangements_def by auto subsection \Properties of Derangements\ lemma derangements_inv: assumes d: "d \ derangements S" shows "inv d \ derangements S" using assms by (auto intro!: derangementsI elim!: derangementsE simp add: permutes_inv permutes_inv_eq) lemma derangements_in_image: assumes "d \ derangements A" "x \ A" shows "d x \ A" using assms by (auto elim: derangementsE simp add: permutes_in_image) lemma derangements_in_image_strong: assumes "d \ derangements A" "x \ A" shows "d x \ A - {x}" using assms by (auto elim: derangementsE simp add: permutes_in_image) lemma derangements_inverse_in_image: assumes "d \ derangements A" "x \ A" shows "inv d x \ A" using assms by (auto intro: derangements_in_image derangements_inv) lemma derangements_fixpoint: assumes "d \ derangements A" "x \ A" shows "d x = x" using assms by (auto elim!: derangementsE simp add: permutes_def) lemma derangements_no_fixpoint: assumes "d \ derangements A" "x \ A" shows "d x \ x" using assms by (auto elim: derangementsE) lemma finite_derangements: assumes "finite A" shows "finite (derangements A)" using assms unfolding derangements_def by (auto simp add: finite_permutations) subsection \Construction of Derangements\ lemma derangements_empty[simp]: "derangements {} = {id}" unfolding derangements_def by auto lemma derangements_singleton[simp]: "derangements {x} = {}" unfolding derangements_def by auto lemma derangements_swap: assumes "d \ derangements S" "x \ S" "y \ S" "x \ y" shows "Fun.swap x y d \ derangements (insert x (insert y S))" proof (rule derangementsI) from assms show "Fun.swap x y d permutes (insert x (insert y S))" by (auto intro: permutes_swap elim: derangementsE) from assms have s: "d x = x" "d y = y" by (auto intro: derangements_fixpoint) { fix x' assume "x' : insert x (insert y S)" from s assms \x \ y\ this show "Fun.swap x y d x' \ x'" by (cases "x' = x"; cases "x' = y") (auto dest: derangements_no_fixpoint) } qed lemma derangements_skip_one: assumes d: "d \ derangements S" and "x \ S" "d (d x) \ x" shows "d(x := x, inv d x := d x) \ derangements (S - {x})" proof - from d have bij: "bij d" by (auto elim: derangementsE simp add: permutes_imp_bij') from d \x : S\ have that: "d x : S - {x}" by (auto dest: derangements_in_image derangements_no_fixpoint) from d \d (d x) \ x\ bij have "\x'\S - {x}. (d(x := x, inv d x := d x)) x' \ x'" by (auto elim!: derangementsE simp add: bij_inv_eq_iff) from d \x : S\ this show derangements: "d(x:=x, inv d x:= d x) : derangements (S - {x})" by (meson derangementsE derangementsI permutations_skip_one) qed lemma derangements_add_one: assumes "d \ derangements S" "x \ S" "y \ S" shows "d(x := y, inv d y := x) \ derangements (insert x S)" proof (rule derangementsI) from assms show "d(x := y, inv d y := x) permutes (insert x S)" by (auto intro: permutes_add_one elim: derangementsE) next fix z assume "z : insert x S" from assms this derangements_inverse_in_image[OF assms(1), of y] show "(d(x := y, inv d y := x)) z \ z" by (auto elim: derangementsE) qed lemma derangements_drop_minimal_cycle: assumes "d \ derangements S" "d (d x) = x" shows "Fun.swap x (d x) d \ derangements (S - {x, d x})" proof (rule derangementsI) from assms show "Fun.swap x (d x) d permutes (S - {x, d x})" by (meson derangementsE permutes_drop_cycle_size_two) next fix y assume "y \ S - {x, d x}" from assms this show "Fun.swap x (d x) d y \ y" by (auto elim: derangementsE) qed subsection \Cardinality of Derangements\ subsubsection \Recursive Characterization\ fun count_derangements :: "nat \ nat" where "count_derangements 0 = 1" | "count_derangements (Suc 0) = 0" | "count_derangements (Suc (Suc n)) = (n + 1) * (count_derangements (Suc n) + count_derangements n)" lemma card_derangements: assumes "finite S" "card S = n" shows "card (derangements S) = count_derangements n" using assms proof (induct n arbitrary: S rule: count_derangements.induct) case 1 from this show ?case by auto next case 2 from this derangements_singleton finite_derangements show ?case using Finite_Set.card_0_eq card_eq_SucD count_derangements.simps(2) by fastforce next case (3 n) from 3(4) obtain x where "x \ S" using card_eq_SucD insertI1 by auto let ?D1 = "(\(y, d). Fun.swap x y d) ` {(y, d). y \ S & y \ x & d : derangements (S - {x, y})}" let ?D2 = "(\(y, f). f(x:=y, inv f y := x)) ` ((S - {x}) \ derangements (S - {x}))" from \x \ S\ have subset1: "?D1 \ derangements S" proof (auto) fix y d assume "y \ S" "y \ x" assume d: "d \ derangements (S - {x, y})" from \x : S\ \y : S\ have S: "S = insert x (insert y (S - {x, y}))" by auto from d \x : S\ \y : S\ \y \ x\ show "Fun.swap x y d \ derangements S" by (subst S) (auto intro!: derangements_swap) qed have subset2: "?D2 \ derangements S" proof (rule subsetI, erule imageE, simp split: prod.split_asm, (erule conjE)+) fix d y assume "d : derangements (S - {x})" "y : S" "y \ x" from this have "d(x := y, inv d y := x) \ derangements (insert x (S - {x}))" by (intro derangements_add_one) auto from this \x : S\ show "d(x := y, inv d y := x) \ derangements S" using insert_Diff by fastforce qed have split: "derangements S = ?D1 \ ?D2" proof from subset1 subset2 show "?D1 \ ?D2 \ derangements S" by simp next show "derangements S \ ?D1 \ ?D2" proof fix d assume d: "d : derangements S" show "d : ?D1 \ ?D2" proof (cases "d (d x) = x") case True from \x : S\ d have "d x \ S" "d x \ x" by (auto simp add: derangements_in_image derangements_no_fixpoint) from d True have "Fun.swap x (d x) d \ derangements (S - {x, d x})" by (rule derangements_drop_minimal_cycle) from \d x \ S\ \d x \ x\ this have "d : ?D1" by (auto intro: image_eqI[where x = "(d x, Fun.swap x (d x) d)"]) from this show ?thesis by auto next case False from d have bij: "bij d" by (auto elim: derangementsE simp add: permutes_imp_bij') from d \x : S\ have that: "d x : S - {x}" by (intro derangements_in_image_strong) from d \x : S\ False have derangements: "d(x:=x, inv d x:= d x) : derangements (S - {x})" by (auto intro: derangements_skip_one) from this have "bij (d(x := x, inv d x:= d x))" by (metis derangementsE permutes_imp_bij')+ from this have a: "inv (d(x := x, inv d x := d x)) (d x) = inv d x" by (metis bij_inv_eq_iff fun_upd_same) from bij have x: "d (inv d x) = x" by (meson bij_inv_eq_iff) from d derangements_inv[of d] \x : S\ have "inv d x \ x" "d x \ x" by (auto dest: derangements_no_fixpoint) from this a x have d_eq: "d = d(inv d x := d x, x := d x, inv (d(x := x, inv d x := d x)) (d x) := x)" by auto from derangements that have "(d x, d(x:=x, inv d x:=d x)) : ((S - {x}) \ derangements (S - {x}))" by auto from d_eq this have "d : ?D2" by (auto intro: image_eqI[where x = "(d x, d(x:=x, inv d x:=d x))"]) from this show ?thesis by auto qed qed qed have no_intersect: "?D1 \ ?D2 = {}" proof - have that: "\d. d \ ?D1 \ d (d x) = x" using Diff_iff Diff_insert2 derangements_fixpoint insertI1 swap_apply(2) by fastforce have "\d. d \ ?D2 \ d (d x) \ x" proof - fix d assume a: "d \ ?D2" from a obtain y d' where d: "d = d'(x := y, inv d' y := x)" "d' \ derangements (S - {x})" "y \ S - {x}" by auto from d(2) have inv: "inv d' \ derangements (S - {x})" by (rule derangements_inv) from d have inv_x: "inv d' y \ x" by (auto dest: derangements_inverse_in_image) from inv have inv_y: "inv d' y \ y" using d(3) derangements_no_fixpoint by blast from d inv_x have 1: "d x = y" by auto from d inv_y have 2: "d y = d' y" by auto from d(2, 3) have 3: "d' y \ S - {x}" by (auto dest: derangements_in_image) from 1 2 3 show "d (d x) \ x" by auto qed from this that show ?thesis by blast qed have inj: "inj_on (\(y, f). Fun.swap x y f) {(y, f). y \ S & y \ x & f : derangements (S - {x, y})}" unfolding inj_on_def by (clarify; metis DiffD2 derangements_fixpoint insertI1 insert_commute swap_apply(1) swap_nilpotent) have eq: "{(y, f). y \ S & y \ x & f : derangements (S - {x, y})} = {(y, f). y \ S - {x} & f : derangements (S - {x, y})}" by simp have eq': "{(y, f). y \ S & y \ x & f : derangements (S - {x, y})} = Sigma (S - {x}) (%y. derangements (S - {x, y}))" unfolding Sigma_def by auto have card: "\ y. y \ S - {x} \ card (derangements (S - {x, y})) = count_derangements n" proof - fix y assume "y \ S - {x}" from 3(3, 4) \x \ S\ this have "card (S - {x, y}) = n" by (metis Diff_insert2 card_Diff_singleton diff_Suc_1 finite_Diff) from 3(3) 3(2)[OF _ this] show "card (derangements (S - {x, y})) = count_derangements n" by auto qed from 3(3, 4) \x : S\ have card2: "card (S - {x}) = Suc n" by (simp add: card.insert_remove insert_absorb) from inj have "card ?D1 = card {(y, f). y \ S - {x} \ f \ derangements (S - {x, y})}" by (simp add: card_image) also from 3(3) have "... = (\y\S - {x}. card (derangements (S - {x, y})))" by (subst card_product_dependent) (simp add: finite_derangements)+ finally have card_1: "card ?D1 =(Suc n) * count_derangements n" using card card2 by auto have inj_2: "inj_on (\(y, f). f(x := y, inv f y := x)) ((S - {x}) \ derangements (S - {x}))" proof - { fix d d' y y' assume 1: "d \ derangements (S - {x})" "d' \ derangements (S - {x})" assume 2: "y \ S" "y \ x" "y' \ S" "y' \ x" assume eq: "d(x := y, inv d y := x) = d'(x := y', inv d' y' := x)" from 1 2 eq \x \ S\ have y: "y = y'" by (metis Diff_insert_absorb derangements_in_image derangements_inv fun_upd_same fun_upd_twist insert_iff mk_disjoint_insert) have "d = d'" proof fix z from 1 have d_x: "d x = d' x" by (auto dest!: derangements_fixpoint) from 1 have bij: "bij d" "bij d'" by (metis derangementsE permutes_imp_bij')+ from this have d_d: "d (inv d y) = y" "d' (inv d' y') = y'" by (simp add: bij_is_surj surj_f_inv_f)+ from 1 2 bij have neq: "d (inv d' y') \ x \ d' (inv d y) \ x" by (metis Diff_iff bij_inv_eq_iff derangements_fixpoint singletonI) from eq have "(d(x := y, inv d y := x)) z = (d'(x := y', inv d' y' := x)) z" by auto from y d_x d_d neq this show "d z = d' z" by (auto split: if_split_asm) qed from y this have "y = y' & d = d'" by auto } from this show ?thesis unfolding inj_on_def by auto qed from 3(3) 3(1)[OF _ card2] have card3: "card (derangements (S - {x})) = count_derangements (Suc n)" by auto from 3(3) inj_2 have card_2: "card ?D2 = (Suc n) * count_derangements (Suc n)" by (simp only: card_image) (auto simp add: card_cartesian_product card3 card2) from \finite S\have finite1: "finite ?D1" unfolding eq' by (auto intro: finite_derangements) from \finite S\ have finite2: "finite ?D2" by (auto intro: finite_cartesian_product finite_derangements) show ?case by (simp add: split card_Un_disjoint finite1 finite2 no_intersect card_1 card_2 field_simps) qed subsubsection \Closed-Form Characterization\ lemma count_derangements: "real (count_derangements n) = fact n * (\k \ {0..n}. (-1) ^ k / fact k)" proof (induct n rule: count_derangements.induct) case (3 n) let ?f = "\n. fact n * (\k = 0..n. (- 1) ^ k / fact k)" have "real (count_derangements (Suc (Suc n))) = (n + 1) * (count_derangements (n + 1) + count_derangements n)" unfolding count_derangements.simps by simp also have "... = real (n + 1) * (real (count_derangements (n + 1)) + real (count_derangements n))" by (simp only: of_nat_mult of_nat_add) also have "... = (n + 1) * (?f (n + 1) + ?f n)" unfolding 3(2) 3(1)[unfolded Suc_eq_plus1] .. also have "(n + 1) * (?f (n + 1) + ?f n) = ?f (n + 2)" proof - define f where "f n = ((fact n) :: real) * (\k = 0..n. (- 1) ^ k / fact k)" for n have f_eq: "\n. f (n + 1) = (n + 1) * f n + (- 1) ^ (n + 1)" proof - fix n have "?f (n + 1) = (n + 1) * fact n * (\k = 0..n. (- 1) ^ k / fact k) + fact (n + 1) * ((- 1) ^ (n + 1) / fact (n + 1))" by (simp add: field_simps) also have "... = (n + 1) * ?f n + (- 1) ^ (n + 1)" by (simp del: One_nat_def) finally show "?thesis n" unfolding f_def by simp qed from this have f_eq': "\ n. (n + 1) * f n = f (n + 1) + (- 1) ^ n" by auto from f_eq'[of n] have "(n + 1) * (f (n + 1) + f n) = ((n + 1) * f (n + 1)) + f (n + 1) + (- 1) ^ n" by (simp only: distrib_left of_nat_add of_nat_1) also have "... = (n + 2) * f (n + 1) + (- 1) ^ (n + 2)" by (simp del: One_nat_def add_2_eq_Suc' add: algebra_simps) simp also from f_eq[of "n + 1"] have "... = f (n + 2)" by simp finally show ?thesis unfolding f_def by simp qed finally show ?case by simp qed (auto) subsubsection \Approximation of Cardinality\ lemma two_power_fact_le_fact: assumes "n \ 1" shows "2^k * fact n \ (fact (n + k) :: 'a :: {semiring_char_0,linordered_semidom})" proof (induction k) case (Suc k) have "2 ^ Suc k * fact n = 2 * (2 ^ k * fact n)" by (simp add: algebra_simps) also note Suc.IH also from assms have "of_nat 1 + of_nat 1 \ of_nat n + (of_nat (Suc k) :: 'a)" by (intro add_mono) (unfold of_nat_le_iff, simp_all) hence "2 * (fact (n + k) :: 'a) \ of_nat (n + Suc k) * fact (n + k)" by (intro mult_right_mono) (simp_all add: add_ac) also have "\ = fact (n + Suc k)" by simp finally show ?case by - (simp add: mult_left_mono) qed simp_all lemma exp1_approx: assumes "n > 0" shows "exp (1::real) - (\k {0..2 / fact n}" proof (unfold atLeastAtMost_iff, safe) have "(\k. 1^k / fact k) sums exp (1::real)" using exp_converges[of "1::real"] by (simp add: field_simps) from sums_split_initial_segment[OF this, of n] have sums: "(\k. 1 / fact (n+k)) sums (exp 1 - (\kk 0" by (intro sums_le[OF _ sums_zero sums]) auto show "(exp 1 - (\k 2 / fact n" proof (rule sums_le) from assms have "(\k. (1 / fact n) * (1 / 2)^k :: real) sums ((1 / fact n) * (1 / (1 - 1 / 2)))" by (intro sums_mult geometric_sums) simp_all also have "\ = 2 / fact n" by simp finally show "(\k. 1 / fact n * (1 / 2) ^ k) sums (2 / fact n :: real)" . next fix k show "(1 / fact (n+k)) \ (1 / fact n) * (1 / 2 :: real)^k" for k using two_power_fact_le_fact[of n k] assms by (auto simp: field_simps) qed fact+ qed lemma exp1_bounds: "exp 1 \ {8 / 3..11 / 4 :: real}" using exp1_approx[of 4] by (simp add: eval_nat_numeral) lemma count_derangements_approximation: assumes "n \ 0" shows "abs(real (count_derangements n) - fact n / exp 1) < 1 / 2" proof (cases "n \ 5") case False from assms this have n: "n = 1 \ n = 2 \ n = 3 \ n = 4" by auto from exp1_bounds have 1: "abs(real (count_derangements 1) - fact 1 / exp 1) < 1 / 2" by simp from exp1_bounds have 2: "abs(real (count_derangements 2) - fact 2 / exp 1) < 1 / 2" by (auto simp add: eval_nat_numeral abs_real_def) from exp1_bounds have 3: "abs(real (count_derangements 3) - fact 3 / exp 1) < 1 / 2" by (auto simp add: eval_nat_numeral abs_if field_simps) from exp1_bounds have 4: "abs(real (count_derangements 4) - fact 4 / exp 1) < 1 / 2" by (auto simp: abs_if field_simps eval_nat_numeral) from 1 2 3 4 n show ?thesis by auto next case True from Maclaurin_exp_le[of "- 1" "n + 1"] obtain t where t: "abs (t :: real) \ abs (- 1)" and exp: "exp (- 1) = (\mk = 0..n. (- 1) ^ k / fact k) = exp (- 1) - exp t / fact (n + 1) * (- 1) ^ (n + 1)" by (simp add: atLeast0AtMost ivl_disj_un(2)[symmetric]) have fact_plus1: "fact (n + 1) = (n + 1) * fact n" by simp have eq: "\real (count_derangements n) - fact n / exp 1\ = exp t / (n + 1)" by (simp del: One_nat_def add: count_derangements sum_eq_exp exp_minus inverse_eq_divide algebra_simps abs_mult) (simp add: fact_plus1) from t have "exp t \ exp 1" by simp also from exp1_bounds have "\ < 3" by simp finally show ?thesis using \n \ 5\ by (simp add: eq) qed theorem derangements_formula: assumes "n \ 0" "finite S" "card S = n" shows "int (card (derangements S)) = round (fact n / exp 1 :: real)" using count_derangements_approximation[of n] assms by (intro round_unique' [symmetric]) (auto simp: card_derangements abs_minus_commute) theorem derangements_formula': assumes "n \ 0" "finite S" "card S = n" shows "card (derangements S) = nat (round (fact n / exp 1 :: real))" using derangements_formula[OF assms] by simp end diff --git a/thys/Groebner_Bases/Algorithm_Schema.thy b/thys/Groebner_Bases/Algorithm_Schema.thy --- a/thys/Groebner_Bases/Algorithm_Schema.thy +++ b/thys/Groebner_Bases/Algorithm_Schema.thy @@ -1,4390 +1,4390 @@ (* Author: Alexander Maletzky *) section \A General Algorithm Schema for Computing Gr\"obner Bases\ theory Algorithm_Schema imports General Groebner_Bases begin text \This theory formalizes a general algorithm schema for computing Gr\"obner bases, generalizing Buchberger's original critical-pair/completion algorithm. The algorithm schema depends on several functional parameters that can be instantiated by a variety of concrete functions. Possible instances yield Buchberger's algorithm, Faug\`ere's F4 algorithm, and (as far as we can tell) even his F5 algorithm.\ subsection \@{term processed}\ definition minus_pairs (infixl "-\<^sub>p" 65) where "minus_pairs A B = A - (B \ prod.swap ` B)" definition Int_pairs (infixl "\\<^sub>p" 65) where "Int_pairs A B = A \ (B \ prod.swap ` B)" definition in_pair (infix "\\<^sub>p" 50) where "in_pair p A \ (p \ A \ prod.swap ` A)" definition subset_pairs (infix "\\<^sub>p" 50) where "subset_pairs A B \ (\x. x \\<^sub>p A \ x \\<^sub>p B)" abbreviation not_in_pair (infix "\\<^sub>p" 50) where "not_in_pair p A \ \ p \\<^sub>p A" lemma in_pair_alt: "p \\<^sub>p A \ (p \ A \ prod.swap p \ A)" by (metis (mono_tags, lifting) UnCI UnE image_iff in_pair_def prod.collapse swap_simp) lemma in_pair_iff: "(a, b) \\<^sub>p A \ ((a, b) \ A \ (b, a) \ A)" by (simp add: in_pair_alt) lemma in_pair_minus_pairs [simp]: "p \\<^sub>p A -\<^sub>p B \ (p \\<^sub>p A \ p \\<^sub>p B)" by (metis Diff_iff in_pair_def in_pair_iff minus_pairs_def prod.collapse) lemma in_minus_pairs [simp]: "p \ A -\<^sub>p B \ (p \ A \ p \\<^sub>p B)" by (metis Diff_iff in_pair_def minus_pairs_def) lemma in_pair_Int_pairs [simp]: "p \\<^sub>p A \\<^sub>p B \ (p \\<^sub>p A \ p \\<^sub>p B)" by (metis (no_types, hide_lams) Int_iff Int_pairs_def in_pair_alt in_pair_def old.prod.exhaust swap_simp) lemma in_pair_Un [simp]: "p \\<^sub>p A \ B \ (p \\<^sub>p A \ p \\<^sub>p B)" by (metis (mono_tags, lifting) UnE UnI1 UnI2 image_Un in_pair_def) lemma in_pair_trans [trans]: assumes "p \\<^sub>p A" and "A \ B" shows "p \\<^sub>p B" using assms by (auto simp: in_pair_def) lemma in_pair_same [simp]: "p \\<^sub>p A \ A \ p \ A \ A" - by (auto simp: in_pair_def swap_def) + by (auto simp: in_pair_def) lemma subset_pairsI [intro]: assumes "\x. x \\<^sub>p A \ x \\<^sub>p B" shows "A \\<^sub>p B" unfolding subset_pairs_def using assms by blast lemma subset_pairsD [trans]: assumes "x \\<^sub>p A" and "A \\<^sub>p B" shows "x \\<^sub>p B" using assms unfolding subset_pairs_def by blast definition processed :: "('a \ 'a) \ 'a list \ ('a \ 'a) list \ bool" where "processed p xs ps \ p \ set xs \ set xs \ p \\<^sub>p set ps" lemma processed_alt: "processed (a, b) xs ps \ ((a \ set xs) \ (b \ set xs) \ (a, b) \\<^sub>p set ps)" unfolding processed_def by auto lemma processedI: assumes "a \ set xs" and "b \ set xs" and "(a, b) \\<^sub>p set ps" shows "processed (a, b) xs ps" unfolding processed_alt using assms by simp lemma processedD1: assumes "processed (a, b) xs ps" shows "a \ set xs" using assms by (simp add: processed_alt) lemma processedD2: assumes "processed (a, b) xs ps" shows "b \ set xs" using assms by (simp add: processed_alt) lemma processedD3: assumes "processed (a, b) xs ps" shows "(a, b) \\<^sub>p set ps" using assms by (simp add: processed_alt) lemma processed_Nil: "processed (a, b) xs [] \ (a \ set xs \ b \ set xs)" by (simp add: processed_alt in_pair_iff) lemma processed_Cons: assumes "processed (a, b) xs ps" and a1: "a = p \ b = q \ thesis" and a2: "a = q \ b = p \ thesis" and a3: "processed (a, b) xs ((p, q) # ps) \ thesis" shows thesis proof - from assms(1) have "a \ set xs" and "b \ set xs" and "(a, b) \\<^sub>p set ps" by (simp_all add: processed_alt) show ?thesis proof (cases "(a, b) = (p, q)") case True hence "a = p" and "b = q" by simp_all thus ?thesis by (rule a1) next case False with \(a, b) \\<^sub>p set ps\ have *: "(a, b) \ set ((p, q) # ps)" by (auto simp: in_pair_iff) show ?thesis proof (cases "(b, a) = (p, q)") case True hence "a = q" and "b = p" by simp_all thus ?thesis by (rule a2) next case False with \(a, b) \\<^sub>p set ps\ have "(b, a) \ set ((p, q) # ps)" by (auto simp: in_pair_iff) with * have "(a, b) \\<^sub>p set ((p, q) # ps)" by (simp add: in_pair_iff) with \a \ set xs\ \b \ set xs\ have "processed (a, b) xs ((p, q) # ps)" by (rule processedI) thus ?thesis by (rule a3) qed qed qed lemma processed_minus: assumes "processed (a, b) xs (ps -- qs)" and a1: "(a, b) \\<^sub>p set qs \ thesis" and a2: "processed (a, b) xs ps \ thesis" shows thesis proof - from assms(1) have "a \ set xs" and "b \ set xs" and "(a, b) \\<^sub>p set (ps -- qs)" by (simp_all add: processed_alt) show ?thesis proof (cases "(a, b) \\<^sub>p set qs") case True thus ?thesis by (rule a1) next case False with \(a, b) \\<^sub>p set (ps -- qs)\ have "(a, b) \\<^sub>p set ps" by (auto simp: set_diff_list in_pair_iff) with \a \ set xs\ \b \ set xs\ have "processed (a, b) xs ps" by (rule processedI) thus ?thesis by (rule a2) qed qed subsection \Algorithm Schema\ subsubsection \\const_lt_component\\ context ordered_term begin definition const_lt_component :: "('t \\<^sub>0 'b::zero) \ 'k option" where "const_lt_component p = (let v = lt p in if pp_of_term v = 0 then Some (component_of_term v) else None)" lemma const_lt_component_SomeI: assumes "lp p = 0" and "component_of_term (lt p) = cmp" shows "const_lt_component p = Some cmp" using assms by (simp add: const_lt_component_def) lemma const_lt_component_SomeD1: assumes "const_lt_component p = Some cmp" shows "lp p = 0" using assms by (simp add: const_lt_component_def Let_def split: if_split_asm) lemma const_lt_component_SomeD2: assumes "const_lt_component p = Some cmp" shows "component_of_term (lt p) = cmp" using assms by (simp add: const_lt_component_def Let_def split: if_split_asm) lemma const_lt_component_subset: "const_lt_component ` (B - {0}) - {None} \ Some ` component_of_term ` Keys B" proof fix k assume "k \ const_lt_component ` (B - {0}) - {None}" hence "k \ const_lt_component ` (B - {0})" and "k \ None" by simp_all from this(1) obtain p where "p \ B - {0}" and "k = const_lt_component p" .. moreover from \k \ None\ obtain k' where "k = Some k'" by blast ultimately have "const_lt_component p = Some k'" and "p \ B" and "p \ 0" by simp_all from this(1) have "component_of_term (lt p) = k'" by (rule const_lt_component_SomeD2) moreover have "lt p \ Keys B" by (rule in_KeysI, rule lt_in_keys, fact+) ultimately have "k' \ component_of_term ` Keys B" by fastforce thus "k \ Some ` component_of_term ` Keys B" by (simp add: \k = Some k'\) qed corollary card_const_lt_component_le: assumes "finite B" shows "card (const_lt_component ` (B - {0}) - {None}) \ card (component_of_term ` Keys B)" proof (rule surj_card_le) show "finite (component_of_term ` Keys B)" by (intro finite_imageI finite_Keys, fact) next show "const_lt_component ` (B - {0}) - {None} \ Some ` component_of_term ` Keys B" by (fact const_lt_component_subset) qed end (* ordered_term *) subsubsection \Type synonyms\ type_synonym ('a, 'b, 'c) pdata' = "('a \\<^sub>0 'b) \ 'c" type_synonym ('a, 'b, 'c) pdata = "('a \\<^sub>0 'b) \ nat \ 'c" type_synonym ('a, 'b, 'c) pdata_pair = "('a, 'b, 'c) pdata \ ('a, 'b, 'c) pdata" type_synonym ('a, 'b, 'c, 'd) selT = "('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata_pair list \ nat \ 'd \ ('a, 'b, 'c) pdata_pair list" type_synonym ('a, 'b, 'c, 'd) complT = "('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata_pair list \ ('a, 'b, 'c) pdata_pair list \ nat \ 'd \ (('a, 'b, 'c) pdata' list \ 'd)" type_synonym ('a, 'b, 'c, 'd) apT = "('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata_pair list \ ('a, 'b, 'c) pdata list \ nat \ 'd \ ('a, 'b, 'c) pdata_pair list" type_synonym ('a, 'b, 'c, 'd) abT = "('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata list \ ('a, 'b, 'c) pdata list \ nat \ 'd \ ('a, 'b, 'c) pdata list" subsubsection \Specification of the @{emph \selector\} parameter\ definition sel_spec :: "('a, 'b, 'c, 'd) selT \ bool" where "sel_spec sel \ (\gs bs ps data. ps \ [] \ (sel gs bs ps data \ [] \ set (sel gs bs ps data) \ set ps))" lemma sel_specI: assumes "\gs bs ps data. ps \ [] \ (sel gs bs ps data \ [] \ set (sel gs bs ps data) \ set ps)" shows "sel_spec sel" unfolding sel_spec_def using assms by blast lemma sel_specD1: assumes "sel_spec sel" and "ps \ []" shows "sel gs bs ps data \ []" using assms unfolding sel_spec_def by blast lemma sel_specD2: assumes "sel_spec sel" and "ps \ []" shows "set (sel gs bs ps data) \ set ps" using assms unfolding sel_spec_def by blast subsubsection \Specification of the @{emph \add-basis\} parameter\ definition ab_spec :: "('a, 'b, 'c, 'd) abT \ bool" where "ab_spec ab \ (\gs bs ns data. ns \ [] \ set (ab gs bs ns data) = set bs \ set ns) \ (\gs bs data. ab gs bs [] data = bs)" lemma ab_specI: assumes "\gs bs ns data. ns \ [] \ set (ab gs bs ns data) = set bs \ set ns" and "\gs bs data. ab gs bs [] data = bs" shows "ab_spec ab" unfolding ab_spec_def using assms by blast lemma ab_specD1: assumes "ab_spec ab" shows "set (ab gs bs ns data) = set bs \ set ns" using assms unfolding ab_spec_def by (metis empty_set sup_bot.right_neutral) lemma ab_specD2: assumes "ab_spec ab" shows "ab gs bs [] data = bs" using assms unfolding ab_spec_def by blast subsubsection \Specification of the @{emph \add-pairs\} parameter\ definition unique_idx :: "('t, 'b, 'c) pdata list \ (nat \ 'd) \ bool" where "unique_idx bs data \ (\f\set bs. \g\set bs. fst (snd f) = fst (snd g) \ f = g) \ (\f\set bs. fst (snd f) < fst data)" lemma unique_idxI: assumes "\f g. f \ set bs \ g \ set bs \ fst (snd f) = fst (snd g) \ f = g" and "\f. f \ set bs \ fst (snd f) < fst data" shows "unique_idx bs data" unfolding unique_idx_def using assms by blast lemma unique_idxD1: assumes "unique_idx bs data" and "f \ set bs" and "g \ set bs" and "fst (snd f) = fst (snd g)" shows "f = g" using assms unfolding unique_idx_def by blast lemma unique_idxD2: assumes "unique_idx bs data" and "f \ set bs" shows "fst (snd f) < fst data" using assms unfolding unique_idx_def by blast lemma unique_idx_Nil: "unique_idx [] data" by (simp add: unique_idx_def) lemma unique_idx_subset: assumes "unique_idx bs data" and "set bs' \ set bs" shows "unique_idx bs' data" proof (rule unique_idxI) fix f g assume "f \ set bs'" and "g \ set bs'" with assms have "unique_idx bs data" and "f \ set bs" and "g \ set bs" by auto moreover assume "fst (snd f) = fst (snd g)" ultimately show "f = g" by (rule unique_idxD1) next fix f assume "f \ set bs'" with assms(2) have "f \ set bs" by auto with assms(1) show "fst (snd f) < fst data" by (rule unique_idxD2) qed context gd_term begin definition ap_spec :: "('t, 'b::field, 'c, 'd) apT \ bool" where "ap_spec ap \ (\gs bs ps hs data. set (ap gs bs ps hs data) \ set ps \ (set hs \ (set gs \ set bs \ set hs)) \ (\B d m. \h\set hs. \g\set gs \ set bs \ set hs. dickson_grading d \ set gs \ set bs \ set hs \ B \ fst ` B \ dgrad_p_set d m \ set ps \ set bs \ (set gs \ set bs) \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ h \ g \ fst h \ 0 \ fst g \ 0 \ (\a b. (a, b) \\<^sub>p set (ap gs bs ps hs data) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)) \ (\a b. a \ set gs \ set bs \ b \ set gs \ set bs \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)) \ crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)) \ (\B d m. \h g. dickson_grading d \ set gs \ set bs \ set hs \ B \ fst ` B \ dgrad_p_set d m \ set ps \ set bs \ (set gs \ set bs) \ (set gs \ set bs) \ set hs = {} \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ h \ g \ fst h \ 0 \ fst g \ 0 \ (h, g) \ set ps -\<^sub>p set (ap gs bs ps hs data) \ (\a b. (a, b) \\<^sub>p set (ap gs bs ps hs data) \ (a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)) \ crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)))" text \Informally, \ap_spec ap\ means that, for suitable arguments \gs\, \bs\, \ps\ and \hs\, the value of \ap gs bs ps hs\ is a list of pairs \ps'\ such that for every element \(a, b)\ missing in \ps'\ there exists a set of pairs \C\ by reference to which \(a, b)\ can be discarded, i.\,e. as soon as all critical pairs of the elements in \C\ can be connected below some set \B\, the same is true for the critical pair of \(a, b)\.\ lemma ap_specI: assumes "\gs bs ps hs data. set (ap gs bs ps hs data) \ set ps \ (set hs \ (set gs \ set bs \ set hs))" assumes "\gs bs ps hs data B d m h g. dickson_grading d \ set gs \ set bs \ set hs \ B \ fst ` B \ dgrad_p_set d m \ h \ set hs \ g \ set gs \ set bs \ set hs \ set ps \ set bs \ (set gs \ set bs) \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ h \ g \ fst h \ 0 \ fst g \ 0 \ (\a b. (a, b) \\<^sub>p set (ap gs bs ps hs data) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)) \ (\a b. a \ set gs \ set bs \ b \ set gs \ set bs \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)) \ crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)" assumes "\gs bs ps hs data B d m h g. dickson_grading d \ set gs \ set bs \ set hs \ B \ fst ` B \ dgrad_p_set d m \ set ps \ set bs \ (set gs \ set bs) \ (set gs \ set bs) \ set hs = {} \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ h \ g \ fst h \ 0 \ fst g \ 0 \ (h, g) \ set ps -\<^sub>p set (ap gs bs ps hs data) \ (\a b. (a, b) \\<^sub>p set (ap gs bs ps hs data) \ (a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)) \ crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)" shows "ap_spec ap" unfolding ap_spec_def apply (intro allI conjI impI) subgoal by (rule assms(1)) subgoal by (intro ballI impI, rule assms(2), blast+) subgoal by (rule assms(3), blast+) done lemma ap_specD1: assumes "ap_spec ap" shows "set (ap gs bs ps hs data) \ set ps \ (set hs \ (set gs \ set bs \ set hs))" using assms unfolding ap_spec_def by (elim allE conjE) (assumption) lemma ap_specD2: assumes "ap_spec ap" and "dickson_grading d" and "set gs \ set bs \ set hs \ B" and "fst ` B \ dgrad_p_set d m" and "(h, g) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" and "set ps \ set bs \ (set gs \ set bs)" and "unique_idx (gs @ bs @ hs) data" and "is_Groebner_basis (fst ` set gs)" and "h \ g" and "fst h \ 0" and "fst g \ 0" and "\a b. (a, b) \\<^sub>p set (ap gs bs ps hs data) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" and "\a b. a \ set gs \ set bs \ b \ set gs \ set bs \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" shows "crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)" proof - from assms(5) have "(h, g) \ set hs \ (set gs \ set bs \ set hs) \ (g, h) \ set hs \ (set gs \ set bs \ set hs)" by (simp only: in_pair_iff) thus ?thesis proof assume "(h, g) \ set hs \ (set gs \ set bs \ set hs)" hence "h \ set hs" and "g \ set gs \ set bs \ set hs" by simp_all from assms(1)[unfolded ap_spec_def, rule_format, of gs bs ps hs data] assms(2-4) this assms (6-) show ?thesis by metis next assume "(g, h) \ set hs \ (set gs \ set bs \ set hs)" hence "g \ set hs" and "h \ set gs \ set bs \ set hs" by simp_all hence "crit_pair_cbelow_on d m (fst ` B) (fst g) (fst h)" using assms(1)[unfolded ap_spec_def, rule_format, of gs bs ps hs data] assms(2,3,4,6,7,8,10,11,12,13) assms(9)[symmetric] by metis thus ?thesis by (rule crit_pair_cbelow_sym) qed qed lemma ap_specD3: assumes "ap_spec ap" and "dickson_grading d" and "set gs \ set bs \ set hs \ B" and "fst ` B \ dgrad_p_set d m" and "set ps \ set bs \ (set gs \ set bs)" and "(set gs \ set bs) \ set hs = {}" and "unique_idx (gs @ bs @ hs) data" and "is_Groebner_basis (fst ` set gs)" and "h \ g" and "fst h \ 0" and "fst g \ 0" and "(h, g) \\<^sub>p set ps -\<^sub>p set (ap gs bs ps hs data)" and "\a b. a \ set hs \ b \ set gs \ set bs \ set hs \ (a, b) \\<^sub>p set (ap gs bs ps hs data) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" shows "crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)" proof - have *: "crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" if 1: "(a, b) \\<^sub>p set (ap gs bs ps hs data)" and 2: "(a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" and 3: "fst a \ 0" and 4: "fst b \ 0" for a b proof - from 2 have "(a, b) \ set hs \ (set gs \ set bs \ set hs) \ (b, a) \ set hs \ (set gs \ set bs \ set hs)" by (simp only: in_pair_iff) thus ?thesis proof assume "(a, b) \ set hs \ (set gs \ set bs \ set hs)" hence "a \ set hs" and "b \ set gs \ set bs \ set hs" by simp_all thus ?thesis using 1 3 4 by (rule assms(13)) next assume "(b, a) \ set hs \ (set gs \ set bs \ set hs)" hence "b \ set hs" and "a \ set gs \ set bs \ set hs" by simp_all moreover from 1 have "(b, a) \\<^sub>p set (ap gs bs ps hs data)" by (auto simp: in_pair_iff) ultimately have "crit_pair_cbelow_on d m (fst ` B) (fst b) (fst a)" using 4 3 by (rule assms(13)) thus ?thesis by (rule crit_pair_cbelow_sym) qed qed from assms(12) have "(h, g) \ set ps -\<^sub>p set (ap gs bs ps hs data) \ (g, h) \ set ps -\<^sub>p set (ap gs bs ps hs data)" by (simp only: in_pair_iff) thus ?thesis proof assume "(h, g) \ set ps -\<^sub>p set (ap gs bs ps hs data)" with assms(1)[unfolded ap_spec_def, rule_format, of gs bs ps hs data] assms(2-11) show ?thesis using assms(10) * by metis next assume "(g, h) \ set ps -\<^sub>p set (ap gs bs ps hs data)" with assms(1)[unfolded ap_spec_def, rule_format, of gs bs ps hs data] assms(2-11) have "crit_pair_cbelow_on d m (fst ` B) (fst g) (fst h)" using assms(10) * by metis thus ?thesis by (rule crit_pair_cbelow_sym) qed qed lemma ap_spec_Nil_subset: assumes "ap_spec ap" shows "set (ap gs bs ps [] data) \ set ps" using ap_specD1[OF assms] by fastforce lemma ap_spec_fst_subset: assumes "ap_spec ap" shows "fst ` set (ap gs bs ps hs data) \ fst ` set ps \ set hs" proof - from ap_specD1[OF assms] have "fst ` set (ap gs bs ps hs data) \ fst ` (set ps \ set hs \ (set gs \ set bs \ set hs))" by (rule image_mono) thus ?thesis by auto qed lemma ap_spec_snd_subset: assumes "ap_spec ap" shows "snd ` set (ap gs bs ps hs data) \ snd ` set ps \ set gs \ set bs \ set hs" proof - from ap_specD1[OF assms] have "snd ` set (ap gs bs ps hs data) \ snd ` (set ps \ set hs \ (set gs \ set bs \ set hs))" by (rule image_mono) thus ?thesis by auto qed lemma ap_spec_inE: assumes "ap_spec ap" and "(p, q) \ set (ap gs bs ps hs data)" assumes 1: "(p, q) \ set ps \ thesis" assumes 2: "p \ set hs \ q \ set gs \ set bs \ set hs \ thesis" shows thesis proof - from assms(2) ap_specD1[OF assms(1)] have "(p, q) \ set ps \ set hs \ (set gs \ set bs \ set hs)" .. thus ?thesis proof assume "(p, q) \ set ps" thus ?thesis by (rule 1) next assume "(p, q) \ set hs \ (set gs \ set bs \ set hs)" hence "p \ set hs" and "q \ set gs \ set bs \ set hs" by blast+ thus ?thesis by (rule 2) qed qed lemma subset_Times_ap: assumes "ap_spec ap" and "ab_spec ab" and "set ps \ set bs \ (set gs \ set bs)" shows "set (ap gs bs (ps -- sps) hs data) \ set (ab gs bs hs data) \ (set gs \ set (ab gs bs hs data))" proof fix p q assume "(p, q) \ set (ap gs bs (ps -- sps) hs data)" with assms(1) show "(p, q) \ set (ab gs bs hs data) \ (set gs \ set (ab gs bs hs data))" proof (rule ap_spec_inE) assume "(p, q) \ set (ps -- sps)" hence "(p, q) \ set ps" by (simp add: set_diff_list) from this assms(3) have "(p, q) \ set bs \ (set gs \ set bs)" .. hence "p \ set bs" and "q \ set gs \ set bs" by blast+ thus ?thesis by (auto simp add: ab_specD1[OF assms(2)]) next assume "p \ set hs" and "q \ set gs \ set bs \ set hs" thus ?thesis by (simp add: ab_specD1[OF assms(2)]) qed qed subsubsection \Function \args_to_set\\ definition args_to_set :: "('t, 'b::field, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list \ ('t \\<^sub>0 'b) set" where "args_to_set x = fst ` (set (fst x) \ set (fst (snd x)) \ fst ` set (snd (snd x)) \ snd ` set (snd (snd x)))" lemma args_to_set_alt: "args_to_set (gs, bs, ps) = fst ` set gs \ fst ` set bs \ fst ` fst ` set ps \ fst ` snd ` set ps" by (simp add: args_to_set_def image_Un) lemma args_to_set_subset_Times: assumes "set ps \ set bs \ (set gs \ set bs)" shows "args_to_set (gs, bs, ps) = fst ` set gs \ fst ` set bs" unfolding args_to_set_alt using assms by auto lemma args_to_set_subset: assumes "ap_spec ap" and "ab_spec ab" shows "args_to_set (gs, ab gs bs hs data, ap gs bs ps hs data) \ fst ` (set gs \ set bs \ fst ` set ps \ snd ` set ps \ set hs)" (is "?l \ fst ` ?r") proof (simp only: args_to_set_alt Un_subset_iff, intro conjI image_mono) show "set (ab gs bs hs data) \ ?r" by (auto simp add: ab_specD1[OF assms(2)]) next from assms(1) have "fst ` set (ap gs bs ps hs data) \ fst ` set ps \ set hs" by (rule ap_spec_fst_subset) thus "fst ` set (ap gs bs ps hs data) \ ?r" by blast next from assms(1) have "snd ` set (ap gs bs ps hs data) \ snd ` set ps \ set gs \ set bs \ set hs" by (rule ap_spec_snd_subset) thus "snd ` set (ap gs bs ps hs data) \ ?r" by blast qed blast lemma args_to_set_alt2: assumes "ap_spec ap" and "ab_spec ab" and "set ps \ set bs \ (set gs \ set bs)" shows "args_to_set (gs, ab gs bs hs data, ap gs bs (ps -- sps) hs data) = fst ` (set gs \ set bs \ set hs)" (is "?l = fst ` ?r") proof from assms(1, 2) have "?l \ fst ` (set gs \ set bs \ fst ` set (ps -- sps) \ snd ` set (ps -- sps) \ set hs)" by (rule args_to_set_subset) also have "... \ fst ` ?r" proof (rule image_mono) have "set gs \ set bs \ fst ` set (ps -- sps) \ snd ` set (ps -- sps) \ set hs \ set gs \ set bs \ fst ` set ps \ snd ` set ps \ set hs" by (auto simp: set_diff_list) also from assms(3) have "... \ ?r" by fastforce finally show "set gs \ set bs \ fst ` set (ps -- sps) \ snd ` set (ps -- sps) \ set hs \ ?r" . qed finally show "?l \ fst ` ?r" . next from assms(2) have eq: "set (ab gs bs hs data) = set bs \ set hs" by (rule ab_specD1) have "fst ` ?r \ fst ` set gs \ fst ` set (ab gs bs hs data)" unfolding eq using assms(3) by fastforce also have "... \ ?l" unfolding args_to_set_alt by fastforce finally show "fst ` ?r \ ?l" . qed lemma args_to_set_subset1: assumes "set gs1 \ set gs2" shows "args_to_set (gs1, bs, ps) \ args_to_set (gs2, bs, ps)" using assms by (auto simp add: args_to_set_alt) lemma args_to_set_subset2: assumes "set bs1 \ set bs2" shows "args_to_set (gs, bs1, ps) \ args_to_set (gs, bs2, ps)" using assms by (auto simp add: args_to_set_alt) lemma args_to_set_subset3: assumes "set ps1 \ set ps2" shows "args_to_set (gs, bs, ps1) \ args_to_set (gs, bs, ps2)" using assms unfolding args_to_set_alt by blast subsubsection \Functions \count_const_lt_components\, \count_rem_comps\ and \full_gb\\ definition rem_comps_spec :: "('t, 'b::zero, 'c) pdata list \ nat \ 'd \ bool" where "rem_comps_spec bs data \ (card (component_of_term ` Keys (fst ` set bs)) = fst data + card (const_lt_component ` (fst ` set bs - {0}) - {None}))" definition count_const_lt_components :: "('t, 'b::zero, 'c) pdata' list \ nat" where "count_const_lt_components hs = length (remdups (filter (\x. x \ None) (map (const_lt_component \ fst) hs)))" definition count_rem_components :: "('t, 'b::zero, 'c) pdata' list \ nat" where "count_rem_components bs = length (remdups (map component_of_term (Keys_to_list (map fst bs)))) - count_const_lt_components [b\bs . fst b \ 0]" lemma count_const_lt_components_alt: "count_const_lt_components hs = card (const_lt_component ` fst ` set hs - {None})" by (simp add: count_const_lt_components_def card_set[symmetric] set_diff_eq image_comp del: not_None_eq) lemma count_rem_components_alt: "count_rem_components bs + card (const_lt_component ` (fst ` set bs - {0}) - {None}) = card (component_of_term ` Keys (fst ` set bs))" proof - have eq: "fst ` {x \ set bs. fst x \ 0} = fst ` set bs - {0}" by fastforce have "card (const_lt_component ` (fst ` set bs - {0}) - {None}) \ card (component_of_term ` Keys (fst ` set bs))" by (rule card_const_lt_component_le, rule finite_imageI, fact finite_set) thus ?thesis by (simp add: count_rem_components_def card_set[symmetric] set_Keys_to_list count_const_lt_components_alt eq) qed lemma rem_comps_spec_count_rem_components: "rem_comps_spec bs (count_rem_components bs, data)" by (simp only: rem_comps_spec_def fst_conv count_rem_components_alt) definition full_gb :: "('t, 'b, 'c) pdata list \ ('t, 'b::zero_neq_one, 'c::default) pdata list" where "full_gb bs = map (\k. (monomial 1 (term_of_pair (0, k)), 0, default)) (remdups (map component_of_term (Keys_to_list (map fst bs))))" lemma fst_set_full_gb: "fst ` set (full_gb bs) = (\v. monomial 1 (term_of_pair (0, component_of_term v))) ` Keys (fst ` set bs)" by (simp add: full_gb_def set_Keys_to_list image_comp) lemma Keys_full_gb: "Keys (fst ` set (full_gb bs)) = (\v. term_of_pair (0, component_of_term v)) ` Keys (fst ` set bs)" by (auto simp add: fst_set_full_gb Keys_def image_image) lemma pps_full_gb: "pp_of_term ` Keys (fst ` set (full_gb bs)) \ {0}" by (simp add: Keys_full_gb image_comp image_subset_iff term_simps) lemma components_full_gb: "component_of_term ` Keys (fst ` set (full_gb bs)) = component_of_term ` Keys (fst ` set bs)" by (simp add: Keys_full_gb image_comp, rule image_cong, fact refl, simp add: term_simps) lemma full_gb_is_full_pmdl: "is_full_pmdl (fst ` set (full_gb bs))" for bs::"('t, 'b::field, 'c::default) pdata list" proof (rule is_full_pmdlI_lt_finite) from finite_set show "finite (fst ` set (full_gb bs))" by (rule finite_imageI) next fix k assume "k \ component_of_term ` Keys (fst ` set (full_gb bs))" then obtain v where "v \ Keys (fst ` set (full_gb bs))" and k: "k = component_of_term v" .. from this(1) obtain b where "b \ fst ` set (full_gb bs)" and "v \ keys b" by (rule in_KeysE) from this(1) obtain u where "u \ Keys (fst ` set bs)" and b: "b = monomial 1 (term_of_pair (0, component_of_term u))" unfolding fst_set_full_gb .. have lt: "lt b = term_of_pair (0, component_of_term u)" by (simp add: b lt_monomial) from \v \ keys b\ have v: "v = term_of_pair (0, component_of_term u)" by (simp add: b) show "\b\fst ` set (full_gb bs). b \ 0 \ component_of_term (lt b) = k \ lp b = 0" proof (intro bexI conjI) show "b \ 0" by (simp add: b monomial_0_iff) next show "component_of_term (lt b) = k" by (simp add: lt term_simps k v) next show "lp b = 0" by (simp add: lt term_simps) qed fact qed text \In fact, @{thm full_gb_is_full_pmdl} also holds if @{typ 'b} is no field.\ lemma full_gb_isGB: "is_Groebner_basis (fst ` set (full_gb bs))" proof (rule Buchberger_criterion_finite) from finite_set show "finite (fst ` set (full_gb bs))" by (rule finite_imageI) next fix p q :: "'t \\<^sub>0 'b" assume "p \ fst ` set (full_gb bs)" then obtain v where p: "p = monomial 1 (term_of_pair (0, component_of_term v))" unfolding fst_set_full_gb .. hence lt: "component_of_term (lt p) = component_of_term v" by (simp add: lt_monomial term_simps) assume "q \ fst ` set (full_gb bs)" then obtain u where q: "q = monomial 1 (term_of_pair (0, component_of_term u))" unfolding fst_set_full_gb .. hence lq: "component_of_term (lt q) = component_of_term u" by (simp add: lt_monomial term_simps) assume "component_of_term (lt p) = component_of_term (lt q)" hence "component_of_term v = component_of_term u" by (simp only: lt lq) hence "p = q" by (simp only: p q) moreover assume "p \ q" ultimately show "(red (fst ` set (full_gb bs)))\<^sup>*\<^sup>* (spoly p q) 0" by (simp only:) qed subsubsection \Specification of the @{emph \completion\} parameter\ definition compl_struct :: "('t, 'b::field, 'c, 'd) complT \ bool" where "compl_struct compl \ (\gs bs ps sps data. sps \ [] \ set sps \ set ps \ (\d. dickson_grading d \ dgrad_p_set_le d (fst ` (set (fst (compl gs bs (ps -- sps) sps data)))) (args_to_set (gs, bs, ps))) \ component_of_term ` Keys (fst ` (set (fst (compl gs bs (ps -- sps) sps data)))) \ component_of_term ` Keys (args_to_set (gs, bs, ps)) \ 0 \ fst ` set (fst (compl gs bs (ps -- sps) sps data)) \ (\h\set (fst (compl gs bs (ps -- sps) sps data)). \b\set gs \ set bs. fst b \ 0 \ \ lt (fst b) adds\<^sub>t lt (fst h)))" lemma compl_structI: assumes "\d gs bs ps sps data. dickson_grading d \ sps \ [] \ set sps \ set ps \ dgrad_p_set_le d (fst ` (set (fst (compl gs bs (ps -- sps) sps data)))) (args_to_set (gs, bs, ps))" assumes "\gs bs ps sps data. sps \ [] \ set sps \ set ps \ component_of_term ` Keys (fst ` (set (fst (compl gs bs (ps -- sps) sps data)))) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" assumes "\gs bs ps sps data. sps \ [] \ set sps \ set ps \ 0 \ fst ` set (fst (compl gs bs (ps -- sps) sps data))" assumes "\gs bs ps sps h b data. sps \ [] \ set sps \ set ps \ h \ set (fst (compl gs bs (ps -- sps) sps data)) \ b \ set gs \ set bs \ fst b \ 0 \ \ lt (fst b) adds\<^sub>t lt (fst h)" shows "compl_struct compl" unfolding compl_struct_def using assms by auto lemma compl_structD1: assumes "compl_struct compl" and "dickson_grading d" and "sps \ []" and "set sps \ set ps" shows "dgrad_p_set_le d (fst ` (set (fst (compl gs bs (ps -- sps) sps data)))) (args_to_set (gs, bs, ps))" using assms unfolding compl_struct_def by blast lemma compl_structD2: assumes "compl_struct compl" and "sps \ []" and "set sps \ set ps" shows "component_of_term ` Keys (fst ` (set (fst (compl gs bs (ps -- sps) sps data)))) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" using assms unfolding compl_struct_def by blast lemma compl_structD3: assumes "compl_struct compl" and "sps \ []" and "set sps \ set ps" shows "0 \ fst ` set (fst (compl gs bs (ps -- sps) sps data))" using assms unfolding compl_struct_def by blast lemma compl_structD4: assumes "compl_struct compl" and "sps \ []" and "set sps \ set ps" and "h \ set (fst (compl gs bs (ps -- sps) sps data))" and "b \ set gs \ set bs" and "fst b \ 0" shows "\ lt (fst b) adds\<^sub>t lt (fst h)" using assms unfolding compl_struct_def by blast definition struct_spec :: "('t, 'b::field, 'c, 'd) selT \ ('t, 'b, 'c, 'd) apT \ ('t, 'b, 'c, 'd) abT \ ('t, 'b, 'c, 'd) complT \ bool" where "struct_spec sel ap ab compl \ (sel_spec sel \ ap_spec ap \ ab_spec ab \ compl_struct compl)" lemma struct_specI: assumes "sel_spec sel" and "ap_spec ap" and "ab_spec ab" and "compl_struct compl" shows "struct_spec sel ap ab compl" unfolding struct_spec_def using assms by (intro conjI) lemma struct_specD1: assumes "struct_spec sel ap ab compl" shows "sel_spec sel" using assms unfolding struct_spec_def by (elim conjE) lemma struct_specD2: assumes "struct_spec sel ap ab compl" shows "ap_spec ap" using assms unfolding struct_spec_def by (elim conjE) lemma struct_specD3: assumes "struct_spec sel ap ab compl" shows "ab_spec ab" using assms unfolding struct_spec_def by (elim conjE) lemma struct_specD4: assumes "struct_spec sel ap ab compl" shows "compl_struct compl" using assms unfolding struct_spec_def by (elim conjE) lemmas struct_specD = struct_specD1 struct_specD2 struct_specD3 struct_specD4 definition compl_pmdl :: "('t, 'b::field, 'c, 'd) complT \ bool" where "compl_pmdl compl \ (\gs bs ps sps data. is_Groebner_basis (fst ` set gs) \ sps \ [] \ set sps \ set ps \ unique_idx (gs @ bs) data \ fst ` (set (fst (compl gs bs (ps -- sps) sps data))) \ pmdl (args_to_set (gs, bs, ps)))" lemma compl_pmdlI: assumes "\gs bs ps sps data. is_Groebner_basis (fst ` set gs) \ sps \ [] \ set sps \ set ps \ unique_idx (gs @ bs) data \ fst ` (set (fst (compl gs bs (ps -- sps) sps data))) \ pmdl (args_to_set (gs, bs, ps))" shows "compl_pmdl compl" unfolding compl_pmdl_def using assms by blast lemma compl_pmdlD: assumes "compl_pmdl compl" and "is_Groebner_basis (fst ` set gs)" and "sps \ []" and "set sps \ set ps" and "unique_idx (gs @ bs) data" shows "fst ` (set (fst (compl gs bs (ps -- sps) sps data))) \ pmdl (args_to_set (gs, bs, ps))" using assms unfolding compl_pmdl_def by blast definition compl_conn :: "('t, 'b::field, 'c, 'd) complT \ bool" where "compl_conn compl \ (\d m gs bs ps sps p q data. dickson_grading d \ fst ` set gs \ dgrad_p_set d m \ is_Groebner_basis (fst ` set gs) \ fst ` set bs \ dgrad_p_set d m \ set ps \ set bs \ (set gs \ set bs) \ sps \ [] \ set sps \ set ps \ unique_idx (gs @ bs) data \ (p, q) \ set sps \ fst p \ 0 \ fst q \ 0 \ crit_pair_cbelow_on d m (fst ` (set gs \ set bs) \ fst ` set (fst (compl gs bs (ps -- sps) sps data))) (fst p) (fst q))" text \Informally, \compl_conn compl\ means that, for suitable arguments \gs\, \bs\, \ps\ and \sps\, the value of \compl gs bs ps sps\ is a list \hs\ such that the critical pairs of all elements in \sps\ can be connected modulo \set gs \ set bs \ set hs\.\ lemma compl_connI: assumes "\d m gs bs ps sps p q data. dickson_grading d \ fst ` set gs \ dgrad_p_set d m \ is_Groebner_basis (fst ` set gs) \ fst ` set bs \ dgrad_p_set d m \ set ps \ set bs \ (set gs \ set bs) \ sps \ [] \ set sps \ set ps \ unique_idx (gs @ bs) data \ (p, q) \ set sps \ fst p \ 0 \ fst q \ 0 \ crit_pair_cbelow_on d m (fst ` (set gs \ set bs) \ fst ` set (fst (compl gs bs (ps -- sps) sps data))) (fst p) (fst q)" shows "compl_conn compl" unfolding compl_conn_def using assms by presburger lemma compl_connD: assumes "compl_conn compl" and "dickson_grading d" and "fst ` set gs \ dgrad_p_set d m" and "is_Groebner_basis (fst ` set gs)" and "fst ` set bs \ dgrad_p_set d m" and "set ps \ set bs \ (set gs \ set bs)" and "sps \ []" and "set sps \ set ps" and "unique_idx (gs @ bs) data" and "(p, q) \ set sps" and "fst p \ 0" and "fst q \ 0" shows "crit_pair_cbelow_on d m (fst ` (set gs \ set bs) \ fst ` set (fst (compl gs bs (ps -- sps) sps data))) (fst p) (fst q)" using assms unfolding compl_conn_def Un_assoc by blast subsubsection \Function \gb_schema_dummy\\ definition (in -) add_indices :: "(('a, 'b, 'c) pdata' list \ 'd) \ (nat \ 'd) \ (('a, 'b, 'c) pdata list \ nat \ 'd)" where [code del]: "add_indices ns data = (map_idx (\h i. (fst h, i, snd h)) (fst ns) (fst data), fst data + length (fst ns), snd ns)" lemma (in -) add_indices_code [code]: "add_indices (ns, data) (n, data') = (map_idx (\(h, d) i. (h, i, d)) ns n, n + length ns, data)" by (simp add: add_indices_def case_prod_beta') lemma fst_add_indices: "map fst (fst (add_indices ns data')) = map fst (fst ns)" by (simp add: add_indices_def map_map_idx map_idx_no_idx) corollary fst_set_add_indices: "fst ` set (fst (add_indices ns data')) = fst ` set (fst ns)" using fst_add_indices by (metis set_map) lemma in_set_add_indicesE: assumes "f \ set (fst (add_indices aux data))" obtains i where "i < length (fst aux)" and "f = (fst ((fst aux) ! i), fst data + i, snd ((fst aux) ! i))" proof - let ?hs = "fst (add_indices aux data)" from assms obtain i where "i < length ?hs" and "f = ?hs ! i" by (metis in_set_conv_nth) from this(1) have "i < length (fst aux)" by (simp add: add_indices_def) hence "?hs ! i = (fst ((fst aux) ! i), fst data + i, snd ((fst aux) ! i))" unfolding add_indices_def fst_conv by (rule map_idx_nth) hence "f = (fst ((fst aux) ! i), fst data + i, snd ((fst aux) ! i))" by (simp add: \f = ?hs ! i\) with \i < length (fst aux)\ show ?thesis .. qed definition gb_schema_aux_term1 :: "((('t, 'b::field, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list) \ (('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list)) set" where "gb_schema_aux_term1 = {(a, b::('t, 'b, 'c) pdata list). (fst ` set a) \p (fst ` set b)} <*lex*> (measure (card \ set))" definition gb_schema_aux_term2 :: "('a \ nat) \ ('t, 'b::field, 'c) pdata list \ ((('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list) \ (('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list)) set" where "gb_schema_aux_term2 d gs = {(a, b). dgrad_p_set_le d (args_to_set (gs, a)) (args_to_set (gs, b)) \ component_of_term ` Keys (args_to_set (gs, a)) \ component_of_term ` Keys (args_to_set (gs, b))}" definition gb_schema_aux_term where "gb_schema_aux_term d gs = gb_schema_aux_term1 \ gb_schema_aux_term2 d gs" text \@{const gb_schema_aux_term} is needed for proving termination of function \gb_schema_aux\.\ lemma gb_schema_aux_term1_wf_on: assumes "dickson_grading d" and "finite K" shows "wfp_on (\x y. (x, y) \ gb_schema_aux_term1) {x::(('t, 'b, 'c) pdata list) \ ((('t, 'b::field, 'c) pdata_pair list)). args_to_set (gs, x) \ dgrad_p_set d m \ component_of_term ` Keys (args_to_set (gs, x)) \ K}" proof (rule wfp_onI_min) let ?B = "dgrad_p_set d m" let ?A = "{x::(('t, 'b, 'c) pdata list) \ ((('t, 'b, 'c) pdata_pair list)). args_to_set (gs, x) \ ?B \ component_of_term ` Keys (args_to_set (gs, x)) \ K}" let ?C = "Pow ?B \ {F. component_of_term ` Keys F \ K}" have A_sub_Pow: "(image fst) ` set ` fst ` ?A \ ?C" proof fix x assume "x \ (image fst) ` set ` fst ` ?A" then obtain x1 where "x1 \ set ` fst ` ?A" and x: "x = fst ` x1" by auto from this(1) obtain x2 where "x2 \ fst ` ?A" and x1: "x1 = set x2" by auto from this(1) obtain x3 where "x3 \ ?A" and x2: "x2 = fst x3" by auto from this(1) have "args_to_set (gs, x3) \ ?B" and "component_of_term ` Keys (args_to_set (gs, x3)) \ K" by simp_all thus "x \ ?C" by (simp add: args_to_set_def x x1 x2 image_Un Keys_Un) qed fix x Q assume "x \ Q" and "Q \ ?A" have Q_sub_A: "(image fst) ` set ` fst ` Q \ (image fst) ` set ` fst ` ?A" by ((rule image_mono)+, fact) from assms have "wfp_on (\p) ?C" by (rule red_supset_wf_on) moreover have "fst ` set (fst x) \ (image fst) ` set ` fst ` Q" by (rule, fact refl, rule, fact refl, rule, fact refl, simp add: \x \ Q\) moreover from Q_sub_A A_sub_Pow have "(image fst) ` set ` fst ` Q \ ?C" by (rule subset_trans) ultimately obtain z1 where "z1 \ (image fst) ` set ` fst ` Q" and 2: "\y. y \p z1 \ y \ (image fst) ` set ` fst ` Q" by (rule wfp_onE_min, auto) from this(1) obtain x1 where "x1 \ Q" and z1: "z1 = fst ` set (fst x1)" by auto let ?Q2 = "{q \ Q. fst ` set (fst q) = z1}" have "snd x1 \ snd ` ?Q2" by (rule, fact refl, simp add: \x1 \ Q\ z1) with wf_measure obtain z2 where "z2 \ snd ` ?Q2" and 3: "\y. (y, z2) \ measure (card \ set) \ y \ snd ` ?Q2" by (rule wfE_min, blast) from this(1) obtain z where "z \ ?Q2" and z2: "z2 = snd z" .. from this(1) have "z \ Q" and eq1: "fst ` set (fst z) = z1" by blast+ from this(1) show "\z\Q. \y\?A. (y, z) \ gb_schema_aux_term1 \ y \ Q" proof show "\y\?A. (y, z) \ gb_schema_aux_term1 \ y \ Q" proof (intro ballI impI) fix y assume "y \ ?A" assume "(y, z) \ gb_schema_aux_term1" hence "(fst ` set (fst y) \p z1 \ (fst y = fst z \ (snd y, z2) \ measure (card \ set)))" by (simp add: gb_schema_aux_term1_def eq1[symmetric] z2 in_lex_prod_alt) thus "y \ Q" proof (elim disjE conjE) assume "fst ` set (fst y) \p z1" hence "fst ` set (fst y) \ (image fst) ` set ` fst ` Q" by (rule 2) thus ?thesis by auto next assume "(snd y, z2) \ measure (card \ set)" hence "snd y \ snd ` ?Q2" by (rule 3) hence "y \ ?Q2" by blast moreover assume "fst y = fst z" ultimately show ?thesis by (simp add: eq1) qed qed qed qed lemma gb_schema_aux_term_wf: assumes "dickson_grading d" shows "wf (gb_schema_aux_term d gs)" proof (rule wfI_min) fix x::"(('t, 'b, 'c) pdata list) \ (('t, 'b, 'c) pdata_pair list)" and Q assume "x \ Q" let ?A = "args_to_set (gs, x)" have "finite ?A" by (simp add: args_to_set_def) then obtain m where A: "?A \ dgrad_p_set d m" by (rule dgrad_p_set_exhaust) define K where "K = component_of_term ` Keys ?A" from \finite ?A\ have "finite K" unfolding K_def by (rule finite_imp_finite_component_Keys) let ?B = "dgrad_p_set d m" let ?Q = "{q \ Q. args_to_set (gs, q) \ ?B \ component_of_term ` Keys (args_to_set (gs, q)) \ K}" from assms \finite K\ have "wfp_on (\x y. (x, y) \ gb_schema_aux_term1) {x. args_to_set (gs, x) \ ?B \ component_of_term ` Keys (args_to_set (gs, x)) \ K}" by (rule gb_schema_aux_term1_wf_on) moreover from \x \ Q\ A have "x \ ?Q" by (simp add: K_def) moreover have "?Q \ {x. args_to_set (gs, x) \ ?B \ component_of_term ` Keys (args_to_set (gs, x)) \ K}" by auto ultimately obtain z where "z \ ?Q" and *: "\y. (y, z) \ gb_schema_aux_term1 \ y \ ?Q" by (rule wfp_onE_min, blast) from this(1) have "z \ Q" and a: "args_to_set (gs, z) \ ?B" and b: "component_of_term ` Keys (args_to_set (gs, z)) \ K" by simp_all from this(1) show "\z\Q. \y. (y, z) \ gb_schema_aux_term d gs \ y \ Q" proof show "\y. (y, z) \ gb_schema_aux_term d gs \ y \ Q" proof (intro allI impI) fix y assume "(y, z) \ gb_schema_aux_term d gs" hence "(y, z) \ gb_schema_aux_term1" and "(y, z) \ gb_schema_aux_term2 d gs" by (simp_all add: gb_schema_aux_term_def) from this(2) have "dgrad_p_set_le d (args_to_set (gs, y)) (args_to_set (gs, z))" and comp_sub: "component_of_term ` Keys (args_to_set (gs, y)) \ component_of_term ` Keys (args_to_set (gs, z))" by (simp_all add: gb_schema_aux_term2_def) from this(1) \args_to_set (gs, z) \ ?B\ have "args_to_set (gs, y) \ ?B" by (rule dgrad_p_set_le_dgrad_p_set) moreover from comp_sub b have "component_of_term ` Keys (args_to_set (gs, y)) \ K" by (rule subset_trans) moreover from \(y, z) \ gb_schema_aux_term1\ have "y \ ?Q" by (rule *) ultimately show "y \ Q" by simp qed qed qed lemma dgrad_p_set_le_args_to_set_ab: assumes "dickson_grading d" and "ap_spec ap" and "ab_spec ab" and "compl_struct compl" assumes "sps \ []" and "set sps \ set ps" and "hs = fst (add_indices (compl gs bs (ps -- sps) sps data) data)" shows "dgrad_p_set_le d (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) (args_to_set (gs, bs, ps))" (is "dgrad_p_set_le _ ?l ?r") proof - have "dgrad_p_set_le d ?l (fst ` (set gs \ set bs \ fst ` set (ps -- sps) \ snd ` set (ps -- sps) \ set hs))" by (rule dgrad_p_set_le_subset, rule args_to_set_subset[OF assms(2, 3)]) also have "dgrad_p_set_le d ... ?r" unfolding image_Un proof (intro dgrad_p_set_leI_Un) show "dgrad_p_set_le d (fst ` set gs) (args_to_set (gs, bs, ps))" by (rule dgrad_p_set_le_subset, auto simp add: args_to_set_def) next show "dgrad_p_set_le d (fst ` set bs) (args_to_set (gs, bs, ps))" by (rule dgrad_p_set_le_subset, auto simp add: args_to_set_def) next show "dgrad_p_set_le d (fst ` fst ` set (ps -- sps)) (args_to_set (gs, bs, ps))" by (rule dgrad_p_set_le_subset, auto simp add: args_to_set_def set_diff_list) next show "dgrad_p_set_le d (fst ` snd ` set (ps -- sps)) (args_to_set (gs, bs, ps))" by (rule dgrad_p_set_le_subset, auto simp add: args_to_set_def set_diff_list) next from assms(4, 1, 5, 6) show "dgrad_p_set_le d (fst ` set hs) (args_to_set (gs, bs, ps))" unfolding assms(7) fst_set_add_indices by (rule compl_structD1) qed finally show ?thesis . qed corollary dgrad_p_set_le_args_to_set_struct: assumes "dickson_grading d" and "struct_spec sel ap ab compl" and "ps \ []" assumes "sps = sel gs bs ps data" and "hs = fst (add_indices (compl gs bs (ps -- sps) sps data) data)" shows "dgrad_p_set_le d (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) (args_to_set (gs, bs, ps))" proof - from assms(2) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ from sel assms(3) have "sps \ []" and "set sps \ set ps" unfolding assms(4) by (rule sel_specD1, rule sel_specD2) from assms(1) ap ab compl this assms(5) show ?thesis by (rule dgrad_p_set_le_args_to_set_ab) qed lemma components_subset_ab: assumes "ap_spec ap" and "ab_spec ab" and "compl_struct compl" assumes "sps \ []" and "set sps \ set ps" and "hs = fst (add_indices (compl gs bs (ps -- sps) sps data) data)" shows "component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" (is "?l \ ?r") proof - have "?l \ component_of_term ` Keys (fst ` (set gs \ set bs \ fst ` set (ps -- sps) \ snd ` set (ps -- sps) \ set hs))" by (rule image_mono, rule Keys_mono, rule args_to_set_subset[OF assms(1, 2)]) also have "... \ ?r" unfolding image_Un Keys_Un Un_subset_iff proof (intro conjI) show "component_of_term ` Keys (fst ` set gs) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" by (rule image_mono, rule Keys_mono, auto simp add: args_to_set_def) next show "component_of_term ` Keys (fst ` set bs) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" by (rule image_mono, rule Keys_mono, auto simp add: args_to_set_def) next show "component_of_term ` Keys (fst ` fst ` set (ps -- sps)) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" by (rule image_mono, rule Keys_mono, auto simp add: set_diff_list args_to_set_def) next show "component_of_term ` Keys (fst ` snd ` set (ps -- sps)) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" by (rule image_mono, rule Keys_mono, auto simp add: args_to_set_def set_diff_list) next from assms(3, 4, 5) show "component_of_term ` Keys (fst ` set hs) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" unfolding assms(6) fst_set_add_indices by (rule compl_structD2) qed finally show ?thesis . qed corollary components_subset_struct: assumes "struct_spec sel ap ab compl" and "ps \ []" assumes "sps = sel gs bs ps data" and "hs = fst (add_indices (compl gs bs (ps -- sps) sps data) data)" shows "component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" proof - from assms(1) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ from sel assms(2) have "sps \ []" and "set sps \ set ps" unfolding assms(3) by (rule sel_specD1, rule sel_specD2) from ap ab compl this assms(4) show ?thesis by (rule components_subset_ab) qed corollary components_struct: assumes "struct_spec sel ap ab compl" and "ps \ []" and "set ps \ set bs \ (set gs \ set bs)" assumes "sps = sel gs bs ps data" and "hs = fst (add_indices (compl gs bs (ps -- sps) sps data) data)" shows "component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) = component_of_term ` Keys (args_to_set (gs, bs, ps))" (is "?l = ?r") proof from assms(1, 2, 4, 5) show "?l \ ?r" by (rule components_subset_struct) next from assms(1) have ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ from ap ab assms(3) have sub: "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) show "?r \ ?l" by (simp add: args_to_set_subset_Times[OF sub] args_to_set_subset_Times[OF assms(3)] ab_specD1[OF ab], rule image_mono, rule Keys_mono, blast) qed lemma struct_spec_red_supset: assumes "struct_spec sel ap ab compl" and "ps \ []" and "sps = sel gs bs ps data" and "hs = fst (add_indices (compl gs bs (ps -- sps) sps data) data)" and "hs \ []" shows "(fst ` set (ab gs bs hs data')) \p (fst ` set bs)" proof - from assms(5) have "set hs \ {}" by simp then obtain h' where "h' \ set hs" by fastforce let ?h = "fst h'" let ?m = "monomial (lc ?h) (lt ?h)" from \h' \ set hs\ have h_in: "?h \ fst ` set hs" by simp hence "?h \ fst ` set (fst (compl gs bs (ps -- sps) sps data))" by (simp only: assms(4) fst_set_add_indices) then obtain h'' where h''_in: "h'' \ set (fst (compl gs bs (ps -- sps) sps data))" and "?h = fst h''" .. from assms(1) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ from sel assms(2) have "sps \ []" and "set sps \ set ps" unfolding assms(3) by (rule sel_specD1, rule sel_specD2) from h_in compl_structD3[OF compl this] have "?h \ 0" unfolding assms(4) fst_set_add_indices by metis show ?thesis proof (simp add: ab_specD1[OF ab] image_Un, rule) fix q assume "is_red (fst ` set bs) q" moreover have "fst ` set bs \ fst ` set bs \ fst ` set hs" by simp ultimately show "is_red (fst ` set bs \ fst ` set hs) q" by (rule is_red_subset) next from \?h \ 0\ have "lc ?h \ 0" by (rule lc_not_0) moreover have "?h \ {?h}" .. ultimately have "is_red {?h} ?m" using \?h \ 0\ adds_term_refl by (rule is_red_monomialI) moreover have "{?h} \ fst ` set bs \ fst ` set hs" using h_in by simp ultimately show "is_red (fst ` set bs \ fst ` set hs) ?m" by (rule is_red_subset) next show "\ is_red (fst ` set bs) ?m" proof assume "is_red (fst ` set bs) ?m" then obtain b' where "b' \ fst ` set bs" and "b' \ 0" and "lt b' adds\<^sub>t lt ?h" by (rule is_red_monomialE) from this(1) obtain b where "b \ set bs" and b': "b' = fst b" .. from this(1) have "b \ set gs \ set bs" by simp from \b' \ 0\ have "fst b \ 0" by (simp add: b') with compl \sps \ []\ \set sps \ set ps\ h''_in \b \ set gs \ set bs\ have "\ lt (fst b) adds\<^sub>t lt ?h" unfolding \?h = fst h''\ by (rule compl_structD4) from this \lt b' adds\<^sub>t lt ?h\ show False by (simp add: b') qed qed qed lemma unique_idx_append: assumes "unique_idx gs data" and "(hs, data') = add_indices aux data" shows "unique_idx (gs @ hs) data'" proof - from assms(2) have hs: "hs = fst (add_indices aux data)" and data': "data' = snd (add_indices aux data)" by (metis fst_conv, metis snd_conv) have len: "length hs = length (fst aux)" by (simp add: hs add_indices_def) have eq: "fst data' = fst data + length hs" by (simp add: data' add_indices_def hs) show ?thesis proof (rule unique_idxI) fix f g assume "f \ set (gs @ hs)" and "g \ set (gs @ hs)" hence d1: "f \ set gs \ set hs" and d2: "g \ set gs \ set hs" by simp_all assume id_eq: "fst (snd f) = fst (snd g)" from d1 show "f = g" proof assume "f \ set gs" from d2 show ?thesis proof assume "g \ set gs" from assms(1) \f \ set gs\ this id_eq show ?thesis by (rule unique_idxD1) next assume "g \ set hs" then obtain j where "g = (fst (fst aux ! j), fst data + j, snd (fst aux ! j))" unfolding hs by (rule in_set_add_indicesE) hence "fst (snd g) = fst data + j" by simp moreover from assms(1) \f \ set gs\ have "fst (snd f) < fst data" by (rule unique_idxD2) ultimately show ?thesis by (simp add: id_eq) qed next assume "f \ set hs" then obtain i where f: "f = (fst (fst aux ! i), fst data + i, snd (fst aux ! i))" unfolding hs by (rule in_set_add_indicesE) hence *: "fst (snd f) = fst data + i" by simp from d2 show ?thesis proof assume "g \ set gs" with assms(1) have "fst (snd g) < fst data" by (rule unique_idxD2) with * show ?thesis by (simp add: id_eq) next assume "g \ set hs" then obtain j where g: "g = (fst (fst aux ! j), fst data + j, snd (fst aux ! j))" unfolding hs by (rule in_set_add_indicesE) hence "fst (snd g) = fst data + j" by simp with * have "i = j" by (simp add: id_eq) thus ?thesis by (simp add: f g) qed qed next fix f assume "f \ set (gs @ hs)" hence "f \ set gs \ set hs" by simp thus "fst (snd f) < fst data'" proof assume "f \ set gs" with assms(1) have "fst (snd f) < fst data" by (rule unique_idxD2) also have "... \ fst data'" by (simp add: eq) finally show ?thesis . next assume "f \ set hs" then obtain i where "i < length (fst aux)" and "f = (fst (fst aux ! i), fst data + i, snd (fst aux ! i))" unfolding hs by (rule in_set_add_indicesE) from this(2) have "fst (snd f) = fst data + i" by simp also from \i < length (fst aux)\ have "... < fst data + length (fst aux)" by simp finally show ?thesis by (simp only: eq len) qed qed qed corollary unique_idx_ab: assumes "ab_spec ab" and "unique_idx (gs @ bs) data" and "(hs, data') = add_indices aux data" shows "unique_idx (gs @ ab gs bs hs data') data'" proof - from assms(2, 3) have "unique_idx ((gs @ bs) @ hs) data'" by (rule unique_idx_append) thus ?thesis by (simp add: unique_idx_def ab_specD1[OF assms(1)]) qed lemma rem_comps_spec_struct: assumes "struct_spec sel ap ab compl" and "rem_comps_spec (gs @ bs) data" and "ps \ []" and "set ps \ (set bs) \ (set gs \ set bs)" and "sps = sel gs bs ps (snd data)" and "aux = compl gs bs (ps -- sps) sps (snd data)" and "(hs, data') = add_indices aux (snd data)" shows "rem_comps_spec (gs @ ab gs bs hs data') (fst data - count_const_lt_components (fst aux), data')" proof - from assms(1) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ from ap ab assms(4) have sub: "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) have hs: "hs = fst (add_indices aux (snd data))" by (simp add: assms(7)[symmetric]) from sel assms(3) have "sps \ []" and "set sps \ set ps" unfolding assms(5) by (rule sel_specD1, rule sel_specD2) have eq0: "fst ` set (fst aux) - {0} = fst ` set (fst aux)" by (rule Diff_triv, simp add: Int_insert_right assms(6), rule compl_structD3, fact+) have "component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data')) = component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data'))" by (simp add: args_to_set_subset_Times[OF sub] image_Un) also from assms(1, 3, 4, 5) hs have "... = component_of_term ` Keys (args_to_set (gs, bs, ps))" unfolding assms(6) by (rule components_struct) also have "... = component_of_term ` Keys (fst ` set (gs @ bs))" by (simp add: args_to_set_subset_Times[OF assms(4)] image_Un) finally have eq: "component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data')) = component_of_term ` Keys (fst ` set (gs @ bs))" . from assms(2) have eq2: "card (component_of_term ` Keys (fst ` set (gs @ bs))) = fst data + card (const_lt_component ` (fst ` set (gs @ bs) - {0}) - {None})" (is "?a = _ + ?b") by (simp only: rem_comps_spec_def) have eq3: "card (const_lt_component ` (fst ` set (gs @ ab gs bs hs data') - {0}) - {None}) = ?b + count_const_lt_components (fst aux)" (is "?c = _") proof (simp add: ab_specD1[OF ab] image_Un Un_assoc[symmetric] Un_Diff count_const_lt_components_alt hs fst_set_add_indices eq0, rule card_Un_disjoint) show "finite (const_lt_component ` (fst ` set gs - {0}) - {None} \ (const_lt_component ` (fst ` set bs - {0}) - {None}))" by (intro finite_UnI finite_Diff finite_imageI finite_set) next show "finite (const_lt_component ` fst ` set (fst aux) - {None})" by (rule finite_Diff, intro finite_imageI, fact finite_set) next have "(const_lt_component ` (fst ` (set gs \ set bs) - {0}) - {None}) \ (const_lt_component ` fst ` set (fst aux) - {None}) = (const_lt_component ` (fst ` (set gs \ set bs) - {0}) \ const_lt_component ` fst ` set (fst aux)) - {None}" by blast also have "... = {}" proof (simp, rule, simp, elim conjE) fix k assume "k \ const_lt_component ` (fst ` (set gs \ set bs) - {0})" then obtain b where "b \ set gs \ set bs" and "fst b \ 0" and k1: "k = const_lt_component (fst b)" by blast assume "k \ const_lt_component ` fst ` set (fst aux)" then obtain h where "h \ set (fst aux)" and k2: "k = const_lt_component (fst h)" by blast show "k = None" proof (rule ccontr, simp, elim exE) fix k' assume "k = Some k'" hence "lp (fst b) = 0" and "component_of_term (lt (fst b)) = k'" unfolding k1 by (rule const_lt_component_SomeD1, rule const_lt_component_SomeD2) moreover from \k = Some k'\ have "lp (fst h) = 0" and "component_of_term (lt (fst h)) = k'" unfolding k2 by (rule const_lt_component_SomeD1, rule const_lt_component_SomeD2) ultimately have "lt (fst b) adds\<^sub>t lt (fst h)" by (simp add: adds_term_def) moreover from compl \sps \ []\ \set sps \ set ps\ \h \ set (fst aux)\ \b \ set gs \ set bs\ \fst b \ 0\ have "\ lt (fst b) adds\<^sub>t lt (fst h)" unfolding assms(6) by (rule compl_structD4) ultimately show False by simp qed qed finally show "(const_lt_component ` (fst ` set gs - {0}) - {None} \ (const_lt_component ` (fst ` set bs - {0}) - {None})) \ (const_lt_component ` fst ` set (fst aux) - {None}) = {}" by (simp only: Un_Diff image_Un) qed have "?c \ ?a" unfolding eq[symmetric] by (rule card_const_lt_component_le, rule finite_imageI, fact finite_set) hence le: "count_const_lt_components (fst aux) \ fst data" by (simp only: eq2 eq3) show ?thesis by (simp only: rem_comps_spec_def eq eq2 eq3, simp add: le) qed lemma pmdl_struct: assumes "struct_spec sel ap ab compl" and "compl_pmdl compl" and "is_Groebner_basis (fst ` set gs)" and "ps \ []" and "set ps \ (set bs) \ (set gs \ set bs)" and "unique_idx (gs @ bs) (snd data)" and "sps = sel gs bs ps (snd data)" and "aux = compl gs bs (ps -- sps) sps (snd data)" and "(hs, data') = add_indices aux (snd data)" shows "pmdl (fst ` set (gs @ ab gs bs hs data')) = pmdl (fst ` set (gs @ bs))" proof - have hs: "hs = fst (add_indices aux (snd data))" by (simp add: assms(9)[symmetric]) from assms(1) have sel: "sel_spec sel" and ab: "ab_spec ab" by (rule struct_specD)+ have eq: "fst ` (set gs \ set (ab gs bs hs data')) = fst ` (set gs \ set bs) \ fst ` set hs" by (auto simp add: ab_specD1[OF ab]) show ?thesis proof (simp add: eq, rule) show "pmdl (fst ` (set gs \ set bs) \ fst ` set hs) \ pmdl (fst ` (set gs \ set bs))" proof (rule pmdl.span_subset_spanI, simp only: Un_subset_iff, rule) show "fst ` (set gs \ set bs) \ pmdl (fst ` (set gs \ set bs))" by (fact pmdl.span_superset) next from sel assms(4) have "sps \ []" and "set sps \ set ps" unfolding assms(7) by (rule sel_specD1, rule sel_specD2) with assms(2, 3) have "fst ` set hs \ pmdl (args_to_set (gs, bs, ps))" unfolding hs assms(8) fst_set_add_indices using assms(6) by (rule compl_pmdlD) thus "fst ` set hs \ pmdl (fst ` (set gs \ set bs))" by (simp only: args_to_set_subset_Times[OF assms(5)] image_Un) qed next show "pmdl (fst ` (set gs \ set bs)) \ pmdl (fst ` (set gs \ set bs) \ fst ` set hs)" by (rule pmdl.span_mono, blast) qed qed lemma discarded_subset: assumes "ab_spec ab" and "D' = D \ (set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps) -\<^sub>p set (ap gs bs (ps -- sps) hs data'))" and "set ps \ set bs \ (set gs \ set bs)" and "D \ (set gs \ set bs) \ (set gs \ set bs)" shows "D' \ (set gs \ set (ab gs bs hs data')) \ (set gs \ set (ab gs bs hs data'))" proof - from assms(1) have eq: "set (ab gs bs hs data') = set bs \ set hs" by (rule ab_specD1) from assms(4) have "D \ (set gs \ (set bs \ set hs)) \ (set gs \ (set bs \ set hs))" by fastforce moreover have "set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps) -\<^sub>p set (ap gs bs (ps -- sps) hs data') \ (set gs \ (set bs \ set hs)) \ (set gs \ (set bs \ set hs))" (is "?l \ ?r") proof (rule subset_trans) show "?l \ set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps)" by (simp add: minus_pairs_def) next have "set hs \ (set gs \ set bs \ set hs) \ ?r" by fastforce moreover have "set (ps -- sps) \ ?r" proof (rule subset_trans) show "set (ps -- sps) \ set ps" by (auto simp: set_diff_list) next from assms(3) show "set ps \ ?r" by fastforce qed ultimately show "set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps) \ ?r" by (rule Un_least) qed ultimately show ?thesis unfolding eq assms(2) by (rule Un_least) qed lemma compl_struct_disjoint: assumes "compl_struct compl" and "sps \ []" and "set sps \ set ps" shows "fst ` set (fst (compl gs bs (ps -- sps) sps data)) \ fst ` (set gs \ set bs) = {}" proof (rule, rule) fix x assume "x \ fst ` set (fst (compl gs bs (ps -- sps) sps data)) \ fst ` (set gs \ set bs)" hence x_in: "x \ fst ` set (fst (compl gs bs (ps -- sps) sps data))" and "x \ fst ` (set gs \ set bs)" by simp_all from x_in obtain h where h_in: "h \ set (fst (compl gs bs (ps -- sps) sps data))" and x1: "x = fst h" .. from compl_structD3[OF assms, of gs bs data] x_in have "x \ 0" by auto from \x \ fst ` (set gs \ set bs)\ obtain b where b_in: "b \ set gs \ set bs" and x2: "x = fst b" .. from \x \ 0\ have "fst b \ 0" by (simp add: x2) with assms h_in b_in have "\ lt (fst b) adds\<^sub>t lt (fst h)" by (rule compl_structD4) hence "\ lt x adds\<^sub>t lt x" by (simp add: x1[symmetric] x2) from this adds_term_refl show "x \ {}" .. qed simp context fixes sel::"('t, 'b::field, 'c::default, 'd) selT" and ap::"('t, 'b, 'c, 'd) apT" and ab::"('t, 'b, 'c, 'd) abT" and compl::"('t, 'b, 'c, 'd) complT" and gs::"('t, 'b, 'c) pdata list" begin function (domintros) gb_schema_dummy :: "nat \ nat \ 'd \ ('t, 'b, 'c) pdata_pair set \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list \ (('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair set)" where "gb_schema_dummy data D bs ps = (if ps = [] then (gs @ bs, D) else (let sps = sel gs bs ps (snd data); ps0 = ps -- sps; aux = compl gs bs ps0 sps (snd data); remcomps = fst (data) - count_const_lt_components (fst aux) in (if remcomps = 0 then (full_gb (gs @ bs), D) else let (hs, data') = add_indices aux (snd data) in gb_schema_dummy (remcomps, data') (D \ ((set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps)) -\<^sub>p set (ap gs bs ps0 hs data'))) (ab gs bs hs data') (ap gs bs ps0 hs data') ) ) )" by pat_completeness auto lemma gb_schema_dummy_domI1: "gb_schema_dummy_dom (data, D, bs, [])" by (rule gb_schema_dummy.domintros, simp) lemma gb_schema_dummy_domI2: assumes "struct_spec sel ap ab compl" shows "gb_schema_dummy_dom (data, D, args)" proof - from assms have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" by (rule struct_specD)+ from ex_dgrad obtain d::"'a \ nat" where dg: "dickson_grading d" .. let ?R = "(gb_schema_aux_term d gs)" from dg have "wf ?R" by (rule gb_schema_aux_term_wf) thus ?thesis proof (induct args arbitrary: data D rule: wf_induct_rule) fix x data D assume IH: "\y data' D'. (y, x) \ ?R \ gb_schema_dummy_dom (data', D', y)" obtain bs ps where x: "x = (bs, ps)" by (meson case_prodE case_prodI2) show "gb_schema_dummy_dom (data, D, x)" unfolding x proof (rule gb_schema_dummy.domintros) fix rc0 n0 data0 hs n1 data1 assume "ps \ []" and hs_data': "(hs, n1, data1) = add_indices (compl gs bs (ps -- sel gs bs ps (n0, data0)) (sel gs bs ps (n0, data0)) (n0, data0)) (n0, data0)" and data: "data = (rc0, n0, data0)" define sps where "sps = sel gs bs ps (n0, data0)" define data' where "data' = (n1, data1)" define D' where "D' = D \ (set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps) -\<^sub>p set (ap gs bs (ps -- sps) hs data'))" define rc where "rc = rc0 - count_const_lt_components (fst (compl gs bs (ps -- sel gs bs ps (n0, data0)) (sel gs bs ps (n0, data0)) (n0, data0)))" from hs_data' have hs: "hs = fst (add_indices (compl gs bs (ps -- sps) sps (snd data)) (snd data))" unfolding sps_def data snd_conv by (metis fstI) show "gb_schema_dummy_dom ((rc, data'), D', ab gs bs hs data', ap gs bs (ps -- sps) hs data')" proof (rule IH, simp add: x gb_schema_aux_term_def gb_schema_aux_term1_def gb_schema_aux_term2_def, intro conjI) show "fst ` set (ab gs bs hs data') \p fst ` set bs \ ab gs bs hs data' = bs \ card (set (ap gs bs (ps -- sps) hs data')) < card (set ps)" proof (cases "hs = []") case True have "ab gs bs hs data' = bs \ card (set (ap gs bs (ps -- sps) hs data')) < card (set ps)" proof (simp only: True, rule) from ab show "ab gs bs [] data' = bs" by (rule ab_specD2) next from sel \ps \ []\ have "sps \ []" and "set sps \ set ps" unfolding sps_def by (rule sel_specD1, rule sel_specD2) moreover from sel_specD1[OF sel \ps \ []\] have "set sps \ {}" by (simp add: sps_def) ultimately have "set ps \ set sps \ {}" by (simp add: inf.absorb_iff2) hence "set (ps -- sps) \ set ps" unfolding set_diff_list by fastforce hence "card (set (ps -- sps)) < card (set ps)" by (simp add: psubset_card_mono) moreover have "card (set (ap gs bs (ps -- sps) [] data')) \ card (set (ps -- sps))" by (rule card_mono, fact finite_set, rule ap_spec_Nil_subset, fact ap) ultimately show "card (set (ap gs bs (ps -- sps) [] data')) < card (set ps)" by simp qed thus ?thesis .. next case False with assms \ps \ []\ sps_def hs have "fst ` set (ab gs bs hs data') \p fst ` set bs" unfolding data snd_conv by (rule struct_spec_red_supset) thus ?thesis .. qed next from dg assms \ps \ []\ sps_def hs show "dgrad_p_set_le d (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) (args_to_set (gs, bs, ps))" unfolding data snd_conv by (rule dgrad_p_set_le_args_to_set_struct) next from assms \ps \ []\ sps_def hs show "component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" unfolding data snd_conv by (rule components_subset_struct) qed qed qed qed lemmas gb_schema_dummy_simp = gb_schema_dummy.psimps[OF gb_schema_dummy_domI2] lemma gb_schema_dummy_Nil [simp]: "gb_schema_dummy data D bs [] = (gs @ bs, D)" by (simp add: gb_schema_dummy.psimps[OF gb_schema_dummy_domI1]) lemma gb_schema_dummy_not_Nil: assumes "struct_spec sel ap ab compl" and "ps \ []" shows "gb_schema_dummy data D bs ps = (let sps = sel gs bs ps (snd data); ps0 = ps -- sps; aux = compl gs bs ps0 sps (snd data); remcomps = fst (data) - count_const_lt_components (fst aux) in (if remcomps = 0 then (full_gb (gs @ bs), D) else let (hs, data') = add_indices aux (snd data) in gb_schema_dummy (remcomps, data') (D \ ((set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps)) -\<^sub>p set (ap gs bs ps0 hs data'))) (ab gs bs hs data') (ap gs bs ps0 hs data') ) )" by (simp add: gb_schema_dummy_simp[OF assms(1)] assms(2)) lemma gb_schema_dummy_induct [consumes 1, case_names base rec1 rec2]: assumes "struct_spec sel ap ab compl" assumes base: "\bs data D. P data D bs [] (gs @ bs, D)" and rec1: "\bs ps sps data D. ps \ [] \ sps = sel gs bs ps (snd data) \ fst (data) \ count_const_lt_components (fst (compl gs bs (ps -- sps) sps (snd data))) \ P data D bs ps (full_gb (gs @ bs), D)" and rec2: "\bs ps sps aux hs rc data data' D D'. ps \ [] \ sps = sel gs bs ps (snd data) \ aux = compl gs bs (ps -- sps) sps (snd data) \ (hs, data') = add_indices aux (snd data) \ rc = fst data - count_const_lt_components (fst aux) \ 0 < rc \ D' = (D \ ((set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps)) -\<^sub>p set (ap gs bs (ps -- sps) hs data'))) \ P (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data') (gb_schema_dummy (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data')) \ P data D bs ps (gb_schema_dummy (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data'))" shows "P data D bs ps (gb_schema_dummy data D bs ps)" proof - from assms(1) have "gb_schema_dummy_dom (data, D, bs, ps)" by (rule gb_schema_dummy_domI2) thus ?thesis proof (induct data D bs ps rule: gb_schema_dummy.pinduct) case (1 data D bs ps) show ?case proof (cases "ps = []") case True show ?thesis by (simp add: True, rule base) next case False show ?thesis proof (simp only: gb_schema_dummy_not_Nil[OF assms(1) False] Let_def split: if_split, intro conjI impI) define sps where "sps = sel gs bs ps (snd data)" assume "fst data - count_const_lt_components (fst (compl gs bs (ps -- sps) sps (snd data))) = 0" hence "fst data \ count_const_lt_components (fst (compl gs bs (ps -- sps) sps (snd data)))" by simp with False sps_def show "P data D bs ps (full_gb (gs @ bs), D)" by (rule rec1) next define sps where "sps = sel gs bs ps (snd data)" define aux where "aux = compl gs bs (ps -- sps) sps (snd data)" define hs where "hs = fst (add_indices aux (snd data))" define data' where "data' = snd (add_indices aux (snd data))" define rc where "rc = fst data - count_const_lt_components (fst aux)" define D' where "D' = (D \ ((set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps)) -\<^sub>p set (ap gs bs (ps -- sps) hs data')))" have eq: "add_indices aux (snd data) = (hs, data')" by (simp add: hs_def data'_def) assume "rc \ 0" hence "0 < rc" by simp show "P data D bs ps (case add_indices aux (snd data) of (hs, data') \ gb_schema_dummy (rc, data') (D \ (set hs \ (set gs \ set bs \ set hs) \ set (ps -- sps) -\<^sub>p set (ap gs bs (ps -- sps) hs data'))) (ab gs bs hs data') (ap gs bs (ps -- sps) hs data'))" unfolding eq prod.case D'_def[symmetric] using False sps_def aux_def eq[symmetric] rc_def \0 < rc\ D'_def proof (rule rec2) show "P (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data') (gb_schema_dummy (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data'))" unfolding D'_def using False sps_def refl aux_def rc_def \rc \ 0\ eq[symmetric] refl by (rule 1) qed qed qed qed qed lemma fst_gb_schema_dummy_dgrad_p_set_le: assumes "dickson_grading d" and "struct_spec sel ap ab compl" shows "dgrad_p_set_le d (fst ` set (fst (gb_schema_dummy data D bs ps))) (args_to_set (gs, bs, ps))" using assms(2) proof (induct rule: gb_schema_dummy_induct) case (base bs data D) show ?case by (simp add: args_to_set_def, rule dgrad_p_set_le_subset, fact subset_refl) next case (rec1 bs ps sps data D) show ?case proof (cases "fst ` set gs \ fst ` set bs \ {0}") case True hence "Keys (fst ` set (gs @ bs)) = {}" by (auto simp add: image_Un Keys_def) hence "component_of_term ` Keys (fst ` set (full_gb (gs @ bs))) = {}" by (simp add: components_full_gb) hence "Keys (fst ` set (full_gb (gs @ bs))) = {}" by simp thus ?thesis by (simp add: dgrad_p_set_le_def dgrad_set_le_def) next case False from pps_full_gb have "dgrad_set_le d (pp_of_term ` Keys (fst ` set (full_gb (gs @ bs)))) {0}" by (rule dgrad_set_le_subset) also have "dgrad_set_le d ... (pp_of_term ` Keys (args_to_set (gs, bs, ps)))" proof (rule dgrad_set_leI, simp) from False have "Keys (args_to_set (gs, bs, ps)) \ {}" by (simp add: args_to_set_alt Keys_Un, metis Keys_not_empty singletonI subsetI) then obtain v where "v \ Keys (args_to_set (gs, bs, ps))" by blast moreover have "d 0 \ d (pp_of_term v)" by (simp add: assms(1) dickson_grading_adds_imp_le) ultimately show "\t\Keys (args_to_set (gs, bs, ps)). d 0 \ d (pp_of_term t)" .. qed finally show ?thesis by (simp add: dgrad_p_set_le_def) qed next case (rec2 bs ps sps aux hs rc data data' D D') from rec2(4) have "hs = fst (add_indices (compl gs bs (ps -- sps) sps (snd data)) (snd data))" unfolding rec2(3) by (metis fstI) with assms rec2(1, 2) have "dgrad_p_set_le d (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) (args_to_set (gs, bs, ps))" by (rule dgrad_p_set_le_args_to_set_struct) with rec2(8) show ?case by (rule dgrad_p_set_le_trans) qed lemma fst_gb_schema_dummy_components: assumes "struct_spec sel ap ab compl" and "set ps \ (set bs) \ (set gs \ set bs)" shows "component_of_term ` Keys (fst ` set (fst (gb_schema_dummy data D bs ps))) = component_of_term ` Keys (args_to_set (gs, bs, ps))" using assms proof (induct rule: gb_schema_dummy_induct) case (base bs data D) show ?case by (simp add: args_to_set_def) next case (rec1 bs ps sps data D) have "component_of_term ` Keys (fst ` set (full_gb (gs @ bs))) = component_of_term ` Keys (fst ` set (gs @ bs))" by (fact components_full_gb) also have "... = component_of_term ` Keys (args_to_set (gs, bs, ps))" by (simp add: args_to_set_subset_Times[OF rec1.prems] image_Un) finally show ?case by simp next case (rec2 bs ps sps aux hs rc data data' D D') from assms(1) have ap: "ap_spec ap" and ab: "ab_spec ab" by (rule struct_specD)+ from this rec2.prems have sub: "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) from rec2(4) have hs: "hs = fst (add_indices (compl gs bs (ps -- sps) sps (snd data)) (snd data))" unfolding rec2(3) by (metis fstI) have "component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) = component_of_term ` Keys (args_to_set (gs, bs, ps))" (is "?l = ?r") proof from assms(1) rec2(1, 2) hs show "?l \ ?r" by (rule components_subset_struct) next show "?r \ ?l" by (simp add: args_to_set_subset_Times[OF rec2.prems] args_to_set_alt2[OF ap ab rec2.prems] image_Un, rule image_mono, rule Keys_mono, blast) qed with rec2.hyps(8)[OF sub] show ?case by (rule trans) qed lemma fst_gb_schema_dummy_pmdl: assumes "struct_spec sel ap ab compl" and "compl_pmdl compl" and "is_Groebner_basis (fst ` set gs)" and "set ps \ set bs \ (set gs \ set bs)" and "unique_idx (gs @ bs) (snd data)" and "rem_comps_spec (gs @ bs) data" shows "pmdl (fst ` set (fst (gb_schema_dummy data D bs ps))) = pmdl (fst ` set (gs @ bs))" proof - from assms(1) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ from assms(1, 4, 5, 6) show ?thesis proof (induct bs ps rule: gb_schema_dummy_induct) case (base bs data D) show ?case by simp next case (rec1 bs ps sps data D) define aux where "aux = compl gs bs (ps -- sps) sps (snd data)" define data' where "data' = snd (add_indices aux (snd data))" define hs where "hs = fst (add_indices aux (snd data))" have hs_data': "(hs, data') = add_indices aux (snd data)" by (simp add: hs_def data'_def) have eq: "set (gs @ ab gs bs hs data') = set (gs @ bs @ hs)" by (simp add: ab_specD1[OF ab]) from sel rec1(1) have "sps \ []" and "set sps \ set ps" unfolding rec1(2) by (rule sel_specD1, rule sel_specD2) from full_gb_is_full_pmdl have "pmdl (fst ` set (full_gb (gs @ bs))) = pmdl (fst ` set (gs @ ab gs bs hs data'))" proof (rule is_full_pmdl_eq) show "is_full_pmdl (fst ` set (gs @ ab gs bs hs data'))" proof (rule is_full_pmdlI_lt_finite) from finite_set show "finite (fst ` set (gs @ ab gs bs hs data'))" by (rule finite_imageI) next fix k assume "k \ component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data'))" hence "Some k \ Some ` component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data'))" by simp also have "... = const_lt_component ` (fst ` set (gs @ ab gs bs hs data') - {0}) - {None}" (is "?A = ?B") proof (rule card_seteq[symmetric]) show "finite ?A" by (intro finite_imageI finite_Keys, fact finite_set) next have "rem_comps_spec (gs @ ab gs bs hs data') (fst data - count_const_lt_components (fst aux), data')" using assms(1) rec1.prems(3) rec1.hyps(1) rec1.prems(1) rec1.hyps(2) aux_def hs_data' by (rule rem_comps_spec_struct) also have "... = (0, data')" by (simp add: aux_def rec1.hyps(3)) finally have "card (const_lt_component ` (fst ` set (gs @ ab gs bs hs data') - {0}) - {None}) = card (component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data')))" by (simp add: rem_comps_spec_def) also have "... = card (Some ` component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data')))" by (rule card_image[symmetric], simp) finally show "card ?A \ card ?B" by simp qed (fact const_lt_component_subset) finally have "Some k \ const_lt_component ` (fst ` set (gs @ ab gs bs hs data') - {0})" by simp then obtain b where "b \ fst ` set (gs @ ab gs bs hs data')" and "b \ 0" and *: "const_lt_component b = Some k" by fastforce show "\b\fst ` set (gs @ ab gs bs hs data'). b \ 0 \ component_of_term (lt b) = k \ lp b = 0" proof (intro bexI conjI) from * show "component_of_term (lt b) = k" by (rule const_lt_component_SomeD2) next from * show "lp b = 0" by (rule const_lt_component_SomeD1) qed fact+ qed next from compl \sps \ []\ \set sps \ set ps\ have "component_of_term ` Keys (fst ` set hs) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" unfolding hs_def aux_def fst_set_add_indices by (rule compl_structD2) hence sub: "component_of_term ` Keys (fst ` set hs) \ component_of_term ` Keys (fst ` set (gs @ bs))" by (simp add: args_to_set_subset_Times[OF rec1.prems(1)] image_Un) have "component_of_term ` Keys (fst ` set (full_gb (gs @ bs))) = component_of_term ` Keys (fst ` set (gs @ bs))" by (fact components_full_gb) also have "... = component_of_term ` Keys (fst ` set ((gs @ bs) @ hs))" by (simp only: set_append[of _ hs] image_Un Keys_Un Un_absorb2 sub) finally show "component_of_term ` Keys (fst ` set (full_gb (gs @ bs))) = component_of_term ` Keys (fst ` set (gs @ ab gs bs hs data'))" by (simp only: eq append_assoc) qed also have "... = pmdl (fst ` set (gs @ bs))" using assms(1, 2, 3) rec1.hyps(1) rec1.prems(1, 2) rec1.hyps(2) aux_def hs_data' by (rule pmdl_struct) finally show ?case by simp next case (rec2 bs ps sps aux hs rc data data' D D') from rec2(4) have hs: "hs = fst (add_indices aux (snd data))" by (metis fstI) have "pmdl (fst ` set (fst (gb_schema_dummy (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data')))) = pmdl (fst ` set (gs @ ab gs bs hs data'))" proof (rule rec2.hyps(8)) from ap ab rec2.prems(1) show "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) next from ab rec2.prems(2) rec2(4) show "unique_idx (gs @ ab gs bs hs data') (snd (rc, data'))" unfolding snd_conv by (rule unique_idx_ab) next show "rem_comps_spec (gs @ ab gs bs hs data') (rc, data')" unfolding rec2.hyps(5) using assms(1) rec2.prems(3) rec2.hyps(1) rec2.prems(1) rec2.hyps(2, 3, 4) by (rule rem_comps_spec_struct) qed also have "... = pmdl (fst ` set (gs @ bs))" using assms(1, 2, 3) rec2.hyps(1) rec2.prems(1, 2) rec2.hyps(2, 3, 4) by (rule pmdl_struct) finally show ?case . qed qed lemma snd_gb_schema_dummy_subset: assumes "struct_spec sel ap ab compl" and "set ps \ set bs \ (set gs \ set bs)" and "D \ (set gs \ set bs) \ (set gs \ set bs)" and "res = gb_schema_dummy data D bs ps" shows "snd res \ set (fst res) \ set (fst res) \ (\xs. fst (res) = full_gb xs)" using assms proof (induct data D bs ps rule: gb_schema_dummy_induct) case (base bs data D) from base(2) show ?case by (simp add: base(3)) next case (rec1 bs ps sps data D) have "\xs. fst res = full_gb xs" by (auto simp: rec1(6)) thus ?case .. next case (rec2 bs ps sps aux hs rc data data' D D') from assms(1) have ab: "ab_spec ab" and ap: "ap_spec ap" by (rule struct_specD)+ from _ _ rec2.prems(3) show ?case proof (rule rec2.hyps(8)) from ap ab rec2.prems(1) show "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) next from ab rec2.hyps(7) rec2.prems(1) rec2.prems(2) show "D' \ (set gs \ set (ab gs bs hs data')) \ (set gs \ set (ab gs bs hs data'))" by (rule discarded_subset) qed qed lemma gb_schema_dummy_connectible1: assumes "struct_spec sel ap ab compl" and "compl_conn compl" and "dickson_grading d" and "fst ` set gs \ dgrad_p_set d m" and "is_Groebner_basis (fst ` set gs)" and "fst ` set bs \ dgrad_p_set d m" and "set ps \ set bs \ (set gs \ set bs)" and "unique_idx (gs @ bs) (snd data)" and "\p q. processed (p, q) (gs @ bs) ps \ (p, q) \\<^sub>p D \ fst p \ 0 \ fst q \ 0 \ crit_pair_cbelow_on d m (fst ` (set gs \ set bs)) (fst p) (fst q)" and "\(\xs. fst (gb_schema_dummy data D bs ps) = full_gb xs)" assumes "f \ set (fst (gb_schema_dummy data D bs ps))" and "g \ set (fst (gb_schema_dummy data D bs ps))" and "(f, g) \\<^sub>p snd (gb_schema_dummy data D bs ps)" and "fst f \ 0" and "fst g \ 0" shows "crit_pair_cbelow_on d m (fst ` set (fst (gb_schema_dummy data D bs ps))) (fst f) (fst g)" using assms(1, 6, 7, 8, 9, 10, 11, 12, 13) proof (induct data D bs ps rule: gb_schema_dummy_induct) case (base bs data D) show ?case proof (cases "f \ set gs") case True show ?thesis proof (cases "g \ set gs") case True note assms(3, 4, 5) moreover from \f \ set gs\ have "fst f \ fst ` set gs" by simp moreover from \g \ set gs\ have "fst g \ fst ` set gs" by simp ultimately have "crit_pair_cbelow_on d m (fst ` set gs) (fst f) (fst g)" using assms(14, 15) by (rule GB_imp_crit_pair_cbelow_dgrad_p_set) moreover have "fst ` set gs \ fst ` set (fst (gs @ bs, D))" by auto ultimately show ?thesis by (rule crit_pair_cbelow_mono) next case False from this base(6, 7) have "processed (g, f) (gs @ bs) []" by (simp add: processed_Nil) moreover from base.prems(8) have "(g, f) \\<^sub>p D" by (simp add: in_pair_iff) ultimately have "crit_pair_cbelow_on d m (fst ` set (gs @ bs)) (fst g) (fst f)" using \fst g \ 0\ \fst f \ 0\ unfolding set_append by (rule base(4)) thus ?thesis unfolding fst_conv by (rule crit_pair_cbelow_sym) qed next case False from this base(6, 7) have "processed (f, g) (gs @ bs) []" by (simp add: processed_Nil) moreover from base.prems(8) have "(f, g) \\<^sub>p D" by simp ultimately show ?thesis unfolding fst_conv set_append using \fst f \ 0\ \fst g \ 0\ by (rule base(4)) qed next case (rec1 bs ps sps data D) from rec1.prems(5) show ?case by auto next case (rec2 bs ps sps aux hs rc data data' D D') from rec2.hyps(4) have hs: "hs = fst (add_indices aux (snd data))" by (metis fstI) from assms(1) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD1, rule struct_specD2, rule struct_specD3, rule struct_specD4) from sel rec2.hyps(1) have "sps \ []" and "set sps \ set ps" unfolding rec2.hyps(2) by (rule sel_specD1, rule sel_specD2) from ap ab rec2.prems(2) have ap_sub: "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) have ns_sub: "fst ` set hs \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) from compl assms(3) \sps \ []\ \set sps \ set ps\ show "dgrad_p_set_le d (fst ` set hs) (args_to_set (gs, bs, ps))" unfolding hs rec2.hyps(3) fst_set_add_indices by (rule compl_structD1) next from assms(4) rec2.prems(1) show "args_to_set (gs, bs, ps) \ dgrad_p_set d m" by (simp add: args_to_set_subset_Times[OF rec2.prems(2)]) qed with rec2.prems(1) have ab_sub: "fst ` set (ab gs bs hs data') \ dgrad_p_set d m" by (auto simp add: ab_specD1[OF ab]) have cpq: "(p, q) \\<^sub>p set sps \ fst p \ 0 \ fst q \ 0 \ crit_pair_cbelow_on d m (fst ` (set gs \ set (ab gs bs hs data'))) (fst p) (fst q)" for p q proof - assume "(p, q) \\<^sub>p set sps" and "fst p \ 0" and "fst q \ 0" from this(1) have "(p, q) \ set sps \ (q, p) \ set sps" by (simp only: in_pair_iff) hence "crit_pair_cbelow_on d m (fst ` (set gs \ set bs) \ fst ` set (fst (compl gs bs (ps -- sps) sps (snd data)))) (fst p) (fst q)" proof assume "(p, q) \ set sps" from assms(2, 3, 4, 5) rec2.prems(1, 2) \sps \ []\ \set sps \ set ps\ rec2.prems(3) this \fst p \ 0\ \fst q \ 0\ show ?thesis by (rule compl_connD) next assume "(q, p) \ set sps" from assms(2, 3, 4, 5) rec2.prems(1, 2) \sps \ []\ \set sps \ set ps\ rec2.prems(3) this \fst q \ 0\ \fst p \ 0\ have "crit_pair_cbelow_on d m (fst ` (set gs \ set bs) \ fst ` set (fst (compl gs bs (ps -- sps) sps (snd data)))) (fst q) (fst p)" by (rule compl_connD) thus ?thesis by (rule crit_pair_cbelow_sym) qed thus "crit_pair_cbelow_on d m (fst ` (set gs \ set (ab gs bs hs data'))) (fst p) (fst q)" by (simp add: ab_specD1[OF ab] hs rec2.hyps(3) fst_set_add_indices image_Un Un_assoc) qed from ab_sub ap_sub _ _ rec2.prems(5, 6, 7, 8) show ?case proof (rule rec2.hyps(8)) from ab rec2.prems(3) rec2(4) show "unique_idx (gs @ ab gs bs hs data') (snd (rc, data'))" unfolding snd_conv by (rule unique_idx_ab) next fix p q :: "('t, 'b, 'c) pdata" define ps' where "ps' = ap gs bs (ps -- sps) hs data'" assume "fst p \ 0" and "fst q \ 0" and "(p, q) \\<^sub>p D'" assume "processed (p, q) (gs @ ab gs bs hs data') ps'" hence p_in: "p \ set gs \ set bs \ set hs" and q_in: "q \ set gs \ set bs \ set hs" and "(p, q) \\<^sub>p set ps'" by (simp_all add: processed_alt ab_specD1[OF ab]) from this(3) \(p, q) \\<^sub>p D'\ have "(p, q) \\<^sub>p D" and "(p, q) \\<^sub>p set (ps -- sps)" and "(p, q) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" by (auto simp: in_pair_iff rec2.hyps(7) ps'_def) from this(3) p_in q_in have "p \ set gs \ set bs" and "q \ set gs \ set bs" by (meson SigmaI UnE in_pair_iff)+ show "crit_pair_cbelow_on d m (fst ` (set gs \ set (ab gs bs hs data'))) (fst p) (fst q)" proof (cases "component_of_term (lt (fst p)) = component_of_term (lt (fst q))") case True show ?thesis proof (cases "(p, q) \\<^sub>p set sps") case True from this \fst p \ 0\ \fst q \ 0\ show ?thesis by (rule cpq) next case False with \(p, q) \\<^sub>p set (ps -- sps)\ have "(p, q) \\<^sub>p set ps" by (auto simp: in_pair_iff set_diff_list) with \p \ set gs \ set bs\ \q \ set gs \ set bs\ have "processed (p, q) (gs @ bs) ps" by (simp add: processed_alt) from this \(p, q) \\<^sub>p D\ \fst p \ 0\ \fst q \ 0\ have "crit_pair_cbelow_on d m (fst ` (set gs \ set bs)) (fst p) (fst q)" by (rule rec2.prems(4)) moreover have "fst ` (set gs \ set bs) \ fst ` (set gs \ set (ab gs bs hs data'))" by (auto simp: ab_specD1[OF ab]) ultimately show ?thesis by (rule crit_pair_cbelow_mono) qed next case False thus ?thesis by (rule crit_pair_cbelow_distinct_component) qed qed qed lemma gb_schema_dummy_connectible2: assumes "struct_spec sel ap ab compl" and "compl_conn compl" and "dickson_grading d" and "fst ` set gs \ dgrad_p_set d m" and "is_Groebner_basis (fst ` set gs)" and "fst ` set bs \ dgrad_p_set d m" and "set ps \ set bs \ (set gs \ set bs)" and "D \ (set gs \ set bs) \ (set gs \ set bs)" and "set ps \\<^sub>p D = {}" and "unique_idx (gs @ bs) (snd data)" and "\B a b. set gs \ set bs \ B \ fst ` B \ dgrad_p_set d m \ (a, b) \\<^sub>p D \ fst a \ 0 \ fst b \ 0 \ (\x y. x \ set gs \ set bs \ y \ set gs \ set bs \ \ (x, y) \\<^sub>p D \ fst x \ 0 \ fst y \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)) \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" and "\x y. x \ set (fst (gb_schema_dummy data D bs ps)) \ y \ set (fst (gb_schema_dummy data D bs ps)) \ (x, y) \\<^sub>p snd (gb_schema_dummy data D bs ps) \ fst x \ 0 \ fst y \ 0 \ crit_pair_cbelow_on d m (fst ` set (fst (gb_schema_dummy data D bs ps))) (fst x) (fst y)" and "\(\xs. fst (gb_schema_dummy data D bs ps) = full_gb xs)" assumes "(f, g) \\<^sub>p snd (gb_schema_dummy data D bs ps)" and "fst f \ 0" and "fst g \ 0" shows "crit_pair_cbelow_on d m (fst ` set (fst (gb_schema_dummy data D bs ps))) (fst f) (fst g)" using assms(1, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16) proof (induct data D bs ps rule: gb_schema_dummy_induct) case (base bs data D) have "set gs \ set bs \ set (fst (gs @ bs, D))" by simp moreover from assms(4) base.prems(1) have "fst ` set (fst (gs @ bs, D)) \ dgrad_p_set d m" by auto moreover from base.prems(9) have "(f, g) \\<^sub>p D" by simp moreover note assms(15, 16) ultimately show ?case proof (rule base.prems(6)) fix x y assume "x \ set gs \ set bs" and "y \ set gs \ set bs" and "(x, y) \\<^sub>p D" hence "x \ set (fst (gs @ bs, D))" and "y \ set (fst (gs @ bs, D))" and "(x, y) \\<^sub>p snd (gs @ bs, D)" by simp_all moreover assume "fst x \ 0" and "fst y \ 0" ultimately show "crit_pair_cbelow_on d m (fst ` set (fst (gs @ bs, D))) (fst x) (fst y)" by (rule base.prems(7)) qed next case (rec1 bs ps sps data D) from rec1.prems(8) show ?case by auto next case (rec2 bs ps sps aux hs rc data data' D D') from rec2.hyps(4) have hs: "hs = fst (add_indices aux (snd data))" by (metis fstI) from assms(1) have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" and compl: "compl_struct compl" by (rule struct_specD)+ let ?X = "set (ps -- sps) \ set hs \ (set gs \ set bs \ set hs)" from sel rec2.hyps(1) have "sps \ []" and "set sps \ set ps" unfolding rec2.hyps(2) by (rule sel_specD1, rule sel_specD2) have "fst ` set hs \ fst ` (set gs \ set bs) = {}" unfolding hs fst_set_add_indices rec2.hyps(3) using compl \sps \ []\ \set sps \ set ps\ by (rule compl_struct_disjoint) hence disj1: "(set gs \ set bs) \ set hs = {}" by fastforce have disj2: "set (ap gs bs (ps -- sps) hs data') \\<^sub>p D' = {}" proof (rule, rule) fix x y assume "(x, y) \ set (ap gs bs (ps -- sps) hs data') \\<^sub>p D'" hence "(x, y) \\<^sub>p set (ap gs bs (ps -- sps) hs data') \\<^sub>p D'" by (simp add: in_pair_alt) hence 1: "(x, y) \\<^sub>p set (ap gs bs (ps -- sps) hs data')" and "(x, y) \\<^sub>p D'" by simp_all hence "(x, y) \\<^sub>p D" by (simp add: rec2.hyps(7)) from this rec2.prems(3) have "x \ set gs \ set bs" and "y \ set gs \ set bs" by (auto simp: in_pair_iff) from 1 ap_specD1[OF ap] have "(x, y) \\<^sub>p ?X" by (rule in_pair_trans) thus "(x, y) \ {}" unfolding in_pair_Un proof assume "(x, y) \\<^sub>p set (ps -- sps)" also have "... \ set ps" by (auto simp: set_diff_list) finally have "(x, y) \\<^sub>p set ps \\<^sub>p D" using \(x, y) \\<^sub>p D\ by simp also have "... = {}" by (fact rec2.prems(4)) finally show ?thesis by (simp add: in_pair_iff) next assume "(x, y) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" hence "x \ set hs \ y \ set hs" by (auto simp: in_pair_iff) thus ?thesis proof assume "x \ set hs" with \x \ set gs \ set bs\ have "x \ (set gs \ set bs) \ set hs" .. thus ?thesis by (simp add: disj1) next assume "y \ set hs" with \y \ set gs \ set bs\ have "y \ (set gs \ set bs) \ set hs" .. thus ?thesis by (simp add: disj1) qed qed qed simp have hs_sub: "fst ` set hs \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) from compl assms(3) \sps \ []\ \set sps \ set ps\ show "dgrad_p_set_le d (fst ` set hs) (args_to_set (gs, bs, ps))" unfolding hs rec2.hyps(3) fst_set_add_indices by (rule compl_structD1) next from assms(4) rec2.prems(1) show "args_to_set (gs, bs, ps) \ dgrad_p_set d m" by (simp add: args_to_set_subset_Times[OF rec2.prems(2)]) qed with rec2.prems(1) have ab_sub: "fst ` set (ab gs bs hs data') \ dgrad_p_set d m" by (auto simp add: ab_specD1[OF ab]) moreover from ap ab rec2.prems(2) have ap_sub: "set (ap gs bs (ps -- sps) hs data') \ set (ab gs bs hs data') \ (set gs \ set (ab gs bs hs data'))" by (rule subset_Times_ap) moreover from ab rec2.hyps(7) rec2.prems(2) rec2.prems(3) have "D' \ (set gs \ set (ab gs bs hs data')) \ (set gs \ set (ab gs bs hs data'))" by (rule discarded_subset) moreover note disj2 moreover from ab rec2.prems(5) rec2.hyps(4) have uid: "unique_idx (gs @ ab gs bs hs data') (snd (rc, data'))" unfolding snd_conv by (rule unique_idx_ab) ultimately show ?case using _ _ rec2.prems(8, 9, 10, 11) proof (rule rec2.hyps(8), simp only: ab_specD1[OF ab] Un_assoc[symmetric]) define ps' where "ps' = ap gs bs (ps -- sps) hs data'" fix B a b assume B_sup: "set gs \ set bs \ set hs \ B" hence "set gs \ set bs \ B" and "set hs \ B" by simp_all assume "(a, b) \\<^sub>p D'" hence ab_cases: "(a, b) \\<^sub>p D \ (a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs) -\<^sub>p set ps' \ (a, b) \\<^sub>p set (ps -- sps) -\<^sub>p set ps'" by (auto simp: rec2.hyps(7) ps'_def) assume B_sub: "fst ` B \ dgrad_p_set d m" and "fst a \ 0" and "fst b \ 0" assume *: "\x y. x \ set gs \ set bs \ set hs \ y \ set gs \ set bs \ set hs \ (x, y) \\<^sub>p D' \ fst x \ 0 \ fst y \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" from rec2.prems(2) have ps_sps_sub: "set (ps -- sps) \ set bs \ (set gs \ set bs)" by (auto simp: set_diff_list) from uid have uid': "unique_idx (gs @ bs @ hs) data'" by (simp add: unique_idx_def ab_specD1[OF ab]) have a: "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" if "fst x \ 0" and "fst y \ 0" and xy_in: "(x, y) \\<^sub>p set (ps -- sps) -\<^sub>p set ps'" for x y proof (cases "x = y") case True from xy_in rec2.prems(2) have "y \ set gs \ set bs" unfolding in_pair_minus_pairs unfolding True in_pair_iff set_diff_list by auto hence "fst y \ fst ` set gs \ fst ` set bs" by fastforce from this assms(4) rec2.prems(1) have "fst y \ dgrad_p_set d m" by blast with assms(3) show ?thesis unfolding True by (rule crit_pair_cbelow_same) next case False from ap assms(3) B_sup B_sub ps_sps_sub disj1 uid' assms(5) False \fst x \ 0\ \fst y \ 0\ xy_in show ?thesis unfolding ps'_def proof (rule ap_specD3) fix a1 b1 :: "('t, 'b, 'c) pdata" assume "fst a1 \ 0" and "fst b1 \ 0" assume "a1 \ set hs" and b1_in: "b1 \ set gs \ set bs \ set hs" hence a1_in: "a1 \ set gs \ set bs \ set hs" by fastforce assume "(a1, b1) \\<^sub>p set (ap gs bs (ps -- sps) hs data')" hence "(a1, b1) \\<^sub>p set ps'" by (simp only: ps'_def) with disj2 have "(a1, b1) \\<^sub>p D'" unfolding ps'_def by (metis empty_iff in_pair_Int_pairs in_pair_alt) with a1_in b1_in show "crit_pair_cbelow_on d m (fst ` B) (fst a1) (fst b1)" using \fst a1 \ 0\ \fst b1 \ 0\ by (rule *) qed qed have b: "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" if "(x, y) \\<^sub>p D" and "fst x \ 0" and "fst y \ 0" for x y using \set gs \ set bs \ B\ B_sub that proof (rule rec2.prems(6)) fix a1 b1 :: "('t, 'b, 'c) pdata" assume "a1 \ set gs \ set bs" and "b1 \ set gs \ set bs" hence a1_in: "a1 \ set gs \ set bs \ set hs" and b1_in: "b1 \ set gs \ set bs \ set hs" by fastforce+ assume "(a1, b1) \\<^sub>p D" and "fst a1 \ 0" and "fst b1 \ 0" show "crit_pair_cbelow_on d m (fst ` B) (fst a1) (fst b1)" proof (cases "(a1, b1) \\<^sub>p ?X -\<^sub>p set ps'") case True moreover from \a1 \ set gs \ set bs\ \b1 \ set gs \ set bs\ disj1 have "(a1, b1) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" - by (auto simp: in_pair_def swap_def) + by (auto simp: in_pair_def) ultimately have "(a1, b1) \\<^sub>p set (ps -- sps) -\<^sub>p set ps'" by auto with \fst a1 \ 0\ \fst b1 \ 0\ show ?thesis by (rule a) next case False with \(a1, b1) \\<^sub>p D\ have "(a1, b1) \\<^sub>p D'" by (auto simp: rec2.hyps(7) ps'_def) with a1_in b1_in show ?thesis using \fst a1 \ 0\ \fst b1 \ 0\ by (rule *) qed qed have c: "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" if x_in: "x \ set gs \ set bs \ set hs" and y_in: "y \ set gs \ set bs \ set hs" and xy: "(x, y) \\<^sub>p (?X -\<^sub>p set ps')" and "fst x \ 0" and "fst y \ 0" for x y proof (cases "(x, y) \\<^sub>p D") case True thus ?thesis using \fst x \ 0\ \fst y \ 0\ by (rule b) next case False with xy have "(x, y) \\<^sub>p D'" unfolding rec2.hyps(7) ps'_def by auto with x_in y_in show ?thesis using \fst x \ 0\ \fst y \ 0\ by (rule *) qed from ab_cases show "crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" proof (elim disjE) assume "(a, b) \\<^sub>p D" thus ?thesis using \fst a \ 0\ \fst b \ 0\ by (rule b) next assume ab_in: "(a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs) -\<^sub>p set ps'" hence ab_in': "(a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" and "(a, b) \\<^sub>p set ps'" by simp_all show ?thesis proof (cases "a = b") case True from ab_in' rec2.prems(2) have "b \ set hs" unfolding True in_pair_iff set_diff_list by auto hence "fst b \ fst ` set hs" by fastforce from this hs_sub have "fst b \ dgrad_p_set d m" .. with assms(3) show ?thesis unfolding True by (rule crit_pair_cbelow_same) next case False from ap assms(3) B_sup B_sub ab_in' ps_sps_sub uid' assms(5) False \fst a \ 0\ \fst b \ 0\ show ?thesis proof (rule ap_specD2) fix x y :: "('t, 'b, 'c) pdata" assume "(x, y) \\<^sub>p set (ap gs bs (ps -- sps) hs data')" also from ap_sub have "... \ (set bs \ set hs) \ (set gs \ set bs \ set hs)" by (simp only: ab_specD1[OF ab] Un_assoc) also have "... \ (set gs \ set bs \ set hs) \ (set gs \ set bs \ set hs)" by fastforce finally have "(x, y) \ (set gs \ set bs \ set hs) \ (set gs \ set bs \ set hs)" unfolding in_pair_same . hence "x \ set gs \ set bs \ set hs" and "y \ set gs \ set bs \ set hs" by simp_all moreover from \(x, y) \\<^sub>p set (ap gs bs (ps -- sps) hs data')\ have "(x, y) \\<^sub>p ?X -\<^sub>p set ps'" by (simp add: ps'_def) moreover assume "fst x \ 0" and "fst y \ 0" ultimately show "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" by (rule c) next fix x y :: "('t, 'b, 'c) pdata" assume "fst x \ 0" and "fst y \ 0" assume 1: "x \ set gs \ set bs" and 2: "y \ set gs \ set bs" hence x_in: "x \ set gs \ set bs \ set hs" and y_in: "y \ set gs \ set bs \ set hs" by simp_all show "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" proof (cases "(x, y) \\<^sub>p set (ps -- sps) -\<^sub>p set ps'") case True with \fst x \ 0\ \fst y \ 0\ show ?thesis by (rule a) next case False have "(x, y) \\<^sub>p set (ps -- sps) \ set hs \ (set gs \ set bs \ set hs) -\<^sub>p set ps'" proof assume "(x, y) \\<^sub>p set (ps -- sps) \ set hs \ (set gs \ set bs \ set hs) -\<^sub>p set ps'" hence "(x, y) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" using False by simp hence "x \ set hs \ y \ set hs" by (auto simp: in_pair_iff) with 1 2 disj1 show False by blast qed with x_in y_in show ?thesis using \fst x \ 0\ \fst y \ 0\ by (rule c) qed qed qed next assume "(a, b) \\<^sub>p set (ps -- sps) -\<^sub>p set ps'" with \fst a \ 0\ \fst b \ 0\ show ?thesis by (rule a) qed next fix x y :: "('t, 'b, 'c) pdata" let ?res = "gb_schema_dummy (rc, data') D' (ab gs bs hs data') (ap gs bs (ps -- sps) hs data')" assume "x \ set (fst ?res)" and "y \ set (fst ?res)" and "(x, y) \\<^sub>p snd ?res" and "fst x \ 0" and "fst y \ 0" thus "crit_pair_cbelow_on d m (fst ` set (fst ?res)) (fst x) (fst y)" by (rule rec2.prems(7)) qed qed corollary gb_schema_dummy_connectible: assumes "struct_spec sel ap ab compl" and "compl_conn compl" and "dickson_grading d" and "fst ` set gs \ dgrad_p_set d m" and "is_Groebner_basis (fst ` set gs)" and "fst ` set bs \ dgrad_p_set d m" and "set ps \ set bs \ (set gs \ set bs)" and "D \ (set gs \ set bs) \ (set gs \ set bs)" and "set ps \\<^sub>p D = {}" and "unique_idx (gs @ bs) (snd data)" and "\p q. processed (p, q) (gs @ bs) ps \ (p, q) \\<^sub>p D \ fst p \ 0 \ fst q \ 0 \ crit_pair_cbelow_on d m (fst ` (set gs \ set bs)) (fst p) (fst q)" and "\B a b. set gs \ set bs \ B \ fst ` B \ dgrad_p_set d m \ (a, b) \\<^sub>p D \ fst a \ 0 \ fst b \ 0 \ (\x y. x \ set gs \ set bs \ y \ set gs \ set bs \ \ (x, y) \\<^sub>p D \ fst x \ 0 \ fst y \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)) \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" assumes "f \ set (fst (gb_schema_dummy data D bs ps))" and "g \ set (fst (gb_schema_dummy data D bs ps))" and "fst f \ 0" and "fst g \ 0" shows "crit_pair_cbelow_on d m (fst ` set (fst (gb_schema_dummy data D bs ps))) (fst f) (fst g)" proof (cases "\xs. fst (gb_schema_dummy data D bs ps) = full_gb xs") case True then obtain xs where xs: "fst (gb_schema_dummy data D bs ps) = full_gb xs" .. note assms(3) moreover have "fst ` set (full_gb xs) \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) have "dgrad_p_set_le d (fst ` set (full_gb xs)) (args_to_set (gs, bs, ps))" unfolding xs[symmetric] using assms(3, 1) by (rule fst_gb_schema_dummy_dgrad_p_set_le) also from assms(7) have "... = fst ` set gs \ fst ` set bs" by (rule args_to_set_subset_Times) finally show "dgrad_p_set_le d (fst ` set (full_gb xs)) (fst ` set gs \ fst ` set bs)" . next from assms(4, 6) show "fst ` set gs \ fst ` set bs \ dgrad_p_set d m" by blast qed moreover note full_gb_isGB moreover from assms(13) have "fst f \ fst ` set (full_gb xs)" by (simp add: xs) moreover from assms(14) have "fst g \ fst ` set (full_gb xs)" by (simp add: xs) ultimately show ?thesis using assms(15, 16) unfolding xs by (rule GB_imp_crit_pair_cbelow_dgrad_p_set) next case not_full: False show ?thesis proof (cases "(f, g) \\<^sub>p snd (gb_schema_dummy data D bs ps)") case True from assms(1-10,12) _ not_full True assms(15,16) show ?thesis proof (rule gb_schema_dummy_connectible2) fix x y assume "x \ set (fst (gb_schema_dummy data D bs ps))" and "y \ set (fst (gb_schema_dummy data D bs ps))" and "(x, y) \\<^sub>p snd (gb_schema_dummy data D bs ps)" and "fst x \ 0" and "fst y \ 0" with assms(1-7,10,11) not_full show "crit_pair_cbelow_on d m (fst ` set (fst (gb_schema_dummy data D bs ps))) (fst x) (fst y)" by (rule gb_schema_dummy_connectible1) qed next case False from assms(1-7,10,11) not_full assms(13,14) False assms(15,16) show ?thesis by (rule gb_schema_dummy_connectible1) qed qed lemma fst_gb_schema_dummy_dgrad_p_set_le_init: assumes "dickson_grading d" and "struct_spec sel ap ab compl" shows "dgrad_p_set_le d (fst ` set (fst (gb_schema_dummy data D (ab gs [] bs (snd data)) (ap gs [] [] bs (snd data))))) (fst ` (set gs \ set bs))" proof - let ?bs = "ab gs [] bs (snd data)" from assms(2) have ap: "ap_spec ap" and ab: "ab_spec ab" by (rule struct_specD)+ from ap_specD1[OF ap, of gs "[]" "[]" bs] have *: "set (ap gs [] [] bs (snd data)) \ set ?bs \ (set gs \ set ?bs)" by (simp add: ab_specD1[OF ab]) from assms have "dgrad_p_set_le d (fst ` set (fst (gb_schema_dummy data D ?bs (ap gs [] [] bs (snd data))))) (args_to_set (gs, ?bs, (ap gs [] [] bs (snd data))))" by (rule fst_gb_schema_dummy_dgrad_p_set_le) also have "... = fst ` (set gs \ set bs)" by (simp add: args_to_set_subset_Times[OF *] image_Un ab_specD1[OF ab]) finally show ?thesis . qed corollary fst_gb_schema_dummy_dgrad_p_set_init: assumes "dickson_grading d" and "struct_spec sel ap ab compl" and "fst ` (set gs \ set bs) \ dgrad_p_set d m" shows "fst ` set (fst (gb_schema_dummy (rc, data) D (ab gs [] bs data) (ap gs [] [] bs data))) \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) let ?data = "(rc, data)" from assms(1, 2) have "dgrad_p_set_le d (fst ` set (fst (gb_schema_dummy ?data D (ab gs [] bs (snd ?data)) (ap gs [] [] bs (snd ?data))))) (fst ` (set gs \ set bs))" by (rule fst_gb_schema_dummy_dgrad_p_set_le_init) thus "dgrad_p_set_le d (fst ` set (fst (gb_schema_dummy ?data D (ab gs [] bs data) (ap gs [] [] bs data)))) (fst ` (set gs \ set bs))" by (simp only: snd_conv) qed fact lemma fst_gb_schema_dummy_components_init: fixes bs data defines "bs0 \ ab gs [] bs data" defines "ps0 \ ap gs [] [] bs data" assumes "struct_spec sel ap ab compl" shows "component_of_term ` Keys (fst ` set (fst (gb_schema_dummy (rc, data) D bs0 ps0))) = component_of_term ` Keys (fst ` set (gs @ bs))" (is "?l = ?r") proof - from assms(3) have ap: "ap_spec ap" and ab: "ab_spec ab" by (rule struct_specD)+ from ap_specD1[OF ap, of gs "[]" "[]" bs] have *: "set ps0 \ set bs0 \ (set gs \ set bs0)" by (simp add: ps0_def bs0_def ab_specD1[OF ab]) with assms(3) have "?l = component_of_term ` Keys (args_to_set (gs, bs0, ps0))" by (rule fst_gb_schema_dummy_components) also have "... = ?r" by (simp only: args_to_set_subset_Times[OF *], simp add: ab_specD1[OF ab] bs0_def image_Un) finally show ?thesis . qed lemma fst_gb_schema_dummy_pmdl_init: fixes bs data defines "bs0 \ ab gs [] bs data" defines "ps0 \ ap gs [] [] bs data" assumes "struct_spec sel ap ab compl" and "compl_pmdl compl" and "is_Groebner_basis (fst ` set gs)" and "unique_idx (gs @ bs0) data" and "rem_comps_spec (gs @ bs0) (rc, data)" shows "pmdl (fst ` set (fst (gb_schema_dummy (rc, data) D bs0 ps0))) = pmdl (fst ` (set (gs @ bs)))" (is "?l = ?r") proof - from assms(3) have ab: "ab_spec ab" by (rule struct_specD3) let ?data = "(rc, data)" from assms(6) have "unique_idx (gs @ bs0) (snd ?data)" by (simp only: snd_conv) from assms(3, 4, 5) _ this assms(7) have "?l = pmdl (fst ` (set (gs @ bs0)))" proof (rule fst_gb_schema_dummy_pmdl) from assms(3) have "ap_spec ap" by (rule struct_specD2) from ap_specD1[OF this, of gs "[]" "[]" bs] show "set ps0 \ set bs0 \ (set gs \ set bs0)" by (simp add: ps0_def bs0_def ab_specD1[OF ab]) qed also have "... = ?r" by (simp add: bs0_def ab_specD1[OF ab]) finally show ?thesis . qed lemma fst_gb_schema_dummy_isGB_init: fixes bs data defines "bs0 \ ab gs [] bs data" defines "ps0 \ ap gs [] [] bs data" defines "D0 \ set bs \ (set gs \ set bs) -\<^sub>p set ps0" assumes "struct_spec sel ap ab compl" and "compl_conn compl" and "is_Groebner_basis (fst ` set gs)" and "unique_idx (gs @ bs0) data" and "rem_comps_spec (gs @ bs0) (rc, data)" shows "is_Groebner_basis (fst ` set (fst (gb_schema_dummy (rc, data) D0 bs0 ps0)))" proof - let ?data = "(rc, data)" let ?res = "gb_schema_dummy ?data D0 bs0 ps0" from assms(4) have ap: "ap_spec ap" and ab: "ab_spec ab" by (rule struct_specD2, rule struct_specD3) have set_bs0: "set bs0 = set bs" by (simp add: bs0_def ab_specD1[OF ab]) from ap_specD1[OF ap, of gs "[]" "[]" bs] have ps0_sub: "set ps0 \ set bs0 \ (set gs \ set bs0)" by (simp add: ps0_def set_bs0) from ex_dgrad obtain d::"'a \ nat" where dg: "dickson_grading d" .. have "finite (fst ` (set gs \ set bs))" by (rule, rule finite_UnI, fact finite_set, fact finite_set) then obtain m where gs_bs_sub: "fst ` (set gs \ set bs) \ dgrad_p_set d m" by (rule dgrad_p_set_exhaust) with dg assms(4) have "fst ` set (fst ?res) \ dgrad_p_set d m" unfolding bs0_def ps0_def by (rule fst_gb_schema_dummy_dgrad_p_set_init) with dg show ?thesis proof (rule crit_pair_cbelow_imp_GB_dgrad_p_set) fix p0 q0 assume p0_in: "p0 \ fst ` set (fst ?res)" and q0_in: "q0 \ fst ` set (fst ?res)" assume "p0 \ 0" and "q0 \ 0" from \fst ` (set gs \ set bs) \ dgrad_p_set d m\ have "fst ` set gs \ dgrad_p_set d m" and "fst ` set bs \ dgrad_p_set d m" by (simp_all add: image_Un) from p0_in obtain p where p_in: "p \ set (fst ?res)" and p0: "p0 = fst p" .. from q0_in obtain q where q_in: "q \ set (fst ?res)" and q0: "q0 = fst q" .. from assms(7) have "unique_idx (gs @ bs0) (snd ?data)" by (simp only: snd_conv) from assms(4, 5) dg \fst ` set gs \ dgrad_p_set d m\ assms(6) _ ps0_sub _ _ this _ _ p_in q_in \p0 \ 0\ \q0 \ 0\ show "crit_pair_cbelow_on d m (fst ` set (fst ?res)) p0 q0" unfolding p0 q0 proof (rule gb_schema_dummy_connectible) from \fst ` set bs \ dgrad_p_set d m\ show "fst ` set bs0 \ dgrad_p_set d m" by (simp only: set_bs0) next have "D0 \ set bs \ (set gs \ set bs)" by (auto simp: assms(3) minus_pairs_def) also have "... \ (set gs \ set bs) \ (set gs \ set bs)" by fastforce finally show "D0 \ (set gs \ set bs0) \ (set gs \ set bs0)" by (simp only: set_bs0) next show "set ps0 \\<^sub>p D0 = {}" proof show "set ps0 \\<^sub>p D0 \ {}" proof fix x assume "x \ set ps0 \\<^sub>p D0" hence "x \\<^sub>p set ps0 \\<^sub>p D0" by (simp add: in_pair_alt) thus "x \ {}" by (auto simp: assms(3)) qed qed simp next fix p' q' assume "processed (p', q') (gs @ bs0) ps0" hence proc: "processed (p', q') (gs @ bs) ps0" by (simp add: set_bs0 processed_alt) hence "p' \ set gs \ set bs" and "q' \ set gs \ set bs" and "(p', q') \\<^sub>p set ps0" by (auto dest: processedD1 processedD2 processedD3) assume "(p', q') \\<^sub>p D0" and "fst p' \ 0" and "fst q' \ 0" have "crit_pair_cbelow_on d m (fst ` (set gs \ set bs)) (fst p') (fst q')" proof (cases "p' = q'") case True from dg show ?thesis unfolding True proof (rule crit_pair_cbelow_same) from \q' \ set gs \ set bs\ have "fst q' \ fst ` (set gs \ set bs)" by simp from this \fst ` (set gs \ set bs) \ dgrad_p_set d m\ show "fst q' \ dgrad_p_set d m" .. qed next case False show ?thesis proof (cases "component_of_term (lt (fst p')) = component_of_term (lt (fst q'))") case True show ?thesis proof (cases "p' \ set gs \ q' \ set gs") case True note dg \fst ` set gs \ dgrad_p_set d m\ assms(6) moreover from True have "fst p' \ fst ` set gs" and "fst q' \ fst ` set gs" by simp_all ultimately have "crit_pair_cbelow_on d m (fst ` set gs) (fst p') (fst q')" using \fst p' \ 0\ \fst q' \ 0\ by (rule GB_imp_crit_pair_cbelow_dgrad_p_set) moreover have "fst ` set gs \ fst ` (set gs \ set bs)" by blast ultimately show ?thesis by (rule crit_pair_cbelow_mono) next case False with \p' \ set gs \ set bs\ \q' \ set gs \ set bs\ have "(p', q') \\<^sub>p set bs \ (set gs \ set bs)" by (auto simp: in_pair_iff) with \(p', q') \\<^sub>p D0\ have "(p', q') \\<^sub>p set ps0" by (simp add: assms(3)) with \(p', q') \\<^sub>p set ps0\ show ?thesis .. qed next case False thus ?thesis by (rule crit_pair_cbelow_distinct_component) qed qed thus "crit_pair_cbelow_on d m (fst ` (set gs \ set bs0)) (fst p') (fst q')" by (simp only: set_bs0) next fix B a b assume "set gs \ set bs0 \ B" hence B_sup: "set gs \ set bs \ B" by (simp only: set_bs0) assume B_sub: "fst ` B \ dgrad_p_set d m" assume "(a, b) \\<^sub>p D0" hence ab_in: "(a, b) \\<^sub>p set bs \ (set gs \ set bs)" and "(a, b) \\<^sub>p set ps0" by (simp_all add: assms(3)) assume "fst a \ 0" and "fst b \ 0" assume *: "\x y. x \ set gs \ set bs0 \ y \ set gs \ set bs0 \ (x, y) \\<^sub>p D0 \ fst x \ 0 \ fst y \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" show "crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" proof (cases "a = b") case True from ab_in have "b \ set gs \ set bs" unfolding True in_pair_iff set_diff_list by auto hence "fst b \ fst ` (set gs \ set bs)" by fastforce from this gs_bs_sub have "fst b \ dgrad_p_set d m" .. with dg show ?thesis unfolding True by (rule crit_pair_cbelow_same) next case False note ap dg moreover from B_sup have B_sup': "set gs \ set [] \ set bs \ B" by simp moreover note B_sub moreover from ab_in have "(a, b) \\<^sub>p set bs \ (set gs \ set [] \ set bs)" by simp moreover have "set [] \ set [] \ (set gs \ set [])" by simp moreover from assms(7) have "unique_idx (gs @ [] @ bs) data" by (simp add: unique_idx_def set_bs0) ultimately show ?thesis using assms(6) False \fst a \ 0\ \fst b \ 0\ proof (rule ap_specD2) fix x y :: "('t, 'b, 'c) pdata" assume "(x, y) \\<^sub>p set (ap gs [] [] bs data)" hence "(x, y) \\<^sub>p set ps0" by (simp only: ps0_def) also have "... \ set bs0 \ (set gs \ set bs0)" by (fact ps0_sub) also have "... \ (set gs \ set bs0) \ (set gs \ set bs0)" by fastforce finally have "(x, y) \ (set gs \ set bs0) \ (set gs \ set bs0)" by (simp only: in_pair_same) hence "x \ set gs \ set bs0" and "y \ set gs \ set bs0" by simp_all moreover from \(x, y) \\<^sub>p set ps0\ have "(x, y) \\<^sub>p D0" by (simp add: D0_def) moreover assume "fst x \ 0" and "fst y \ 0" ultimately show "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" by (rule *) next fix x y :: "('t, 'b, 'c) pdata" assume "x \ set gs \ set []" and "y \ set gs \ set []" hence "fst x \ fst ` set gs" and "fst y \ fst ` set gs" by simp_all assume "fst x \ 0" and "fst y \ 0" with dg \fst ` set gs \ dgrad_p_set d m\ assms(6) \fst x \ fst ` set gs\ \fst y \ fst ` set gs\ have "crit_pair_cbelow_on d m (fst ` set gs) (fst x) (fst y)" by (rule GB_imp_crit_pair_cbelow_dgrad_p_set) moreover from B_sup have "fst ` set gs \ fst ` B" by fastforce ultimately show "crit_pair_cbelow_on d m (fst ` B) (fst x) (fst y)" by (rule crit_pair_cbelow_mono) qed qed qed qed qed subsubsection \Function \gb_schema_aux\\ function (domintros) gb_schema_aux :: "nat \ nat \ 'd \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata_pair list \ ('t, 'b, 'c) pdata list" where "gb_schema_aux data bs ps = (if ps = [] then gs @ bs else (let sps = sel gs bs ps (snd data); ps0 = ps -- sps; aux = compl gs bs ps0 sps (snd data); remcomps = fst (data) - count_const_lt_components (fst aux) in (if remcomps = 0 then full_gb (gs @ bs) else let (hs, data') = add_indices aux (snd data) in gb_schema_aux (remcomps, data') (ab gs bs hs data') (ap gs bs ps0 hs data') ) ) )" by pat_completeness auto text \The \data\ parameter of @{const gb_schema_aux} is a triple \(c, i, d)\, where \c\ is the number of components \cmp\ of the input list for which the current basis \gs @ bs\ does @{emph \not\} yet contain an element whose leading power-product is \0\ and has component \cmp\. As soon as \c\ gets \0\, the function can return a trivial Gr\"obner basis, since then the submodule generated by the input list is just the full module. This idea generalizes the well-known fact that if a set of scalar polynomials contains a non-zero constant, the ideal generated by that set is the whole ring. \i\ is the total number of polynomials generated during the execution of the function so far; it is used to attach unique indices to the polynomials for fast equality tests. \d\, finally, is some arbitrary data-field that may be used by concrete instances of @{const gb_schema_aux} for storing information.\ lemma gb_schema_aux_domI1: "gb_schema_aux_dom (data, bs, [])" by (rule gb_schema_aux.domintros, simp) lemma gb_schema_aux_domI2: assumes "struct_spec sel ap ab compl" shows "gb_schema_aux_dom (data, args)" proof - from assms have sel: "sel_spec sel" and ap: "ap_spec ap" and ab: "ab_spec ab" by (rule struct_specD)+ from ex_dgrad obtain d::"'a \ nat" where dg: "dickson_grading d" .. let ?R = "gb_schema_aux_term d gs" from dg have "wf ?R" by (rule gb_schema_aux_term_wf) thus ?thesis proof (induct args arbitrary: data rule: wf_induct_rule) fix x data assume IH: "\y data'. (y, x) \ ?R \ gb_schema_aux_dom (data', y)" obtain bs ps where x: "x = (bs, ps)" by (meson case_prodE case_prodI2) show "gb_schema_aux_dom (data, x)" unfolding x proof (rule gb_schema_aux.domintros) fix rc0 n0 data0 hs n1 data1 assume "ps \ []" and hs_data': "(hs, n1, data1) = add_indices (compl gs bs (ps -- sel gs bs ps (n0, data0)) (sel gs bs ps (n0, data0)) (n0, data0)) (n0, data0)" and data: "data = (rc0, n0, data0)" define sps where "sps = sel gs bs ps (n0, data0)" define data' where "data' = (n1, data1)" define rc where "rc = rc0 - count_const_lt_components (fst (compl gs bs (ps -- sel gs bs ps (n0, data0)) (sel gs bs ps (n0, data0)) (n0, data0)))" from hs_data' have hs: "hs = fst (add_indices (compl gs bs (ps -- sps) sps (snd data)) (snd data))" unfolding sps_def data snd_conv by (metis fstI) show "gb_schema_aux_dom ((rc, data'), ab gs bs hs data', ap gs bs (ps -- sps) hs data')" proof (rule IH, simp add: x gb_schema_aux_term_def gb_schema_aux_term1_def gb_schema_aux_term2_def, intro conjI) show "fst ` set (ab gs bs hs data') \p fst ` set bs \ ab gs bs hs data' = bs \ card (set (ap gs bs (ps -- sps) hs data')) < card (set ps)" proof (cases "hs = []") case True have "ab gs bs hs data' = bs \ card (set (ap gs bs (ps -- sps) hs data')) < card (set ps)" proof (simp only: True, rule) from ab show "ab gs bs [] data' = bs" by (rule ab_specD2) next from sel \ps \ []\ have "sps \ []" and "set sps \ set ps" unfolding sps_def by (rule sel_specD1, rule sel_specD2) moreover from sel_specD1[OF sel \ps \ []\] have "set sps \ {}" by (simp add: sps_def) ultimately have "set ps \ set sps \ {}" by (simp add: inf.absorb_iff2) hence "set (ps -- sps) \ set ps" unfolding set_diff_list by fastforce hence "card (set (ps -- sps)) < card (set ps)" by (simp add: psubset_card_mono) moreover have "card (set (ap gs bs (ps -- sps) [] data')) \ card (set (ps -- sps))" by (rule card_mono, fact finite_set, rule ap_spec_Nil_subset, fact ap) ultimately show "card (set (ap gs bs (ps -- sps) [] data')) < card (set ps)" by simp qed thus ?thesis .. next case False with assms \ps \ []\ sps_def hs have "fst ` set (ab gs bs hs data') \p fst ` set bs" unfolding data snd_conv by (rule struct_spec_red_supset) thus ?thesis .. qed next from dg assms \ps \ []\ sps_def hs show "dgrad_p_set_le d (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) (args_to_set (gs, bs, ps))" unfolding data snd_conv by (rule dgrad_p_set_le_args_to_set_struct) next from assms \ps \ []\ sps_def hs show "component_of_term ` Keys (args_to_set (gs, ab gs bs hs data', ap gs bs (ps -- sps) hs data')) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" unfolding data snd_conv by (rule components_subset_struct) qed qed qed qed lemma gb_schema_aux_Nil [simp, code]: "gb_schema_aux data bs [] = gs @ bs" by (simp add: gb_schema_aux.psimps[OF gb_schema_aux_domI1]) lemmas gb_schema_aux_simps = gb_schema_aux.psimps[OF gb_schema_aux_domI2] lemma gb_schema_aux_induct [consumes 1, case_names base rec1 rec2]: assumes "struct_spec sel ap ab compl" assumes base: "\bs data. P data bs [] (gs @ bs)" and rec1: "\bs ps sps data. ps \ [] \ sps = sel gs bs ps (snd data) \ fst (data) \ count_const_lt_components (fst (compl gs bs (ps -- sps) sps (snd data))) \ P data bs ps (full_gb (gs @ bs))" and rec2: "\bs ps sps aux hs rc data data'. ps \ [] \ sps = sel gs bs ps (snd data) \ aux = compl gs bs (ps -- sps) sps (snd data) \ (hs, data') = add_indices aux (snd data) \ rc = fst data - count_const_lt_components (fst aux) \ 0 < rc \ P (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data') (gb_schema_aux (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data')) \ P data bs ps (gb_schema_aux (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data'))" shows "P data bs ps (gb_schema_aux data bs ps)" proof - from assms(1) have "gb_schema_aux_dom (data, bs, ps)" by (rule gb_schema_aux_domI2) thus ?thesis proof (induct data bs ps rule: gb_schema_aux.pinduct) case (1 data bs ps) show ?case proof (cases "ps = []") case True show ?thesis by (simp add: True, rule base) next case False show ?thesis proof (simp add: gb_schema_aux_simps[OF assms(1), of data bs ps] False Let_def split: if_split, intro conjI impI) define sps where "sps = sel gs bs ps (snd data)" assume "fst data \ count_const_lt_components (fst (compl gs bs (ps -- sps) sps (snd data)))" with False sps_def show "P data bs ps (full_gb (gs @ bs))" by (rule rec1) next define sps where "sps = sel gs bs ps (snd data)" define aux where "aux = compl gs bs (ps -- sps) sps (snd data)" define hs where "hs = fst (add_indices aux (snd data))" define data' where "data' = snd (add_indices aux (snd data))" define rc where "rc = fst data - count_const_lt_components (fst aux)" have eq: "add_indices aux (snd data) = (hs, data')" by (simp add: hs_def data'_def) assume "\ fst data \ count_const_lt_components (fst aux)" hence "0 < rc" by (simp add: rc_def) hence "rc \ 0" by simp show "P data bs ps (case add_indices aux (snd data) of (hs, data') \ gb_schema_aux (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data'))" unfolding eq prod.case using False sps_def aux_def eq[symmetric] rc_def \0 < rc\ proof (rule rec2) show "P (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data') (gb_schema_aux (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data'))" using False sps_def refl aux_def rc_def \rc \ 0\ eq[symmetric] refl by (rule 1) qed qed qed qed qed lemma gb_schema_dummy_eq_gb_schema_aux: assumes "struct_spec sel ap ab compl" shows "fst (gb_schema_dummy data D bs ps) = gb_schema_aux data bs ps" using assms proof (induct data D bs ps rule: gb_schema_dummy_induct) case (base bs data D) show ?case by simp next case (rec1 bs ps sps data D) thus ?case by (simp add: gb_schema_aux.psimps[OF gb_schema_aux_domI2, OF assms]) next case (rec2 bs ps sps aux hs rc data data' D D') note rec2.hyps(8) also from rec2.hyps(1, 2, 3) rec2.hyps(4)[symmetric] rec2.hyps(5, 6, 7) have "gb_schema_aux (rc, data') (ab gs bs hs data') (ap gs bs (ps -- sps) hs data') = gb_schema_aux data bs ps" by (simp add: gb_schema_aux.psimps[OF gb_schema_aux_domI2, OF assms, of data] Let_def) finally show ?case . qed corollary gb_schema_aux_dgrad_p_set_le: assumes "dickson_grading d" and "struct_spec sel ap ab compl" shows "dgrad_p_set_le d (fst ` set (gb_schema_aux data bs ps)) (args_to_set (gs, bs, ps))" using fst_gb_schema_dummy_dgrad_p_set_le[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(2)] . corollary gb_schema_aux_components: assumes "struct_spec sel ap ab compl" and "set ps \ set bs \ (set gs \ set bs)" shows "component_of_term ` Keys (fst ` set (gb_schema_aux data bs ps)) = component_of_term ` Keys (args_to_set (gs, bs, ps))" using fst_gb_schema_dummy_components[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(1)] . lemma gb_schema_aux_pmdl: assumes "struct_spec sel ap ab compl" and "compl_pmdl compl" and "is_Groebner_basis (fst ` set gs)" and "set ps \ set bs \ (set gs \ set bs)" and "unique_idx (gs @ bs) (snd data)" and "rem_comps_spec (gs @ bs) data" shows "pmdl (fst ` set (gb_schema_aux data bs ps)) = pmdl (fst ` set (gs @ bs))" using fst_gb_schema_dummy_pmdl[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(1)] . corollary gb_schema_aux_dgrad_p_set_le_init: assumes "dickson_grading d" and "struct_spec sel ap ab compl" shows "dgrad_p_set_le d (fst ` set (gb_schema_aux data (ab gs [] bs (snd data)) (ap gs [] [] bs (snd data)))) (fst ` (set gs \ set bs))" using fst_gb_schema_dummy_dgrad_p_set_le_init[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(2)] . corollary gb_schema_aux_dgrad_p_set_init: assumes "dickson_grading d" and "struct_spec sel ap ab compl" and "fst ` (set gs \ set bs) \ dgrad_p_set d m" shows "fst ` set (gb_schema_aux (rc, data) (ab gs [] bs data) (ap gs [] [] bs data)) \ dgrad_p_set d m" using fst_gb_schema_dummy_dgrad_p_set_init[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(2)] . corollary gb_schema_aux_components_init: assumes "struct_spec sel ap ab compl" shows "component_of_term ` Keys (fst ` set (gb_schema_aux (rc, data) (ab gs [] bs data) (ap gs [] [] bs data))) = component_of_term ` Keys (fst ` set (gs @ bs))" using fst_gb_schema_dummy_components_init[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms] . corollary gb_schema_aux_pmdl_init: assumes "struct_spec sel ap ab compl" and "compl_pmdl compl" and "is_Groebner_basis (fst ` set gs)" and "unique_idx (gs @ ab gs [] bs data) data" and "rem_comps_spec (gs @ ab gs [] bs data) (rc, data)" shows "pmdl (fst ` set (gb_schema_aux (rc, data) (ab gs [] bs data) (ap gs [] [] bs data))) = pmdl (fst ` (set (gs @ bs)))" using fst_gb_schema_dummy_pmdl_init[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(1)] . lemma gb_schema_aux_isGB_init: assumes "struct_spec sel ap ab compl" and "compl_conn compl" and "is_Groebner_basis (fst ` set gs)" and "unique_idx (gs @ ab gs [] bs data) data" and "rem_comps_spec (gs @ ab gs [] bs data) (rc, data)" shows "is_Groebner_basis (fst ` set (gb_schema_aux (rc, data) (ab gs [] bs data) (ap gs [] [] bs data)))" using fst_gb_schema_dummy_isGB_init[OF assms] unfolding gb_schema_dummy_eq_gb_schema_aux[OF assms(1)] . end subsubsection \Functions \gb_schema_direct\ and \term gb_schema_incr\\ definition gb_schema_direct :: "('t, 'b, 'c, 'd) selT \ ('t, 'b, 'c, 'd) apT \ ('t, 'b, 'c, 'd) abT \ ('t, 'b, 'c, 'd) complT \ ('t, 'b, 'c) pdata' list \ 'd \ ('t, 'b::field, 'c::default) pdata' list" where "gb_schema_direct sel ap ab compl bs0 data0 = (let data = (length bs0, data0); bs1 = fst (add_indices (bs0, data0) (0, data0)); bs = ab [] [] bs1 data in map (\(f, _, d). (f, d)) (gb_schema_aux sel ap ab compl [] (count_rem_components bs, data) bs (ap [] [] [] bs1 data)) )" primrec gb_schema_incr :: "('t, 'b, 'c, 'd) selT \ ('t, 'b, 'c, 'd) apT \ ('t, 'b, 'c, 'd) abT \ ('t, 'b, 'c, 'd) complT \ (('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata \ 'd \ 'd) \ ('t, 'b, 'c) pdata' list \ 'd \ ('t, 'b::field, 'c::default) pdata' list" where "gb_schema_incr _ _ _ _ _ [] _ = []"| "gb_schema_incr sel ap ab compl upd (b0 # bs) data = (let (gs, n, data') = add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data); b = (fst b0, n, snd b0); data'' = upd gs b data' in map (\(f, _, d). (f, d)) (gb_schema_aux sel ap ab compl gs (count_rem_components (b # gs), Suc n, data'') (ab gs [] [b] (Suc n, data'')) (ap gs [] [] [b] (Suc n, data''))) )" lemma (in -) fst_set_drop_indices: "fst ` (\(f, _, d). (f, d)) ` A = fst ` A" for A::"('x \ 'y \ 'z) set" by (simp add: image_image, rule image_cong, fact refl, simp add: prod.case_eq_if) lemma fst_gb_schema_direct: "fst ` set (gb_schema_direct sel ap ab compl bs0 data0) = (let data = (length bs0, data0); bs1 = fst (add_indices (bs0, data0) (0, data0)); bs = ab [] [] bs1 data in fst ` set (gb_schema_aux sel ap ab compl [] (count_rem_components bs, data) bs (ap [] [] [] bs1 data)) )" by (simp add: gb_schema_direct_def Let_def fst_set_drop_indices) lemma gb_schema_direct_dgrad_p_set: assumes "dickson_grading d" and "struct_spec sel ap ab compl" and "fst ` set bs \ dgrad_p_set d m" shows "fst ` set (gb_schema_direct sel ap ab compl bs data) \ dgrad_p_set d m" unfolding fst_gb_schema_direct Let_def using assms(1, 2) proof (rule gb_schema_aux_dgrad_p_set_init) show "fst ` (set [] \ set (fst (add_indices (bs, data) (0, data)))) \ dgrad_p_set d m" using assms(3) by (simp add: image_Un fst_set_add_indices) qed theorem gb_schema_direct_isGB: assumes "struct_spec sel ap ab compl" and "compl_conn compl" shows "is_Groebner_basis (fst ` set (gb_schema_direct sel ap ab compl bs data))" unfolding fst_gb_schema_direct Let_def using assms proof (rule gb_schema_aux_isGB_init) from is_Groebner_basis_empty show "is_Groebner_basis (fst ` set [])" by simp next let ?data = "(length bs, data)" from assms(1) have "ab_spec ab" by (rule struct_specD) moreover have "unique_idx ([] @ []) (0, data)" by (simp add: unique_idx_Nil) ultimately show "unique_idx ([] @ ab [] [] (fst (add_indices (bs, data) (0, data))) ?data) ?data" proof (rule unique_idx_ab) show "(fst (add_indices (bs, data) (0, data)), length bs, data) = add_indices (bs, data) (0, data)" by (simp add: add_indices_def) qed qed (simp add: rem_comps_spec_count_rem_components) theorem gb_schema_direct_pmdl: assumes "struct_spec sel ap ab compl" and "compl_pmdl compl" shows "pmdl (fst ` set (gb_schema_direct sel ap ab compl bs data)) = pmdl (fst ` set bs)" proof - have "pmdl (fst ` set (gb_schema_direct sel ap ab compl bs data)) = pmdl (fst ` set ([] @ (fst (add_indices (bs, data) (0, data)))))" unfolding fst_gb_schema_direct Let_def using assms proof (rule gb_schema_aux_pmdl_init) from is_Groebner_basis_empty show "is_Groebner_basis (fst ` set [])" by simp next let ?data = "(length bs, data)" from assms(1) have "ab_spec ab" by (rule struct_specD) moreover have "unique_idx ([] @ []) (0, data)" by (simp add: unique_idx_Nil) ultimately show "unique_idx ([] @ ab [] [] (fst (add_indices (bs, data) (0, data))) ?data) ?data" proof (rule unique_idx_ab) show "(fst (add_indices (bs, data) (0, data)), length bs, data) = add_indices (bs, data) (0, data)" by (simp add: add_indices_def) qed qed (simp add: rem_comps_spec_count_rem_components) thus ?thesis by (simp add: fst_set_add_indices) qed lemma fst_gb_schema_incr: "fst ` set (gb_schema_incr sel ap ab compl upd (b0 # bs) data) = (let (gs, n, data') = add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data); b = (fst b0, n, snd b0); data'' = upd gs b data' in fst ` set (gb_schema_aux sel ap ab compl gs (count_rem_components (b # gs), Suc n, data'') (ab gs [] [b] (Suc n, data'')) (ap gs [] [] [b] (Suc n, data''))) )" by (simp only: gb_schema_incr.simps Let_def prod.case_distrib[of set] prod.case_distrib[of "image fst"] set_map fst_set_drop_indices) lemma gb_schema_incr_dgrad_p_set: assumes "dickson_grading d" and "struct_spec sel ap ab compl" and "fst ` set bs \ dgrad_p_set d m" shows "fst ` set (gb_schema_incr sel ap ab compl upd bs data) \ dgrad_p_set d m" using assms(3) proof (induct bs) case Nil show ?case by simp next case (Cons b0 bs) from Cons(2) have 1: "fst b0 \ dgrad_p_set d m" and 2: "fst ` set bs \ dgrad_p_set d m" by simp_all show ?case proof (simp only: fst_gb_schema_incr Let_def split: prod.splits, simp, intro allI impI) fix gs n data' assume "add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data) = (gs, n, data')" hence gs: "gs = fst (add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data))" by simp define b where "b = (fst b0, n, snd b0)" define data'' where "data'' = upd gs b data'" from assms(1, 2) show "fst ` set (gb_schema_aux sel ap ab compl gs (count_rem_components (b # gs), Suc n, data'') (ab gs [] [b] (Suc n, data'')) (ap gs [] [] [b] (Suc n, data''))) \ dgrad_p_set d m" proof (rule gb_schema_aux_dgrad_p_set_init) from 1 Cons(1)[OF 2] show "fst ` (set gs \ set [b]) \ dgrad_p_set d m" by (simp add: gs fst_set_add_indices b_def) qed qed qed theorem gb_schema_incr_dgrad_p_set_isGB: assumes "struct_spec sel ap ab compl" and "compl_conn compl" shows "is_Groebner_basis (fst ` set (gb_schema_incr sel ap ab compl upd bs data))" proof (induct bs) case Nil from is_Groebner_basis_empty show ?case by simp next case (Cons b0 bs) show ?case proof (simp only: fst_gb_schema_incr Let_def split: prod.splits, simp, intro allI impI) fix gs n data' assume *: "add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data) = (gs, n, data')" hence gs: "gs = fst (add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data))" by simp define b where "b = (fst b0, n, snd b0)" define data'' where "data'' = upd gs b data'" from assms(1) have ab: "ab_spec ab" by (rule struct_specD3) from Cons have "is_Groebner_basis (fst ` set gs)" by (simp add: gs fst_set_add_indices) with assms show "is_Groebner_basis (fst ` set (gb_schema_aux sel ap ab compl gs (count_rem_components (b # gs), Suc n, data'') (ab gs [] [b] (Suc n, data'')) (ap gs [] [] [b] (Suc n, data''))))" proof (rule gb_schema_aux_isGB_init) from ab show "unique_idx (gs @ ab gs [] [b] (Suc n, data'')) (Suc n, data'')" proof (rule unique_idx_ab) from unique_idx_Nil *[symmetric] have "unique_idx ([] @ gs) (n, data')" by (rule unique_idx_append) thus "unique_idx (gs @ []) (n, data')" by simp next show "([b], Suc n, data'') = add_indices ([b0], data'') (n, data')" by (simp add: add_indices_def b_def) qed next have "rem_comps_spec (b # gs) (count_rem_components (b # gs), Suc n, data'')" by (fact rem_comps_spec_count_rem_components) moreover have "set (b # gs) = set (gs @ ab gs [] [b] (Suc n, data''))" by (simp add: ab_specD1[OF ab]) ultimately show "rem_comps_spec (gs @ ab gs [] [b] (Suc n, data'')) (count_rem_components (b # gs), Suc n, data'')" by (simp only: rem_comps_spec_def) qed qed qed theorem gb_schema_incr_pmdl: assumes "struct_spec sel ap ab compl" and "compl_conn compl" "compl_pmdl compl" shows "pmdl (fst ` set (gb_schema_incr sel ap ab compl upd bs data)) = pmdl (fst ` set bs)" proof (induct bs) case Nil show ?case by simp next case (Cons b0 bs) show ?case proof (simp only: fst_gb_schema_incr Let_def split: prod.splits, simp, intro allI impI) fix gs n data' assume *: "add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data) = (gs, n, data')" hence gs: "gs = fst (add_indices (gb_schema_incr sel ap ab compl upd bs data, data) (0, data))" by simp define b where "b = (fst b0, n, snd b0)" define data'' where "data'' = upd gs b data'" from assms(1) have ab: "ab_spec ab" by (rule struct_specD3) from assms(1, 2) have "is_Groebner_basis (fst ` set gs)" unfolding gs fst_conv fst_set_add_indices by (rule gb_schema_incr_dgrad_p_set_isGB) with assms(1, 3) have eq: "pmdl (fst ` set (gb_schema_aux sel ap ab compl gs (count_rem_components (b # gs), Suc n, data'') (ab gs [] [b] (Suc n, data'')) (ap gs [] [] [b] (Suc n, data'')))) = pmdl (fst ` set (gs @ [b]))" proof (rule gb_schema_aux_pmdl_init) from ab show "unique_idx (gs @ ab gs [] [b] (Suc n, data'')) (Suc n, data'')" proof (rule unique_idx_ab) from unique_idx_Nil *[symmetric] have "unique_idx ([] @ gs) (n, data')" by (rule unique_idx_append) thus "unique_idx (gs @ []) (n, data')" by simp next show "([b], Suc n, data'') = add_indices ([b0], data'') (n, data')" by (simp add: add_indices_def b_def) qed next have "rem_comps_spec (b # gs) (count_rem_components (b # gs), Suc n, data'')" by (fact rem_comps_spec_count_rem_components) moreover have "set (b # gs) = set (gs @ ab gs [] [b] (Suc n, data''))" by (simp add: ab_specD1[OF ab]) ultimately show "rem_comps_spec (gs @ ab gs [] [b] (Suc n, data'')) (count_rem_components (b # gs), Suc n, data'')" by (simp only: rem_comps_spec_def) qed also have "... = pmdl (insert (fst b) (fst ` set gs))" by simp also from Cons have "... = pmdl (insert (fst b) (fst ` set bs))" unfolding gs fst_conv fst_set_add_indices by (rule pmdl.span_insert_cong) finally show "pmdl (fst ` set (gb_schema_aux sel ap ab compl gs (count_rem_components (b # gs), Suc n, data'') (ab gs [] [b] (Suc n, data'')) (ap gs [] [] [b] (Suc n, data'')))) = pmdl (insert (fst b0) (fst ` set bs))" by (simp add: b_def) qed qed subsection \Suitable Instances of the @{emph \add-pairs\} Parameter\ subsubsection \Specification of the @{emph \crit\} parameters\ type_synonym (in -) ('t, 'b, 'c, 'd) icritT = "nat \ 'd \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata \ ('t, 'b, 'c) pdata \ bool" type_synonym (in -) ('t, 'b, 'c, 'd) ncritT = "nat \ 'd \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ bool \ (bool \ ('t, 'b, 'c) pdata_pair) list \ ('t, 'b, 'c) pdata \ ('t, 'b, 'c) pdata \ bool" type_synonym (in -) ('t, 'b, 'c, 'd) ocritT = "nat \ 'd \ ('t, 'b, 'c) pdata list \ (bool \ ('t, 'b, 'c) pdata_pair) list \ ('t, 'b, 'c) pdata \ ('t, 'b, 'c) pdata \ bool" definition icrit_spec :: "('t, 'b::field, 'c, 'd) icritT \ bool" where "icrit_spec crit \ (\d m data gs bs hs p q. dickson_grading d \ fst ` (set gs \ set bs \ set hs) \ dgrad_p_set d m \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ p \ set hs \ q \ set gs \ set bs \ set hs \ fst p \ 0 \ fst q \ 0 \ crit data gs bs hs p q \ crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q))" text \Criteria satisfying @{const icrit_spec} can be used for discarding pairs @{emph \instantly\}, without reference to any other pairs. The product criterion for scalar polynomials satisfies @{const icrit_spec}, and so does the component criterion (which checks whether the component-indices of the leading terms of two polynomials are identical).\ definition ncrit_spec :: "('t, 'b::field, 'c, 'd) ncritT \ bool" where "ncrit_spec crit \ (\d m data gs bs hs ps B q_in_bs p q. dickson_grading d \ set gs \ set bs \ set hs \ B \ fst ` B \ dgrad_p_set d m \ snd ` set ps \ set hs \ (set gs \ set bs \ set hs) \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ (q_in_bs \ (q \ set gs \ set bs)) \ (\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')) \ (\p' q'. p' \ set gs \ set bs \ q' \ set gs \ set bs \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')) \ p \ set hs \ q \ set gs \ set bs \ set hs \ fst p \ 0 \ fst q \ 0 \ crit data gs bs hs q_in_bs ps p q \ crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q))" definition ocrit_spec :: "('t, 'b::field, 'c, 'd) ocritT \ bool" where "ocrit_spec crit \ (\d m data hs ps B p q. dickson_grading d \ set hs \ B \ fst ` B \ dgrad_p_set d m \ unique_idx (p # q # hs @ (map (fst \ snd) ps) @ (map (snd \ snd) ps)) data \ (\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')) \ p \ B \ q \ B \ fst p \ 0 \ fst q \ 0 \ crit data hs ps p q \ crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q))" text \Criteria satisfying @{const ncrit_spec} can be used for discarding new pairs by reference to new and old elements, whereas criteria satisfying @{const ocrit_spec} can be used for discarding old pairs by reference to new elements @{emph \only\} (no existing ones!). The chain criterion satisfies both @{const ncrit_spec} and @{const ocrit_spec}.\ lemma icrit_specI: assumes "\d m data gs bs hs p q. dickson_grading d \ fst ` (set gs \ set bs \ set hs) \ dgrad_p_set d m \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ p \ set hs \ q \ set gs \ set bs \ set hs \ fst p \ 0 \ fst q \ 0 \ crit data gs bs hs p q \ crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q)" shows "icrit_spec crit" unfolding icrit_spec_def using assms by auto lemma icrit_specD: assumes "icrit_spec crit" and "dickson_grading d" and "fst ` (set gs \ set bs \ set hs) \ dgrad_p_set d m" and "unique_idx (gs @ bs @ hs) data" and "is_Groebner_basis (fst ` set gs)" and "p \ set hs" and "q \ set gs \ set bs \ set hs" and "fst p \ 0" and "fst q \ 0" and "crit data gs bs hs p q" shows "crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q)" using assms unfolding icrit_spec_def by blast lemma ncrit_specI: assumes "\d m data gs bs hs ps B q_in_bs p q. dickson_grading d \ set gs \ set bs \ set hs \ B \ fst ` B \ dgrad_p_set d m \ snd ` set ps \ set hs \ (set gs \ set bs \ set hs) \ unique_idx (gs @ bs @ hs) data \ is_Groebner_basis (fst ` set gs) \ (q_in_bs \ q \ set gs \ set bs) \ (\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')) \ (\p' q'. p' \ set gs \ set bs \ q' \ set gs \ set bs \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')) \ p \ set hs \ q \ set gs \ set bs \ set hs \ fst p \ 0 \ fst q \ 0 \ crit data gs bs hs q_in_bs ps p q \ crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" shows "ncrit_spec crit" unfolding ncrit_spec_def by (intro allI impI, rule assms, assumption+, meson, meson, assumption+) lemma ncrit_specD: assumes "ncrit_spec crit" and "dickson_grading d" and "set gs \ set bs \ set hs \ B" and "fst ` B \ dgrad_p_set d m" and "snd ` set ps \ set hs \ (set gs \ set bs \ set hs)" and "unique_idx (gs @ bs @ hs) data" and "is_Groebner_basis (fst ` set gs)" and "q_in_bs \ q \ set gs \ set bs" and "\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and "\p' q'. p' \ set gs \ set bs \ q' \ set gs \ set bs \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and "p \ set hs" and "q \ set gs \ set bs \ set hs" and "fst p \ 0" and "fst q \ 0" and "crit data gs bs hs q_in_bs ps p q" shows "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" using assms unfolding ncrit_spec_def by blast lemma ocrit_specI: assumes "\d m data hs ps B p q. dickson_grading d \ set hs \ B \ fst ` B \ dgrad_p_set d m \ unique_idx (p # q # hs @ (map (fst \ snd) ps) @ (map (snd \ snd) ps)) data \ (\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')) \ p \ B \ q \ B \ fst p \ 0 \ fst q \ 0 \ crit data hs ps p q \ crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" shows "ocrit_spec crit" unfolding ocrit_spec_def by (intro allI impI, rule assms, assumption+, meson, assumption+) lemma ocrit_specD: assumes "ocrit_spec crit" and "dickson_grading d" and "set hs \ B" and "fst ` B \ dgrad_p_set d m" and "unique_idx (p # q # hs @ (map (fst \ snd) ps) @ (map (snd \ snd) ps)) data" and "\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and "p \ B" and "q \ B" and "fst p \ 0" and "fst q \ 0" and "crit data hs ps p q" shows "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" using assms unfolding ocrit_spec_def by blast subsubsection \Suitable instances of the @{emph \crit\} parameters\ definition component_crit :: "('t, 'b::zero, 'c, 'd) icritT" where "component_crit data gs bs hs p q \ (component_of_term (lt (fst p)) \ component_of_term (lt (fst q)))" lemma icrit_spec_component_crit: "icrit_spec (component_crit::('t, 'b::field, 'c, 'd) icritT)" proof (rule icrit_specI) fix d m and data::"nat \ 'd" and gs bs hs and p q::"('t, 'b, 'c) pdata" assume "component_crit data gs bs hs p q" hence "component_of_term (lt (fst p)) \ component_of_term (lt (fst q))" by (simp add: component_crit_def) thus "crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q)" by (rule crit_pair_cbelow_distinct_component) qed text \The product criterion is only applicable to scalar polynomials.\ definition product_crit :: "('a, 'b::zero, 'c, 'd) icritT" where "product_crit data gs bs hs p q \ (gcs (punit.lt (fst p)) (punit.lt (fst q)) = 0)" lemma (in gd_term) icrit_spec_product_crit: "punit.icrit_spec (product_crit::('a, 'b::field, 'c, 'd) icritT)" proof (rule punit.icrit_specI) fix d m and data::"nat \ 'd" and gs bs hs and p q::"('a, 'b, 'c) pdata" assume "product_crit data gs bs hs p q" hence *: "gcs (punit.lt (fst p)) (punit.lt (fst q)) = 0" by (simp only: product_crit_def) assume "p \ set hs" and q_in: "q \ set gs \ set bs \ set hs" (is "_ \ ?B") assume "dickson_grading d" and sub: "fst ` (set gs \ set bs \ set hs) \ punit.dgrad_p_set d m" moreover from \p \ set hs\ have "fst p \ fst ` ?B" by simp moreover from q_in have "fst q \ fst ` ?B" by simp moreover assume "fst p \ 0" and "fst q \ 0" ultimately show "punit.crit_pair_cbelow_on d m (fst ` ?B) (fst p) (fst q)" using * by (rule product_criterion) qed text \@{const component_crit} and @{const product_crit} ignore the \data\ parameter.\ fun (in -) pair_in_list :: "(bool \ ('a, 'b, 'c) pdata_pair) list \ nat \ nat \ bool" where "pair_in_list [] _ _ = False" |"pair_in_list ((_, (_, i', _), (_, j', _)) # ps) i j = ((i = i' \ j = j') \ (i = j' \ j = i') \ pair_in_list ps i j)" lemma (in -) pair_in_listE: assumes "pair_in_list ps i j" obtains p q a b where "((p, i, a), (q, j, b)) \\<^sub>p snd ` set ps" using assms proof (induct ps i j arbitrary: thesis rule: pair_in_list.induct) case (1 i j) from 1(2) show ?case by simp next case (2 c p i' a q j' b ps i j) from 2(3) have "(i = i' \ j = j') \ (i = j' \ j = i') \ pair_in_list ps i j" by simp thus ?case proof (elim disjE conjE) assume "i = i'" and "j = j'" have "((p, i, a), (q, j, b)) \\<^sub>p snd ` set ((c, (p, i', a), q, j', b) # ps)" unfolding \i = i'\ \j = j'\ in_pair_iff by fastforce thus ?thesis by (rule 2(2)) next assume "i = j'" and "j = i'" have "((q, i, b), (p, j, a)) \\<^sub>p snd ` set ((c, (p, i', a), q, j', b) # ps)" unfolding \i = j'\ \j = i'\ in_pair_iff by fastforce thus ?thesis by (rule 2(2)) next assume "pair_in_list ps i j" obtain p' q' a' b' where "((p', i, a'), (q', j, b')) \\<^sub>p snd ` set ps" by (rule 2(1), assumption, rule \pair_in_list ps i j\) also have "... \ snd ` set ((c, (p, i', a), q, j', b) # ps)" by auto finally show ?thesis by (rule 2(2)) qed qed definition chain_ncrit :: "('t, 'b::zero, 'c, 'd) ncritT" where "chain_ncrit data gs bs hs q_in_bs ps p q \ (let v = lt (fst p); l = term_of_pair (lcs (pp_of_term v) (lp (fst q)), component_of_term v); i = fst (snd p); j = fst (snd q) in (\r\set gs. let k = fst (snd r) in k \ i \ k \ j \ lt (fst r) adds\<^sub>t l \ pair_in_list ps i k \ (q_in_bs \ pair_in_list ps j k) \ fst r \ 0) \ (\r\set bs. let k = fst (snd r) in k \ i \ k \ j \ lt (fst r) adds\<^sub>t l \ pair_in_list ps i k \ (q_in_bs \ pair_in_list ps j k) \ fst r \ 0) \ (\h\set hs. let k = fst (snd h) in k \ i \ k \ j \ lt (fst h) adds\<^sub>t l \ pair_in_list ps i k \ pair_in_list ps j k \ fst h \ 0))" definition chain_ocrit :: "('t, 'b::zero, 'c, 'd) ocritT" where "chain_ocrit data hs ps p q \ (let v = lt (fst p); l = term_of_pair (lcs (pp_of_term v) (lp (fst q)), component_of_term v); i = fst (snd p); j = fst (snd q) in (\h\set hs. let k = fst (snd h) in k \ i \ k \ j \ lt (fst h) adds\<^sub>t l \ pair_in_list ps i k \ pair_in_list ps j k \ fst h \ 0))" text \@{const chain_ncrit} and @{const chain_ocrit} ignore the \data\ parameter.\ lemma chain_ncritE: assumes "chain_ncrit data gs bs hs q_in_bs ps p q" and "snd ` set ps \ set hs \ (set gs \ set bs \ set hs)" and "unique_idx (gs @ bs @ hs) data" and "p \ set hs" and "q \ set gs \ set bs \ set hs" obtains r where "r \ set gs \ set bs \ set hs" and "fst r \ 0" and "r \ p" and "r \ q" and "lt (fst r) adds\<^sub>t term_of_pair (lcs (lp (fst p)) (lp (fst q)), component_of_term (lt (fst p)))" and "(p, r) \\<^sub>p snd ` set ps" and "(r \ set gs \ set bs \ q_in_bs) \ (q, r) \\<^sub>p snd ` set ps" proof - let ?l = "term_of_pair (lcs (lp (fst p)) (lp (fst q)), component_of_term (lt (fst p)))" let ?i = "fst (snd p)" let ?j = "fst (snd q)" let ?xs = "gs @ bs @ hs" have 3: "x \ set ?xs" if "(x, y) \\<^sub>p snd ` set ps" for x y proof - note that also have "snd ` set ps \ set hs \ (set gs \ set bs \ set hs)" by (fact assms(2)) also have "... \ (set gs \ set bs \ set hs) \ (set gs \ set bs \ set hs)" by fastforce finally have "(x, y) \ (set gs \ set bs \ set hs) \ (set gs \ set bs \ set hs)" by (simp only: in_pair_same) thus ?thesis by simp qed have 4: "x \ set ?xs" if "(y, x) \\<^sub>p snd ` set ps" for x y proof - from that have "(x, y) \\<^sub>p snd ` set ps" by (simp add: in_pair_iff disj_commute) thus ?thesis by (rule 3) qed from assms(1) have "\r \ set gs \ set bs \ set hs. let k = fst (snd r) in k \ ?i \ k \ ?j \ lt (fst r) adds\<^sub>t ?l \ pair_in_list ps ?i k \ ((r \ set gs \ set bs \ q_in_bs) \ pair_in_list ps ?j k) \ fst r \ 0" by (smt UnI1 chain_ncrit_def sup_commute) then obtain r where r_in: "r \ set gs \ set bs \ set hs" and "fst r \ 0" and rp: "fst (snd r) \ ?i" and rq: "fst (snd r) \ ?j" and "lt (fst r) adds\<^sub>t ?l" and 1: "pair_in_list ps ?i (fst (snd r))" and 2: "(r \ set gs \ set bs \ q_in_bs) \ pair_in_list ps ?j (fst (snd r))" unfolding Let_def by blast let ?k = "fst (snd r)" note r_in \fst r \ 0\ moreover from rp have "r \ p" by auto moreover from rq have "r \ q" by auto ultimately show ?thesis using \lt (fst r) adds\<^sub>t ?l\ proof from 1 obtain p' r' a b where *: "((p', ?i, a), (r', ?k, b)) \\<^sub>p snd ` set ps" by (rule pair_in_listE) note assms(3) moreover from * have "(p', ?i, a) \ set ?xs" by (rule 3) moreover from assms(4) have "p \ set ?xs" by simp moreover have "fst (snd (p', ?i, a)) = ?i" by simp ultimately have p': "(p', ?i, a) = p" by (rule unique_idxD1) note assms(3) moreover from * have "(r', ?k, b) \ set ?xs" by (rule 4) moreover from r_in have "r \ set ?xs" by simp moreover have "fst (snd (r', ?k, b)) = ?k" by simp ultimately have r': "(r', ?k, b) = r" by (rule unique_idxD1) from * show "(p, r) \\<^sub>p snd ` set ps" by (simp only: p' r') next from 2 show "(r \ set gs \ set bs \ q_in_bs) \ (q, r) \\<^sub>p snd ` set ps" proof assume "r \ set gs \ set bs \ q_in_bs" thus ?thesis .. next assume "pair_in_list ps ?j ?k" then obtain q' r' a b where *: "((q', ?j, a), (r', ?k, b)) \\<^sub>p snd ` set ps" by (rule pair_in_listE) note assms(3) moreover from * have "(q', ?j, a) \ set ?xs" by (rule 3) moreover from assms(5) have "q \ set ?xs" by simp moreover have "fst (snd (q', ?j, a)) = ?j" by simp ultimately have q': "(q', ?j, a) = q" by (rule unique_idxD1) note assms(3) moreover from * have "(r', ?k, b) \ set ?xs" by (rule 4) moreover from r_in have "r \ set ?xs" by simp moreover have "fst (snd (r', ?k, b)) = ?k" by simp ultimately have r': "(r', ?k, b) = r" by (rule unique_idxD1) from * have "(q, r) \\<^sub>p snd ` set ps" by (simp only: q' r') thus ?thesis .. qed qed qed lemma chain_ocritE: assumes "chain_ocrit data hs ps p q" and "unique_idx (p # q # hs @ (map (fst \ snd) ps) @ (map (snd \ snd) ps)) data" (is "unique_idx ?xs _") obtains h where "h \ set hs" and "fst h \ 0" and "h \ p" and "h \ q" and "lt (fst h) adds\<^sub>t term_of_pair (lcs (lp (fst p)) (lp (fst q)), component_of_term (lt (fst p)))" and "(p, h) \\<^sub>p snd ` set ps" and "(q, h) \\<^sub>p snd ` set ps" proof - let ?l = "term_of_pair (lcs (lp (fst p)) (lp (fst q)), component_of_term (lt (fst p)))" have 3: "x \ set ?xs" if "(x, y) \\<^sub>p snd ` set ps" for x y proof - from that have "(x, y) \ snd ` set ps \ (y, x) \ snd ` set ps" by (simp only: in_pair_iff) thus ?thesis proof assume "(x, y) \ snd ` set ps" hence "fst (x, y) \ fst ` snd ` set ps" by fastforce thus ?thesis by (simp add: image_comp) next assume "(y, x) \ snd ` set ps" hence "snd (y, x) \ snd ` snd ` set ps" by fastforce thus ?thesis by (simp add: image_comp) qed qed have 4: "x \ set ?xs" if "(y, x) \\<^sub>p snd ` set ps" for x y proof - from that have "(x, y) \\<^sub>p snd ` set ps" by (simp add: in_pair_iff disj_commute) thus ?thesis by (rule 3) qed from assms(1) obtain h where "h \ set hs" and "fst h \ 0" and hp: "fst (snd h) \ fst (snd p)" and hq: "fst (snd h) \ fst (snd q)" and "lt (fst h) adds\<^sub>t ?l" and 1: "pair_in_list ps (fst (snd p)) (fst (snd h))" and 2: "pair_in_list ps (fst (snd q)) (fst (snd h))" unfolding chain_ocrit_def Let_def by blast let ?i = "fst (snd p)" let ?j = "fst (snd q)" let ?k = "fst (snd h)" note \h \ set hs\ \fst h \ 0\ moreover from hp have "h \ p" by auto moreover from hq have "h \ q" by auto ultimately show ?thesis using \lt (fst h) adds\<^sub>t ?l\ proof from 1 obtain p' h' a b where *: "((p', ?i, a), (h', ?k, b)) \\<^sub>p snd ` set ps" by (rule pair_in_listE) note assms(2) moreover from * have "(p', ?i, a) \ set ?xs" by (rule 3) moreover have "p \ set ?xs" by simp moreover have "fst (snd (p', ?i, a)) = ?i" by simp ultimately have p': "(p', ?i, a) = p" by (rule unique_idxD1) note assms(2) moreover from * have "(h', ?k, b) \ set ?xs" by (rule 4) moreover from \h \ set hs\ have "h \ set ?xs" by simp moreover have "fst (snd (h', ?k, b)) = ?k" by simp ultimately have h': "(h', ?k, b) = h" by (rule unique_idxD1) from * show "(p, h) \\<^sub>p snd ` set ps" by (simp only: p' h') next from 2 obtain q' h' a b where *: "((q', ?j, a), (h', ?k, b)) \\<^sub>p snd ` set ps" by (rule pair_in_listE) note assms(2) moreover from * have "(q', ?j, a) \ set ?xs" by (rule 3) moreover have "q \ set ?xs" by simp moreover have "fst (snd (q', ?j, a)) = ?j" by simp ultimately have q': "(q', ?j, a) = q" by (rule unique_idxD1) note assms(2) moreover from * have "(h', ?k, b) \ set ?xs" by (rule 4) moreover from \h \ set hs\ have "h \ set ?xs" by simp moreover have "fst (snd (h', ?k, b)) = ?k" by simp ultimately have h': "(h', ?k, b) = h" by (rule unique_idxD1) from * show "(q, h) \\<^sub>p snd ` set ps" by (simp only: q' h') qed qed lemma ncrit_spec_chain_ncrit: "ncrit_spec (chain_ncrit::('t, 'b::field, 'c, 'd) ncritT)" proof (rule ncrit_specI) fix d m and data::"nat \ 'd" and gs bs hs and ps::"(bool \ ('t, 'b, 'c) pdata_pair) list" and B q_in_bs and p q::"('t, 'b, 'c) pdata" assume dg: "dickson_grading d" and B_sup: "set gs \ set bs \ set hs \ B" and B_sub: "fst ` B \ dgrad_p_set d m" and q_in_bs: "q_in_bs \ q \ set gs \ set bs" and 1: "\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and 2: "\p' q'. p' \ set gs \ set bs \ q' \ set gs \ set bs \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and "fst p \ 0" and "fst q \ 0" let ?l = "term_of_pair (lcs (lp (fst p)) (lp (fst q)), component_of_term (lt (fst p)))" assume "chain_ncrit data gs bs hs q_in_bs ps p q" and "snd ` set ps \ set hs \ (set gs \ set bs \ set hs)" and "unique_idx (gs @ bs @ hs) data" and "p \ set hs" and "q \ set gs \ set bs \ set hs" then obtain r where "r \ set gs \ set bs \ set hs" and "fst r \ 0" and "r \ p" and "r \ q" and adds: "lt (fst r) adds\<^sub>t ?l" and "(p, r) \\<^sub>p snd ` set ps" and disj: "(r \ set gs \ set bs \ q_in_bs) \ (q, r) \\<^sub>p snd ` set ps" by (rule chain_ncritE) note dg B_sub moreover from \p \ set hs\ \q \ set gs \ set bs \ set hs\ B_sup have "fst p \ fst ` B" and "fst q \ fst ` B" by auto moreover note \fst p \ 0\ \fst q \ 0\ moreover from adds have "lp (fst r) adds lcs (lp (fst p)) (lp (fst q))" by (simp add: adds_term_def term_simps) moreover from adds have "component_of_term (lt (fst r)) = component_of_term (lt (fst p))" by (simp add: adds_term_def term_simps) ultimately show "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" proof (rule chain_criterion) from \(p, r) \\<^sub>p snd ` set ps\ \fst p \ 0\ \fst r \ 0\ show "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst r)" by (rule 1) next from disj show "crit_pair_cbelow_on d m (fst ` B) (fst r) (fst q)" proof assume "r \ set gs \ set bs \ q_in_bs" hence "r \ set gs \ set bs" and q_in_bs by simp_all from q_in_bs this(2) have "q \ set gs \ set bs" .. with \r \ set gs \ set bs\ show ?thesis using \fst r \ 0\ \fst q \ 0\ by (rule 2) next assume "(q, r) \\<^sub>p snd ` set ps" hence "(r, q) \\<^sub>p snd ` set ps" by (simp only: in_pair_iff disj_commute) thus ?thesis using \fst r \ 0\ \fst q \ 0\ by (rule 1) qed qed qed lemma ocrit_spec_chain_ocrit: "ocrit_spec (chain_ocrit::('t, 'b::field, 'c, 'd) ocritT)" proof (rule ocrit_specI) fix d m and data::"nat \ 'd" and hs::"('t, 'b, 'c) pdata list" and ps::"(bool \ ('t, 'b, 'c) pdata_pair) list" and B and p q::"('t, 'b, 'c) pdata" assume dg: "dickson_grading d" and B_sup: "set hs \ B" and B_sub: "fst ` B \ dgrad_p_set d m" and 1: "\p' q'. (p', q') \\<^sub>p snd ` set ps \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and "fst p \ 0" and "fst q \ 0" and "p \ B" and "q \ B" let ?l = "term_of_pair (lcs (lp (fst p)) (lp (fst q)), component_of_term (lt (fst p)))" assume "chain_ocrit data hs ps p q" and "unique_idx (p # q # hs @ map (fst \ snd) ps @ map (snd \ snd) ps) data" then obtain h where "h \ set hs" and "fst h \ 0" and "h \ p" and "h \ q" and adds: "lt (fst h) adds\<^sub>t ?l" and "(p, h) \\<^sub>p snd ` set ps" and "(q, h) \\<^sub>p snd ` set ps" by (rule chain_ocritE) note dg B_sub moreover from \p \ B\ \q \ B\ B_sup have "fst p \ fst ` B" and "fst q \ fst ` B" by auto moreover note \fst p \ 0\ \fst q \ 0\ moreover from adds have "lp (fst h) adds lcs (lp (fst p)) (lp (fst q))" by (simp add: adds_term_def term_simps) moreover from adds have "component_of_term (lt (fst h)) = component_of_term (lt (fst p))" by (simp add: adds_term_def term_simps) ultimately show "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" proof (rule chain_criterion) from \(p, h) \\<^sub>p snd ` set ps\ \fst p \ 0\ \fst h \ 0\ show "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst h)" by (rule 1) next from \(q, h) \\<^sub>p snd ` set ps\ have "(h, q) \\<^sub>p snd ` set ps" by (simp only: in_pair_iff disj_commute) thus "crit_pair_cbelow_on d m (fst ` B) (fst h) (fst q)" using \fst h \ 0\ \fst q \ 0\ by (rule 1) qed qed lemma icrit_spec_no_crit: "icrit_spec ((\_ _ _ _ _ _. False)::('t, 'b::field, 'c, 'd) icritT)" by (rule icrit_specI, simp) lemma ncrit_spec_no_crit: "ncrit_spec ((\_ _ _ _ _ _ _ _. False)::('t, 'b::field, 'c, 'd) ncritT)" by (rule ncrit_specI, simp) lemma ocrit_spec_no_crit: "ocrit_spec ((\_ _ _ _ _. False)::('t, 'b::field, 'c, 'd) ocritT)" by (rule ocrit_specI, simp) subsubsection \Creating Initial List of New Pairs\ type_synonym (in -) ('t, 'b, 'c) apsT = "bool \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata \ (bool \ ('t, 'b, 'c) pdata_pair) list \ (bool \ ('t, 'b, 'c) pdata_pair) list" type_synonym (in -) ('t, 'b, 'c, 'd) npT = "('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ nat \ 'd \ (bool \ ('t, 'b, 'c) pdata_pair) list" definition np_spec :: "('t, 'b, 'c, 'd) npT \ bool" where "np_spec np \ (\gs bs hs data. snd ` set (np gs bs hs data) \ set hs \ (set gs \ set bs \ set hs) \ set hs \ (set gs \ set bs) \ snd ` set (np gs bs hs data) \ (\a b. a \ set hs \ b \ set hs \ a \ b \ (a, b) \\<^sub>p snd ` set (np gs bs hs data)) \ (\p q. (True, p, q) \ set (np gs bs hs data) \ q \ set gs \ set bs))" lemma np_specI: assumes "\gs bs hs data. snd ` set (np gs bs hs data) \ set hs \ (set gs \ set bs \ set hs) \ set hs \ (set gs \ set bs) \ snd ` set (np gs bs hs data) \ (\a b. a \ set hs \ b \ set hs \ a \ b \ (a, b) \\<^sub>p snd ` set (np gs bs hs data)) \ (\p q. (True, p, q) \ set (np gs bs hs data) \ q \ set gs \ set bs)" shows "np_spec np" unfolding np_spec_def using assms by meson lemma np_specD1: assumes "np_spec np" shows "snd ` set (np gs bs hs data) \ set hs \ (set gs \ set bs \ set hs)" using assms[unfolded np_spec_def, rule_format, of gs bs hs data] .. lemma np_specD2: assumes "np_spec np" shows "set hs \ (set gs \ set bs) \ snd ` set (np gs bs hs data)" using assms[unfolded np_spec_def, rule_format, of gs bs hs data] by auto lemma np_specD3: assumes "np_spec np" and "a \ set hs" and "b \ set hs" and "a \ b" shows "(a, b) \\<^sub>p snd ` set (np gs bs hs data)" using assms(1)[unfolded np_spec_def, rule_format, of gs bs hs data] assms(2,3,4) by blast lemma np_specD4: assumes "np_spec np" and "(True, p, q) \ set (np gs bs hs data)" shows "q \ set gs \ set bs" using assms(1)[unfolded np_spec_def, rule_format, of gs bs hs data] assms(2) by blast lemma np_specE: assumes "np_spec np" and "p \ set hs" and "q \ set gs \ set bs \ set hs" and "p \ q" assumes 1: "\q_in_bs. (q_in_bs, p, q) \ set (np gs bs hs data) \ thesis" assumes 2: "\p_in_bs. (p_in_bs, q, p) \ set (np gs bs hs data) \ thesis" shows thesis proof (cases "q \ set gs \ set bs") case True with assms(2) have "(p, q) \ set hs \ (set gs \ set bs)" by simp also from assms(1) have "... \ snd ` set (np gs bs hs data)" by (rule np_specD2) finally obtain q_in_bs where "(q_in_bs, p, q) \ set (np gs bs hs data)" by fastforce thus ?thesis by (rule 1) next case False with assms(3) have "q \ set hs" by simp from assms(1,2) this assms(4) have "(p, q) \\<^sub>p snd ` set (np gs bs hs data)" by (rule np_specD3) hence "(p, q) \ snd ` set (np gs bs hs data) \ (q, p) \ snd ` set (np gs bs hs data)" by (simp only: in_pair_iff) thus ?thesis proof assume "(p, q) \ snd ` set (np gs bs hs data)" then obtain q_in_bs where "(q_in_bs, p, q) \ set (np gs bs hs data)" by fastforce thus ?thesis by (rule 1) next assume "(q, p) \ snd ` set (np gs bs hs data)" then obtain p_in_bs where "(p_in_bs, q, p) \ set (np gs bs hs data)" by fastforce thus ?thesis by (rule 2) qed qed definition add_pairs_single_naive :: "'d \ ('t, 'b::zero, 'c) apsT" where "add_pairs_single_naive data flag gs bs h ps = ps @ (map (\g. (flag, h, g)) gs) @ (map (\b. (flag, h, b)) bs)" lemma set_add_pairs_single_naive: "set (add_pairs_single_naive data flag gs bs h ps) = set ps \ Pair flag ` ({h} \ (set gs \ set bs))" by (auto simp add: add_pairs_single_naive_def Let_def) fun add_pairs_single_sorted :: "((bool \ ('t, 'b, 'c) pdata_pair) \ (bool \ ('t, 'b, 'c) pdata_pair) \ bool) \ ('t, 'b::zero, 'c) apsT" where "add_pairs_single_sorted _ _ [] [] _ ps = ps"| "add_pairs_single_sorted rel flag [] (b # bs) h ps = add_pairs_single_sorted rel flag [] bs h (insort_wrt rel (flag, h, b) ps)"| "add_pairs_single_sorted rel flag (g # gs) bs h ps = add_pairs_single_sorted rel flag gs bs h (insort_wrt rel (flag, h, g) ps)" lemma set_add_pairs_single_sorted: "set (add_pairs_single_sorted rel flag gs bs h ps) = set ps \ Pair flag ` ({h} \ (set gs \ set bs))" proof (induct gs arbitrary: ps) case Nil show ?case proof (induct bs arbitrary: ps) case Nil show ?case by simp next case (Cons b bs) show ?case by (simp add: Cons) qed next case (Cons g gs) show ?case by (simp add: Cons) qed primrec (in -) pairs :: "('t, 'b, 'c) apsT \ bool \ ('t, 'b, 'c) pdata list \ (bool \ ('t, 'b, 'c) pdata_pair) list" where "pairs _ _ [] = []"| "pairs aps flag (x # xs) = aps flag [] xs x (pairs aps flag xs)" lemma pairs_subset: assumes "\gs bs h ps. set (aps flag gs bs h ps) = set ps \ Pair flag ` ({h} \ (set gs \ set bs))" shows "set (pairs aps flag xs) \ Pair flag ` (set xs \ set xs)" proof (induct xs) case Nil show ?case by simp next case (Cons x xs) from Cons have "set (pairs aps flag xs) \ Pair flag ` (set (x # xs) \ set (x # xs))" by fastforce moreover have "{x} \ set xs \ set (x # xs) \ set (x # xs)" by fastforce ultimately show ?case by (auto simp add: assms) qed lemma in_pairsI: assumes "\gs bs h ps. set (aps flag gs bs h ps) = set ps \ Pair flag ` ({h} \ (set gs \ set bs))" and "a \ b" and "a \ set xs" and "b \ set xs" shows "(flag, a, b) \ set (pairs aps flag xs) \ (flag, b, a) \ set (pairs aps flag xs)" using assms(3, 4) proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) from Cons(3) have d: "b = x \ b \ set xs" by simp from Cons(2) have "a = x \ a \ set xs" by simp thus ?case proof assume "a = x" with assms(2) have "b \ x" by simp with d have "b \ set xs" by simp hence "(flag, a, b) \ set (pairs aps flag (x # xs))" by (simp add: \a = x\ assms(1)) thus ?thesis by simp next assume "a \ set xs" from d show ?thesis proof assume "b = x" from \a \ set xs\ have "(flag, b, a) \ set (pairs aps flag (x # xs))" by (simp add: \b = x\ assms(1)) thus ?thesis by simp next assume "b \ set xs" with \a \ set xs\ have "(flag, a, b) \ set (pairs aps flag xs) \ (flag, b, a) \ set (pairs aps flag xs)" by (rule Cons(1)) thus ?thesis by (auto simp: assms(1)) qed qed qed corollary in_pairsI': assumes "\gs bs h ps. set (aps flag gs bs h ps) = set ps \ Pair flag ` ({h} \ (set gs \ set bs))" and "a \ set xs" and "b \ set xs" and "a \ b" shows "(a, b) \\<^sub>p snd ` set (pairs aps flag xs)" proof - from assms(1,4,2,3) have "(flag, a, b) \ set (pairs aps flag xs) \ (flag, b, a) \ set (pairs aps flag xs)" by (rule in_pairsI) thus ?thesis proof assume "(flag, a, b) \ set (pairs aps flag xs)" hence "snd (flag, a, b) \ snd ` set (pairs aps flag xs)" by fastforce thus ?thesis by (simp add: in_pair_iff) next assume "(flag, b, a) \ set (pairs aps flag xs)" hence "snd (flag, b, a) \ snd ` set (pairs aps flag xs)" by fastforce thus ?thesis by (simp add: in_pair_iff) qed qed definition new_pairs_naive :: "('t, 'b::zero, 'c, 'd) npT" where "new_pairs_naive gs bs hs data = fold (add_pairs_single_naive data True gs bs) hs (pairs (add_pairs_single_naive data) False hs)" definition new_pairs_sorted :: "(nat \ 'd \ (bool \ ('t, 'b, 'c) pdata_pair) \ (bool \ ('t, 'b, 'c) pdata_pair) \ bool) \ ('t, 'b::zero, 'c, 'd) npT" where "new_pairs_sorted rel gs bs hs data = fold (add_pairs_single_sorted (rel data) True gs bs) hs (pairs (add_pairs_single_sorted (rel data)) False hs)" lemma set_fold_aps: assumes "\gs bs h ps. set (aps flag gs bs h ps) = set ps \ Pair flag ` ({h} \ (set gs \ set bs))" shows "set (fold (aps flag gs bs) hs ps) = Pair flag ` (set hs \ (set gs \ set bs)) \ set ps" proof (induct hs arbitrary: ps) case Nil show ?case by simp next case (Cons h hs) show ?case by (auto simp add: Cons assms) qed lemma set_new_pairs_naive: "set (new_pairs_naive gs bs hs data) = Pair True ` (set hs \ (set gs \ set bs)) \ set (pairs (add_pairs_single_naive data) False hs)" proof - have "set (new_pairs_naive gs bs hs data) = Pair True ` (set hs \ (set gs \ set bs)) \ set (pairs (add_pairs_single_naive data) False hs)" unfolding new_pairs_naive_def by (rule set_fold_aps, fact set_add_pairs_single_naive) thus ?thesis by (simp add: ac_simps) qed lemma set_new_pairs_sorted: "set (new_pairs_sorted rel gs bs hs data) = Pair True ` (set hs \ (set gs \ set bs)) \ set (pairs (add_pairs_single_sorted (rel data)) False hs)" proof - have "set (new_pairs_sorted rel gs bs hs data) = Pair True ` (set hs \ (set gs \ set bs)) \ set (pairs (add_pairs_single_sorted (rel data)) False hs)" unfolding new_pairs_sorted_def by (rule set_fold_aps, fact set_add_pairs_single_sorted) thus ?thesis by (simp add: set_merge_wrt ac_simps) qed lemma (in -) fst_snd_Pair [simp]: shows "fst \ Pair x = (\_. x)" and "snd \ Pair x = id" by auto lemma np_spec_new_pairs_naive: "np_spec new_pairs_naive" proof (rule np_specI) fix gs bs hs :: "('t, 'b, 'c) pdata list" and data::"nat \ 'd" have 1: "set hs \ (set gs \ set bs) \ set hs \ (set gs \ set bs \ set hs)" by fastforce have "set (pairs (add_pairs_single_naive data) False hs) \ Pair False ` (set hs \ set hs)" by (rule pairs_subset, simp add: set_add_pairs_single_naive) hence "snd ` set (pairs (add_pairs_single_naive data) False hs) \ snd ` Pair False ` (set hs \ set hs)" by (rule image_mono) also have "... = set hs \ set hs" by (simp add: image_comp) finally have 2: "snd ` set (pairs (add_pairs_single_naive data) False hs) \ set hs \ (set gs \ set bs \ set hs)" by fastforce show "snd ` set (new_pairs_naive gs bs hs data) \ set hs \ (set gs \ set bs \ set hs) \ set hs \ (set gs \ set bs) \ snd ` set (new_pairs_naive gs bs hs data) \ (\a b. a \ set hs \ b \ set hs \ a \ b \ (a, b) \\<^sub>p snd ` set (new_pairs_naive gs bs hs data)) \ (\p q. (True, p, q) \ set (new_pairs_naive gs bs hs data) \ q \ set gs \ set bs)" proof (intro conjI allI impI) show "snd ` set (new_pairs_naive gs bs hs data) \ set hs \ (set gs \ set bs \ set hs)" by (simp add: set_new_pairs_naive image_Un image_comp 1 2) next show "set hs \ (set gs \ set bs) \ snd ` set (new_pairs_naive gs bs hs data)" by (simp add: set_new_pairs_naive image_Un image_comp) next fix a b assume "a \ set hs" and "b \ set hs" and "a \ b" with set_add_pairs_single_naive have "(a, b) \\<^sub>p snd ` set (pairs (add_pairs_single_naive data) False hs)" by (rule in_pairsI') thus "(a, b) \\<^sub>p snd ` set (new_pairs_naive gs bs hs data)" by (simp add: set_new_pairs_naive image_Un) next fix p q assume "(True, p, q) \ set (new_pairs_naive gs bs hs data)" hence "q \ set gs \ set bs \ (True, p, q) \ set (pairs (add_pairs_single_naive data) False hs)" by (auto simp: set_new_pairs_naive) thus "q \ set gs \ set bs" proof assume "(True, p, q) \ set (pairs (add_pairs_single_naive data) False hs)" also from set_add_pairs_single_naive have "... \ Pair False ` (set hs \ set hs)" by (rule pairs_subset) finally show ?thesis by auto qed qed qed lemma np_spec_new_pairs_sorted: "np_spec (new_pairs_sorted rel)" proof (rule np_specI) fix gs bs hs :: "('t, 'b, 'c) pdata list" and data::"nat \ 'd" have 1: "set hs \ (set gs \ set bs) \ set hs \ (set gs \ set bs \ set hs)" by fastforce have "set (pairs (add_pairs_single_sorted (rel data)) False hs) \ Pair False ` (set hs \ set hs)" by (rule pairs_subset, simp add: set_add_pairs_single_sorted) hence "snd ` set (pairs (add_pairs_single_sorted (rel data)) False hs) \ snd ` Pair False ` (set hs \ set hs)" by (rule image_mono) also have "... = set hs \ set hs" by (simp add: image_comp) finally have 2: "snd ` set (pairs (add_pairs_single_sorted (rel data)) False hs) \ set hs \ (set gs \ set bs \ set hs)" by fastforce show "snd ` set (new_pairs_sorted rel gs bs hs data) \ set hs \ (set gs \ set bs \ set hs) \ set hs \ (set gs \ set bs) \ snd ` set (new_pairs_sorted rel gs bs hs data) \ (\a b. a \ set hs \ b \ set hs \ a \ b \ (a, b) \\<^sub>p snd ` set (new_pairs_sorted rel gs bs hs data)) \ (\p q. (True, p, q) \ set (new_pairs_sorted rel gs bs hs data) \ q \ set gs \ set bs)" proof (intro conjI allI impI) show "snd ` set (new_pairs_sorted rel gs bs hs data) \ set hs \ (set gs \ set bs \ set hs)" by (simp add: set_new_pairs_sorted image_Un image_comp 1 2) next show "set hs \ (set gs \ set bs) \ snd ` set (new_pairs_sorted rel gs bs hs data)" by (simp add: set_new_pairs_sorted image_Un image_comp) next fix a b assume "a \ set hs" and "b \ set hs" and "a \ b" with set_add_pairs_single_sorted have "(a, b) \\<^sub>p snd ` set (pairs (add_pairs_single_sorted (rel data)) False hs)" by (rule in_pairsI') thus "(a, b) \\<^sub>p snd ` set (new_pairs_sorted rel gs bs hs data)" by (simp add: set_new_pairs_sorted image_Un) next fix p q assume "(True, p, q) \ set (new_pairs_sorted rel gs bs hs data)" hence "q \ set gs \ set bs \ (True, p, q) \ set (pairs (add_pairs_single_sorted (rel data)) False hs)" by (auto simp: set_new_pairs_sorted) thus "q \ set gs \ set bs" proof assume "(True, p, q) \ set (pairs (add_pairs_single_sorted (rel data)) False hs)" also from set_add_pairs_single_sorted have "... \ Pair False ` (set hs \ set hs)" by (rule pairs_subset) finally show ?thesis by auto qed qed qed text \@{term "new_pairs_naive gs bs hs data"} and @{term "new_pairs_sorted rel gs bs hs data"} return lists of triples @{term "(q_in_bs, p, q)"}, where \q_in_bs\ indicates whether \q\ is contained in the list @{term "gs @ bs"} or in the list \hs\. \p\ is always contained in \hs\.\ definition canon_pair_order_aux :: "('t, 'b::zero, 'c) pdata_pair \ ('t, 'b, 'c) pdata_pair \ bool" where "canon_pair_order_aux p q \ (lcs (lp (fst (fst p))) (lp (fst (snd p))) \ lcs (lp (fst (fst q))) (lp (fst (snd q))))" abbreviation "canon_pair_order data p q \ canon_pair_order_aux (snd p) (snd q)" abbreviation "canon_pair_comb \ merge_wrt canon_pair_order_aux" subsubsection \Applying Criteria to New Pairs\ definition apply_icrit :: "('t, 'b, 'c, 'd) icritT \ (nat \ 'd) \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ (bool \ ('t, 'b, 'c) pdata_pair) list \ (bool \ bool \ ('t, 'b, 'c) pdata_pair) list" where "apply_icrit crit data gs bs hs ps = (let c = crit data gs bs hs in map (\(q_in_bs, p, q). (c p q, q_in_bs, p, q)) ps)" lemma fst_apply_icrit: assumes "icrit_spec crit" and "dickson_grading d" and "fst ` (set gs \ set bs \ set hs) \ dgrad_p_set d m" and "unique_idx (gs @ bs @ hs) data" and "is_Groebner_basis (fst ` set gs)" and "p \ set hs" and "q \ set gs \ set bs \ set hs" and "fst p \ 0" and "fst q \ 0" and "(True, q_in_bs, p, q) \ set (apply_icrit crit data gs bs hs ps)" shows "crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q)" proof - from assms(10) have "crit data gs bs hs p q" by (auto simp: apply_icrit_def) with assms(1-9) show ?thesis by (rule icrit_specD) qed lemma snd_apply_icrit [simp]: "map snd (apply_icrit crit data gs bs hs ps) = ps" by (auto simp add: apply_icrit_def case_prod_beta' intro: nth_equalityI) lemma set_snd_apply_icrit [simp]: "snd ` set (apply_icrit crit data gs bs hs ps) = set ps" proof - have "snd ` set (apply_icrit crit data gs bs hs ps) = set (map snd (apply_icrit crit data gs bs hs ps))" by (simp del: snd_apply_icrit) also have "... = set ps" by (simp only: snd_apply_icrit) finally show ?thesis . qed definition apply_ncrit :: "('t, 'b, 'c, 'd) ncritT \ (nat \ 'd) \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ ('t, 'b, 'c) pdata list \ (bool \ bool \ ('t, 'b, 'c) pdata_pair) list \ (bool \ ('t, 'b, 'c) pdata_pair) list" where "apply_ncrit crit data gs bs hs ps = (let c = crit data gs bs hs in rev (fold (\(ic, q_in_bs, p, q). \ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps []))" lemma apply_ncrit_append: "apply_ncrit crit data gs bs hs (xs @ ys) = rev (fold (\(ic, q_in_bs, p, q). \ps'. if \ ic \ crit data gs bs hs q_in_bs ps' p q then ps' else (ic, p, q) # ps') ys (rev (apply_ncrit crit data gs bs hs xs)))" by (simp add: apply_ncrit_def Let_def) lemma fold_superset: "set acc \ set (fold (\(ic, q_in_bs, p, q). \ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps acc)" proof (induct ps arbitrary: acc) case Nil show ?case by simp next case (Cons x ps) obtain ic' q_in_bs' p' q' where x: "x = (ic', q_in_bs', p', q')" using prod_cases4 by blast have 1: "set acc0 \ set (fold (\(ic, q_in_bs, p, q) ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps acc0)" for acc0 by (rule Cons) have "set acc \ set ((ic', p', q') # acc)" by fastforce also have "... \ set (fold (\(ic, q_in_bs, p, q) ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps ((ic', p', q') # acc))" by (fact 1) finally have 2: "set acc \ set (fold (\(ic, q_in_bs, p, q) ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps ((ic', p', q') # acc))" . show ?case by (simp add: x 1 2) qed lemma apply_ncrit_superset: "set (apply_ncrit crit data gs bs hs ps) \ set (apply_ncrit crit data gs bs hs (ps @ qs))" (is "?l \ ?r") proof - have "?l = set (rev (apply_ncrit crit data gs bs hs ps))" by simp also have "... \ set (fold (\(ic, q_in_bs, p, q) ps'. if \ ic \ crit data gs bs hs q_in_bs ps' p q then ps' else (ic, p, q) # ps') qs (rev (apply_ncrit crit data gs bs hs ps)))" by (fact fold_superset) also have "... = ?r" by (simp add: apply_ncrit_append) finally show ?thesis . qed lemma apply_ncrit_subset_aux: assumes "(ic, p, q) \ set (fold (\(ic, q_in_bs, p, q). \ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps acc)" shows "(ic, p, q) \ set acc \ (\q_in_bs. (ic, q_in_bs, p, q) \ set ps)" using assms proof (induct ps arbitrary: acc) case Nil thus ?case by simp next case (Cons x ps) obtain ic' q_in_bs' p' q' where x: "x = (ic', q_in_bs', p', q')" using prod_cases4 by blast from Cons(2) have "(ic, p, q) \ set (fold (\(ic, q_in_bs, p, q) ps'. if \ ic \ c q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps (if \ ic' \ c q_in_bs' acc p' q' then acc else (ic', p', q') # acc))" by (simp add: x) hence "(ic, p, q) \ set (if \ ic' \ c q_in_bs' acc p' q' then acc else (ic', p', q') # acc) \ (\q_in_bs. (ic, q_in_bs, p, q) \ set ps)" by (rule Cons(1)) hence "(ic, p, q) \ set acc \ (ic, p, q) = (ic', p', q') \ (\q_in_bs. (ic, q_in_bs, p, q) \ set ps)" by (auto split: if_splits) thus ?case proof (elim disjE) assume "(ic, p, q) \ set acc" thus ?thesis .. next assume "(ic, p, q) = (ic', p', q')" hence "x = (ic, q_in_bs', p, q)" by (simp add: x) thus ?thesis by auto next assume "\q_in_bs. (ic, q_in_bs, p, q) \ set ps" then obtain q_in_bs where "(ic, q_in_bs, p, q) \ set ps" .. thus ?thesis by auto qed qed corollary apply_ncrit_subset: assumes "(ic, p, q) \ set (apply_ncrit crit data gs bs hs ps)" obtains q_in_bs where "(ic, q_in_bs, p, q) \ set ps" proof - from assms have "(ic, p, q) \ set (fold (\(ic, q_in_bs, p, q). \ps'. if \ ic \ crit data gs bs hs q_in_bs ps' p q then ps' else (ic, p, q) # ps') ps [])" by (simp add: apply_ncrit_def) hence "(ic, p, q) \ set [] \ (\q_in_bs. (ic, q_in_bs, p, q) \ set ps)" by (rule apply_ncrit_subset_aux) hence "\q_in_bs. (ic, q_in_bs, p, q) \ set ps" by simp then obtain q_in_bs where "(ic, q_in_bs, p, q) \ set ps" .. thus ?thesis .. qed corollary apply_ncrit_subset': "snd ` set (apply_ncrit crit data gs bs hs ps) \ snd ` snd ` set ps" proof fix p q assume "(p, q) \ snd ` set (apply_ncrit crit data gs bs hs ps)" then obtain ic where "(ic, p, q) \ set (apply_ncrit crit data gs bs hs ps)" by fastforce then obtain q_in_bs where "(ic, q_in_bs, p, q) \ set ps" by (rule apply_ncrit_subset) thus "(p, q) \ snd ` snd ` set ps" by force qed lemma not_in_apply_ncrit: assumes "(ic, p, q) \ set (apply_ncrit crit data gs bs hs (xs @ ((ic, q_in_bs, p, q) # ys)))" shows "crit data gs bs hs q_in_bs (rev (apply_ncrit crit data gs bs hs xs)) p q" using assms proof (simp add: apply_ncrit_append split: if_splits) assume "(ic, p, q) \ set (fold (\(ic, q_in_bs, p, q) ps'. if \ ic \ crit data gs bs hs q_in_bs ps' p q then ps' else (ic, p, q) # ps') ys ((ic, p, q) # rev (apply_ncrit crit data gs bs hs xs)))" (is "_ \ ?A") have "(ic, p, q) \ set ((ic, p, q) # rev (apply_ncrit crit data gs bs hs xs))" by simp also have "... \ ?A" by (rule fold_superset) finally have "(ic, p, q) \ ?A" . with \(ic, p, q) \ ?A\ show ?thesis .. qed lemma (in -) setE: assumes "x \ set xs" obtains ys zs where "xs = ys @ (x # zs)" using assms proof (induct xs arbitrary: thesis) case Nil from Nil(2) show ?case by simp next case (Cons a xs) from Cons(3) have "x = a \ x \ set xs" by simp thus ?case proof assume "x = a" show ?thesis by (rule Cons(2)[of "[]" xs], simp add: \x = a\) next assume "x \ set xs" then obtain ys zs where "xs = ys @ (x # zs)" by (meson Cons(1)) show ?thesis by (rule Cons(2)[of "a # ys" zs], simp add: \xs = ys @ (x # zs)\) qed qed lemma apply_ncrit_connectible: assumes "ncrit_spec crit" and "dickson_grading d" and "set gs \ set bs \ set hs \ B" and "fst ` B \ dgrad_p_set d m" and "snd ` snd ` set ps \ set hs \ (set gs \ set bs \ set hs)" and "unique_idx (gs @ bs @ hs) data" and "is_Groebner_basis (fst ` set gs)" and "\p' q'. (p', q') \ snd ` set (apply_ncrit crit data gs bs hs ps) \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" and "\p' q'. p' \ set gs \ set bs \ q' \ set gs \ set bs \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" assumes "(ic, q_in_bs, p, q) \ set ps" and "fst p \ 0" and "fst q \ 0" and "q_in_bs \ (q \ set gs \ set bs)" shows "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" proof (cases "(p, q) \ snd ` set (apply_ncrit crit data gs bs hs ps)") case True thus ?thesis using assms(11,12) by (rule assms(8)) next case False from assms(10) have "(p, q) \ snd ` snd ` set ps" by force also have "... \ set hs \ (set gs \ set bs \ set hs)" by (fact assms(5)) finally have "p \ set hs" and "q \ set gs \ set bs \ set hs" by simp_all from \(ic, q_in_bs, p, q) \ set ps\ obtain xs ys where ps: "ps = xs @ ((ic, q_in_bs, p, q) # ys)" by (rule setE) let ?ps = "rev (apply_ncrit crit data gs bs hs xs)" have "snd ` set ?ps \ snd ` snd ` set xs" by (simp add: apply_ncrit_subset') also have "... \ snd ` snd ` set ps" unfolding ps by fastforce finally have sub: "snd ` set ?ps \ set hs \ (set gs \ set bs \ set hs)" using assms(5) by (rule subset_trans) from False have "(p, q) \ snd ` set (apply_ncrit crit data gs bs hs ps)" by (simp add: in_pair_iff) hence "(ic, p, q) \ set (apply_ncrit crit data gs bs hs (xs @ ((ic, q_in_bs, p, q) # ys)))" unfolding ps by force hence "crit data gs bs hs q_in_bs ?ps p q" by (rule not_in_apply_ncrit) with assms(1-4) sub assms(6,7,13) _ _ \p \ set hs\ \q \ set gs \ set bs \ set hs\ assms(11,12) show ?thesis proof (rule ncrit_specD) fix p' q' assume "(p', q') \\<^sub>p snd ` set ?ps" also have "... \ snd ` set (apply_ncrit crit data gs bs hs ps)" by (rule image_mono, simp add: ps apply_ncrit_superset) finally have disj: "(p', q') \ snd ` set (apply_ncrit crit data gs bs hs ps) \ (q', p') \ snd ` set (apply_ncrit crit data gs bs hs ps)" by (simp only: in_pair_iff) assume "fst p' \ 0" and "fst q' \ 0" from disj show "crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" proof assume "(p', q') \ snd ` set (apply_ncrit crit data gs bs hs ps)" thus ?thesis using \fst p' \ 0\ \fst q' \ 0\ by (rule assms(8)) next assume "(q', p') \ snd ` set (apply_ncrit crit data gs bs hs ps)" hence "crit_pair_cbelow_on d m (fst ` B) (fst q') (fst p')" using \fst q' \ 0\ \fst p' \ 0\ by (rule assms(8)) thus ?thesis by (rule crit_pair_cbelow_sym) qed qed (assumption, fact assms(9)) qed subsubsection \Applying Criteria to Old Pairs\ definition apply_ocrit :: "('t, 'b, 'c, 'd) ocritT \ (nat \ 'd) \ ('t, 'b, 'c) pdata list \ (bool \ ('t, 'b, 'c) pdata_pair) list \ ('t, 'b, 'c) pdata_pair list \ ('t, 'b, 'c) pdata_pair list" where "apply_ocrit crit data hs ps' ps = (let c = crit data hs ps' in [(p, q)\ps . \ c p q])" lemma set_apply_ocrit: "set (apply_ocrit crit data hs ps' ps) = {(p, q) | p q. (p, q) \ set ps \ \ crit data hs ps' p q}" by (auto simp: apply_ocrit_def) corollary set_apply_ocrit_iff: "(p, q) \ set (apply_ocrit crit data hs ps' ps) \ ((p, q) \ set ps \ \ crit data hs ps' p q)" by (auto simp: apply_ocrit_def) lemma apply_ocrit_connectible: assumes "ocrit_spec crit" and "dickson_grading d" and "set hs \ B" and "fst ` B \ dgrad_p_set d m" and "unique_idx (p # q # hs @ (map (fst \ snd) ps') @ (map (snd \ snd) ps')) data" and "\p' q'. (p', q') \ snd ` set ps' \ fst p' \ 0 \ fst q' \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" assumes "p \ B" and "q \ B" and "fst p \ 0" and "fst q \ 0" and "(p, q) \ set ps" and "(p, q) \ set (apply_ocrit crit data hs ps' ps)" shows "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" proof - from assms(11,12) have "crit data hs ps' p q" by (simp add: set_apply_ocrit_iff) with assms(1-5) _ assms(7-10) show ?thesis proof (rule ocrit_specD) fix p' q' assume "(p', q') \\<^sub>p snd ` set ps'" hence disj: "(p', q') \ snd ` set ps' \ (q', p') \ snd ` set ps'" by (simp only: in_pair_iff) assume "fst p' \ 0" and "fst q' \ 0" from disj show "crit_pair_cbelow_on d m (fst ` B) (fst p') (fst q')" proof assume "(p', q') \ snd ` set ps'" thus ?thesis using \fst p' \ 0\ \fst q' \ 0\ by (rule assms(6)) next assume "(q', p') \ snd ` set ps'" hence "crit_pair_cbelow_on d m (fst ` B) (fst q') (fst p')" using \fst q' \ 0\ \fst p' \ 0\ by (rule assms(6)) thus ?thesis by (rule crit_pair_cbelow_sym) qed qed qed subsubsection \Creating Final List of Pairs\ context fixes np::"('t, 'b::field, 'c, 'd) npT" and icrit::"('t, 'b, 'c, 'd) icritT" and ncrit::"('t, 'b, 'c, 'd) ncritT" and ocrit::"('t, 'b, 'c, 'd) ocritT" and comb::"('t, 'b, 'c) pdata_pair list \ ('t, 'b, 'c) pdata_pair list \ ('t, 'b, 'c) pdata_pair list" begin definition add_pairs :: "('t, 'b, 'c, 'd) apT" where "add_pairs gs bs ps hs data = (let ps1 = apply_ncrit ncrit data gs bs hs (apply_icrit icrit data gs bs hs (np gs bs hs data)); ps2 = apply_ocrit ocrit data hs ps1 ps in comb (map snd [x\ps1 . \ fst x]) ps2)" lemma set_add_pairs: assumes "\xs ys. set (comb xs ys) = set xs \ set ys" assumes "ps1 = apply_ncrit ncrit data gs bs hs (apply_icrit icrit data gs bs hs (np gs bs hs data))" shows "set (add_pairs gs bs ps hs data) = {(p, q) | p q. (False, p, q) \ set ps1 \ ((p, q) \ set ps \ \ ocrit data hs ps1 p q)}" proof - have eq: "snd ` {x \ set ps1. \ fst x} = {(p, q) | p q. (False, p, q) \ set ps1}" by force thus ?thesis by (auto simp: add_pairs_def Let_def assms(1) assms(2)[symmetric] set_apply_ocrit) qed lemma set_add_pairs_iff: assumes "\xs ys. set (comb xs ys) = set xs \ set ys" assumes "ps1 = apply_ncrit ncrit data gs bs hs (apply_icrit icrit data gs bs hs (np gs bs hs data))" shows "((p, q) \ set (add_pairs gs bs ps hs data)) \ ((False, p, q) \ set ps1 \ ((p, q) \ set ps \ \ ocrit data hs ps1 p q))" proof - from assms have eq: "set (add_pairs gs bs ps hs data) = {(p, q) | p q. (False, p, q) \ set ps1 \ ((p, q) \ set ps \ \ ocrit data hs ps1 p q)}" by (rule set_add_pairs) obtain a aa b where p: "p = (a, aa, b)" using prod_cases3 by blast obtain ab ac ba where q: "q = (ab, ac, ba)" using prod_cases3 by blast show ?thesis by (simp add: eq p q) qed lemma ap_spec_add_pairs: assumes "np_spec np" and "icrit_spec icrit" and "ncrit_spec ncrit" and "ocrit_spec ocrit" and "\xs ys. set (comb xs ys) = set xs \ set ys" shows "ap_spec add_pairs" proof (rule ap_specI) fix gs bs :: "('t, 'b, 'c) pdata list" and ps hs and data::"nat \ 'd" define ps1 where "ps1 = apply_ncrit ncrit data gs bs hs (apply_icrit icrit data gs bs hs (np gs bs hs data))" show "set (add_pairs gs bs ps hs data) \ set ps \ set hs \ (set gs \ set bs \ set hs)" proof fix p q assume "(p, q) \ set (add_pairs gs bs ps hs data)" with assms(5) ps1_def have "(False, p, q) \ set ps1 \ ((p, q) \ set ps \ \ ocrit data hs ps1 p q)" by (simp add: set_add_pairs_iff) thus "(p, q) \ set ps \ set hs \ (set gs \ set bs \ set hs)" proof assume "(False, p, q) \ set ps1" hence "snd (False, p, q) \ snd ` set ps1" by fastforce hence "(p, q) \ snd ` set ps1" by simp also have "... \ snd ` snd ` set (apply_icrit icrit data gs bs hs (np gs bs hs data))" unfolding ps1_def by (fact apply_ncrit_subset') also have "... = snd ` set (np gs bs hs data)" by simp also from assms(1) have "... \ set hs \ (set gs \ set bs \ set hs)" by (rule np_specD1) finally show ?thesis .. next assume "(p, q) \ set ps \ \ ocrit data hs ps1 p q" thus ?thesis by simp qed qed next fix gs bs :: "('t, 'b, 'c) pdata list" and ps hs and data::"nat \ 'd" and B and d::"'a \ nat" and m h g assume dg: "dickson_grading d" and B_sup: "set gs \ set bs \ set hs \ B" and B_sub: "fst ` B \ dgrad_p_set d m" and h_in: "h \ set hs" and g_in: "g \ set gs \ set bs \ set hs" and ps_sub: "set ps \ set bs \ (set gs \ set bs)" and uid: "unique_idx (gs @ bs @ hs) data" and gb: "is_Groebner_basis (fst ` set gs)" and "h \ g" and "fst h \ 0" and "fst g \ 0" assume a: "\a b. (a, b) \\<^sub>p set (add_pairs gs bs ps hs data) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" assume b: "\a b. a \ set gs \ set bs \ b \ set gs \ set bs \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" define ps0 where "ps0 = apply_icrit icrit data gs bs hs (np gs bs hs data)" define ps1 where "ps1 = apply_ncrit ncrit data gs bs hs ps0" have "snd ` snd ` set ps0 = snd ` set (np gs bs hs data)" by (simp add: ps0_def) also from assms(1) have "... \ set hs \ (set gs \ set bs \ set hs)" by (rule np_specD1) finally have ps0_sub: "snd ` snd ` set ps0 \ set hs \ (set gs \ set bs \ set hs)" . have "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" if "(p, q) \ snd ` set ps1" and "fst p \ 0" and "fst q \ 0" for p q proof - from \(p, q) \ snd ` set ps1\ obtain ic where "(ic, p, q) \ set ps1" by fastforce show ?thesis proof (cases "ic") case True from \(ic, p, q) \ set ps1\ obtain q_in_bs where "(ic, q_in_bs, p, q) \ set ps0" unfolding ps1_def by (rule apply_ncrit_subset) with True have "(True, q_in_bs, p, q) \ set ps0" by simp hence "snd (snd (True, q_in_bs, p, q)) \ snd ` snd ` set ps0" by fastforce hence "(p, q) \ snd ` snd ` set ps0" by simp also have "... \ set hs \ (set gs \ set bs \ set hs)" by (fact ps0_sub) finally have "p \ set hs" and "q \ set gs \ set bs \ set hs" by simp_all from B_sup have B_sup': "fst ` (set gs \ set bs \ set hs) \ fst ` B" by (rule image_mono) hence "fst ` (set gs \ set bs \ set hs) \ dgrad_p_set d m" using B_sub by (rule subset_trans) from assms(2) dg this uid gb \p \ set hs\ \q \ set gs \ set bs \ set hs\ \fst p \ 0\ \fst q \ 0\ \(True, q_in_bs, p, q) \ set ps0\ have "crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q)" unfolding ps0_def by (rule fst_apply_icrit) thus ?thesis using B_sup' by (rule crit_pair_cbelow_mono) next case False with \(ic, p, q) \ set ps1\ have "(False, p, q) \ set ps1" by simp with assms(5) ps1_def have "(p, q) \ set (add_pairs gs bs ps hs data)" by (simp add: set_add_pairs_iff ps0_def) hence "(p, q) \\<^sub>p set (add_pairs gs bs ps hs data)" by (simp add: in_pair_iff) thus ?thesis using \fst p \ 0\ \fst q \ 0\ by (rule a) qed qed with assms(3) dg B_sup B_sub ps0_sub uid gb have *: "(ic, q_in_bs, p, q) \ set ps0 \ fst p \ 0 \ fst q \ 0 \ (q_in_bs \ q \ set gs \ set bs) \ crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" for ic q_in_bs p q using b unfolding ps1_def by (rule apply_ncrit_connectible) show "crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)" proof (cases "h = g") case True from g_in B_sup have "g \ B" .. hence "fst g \ fst ` B" by simp hence "fst g \ dgrad_p_set d m" using B_sub .. with dg show ?thesis unfolding True by (rule crit_pair_cbelow_same) next case False with assms(1) h_in g_in show ?thesis proof (rule np_specE) fix g_in_bs assume "(g_in_bs, h, g) \ set (np gs bs hs data)" also have "... = snd ` set ps0" by (simp add: ps0_def) finally obtain ic where "(ic, g_in_bs, h, g) \ set ps0" by fastforce moreover note \fst h \ 0\ \fst g \ 0\ moreover from assms(1) have "g \ set gs \ set bs" if "g_in_bs" proof (rule np_specD4) from \(g_in_bs, h, g) \ set (np gs bs hs data)\ that show "(True, h, g) \ set (np gs bs hs data)" by simp qed ultimately show ?thesis by (rule *) next fix h_in_bs assume "(h_in_bs, g, h) \ set (np gs bs hs data)" also have "... = snd ` set ps0" by (simp add: ps0_def) finally obtain ic where "(ic, h_in_bs, g, h) \ set ps0" by fastforce moreover note \fst g \ 0\ \fst h \ 0\ moreover from assms(1) have "h \ set gs \ set bs" if "h_in_bs" proof (rule np_specD4) from \(h_in_bs, g, h) \ set (np gs bs hs data)\ that show "(True, g, h) \ set (np gs bs hs data)" by simp qed ultimately have "crit_pair_cbelow_on d m (fst ` B) (fst g) (fst h)" by (rule *) thus ?thesis by (rule crit_pair_cbelow_sym) qed qed next fix gs bs :: "('t, 'b, 'c) pdata list" and ps hs and data::"nat \ 'd" and B and d::"'a \ nat" and m h g define ps1 where "ps1 = apply_ncrit ncrit data gs bs hs (apply_icrit icrit data gs bs hs (np gs bs hs data))" assume "(h, g) \ set ps -\<^sub>p set (add_pairs gs bs ps hs data)" hence "(h, g) \ set ps" and "(h, g) \\<^sub>p set (add_pairs gs bs ps hs data)" by simp_all from this(2) have "(h, g) \ set (add_pairs gs bs ps hs data)" by (simp add: in_pair_iff) assume dg: "dickson_grading d" and B_sup: "set gs \ set bs \ set hs \ B" and B_sub: "fst ` B \ dgrad_p_set d m" and ps_sub: "set ps \ set bs \ (set gs \ set bs)" and "(set gs \ set bs) \ set hs = {}" \ \unused\ and uid: "unique_idx (gs @ bs @ hs) data" and gb: "is_Groebner_basis (fst ` set gs)" and "h \ g" and "fst h \ 0" and "fst g \ 0" assume *: "\a b. (a, b) \\<^sub>p set (add_pairs gs bs ps hs data) \ (a, b) \\<^sub>p set hs \ (set gs \ set bs \ set hs) \ fst a \ 0 \ fst b \ 0 \ crit_pair_cbelow_on d m (fst ` B) (fst a) (fst b)" have "snd ` set ps1 \ snd ` snd ` set (apply_icrit icrit data gs bs hs (np gs bs hs data))" unfolding ps1_def by (rule apply_ncrit_subset') also have "... = snd ` set (np gs bs hs data)" by simp also from assms(1) have "... \ set hs \ (set gs \ set bs \ set hs)" by (rule np_specD1) finally have ps1_sub: "snd ` set ps1 \ set hs \ (set gs \ set bs \ set hs)" . from \(h, g) \ set ps\ ps_sub have h_in: "h \ set gs \ set bs" and g_in: "g \ set gs \ set bs" by fastforce+ with B_sup have "h \ B" and "g \ B" by auto with assms(4) dg _ B_sub _ _ show "crit_pair_cbelow_on d m (fst ` B) (fst h) (fst g)" using \fst h \ 0\ \fst g \ 0\ \(h, g) \ set ps\ proof (rule apply_ocrit_connectible) from B_sup show "set hs \ B" by simp next from ps1_sub h_in g_in have "set (h # g # hs @ map (fst \ snd) ps1 @ map (snd \ snd) ps1) \ set (gs @ bs @ hs)" by fastforce with uid show "unique_idx (h # g # hs @ map (fst \ snd) ps1 @ map (snd \ snd) ps1) data" by (rule unique_idx_subset) next fix p q assume "(p, q) \ snd ` set ps1" hence pq_in: "(p, q) \ set hs \ (set gs \ set bs \ set hs)" using ps1_sub .. hence p_in: "p \ set hs" and q_in: "q \ set gs \ set bs \ set hs" by simp_all assume "fst p \ 0" and "fst q \ 0" from \(p, q) \ snd ` set ps1\ obtain ic where "(ic, p, q) \ set ps1" by fastforce show "crit_pair_cbelow_on d m (fst ` B) (fst p) (fst q)" proof (cases "ic") case True hence "ic = True" by simp from B_sup have B_sup': "fst ` (set gs \ set bs \ set hs) \ fst ` B" by (rule image_mono) note assms(2) dg moreover from B_sup' B_sub have "fst ` (set gs \ set bs \ set hs) \ dgrad_p_set d m" by (rule subset_trans) moreover note uid gb p_in q_in \fst p \ 0\ \fst q \ 0\ moreover from \(ic, p, q) \ set ps1\ obtain q_in_bs where "(True, q_in_bs, p, q) \ set (apply_icrit icrit data gs bs hs (np gs bs hs data))" unfolding ps1_def \ic = True\ by (rule apply_ncrit_subset) ultimately have "crit_pair_cbelow_on d m (fst ` (set gs \ set bs \ set hs)) (fst p) (fst q)" by (rule fst_apply_icrit) thus ?thesis using B_sup' by (rule crit_pair_cbelow_mono) next case False with \(ic, p, q) \ set ps1\ have "(False, p, q) \ set ps1" by simp with assms(5) ps1_def have "(p, q) \ set (add_pairs gs bs ps hs data)" by (simp add: set_add_pairs_iff) hence "(p, q) \\<^sub>p set (add_pairs gs bs ps hs data)" by (simp add: in_pair_iff) moreover from pq_in have "(p, q) \\<^sub>p set hs \ (set gs \ set bs \ set hs)" by (simp add: in_pair_iff) ultimately show ?thesis using \fst p \ 0\ \fst q \ 0\ by (rule *) qed next show "(h, g) \ set (apply_ocrit ocrit data hs ps1 ps)" proof assume "(h, g) \ set (apply_ocrit ocrit data hs ps1 ps)" hence "(h, g) \ set (add_pairs gs bs ps hs data)" by (simp add: add_pairs_def assms(5) Let_def ps1_def) with \(h, g) \ set (add_pairs gs bs ps hs data)\ show False .. qed qed qed end abbreviation "add_pairs_canon \ add_pairs (new_pairs_sorted canon_pair_order) component_crit chain_ncrit chain_ocrit canon_pair_comb" lemma ap_spec_add_pairs_canon: "ap_spec add_pairs_canon" using np_spec_new_pairs_sorted icrit_spec_component_crit ncrit_spec_chain_ncrit ocrit_spec_chain_ocrit set_merge_wrt by (rule ap_spec_add_pairs) subsection \Suitable Instances of the @{emph \completion\} Parameter\ definition rcp_spec :: "('t, 'b::field, 'c, 'd) complT \ bool" where "rcp_spec rcp \ (\gs bs ps sps data. 0 \ fst ` set (fst (rcp gs bs ps sps data)) \ (\h b. h \ set (fst (rcp gs bs ps sps data)) \ b \ set gs \ set bs \ fst b \ 0 \ \ lt (fst b) adds\<^sub>t lt (fst h)) \ (\d. dickson_grading d \ dgrad_p_set_le d (fst ` set (fst (rcp gs bs ps sps data))) (args_to_set (gs, bs, sps))) \ component_of_term ` Keys (fst ` (set (fst (rcp gs bs ps sps data)))) \ component_of_term ` Keys (args_to_set (gs, bs, sps)) \ (is_Groebner_basis (fst ` set gs) \ unique_idx (gs @ bs) data \ (fst ` set (fst (rcp gs bs ps sps data)) \ pmdl (args_to_set (gs, bs, sps)) \ (\(p, q)\set sps. set sps \ set bs \ (set gs \ set bs) \ (red (fst ` (set gs \ set bs) \ fst ` set (fst (rcp gs bs ps sps data))))\<^sup>*\<^sup>* (spoly (fst p) (fst q)) 0))))" text \Informally, \rcp_spec rcp\ expresses that, for suitable \gs\, \bs\ and \sps\, the value of \rcp gs bs ps sps\ \begin{itemize} \item is a list consisting exclusively of non-zero polynomials contained in the module generated by \set bs \ set gs\, whose leading terms are not divisible by the leading term of any non-zero @{prop "b \ set bs"}, and \item contains sufficiently many new polynomials such that all S-polynomials originating from \sps\ can be reduced to \0\ modulo the enlarged list of polynomials. \end{itemize}\ lemma rcp_specI: assumes "\gs bs ps sps data. 0 \ fst ` set (fst (rcp gs bs ps sps data))" assumes "\gs bs ps sps h b data. h \ set (fst (rcp gs bs ps sps data)) \ b \ set gs \ set bs \ fst b \ 0 \ \ lt (fst b) adds\<^sub>t lt (fst h)" assumes "\gs bs ps sps d data. dickson_grading d \ dgrad_p_set_le d (fst ` set (fst (rcp gs bs ps sps data))) (args_to_set (gs, bs, sps))" assumes "\gs bs ps sps data. component_of_term ` Keys (fst ` (set (fst (rcp gs bs ps sps data)))) \ component_of_term ` Keys (args_to_set (gs, bs, sps))" assumes "\gs bs ps sps data. is_Groebner_basis (fst ` set gs) \ unique_idx (gs @ bs) data \ (fst ` set (fst (rcp gs bs ps sps data)) \ pmdl (args_to_set (gs, bs, sps)) \ (\(p, q)\set sps. set sps \ set bs \ (set gs \ set bs) \ (red (fst ` (set gs \ set bs) \ fst ` set (fst (rcp gs bs ps sps data))))\<^sup>*\<^sup>* (spoly (fst p) (fst q)) 0))" shows "rcp_spec rcp" unfolding rcp_spec_def using assms by auto lemma rcp_specD1: assumes "rcp_spec rcp" shows "0 \ fst ` set (fst (rcp gs bs ps sps data))" using assms unfolding rcp_spec_def by (elim allE conjE) lemma rcp_specD2: assumes "rcp_spec rcp" and "h \ set (fst (rcp gs bs ps sps data))" and "b \ set gs \ set bs" and "fst b \ 0" shows "\ lt (fst b) adds\<^sub>t lt (fst h)" using assms unfolding rcp_spec_def by (elim allE conjE, blast) lemma rcp_specD3: assumes "rcp_spec rcp" and "dickson_grading d" shows "dgrad_p_set_le d (fst ` set (fst (rcp gs bs ps sps data))) (args_to_set (gs, bs, sps))" using assms unfolding rcp_spec_def by (elim allE conjE, blast) lemma rcp_specD4: assumes "rcp_spec rcp" shows "component_of_term ` Keys (fst ` (set (fst (rcp gs bs ps sps data)))) \ component_of_term ` Keys (args_to_set (gs, bs, sps))" using assms unfolding rcp_spec_def by (elim allE conjE) lemma rcp_specD5: assumes "rcp_spec rcp" and "is_Groebner_basis (fst ` set gs)" and "unique_idx (gs @ bs) data" shows "fst ` set (fst (rcp gs bs ps sps data)) \ pmdl (args_to_set (gs, bs, sps))" using assms unfolding rcp_spec_def by blast lemma rcp_specD6: assumes "rcp_spec rcp" and "is_Groebner_basis (fst ` set gs)" and "unique_idx (gs @ bs) data" and "set sps \ set bs \ (set gs \ set bs)" and "(p, q) \ set sps" shows "(red (fst ` (set gs \ set bs) \ fst ` set (fst (rcp gs bs ps sps data))))\<^sup>*\<^sup>* (spoly (fst p) (fst q)) 0" using assms unfolding rcp_spec_def by blast lemma compl_struct_rcp: assumes "rcp_spec rcp" shows "compl_struct rcp" proof (rule compl_structI) fix d::"'a \ nat" and gs bs ps and sps::"('t, 'b, 'c) pdata_pair list" and data::"nat \ 'd" assume "dickson_grading d" and "set sps \ set ps" from assms this(1) have "dgrad_p_set_le d (fst ` set (fst (rcp gs bs (ps -- sps) sps data))) (args_to_set (gs, bs, sps))" by (rule rcp_specD3) also have "dgrad_p_set_le d ... (args_to_set (gs, bs, ps))" by (rule dgrad_p_set_le_subset, rule args_to_set_subset3, fact \set sps \ set ps\) finally show "dgrad_p_set_le d (fst ` set (fst (rcp gs bs (ps -- sps) sps data))) (args_to_set (gs, bs, ps))" . next fix gs bs ps and sps::"('t, 'b, 'c) pdata_pair list" and data::"nat \ 'd" from assms show "0 \ fst ` set (fst (rcp gs bs (ps -- sps) sps data))" by (rule rcp_specD1) next fix gs bs ps sps h b data assume "h \ set (fst (rcp gs bs (ps -- sps) sps data))" and "b \ set gs \ set bs" and "fst b \ 0" with assms show "\ lt (fst b) adds\<^sub>t lt (fst h)" by (rule rcp_specD2) next fix gs bs ps and sps::"('t, 'b, 'c) pdata_pair list" and data::"nat \ 'd" assume "set sps \ set ps" from assms have "component_of_term ` Keys (fst ` set (fst (rcp gs bs (ps -- sps) sps data))) \ component_of_term ` Keys (args_to_set (gs, bs, sps))" by (rule rcp_specD4) also have "... \ component_of_term ` Keys (args_to_set (gs, bs, ps))" by (rule image_mono, rule Keys_mono, rule args_to_set_subset3, fact \set sps \ set ps\) finally show "component_of_term ` Keys (fst ` set (fst (rcp gs bs (ps -- sps) sps data))) \ component_of_term ` Keys (args_to_set (gs, bs, ps))" . qed lemma compl_pmdl_rcp: assumes "rcp_spec rcp" shows "compl_pmdl rcp" proof (rule compl_pmdlI) fix gs bs :: "('t, 'b, 'c) pdata list" and ps sps :: "('t, 'b, 'c) pdata_pair list" and data::"nat \ 'd" assume gb: "is_Groebner_basis (fst ` set gs)" and "set sps \ set ps" and un: "unique_idx (gs @ bs) data" let ?res = "fst (rcp gs bs (ps -- sps) sps data)" from assms gb un have "fst ` set ?res \ pmdl (args_to_set (gs, bs, sps))" by (rule rcp_specD5) also have "... \ pmdl (args_to_set (gs, bs, ps))" by (rule pmdl.span_mono, rule args_to_set_subset3, fact \set sps \ set ps\) finally show "fst ` set ?res \ pmdl (args_to_set (gs, bs, ps))" . qed lemma compl_conn_rcp: assumes "rcp_spec rcp" shows "compl_conn rcp" proof (rule compl_connI) fix d::"'a \ nat" and m gs bs ps sps p and q::"('t, 'b, 'c) pdata" and data::"nat \ 'd" assume dg: "dickson_grading d" and gs_sub: "fst ` set gs \ dgrad_p_set d m" and gb: "is_Groebner_basis (fst ` set gs)" and bs_sub: "fst ` set bs \ dgrad_p_set d m" and ps_sub: "set ps \ set bs \ (set gs \ set bs)" and "set sps \ set ps" and uid: "unique_idx (gs @ bs) data" and "(p, q) \ set sps" and "fst p \ 0" and "fst q \ 0" from \set sps \ set ps\ ps_sub have sps_sub: "set sps \ set bs \ (set gs \ set bs)" by (rule subset_trans) let ?res = "fst (rcp gs bs (ps -- sps) sps data)" have "fst ` set ?res \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set, rule rcp_specD3, fact+) show "args_to_set (gs, bs, sps) \ dgrad_p_set d m" by (simp add: args_to_set_subset_Times[OF sps_sub], rule, fact+) qed moreover have gs_bs_sub: "fst ` (set gs \ set bs) \ dgrad_p_set d m" by (simp add: image_Un, rule, fact+) ultimately have res_sub: "fst ` (set gs \ set bs) \ fst ` set ?res \ dgrad_p_set d m" by simp from \(p, q) \ set sps\ \set sps \ set ps\ ps_sub have "fst p \ fst ` set bs" and "fst q \ fst ` (set gs \ set bs)" by auto with \fst ` set bs \ dgrad_p_set d m\ gs_bs_sub have "fst p \ dgrad_p_set d m" and "fst q \ dgrad_p_set d m" by auto with dg res_sub show "crit_pair_cbelow_on d m (fst ` (set gs \ set bs) \ fst ` set ?res) (fst p) (fst q)" using \fst p \ 0\ \fst q \ 0\ proof (rule spoly_red_zero_imp_crit_pair_cbelow_on) from assms gb uid sps_sub \(p, q) \ set sps\ show "(red (fst ` (set gs \ set bs) \ fst ` set (fst (rcp gs bs (ps -- sps) sps data))))\<^sup>*\<^sup>* (spoly (fst p) (fst q)) 0" by (rule rcp_specD6) qed qed end (* gd_term *) subsection \Suitable Instances of the @{emph \add-basis\} Parameter\ definition add_basis_naive :: "('a, 'b, 'c, 'd) abT" where "add_basis_naive gs bs ns data = bs @ ns" lemma ab_spec_add_basis_naive: "ab_spec add_basis_naive" by (rule ab_specI, simp_all add: add_basis_naive_def) definition add_basis_sorted :: "(nat \ 'd \ ('a, 'b, 'c) pdata \ ('a, 'b, 'c) pdata \ bool) \ ('a, 'b, 'c, 'd) abT" where "add_basis_sorted rel gs bs ns data = merge_wrt (rel data) bs ns" lemma ab_spec_add_basis_sorted: "ab_spec (add_basis_sorted rel)" by (rule ab_specI, simp_all add: add_basis_sorted_def set_merge_wrt) definition card_keys :: "('a \\<^sub>0 'b::zero) \ nat" where "card_keys = card \ keys" definition (in ordered_term) canon_basis_order :: "'d \ ('t, 'b::zero, 'c) pdata \ ('t, 'b, 'c) pdata \ bool" where "canon_basis_order data p q \ (let cp = card_keys (fst p); cq = card_keys (fst q) in cp < cq \ (cp = cq \ lt (fst p) \\<^sub>t lt (fst q)))" abbreviation (in ordered_term) "add_basis_canon \ add_basis_sorted canon_basis_order" subsection \Special Case: Scalar Polynomials\ context gd_powerprod begin lemma remdups_map_component_of_term_punit: "remdups (map (\_. ()) (punit.Keys_to_list (map fst bs))) = (if (\b\set bs. fst b = 0) then [] else [()])" proof (split if_split, intro conjI impI) assume "\b\set bs. fst b = 0" hence "fst ` set bs \ {0}" by blast hence "Keys (fst ` set bs) = {}" by (metis Keys_empty Keys_zero subset_singleton_iff) hence "punit.Keys_to_list (map fst bs) = []" by (simp add: set_empty[symmetric] punit.set_Keys_to_list del: set_empty) thus "remdups (map (\_. ()) (punit.Keys_to_list (map fst bs))) = []" by simp next assume "\ (\b\set bs. fst b = 0)" hence "\b\set bs. fst b \ 0" by simp then obtain b where "b \ set bs" and "fst b \ 0" .. hence "Keys (fst ` set bs) \ {}" by (meson Keys_not_empty \fst b \ 0\ imageI) hence "set (punit.Keys_to_list (map fst bs)) \ {}" by (simp add: punit.set_Keys_to_list) hence "punit.Keys_to_list (map fst bs) \ []" by simp thus "remdups (map (\_. ()) (punit.Keys_to_list (map fst bs))) = [()]" by (metis (full_types) old.unit.exhaust sorted.cases Nil_is_map_conv \punit.Keys_to_list (map fst bs) \ []\ distinct_length_2_or_more distinct_remdups remdups_eq_nil_right_iff) (*SLOW: 13s*) qed lemma count_const_lt_components_punit [code]: "punit.count_const_lt_components hs = (if (\h\set hs. punit.const_lt_component (fst h) = Some ()) then 1 else 0)" proof (simp add: punit.count_const_lt_components_def cong del: image_cong_simp, simp add: card_set [symmetric] cong del: image_cong_simp, rule) assume "\h\set hs. punit.const_lt_component (fst h) = Some ()" then obtain h where "h \ set hs" and "punit.const_lt_component (fst h) = Some ()" .. from this(2) have "(punit.const_lt_component \ fst) h = Some ()" by simp with \h \ set hs\ have "Some () \ (punit.const_lt_component \ fst) ` set hs" by (metis rev_image_eqI) hence "{x. x = Some () \ x \ (punit.const_lt_component \ fst) ` set hs} = {Some ()}" by auto thus "card {x. x = Some () \ x \ (punit.const_lt_component \ fst) ` set hs} = Suc 0" by simp qed lemma count_rem_components_punit [code]: "punit.count_rem_components bs = (if (\b\set bs. fst b = 0) then 0 else if (\b\set bs. fst b \ 0 \ punit.const_lt_component (fst b) = Some ()) then 0 else 1)" proof (cases "\b\set bs. fst b = 0") case True thus ?thesis by (simp add: punit.count_rem_components_def remdups_map_component_of_term_punit) next case False have eq: "(\b\set [b\bs . fst b \ 0]. punit.const_lt_component (fst b) = Some ()) = (\b\set bs. fst b \ 0 \ punit.const_lt_component (fst b) = Some ())" by (metis (mono_tags, lifting) filter_set member_filter) show ?thesis by (simp only: False punit.count_rem_components_def eq if_False remdups_map_component_of_term_punit count_const_lt_components_punit punit_component_of_term, simp) qed lemma full_gb_punit [code]: "punit.full_gb bs = (if (\b\set bs. fst b = 0) then [] else [(1, 0, default)])" by (simp add: punit.full_gb_def remdups_map_component_of_term_punit) abbreviation "add_pairs_punit_canon \ punit.add_pairs (punit.new_pairs_sorted punit.canon_pair_order) punit.product_crit punit.chain_ncrit punit.chain_ocrit punit.canon_pair_comb" lemma ap_spec_add_pairs_punit_canon: "punit.ap_spec add_pairs_punit_canon" using punit.np_spec_new_pairs_sorted punit.icrit_spec_product_crit punit.ncrit_spec_chain_ncrit punit.ocrit_spec_chain_ocrit set_merge_wrt by (rule punit.ap_spec_add_pairs) end (* gd_powerprod *) end (* theory *) diff --git a/thys/Jordan_Normal_Form/Determinant.thy b/thys/Jordan_Normal_Form/Determinant.thy --- a/thys/Jordan_Normal_Form/Determinant.thy +++ b/thys/Jordan_Normal_Form/Determinant.thy @@ -1,2429 +1,2429 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Determinants\ text \Most of the following definitions and proofs on determinants have been copied and adapted from ~~/src/HOL/Multivariate-Analysis/Determinants.thy. Exceptions are \emph{det-identical-rows}. We further generalized some lemmas, e.g., that the determinant is 0 iff the kernel of a matrix is non-empty is available for integral domains, not just for fields.\ theory Determinant imports Missing_Permutations Column_Operations "HOL-Computational_Algebra.Polynomial_Factorial" (* Only for to_fract. Probably not the right place. *) Polynomial_Interpolation.Ring_Hom Polynomial_Interpolation.Missing_Unsorted begin definition det:: "'a mat \ 'a :: comm_ring_1" where "det A = (if dim_row A = dim_col A then (\ p \ {p. p permutes {0 ..< dim_row A}}. signof p * (\ i = 0 ..< dim_row A. A $$ (i, p i))) else 0)" lemma(in ring_hom) hom_signof[simp]: "hom (signof p) = signof p" unfolding signof_def by (auto simp: hom_distribs) lemma(in comm_ring_hom) hom_det[simp]: "det (map_mat hom A) = hom (det A)" unfolding det_def by (auto simp: hom_distribs) lemma det_def': "A \ carrier_mat n n \ det A = (\ p \ {p. p permutes {0 ..< n}}. signof p * (\ i = 0 ..< n. A $$ (i, p i)))" unfolding det_def by auto lemma det_smult[simp]: "det (a \\<^sub>m A) = a ^ dim_col A * det A" proof - have [simp]: "(\i = 0.. carrier_mat n n" shows "det (transpose_mat A) = det A" proof - let ?di = "\A i j. A $$ (i,j)" let ?U = "{0 ..< n}" have fU: "finite ?U" by simp let ?inv = "Hilbert_Choice.inv" { fix p assume p: "p \ {p. p permutes ?U}" from p have pU: "p permutes ?U" by blast have sth: "signof (?inv p) = signof p" by (rule signof_inv[OF _ pU], simp) from permutes_inj[OF pU] have pi: "inj_on p ?U" by (blast intro: subset_inj_on) let ?f = "\i. transpose_mat A $$ (i, ?inv p i)" note pU_U = permutes_image[OF pU] note [simp] = permutes_less[OF pU] have "prod ?f ?U = prod ?f (p ` ?U)" using pU_U by simp also have "\ = prod (?f \ p) ?U" by (rule prod.reindex[OF pi]) also have "\ = prod (\i. A $$ (i, p i)) ?U" by (rule prod.cong, insert A, auto) finally have "signof (?inv p) * prod ?f ?U = signof p * prod (\i. A $$ (i, p i)) ?U" unfolding sth by simp } then show ?thesis unfolding det_def using A by (simp, subst sum_permutations_inverse, intro sum.cong, auto) qed lemma det_col: assumes A: "A \ carrier_mat n n" shows "det A = (\ p | p permutes {0 ..< n}. signof p * (\jp. _ * ?prod p) ?P)") proof - let ?i = "Hilbert_Choice.inv" let ?N = "{0 ..< n}" let ?f = "\p. signof p * ?prod p" let ?prod' = "\p. \jx y. A $$ (x,y)"] by auto have [simp]: "signof p = signof (?i p)" apply(rule signof_inv[symmetric]) using p by auto show "?f p = ?f' p" by auto qed also have "... = sum ?f' ?P'" by (rule sum.cong[OF image_inverse_permutations[symmetric]],auto) also have "... = sum ?f'' ?P" unfolding sum.reindex[OF inv_inj_on_permutes,unfolded image_Collect] unfolding o_def apply (rule sum.cong[OF refl]) using inv_inv_eq[OF permutes_bij] by force finally show ?thesis unfolding det_def'[OF A] by auto qed lemma mat_det_left_def: assumes A: "A \ carrier_mat n n" shows "det A = (\p\{p. p permutes {0..i = 0 ..< dim_row A. A $$ (p i, i)))" proof - have cong: "\ a b c. b = c \ a * b = a * c" by simp show ?thesis unfolding det_transpose[OF A, symmetric] unfolding det_def index_transpose_mat using A by simp qed lemma det_upper_triangular: assumes ut: "upper_triangular A" and m: "A \ carrier_mat n n" shows "det A = prod_list (diag_mat A)" proof - note det_def = det_def'[OF m] let ?U = "{0.. ?PU" by (auto simp add: permutes_id) { fix p assume p: "p \ ?PU - {id}" from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ from permutes_natset_ge[OF pU] pid obtain i where i: "p i < i" and "i < n" by fastforce from upper_triangularD[OF ut i] \i < n\ m have ex:"\i \ ?U. A $$ (i,p i) = 0" by auto have "(\ i = 0 ..< n. A $$ (i, p i)) = 0" by (rule prod_zero[OF fU ex]) hence "?pp p = 0" by simp } then have p0: "\ p. p \ ?PU - {id} \ ?pp p = 0" by blast from m have dim: "dim_row A = n" by simp have "det A = (\ p \ ?PU. ?pp p)" unfolding det_def by auto also have "\ = ?pp id + (\ p \ ?PU - {id}. ?pp p)" by (rule sum.remove, insert id0 fPU m, auto simp: p0) also have "(\ p \ ?PU - {id}. ?pp p) = 0" by (rule sum.neutral, insert fPU, auto simp: p0) finally show ?thesis using m by (auto simp: prod_list_diag_prod) qed lemma det_one[simp]: "det (1\<^sub>m n) = 1" proof - have "det (1\<^sub>m n) = prod_list (diag_mat (1\<^sub>m n))" by (rule det_upper_triangular[of _ n], auto) also have "\ = 1" by (induct n, auto) finally show ?thesis . qed lemma det_zero[simp]: assumes "n > 0" shows "det (0\<^sub>m n n) = 0" proof - have "det (0\<^sub>m n n) = prod_list (diag_mat (0\<^sub>m n n))" by (rule det_upper_triangular[of _ n], auto) also have "\ = 0" using \n > 0\ by (cases n, auto) finally show ?thesis . qed lemma det_dim_zero[simp]: "A \ carrier_mat 0 0 \ det A = 1" unfolding det_def carrier_mat_def signof_def sign_def by auto lemma det_lower_triangular: assumes ld: "\i j. i < j \ j < n \ A $$ (i,j) = 0" and m: "A \ carrier_mat n n" shows "det A = prod_list (diag_mat A)" proof - have "det A = det (transpose_mat A)" using det_transpose[OF m] by simp also have "\ = prod_list (diag_mat (transpose_mat A))" by (rule det_upper_triangular, insert m ld, auto) finally show ?thesis using m by simp qed lemma det_permute_rows: assumes A: "A \ carrier_mat n n" and p: "p permutes {0 ..< (n :: nat)}" shows "det (mat n n (\ (i,j). A $$ (p i, j))) = signof p * det A" proof - let ?U = "{0 ..< (n :: nat)}" have cong: "\ a b c. b = c \ a * b = a * c" by auto have "det (mat n n (\ (i,j). A $$ (p i, j))) = (\ q \ {q . q permutes ?U}. signof q * (\ i \ ?U. A $$ (p i, q i)))" unfolding det_def using A p by auto also have "\ = (\ q \ {q . q permutes ?U}. signof (q \ p) * (\ i \ ?U. A $$ (p i, (q \ p) i)))" by (rule sum_permutations_compose_right[OF p]) finally have 1: "det (mat n n (\ (i,j). A $$ (p i, j))) = (\ q \ {q . q permutes ?U}. signof (q \ p) * (\ i \ ?U. A $$ (p i, (q \ p) i)))" . have 2: "signof p * det A = (\ q\{q. q permutes ?U}. signof p * signof q * (\i\ ?U. A $$ (i, q i)))" unfolding det_def'[OF A] sum_distrib_left by (simp add: ac_simps) show ?thesis unfolding 1 2 proof (rule sum.cong, insert p A, auto) fix q assume q: "q permutes ?U" let ?inv = "Hilbert_Choice.inv" from permutes_inv[OF p] have ip: "?inv p permutes ?U" . have "prod (\i. A $$ (p i, (q \ p) i)) ?U = prod (\i. A $$ ((p \ ?inv p) i, (q \ (p \ ?inv p)) i)) ?U" unfolding o_def by (rule trans[OF prod.permute[OF ip] prod.cong], insert A p q, auto) also have "\ = prod (\i. A$$(i,q i)) ?U" by (simp only: o_def permutes_inverses[OF p]) finally have thp: "prod (\i. A $$ (p i, (q \ p) i)) ?U = prod (\i. A$$(i,q i)) ?U" . show "signof (q \ p) * (\i\{0..i\{0..i\{0..i\{0..i\{0..i\{0.. l < n \ swaprows_mat n k l = mat n n (\(i, j). 1\<^sub>m n $$ (Fun.swap k l id i, j))" - by (rule eq_matI, auto simp: swap_def) + by (rule eq_matI) (auto simp add: swap_id_eq) lemma det_swaprows_mat: assumes k: "k < n" and l: "l < n" and kl: "k \ l" shows "det (swaprows_mat n k l) = - 1" proof - let ?n = "{0 ..< (n :: nat)}" let ?p = "Fun.swap k l id" have p: "?p permutes ?n" by (rule permutes_swap_id, insert k l, auto) show ?thesis by (rule trans[OF trans[OF _ det_permute_rows[OF one_carrier_mat[of n] p]]], subst swap_rows_mat_eq_permute[OF k l], auto simp: signof_def sign_swap_id kl) qed lemma det_addrow_mat: assumes l: "k \ l" shows "det (addrow_mat n a k l) = 1" proof - have "det (addrow_mat n a k l) = prod_list (diag_mat (addrow_mat n a k l))" proof (cases "k < l") case True show ?thesis by (rule det_upper_triangular[of _ n], insert True, auto intro!: upper_triangularI) next case False show ?thesis by (rule det_lower_triangular[of n], insert False, auto) qed also have "\ = 1" unfolding prod_list_diag_prod by (rule prod.neutral, insert l, auto) finally show ?thesis . qed text \The following proof is new, as it does not use $2 \neq 0$ as in Multivariate-Analysis.\ lemma det_identical_rows: assumes A: "A \ carrier_mat n n" and ij: "i \ j" and i: "i < n" and j: "j < n" and r: "row A i = row A j" shows "det A = 0" proof- let ?p = "Fun.swap i j id" let ?n = "{0 ..< n}" have sp: "signof ?p = - 1" "sign ?p = -1" unfolding signof_def using ij by (auto simp add: sign_swap_id) let ?f = "\ p. signof p * (\i\?n. A $$ (p i, i))" let ?all = "{p. p permutes ?n}" let ?one = "{p. p permutes ?n \ sign p = 1}" let ?none = "{p. p permutes ?n \ sign p \ 1}" let ?pone = "(\ p. ?p o p) ` ?one" have split: "?one \ ?none = ?all" by auto have p: "?p permutes ?n" by (rule permutes_swap_id, insert i j, auto) from permutes_inj[OF p] have injp: "inj ?p" by auto { fix q assume q: "q permutes ?n" have "(\k\?n. A $$ (?p (q k), k)) = (\k\?n. A $$ (q k, k))" proof (rule prod.cong) fix k assume k: "k \ ?n" from r have row: "row A i $ k = row A j $ k" by simp hence "A $$ (i,k) = A $$ (j,k)" using k i j A by auto thus "A $$ (?p (q k), k) = A $$ (q k, k)" by (cases "q k = i", auto, cases "q k = j", auto) qed (insert A q, auto) } note * = this have pp: "\ q. q permutes ?n \ permutation q" unfolding permutation_permutes by auto have "det A = (\p\ ?one \ ?none. ?f p)" using A unfolding mat_det_left_def[OF A] split by simp also have "\ = (\p\ ?one. ?f p) + (\p\ ?none. ?f p)" by (rule sum.union_disjoint, insert A, auto simp: finite_permutations) also have "?none = ?pone" proof - { fix q assume "q \ ?none" hence q: "q permutes ?n" and sq: "sign q = -1" unfolding sign_def by auto from permutes_compose[OF q p] sign_compose[OF pp[OF p] pp[OF q], unfolded sp sq] have "?p o q \ ?one" by auto hence "?p o (?p o q) \ ?pone" by auto - also have "?p o (?p o q) = q" unfolding o_def - by (intro ext, auto simp: swap_def) + also have "?p o (?p o q) = q" + by (auto simp: swap_id_eq) finally have "q \ ?pone" . } moreover { fix pq assume "pq \ ?pone" then obtain q where q: "q \ ?one" and pq: "pq = ?p o q" by auto from q have q: "q permutes ?n" and sq: "sign q = 1" by auto from sign_compose[OF pp[OF p] pp[OF q], unfolded sq sp] have spq: "sign pq = -1" unfolding pq by auto from permutes_compose[OF q p] have pq: "pq permutes ?n" unfolding pq by auto from pq spq have "pq \ ?none" by auto } ultimately show ?thesis by blast qed also have "(\p\ ?pone. ?f p) = (\p\ ?one. ?f (?p o p))" proof (rule trans[OF sum.reindex]) show "inj_on ((\) ?p) ?one" using fun.inj_map[OF injp] unfolding inj_on_def by auto qed simp also have "(\p\ ?one. ?f p) + (\p\ ?one. ?f (?p o p)) = (\p\ ?one. ?f p + ?f (?p o p))" by (rule sum.distrib[symmetric]) also have "\ = 0" by (rule sum.neutral, insert A, auto simp: sp sign_compose[OF pp[OF p] pp] ij signof_def finite_permutations *) finally show ?thesis . qed lemma det_row_0: assumes k: "k < n" and c: "c \ {0 ..< n} \ carrier_vec n" shows "det (mat\<^sub>r n n (\i. if i = k then 0\<^sub>v n else c i)) = 0" proof - { fix p assume p: "p permutes {0 ..< n}" have "(\i\{0..r n n (\i. if i = k then 0\<^sub>v n else c i) $$ (i, p i)) = 0" by (rule prod_zero[OF _ bexI[of _ k]], insert k p c[unfolded carrier_vec_def], auto) } thus ?thesis unfolding det_def by simp qed lemma det_row_add: assumes abc: "a k \ carrier_vec n" "b k \ carrier_vec n" "c \ {0.. carrier_vec n" and k: "k < n" shows "det(mat\<^sub>r n n (\ i. if i = k then a i + b i else c i)) = det(mat\<^sub>r n n (\ i. if i = k then a i else c i)) + det(mat\<^sub>r n n (\ i. if i = k then b i else c i))" (is "?lhs = ?rhs") proof - let ?n = "{0..p\{p. p permutes ?n}. signof p * (\i\?n. ?m a c p i)) + (\p\{p. p permutes ?n}. signof p * (\i\?n. ?m b c p i))" unfolding det_def by simp also have "\ = (\p\{p. p permutes ?n}. signof p * (\i\?n. ?m a c p i) + signof p * (\i\?n. ?m b c p i))" by (rule sum.distrib[symmetric]) also have "\ = (\p\{p. p permutes ?n}. signof p * (\i\?n. ?m ?ab c p i))" proof (rule sum.cong, force) fix p assume "p \ {p. p permutes ?n}" hence p: "p permutes ?n" by simp show "signof p * (\i\?n. ?m a c p i) + signof p * (\i\?n. ?m b c p i) = signof p * (\i\?n. ?m ?ab c p i)" unfolding distrib_left[symmetric] proof (rule arg_cong[of _ _ "\ a. signof p * a"]) from k have f: "finite ?n" and k': "k \ ?n" by auto let ?nk = "?n - {k}" note split = prod.remove[OF f k'] have id1: "(\i\?n. ?m a c p i) = ?m a c p k * (\i\?nk. ?m a c p i)" by (rule split) have id2: "(\i\?n. ?m b c p i) = ?m b c p k * (\i\?nk. ?m b c p i)" by (rule split) have id3: "(\i\?n. ?m ?ab c p i) = ?m ?ab c p k * (\i\?nk. ?m ?ab c p i)" by (rule split) have id: "\ a. (\i\?nk. ?m a c p i) = (\i\?nk. ?c p i)" by (rule prod.cong, insert abc k p, auto intro!: intros) have ab: "?ab k \ carrier_vec n" using abc by (auto intro: intros) { fix f assume "f k \ (carrier_vec n :: 'a vec set)" hence "mat\<^sub>r n n (\i. if i = k then f i else c i) $$ (k, p k) = f k $ p k" by (insert p k abc, auto) } note first = this note id' = id1 id2 id3 have dist: "(a k + b k) $ p k = a k $ p k + b k $ p k" by (rule index_add_vec(1), insert p k abc, force) show "(\i\?n. ?m a c p i) + (\i\?n. ?m b c p i) = (\i\?n. ?m ?ab c p i)" unfolding id' id first[of a, OF abc(1)] first[of b, OF abc(2)] first[of ?ab, OF ab] dist by (rule distrib_right[symmetric]) qed qed also have "\ = ?lhs" unfolding det_def by simp finally show ?thesis by simp qed lemma det_linear_row_finsum: assumes fS: "finite S" and c: "c \ {0.. carrier_vec n" and k: "k < n" and a: "a k \ S \ carrier_vec n" shows "det (mat\<^sub>r n n (\ i. if i = k then finsum_vec TYPE('a :: comm_ring_1) n (a i) S else c i)) = sum (\j. det (mat\<^sub>r n n (\ i. if i = k then a i j else c i))) S" proof - let ?sum = "finsum_vec TYPE('a) n" show ?thesis using a proof (induct rule: finite_induct[OF fS]) case 1 show ?case by (simp, unfold finsum_vec_empty, rule det_row_0[OF k c]) next case (2 x F) from 2(4) have ak: "a k \ F \ carrier_vec n" and akx: "a k x \ carrier_vec n" by auto { fix i note if_cong[OF refl finsum_vec_insert[OF 2(1-2)], of _ "a i" n "c i" "c i"] } note * = this show ?case proof (subst *) show "det (mat\<^sub>r n n (\i. if i = k then a i x + ?sum (a i) F else c i)) = (\j\insert x F. det (mat\<^sub>r n n (\i. if i = k then a i j else c i)))" proof (subst det_row_add) show "det (mat\<^sub>r n n (\i. if i = k then a i x else c i)) + det (mat\<^sub>r n n (\i. if i = k then ?sum (a i) F else c i)) = (\j\insert x F. det (mat\<^sub>r n n (\i. if i = k then a i j else c i)))" unfolding 2(3)[OF ak] sum.insert[OF 2(1-2)] by simp qed (insert c k ak akx 2(1), auto intro!: finsum_vec_closed) qed (insert akx ak, force+) qed qed lemma det_linear_rows_finsum_lemma: assumes fS: "finite S" and fT: "finite T" and c: "c \ {0.. carrier_vec n" and T: "T \ {0 ..< n}" and a: "a \ T \ S \ carrier_vec n" shows "det (mat\<^sub>r n n (\ i. if i \ T then finsum_vec TYPE('a :: comm_ring_1) n (a i) S else c i)) = sum (\f. det(mat\<^sub>r n n (\ i. if i \ T then a i (f i) else c i))) {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" proof - let ?sum = "finsum_vec TYPE('a) n" show ?thesis using fT c a T proof (induct T arbitrary: a c set: finite) case empty let ?f = "(\ i. i) :: nat \ nat" have [simp]: "{f. \i. f i = i} = {?f}" by auto show ?case by simp next case (insert z T a c) hence z: "z < n" and azS: "a z \ S \ carrier_vec n" by auto let ?F = "\T. {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" let ?h = "\(y,g) i. if i = z then y else g i" let ?k = "\h. (h(z),(\i. if i = z then i else h i))" let ?s = "\ k a c f. det(mat\<^sub>r n n (\ i. if i \ T then a i (f i) else c i))" let ?c = "\j i. if i = z then a i j else c i" have thif: "\a b c d. (if a \ b then c else d) = (if a then c else if b then c else d)" by simp have thif2: "\a b c d e. (if a then b else if c then d else e) = (if c then (if a then b else d) else (if a then b else e))" by simp from \z \ T\ have nz: "\i. i \ T \ i = z \ False" by auto from insert have c: "\ i. i < n \ c i \ carrier_vec n" by auto have fin: "finite {f. (\i\T. f i \ S) \ (\i. i \ T \ f i = i)}" by (rule finite_bounded_functions[OF fS insert(1)]) have "det (mat\<^sub>r n n (\ i. if i \ insert z T then ?sum (a i) S else c i)) = det (mat\<^sub>r n n (\ i. if i = z then ?sum (a i) S else if i \ T then ?sum (a i) S else c i))" unfolding insert_iff thif .. also have "\ = (\j\S. det (mat\<^sub>r n n (\ i. if i \ T then ?sum (a i) S else if i = z then a i j else c i)))" apply (subst det_linear_row_finsum[OF fS _ z]) prefer 3 apply (subst thif2) using nz apply (simp cong del: if_weak_cong cong add: if_cong) apply (insert azS c fS insert(5), (force intro!: finsum_vec_closed)+) done also have "\ = (sum (\ (j, f). det (mat\<^sub>r n n (\ i. if i \ T then a i (f i) else if i = z then a i j else c i))) (S \ ?F T))" unfolding sum.cartesian_product[symmetric] by (rule sum.cong[OF refl], subst insert.hyps(3), insert azS c fin z insert(5-6), auto) finally have tha: "det (mat\<^sub>r n n (\ i. if i \ insert z T then ?sum (a i) S else c i)) = (sum (\ (j, f). det (mat\<^sub>r n n (\ i. if i \ T then a i (f i) else if i = z then a i j else c i))) (S \ ?F T))" . show ?case unfolding tha by (rule sum.reindex_bij_witness[where i="?k" and j="?h"], insert \z \ T\ azS c fS insert(5-6) z fin, auto intro!: arg_cong[of _ _ det]) qed qed lemma det_linear_rows_sum: assumes fS: "finite S" and a: "a \ {0.. S \ carrier_vec n" shows "det (mat\<^sub>r n n (\ i. finsum_vec TYPE('a :: comm_ring_1) n (a i) S)) = sum (\f. det (mat\<^sub>r n n (\ i. a i (f i)))) {f. (\i\{0.. S) \ (\i. i \ {0.. f i = i)}" proof - let ?T = "{0..x y. mat\<^sub>r n n (\ i. if i \ ?T then x i else y i) = mat\<^sub>r n n (\ i. x i)" by (rule eq_rowI, auto) have c: "(\ _. 0\<^sub>v n) \ ?T \ carrier_vec n" by auto show ?thesis by (rule det_linear_rows_finsum_lemma[OF fS fT c subset_refl a, unfolded th0]) qed lemma det_rows_mul: assumes a: "a \ {0.. carrier_vec n" shows "det(mat\<^sub>r n n (\ i. c i \\<^sub>v a i)) = prod c {0..r n n (\ i. a i))" proof - have A: "mat\<^sub>r n n (\ i. c i \\<^sub>v a i) \ carrier_mat n n" and A': "mat\<^sub>r n n (\ i. a i) \ carrier_mat n n" using a unfolding carrier_mat_def by auto show ?thesis unfolding det_def'[OF A] det_def'[OF A'] proof (rule trans[OF sum.cong sum_distrib_left[symmetric]]) fix p assume p: "p \ {p. p permutes {0..ia\{0..r n n (\i. c i \\<^sub>v a i) $$ (ia, p ia)) = prod c {0..ia\{0..r n n a $$ (ia, p ia))" unfolding prod.distrib[symmetric] by (rule prod.cong, insert p a, force+) show "signof p * (\ia\{0..r n n (\i. c i \\<^sub>v a i) $$ (ia, p ia)) = prod c {0..ia\{0..r n n a $$ (ia, p ia)))" unfolding id by auto qed simp qed lemma mat_mul_finsum_alt: assumes A: "A \ carrier_mat nr n" and B: "B \ carrier_mat n nc" shows "A * B = mat\<^sub>r nr nc (\ i. finsum_vec TYPE('a :: semiring_0) nc (\k. A $$ (i,k) \\<^sub>v row B k) {0 ..< n})" by (rule eq_matI, insert A B, auto, subst index_finsum_vec, auto simp: scalar_prod_def intro: sum.cong) lemma det_mult: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "det (A * B) = det A * det (B :: 'a :: comm_ring_1 mat)" proof - let ?U = "{0 ..< n}" let ?F = "{f. (\i\ ?U. f i \ ?U) \ (\i. i \ ?U \ f i = i)}" let ?PU = "{p. p permutes ?U}" have fU: "finite ?U" by blast have fF: "finite ?F" by (rule finite_bounded_functions, auto) { fix p assume p: "p permutes ?U" have "p \ ?F" unfolding mem_Collect_eq permutes_in_image[OF p] using p[unfolded permutes_def] by simp } then have PUF: "?PU \ ?F" by blast { fix f assume fPU: "f \ ?F - ?PU" have fUU: "f ` ?U \ ?U" using fPU by auto from fPU have f: "\i \ ?U. f i \ ?U" "\i. i \ ?U \ f i = i" "\(\y. \!x. f x = y)" unfolding permutes_def by auto let ?A = "mat\<^sub>r n n (\ i. A $$ (i, f i) \\<^sub>v row B (f i))" let ?B = "mat\<^sub>r n n (\ i. row B (f i))" have B': "?B \ carrier_mat n n" by (intro mat_row_carrierI) { assume fi: "inj_on f ?U" from inj_on_nat_permutes[OF fi] f have "f permutes ?U" by auto with fPU have False by simp } hence fni: "\ inj_on f ?U" by auto then obtain i j where ij: "f i = f j" "i \ j" "i < n" "j < n" unfolding inj_on_def by auto from ij have rth: "row ?B i = row ?B j" by auto have "det ?A = 0" by (subst det_rows_mul, unfold det_identical_rows[OF B' ij(2-4) rth], insert f A B, auto) } then have zth: "\ f. f \ ?F - ?PU \ det (mat\<^sub>r n n (\ i. A $$ (i, f i) \\<^sub>v row B (f i))) = 0" by simp { fix p assume pU: "p \ ?PU" from pU have p: "p permutes ?U" by blast let ?s = "\p. (signof p) :: 'a" let ?f = "\q. ?s p * (\ i\ ?U. A $$ (i,p i)) * (?s q * (\i\ ?U. B $$ (i, q i)))" have "(sum (\q. ?s q * (\i\ ?U. mat\<^sub>r n n (\ i. A $$ (i, p i) \\<^sub>v row B (p i)) $$ (i, q i))) ?PU) = (sum (\q. ?s p * (\ i\ ?U. A $$ (i,p i)) * (?s q * (\ i\ ?U. B $$ (i, q i)))) ?PU)" unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] proof (rule sum.cong[OF refl]) fix q assume "q \ {q. q permutes ?U}" hence q: "q permutes ?U" by simp from p q have pp: "permutation p" and pq: "permutation q" unfolding permutation_permutes by auto note sign = signof_compose[OF q permutes_inv[OF p], unfolded signof_inv[OF fU p]] let ?inv = "Hilbert_Choice.inv" have th001: "prod (\i. B$$ (i, q (?inv p i))) ?U = prod ((\i. B$$ (i, q (?inv p i))) \ p) ?U" by (rule prod.permute[OF p]) have thp: "prod (\i. mat\<^sub>r n n (\ i. A$$(i,p i) \\<^sub>v row B (p i)) $$ (i, q i)) ?U = prod (\i. A$$(i,p i)) ?U * prod (\i. B$$ (i, q (?inv p i))) ?U" unfolding th001 o_def permutes_inverses[OF p] by (subst prod.distrib[symmetric], insert A p q B, auto intro: prod.cong) define AA where "AA = (\i\?U. A $$ (i, p i))" define BB where "BB = (\ia\{0..ia\{0..r n n (\i. A $$ (i, p i) \\<^sub>v row B (p i)) $$ (ia, q ia)) = ?s p * (\i\{0.. ?inv p) * (\ia\{0..i = 0..r n n (\i. A $$ (i, p i) \\<^sub>v row B (p i)) $$ (i, q i)) = ?s p * (\i = 0.. ?inv p) * (\i = 0.. ?inv p) i)))" by simp qed } note * = this have th2: "sum (\f. det (mat\<^sub>r n n (\ i. A$$(i,f i) \\<^sub>v row B (f i)))) ?PU = det A * det B" unfolding det_def'[OF A] det_def'[OF B] det_def'[OF mat_row_carrierI] unfolding sum_product dim_row_mat by (rule sum.cong, insert A, force, subst *, insert A B, auto) let ?f = "\ f. det (mat\<^sub>r n n (\ i. A $$ (i, f i) \\<^sub>v row B (f i)))" have "det (A * B) = sum ?f ?F" unfolding mat_mul_finsum_alt[OF A B] by (rule det_linear_rows_sum[OF fU], insert A B, auto) also have "\ = sum ?f ((?F - ?PU) \ (?F \ ?PU))" by (rule arg_cong[where f = "sum ?f"], auto) also have "\ = sum ?f (?F - ?PU) + sum ?f (?F \ ?PU)" by (rule sum.union_disjoint, insert A B finite_bounded_functions[OF fU fU], auto) also have "sum ?f (?F - ?PU) = 0" by (rule sum.neutral, insert zth, auto) also have "?F \ ?PU = ?PU" unfolding permutes_def by fastforce also have "sum ?f ?PU = det A * det B" unfolding th2 .. finally show ?thesis by simp qed lemma unit_imp_det_non_zero: assumes "A \ Units (ring_mat TYPE('a :: comm_ring_1) n b)" shows "det A \ 0" proof - from assms[unfolded Units_def ring_mat_def] obtain B where A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and BA: "B * A = 1\<^sub>m n" by auto from arg_cong[OF BA, of det, unfolded det_mult[OF B A] det_one] show ?thesis by auto qed text \The following proof is based on the Gauss-Jordan algorithm.\ lemma det_non_zero_imp_unit: assumes A: "A \ carrier_mat n n" and dA: "det A \ (0 :: 'a :: field)" shows "A \ Units (ring_mat TYPE('a) n b)" proof (rule ccontr) let ?g = "gauss_jordan A (0\<^sub>m n 0)" let ?B = "fst ?g" obtain B C where B: "?g = (B,C)" by (cases ?g) assume "\ ?thesis" from this[unfolded gauss_jordan_check_invertable[OF A zero_carrier_mat[of n 0]] B] have "B \ 1\<^sub>m n" by auto with row_echelon_form_imp_1_or_0_row[OF gauss_jordan_carrier(1)[OF A _ B] gauss_jordan_row_echelon[OF A B], of 0] have n: "0 < n" and row: "row B (n - 1) = 0\<^sub>v n" by auto let ?n = "n - 1" from n have n1: "?n < n" by auto from gauss_jordan_transform[OF A _ B, of 0 b] obtain P where P: "P\Units (ring_mat TYPE('a) n b)" and PA: "B = P * A" by auto from unit_imp_det_non_zero[OF P] have dP: "det P \ 0" by auto from P have P: "P \ carrier_mat n n" unfolding Units_def ring_mat_def by auto from det_mult[OF P A] dP dA have "det B \ 0" unfolding PA by simp also have "det B = 0" proof - from gauss_jordan_carrier[OF A _ B, of 0] have B: "B \ carrier_mat n n" by auto { fix j assume j: "j < n" from index_row(1)[symmetric, of ?n B j, unfolded row] B have "B $$ (?n, j) = 0" using B n j by auto } hence "B = mat\<^sub>r n n (\i. if i = ?n then 0\<^sub>v n else row B i)" by (intro eq_matI, insert B, auto) also have "det \ = 0" by (rule det_row_0[OF n1], insert B, auto) finally show "det B = 0" . qed finally show False by simp qed lemma mat_mult_left_right_inverse: assumes A: "(A :: 'a :: field mat) \ carrier_mat n n" and B: "B \ carrier_mat n n" and AB: "A * B = 1\<^sub>m n" shows "B * A = 1\<^sub>m n" proof - let ?R = "ring_mat TYPE('a) n undefined" from det_mult[OF A B, unfolded AB] have "det A \ 0" "det B \ 0" by auto from det_non_zero_imp_unit[OF A this(1)] det_non_zero_imp_unit[OF B this(2)] have U: "A \ Units ?R" "B \ Units ?R" . interpret ring ?R by (rule ring_mat) from Units_inv_comm[unfolded ring_mat_simps, OF AB U] show ?thesis . qed lemma det_zero_imp_zero_row: assumes A: "(A :: 'a :: field mat) \ carrier_mat n n" and det: "det A = 0" shows "\ P. P \ Units (ring_mat TYPE('a) n b) \ row (P * A) (n - 1) = 0\<^sub>v n \ 0 < n \ row_echelon_form (P * A)" proof - let ?R = "ring_mat TYPE('a) n b" let ?U = "Units ?R" interpret m: ring ?R by (rule ring_mat) let ?g = "gauss_jordan A A" obtain A' B' where g: "?g = (A', B')" by (cases ?g) from det unit_imp_det_non_zero[of A n b] have AU: "A \ ?U" by auto with gauss_jordan_inverse_one_direction(1)[OF A A, of _ b] have A'1: "A' \ 1\<^sub>m n" using g by auto from gauss_jordan_carrier(1)[OF A A g] have A': "A' \ carrier_mat n n" by auto from gauss_jordan_row_echelon[OF A g] have re: "row_echelon_form A'" . from row_echelon_form_imp_1_or_0_row[OF A' this] A'1 have n: "0 < n" and row: "row A' (n - 1) = 0\<^sub>v n" by auto from gauss_jordan_transform[OF A A g, of b] obtain P where P: "P \ ?U" and A': "A' = P * A" by auto thus ?thesis using n row re by auto qed lemma det_0_iff_vec_prod_zero_field: assumes A: "(A :: 'a :: field mat) \ carrier_mat n n" shows "det A = 0 \ (\ v. v \ carrier_vec n \ v \ 0\<^sub>v n \ A *\<^sub>v v = 0\<^sub>v n)" (is "?l = (\ v. ?P v)") proof - let ?R = "ring_mat TYPE('a) n ()" let ?U = "Units ?R" interpret m: ring ?R by (rule ring_mat) show ?thesis proof (cases "det A = 0") case False from det_non_zero_imp_unit[OF A this, of "()"] have "A \ ?U" . then obtain B where unit: "B * A = 1\<^sub>m n" and B: "B \ carrier_mat n n" unfolding Units_def ring_mat_def by auto { fix v assume "?P v" hence v: "v \ carrier_vec n" "v \ 0\<^sub>v n" "A *\<^sub>v v = 0\<^sub>v n" by auto have "v = (B * A) *\<^sub>v v" using v B unfolding unit by auto also have "\ = B *\<^sub>v (A *\<^sub>v v)" using B A v by simp also have "\ = B *\<^sub>v 0\<^sub>v n" unfolding v .. also have "\ = 0\<^sub>v n" using B by auto finally have False using v by simp } with False show ?thesis by blast next case True let ?n = "n - 1" from det_zero_imp_zero_row[OF A True, of "()"] obtain P where PU: "P \ ?U" and row: "row (P * A) ?n = 0\<^sub>v n" and n: "0 < n" "?n < n" and re: "row_echelon_form (P * A)" by auto define PA where "PA = P * A" note row = row[folded PA_def] note re = re[folded PA_def] from PU obtain Q where P: "P \ carrier_mat n n" and Q: "Q \ carrier_mat n n" and unit: "Q * P = 1\<^sub>m n" "P * Q = 1\<^sub>m n" unfolding Units_def ring_mat_def by auto from P A have PA: "PA \ carrier_mat n n" and dimPA: "dim_row PA = n" unfolding PA_def by auto from re[unfolded row_echelon_form_def] obtain p where p: "pivot_fun PA p n" using PA by auto note piv = pivot_positions[OF PA p] note pivot = pivot_funD[OF dimPA p n(2)] { assume "p ?n < n" with pivot(4)[OF this] n arg_cong[OF row, of "\ v. v $ p ?n"] have False using PA by auto } with pivot(1) have pn: "p ?n = n" by fastforce with piv(1) have "set (pivot_positions PA) \ {(i, p i) |i. i < n \ p i \ n} - {(?n,p ?n)}" by auto also have "\ \ {(i, p i) | i. i < ?n}" using n by force finally have "card (set (pivot_positions PA)) \ card {(i, p i) | i. i < ?n}" by (intro card_mono, auto) also have "{(i, p i) | i. i < ?n} = (\ i. (i, p i)) ` {0 ..< ?n}" by auto also have "card \ = card {0 ..< ?n}" by (rule card_image, auto simp: inj_on_def) also have "\ < n" using n by simp finally have "card (set (pivot_positions PA)) < n" . hence "card (snd ` (set (pivot_positions PA))) < n" using card_image_le[OF finite_set, of snd "pivot_positions PA"] by auto hence neq: "snd ` (set (pivot_positions PA)) \ {0 ..< n}" by auto from find_base_vector[OF re PA neq] obtain v where v: "v \ carrier_vec n" and v0: "v \ 0\<^sub>v n" and pav: "PA *\<^sub>v v = 0\<^sub>v n" by auto have "A *\<^sub>v v = Q * P *\<^sub>v (A *\<^sub>v v)" unfolding unit using A v by auto also have "\ = Q *\<^sub>v (PA *\<^sub>v v)" unfolding PA_def using Q P A v by auto also have "PA *\<^sub>v v = 0\<^sub>v n" unfolding pav .. also have "Q *\<^sub>v 0\<^sub>v n = 0\<^sub>v n" using Q by auto finally have Av: "A *\<^sub>v v = 0\<^sub>v n" by auto show ?thesis unfolding True using Av v0 v by auto qed qed text \In order to get the result for integral domains, we embed the domain in its fraction field, and then apply the result for fields.\ lemma det_0_iff_vec_prod_zero: assumes A: "(A :: 'a :: idom mat) \ carrier_mat n n" shows "det A = 0 \ (\ v. v \ carrier_vec n \ v \ 0\<^sub>v n \ A *\<^sub>v v = 0\<^sub>v n)" proof - let ?h = "to_fract :: 'a \ 'a fract" let ?A = "map_mat ?h A" have A': "?A \ carrier_mat n n" using A by auto interpret inj_comm_ring_hom ?h by (unfold_locales, auto) have "(det A = 0) = (?h (det A) = ?h 0)" by auto also have "\ = (det ?A = 0)" unfolding hom_zero hom_det .. also have "\ = ((\ v. v \ carrier_vec n \ v \ 0\<^sub>v n \ ?A *\<^sub>v v = 0\<^sub>v n))" unfolding det_0_iff_vec_prod_zero_field[OF A'] .. also have "\ = ((\ v. v \ carrier_vec n \ v \ 0\<^sub>v n \ A *\<^sub>v v = 0\<^sub>v n))" (is "?l = ?r") proof assume ?r then obtain v where v: "v \ carrier_vec n" "v \ 0\<^sub>v n" "A *\<^sub>v v = 0\<^sub>v n" by auto show ?l by (rule exI[of _ "map_vec ?h v"], insert v, auto simp: mult_mat_vec_hom[symmetric, OF A v(1)]) next assume ?l then obtain v where v: "v \ carrier_vec n" and v0: "v \ 0\<^sub>v n" and Av: "?A *\<^sub>v v = 0\<^sub>v n" by auto have "\ i. \ a b. v $ i = Fraction_Field.Fract a b \ b \ 0" using Fract_cases[of "v $ i" for i] by metis from choice[OF this] obtain a where "\ i. \ b. v $ i = Fraction_Field.Fract (a i) b \ b \ 0" by metis from choice[OF this] obtain b where vi: "\ i. v $ i = Fraction_Field.Fract (a i) (b i)" and bi: "\ i. b i \ 0" by auto define m where "m = prod_list (map b [0.. 0" unfolding m_def hom_0_iff prod_list_zero_iff using bi by auto from v0[unfolded vec_eq_iff] v obtain i where i: "i < n" "v $ i \ 0" by auto { fix i assume "i < n" hence "b i \ set (map b [0 ..< n])" by auto from split_list[OF this] obtain ys zs where "map b [0.. c. ?m * v $ i = ?h c" .. } hence "\ i. \ c. i < n \ ?m * v $ i = ?h c" by auto from choice[OF this] obtain c where c: "\ i. i < n \ ?m * v $ i = ?h (c i)" by auto define w where "w = vec n c" have w: "w \ carrier_vec n" unfolding w_def by simp have mvw: "?m \\<^sub>v v = map_vec ?h w" unfolding w_def using c v by (intro eq_vecI, auto) with m0 i c[OF i(1)] have "w $ i \ 0" unfolding w_def by auto with i w have w0: "w \ 0\<^sub>v n" by auto from arg_cong[OF Av, of "\ v. ?m \\<^sub>v v"] have "?m \\<^sub>v (?A *\<^sub>v v) = map_vec ?h (0\<^sub>v n)" by auto also have "?m \\<^sub>v (?A *\<^sub>v v) = ?A *\<^sub>v (?m \\<^sub>v v)" using A v by auto also have "\ = ?A *\<^sub>v (map_vec ?h w)" unfolding mvw .. also have "\ = map_vec ?h (A *\<^sub>v w)" unfolding mult_mat_vec_hom[OF A w] .. finally have "A *\<^sub>v w = 0\<^sub>v n" by (rule vec_hom_inj) with w w0 show ?r by blast qed finally show ?thesis . qed lemma det_0_negate: assumes A: "(A :: 'a :: field mat) \ carrier_mat n n" shows "(det (- A) = 0) = (det A = 0)" proof - from A have mA: "- A \ carrier_mat n n" by auto { fix v :: "'a vec" assume v: "v \ carrier_vec n" hence Av: "A *\<^sub>v v \ carrier_vec n" using A by auto have id: "- A *\<^sub>v v = - (A *\<^sub>v v)" using v A by simp have "(- A *\<^sub>v v = 0\<^sub>v n) = (A *\<^sub>v v = 0\<^sub>v n)" unfolding id unfolding uminus_zero_vec_eq[OF Av] .. } thus ?thesis unfolding det_0_iff_vec_prod_zero[OF A] det_0_iff_vec_prod_zero[OF mA] by auto qed lemma det_multrow: assumes k: "k < n" and A: "A \ carrier_mat n n" shows "det (multrow k a A) = a * det A" proof - have "multrow k a A = multrow_mat n k a * A" by (rule multrow_mat[OF A]) also have "det (multrow_mat n k a * A) = det (multrow_mat n k a) * det A" by (rule det_mult[OF _ A], auto) also have "det (multrow_mat n k a) = a" by (rule det_multrow_mat[OF k]) finally show ?thesis . qed lemma det_multrow_div: assumes k: "k < n" and A: "A \ carrier_mat n n" and a0: "a \ 0" shows "det (multrow k a A :: 'a :: idom_divide mat) div a = det A" proof - have "det (multrow k a A) div a = a * det A div a" using k A by (simp add: det_multrow) also have "... = det A" using a0 by auto finally show ?thesis. qed lemma det_addrow: assumes l: "l < n" and k: "k \ l" and A: "A \ carrier_mat n n" shows "det (addrow a k l A) = det A" proof - have "addrow a k l A = addrow_mat n a k l * A" by (rule addrow_mat[OF A l]) also have "det (addrow_mat n a k l * A) = det (addrow_mat n a k l) * det A" by (rule det_mult[OF _ A], auto) also have "det (addrow_mat n a k l) = 1" by (rule det_addrow_mat[OF k]) finally show ?thesis using A by simp qed lemma det_swaprows: assumes *: "k < n" "l < n" and k: "k \ l" and A: "A \ carrier_mat n n" shows "det (swaprows k l A) = - det A" proof - have "swaprows k l A = swaprows_mat n k l * A" by (rule swaprows_mat[OF A *]) also have "det (swaprows_mat n k l * A) = det (swaprows_mat n k l) * det A" by (rule det_mult[OF _ A], insert A, auto) also have "det (swaprows_mat n k l) = - 1" by (rule det_swaprows_mat[OF * k]) finally show ?thesis using A by simp qed lemma det_similar: assumes "similar_mat A B" shows "det A = det B" proof - from similar_matD[OF assms] obtain n P Q where carr: "{A, B, P, Q} \ carrier_mat n n" (is "_ \ ?C") and PQ: "P * Q = 1\<^sub>m n" and AB: "A = P * B * Q" by blast hence A: "A \ ?C" and B: "B \ ?C" and P: "P \ ?C" and Q: "Q \ ?C" by auto from det_mult[OF P Q, unfolded PQ] have PQ: "det P * det Q = 1" by auto from det_mult[OF _ Q, of "P * B", unfolded det_mult[OF P B] AB[symmetric]] P B have "det A = det P * det B * det Q" by auto also have "\ = (det P * det Q) * det B" by (simp add: ac_simps) also have "\ = det B" unfolding PQ by simp finally show ?thesis . qed lemma det_four_block_mat_upper_right_zero_col: assumes A1: "A1 \ carrier_mat n n" and A20: "A2 = (0\<^sub>m n 1)" and A3: "A3 \ carrier_mat 1 n" and A4: "A4 \ carrier_mat 1 1" shows "det (four_block_mat A1 A2 A3 A4) = det A1 * det A4" (is "det ?A = _") proof - let ?A = "four_block_mat A1 A2 A3 A4" from A20 have A2: "A2 \ carrier_mat n 1" by auto define A where "A = ?A" from four_block_carrier_mat[OF A1 A4] A1 have A: "A \ carrier_mat (Suc n) (Suc n)" and dim: "dim_row A1 = n" unfolding A_def by auto let ?Pn = "\ p. p permutes {0 ..< n}" let ?Psn = "\ p. p permutes {0 ..< Suc n}" let ?perm = "{p. ?Psn p}" let ?permn = "{p. ?Pn p}" let ?prod = "\ p. signof p * (\i = 0.. i \ {0..< n}. A $$ (p i, i)))" by (subst prod.remove[of _ n], auto) also have "\ = A $$ (p n, n) * signof p * (\ i \ {0..< n}. A $$ (p i, i))" by simp finally have "?prod p = ?prod' p" . } note prod_id = this define prod' where "prod' = ?prod'" { fix i q assume i: "i \ {0..< n}" "q permutes {0 ..< n}" hence "Fun.swap n i id (q n) < n" unfolding permutes_def by auto hence "A $$ (Fun.swap n i id (q n), n) = 0" unfolding A_def using A1 A20 A3 A4 by auto hence "prod' (Fun.swap n i id \ q) = 0" unfolding prod'_def by simp } note zero = this have cong: "\ a b c. b = c \ a * b = a * c" by auto have "det ?A = sum ?prod ?perm" unfolding A_def[symmetric] using mat_det_left_def[OF A] A by simp also have "\ = sum prod' ?perm" unfolding prod'_def by (rule sum.cong[OF refl], insert prod_id, auto) also have "{0 ..< Suc n} = insert n {0 ..< n}" by auto also have "sum prod' {p. p permutes \} = (\i\insert n {0..q\?permn. prod' (Fun.swap n i id \ q))" by (subst sum_over_permutations_insert, auto) also have "\ = (\q\?permn. prod' q) + (\i\{0..q\?permn. prod' (Fun.swap n i id \ q))" by (subst sum.insert, auto) also have "(\i\{0..q\?permn. prod' (Fun.swap n i id \ q)) = 0" by (rule sum.neutral, intro ballI, rule sum.neutral, intro ballI, rule zero, auto) also have "(\q\ ?permn. prod' q) = A $$ (n,n) * (\q\ ?permn. ?prod'' q)" unfolding prod'_def by (subst sum_distrib_left, rule sum.cong[OF refl], auto simp: permutes_def ac_simps) also have "A $$ (n,n) = A4 $$ (0,0)" unfolding A_def using A1 A2 A3 A4 by auto also have "(\q\ ?permn. ?prod'' q) = (\q\ ?permn. ?prod''' q)" by (rule sum.cong[OF refl], rule cong, rule prod.cong, insert A1 A2 A3 A4, auto simp: permutes_def A_def) also have "\ = det A1" unfolding mat_det_left_def[OF A1] dim by auto also have "A4 $$ (0,0) = det A4" using A4 unfolding det_def[of A4] by (auto simp: signof_def sign_def) finally show ?thesis by simp qed lemma det_swap_initial_rows: assumes A: "A \ carrier_mat m m" and lt: "k + n \ m" shows "det A = (- 1) ^ (k * n) * det (mat m m (\(i, j). A $$ (if i < n then i + k else if i < k + n then i - n else i, j)))" proof - define sw where "sw = (\ (A :: 'a mat) xs. fold (\ (i,j). swaprows i j) xs A)" have dim_sw[simp]: "dim_row (sw A xs) = dim_row A" "dim_col (sw A xs) = dim_col A" for xs A unfolding sw_def by (induct xs arbitrary: A, auto) { fix xs and A :: "'a mat" assume "dim_row A = dim_col A" "\ i j. (i,j) \ set xs \ i < dim_col A \ j < dim_col A \ i \ j" hence "det (sw A xs) = (-1)^(length xs) * det A" unfolding sw_def proof (induct xs arbitrary: A) case (Cons xy xs A) obtain x y where xy: "xy = (x,y)" by force from Cons(3)[unfolded xy, of x y] Cons(2) have [simp]: "det (swaprows x y A) = - det A" by (intro det_swaprows, auto) show ?case unfolding xy by (simp, insert Cons(2-), (subst Cons(1), auto)+) qed simp } note sw = this define swb where "swb = (\ A i n. sw A (map (\ j. (j,Suc j)) [i ..< i + n]))" { fix k n and A :: "'a mat" assume k_n: "k + n < dim_row A" hence "swb A k n = mat (dim_row A) (dim_col A) (\ (i,j). let r = (if i < k \ i > k + n then i else if i = k + n then k else Suc i) in A $$ (r,j))" proof (induct n) case 0 show ?case unfolding swb_def sw_def by (rule eq_matI, auto) next case (Suc n) hence dim: "k + n < dim_row A" by auto have id: "swb A k (Suc n) = swaprows (k + n) (Suc k + n) (swb A k n)" unfolding swb_def sw_def by simp show ?case unfolding id Suc(1)[OF dim] by (rule eq_matI, insert Suc(2), auto) qed } note swb = this define swbl where "swbl = (\ A k n. fold (\ i A. swb A i n) (rev [0 ..< k]) A)" { fix k n and A :: "'a mat" assume k_n: "k + n \ dim_row A" hence "swbl A k n = mat (dim_row A) (dim_col A) (\ (i,j). let r = (if i < n then i + k else if i < k + n then i - n else i) in A $$ (r,j))" proof (induct k arbitrary: A) case 0 thus ?case unfolding swbl_def by (intro eq_matI, auto simp: swb) next case (Suc k) hence dim: "k + n < dim_row A" by auto have id: "swbl A (Suc k) n = swbl (swb A k n) k n" unfolding swbl_def by simp show ?case unfolding id swb[OF dim] by (subst Suc(1), insert dim, force, intro eq_matI, auto simp: less_Suc_eq_le) qed } note swbl = this { fix k n and A :: "'a mat" assume k_n: "k + n \ dim_col A" "dim_row A = dim_col A" hence "det (swbl A k n) = (-1)^(k*n) * det A" proof (induct k arbitrary: A) case 0 thus ?case unfolding swbl_def by auto next case (Suc k) hence dim: "k + n < dim_row A" by auto have id: "swbl A (Suc k) n = swbl (swb A k n) k n" unfolding swbl_def by simp have det: "det (swb A k n) = (-1)^n * det A" unfolding swb_def by (subst sw, insert Suc(2-), auto) show ?case unfolding id by (subst Suc(1), insert Suc(2-), auto simp: det, auto simp: swb power_add) qed } note det_swbl = this from assms have dim: "dim_row A = dim_col A" "k + n \ dim_col A" "k + n \ dim_row A" "dim_col A = m" by auto from arg_cong[OF det_swbl[OF dim(2,1), unfolded swbl[OF dim(3)], unfolded Let_def dim], of "\ x. (-1)^(k*n) * x"] show ?thesis by simp qed lemma det_swap_rows: assumes A: "A \ carrier_mat (k + n) (k + n)" shows "det A = (-1)^(k * n) * det (mat (k + n) (k + n) (\ (i,j). A $$ ((if i < k then i + n else i - k),j)))" proof - have le: "n + k \ k + n" by simp show ?thesis unfolding det_swap_initial_rows[OF A le] by (intro arg_cong2[of _ _ _ _ "\ x y. ((-1)^x * det y)"], force, intro eq_matI, auto) qed lemma det_swap_final_rows: assumes A: "A \ carrier_mat m m" and m: "m = l + k + n" shows "det A = (- 1) ^ (k * n) * det (mat m m (\(i, j). A $$ (if i < l then i else if i < l + n then i + k else i - n, j)))" (is "_ = _ * det ?M") proof - (* l k n -swap-rows\ k n l -swap-initial\ n k l -swap-rows\ l n k *) have m1: "m = (k + n) + l" using m by simp have m2: "k + n \ m" using m by simp have m3: "m = l + (n + k)" using m by simp define M where "M = ?M" let ?M1 = "mat m m (\(i, j). A $$ (if i < k + n then i + l else i - (k + n), j))" let ?M2 = "mat m m (\(i, j). A $$ (if i < n then i + k + l else if i < k + n then i - n + l else i - (k + n), j))" have M2: "?M2 \ carrier_mat m m" by auto have "det A = (- 1) ^ ((k + n) * l) * det ?M1" unfolding det_swap_rows[OF A[unfolded m1]] m1[symmetric] by simp also have "det ?M1 = (- 1) ^ (k * n) * det ?M2" by (subst det_swap_initial_rows[OF _ m2], force, rule arg_cong[of _ _ "\ x. _ * det x"], rule eq_matI, auto simp: m) also have "det ?M2 = (- 1) ^ (l * (n + k)) * det M" unfolding M_def det_swap_rows[OF M2[unfolded m3], folded m3] by (rule arg_cong[of _ _ "\ x. _ * det x"], rule eq_matI, auto simp: m) finally have "det A = (-1) ^ ((k + n) * l + (k * n) + l * (n + k)) * det M" (is "_ = ?b ^ _ * _") by (simp add: power_add) also have "(k + n) * l + (k * n) + l * (n + k) = 2 * (l * (n + k)) + k * n" by simp also have "?b ^ \ = ?b ^ (k * n)" by (simp add: power_add) finally show ?thesis unfolding M_def . qed lemma det_swap_final_cols: assumes A: "A \ carrier_mat m m" and m: "m = l + k + n" shows "det A = (- 1) ^ (k * n) * det (mat m m (\(i, j). A $$ (i, if j < l then j else if j < l + n then j + k else j - n)))" proof - have "det A = det (A\<^sup>T)" unfolding det_transpose[OF A] .. also have "\ = (- 1) ^ (k * n) * det (mat m m (\(i, j). A\<^sup>T $$ (if i < l then i else if i < l + n then i + k else i - n, j)))" (is "_ = _ * det ?M") by (rule det_swap_final_rows[OF _ m], insert A, auto) also have "det ?M = det (?M\<^sup>T)" by (subst det_transpose, auto) also have "?M\<^sup>T = mat m m (\(i, j). A $$ (i, if j < l then j else if j < l + n then j + k else j - n))" unfolding transpose_mat_def using A m by (intro eq_matI, auto) finally show ?thesis . qed lemma det_swap_initial_cols: assumes A: "A \ carrier_mat m m" and lt: "k + n \ m" shows "det A = (- 1) ^ (k * n) * det (mat m m (\(i, j). A $$ (i, if j < n then j + k else if j < k + n then j - n else j)))" proof - have "det A = det (A\<^sup>T)" unfolding det_transpose[OF A] .. also have "\ = (- 1) ^ (k * n) * det (mat m m (\(j, i). A\<^sup>T $$ (if j < n then j + k else if j < k + n then j - n else j,i)))" (is "_ = _ * det ?M") by (rule det_swap_initial_rows[OF _ lt], insert A, auto) also have "det ?M = det (?M\<^sup>T)" by (subst det_transpose, auto) also have "?M\<^sup>T = mat m m (\(i, j). A $$ (i, if j < n then j + k else if j < k + n then j - n else j))" unfolding transpose_mat_def using A lt by (intro eq_matI, auto) finally show ?thesis . qed lemma det_swap_cols: assumes A: "A \ carrier_mat (k + n) (k + n)" shows "det A = (-1)^(k * n) * det (mat (k + n) (k + n) (\ (i,j). A $$ (i,(if j < k then j + n else j - k))))" (is "_ = _ * det ?B") proof - have le: "n + k \ k + n" by simp show ?thesis unfolding det_swap_initial_cols[OF A le] by (intro arg_cong2[of _ _ _ _ "\ x y. ((-1)^x * det y)"], force, intro eq_matI, auto) qed lemma det_four_block_mat_upper_right_zero: fixes A1 :: "'a :: idom mat" assumes A1: "A1 \ carrier_mat n n" and A20: "A2 = (0\<^sub>m n m)" and A3: "A3 \ carrier_mat m n" and A4: "A4 \ carrier_mat m m" shows "det (four_block_mat A1 A2 A3 A4) = det A1 * det A4" using assms(2-) proof (induct m arbitrary: A2 A3 A4) case (0 A2 A3 A4) hence *: "four_block_mat A1 A2 A3 A4 = A1" using A1 by (intro eq_matI, auto) from 0 have 4: "A4 = 1\<^sub>m 0" by auto show ?case unfolding * unfolding 4 by simp next case (Suc m A2 A3 A4) let ?m = "Suc m" from Suc have A2: "A2 \ carrier_mat n ?m" by auto note A20 = Suc(2) note A34 = Suc(3-4) let ?A = "four_block_mat A1 A2 A3 A4" let ?P = "\ B3 B4 v k. v \ 0 \ v * det ?A = det (four_block_mat A1 A2 B3 B4) \ v * det A4 = det B4 \ B3 \ carrier_mat ?m n \ B4 \ carrier_mat ?m ?m \ (\ i < k. B4 $$ (i,m) = 0)" have "k \ m \ \ B3 B4 v. ?P B3 B4 v k" for k proof (induct k) case 0 have "?P A3 A4 1 0" using A34 by auto thus ?case by blast next case (Suc k) then obtain B3 B4 v where v: "v \ 0" and det: "v * det ?A = det (four_block_mat A1 A2 B3 B4)" "v * det A4 = det B4" and B3: "B3 \ carrier_mat ?m n" and B4: "B4 \ carrier_mat ?m ?m" and 0: "\ i < k. B4 $$ (i,m) = 0" by auto show ?case proof (cases "B4 $$ (k,m) = 0") case True with 0 have 0: "\ i < Suc k. B4 $$ (i,m) = 0" using less_Suc_eq by auto with v det B3 B4 have "?P B3 B4 v (Suc k)" by auto thus ?thesis by blast next case Bk: False let ?k = "Suc k" from Suc(2) have k: "k < ?m" "Suc k < ?m" "k \ Suc k" by auto show ?thesis proof (cases "B4 $$ (?k,m) = 0") case True let ?B4 = "swaprows k (Suc k) B4" let ?B3 = "swaprows k (Suc k) B3" let ?B = "four_block_mat A1 A2 ?B3 ?B4" let ?v = "-v" from det_swaprows[OF k B4] det have det1: "?v * det A4 = det ?B4" by simp from v have v: "?v \ 0" by auto from B3 have B3': "?B3 \ carrier_mat ?m n" by auto from B4 have B4': "?B4 \ carrier_mat ?m ?m" by auto have "?v * det ?A = - det (four_block_mat A1 A2 B3 B4)" using det by simp also have "\ = det (swaprows (n + k) (n + ?k) (four_block_mat A1 A2 B3 B4))" by (rule sym, rule det_swaprows[of _ "n + ?m"], insert A1 A2 B3 B4 k, auto) also have "swaprows (n + k) (n + ?k) (four_block_mat A1 A2 B3 B4) = ?B" proof (rule eq_matI, unfold index_mat_four_block index_mat_swaprows, goal_cases) case (1 i j) show ?case proof (cases "i < n") case True thus ?thesis using 1(2) A1 A2 B3 B4 by simp next case False hence "i = n + (i - n)" by simp then obtain d where "i = n + d" by blast thus ?thesis using 1 A1 A2 B3 B4 k(2) by simp qed qed auto finally have det2: "?v * det ?A = det ?B" . from True 0 B4 k(2) have "\ i < Suc k. ?B4 $$ (i,m) = 0" unfolding less_Suc_eq by auto with det1 det2 B3' B4' v have "?P ?B3 ?B4 ?v (Suc k)" by auto thus ?thesis by blast next case False let ?bk = "B4 $$ (?k,m)" let ?b = "B4 $$ (k,m)" let ?v = "v * ?bk" let ?B3 = "addrow (- ?b) k ?k (multrow k ?bk B3)" let ?B4 = "addrow (- ?b) k ?k (multrow k ?bk B4)" have *: "det ?B4 = ?bk * det B4" by (subst det_addrow[OF k(2-3)], force simp: B4, rule det_multrow[OF k(1) B4]) with det(2)[symmetric] have det2: "?v * det A4 = det ?B4" by (auto simp: ac_simps) from 0 k(2) B4 have 0: "\ i < Suc k. ?B4 $$ (i,m) = 0" unfolding less_Suc_eq by auto from False v have v: "?v \ 0" by auto from B3 have B3': "?B3 \ carrier_mat ?m n" by auto from B4 have B4': "?B4 \ carrier_mat ?m ?m" by auto let ?B' = "multrow (n + k) ?bk (four_block_mat A1 A2 B3 B4)" have B': "?B' \ carrier_mat (n + ?m) (n + ?m)" using A1 A2 B3 B4 k by auto let ?B = "four_block_mat A1 A2 ?B3 ?B4" have "?v * det ?A = ?bk * det (four_block_mat A1 A2 B3 B4)" using det by simp also have "\ = det (addrow (- ?b) (n + k) (n + ?k) ?B')" by (subst det_addrow[OF _ _ B'], insert k(2), force, force, rule sym, rule det_multrow[of _ "n + ?m"], insert A1 A2 B3 B4 k, auto) also have "addrow (- ?b) (n + k) (n + ?k) ?B' = ?B" proof (rule eq_matI, unfold index_mat_four_block index_mat_multrow index_mat_addrow, goal_cases) case (1 i j) show ?case proof (cases "i < n") case True thus ?thesis using 1(2) A1 A2 B3 B4 by simp next case False hence "i = n + (i - n)" by simp then obtain d where "i = n + d" by blast thus ?thesis using 1 A1 A2 B3 B4 k(2) by simp qed qed auto finally have det1: "?v * det ?A = det ?B" . from det1 det2 B3' B4' v 0 have "?P ?B3 ?B4 ?v (Suc k)" by auto thus ?thesis by blast qed qed qed from this[OF le_refl] obtain B3 B4 v where P: "?P B3 B4 v m" by blast let ?B = "four_block_mat A1 A2 B3 B4" from P have v: "v \ 0" and det: "v * det ?A = det ?B" "v * det A4 = det B4" and B3: "B3 \ carrier_mat ?m n" and B4: "B4 \ carrier_mat ?m ?m" and 0: "\ i. i < m \ B4 $$ (i, m) = 0" by auto let ?A2 = "0\<^sub>m n m" let ?A3 = "mat m n (\ ij. B3 $$ ij)" let ?A4 = "mat m m (\ ij. B4 $$ ij)" let ?B1 = "four_block_mat A1 ?A2 ?A3 ?A4" let ?B2 = "0\<^sub>m (n + m) 1" let ?B3 = "mat 1 (n + m) (\ (i,j). if j < n then B3 $$ (m,j) else B4 $$ (m,j - n))" let ?B4 = "mat 1 1 (\ _. B4 $$ (m,m))" have B44: "B4 = four_block_mat ?A4 (0\<^sub>m m 1) (mat 1 m (\ (i,j). B4 $$ (m,j))) ?B4" proof (rule eq_matI, unfold index_mat_four_block dim_col_mat dim_row_mat, goal_cases) case (1 i j) hence [simp]: "\ i < m \ i = m" "\ j < m \ j = m" by auto from 1 show ?case using B4 0 by auto qed (insert B4, auto) have "?B = four_block_mat ?B1 ?B2 ?B3 ?B4" proof (rule eq_matI, unfold index_mat_four_block dim_col_mat dim_row_mat, goal_cases) case (1 i j) then consider (UL) "i < n + m" "j < n + m" | (UR) "i < n + m" "j = n + m" | (LL) "i = n + m" "j < n + m" | (LR) "i = n + m" "j = n + m" using A1 by auto linarith thus ?case proof cases case UL hence [simp]: "\ i < n \ i - n < m" "\ j < n \ j - n < m" "\ j < n \ j - n < Suc m" by auto from UL show ?thesis using A1 A20 B3 B4 by simp next case LL hence [simp]: "\ j < n \ j - n < m" "\ j < n \ j - n < Suc m" by auto from LL show ?thesis using A1 A2 B3 B4 by simp next case LR thus ?thesis using A1 A2 B3 B4 by simp next case UR hence [simp]: "\ i < n \ i - n < m" by auto from UR show ?thesis using A1 A20 0 B3 B4 by simp qed qed (insert B4, auto) hence "det ?B = det (four_block_mat ?B1 ?B2 ?B3 ?B4)" by simp also have "\ = det ?B1 * det ?B4" by (rule det_four_block_mat_upper_right_zero_col[of _ "n + m"], insert A1 A2 B3 B4, auto) also have "det ?B1 = det A1 * det (mat m m (($$) B4))" by (rule Suc(1), insert B3 B4, auto) also have "\ * det ?B4 = det A1 * (det (mat m m (($$) B4)) * det ?B4)" by simp also have "det (mat m m (($$) B4)) * det ?B4 = det B4" unfolding arg_cong[OF B44, of det] by (subst det_four_block_mat_upper_right_zero_col[OF _ refl], auto) finally have id: "det ?B = det A1 * det B4" . from this[folded det] have "v * det ?A = v * (det A1 * det A4)" by simp with v show "det ?A = det A1 * det A4" by simp qed lemma det_swapcols: assumes *: "k < n" "l < n" "k \ l" and A: "A \ carrier_mat n n" shows "det (swapcols k l A) = - det A" proof - let ?B = "transpose_mat A" let ?C = "swaprows k l ?B" let ?D = "transpose_mat ?C" have C: "?C \ carrier_mat n n" and B: "?B \ carrier_mat n n" unfolding transpose_carrier_mat swaprows_carrier using A by auto show ?thesis unfolding swapcols_is_transp_swap_rows[OF A *(1-2)] det_transpose[OF C] det_swaprows[OF * B] det_transpose[OF A] .. qed lemma swap_row_to_front_det: "A \ carrier_mat n n \ I < n \ det (swap_row_to_front A I) = (-1)^I * det A" proof (induct I arbitrary: A) case (Suc I A) from Suc(3) have I: "I < n" by auto let ?I = "Suc I" let ?A = "swaprows I ?I A" have AA: "?A \ carrier_mat n n" using Suc(2) by simp have "det (swap_row_to_front A (Suc I)) = det (swap_row_to_front ?A I)" by simp also have "\ = (-1)^I * det ?A" by (rule Suc(1)[OF AA I]) also have "det ?A = -1 * det A" using det_swaprows[OF I Suc(3) _ Suc(2)] by simp finally show ?case by simp qed simp lemma swap_col_to_front_det: "A \ carrier_mat n n \ I < n \ det (swap_col_to_front A I) = (-1)^I * det A" proof (induct I arbitrary: A) case (Suc I A) from Suc(3) have I: "I < n" by auto let ?I = "Suc I" let ?A = "swapcols I ?I A" have AA: "?A \ carrier_mat n n" using Suc(2) by simp have "det (swap_col_to_front A (Suc I)) = det (swap_col_to_front ?A I)" by simp also have "\ = (-1)^I * det ?A" by (rule Suc(1)[OF AA I]) also have "det ?A = -1 * det A" using det_swapcols[OF I Suc(3) _ Suc(2)] by simp finally show ?case by simp qed simp lemma swap_row_to_front_four_block: assumes A1: "A1 \ carrier_mat n m1" and A2: "A2 \ carrier_mat n m2" and A3: "A3 \ carrier_mat 1 m1" and A4: "A4 \ carrier_mat 1 m2" shows "swap_row_to_front (four_block_mat A1 A2 A3 A4) n = four_block_mat A3 A4 A1 A2" by (subst swap_row_to_front_result[OF four_block_carrier_mat[OF A1 A4]], force, rule eq_matI, insert A1 A2 A3 A4, auto) lemma swap_col_to_front_four_block: assumes A1: "A1 \ carrier_mat n1 m" and A2: "A2 \ carrier_mat n1 1" and A3: "A3 \ carrier_mat n2 m" and A4: "A4 \ carrier_mat n2 1" shows "swap_col_to_front (four_block_mat A1 A2 A3 A4) m = four_block_mat A2 A1 A4 A3" by (subst swap_col_to_front_result[OF four_block_carrier_mat[OF A1 A4]], force, rule eq_matI, insert A1 A2 A3 A4, auto) lemma det_four_block_mat_lower_right_zero_col: assumes A1: "A1 \ carrier_mat 1 n" and A2: "A2 \ carrier_mat 1 1" and A3: "A3 \ carrier_mat n n" and A40: "A4 = (0\<^sub>m n 1)" shows "det (four_block_mat A1 A2 A3 A4) = (-1)^n * det A2 * det A3" (is "det ?A = _") proof - let ?B = "four_block_mat A3 A4 A1 A2" from four_block_carrier_mat[OF A3 A2] have B: "?B \ carrier_mat (Suc n) (Suc n)" by simp from A40 have A4: "A4 \ carrier_mat n 1" by auto from arg_cong[OF swap_row_to_front_four_block[OF A3 A4 A1 A2], of det] swap_row_to_front_det[OF B, of n] have "det ?A = (-1)^n * det ?B" by auto also have "det ?B = det A3 * det A2" by (rule det_four_block_mat_upper_right_zero_col[OF A3 A40 A1 A2]) finally show ?thesis by simp qed lemma det_four_block_mat_lower_left_zero_col: assumes A1: "A1 \ carrier_mat 1 1" and A2: "A2 \ carrier_mat 1 n" and A30: "A3 = (0\<^sub>m n 1)" and A4: "A4 \ carrier_mat n n" shows "det (four_block_mat A1 A2 A3 A4) = det A1 * det A4" (is "det ?A = _") proof - from A30 have A3: "A3 \ carrier_mat n 1" by auto let ?B = "four_block_mat A2 A1 A4 A3" from four_block_carrier_mat[OF A2 A3] have B: "?B \ carrier_mat (Suc n) (Suc n)" by simp from arg_cong[OF swap_col_to_front_four_block[OF A2 A1 A4 A3], of det] swap_col_to_front_det[OF B, of n] have "det ?A = (-1)^n * det ?B" by auto also have "det ?B = (- 1) ^ n * det A1 * det A4" by (rule det_four_block_mat_lower_right_zero_col[OF A2 A1 A4 A30]) also have "(-1)^n * \ = (-1 * -1)^n * det A1 * det A4" unfolding power_mult_distrib by (simp add: ac_simps) finally show ?thesis by simp qed lemma det_addcol[simp]: assumes l: "l < n" and k: "k \ l" and A: "A \ carrier_mat n n" shows "det (addcol a k l A) = det A" proof - have "addcol a k l A = A * addrow_mat n a l k" using addcol_mat[OF A l]. also have "det (A * addrow_mat n a l k) = det A * det (addrow_mat n a l k)" by(rule det_mult[OF A], auto) also have "det (addrow_mat n a l k) = 1" using det_addrow_mat[OF k[symmetric]]. finally show ?thesis using A by simp qed definition "insert_index i \ \i'. if i' < i then i' else Suc i'" definition "delete_index i \ \i'. if i' < i then i' else i' - Suc 0" lemma insert_index[simp]: "i' < i \ insert_index i i' = i'" "i' \ i \ insert_index i i' = Suc i'" unfolding insert_index_def by auto lemma delete_insert_index[simp]: "delete_index i (insert_index i i') = i'" unfolding insert_index_def delete_index_def by auto lemma insert_delete_index: assumes i'i: "i' \ i" shows "insert_index i (delete_index i i') = i'" unfolding insert_index_def delete_index_def using i'i by auto definition "delete_dom p i \ \i'. p (insert_index i i')" definition "delete_ran p j \ \i. delete_index j (p i)" definition "permutation_delete p i = delete_ran (delete_dom p i) (p i)" definition "insert_ran p j \ \i. insert_index j (p i)" definition "insert_dom p i j \ \i'. if i' < i then p i' else if i' = i then j else p (i'-1)" definition "permutation_insert i j p \ insert_dom (insert_ran p j) i j" lemmas permutation_delete_expand = permutation_delete_def[unfolded delete_dom_def delete_ran_def insert_index_def delete_index_def] lemmas permutation_insert_expand = permutation_insert_def[unfolded insert_dom_def insert_ran_def insert_index_def delete_index_def] lemma permutation_insert_inserted[simp]: "permutation_insert (i::nat) j p i = j" unfolding permutation_insert_expand by auto lemma permutation_insert_base: assumes p: "p permutes {0.. Fun.swap i (Suc i) id = permutation_insert i j p" (is "?l = ?r") proof (rule ext) fix x show "?l x = ?r x" - apply (cases rule: linorder_cases[of "x" "i"]) - unfolding permutation_insert_expand swap_def by auto + by (cases rule: linorder_cases[of "x" "i"]) + (auto simp add: swap_id_eq permutation_insert_expand) qed lemma permutation_insert_column_step: assumes p: "p permutes {0.. (permutation_insert i (Suc j) p) = permutation_insert i j p" (is "?l = ?r") proof (rule ext) fix x show "?l x = ?r x" proof (cases rule: linorder_cases[of "x" "i"]) case less note x = this show ?thesis apply (cases rule: linorder_cases[of "p x" "j"]) unfolding permutation_insert_expand using x by simp+ next case equal thus ?thesis by simp next case greater note x = this show ?thesis apply (cases rule: linorder_cases[of "p (x-1)" "j"]) unfolding permutation_insert_expand using x by simp+ qed qed lemma delete_dom_image: assumes i: "i \ {0.. ?N") assumes iff: "\i' \ ?N. f i' = f i \ i' = i" shows "delete_dom f i ` {0.. ?L" then obtain i' where i': "i' \ {0.. ?R" proof(cases "i' < i") case True show ?thesis unfolding image_def unfolding Diff_iff unfolding mem_Collect_eq singleton_iff proof(intro conjI bexI) show "j' \ f i" proof assume j': "j' = f i" hence "f i' = f i" using dj'[unfolded delete_dom_def insert_index_def] using True by simp thus "False" using iff i True by auto qed show "j' = f i'" using dj' True unfolding delete_dom_def insert_index_def by simp qed (insert i',simp) next case False show ?thesis unfolding image_def unfolding Diff_iff unfolding mem_Collect_eq singleton_iff proof(intro conjI bexI) show Si': "Suc i' \ ?N" using i' by auto show "j' \ f i" proof assume j': "j' = f i" hence "f (Suc i') = f i" using dj'[unfolded delete_dom_def insert_index_def] j' False by simp thus "False" using iff Si' False by auto qed show "j' = f (Suc i')" using dj' False unfolding delete_dom_def insert_index_def by simp qed qed } { assume R: "j' \ ?R" then obtain i' where i': "i' \ ?N" and j'fi: "j' \ f i" and j'fi': "j' = f i'" by auto hence i'i: "i' \ i" using iff by auto hence n: "n > 0" using i i' by auto show "j' \ ?L" proof (cases "i' < i") case True show ?thesis proof show "j' = delete_dom f i i'" unfolding delete_dom_def insert_index_def using True j'fi' by simp qed (insert True i, simp) next case False show ?thesis proof show "i'-1 \ {0.. {0.. ?N") assumes fimg: "f ` {0.. ?L" then obtain i where i: "i \ {0.. ?N - {j}" using fimg i by blast thus "j' \ ?R" using ij' j unfolding delete_ran_def delete_index_def by auto } { assume R: "j' \ ?R" show "j' \ ?L" proof (cases "j' < j") case True hence "j' \ ?N - {j}" using R by auto then obtain i where fij': "f i = j'" and i: "i \ {0.. ?N - {j}" using R by auto then obtain i where fij': "f i = Suc j'" and i: "i \ {0.. S" shows "inj_on (delete_index i) S" proof(intro inj_onI) fix x y assume eq: "delete_index i x = delete_index i y" and x: "x \ S" and y: "y \ S" have "x \ i" "y \ i" using x y iS by auto thus "x = y" using eq unfolding delete_index_def by(cases "x < i"; cases "y < i";simp) qed lemma insert_index_inj_on: shows "inj_on (insert_index i) S" proof(intro inj_onI) fix x y assume eq: "insert_index i x = insert_index i y" and x: "x \ S" and y: "y \ S" show "x = y" using eq unfolding insert_index_def by(cases "x < i"; cases "y < i";simp) qed lemma delete_dom_inj_on: assumes i: "i \ {0.. ?N") assumes inj: "inj_on f ?N" shows "inj_on (delete_dom f i) {0.. {0.. ?N") assumes img: "f ` {0.. {0 ..< Suc n}" (is "_ \ ?N") assumes bij: "bij_betw p ?N ?N" shows "bij_betw (permutation_delete p i) {0.. ?N" using i by auto have "\i'\?N. p i' = p i \ i' = i" using inj i unfolding inj_on_def by auto from delete_dom_image[OF i this] have "delete_dom p i ` {0.. {0.. n" by simp show "?p x = x" proof(cases "x < i") case True thus ?thesis unfolding permutation_delete_def using x i by simp next case False hence "p (Suc x) = Suc x" using x permutes_others[OF p] by auto thus ?thesis unfolding permutation_delete_expand using False pi x by simp qed qed (insert i,auto) lemma permutation_insert_delete: assumes p: "p permutes {0.. i" by auto hence cond: "\ i' - 1 < i" using i'i by simp show ?thesis proof (cases rule: linorder_cases[of "p i'" "p i"]) case less hence pd: "permutation_delete p i (i'-1) = p i'" unfolding permutation_delete_expand using i'i cond by auto show ?thesis unfolding permutation_insert_expand pd using i'i less by simp next case equal hence "i = i'" using permutes_inj[OF p] injD by metis hence False using i'i by auto thus ?thesis by auto next case greater hence pd: "permutation_delete p i (i'-1) = p i' - 1" unfolding permutation_delete_expand using i'i cond by simp show ?thesis unfolding permutation_insert_expand pd using i'i greater by auto qed qed qed lemma insert_index_exclude[simp]: "insert_index i i' \ i" unfolding insert_index_def by auto lemma insert_index_image: assumes i: "i < Suc n" shows "insert_index i ` {0.. ?L" then obtain i'' where ins: "i' = insert_index i i''" and i'': "i'' \ {0.. ?N - {i}" proof(rule DiffI) show "i' \ ?N" using ins unfolding insert_index_def using i'' by auto show "i' \ {i}" unfolding singleton_iff unfolding ins unfolding insert_index_def by auto qed } { assume R: "i' \ ?R" show "i' \ ?L" proof(cases rule: linorder_cases[of "i'" "i"]) case less hence i': "i' \ {0.. {0..i. insert_index j (f i)) ` {0.. f) ` {0.. ?f ` ?N" then obtain i' where i': "i' \ ?N" and j': "j' = ?f i'" by auto show "j' \ ?N" proof (cases rule:linorder_cases[of "i'" "i"]) case less hence "i' \ {0.. {0.. ?N" show "j' \ ?f ` ?N" proof (cases "j' = j") case True hence "?f i = j'" unfolding insert_dom_def by auto thus ?thesis using i by auto next case False hence j': "j' \ ?N - {j}" using j j' by auto then obtain i' where j'fi: "j' = f i'" and i': "i' \ {0.. {0.. {0.. f i' < j" apply (rule ccontr) using eq2 False by auto ultimately show ?thesis using eq2 by auto qed from inj_onD[OF inj this i i'] show "i = i'". qed lemma insert_dom_inj_on: assumes inj: "inj_on f {0.. ?N" moreover hence "q (i'-1) = i'-1" using permutes_others[OF q] by auto ultimately show "?p i' = i'" unfolding permutation_insert_expand using i j by auto qed lemma permutation_fix: assumes i: "i < Suc n" and j: "j < Suc n" shows "{ p. p permutes {0.. p i = j } = permutation_insert i j ` { q. q permutes {0.. ?L" hence p: "p permutes ?N" and pij: "p i = j" by auto show "p \ ?R" unfolding mem_Collect_eq using permutation_delete_permutes[OF p i] using permutation_insert_delete[OF p i,symmetric] unfolding pij by auto } { assume "p \ ?R" then obtain q where pq: "p = permutation_insert i j q" and q: "q permutes {0.. ?L" using pq permutation_insert_permutes[OF q i j] by auto } qed lemma permutation_split_ran: assumes j: "j \ S" shows "{ p. p permutes S } = (\i \ S. { p. p permutes S \ p i = j })" (is "?L = ?R") unfolding set_eq_iff proof(intro allI iffI) fix p { assume "p \ ?L" hence p: "p permutes S" by auto obtain i where i: "i \ S" and pij: "p i = j" using j permutes_image[OF p] by force thus "p \ ?R" using p by auto } { assume "p \ ?R" then obtain i where p: "p permutes S" and i: "i \ S" and pij: "p i = j" by auto show "p \ ?L" unfolding mem_Collect_eq using p. } qed lemma permutation_disjoint_dom: assumes i: "i \ S" and i': "i' \ S" and j: "j \ S" and ii': "i \ i'" shows "{ p. p permutes S \ p i = j } \ { p. p permutes S \ p i' = j } = {}" (is "?L \ ?R = {}") proof - { fix p assume "p \ ?L \ ?R" hence p: "p permutes S" and "p i = j" and "p i' = j" by auto hence "p i = p i'" by auto note injD[OF permutes_inj[OF p] this] hence False using ii' by auto } thus ?thesis by auto qed lemma permutation_disjoint_ran: assumes i: "i \ S" and j: "j \ S" and j': "j' \ S" and jj': "j \ j'" shows "{ p. p permutes S \ p i = j } \ { p. p permutes S \ p i = j' } = {}" (is "?L \ ?R = {}") proof - { fix p assume "p \ ?L \ ?R" hence "p permutes S" and "p i = j" and "p i = j'" by auto hence False using jj' by auto } thus ?thesis by auto qed lemma permutation_insert_inj_on: assumes "i < Suc n" assumes "j < Suc n" shows "inj_on (permutation_insert i j) { q. q permutes {0.. ?S" "q' \ ?S" hence q: "q permutes {0.. n" hence "signof (permutation_insert n (n-j) p) = (-1::'a)^(n+(n-j)) * signof p" proof(induct "j") case 0 show ?case using permutation_insert_base[OF p] by (simp add: mult_2[symmetric]) next case (Suc j) hence Sjn: "Suc j \ n" and j: "j < n" and Sj: "n - Suc j < n" by auto hence n0: "n > 0" by auto have ease: "Suc (n - Suc j) = n - j" using j by auto let ?swap = "Fun.swap (n - Suc j) (n - j) id" let ?prev = "permutation_insert n (n - j) p" have "signof (permutation_insert n (n - Suc j) p) = signof (?swap \ ?prev)" unfolding permutation_insert_column_step[OF p Sj, unfolded ease].. also have "... = signof ?swap * signof ?prev" proof(rule signof_compose) show "?swap permutes {0.. n" using j by auto have row_base: "signof (permutation_insert n j p) = (-1::'a)^(n+j) * signof p" using col[OF nj] using j by simp { fix i assume "i \ n" hence "signof (permutation_insert (n-i) j p) = (-1::'a)^((n-i)+j) * signof p" proof (induct i) case 0 show ?case using row_base by auto next case (Suc i) hence Sin: "Suc i \ n" and i: "i \ n" and Si: "n - Suc i < n" by auto have ease: "Suc (n - Suc i) = n - i" using Sin by auto let ?prev = "permutation_insert (n-i) j p" let ?swap = "Fun.swap (n - Suc i) (n-i) id" have "signof (permutation_insert (n - Suc i) j p) = signof (?prev \ ?swap)" using permutation_insert_row_step[of "n - Suc i"] unfolding ease by auto also have "... = signof ?prev * signof ?swap" proof(rule signof_compose) show "?swap permutes {0.. n" using i by auto show ?thesis using row[OF ni] using i by simp qed lemma foo: assumes i: "i < Suc n" and j: "j < Suc n" assumes q: "q permutes {0.. {0.. ?L" then obtain i' where ij: "ij = (i', permutation_insert i j q i')" and i': "i' < Suc n" and i'i: "i' \ i" by auto show "ij \ ?R" unfolding mem_Collect_eq proof(intro exI conjI) show "ij = (insert_index i (delete_index i i'), insert_index j (q (delete_index i i')))" using ij unfolding insert_delete_index[OF i'i] using i'i unfolding permutation_insert_expand insert_index_def delete_index_def by auto show "delete_index i i' < n" using i' i i'i unfolding delete_index_def by auto qed } { assume "ij \ ?R" then obtain i'' where ij: "ij = (insert_index i i'', insert_index j (q i''))" and i'': "i'' < n" by auto show "ij \ ?L" unfolding mem_Collect_eq proof(intro exI conjI) show "insert_index i i'' \ {0.. mat (dim_row A - 1) (dim_col A - 1) (\(i',j'). A $$ (if i' < i then i' else Suc i', if j' < j then j' else Suc j'))" lemma mat_delete_dim[simp]: "dim_row (mat_delete A i j) = dim_row A - 1" "dim_col (mat_delete A i j) = dim_col A - 1" unfolding mat_delete_def by auto lemma mat_delete_carrier: assumes A: "A \ carrier_mat m n" shows "mat_delete A i j \ carrier_mat (m-1) (n-1)" unfolding mat_delete_def using A by auto lemma "mat_delete_index": assumes A: "A \ carrier_mat (Suc n) (Suc n)" and i: "i < Suc n" and j: "j < Suc n" and i': "i' < n" and j': "j' < n" shows "A $$ (insert_index i i', insert_index j j') = mat_delete A i j $$ (i', j')" unfolding mat_delete_def unfolding permutation_insert_expand unfolding insert_index_def using A i j i' j' by auto definition "cofactor A i j = (-1)^(i+j) * det (mat_delete A i j)" lemma laplace_expansion_column: assumes A: "(A :: 'a :: comm_ring_1 mat) \ carrier_mat n n" and j: "j < n" shows "det A = (\i carrier_mat (Suc l) (Suc l)" and jl: "j < Suc l" using A j unfolding l_def by auto let ?N = "{0 ..< Suc l}" define f where "f = (\p i. A $$ (i, p i))" define g where "g = (\p. prod (f p) ?N)" define h where "h = (\p. signof p * g p)" define Q where "Q = { q . q permutes {0.. ?N" using jl by auto have disj: "\i \ ?N. \i' \ ?N. i \ i' \ {p. p permutes ?N \ p i = j} \ {p. p permutes ?N \ p i' = j} = {}" using permutation_disjoint_dom[OF _ _ jN] by auto have fin: "\i\?N. finite {p. p permutes ?N \ p i = j}" using finite_permutations[of ?N] by auto have "det A = sum h { p. p permutes ?N }" using det_def'[OF A] unfolding h_def g_def f_def using atLeast0LessThan by auto also have "... = sum h (\i\?N. {p. p permutes ?N \ p i = j})" unfolding permutation_split_ran[OF jN].. also have "... = (\i\?N. sum h { p | p. p permutes ?N \ p i = j})" using sum.UNION_disjoint[OF _ fin disj] by auto also { fix i assume "i \ ?N" hence i: "i < Suc l" by auto have "sum h { p | p. p permutes ?N \ p i = j} = sum h (permutation_insert i j ` Q)" using permutation_fix[OF i jl] unfolding Q_def by auto also have "... = sum (h \ permutation_insert i j) Q" unfolding Q_def using sum.reindex[OF permutation_insert_inj_on[OF i jl]]. also have "... = (\ q \ Q. signof (permutation_insert i j q) * prod (f (permutation_insert i j q)) ?N)" unfolding h_def g_def Q_def by simp also { fix q assume "q \ Q" hence q: "q permutes {0.. ?N - {i}" by auto have close: "insert i (?N - {i}) = ?N" using notin i by auto have "prod (f ?p) ?N = f ?p i * prod (f ?p) (?N-{i})" unfolding prod.insert[OF fin notin, unfolded close] by auto also have "... = A $$ (i, j) * prod (f ?p) (?N-{i})" unfolding f_def Q_def using permutation_insert_inserted by simp also have "prod (f ?p) (?N-{i}) = prod (\i'. A $$ (i', permutation_insert i j q i')) (?N-{i})" unfolding f_def.. also have "... = prod (\ij. A $$ ij) ((\i'. (i', permutation_insert i j q i')) ` (?N-{i}))" (is "_ = prod _ ?part") unfolding prod.reindex[OF inj_on_convol_ident] o_def.. also have "?part = {(i', permutation_insert i j q i') | i'. i' \ ?N-{i} }" unfolding image_def by metis also have "... = {(insert_index i i'', insert_index j (q i'')) | i''. i'' < l}" unfolding foo[OF i jl q].. also have "... = ((\i''. (insert_index i i'', insert_index j (q i''))) ` {0..ij. A $$ ij)... = prod ((\ij. A $$ ij) \ (\i''. (insert_index i i'', insert_index j (q i'')))) {0..i''. (i'', insert_index j (q i'')))" using inj_on_convol_ident. have 2: "inj (\(i'',j). (insert_index i i'', j))" apply (intro injI) using injD[OF insert_index_inj_on[of _ UNIV]] by auto have "inj (\i''. (insert_index i i'', insert_index j (q i'')))" using inj_compose[OF 2 1] unfolding o_def by auto thus "inj_on (\i''. (insert_index i i'', insert_index j (q i''))) {0..i''. A $$ (insert_index i i'', insert_index j (q i''))) {0..i''. mat_delete A i j $$ (i'', q i'')) {0..i'' = 0..< l. mat_delete A i j $$ (i'', q i''))" by auto hence "signof ?p * prod (f ?p) ?N = (-1::'a)^(i+j) * signof q * ..." unfolding signof_permutation_insert[OF q i jl] by auto } hence "... = (\ q \ Q. (-1)^(i+j) * signof q * A $$ (i, j) * (\i'' = 0 ..< l. mat_delete A i j $$ (i'', q i'')))" by(intro sum.cong[OF refl],auto) also have "... = ( \ q \ Q. A $$ (i, j) * (-1)^(i+j) * ( signof q * (\i'' = 0..< l. mat_delete A i j $$ (i'', q i'')) ) )" by (intro sum.cong[OF refl],auto) also have "... = A $$ (i, j) * (-1)^(i+j) * ( \ q \ Q. signof q * (\i''= 0 ..< l. mat_delete A i j $$ (i'', q i'')) )" unfolding sum_distrib_left by auto also have "... = (A $$ (i, j) * (-1)^(i+j) * det (mat_delete A i j))" unfolding det_def'[OF mat_delete_carrier[OF A]] unfolding Q_def by auto finally have "sum h {p | p. p permutes ?N \ p i = j} = A $$ (i, j) * cofactor A i j" unfolding cofactor_def by auto } hence "... = (\i\?N. A $$ (i,j) * cofactor A i j)" by auto finally show ?thesis unfolding atLeast0LessThan using A j unfolding l_def by auto qed lemma laplace_expansion_row: assumes A: "(A :: 'a :: comm_ring_1 mat) \ carrier_mat n n" and i: "i < n" shows "det A = (\jT)" using det_transpose[OF A] by simp also have "\ = (\jT $$ (j, i) * cofactor A\<^sup>T j i)" by (rule laplace_expansion_column[OF _ i], insert A, auto) also have "\ = (\j x y. x * y"], goal_cases) case (1 j) thus ?case using A i by auto next case (2 j) have "det (mat_delete A\<^sup>T j i) = det ((mat_delete A\<^sup>T j i)\<^sup>T)" by (subst det_transpose, insert A, auto simp: mat_delete_def) also have "(mat_delete A\<^sup>T j i)\<^sup>T = mat_delete A i j" unfolding mat_delete_def using A by auto finally show ?case by (simp add: ac_simps) qed finally show ?thesis . qed lemma degree_det_le: assumes "\ i j. i < n \ j < n \ degree (A $$ (i,j)) \ k" and A: "A \ carrier_mat n n" shows "degree (det A) \ k * n" proof - { fix p assume p: "p permutes {0..x = 0.. (\x = 0.. = k * n" unfolding sum_constant by simp also note calculation } note * = this show ?thesis unfolding det_def'[OF A] by (rule degree_sum_le, insert *, auto simp: finite_permutations signof_def intro!: order.trans[OF degree_prod_sum_le]) qed lemma upper_triangular_imp_det_eq_0_iff: fixes A :: "'a :: idom mat" assumes "A \ carrier_mat n n" and "upper_triangular A" shows "det A = 0 \ 0 \ set (diag_mat A)" using assms by (auto simp: det_upper_triangular) lemma det_identical_columns: assumes A: "A \ carrier_mat n n" and ij: "i \ j" and i: "i < n" and j: "j < n" and r: "col A i = col A j" shows "det A = 0" proof- have "det A = det A\<^sup>T" using det_transpose[OF A] .. also have "... = 0" proof (rule det_identical_rows[of _ n i j]) show "row (transpose_mat A) i = row (transpose_mat A) j" using A i j r by auto qed (auto simp add: assms) finally show ?thesis . qed definition adj_mat :: "'a :: comm_ring_1 mat \ 'a mat" where "adj_mat A = mat (dim_row A) (dim_col A) (\ (i,j). cofactor A j i)" lemma adj_mat: assumes A: "A \ carrier_mat n n" shows "adj_mat A \ carrier_mat n n" "A * adj_mat A = det A \\<^sub>m 1\<^sub>m n" "adj_mat A * A = det A \\<^sub>m 1\<^sub>m n" proof - from A have dims: "dim_row A = n" "dim_col A = n" by auto show aA: "adj_mat A \ carrier_mat n n" unfolding adj_mat_def dims by simp { fix i j assume ij: "i < n" "j < n" define B where "B = mat n n (\ (i',j'). if i' = j then A $$ (i,j') else A $$ (i',j'))" have "(A * adj_mat A) $$ (i,j) = (\ k < n. A $$ (i,k) * cofactor A j k)" unfolding times_mat_def scalar_prod_def adj_mat_def using ij A by (auto intro: sum.cong) also have "\ = (\ k < n. A $$ (i,k) * (-1)^(j + k) * det (mat_delete A j k))" unfolding cofactor_def by (auto intro: sum.cong) also have "\ = (\ k < n. B $$ (j,k) * (-1)^(j + k) * det (mat_delete B j k))" by (rule sum.cong[OF refl], intro arg_cong2[of _ _ _ _ "\ x y. y * _ * det x"], insert A ij, auto simp: B_def mat_delete_def) also have "\ = (\ k < n. B $$ (j,k) * cofactor B j k)" unfolding cofactor_def by (simp add: ac_simps) also have "\ = det B" by (rule laplace_expansion_row[symmetric], insert ij, auto simp: B_def) also have "\ = (if i = j then det A else 0)" proof (cases "i = j") case True hence "B = A" using A by (auto simp add: B_def) with True show ?thesis by simp next case False have "det B = 0" by (rule Determinant.det_identical_rows[OF _ False ij], insert A ij, auto simp: B_def) with False show ?thesis by simp qed also have "\ = (det A \\<^sub>m 1\<^sub>m n) $$ (i,j)" using ij by auto finally have "(A * adj_mat A) $$ (i, j) = (det A \\<^sub>m 1\<^sub>m n) $$ (i, j)" . } note main = this show "A * adj_mat A = det A \\<^sub>m 1\<^sub>m n" by (rule eq_matI[OF main], insert A aA, auto) (* now the completely symmetric version *) { fix i j assume ij: "i < n" "j < n" define B where "B = mat n n (\ (i',j'). if j' = i then A $$ (i',j) else A $$ (i',j'))" have "(adj_mat A * A) $$ (i,j) = (\ k < n. A $$ (k,j) * cofactor A k i)" unfolding times_mat_def scalar_prod_def adj_mat_def using ij A by (auto intro: sum.cong) also have "\ = (\ k < n. A $$ (k,j) * (-1)^(k + i) * det (mat_delete A k i))" unfolding cofactor_def by (auto intro: sum.cong) also have "\ = (\ k < n. B $$ (k,i) * (-1)^(k + i) * det (mat_delete B k i))" by (rule sum.cong[OF refl], intro arg_cong2[of _ _ _ _ "\ x y. y * _ * det x"], insert A ij, auto simp: B_def mat_delete_def) also have "\ = (\ k < n. B $$ (k,i) * cofactor B k i)" unfolding cofactor_def by (simp add: ac_simps) also have "\ = det B" by (rule laplace_expansion_column[symmetric], insert ij, auto simp: B_def) also have "\ = (if i = j then det A else 0)" proof (cases "i = j") case True hence "B = A" using A by (auto simp add: B_def) with True show ?thesis by simp next case False have "det B = 0" by (rule Determinant.det_identical_columns[OF _ False ij], insert A ij, auto simp: B_def) with False show ?thesis by simp qed also have "\ = (det A \\<^sub>m 1\<^sub>m n) $$ (i,j)" using ij by auto finally have "(adj_mat A * A) $$ (i, j) = (det A \\<^sub>m 1\<^sub>m n) $$ (i, j)" . } note main = this show "adj_mat A * A = det A \\<^sub>m 1\<^sub>m n" by (rule eq_matI[OF main], insert A aA, auto) qed definition "replace_col A b k = mat (dim_row A) (dim_col A) (\ (i,j). if j = k then b $ i else A $$ (i,j))" lemma cramer_lemma_mat: assumes A: "A \ carrier_mat n n" and x: "x \ carrier_vec n" and k: "k < n" shows "det (replace_col A (A *\<^sub>v x) k) = x $ k * det A" proof - define b where "b = A *\<^sub>v x" have b: "b \ carrier_vec n" using A x unfolding b_def by auto let ?Ab = "replace_col A b k" have Ab: "?Ab \ carrier_mat n n" using A by (auto simp: replace_col_def) have "x $ k * det A = (det A \\<^sub>v x) $ k" using A k x by auto also have "det A \\<^sub>v x = det A \\<^sub>v (1\<^sub>m n *\<^sub>v x)" using x by auto also have "\ = (det A \\<^sub>m 1\<^sub>m n) *\<^sub>v x" using A x by auto also have "\ = (adj_mat A * A) *\<^sub>v x" using adj_mat[OF A] by simp also have "\ = adj_mat A *\<^sub>v b" using adj_mat[OF A] A x unfolding b_def by (metis assoc_mult_mat_vec) also have "\ $ k = row (adj_mat A) k \ b" using adj_mat[OF A] b k by auto also have "\ = det (replace_col A b k)" unfolding scalar_prod_def using b k A by (subst laplace_expansion_column[OF Ab k], auto intro!: sum.cong arg_cong[of _ _ det] arg_cong[of _ _ "\ x. _ * x"] eq_matI simp: replace_col_def adj_mat_def Matrix.row_def cofactor_def mat_delete_def ac_simps) finally show ?thesis unfolding b_def by simp qed end \ No newline at end of file diff --git a/thys/List-Index/List_Index.thy b/thys/List-Index/List_Index.thy --- a/thys/List-Index/List_Index.thy +++ b/thys/List-Index/List_Index.thy @@ -1,561 +1,561 @@ (* Author: Tobias Nipkow, with contributions from Dmitriy Traytel, Lukas Bulwahn, and Peter Lammich *) section \Index-based manipulation of lists\ theory List_Index imports Main begin text \\noindent This theory collects functions for index-based manipulation of lists. \ subsection \Finding an index\ text \ This subsection defines three functions for finding the index of items in a list: \begin{description} \item[\find_index P xs\] finds the index of the first element in \xs\ that satisfies \P\. \item[\index xs x\] finds the index of the first occurrence of \x\ in \xs\. \item[\last_index xs x\] finds the index of the last occurrence of \x\ in \xs\. \end{description} All functions return @{term "length xs"} if \xs\ does not contain a suitable element. The argument order of \find_index\ follows the function of the same name in the Haskell standard library. For \index\ (and \last_index\) the order is intentionally reversed: \index\ maps lists to a mapping from elements to their indices, almost the inverse of function \nth\.\ primrec find_index :: "('a \ bool) \ 'a list \ nat" where "find_index _ [] = 0" | "find_index P (x#xs) = (if P x then 0 else find_index P xs + 1)" definition index :: "'a list \ 'a \ nat" where "index xs = (\a. find_index (\x. x=a) xs)" definition last_index :: "'a list \ 'a \ nat" where "last_index xs x = (let i = index (rev xs) x; n = size xs in if i = n then i else n - (i+1))" lemma find_index_append: "find_index P (xs @ ys) = (if \x\set xs. P x then find_index P xs else size xs + find_index P ys)" by (induct xs) simp_all lemma find_index_le_size: "find_index P xs <= size xs" by(induct xs) simp_all lemma index_le_size: "index xs x <= size xs" by(simp add: index_def find_index_le_size) lemma last_index_le_size: "last_index xs x <= size xs" by(simp add: last_index_def Let_def index_le_size) lemma index_Nil[simp]: "index [] a = 0" by(simp add: index_def) lemma index_Cons[simp]: "index (x#xs) a = (if x=a then 0 else index xs a + 1)" by(simp add: index_def) lemma index_append: "index (xs @ ys) x = (if x : set xs then index xs x else size xs + index ys x)" by (induct xs) simp_all lemma index_conv_size_if_notin[simp]: "x \ set xs \ index xs x = size xs" by (induct xs) auto lemma find_index_eq_size_conv: "size xs = n \ (find_index P xs = n) = (\x \ set xs. ~ P x)" by(induct xs arbitrary: n) auto lemma size_eq_find_index_conv: "size xs = n \ (n = find_index P xs) = (\x \ set xs. ~ P x)" by(metis find_index_eq_size_conv) lemma index_size_conv: "size xs = n \ (index xs x = n) = (x \ set xs)" by(auto simp: index_def find_index_eq_size_conv) lemma size_index_conv: "size xs = n \ (n = index xs x) = (x \ set xs)" by (metis index_size_conv) lemma last_index_size_conv: "size xs = n \ (last_index xs x = n) = (x \ set xs)" apply(auto simp: last_index_def index_size_conv) apply(drule length_pos_if_in_set) apply arith done lemma size_last_index_conv: "size xs = n \ (n = last_index xs x) = (x \ set xs)" by (metis last_index_size_conv) lemma find_index_less_size_conv: "(find_index P xs < size xs) = (\x \ set xs. P x)" by (induct xs) auto lemma index_less_size_conv: "(index xs x < size xs) = (x \ set xs)" by(auto simp: index_def find_index_less_size_conv) lemma last_index_less_size_conv: "(last_index xs x < size xs) = (x : set xs)" by(simp add: last_index_def Let_def index_size_conv length_pos_if_in_set del:length_greater_0_conv) lemma index_less[simp]: "x : set xs \ size xs <= n \ index xs x < n" apply(induct xs) apply auto apply (metis index_less_size_conv less_eq_Suc_le less_trans_Suc) done lemma last_index_less[simp]: "x : set xs \ size xs <= n \ last_index xs x < n" by(simp add: last_index_less_size_conv[symmetric]) lemma last_index_Cons: "last_index (x#xs) y = (if x=y then if x \ set xs then last_index xs y + 1 else 0 else last_index xs y + 1)" using index_le_size[of "rev xs" y] apply(auto simp add: last_index_def index_append Let_def) apply(simp add: index_size_conv) done lemma last_index_append: "last_index (xs @ ys) x = (if x : set ys then size xs + last_index ys x else if x : set xs then last_index xs x else size xs + size ys)" by (induct xs) (simp_all add: last_index_Cons last_index_size_conv) lemma last_index_Snoc[simp]: "last_index (xs @ [x]) y = (if x=y then size xs else if y : set xs then last_index xs y else size xs + 1)" by(simp add: last_index_append last_index_Cons) lemma nth_find_index: "find_index P xs < size xs \ P(xs ! find_index P xs)" by (induct xs) auto lemma nth_index[simp]: "x \ set xs \ xs ! index xs x = x" by (induct xs) auto lemma nth_last_index[simp]: "x \ set xs \ xs ! last_index xs x = x" by(simp add:last_index_def index_size_conv Let_def rev_nth[symmetric]) lemma index_rev: "\ distinct xs; x \ set xs \ \ index (rev xs) x = length xs - index xs x - 1" by (induct xs) (auto simp: index_append) lemma index_nth_id: "\ distinct xs; n < length xs \ \ index xs (xs ! n) = n" by (metis in_set_conv_nth index_less_size_conv nth_eq_iff_index_eq nth_index) lemma index_upt[simp]: "m \ i \ i < n \ index [m.. set xs \ y \ set xs \ (index xs x = index xs y) = (x = y)" by (induct xs) auto lemma last_index_eq_index_conv[simp]: "x \ set xs \ y \ set xs \ (last_index xs x = last_index xs y) = (x = y)" by (induct xs) (auto simp:last_index_Cons) lemma inj_on_index: "inj_on (index xs) (set xs)" by (simp add:inj_on_def) lemma inj_on_index2: "I \ set xs \ inj_on (index xs) I" by (rule inj_onI) auto lemma inj_on_last_index: "inj_on (last_index xs) (set xs)" by (simp add:inj_on_def) lemma find_index_conv_takeWhile: "find_index P xs = size(takeWhile (Not o P) xs)" by(induct xs) auto lemma index_conv_takeWhile: "index xs x = size(takeWhile (\y. x\y) xs)" by(induct xs) auto lemma find_index_first: "i < find_index P xs \ \P (xs!i)" unfolding find_index_conv_takeWhile by (metis comp_apply nth_mem set_takeWhileD takeWhile_nth) lemma index_first: "i x\xs!i" using find_index_first unfolding index_def by blast lemma find_index_eqI: assumes "i\length xs" assumes "\jP (xs!j)" assumes "i P (xs!i)" shows "find_index P xs = i" by (metis (mono_tags, lifting) antisym_conv2 assms find_index_eq_size_conv find_index_first find_index_less_size_conv linorder_neqE_nat nth_find_index) lemma find_index_eq_iff: "find_index P xs = i \ (i\length xs \ (\jP (xs!j)) \ (i P (xs!i)))" by (auto intro: find_index_eqI simp: nth_find_index find_index_le_size find_index_first) lemma index_eqI: assumes "i\length xs" assumes "\j x" assumes "i xs!i = x" shows "index xs x = i" unfolding index_def by (simp add: find_index_eqI assms) lemma index_eq_iff: "index xs x = i \ (i\length xs \ (\j x) \ (i xs!i = x))" by (auto intro: index_eqI simp: index_le_size index_less_size_conv dest: index_first) lemma index_take: "index xs x >= i \ x \ set(take i xs)" apply(subst (asm) index_conv_takeWhile) apply(subgoal_tac "set(take i xs) <= set(takeWhile ((\) x) xs)") apply(blast dest: set_takeWhileD) apply(metis set_take_subset_set_take takeWhile_eq_take) done lemma last_index_drop: "last_index xs x < i \ x \ set(drop i xs)" apply(subgoal_tac "set(drop i xs) = set(take (size xs - i) (rev xs))") apply(simp add: last_index_def index_take Let_def split:if_split_asm) apply (metis rev_drop set_rev) done lemma set_take_if_index: assumes "index xs x < i" and "i \ length xs" shows "x \ set (take i xs)" proof - have "index (take i xs @ drop i xs) x < i" using append_take_drop_id[of i xs] assms(1) by simp thus ?thesis using assms(2) by(simp add:index_append del:append_take_drop_id split: if_splits) qed lemma index_take_if_index: assumes "index xs x \ n" shows "index (take n xs) x = index xs x" proof cases assume "x : set(take n xs)" with assms show ?thesis by (metis append_take_drop_id index_append) next assume "x \ set(take n xs)" with assms show ?thesis by (metis order_le_less set_take_if_index le_cases length_take min_def size_index_conv take_all) qed lemma index_take_if_set: "x : set(take n xs) \ index (take n xs) x = index xs x" by (metis index_take index_take_if_index linear) lemma index_last[simp]: "xs \ [] \ distinct xs \ index xs (last xs) = length xs - 1" by (induction xs) auto lemma index_update_if_diff2: "n < length xs \ x \ xs!n \ x \ y \ index (xs[n := y]) x = index xs x" by(subst (2) id_take_nth_drop[of n xs]) (auto simp: upd_conv_take_nth_drop index_append min_def) lemma set_drop_if_index: "distinct xs \ index xs x < i \ x \ set(drop i xs)" by (metis in_set_dropD index_nth_id last_index_drop last_index_less_size_conv nth_last_index) lemma index_swap_if_distinct: assumes "distinct xs" "i < size xs" "j < size xs" shows "index (xs[i := xs!j, j := xs!i]) x = (if x = xs!i then j else if x = xs!j then i else index xs x)" proof- have "distinct(xs[i := xs!j, j := xs!i])" using assms by simp with assms show ?thesis - apply (auto simp: swap_def simp del: distinct_swap) + apply (auto simp: simp del: distinct_swap) apply (metis index_nth_id list_update_same_conv) apply (metis (erased, hide_lams) index_nth_id length_list_update list_update_swap nth_list_update_eq) apply (metis index_nth_id length_list_update nth_list_update_eq) by (metis index_update_if_diff2 length_list_update nth_list_update) qed lemma bij_betw_index: "distinct xs \ X = set xs \ l = size xs \ bij_betw (index xs) X {0.. set xs = X \ index xs ` X = {0.. inj_on f S; y \ S; set xs \ S \ \ index (map f xs) (f y) = index xs y" by (induct xs) (auto simp: inj_on_eq_iff) lemma index_map_inj: "inj f \ index (map f xs) (f y) = index xs y" by (simp add: index_map_inj_on[where S=UNIV]) subsection \Map with index\ primrec map_index' :: "nat \ (nat \ 'a \ 'b) \ 'a list \ 'b list" where "map_index' n f [] = []" | "map_index' n f (x#xs) = f n x # map_index' (Suc n) f xs" lemma length_map_index'[simp]: "length (map_index' n f xs) = length xs" by (induct xs arbitrary: n) auto lemma map_index'_map_zip: "map_index' n f xs = map (case_prod f) (zip [n ..< n + length xs] xs)" proof (induct xs arbitrary: n) case (Cons x xs) hence "map_index' n f (x#xs) = f n x # map (case_prod f) (zip [Suc n ..< n + length (x # xs)] xs)" by simp also have "\ = map (case_prod f) (zip (n # [Suc n ..< n + length (x # xs)]) (x # xs))" by simp also have "(n # [Suc n ..< n + length (x # xs)]) = [n ..< n + length (x # xs)]" by (induct xs) auto finally show ?case by simp qed simp abbreviation "map_index \ map_index' 0" lemmas map_index = map_index'_map_zip[of 0, simplified] lemma take_map_index: "take p (map_index f xs) = map_index f (take p xs)" unfolding map_index by (auto simp: min_def take_map take_zip) lemma drop_map_index: "drop p (map_index f xs) = map_index' p f (drop p xs)" unfolding map_index'_map_zip by (cases "p < length xs") (auto simp: drop_map drop_zip) lemma map_map_index[simp]: "map g (map_index f xs) = map_index (\n x. g (f n x)) xs" unfolding map_index by auto lemma map_index_map[simp]: "map_index f (map g xs) = map_index (\n x. f n (g x)) xs" unfolding map_index by (auto simp: map_zip_map2) lemma set_map_index[simp]: "x \ set (map_index f xs) = (\i < length xs. f i (xs ! i) = x)" unfolding map_index by (auto simp: set_zip intro!: image_eqI[of _ "case_prod f"]) lemma set_map_index'[simp]: "x\set (map_index' n f xs) \ (\i map_index f xs ! p = f p (xs ! p)" unfolding map_index by auto lemma map_index_cong: "\p < length xs. f p (xs ! p) = g p (xs ! p) \ map_index f xs = map_index g xs" unfolding map_index by (auto simp: set_zip) lemma map_index_id: "map_index (curry snd) xs = xs" unfolding map_index by auto lemma map_index_no_index[simp]: "map_index (\n x. f x) xs = map f xs" unfolding map_index by (induct xs rule: rev_induct) auto lemma map_index_congL: "\p < length xs. f p (xs ! p) = xs ! p \ map_index f xs = xs" by (rule trans[OF map_index_cong map_index_id]) auto lemma map_index'_is_NilD: "map_index' n f xs = [] \ xs = []" by (induct xs) auto declare map_index'_is_NilD[of 0, dest!] lemma map_index'_is_ConsD: "map_index' n f xs = y # ys \ \z zs. xs = z # zs \ f n z = y \ map_index' (n + 1) f zs = ys" by (induct xs arbitrary: n) auto lemma map_index'_eq_imp_length_eq: "map_index' n f xs = map_index' n g ys \ length xs = length ys" proof (induct ys arbitrary: xs n) case (Cons y ys) thus ?case by (cases xs) auto qed (auto dest!: map_index'_is_NilD) lemmas map_index_eq_imp_length_eq = map_index'_eq_imp_length_eq[of 0] lemma map_index'_comp[simp]: "map_index' n f (map_index' n g xs) = map_index' n (\n. f n o g n) xs" by (induct xs arbitrary: n) auto lemma map_index'_append[simp]: "map_index' n f (a @ b) = map_index' n f a @ map_index' (n + length a) f b" by (induct a arbitrary: n) auto lemma map_index_append[simp]: "map_index f (a @ b) = map_index f a @ map_index' (length a) f b" using map_index'_append[where n=0] by (simp del: map_index'_append) subsection \Insert at position\ primrec insert_nth :: "nat \ 'a \ 'a list \ 'a list" where "insert_nth 0 x xs = x # xs" | "insert_nth (Suc n) x xs = (case xs of [] \ [x] | y # ys \ y # insert_nth n x ys)" lemma insert_nth_take_drop[simp]: "insert_nth n x xs = take n xs @ [x] @ drop n xs" proof (induct n arbitrary: xs) case Suc thus ?case by (cases xs) auto qed simp lemma length_insert_nth: "length (insert_nth n x xs) = Suc (length xs)" by (induct xs) auto lemma set_insert_nth: "set (insert_nth i x xs) = insert x (set xs)" by (simp add: set_append[symmetric]) lemma distinct_insert_nth: assumes "distinct xs" assumes "x \ set xs" shows "distinct (insert_nth i x xs)" using assms proof (induct xs arbitrary: i) case Nil then show ?case by (cases i) auto next case (Cons a xs) then show ?case by (cases i) (auto simp add: set_insert_nth simp del: insert_nth_take_drop) qed lemma nth_insert_nth_front: assumes "i < j" "j \ length xs" shows "insert_nth j x xs ! i = xs ! i" using assms by (simp add: nth_append) lemma nth_insert_nth_index_eq: assumes "i \ length xs" shows "insert_nth i x xs ! i = x" using assms by (simp add: nth_append) lemma nth_insert_nth_back: assumes "j < i" "i \ length xs" shows "insert_nth j x xs ! i = xs ! (i - 1)" using assms by (cases i) (auto simp add: nth_append min_def) lemma nth_insert_nth: assumes "i \ length xs" "j \ length xs" shows "insert_nth j x xs ! i = (if i = j then x else if i < j then xs ! i else xs ! (i - 1))" using assms by (simp add: nth_insert_nth_front nth_insert_nth_index_eq nth_insert_nth_back del: insert_nth_take_drop) lemma insert_nth_inverse: assumes "j \ length xs" "j' \ length xs'" assumes "x \ set xs" "x \ set xs'" assumes "insert_nth j x xs = insert_nth j' x xs'" shows "j = j'" proof - from assms(1,3) have "\i\length xs. insert_nth j x xs ! i = x \ i = j" by (auto simp add: nth_insert_nth simp del: insert_nth_take_drop) moreover from assms(2,4) have "\i\length xs'. insert_nth j' x xs' ! i = x \ i = j'" by (auto simp add: nth_insert_nth simp del: insert_nth_take_drop) ultimately show "j = j'" using assms(1,2,5) by (metis dual_order.trans nat_le_linear) qed text \Insert several elements at given (ascending) positions\ lemma length_fold_insert_nth: "length (fold (\(p, b). insert_nth p b) pxs xs) = length xs + length pxs" by (induct pxs arbitrary: xs) auto lemma invar_fold_insert_nth: "\\x\set pxs. p < fst x; p < length xs; xs ! p = b\ \ fold (\(x, y). insert_nth x y) pxs xs ! p = b" by (induct pxs arbitrary: xs) (auto simp: nth_append) lemma nth_fold_insert_nth: "\sorted (map fst pxs); distinct (map fst pxs); \(p, b) \ set pxs. p < length xs + length pxs; i < length pxs; pxs ! i = (p, b)\ \ fold (\(p, b). insert_nth p b) pxs xs ! p = b" proof (induct pxs arbitrary: xs i p b) case (Cons pb pxs) show ?case proof (cases i) case 0 with Cons.prems have "p < Suc (length xs)" proof (induct pxs rule: rev_induct) case (snoc pb' pxs) then obtain p' b' where "pb' = (p', b')" by auto with snoc.prems have "\p \ fst ` set pxs. p < p'" "p' \ Suc (length xs + length pxs)" by (auto simp: image_iff sorted_append le_eq_less_or_eq) with snoc.prems show ?case by (intro snoc(1)) (auto simp: sorted_append) qed auto with 0 Cons.prems show ?thesis unfolding fold.simps o_apply by (intro invar_fold_insert_nth) (auto simp: image_iff le_eq_less_or_eq nth_append) next case (Suc n) with Cons.prems show ?thesis unfolding fold.simps by (auto intro!: Cons(1)) qed qed simp subsection \Remove at position\ fun remove_nth :: "nat \ 'a list \ 'a list" where "remove_nth i [] = []" | "remove_nth 0 (x # xs) = xs" | "remove_nth (Suc i) (x # xs) = x # remove_nth i xs" lemma remove_nth_take_drop: "remove_nth i xs = take i xs @ drop (Suc i) xs" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) auto qed lemma remove_nth_insert_nth: assumes "i \ length xs" shows "remove_nth i (insert_nth i x xs) = xs" using assms proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) auto qed lemma insert_nth_remove_nth: assumes "i < length xs" shows "insert_nth i (xs ! i) (remove_nth i xs) = xs" using assms proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) auto qed lemma length_remove_nth: assumes "i < length xs" shows "length (remove_nth i xs) = length xs - 1" using assms unfolding remove_nth_take_drop by simp lemma set_remove_nth_subset: "set (remove_nth j xs) \ set xs" proof (induct xs arbitrary: j) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases j) auto qed lemma set_remove_nth: assumes "distinct xs" "j < length xs" shows "set (remove_nth j xs) = set xs - {xs ! j}" using assms proof (induct xs arbitrary: j) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases j) auto qed lemma distinct_remove_nth: assumes "distinct xs" shows "distinct (remove_nth i xs)" using assms proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) (auto simp add: set_remove_nth_subset rev_subsetD) qed end diff --git a/thys/Perron_Frobenius/HMA_Connect.thy b/thys/Perron_Frobenius/HMA_Connect.thy --- a/thys/Perron_Frobenius/HMA_Connect.thy +++ b/thys/Perron_Frobenius/HMA_Connect.thy @@ -1,749 +1,751 @@ (* Authors: J. Divasón, R. Thiemann, A. Yamada, O. Kunčar *) subsection \Transfer rules to convert theorems from JNF to HMA and vice-versa.\ theory HMA_Connect imports Jordan_Normal_Form.Spectral_Radius "HOL-Analysis.Determinants" "HOL-Analysis.Cartesian_Euclidean_Space" Bij_Nat Cancel_Card_Constraint "HOL-Eisbach.Eisbach" begin text \Prefer certain constants and lemmas without prefix.\ hide_const (open) Matrix.mat hide_const (open) Matrix.row hide_const (open) Determinant.det lemmas mat_def = Finite_Cartesian_Product.mat_def lemmas det_def = Determinants.det_def lemmas row_def = Finite_Cartesian_Product.row_def notation vec_index (infixl "$v" 90) notation vec_nth (infixl "$h" 90) text \Forget that @{typ "'a mat"}, @{typ "'a Matrix.vec"}, and @{typ "'a poly"} have been defined via lifting\ (* TODO: add to end of matrix theory, stores lifting + transfer setup *) lifting_forget vec.lifting lifting_forget mat.lifting lifting_forget poly.lifting text \Some notions which we did not find in the HMA-world.\ definition eigen_vector :: "'a::comm_ring_1 ^ 'n ^ 'n \ 'a ^ 'n \ 'a \ bool" where "eigen_vector A v ev = (v \ 0 \ A *v v = ev *s v)" definition eigen_value :: "'a :: comm_ring_1 ^ 'n ^ 'n \ 'a \ bool" where "eigen_value A k = (\ v. eigen_vector A v k)" definition similar_matrix_wit :: "'a :: semiring_1 ^ 'n ^ 'n \ 'a ^ 'n ^ 'n \ 'a ^ 'n ^ 'n \ 'a ^ 'n ^ 'n \ bool" where "similar_matrix_wit A B P Q = (P ** Q = mat 1 \ Q ** P = mat 1 \ A = P ** B ** Q)" definition similar_matrix :: "'a :: semiring_1 ^ 'n ^ 'n \ 'a ^ 'n ^ 'n \ bool" where "similar_matrix A B = (\ P Q. similar_matrix_wit A B P Q)" definition spectral_radius :: "complex ^ 'n ^ 'n \ real" where "spectral_radius A = Max { norm ev | v ev. eigen_vector A v ev}" definition Spectrum :: "'a :: field ^ 'n ^ 'n \ 'a set" where "Spectrum A = Collect (eigen_value A)" definition vec_elements_h :: "'a ^ 'n \ 'a set" where "vec_elements_h v = range (vec_nth v)" lemma vec_elements_h_def': "vec_elements_h v = {v $h i | i. True}" unfolding vec_elements_h_def by auto definition elements_mat_h :: "'a ^ 'nc ^ 'nr \ 'a set" where "elements_mat_h A = range (\ (i,j). A $h i $h j)" lemma elements_mat_h_def': "elements_mat_h A = {A $h i $h j | i j. True}" unfolding elements_mat_h_def by auto definition map_vector :: "('a \ 'b) \ 'a ^'n \ 'b ^'n" where "map_vector f v \ \ i. f (v $h i)" definition map_matrix :: "('a \ 'b) \ 'a ^ 'n ^ 'm \ 'b ^ 'n ^ 'm" where "map_matrix f A \ \ i. map_vector f (A $h i)" definition normbound :: "'a :: real_normed_field ^ 'nc ^ 'nr \ real \ bool" where "normbound A b \ \ x \ elements_mat_h A. norm x \ b" lemma spectral_radius_ev_def: "spectral_radius A = Max (norm ` (Collect (eigen_value A)))" unfolding spectral_radius_def eigen_value_def[abs_def] by (rule arg_cong[where f = Max], auto) lemma elements_mat: "elements_mat A = {A $$ (i,j) | i j. i < dim_row A \ j < dim_col A}" unfolding elements_mat_def by force definition vec_elements :: "'a Matrix.vec \ 'a set" where "vec_elements v = set [v $ i. i <- [0 ..< dim_vec v]]" lemma vec_elements: "vec_elements v = { v $ i | i. i < dim_vec v}" unfolding vec_elements_def by auto (* TODO: restore a bundle, for e.g., for matrix_impl *) context includes vec.lifting begin end definition from_hma\<^sub>v :: "'a ^ 'n \ 'a Matrix.vec" where "from_hma\<^sub>v v = Matrix.vec CARD('n) (\ i. v $h from_nat i)" definition from_hma\<^sub>m :: "'a ^ 'nc ^ 'nr \ 'a Matrix.mat" where "from_hma\<^sub>m a = Matrix.mat CARD('nr) CARD('nc) (\ (i,j). a $h from_nat i $h from_nat j)" definition to_hma\<^sub>v :: "'a Matrix.vec \ 'a ^ 'n" where "to_hma\<^sub>v v = (\ i. v $v to_nat i)" definition to_hma\<^sub>m :: "'a Matrix.mat \ 'a ^ 'nc ^ 'nr " where "to_hma\<^sub>m a = (\ i j. a $$ (to_nat i, to_nat j))" declare vec_lambda_eta[simp] lemma to_hma_from_hma\<^sub>v[simp]: "to_hma\<^sub>v (from_hma\<^sub>v v) = v" by (auto simp: to_hma\<^sub>v_def from_hma\<^sub>v_def to_nat_less_card) lemma to_hma_from_hma\<^sub>m[simp]: "to_hma\<^sub>m (from_hma\<^sub>m v) = v" by (auto simp: to_hma\<^sub>m_def from_hma\<^sub>m_def to_nat_less_card) lemma from_hma_to_hma\<^sub>v[simp]: "v \ carrier_vec (CARD('n)) \ from_hma\<^sub>v (to_hma\<^sub>v v :: 'a ^ 'n) = v" by (auto simp: to_hma\<^sub>v_def from_hma\<^sub>v_def to_nat_from_nat_id) lemma from_hma_to_hma\<^sub>m[simp]: "A \ carrier_mat (CARD('nr)) (CARD('nc)) \ from_hma\<^sub>m (to_hma\<^sub>m A :: 'a ^ 'nc ^ 'nr) = A" by (auto simp: to_hma\<^sub>m_def from_hma\<^sub>m_def to_nat_from_nat_id) lemma from_hma\<^sub>v_inj[simp]: "from_hma\<^sub>v x = from_hma\<^sub>v y \ x = y" by (intro iffI, insert to_hma_from_hma\<^sub>v[of x], auto) lemma from_hma\<^sub>m_inj[simp]: "from_hma\<^sub>m x = from_hma\<^sub>m y \ x = y" by(intro iffI, insert to_hma_from_hma\<^sub>m[of x], auto) definition HMA_V :: "'a Matrix.vec \ 'a ^ 'n \ bool" where "HMA_V = (\ v w. v = from_hma\<^sub>v w)" definition HMA_M :: "'a Matrix.mat \ 'a ^ 'nc ^ 'nr \ bool" where "HMA_M = (\ a b. a = from_hma\<^sub>m b)" definition HMA_I :: "nat \ 'n :: finite \ bool" where "HMA_I = (\ i a. i = to_nat a)" context includes lifting_syntax begin lemma Domainp_HMA_V [transfer_domain_rule]: "Domainp (HMA_V :: 'a Matrix.vec \ 'a ^ 'n \ bool) = (\ v. v \ carrier_vec (CARD('n )))" by(intro ext iffI, insert from_hma_to_hma\<^sub>v[symmetric], auto simp: from_hma\<^sub>v_def HMA_V_def) lemma Domainp_HMA_M [transfer_domain_rule]: "Domainp (HMA_M :: 'a Matrix.mat \ 'a ^ 'nc ^ 'nr \ bool) = (\ A. A \ carrier_mat CARD('nr) CARD('nc))" by (intro ext iffI, insert from_hma_to_hma\<^sub>m[symmetric], auto simp: from_hma\<^sub>m_def HMA_M_def) lemma Domainp_HMA_I [transfer_domain_rule]: "Domainp (HMA_I :: nat \ 'n :: finite \ bool) = (\ i. i < CARD('n))" (is "?l = ?r") proof (intro ext) fix i :: nat show "?l i = ?r i" unfolding HMA_I_def Domainp_iff by (auto intro: exI[of _ "from_nat i"] simp: to_nat_from_nat_id to_nat_less_card) qed lemma bi_unique_HMA_V [transfer_rule]: "bi_unique HMA_V" "left_unique HMA_V" "right_unique HMA_V" unfolding HMA_V_def bi_unique_def left_unique_def right_unique_def by auto lemma bi_unique_HMA_M [transfer_rule]: "bi_unique HMA_M" "left_unique HMA_M" "right_unique HMA_M" unfolding HMA_M_def bi_unique_def left_unique_def right_unique_def by auto lemma bi_unique_HMA_I [transfer_rule]: "bi_unique HMA_I" "left_unique HMA_I" "right_unique HMA_I" unfolding HMA_I_def bi_unique_def left_unique_def right_unique_def by auto lemma right_total_HMA_V [transfer_rule]: "right_total HMA_V" unfolding HMA_V_def right_total_def by simp lemma right_total_HMA_M [transfer_rule]: "right_total HMA_M" unfolding HMA_M_def right_total_def by simp lemma right_total_HMA_I [transfer_rule]: "right_total HMA_I" unfolding HMA_I_def right_total_def by simp lemma HMA_V_index [transfer_rule]: "(HMA_V ===> HMA_I ===> (=)) ($v) ($h)" unfolding rel_fun_def HMA_V_def HMA_I_def from_hma\<^sub>v_def by (auto simp: to_nat_less_card) text \We introduce the index function to have pointwise access to HMA-matrices by a constant. Otherwise, the transfer rule with @{term "\ A i j. A $h i $h j"} instead of index is not applicable.\ definition "index_hma A i j \ A $h i $h j" lemma HMA_M_index [transfer_rule]: "(HMA_M ===> HMA_I ===> HMA_I ===> (=)) (\ A i j. A $$ (i,j)) index_hma" by (intro rel_funI, simp add: index_hma_def to_nat_less_card HMA_M_def HMA_I_def from_hma\<^sub>m_def) lemma HMA_V_0 [transfer_rule]: "HMA_V (0\<^sub>v CARD('n)) (0 :: 'a :: zero ^ 'n)" unfolding HMA_V_def from_hma\<^sub>v_def by auto lemma HMA_M_0 [transfer_rule]: "HMA_M (0\<^sub>m CARD('nr) CARD('nc)) (0 :: 'a :: zero ^ 'nc ^ 'nr )" unfolding HMA_M_def from_hma\<^sub>m_def by auto lemma HMA_M_1[transfer_rule]: "HMA_M (1\<^sub>m (CARD('n))) (mat 1 :: 'a::{zero,one}^'n^'n)" unfolding HMA_M_def by (auto simp add: mat_def from_hma\<^sub>m_def from_nat_inj) lemma from_hma\<^sub>v_add: "from_hma\<^sub>v v + from_hma\<^sub>v w = from_hma\<^sub>v (v + w)" unfolding from_hma\<^sub>v_def by auto lemma HMA_V_add [transfer_rule]: "(HMA_V ===> HMA_V ===> HMA_V) (+) (+) " unfolding rel_fun_def HMA_V_def by (auto simp: from_hma\<^sub>v_add) lemma from_hma\<^sub>v_diff: "from_hma\<^sub>v v - from_hma\<^sub>v w = from_hma\<^sub>v (v - w)" unfolding from_hma\<^sub>v_def by auto lemma HMA_V_diff [transfer_rule]: "(HMA_V ===> HMA_V ===> HMA_V) (-) (-)" unfolding rel_fun_def HMA_V_def by (auto simp: from_hma\<^sub>v_diff) lemma from_hma\<^sub>m_add: "from_hma\<^sub>m a + from_hma\<^sub>m b = from_hma\<^sub>m (a + b)" unfolding from_hma\<^sub>m_def by auto lemma HMA_M_add [transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (+) (+) " unfolding rel_fun_def HMA_M_def by (auto simp: from_hma\<^sub>m_add) lemma from_hma\<^sub>m_diff: "from_hma\<^sub>m a - from_hma\<^sub>m b = from_hma\<^sub>m (a - b)" unfolding from_hma\<^sub>m_def by auto lemma HMA_M_diff [transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (-) (-) " unfolding rel_fun_def HMA_M_def by (auto simp: from_hma\<^sub>m_diff) lemma scalar_product: fixes v :: "'a :: semiring_1 ^ 'n " shows "scalar_prod (from_hma\<^sub>v v) (from_hma\<^sub>v w) = scalar_product v w" unfolding scalar_product_def scalar_prod_def from_hma\<^sub>v_def dim_vec by (simp add: sum.reindex[OF inj_to_nat, unfolded range_to_nat]) lemma [simp]: "from_hma\<^sub>m (y :: 'a ^ 'nc ^ 'nr) \ carrier_mat (CARD('nr)) (CARD('nc))" "dim_row (from_hma\<^sub>m (y :: 'a ^ 'nc ^ 'nr )) = CARD('nr)" "dim_col (from_hma\<^sub>m (y :: 'a ^ 'nc ^ 'nr )) = CARD('nc)" unfolding from_hma\<^sub>m_def by simp_all lemma [simp]: "from_hma\<^sub>v (y :: 'a ^ 'n) \ carrier_vec (CARD('n))" "dim_vec (from_hma\<^sub>v (y :: 'a ^ 'n)) = CARD('n)" unfolding from_hma\<^sub>v_def by simp_all declare rel_funI [intro!] lemma HMA_scalar_prod [transfer_rule]: "(HMA_V ===> HMA_V ===> (=)) scalar_prod scalar_product" by (auto simp: HMA_V_def scalar_product) lemma HMA_row [transfer_rule]: "(HMA_I ===> HMA_M ===> HMA_V) (\ i a. Matrix.row a i) row" unfolding HMA_M_def HMA_I_def HMA_V_def by (auto simp: from_hma\<^sub>m_def from_hma\<^sub>v_def to_nat_less_card row_def) lemma HMA_col [transfer_rule]: "(HMA_I ===> HMA_M ===> HMA_V) (\ i a. col a i) column" unfolding HMA_M_def HMA_I_def HMA_V_def by (auto simp: from_hma\<^sub>m_def from_hma\<^sub>v_def to_nat_less_card column_def) definition mk_mat :: "('i \ 'j \ 'c) \ 'c^'j^'i" where "mk_mat f = (\ i j. f i j)" definition mk_vec :: "('i \ 'c) \ 'c^'i" where "mk_vec f = (\ i. f i)" lemma HMA_M_mk_mat[transfer_rule]: "((HMA_I ===> HMA_I ===> (=)) ===> HMA_M) (\ f. Matrix.mat (CARD('nr)) (CARD('nc)) (\ (i,j). f i j)) (mk_mat :: (('nr \ 'nc \ 'a) \ 'a^'nc^'nr))" proof- { fix x y i j assume id: "\ (ya :: 'nr) (yb :: 'nc). (x (to_nat ya) (to_nat yb) :: 'a) = y ya yb" and i: "i < CARD('nr)" and j: "j < CARD('nc)" from to_nat_from_nat_id[OF i] to_nat_from_nat_id[OF j] id[rule_format, of "from_nat i" "from_nat j"] have "x i j = y (from_nat i) (from_nat j)" by auto } thus ?thesis unfolding rel_fun_def mk_mat_def HMA_M_def HMA_I_def from_hma\<^sub>m_def by auto qed lemma HMA_M_mk_vec[transfer_rule]: "((HMA_I ===> (=)) ===> HMA_V) (\ f. Matrix.vec (CARD('n)) (\ i. f i)) (mk_vec :: (('n \ 'a) \ 'a^'n))" proof- { fix x y i assume id: "\ (ya :: 'n). (x (to_nat ya) :: 'a) = y ya" and i: "i < CARD('n)" from to_nat_from_nat_id[OF i] id[rule_format, of "from_nat i"] have "x i = y (from_nat i)" by auto } thus ?thesis unfolding rel_fun_def mk_vec_def HMA_V_def HMA_I_def from_hma\<^sub>v_def by auto qed lemma mat_mult_scalar: "A ** B = mk_mat (\ i j. scalar_product (row i A) (column j B))" unfolding vec_eq_iff matrix_matrix_mult_def scalar_product_def mk_mat_def by (auto simp: row_def column_def) lemma mult_mat_vec_scalar: "A *v v = mk_vec (\ i. scalar_product (row i A) v)" unfolding vec_eq_iff matrix_vector_mult_def scalar_product_def mk_mat_def mk_vec_def by (auto simp: row_def column_def) lemma dim_row_transfer_rule: "HMA_M A (A' :: 'a ^ 'nc ^ 'nr) \ (=) (dim_row A) (CARD('nr))" unfolding HMA_M_def by auto lemma dim_col_transfer_rule: "HMA_M A (A' :: 'a ^ 'nc ^ 'nr) \ (=) (dim_col A) (CARD('nc))" unfolding HMA_M_def by auto lemma HMA_M_mult [transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) ((*)) ((**))" proof - { fix A B :: "'a :: semiring_1 mat" and A' :: "'a ^ 'n ^ 'nr" and B' :: "'a ^ 'nc ^ 'n" assume 1[transfer_rule]: "HMA_M A A'" "HMA_M B B'" note [transfer_rule] = dim_row_transfer_rule[OF 1(1)] dim_col_transfer_rule[OF 1(2)] have "HMA_M (A * B) (A' ** B')" unfolding times_mat_def mat_mult_scalar by (transfer_prover_start, transfer_step+, transfer, auto) } thus ?thesis by blast qed lemma HMA_V_smult [transfer_rule]: "((=) ===> HMA_V ===> HMA_V) (\\<^sub>v) ((*s))" unfolding smult_vec_def unfolding rel_fun_def HMA_V_def from_hma\<^sub>v_def by auto lemma HMA_M_mult_vec [transfer_rule]: "(HMA_M ===> HMA_V ===> HMA_V) ((*\<^sub>v)) ((*v))" proof - { fix A :: "'a :: semiring_1 mat" and v :: "'a Matrix.vec" and A' :: "'a ^ 'nc ^ 'nr" and v' :: "'a ^ 'nc" assume 1[transfer_rule]: "HMA_M A A'" "HMA_V v v'" note [transfer_rule] = dim_row_transfer_rule have "HMA_V (A *\<^sub>v v) (A' *v v')" unfolding mult_mat_vec_def mult_mat_vec_scalar by (transfer_prover_start, transfer_step+, transfer, auto) } thus ?thesis by blast qed lemma HMA_det [transfer_rule]: "(HMA_M ===> (=)) Determinant.det (det :: 'a :: comm_ring_1 ^ 'n ^ 'n \ 'a)" proof - { fix a :: "'a ^ 'n ^ 'n" let ?tn = "to_nat :: 'n :: finite \ nat" let ?fn = "from_nat :: nat \ 'n" let ?zn = "{0..< CARD('n)}" let ?U = "UNIV :: 'n set" let ?p1 = "{p. p permutes ?zn}" let ?p2 = "{p. p permutes ?U}" let ?f= "\ p i. if i \ ?U then ?fn (p (?tn i)) else i" let ?g = "\ p i. ?fn (p (?tn i))" have fg: "\ a b c. (if a \ ?U then b else c) = b" by auto have "?p2 = ?f ` ?p1" by (rule permutes_bij', auto simp: to_nat_less_card to_nat_from_nat_id) hence id: "?p2 = ?g ` ?p1" by simp have inj_g: "inj_on ?g ?p1" unfolding inj_on_def proof (intro ballI impI ext, auto) fix p q i assume p: "p permutes ?zn" and q: "q permutes ?zn" and id: "(\ i. ?fn (p (?tn i))) = (\ i. ?fn (q (?tn i)))" { fix i from permutes_in_image[OF p] have pi: "p (?tn i) < CARD('n)" by (simp add: to_nat_less_card) from permutes_in_image[OF q] have qi: "q (?tn i) < CARD('n)" by (simp add: to_nat_less_card) from fun_cong[OF id] have "?fn (p (?tn i)) = from_nat (q (?tn i))" . from arg_cong[OF this, of ?tn] have "p (?tn i) = q (?tn i)" by (simp add: to_nat_from_nat_id pi qi) } note id = this show "p i = q i" proof (cases "i < CARD('n)") case True hence "?tn (?fn i) = i" by (simp add: to_nat_from_nat_id) from id[of "?fn i", unfolded this] show ?thesis . next case False thus ?thesis using p q unfolding permutes_def by simp qed qed have mult_cong: "\ a b c d. a = b \ c = d \ a * c = b * d" by simp have "sum (\ p. signof p * (\i\?zn. a $h ?fn i $h ?fn (p i))) ?p1 = sum (\ p. of_int (sign p) * (\i\UNIV. a $h i $h p i)) ?p2" unfolding id sum.reindex[OF inj_g] proof (rule sum.cong[OF refl], unfold mem_Collect_eq o_def, rule mult_cong) fix p assume p: "p permutes ?zn" let ?q = "\ i. ?fn (p (?tn i))" from id p have q: "?q permutes ?U" by auto from p have pp: "permutation p" unfolding permutation_permutes by auto let ?ft = "\ p i. ?fn (p (?tn i))" have fin: "finite ?zn" by simp have "sign p = sign ?q \ p permutes ?zn" using p fin proof (induction rule: permutes_induct) case id show ?case by (auto simp: sign_id[unfolded id_def] permutes_id[unfolded id_def]) next case (swap a b p) + then have \permutation p\ + by (auto intro: permutes_imp_permutation) let ?sab = "Fun.swap a b id" let ?sfab = "Fun.swap (?fn a) (?fn b) id" have p_sab: "permutation ?sab" by (rule permutation_swap_id) have p_sfab: "permutation ?sfab" by (rule permutation_swap_id) from swap(4) have IH1: "p permutes ?zn" and IH2: "sign p = sign (?ft p)" by auto have sab_perm: "?sab permutes ?zn" using swap(1-2) by (rule permutes_swap_id) from permutes_compose[OF IH1 this] have perm1: "?sab o p permutes ?zn" . from IH1 have p_p1: "p \ ?p1" by simp hence "?ft p \ ?ft ` ?p1" by (rule imageI) from this[folded id] have "?ft p permutes ?U" by simp hence p_ftp: "permutation (?ft p)" unfolding permutation_permutes by auto { fix a b assume a: "a \ ?zn" and b: "b \ ?zn" hence "(?fn a = ?fn b) = (a = b)" using swap(1-2) by (auto simp: from_nat_inj) } note inj = this from inj[OF swap(1-2)] have id2: "sign ?sfab = sign ?sab" unfolding sign_swap_id by simp have id: "?ft (Fun.swap a b id \ p) = Fun.swap (?fn a) (?fn b) id \ ?ft p" proof fix c show "?ft (Fun.swap a b id \ p) c = (Fun.swap (?fn a) (?fn b) id \ ?ft p) c" proof (cases "p (?tn c) = a \ p (?tn c) = b") case True - thus ?thesis by (cases, auto simp add: o_def swap_def) + thus ?thesis by (cases, auto simp add: swap_id_eq) next case False hence neq: "p (?tn c) \ a" "p (?tn c) \ b" by auto have pc: "p (?tn c) \ ?zn" unfolding permutes_in_image[OF IH1] by (simp add: to_nat_less_card) from neq[folded inj[OF pc swap(1)] inj[OF pc swap(2)]] have "?fn (p (?tn c)) \ ?fn a" "?fn (p (?tn c)) \ ?fn b" . - with neq show ?thesis by (auto simp: o_def swap_def) + with neq show ?thesis by (auto simp: swap_id_eq) qed qed - show ?case unfolding IH2 id sign_compose[OF p_sab swap(3)] sign_compose[OF p_sfab p_ftp] id2 + show ?case unfolding IH2 id sign_compose[OF p_sab \permutation p\] sign_compose[OF p_sfab p_ftp] id2 by (rule conjI[OF refl perm1]) qed thus "signof p = of_int (sign ?q)" unfolding signof_def sign_def by auto show "(\i = 0..i\UNIV. a $h i $h ?q i)" unfolding range_to_nat[symmetric] prod.reindex[OF inj_to_nat] by (rule prod.cong[OF refl], unfold o_def, simp) qed } thus ?thesis unfolding HMA_M_def by (auto simp: from_hma\<^sub>m_def Determinant.det_def det_def) qed lemma HMA_mat[transfer_rule]: "((=) ===> HMA_M) (\ k. k \\<^sub>m 1\<^sub>m CARD('n)) (Finite_Cartesian_Product.mat :: 'a::semiring_1 \ 'a^'n^'n)" unfolding Finite_Cartesian_Product.mat_def[abs_def] rel_fun_def HMA_M_def by (auto simp: from_hma\<^sub>m_def from_nat_inj) lemma HMA_mat_minus[transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (\ A B. A + map_mat uminus B) ((-) :: 'a :: group_add ^'nc^'nr \ 'a^'nc^'nr \ 'a^'nc^'nr)" unfolding rel_fun_def HMA_M_def from_hma\<^sub>m_def by auto definition mat2matofpoly where "mat2matofpoly A = (\ i j. [: A $ i $ j :])" definition charpoly where charpoly_def: "charpoly A = det (mat (monom 1 (Suc 0)) - mat2matofpoly A)" definition erase_mat :: "'a :: zero ^ 'nc ^ 'nr \ 'nr \ 'nc \ 'a ^ 'nc ^ 'nr" where "erase_mat A i j = (\ i'. \ j'. if i' = i \ j' = j then 0 else A $ i' $ j')" definition sum_UNIV_type :: "('n :: finite \ 'a :: comm_monoid_add) \ 'n itself \ 'a" where "sum_UNIV_type f _ = sum f UNIV" definition sum_UNIV_set :: "(nat \ 'a :: comm_monoid_add) \ nat \ 'a" where "sum_UNIV_set f n = sum f {.. 'n :: finite itself \ bool" where "HMA_T n _ = (n = CARD('n))" lemma HMA_mat2matofpoly[transfer_rule]: "(HMA_M ===> HMA_M) (\x. map_mat (\a. [:a:]) x) mat2matofpoly" unfolding rel_fun_def HMA_M_def from_hma\<^sub>m_def mat2matofpoly_def by auto lemma HMA_char_poly [transfer_rule]: "((HMA_M :: ('a:: comm_ring_1 mat \ 'a^'n^'n \ bool)) ===> (=)) char_poly charpoly" proof - { fix A :: "'a mat" and A' :: "'a^'n^'n" assume [transfer_rule]: "HMA_M A A'" hence [simp]: "dim_row A = CARD('n)" by (simp add: HMA_M_def) have [simp]: "monom 1 (Suc 0) = [:0, 1 :: 'a :]" by (simp add: monom_Suc) have [simp]: "map_mat uminus (map_mat (\a. [:a:]) A) = map_mat (\a. [:-a:]) A" by (rule eq_matI, auto) have "char_poly A = charpoly A'" unfolding char_poly_def[abs_def] char_poly_matrix_def charpoly_def[abs_def] by (transfer, simp) } thus ?thesis by blast qed lemma HMA_eigen_vector [transfer_rule]: "(HMA_M ===> HMA_V ===> (=)) eigenvector eigen_vector" proof - { fix A :: "'a mat" and v :: "'a Matrix.vec" and A' :: "'a ^ 'n ^ 'n" and v' :: "'a ^ 'n" and k :: 'a assume 1[transfer_rule]: "HMA_M A A'" and 2[transfer_rule]: "HMA_V v v'" hence [simp]: "dim_row A = CARD('n)" "dim_vec v = CARD('n)" by (auto simp add: HMA_V_def HMA_M_def) have [simp]: "v \ carrier_vec CARD('n)" using 2 unfolding HMA_V_def by simp have "eigenvector A v = eigen_vector A' v'" unfolding eigenvector_def[abs_def] eigen_vector_def[abs_def] by (transfer, simp) } thus ?thesis by blast qed lemma HMA_eigen_value [transfer_rule]: "(HMA_M ===> (=) ===> (=)) eigenvalue eigen_value" proof - { fix A :: "'a mat" and A' :: "'a ^ 'n ^ 'n" and k assume 1[transfer_rule]: "HMA_M A A'" hence [simp]: "dim_row A = CARD('n)" by (simp add: HMA_M_def) note [transfer_rule] = dim_row_transfer_rule[OF 1(1)] have "(eigenvalue A k) = (eigen_value A' k)" unfolding eigenvalue_def[abs_def] eigen_value_def[abs_def] by (transfer, auto simp add: eigenvector_def) } thus ?thesis by blast qed lemma HMA_spectral_radius [transfer_rule]: "(HMA_M ===> (=)) Spectral_Radius.spectral_radius spectral_radius" unfolding Spectral_Radius.spectral_radius_def[abs_def] spectrum_def spectral_radius_ev_def[abs_def] by transfer_prover lemma HMA_elements_mat[transfer_rule]: "((HMA_M :: ('a mat \ 'a ^ 'nc ^ 'nr \ bool)) ===> (=)) elements_mat elements_mat_h" proof - { fix y :: "'a ^ 'nc ^ 'nr" and i j :: nat assume i: "i < CARD('nr)" and j: "j < CARD('nc)" hence "from_hma\<^sub>m y $$ (i, j) \ range (\(i, ya). y $h i $h ya)" using to_nat_from_nat_id[OF i] to_nat_from_nat_id[OF j] by (auto simp: from_hma\<^sub>m_def) } moreover { fix y :: "'a ^ 'nc ^ 'nr" and a b have "\i j. y $h a $h b = from_hma\<^sub>m y $$ (i, j) \ i < CARD('nr) \ j < CARD('nc)" unfolding from_hma\<^sub>m_def by (rule exI[of _ "Bij_Nat.to_nat a"], rule exI[of _ "Bij_Nat.to_nat b"], auto simp: to_nat_less_card) } ultimately show ?thesis unfolding elements_mat[abs_def] elements_mat_h_def[abs_def] HMA_M_def by auto qed lemma HMA_vec_elements[transfer_rule]: "((HMA_V :: ('a Matrix.vec \ 'a ^ 'n \ bool)) ===> (=)) vec_elements vec_elements_h" proof - { fix y :: "'a ^ 'n" and i :: nat assume i: "i < CARD('n)" hence "from_hma\<^sub>v y $ i \ range (vec_nth y)" using to_nat_from_nat_id[OF i] by (auto simp: from_hma\<^sub>v_def) } moreover { fix y :: "'a ^ 'n" and a have "\i. y $h a = from_hma\<^sub>v y $ i \ i < CARD('n)" unfolding from_hma\<^sub>v_def by (rule exI[of _ "Bij_Nat.to_nat a"], auto simp: to_nat_less_card) } ultimately show ?thesis unfolding vec_elements[abs_def] vec_elements_h_def[abs_def] rel_fun_def HMA_V_def by auto qed lemma norm_bound_elements_mat: "norm_bound A b = (\ x \ elements_mat A. norm x \ b)" unfolding norm_bound_def elements_mat by auto lemma HMA_normbound [transfer_rule]: "((HMA_M :: 'a :: real_normed_field mat \ 'a ^ 'nc ^ 'nr \ bool) ===> (=) ===> (=)) norm_bound normbound" unfolding normbound_def[abs_def] norm_bound_elements_mat[abs_def] by (transfer_prover) lemma HMA_map_matrix [transfer_rule]: "((=) ===> HMA_M ===> HMA_M) map_mat map_matrix" unfolding map_vector_def map_matrix_def[abs_def] map_mat_def[abs_def] HMA_M_def from_hma\<^sub>m_def by auto lemma HMA_transpose_matrix [transfer_rule]: "(HMA_M ===> HMA_M) transpose_mat transpose" unfolding transpose_mat_def transpose_def HMA_M_def from_hma\<^sub>m_def by auto lemma HMA_map_vector [transfer_rule]: "((=) ===> HMA_V ===> HMA_V) map_vec map_vector" unfolding map_vector_def[abs_def] map_vec_def[abs_def] HMA_V_def from_hma\<^sub>v_def by auto lemma HMA_similar_mat_wit [transfer_rule]: "((HMA_M :: _ \ 'a :: comm_ring_1 ^ 'n ^ 'n \ _) ===> HMA_M ===> HMA_M ===> HMA_M ===> (=)) similar_mat_wit similar_matrix_wit" proof (intro rel_funI, goal_cases) case (1 a A b B c C d D) note [transfer_rule] = this hence id: "dim_row a = CARD('n)" by (auto simp: HMA_M_def) have *: "(c * d = 1\<^sub>m (dim_row a) \ d * c = 1\<^sub>m (dim_row a) \ a = c * b * d) = (C ** D = mat 1 \ D ** C = mat 1 \ A = C ** B ** D)" unfolding id by (transfer, simp) show ?case unfolding similar_mat_wit_def Let_def similar_matrix_wit_def * using 1 by (auto simp: HMA_M_def) qed lemma HMA_similar_mat [transfer_rule]: "((HMA_M :: _ \ 'a :: comm_ring_1 ^ 'n ^ 'n \ _) ===> HMA_M ===> (=)) similar_mat similar_matrix" proof (intro rel_funI, goal_cases) case (1 a A b B) note [transfer_rule] = this hence id: "dim_row a = CARD('n)" by (auto simp: HMA_M_def) { fix c d assume "similar_mat_wit a b c d" hence "{c,d} \ carrier_mat CARD('n) CARD('n)" unfolding similar_mat_wit_def id Let_def by auto } note * = this show ?case unfolding similar_mat_def similar_matrix_def by (transfer, insert *, blast) qed lemma HMA_spectrum[transfer_rule]: "(HMA_M ===> (=)) spectrum Spectrum" unfolding spectrum_def[abs_def] Spectrum_def[abs_def] by transfer_prover lemma HMA_M_erase_mat[transfer_rule]: "(HMA_M ===> HMA_I ===> HMA_I ===> HMA_M) mat_erase erase_mat" unfolding mat_erase_def[abs_def] erase_mat_def[abs_def] by (auto simp: HMA_M_def HMA_I_def from_hma\<^sub>m_def to_nat_from_nat_id intro!: eq_matI) lemma HMA_M_sum_UNIV[transfer_rule]: "((HMA_I ===> (=)) ===> HMA_T ===> (=)) sum_UNIV_set sum_UNIV_type" unfolding rel_fun_def proof (clarify, rename_tac f fT n nT) fix f and fT :: "'b \ 'a" and n and nT :: "'b itself" assume f: "\x y. HMA_I x y \ f x = fT y" and n: "HMA_T n nT" let ?f = "from_nat :: nat \ 'b" let ?t = "to_nat :: 'b \ nat" from n[unfolded HMA_T_def] have n: "n = CARD('b)" . from to_nat_from_nat_id[where 'a = 'b, folded n] have tf: "i < n \ ?t (?f i) = i" for i by auto have "sum_UNIV_set f n = sum f (?t ` ?f ` {.. = sum (f \ ?t) (?f ` {.. ?t) i = fT i" unfolding o_def by (rule f[rule_format], auto simp: HMA_I_def) qed also have "\ = sum_UNIV_type fT nT" unfolding sum_UNIV_type_def .. finally show "sum_UNIV_set f n = sum_UNIV_type fT nT" . qed end text \Setup a method to easily convert theorems from JNF into HMA.\ method transfer_hma uses rule = ( (fold index_hma_def)?, (* prepare matrix access for transfer *) transfer, rule rule, (unfold carrier_vec_def carrier_mat_def)?, auto) text \Now it becomes easy to transfer results which are not yet proven in HMA, such as:\ lemma matrix_add_vect_distrib: "(A + B) *v v = A *v v + B *v v" by (transfer_hma rule: add_mult_distrib_mat_vec) lemma matrix_vector_right_distrib: "M *v (v + w) = M *v v + M *v w" by (transfer_hma rule: mult_add_distrib_mat_vec) lemma matrix_vector_right_distrib_diff: "(M :: 'a :: ring_1 ^ 'nr ^ 'nc) *v (v - w) = M *v v - M *v w" by (transfer_hma rule: mult_minus_distrib_mat_vec) lemma eigen_value_root_charpoly: "eigen_value A k \ poly (charpoly (A :: 'a :: field ^ 'n ^ 'n)) k = 0" by (transfer_hma rule: eigenvalue_root_char_poly) lemma finite_spectrum: fixes A :: "'a :: field ^ 'n ^ 'n" shows "finite (Collect (eigen_value A))" by (transfer_hma rule: card_finite_spectrum(1)[unfolded spectrum_def]) lemma non_empty_spectrum: fixes A :: "complex ^ 'n ^ 'n" shows "Collect (eigen_value A) \ {}" by (transfer_hma rule: spectrum_non_empty[unfolded spectrum_def]) lemma charpoly_transpose: "charpoly (transpose A :: 'a :: field ^ 'n ^ 'n) = charpoly A" by (transfer_hma rule: char_poly_transpose_mat) lemma eigen_value_transpose: "eigen_value (transpose A :: 'a :: field ^ 'n ^ 'n) v = eigen_value A v" unfolding eigen_value_root_charpoly charpoly_transpose by simp lemma matrix_diff_vect_distrib: "(A - B) *v v = A *v v - B *v (v :: 'a :: ring_1 ^ 'n)" by (transfer_hma rule: minus_mult_distrib_mat_vec) lemma similar_matrix_charpoly: "similar_matrix A B \ charpoly A = charpoly B" by (transfer_hma rule: char_poly_similar) lemma pderiv_char_poly_erase_mat: fixes A :: "'a :: idom ^ 'n ^ 'n" shows "monom 1 1 * pderiv (charpoly A) = sum (\ i. charpoly (erase_mat A i i)) UNIV" proof - let ?A = "from_hma\<^sub>m A" let ?n = "CARD('n)" have tA[transfer_rule]: "HMA_M ?A A" unfolding HMA_M_def by simp have tN[transfer_rule]: "HMA_T ?n TYPE('n)" unfolding HMA_T_def by simp have A: "?A \ carrier_mat ?n ?n" unfolding from_hma\<^sub>m_def by auto have id: "sum (\ i. charpoly (erase_mat A i i)) UNIV = sum_UNIV_type (\ i. charpoly (erase_mat A i i)) TYPE('n)" unfolding sum_UNIV_type_def .. show ?thesis unfolding id by (transfer, insert pderiv_char_poly_mat_erase[OF A], simp add: sum_UNIV_set_def) qed lemma degree_monic_charpoly: fixes A :: "'a :: comm_ring_1 ^ 'n ^ 'n" shows "degree (charpoly A) = CARD('n) \ monic (charpoly A)" proof (transfer, goal_cases) case 1 from degree_monic_char_poly[OF 1] show ?case by auto qed end diff --git a/thys/Perron_Frobenius/Spectral_Radius_Theory.thy b/thys/Perron_Frobenius/Spectral_Radius_Theory.thy --- a/thys/Perron_Frobenius/Spectral_Radius_Theory.thy +++ b/thys/Perron_Frobenius/Spectral_Radius_Theory.thy @@ -1,271 +1,270 @@ section \Combining Spectral Radius Theory with Perron Frobenius theorem\ theory Spectral_Radius_Theory imports Polynomial_Factorization.Square_Free_Factorization Jordan_Normal_Form.Spectral_Radius Jordan_Normal_Form.Char_Poly Perron_Frobenius "HOL-Computational_Algebra.Field_as_Ring" begin abbreviation spectral_radius where "spectral_radius \ Spectral_Radius.spectral_radius" hide_const (open) Module.smult text \Via JNFs it has been proven that the growth of $A^k$ is polynomially bounded, if all complex eigenvalues have a norm at most 1, i.e., the spectral radius must be at most 1. Moreover, the degree of the polynomial growth can be bounded by the order of those roots which have norm 1, cf. @{thm spectral_radius_poly_bound}.\ text \Perron Frobenius theorem tells us that for a real valued non negative matrix, the largest eigenvalue is a real non-negative one. Hence, we only have to check, that all real eigenvalues are at most one.\ text \We combine both theorems in the following. To be more precise, the set-based complexity results from JNFs with the type-based Perron Frobenius theorem in HMA are connected to obtain a set based complexity criterion for real-valued non-negative matrices, where one only investigated the real valued eigenvalues for checking the eigenvalue-at-most-1 condition. Here, in the precondition of the roots of the polynomial, the type-system ensures that we only have to look at real-valued eigenvalues, and can ignore the complex-valued ones. The linkage between set-and type-based is performed via HMA-connect.\ lemma perron_frobenius_spectral_radius_complex: fixes A :: "complex mat" assumes A: "A \ carrier_mat n n" and real_nonneg: "real_nonneg_mat A" and ev_le_1: "\ x. poly (char_poly (map_mat Re A)) x = 0 \ x \ 1" and ev_order: "\ x. norm x = 1 \ order x (char_poly A) \ d" shows "\c1 c2. \k. norm_bound (A ^\<^sub>m k) (c1 + c2 * real k ^ (d - 1))" proof (cases "n = 0") case False hence n: "n > 0" "n \ 0" by auto define sr where "sr = spectral_radius A" note sr = spectral_radius_mem_max[OF A n(1), folded sr_def] show ?thesis proof (rule spectral_radius_poly_bound[OF A], unfold sr_def[symmetric]) let ?cr = "complex_of_real" text \here is the transition from type-based perron-frobenius to set-based\ from perron_frobenius[untransferred, cancel_card_constraint, OF A real_nonneg n(2)] obtain v where v: "v \ carrier_vec n" and ev: "eigenvector A v (?cr sr)" and rnn: "real_nonneg_vec v" unfolding sr_def by auto define B where "B = map_mat Re A" let ?A = "map_mat ?cr B" have AB: "A = ?A" unfolding B_def by (rule eq_matI, insert real_nonneg[unfolded real_nonneg_mat_def elements_mat_def], auto) define w where "w = map_vec Re v" let ?v = "map_vec ?cr w" have vw: "v = ?v" unfolding w_def by (rule eq_vecI, insert rnn[unfolded real_nonneg_vec_def vec_elements_def], auto) have B: "B \ carrier_mat n n" unfolding B_def using A by auto from AB vw ev have ev: "eigenvector ?A ?v (?cr sr)" by simp have "eigenvector B w sr" by (rule of_real_hom.eigenvector_hom_rev[OF B ev]) hence "eigenvalue B sr" unfolding eigenvalue_def by blast from ev_le_1[folded B_def, OF this[unfolded eigenvalue_root_char_poly[OF B]]] show "sr \ 1" . next fix ev assume "cmod ev = 1" thus "order ev (char_poly A) \ d" by (rule ev_order) qed next case True with A show ?thesis by (intro exI[of _ 0], auto simp: norm_bound_def) qed text \The following lemma is the same as @{thm perron_frobenius_spectral_radius_complex}, except that now the type @{typ real} is used instead of @{typ complex}.\ lemma perron_frobenius_spectral_radius: fixes A :: "real mat" assumes A: "A \ carrier_mat n n" and nonneg: "nonneg_mat A" and ev_le_1: "\ x. poly (char_poly A) x = 0 \ x \ 1" and ev_order: "\ x :: complex. norm x = 1 \ order x (map_poly of_real (char_poly A)) \ d" shows "\c1 c2. \k a. a \ elements_mat (A ^\<^sub>m k) \ abs a \ (c1 + c2 * real k ^ (d - 1))" proof - let ?cr = "complex_of_real" let ?B = "map_mat ?cr A" have B: "?B \ carrier_mat n n" using A by auto have rnn: "real_nonneg_mat ?B" using nonneg unfolding real_nonneg_mat_def nonneg_mat_def by (auto simp: elements_mat_def) have id: "map_mat Re ?B = A" by (rule eq_matI, auto) have "\c1 c2. \k. norm_bound (?B ^\<^sub>m k) (c1 + c2 * real k ^ (d - 1))" by (rule perron_frobenius_spectral_radius_complex[OF B rnn], unfold id, insert ev_le_1 ev_order, auto simp: of_real_hom.char_poly_hom[OF A]) then obtain c1 c2 where nb: "\ k. norm_bound (?B ^\<^sub>m k) (c1 + c2 * real k ^ (d - 1))" by auto show ?thesis proof (rule exI[of _ c1], rule exI[of _ c2], intro allI impI) fix k a assume "a \ elements_mat (A ^\<^sub>m k)" with pow_carrier_mat[OF A] obtain i j where a: "a = (A ^\<^sub>m k) $$ (i,j)" and ij: "i < n" "j < n" unfolding elements_mat by force from ij nb[of k] A have "norm ((?B ^\<^sub>m k) $$ (i,j)) \ c1 + c2 * real k ^ (d - 1)" unfolding norm_bound_def by auto also have "(?B ^\<^sub>m k) $$ (i,j) = ?cr a" unfolding of_real_hom.mat_hom_pow[OF A, symmetric] a using ij A by auto also have "norm (?cr a) = abs a" by auto finally show "abs a \ (c1 + c2 * real k ^ (d - 1))" . qed qed text \We can also convert the set-based lemma @{thm perron_frobenius_spectral_radius} to a type-based version.\ lemma perron_frobenius_spectral_type_based: assumes "non_neg_mat (A :: real ^ 'n ^ 'n)" and "\ x. poly (charpoly A) x = 0 \ x \ 1" and "\ x :: complex. norm x = 1 \ order x (map_poly of_real (charpoly A)) \ d" shows "\c1 c2. \k a. a \ elements_mat_h (matpow A k) \ abs a \ (c1 + c2 * real k ^ (d - 1))" using assms perron_frobenius_spectral_radius by (transfer, blast) text \And of course, we can also transfer the type-based lemma back to a set-based setting, only that -- without further case-analysis -- we get the additional assumption @{term "(n :: nat) \ 0"}.\ lemma assumes "A \ carrier_mat n n" and "nonneg_mat A" and "\ x. poly (char_poly A) x = 0 \ x \ 1" and "\ x :: complex. norm x = 1 \ order x (map_poly of_real (char_poly A)) \ d" and "n \ 0" shows "\c1 c2. \k a. a \ elements_mat (A ^\<^sub>m k) \ abs a \ (c1 + c2 * real k ^ (d - 1))" using perron_frobenius_spectral_type_based[untransferred, cancel_card_constraint, OF assms] . text \Note that the precondition eigenvalue-at-most-1 can easily be formulated as a cardinality constraints which can be decided by Sturm's theorem. And in order to obtain a bound on the order, one can perform a square-free-factorization (via Yun's factorization algorithm) of the characteristic polynomial into $f_1^1 \cdot \ldots f_d^d$ where each $f_i$ has precisely the roots of order $i$.\ context fixes A :: "real mat" and c :: real and fis and n :: nat assumes A: "A \ carrier_mat n n" and nonneg: "nonneg_mat A" and yun: "yun_factorization gcd (char_poly A) = (c,fis)" and ev_le_1: "card {x. poly (char_poly A) x = 0 \ x > 1} = 0" begin text \Note that @{const yun_factorization} has an offset by 1, so the pair @{term "(f\<^sub>i,i) \ set fis"} encodes @{term "f\<^sub>i^(Suc i)"}.\ lemma perron_frobenius_spectral_radius_yun: assumes bnd: "\ f\<^sub>i i. (f\<^sub>i,i) \ set fis \ (\ x :: complex. poly (map_poly of_real f\<^sub>i) x = 0 \ norm x = 1) \ Suc i \ d" shows "\c1 c2. \k a. a \ elements_mat (A ^\<^sub>m k) \ abs a \ (c1 + c2 * real k ^ (d - 1))" proof (rule perron_frobenius_spectral_radius[OF A nonneg]; intro allI impI) let ?cr = complex_of_real let ?cp = "map_poly ?cr (char_poly A)" fix x :: complex assume x: "norm x = 1" have A0: "char_poly A \ 0" using degree_monic_char_poly[OF A] by auto interpret field_hom_0' ?cr by (standard, auto) from A0 have cp0: "?cp \ 0" by auto obtain ox where ox: "order x ?cp = ox" by blast note sff = square_free_factorization_order_root[OF yun_factorization(1)[OF yun_factorization_hom[of "char_poly A", unfolded yun map_prod_def split]] cp0, of x ox, unfolded ox] show "order x ?cp \ d" unfolding ox proof (cases ox) case (Suc oo) with sff obtain fi where mem: "(fi,oo) \ set fis" and rt: "poly (map_poly ?cr fi) x = 0" by auto from bnd[OF mem exI[of _ x], OF conjI[OF rt x]] show "ox \ d" unfolding Suc . qed auto next let ?L = "{x. poly (char_poly A) x = 0 \ x > 1}" fix x :: real assume rt: "poly (char_poly A) x = 0" have "finite ?L" by (rule finite_subset[OF _ poly_roots_finite[of "char_poly A"]], insert degree_monic_char_poly[OF A], auto) with ev_le_1 have "?L = {}" by simp with rt show "x \ 1" by auto qed text \Note that the only remaining problem in applying @{thm perron_frobenius_spectral_radius_yun} is to check the condition @{term "\ x :: complex. poly (map_poly of_real f\<^sub>i) x = 0 \ norm x = 1"}. Here, there are at least three possibilities. First, one can just ignore this precondition and weaken the statement. Second, one can apply Sturm's theorem to determine whether all roots are real. This can be done by comparing the number of distinct real roots with the degree of @{term f\<^sub>i}, since @{term f\<^sub>i} is square-free. If all roots are real, then one can decide the criterion by checking the only two possible real roots with norm equal to 1, namely 1 and -1. If on the other hand there are complex roots, then we loose precision at this point. Third, one uses a factorization algorithm (e.g., via complex algebraic numbers) to precisely determine the complex roots and decide the condition. The second approach is illustrated in the following theorem. Note that all preconditions -- including the ones from the context -- can easily be checked with the help of Sturm's method. This method is used as a fast approximative technique in CeTA \cite{CeTA}. Only if the desired degree cannot be ensured by this method, the more costly complex algebraic number based factorization is applied.\ lemma perron_frobenius_spectral_radius_yun_real_roots: assumes bnd: "\ f\<^sub>i i. (f\<^sub>i,i) \ set fis \ card { x. poly f\<^sub>i x = 0} \ degree f\<^sub>i \ poly f\<^sub>i 1 = 0 \ poly f\<^sub>i (-1) = 0 \ Suc i \ d" shows "\c1 c2. \k a. a \ elements_mat (A ^\<^sub>m k) \ abs a \ (c1 + c2 * real k ^ (d - 1))" proof (rule perron_frobenius_spectral_radius_yun) fix fi i let ?cr = complex_of_real let ?cp = "map_poly ?cr" assume fi: "(fi, i) \ set fis" and "\ x. poly (map_poly ?cr fi) x = 0 \ norm x = 1" then obtain x where rt: "poly (?cp fi) x = 0" and x: "norm x = 1" by auto show "Suc i \ d" proof (rule bnd[OF fi]) consider (c) "x \ \" | (1) "x = 1" | (m1) "x = -1" | (r) "x \ \" "x \ {1, -1}" by (cases "x \ \"; auto) thus "card {x. poly fi x = 0} \ degree fi \ poly fi 1 = 0 \ poly fi (- 1) = 0" proof (cases) case 1 from rt have "poly fi 1 = 0" unfolding 1 by simp thus ?thesis by simp next case m1 have id: "-1 = ?cr (-1)" by simp from rt have "poly fi (-1) = 0" unfolding m1 id of_real_hom.hom_zero[where 'a=complex,symmetric] of_real_hom.poly_map_poly by simp thus ?thesis by simp next case r then obtain y where xy: "x = of_real y" unfolding Reals_def by auto from r(2)[unfolded xy] have y: "y \ {1,-1}" by auto from x[unfolded xy] have "abs y = 1" by auto with y have False by auto thus ?thesis .. next case c from yun_factorization(2)[OF yun] fi have "monic fi" by auto hence fi: "?cp fi \ 0" by auto hence fin: "finite {x. poly (?cp fi) x = 0}" by (rule poly_roots_finite) have "?cr ` {x. poly (?cp fi) (?cr x) = 0} \ {x. poly (?cp fi) x = 0}" (is "?l \ ?r") proof (rule, force) have "x \ ?r" using rt by auto moreover have "x \ ?l" using c unfolding Reals_def by auto ultimately show "?l \ ?r" by blast qed from psubset_card_mono[OF fin this] have "card ?l < card ?r" . also have "\ \ degree (?cp fi)" by (rule poly_roots_degree[OF fi]) also have "\ = degree fi" by simp also have "?l = ?cr ` {x. poly fi x = 0}" by auto also have "card \ = card {x. poly fi x = 0}" by (rule card_image, auto simp: inj_on_def) finally have "card {x. poly fi x = 0} \ degree fi" by simp thus ?thesis by auto qed qed qed end -thm perron_frobenius_spectral_radius_yun_real_roots end diff --git a/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy b/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy --- a/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy +++ b/thys/Planarity_Certificates/Planarity/Planar_Subgraph.thy @@ -1,1857 +1,1857 @@ theory Planar_Subgraph imports Graph_Genus Permutations_2 "HOL-Library.FuncSet" "HOL-Library.Simps_Case_Conv" begin section \Combinatorial Planarity and Subgraphs\ lemma out_arcs_emptyD_dominates: assumes "out_arcs G x = {}" shows "\x \\<^bsub>G\<^esub> y" using assms by (auto simp: out_arcs_def) lemma (in wf_digraph) reachable_refl_iff: "u \\<^sup>* u \ u \ verts G" by (auto simp: reachable_in_verts) context digraph_map begin lemma face_cycle_set_succ[simp]: "face_cycle_set (face_cycle_succ a) = face_cycle_set a" by (metis face_cycle_eq face_cycle_set_self face_cycle_succ_inD) lemma face_cycle_succ_funpow_in[simp]: "(face_cycle_succ ^^ n) a \ arcs G \ a \ arcs G" by (induct n) auto lemma segment_face_cycle_x_x_eq: "segment face_cycle_succ x x = face_cycle_set x - {x}" unfolding face_cycle_set_def using face_cycle_succ_permutes finite_arcs permutation_permutes by (intro segment_x_x_eq) blast lemma fcs_x_eq_x: "face_cycle_succ x = x \ face_cycle_set x = {x}" (is "?L \ ?R") unfolding face_cycle_set_def orbit_eq_singleton_iff .. end lemma (in bidirected_digraph) bidirected_digraph_del_arc: "bidirected_digraph (pre_digraph.del_arc (pre_digraph.del_arc G (arev a)) a) (perm_restrict arev (arcs G - {a , arev a}))" proof unfold_locales fix b assume A: "b \ arcs (pre_digraph.del_arc (del_arc (arev a)) a)" have "arev b \ b \ b \ arev a \ b \ a \ perm_restrict arev (arcs G - {a, arev a}) (arev b) = b" using bij_arev arev_dom by (subst perm_restrict_simps) (auto simp: bij_iff) then show "perm_restrict arev (arcs G - {a, arev a}) (perm_restrict arev (arcs G - {a, arev a}) b) = b" using A by (auto simp: pre_digraph.del_arc_simps perm_restrict_simps arev_dom) qed (auto simp: pre_digraph.del_arc_simps perm_restrict_simps arev_dom) lemma (in bidirected_digraph) bidirected_digraph_del_vert: "bidirected_digraph (del_vert u) (perm_restrict arev (arcs (del_vert u)))" by unfold_locales (auto simp: del_vert_simps perm_restrict_simps arev_dom) lemma (in pre_digraph) ends_del_arc: "arc_to_ends (del_arc u) = arc_to_ends G" by (simp add: arc_to_ends_def fun_eq_iff) lemma (in pre_digraph) dominates_arcsD: assumes "v \\<^bsub>del_arc u\<^esub> w" shows "v \\<^bsub>G\<^esub> w" using assms by (auto simp: arcs_ends_def ends_del_arc) lemma (in wf_digraph) reachable_del_arcD: assumes "v \\<^sup>*\<^bsub>del_arc u\<^esub> w" shows "v \\<^sup>*\<^bsub>G\<^esub> w" proof - interpret H: wf_digraph "del_arc u" by (rule wf_digraph_del_arc) from assms show ?thesis by (induct) (auto dest: dominates_arcsD intro: adj_reachable_trans) qed lemma (in fin_digraph) finite_isolated_verts[intro!]: "finite isolated_verts" by (auto simp: isolated_verts_def) lemma (in wf_digraph) isolated_verts_in_sccs: assumes "u \ isolated_verts" shows "{u} \ sccs_verts" proof - have "v = u" if "u \\<^sup>*\<^bsub>G\<^esub> v" for v using that assms by induct (auto simp: arcs_ends_def arc_to_ends_def isolated_verts_def) with assms show ?thesis by (auto simp: sccs_verts_def isolated_verts_def) qed lemma (in digraph_map) in_face_cycle_sets: "a \ arcs G \ face_cycle_set a \ face_cycle_sets" by (auto simp: face_cycle_sets_def) lemma (in digraph_map) heads_face_cycle_set: assumes "a \ arcs G" shows "head G ` face_cycle_set a = tail G ` face_cycle_set a" (is "?L = ?R") proof (intro set_eqI iffI) fix u assume "u \ ?L" then obtain b where "b \ face_cycle_set a" "head G b = u" by blast then have "face_cycle_succ b \ face_cycle_set a" "tail G (face_cycle_succ b) = u" using assms by (auto simp: tail_face_cycle_succ face_cycle_succ_inI in_face_cycle_setD) then show "u \ ?R" by auto next fix u assume "u \ ?R" then obtain b where "b \ face_cycle_set a" "tail G b = u" by blast moreover then obtain c where "b = face_cycle_succ c" by (metis face_cycle_succ_pred) ultimately have "c \ face_cycle_set a" "head G c = u" by (auto dest: face_cycle_succ_inD) (metis assms face_cycle_succ_no_arc in_face_cycle_setD tail_face_cycle_succ) then show "u \ ?L" by auto qed lemma (in pre_digraph) casI_nth: assumes "p \ []" "u = tail G (hd p)" "v = head G (last p)" "\i. Suc i < length p \ head G (p ! i) = tail G (p ! Suc i)" shows "cas u p v" using assms proof (induct p arbitrary: u) case Nil then show ?case by simp next case (Cons a p) have "cas (head G a) p v" proof (cases "p = []") case False then show ?thesis using Cons.prems(1-3) Cons.prems(4)[of 0] Cons.prems(4)[of "Suc i" for i] by (intro Cons) (simp_all add: hd_conv_nth) qed (simp add: Cons) with Cons show ?case by simp qed lemma (in digraph_map) obtain_trail_in_fcs: assumes "a \ arcs G" "a0 \ face_cycle_set a" "an \ face_cycle_set a" obtains p where "trail (tail G a0) p (head G an)" "p \ []" "hd p = a0" "last p = an" "set p \ face_cycle_set a" proof - have fcs_a: "face_cycle_set a = orbit face_cycle_succ a0" using assms face_cycle_eq by (simp add: face_cycle_set_def) have "a0 = (face_cycle_succ ^^ 0) a0" by simp have "an = (face_cycle_succ ^^ funpow_dist face_cycle_succ a0 an) a0" using assms by (simp add: fcs_a funpow_dist_prop) define p where "p = map (\n. (face_cycle_succ ^^ n) a0) [0..i. i < length p \ p ! i = (face_cycle_succ ^^ i) a0" by (auto simp: p_def simp del: upt_Suc) have P2: "p \ []" by (simp add: p_def) have P3: "hd p = a0" using \a0 = _\ by (auto simp: p_def hd_map simp del: upt_Suc) have P4: "last p = an" using \an = _\ by (simp add: p_def) have P5: "set p \ face_cycle_set a" unfolding p_def fcs_a orbit_altdef_permutation[OF permutation_face_cycle_succ] by auto have P1: "trail (tail G a0) p (head G an)" proof - have "distinct p" proof - have "an \ orbit face_cycle_succ a0" using assms by (simp add: fcs_a) then have "inj_on (\n. (face_cycle_succ ^^ n) a0) {0..funpow_dist face_cycle_succ a0 an}" by (rule inj_on_funpow_dist) also have "{0..funpow_dist face_cycle_succ a0 an} = (set [0..n. (face_cycle_succ ^^ n) a0) (set [0.. arcs G" by (metis assms(1-2) in_face_cycle_setD) then have "tail G a0 \ verts G" by simp moreover have "set p \ arcs G" using P5 by (metis assms(1) in_face_cycle_setD subset_code(1)) moreover then have "\i. Suc i < length p \ p ! Suc i \ arcs G" by auto then have "\i. Suc i < length p \ head G (p ! i) = tail G (p ! Suc i)" by (auto simp: p_nth tail_face_cycle_succ) ultimately show ?thesis using P2 P3 P4 unfolding trail_def awalk_def by (auto intro: casI_nth) qed from P1 P2 P3 P4 P5 show ?thesis .. qed lemma (in digraph_map) obtain_trail_in_fcs': assumes "a \ arcs G" "u \ tail G ` face_cycle_set a" "v \ tail G ` face_cycle_set a" obtains p where "trail u p v" "set p \ face_cycle_set a" proof - from assms obtain a0 where "tail G a0 = u" "a0 \ face_cycle_set a" by auto moreover from assms obtain an where "head G an = v" "an \ face_cycle_set a" by (auto simp: heads_face_cycle_set[symmetric]) ultimately obtain p where "trail u p v" "set p \ face_cycle_set a" using \a \ arcs G\ by (metis obtain_trail_in_fcs) then show ?thesis .. qed subsection \Deleting an isolated vertex\ locale del_vert_props = digraph_map + fixes u assumes u_in: "u \ verts G" assumes u_isolated: "out_arcs G u = {}" begin lemma u_isolated_in: "in_arcs G u = {}" using u_isolated by (simp add: in_arcs_eq) lemma arcs_dv: "arcs (del_vert u) = arcs G" using u_isolated u_isolated_in by (auto simp: del_vert_simps) lemma out_arcs_dv: "out_arcs (del_vert u) = out_arcs G" by (auto simp: fun_eq_iff arcs_dv tail_del_vert) lemma digraph_map_del_vert: shows "digraph_map (del_vert u) M" proof - have "perm_restrict (edge_rev M) (arcs (del_vert u)) = edge_rev M" using has_dom_arev arcs_dv by (auto simp: perm_restrict_dom_subset) then interpret H: bidirected_digraph "del_vert u" "edge_rev M" using bidirected_digraph_del_vert[of u] by simp show ?thesis by unfold_locales (auto simp: arcs_dv edge_succ_permutes out_arcs_dv edge_succ_cyclic verts_del_vert) qed end sublocale del_vert_props \ H: digraph_map "del_vert u" M by (rule digraph_map_del_vert) context del_vert_props begin lemma card_verts_dv: "card (verts G) = Suc (card (verts (del_vert u)))" by (auto simp: verts_del_vert) (rule card.remove, auto simp: u_in) lemma card_arcs_dv: "card (arcs (del_vert u)) = card (arcs G)" using u_isolated by (auto simp add: arcs_dv in_arcs_eq) lemma isolated_verts_dv: "H.isolated_verts = isolated_verts - {u}" by (auto simp: isolated_verts_def H.isolated_verts_def verts_del_vert out_arcs_dv) lemma u_in_isolated_verts: "u \ isolated_verts" using u_in u_isolated by (auto simp: isolated_verts_def) lemma card_isolated_verts_dv: "card isolated_verts = Suc (card H.isolated_verts)" by (simp add: isolated_verts_dv) (rule card.remove, auto simp: u_in_isolated_verts) lemma face_cycles_dv: "H.face_cycle_sets = face_cycle_sets" unfolding H.face_cycle_sets_def face_cycle_sets_def arcs_dv .. lemma euler_char_dv: "euler_char = 1 + H.euler_char" by (auto simp: euler_char_def H.euler_char_def card_arcs_dv card_verts_dv face_cycles_dv) lemma adj_dv: "v \\<^bsub>del_vert u\<^esub> w \ v \\<^bsub>G\<^esub> w" by (auto simp: arcs_ends_def arcs_dv ends_del_vert) lemma reachable_del_vertD: assumes "v \\<^sup>*\<^bsub>del_vert u\<^esub> w" shows "v \\<^sup>*\<^bsub>G\<^esub> w" using assms by induct (auto simp: verts_del_vert adj_dv intro: adj_reachable_trans) lemma reachable_del_vertI: assumes "v \\<^sup>*\<^bsub>G\<^esub> w" "u \ v \ u \ w" shows "v \\<^sup>*\<^bsub>del_vert u\<^esub> w" using assms proof induct case (step x y) from \x \\<^bsub>G\<^esub> y\ obtain a where "a \ arcs G" "head G a = y" by auto then have "a \ in_arcs G y" by auto then have "y \ u" using u_isolated in_arcs_eq[of u] by auto with step show ?case by (auto simp: adj_dv intro: H.adj_reachable_trans) qed (auto simp: verts_del_vert) lemma G_reach_conv: "v \\<^sup>*\<^bsub>G\<^esub> w \ v \\<^sup>*\<^bsub>del_vert u\<^esub> w \ (v = u \ w = u)" by (auto dest: reachable_del_vertI reachable_del_vertD intro: u_in) lemma sccs_verts_dv: "H.sccs_verts = sccs_verts - {{u}}" (is "?L = ?R") proof - have *:"\S x. S \ sccs_verts \ S \ H.sccs_verts \ x \ S \ x = u" by (simp add: H.in_sccs_verts_conv_reachable in_sccs_verts_conv_reachable G_reach_conv) (meson H.reachable_trans) show ?thesis by (auto dest: *) (auto simp: H.in_sccs_verts_conv_reachable in_sccs_verts_conv_reachable G_reach_conv H.reachable_refl_iff verts_del_vert) qed lemma card_sccs_verts_dv: "card sccs_verts = Suc (card H.sccs_verts)" unfolding sccs_verts_dv by (rule card.remove) (auto simp: isolated_verts_in_sccs u_in_isolated_verts finite_sccs_verts) lemma card_sccs_dv: "card sccs = Suc (card H.sccs)" using card_sccs_verts_dv by (simp add: card_sccs_verts H.card_sccs_verts) lemma euler_genus_eq: "H.euler_genus = euler_genus" by (auto simp: pre_digraph_map.euler_genus_def card_sccs_dv card_isolated_verts_dv euler_char_dv) end subsection \Deleting an arc pair\ locale bidel_arc = G: digraph_map + fixes a assumes a_in: "a \ arcs G" begin abbreviation "a' \ edge_rev M a" definition H :: "('a,'b) pre_digraph" where "H \ pre_digraph.del_arc (pre_digraph.del_arc G a') a" definition HM :: "'b pre_map" where "HM = \ edge_rev = perm_restrict (edge_rev M) (arcs G - {a, a'}) , edge_succ = perm_rem a (perm_rem a' (edge_succ M)) \" lemma verts_H: "verts H = verts G" and arcs_H: "arcs H = arcs G - {a, a'}" and tail_H: "tail H = tail G" and head_H: "head H = head G" and ends_H: "arc_to_ends H = arc_to_ends G"and arcs_in: "{a,a'} \ arcs G" and ends_in: "{tail G a, head G a} \ verts G" by (auto simp: H_def pre_digraph.del_arc_simps a_in arc_to_ends_def) lemma cyclic_on_edge_succ: assumes "x \ verts H" "out_arcs H x \ {}" shows "cyclic_on (edge_succ HM) (out_arcs H x)" proof - have oa_H: "out_arcs H x = (out_arcs G x - {a'}) - {a}" by (auto simp: arcs_H tail_H) have "cyclic_on (perm_rem a (perm_rem a' (edge_succ M))) (out_arcs G x - {a'} - {a})" using assms by (intro cyclic_on_perm_rem G.edge_succ_cyclic) (auto simp: oa_H G.bij_edge_succ G.edge_succ_cyclic) then show ?thesis by (simp add: HM_def oa_H) qed lemma digraph_map: "digraph_map H HM" proof - interpret fin_digraph H unfolding H_def by (rule fin_digraph.fin_digraph_del_arc) (rule G.fin_digraph_del_arc) interpret bidirected_digraph H "edge_rev HM" unfolding H_def using G.bidirected_digraph_del_arc[of a] by (auto simp: HM_def) have *: "insert a' (insert a (arcs H)) = arcs G" using a_in by (auto simp: arcs_H) have "edge_succ HM permutes arcs H" unfolding HM_def by (auto simp: * intro!: perm_rem_permutes G.edge_succ_permutes) moreover { fix v assume "v \ verts H" "out_arcs H v \ {}" then have "cyclic_on (edge_succ HM) (out_arcs H v)" by (rule cyclic_on_edge_succ) } ultimately show ?thesis by unfold_locales qed lemma rev_H: "bidel_arc.H G M a' = H" (is ?t1) and rev_HM: "bidel_arc.HM G M a' = HM" (is ?t2) proof - interpret rev: bidel_arc G M a' using a_in by unfold_locales simp show ?t1 by (rule pre_digraph.equality) (auto simp: rev.verts_H verts_H rev.arcs_H arcs_H rev.tail_H tail_H rev.head_H head_H) show ?t2 using G.edge_succ_permutes by (intro pre_map.equality) (auto simp: HM_def rev.HM_def insert_commute perm_rem_commutes permutes_conv_has_dom) qed end sublocale bidel_arc \ H: digraph_map H HM by (rule digraph_map) context bidel_arc begin lemma a_neq_a': "a \ a'" by (metis G.arev_neq a_in) lemma arcs_G: "arcs G = insert a (insert a' (arcs H))" and arcs_not_in: "{a,a'} \ arcs H = {}" using arcs_in by (auto simp: arcs_H) lemma card_arcs_da: "card (arcs G) = 2 + card (arcs H)" using arcs_G arcs_not_in a_neq_a' by (auto simp: card_insert_if) lemma cas_da: "H.cas = G.cas" proof - { fix u p v have "H.cas u p v = G.cas u p v" by (induct p arbitrary: u) (simp_all add: tail_H head_H) } then show ?thesis by (simp add: fun_eq_iff) qed lemma reachable_daD: assumes "v \\<^sup>*\<^bsub>H\<^esub> w" shows "v \\<^sup>*\<^bsub>G\<^esub> w" apply (rule G.reachable_del_arcD) apply (rule wf_digraph.reachable_del_arcD) apply (rule G.wf_digraph_del_arc) using assms unfolding H_def by assumption lemma not_G_isolated_a: "{tail G a, head G a} \ G.isolated_verts = {}" using a_in G.in_arcs_eq[of "head G a"] by (auto simp: G.isolated_verts_def) lemma isolated_other_da: assumes "u \ {tail G a, head G a}" shows "u \ H.isolated_verts \ u \ G.isolated_verts" using assms by (auto simp: pre_digraph.isolated_verts_def verts_H arcs_H tail_H out_arcs_def) lemma isolated_da_pre: "H.isolated_verts = G.isolated_verts \ (if tail G a \ H.isolated_verts then {tail G a} else {}) \ (if head G a \ H.isolated_verts then {head G a} else {})" (is "?L = ?R") proof (intro set_eqI iffI) fix x assume "x \ ?L" then show "x \ ?R" by (cases "x \ {tail G a, head G a}") (auto simp:isolated_other_da) next fix x assume "x \ ?R" then show "x \ ?L" using not_G_isolated_a by (cases "x \ {tail G a, head G a}") (auto simp:isolated_other_da split: if_splits) qed lemma card_isolated_verts_da0: "card H.isolated_verts = card G.isolated_verts + card ({tail G a, head G a} \ H.isolated_verts)" using not_G_isolated_a by (subst isolated_da_pre) (auto simp: card_insert_if G.finite_isolated_verts) lemma segments_neq: assumes "segment G.face_cycle_succ a' a \ {} \ segment G.face_cycle_succ a a' \ {}" shows "segment G.face_cycle_succ a a' \ segment G.face_cycle_succ a' a" proof - have bij_fcs: "bij G.face_cycle_succ" using G.face_cycle_succ_permutes by (auto simp: permutes_conv_has_dom) show ?thesis using segment_disj[OF a_neq_a' bij_fcs] assms by auto qed lemma H_fcs_eq_G_fcs: assumes "b \ arcs G" "{b,G.face_cycle_succ b} \ {a,a'} = {}" shows "H.face_cycle_succ b = G.face_cycle_succ b" proof - have "edge_rev M b \ {a,a'}" using assms by auto (metis G.arev_arev) then show ?thesis using assms unfolding G.face_cycle_succ_def H.face_cycle_succ_def by (auto simp: HM_def perm_restrict_simps perm_rem_simps G.bij_edge_succ) qed lemma face_cycle_set_other_da: assumes "{a, a'} \ G.face_cycle_set b = {}" "b \ arcs G" shows "H.face_cycle_set b = G.face_cycle_set b" proof - have "\s. s \ G.face_cycle_set b \ b \ arcs G \ a \ G.face_cycle_set b \ a' \ G.face_cycle_set b \ pre_digraph_map.face_cycle_succ HM s = G.face_cycle_succ s" by (subst H_fcs_eq_G_fcs) (auto simp: G.in_face_cycle_setD G.face_cycle_succ_inI) then show ?thesis using assms unfolding pre_digraph_map.face_cycle_set_def by (intro orbit_cong) (auto simp add: pre_digraph_map.face_cycle_set_def[symmetric]) qed lemma in_face_cycle_set_other: assumes "S \ G.face_cycle_sets" "{a, a'} \ S = {}" shows "S \ H.face_cycle_sets" proof - from assms obtain b where "S = G.face_cycle_set b" "b \ arcs G" by (auto simp: G.face_cycle_sets_def) with assms have "S = H.face_cycle_set b" by (simp add: face_cycle_set_other_da) moreover with assms have "b \ arcs H" using \b \ arcs G\ by (auto simp: arcs_H) ultimately show ?thesis by (auto simp: H.face_cycle_sets_def) qed lemma H_fcs_in_G_fcs: assumes "b \ arcs H - (G.face_cycle_set a \ G.face_cycle_set a')" shows "H.face_cycle_set b \ G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'}" proof - have "H.face_cycle_set b = G.face_cycle_set b" using assms by (intro face_cycle_set_other_da) (auto simp: arcs_H G.face_cycle_eq) moreover have "G.face_cycle_set b \ {G.face_cycle_set a, G.face_cycle_set a'}" "b \ arcs G" using G.face_cycle_eq assms by (auto simp: arcs_H) ultimately show ?thesis by (auto simp: G.face_cycle_sets_def) qed lemma face_cycle_sets_da0: "H.face_cycle_sets = G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'} \ H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a,a'})" (is "?L = ?R") proof (intro set_eqI iffI) fix S assume "S \ ?L" then obtain b where "S = H.face_cycle_set b" "b \ arcs H" by (auto simp: H.face_cycle_sets_def) then show "S \ ?R" using arcs_not_in H_fcs_in_G_fcs by (cases "b \ G.face_cycle_set a \ G.face_cycle_set a'") auto next fix S assume "S \ ?R" show "S \ ?L" proof (cases "S \ G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'}") case True then have "S \ {a,a'} = {}" using G.face_cycle_set_parts by (auto simp: G.face_cycle_sets_def) with True show ?thesis by (intro in_face_cycle_set_other) auto next case False then have "S \ H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a,a'})" using \S \ ?R\ by blast moreover have "(G.face_cycle_set a \ G.face_cycle_set a') - {a,a'} \ arcs H" using a_in by (auto simp: arcs_H dest: G.in_face_cycle_setD) ultimately show ?thesis by (auto simp: H.face_cycle_sets_def) qed qed lemma card_fcs_aa'_le: "card {G.face_cycle_set a, G.face_cycle_set a'} \ card G.face_cycle_sets" using a_in by (intro card_mono) (auto simp: G.face_cycle_sets_def) lemma card_face_cycle_sets_da0: "card H.face_cycle_sets = card G.face_cycle_sets - card {G.face_cycle_set a, G.face_cycle_set a'} + card (H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a,a'}))" proof - have face_cycle_sets_inter: "(G.face_cycle_sets - {G.face_cycle_set a, G.face_cycle_set a'}) \ H.face_cycle_set ` ((G.face_cycle_set a \ G.face_cycle_set a') - {a, a'}) = {}" (is "?L \ ?R = _") proof - define L R P where "L = ?L" and "R = ?R" and "P x \ x \ (G.face_cycle_set a \ G.face_cycle_set a') = {}" for x then have "\x. x \ L \ P x" "\x. x \ R \ \P x" using G.face_cycle_set_parts by (auto simp: G.face_cycle_sets_def) then have "L \ R = {}" by blast then show ?thesis unfolding L_def R_def . qed then show ?thesis using arcs_G by (simp add: card_Diff_subset[symmetric] card_Un_disjoint[symmetric] G.in_face_cycle_sets face_cycle_sets_da0) qed end locale bidel_arc_same_face = bidel_arc + assumes same_face: "G.face_cycle_set a' = G.face_cycle_set a" begin lemma a_in_o: "a \ orbit G.face_cycle_succ a'" unfolding G.face_cycle_set_def[symmetric] by (simp add: same_face) lemma segment_a'_a_in: "segment G.face_cycle_succ a' a \ arcs H" (is "?seg \ _") proof - have "?seg \ G.face_cycle_set a'" by (auto simp: G.face_cycle_set_def segmentD_orbit) moreover have "G.face_cycle_set a' \ arcs G" by (auto simp: G.face_cycle_set_altdef a_in) ultimately show ?thesis using a_in_o by (auto simp: arcs_H a_in not_in_segment1 not_in_segment2) qed lemma segment_a'_a_neD: assumes "segment G.face_cycle_succ a' a \ {}" shows "segment G.face_cycle_succ a' a \ H.face_cycle_sets" (is "?seg \ _") proof - let ?b = "G.face_cycle_succ a'" have fcs_a_neq_a': "G.face_cycle_succ a' \ a" by (metis assms segment1_empty) have in_aG: "\x. x \ segment G.face_cycle_succ a' a \ x \ arcs G - {a,a'}" using not_in_segment1 not_in_segment2 segment_a'_a_in by (auto simp: arcs_H) { fix x assume A: "x \ segment G.face_cycle_succ a' a" and B: "G.face_cycle_succ x \ a" from A have "G.face_cycle_succ x \ a'" proof induct case base then show ?case by (metis a_neq_a' G.face_cycle_set_self not_in_segment1 G.face_cycle_set_def same_face segment.intros) next case step then show ?case by (metis a_in_o a_neq_a' not_in_segment1 segment.step) qed with A B have "{x, G.face_cycle_succ x} \ {a, a'} = {}" using not_in_segment1[OF a_in_o] not_in_segment2[of a G.face_cycle_succ a'] by safe with in_aG have "H.face_cycle_succ x = G.face_cycle_succ x" by (intro H_fcs_eq_G_fcs) (auto intro: A) } note fcs_x_eq = this { fix x assume A: "x \ segment G.face_cycle_succ a' a" and B: "G.face_cycle_succ x = a" have "G.face_cycle_succ a \ a" using B in_aG[OF A] G.bij_face_cycle_succ by (auto simp: bij_eq_iff) then have "edge_succ M a \ edge_rev M a" by (metis a_in_o G.arev_arev comp_apply G.face_cycle_succ_def not_in_segment1 segment.base) then have "H.face_cycle_succ x = G.face_cycle_succ a'" using in_aG[OF A] B G.bij_edge_succ unfolding H.face_cycle_succ_def G.face_cycle_succ_def by (auto simp: HM_def perm_restrict_simps perm_rem_conv G.arev_eq_iff) } note fcs_last_x_eq = this have "segment G.face_cycle_succ a' a = H.face_cycle_set ?b" proof (intro set_eqI iffI) fix x assume "x \ segment G.face_cycle_succ a' a" then show "x \ H.face_cycle_set ?b" proof induct case base then show ?case by auto next case (step x) then show ?case by (subst fcs_x_eq[symmetric]) (auto simp: H.face_cycle_succ_inI) qed next fix x assume A: "x \ H.face_cycle_set ?b" then show "x \ segment G.face_cycle_succ a' a" proof induct case base then show ?case by (intro segment.base fcs_a_neq_a') next case (step x) then show ?case using fcs_a_neq_a' by (cases "G.face_cycle_succ x = a") (auto simp: fcs_last_x_eq fcs_x_eq intro: segment.intros) qed qed then show ?thesis using segment_a'_a_in by (auto simp: H.face_cycle_sets_def) qed lemma segment_a_a'_neD: assumes "segment G.face_cycle_succ a a' \ {}" shows "segment G.face_cycle_succ a a' \ H.face_cycle_sets" proof - interpret rev: bidel_arc_same_face G M a' using a_in same_face by unfold_locales simp_all from assms show ?thesis using rev.segment_a'_a_neD by (simp add: rev_H rev_HM) qed lemma H_fcs_full: assumes "SS \ H.face_cycle_sets" shows "H.face_cycle_set ` (\SS) = SS" proof - { fix x assume "x \ \SS" then obtain S where "S \ SS" "x \ S" "S \ H.face_cycle_sets" using assms by auto then have "H.face_cycle_set x = S" using H.face_cycle_set_parts by (auto simp: H.face_cycle_sets_def) then have "H.face_cycle_set x \ SS" using \S \ SS\ by auto } moreover { fix S assume "S \ SS" then obtain x where "x \ arcs H" "S = H.face_cycle_set x" "x \ S" using assms by (auto simp: H.face_cycle_sets_def) then have "S \ H.face_cycle_set ` \SS" using \S \ SS\ by auto } ultimately show ?thesis by auto qed lemma card_fcs_gt_0: "0 < card G.face_cycle_sets" using a_in by (auto simp: card_gt_0_iff dest: G.in_face_cycle_sets) lemma card_face_cycle_sets_da': "card H.face_cycle_sets = card G.face_cycle_sets - 1 + card ({segment G.face_cycle_succ a a', segment G.face_cycle_succ a' a, {}} - {{}})" proof - have "G.face_cycle_set a = {a,a'} \ segment G.face_cycle_succ a a' \ segment G.face_cycle_succ a' a" using a_neq_a' same_face by (intro cyclic_split_segment) (auto simp: G.face_cycle_succ_cyclic) then have *: "G.face_cycle_set a \ G.face_cycle_set a' - {a, a'} = segment G.face_cycle_succ a a' \ segment G.face_cycle_succ a' a" by (auto simp: same_face G.face_cycle_set_def[symmetric] not_in_segment1 not_in_segment2) have **: "H.face_cycle_set ` (G.face_cycle_set a \ G.face_cycle_set a' - {a, a'}) = (if segment G.face_cycle_succ a a' \ {} then {segment G.face_cycle_succ a a'} else {}) \ (if segment G.face_cycle_succ a' a \ {} then {segment G.face_cycle_succ a' a} else {})" unfolding * using H_fcs_full[of "{segment G.face_cycle_succ a a', segment G.face_cycle_succ a' a}"] using H_fcs_full[of "{segment G.face_cycle_succ a a'}"] using H_fcs_full[of "{segment G.face_cycle_succ a' a}"] by (auto simp add: segment_a_a'_neD segment_a'_a_neD) show ?thesis unfolding card_face_cycle_sets_da0 ** by (simp add: same_face card_insert_if) qed end locale bidel_arc_diff_face = bidel_arc + assumes diff_face: "G.face_cycle_set a' \ G.face_cycle_set a" begin definition S :: "'b set" where "S \ segment G.face_cycle_succ a a \ segment G.face_cycle_succ a' a'" lemma diff_face_not_in: "a \ G.face_cycle_set a'" "a' \ G.face_cycle_set a" using diff_face G.face_cycle_eq by auto lemma H_fcs_eq_for_a: assumes "b \ arcs H \ G.face_cycle_set a" shows "H.face_cycle_set b = S" (is "?L = ?R") proof (intro set_eqI iffI) fix c assume "c \ ?L" then have "c \ arcs H" using assms by (auto dest: H.in_face_cycle_setD) moreover have "c \ G.face_cycle_set a \ G.face_cycle_set a'" proof (rule ccontr) assume A: "\?thesis" then have "G.face_cycle_set c \ (G.face_cycle_set a \ G.face_cycle_set a') = {}" using G.face_cycle_set_parts by (auto simp: arcs_H) also then have "G.face_cycle_set c = H.face_cycle_set c" using \c \ arcs H\ by (subst face_cycle_set_other_da) (auto simp: arcs_H) also have "\ = H.face_cycle_set b" using \c \ ?L\ using H.face_cycle_set_parts by auto finally show False using assms by auto qed ultimately show "c \ ?R" unfolding S_def arcs_H G.segment_face_cycle_x_x_eq by auto next fix x assume "x \ ?R" from assms have "a \ b" by (auto simp: arcs_H) from assms have b_in: "b \ segment G.face_cycle_succ a a" using G.segment_face_cycle_x_x_eq by (auto simp: arcs_H) have fcs_a_neq_a: "G.face_cycle_succ a \ a" using assms \a \ b\ by (auto simp add: G.segment_face_cycle_x_x_eq G.fcs_x_eq_x) have split_seg: "segment G.face_cycle_succ a a = segment G.face_cycle_succ a b \ {b} \ segment G.face_cycle_succ b a" using b_in by (intro segment_split) have a_in_orb_a: "a \ orbit G.face_cycle_succ a" by (simp add: G.face_cycle_set_def[symmetric]) define c where "c = inv G.face_cycle_succ a" have c_succ: "G.face_cycle_succ c = a" unfolding c_def by (meson bij_inv_eq_iff permutation_bijective G.permutation_face_cycle_succ) have c_in_aa: "c \ segment G.face_cycle_succ a a" unfolding G.segment_face_cycle_x_x_eq c_def using fcs_a_neq_a c_succ c_def by force have c_in: "c \ {b} \ segment G.face_cycle_succ b a" using split_seg b_in c_succ c_in_aa by (auto dest: not_in_segment1[OF segmentD_orbit] intro: segment.intros) from c_in_aa have "c \ arcs H" unfolding G.segment_face_cycle_x_x_eq using arcs_in c_succ diff_face by (auto simp: arcs_H G.face_cycle_eq[of a']) have b_in_L: "b \ ?L" by auto moreover { fix x assume "x \ segment G.face_cycle_succ b a" then have "x \ ?L" proof induct case base then show ?case using assms diff_face_not_in(2) by (subst H_fcs_eq_G_fcs[symmetric]) (auto simp: arcs_H intro: H.face_cycle_succ_inI G.face_cycle_succ_inI) next case (step x) have "G.face_cycle_succ x \ G.face_cycle_set a \ b \ G.face_cycle_set a \ False" using step(1) by (metis G.face_cycle_eq G.face_cycle_succ_inI pre_digraph_map.face_cycle_set_def segmentD_orbit) moreover have "x \ arcs G" using step assms H.in_face_cycle_setD arcs_H by auto moreover then have "(G.face_cycle_succ x \ G.face_cycle_set a \ b \ G.face_cycle_set a \ False) \ {x, G.face_cycle_succ x} \ {a, a'} = {}" using step(2,3) assms diff_face_not_in(2) H.in_face_cycle_setD arcs_H by safe auto ultimately show ?case using step by (subst H_fcs_eq_G_fcs[symmetric]) (auto intro: H.face_cycle_succ_inI) qed } note sba_in_L = this moreover { fix x assume A: "x \ segment G.face_cycle_succ a' a'" then have "x \ ?L" proof - from c_in have "c \ ?L" using b_in_L sba_in_L by blast have "G.face_cycle_succ a' \ a'" using A by (auto simp add: G.segment_face_cycle_x_x_eq G.fcs_x_eq_x) then have *: "G.face_cycle_succ a' = H.face_cycle_succ c" using a_neq_a' c_succ \c \ arcs H\ unfolding G.face_cycle_succ_def H.face_cycle_succ_def arcs_H by (auto simp: HM_def perm_restrict_simps perm_rem_conv G.bij_edge_succ G.arev_eq_iff) from A have "x \ H.face_cycle_set c" proof induct case base then show ?case by (simp add: * H.face_cycle_succ_inI) next case (step x) have "x \ arcs G" using\c \ arcs H\ step.hyps(2) by (auto simp: arcs_H dest: H.in_face_cycle_setD) moreover have "G.face_cycle_succ x \ a' \ {x, G.face_cycle_succ x} \ {a, a'} = {}" using step(1) diff_face_not_in(1) G.face_cycle_succ_inI G.segment_face_cycle_x_x_eq by (auto simp: not_in_segment2) ultimately show ?case using step by (subst H_fcs_eq_G_fcs[symmetric]) (auto intro: H.face_cycle_succ_inI) qed also have "H.face_cycle_set c = ?L" using \c \ ?L\ H.face_cycle_set_parts by auto finally show ?thesis . qed } note sa'a'_in_L = this moreover { assume A: "x \ segment G.face_cycle_succ a b" obtain d where "d \ ?L" and d_succ: "H.face_cycle_succ d = G.face_cycle_succ a" proof (cases "G.face_cycle_succ a' = a'") case True from c_in have "c \ ?L" using b_in_L sba_in_L by blast moreover have "H.face_cycle_succ c = G.face_cycle_succ a" using fcs_a_neq_a c_succ a_neq_a' True \c \ arcs H\ unfolding G.face_cycle_succ_def H.face_cycle_succ_def arcs_H by (auto simp: HM_def perm_restrict_simps arcs_H perm_rem_conv G.bij_edge_succ G.arev_eq_iff) ultimately show ?thesis .. next case False define d where "d = inv G.face_cycle_succ a'" have d_succ: "G.face_cycle_succ d = a'" unfolding d_def by (meson bij_inv_eq_iff permutation_bijective G.permutation_face_cycle_succ) have *: "d \ ?L" using sa'a'_in_L False by (metis DiffI d_succ empty_iff G.face_cycle_set_self G.face_cycle_set_succ insert_iff G.permutation_face_cycle_succ pre_digraph_map.face_cycle_set_def segment_x_x_eq) then have "d \ arcs H" using assms by (auto dest: H.in_face_cycle_setD) have "H.face_cycle_succ d = G.face_cycle_succ a" using fcs_a_neq_a a_neq_a' \d \ arcs H\ d_succ unfolding G.face_cycle_succ_def H.face_cycle_succ_def arcs_H by (auto simp: HM_def perm_restrict_simps arcs_H perm_rem_conv G.bij_edge_succ G.arev_eq_iff) with * show ?thesis .. qed then have "d \ arcs H" using assms by - (drule H.in_face_cycle_setD, auto) from A have "x \ H.face_cycle_set d" proof induct case base then show ?case by (simp add: d_succ[symmetric] H.face_cycle_succ_inI) next case (step x) moreover have "x \ arcs G" using \d \ arcs H\ arcs_H digraph_map.in_face_cycle_setD step.hyps(2) by fastforce moreover have " {x, G.face_cycle_succ x} \ {a, a'} = {}" proof - have "a \ x" using step(2) H.in_face_cycle_setD \d \ arcs H\ arcs_not_in by blast moreover have "a \ G.face_cycle_succ x" by (metis b_in not_in_segment1 segment.step segmentD_orbit step(1)) moreover have "a' \ x" "a' \ G.face_cycle_succ x" using step(1) diff_face_not_in(2) by (auto simp: G.face_cycle_set_def dest!: segmentD_orbit intro: orbit.step) ultimately show ?thesis by auto qed ultimately show ?case using step by (subst H_fcs_eq_G_fcs[symmetric]) (auto intro: H.face_cycle_succ_inI) qed also have "H.face_cycle_set d = ?L" using \d \ ?L\ H.face_cycle_set_parts by auto finally have "x \ ?L" . } ultimately show "x \ ?L" using \x \ ?R\ unfolding S_def split_seg by blast qed lemma HJ_fcs_eq_for_a': assumes "b \ arcs H \ G.face_cycle_set a'" shows "H.face_cycle_set b = S" proof - interpret rev: bidel_arc_diff_face G M a' using arcs_in diff_face by unfold_locales simp_all show ?thesis using rev.H_fcs_eq_for_a assms by (auto simp: rev_H rev_HM S_def rev.S_def) qed lemma card_face_cycle_sets_da': "card H.face_cycle_sets = card G.face_cycle_sets - card {G.face_cycle_set a, G.face_cycle_set a'} + (if S = {} then 0 else 1)" proof - have "S = G.face_cycle_set a \ G.face_cycle_set a' - {a, a'}" unfolding S_def using diff_face_not_in by (auto simp: segment_x_x_eq G.permutation_face_cycle_succ G.face_cycle_set_def) moreover { fix x assume "x \ S" then have "x \ arcs H \ (G.face_cycle_set a \ G.face_cycle_set a' - {a, a'})" unfolding \S = _\ arcs_H using a_in by (auto intro: G.in_face_cycle_setD) then have "H.face_cycle_set x = S" using H_fcs_eq_for_a HJ_fcs_eq_for_a' by blast } then have "H.face_cycle_set ` S = (if S = {} then {} else {S})" by auto ultimately show ?thesis by (simp add: card_face_cycle_sets_da0) qed end locale bidel_arc_biconnected = bidel_arc + assumes reach_a: "tail G a \\<^sup>*\<^bsub>H\<^esub> head G a" begin lemma reach_a': "tail G a' \\<^sup>*\<^bsub>H\<^esub> head G a'" using reach_a a_in by (simp add: symmetric_reachable H.sym_arcs) lemma tail_a': "tail G a' = head G a" and head_a': "head G a' = tail G a" using a_in by simp_all lemma reachable_daI: assumes "v \\<^sup>*\<^bsub>G\<^esub> w" shows "v \\<^sup>*\<^bsub>H\<^esub> w" proof - have *: "\v w. v \\<^bsub>G\<^esub> w \ v \\<^sup>*\<^bsub>H\<^esub> w" using reach_a reach_a' by (auto simp: arcs_ends_def ends_H arcs_G arc_to_ends_def tail_a') show ?thesis using assms by induct (auto simp: verts_H intro: * H.reachable_trans) qed lemma reachable_da: "v \\<^sup>*\<^bsub>H\<^esub> w \ v \\<^sup>*\<^bsub>G\<^esub> w" by (metis reachable_daD reachable_daI) lemma sccs_verts_da: "H.sccs_verts = G.sccs_verts" by (auto simp: G.in_sccs_verts_conv_reachable H.in_sccs_verts_conv_reachable reachable_da) lemma card_sccs_da: "card H.sccs = card G.sccs" by (simp add: G.card_sccs_verts[symmetric] H.card_sccs_verts[symmetric] sccs_verts_da) end locale bidel_arc_not_biconnected = bidel_arc + assumes not_reach_a: "\tail G a \\<^sup>*\<^bsub>H\<^esub> head G a" begin lemma H_awalkI: "G.awalk u p v \ {a,a'} \ set p = {} \ H.awalk u p v" by (auto simp: pre_digraph.apath_def pre_digraph.awalk_def verts_H arcs_H cas_da) lemma tail_neq_head: "tail G a \ head G a" using not_reach_a a_in by (metis H.reachable_refl G.head_in_verts verts_H) lemma scc_of_tail_neq_head: "H.scc_of (tail G a) \ H.scc_of (head G a)" proof - have "tail G a \ H.scc_of (tail G a)" "head G a \ H.scc_of (head G a)" using ends_in by (auto simp: H.in_scc_of_self verts_H) with not_reach_a show ?thesis by (auto simp: H.scc_of_def) qed lemma scc_of_G_tail: assumes "u \ G.scc_of (tail G a)" shows "H.scc_of u = H.scc_of (tail G a) \ H.scc_of u = H.scc_of (head G a)" proof - from assms have "u \\<^sup>*\<^bsub>G\<^esub> tail G a" by (auto simp: G.scc_of_def) then obtain p where p: "G.apath u p (tail G a)" by (auto simp: G.reachable_apath) show ?thesis proof (cases "head G a \ set (G.awalk_verts u p)") case True with p obtain p' q where "p = p' @ q" "G.awalk (head G a) q (tail G a)" and p': "G.awalk u p' (head G a)" unfolding G.apath_def by (metis G.awalk_decomp) moreover then have "tail G a \ set (tl (G.awalk_verts (head G a) q))" using tail_neq_head apply (cases q) apply (simp add: G.awalk_Nil_iff ) apply (simp add: G.awalk_Cons_iff) by (metis G.awalkE G.awalk_verts_non_Nil last_in_set) ultimately have "tail G a \ set (G.awalk_verts u p')" using G.apath_decomp_disjoint[OF p, of p' q "tail G a"] by auto with p' have "{a,a'} \ set p' = {}" by (auto simp: G.set_awalk_verts G.apath_def) (metis a_in imageI G.head_arev) with p' show ?thesis unfolding G.apath_def by (metis H.scc_ofI_awalk H.scc_of_eq H_awalkI) next case False with p have "{a,a'} \ set p = {}" by (auto simp: G.set_awalk_verts G.apath_def) (metis a_in imageI G.tail_arev) with p show ?thesis unfolding G.apath_def by (metis H.scc_ofI_awalk H.scc_of_eq H_awalkI) qed qed lemma scc_of_other: assumes "u \ G.scc_of (tail G a)" shows "H.scc_of u = G.scc_of u" using assms proof (intro set_eqI iffI) fix v assume "v \ H.scc_of u" then show "v \ G.scc_of u" by (auto simp: H.scc_of_def G.scc_of_def intro: reachable_daD) next fix v assume "v \ G.scc_of u" then obtain p where p: "G.awalk u p v" by (auto simp: G.scc_of_def G.reachable_awalk) moreover have "{a,a'} \ set p = {}" proof - have "\u \\<^sup>*\<^bsub>G\<^esub> tail G a" using assms by (metis G.scc_ofI_reachable) then have "\p. \G.awalk u p (tail G a)" by (metis G.reachable_awalk) then have "tail G a \ set (G.awalk_verts u p)" using p by (auto dest: G.awalk_decomp) with p show ?thesis by (auto simp: G.set_awalk_verts G.apath_def) (metis a_in imageI G.head_arev) qed ultimately have "H.awalk u p v" by (rule H_awalkI) then show "v \ H.scc_of u" by (metis H.scc_ofI_reachable' H.reachable_awalk) qed lemma scc_of_tail_inter: "tail G a \ G.scc_of (tail G a) \ H.scc_of (tail G a)" using ends_in by (auto simp: G.in_scc_of_self H.in_scc_of_self verts_H) lemma scc_of_head_inter: "head G a \ G.scc_of (tail G a) \ H.scc_of (head G a)" proof - have "tail G a \\<^bsub>G\<^esub> head G a" "head G a \\<^bsub>G\<^esub> tail G a" by (metis a_in G.in_arcs_imp_in_arcs_ends) (metis a_in G.graph_symmetric G.in_arcs_imp_in_arcs_ends) then have "tail G a \\<^sup>*\<^bsub>G\<^esub> head G a" "head G a \\<^sup>*\<^bsub>G\<^esub> tail G a" by auto then show ?thesis using ends_in by (auto simp: G.scc_of_def H.in_scc_of_self verts_H) qed lemma G_scc_of_tail_not_in: "G.scc_of (tail G a) \ H.sccs_verts" proof assume A: "G.scc_of (tail G a) \ H.sccs_verts" from A scc_of_tail_inter have "G.scc_of (tail G a) = H.scc_of (tail G a)" by (metis H.scc_of_in_sccs_verts H.sccs_verts_disjoint a_in empty_iff G.tail_in_verts verts_H) moreover from A scc_of_head_inter have "G.scc_of (tail G a) = H.scc_of (head G a)" by (metis H.scc_of_in_sccs_verts H.sccs_verts_disjoint a_in empty_iff G.head_in_verts verts_H) ultimately show False using scc_of_tail_neq_head by blast qed lemma H_scc_of_a_not_in: "H.scc_of (tail G a) \ G.sccs_verts" "H.scc_of (head G a) \ G.sccs_verts" proof safe assume "H.scc_of (tail G a) \ G.sccs_verts" with scc_of_tail_inter have "G.scc_of (tail G a) = H.scc_of (tail G a)" by (metis G.scc_of_in_sccs_verts G.sccs_verts_disjoint a_in empty_iff G.tail_in_verts) with G_scc_of_tail_not_in show False using ends_in by (auto simp: H.scc_of_in_sccs_verts verts_H) next assume "H.scc_of (head G a) \ G.sccs_verts" with scc_of_head_inter have "G.scc_of (tail G a) = H.scc_of (head G a)" by (metis G.scc_of_in_sccs_verts G.sccs_verts_disjoint a_in empty_iff G.tail_in_verts) with G_scc_of_tail_not_in show False using ends_in by (auto simp: H.scc_of_in_sccs_verts verts_H) qed lemma scc_verts_da: "H.sccs_verts = (G.sccs_verts - {G.scc_of (tail G a)}) \ {H.scc_of (tail G a), H.scc_of (head G a)}" (is "?L = ?R") proof (intro set_eqI iffI) fix S assume "S \ ?L" then obtain u where "u \ verts G" "S = H.scc_of u" by (auto simp: verts_H H.sccs_verts_conv_scc_of) moreover then have "G.scc_of (tail G a) \ H.scc_of u" using \S \ ?L\ G_scc_of_tail_not_in by auto ultimately show "S \ ?R" unfolding G.sccs_verts_conv_scc_of by (cases "u \ G.scc_of (tail G a)") (auto dest: scc_of_G_tail scc_of_other) next fix S assume "S \ ?R" show "S \ ?L" proof (cases "S \ G.sccs_verts") case True with \S \ ?R\ obtain u where u: "u \ verts G" "S = G.scc_of u" and "S \ G.scc_of (tail G a)" using H_scc_of_a_not_in by (auto simp: G.sccs_verts_conv_scc_of) then have "G.scc_of u \ G.scc_of (tail G a) = {}" using ends_in by (intro G.sccs_verts_disjoint) (auto simp: G.scc_of_in_sccs_verts) then have "u \ G.scc_of (tail G a)" using u by (auto dest: G.in_scc_of_self) with u show ?thesis using scc_of_other by (auto simp: H.sccs_verts_conv_scc_of verts_H G.sccs_verts_conv_scc_of) next case False with \S \ ?R\ ends_in show ?thesis by (auto simp: H.sccs_verts_conv_scc_of verts_H) qed qed lemma card_sccs_da: "card H.sccs = Suc (card G.sccs)" using H_scc_of_a_not_in ends_in unfolding G.card_sccs_verts[symmetric] H.card_sccs_verts[symmetric] scc_verts_da by (simp add: card_insert_if G.finite_sccs_verts scc_of_tail_neq_head card_Suc_Diff1 G.scc_of_in_sccs_verts del: card_Diff_insert) end sublocale bidel_arc_not_biconnected \ bidel_arc_same_face proof note a_in moreover from a_in have "head G a \ tail G ` G.face_cycle_set a" by (simp add: G.heads_face_cycle_set[symmetric]) moreover have "tail G a \ tail G ` G.face_cycle_set a" by simp ultimately obtain p where p: "G.trail (head G a) p (tail G a)" "set p \ G.face_cycle_set a" by (rule G.obtain_trail_in_fcs') define p' where "p' = G.awalk_to_apath p" from p have p': "G.apath (head G a) p' (tail G a)" "set p' \ G.face_cycle_set a" by (auto simp: G.trail_def p'_def dest: G.apath_awalk_to_apath G.awalk_to_apath_subset) then have "set p' \ arcs G" using a_in by (blast dest: G.in_face_cycle_setD) have "\set p' \ arcs H" proof assume "set p' \ arcs H" then have "H.awalk (head G a) p' (tail G a)" using p' by (auto simp: G.apath_def arcs_H intro: H_awalkI) then show False using not_reach_a by (metis H.symmetric_reachable' H.reachable_awalk) qed then have "set p' \ {a,a'} \ {}" using \set p' \ arcs G\ by (auto simp: arcs_H) moreover have "a \ set p'" proof assume "a \ set p'" then have "head G a \ set (tl (G.awalk_verts (head G a) p'))" using \G.apath _ p' _\ by (cases p') (auto simp: G.set_awalk_verts G.apath_def G.awalk_Cons_iff, metis imageI) moreover have "head G a \ set (tl (G.awalk_verts (head G a) p'))" using \G.apath _ p' _\ by (cases p') (auto simp: G.apath_def) ultimately show False by contradiction qed ultimately have "a' \ G.face_cycle_set a" using p'(2) by auto then show "G.face_cycle_set a' = G.face_cycle_set a" using G.face_cycle_set_parts by auto qed locale bidel_arc_tail_conn = bidel_arc + assumes conn_tail: "tail G a \ H.isolated_verts" locale bidel_arc_head_conn = bidel_arc + assumes conn_head: "head G a \ H.isolated_verts" locale bidel_arc_tail_isolated = bidel_arc + assumes isolated_tail: "tail G a \ H.isolated_verts" locale bidel_arc_head_isolated = bidel_arc + assumes isolated_head: "head G a \ H.isolated_verts" begin lemma G_edge_succ_a'_no_loop: assumes no_loop_a: "head G a \ tail G a" shows G_edge_succ_a': "edge_succ M a' = a'" (is ?t2) proof - have *: "out_arcs G (tail G a') = {a'}" using a_in isolated_head no_loop_a by (auto simp: H.isolated_verts_def verts_H out_arcs_def arcs_H tail_H) obtain "edge_succ M a' \ {a'}" using G.edge_succ_cyclic[of "tail G a'"] apply (rule eq_on_cyclic_on_iff1[where x="a'"]) using * a_in a_neq_a' no_loop_a by simp_all then show ?thesis by auto qed lemma G_face_cycle_succ_a_no_loop: assumes no_loop_a: "head G a \ tail G a" shows "G.face_cycle_succ a = a'" using assms by (auto simp: G.face_cycle_succ_def G_edge_succ_a'_no_loop) end locale bidel_arc_same_face_tail_conn = bidel_arc_same_face + bidel_arc_tail_conn begin definition a_neigh :: 'b where "a_neigh \ SOME b. G.face_cycle_succ b = a" lemma face_cycle_succ_a_neigh: "G.face_cycle_succ a_neigh = a" proof - have "\b. G.face_cycle_succ b = a" by (metis G.face_cycle_succ_pred) then show ?thesis unfolding a_neigh_def by (rule someI_ex) qed lemma a_neigh_in: "a_neigh \ arcs G" using a_in by (metis face_cycle_succ_a_neigh G.face_cycle_succ_closed) lemma a_neigh_neq_a: "a_neigh \ a" proof assume "a_neigh = a" then have "G.face_cycle_set a = {a}" using face_cycle_succ_a_neigh by (simp add: G.fcs_x_eq_x) with a_neq_a' same_face G.face_cycle_set_self[of a'] show False by simp qed lemma a_neigh_neq_a': "a_neigh \ a'" proof assume A: "a_neigh = a'" have a_in_oa: "a \ out_arcs G (tail G a)" using a_in by auto have cyc: "cyclic_on (edge_succ M) (out_arcs G (tail G a))" using a_in by (intro G.edge_succ_cyclic) auto from A have "G.face_cycle_succ a' = a" by (metis face_cycle_succ_a_neigh) then have "edge_succ M a = a " by (auto simp: G.face_cycle_succ_def) then have "card (out_arcs G (tail G a)) = 1" using cyc a_in by (auto elim: eq_on_cyclic_on_iff1) then have "out_arcs G (tail G a) = {a}" using a_in_oa by (auto simp del: in_out_arcs_conv dest: card_eq_SucD) then show False using conn_tail a_in by (auto simp: H.isolated_verts_def arcs_H tail_H verts_H out_arcs_def) qed lemma edge_rev_a_neigh_neq: "edge_rev M a_neigh \ a'" by (metis a_neigh_neq_a G.arev_arev) lemma edge_succ_a_neq: "edge_succ M a \ a'" proof assume "edge_succ M a = a'" then have "G.face_cycle_set a' = {a'}" using face_cycle_succ_a_neigh by (auto simp: G.face_cycle_set_altdef id_funpow_id G.face_cycle_succ_def) with a_neq_a' same_face G.face_cycle_set_self[of a] show False by auto qed lemma H_face_cycle_succ_a_neigh: "H.face_cycle_succ a_neigh = G.face_cycle_succ a'" using face_cycle_succ_a_neigh edge_succ_a_neq edge_rev_a_neigh_neq a_neigh_neq_a a_neigh_neq_a' a_neigh_in unfolding H.face_cycle_succ_def G.face_cycle_succ_def by (auto simp: HM_def perm_restrict_simps perm_rem_conv G.bij_edge_succ) lemma H_fcs_a_neigh: "H.face_cycle_set a_neigh = segment G.face_cycle_succ a' a" (is "?L = ?R") proof - { fix n assume A: "0 < n" "n < funpow_dist1 G.face_cycle_succ a' a" then have *: "(G.face_cycle_succ ^^ n) a' \ segment G.face_cycle_succ a' a" using a_in_o by (auto simp: segment_altdef) then have "(G.face_cycle_succ ^^ n) a' \ {a,a'}" "(G.face_cycle_succ ^^ n) a' \ arcs G" using not_in_segment1[OF a_in_o] not_in_segment2[of a G.face_cycle_succ a'] by (auto simp: segment_altdef a_in_o) } note X = this { fix n assume "0 < n" "n < funpow_dist1 G.face_cycle_succ a' a" then have "(H.face_cycle_succ ^^ n) a_neigh = (G.face_cycle_succ ^^ n) a'" proof (induct n) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases "n=0") case True then show ?thesis by (simp add: H_face_cycle_succ_a_neigh) next case False then have "(H.face_cycle_succ ^^ n) a_neigh = (G.face_cycle_succ ^^ n) a'" using Suc by simp then show ?thesis using X[of "Suc n"] X[of n] False Suc by (simp add: H_fcs_eq_G_fcs) qed qed } note Y = this have fcs_a'_neq_a: "G.face_cycle_succ a' \ a" by (metis (no_types) a_neigh_neq_a' G.face_cycle_pred_succ face_cycle_succ_a_neigh) show ?thesis proof (intro set_eqI iffI) fix b assume "b \ ?L" define m where "m = funpow_dist1 G.face_cycle_succ a' a - 1" have b_in0: "b \ orbit H.face_cycle_succ (a_neigh)" using \b \ ?L\ by (simp add: H.face_cycle_set_def[symmetric]) have "0 < m" by (auto simp: m_def) (metis a_neigh_neq_a' G.face_cycle_pred_succ G.face_cycle_set_def G.face_cycle_set_self G.face_cycle_set_succ face_cycle_succ_a_neigh funpow_dist_0_eq neq0_conv same_face) then have pos_dist: "0 < funpow_dist1 H.face_cycle_succ a_neigh b" by (simp add: m_def) have *: "(G.face_cycle_succ ^^ Suc m) a' = a" using a_in_o by (simp add: m_def funpow_simp_l funpow_dist1_prop del: funpow.simps) have "(H.face_cycle_succ ^^ m) a_neigh = a_neigh" proof - have "a = G.face_cycle_succ ((H.face_cycle_succ ^^ m) a_neigh)" using * \0 < m\ by (simp add: Y m_def) then show ?thesis using face_cycle_succ_a_neigh by (metis G.face_cycle_pred_succ) qed then have "funpow_dist1 H.face_cycle_succ a_neigh b \ m" using \0 < m\ b_in0 by (intro funpow_dist1_le_self) simp_all also have "\ < funpow_dist1 G.face_cycle_succ a' a" unfolding m_def by simp finally have dist_less: "funpow_dist1 H.face_cycle_succ a_neigh b < funpow_dist1 G.face_cycle_succ a' a" . have "b = (H.face_cycle_succ ^^ funpow_dist1 H.face_cycle_succ a_neigh b) a_neigh" using b_in0 by (simp add: funpow_dist1_prop del: funpow.simps) also have "\ = (G.face_cycle_succ ^^ funpow_dist1 H.face_cycle_succ a_neigh b) a'" using pos_dist dist_less by (rule Y) also have "\ \ ?R" using pos_dist dist_less by (simp add: segment_altdef a_in_o del: funpow.simps) finally show "b \ ?R" . next fix b assume "b \ ?R" then show "b \ ?L" using Y by (auto simp: segment_altdef a_in_o H.face_cycle_set_altdef Suc_le_eq) metis qed qed end locale bidel_arc_isolated_loop = bidel_arc_biconnected + bidel_arc_tail_isolated begin lemma loop_a[simp]: "head G a = tail G a" using isolated_tail reach_a by (auto simp: H.isolated_verts_def elim: H.converse_reachable_cases dest: out_arcs_emptyD_dominates) end sublocale bidel_arc_isolated_loop \ bidel_arc_head_isolated using isolated_tail loop_a by unfold_locales simp context bidel_arc_isolated_loop begin text \The edges @{term a} and @{term a'} form a loop on an otherwise isolated vertex \ lemma card_isolated_verts_da: "card H.isolated_verts = Suc (card G.isolated_verts)" by (simp add: card_isolated_verts_da0 isolated_tail) lemma G_edge_succ_a[simp]: "edge_succ M a = a'" (is ?t1) and G_edge_succ_a'[simp]: "edge_succ M a' = a" (is ?t2) proof - have *: "out_arcs G (tail G a) = {a,a'}" using a_in isolated_tail by (auto simp: H.isolated_verts_def verts_H out_arcs_def arcs_H tail_H) obtain "edge_succ M a' \ {a,a'}" "edge_succ M a' \ a'" using G.edge_succ_cyclic[of "tail G a'"] apply (rule eq_on_cyclic_on_iff1[where x="a'"]) using * a_in a_neq_a' loop_a by auto moreover obtain "edge_succ M a \ {a,a'}" "edge_succ M a \ a" using G.edge_succ_cyclic[of "tail G a"] apply (rule eq_on_cyclic_on_iff1[where x="a"]) using * a_in a_neq_a' loop_a by auto ultimately show ?t1 ?t2 by auto qed lemma G_face_cycle_succ_a[simp]: "G.face_cycle_succ a = a" and G_face_cycle_succ_a'[simp]: "G.face_cycle_succ a' = a'" by (auto simp: G.face_cycle_succ_def) lemma G_face_cycle_set_a[simp]: "G.face_cycle_set a = {a}" and G_face_cycle_set_a'[simp]: "G.face_cycle_set a' = {a'}" unfolding G.fcs_x_eq_x[symmetric] by simp_all end sublocale bidel_arc_isolated_loop \ bidel_arc_diff_face using a_neq_a' by unfold_locales auto context bidel_arc_isolated_loop begin lemma card_face_cycle_sets_da: "card G.face_cycle_sets = Suc (Suc (card H.face_cycle_sets))" unfolding card_face_cycle_sets_da' using diff_face card_fcs_aa'_le by (auto simp: card_insert_if S_def G.segment_face_cycle_x_x_eq) lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: card_isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_two_isolated = bidel_arc_not_biconnected + bidel_arc_tail_isolated + bidel_arc_head_isolated begin text \@{term "tail G a"} and @{term "head G a"} form an SCC with @{term a} and @{term a'} as the only arcs.\ lemma no_loop_a: "head G a \ tail G a" using not_reach_a a_in by (auto simp: verts_H) lemma card_isolated_verts_da: "card H.isolated_verts = Suc (Suc (card G.isolated_verts))" using no_loop_a isolated_tail isolated_head by (simp add: card_isolated_verts_da0 card_insert_if) lemma G_edge_succ_a'[simp]: "edge_succ M a' = a'" using G_edge_succ_a'_no_loop no_loop_a by simp lemma G_edge_succ_a[simp]: "edge_succ M a = a" proof - have *: "out_arcs G (tail G a) = {a}" using a_in isolated_tail isolated_head no_loop_a by (auto simp: H.isolated_verts_def verts_H out_arcs_def arcs_H tail_H) obtain "edge_succ M a \ {a}" using G.edge_succ_cyclic[of "tail G a"] apply (rule eq_on_cyclic_on_iff1[where x="a"]) using * a_in a_neq_a' no_loop_a by simp_all then show ?thesis by auto qed lemma G_face_cycle_succ_a[simp]: "G.face_cycle_succ a = a'" and G_face_cycle_succ_a'[simp]: "G.face_cycle_succ a' = a" by (auto simp: G.face_cycle_succ_def) lemma G_face_cycle_set_a[simp]: "G.face_cycle_set a = {a,a'}" (is ?t1) and G_face_cycle_set_a'[simp]: "G.face_cycle_set a' = {a,a'}" (is ?t2) proof - { fix n have "(G.face_cycle_succ ^^ n) a \ {a,a'}" "(G.face_cycle_succ ^^ n) a' \ {a,a'}" by (induct n) auto } then show ?t1 ?t2 by (auto simp: G.face_cycle_set_altdef intro: exI[where x=0] exI[where x=1]) qed lemma card_face_cycle_sets_da: "card G.face_cycle_sets = Suc (card H.face_cycle_sets)" unfolding card_face_cycle_sets_da0 using card_fcs_aa'_le by simp lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: card_isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_tail_not_isol = bidel_arc_not_biconnected + bidel_arc_tail_conn sublocale bidel_arc_tail_not_isol \ bidel_arc_same_face_tail_conn by unfold_locales locale bidel_arc_only_tail_not_isol = bidel_arc_tail_not_isol + bidel_arc_head_isolated context bidel_arc_only_tail_not_isol begin lemma card_isolated_verts_da: "card H.isolated_verts = Suc (card G.isolated_verts)" using isolated_head conn_tail by (simp add: card_isolated_verts_da0) lemma segment_a'_a_ne: "segment G.face_cycle_succ a' a \ {}" unfolding H_fcs_a_neigh[symmetric] by auto lemma segment_a_a'_e: "segment G.face_cycle_succ a a' = {}" proof - have "a' = G.face_cycle_succ a" using tail_neq_head by (simp add: G_face_cycle_succ_a_no_loop) then show ?thesis by (auto simp: segment1_empty) qed lemma card_face_cycle_sets_da: "card H.face_cycle_sets = card G.face_cycle_sets" unfolding card_face_cycle_sets_da' using segment_a'_a_ne segment_a_a'_e card_fcs_gt_0 by (simp add: card_insert_if) lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: card_isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_only_head_not_isol = bidel_arc_not_biconnected + bidel_arc_head_conn + bidel_arc_tail_isolated begin interpretation rev: bidel_arc G M a' using a_in by unfold_locales simp interpretation rev: bidel_arc_only_tail_not_isol G M a' using a_in not_reach_a by unfold_locales (auto simp: rev_H isolated_tail conn_head dest: H.symmetric_reachable') lemma euler_genus_da: "H.euler_genus = G.euler_genus" using rev.euler_genus_da by (simp add: rev_H rev_HM) end locale bidel_arc_two_not_isol = bidel_arc_tail_not_isol + bidel_arc_head_conn begin lemma isolated_verts_da: "H.isolated_verts = G.isolated_verts" using conn_head conn_tail by (subst isolated_da_pre) simp lemma segment_a'_a_ne': "segment G.face_cycle_succ a' a \ {}" unfolding H_fcs_a_neigh[symmetric] by auto interpretation rev: bidel_arc_tail_not_isol G M a' using arcs_in not_reach_a rev_H conn_head by unfold_locales (auto dest: H.symmetric_reachable') lemma segment_a_a'_ne': "segment G.face_cycle_succ a a' \ {}" using rev.H_fcs_a_neigh[symmetric] rev_H rev_HM by auto lemma card_face_cycle_sets_da: "card H.face_cycle_sets = Suc (card G.face_cycle_sets)" unfolding card_face_cycle_sets_da' using segment_a'_a_ne' segment_a_a'_ne' card_fcs_gt_0 by (simp add: segments_neq card_insert_if) lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: isolated_verts_da verts_H card_arcs_da card_face_cycle_sets_da card_sccs_da) end locale bidel_arc_biconnected_non_triv = bidel_arc_biconnected + bidel_arc_tail_conn sublocale bidel_arc_biconnected_non_triv \ bidel_arc_head_conn by unfold_locales (metis (mono_tags) G.in_sccs_verts_conv_reachable G.symmetric_reachable' H.isolated_verts_in_sccs conn_tail empty_iff insert_iff reach_a reachable_daD sccs_verts_da) context bidel_arc_biconnected_non_triv begin lemma isolated_verts_da: "H.isolated_verts = G.isolated_verts" using conn_head conn_tail by (subst isolated_da_pre) simp end locale bidel_arc_biconnected_same = bidel_arc_biconnected_non_triv + bidel_arc_same_face sublocale bidel_arc_biconnected_same \ bidel_arc_same_face_tail_conn by unfold_locales context bidel_arc_biconnected_same begin interpretation rev: bidel_arc_same_face_tail_conn G M a' using arcs_in conn_head by unfold_locales (auto simp: same_face rev_H) lemma card_face_cycle_sets_da: "Suc (card H.face_cycle_sets) \ (card G.face_cycle_sets)" unfolding card_face_cycle_sets_da' using card_fcs_gt_0 by linarith lemma euler_genus_da: "H.euler_genus \ G.euler_genus" using card_face_cycle_sets_da unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: isolated_verts_da verts_H card_arcs_da card_sccs_da ) end locale bidel_arc_biconnected_diff = bidel_arc_biconnected_non_triv + bidel_arc_diff_face begin lemma fcs_not_triv: "G.face_cycle_set a \ {a} \ G.face_cycle_set a' \ {a'}" proof (rule ccontr) assume "\?thesis" then have "G.face_cycle_succ a = a" "G.face_cycle_succ a' = a'" by (auto simp: G.fcs_x_eq_x) then have *: "edge_succ M a = a'" "edge_succ M a' = a" by (auto simp: G.face_cycle_succ_def) then have "(edge_succ M ^^ 2) a = a" by (auto simp: eval_nat_numeral) { fix n have "(edge_succ M ^^ 2) a = a" by (auto simp: * eval_nat_numeral) then have "(edge_succ M ^^ n) a = (edge_succ M ^^ (n mod 2)) a" by (auto simp: funpow_mod_eq) moreover have "n mod 2 = 0 \ n mod 2 = 1" by auto ultimately have "(edge_succ M ^^ n) a \ {a, a'}" by (auto simp: *) } then have "orbit (edge_succ M) a = {a, a'}" by (auto simp: orbit_altdef_permutation[OF G.permutation_edge_succ] exI[where x=0] exI[where x=1] *) have "out_arcs G (tail G a) \ {a,a'}" proof - have "cyclic_on (edge_succ M) (out_arcs G (tail G a))" using arcs_in by (intro G.edge_succ_cyclic) auto then have "orbit (edge_succ M) a = out_arcs G (tail G a)" using arcs_in by (intro orbit_cyclic_eq3) auto then show ?thesis using \orbit _ _ = {_, _}\ by auto qed then have "out_arcs H (tail G a) = {}" by (auto simp: arcs_H tail_H) then have "tail G a \ H.isolated_verts" using arcs_in by (simp add: H.isolated_verts_def verts_H) then show False using conn_tail by contradiction qed lemma S_ne: "S \ {}" using fcs_not_triv by (auto simp: S_def G.segment_face_cycle_x_x_eq) lemma card_face_cycle_sets_da: "card G.face_cycle_sets = Suc (card H.face_cycle_sets)" unfolding card_face_cycle_sets_da' using S_ne diff_face card_fcs_aa'_le by simp lemma euler_genus_da: "H.euler_genus = G.euler_genus" unfolding G.euler_genus_def H.euler_genus_def G.euler_char_def H.euler_char_def by (simp add: isolated_verts_da verts_H card_arcs_da card_sccs_da card_face_cycle_sets_da) end context bidel_arc begin lemma euler_genus_da: "H.euler_genus \ G.euler_genus" proof - let ?biconnected = "tail G a \\<^sup>*\<^bsub>H\<^esub> head G a" let ?isol_tail = "tail G a \ H.isolated_verts" let ?isol_head = "head G a \ H.isolated_verts" let ?same_face = "G.face_cycle_set a' = G.face_cycle_set a" { assume ?biconnected ?isol_tail then interpret EG: bidel_arc_isolated_loop by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume ?biconnected "\?isol_tail" ?same_face then interpret EG: bidel_arc_biconnected_same by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume ?biconnected "\?isol_tail" "\?same_face" then interpret EG: bidel_arc_biconnected_diff by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" ?isol_tail ?isol_head then interpret EG: bidel_arc_two_isolated by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" "\?isol_tail" ?isol_head then interpret EG: bidel_arc_only_tail_not_isol by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" ?isol_tail "\?isol_head" then interpret EG: bidel_arc_only_head_not_isol by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } moreover { assume "\?biconnected" "\?isol_tail" "\?isol_head" then interpret EG: bidel_arc_two_not_isol by unfold_locales have ?thesis by (simp add: EG.euler_genus_da) } ultimately show ?thesis by satx qed end subsection \Modifying @{term edge_rev}\ definition (in pre_digraph_map) rev_swap :: "'b \ 'b \ 'b pre_map" where "rev_swap a b = \ edge_rev = perm_swap a b (edge_rev M), edge_succ = perm_swap a b (edge_succ M) \" context digraph_map begin lemma digraph_map_rev_swap: assumes "arc_to_ends G a = arc_to_ends G b" "{a,b} \ arcs G" shows "digraph_map G (rev_swap a b)" proof let ?M' = "rev_swap a b" have tail_swap: "\x. tail G ((a \\<^sub>F b) x) = tail G x" using assms by (case_tac "x \ {a,b}") (auto simp: arc_to_ends_def) have swap_in_arcs: "\x. (a \\<^sub>F b) x \ arcs G \ x \ arcs G" using assms by (case_tac "x \ {a,b}") auto have es_perm: "edge_succ ?M' permutes arcs G" using assms edge_succ_permutes unfolding permutes_conv_has_dom by (auto simp: rev_swap_def has_dom_perm_swap) { fix x show "(x \ arcs G) = (edge_rev (rev_swap a b) x \ x)" using assms(2) - by (cases "x \ {a,b}") (auto simp: rev_swap_def perm_swap_def arev_dom swap_def split: if_splits) + by (cases "x \ {a,b}") (auto simp: rev_swap_def perm_swap_def arev_dom swap_id_eq split: if_splits) next fix x assume "x \ arcs G" then show "edge_rev ?M' (edge_rev ?M' x) = x" by (auto simp: rev_swap_def perm_swap_comp[symmetric]) next fix x assume "x \ arcs G" then show "tail G (edge_rev ?M' x) = head G x" using assms by (case_tac "x \ {a,b}") (auto simp: rev_swap_def perm_swap_def tail_swap arc_to_ends_def) next show "edge_succ ?M' permutes arcs G" by fact next fix v assume A: "v \ verts G" "out_arcs G v \ {}" then obtain c where "c \ out_arcs G v" by blast have "inj (perm_swap a b (edge_succ M))" by (simp add: bij_is_inj bij_edge_succ) have "out_arcs G v = (a \\<^sub>F b) ` out_arcs G v" by (auto simp: tail_swap swap_swap_id swap_in_arcs intro: image_eqI[where x="(a \\<^sub>F b) y" for y]) also have "(a \\<^sub>F b) ` out_arcs G v = (a \\<^sub>F b) ` orbit (edge_succ M) ((a \\<^sub>F b) c)" using edge_succ_cyclic using A \c \ _\ by (intro arg_cong[where f="(`) (a \\<^sub>F b)"]) (intro orbit_cyclic_eq3[symmetric], auto simp: swap_in_arcs tail_swap) also have "\ = orbit (edge_succ ?M') c" by (simp add: orbit_perm_swap rev_swap_def) finally have oa_orb: "out_arcs G v = orbit (edge_succ ?M') c" . show "cyclic_on (edge_succ ?M') (out_arcs G v)" unfolding oa_orb using es_perm finite_arcs by (rule cyclic_on_orbit) } qed lemma euler_genus_rev_swap: assumes "arc_to_ends G a = arc_to_ends G b" "{a,b} \ arcs G" shows "pre_digraph_map.euler_genus G (rev_swap a b) = euler_genus" proof - let ?M' = "rev_swap a b" interpret G': digraph_map G ?M' using assms by (rule digraph_map_rev_swap) have swap_in_arcs: "\x. (a \\<^sub>F b) x \ arcs G \ x \ arcs G" using assms by (case_tac "x \ {a,b}") auto have G'_fcs: "G'.face_cycle_succ = perm_swap a b face_cycle_succ" unfolding G'.face_cycle_succ_def face_cycle_succ_def by (auto simp: fun_eq_iff rev_swap_def perm_swap_comp) have "\x. G'.face_cycle_set x = (a \\<^sub>F b) ` face_cycle_set ((a \\<^sub>F b) x)" by (auto simp: face_cycle_set_def G'.face_cycle_set_def orbit_perm_swap G'_fcs imageI) then have "G'.face_cycle_sets = (\S. (a \\<^sub>F b) ` S) ` face_cycle_sets" by (auto simp: pre_digraph_map.face_cycle_sets_def swap_in_arcs) (metis swap_swap_id image_eqI swap_in_arcs) then have "card G'.face_cycle_sets = card ((\S. (a \\<^sub>F b) ` S) ` face_cycle_sets)" by simp also have "\ = card face_cycle_sets" by (rule card_image) (rule inj_on_f_imageI[where S="UNIV"], auto) finally show "pre_digraph_map.euler_genus G ?M' = euler_genus" unfolding pre_digraph_map.euler_genus_def pre_digraph_map.euler_char_def by simp qed end subsection \Conclusion\ lemma bidirected_subgraph_obtain: assumes sg: "subgraph H G" "arcs H \ arcs G" assumes fin: "finite (arcs G)" assumes bidir: "\rev. bidirected_digraph G rev" "\rev. bidirected_digraph H rev" obtains a a' where "{a,a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a'= tail G a" proof - obtain a where a: "a \ arcs G - arcs H" using sg by blast obtain rev_G rev_H where rev: "bidirected_digraph G rev_G" "bidirected_digraph H rev_H" using bidir by blast interpret G: bidirected_digraph G rev_G by (rule rev) interpret H: bidirected_digraph H rev_H by (rule rev) have sg_props: "arcs H \ arcs G" "tail H = tail G" "head H = head G" using sg by (auto simp: subgraph_def compatible_def) { fix w1 w2 assume A: "tail G a = w1" "head G a = w2" have "in_arcs H w1 \ out_arcs H w2 = rev_H ` (out_arcs H w1 \ in_arcs H w2)" (is "?Sh = _") unfolding H.in_arcs_eq by (simp add: image_Int image_image H.inj_on_arev) then have "card (in_arcs H w1 \ out_arcs H w2) = card (out_arcs H w1 \ in_arcs H w2)" by (metis card_image H.arev_arev inj_on_inverseI) also have "\ < card (out_arcs G w1 \ in_arcs G w2)" (is "card ?Sh1 < card ?Sg1") proof (rule psubset_card_mono) show "finite ?Sg1" using fin by (auto simp: out_arcs_def) show "?Sh1 \ ?Sg1" using A a sg_props by auto qed also have "?Sg1 = rev_G ` (in_arcs G w1 \ out_arcs G w2)" (is "_ = _ ` ?Sg") unfolding G.in_arcs_eq by (simp add: image_Int image_image G.inj_on_arev) also have "card \ = card ?Sg" by (metis card_image G.arev_arev inj_on_inverseI) finally have card_less: "card ?Sh < card ?Sg" . have S_ss: "?Sh \ ?Sg" using sg_props by auto have ?thesis proof (cases "w1 = w2") case True have "card (?Sh - {a}) = card ?Sh" using a by (intro arg_cong[where f=card]) auto also have "\ < card ?Sg - 1" proof - from True have "even (card ?Sg)" "even (card ?Sh)" by (auto simp: G.even_card_loops H.even_card_loops) then show ?thesis using card_less by simp (metis Suc_pred even_Suc le_neq_implies_less lessE less_Suc_eq_le zero_less_Suc) qed also have "\ = card (?Sg - {a})" using fin a A True by (auto simp: out_arcs_def card_Diff_singleton) finally have card_diff_a_less: "card (?Sh - {a}) < card (?Sg - {a})" . moreover from S_ss have "?Sh - {a} \ ?Sg - {a}" using S_ss by blast ultimately have "?Sh - {a} \ ?Sg - {a}" by (intro card_psubset) auto then obtain a' where "a' \ (?Sg - {a})- ?Sh" by blast then have "{a,a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a'= tail G a" using A a sg_props by auto then show ?thesis .. next case False from card_less S_ss have "?Sh \ ?Sg" by auto then obtain a' where "a' \ ?Sg - ?Sh" by blast then have "{a,a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a'= tail G a" using A a sg_props False by auto then show ?thesis .. qed } then show ?thesis by simp qed lemma subgraph_euler_genus_le: assumes G: "subgraph H G" "digraph_map G GM" and H: "\rev. bidirected_digraph H rev" obtains HM where "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ pre_digraph_map.euler_genus G GM" proof - let ?d = "\G. card (arcs G) + card (verts G) - card (arcs H) - card (verts H)" from H obtain rev_H where "bidirected_digraph H rev_H" by blast then interpret H: bidirected_digraph H rev_H . from G have "\HM. digraph_map H HM \ pre_digraph_map.euler_genus H HM \ pre_digraph_map.euler_genus G GM" proof (induct "?d G" arbitrary: G GM rule: less_induct) case less from less interpret G: digraph_map G GM by - have H_ss: "arcs H \ arcs G" "verts H \ verts G" using \subgraph H G\ by auto then have card_le: "card (arcs H) \ card (arcs G)" "card (verts H) \ card (verts G)" by (auto intro: card_mono) have ends: "tail H = tail G" "head H = head G" using \subgraph H G\ by (auto simp: compatible_def) show ?case proof (cases "?d G = 0") case True then have "card (arcs H) = card (arcs G)" "card (verts H) = card (verts G)" using card_le by linarith+ then have "arcs H = arcs G" "verts H = verts G" using H_ss by (auto simp: card_subset_eq) then have "H = G" using \subgraph H G\ by (auto simp: compatible_def) then have "digraph_map H GM \ pre_digraph_map.euler_genus H GM \ G.euler_genus" by auto then show ?thesis .. next case False then have H_ne: "(arcs G - arcs H) \ {} \ (verts G - verts H) \ {}" using H_ss card_le by auto { assume A: "arcs G - arcs H \ {}" then obtain a a' where aa': "{a, a'} \ arcs G - arcs H" "a' \ a" "tail G a' = head G a" "head G a' = tail G a" using H_ss \subgraph H G\ by (auto intro: bidirected_subgraph_obtain) let ?GM' = "G.rev_swap (edge_rev GM a) a'" interpret G': digraph_map G ?GM' using aa' by (intro G.digraph_map_rev_swap) (auto simp: arc_to_ends_def) interpret G': bidel_arc G ?GM' a using aa' by unfold_locales simp have "edge_rev GM a \ a" using aa' by (intro G.arev_neq) auto then have er_a: "edge_rev ?GM' a = a'" using \a' \ a\ by (auto simp: G.rev_swap_def perm_swap_def swap_id_eq dest: G.arev_neq) then have sg: "subgraph H G'.H" using H_ss aa' by (intro subgraphI) (auto simp: G'.verts_H G'.arcs_H G'.tail_H G'.head_H ends compatible_def intro: H.wf_digraph G'.H.wf_digraph) have "card {a,a'} \ card (arcs G)" using aa' by (intro card_mono) auto then obtain HM where HM: "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ G'.H.euler_genus" using aa' False by atomize_elim (rule less, auto simp: G'.verts_H G'.arcs_H card_insert_if sg er_a) have "G'.H.euler_genus \ G'.euler_genus" by (rule G'.euler_genus_da) also have "G'.euler_genus = G.euler_genus" using aa' by (auto simp: G.euler_genus_rev_swap arc_to_ends_def) finally have ?thesis using HM by auto } moreover { assume A: "arcs G - arcs H = {}" then have A': "verts G - verts H \ {}" and arcs_H: "arcs H = arcs G" using H_ss H_ne by auto then obtain v where v: "v \ verts G - verts H" by auto have card_lt: "card (verts H) < card (verts G)" using A' H_ss by (intro psubset_card_mono) auto have "out_arcs G v = out_arcs H v" using A H_ss by (auto simp: ends) then interpret G: del_vert_props G GM v using v by unfold_locales auto have "?d (G.del_vert v) < ?d G" using card_lt by (simp add: arcs_H G.arcs_dv G.card_verts_dv) moreover have "subgraph H (G.del_vert v)" using H_ss v by (auto simp: subgraph_def arcs_H G.arcs_dv G.verts_del_vert H.wf_digraph G.H.wf_digraph compatible_def G.tail_del_vert G.head_del_vert ends) moreover have "bidirected_digraph (G.del_vert v) (edge_rev GM)" using G.arev_dom by (intro G.H.bidirected_digraphI) (auto simp: G.arcs_dv) ultimately have ?thesis unfolding G.euler_genus_eq[symmetric] by (intro less) auto } ultimately show ?thesis by blast qed qed then obtain HM where "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ pre_digraph_map.euler_genus G GM" by atomize_elim then show ?thesis .. qed lemma (in digraph_map) nonneg_euler_genus: "0 \ euler_genus" proof - define H where "H = \ verts = {}, arcs = {}, tail = tail G, head = head G \" then have H_simps: "verts H = {}" "arcs H = {}" "tail H = tail G" "head H = head G" by (simp_all add: H_def) interpret H: bidirected_digraph H id by unfold_locales (auto simp: H_def) have "wf_digraph H" "wf_digraph G" by unfold_locales then have "subgraph H G" by (intro subgraphI) (auto simp: H_def compatible_def) then obtain HM where "digraph_map H HM" "pre_digraph_map.euler_genus H HM \ euler_genus" by (rule subgraph_euler_genus_le) auto then interpret H: digraph_map H HM by - have "H.sccs = {}" proof - { fix x assume *: "x \ H.sccs_verts" then have "x = {}" by (auto dest: H.sccs_verts_subsets simp: H_simps) with * have False by (auto simp: H.in_sccs_verts_conv_reachable) } then show ?thesis by (auto simp: H.sccs_verts_conv) qed then have "H.euler_genus = 0" by (auto simp: H.euler_genus_def H.euler_char_def H.isolated_verts_def H.face_cycle_sets_def H_simps) then show ?thesis using \H.euler_genus \ _\ by simp qed lemma subgraph_comb_planar: assumes "subgraph G H" "comb_planar H" "\rev. bidirected_digraph G rev" shows "comb_planar G" proof - from \comb_planar H\ obtain HM where "digraph_map H HM" and H_genus: "pre_digraph_map.euler_genus H HM = 0" unfolding comb_planar_def by metis obtain GM where G: "digraph_map G GM" "pre_digraph_map.euler_genus G GM \ pre_digraph_map.euler_genus H HM" using assms(1) \digraph_map H HM\ assms(3) by (rule subgraph_euler_genus_le) interpret G: digraph_map G GM by fact show ?thesis using G H_genus G.nonneg_euler_genus unfolding comb_planar_def by auto qed end diff --git a/thys/Smith_Normal_Form/Mod_Type_Connect.thy b/thys/Smith_Normal_Form/Mod_Type_Connect.thy --- a/thys/Smith_Normal_Form/Mod_Type_Connect.thy +++ b/thys/Smith_Normal_Form/Mod_Type_Connect.thy @@ -1,581 +1,583 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \A new bridge to convert theorems from JNF to HOL Analysis and vice-versa, based on the @{text "mod_type"} class\ theory Mod_Type_Connect imports Perron_Frobenius.HMA_Connect Rank_Nullity_Theorem.Mod_Type Gauss_Jordan.Elementary_Operations begin text \Some lemmas on @{text "Mod_Type.to_nat"} and @{text "Mod_Type.from_nat"} are added to have them with the same names as the analogous ones for @{text "Bij_Nat.to_nat"} and @{text "Bij_Nat.to_nat"}.\ lemma inj_to_nat: "inj to_nat" by (simp add: inj_on_def) lemmas from_nat_inj = from_nat_eq_imp_eq lemma range_to_nat: "range (to_nat :: 'a :: mod_type \ nat) = {0 ..< CARD('a)}" by (simp add: bij_betw_imp_surj_on mod_type_class.bij_to_nat) text \This theory is an adaptation of the one presented in @{text "Perron_Frobenius.HMA_Connect"}, but for matrices and vectors where indexes have the @{text "mod_type"} class restriction. It is worth noting that some definitions still use the old abbreviation for HOL Analysis (HMA, from HOL Multivariate Analysis) instead of HA. This is done to be consistent with the existing names in the Perron-Frobenius development\ context includes vec.lifting begin end definition from_hma\<^sub>v :: "'a ^ 'n :: mod_type \ 'a Matrix.vec" where "from_hma\<^sub>v v = Matrix.vec CARD('n) (\ i. v $h from_nat i)" definition from_hma\<^sub>m :: "'a ^ 'nc :: mod_type ^ 'nr :: mod_type \ 'a Matrix.mat" where "from_hma\<^sub>m a = Matrix.mat CARD('nr) CARD('nc) (\ (i,j). a $h from_nat i $h from_nat j)" definition to_hma\<^sub>v :: "'a Matrix.vec \ 'a ^ 'n :: mod_type" where "to_hma\<^sub>v v = (\ i. v $v to_nat i)" definition to_hma\<^sub>m :: "'a Matrix.mat \ 'a ^ 'nc :: mod_type ^ 'nr :: mod_type " where "to_hma\<^sub>m a = (\ i j. a $$ (to_nat i, to_nat j))" lemma to_hma_from_hma\<^sub>v[simp]: "to_hma\<^sub>v (from_hma\<^sub>v v) = v" by (auto simp: to_hma\<^sub>v_def from_hma\<^sub>v_def to_nat_less_card) lemma to_hma_from_hma\<^sub>m[simp]: "to_hma\<^sub>m (from_hma\<^sub>m v) = v" by (auto simp: to_hma\<^sub>m_def from_hma\<^sub>m_def to_nat_less_card) lemma from_hma_to_hma\<^sub>v[simp]: "v \ carrier_vec (CARD('n)) \ from_hma\<^sub>v (to_hma\<^sub>v v :: 'a ^ 'n :: mod_type) = v" by (auto simp: to_hma\<^sub>v_def from_hma\<^sub>v_def to_nat_from_nat_id) lemma from_hma_to_hma\<^sub>m[simp]: "A \ carrier_mat (CARD('nr)) (CARD('nc)) \ from_hma\<^sub>m (to_hma\<^sub>m A :: 'a ^ 'nc :: mod_type ^ 'nr :: mod_type) = A" by (auto simp: to_hma\<^sub>m_def from_hma\<^sub>m_def to_nat_from_nat_id) lemma from_hma\<^sub>v_inj[simp]: "from_hma\<^sub>v x = from_hma\<^sub>v y \ x = y" by (intro iffI, insert to_hma_from_hma\<^sub>v[of x], auto) lemma from_hma\<^sub>m_inj[simp]: "from_hma\<^sub>m x = from_hma\<^sub>m y \ x = y" by(intro iffI, insert to_hma_from_hma\<^sub>m[of x], auto) definition HMA_V :: "'a Matrix.vec \ 'a ^ 'n :: mod_type \ bool" where "HMA_V = (\ v w. v = from_hma\<^sub>v w)" definition HMA_M :: "'a Matrix.mat \ 'a ^ 'nc :: mod_type ^ 'nr :: mod_type \ bool" where "HMA_M = (\ a b. a = from_hma\<^sub>m b)" definition HMA_I :: "nat \ 'n :: mod_type \ bool" where "HMA_I = (\ i a. i = to_nat a)" context includes lifting_syntax begin lemma Domainp_HMA_V [transfer_domain_rule]: "Domainp (HMA_V :: 'a Matrix.vec \ 'a ^ 'n :: mod_type \ bool) = (\ v. v \ carrier_vec (CARD('n )))" by(intro ext iffI, insert from_hma_to_hma\<^sub>v[symmetric], auto simp: from_hma\<^sub>v_def HMA_V_def) lemma Domainp_HMA_M [transfer_domain_rule]: "Domainp (HMA_M :: 'a Matrix.mat \ 'a ^ 'nc :: mod_type ^ 'nr :: mod_type \ bool) = (\ A. A \ carrier_mat CARD('nr) CARD('nc))" by (intro ext iffI, insert from_hma_to_hma\<^sub>m[symmetric], auto simp: from_hma\<^sub>m_def HMA_M_def) lemma Domainp_HMA_I [transfer_domain_rule]: "Domainp (HMA_I :: nat \ 'n :: mod_type \ bool) = (\ i. i < CARD('n))" (is "?l = ?r") proof (intro ext) fix i :: nat show "?l i = ?r i" unfolding HMA_I_def Domainp_iff by (auto intro: exI[of _ "from_nat i"] simp: to_nat_from_nat_id to_nat_less_card) qed lemma bi_unique_HMA_V [transfer_rule]: "bi_unique HMA_V" "left_unique HMA_V" "right_unique HMA_V" unfolding HMA_V_def bi_unique_def left_unique_def right_unique_def by auto lemma bi_unique_HMA_M [transfer_rule]: "bi_unique HMA_M" "left_unique HMA_M" "right_unique HMA_M" unfolding HMA_M_def bi_unique_def left_unique_def right_unique_def by auto lemma bi_unique_HMA_I [transfer_rule]: "bi_unique HMA_I" "left_unique HMA_I" "right_unique HMA_I" unfolding HMA_I_def bi_unique_def left_unique_def right_unique_def by auto lemma right_total_HMA_V [transfer_rule]: "right_total HMA_V" unfolding HMA_V_def right_total_def by simp lemma right_total_HMA_M [transfer_rule]: "right_total HMA_M" unfolding HMA_M_def right_total_def by simp lemma right_total_HMA_I [transfer_rule]: "right_total HMA_I" unfolding HMA_I_def right_total_def by simp lemma HMA_V_index [transfer_rule]: "(HMA_V ===> HMA_I ===> (=)) ($v) ($h)" unfolding rel_fun_def HMA_V_def HMA_I_def from_hma\<^sub>v_def by (auto simp: to_nat_less_card) lemma HMA_M_index [transfer_rule]: "(HMA_M ===> HMA_I ===> HMA_I ===> (=)) (\ A i j. A $$ (i,j)) index_hma" by (intro rel_funI, simp add: index_hma_def to_nat_less_card HMA_M_def HMA_I_def from_hma\<^sub>m_def) lemma HMA_V_0 [transfer_rule]: "HMA_V (0\<^sub>v CARD('n)) (0 :: 'a :: zero ^ 'n:: mod_type)" unfolding HMA_V_def from_hma\<^sub>v_def by auto lemma HMA_M_0 [transfer_rule]: "HMA_M (0\<^sub>m CARD('nr) CARD('nc)) (0 :: 'a :: zero ^ 'nc:: mod_type ^ 'nr :: mod_type)" unfolding HMA_M_def from_hma\<^sub>m_def by auto lemma HMA_M_1[transfer_rule]: "HMA_M (1\<^sub>m (CARD('n))) (mat 1 :: 'a::{zero,one}^'n:: mod_type^'n:: mod_type)" unfolding HMA_M_def by (auto simp add: mat_def from_hma\<^sub>m_def from_nat_inj) lemma from_hma\<^sub>v_add: "from_hma\<^sub>v v + from_hma\<^sub>v w = from_hma\<^sub>v (v + w)" unfolding from_hma\<^sub>v_def by auto lemma HMA_V_add [transfer_rule]: "(HMA_V ===> HMA_V ===> HMA_V) (+) (+) " unfolding rel_fun_def HMA_V_def by (auto simp: from_hma\<^sub>v_add) lemma from_hma\<^sub>v_diff: "from_hma\<^sub>v v - from_hma\<^sub>v w = from_hma\<^sub>v (v - w)" unfolding from_hma\<^sub>v_def by auto lemma HMA_V_diff [transfer_rule]: "(HMA_V ===> HMA_V ===> HMA_V) (-) (-)" unfolding rel_fun_def HMA_V_def by (auto simp: from_hma\<^sub>v_diff) lemma from_hma\<^sub>m_add: "from_hma\<^sub>m a + from_hma\<^sub>m b = from_hma\<^sub>m (a + b)" unfolding from_hma\<^sub>m_def by auto lemma HMA_M_add [transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (+) (+) " unfolding rel_fun_def HMA_M_def by (auto simp: from_hma\<^sub>m_add) lemma from_hma\<^sub>m_diff: "from_hma\<^sub>m a - from_hma\<^sub>m b = from_hma\<^sub>m (a - b)" unfolding from_hma\<^sub>m_def by auto lemma HMA_M_diff [transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (-) (-) " unfolding rel_fun_def HMA_M_def by (auto simp: from_hma\<^sub>m_diff) lemma scalar_product: fixes v :: "'a :: semiring_1 ^ 'n :: mod_type" shows "scalar_prod (from_hma\<^sub>v v) (from_hma\<^sub>v w) = scalar_product v w" unfolding scalar_product_def scalar_prod_def from_hma\<^sub>v_def dim_vec by (simp add: sum.reindex[OF inj_to_nat, unfolded range_to_nat]) lemma [simp]: "from_hma\<^sub>m (y :: 'a ^ 'nc :: mod_type ^ 'nr:: mod_type) \ carrier_mat (CARD('nr)) (CARD('nc))" "dim_row (from_hma\<^sub>m (y :: 'a ^ 'nc:: mod_type ^ 'nr :: mod_type)) = CARD('nr)" "dim_col (from_hma\<^sub>m (y :: 'a ^ 'nc :: mod_type ^ 'nr:: mod_type )) = CARD('nc)" unfolding from_hma\<^sub>m_def by simp_all lemma [simp]: "from_hma\<^sub>v (y :: 'a ^ 'n:: mod_type) \ carrier_vec (CARD('n))" "dim_vec (from_hma\<^sub>v (y :: 'a ^ 'n:: mod_type)) = CARD('n)" unfolding from_hma\<^sub>v_def by simp_all lemma HMA_scalar_prod [transfer_rule]: "(HMA_V ===> HMA_V ===> (=)) scalar_prod scalar_product" by (auto simp: HMA_V_def scalar_product) lemma HMA_row [transfer_rule]: "(HMA_I ===> HMA_M ===> HMA_V) (\ i a. Matrix.row a i) row" unfolding HMA_M_def HMA_I_def HMA_V_def by (auto simp: from_hma\<^sub>m_def from_hma\<^sub>v_def to_nat_less_card row_def) lemma HMA_col [transfer_rule]: "(HMA_I ===> HMA_M ===> HMA_V) (\ i a. col a i) column" unfolding HMA_M_def HMA_I_def HMA_V_def by (auto simp: from_hma\<^sub>m_def from_hma\<^sub>v_def to_nat_less_card column_def) lemma HMA_M_mk_mat[transfer_rule]: "((HMA_I ===> HMA_I ===> (=)) ===> HMA_M) (\ f. Matrix.mat (CARD('nr)) (CARD('nc)) (\ (i,j). f i j)) (mk_mat :: (('nr \ 'nc \ 'a) \ 'a^'nc:: mod_type^'nr:: mod_type))" proof- { fix x y i j assume id: "\ (ya :: 'nr) (yb :: 'nc). (x (to_nat ya) (to_nat yb) :: 'a) = y ya yb" and i: "i < CARD('nr)" and j: "j < CARD('nc)" from to_nat_from_nat_id[OF i] to_nat_from_nat_id[OF j] id[rule_format, of "from_nat i" "from_nat j"] have "x i j = y (from_nat i) (from_nat j)" by auto } thus ?thesis unfolding rel_fun_def mk_mat_def HMA_M_def HMA_I_def from_hma\<^sub>m_def by auto qed lemma HMA_M_mk_vec[transfer_rule]: "((HMA_I ===> (=)) ===> HMA_V) (\ f. Matrix.vec (CARD('n)) (\ i. f i)) (mk_vec :: (('n \ 'a) \ 'a^'n:: mod_type))" proof- { fix x y i assume id: "\ (ya :: 'n). (x (to_nat ya) :: 'a) = y ya" and i: "i < CARD('n)" from to_nat_from_nat_id[OF i] id[rule_format, of "from_nat i"] have "x i = y (from_nat i)" by auto } thus ?thesis unfolding rel_fun_def mk_vec_def HMA_V_def HMA_I_def from_hma\<^sub>v_def by auto qed lemma mat_mult_scalar: "A ** B = mk_mat (\ i j. scalar_product (row i A) (column j B))" unfolding vec_eq_iff matrix_matrix_mult_def scalar_product_def mk_mat_def by (auto simp: row_def column_def) lemma mult_mat_vec_scalar: "A *v v = mk_vec (\ i. scalar_product (row i A) v)" unfolding vec_eq_iff matrix_vector_mult_def scalar_product_def mk_mat_def mk_vec_def by (auto simp: row_def column_def) lemma dim_row_transfer_rule: "HMA_M A (A' :: 'a ^ 'nc:: mod_type ^ 'nr:: mod_type) \ (=) (dim_row A) (CARD('nr))" unfolding HMA_M_def by auto lemma dim_col_transfer_rule: "HMA_M A (A' :: 'a ^ 'nc:: mod_type ^ 'nr:: mod_type) \ (=) (dim_col A) (CARD('nc))" unfolding HMA_M_def by auto lemma HMA_M_mult [transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (*) (**)" proof - { fix A B :: "'a :: semiring_1 mat" and A' :: "'a ^ 'n :: mod_type ^ 'nr:: mod_type" and B' :: "'a ^ 'nc :: mod_type ^ 'n:: mod_type" assume 1[transfer_rule]: "HMA_M A A'" "HMA_M B B'" note [transfer_rule] = dim_row_transfer_rule[OF 1(1)] dim_col_transfer_rule[OF 1(2)] have "HMA_M (A * B) (A' ** B')" unfolding times_mat_def mat_mult_scalar by (transfer_prover_start, transfer_step+, transfer, auto) } thus ?thesis by blast qed lemma HMA_V_smult [transfer_rule]: "((=) ===> HMA_V ===> HMA_V) (\\<^sub>v) (*s)" unfolding smult_vec_def unfolding rel_fun_def HMA_V_def from_hma\<^sub>v_def by auto lemma HMA_M_mult_vec [transfer_rule]: "(HMA_M ===> HMA_V ===> HMA_V) (*\<^sub>v) (*v)" proof - { fix A :: "'a :: semiring_1 mat" and v :: "'a Matrix.vec" and A' :: "'a ^ 'nc :: mod_type ^ 'nr :: mod_type" and v' :: "'a ^ 'nc :: mod_type" assume 1[transfer_rule]: "HMA_M A A'" "HMA_V v v'" note [transfer_rule] = dim_row_transfer_rule have "HMA_V (A *\<^sub>v v) (A' *v v')" unfolding mult_mat_vec_def mult_mat_vec_scalar by (transfer_prover_start, transfer_step+, transfer, auto) } thus ?thesis by blast qed lemma HMA_det [transfer_rule]: "(HMA_M ===> (=)) Determinant.det (det :: 'a :: comm_ring_1 ^ 'n :: mod_type ^ 'n :: mod_type \ 'a)" proof - { fix a :: "'a ^ 'n :: mod_type^ 'n:: mod_type" let ?tn = "to_nat :: 'n :: mod_type \ nat" let ?fn = "from_nat :: nat \ 'n" let ?zn = "{0..< CARD('n)}" let ?U = "UNIV :: 'n set" let ?p1 = "{p. p permutes ?zn}" let ?p2 = "{p. p permutes ?U}" let ?f= "\ p i. if i \ ?U then ?fn (p (?tn i)) else i" let ?g = "\ p i. ?fn (p (?tn i))" have fg: "\ a b c. (if a \ ?U then b else c) = b" by auto have "?p2 = ?f ` ?p1" by (rule permutes_bij', auto simp: to_nat_less_card to_nat_from_nat_id) hence id: "?p2 = ?g ` ?p1" by simp have inj_g: "inj_on ?g ?p1" unfolding inj_on_def proof (intro ballI impI ext, auto) fix p q i assume p: "p permutes ?zn" and q: "q permutes ?zn" and id: "(\ i. ?fn (p (?tn i))) = (\ i. ?fn (q (?tn i)))" { fix i from permutes_in_image[OF p] have pi: "p (?tn i) < CARD('n)" by (simp add: to_nat_less_card) from permutes_in_image[OF q] have qi: "q (?tn i) < CARD('n)" by (simp add: to_nat_less_card) from fun_cong[OF id] have "?fn (p (?tn i)) = from_nat (q (?tn i))" . from arg_cong[OF this, of ?tn] have "p (?tn i) = q (?tn i)" by (simp add: to_nat_from_nat_id pi qi) } note id = this show "p i = q i" proof (cases "i < CARD('n)") case True hence "?tn (?fn i) = i" by (simp add: to_nat_from_nat_id) from id[of "?fn i", unfolded this] show ?thesis . next case False thus ?thesis using p q unfolding permutes_def by simp qed qed have mult_cong: "\ a b c d. a = b \ c = d \ a * c = b * d" by simp have "sum (\ p. signof p * (\i\?zn. a $h ?fn i $h ?fn (p i))) ?p1 = sum (\ p. of_int (sign p) * (\i\UNIV. a $h i $h p i)) ?p2" unfolding id sum.reindex[OF inj_g] proof (rule sum.cong[OF refl], unfold mem_Collect_eq o_def, rule mult_cong) fix p assume p: "p permutes ?zn" let ?q = "\ i. ?fn (p (?tn i))" from id p have q: "?q permutes ?U" by auto from p have pp: "permutation p" unfolding permutation_permutes by auto let ?ft = "\ p i. ?fn (p (?tn i))" have fin: "finite ?zn" by simp have "sign p = sign ?q \ p permutes ?zn" using p fin proof (induction rule: permutes_induct) case id show ?case by (auto simp: sign_id[unfolded id_def] permutes_id[unfolded id_def]) next case (swap a b p) + then have \permutation p\ + using permutes_imp_permutation by blast let ?sab = "Fun.swap a b id" let ?sfab = "Fun.swap (?fn a) (?fn b) id" have p_sab: "permutation ?sab" by (rule permutation_swap_id) have p_sfab: "permutation ?sfab" by (rule permutation_swap_id) from swap(4) have IH1: "p permutes ?zn" and IH2: "sign p = sign (?ft p)" by auto have sab_perm: "?sab permutes ?zn" using swap(1-2) by (rule permutes_swap_id) from permutes_compose[OF IH1 this] have perm1: "?sab o p permutes ?zn" . from IH1 have p_p1: "p \ ?p1" by simp hence "?ft p \ ?ft ` ?p1" by (rule imageI) from this[folded id] have "?ft p permutes ?U" by simp hence p_ftp: "permutation (?ft p)" unfolding permutation_permutes by auto { fix a b assume a: "a \ ?zn" and b: "b \ ?zn" hence "(?fn a = ?fn b) = (a = b)" using swap(1-2) by (auto simp add: from_nat_eq_imp_eq) } note inj = this from inj[OF swap(1-2)] have id2: "sign ?sfab = sign ?sab" unfolding sign_swap_id by simp have id: "?ft (Fun.swap a b id \ p) = Fun.swap (?fn a) (?fn b) id \ ?ft p" proof fix c show "?ft (Fun.swap a b id \ p) c = (Fun.swap (?fn a) (?fn b) id \ ?ft p) c" proof (cases "p (?tn c) = a \ p (?tn c) = b") case True - thus ?thesis by (cases, auto simp add: o_def swap_def) + thus ?thesis by (cases, auto simp add: o_def swap_id_eq) next case False hence neq: "p (?tn c) \ a" "p (?tn c) \ b" by auto have pc: "p (?tn c) \ ?zn" unfolding permutes_in_image[OF IH1] by (simp add: to_nat_less_card) from neq[folded inj[OF pc swap(1)] inj[OF pc swap(2)]] have "?fn (p (?tn c)) \ ?fn a" "?fn (p (?tn c)) \ ?fn b" . - with neq show ?thesis by (auto simp: o_def swap_def) + with neq show ?thesis by (auto simp: o_def swap_id_eq) qed qed - show ?case unfolding IH2 id sign_compose[OF p_sab swap(3)] sign_compose[OF p_sfab p_ftp] id2 + show ?case unfolding IH2 id sign_compose[OF p_sab \permutation p\] sign_compose[OF p_sfab p_ftp] id2 by (rule conjI[OF refl perm1]) qed thus "signof p = of_int (sign ?q)" unfolding signof_def sign_def by auto show "(\i = 0..i\UNIV. a $h i $h ?q i)" unfolding range_to_nat[symmetric] prod.reindex[OF inj_to_nat] by (rule prod.cong[OF refl], unfold o_def, simp) qed } thus ?thesis unfolding HMA_M_def by (auto simp: from_hma\<^sub>m_def Determinant.det_def det_def) qed lemma HMA_mat[transfer_rule]: "((=) ===> HMA_M) (\ k. k \\<^sub>m 1\<^sub>m CARD('n)) (Finite_Cartesian_Product.mat :: 'a::semiring_1 \ 'a^'n :: mod_type^'n :: mod_type)" unfolding Finite_Cartesian_Product.mat_def[abs_def] rel_fun_def HMA_M_def by (auto simp: from_hma\<^sub>m_def from_nat_inj) lemma HMA_mat_minus[transfer_rule]: "(HMA_M ===> HMA_M ===> HMA_M) (\ A B. A + map_mat uminus B) ((-) :: 'a :: group_add ^'nc:: mod_type^'nr:: mod_type \ 'a^'nc:: mod_type^'nr:: mod_type \ 'a^'nc:: mod_type^'nr:: mod_type)" unfolding rel_fun_def HMA_M_def from_hma\<^sub>m_def by auto lemma HMA_transpose_matrix [transfer_rule]: "(HMA_M ===> HMA_M) transpose_mat transpose" unfolding transpose_mat_def transpose_def HMA_M_def from_hma\<^sub>m_def by auto lemma HMA_invertible_matrix_mod_type[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'n :: mod_type ^ 'n :: mod_type \ _) ===> (=)) invertible_mat invertible" proof (intro rel_funI, goal_cases) case (1 x y) note rel_xy[transfer_rule] = "1" have eq_dim: "dim_col x = dim_row x" using Mod_Type_Connect.dim_col_transfer_rule Mod_Type_Connect.dim_row_transfer_rule rel_xy by fastforce moreover have "\A'. y ** A' = mat 1 \ A' ** y = mat 1" if xB: "x * B = 1\<^sub>m (dim_row x)" and Bx: "B * x = 1\<^sub>m (dim_row B)" for B proof - let ?A' = "Mod_Type_Connect.to_hma\<^sub>m B:: 'a :: comm_ring_1 ^ 'n :: mod_type^ 'n :: mod_type" have rel_BA[transfer_rule]: "Mod_Type_Connect.HMA_M B ?A'" by (metis (no_types, lifting) Bx Mod_Type_Connect.HMA_M_def eq_dim carrier_mat_triv dim_col_mat(1) Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.from_hma_to_hma\<^sub>m index_mult_mat(3) index_one_mat(3) rel_xy xB) have [simp]: "dim_row B = CARD('n)" using Mod_Type_Connect.dim_row_transfer_rule rel_BA by blast have [simp]: "dim_row x = CARD('n)" using Mod_Type_Connect.dim_row_transfer_rule rel_xy by blast have "y ** ?A' = mat 1" using xB by (transfer, simp) moreover have "?A' ** y = mat 1" using Bx by (transfer, simp) ultimately show ?thesis by blast qed moreover have "\B. x * B = 1\<^sub>m (dim_row x) \ B * x = 1\<^sub>m (dim_row B)" if yA: "y ** A' = mat 1" and Ay: "A' ** y = mat 1" for A' proof - let ?B = "(Mod_Type_Connect.from_hma\<^sub>m A')" have [simp]: "dim_row x = CARD('n)" using rel_xy Mod_Type_Connect.dim_row_transfer_rule by blast have [transfer_rule]: "Mod_Type_Connect.HMA_M ?B A'" by (simp add: Mod_Type_Connect.HMA_M_def) hence [simp]: "dim_row ?B = CARD('n)" using dim_row_transfer_rule by auto have "x * ?B = 1\<^sub>m (dim_row x)" using yA by (transfer', auto) moreover have "?B * x = 1\<^sub>m (dim_row ?B)" using Ay by (transfer', auto) ultimately show ?thesis by auto qed ultimately show ?case unfolding invertible_mat_def invertible_def inverts_mat_def by auto qed end text \Some transfer rules for relating the elementary operations are also proved.\ context includes lifting_syntax begin lemma HMA_swaprows[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (Mod_Type_Connect.HMA_I :: _ \'nr :: mod_type \ _ ) ===> (Mod_Type_Connect.HMA_I :: _ \'nr :: mod_type \ _ ) ===> Mod_Type_Connect.HMA_M) (\A a b. swaprows a b A) interchange_rows" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def interchange_rows_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.HMA_I_def to_nat_less_card to_nat_from_nat_id) lemma HMA_swapcols[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (Mod_Type_Connect.HMA_I :: _ \'nc :: mod_type \ _ ) ===> (Mod_Type_Connect.HMA_I :: _ \'nc :: mod_type \ _ ) ===> Mod_Type_Connect.HMA_M) (\A a b. swapcols a b A) interchange_columns" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def interchange_columns_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.HMA_I_def to_nat_less_card to_nat_from_nat_id) lemma HMA_addrow[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (Mod_Type_Connect.HMA_I :: _ \'nr :: mod_type \ _ ) ===> (Mod_Type_Connect.HMA_I :: _ \'nr :: mod_type \ _ ) ===> (=) ===> Mod_Type_Connect.HMA_M) (\A a b q. addrow q a b A) row_add" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def row_add_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.HMA_I_def to_nat_less_card to_nat_from_nat_id) lemma HMA_addcol[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (Mod_Type_Connect.HMA_I :: _ \'nc :: mod_type \ _ ) ===> (Mod_Type_Connect.HMA_I :: _ \'nc :: mod_type \ _ ) ===> (=) ===> Mod_Type_Connect.HMA_M) (\A a b q. addcol q a b A) column_add" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def column_add_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.HMA_I_def to_nat_less_card to_nat_from_nat_id) lemma HMA_multrow[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (Mod_Type_Connect.HMA_I :: _ \'nr :: mod_type \ _ ) ===> (=) ===> Mod_Type_Connect.HMA_M) (\A i q. multrow i q A) mult_row" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def mult_row_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.HMA_I_def to_nat_less_card to_nat_from_nat_id) lemma HMA_multcol[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (Mod_Type_Connect.HMA_I :: _ \'nc :: mod_type \ _ ) ===> (=) ===> Mod_Type_Connect.HMA_M) (\A i q. multcol i q A) mult_column" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def mult_column_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def Mod_Type_Connect.HMA_I_def to_nat_less_card to_nat_from_nat_id) end fun HMA_M3 where "HMA_M3 (P,A,Q) (P' :: 'a :: comm_ring_1 ^ 'nr :: mod_type ^ 'nr :: mod_type, A' :: 'a ^ 'nc :: mod_type ^ 'nr :: mod_type, Q' :: 'a ^ 'nc :: mod_type ^ 'nc :: mod_type) = (Mod_Type_Connect.HMA_M P P' \ Mod_Type_Connect.HMA_M A A' \ Mod_Type_Connect.HMA_M Q Q')" lemma HMA_M3_def: "HMA_M3 A B = (Mod_Type_Connect.HMA_M (fst A) (fst B) \ Mod_Type_Connect.HMA_M (fst (snd A)) (fst (snd B)) \ Mod_Type_Connect.HMA_M (snd (snd A)) (snd (snd B)))" by (smt HMA_M3.simps prod.collapse) context includes lifting_syntax begin lemma Domainp_HMA_M3 [transfer_domain_rule]: "Domainp (HMA_M3 :: _\(_\('a::comm_ring_1^'nc::mod_type^'nr::mod_type)\_)\_) = (\(P,A,Q). P \ carrier_mat CARD('nr) CARD('nr) \ A \ carrier_mat CARD('nr) CARD('nc) \ Q \ carrier_mat CARD('nc) CARD('nc))" proof - let ?HMA_M3 = "HMA_M3::_\(_\('a::comm_ring_1^'nc::mod_type^'nr::mod_type)\_)\_" have 1: "P \ carrier_mat CARD('nr) CARD('nr) \ A \ carrier_mat CARD('nr) CARD('nc) \ Q \ carrier_mat CARD('nc) CARD('nc)" if "Domainp ?HMA_M3 (P,A,Q)" for P A Q using that unfolding Domainp_iff by (auto simp add: Mod_Type_Connect.HMA_M_def) have 2: "Domainp ?HMA_M3 (P,A,Q)" if PAQ: "P \ carrier_mat CARD('nr) CARD('nr) \ A \ carrier_mat CARD('nr) CARD('nc) \Q \ carrier_mat CARD('nc) CARD('nc)" for P A Q proof - let ?P = "Mod_Type_Connect.to_hma\<^sub>m P::'a^'nr::mod_type^'nr::mod_type" let ?A = "Mod_Type_Connect.to_hma\<^sub>m A::'a^'nc::mod_type^'nr::mod_type" let ?Q = "Mod_Type_Connect.to_hma\<^sub>m Q::'a^'nc::mod_type^'nc::mod_type" have "HMA_M3 (P,A,Q) (?P,?A,?Q)" by (auto simp add: Mod_Type_Connect.HMA_M_def PAQ) thus ?thesis unfolding Domainp_iff by auto qed have "fst x \ carrier_mat CARD('nr) CARD('nr) \ fst (snd x) \ carrier_mat CARD('nr) CARD('nc) \ (snd (snd x)) \ carrier_mat CARD('nc) CARD('nc)" if "Domainp ?HMA_M3 x" for x using 1 by (metis (full_types) surjective_pairing that) moreover have "Domainp ?HMA_M3 x" if "fst x \ carrier_mat CARD('nr) CARD('nr) \ fst (snd x) \ carrier_mat CARD('nr) CARD('nc) \ (snd (snd x)) \ carrier_mat CARD('nc) CARD('nc)" for x using 2 by (metis (full_types) surjective_pairing that) ultimately show ?thesis by (intro ext iffI, unfold split_beta, metis+) qed lemma bi_unique_HMA_M3 [transfer_rule]: "bi_unique HMA_M3" "left_unique HMA_M3" "right_unique HMA_M3" unfolding HMA_M3_def bi_unique_def left_unique_def right_unique_def by (auto simp add: Mod_Type_Connect.HMA_M_def) lemma right_total_HMA_M3 [transfer_rule]: "right_total HMA_M3" unfolding HMA_M_def right_total_def by (simp add: Mod_Type_Connect.HMA_M_def) end (* TODO: add more theorems to connect everything from HA to JNF in this setting. *) end