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,425 +1,426 @@ (*<*) theory Blue_Eyes imports "HOL-Combinatorics.Transposition" 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))" + "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 transpose c\<^sub>1 c\<^sub>2 (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_id_eq) + by (cases \c\<^sub>1 = blue\; cases \c\<^sub>2 = blue\) + (auto simp add: try_swap_def valid_def transpose_eq_iff) 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_id_eq) + by (auto simp add: try_swap_def transpose_eq_iff) 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_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_id_eq) + then have [simp]: "\j q. A j (Transposition.transpose i i' (q j)) = A j (q j)" + by (simp add: vec_eq_iff Transposition.transpose_def) let ?p = "\p. of_int (sign p) * (\i\UNIV. A i (p i))" - let ?s = "\q. Fun.swap i i' id \ q" + let ?s = "\q. Transposition.transpose i i' \ 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_id_eq) + by (auto simp: inj_on_def fun_eq_iff Transposition.transpose_def) 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,518 +1,530 @@ (* Author: Lukas Bulwahn *) section \Derangements\ theory Derangements imports Complex_Main "HOL-Combinatorics.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-Combinatorics.Permutations"} Theory\ lemma permutes_imp_bij': assumes "p permutes S" shows "bij p" -using assms by (meson bij_def permutes_inj permutes_surj) +using assms by (fact permutes_bij) 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 + by (rule permutes_compose) (use assms in \simp_all add: assms permutes_imp_permutes_insert permutes_swap_id\) 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))+ 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)) + \p \ Transposition.transpose x (p x) permutes (S - {x, p x})\ + if \p permutes S\ \p (p x) = x\ +proof (rule bij_imp_permutes') + from \p permutes S\ have \bij p\ + by (rule permutes_imp_bij') + then show \bij (p \ Transposition.transpose x (p x))\ + by (simp add: bij_swap_iff) + from \p (p x) = x\ \p permutes S\ have \(p \ Transposition.transpose x (p x)) y = y\ + if \y \ S - {x, p x}\ for y + using that by (cases \y \ {x, p x}\) (auto simp add: permutes_not_in) + then show \\y. y \ S - {x, p x} \ + (p \ Transposition.transpose x (p x)) y = y\ + by blast +qed + 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)"]) + with \d x \ S\ \d x \ x\ 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/Gauss-Jordan-Elim-Fun/Gauss_Jordan_Elim_Fun.thy b/thys/Gauss-Jordan-Elim-Fun/Gauss_Jordan_Elim_Fun.thy --- a/thys/Gauss-Jordan-Elim-Fun/Gauss_Jordan_Elim_Fun.thy +++ b/thys/Gauss-Jordan-Elim-Fun/Gauss_Jordan_Elim_Fun.thy @@ -1,262 +1,265 @@ (* Gauss-Jordan elimination for matrices represented as functions Author: Tobias Nipkow *) section \Gauss-Jordan elimination algorithm\ theory Gauss_Jordan_Elim_Fun imports "HOL-Combinatorics.Transposition" begin text\Matrices are functions:\ type_synonym 'a matrix = "nat \ nat \ 'a" text\In order to restrict to finite matrices, a matrix is usually combined with one or two natural numbers indicating the maximal row and column of the matrix. Gauss-Jordan elimination is parameterized with a natural number \n\. It indicates that the matrix \A\ has \n\ rows and columns. In fact, \A\ is the augmented matrix with \n+1\ columns. Column \n\ is the ``right-hand side'', i.e.\ the constant vector \b\. The result is the unit matrix augmented with the solution in column \n\; see the correctness theorem below.\ fun gauss_jordan :: "('a::field)matrix \ nat \ ('a)matrix option" where "gauss_jordan A 0 = Some(A)" | "gauss_jordan A (Suc m) = (case dropWhile (\i. A i m = 0) [0.. None | p # _ \ (let Ap' = (\j. A p j / A p m); A' = (\i. if i=p then Ap' else (\j. A i j - A i m * Ap' j)) in gauss_jordan (Fun.swap p m A') m))" text\Some auxiliary functions:\ definition solution :: "('a::field)matrix \ nat \ (nat \ 'a) \ bool" where "solution A n x = (\i j=0.. nat \ nat \ bool" where "unit A m n = (\i j::nat. m\j \ j A i j = (if i=j then 1 else 0))" lemma solution_swap: assumes "p1 < n" "p2 < n" shows "solution (Fun.swap p1 p2 A) n x = solution A n x" (is "?L = ?R") proof(cases "p1=p2") case True thus ?thesis by simp next case False show ?thesis proof assume ?R thus ?L using assms False by(simp add: solution_def Fun.swap_def) next assume ?L show ?R proof(auto simp: solution_def) fix i assume "ij = 0..?L\ assms False show ?thesis by(fastforce simp add: solution_def Fun.swap_def) next assume "i\p1" show ?thesis proof cases assume "i=p2" with \?L\ assms False show ?thesis by(fastforce simp add: solution_def Fun.swap_def) next assume "i\p2" with \i\p1\ \?L\ \i assms False show ?thesis by(fastforce simp add: solution_def Fun.swap_def) qed qed qed qed qed (* Converting these apply scripts makes them blow up - see above *) lemma solution_upd1: "c \ 0 \ solution (A(p:=(\j. A p j / c))) n x = solution A n x" apply(cases "p ap = A p; \i j. i\p \ a i j = A i j; p \ solution (\i. if i=p then ap else (\j. a i j - c i * ap j)) n x = solution A n x" apply(clarsimp simp add: solution_def) apply rule prefer 2 apply (simp add: field_simps sum_subtractf sum_distrib_left[symmetric]) apply(clarsimp) apply(case_tac "i=p") apply simp apply (auto simp add: field_simps sum_subtractf sum_distrib_left[symmetric] all_conj_distrib) done subsection\Correctness\ text\The correctness proof:\ lemma gauss_jordan_lemma: "m\n \ unit A m n \ gauss_jordan A m = Some B \ unit B 0 n \ solution A n (\j. B j n)" proof(induct m arbitrary: A B) case 0 { fix a and b c d :: "'a" have "(if a then b else c) * d = (if a then b*d else c*d)" by simp } with 0 show ?case by(simp add: unit_def solution_def sum.If_cases) next case (Suc m) let "?Ap' p" = "(\j. A p j / A p m)" let "?A' p" = "(\i. if i=p then ?Ap' p else (\j. A i j - A i m * ?Ap' p j))" from \gauss_jordan A (Suc m) = Some B\ obtain p ks where "dropWhile (\i. A i m = 0) [0..m" "A p m \ 0" apply(simp_all add: dropWhile_eq_Cons_conv del:upt_Suc) by (metis set_upt atLeast0AtMost atLeastLessThanSuc_atLeastAtMost atMost_iff in_set_conv_decomp) have "m\n" "mSuc m \ n\ by arith+ have "unit (Fun.swap p m (?A' p)) m n" using Suc.prems(2) p unfolding unit_def Fun.swap_def Suc_le_eq by (auto simp: le_less) from Suc.hyps[OF \m\n\ this rec] \m p show ?case - by(simp add: solution_swap solution_upd1 solution_upd_but1[where A = "A(p := ?Ap' p)"]) + by (simp only: solution_swap) (simp_all add: solution_swap solution_upd_but1 [where A = "A(p := ?Ap' p)"] solution_upd1) qed theorem gauss_jordan_correct: "gauss_jordan A n = Some B \ solution A n (\j. B j n)" by(simp add:gauss_jordan_lemma[of n n] unit_def field_simps) definition solution2 :: "('a::field)matrix \ nat \ nat \ (nat \ 'a) \ bool" where "solution2 A m n x = (\i j=0.. solution2 A m n x \ (\y. solution2 A m n y \ (\jp 0" proof(rule ccontr) assume "\(\p 0)" hence 1: "\p. p A p q = 0" by simp { fix y assume 2: "\j. j\q \ y j = x j" { fix i assume "ij = 0..i] 2 have "(\j = 0..q < m\ show False by (simp add: usolution_def) (metis fun_upd_same zero_neq_one) qed lemma lem1: fixes f :: "'a \ 'b::field" shows "(\x\A. f x * (a * g x)) = a * (\x\A. f x * g x)" by (simp add: sum_distrib_left field_simps) lemma lem2: fixes f :: "'a \ 'b::field" shows "(\x\A. f x * (g x * a)) = a * (\x\A. f x * g x)" by (simp add: sum_distrib_left field_simps) subsection\Complete\ lemma gauss_jordan_complete: "m \ n \ usolution A m n x \ \B. gauss_jordan A m = Some B" proof(induction m arbitrary: A) case 0 show ?case by simp next case (Suc m A) from \Suc m \ n\ have "m\n" and "mm] obtain p' where "p' 0" by blast hence "dropWhile (\i. A i m = 0) [0.. []" by (simp add: atLeast0LessThan) (metis lessThan_iff linorder_neqE_nat not_less_eq) then obtain p xs where 1: "dropWhile (\i. A i m = 0) [0..m" "A p m \ 0" by (simp_all add: dropWhile_eq_Cons_conv del: upt_Suc) (metis set_upt atLeast0AtMost atLeastLessThanSuc_atLeastAtMost atMost_iff in_set_conv_decomp) then have p: "p < Suc m" "A p m \ 0" by auto let ?Ap' = "(\j. A p j / A p m)" let ?A' = "(\i. if i=p then ?Ap' else (\j. A i j - A i m * ?Ap' j))" let ?A = "Fun.swap p m ?A'" have A: "solution2 A (Suc m) n x" using Suc.prems(2) by(simp add: usolution_def) { fix i assume le_m: "p < Suc m" "i < Suc m" "A p m \ 0" have "(\j = 0..j = 0..j = 0.. = A i n - A p n * A i m / A p m" using A le_m by (simp add: solution2_def field_simps del: sum.op_ivl_Suc) finally have "(\j = 0..j=0.. p" show ?thesis proof (cases "i = m") assume "i = m" with p \i \ p\ have "p < m" by simp with a[unfolded solution2_def, THEN spec, of p] p(2) have "A p m * (A m m * A p n + A p m * (\j = 0..j = 0..A p m \ 0\ show ?thesis unfolding \i = m\ by simp (simp add: field_simps) next assume "i \ m" then have "i < m" using \i < Suc m\ by simp with a[unfolded solution2_def, THEN spec, of i] p(2) have "A p m * (A i m * A p n + A p m * (\j = 0..j = 0..A p m \ 0\ show ?thesis by simp (simp add: field_simps) qed qed qed with \usolution A (Suc m) n x\ have "\jjm\n\ this] 1 show ?case by(simp) + } ultimately have "usolution ?A m n x" + by (simp add: usolution_def) + note * = Suc.IH [OF \m \ n\ this] + from 1 show ?case + by auto (use * in blast) qed text\Future work: extend the proof to matrix inversion.\ hide_const (open) unit end diff --git a/thys/Gauss_Jordan/Determinants2.thy b/thys/Gauss_Jordan/Determinants2.thy --- a/thys/Gauss_Jordan/Determinants2.thy +++ b/thys/Gauss_Jordan/Determinants2.thy @@ -1,501 +1,503 @@ (* Title: Determinants2.thy Author: Jose Divasón Author: Jesús Aransay *) section\Computing determinants of matrices using the Gauss Jordan algorithm\ theory Determinants2 imports Gauss_Jordan_PA begin subsection\Some previous properties\ subsubsection\Relationships between determinants and elementary row operations\ lemma det_interchange_rows: shows "det (interchange_rows A i j) = of_int (if i = j then 1 else -1) * det A" proof - - have "(interchange_rows A i j) = (\ a. A $ (Fun.swap i j id) a)" unfolding interchange_rows_def Fun.swap_def by vector - hence "det(interchange_rows A i j) = det(\ a. A$(Fun.swap i j id) a)" by simp - also have "... = of_int (sign (Fun.swap i j id)) * det A" by (rule det_permute_rows[of "Fun.swap i j id" A], simp add: permutes_swap_id) + have "(interchange_rows A i j) = (\ a. A $ (Transposition.transpose i j) a)" + unfolding interchange_rows_def Transposition.transpose_def by vector + hence "det(interchange_rows A i j) = det(\ a. A$(Transposition.transpose i j) a)" by simp + also have "... = of_int (sign (Transposition.transpose i j)) * det A" by (rule det_permute_rows[of "Transposition.transpose i j" A], simp add: permutes_swap_id) finally show ?thesis unfolding sign_swap_id . qed corollary det_interchange_different_rows: assumes i_not_j: "i \ j" shows "det (interchange_rows A i j) = - det A" unfolding det_interchange_rows using i_not_j by simp corollary det_interchange_same_rows: assumes i_eq_j: "i = j" shows "det (interchange_rows A i j) = det A" unfolding det_interchange_rows using i_eq_j by simp lemma det_mult_row: shows "det (mult_row A a k) = k * det A" proof - have A_rw: "(\ i. if i = a then A$a else A$i) = A" by vector have "(mult_row A a k) = (\ i. if i = a then k *s A $ a else A $ i)" unfolding mult_row_def by vector hence "det(mult_row A a k) = det(\ i. if i = a then k *s A $ a else A $ i)" by simp also have "... = k * det(\ i. if i = a then A$a else A$i)" unfolding det_row_mul .. also have "... = k * det A" unfolding A_rw .. finally show ?thesis . qed (*The name det_row_add is already used in the Determinants.thy file of the standard library*) lemma det_row_add': assumes i_not_j: "i \ j" shows "det (row_add A i j q) = det A" proof - have "(row_add A i j q) = (\ k. if k = i then row i A + q *s row j A else row k A)" unfolding row_add_def row_def by vector hence "det(row_add A i j q) = det(\ k. if k = i then row i A + q *s row j A else row k A)" by simp also have "... = det A" unfolding det_row_operation[OF i_not_j] .. finally show ?thesis . qed subsubsection\Relationships between determinants and elementary column operations\ lemma det_interchange_columns: shows "det (interchange_columns A i j) = of_int (if i = j then 1 else -1) * det A" proof - -have "(interchange_columns A i j) = (\ a b. A $ a $ (Fun.swap i j id) b)" unfolding interchange_columns_def Fun.swap_def by vector -hence "det(interchange_columns A i j) = det(\ a b. A $ a $ (Fun.swap i j id) b)" by simp -also have "... = of_int (sign (Fun.swap i j id)) * det A" by (rule det_permute_columns[of "Fun.swap i j id" A], simp add: permutes_swap_id) + have "(interchange_columns A i j) = (\ a b. A $ a $ (Transposition.transpose i j) b)" + unfolding interchange_columns_def Transposition.transpose_def by vector +hence "det(interchange_columns A i j) = det(\ a b. A $ a $ (Transposition.transpose i j) b)" by simp +also have "... = of_int (sign (Transposition.transpose i j)) * det A" by (rule det_permute_columns[of "Transposition.transpose i j" A], simp add: permutes_swap_id) finally show ?thesis unfolding sign_swap_id . qed corollary det_interchange_different_columns: assumes i_not_j: "i \ j" shows "det (interchange_columns A i j) = - det A" unfolding det_interchange_columns using i_not_j by simp corollary det_interchange_same_columns: assumes i_eq_j: "i = j" shows "det (interchange_columns A i j) = det A" unfolding det_interchange_columns using i_eq_j by simp lemma det_mult_columns: shows "det (mult_column A a k) = k * det A" proof - have "mult_column A a k = transpose (mult_row (transpose A) a k)" unfolding transpose_def mult_row_def mult_column_def by vector hence "det (mult_column A a k) = det (transpose (mult_row (transpose A) a k))" by simp also have "... = det (mult_row (transpose A) a k)" unfolding det_transpose .. also have "... = k * det (transpose A)" unfolding det_mult_row .. also have "... = k * det A" unfolding det_transpose .. finally show ?thesis . qed lemma det_column_add: assumes i_not_j: "i \ j" shows "det (column_add A i j q) = det A" proof - have "(column_add A i j q) = (transpose (row_add (transpose A) i j q))" unfolding transpose_def column_add_def row_add_def by vector hence "det (column_add A i j q) = det (transpose (row_add (transpose A) i j q))" by simp also have "... = det (row_add (transpose A) i j q)" unfolding det_transpose .. also have "... = det A" unfolding det_row_add'[OF i_not_j] det_transpose .. finally show ?thesis . qed subsection\Proving that the determinant can be computed by means of the Gauss Jordan algorithm\ subsubsection\Previous properties\ lemma det_row_add_iterate_upt_n: fixes A::"'a::{comm_ring_1}^'n::{mod_type}^'n::{mod_type}" assumes n: "n to_nat i" have "det (row_add_iterate (row_add A (from_nat (Suc n)) i (- A $ from_nat (Suc n) $ j)) n i j) = det (row_add A (from_nat (Suc n)) i (- A $ from_nat (Suc n) $ j))" proof (rule Suc.hyps, unfold nrows_def) show " n < CARD('n)" using Suc.prems unfolding nrows_def by auto qed also have "... = det A" proof (rule det_row_add',rule ccontr, simp) assume "from_nat (Suc n) = i" hence "to_nat (from_nat (Suc n)::'n) = to_nat i" by simp hence "(Suc n) = to_nat i" unfolding to_nat_from_nat_id[OF Suc.prems[unfolded nrows_def]] . thus False using Suc_n_not_i by contradiction qed finally show "det (row_add_iterate (row_add A (from_nat (Suc n)) i (- A $ from_nat (Suc n) $ j)) n i j) = det A" . qed qed corollary det_row_add_iterate: fixes A::"'a::{comm_ring_1}^'n::{mod_type}^'n::{mod_type}" shows "det (row_add_iterate A (nrows A - 1) i j) = det A" by (metis det_row_add_iterate_upt_n diff_less neq0_conv nrows_not_0 zero_less_one) lemma det_Gauss_Jordan_in_ij: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" and i j::"'n" defines A': "A'== mult_row (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) i (1 / (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) $ i $ j)" shows "det (Gauss_Jordan_in_ij A i j) = det A' " proof - have nrows_eq: "nrows A' = nrows A" unfolding nrows_def by simp have "row_add_iterate A' (nrows A - 1) i j = Gauss_Jordan_in_ij A i j" using row_add_iterate_eq_Gauss_Jordan_in_ij unfolding A' . hence "det (Gauss_Jordan_in_ij A i j) = det (row_add_iterate A' (nrows A - 1) i j)" by simp also have "... = det A'" by (rule det_row_add_iterate[of A', unfolded nrows_eq]) finally show ?thesis . qed lemma det_Gauss_Jordan_in_ij_1: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" and i j::"'n" defines A': "A'== mult_row (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) i (1 / (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) $ i $ j)" assumes i: "(LEAST n. A $ n $ j \ 0 \ i \ n) = i" shows "det (Gauss_Jordan_in_ij A i j) = 1/(A$i$j) * det A" proof - have "det (Gauss_Jordan_in_ij A i j) = det A' " using det_Gauss_Jordan_in_ij unfolding A' by auto also have "... = 1/(A$i$j) * det A" unfolding A' det_mult_row unfolding i det_interchange_rows by auto finally show ?thesis . qed lemma det_Gauss_Jordan_in_ij_2: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" and i j::"'n" defines A': "A'== mult_row (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) i (1 / (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) $ i $ j)" assumes i: "(LEAST n. A $ n $ j \ 0 \ i \ n) \ i" shows "det (Gauss_Jordan_in_ij A i j) = - 1/(A $ (LEAST n. A $ n $ j \ 0 \ i \ n) $ j) * det A" proof - have "det (Gauss_Jordan_in_ij A i j) = det A' " using det_Gauss_Jordan_in_ij unfolding A' by auto also have "... = - 1/(A$ (LEAST n. A $ n $ j \ 0 \ i \ n) $j) * det A" unfolding A' det_mult_row unfolding det_interchange_rows using i by auto finally show ?thesis . qed subsubsection\Definitions\ text\The following definitions allow the computation of the determinant of a matrix using the Gauss-Jordan algorithm. In the first component the determinant of each transformation is accumulated and the second component contains the matrix transformed into a reduced row echelon form matrix\ definition Gauss_Jordan_in_ij_det_P :: "'a::{semiring_1, inverse, one, uminus}^'m^'n::{finite, ord}=> 'n=>'m=>('a \ ('a^'m^'n::{finite, ord}))" where "Gauss_Jordan_in_ij_det_P A i j = (let n = (LEAST n. A $ n $ j \ 0 \ i \ n) in (if i = n then 1/(A $ i $ j) else - 1/(A $ n $ j), Gauss_Jordan_in_ij A i j))" definition Gauss_Jordan_column_k_det_P where "Gauss_Jordan_column_k_det_P A' k = (let det_P= fst A'; i = fst (snd A'); A = snd (snd A'); from_nat_i = from_nat i; from_nat_k = from_nat k in if (\m\from_nat_i. A $ m $ from_nat_k = 0) \ i = nrows A then (det_P, i, A) else let gauss = Gauss_Jordan_in_ij_det_P A (from_nat_i) (from_nat_k) in (fst gauss * det_P, i + 1, snd gauss))" definition Gauss_Jordan_upt_k_det_P where "Gauss_Jordan_upt_k_det_P A k = (let foldl = foldl Gauss_Jordan_column_k_det_P (1, 0, A) [0..Proofs\ text\This is an equivalent definition created to achieve a more efficient computation.\ lemma Gauss_Jordan_in_ij_det_P_code[code]: shows "Gauss_Jordan_in_ij_det_P A i j = (let n = (LEAST n. A $ n $ j \ 0 \ i \ n); interchange_A = interchange_rows A i n; A' = mult_row interchange_A i (1 / interchange_A $ i $ j) in (if i = n then 1/(A $ i $ j) else - 1/(A $ n $ j), Gauss_Jordan_wrapper i j A' interchange_A))" unfolding Gauss_Jordan_in_ij_det_P_def Gauss_Jordan_in_ij_def Gauss_Jordan_wrapper_def Let_def by auto lemma det_Gauss_Jordan_in_ij_det_P: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" and i j::"'n" shows "(fst (Gauss_Jordan_in_ij_det_P A i j)) * det A = det (snd (Gauss_Jordan_in_ij_det_P A i j))" unfolding Gauss_Jordan_in_ij_det_P_def Let_def fst_conv snd_conv using det_Gauss_Jordan_in_ij_1[of A j i] using det_Gauss_Jordan_in_ij_2[of A j i] by auto lemma det_Gauss_Jordan_column_k_det_P: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" assumes det: "det_P * det B = det A" shows "(fst (Gauss_Jordan_column_k_det_P (det_P,i,A) k)) * det B = det (snd (snd (Gauss_Jordan_column_k_det_P (det_P,i,A) k)))" proof (unfold Gauss_Jordan_column_k_det_P_def Let_def, auto simp add: assms) fix m assume i_not_nrows: "i \ nrows A" and i_less_m: "from_nat i \ m" and Amk_not_0: "A $ m $ from_nat k \ 0" show "fst (Gauss_Jordan_in_ij_det_P A (from_nat i) (from_nat k)) * det_P * det B = det (snd (Gauss_Jordan_in_ij_det_P A (from_nat i) (from_nat k)))" unfolding mult.assoc det unfolding det_Gauss_Jordan_in_ij_det_P .. qed lemma det_Gauss_Jordan_upt_k_det_P: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows "(fst (Gauss_Jordan_upt_k_det_P A k)) * det A = det (snd (Gauss_Jordan_upt_k_det_P A k))" proof (induct k) case 0 show ?case unfolding Gauss_Jordan_upt_k_det_P_def Let_def unfolding fst_conv snd_conv by (simp add:det_Gauss_Jordan_column_k_det_P) next case (Suc k) have suc_rw: "[0..i j. j to_nat j < k \ A $ i $ j = 0)" definition upper_triangular where "upper_triangular A = (\i j. j A $ i $ j = 0)" lemma upper_triangular_upt_imp_upper_triangular: assumes "upper_triangular_upt_k A (nrows A)" shows "upper_triangular A" using assms unfolding upper_triangular_upt_k_def upper_triangular_def nrows_def using to_nat_less_card[where ?'a='b] by blast lemma rref_imp_upper_triagular_upt: fixes A::"'a::{one, zero}^'n::{mod_type}^'n::{mod_type}" assumes "reduced_row_echelon_form A" shows "upper_triangular_upt_k A k" proof (induct k) case 0 show ?case unfolding upper_triangular_upt_k_def by simp next case (Suc k) show ?case unfolding upper_triangular_upt_k_def proof (clarify) fix i j::'n assume j_less_i: "j < i" and j_less_suc_k: "to_nat j < Suc k" show "A $ i $ j = 0" proof (cases "to_nat j < k") case True thus ?thesis using Suc.hyps unfolding upper_triangular_upt_k_def using j_less_i True by auto next case False hence j_eq_k: "to_nat j = k" using j_less_suc_k by simp have rref_suc: "reduced_row_echelon_form_upt_k A (Suc k)" by (metis assms rref_implies_rref_upt) show ?thesis proof (cases "A $ i $ from_nat k = 0") case True have "from_nat k = j" by (metis from_nat_to_nat_id j_eq_k) thus ?thesis using True by simp next case False have zero_i_k: "is_zero_row_upt_k i k A" unfolding is_zero_row_upt_k_def by (metis (hide_lams, mono_tags) Suc.hyps leD le_less_linear less_imp_le j_eq_k j_less_i le_trans to_nat_mono' upper_triangular_upt_k_def) have not_zero_i_suc_k: "\ is_zero_row_upt_k i (Suc k) A" unfolding is_zero_row_upt_k_def using False by (metis j_eq_k lessI to_nat_from_nat) have Least_eq: "(LEAST n. A $ i $ n \ 0) = from_nat k" proof (rule Least_equality) show "A $ i $ from_nat k \ 0" using False by simp show "\y. A $ i $ y \ 0 \ from_nat k \ y" by (metis (full_types) is_zero_row_upt_k_def not_le_imp_less to_nat_le zero_i_k) qed have i_not_k: "i \ from_nat k" by (metis less_irrefl from_nat_to_nat_id j_eq_k j_less_i) show ?thesis using rref_upt_condition4_explicit[OF rref_suc not_zero_i_suc_k i_not_k] unfolding Least_eq using rref_upt_condition1_explicit[OF rref_suc] using Suc.hyps unfolding upper_triangular_upt_k_def by (metis (mono_tags) leD not_le_imp_less is_zero_row_upt_k_def is_zero_row_upt_k_suc j_eq_k j_less_i not_zero_i_suc_k to_nat_from_nat to_nat_mono') qed qed qed qed lemma rref_imp_upper_triagular: assumes "reduced_row_echelon_form A" shows "upper_triangular A" by (metis assms rref_imp_upper_triagular_upt upper_triangular_upt_imp_upper_triangular) lemma det_Gauss_Jordan[code_unfold]: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows "det (Gauss_Jordan A) = prod (\i. (Gauss_Jordan A)$i$i) (UNIV:: 'n set)" using det_upperdiagonal rref_imp_upper_triagular[OF rref_Gauss_Jordan[of A]] unfolding upper_triangular_def by blast lemma snd_Gauss_Jordan_in_ij_det_P_is_snd_Gauss_Jordan_in_ij_PA: shows "snd (Gauss_Jordan_in_ij_det_P A i j) = snd (Gauss_Jordan_in_ij_PA (P,A) i j)" unfolding Gauss_Jordan_in_ij_det_P_def Gauss_Jordan_in_ij_PA_def unfolding Gauss_Jordan_in_ij_def Let_def snd_conv fst_conv .. lemma snd_Gauss_Jordan_column_k_det_P_is_snd_Gauss_Jordan_column_k_PA: shows "snd (Gauss_Jordan_column_k_det_P (n,i,A) k) = snd (Gauss_Jordan_column_k_PA (P,i,A) k)" unfolding Gauss_Jordan_column_k_det_P_def Gauss_Jordan_column_k_PA_def Let_def snd_conv unfolding fst_conv using snd_Gauss_Jordan_in_ij_det_P_is_snd_Gauss_Jordan_in_ij_PA by auto lemma det_fst_row_add_iterate_PA: fixes A::"'a::{comm_ring_1}^'n::{mod_type}^'n::{mod_type}" assumes n: "n 0 \ i \ n)) i (1 / interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n) $ i $ j)" define A' where "A' = mult_row (interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n)) i (1 / interchange_rows A i (LEAST n. A $ n $ j \ 0 \ i \ n) $ i $ j)" have "det (fst (Gauss_Jordan_in_ij_PA (P,A) i j)) = det (fst (row_add_iterate_PA (P',A') (nrows A - 1) i j))" unfolding fst_row_add_iterate_PA_eq_fst_Gauss_Jordan_in_ij_PA[symmetric] A'_def P'_def .. also have "...= det P'" by (rule det_fst_row_add_iterate_PA, simp add: nrows_def) also have "... = (if i = (LEAST n. A $ n $ j \ 0 \ i \ n) then 1 / A $ i $ j else - 1 / A $ (LEAST n. A $ n $ j \ 0 \ i \ n) $ j) * det P" proof (cases "i = (LEAST n. A $ n $ j \ 0 \ i \ n)") case True show ?thesis unfolding if_P[OF True] P'_def unfolding True[symmetric] unfolding interchange_same_rows unfolding det_mult_row .. next case False show ?thesis unfolding if_not_P[OF False] P'_def unfolding det_mult_row unfolding det_interchange_different_rows[OF False] by simp qed also have "... = fst (Gauss_Jordan_in_ij_det_P A i j) * det P" unfolding Gauss_Jordan_in_ij_det_P_def by simp finally show ?thesis .. qed lemma det_fst_Gauss_Jordan_column_k_PA_eq_fst_Gauss_Jordan_column_k_det_P: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows "fst (Gauss_Jordan_column_k_det_P (det P,i,A) k) = det (fst (Gauss_Jordan_column_k_PA (P,i,A) k))" unfolding Gauss_Jordan_column_k_det_P_def Gauss_Jordan_column_k_PA_def Let_def snd_conv fst_conv using det_fst_Gauss_Jordan_in_ij_PA_eq_fst_Gauss_Jordan_in_ij_det_P by auto lemma fst_snd_Gauss_Jordan_column_k_det_P_eq_fst_snd_Gauss_Jordan_column_k_PA: shows "fst (snd (Gauss_Jordan_column_k_det_P (n,i,A) k)) = fst (snd (Gauss_Jordan_column_k_PA (P,i,A) k))" unfolding Gauss_Jordan_column_k_det_P_def Gauss_Jordan_column_k_PA_def Let_def snd_conv fst_conv by auto text\The way of proving the following lemma is very similar to the demonstration of @{thm "rref_and_index_Gauss_Jordan_upt_k"}.\ lemma foldl_Gauss_Jordan_column_k_det_P: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows det_fst_Gauss_Jordan_upt_k_PA_eq_fst_Gauss_Jordan_upt_k_det_P: "fst (Gauss_Jordan_upt_k_det_P A k) = det (fst (Gauss_Jordan_upt_k_PA A k))" and snd_Gauss_Jordan_upt_k_det_P_is_snd_Gauss_Jordan_upt_k_PA: "snd (Gauss_Jordan_upt_k_det_P A k) = snd (Gauss_Jordan_upt_k_PA A k)" and fst_snd_foldl_Gauss_det_P_PA: "fst (snd (foldl Gauss_Jordan_column_k_det_P (1, 0, A) [0.. fst (Gauss_Jordan_upt_k_det_P A (Suc k)) = det (fst (Gauss_Jordan_upt_k_PA A (Suc k)))" unfolding Gauss_Jordan_upt_k_det_P_def unfolding Gauss_Jordan_upt_k_PA_def Let_def fst_conv unfolding list_rw foldl_append unfolding List.foldl.simps unfolding f_def[symmetric] g_def[symmetric] apply (subst f_rw) apply (subst g_rw) unfolding fst_snd snd_snd fst_det by (rule det_fst_Gauss_Jordan_column_k_PA_eq_fst_Gauss_Jordan_column_k_det_P) show "snd (Gauss_Jordan_upt_k_det_P A (Suc k)) = snd (Gauss_Jordan_upt_k_PA A (Suc k))" unfolding Gauss_Jordan_upt_k_det_P_def unfolding Gauss_Jordan_upt_k_PA_def Let_def fst_conv unfolding list_rw foldl_append unfolding List.foldl.simps unfolding f_def[symmetric] g_def[symmetric] apply (subst f_rw) apply (subst g_rw) unfolding fst_snd snd_snd fst_det by (metis fst_snd prod.collapse snd_Gauss_Jordan_column_k_det_P_is_snd_Gauss_Jordan_column_k_PA snd_eqD snd_snd) show "fst (snd (foldl Gauss_Jordan_column_k_det_P (1, 0, A) [0..i. (snd (Gauss_Jordan_det_P A))$i$i) (UNIV:: 'n set)" unfolding snd_Gauss_Jordan_det_P_is_Gauss_Jordan det_Gauss_Jordan .. lemma det_fst_Gauss_Jordan_PA_eq_fst_Gauss_Jordan_det_P: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows "fst (Gauss_Jordan_det_P A) = det (fst (Gauss_Jordan_PA A))" by (unfold Gauss_Jordan_det_P_def Gauss_Jordan_PA_def, rule det_fst_Gauss_Jordan_upt_k_PA_eq_fst_Gauss_Jordan_upt_k_det_P) lemma fst_Gauss_Jordan_det_P_not_0: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows "fst (Gauss_Jordan_det_P A) \ 0" unfolding det_fst_Gauss_Jordan_PA_eq_fst_Gauss_Jordan_det_P by (metis (mono_tags) det_I det_mul invertible_fst_Gauss_Jordan_PA matrix_inv_right mult_zero_left zero_neq_one) lemma det_code_equation[code_unfold]: fixes A::"'a::{field}^'n::{mod_type}^'n::{mod_type}" shows "det A = (let A' = Gauss_Jordan_det_P A in prod (\i. (snd (A'))$i$i) (UNIV::'n set)/(fst (A')))" unfolding Let_def using det_Gauss_Jordan_det_P[of A] unfolding det_snd_Gauss_Jordan_det_P by (simp add: fst_Gauss_Jordan_det_P_not_0 nonzero_eq_divide_eq ac_simps) end 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,2435 +1,2436 @@ (* 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_Misc 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" by (simp add: hom_uminus sign_def) 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 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 add: swap_id_eq) + "k < n \ l < n \ swaprows_mat n k l = mat n n (\(i, j). 1\<^sub>m n $$ (transpose k l i, j))" + by (rule eq_matI) (auto simp add: transpose_def) 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" + let ?p = "transpose k l" 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: 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 ?p = "transpose i j" let ?n = "{0 ..< n}" have sp: "signof ?p = - 1" "sign ?p = (- 1 :: int)" 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 :: int)}" let ?none = "{p. p permutes ?n \ sign p \ (1 :: int)}" 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 :: int)" 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" 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 :: int)" by auto from sign_compose[OF pp[OF p] pp[OF q], unfolded sq sp] have spq: "sign pq = (- 1 :: int)" 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 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: 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") + \permutation_insert (Suc i) j p \ transpose i (Suc i) = permutation_insert i j p\ (is \?l = ?r\) proof (rule ext) - fix x show "?l x = ?r x" - by (cases rule: linorder_cases[of "x" "i"]) - (auto simp add: swap_id_eq permutation_insert_expand) + fix x + consider \x < i\ | \x = i\ | \x = Suc i\ | \Suc i < x\ + using Suc_lessI by (cases rule: linorder_cases [of x i]) blast+ + then show \?l x = ?r x\ + by cases (simp_all add: 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" + shows "transpose j (Suc j) \ 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 ?swap = "transpose (n - Suc j) (n - j)" 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].. + 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" + let ?swap = "transpose (n - Suc i) (n-i)" 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] apply (rule degree_sum_le) apply (simp_all add: finite_permutations) apply (drule *) apply (rule order.trans [OF degree_mult_le]) apply simp apply (rule order.trans [OF degree_prod_sum_le]) apply simp_all done 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/LLL_Basis_Reduction/LLL.thy b/thys/LLL_Basis_Reduction/LLL.thy --- a/thys/LLL_Basis_Reduction/LLL.thy +++ b/thys/LLL_Basis_Reduction/LLL.thy @@ -1,1745 +1,1745 @@ (* Authors: Jose Divasón Maximilian Haslbeck Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \The LLL Algorithm\ text \Soundness of the LLL algorithm is proven in four steps. In the basic version, we do recompute the Gram-Schmidt ortogonal (GSO) basis in every step. This basic version will have a full functional soundness proof, i.e., termination and the property that the returned basis is reduced. Then in LLL-Number-Bounds we will strengthen the invariant and prove that all intermediate numbers stay polynomial in size. Moreover, in LLL-Impl we will refine the basic version, so that the GSO does not need to be recomputed in every step. Finally, in LLL-Complexity, we develop an cost-annotated version of the refined algorithm and prove a polynomial upper bound on the number of arithmetic operations.\ text \This theory provides a basic implementation and a soundness proof of the LLL algorithm to compute a "short" vector in a lattice.\ theory LLL imports Gram_Schmidt_2 Missing_Lemmas Jordan_Normal_Form.Determinant "Abstract-Rewriting.SN_Order_Carrier" begin subsection \Core Definitions, Invariants, and Theorems for Basic Version\ (* Note/TODO by Max Haslbeck: Up to here I refactored the code in Gram_Schmidt_2 and Gram_Schmidt_Int which now makes heavy use of locales. In the future I would also like to do this here (instead of using LLL_invariant everywhere). *) locale LLL = fixes n :: nat (* n-dimensional vectors, *) and m :: nat (* number of vectors *) and fs_init :: "int vec list" (* initial basis *) and \ :: rat (* approximation factor *) begin sublocale vec_module "TYPE(int)" n. abbreviation RAT where "RAT \ map (map_vec rat_of_int)" abbreviation SRAT where "SRAT xs \ set (RAT xs)" abbreviation Rn where "Rn \ carrier_vec n :: rat vec set" sublocale gs: gram_schmidt_fs n "RAT fs_init" . abbreviation lin_indep where "lin_indep fs \ gs.lin_indpt_list (RAT fs)" abbreviation gso where "gso fs \ gram_schmidt_fs.gso n (RAT fs)" abbreviation \ where "\ fs \ gram_schmidt_fs.\ n (RAT fs)" abbreviation reduced where "reduced fs \ gram_schmidt_fs.reduced n (RAT fs) \" abbreviation weakly_reduced where "weakly_reduced fs \ gram_schmidt_fs.weakly_reduced n (RAT fs) \" text \lattice of initial basis\ definition "L = lattice_of fs_init" text \maximum squared norm of initial basis\ definition "N = max_list (map (nat \ sq_norm) fs_init)" text \maximum absolute value in initial basis\ definition "M = Max ({abs (fs_init ! i $ j) | i j. i < m \ j < n} \ {0})" text \This is the core invariant which enables to prove functional correctness.\ definition "\_small fs i = (\ j < i. abs (\ fs i j) \ 1/2)" definition LLL_invariant_weak :: "int vec list \ bool" where "LLL_invariant_weak fs = ( gs.lin_indpt_list (RAT fs) \ lattice_of fs = L \ length fs = m)" lemma LLL_inv_wD: assumes "LLL_invariant_weak fs" shows "lin_indep fs" "length (RAT fs) = m" "set fs \ carrier_vec n" "\ i. i < m \ fs ! i \ carrier_vec n" "\ i. i < m \ gso fs i \ carrier_vec n" "length fs = m" "lattice_of fs = L" proof (atomize (full), goal_cases) case 1 interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs" by (standard) (use assms LLL_invariant_weak_def gs.lin_indpt_list_def in auto) show ?case using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier by (auto simp add: LLL_invariant_weak_def gram_schmidt_fs.reduced_def) qed lemma LLL_inv_wI: assumes "set fs \ carrier_vec n" "length fs = m" "lattice_of fs = L" "lin_indep fs" shows "LLL_invariant_weak fs" unfolding LLL_invariant_weak_def Let_def using assms by auto definition LLL_invariant :: "bool \ nat \ int vec list \ bool" where "LLL_invariant upw i fs = ( gs.lin_indpt_list (RAT fs) \ lattice_of fs = L \ reduced fs i \ i \ m \ length fs = m \ (upw \ \_small fs i) )" lemma LLL_inv_imp_w: "LLL_invariant upw i fs \ LLL_invariant_weak fs" unfolding LLL_invariant_def LLL_invariant_weak_def by blast lemma LLL_invD: assumes "LLL_invariant upw i fs" shows "lin_indep fs" "length (RAT fs) = m" "set fs \ carrier_vec n" "\ i. i < m \ fs ! i \ carrier_vec n" "\ i. i < m \ gso fs i \ carrier_vec n" "length fs = m" "lattice_of fs = L" "weakly_reduced fs i" "i \ m" "reduced fs i" "upw \ \_small fs i" proof (atomize (full), goal_cases) case 1 interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs" by (standard) (use assms LLL_invariant_def gs.lin_indpt_list_def in auto) show ?case using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier by (auto simp add: LLL_invariant_def gram_schmidt_fs.reduced_def) qed lemma LLL_invI: assumes "set fs \ carrier_vec n" "length fs = m" "lattice_of fs = L" "i \ m" "lin_indep fs" "reduced fs i" "upw \ \_small fs i" shows "LLL_invariant upw i fs" unfolding LLL_invariant_def Let_def split using assms by auto end locale fs_int' = fixes n m fs_init fs assumes LLL_inv: "LLL.LLL_invariant_weak n m fs_init fs" sublocale fs_int' \ fs_int_indpt using LLL_inv unfolding LLL.LLL_invariant_weak_def by (unfold_locales) blast context LLL begin lemma gso_cong: assumes "\ i. i \ x \ f1 ! i = f2 ! i" "x < length f1" "x < length f2" shows "gso f1 x = gso f2 x" by (rule gs.gso_cong, insert assms, auto) lemma \_cong: assumes "\ k. j < i \ k \ j \ f1 ! k = f2 ! k" and i: "i < length f1" "i < length f2" and "j < i \ f1 ! i = f2 ! i" shows "\ f1 i j = \ f2 i j" by (rule gs.\_cong, insert assms, auto) definition reduction where "reduction = (4+\)/(4*\)" definition d :: "int vec list \ nat \ int" where "d fs k = gs.Gramian_determinant fs k" definition D :: "int vec list \ nat" where "D fs = nat (\ i < m. d fs i)" definition "d\ gs i j = int_of_rat (of_int (d gs (Suc j)) * \ gs i j)" definition logD :: "int vec list \ nat" where "logD fs = (if \ = 4/3 then (D fs) else nat (floor (log (1 / of_rat reduction) (D fs))))" definition LLL_measure :: "nat \ int vec list \ nat" where "LLL_measure i fs = (2 * logD fs + m - i)" context fixes fs assumes Linv: "LLL_invariant_weak fs" begin interpretation fs: fs_int' n m fs_init fs by (standard) (use Linv in auto) lemma Gramian_determinant: assumes k: "k \ m" shows "of_int (gs.Gramian_determinant fs k) = (\ j 0" (is ?g2) using assms fs.Gramian_determinant LLL_inv_wD[OF Linv] by auto lemma LLL_d_pos [intro]: assumes k: "k \ m" shows "d fs k > 0" unfolding d_def using fs.Gramian_determinant k LLL_inv_wD[OF Linv] by auto lemma LLL_d_Suc: assumes k: "k < m" shows "of_int (d fs (Suc k)) = sq_norm (gso fs k) * of_int (d fs k)" using assms fs.fs_int_d_Suc LLL_inv_wD[OF Linv] unfolding fs.d_def d_def by auto lemma LLL_D_pos: shows "D fs > 0" using fs.fs_int_D_pos LLL_inv_wD[OF Linv] unfolding D_def fs.D_def fs.d_def d_def by auto end text \Condition when we can increase the value of $i$\ lemma increase_i: assumes Linv: "LLL_invariant upw i fs" assumes i: "i < m" and upw: "upw \ i = 0" and red_i: "i \ 0 \ sq_norm (gso fs (i - 1)) \ \ * sq_norm (gso fs i)" shows "LLL_invariant True (Suc i) fs" "LLL_measure i fs > LLL_measure (Suc i) fs" proof - note inv = LLL_invD[OF Linv] from inv(8,10) have red: "weakly_reduced fs i" and sred: "reduced fs i" by (auto) from red red_i i have red: "weakly_reduced fs (Suc i)" unfolding gram_schmidt_fs.weakly_reduced_def by (intro allI impI, rename_tac ii, case_tac "Suc ii = i", auto) from inv(11) upw have sred_i: "\ j. j < i \ \\ fs i j\ \ 1 / 2" unfolding \_small_def by auto from sred sred_i have sred: "reduced fs (Suc i)" unfolding gram_schmidt_fs.reduced_def by (intro conjI[OF red] allI impI, rename_tac ii j, case_tac "ii = i", auto) show "LLL_invariant True (Suc i) fs" by (intro LLL_invI, insert inv red sred i, auto) show "LLL_measure i fs > LLL_measure (Suc i) fs" unfolding LLL_measure_def using i by auto qed text \Standard addition step which makes $\mu_{i,j}$ small\ definition "\_small_row i fs j = (\ j'. j \ j' \ j' < i \ abs (\ fs i j') \ inverse 2)" lemma basis_reduction_add_row_main: assumes Linv: "LLL_invariant_weak fs" and i: "i < m" and j: "j < i" and fs': "fs' = fs[ i := fs ! i - c \\<^sub>v fs ! j]" shows "LLL_invariant_weak fs'" "LLL_invariant True i fs \ LLL_invariant True i fs'" "c = round (\ fs i j) \ \_small_row i fs (Suc j) \ \_small_row i fs' j" (* mu-value at position i j gets small *) "c = round (\ fs i j) \ abs (\ fs' i j) \ 1/2" (* mu-value at position i j gets small *) "LLL_measure i fs' = LLL_measure i fs" (* new values of gso: no change *) "\ i. i < m \ gso fs' i = gso fs i" (* new values of mu *) "\ i' j'. i' < m \ j' < m \ \ fs' i' j' = (if i' = i \ j' \ j then \ fs i j' - of_int c * \ fs j j' else \ fs i' j')" (* new values of d *) "\ ii. ii \ m \ d fs' ii = d fs ii" proof - define bnd :: rat where bnd: "bnd = 4 ^ (m - 1 - Suc j) * of_nat (N ^ (m - 1) * m)" define M where "M = map (\i. map (\ fs i) [0.. i" "j < m" and jstrict: "j < i" and add: "set fs \ carrier_vec n" "i < length fs" "j < length fs" "i \ j" and len: "length fs = m" and indep: "lin_indep fs" using inv j i by auto let ?R = rat_of_int let ?RV = "map_vec ?R" from inv i j have Fij: "fs ! i \ carrier_vec n" "fs ! j \ carrier_vec n" by auto let ?x = "fs ! i - c \\<^sub>v fs ! j" let ?g = "gso fs" let ?g' = "gso fs'" let ?mu = "\ fs" let ?mu' = "\ fs'" from inv j i have Fi:"\ i. i < length (RAT fs) \ (RAT fs) ! i \ carrier_vec n" and gs_carr: "?g j \ carrier_vec n" "?g i \ carrier_vec n" "\ i. i < j \ ?g i \ carrier_vec n" "\ j. j < i \ ?g j \ carrier_vec n" and len': "length (RAT fs) = m" and add':"set (map ?RV fs) \ carrier_vec n" by auto have RAT_F1: "RAT fs' = (RAT fs)[i := (RAT fs) ! i - ?R c \\<^sub>v (RAT fs) ! j]" unfolding fs' proof (rule nth_equalityI[rule_format], goal_cases) case (2 k) show ?case proof (cases "k = i") case False thus ?thesis using 2 by auto next case True hence "?thesis = (?RV (fs ! i - c \\<^sub>v fs ! j) = ?RV (fs ! i) - ?R c \\<^sub>v ?RV (fs ! j))" using 2 add by auto also have "\" by (rule eq_vecI, insert Fij, auto) finally show ?thesis by simp qed qed auto hence RAT_F1_i:"RAT fs' ! i = (RAT fs) ! i - ?R c \\<^sub>v (RAT fs) ! j" (is "_ = _ - ?mui") using i len by auto have uminus: "fs ! i - c \\<^sub>v fs ! j = fs ! i + -c \\<^sub>v fs ! j" by (subst minus_add_uminus_vec, insert Fij, auto) have "lattice_of fs' = lattice_of fs" unfolding fs' uminus by (rule lattice_of_add[OF add, of _ "- c"], auto) with inv have lattice: "lattice_of fs' = L" by auto from add len have "k < length fs \ \ k \ i \ fs' ! k \ carrier_vec n" for k unfolding fs' by (metis (no_types, lifting) nth_list_update nth_mem subset_eq carrier_dim_vec index_minus_vec(2) index_smult_vec(2)) hence "k < length fs \ fs' ! k \ carrier_vec n" for k unfolding fs' using add len by (cases "k \ i",auto) with len have F1: "set fs' \ carrier_vec n" "length fs' = m" unfolding fs' by (auto simp: set_conv_nth) hence F1': "length (RAT fs') = m" "SRAT fs' \ Rn" by auto from indep have dist: "distinct (RAT fs)" by (auto simp: gs.lin_indpt_list_def) have Fij': "(RAT fs) ! i \ Rn" "(RAT fs) ! j \ Rn" using add'[unfolded set_conv_nth] i \j < m\ len by auto have uminus': "(RAT fs) ! i - ?R c \\<^sub>v (RAT fs) ! j = (RAT fs) ! i + - ?R c \\<^sub>v (RAT fs) ! j" by (subst minus_add_uminus_vec[where n = n], insert Fij', auto) have span_F_F1: "gs.span (SRAT fs) = gs.span (SRAT fs')" unfolding RAT_F1 uminus' by (rule gs.add_vec_span, insert len add, auto) have **: "?RV (fs ! i) + - ?R c \\<^sub>v (RAT fs) ! j = ?RV (fs ! i - c \\<^sub>v fs ! j)" by (rule eq_vecI, insert Fij len i j, auto) from i j len have "j < length (RAT fs)" "i < length (RAT fs)" "i \ j" by auto from gs.lin_indpt_list_add_vec[OF this indep, of "- of_int c"] have "gs.lin_indpt_list ((RAT fs) [i := (RAT fs) ! i + - ?R c \\<^sub>v (RAT fs) ! j])" (is "gs.lin_indpt_list ?F1") . also have "?F1 = RAT fs'" unfolding fs' using i len Fij' ** by (auto simp: map_update) finally have indep_F1: "lin_indep fs'" . have conn1: "set (RAT fs) \ carrier_vec n" "length (RAT fs) = m" "distinct (RAT fs)" "gs.lin_indpt (set (RAT fs))" using inv unfolding gs.lin_indpt_list_def by auto have conn2: "set (RAT fs') \ carrier_vec n" "length (RAT fs') = m" "distinct (RAT fs')" "gs.lin_indpt (set (RAT fs'))" using indep_F1 F1' unfolding gs.lin_indpt_list_def by auto interpret gs1: gram_schmidt_fs_lin_indpt n "RAT fs" by (standard) (use inv gs.lin_indpt_list_def in auto) interpret gs2: gram_schmidt_fs_lin_indpt n "RAT fs'" by (standard) (use indep_F1 F1' gs.lin_indpt_list_def in auto) let ?G = "map ?g [0 ..< m]" let ?G' = "map ?g' [0 ..< m]" from gs1.span_gso gs2.span_gso gs1.gso_carrier gs2.gso_carrier conn1 conn2 span_F_F1 len have span_G_G1: "gs.span (set ?G) = gs.span (set ?G')" and lenG: "length ?G = m" and Gi: "i < length ?G \ ?G ! i \ Rn" and G1i: "i < length ?G' \ ?G' ! i \ Rn" for i by auto have eq: "x \ i \ RAT fs' ! x = (RAT fs) ! x" for x unfolding RAT_F1 by auto hence eq_part: "x < i \ ?g' x = ?g x" for x by (intro gs.gso_cong, insert len, auto) have G: "i < m \ (RAT fs) ! i \ Rn" "i < m \ fs ! i \ carrier_vec n" for i by(insert add len', auto) note carr1[intro] = this[OF i] this[OF ji(2)] have "x < m \ ?g x \ Rn" "x < m \ ?g' x \ Rn" "x < m \ dim_vec (gso fs x) = n" "x < m \ dim_vec (gso fs' x) = n" for x using inv G1i by (auto simp:o_def Gi G1i) hence carr2[intro!]:"?g i \ Rn" "?g' i \ Rn" "?g ` {0.. Rn" "?g ` {0.. Rn" using i by auto have F1_RV: "?RV (fs' ! i) = RAT fs' ! i" using i F1 by auto have F_RV: "?RV (fs ! i) = (RAT fs) ! i" using i len by auto from eq_part have span_G1_G: "gs.span (?g' ` {0.. ?rs" using gs2.oc_projection_exist[of i] conn2 i unfolding span_G1_G by auto from \j < i\ have Gj_mem: "(RAT fs) ! j \ (\ x. ((RAT fs) ! x)) ` {0 ..< i}" by auto have id1: "set (take i (RAT fs)) = (\x. ?RV (fs ! x)) ` {0..i < m\ len by (subst nth_image[symmetric], force+) have "(RAT fs) ! j \ ?rs \ (RAT fs) ! j \ gs.span ((\x. ?RV (fs ! x)) ` {0..i < m\ id1 inv by auto also have "(\x. ?RV (fs ! x)) ` {0..x. ((RAT fs) ! x)) ` {0..i < m\ len by force also have "(RAT fs) ! j \ gs.span \" by (rule gs.span_mem[OF _ Gj_mem], insert \i < m\ G, auto) finally have "(RAT fs) ! j \ ?rs" . hence in2:"?mui \ ?rs" apply(intro gs.prod_in_span) by force+ have ineq:"((RAT fs) ! i - ?g' i) + ?mui - ?mui = ((RAT fs) ! i - ?g' i)" using carr1 carr2 by (intro eq_vecI, auto) have cong': "A = B \ A \ C \ B \ C" for A B :: "'a vec" and C by auto have *: "?g ` {0.. Rn" by auto have in_span: "(RAT fs) ! i - ?g' i \ ?rs" by (rule cong'[OF eq_vecI gs.span_add1[OF * in1 in2,unfolded ineq]], insert carr1 carr2, auto) { fix x assume x:"x < i" hence "x < m" "i \ x" using i by auto from gs2.orthogonal this inv assms have "?g' i \ ?g' x = 0" by auto } hence G1_G: "?g' i = ?g i" by (intro gs1.oc_projection_unique) (use inv i eq_part in_span in auto) show eq_fs:"x < m \ ?g' x = ?g x" for x proof(induct x rule:nat_less_induct[rule_format]) case (1 x) hence ind: "m < x \ ?g' m = ?g m" for m by auto { assume "x > i" hence ?case unfolding gs2.gso.simps[of x] gs1.gso.simps[of x] unfolding gs1.\.simps gs2.\.simps using ind eq by (auto intro: cong[OF _ cong[OF refl[of "gs.sumlist"]]]) } note eq_rest = this show ?case by (rule linorder_class.linorder_cases[of x i],insert G1_G eq_part eq_rest,auto) qed hence Hs:"?G' = ?G" by (auto simp:o_def) have red: "weakly_reduced fs i \ weakly_reduced fs' i" using eq_fs \i < m\ unfolding gram_schmidt_fs.weakly_reduced_def by simp let ?Mi = "M ! i ! j" have Gjn: "dim_vec (fs ! j) = n" using Fij(2) carrier_vecD by blast define E where "E = addrow_mat m (- ?R c) i j" define M' where "M' = gs1.M m" define N' where "N' = gs2.M m" have E: "E \ carrier_mat m m" unfolding E_def by simp have M: "M' \ carrier_mat m m" unfolding gs1.M_def M'_def by auto have N: "N' \ carrier_mat m m" unfolding gs2.M_def N'_def by auto let ?mat = "mat_of_rows n" let ?GsM = "?mat ?G" have Gs: "?GsM \ carrier_mat m n" by auto hence GsT: "?GsM\<^sup>T \ carrier_mat n m" by auto have Gnn: "?mat (RAT fs) \ carrier_mat m n" unfolding mat_of_rows_def using len by auto have "?mat (RAT fs') = addrow (- ?R c) i j (?mat (RAT fs))" unfolding RAT_F1 by (rule eq_matI, insert Gjn ji(2), auto simp: len mat_of_rows_def) also have "\ = E * ?mat (RAT fs)" unfolding E_def by (rule addrow_mat, insert j i, auto simp: mat_of_rows_def len) finally have HEG: "?mat (RAT fs') = E * ?mat (RAT fs)" . (* lemma 16.12(i), part 1 *) have "(E * M') * ?mat ?G = E * (M' * ?mat ?G)" by (rule assoc_mult_mat[OF E M Gs]) also have "M' * ?GsM = ?mat (RAT fs)" using gs1.matrix_equality conn1 M'_def by simp also have "E * \ = ?mat (RAT fs')" unfolding HEG .. also have "\ = N' * ?mat ?G'" using gs2.matrix_equality conn2 unfolding N'_def by simp also have "?mat ?G' = ?GsM" unfolding Hs .. finally have "(E * M') * ?GsM = N' * ?GsM" . from arg_cong[OF this, of "\ x. x * ?GsM\<^sup>T"] E M N have EMN: "(E * M') * (?GsM * ?GsM\<^sup>T) = N' * (?GsM * ?GsM\<^sup>T)" by (subst (1 2) assoc_mult_mat[OF _ Gs GsT, of _ m, symmetric], auto) have "det (?GsM * ?GsM\<^sup>T) = gs.Gramian_determinant ?G m" unfolding gs.Gramian_determinant_def by (subst gs.Gramian_matrix_alt_def, auto simp: Let_def) also have "\ > 0" proof - have 1: "gs.lin_indpt_list ?G" using conn1 gs1.orthogonal_gso gs1.gso_carrier by (intro gs.orthogonal_imp_lin_indpt_list) (auto) interpret G: gram_schmidt_fs_lin_indpt n ?G by (standard) (use 1 gs.lin_indpt_list_def in auto) show ?thesis by (intro G.Gramian_determinant) auto qed finally have "det (?GsM * ?GsM\<^sup>T) \ 0" by simp from vec_space.det_nonzero_congruence[OF EMN this _ _ N] Gs E M have EMN: "E * M' = N'" by auto (* lemma 16.12(i), part 2 *) { fix i' j' assume ij: "i' < m" "j' < m" and choice: "i' \ i \ j < j'" have "?mu' i' j' = N' $$ (i',j')" using ij F1 unfolding N'_def gs2.M_def by auto also have "\ = addrow (- ?R c) i j M' $$ (i',j')" unfolding EMN[symmetric] E_def by (subst addrow_mat[OF M], insert ji, auto) also have "\ = (if i = i' then - ?R c * M' $$ (j, j') + M' $$ (i', j') else M' $$ (i', j'))" by (rule index_mat_addrow, insert ij M, auto) also have "\ = M' $$ (i', j')" proof (cases "i = i'") case True with choice have jj: "j < j'" by auto have "M' $$ (j, j') = ?mu j j'" using ij ji len unfolding M'_def gs1.M_def by auto also have "\ = 0" unfolding gs1.\.simps using jj by auto finally show ?thesis using True by auto qed auto also have "\ = ?mu i' j'" using ij len unfolding M'_def gs1.M_def by auto also note calculation } note mu_no_change = this { fix j' assume jj': "j' \ j" with j i have j': "j' < m" by auto have "?mu' i j' = N' $$ (i,j')" using jj' j i F1 unfolding N'_def gs2.M_def by auto also have "\ = addrow (- ?R c) i j M' $$ (i,j')" unfolding EMN[symmetric] E_def by (subst addrow_mat[OF M], insert ji, auto) also have "\ = - ?R c * M' $$ (j, j') + M' $$ (i, j')" by (rule index_mat_addrow, insert j' i M, auto) also have "\ = M' $$ (i, j') - ?R c * M' $$ (j, j')" by simp also have "M' $$ (i, j') = ?mu i j'" using i j' len unfolding M'_def gs1.M_def by auto also have "M' $$ (j, j') = ?mu j j'" using i j j' len unfolding M'_def gs1.M_def by auto finally have "?mu' i j' = ?mu i j' - ?R c * ?mu j j'" by auto } note mu_change = this show mu_update: "i' < m \ j' < m \ ?mu' i' j' = (if i' = i \ j' \ j then ?mu i j' - ?R c * ?mu j j' else ?mu i' j')" for i' j' using mu_change[of j'] mu_no_change[of i' j'] by auto { assume "LLL_invariant True i fs" from LLL_invD[OF this] have "weakly_reduced fs i" and sred: "reduced fs i" by auto from red[OF this(1)] have red: "weakly_reduced fs' i" . have sred: "reduced fs' i" unfolding gram_schmidt_fs.reduced_def proof (intro conjI[OF red] impI allI, goal_cases) case (1 i' j) with mu_no_change[of i' j] sred[unfolded gram_schmidt_fs.reduced_def, THEN conjunct2, rule_format, of i' j] i show ?case by auto qed show "LLL_invariant True i fs'" by (intro LLL_invI[OF F1 lattice \i \ m\ indep_F1 sred], auto) } show Linv': "LLL_invariant_weak fs'" by (intro LLL_inv_wI[OF F1 lattice indep_F1]) have mudiff:"?mu i j - of_int c = ?mu' i j" by (subst mu_change, auto simp: gs1.\.simps) have lin_indpt_list_fs: "gs.lin_indpt_list (RAT fs')" unfolding gs.lin_indpt_list_def using conn2 by auto { assume c: "c = round (\ fs i j)" have small: "abs (?mu i j - of_int c) \ inverse 2" unfolding j c using of_int_round_abs_le by (auto simp add: abs_minus_commute) from this[unfolded mudiff] show mu'_2: "abs (?mu' i j) \ 1 / 2" by simp assume mu_small: "\_small_row i fs (Suc j)" show "\_small_row i fs' j" unfolding \_small_row_def proof (intro allI, goal_cases) case (1 j') show ?case using mu'_2 mu_small[unfolded \_small_row_def, rule_format, of j'] by (cases "j' > j", insert mu_update[of i j'] i, auto) qed } { fix i assume i: "i \ m" have "rat_of_int (d fs' i) = of_int (d fs i)" unfolding d_def Gramian_determinant(1)[OF Linv i] Gramian_determinant(1)[OF Linv' i] by (rule prod.cong[OF refl], subst eq_fs, insert i, auto) thus "d fs' i = d fs i" by simp } note d = this have D: "D fs' = D fs" unfolding D_def by (rule arg_cong[of _ _ nat], rule prod.cong[OF refl], auto simp: d) show "LLL_measure i fs' = LLL_measure i fs" unfolding LLL_measure_def logD_def D .. qed text \Addition step which can be skipped since $\mu$-value is already small\ lemma basis_reduction_add_row_main_0: assumes Linv: "LLL_invariant_weak fs" and i: "i < m" and j: "j < i" and 0: "round (\ fs i j) = 0" and mu_small: "\_small_row i fs (Suc j)" shows "\_small_row i fs j" (is ?g1) proof - note inv = LLL_inv_wD[OF Linv] from inv(5)[OF i] inv(5)[of j] i j have id: "fs[i := fs ! i - 0 \\<^sub>v fs ! j] = fs" by (intro nth_equalityI, insert inv i, auto) show ?g1 using basis_reduction_add_row_main[OF Linv i j _, of fs] 0 id mu_small by auto qed lemma \_small_row_refl: "\_small_row i fs i" unfolding \_small_row_def by auto lemma basis_reduction_add_row_done: assumes Linv: "LLL_invariant True i fs" and i: "i < m" and mu_small: "\_small_row i fs 0" shows "LLL_invariant False i fs" proof - note inv = LLL_invD[OF Linv] from mu_small have mu_small: "\_small fs i" unfolding \_small_row_def \_small_def by auto show ?thesis using i mu_small by (intro LLL_invI[OF inv(3,6,7,9,1,10)], auto) qed (* lemma 16.16 (ii), one case *) lemma d_swap_unchanged: assumes len: "length F1 = m" and i0: "i \ 0" and i: "i < m" and ki: "k \ i" and km: "k \ m" and swap: "F2 = F1[i := F1 ! (i - 1), i - 1 := F1 ! i]" shows "d F1 k = d F2 k" proof - let ?F1_M = "mat k n (\(i, y). F1 ! i $ y)" let ?F2_M = "mat k n (\(i, y). F2 ! i $ y)" have "\ P. P \ carrier_mat k k \ det P \ {-1, 1} \ ?F2_M = P * ?F1_M" proof cases assume ki: "k < i" hence H: "?F2_M = ?F1_M" unfolding swap by (intro eq_matI, auto) let ?P = "1\<^sub>m k" have "?P \ carrier_mat k k" "det ?P \ {-1, 1}" "?F2_M = ?P * ?F1_M" unfolding H by auto thus ?thesis by blast next assume "\ k < i" with ki have ki: "k > i" by auto let ?P = "swaprows_mat k i (i - 1)" from i0 ki have neq: "i \ i - 1" and kmi: "i - 1 < k" by auto have *: "?P \ carrier_mat k k" "det ?P \ {-1, 1}" using det_swaprows_mat[OF ki kmi neq] ki by auto from i len have iH: "i < length F1" "i - 1 < length F1" by auto have "?P * ?F1_M = swaprows i (i - 1) ?F1_M" by (subst swaprows_mat[OF _ ki kmi], auto) also have "\ = ?F2_M" unfolding swap by (intro eq_matI, rename_tac ii jj, case_tac "ii = i", (insert iH, simp add: nth_list_update)[1], case_tac "ii = i - 1", insert iH neq ki, auto simp: nth_list_update) finally show ?thesis using * by metis qed then obtain P where P: "P \ carrier_mat k k" and detP: "det P \ {-1, 1}" and H': "?F2_M = P * ?F1_M" by auto have "d F2 k = det (gs.Gramian_matrix F2 k)" unfolding d_def gs.Gramian_determinant_def by simp also have "\ = det (?F2_M * ?F2_M\<^sup>T)" unfolding gs.Gramian_matrix_def Let_def by simp also have "?F2_M * ?F2_M\<^sup>T = ?F2_M * (?F1_M\<^sup>T * P\<^sup>T)" unfolding H' by (subst transpose_mult[OF P], auto) also have "\ = P * (?F1_M * (?F1_M\<^sup>T * P\<^sup>T))" unfolding H' by (subst assoc_mult_mat[OF P], auto) also have "det \ = det P * det (?F1_M * (?F1_M\<^sup>T * P\<^sup>T))" by (rule det_mult[OF P], insert P, auto) also have "?F1_M * (?F1_M\<^sup>T * P\<^sup>T) = (?F1_M * ?F1_M\<^sup>T) * P\<^sup>T" by (subst assoc_mult_mat, insert P, auto) also have "det \ = det (?F1_M * ?F1_M\<^sup>T) * det P" by (subst det_mult, insert P, auto simp: det_transpose) also have "det (?F1_M * ?F1_M\<^sup>T) = det (gs.Gramian_matrix F1 k)" unfolding gs.Gramian_matrix_def Let_def by simp also have "\ = d F1 k" unfolding d_def gs.Gramian_determinant_def by simp finally have "d F2 k = (det P * det P) * d F1 k" by simp also have "det P * det P = 1" using detP by auto finally show "d F1 k = d F2 k" by simp qed definition base where "base = real_of_rat ((4 * \) / (4 + \))" definition g_bound :: "int vec list \ bool" where "g_bound fs = (\ i < m. sq_norm (gso fs i) \ of_nat N)" end locale LLL_with_assms = LLL + assumes \: "\ \ 4/3" and lin_dep: "lin_indep fs_init" and len: "length fs_init = m" begin lemma \0: "\ > 0" "\ \ 0" using \ by auto lemma fs_init: "set fs_init \ carrier_vec n" using lin_dep[unfolded gs.lin_indpt_list_def] by auto lemma reduction: "0 < reduction" "reduction \ 1" "\ > 4/3 \ reduction < 1" "\ = 4/3 \ reduction = 1" using \ unfolding reduction_def by auto lemma base: "\ > 4/3 \ base > 1" using reduction(1,3) unfolding reduction_def base_def by auto lemma basis_reduction_swap_main: assumes Linvw: "LLL_invariant_weak fs" and small: "LLL_invariant False i fs \ abs (\ fs i (i - 1)) \ 1/2" and i: "i < m" and i0: "i \ 0" and norm_ineq: "sq_norm (gso fs (i - 1)) > \ * sq_norm (gso fs i)" and fs'_def: "fs' = fs[i := fs ! (i - 1), i - 1 := fs ! i]" shows "LLL_invariant_weak fs'" and "LLL_invariant False i fs \ LLL_invariant False (i - 1) fs'" and "LLL_measure i fs > LLL_measure (i - 1) fs'" (* new values of gso *) and "\ k. k < m \ gso fs' k = (if k = i - 1 then gso fs i + \ fs i (i - 1) \\<^sub>v gso fs (i - 1) else if k = i then gso fs (i - 1) - (RAT fs ! (i - 1) \ gso fs' (i - 1) / sq_norm (gso fs' (i - 1))) \\<^sub>v gso fs' (i - 1) else gso fs k)" (is "\ k. _ \ _ = ?newg k") (* new values of norms of gso *) and "\ k. k < m \ sq_norm (gso fs' k) = (if k = i - 1 then sq_norm (gso fs i) + (\ fs i (i - 1) * \ fs i (i - 1)) * sq_norm (gso fs (i - 1)) else if k = i then sq_norm (gso fs i) * sq_norm (gso fs (i - 1)) / sq_norm (gso fs' (i - 1)) else sq_norm (gso fs k))" (is "\ k. _ \ _ = ?new_norm k") (* new values of \-values *) and "\ ii j. ii < m \ j < ii \ \ fs' ii j = ( if ii = i - 1 then \ fs i j else if ii = i then if j = i - 1 then \ fs i (i - 1) * sq_norm (gso fs (i - 1)) / sq_norm (gso fs' (i - 1)) else \ fs (i - 1) j else if ii > i \ j = i then \ fs ii (i - 1) - \ fs i (i - 1) * \ fs ii i else if ii > i \ j = i - 1 then \ fs ii (i - 1) * \ fs' i (i - 1) + \ fs ii i * sq_norm (gso fs i) / sq_norm (gso fs' (i - 1)) else \ fs ii j)" (is "\ ii j. _ \ _ \ _ = ?new_mu ii j") (* new d-values *) and "\ ii. ii \ m \ of_int (d fs' ii) = (if ii = i then sq_norm (gso fs' (i - 1)) / sq_norm (gso fs (i - 1)) * of_int (d fs i) else of_int (d fs ii))" proof - note inv = LLL_inv_wD[OF Linvw] interpret fs: fs_int' n m fs_init fs by (standard) (use Linvw in auto) let ?mu1 = "\ fs" let ?mu2 = "\ fs'" let ?g1 = "gso fs" let ?g2 = "gso fs'" have m12: "\?mu1 i (i - 1)\ \ inverse 2" using small proof assume "LLL_invariant False i fs" from LLL_invD(11)[OF this] i0 show ?thesis unfolding \_small_def by auto qed auto note d = d_def note Gd = Gramian_determinant(1) note Gd12 = Gd[OF Linvw] let ?x = "?g1 (i - 1)" let ?y = "?g1 i" let ?cond = "\ * sq_norm ?y < sq_norm ?x" from inv have len: "length fs = m" and HC: "set fs \ carrier_vec n" and L: "lattice_of fs = L" using i by auto from i0 inv i have swap: "set fs \ carrier_vec n" "i < length fs" "i - 1 < length fs" "i \ i - 1" unfolding Let_def by auto have RAT_fs': "RAT fs' = (RAT fs)[i := (RAT fs) ! (i - 1), i - 1 := (RAT fs) ! i]" unfolding fs'_def using swap by (intro nth_equalityI, auto simp: nth_list_update) have span': "gs.span (SRAT fs) = gs.span (SRAT fs')" unfolding fs'_def by (rule arg_cong[of _ _ gs.span], insert swap, auto) have lfs': "lattice_of fs' = lattice_of fs" unfolding fs'_def by (rule lattice_of_swap[OF swap refl]) with inv have lattice: "lattice_of fs' = L" by auto have len': "length fs' = m" using inv unfolding fs'_def by auto have fs': "set fs' \ carrier_vec n" using swap unfolding fs'_def set_conv_nth by (auto, rename_tac k, case_tac "k = i", force, case_tac "k = i - 1", auto) let ?rv = "map_vec rat_of_int" from inv(1) have indepH: "lin_indep fs" . from i i0 len have "i < length (RAT fs)" "i - 1 < length (RAT fs)" by auto with distinct_swap[OF this] len have "distinct (RAT fs') = distinct (RAT fs)" unfolding RAT_fs' by (auto simp: map_update) with len' fs' span' indepH have indepH': "lin_indep fs'" unfolding fs'_def using i i0 by (auto simp: gs.lin_indpt_list_def) have lenR': "length (RAT fs') = m" using len' by auto have conn1: "set (RAT fs) \ carrier_vec n" "length (RAT fs) = m" "distinct (RAT fs)" "gs.lin_indpt (set (RAT fs))" using inv unfolding gs.lin_indpt_list_def by auto have conn2: "set (RAT fs') \ carrier_vec n" "length (RAT fs') = m" "distinct (RAT fs')" "gs.lin_indpt (set (RAT fs'))" using indepH' lenR' unfolding gs.lin_indpt_list_def by auto interpret gs2: gram_schmidt_fs_lin_indpt n "RAT fs'" by (standard) (use indepH' lenR' gs.lin_indpt_list_def in auto) have fs'_fs: "k < i - 1 \ fs' ! k = fs ! k" for k unfolding fs'_def by auto { fix k assume ki: "k < i - 1" with i have kn: "k < m" by simp have "?g2 k = ?g1 k" by (rule gs.gso_cong, insert ki kn len, auto simp: fs'_def) } note G2_G = this have take_eq: "take (Suc i - 1 - 1) fs' = take (Suc i - 1 - 1) fs" by (intro nth_equalityI, insert len len' i swap(2-), auto intro!: fs'_fs) have i1n: "i - 1 < m" using i by auto let ?R = rat_of_int let ?RV = "map_vec ?R" let ?f1 = "\ i. RAT fs ! i" let ?f2 = "\ i. RAT fs' ! i" let ?n1 = "\ i. sq_norm (?g1 i)" let ?n2 = "\ i. sq_norm (?g2 i)" have heq:"fs ! (i - 1) = fs' ! i" "take (i-1) fs = take (i-1) fs'" "?f2 (i - 1) = ?f1 i" "?f2 i = ?f1 (i - 1)" unfolding fs'_def using i len i0 by auto have norm_pos2: "j < m \ ?n2 j > 0" for j using gs2.sq_norm_pos len' by simp have norm_pos1: "j < m \ ?n1 j > 0" for j using fs.gs.sq_norm_pos inv by simp have norm_zero2: "j < m \ ?n2 j \ 0" for j using norm_pos2[of j] by linarith have norm_zero1: "j < m \ ?n1 j \ 0" for j using norm_pos1[of j] by linarith have gs: "\ j. j < m \ ?g1 j \ Rn" using inv by blast have gs2: "\ j. j < m \ ?g2 j \ Rn" using fs.gs.gso_carrier conn2 by auto have g: "\ j. j < m \ ?f1 j \ Rn" using inv by auto have g2: "\ j. j < m \ ?f2 j \ Rn" using gs2.f_carrier conn2 by blast let ?fs1 = "?f1 ` {0..< (i - 1)}" have G: "?fs1 \ Rn" using g i by auto let ?gs1 = "?g1 ` {0..< (i - 1)}" have G': "?gs1 \ Rn" using gs i by auto let ?S = "gs.span ?fs1" let ?S' = "gs.span ?gs1" have S'S: "?S' = ?S" by (rule fs.gs.partial_span', insert conn1 i, auto) have "gs.is_oc_projection (?g2 (i - 1)) (gs.span (?g2 ` {0..< (i - 1)})) (?f2 (i - 1))" using i len' by (intro gs2.gso_oc_projection_span(2)) auto also have "?f2 (i - 1) = ?f1 i" unfolding fs'_def using len i by auto also have "gs.span (?g2 ` {0 ..< (i - 1)}) = gs.span (?f2 ` {0 ..< (i - 1)})" using i len' by (intro gs2.partial_span') auto also have "?f2 ` {0 ..< (i - 1)} = ?fs1" by (rule image_cong[OF refl], insert len i, auto simp: fs'_def) finally have claim1: "gs.is_oc_projection (?g2 (i - 1)) ?S (?f1 i)" . have list_id: "[0..j. ?mu1 i j \\<^sub>v ?g1 j) [0 ..< i]) + ?g1 i" (is "_ = ?sum + _") apply(subst fs.gs.fi_is_sum_of_mu_gso, insert len i, force) unfolding map_append list_id by (subst gs.M.sumlist_snoc, insert i gs conn1, auto simp: fs.gs.\.simps) have f1im1_sum: "?f1 (i - 1) = gs.sumlist (map (\j. ?mu1 (i - 1) j \\<^sub>v ?g1 j) [0...simps) have sum: "?sum \ Rn" by (rule gs.sumlist_carrier, insert gs i, auto) have sum1: "?sum1 \ Rn" by (rule gs.sumlist_carrier, insert gs i, auto) from gs.span_closed[OF G] have S: "?S \ Rn" by auto from gs i have gs': "\ j. j < i - 1 \ ?g1 j \ Rn" and gsi: "?g1 (i - 1) \ Rn" by auto have "[0 ..< i] = [0 ..< Suc (i - 1)]" using i0 by simp also have "\ = [0 ..< i - 1] @ [i - 1]" by simp finally have list: "[0 ..< i] = [0 ..< i - 1] @ [i - 1]" . { (* d does not change for k \ i *) fix k assume kn: "k \ m" and ki: "k \ i" from d_swap_unchanged[OF len i0 i ki kn fs'_def] have "d fs k = d fs' k" by simp } note d = this (* new value of g (i-1) *) have g2_im1: "?g2 (i - 1) = ?g1 i + ?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1)" (is "_ = _ + ?mu_f1") proof (rule gs.is_oc_projection_eq[OF claim1 _ S g[OF i]]) show "gs.is_oc_projection (?g1 i + ?mu_f1) ?S (?f1 i)" unfolding gs.is_oc_projection_def proof (intro conjI allI impI) let ?sum' = "gs.sumlist (map (\j. ?mu1 i j \\<^sub>v ?g1 j) [0 ..< i - 1])" have sum': "?sum' \ Rn" by (rule gs.sumlist_carrier, insert gs i, auto) show inRn: "(?g1 i + ?mu_f1) \ Rn" using gs[OF i] gsi i by auto have carr: "?sum \ Rn" "?g1 i \ Rn" "?mu_f1 \ Rn" "?sum' \ Rn" using sum' sum gs[OF i] gsi i by auto have "?f1 i - (?g1 i + ?mu_f1) = (?sum + ?g1 i) - (?g1 i + ?mu_f1)" unfolding f1i_sum by simp also have "\ = ?sum - ?mu_f1" using carr by auto also have "?sum = gs.sumlist (map (\j. ?mu1 i j \\<^sub>v ?g1 j) [0 ..< i - 1] @ [?mu_f1])" unfolding list by simp also have "\ = ?sum' + ?mu_f1" by (subst gs.sumlist_append, insert gs' gsi, auto) also have "\ - ?mu_f1 = ?sum'" using sum' gsi by auto finally have id: "?f1 i - (?g1 i + ?mu_f1) = ?sum'" . show "?f1 i - (?g1 i + ?mu_f1) \ gs.span ?S" unfolding id gs.span_span[OF G] proof (rule gs.sumlist_in_span[OF G]) fix v assume "v \ set (map (\j. ?mu1 i j \\<^sub>v ?g1 j) [0 ..< i - 1])" then obtain j where j: "j < i - 1" and v: "v = ?mu1 i j \\<^sub>v ?g1 j" by auto show "v \ ?S" unfolding v by (rule gs.smult_in_span[OF G], unfold S'S[symmetric], rule gs.span_mem, insert gs i j, auto) qed fix x assume "x \ ?S" hence x: "x \ ?S'" using S'S by simp show "(?g1 i + ?mu_f1) \ x = 0" proof (rule gs.orthocompl_span[OF _ G' inRn x]) fix x assume "x \ ?gs1" then obtain j where j: "j < i - 1" and x_id: "x = ?g1 j" by auto from j i x_id gs[of j] have x: "x \ Rn" by auto { fix k assume k: "k > j" "k < m" have "?g1 k \ x = 0" unfolding x_id by (rule fs.gs.orthogonal, insert conn1 k, auto) } from this[of i] this[of "i - 1"] j i have main: "?g1 i \ x = 0" "?g1 (i - 1) \ x = 0" by auto have "(?g1 i + ?mu_f1) \ x = ?g1 i \ x + ?mu_f1 \ x" by (rule add_scalar_prod_distrib[OF gs[OF i] _ x], insert gsi, auto) also have "\ = 0" using main by (subst smult_scalar_prod_distrib[OF gsi x], auto) finally show "(?g1 i + ?mu_f1) \ x = 0" . qed qed qed { (* 16.13 (i): for g, only g_i and g_{i-1} can change *) fix k assume kn: "k < m" and ki: "k \ i" "k \ i - 1" have "?g2 k = gs.oc_projection (gs.span (?g2 ` {0..i") case True hence "k < i - 1" using ki by auto then show ?thesis apply(intro image_cong) unfolding fs'_def using len i by auto next case False - have "?f2 ` {0.. transpose i (i - 1)) ` {0.. = ?f1 ` {0.. = gs.span (?g1 ` {0.. = ?g1 k" by (subst fs.gs.gso_oc_projection_span, insert kn conn1, auto) finally have "?g2 k = ?g1 k" . } note g2_g1_identical = this (* calculation of new mu-values *) { (* no change of mu for lines before line i - 1 *) fix jj ii assume ii: "ii < i - 1" have "?mu2 ii jj = ?mu1 ii jj" using ii i len by (subst gs.\_cong[of _ _ "RAT fs" "RAT fs'"], auto simp: fs'_def) } note mu'_mu_small_i = this { (* swap of mu-values in lines i - 1 and i for j < i - 1 *) fix jj assume jj: "jj < i - 1" hence id1: "jj < i - 1 \ True" "jj < i \ True" by auto have id2: "?g2 jj = ?g1 jj" by (subst g2_g1_identical, insert jj i, auto) have "?mu2 i jj = ?mu1 (i - 1) jj" "?mu2 (i - 1) jj = ?mu1 i jj" unfolding gs2.\.simps fs.gs.\.simps id1 id2 if_True using len i i0 by (auto simp: fs'_def) } note mu'_mu_i_im1_j = this have im1: "i - 1 < m" using i by auto (* calculation of new value of g_i *) let ?g2_im1 = "?g2 (i - 1)" have g2_im1_Rn: "?g2_im1 \ Rn" using i conn2 by (auto intro!: fs.gs.gso_carrier) { let ?mu2_f2 = "\ j. - ?mu2 i j \\<^sub>v ?g2 j" let ?sum = "gs.sumlist (map (\j. - ?mu1 (i - 1) j \\<^sub>v ?g1 j) [0 ..< i - 1])" have mhs: "?mu2_f2 (i - 1) \ Rn" using i conn2 by (auto intro!: fs.gs.gso_carrier) have sum': "?sum \ Rn" by (rule gs.sumlist_carrier, insert gs i, auto) have gim1: "?f1 (i - 1) \ Rn" using g i by auto have "?g2 i = ?f2 i + gs.sumlist (map ?mu2_f2 [0 ..< i-1] @ [?mu2_f2 (i-1)])" unfolding gs2.gso.simps[of i] list by simp also have "?f2 i = ?f1 (i - 1)" unfolding fs'_def using len i i0 by auto also have "map ?mu2_f2 [0 ..< i-1] = map (\j. - ?mu1 (i - 1) j \\<^sub>v ?g1 j) [0 ..< i - 1]" by (rule map_cong[OF refl], subst g2_g1_identical, insert i, auto simp: mu'_mu_i_im1_j) also have "gs.sumlist (\ @ [?mu2_f2 (i - 1)]) = ?sum + ?mu2_f2 (i - 1)" by (subst gs.sumlist_append, insert gs i mhs, auto) also have "?f1 (i - 1) + \ = (?f1 (i - 1) + ?sum) + ?mu2_f2 (i - 1)" using gim1 sum' mhs by auto also have "?f1 (i - 1) + ?sum = ?g1 (i - 1)" unfolding fs.gs.gso.simps[of "i - 1"] by simp also have "?mu2_f2 (i - 1) = - (?f2 i \ ?g2_im1 / sq_norm ?g2_im1) \\<^sub>v ?g2_im1" unfolding gs2.\.simps using i0 by simp also have "\ = - ((?f2 i \ ?g2_im1 / sq_norm ?g2_im1) \\<^sub>v ?g2_im1)" by auto also have "?g1 (i - 1) + \ = ?g1 (i - 1) - ((?f2 i \ ?g2_im1 / sq_norm ?g2_im1) \\<^sub>v ?g2_im1)" by (rule sym, rule minus_add_uminus_vec[of _ n], insert gsi g2_im1_Rn, auto) also have "?f2 i = ?f1 (i - 1)" by fact finally have "?g2 i = ?g1 (i - 1) - (?f1 (i - 1) \ ?g2 (i - 1) / sq_norm (?g2 (i - 1))) \\<^sub>v ?g2 (i - 1)" . } note g2_i = this let ?n1 = "\ i. sq_norm (?g1 i)" let ?n2 = "\ i. sq_norm (?g2 i)" (* calculation of new norms *) { (* norm of g (i - 1) *) have "?n2 (i - 1) = sq_norm (?g1 i + ?mu_f1)" unfolding g2_im1 by simp also have "\ = (?g1 i + ?mu_f1) \ (?g1 i + ?mu_f1)" by (simp add: sq_norm_vec_as_cscalar_prod) also have "\ = (?g1 i + ?mu_f1) \ ?g1 i + (?g1 i + ?mu_f1) \ ?mu_f1" by (rule scalar_prod_add_distrib, insert gs i, auto) also have "(?g1 i + ?mu_f1) \ ?g1 i = ?g1 i \ ?g1 i + ?mu_f1 \ ?g1 i" by (rule add_scalar_prod_distrib, insert gs i, auto) also have "(?g1 i + ?mu_f1) \ ?mu_f1 = ?g1 i \ ?mu_f1 + ?mu_f1 \ ?mu_f1" by (rule add_scalar_prod_distrib, insert gs i, auto) also have "?mu_f1 \ ?g1 i = ?g1 i \ ?mu_f1" by (rule comm_scalar_prod, insert gs i, auto) also have "?g1 i \ ?g1 i = sq_norm (?g1 i)" by (simp add: sq_norm_vec_as_cscalar_prod) also have "?g1 i \ ?mu_f1 = ?mu1 i (i - 1) * (?g1 i \ ?g1 (i - 1))" by (rule scalar_prod_smult_right, insert gs[OF i] gs[OF \i - 1 < m\], auto) also have "?g1 i \ ?g1 (i - 1) = 0" using orthogonalD[OF fs.gs.orthogonal_gso, of i "i - 1"] i len i0 by (auto simp: o_def) also have "?mu_f1 \ ?mu_f1 = ?mu1 i (i - 1) * (?mu_f1 \ ?g1 (i - 1))" by (rule scalar_prod_smult_right, insert gs[OF i] gs[OF \i - 1 < m\], auto) also have "?mu_f1 \ ?g1 (i - 1) = ?mu1 i (i - 1) * (?g1 (i - 1) \ ?g1 (i - 1))" by (rule scalar_prod_smult_left, insert gs[OF i] gs[OF \i - 1 < m\], auto) also have "?g1 (i - 1) \ ?g1 (i - 1) = sq_norm (?g1 (i - 1))" by (simp add: sq_norm_vec_as_cscalar_prod) finally have "?n2 (i - 1) = ?n1 i + (?mu1 i (i - 1) * ?mu1 i (i - 1)) * ?n1 (i - 1)" by (simp add: ac_simps o_def) } note sq_norm_g2_im1 = this from norm_pos1[OF i] norm_pos1[OF im1] norm_pos2[OF i] norm_pos2[OF im1] have norm0: "?n1 i \ 0" "?n1 (i - 1) \ 0" "?n2 i \ 0" "?n2 (i - 1) \ 0" by auto hence norm0': "?n2 (i - 1) \ 0" using i by auto { (* new norm of g i *) have si: "Suc i \ m" and im1: "i - 1 \ m" using i by auto have det1: "gs.Gramian_determinant (RAT fs) (Suc i) = (\jfs.gs.gso j\\<^sup>2)" using fs.gs.Gramian_determinant si len by auto have det2: "gs.Gramian_determinant (RAT fs') (Suc i) = (\jgs2.gso j\\<^sup>2)" using gs2.Gramian_determinant si len' by auto from norm_zero1[OF less_le_trans[OF _ im1]] have 0: "(\j < i-1. ?n1 j) \ 0" by (subst prod_zero_iff, auto) have "rat_of_int (d fs' (Suc i)) = rat_of_int (d fs (Suc i))" using d_swap_unchanged[OF len i0 i _ si fs'_def] by auto also have "rat_of_int (d fs' (Suc i)) = gs.Gramian_determinant (RAT fs') (Suc i)" unfolding d_def by (subst fs.of_int_Gramian_determinant[symmetric], insert conn2 i g fs', auto simp: set_conv_nth) also have "\ = (\j = (\jj\ ?set. ?n2 j) = ?n2 i * ?n2 (i - 1) * (\j < i-1. ?n2 j)" using i0 by (subst prod.insert; (subst prod.insert)?; auto) also have "(\j\ ?set. ?n1 j) = ?n1 i * ?n1 (i - 1) * (\j < i-1. ?n1 j)" using i0 by (subst prod.insert; (subst prod.insert)?; auto) also have "(\j < i-1. ?n2 j) = (\j < i-1. ?n1 j)" by (rule prod.cong, insert G2_G, auto) finally have "?n2 i = ?n1 i * ?n1 (i - 1) / ?n2 (i - 1)" using 0 norm0' by (auto simp: field_simps) } note sq_norm_g2_i = this (* mu values in rows > i do not change with j \ {i, i - 1} *) { fix ii j assume ii: "ii > i" "ii < m" and ji: "j \ i" "j \ i - 1" { assume j: "j < ii" have "?mu2 ii j = (?f2 ii \ ?g2 j) / sq_norm (?g2 j)" unfolding gs2.\.simps using j by auto also have "?f2 ii = ?f1 ii" using ii len unfolding fs'_def by auto also have "?g2 j = ?g1 j" using g2_g1_identical[of j] j ii ji by auto finally have "?mu2 ii j = ?mu1 ii j" unfolding fs.gs.\.simps using j by auto } hence "?mu2 ii j = ?mu1 ii j" by (cases "j < ii", auto simp: gs2.\.simps fs.gs.\.simps) } note mu_no_change_large_row = this { (* the new value of mu i (i - 1) *) have "?mu2 i (i - 1) = (?f2 i \ ?g2 (i - 1)) / ?n2 (i - 1)" unfolding gs2.\.simps using i0 by auto also have "?f2 i \ ?g2 (i - 1) = ?f1 (i - 1) \ ?g2 (i - 1)" using len i i0 unfolding fs'_def by auto also have "\ = ?f1 (i - 1) \ (?g1 i + ?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1))" unfolding g2_im1 by simp also have "\ = ?f1 (i - 1) \ ?g1 i + ?f1 (i - 1) \ (?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1))" by (rule scalar_prod_add_distrib[of _ n], insert i gs g, auto) also have "?f1 (i - 1) \ ?g1 i = 0" by (subst fs.gs.fi_scalar_prod_gso, insert conn1 im1 i i0, auto simp: fs.gs.\.simps fs.gs.\.simps) also have "?f1 (i - 1) \ (?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1)) = ?mu1 i (i - 1) * (?f1 (i - 1) \ ?g1 (i - 1))" by (rule scalar_prod_smult_distrib, insert gs g i, auto) also have "?f1 (i - 1) \ ?g1 (i - 1) = ?n1 (i - 1)" by (subst fs.gs.fi_scalar_prod_gso, insert conn1 im1, auto simp: fs.gs.\.simps) finally have "?mu2 i (i - 1) = ?mu1 i (i - 1) * ?n1 (i - 1) / ?n2 (i - 1)" by (simp add: sq_norm_vec_as_cscalar_prod) } note mu'_mu_i_im1 = this { (* the new values of mu ii (i - 1) for ii > i *) fix ii assume iii: "ii > i" and ii: "ii < m" hence iii1: "i - 1 < ii" by auto have "?mu2 ii (i - 1) = (?f2 ii \ ?g2 (i - 1)) / ?n2 (i - 1)" unfolding gs2.\.simps using i0 iii1 by auto also have "?f2 ii \ ?g2 (i-1) = ?f1 ii \ ?g2 (i - 1)" using len i i0 iii ii unfolding fs'_def by auto also have "\ = ?f1 ii \ (?g1 i + ?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1))" unfolding g2_im1 by simp also have "\ = ?f1 ii \ ?g1 i + ?f1 ii \ (?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1))" by (rule scalar_prod_add_distrib[of _ n], insert i ii gs g, auto) also have "?f1 ii \ ?g1 i = ?mu1 ii i * ?n1 i" by (rule fs.gs.fi_scalar_prod_gso, insert conn1 ii i, auto) also have "?f1 ii \ (?mu1 i (i - 1) \\<^sub>v ?g1 (i - 1)) = ?mu1 i (i - 1) * (?f1 ii \ ?g1 (i - 1))" by (rule scalar_prod_smult_distrib, insert gs g i ii, auto) also have "?f1 ii \ ?g1 (i - 1) = ?mu1 ii (i - 1) * ?n1 (i - 1)" by (rule fs.gs.fi_scalar_prod_gso, insert conn1 ii im1, auto) finally have "?mu2 ii (i - 1) = ?mu1 ii (i - 1) * ?mu2 i (i - 1) + ?mu1 ii i * ?n1 i / ?n2 (i - 1)" unfolding mu'_mu_i_im1 using norm0 by (auto simp: field_simps) } note mu'_mu_large_row_im1 = this { (* the new values of mu ii i for ii > i *) fix ii assume iii: "ii > i" and ii: "ii < m" have "?mu2 ii i = (?f2 ii \ ?g2 i) / ?n2 i" unfolding gs2.\.simps using i0 iii by auto also have "?f2 ii \ ?g2 i = ?f1 ii \ ?g2 i" using len i i0 iii ii unfolding fs'_def by auto also have "\ = ?f1 ii \ (?g1 (i - 1) - (?f1 (i - 1) \ ?g2 (i - 1) / ?n2 (i - 1)) \\<^sub>v ?g2 (i - 1))" unfolding g2_i by simp also have "?f1 (i - 1) = ?f2 i" using i i0 len unfolding fs'_def by auto also have "?f2 i \ ?g2 (i - 1) / ?n2 (i - 1) = ?mu2 i (i - 1)" unfolding gs2.\.simps using i i0 by auto also have "?f1 ii \ (?g1 (i - 1) - ?mu2 i (i - 1) \\<^sub>v ?g2 (i - 1)) = ?f1 ii \ ?g1 (i - 1) - ?f1 ii \ (?mu2 i (i - 1) \\<^sub>v ?g2 (i - 1))" by (rule scalar_prod_minus_distrib[OF g gs], insert gs2 ii i, auto) also have "?f1 ii \ ?g1 (i - 1) = ?mu1 ii (i - 1) * ?n1 (i - 1)" by (rule fs.gs.fi_scalar_prod_gso, insert conn1 ii im1, auto) also have "?f1 ii \ (?mu2 i (i - 1) \\<^sub>v ?g2 (i - 1)) = ?mu2 i (i - 1) * (?f1 ii \ ?g2 (i - 1))" by (rule scalar_prod_smult_distrib, insert gs gs2 g i ii, auto) also have "?f1 ii \ ?g2 (i - 1) = (?f1 ii \ ?g2 (i - 1) / ?n2 (i - 1)) * ?n2 (i - 1)" using norm0 by (auto simp: field_simps) also have "?f1 ii \ ?g2 (i - 1) = ?f2 ii \ ?g2 (i - 1)" using len ii iii unfolding fs'_def by auto also have "\ / ?n2 (i - 1) = ?mu2 ii (i - 1)" unfolding gs2.\.simps using iii by auto finally have "?mu2 ii i = (?mu1 ii (i - 1) * ?n1 (i - 1) - ?mu2 i (i - 1) * ?mu2 ii (i - 1) * ?n2 (i - 1)) / ?n2 i" by simp also have "\ = (?mu1 ii (i - 1) - ?mu1 i (i - 1) * ?mu2 ii (i - 1)) * ?n2 (i - 1) / ?n1 i" unfolding sq_norm_g2_i mu'_mu_i_im1 using norm0 by (auto simp: field_simps) also have "\ = (?mu1 ii (i - 1) * ?n2 (i - 1) - ?mu1 i (i - 1) * ((?mu1 ii i * ?n1 i + ?mu1 i (i - 1) * ?mu1 ii (i - 1) * ?n1 (i - 1)))) / ?n1 i" unfolding mu'_mu_large_row_im1[OF iii ii] mu'_mu_i_im1 using norm0 by (auto simp: field_simps) also have "\ = ?mu1 ii (i - 1) - ?mu1 i (i - 1) * ?mu1 ii i" unfolding sq_norm_g2_im1 using norm0 by (auto simp: field_simps) finally have "?mu2 ii i = ?mu1 ii (i - 1) - ?mu1 i (i - 1) * ?mu1 ii i" . } note mu'_mu_large_row_i = this { fix k assume k: "k < m" show "?g2 k = ?newg k" unfolding g2_i[symmetric] unfolding g2_im1[symmetric] using g2_g1_identical[OF k] by auto show "?n2 k = ?new_norm k" unfolding sq_norm_g2_i[symmetric] unfolding sq_norm_g2_im1[symmetric] using g2_g1_identical[OF k] by auto fix j assume jk: "j < k" hence j: "j < m" using k by auto have "k < i - 1 \ k = i - 1 \ k = i \ k > i" by linarith thus "?mu2 k j = ?new_mu k j" unfolding mu'_mu_i_im1[symmetric] using mu'_mu_large_row_i[OF _ k] mu'_mu_large_row_im1 [OF _ k] mu_no_change_large_row[OF _ k, of j] mu'_mu_small_i mu'_mu_i_im1_j jk j k by auto } note new_g = this { (* 16.13 (ii) : norm of g (i - 1) decreases by reduction factor *) note sq_norm_g2_im1 also have "?n1 i + (?mu1 i (i - 1) * ?mu1 i (i - 1)) * ?n1 (i - 1) < 1/\ * (?n1 (i - 1)) + (1/2 * 1/2) * (?n1 (i - 1))" proof (rule add_less_le_mono[OF _ mult_mono]) from norm_ineq[unfolded mult.commute[of \], THEN linordered_field_class.mult_imp_less_div_pos[OF \0(1)]] show "?n1 i < 1/\ * ?n1 (i - 1)" using len i by auto from m12 have abs: "abs (?mu1 i (i - 1)) \ 1/2" by auto have "?mu1 i (i - 1) * ?mu1 i (i - 1) \ abs (?mu1 i (i - 1)) * abs (?mu1 i (i - 1))" by auto also have "\ \ 1/2 * 1/2" using mult_mono[OF abs abs] by auto finally show "?mu1 i (i - 1) * ?mu1 i (i - 1) \ 1/2 * 1/2" by auto qed auto also have "\ = reduction * sq_norm (?g1 (i - 1))" unfolding reduction_def using \0 by (simp add: ring_distribs add_divide_distrib) finally have "?n2 (i - 1) < reduction * ?n1 (i - 1)" . } note g_reduction = this (* Lemma 16.13 (ii) *) have lin_indpt_list_fs': "gs.lin_indpt_list (RAT fs')" unfolding gs.lin_indpt_list_def using conn2 by auto { (* stay reduced *) assume "LLL_invariant False i fs" note inv = LLL_invD[OF this] from inv have "weakly_reduced fs i" by auto hence "weakly_reduced fs (i - 1)" unfolding gram_schmidt_fs.weakly_reduced_def by auto hence red: "weakly_reduced fs' (i - 1)" unfolding gram_schmidt_fs.weakly_reduced_def using i G2_G by simp from inv have sred: "reduced fs i" by auto have sred: "reduced fs' (i - 1)" unfolding gram_schmidt_fs.reduced_def proof (intro conjI[OF red] allI impI, goal_cases) case (1 i' j) with sred have "\?mu1 i' j\ \ 1 / 2" unfolding gram_schmidt_fs.reduced_def by auto thus ?case using mu'_mu_small_i[OF 1(1)] by simp qed have mu_small: "\_small fs' (i - 1)" unfolding \_small_def proof (intro allI impI, goal_cases) case (1 j) thus ?case using inv(11) unfolding mu'_mu_i_im1_j[OF 1] \_small_def by auto qed show "LLL_invariant False (i - 1) fs'" by (rule LLL_invI, insert lin_indpt_list_fs' conn2 mu_small span' lattice fs' sred i, auto) } (* invariant is established *) show newInvw: "LLL_invariant_weak fs'" by (rule LLL_inv_wI, insert lin_indpt_list_fs' conn2 span' lattice fs', auto) (* show decrease in measure *) { (* 16.16 (ii), the decreasing case *) have ile: "i \ m" using i by auto from Gd[OF newInvw, folded d_def, OF ile] have "?R (d fs' i) = (\j = prod ?n2 ({0 ..< i-1} \ {i - 1})" by (rule sym, rule prod.cong, (insert i0, auto)[1], insert i, auto) also have "\ = ?n2 (i - 1) * prod ?n2 ({0 ..< i-1})" by simp also have "prod ?n2 ({0 ..< i-1}) = prod ?n1 ({0 ..< i-1})" by (rule prod.cong[OF refl], subst g2_g1_identical, insert i, auto) also have "\ = (prod ?n1 ({0 ..< i-1} \ {i - 1})) / ?n1 (i - 1)" by (subst prod.union_disjoint, insert norm_pos1[OF im1], auto) also have "prod ?n1 ({0 ..< i-1} \ {i - 1}) = prod ?n1 {0.. = (\j = ?R (d fs i)" unfolding d_def Gd[OF Linvw ile] by (rule prod.cong[OF refl], insert i, auto) finally have new_di: "?R (d fs' i) = ?n2 (i - 1) / ?n1 (i - 1) * ?R (d fs i)" by simp also have "\ < (reduction * ?n1 (i - 1)) / ?n1 (i - 1) * ?R (d fs i)" by (rule mult_strict_right_mono[OF divide_strict_right_mono[OF g_reduction norm_pos1[OF im1]]], insert LLL_d_pos[OF Linvw] i, auto) also have "\ = reduction * ?R (d fs i)" using norm_pos1[OF im1] by auto finally have "d fs' i < real_of_rat reduction * d fs i" using of_rat_less of_rat_mult of_rat_of_int_eq by metis note this new_di } note d_i = this show "ii \ m \ ?R (d fs' ii) = (if ii = i then ?n2 (i - 1) / ?n1 (i - 1) * ?R (d fs i) else ?R (d fs ii))" for ii using d_i d by auto have pos: "k < m \ 0 < d fs' k" "k < m \ 0 \ d fs' k" for k using LLL_d_pos[OF newInvw, of k] by auto have prodpos:"0< (\ix\{0.. (\x\{0..iaii j \ {0 ..< m} - {i} \ {i}. d fs' j)" by (rule prod.cong, insert i, auto) also have "real_of_int \ = real_of_int (\ j \ {0 ..< m} - {i}. d fs' j) * real_of_int (d fs' i)" by (subst prod.union_disjoint, auto) also have "\ < (\ j \ {0 ..< m} - {i}. d fs' j) * (of_rat reduction * d fs i)" by(rule mult_strict_left_mono[OF d_i(1)],insert prod_pos',auto) also have "(\ j \ {0 ..< m} - {i}. d fs' j) = (\ j \ {0 ..< m} - {i}. d fs j)" by (rule prod.cong, insert d, auto) also have "\ * (of_rat reduction * d fs i) = of_rat reduction * (\ j \ {0 ..< m} - {i} \ {i}. d fs j)" by (subst prod.union_disjoint, auto) also have "(\ j \ {0 ..< m} - {i} \ {i}. d fs j) = (\ j = 4/3") case True show ?thesis using D unfolding reduction(4)[OF True] logD_def unfolding True by simp next case False hence False': "\ = 4/3 \ False" by simp from False \ have "\ > 4/3" by simp with reduction have reduction1: "reduction < 1" by simp let ?new = "real (D fs')" let ?old = "real (D fs)" let ?log = "log (1/of_rat reduction)" note pos = LLL_D_pos[OF newInvw] LLL_D_pos[OF Linvw] from reduction have "real_of_rat reduction > 0" by auto hence gediv:"1/real_of_rat reduction > 0" by auto have "(1/of_rat reduction) * ?new \ ((1/of_rat reduction) * of_rat reduction) * ?old" unfolding mult.assoc mult_le_cancel_iff2[OF gediv] using D by simp also have "(1/of_rat reduction) * of_rat reduction = 1" using reduction by auto finally have "(1/of_rat reduction) * ?new \ ?old" by auto hence "?log ((1/of_rat reduction) * ?new) \ ?log ?old" by (subst log_le_cancel_iff, auto simp: pos reduction1 reduction) hence "floor (?log ((1/of_rat reduction) * ?new)) \ floor (?log ?old)" by (rule floor_mono) hence "nat (floor (?log ((1/of_rat reduction) * ?new))) \ nat (floor (?log ?old))" by simp also have "\ = logD fs" unfolding logD_def False' by simp also have "?log ((1/of_rat reduction) * ?new) = 1 + ?log ?new" by (subst log_mult, insert reduction reduction1, auto simp: pos ) also have "floor (1 + ?log ?new) = 1 + floor (?log ?new)" by simp also have "nat (1 + floor (?log ?new)) = 1 + nat (floor (?log ?new))" by (subst nat_add_distrib, insert pos reduction reduction1, auto) also have "nat (floor (?log ?new)) = logD fs'" unfolding logD_def False' by simp finally show "logD fs' < logD fs" by simp qed show "LLL_measure i fs > LLL_measure (i - 1) fs'" unfolding LLL_measure_def using i logD by simp qed lemma LLL_inv_initial_state: "LLL_invariant True 0 fs_init" proof - from lin_dep[unfolded gs.lin_indpt_list_def] have "set (RAT fs_init) \ Rn" by auto hence fs_init: "set fs_init \ carrier_vec n" by auto show ?thesis by (rule LLL_invI[OF fs_init len _ _ lin_dep], auto simp: L_def gs.reduced_def gs.weakly_reduced_def) qed lemma LLL_inv_m_imp_reduced: assumes "LLL_invariant True m fs" shows "reduced fs m" using LLL_invD[OF assms] by blast lemma basis_reduction_short_vector: assumes LLL_inv: "LLL_invariant True m fs" and v: "v = hd fs" and m0: "m \ 0" shows "v \ carrier_vec n" "v \ L - {0\<^sub>v n}" "h \ L - {0\<^sub>v n} \ rat_of_int (sq_norm v) \ \ ^ (m - 1) * rat_of_int (sq_norm h)" "v \ 0\<^sub>v j" proof - let ?L = "lattice_of fs_init" have a1: "\ \ 1" using \ by auto from LLL_invD[OF LLL_inv] have L: "lattice_of fs = L" and red: "gram_schmidt_fs.weakly_reduced n (RAT fs) \ (length (RAT fs))" and basis: "lin_indep fs" and lenH: "length fs = m" and H: "set fs \ carrier_vec n" by (auto simp: gs.lin_indpt_list_def gs.reduced_def) from lin_dep have G: "set fs_init \ carrier_vec n" unfolding gs.lin_indpt_list_def by auto with m0 len have "dim_vec (hd fs_init) = n" by (cases fs_init, auto) from v m0 lenH v have v: "v = fs ! 0" by (cases fs, auto) interpret gs1: gram_schmidt_fs_lin_indpt n "RAT fs" by (standard) (use assms LLL_invariant_def gs.lin_indpt_list_def in auto) let ?r = "rat_of_int" let ?rv = "map_vec ?r" let ?F = "RAT fs" let ?h = "?rv h" { assume h:"h \ L - {0\<^sub>v n}" (is ?h_req) from h[folded L] have h: "h \ lattice_of fs" "h \ 0\<^sub>v n" by auto { assume f: "?h = 0\<^sub>v n" have "?h = ?rv (0\<^sub>v n)" unfolding f by (intro eq_vecI, auto) hence "h = 0\<^sub>v n" using of_int_hom.vec_hom_zero_iff[of h] of_int_hom.vec_hom_inj by auto with h have False by simp } hence h0: "?h \ 0\<^sub>v n" by auto with lattice_of_of_int[OF H h(1)] have "?h \ gs.lattice_of ?F - {0\<^sub>v n}" by auto } from gs1.weakly_reduced_imp_short_vector[OF red this a1] lenH show "h \ L - {0\<^sub>v n} \ ?r (sq_norm v) \ \ ^ (m - 1) * ?r (sq_norm h)" using basis unfolding L v gs.lin_indpt_list_def by (auto simp: sq_norm_of_int) from m0 H lenH show vn: "v \ carrier_vec n" unfolding v by (cases fs, auto) have vL: "v \ L" unfolding L[symmetric] v using m0 H lenH by (intro basis_in_latticeI, cases fs, auto) { assume "v = 0\<^sub>v n" hence "hd ?F = 0\<^sub>v n" unfolding v using m0 lenH by (cases fs, auto) with gs.lin_indpt_list_nonzero[OF basis] have False using m0 lenH by (cases fs, auto) } with vL show v: "v \ L - {0\<^sub>v n}" by auto have jn:"0\<^sub>v j \ carrier_vec n \ j = n" unfolding zero_vec_def carrier_vec_def by auto with v vn show "v \ 0\<^sub>v j" by auto qed lemma LLL_mu_d_Z: assumes inv: "LLL_invariant_weak fs" and j: "j \ ii" and ii: "ii < m" shows "of_int (d fs (Suc j)) * \ fs ii j \ \" proof - interpret fs: fs_int' n m fs_init fs by standard (use inv in auto) show ?thesis using assms fs.fs_int_mu_d_Z LLL_inv_wD[OF inv] unfolding d_def fs.d_def by auto qed context fixes fs assumes Linv: "LLL_invariant_weak fs" and gbnd: "g_bound fs" begin interpretation gs1: gram_schmidt_fs_lin_indpt n "RAT fs" by (standard) (use Linv LLL_invariant_weak_def gs.lin_indpt_list_def in auto) lemma LLL_inv_N_pos: assumes m: "m \ 0" shows "N > 0" proof - let ?r = rat_of_int note inv = LLL_inv_wD[OF Linv] from inv have F: "RAT fs ! 0 \ Rn" "fs ! 0 \ carrier_vec n" using m by auto from m have upt: "[0..< m] = 0 # [1 ..< m]" using upt_add_eq_append[of 0 1 "m - 1"] by auto from inv(6) m have "map_vec ?r (fs ! 0) \ 0\<^sub>v n" using gs.lin_indpt_list_nonzero[OF inv(1)] unfolding set_conv_nth by force hence F0: "fs ! 0 \ 0\<^sub>v n" by auto hence "sq_norm (fs ! 0) \ 0" using F by simp hence 1: "sq_norm (fs ! 0) \ 1" using sq_norm_vec_ge_0[of "fs ! 0"] by auto from gbnd m have "sq_norm (gso fs 0) \ of_nat N" unfolding g_bound_def by auto also have "gso fs 0 = RAT fs ! 0" unfolding upt using F by (simp add: gs1.gso.simps[of 0]) also have "RAT fs ! 0 = map_vec ?r (fs ! 0)" using inv(6) m by auto also have "sq_norm \ = ?r (sq_norm (fs ! 0))" by (simp add: sq_norm_of_int) finally show ?thesis using 1 by (cases N, auto) qed (* equation (3) in front of Lemma 16.18 *) lemma d_approx_main: assumes i: "ii \ m" "m \ 0" shows "rat_of_int (d fs ii) \ rat_of_nat (N^ii)" proof - note inv = LLL_inv_wD[OF Linv] from LLL_inv_N_pos i have A: "0 < N" by auto note main = inv(2)[unfolded gram_schmidt_int_def gram_schmidt_wit_def] have "rat_of_int (d fs ii) = (\jgso fs j\\<^sup>2)" unfolding d_def using i by (auto simp: Gramian_determinant [OF Linv]) also have "\ \ (\j = (of_nat N)^ii" unfolding prod_constant by simp also have "\ = of_nat (N^ii)" by simp finally show ?thesis by simp qed lemma d_approx: assumes i: "ii < m" shows "rat_of_int (d fs ii) \ rat_of_nat (N^ii)" using d_approx_main[of ii] assms by auto lemma d_bound: assumes i: "ii < m" shows "d fs ii \ N^ii" using d_approx[OF assms] unfolding d_def by linarith lemma D_approx: "D fs \ N ^ (m * m)" proof - note inv = LLL_inv_wD[OF Linv] from LLL_inv_N_pos have N: "m \ 0 \ 0 < N" by auto note main = inv(2)[unfolded gram_schmidt_int_def gram_schmidt_wit_def] have "rat_of_int (\ii \ (\i \ (\i = (of_nat N)^(m * m)" unfolding prod_constant power_mult by simp also have "\ = of_nat (N ^ (m * m))" by simp finally have "(\i N ^ (m * m)" by linarith also have "(\i N ^ (m * m)" by linarith qed lemma LLL_measure_approx: assumes "\ > 4/3" "m \ 0" shows "LLL_measure i fs \ m + 2 * m * m * log base N" proof - have b1: "base > 1" using base assms by auto have id: "base = 1 / real_of_rat reduction" unfolding base_def reduction_def using \0 by (auto simp: field_simps of_rat_divide) from LLL_D_pos[OF Linv] have D1: "real (D fs) \ 1" by auto note invD = LLL_inv_wD[OF Linv] from invD have F: "set fs \ carrier_vec n" and len: "length fs = m" by auto have N0: "N > 0" using LLL_inv_N_pos[OF assms(2)] . from D_approx have D: "D fs \ N ^ (m * m)" . hence "real (D fs) \ real (N ^ (m * m))" by linarith also have "\ = real N ^ (m * m)" by simp finally have log: "log base (real (D fs)) \ log base (real N ^ (m * m))" by (subst log_le_cancel_iff[OF b1], insert D1 N0, auto) have "real (logD fs) = real (nat \log base (real (D fs))\)" unfolding logD_def id using assms by auto also have "\ \ log base (real (D fs))" using b1 D1 by auto also have "\ \ log base (real N ^ (m * m))" by fact also have "\ = (m * m) * log base (real N)" by (rule log_nat_power, insert N0, auto) finally have main: "logD fs \ m * m * log base N" by simp have "real (LLL_measure i fs) = real (2 * logD fs + m - i)" unfolding LLL_measure_def split invD(1) by simp also have "\ \ 2 * real (logD fs) + m" using invD by simp also have "\ \ 2 * (m * m * log base N) + m" using main by auto finally show ?thesis by simp qed end lemma g_bound_fs_init: "g_bound fs_init" proof - { fix i assume i: "i < m" let ?N = "map (nat o sq_norm) fs_init" let ?r = rat_of_int from i have mem: "nat (sq_norm (fs_init ! i)) \ set ?N" using fs_init len unfolding set_conv_nth by force interpret gs: gram_schmidt_fs_lin_indpt n "RAT fs_init" by (standard) (use len lin_dep LLL_invariant_def gs.lin_indpt_list_def in auto) from mem_set_imp_le_max_list[OF _ mem] have FN: "nat (sq_norm (fs_init ! i)) \ N" unfolding N_def by force hence "\fs_init ! i\\<^sup>2 \ int N" using i by auto also have "\ \ int (N * m)" using i by fastforce finally have f_bnd: "\fs_init ! i\\<^sup>2 \ int (N * m)" . from FN have "rat_of_nat (nat (sq_norm (fs_init ! i))) \ rat_of_nat N" by simp also have "rat_of_nat (nat (sq_norm (fs_init ! i))) = ?r (sq_norm (fs_init ! i))" using sq_norm_vec_ge_0[of "fs_init ! i"] by auto also have "\ = sq_norm (RAT fs_init ! i)" unfolding sq_norm_of_int[symmetric] using fs_init len i by auto finally have "sq_norm (RAT fs_init ! i) \ rat_of_nat N" . with gs.sq_norm_gso_le_f i len lin_dep have g_bnd: "\gs.gso i\\<^sup>2 \ rat_of_nat N" unfolding gs.lin_indpt_list_def by fastforce note f_bnd g_bnd } thus "g_bound fs_init" unfolding g_bound_def by auto qed lemma LLL_measure_approx_fs_init: "LLL_invariant upw i fs_init \ 4 / 3 < \ \ m \ 0 \ real (LLL_measure i fs_init) \ real m + real (2 * m * m) * log base (real N)" using LLL_measure_approx[OF LLL_inv_imp_w g_bound_fs_init] . lemma N_le_MMn: assumes m0: "m \ 0" shows "N \ nat M * nat M * n" unfolding N_def proof (rule max_list_le, unfold set_map o_def) fix ni assume "ni \ (\x. nat \x\\<^sup>2) ` set fs_init" then obtain fi where ni: "ni = nat (\fi\\<^sup>2)" and fi: "fi \ set fs_init" by auto from fi len obtain i where fii: "fi = fs_init ! i" and i: "i < m" unfolding set_conv_nth by auto from fi fs_init have fi: "fi \ carrier_vec n" by auto let ?set = "{\fs_init ! i $ j\ |i j. i < m \ j < n} \ {0}" have id: "?set = (\ (i,j). abs (fs_init ! i $ j)) ` ({0.. {0.. {0}" by force have fin: "finite ?set" unfolding id by auto { fix j assume "j < n" hence "M \ \fs_init ! i $ j\" unfolding M_def using i by (intro Max_ge[of _ "abs (fs_init ! i $ j)"], intro fin, auto) } note M = this from Max_ge[OF fin, of 0] have M0: "M \ 0" unfolding M_def by auto have "ni = nat (\fi\\<^sup>2)" unfolding ni by auto also have "\ \ nat (int n * \fi\\<^sub>\\<^sup>2)" using sq_norm_vec_le_linf_norm[OF fi] by (intro nat_mono, auto) also have "\ = n * nat (\fi\\<^sub>\\<^sup>2)" by (simp add: nat_mult_distrib) also have "\ \ n * nat (M^2)" proof (rule mult_left_mono[OF nat_mono]) have fi: "\fi\\<^sub>\ \ M" unfolding linf_norm_vec_def proof (rule max_list_le, unfold set_append set_map, rule ccontr) fix x assume "x \ abs ` set (list_of_vec fi) \ set [0]" and xM: "\ x \ M" with M0 obtain fij where fij: "fij \ set (list_of_vec fi)" and x: "x = abs fij" by auto from fij fi obtain j where j: "j < n" and fij: "fij = fi $ j" unfolding set_list_of_vec vec_set_def by auto from M[OF j] xM[unfolded x fij fii] show False by auto qed auto show "\fi\\<^sub>\\<^sup>2 \ M^2" unfolding abs_le_square_iff[symmetric] using fi using linf_norm_vec_ge_0[of fi] by auto qed auto finally show "ni \ nat M * nat M * n" using M0 by (subst nat_mult_distrib[symmetric], auto simp: power2_eq_square ac_simps) qed (insert m0 len, auto) subsection \Basic LLL implementation based on previous results\ text \We now assemble a basic implementation of the LLL algorithm, where only the lattice basis is updated, and where the GSO and the $\mu$-values are always computed from scratch. This enables a simple soundness proof and permits to separate an efficient implementation from the soundness reasoning.\ fun basis_reduction_add_rows_loop where "basis_reduction_add_rows_loop i fs 0 = fs" | "basis_reduction_add_rows_loop i fs (Suc j) = ( let c = round (\ fs i j); fs' = (if c = 0 then fs else fs[ i := fs ! i - c \\<^sub>v fs ! j]) in basis_reduction_add_rows_loop i fs' j)" definition basis_reduction_add_rows where "basis_reduction_add_rows upw i fs = (if upw then basis_reduction_add_rows_loop i fs i else fs)" definition basis_reduction_swap where "basis_reduction_swap i fs = (False, i - 1, fs[i := fs ! (i - 1), i - 1 := fs ! i])" definition basis_reduction_step where "basis_reduction_step upw i fs = (if i = 0 then (True, Suc i, fs) else let fs' = basis_reduction_add_rows upw i fs in if sq_norm (gso fs' (i - 1)) \ \ * sq_norm (gso fs' i) then (True, Suc i, fs') else basis_reduction_swap i fs')" function basis_reduction_main where "basis_reduction_main (upw,i,fs) = (if i < m \ LLL_invariant upw i fs then basis_reduction_main (basis_reduction_step upw i fs) else fs)" by pat_completeness auto definition "reduce_basis = basis_reduction_main (True, 0, fs_init)" definition "short_vector = hd reduce_basis" text \Soundness of this implementation is easily proven\ lemma basis_reduction_add_rows_loop: assumes inv: "LLL_invariant True i fs" and mu_small: "\_small_row i fs j" and res: "basis_reduction_add_rows_loop i fs j = fs'" and i: "i < m" and j: "j \ i" shows "LLL_invariant False i fs'" "LLL_measure i fs' = LLL_measure i fs" proof (atomize(full), insert assms, induct j arbitrary: fs) case (0 fs) thus ?case using basis_reduction_add_row_done[of i fs] by auto next case (Suc j fs) hence j: "j < i" by auto let ?c = "round (\ fs i j)" show ?case proof (cases "?c = 0") case True thus ?thesis using Suc(1)[OF Suc(2) basis_reduction_add_row_main_0[OF LLL_inv_imp_w[OF Suc(2)] i j True Suc(3)]] Suc(2-) by auto next case False note step = basis_reduction_add_row_main(2-)[OF LLL_inv_imp_w[OF Suc(2)] i j refl] note step = step(1)[OF Suc(2)] step(2-) show ?thesis using Suc(1)[OF step(1-2)] False Suc(2-) step(4) by simp qed qed lemma basis_reduction_add_rows: assumes inv: "LLL_invariant upw i fs" and res: "basis_reduction_add_rows upw i fs = fs'" and i: "i < m" shows "LLL_invariant False i fs'" "LLL_measure i fs' = LLL_measure i fs" proof (atomize(full), goal_cases) case 1 note def = basis_reduction_add_rows_def show ?case proof (cases upw) case False with res inv show ?thesis by (simp add: def) next case True with inv have "LLL_invariant True i fs" by auto note start = this \_small_row_refl[of i fs] from res[unfolded def] True have "basis_reduction_add_rows_loop i fs i = fs'" by auto from basis_reduction_add_rows_loop[OF start this i] show ?thesis by auto qed qed lemma basis_reduction_swap: assumes inv: "LLL_invariant False i fs" and res: "basis_reduction_swap i fs = (upw',i',fs')" and cond: "sq_norm (gso fs (i - 1)) > \ * sq_norm (gso fs i)" and i: "i < m" "i \ 0" shows "LLL_invariant upw' i' fs'" (is ?g1) "LLL_measure i' fs' < LLL_measure i fs" (is ?g2) proof - note invw = LLL_inv_imp_w[OF inv] note def = basis_reduction_swap_def from res[unfolded basis_reduction_swap_def] have id: "upw' = False" "i' = i - 1" "fs' = fs[i := fs ! (i - 1), i - 1 := fs ! i]" by auto from basis_reduction_swap_main(2-3)[OF invw _ i cond id(3)] inv show ?g1 ?g2 unfolding id by auto qed lemma basis_reduction_step: assumes inv: "LLL_invariant upw i fs" and res: "basis_reduction_step upw i fs = (upw',i',fs')" and i: "i < m" shows "LLL_invariant upw' i' fs'" "LLL_measure i' fs' < LLL_measure i fs" proof (atomize(full), goal_cases) case 1 note def = basis_reduction_step_def note invw = LLL_inv_imp_w[OF inv] obtain fs'' where fs'': "basis_reduction_add_rows upw i fs = fs''" by auto show ?case proof (cases "i = 0") case True from increase_i[OF inv i] True res show ?thesis by (auto simp: def) next case False hence id: "(i = 0) = False" by auto note res = res[unfolded def id if_False fs'' Let_def] let ?x = "sq_norm (gso fs'' (i - 1))" let ?y = "\ * sq_norm (gso fs'' i)" from basis_reduction_add_rows[OF inv fs'' i] have inv: "LLL_invariant False i fs''" and meas: "LLL_measure i fs'' = LLL_measure i fs" by auto note invw = LLL_inv_imp_w[OF inv] show ?thesis proof (cases "?x \ ?y") case True from increase_i[OF inv i] id True res meas show ?thesis by simp next case gt: False hence "?x > ?y" by auto from basis_reduction_swap[OF inv _ this i False] gt res meas show ?thesis by auto qed qed qed termination by (relation "measure (\ (upw,i,fs). LLL_measure i fs)", insert basis_reduction_step, auto split: prod.splits) declare basis_reduction_main.simps[simp del] lemma basis_reduction_main: assumes "LLL_invariant upw i fs" and res: "basis_reduction_main (upw,i,fs) = fs'" shows "LLL_invariant True m fs'" using assms proof (induct "LLL_measure i fs" arbitrary: i fs upw rule: less_induct) case (less i fs upw) have id: "LLL_invariant upw i fs = True" using less by auto note res = less(3)[unfolded basis_reduction_main.simps[of upw i fs] id] note inv = less(2) note IH = less(1) show ?case proof (cases "i < m") case i: True obtain i' fs' upw' where step: "basis_reduction_step upw i fs = (upw',i',fs')" (is "?step = _") by (cases ?step, auto) from IH[OF basis_reduction_step(2,1)[OF inv step i]] res[unfolded step] i show ?thesis by auto next case False with LLL_invD[OF inv] have i: "i = m" by auto with False res inv have "LLL_invariant upw m fs'" by auto thus "LLL_invariant True m fs'" unfolding LLL_invariant_def by auto qed qed lemma reduce_basis_inv: assumes res: "reduce_basis = fs" shows "LLL_invariant True m fs" using basis_reduction_main[OF LLL_inv_initial_state res[unfolded reduce_basis_def]] . lemma reduce_basis: assumes res: "reduce_basis = fs" shows "lattice_of fs = L" "reduced fs m" "lin_indep fs" "length fs = m" using LLL_invD[OF reduce_basis_inv[OF res]] by blast+ lemma short_vector: assumes res: "short_vector = v" and m0: "m \ 0" shows "v \ carrier_vec n" "v \ L - {0\<^sub>v n}" "h \ L - {0\<^sub>v n} \ rat_of_int (sq_norm v) \ \ ^ (m - 1) * rat_of_int (sq_norm h)" "v \ 0\<^sub>v j" using basis_reduction_short_vector[OF reduce_basis_inv[OF refl] res[symmetric, unfolded short_vector_def] m0] by blast+ end 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,751 +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" + let ?sab = "Transposition.transpose a b" + let ?sfab = "Transposition.transpose (?fn a) (?fn b)" 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" + have id: "?ft (Transposition.transpose a b \ p) = Transposition.transpose (?fn a) (?fn b) \ ?ft p" proof fix c - show "?ft (Fun.swap a b id \ p) c = (Fun.swap (?fn a) (?fn b) id \ ?ft p) c" + show "?ft (Transposition.transpose a b \ p) c = (Transposition.transpose (?fn a) (?fn b) \ ?ft p) c" proof (cases "p (?tn c) = a \ p (?tn c) = b") case True 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: swap_id_eq) qed qed 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 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/Planarity_Certificates/Planarity/Permutations_2.thy b/thys/Planarity_Certificates/Planarity/Permutations_2.thy --- a/thys/Planarity_Certificates/Planarity/Permutations_2.thy +++ b/thys/Planarity_Certificates/Planarity/Permutations_2.thy @@ -1,153 +1,149 @@ theory Permutations_2 imports "HOL-Combinatorics.Permutations" Graph_Theory.Auxiliary Executable_Permutations begin section \Modifying Permutations\ abbreviation funswapid :: "'a \ 'a \ 'a \ 'a" (infix "\\<^sub>F" 90) where - "x \\<^sub>F y \ Fun.swap x y id" + "x \\<^sub>F y \ transpose x y" definition perm_swap :: "'a \ 'a \ ('a \ 'a) \ ('a \ 'a)" where "perm_swap x y f \ x \\<^sub>F y o f o x \\<^sub>F y" definition perm_rem :: "'a \ ('a \ 'a) \ ('a \ 'a)" where "perm_rem x f \ if f x \ x then x \\<^sub>F f x o f else f" text \ An example: @{lemma "perm_rem (2 :: nat) (list_succ [1,2,3,4]) x = list_succ [1,3,4] x" by (auto simp: perm_rem_def Fun.swap_def list_succ_def)} \ lemma perm_swap_id[simp]: "perm_swap a b id = id" by (auto simp: perm_swap_def) lemma perm_rem_permutes: assumes "f permutes S \ {x}" shows "perm_rem x f permutes S" - using assms by (auto simp: permutes_def perm_rem_def) (metis swap_id_eq)+ + using assms by (auto simp: permutes_def perm_rem_def) (metis transpose_def)+ lemma perm_rem_same: assumes "bij f" "f y = y" shows "perm_rem x f y = f y" - using assms by (auto simp: perm_rem_def swap_id_eq bij_iff) + using assms by (auto simp: perm_rem_def bij_iff transpose_def) lemma perm_rem_simps: assumes "bij f" shows "x = y \ perm_rem x f y = x" "f y = x \ perm_rem x f y = f x" "y \ x \ f y \ x \ perm_rem x f y = f y" - using assms - apply (auto simp: perm_rem_def ) - by (metis bij_iff id_apply swap_apply(3)) + using assms by (auto simp: perm_rem_def transpose_def bij_iff) lemma bij_swap_compose: "bij (x \\<^sub>F y o f) \ bij f" by (metis UNIV_I bij_betw_comp_iff2 bij_betw_id bij_swap_iff subsetI) lemma bij_perm_rem[simp]: "bij (perm_rem x f) \ bij f" by (simp add: perm_rem_def bij_swap_compose) lemma perm_rem_conv: "\f x y. bij f \ perm_rem x f y = ( if x = y then x else if f y = x then f (f y) else f y)" by (auto simp: perm_rem_simps) lemma perm_rem_commutes: assumes "bij f" shows "perm_rem a (perm_rem b f) = perm_rem b (perm_rem a f)" proof - have bij_simp: "\x y. f x = f y \ x = y" using assms by (auto simp: bij_iff) show ?thesis using assms by (auto simp: perm_rem_conv bij_simp fun_eq_iff) qed lemma perm_rem_id[simp]: "perm_rem a id = id" by (simp add: perm_rem_def) lemma bij_eq_iff: assumes "bij f" shows "f x = f y \ x = y" using assms by (metis bij_iff) lemma swap_swap_id[simp]: "(x \\<^sub>F y) ((x \\<^sub>F y) z) = z" by (simp add: swap_id_eq) lemma in_funswapid_image_iff: "\a b x S. x \ (a \\<^sub>F b) ` S \ (a \\<^sub>F b) x \ S" - by (metis bij_def bij_id bij_swap_iff inj_image_mem_iff swap_swap_id) - + by (metis inj_image_mem_iff inj_transpose swap_swap_id) + lemma perm_swap_comp: "perm_swap a b (f \ g) x = perm_swap a b f (perm_swap a b g x)" by (auto simp: perm_swap_def) lemma bij_perm_swap_iff[simp]: "bij (perm_swap a b f) \ bij f" - by (auto simp: perm_swap_def bij_swap_compose bij_comp comp_swap) - + by (simp add: bij_swap_compose bij_swap_iff perm_swap_def) + lemma funpow_perm_swap: "perm_swap a b f ^^ n = perm_swap a b (f ^^ n)" by (induct n) (auto simp: perm_swap_def fun_eq_iff) lemma orbit_perm_swap: "orbit (perm_swap a b f) x = (a \\<^sub>F b) ` orbit f ((a \\<^sub>F b) x)" by (auto simp: orbit_altdef funpow_perm_swap) (auto simp: perm_swap_def) lemma has_dom_perm_swap: "has_dom (perm_swap a b f) S = has_dom f ((a \\<^sub>F b) ` S)" by (auto simp: has_dom_def perm_swap_def inj_image_mem_iff) (metis image_iff swap_swap_id) lemma perm_restrict_dom_subset: assumes "has_dom f A" shows "perm_restrict f A = f" proof - from assms have "\x. x \ A \ f x = x" by (auto simp: has_dom_def) then show ?thesis by (auto simp: perm_restrict_def fun_eq_iff) qed lemma has_domD: "has_dom f S \ x \ S \ f x = x" by (auto simp: has_dom_def) lemma has_domI: "(\x. x \ S \ f x = x) \ has_dom f S" by (auto simp: has_dom_def) lemma perm_swap_permutes2: assumes "f permutes ((x \\<^sub>F y) ` S)" shows "perm_swap x y f permutes S" using assms - by (auto simp: perm_swap_def permutes_conv_has_dom has_dom_perm_swap[unfolded perm_swap_def]) - (metis bij_swap_iff bij_swap_compose_bij comp_id comp_swap) + by (auto simp: perm_swap_def permutes_conv_has_dom has_dom_perm_swap [unfolded perm_swap_def] intro!: bij_comp) + section \Cyclic Permutations\ lemma cyclic_on_perm_swap: assumes "cyclic_on f S" shows "cyclic_on (perm_swap x y f) ((x \\<^sub>F y) ` S)" using assms by (rule cyclic_on_image) (auto simp: perm_swap_def) lemma orbit_perm_rem: assumes "bij f" "x \ y" shows "orbit (perm_rem y f) x = orbit f x - {y}" (is "?L = ?R") proof (intro set_eqI iffI) fix z assume "z \ ?L" then show "z \ ?R" using assms by induct (auto simp: perm_rem_conv bij_iff intro: orbit.intros) next fix z assume A: "z \ ?R" { assume "z \ orbit f x" then have "(z \ y \ z \ ?L) \ (z = y \ f z \ ?L)" proof induct case base with assms show ?case by (auto intro: orbit_eqI(1) simp: perm_rem_conv) next case (step z) then show ?case using assms by (cases "y = z") (auto intro: orbit_eqI simp: perm_rem_conv) qed } with A show "z \ ?L" by auto qed lemma orbit_perm_rem_eq: assumes "bij f" shows "orbit (perm_rem y f) x = (if x = y then {y} else orbit f x - {y})" using assms by (simp add: orbit_eq_singleton_iff orbit_perm_rem perm_rem_simps) lemma cyclic_on_perm_rem: assumes "cyclic_on f S" "bij f" "S \ {x}" shows "cyclic_on (perm_rem x f) (S - {x})" using assms[unfolded cyclic_on_alldef] by (simp add: cyclic_on_def orbit_perm_rem_eq) auto - - end diff --git a/thys/Planarity_Certificates/Planarity/Planar_Subdivision.thy b/thys/Planarity_Certificates/Planarity/Planar_Subdivision.thy --- a/thys/Planarity_Certificates/Planarity/Planar_Subdivision.thy +++ b/thys/Planarity_Certificates/Planarity/Planar_Subdivision.thy @@ -1,483 +1,484 @@ theory Planar_Subdivision imports Graph_Genus Reachablen Permutations_2 begin section \Combinatorial Planarity and Subdivisions\ locale subdiv1_contr = subdiv_step + fixes HM assumes H_map: "digraph_map H HM" assumes edge_rev_conv: "edge_rev HM = rev_H" sublocale subdiv1_contr \ H: digraph_map H HM rewrites "edge_rev HM = rev_H" by (intro H_map edge_rev_conv)+ sublocale subdiv1_contr \ G: fin_digraph G by unfold_locales (auto simp: arcs_G verts_G) context subdiv1_contr begin definition GM :: "'b pre_map" where "GM \ \ edge_rev = rev_G , edge_succ = perm_swap uw uv (perm_swap vw vu (fold perm_rem [wu, wv] (edge_succ HM))) \" lemma edge_rev_GM: "edge_rev GM = rev_G" by (simp add: GM_def) lemma edge_succ_GM: "edge_succ GM = perm_swap uw uv (perm_swap vw (rev_G uv) (fold perm_rem [wu, wv] (edge_succ HM)))" by (simp add: GM_def) lemma rev_H_eq_rev_G: assumes "x \ arcs G - {uv,vu}" shows "rev_H x = rev_G x" proof - have "perm_restrict rev_H (arcs G) = perm_restrict rev_G (arcs H)" using subdiv_step by (auto simp: subdivision_step_def) with assms show ?thesis unfolding arcs_H by (auto simp: perm_restrict_def fun_eq_iff split: if_splits) qed lemma edge_succ_permutes: "edge_succ GM permutes arcs G" proof - have "arcs H \ (vw \\<^sub>F rev_G uv) ` (uw \\<^sub>F uv) ` arcs G \ {wv} \ {wu}" using subdiv_distinct_arcs in_arcs_G - by (auto simp: arcs_H in_funswapid_image_iff swap_id_eq split: if_splits) + by (auto simp: arcs_H in_funswapid_image_iff transpose_def split: if_splits) then have "perm_swap uw uv (perm_swap vw (rev_G uv) (perm_rem (wv) (perm_rem (wu) (edge_succ HM)))) permutes arcs G" by (blast intro: perm_rem_permutes perm_swap_permutes2 permutes_subset H.edge_succ_permutes) then show ?thesis by (auto simp: edge_succ_GM) qed lemma out_arcs_empty: assumes "x \ verts G" shows "out_arcs G x = {} \ out_arcs H x = {}" proof assume A: "out_arcs H x = {}" have tail_eqI: "\a. tail H a = tail G a" by (simp only: tail_eq) { fix a assume "a \ out_arcs G x" moreover have "a \ arcs H \ a \ uv" "a \ arcs H \ a \ vu" using not_in_arcs_H by auto ultimately have "(uw \\<^sub>F uv) ((vw \\<^sub>F vu) a) \ out_arcs H x" using subdiv_distinct_arcs in_arcs_H not_in_arcs_H by (auto simp: arcs_G intro: tail_eqI) } then show "out_arcs G x = {}" using A by (auto simp del: in_out_arcs_conv) next assume A: "out_arcs G x = {}" have tail_eqI: "\a. tail H a = tail G a" by (simp only: tail_eq) { fix a assume "a \ out_arcs H x" moreover have "x \ w" using assms not_in_verts_G by blast ultimately have "(uw \\<^sub>F uv) ((vw \\<^sub>F vu) a) \ out_arcs G x" using subdiv_distinct_arcs in_arcs_G not_in_arcs_G - by (auto simp: arcs_H ) (auto simp: swap_id_eq intro: tail_eqI[symmetric]) + by (auto simp: arcs_H ) (auto simp: transpose_def intro: tail_eqI[symmetric]) } then show "out_arcs H x = {}" using A by (auto simp del: in_out_arcs_conv) qed lemma cyclic_on_edge_succ: assumes "x \ verts G" "out_arcs G x \ {}" shows "cyclic_on (edge_succ GM) (out_arcs G x)" proof - have oa_Gx: "out_arcs G x = (uw \\<^sub>F uv) ` (vw \\<^sub>F vu) ` (out_arcs H x - {wu} - {wv})" using subdiv_distinct_arcs not_in_arcs_G in_arcs_G - by (auto simp: in_funswapid_image_iff arcs_H swap_id_eq tail_eq[symmetric] split: if_splits) + by (auto simp: in_funswapid_image_iff arcs_H transpose_def tail_eq[symmetric] split: if_splits) have "cyclic_on (perm_swap uw uv (perm_swap vw (rev_G uv) (perm_rem (wv) (perm_rem (wu) (edge_succ HM))))) (out_arcs G x)" unfolding oa_Gx proof (intro cyclic_on_perm_swap cyclic_on_perm_rem) show "cyclic_on (edge_succ HM) (out_arcs H x)" using assms by (auto simp: out_arcs_empty verts_H intro: H.edge_succ_cyclic) show "bij (edge_succ HM)" by (simp add: H.bij_edge_succ) show "bij (perm_rem (wu) (edge_succ HM))" by (simp add: H.bij_edge_succ) have "x \ w" using assms not_in_verts_G by auto then have "wu \ out_arcs H x" "wv \ out_arcs H x" by (auto simp: arc_to_ends_def) then show "out_arcs H x - {wu} \ {wv}" "out_arcs H x \ {wu}" by blast+ qed then show ?thesis by (simp add: edge_succ_GM) qed lemma digraph_map_GM: shows "digraph_map G GM" by unfold_locales (auto simp: edge_rev_GM G.arev_dom edge_succ_permutes cyclic_on_edge_succ verts_G) end sublocale subdiv1_contr \ GM: digraph_map G GM by (rule digraph_map_GM) context subdiv1_contr begin lemma reachableGD: assumes "x \\<^sup>*\<^bsub>G\<^esub> y" shows "x \\<^sup>*\<^bsub>H\<^esub> y" using assms proof induct case base then show ?case by (auto simp: verts_H) next case (step x z) moreover have "u \\<^sup>*\<^bsub>H\<^esub> v" "v \\<^sup>*\<^bsub>H\<^esub> u" using adj_with_w by auto moreover { assume A: "(x,z) \ (u,v)" "(x,z) \ (v,u)" from \x \\<^bsub>G\<^esub> z\ obtain a where "a \ arcs G" "tail G a = x" "head G a = z" by auto with A have "a \ arcs H" "arc_to_ends H a = (x,z)" "tail H a = x" "head G a = z" by (auto simp: arcs_H tail_eq head_eq arc_to_ends_def fun_eq_iff) then have "x \\<^bsub>H\<^esub> z" by (auto simp: arcs_ends_def intro: rev_image_eqI) } ultimately show ?case by (auto intro: H.reachable_trans) qed definition proj_verts_H :: "'a \ 'a" where "proj_verts_H x \ if x = w then u else x" lemma proj_verts_H_in_G: "x \ verts H \ proj_verts_H x \ verts G" using in_verts_G by (auto simp: proj_verts_H_def verts_H) lemma dominatesHD: assumes "x \\<^bsub>H\<^esub> y" shows "proj_verts_H x \\<^sup>*\<^bsub>G\<^esub> proj_verts_H y" proof - have X1: "\a. (w, y) = arc_to_ends G a \ a \ arcs G" by (metis G.adj_in_verts(1) G.dominatesI not_in_verts_G) have X2: "\a. (x, w) = arc_to_ends G a \ a \ arcs G" by (metis G.adj_in_verts(2) G.dominatesI not_in_verts_G) show ?thesis using assms subdiv_ate_H_rev subdiv_ate in_verts_G by (auto simp: arcs_ends_def arcs_H arc_to_ends_eq proj_verts_H_def G_reach dest: X1 X2) qed lemma reachableHD: assumes reach:"x \\<^sup>*\<^bsub>H\<^esub> y" shows "proj_verts_H x \\<^sup>*\<^bsub>G\<^esub> proj_verts_H y" using assms by induct (blast intro: proj_verts_H_in_G G.reachable_trans dominatesHD)+ lemma H_reach_conv: "\x y. x \\<^sup>*\<^bsub>H\<^esub> y \ proj_verts_H x \\<^sup>*\<^bsub>G\<^esub> proj_verts_H y" using w_reach by (auto simp: reachableHD) (auto simp: proj_verts_H_def verts_H split: if_splits dest: reachableGD intro: H.reachable_trans) lemma sccs_eq: "G.sccs_verts = (`) proj_verts_H ` H.sccs_verts" (is "?L = ?R") proof (intro set_eqI iffI) fix S assume "S \ ?L" then have "w \ S" using G.sccs_verts_subsets not_in_verts_G by blast then have S_eq: "proj_verts_H ` proj_verts_H -` S = S" by (auto simp: proj_verts_H_def intro: range_eqI) then have "proj_verts_H -` S \ {}" using \S \ ?L\ by safe (auto simp: G.sccs_verts_def) with \S \ ?L\ have "proj_verts_H -` S \ H.sccs_verts" by (auto simp: G.in_sccs_verts_conv_reachable H.in_sccs_verts_conv_reachable H_reach_conv) then have "proj_verts_H ` proj_verts_H -` S \ (`) proj_verts_H ` H.sccs_verts" by (rule imageI) then show "S \ ?R" by (simp only: S_eq) next fix S assume "S \ ?R" have X: "\v x. v \ proj_verts_H ` x \ v = w \ (\y. v = proj_verts_H y \ y \ x)" by (auto simp: proj_verts_H_def split: if_splits) from \S \ ?R\ show "S \ ?L" using not_in_verts_G by (fastforce simp: G.reachable_in_verts G.in_sccs_verts_conv_reachable H.in_sccs_verts_conv_reachable H_reach_conv dest: X) qed lemma inj_on_proj_verts_H: "inj_on ((`) proj_verts_H) (pre_digraph.sccs_verts H)" proof (rule inj_onI) fix S T assume A: "S \ H.sccs_verts" "T \ H.sccs_verts" "proj_verts_H ` S = proj_verts_H ` T" have "\x. w \ x \ proj_verts_H ` x = x" by (auto simp: proj_verts_H_def) with A have "S \ T \ S \ T \ {}" by (metis H.in_sccs_verts_conv_reachable Int_iff empty_iff image_eqI proj_verts_H_def w_reach(1,2)) then show "S = T" using H.sccs_verts_disjoint[OF A(1,2)] by metis qed lemma card_sccs_verts: "card G.sccs_verts = card H.sccs_verts" unfolding sccs_eq by (intro card_image inj_on_proj_verts_H) lemma card_sccs_eq: "card G.sccs = card H.sccs" using card_sccs_verts G.inj_on_verts_sccs H.inj_on_verts_sccs by (auto simp: G.sccs_verts_conv H.sccs_verts_conv card_image) lemma isolated_verts_eq: "G.isolated_verts = H.isolated_verts" by (auto simp: G.isolated_verts_def H.isolated_verts_def verts_H out_arcs_w dest: out_arcs_empty) lemma card_verts: "card (verts H) = card (verts G) + 1" unfolding verts_H using not_in_verts_G by auto lemma card_arcs: "card (arcs H) = card (arcs G) + 2" unfolding arcs_H using not_in_arcs_G subdiv_distinct_arcs in_arcs_G by (auto simp: card_insert_if) lemma edge_succ_wu: "edge_succ HM wu = wv" using out_arcs_w out_degree_w edge_succ_permutes H.edge_succ_cyclic[of w] by (auto elim: eq_on_cyclic_on_iff1[where x="wu"] simp: verts_H out_degree_def) lemma edge_succ_wv: "edge_succ HM wv = wu" using out_arcs_w out_degree_w edge_succ_permutes H.edge_succ_cyclic[of w] by (auto elim: eq_on_cyclic_on_iff1[where x="wv"] simp: verts_H out_degree_def) lemmas edge_succ_w = edge_succ_wu edge_succ_wv lemma H_face_cycle_succ: "H.face_cycle_succ uw = wv" "H.face_cycle_succ vw = wu" unfolding H.face_cycle_succ_def by (auto simp: edge_succ_w) lemma H_edge_succ_tail_eqD: assumes "edge_succ HM a = b" shows "tail H a = tail H b" using assms H.tail_edge_succ[of a] by auto lemma YYY: "(wu \\<^sub>F wv) (edge_succ HM vw) = (edge_succ HM vw)" "(wu \\<^sub>F wv) (edge_succ HM uw) = (edge_succ HM uw)" - using H.edge_succ_cyclic[of w] subdiv_distinct_verts0 by (auto simp: swap_id_eq dest: H_edge_succ_tail_eqD) + using H.edge_succ_cyclic[of w] subdiv_distinct_verts0 + by (auto simp: Transposition.transpose_def dest: H_edge_succ_tail_eqD) text \Project arcs of @{term H} to corresponding arcs of @{term G}\ definition proj_arcs_H :: "'b \ 'b" where "proj_arcs_H x \ if x = uw \ x = wv then uv else if x = vw \ x = wu then vu else x" text \Project arcs of @{term G} to corresponding arcs of @{term H}\ definition proj_arcs_G :: "'b \ 'b" where "proj_arcs_G x \ if x = uv then uw else if x = vu then vw else x" lemma proj_arcs_H_simps[simp]: "proj_arcs_H uw = uv" "proj_arcs_H wv = uv" "proj_arcs_H vw = vu" "proj_arcs_H wu = vu" "x \ {uw,vw,wu,wv} \ proj_arcs_H x = x" "a \ arcs G \ proj_arcs_H a = a" using subdiv_distinct_arcs not_in_arcs_G by (auto simp: proj_arcs_H_def) lemma proj_arcs_H_in_arcs_G: "a \ arcs H \ proj_arcs_H a \ arcs G" using subdiv_distinct_arcs in_arcs_G by (auto simp: proj_arcs_H_def arcs_H) lemma proj_arcs_eq_swap: assumes "a \ {uv,vu,wu,wv}" shows "proj_arcs_H a = (uw \\<^sub>F uv o vw \\<^sub>F vu) a" using assms subdiv_distinct_arcs by (cases "a \ {uw,vw}") auto lemma proj_arcs_G_simps: "proj_arcs_G uv = uw" "proj_arcs_G vu = vw" "a \ {uv,vu} \ proj_arcs_G a = a" using subdiv_distinct_arcs not_in_arcs_G by (auto simp: swap_id_eq proj_arcs_G_def) lemma proj_arcs_G_in_arcs_H: assumes "a \ arcs G" shows "proj_arcs_G a \ arcs H" using assms subdiv_distinct_arcs by (auto simp: proj_arcs_G_def arcs_H) lemma proj_arcs_HG: "a \ arcs G \ proj_arcs_H (proj_arcs_G a) = a" by (auto simp: proj_arcs_G_def) lemma fcs_proj_arcs_GH: assumes "a \ arcs H" shows "H.face_cycle_set (proj_arcs_G (proj_arcs_H a)) = H.face_cycle_set a" proof - have "H.face_cycle_set vw = H.face_cycle_set wu" "H.face_cycle_set uw = H.face_cycle_set wv" unfolding H.face_cycle_set_def by (auto simp add: H_face_cycle_succ[symmetric] self_in_orbit_step H.permutation_face_cycle_succ permutation_self_in_orbit) then show ?thesis using assms not_in_arcs_H by (cases "a \ {uv,vu,uw,wu,vw,wv}") (auto simp: proj_arcs_G_simps) qed lemma H_face_cycle_succ_neq_uv: "a \ {uv,vu} \ H.face_cycle_succ a \ {uv,vu}" using not_in_arcs_H by (cases "a \ arcs H") (auto dest: H.face_cycle_succ_in_arcsI) lemma face_cycle_succ_choose_inter: "{H.face_cycle_succ uw, H.face_cycle_succ vw, H.face_cycle_succ wu, H.face_cycle_succ wv} \ {uv,vu} = {}" using subdiv_distinct_arcs H_face_cycle_succ_neq_uv by safe (simp_all, metis+) lemma face_cycle_succ_choose_neq: "H.face_cycle_succ wu \ {wu,wv}" "H.face_cycle_succ wv \ {wu,wv}" using subdiv_distinct_verts0 in_arcs_H by (auto simp del: H.edge_rev_in_arcs dest: H.tail_face_cycle_succ) lemma H_face_cycle_succ_G_not_in: assumes "a \ arcs G" shows "H.face_cycle_succ a \ {wu,wv}" proof (cases "a \ {uv,vu}") case True with assms show ?thesis using subdiv_distinct_arcs by (auto simp: arcs_H) next case False with assms have "a \ arcs H" by (auto simp: arcs_H) from assms have "head H a \ w" by (auto simp: head_eq verts_G arcs_H dest: G.head_in_verts) then show ?thesis using H.tail_face_cycle_succ[OF \a \ arcs H\] by auto qed lemma face_cycle_succ_uv: "GM.face_cycle_succ uv = proj_arcs_H (H.face_cycle_succ wv)" and face_cycle_succ_vu: "GM.face_cycle_succ vu = proj_arcs_H (H.face_cycle_succ wu)" unfolding GM.face_cycle_succ_def edge_rev_GM edge_succ_GM using face_cycle_succ_choose_neq face_cycle_succ_choose_inter subdiv_distinct_arcs apply (auto simp: fun_eq_iff perm_swap_def) apply (auto simp: perm_rem_def edge_succ_w H.face_cycle_succ_def YYY proj_arcs_H_def) done lemma face_cycle_succ_not_uv: assumes "a \ arcs G" "a \ {uv,vu}" shows "GM.face_cycle_succ a = proj_arcs_H (H.face_cycle_succ a)" proof - have "GM.face_cycle_succ a = (uw \\<^sub>F uv) ((vw \\<^sub>F rev_G uv) (perm_rem (wv) (perm_rem (wu) (edge_succ HM)) (((vw \\<^sub>F vu) ((uw \\<^sub>F uv) (rev_G a))))))" by (simp add: GM.face_cycle_succ_def perm_swap_def edge_succ_GM edge_rev_GM) also have "(vw \\<^sub>F vu) ((uw \\<^sub>F uv) (rev_G a)) = rev_G a" - using assms not_in_arcs_G by (auto simp: swap_id_eq G.arev_eq_iff) + using assms not_in_arcs_G by (auto simp: transpose_def G.arev_eq_iff) also have "perm_rem (wv) (perm_rem (wu) (edge_succ HM)) (rev_G a) = edge_succ HM (rev_G a)" proof - - have *: "\a. tail H a \ w \ (wu \\<^sub>F wv) a = a" by (auto simp: swap_id_eq) + have *: "\a. tail H a \ w \ (wu \\<^sub>F wv) a = a" by (auto simp: transpose_def) from assms have "head H a \ w" "tail H (rev_G a) = head H a" by (auto simp: tail_eq head_eq verts_G dest: G.head_in_verts) then have "((wu \\<^sub>F wv) (edge_succ HM (rev_G a))) = edge_succ HM (rev_G a)" by (intro *) auto then show ?thesis by (auto simp: perm_rem_def edge_succ_w) qed also have "edge_succ HM (rev_G a) = H.face_cycle_succ a" using assms unfolding H.face_cycle_succ_def by (simp add: rev_H_eq_rev_G) also have "(uw \\<^sub>F uv) ((vw \\<^sub>F rev_G uv) (H.face_cycle_succ a)) = proj_arcs_H (H.face_cycle_succ a)" proof - from assms have "a \ arcs H" by (auto simp: arcs_H) then have fcs_not_in: "H.face_cycle_succ a \ {uv, vu, wu, wv}" using assms H_face_cycle_succ_G_not_in in_arcs_G not_in_arcs_H by (auto simp del: G.arev_in_arcs dest: H.face_cycle_succ_closed[THEN iffD2]) then show ?thesis by (auto simp add: proj_arcs_eq_swap) qed finally show ?thesis . qed lemmas G_face_cycle_succ = face_cycle_succ_uv face_cycle_succ_vu face_cycle_succ_not_uv lemma in_G_fcs_in_H_fcs: assumes "a \ arcs G" assumes "x \ GM.face_cycle_set a" shows "x \ proj_arcs_H ` H.face_cycle_set (proj_arcs_G a)" using \x \ _\ proof induct case base show ?case by (rule rev_image_eqI[where x="proj_arcs_G a"]) (auto simp: \a \ arcs G\ proj_arcs_G_def) next case (step b) { fix x assume "x \ H.face_cycle_set (proj_arcs_G a)" then have "x \ arcs H" using \a \ arcs G\ by (auto dest: H.in_face_cycle_setD simp: proj_arcs_G_in_arcs_H) moreover then have "x \ {uv,vu}" "x \ {uw,wu,vw,wv} \ x \ arcs G" using \a \ arcs G\ by (auto simp: arcs_H dest: H.in_face_cycle_setD) ultimately have "GM.face_cycle_succ (proj_arcs_H x) \ {proj_arcs_H (H.face_cycle_succ x), proj_arcs_H (H.face_cycle_succ (H.face_cycle_succ x))}" by (cases "x \ {uw,vw,wu,wv}") (auto simp: G_face_cycle_succ H_face_cycle_succ) } moreover have "b \ arcs G" using step(1) \a \ arcs G\ by (simp add: GM.in_face_cycle_setD GM.face_cycle_set_def) ultimately show ?case using \b \ arcs G\ step(2) by (auto intro: H.face_cycle_succ_inI) qed lemma in_H_fcs_in_G_fcs: assumes "a \ arcs H" assumes "x \ H.face_cycle_set a" shows "x \ proj_arcs_H -` GM.face_cycle_set (proj_arcs_H a)" using \x \ _\ proof induct case base then show ?case by auto next case (step y) then have "y \ arcs H" using \a \ arcs H\ by (auto dest: H.in_face_cycle_setD) moreover then have "y \ {uv,vu}" by (fastforce simp: arcs_H) ultimately have "proj_arcs_H (H.face_cycle_succ y) = GM.face_cycle_succ (proj_arcs_H y) \ proj_arcs_H (H.face_cycle_succ y) = proj_arcs_H y" by (cases "y \ {uw,vw,wv,wu}") (auto simp: H_face_cycle_succ G_face_cycle_succ arcs_G) with step show ?case by (auto intro: GM.face_cycle_succ_inI) qed lemma G_fcs_eq: assumes "a \ arcs G" shows "GM.face_cycle_set a = proj_arcs_H ` H.face_cycle_set (proj_arcs_G a)" (is "?L = ?R") using assms by (auto dest: in_H_fcs_in_G_fcs[rotated] in_G_fcs_in_H_fcs[rotated] simp: proj_arcs_G_in_arcs_H proj_arcs_HG) lemma H_fcs_eq: assumes "a \ arcs H" shows "proj_arcs_H ` H.face_cycle_set a = GM.face_cycle_set (proj_arcs_H a)" using assms by (auto dest: in_H_fcs_in_G_fcs[rotated] in_G_fcs_in_H_fcs[rotated] simp: proj_arcs_H_in_arcs_G fcs_proj_arcs_GH) lemma face_cycle_sets: shows "GM.face_cycle_sets = (`) proj_arcs_H ` H.face_cycle_sets" (is "?L = ?R") unfolding GM.face_cycle_sets_def H.face_cycle_sets_def by (blast intro!: H_fcs_eq G_fcs_eq proj_arcs_G_in_arcs_H proj_arcs_H_in_arcs_G) lemma inj_on_proj_arcs_H: "inj_on ((`) proj_arcs_H) H.face_cycle_sets" proof (rule inj_onI) fix A B assume fcs: "A \ H.face_cycle_sets" "B \ H.face_cycle_sets" and pa_eq: "proj_arcs_H ` A = proj_arcs_H ` B" have xw_iff_wy: "\X. X \ H.face_cycle_sets \ uw \ X \ wv \ X" "\X. X \ H.face_cycle_sets \ vw \ X \ wu \ X" using H_face_cycle_succ by (auto simp: H.face_cycle_sets_def dest: H.face_cycle_succ_inI intro: H.face_cycle_succ_inD) have not_in_A: "uv \ A" "vu \ A" and not_in_B: "vu \ B" "uv \ B" using fcs not_in_arcs_H by (auto dest: H.in_face_cycle_setsD) have "A = proj_arcs_H -` (proj_arcs_H ` A) - {uv,vu}" using subdiv_distinct_arcs not_in_A by (auto simp: proj_arcs_H_def xw_iff_wy[OF fcs(1)] split: if_splits) also have "\ = proj_arcs_H -` (proj_arcs_H ` B) - {uv,vu}" by (simp add: pa_eq) also have "\ = B" using subdiv_distinct_arcs not_in_B by (auto simp: proj_arcs_H_def xw_iff_wy[OF fcs(2)] split: if_splits) finally show "A = B" . qed lemma card_face_cycle_sets: "card GM.face_cycle_sets = card H.face_cycle_sets" unfolding face_cycle_sets using inj_on_proj_arcs_H by (rule card_image) lemma euler_char_eq: "GM.euler_char = H.euler_char" by (auto simp: GM.euler_char_def H.euler_char_def card_verts card_arcs card_face_cycle_sets) lemma euler_genus_eq: "GM.euler_genus = H.euler_genus" by (auto simp: GM.euler_genus_def H.euler_genus_def euler_char_eq card_sccs_eq isolated_verts_eq) end lemma subdivision_genus_same_rev: assumes "subdivision (G, rev_G) (H, edge_rev HM)" "digraph_map H HM" "pre_digraph_map.euler_genus H HM = m" shows "\GM. digraph_map G GM \ pre_digraph_map.euler_genus G GM = m \ edge_rev GM = rev_G" proof - from assms show ?thesis proof (induction rev_H\"edge_rev HM" arbitrary: HM) case base then show ?case by auto next case (divide I rev_I H u v w uv uw vw) then interpret subdiv_step I rev_I H "edge_rev HM" u v w uv uw vw by unfold_locales simp interpret H: digraph_map H HM using \digraph_map H HM\ . interpret IH: subdiv1_contr I rev_I H "edge_rev HM" u v w uv uw vw HM by unfold_locales simp have eulerI: "IH.GM.euler_genus = m" by (auto simp: IH.euler_genus_eq divide) with _ IH.digraph_map_GM show ?case by (rule divide) (simp add: IH.edge_rev_GM) qed qed lemma subdivision_genus: assumes "subdivision (G, rev_G) (H, rev_H)" "digraph_map H HM" "pre_digraph_map.euler_genus H HM = m" shows "\GM. digraph_map G GM \ pre_digraph_map.euler_genus G GM = m" proof - interpret H: digraph_map H HM by fact show ?thesis using subdivision_genus_same_rev subdivision_choose_rev assms H.bidirected_digraph by metis qed lemma subdivision_comb_planar: assumes "subdivision (G, rev_G) (H, rev_H)" "comb_planar H" shows "comb_planar G" using assms unfolding comb_planar_def by (metis subdivision_genus) 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,1858 +1,1858 @@ 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 (metis G.arev_arev_raw G.face_cycle_succ_def G.fcs_x_eq_x a_in comp_apply singletonD) 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_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_id_eq split: if_splits) + by (cases "x \ {a,b}") (auto simp: rev_swap_def perm_swap_def arev_dom Transposition.transpose_def 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,583 +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" + let ?sab = "Transposition.transpose a b" + let ?sfab = "Transposition.transpose (?fn a) (?fn b)" 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" + have id: "?ft (Transposition.transpose a b \ p) = Transposition.transpose (?fn a) (?fn b) \ ?ft p" proof fix c - show "?ft (Fun.swap a b id \ p) c = (Fun.swap (?fn a) (?fn b) id \ ?ft p) c" + show "?ft (Transposition.transpose a b \ p) c = (Transposition.transpose (?fn a) (?fn b) \ ?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_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_id_eq) qed qed 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)" by simp 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" + \ _) ===> (=)) 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 diff --git a/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy b/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy --- a/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy +++ b/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy @@ -1,2890 +1,2890 @@ (* File: Symmetric_Polynomials.thy Author: Manuel Eberl (TU München) The definition of symmetric polynomials and the elementary symmetric polynomials. Proof of the fundamental theorem of symmetric polynomials and its corollaries. *) section \Symmetric Polynomials\ theory Symmetric_Polynomials imports Vieta "Polynomials.More_MPoly_Type" "HOL-Combinatorics.Permutations" begin subsection \Auxiliary facts\ (* TODO: Many of these facts and definitions should be moved elsewhere, especially the ones on polynomials (leading monomial, mapping, insertion etc.) *) text \ An infinite set has infinitely many infinite subsets. \ lemma infinite_infinite_subsets: assumes "infinite A" shows "infinite {X. X \ A \ infinite X}" proof - have "\k. \X. X \ A \ infinite X \ card (A - X) = k" for k :: nat proof fix k :: nat obtain Y where "finite Y" "card Y = k" "Y \ A" using infinite_arbitrarily_large[of A k] assms by auto moreover from this have "A - (A - Y) = Y" by auto ultimately show "\X. X \ A \ infinite X \ card (A - X) = k" using assms by (intro exI[of _ "A - Y"]) auto qed from choice[OF this] obtain f where f: "\k. f k \ A \ infinite (f k) \ card (A - f k) = k" by blast have "k = l" if "f k = f l" for k l proof (rule ccontr) assume "k \ l" hence "card (A - f k) \ card (A - f l)" using f[of k] f[of l] by auto with \f k = f l\ show False by simp qed hence "inj f" by (auto intro: injI) moreover have "range f \ {X. X \ A \ infinite X}" using f by auto ultimately show ?thesis by (subst infinite_iff_countable_subset) auto qed text \ An infinite set contains infinitely many finite subsets of any fixed nonzero cardinality. \ lemma infinite_card_subsets: assumes "infinite A" "k > 0" shows "infinite {X. X \ A \ finite X \ card X = k}" proof - obtain B where B: "B \ A" "finite B" "card B = k - 1" using infinite_arbitrarily_large[OF assms(1), of "k - 1"] by blast define f where "f = (\x. insert x B)" have "f ` (A - B) \ {X. X \ A \ finite X \ card X = k}" using assms B by (auto simp: f_def) moreover have "inj_on f (A - B)" by (auto intro!: inj_onI simp: f_def) hence "infinite (f ` (A - B))" using assms B by (subst finite_image_iff) auto ultimately show ?thesis by (rule infinite_super) qed lemma comp_bij_eq_iff: assumes "bij f" shows "g \ f = h \ f \ g = h" proof assume *: "g \ f = h \ f" show "g = h" proof fix x obtain y where [simp]: "x = f y" using bij_is_surj[OF assms] by auto have "(g \ f) y = (h \ f) y" by (simp only: *) thus "g x = h x" by simp qed qed auto lemma sum_list_replicate [simp]: "sum_list (replicate n x) = of_nat n * (x :: 'a :: semiring_1)" by (induction n) (auto simp: algebra_simps) lemma ex_subset_of_card: assumes "finite A" "card A \ k" shows "\B. B \ A \ card B = k" using assms proof (induction arbitrary: k rule: finite_induct) case empty thus ?case by auto next case (insert x A k) show ?case proof (cases "k = 0") case True thus ?thesis by (intro exI[of _ "{}"]) auto next case False from insert have "\B\A. card B = k - 1" by (intro insert.IH) auto then obtain B where B: "B \ A" "card B = k - 1" by auto with insert have [simp]: "x \ B" by auto have "insert x B \ insert x A" using B insert by auto moreover have "card (insert x B) = k" using insert B finite_subset[of B A] False by (subst card.insert_remove) auto ultimately show ?thesis by blast qed qed lemma length_sorted_list_of_set [simp]: "length (sorted_list_of_set A) = card A" using distinct_card[of "sorted_list_of_set A"] by (cases "finite A") simp_all lemma upt_add_eq_append': "i \ j \ j \ k \ [i..Subrings and ring homomorphisms\ locale ring_closed = fixes A :: "'a :: comm_ring_1 set" assumes zero_closed [simp]: "0 \ A" assumes one_closed [simp]: "1 \ A" assumes add_closed [simp]: "x \ A \ y \ A \ (x + y) \ A" assumes mult_closed [simp]: "x \ A \ y \ A \ (x * y) \ A" assumes uminus_closed [simp]: "x \ A \ -x \ A" begin lemma minus_closed [simp]: "x \ A \ y \ A \ x - y \ A" using add_closed[of x "-y"] uminus_closed[of y] by auto lemma sum_closed [intro]: "(\x. x \ X \ f x \ A) \ sum f X \ A" by (induction X rule: infinite_finite_induct) auto lemma power_closed [intro]: "x \ A \ x ^ n \ A" by (induction n) auto lemma Sum_any_closed [intro]: "(\x. f x \ A) \ Sum_any f \ A" unfolding Sum_any.expand_set by (rule sum_closed) lemma prod_closed [intro]: "(\x. x \ X \ f x \ A) \ prod f X \ A" by (induction X rule: infinite_finite_induct) auto lemma Prod_any_closed [intro]: "(\x. f x \ A) \ Prod_any f \ A" unfolding Prod_any.expand_set by (rule prod_closed) lemma prod_fun_closed [intro]: "(\x. f x \ A) \ (\x. g x \ A) \ prod_fun f g x \ A" by (auto simp: prod_fun_def when_def intro!: Sum_any_closed mult_closed) lemma of_nat_closed [simp, intro]: "of_nat n \ A" by (induction n) auto lemma of_int_closed [simp, intro]: "of_int n \ A" by (induction n) auto end locale ring_homomorphism = fixes f :: "'a :: comm_ring_1 \ 'b :: comm_ring_1" assumes add[simp]: "f (x + y) = f x + f y" assumes uminus[simp]: "f (-x) = -f x" assumes mult[simp]: "f (x * y) = f x * f y" assumes zero[simp]: "f 0 = 0" assumes one [simp]: "f 1 = 1" begin lemma diff [simp]: "f (x - y) = f x - f y" using add[of x "-y"] by (simp del: add) lemma power [simp]: "f (x ^ n) = f x ^ n" by (induction n) auto lemma sum [simp]: "f (sum g A) = (\x\A. f (g x))" by (induction A rule: infinite_finite_induct) auto lemma prod [simp]: "f (prod g A) = (\x\A. f (g x))" by (induction A rule: infinite_finite_induct) auto end lemma ring_homomorphism_id [intro]: "ring_homomorphism id" by standard auto lemma ring_homomorphism_id' [intro]: "ring_homomorphism (\x. x)" by standard auto lemma ring_homomorphism_of_int [intro]: "ring_homomorphism of_int" by standard auto subsection \Various facts about multivariate polynomials\ lemma poly_mapping_nat_ge_0 [simp]: "(m :: nat \\<^sub>0 nat) \ 0" proof (cases "m = 0") case False hence "Poly_Mapping.lookup m \ Poly_Mapping.lookup 0" by transfer auto hence "\k. Poly_Mapping.lookup m k \ 0" by (auto simp: fun_eq_iff) from LeastI_ex[OF this] Least_le[of "\k. Poly_Mapping.lookup m k \ 0"] show ?thesis by (force simp: less_eq_poly_mapping_def less_fun_def) qed auto lemma poly_mapping_nat_le_0 [simp]: "(m :: nat \\<^sub>0 nat) \ 0 \ m = 0" unfolding less_eq_poly_mapping_def poly_mapping_eq_iff less_fun_def by auto lemma of_nat_diff_poly_mapping_nat: assumes "m \ n" shows "of_nat (m - n) = (of_nat m - of_nat n :: 'a :: monoid_add \\<^sub>0 nat)" by (auto intro!: poly_mapping_eqI simp: lookup_of_nat lookup_minus when_def) lemma mpoly_coeff_transfer [transfer_rule]: "rel_fun cr_mpoly (=) poly_mapping.lookup MPoly_Type.coeff" unfolding MPoly_Type.coeff_def by transfer_prover lemma mapping_of_sum: "(\x\A. mapping_of (f x)) = mapping_of (sum f A)" by (induction A rule: infinite_finite_induct) (auto simp: plus_mpoly.rep_eq zero_mpoly.rep_eq) lemma mapping_of_eq_0_iff [simp]: "mapping_of p = 0 \ p = 0" by transfer auto lemma Sum_any_mapping_of: "Sum_any (\x. mapping_of (f x)) = mapping_of (Sum_any f)" by (simp add: Sum_any.expand_set mapping_of_sum) lemma Sum_any_parametric_cr_mpoly [transfer_rule]: "(rel_fun (rel_fun (=) cr_mpoly) cr_mpoly) Sum_any Sum_any" by (auto simp: rel_fun_def cr_mpoly_def Sum_any_mapping_of) lemma lookup_mult_of_nat [simp]: "lookup (of_nat n * m) k = n * lookup m k" proof - have "of_nat n * m = (\i k = (\i = n * lookup m k" by simp finally show ?thesis . qed lemma mpoly_eqI: assumes "\mon. MPoly_Type.coeff p mon = MPoly_Type.coeff q mon" shows "p = q" using assms by (transfer, transfer) (auto simp: fun_eq_iff) lemma coeff_mpoly_times: "MPoly_Type.coeff (p * q) mon = prod_fun (MPoly_Type.coeff p) (MPoly_Type.coeff q) mon" by (transfer', transfer') auto lemma (in ring_closed) coeff_mult_closed [intro]: "(\x. coeff p x \ A) \ (\x. coeff q x \ A) \ coeff (p * q) x \ A" by (auto simp: coeff_mpoly_times prod_fun_closed) lemma coeff_notin_vars: assumes "\(keys m \ vars p)" shows "coeff p m = 0" using assms unfolding vars_def by transfer' (auto simp: in_keys_iff) lemma finite_coeff_support [intro]: "finite {m. coeff p m \ 0}" by transfer simp lemma insertion_altdef: "insertion f p = Sum_any (\m. coeff p m * Prod_any (\i. f i ^ lookup m i))" by (transfer', transfer') (simp add: insertion_fun_def) lemma mpoly_coeff_uminus [simp]: "coeff (-p) m = -coeff p m" by transfer auto lemma Sum_any_uminus: "Sum_any (\x. -f x :: 'a :: ab_group_add) = -Sum_any f" by (simp add: Sum_any.expand_set sum_negf) lemma insertion_uminus [simp]: "insertion f (-p :: 'a :: comm_ring_1 mpoly) = -insertion f p" by (simp add: insertion_altdef Sum_any_uminus) lemma Sum_any_lookup: "finite {x. g x \ 0} \ Sum_any (\x. lookup (g x) y) = lookup (Sum_any g) y" by (auto simp: Sum_any.expand_set lookup_sum intro!: sum.mono_neutral_left) lemma Sum_any_diff: assumes "finite {x. f x \ 0}" assumes "finite {x. g x \ 0}" shows "Sum_any (\x. f x - g x :: 'a :: ab_group_add) = Sum_any f - Sum_any g" proof - have "{x. f x - g x \ 0} \ {x. f x \ 0} \ {x. g x \ 0}" by auto moreover have "finite ({x. f x \ 0} \ {x. g x \ 0})" by (subst finite_Un) (insert assms, auto) ultimately have "finite {x. f x - g x \ 0}" by (rule finite_subset) with assms show ?thesis by (simp add: algebra_simps Sum_any.distrib [symmetric]) qed lemma insertion_diff: "insertion f (p - q :: 'a :: comm_ring_1 mpoly) = insertion f p - insertion f q" proof (transfer, transfer) fix f :: "nat \ 'a" and p q :: "(nat \\<^sub>0 nat) \ 'a" assume fin: "finite {x. p x \ 0}" "finite {x. q x \ 0}" have "insertion_fun f (\x. p x - q x) = (\m. p m * (\v. f v ^ lookup m v) - q m * (\v. f v ^ lookup m v))" by (simp add: insertion_fun_def algebra_simps Sum_any_diff) also have "\ = (\m. p m * (\v. f v ^ lookup m v)) - (\m. q m * (\v. f v ^ lookup m v))" by (subst Sum_any_diff) (auto intro: finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)]) also have "\ = insertion_fun f p - insertion_fun f q" by (simp add: insertion_fun_def) finally show "insertion_fun f (\x. p x - q x) = \" . qed lemma insertion_power: "insertion f (p ^ n) = insertion f p ^ n" by (induction n) (simp_all add: insertion_mult) lemma insertion_sum: "insertion f (sum g A) = (\x\A. insertion f (g x))" by (induction A rule: infinite_finite_induct) (auto simp: insertion_add) lemma insertion_prod: "insertion f (prod g A) = (\x\A. insertion f (g x))" by (induction A rule: infinite_finite_induct) (auto simp: insertion_mult) lemma coeff_Var: "coeff (Var i) m = (1 when m = Poly_Mapping.single i 1)" by transfer' (auto simp: Var\<^sub>0_def lookup_single when_def) lemma vars_Var: "vars (Var i :: 'a :: {one,zero} mpoly) = (if (0::'a) = 1 then {} else {i})" unfolding vars_def by (auto simp: Var.rep_eq Var\<^sub>0_def) lemma insertion_Var [simp]: "insertion f (Var i) = f i" proof - have "insertion f (Var i) = (\m. (1 when m = Poly_Mapping.single i 1) * (\i. f i ^ lookup m i))" by (simp add: insertion_altdef coeff_Var) also have "\ = (\j. f j ^ lookup (Poly_Mapping.single i 1) j)" by (subst Sum_any.expand_superset[of "{Poly_Mapping.single i 1}"]) (auto simp: when_def) also have "\ = f i" by (subst Prod_any.expand_superset[of "{i}"]) (auto simp: when_def lookup_single) finally show ?thesis . qed lemma insertion_Sum_any: assumes "finite {x. g x \ 0}" shows "insertion f (Sum_any g) = Sum_any (\x. insertion f (g x))" unfolding Sum_any.expand_set insertion_sum by (intro sum.mono_neutral_right) (auto intro!: finite_subset[OF _ assms]) lemma keys_diff_subset: "keys (f - g) \ keys f \ keys g" by transfer auto lemma keys_empty_iff [simp]: "keys p = {} \ p = 0" by transfer auto lemma mpoly_coeff_0 [simp]: "MPoly_Type.coeff 0 m = 0" by transfer auto lemma lookup_1: "lookup 1 m = (if m = 0 then 1 else 0)" by transfer (simp add: when_def) lemma mpoly_coeff_1: "MPoly_Type.coeff 1 m = (if m = 0 then 1 else 0)" by (simp add: MPoly_Type.coeff_def one_mpoly.rep_eq lookup_1) lemma lookup_Const\<^sub>0: "lookup (Const\<^sub>0 c) m = (if m = 0 then c else 0)" unfolding Const\<^sub>0_def by (simp add: lookup_single when_def) lemma mpoly_coeff_Const: "MPoly_Type.coeff (Const c) m = (if m = 0 then c else 0)" by (simp add: MPoly_Type.coeff_def Const.rep_eq lookup_Const\<^sub>0) lemma coeff_smult [simp]: "coeff (smult c p) m = (c :: 'a :: mult_zero) * coeff p m" by transfer (auto simp: map_lookup) lemma in_keys_mapI: "x \ keys m \ f (lookup m x) \ 0 \ x \ keys (Poly_Mapping.map f m)" by transfer auto lemma keys_uminus [simp]: "keys (-m) = keys m" by transfer auto lemma vars_uminus [simp]: "vars (-p) = vars p" unfolding vars_def by transfer' auto lemma vars_smult: "vars (smult c p) \ vars p" unfolding vars_def by (transfer', transfer') auto lemma vars_0 [simp]: "vars 0 = {}" unfolding vars_def by transfer' simp lemma vars_1 [simp]: "vars 1 = {}" unfolding vars_def by transfer' simp lemma vars_sum: "vars (sum f A) \ (\x\A. vars (f x))" using vars_add by (induction A rule: infinite_finite_induct) auto lemma vars_prod: "vars (prod f A) \ (\x\A. vars (f x))" using vars_mult by (induction A rule: infinite_finite_induct) auto lemma vars_Sum_any: "vars (Sum_any h) \ (\i. vars (h i))" unfolding Sum_any.expand_set by (intro order.trans[OF vars_sum]) auto lemma vars_Prod_any: "vars (Prod_any h) \ (\i. vars (h i))" unfolding Prod_any.expand_set by (intro order.trans[OF vars_prod]) auto lemma vars_power: "vars (p ^ n) \ vars p" using vars_mult by (induction n) auto lemma vars_diff: "vars (p1 - p2) \ vars p1 \ vars p2" unfolding vars_def proof transfer' fix p1 p2 :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "\ (keys ` keys (p1 - p2)) \ \(keys ` (keys p1)) \ \(keys ` (keys p2))" using keys_diff_subset[of p1 p2] by (auto simp flip: not_in_keys_iff_lookup_eq_zero) qed lemma insertion_smult [simp]: "insertion f (smult c p) = c * insertion f p" unfolding insertion_altdef by (subst Sum_any_right_distrib) (auto intro: finite_subset[OF _ finite_coeff_support[of p]] simp: mult.assoc) lemma coeff_add [simp]: "coeff (p + q) m = coeff p m + coeff q m" by transfer' (simp add: lookup_add) lemma coeff_diff [simp]: "coeff (p - q) m = coeff p m - coeff q m" by transfer' (simp add: lookup_minus) lemma insertion_monom [simp]: "insertion f (monom m c) = c * Prod_any (\x. f x ^ lookup m x)" proof - have "insertion f (monom m c) = (\m'. (c when m = m') * (\v. f v ^ lookup m' v))" by (simp add: insertion_def insertion_aux_def insertion_fun_def lookup_single) also have "\ = c * (\v. f v ^ lookup m v)" by (subst Sum_any.expand_superset[of "{m}"]) (auto simp: when_def) finally show ?thesis . qed lemma insertion_aux_Const\<^sub>0 [simp]: "insertion_aux f (Const\<^sub>0 c) = c" proof - have "insertion_aux f (Const\<^sub>0 c) = (\m. (c when m = 0) * (\v. f v ^ lookup m v))" by (simp add: Const\<^sub>0_def insertion_aux_def insertion_fun_def lookup_single) also have "\ = (\m\{0}. (c when m = 0) * (\v. f v ^ lookup m v))" by (intro Sum_any.expand_superset) (auto simp: when_def) also have "\ = c" by simp finally show ?thesis . qed lemma insertion_Const [simp]: "insertion f (Const c) = c" by (simp add: insertion_def Const.rep_eq) lemma coeffs_0 [simp]: "coeffs 0 = {}" by transfer auto lemma coeffs_1 [simp]: "coeffs 1 = {1}" by transfer auto lemma coeffs_Const: "coeffs (Const c) = (if c = 0 then {} else {c})" unfolding Const_def Const\<^sub>0_def by transfer' auto lemma coeffs_subset: "coeffs (Const c) \ {c}" by (auto simp: coeffs_Const) lemma keys_Const\<^sub>0: "keys (Const\<^sub>0 c) = (if c = 0 then {} else {0})" unfolding Const\<^sub>0_def by transfer' auto lemma vars_Const [simp]: "vars (Const c) = {}" unfolding vars_def by transfer' (auto simp: keys_Const\<^sub>0) lemma prod_fun_compose_bij: assumes "bij f" and f: "\x y. f (x + y) = f x + f y" shows "prod_fun m1 m2 (f x) = prod_fun (m1 \ f) (m2 \ f) x" proof - have [simp]: "f x = f y \ x = y" for x y using \bij f\ by (auto dest!: bij_is_inj inj_onD) have "prod_fun (m1 \ f) (m2 \ f) x = Sum_any ((\l. m1 l * Sum_any ((\q. m2 q when f x = l + q) \ f)) \ f)" by (simp add: prod_fun_def f(1) [symmetric] o_def) also have "\ = Sum_any ((\l. m1 l * Sum_any ((\q. m2 q when f x = l + q))))" by (simp only: Sum_any.reindex_cong[OF assms(1) refl, symmetric]) also have "\ = prod_fun m1 m2 (f x)" by (simp add: prod_fun_def) finally show ?thesis .. qed lemma add_nat_poly_mapping_zero_iff [simp]: "(a + b :: 'a \\<^sub>0 nat) = 0 \ a = 0 \ b = 0" by transfer (auto simp: fun_eq_iff) lemma prod_fun_nat_0: fixes f g :: "('a \\<^sub>0 nat) \ 'b::semiring_0" shows "prod_fun f g 0 = f 0 * g 0" proof - have "prod_fun f g 0 = (\l. f l * (\q. g q when 0 = l + q))" unfolding prod_fun_def .. also have "(\l. \q. g q when 0 = l + q) = (\l. \q\{0}. g q when 0 = l + q)" by (intro ext Sum_any.expand_superset) (auto simp: when_def) also have "(\l. f l * \ l) = (\l\{0}. f l * \ l)" by (intro ext Sum_any.expand_superset) (auto simp: when_def) finally show ?thesis by simp qed lemma mpoly_coeff_times_0: "coeff (p * q) 0 = coeff p 0 * coeff q 0" by (simp add: coeff_mpoly_times prod_fun_nat_0) lemma mpoly_coeff_prod_0: "coeff (\x\A. f x) 0 = (\x\A. coeff (f x) 0)" by (induction A rule: infinite_finite_induct) (auto simp: mpoly_coeff_times_0 mpoly_coeff_1) lemma mpoly_coeff_power_0: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induction n) (auto simp: mpoly_coeff_times_0 mpoly_coeff_1) lemma prod_fun_max: fixes f g :: "'a::{linorder, ordered_cancel_comm_monoid_add} \ 'b::semiring_0" assumes zero: "\m. m > a \ f m = 0" "\m. m > b \ g m = 0" assumes fin: "finite {m. f m \ 0}" "finite {m. g m \ 0}" shows "prod_fun f g (a + b) = f a * g b" proof - note fin' = finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)] have "prod_fun f g (a + b) = (\l. f l * (\q. g q when a + b = l + q))" by (simp add: prod_fun_def Sum_any_right_distrib) also have "\ = (\l. \q. f l * g q when a + b = l + q)" by (subst Sum_any_right_distrib) (auto intro!: Sum_any.cong fin'(2) simp: when_def) also { fix l q assume lq: "a + b = l + q" "(a, b) \ (l, q)" and nz: "f l * g q \ 0" from nz and zero have "l \ a" "q \ b" by (auto intro: leI) moreover from this and lq(2) have "l < a \ q < b" by auto ultimately have "l + q < a + b" by (auto intro: add_less_le_mono add_le_less_mono) with lq(1) have False by simp } hence "(\l. \q. f l * g q when a + b = l + q) = (\l. \q. f l * g q when (a, b) = (l, q))" by (intro Sum_any.cong refl) (auto simp: when_def) also have "\ = (\(l,q). f l * g q when (a, b) = (l, q))" by (intro Sum_any.cartesian_product[of "{(a, b)}"]) auto also have "\ = (\(l,q)\{(a,b)}. f l * g q when (a, b) = (l, q))" by (intro Sum_any.expand_superset) auto also have "\ = f a * g b" by simp finally show ?thesis . qed lemma prod_fun_gt_max_eq_zero: fixes f g :: "'a::{linorder, ordered_cancel_comm_monoid_add} \ 'b::semiring_0" assumes "m > a + b" assumes zero: "\m. m > a \ f m = 0" "\m. m > b \ g m = 0" assumes fin: "finite {m. f m \ 0}" "finite {m. g m \ 0}" shows "prod_fun f g m = 0" proof - note fin' = finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)] have "prod_fun f g m = (\l. f l * (\q. g q when m = l + q))" by (simp add: prod_fun_def Sum_any_right_distrib) also have "\ = (\l. \q. f l * g q when m = l + q)" by (subst Sum_any_right_distrib) (auto intro!: Sum_any.cong fin'(2) simp: when_def) also { fix l q assume lq: "m = l + q" and nz: "f l * g q \ 0" from nz and zero have "l \ a" "q \ b" by (auto intro: leI) hence "l + q \ a + b" by (intro add_mono) also have "\ < m" by fact finally have "l + q < m" . } hence "(\l. \q. f l * g q when m = l + q) = (\l. \q. f l * g q when False)" by (intro Sum_any.cong refl) (auto simp: when_def) also have "\ = 0" by simp finally show ?thesis . qed subsection \Restricting a monomial to a subset of variables\ lift_definition restrictpm :: "'a set \ ('a \\<^sub>0 'b :: zero) \ ('a \\<^sub>0 'b)" is "\A f x. if x \ A then f x else 0" by (erule finite_subset[rotated]) auto lemma lookup_restrictpm: "lookup (restrictpm A m) x = (if x \ A then lookup m x else 0)" by transfer auto lemma lookup_restrictpm_in [simp]: "x \ A \ lookup (restrictpm A m) x = lookup m x" and lookup_restrict_pm_not_in [simp]: "x \ A \ lookup (restrictpm A m) x = 0" by (simp_all add: lookup_restrictpm) lemma keys_restrictpm [simp]: "keys (restrictpm A m) = keys m \ A" by transfer auto lemma restrictpm_add: "restrictpm X (m1 + m2) = restrictpm X m1 + restrictpm X m2" by transfer auto lemma restrictpm_id [simp]: "keys m \ X \ restrictpm X m = m" by transfer (auto simp: fun_eq_iff) lemma restrictpm_orthogonal [simp]: "keys m \ -X \ restrictpm X m = 0" by transfer (auto simp: fun_eq_iff) lemma restrictpm_add_disjoint: "X \ Y = {} \ restrictpm X m + restrictpm Y m = restrictpm (X \ Y) m" by transfer (auto simp: fun_eq_iff) lemma restrictpm_add_complements: "restrictpm X m + restrictpm (-X) m = m" "restrictpm (-X) m + restrictpm X m = m" by (subst restrictpm_add_disjoint; force)+ subsection \Mapping over a polynomial\ lift_definition map_mpoly :: "('a :: zero \ 'b :: zero) \ 'a mpoly \ 'b mpoly" is "\(f :: 'a \ 'b) (p :: (nat \\<^sub>0 nat) \\<^sub>0 'a). Poly_Mapping.map f p" . lift_definition mapm_mpoly :: "((nat \\<^sub>0 nat) \ 'a :: zero \ 'b :: zero) \ 'a mpoly \ 'b mpoly" is "\(f :: (nat \\<^sub>0 nat) \ 'a \ 'b) (p :: (nat \\<^sub>0 nat) \\<^sub>0 'a). Poly_Mapping.mapp f p" . lemma poly_mapping_map_conv_mapp: "Poly_Mapping.map f = Poly_Mapping.mapp (\_. f)" by (auto simp: Poly_Mapping.mapp_def Poly_Mapping.map_def map_fun_def o_def fun_eq_iff when_def in_keys_iff cong: if_cong) lemma map_mpoly_conv_mapm_mpoly: "map_mpoly f = mapm_mpoly (\_. f)" by transfer' (auto simp: poly_mapping_map_conv_mapp) lemma map_mpoly_comp: "f 0 = 0 \ map_mpoly f (map_mpoly g p) = map_mpoly (f \ g) p" by (transfer', transfer') (auto simp: when_def fun_eq_iff) lemma mapp_mapp: "(\x. f x 0 = 0) \ Poly_Mapping.mapp f (Poly_Mapping.mapp g m) = Poly_Mapping.mapp (\x y. f x (g x y)) m" by transfer' (auto simp: fun_eq_iff lookup_mapp in_keys_iff) lemma mapm_mpoly_comp: "(\x. f x 0 = 0) \ mapm_mpoly f (mapm_mpoly g p) = mapm_mpoly (\m c. f m (g m c)) p" by transfer' (simp add: mapp_mapp) lemma coeff_map_mpoly: "coeff (map_mpoly f p) m = (if coeff p m = 0 then 0 else f (coeff p m))" by (transfer, transfer) auto lemma coeff_map_mpoly' [simp]: "f 0 = 0 \ coeff (map_mpoly f p) m = f (coeff p m)" by (subst coeff_map_mpoly) auto lemma coeff_mapm_mpoly: "coeff (mapm_mpoly f p) m = (if coeff p m = 0 then 0 else f m (coeff p m))" by (transfer, transfer') (auto simp: in_keys_iff) lemma coeff_mapm_mpoly' [simp]: "(\m. f m 0 = 0) \ coeff (mapm_mpoly f p) m = f m (coeff p m)" by (subst coeff_mapm_mpoly) auto lemma vars_map_mpoly_subset: "vars (map_mpoly f p) \ vars p" unfolding vars_def by (transfer', transfer') (auto simp: map_mpoly.rep_eq) lemma coeff_sum [simp]: "coeff (sum f A) m = (\x\A. coeff (f x) m)" by (induction A rule: infinite_finite_induct) auto lemma coeff_Sum_any: "finite {x. f x \ 0} \ coeff (Sum_any f) m = Sum_any (\x. coeff (f x) m)" by (auto simp add: Sum_any.expand_set intro!: sum.mono_neutral_right) lemma Sum_any_zeroI: "(\x. f x = 0) \ Sum_any f = 0" by (auto simp: Sum_any.expand_set) lemma insertion_Prod_any: "finite {x. g x \ 1} \ insertion f (Prod_any g) = Prod_any (\x. insertion f (g x))" by (auto simp: Prod_any.expand_set insertion_prod intro!: prod.mono_neutral_right) lemma insertion_insertion: "insertion g (insertion k p) = insertion (\x. insertion g (k x)) (map_mpoly (insertion g) p)" (is "?lhs = ?rhs") proof - have "insertion g (insertion k p) = (\x. insertion g (coeff p x) * insertion g (\i. k i ^ lookup x i))" unfolding insertion_altdef[of k p] by (subst insertion_Sum_any) (auto intro: finite_subset[OF _ finite_coeff_support[of p]] simp: insertion_mult) also have "\ = (\x. insertion g (coeff p x) * (\i. insertion g (k i) ^ lookup x i))" proof (intro Sum_any.cong) fix x show "insertion g (coeff p x) * insertion g (\i. k i ^ lookup x i) = insertion g (coeff p x) * (\i. insertion g (k i) ^ lookup x i)" by (subst insertion_Prod_any) (auto simp: insertion_power intro!: finite_subset[OF _ finite_lookup[of x]] Nat.gr0I) qed also have "\ = insertion (\x. insertion g (k x)) (map_mpoly (insertion g) p)" unfolding insertion_altdef[of _ "map_mpoly f p" for f] by auto finally show ?thesis . qed lemma insertion_substitute_linear: "insertion (\i. c i * f i) p = insertion f (mapm_mpoly (\m d. Prod_any (\i. c i ^ lookup m i) * d) p)" unfolding insertion_altdef proof (intro Sum_any.cong, goal_cases) case (1 m) have "coeff (mapm_mpoly (\m. (*) (\i. c i ^ lookup m i)) p) m * (\i. f i ^ lookup m i) = MPoly_Type.coeff p m * ((\i. c i ^ lookup m i) * (\i. f i ^ lookup m i))" by (simp add: mult_ac) also have "(\i. c i ^ lookup m i) * (\i. f i ^ lookup m i) = (\i. (c i * f i) ^ lookup m i)" by (subst Prod_any.distrib [symmetric]) (auto simp: power_mult_distrib intro!: finite_subset[OF _ finite_lookup[of m]] Nat.gr0I) finally show ?case by simp qed lemma vars_mapm_mpoly_subset: "vars (mapm_mpoly f p) \ vars p" unfolding vars_def using keys_mapp_subset[of f] by (auto simp: mapm_mpoly.rep_eq) lemma map_mpoly_cong: assumes "\m. f (coeff p m) = g (coeff p m)" "p = q" shows "map_mpoly f p = map_mpoly g q" using assms by (intro mpoly_eqI) (auto simp: coeff_map_mpoly) subsection \The leading monomial and leading coefficient\ text \ The leading monomial of a multivariate polynomial is the one with the largest monomial w.\,r.\,t.\ the monomial ordering induced by the standard variable ordering. The leading coefficient is the coefficient of the leading monomial. As a convention, the leading monomial of the zero polynomial is defined to be the same as that of any non-constant zero polynomial, i.\,e.\ the monomial $X_1^0 \ldots X_n^0$. \ lift_definition lead_monom :: "'a :: zero mpoly \ (nat \\<^sub>0 nat)" is "\f :: (nat \\<^sub>0 nat) \\<^sub>0 'a. Max (insert 0 (keys f))" . lemma lead_monom_geI [intro]: assumes "coeff p m \ 0" shows "m \ lead_monom p" using assms by (auto simp: lead_monom_def coeff_def in_keys_iff) lemma coeff_gt_lead_monom_zero [simp]: assumes "m > lead_monom p" shows "coeff p m = 0" using lead_monom_geI[of p m] assms by force lemma lead_monom_nonzero_eq: assumes "p \ 0" shows "lead_monom p = Max (keys (mapping_of p))" using assms by transfer (simp add: max_def) lemma lead_monom_0 [simp]: "lead_monom 0 = 0" by (simp add: lead_monom_def zero_mpoly.rep_eq) lemma lead_monom_1 [simp]: "lead_monom 1 = 0" by (simp add: lead_monom_def one_mpoly.rep_eq) lemma lead_monom_Const [simp]: "lead_monom (Const c) = 0" by (simp add: lead_monom_def Const.rep_eq Const\<^sub>0_def) lemma lead_monom_uminus [simp]: "lead_monom (-p) = lead_monom p" by (simp add: lead_monom_def uminus_mpoly.rep_eq) lemma keys_mult_const [simp]: fixes c :: "'a :: {semiring_0, semiring_no_zero_divisors}" assumes "c \ 0" shows "keys (Poly_Mapping.map ((*) c) p) = keys p" using assms by transfer auto lemma lead_monom_eq_0_iff: "lead_monom p = 0 \ vars p = {}" unfolding vars_def by transfer' (auto simp: Max_eq_iff) lemma lead_monom_monom: "lead_monom (monom m c) = (if c = 0 then 0 else m)" by (auto simp add: lead_monom_def monom.rep_eq Const\<^sub>0_def max_def ) lemma lead_monom_monom' [simp]: "c \ 0 \ lead_monom (monom m c) = m" by (simp add: lead_monom_monom) lemma lead_monom_numeral [simp]: "lead_monom (numeral n) = 0" unfolding monom_numeral[symmetric] by (subst lead_monom_monom) auto lemma lead_monom_add: "lead_monom (p + q) \ max (lead_monom p) (lead_monom q)" proof transfer fix p q :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "Max (insert 0 (keys (p + q))) \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof (rule Max.boundedI) fix m assume m: "m \ insert 0 (keys (p + q))" thus "m \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof assume "m \ keys (p + q)" with keys_add[of p q] have "m \ keys p \ m \ keys q" by (auto simp: in_keys_iff plus_poly_mapping.rep_eq) thus ?thesis by (auto simp: le_max_iff_disj) qed auto qed auto qed lemma lead_monom_diff: "lead_monom (p - q) \ max (lead_monom p) (lead_monom q)" proof transfer fix p q :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "Max (insert 0 (keys (p - q))) \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof (rule Max.boundedI) fix m assume m: "m \ insert 0 (keys (p - q))" thus "m \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof assume "m \ keys (p - q)" with keys_diff_subset[of p q] have "m \ keys p \ m \ keys q" by auto thus ?thesis by (auto simp: le_max_iff_disj) qed auto qed auto qed definition lead_coeff where "lead_coeff p = coeff p (lead_monom p)" lemma vars_empty_iff: "vars p = {} \ p = Const (lead_coeff p)" proof assume "vars p = {}" hence [simp]: "lead_monom p = 0" by (simp add: lead_monom_eq_0_iff) have [simp]: "mon \ 0 \ (mon > (0 :: nat \\<^sub>0 nat))" for mon by (auto simp: order.strict_iff_order) thus "p = Const (lead_coeff p)" by (intro mpoly_eqI) (auto simp: mpoly_coeff_Const lead_coeff_def) next assume "p = Const (lead_coeff p)" also have "vars \ = {}" by simp finally show "vars p = {}" . qed lemma lead_coeff_0 [simp]: "lead_coeff 0 = 0" by (simp add: lead_coeff_def) lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by (simp add: lead_coeff_def mpoly_coeff_1) lemma lead_coeff_Const [simp]: "lead_coeff (Const c) = c" by (simp add: lead_coeff_def mpoly_coeff_Const) lemma lead_coeff_monom [simp]: "lead_coeff (monom p c) = c" by (simp add: lead_coeff_def coeff_monom when_def lead_monom_monom) lemma lead_coeff_nonzero [simp]: "p \ 0 \ lead_coeff p \ 0" unfolding lead_coeff_def lead_monom_def by (cases "keys (mapping_of p) = {}") (auto simp: coeff_def max_def) lemma fixes c :: "'a :: semiring_0" assumes "c * lead_coeff p \ 0" shows lead_monom_smult [simp]: "lead_monom (smult c p) = lead_monom p" and lead_coeff_smult [simp]: "lead_coeff (smult c p) = c * lead_coeff p" proof - from assms have *: "keys (mapping_of p) \ {}" by auto from assms have "coeff (MPoly_Type.smult c p) (lead_monom p) \ 0" by (simp add: lead_coeff_def) hence smult_nz: "MPoly_Type.smult c p \ 0" by (auto simp del: coeff_smult) with assms have **: "keys (mapping_of (smult c p)) \ {}" by simp have "Max (keys (mapping_of (smult c p))) = Max (keys (mapping_of p))" proof (safe intro!: antisym Max.coboundedI) have "lookup (mapping_of p) (Max (keys (mapping_of p))) = lead_coeff p" using * by (simp add: lead_coeff_def lead_monom_def max_def coeff_def) with assms show "Max (keys (mapping_of p)) \ keys (mapping_of (smult c p))" using * by (auto simp: smult.rep_eq intro!: in_keys_mapI) from smult_nz have "lead_coeff (smult c p) \ 0" by (intro lead_coeff_nonzero) auto hence "coeff p (Max (keys (mapping_of (smult c p)))) \ 0" using assms * ** by (auto simp: lead_coeff_def lead_monom_def max_def) thus "Max (keys (mapping_of (smult c p))) \ keys (mapping_of p)" by (auto simp: smult.rep_eq coeff_def in_keys_iff) qed auto with * ** show "lead_monom (smult c p) = lead_monom p" by (simp add: lead_monom_def max_def) thus "lead_coeff (smult c p) = c * lead_coeff p" by (simp add: lead_coeff_def) qed lemma lead_coeff_mult_aux: "coeff (p * q) (lead_monom p + lead_monom q) = lead_coeff p * lead_coeff q" proof (cases "p = 0 \ q = 0") case False define a b where "a = lead_monom p" and "b = lead_monom q" have "coeff (p * q) (a + b) = coeff p a * coeff q b" unfolding coeff_mpoly_times by (rule prod_fun_max) (insert False, auto simp: a_def b_def) thus ?thesis by (simp add: a_def b_def lead_coeff_def) qed auto lemma lead_monom_mult_le: "lead_monom (p * q) \ lead_monom p + lead_monom q" proof (cases "p * q = 0") case False show ?thesis proof (intro leI notI) assume "lead_monom p + lead_monom q < lead_monom (p * q)" hence "lead_coeff (p * q) = 0" unfolding lead_coeff_def coeff_mpoly_times by (rule prod_fun_gt_max_eq_zero) auto with False show False by simp qed qed auto lemma lead_monom_mult: assumes "lead_coeff p * lead_coeff q \ 0" shows "lead_monom (p * q) = lead_monom p + lead_monom q" by (intro antisym lead_monom_mult_le lead_monom_geI) (insert assms, auto simp: lead_coeff_mult_aux) lemma lead_coeff_mult: assumes "lead_coeff p * lead_coeff q \ 0" shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q" using assms by (simp add: lead_monom_mult lead_coeff_mult_aux lead_coeff_def) lemma keys_lead_monom_subset: "keys (lead_monom p) \ vars p" proof (cases "p = 0") case False hence "lead_coeff p \ 0" by simp hence "coeff p (lead_monom p) \ 0" unfolding lead_coeff_def . thus ?thesis unfolding vars_def by transfer' (auto simp: max_def in_keys_iff) qed auto lemma assumes "(\i\A. lead_coeff (f i)) \ 0" shows lead_monom_prod: "lead_monom (\i\A. f i) = (\i\A. lead_monom (f i))" (is ?th1) and lead_coeff_prod: "lead_coeff (\i\A. f i) = (\i\A. lead_coeff (f i))" (is ?th2) proof - have "?th1 \ ?th2" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) from insert have nz: "lead_coeff (f x) \ 0" "(\i\A. lead_coeff (f i)) \ 0" by auto note IH = insert.IH[OF this(2)] from insert have nz': "lead_coeff (f x) * lead_coeff (\i\A. f i) \ 0" by (subst IH) auto from insert.prems insert.hyps nz nz' show ?case by (auto simp: lead_monom_mult lead_coeff_mult IH) qed auto thus ?th1 ?th2 by blast+ qed lemma lead_monom_sum_le: "(\x. x \ X \ lead_monom (h x) \ ub) \ lead_monom (sum h X) \ ub" by (induction X rule: infinite_finite_induct) (auto intro!: order.trans[OF lead_monom_add]) text \ The leading monomial of a sum where the leading monomial the summands are distinct is simply the maximum of the leading monomials. \ lemma lead_monom_sum: assumes "inj_on (lead_monom \ h) X" and "finite X" and "X \ {}" and "\x. x \ X \ h x \ 0" defines "m \ Max ((lead_monom \ h) ` X)" shows "lead_monom (\x\X. h x) = m" proof (rule antisym) show "lead_monom (sum h X) \ m" unfolding m_def using assms by (intro lead_monom_sum_le Max_ge finite_imageI) auto next from assms have "m \ (lead_monom \ h) ` X" unfolding m_def by (intro Max_in finite_imageI) auto then obtain x where x: "x \ X" "m = lead_monom (h x)" by auto have "coeff (\x\X. h x) m = (\x\X. coeff (h x) m)" by simp also have "\ = (\x\{x}. coeff (h x) m)" proof (intro sum.mono_neutral_right ballI) fix y assume y: "y \ X - {x}" hence "(lead_monom \ h) y \ m" using assms unfolding m_def by (intro Max_ge finite_imageI) auto moreover have "(lead_monom \ h) y \ (lead_monom \ h) x" using \x \ X\ y inj_onD[OF assms(1), of x y] by auto ultimately have "lead_monom (h y) < m" using x by auto thus "coeff (h y) m = 0" by simp qed (insert x assms, auto) also have "\ = coeff (h x) m" by simp also have "\ = lead_coeff (h x)" using x by (simp add: lead_coeff_def) also have "\ \ 0" using assms x by auto finally show "lead_monom (sum h X) \ m" by (intro lead_monom_geI) qed lemma lead_coeff_eq_0_iff [simp]: "lead_coeff p = 0 \ p = 0" by (cases "p = 0") auto lemma fixes f :: "_ \ 'a :: semidom mpoly" assumes "\i. i \ A \ f i \ 0" shows lead_monom_prod' [simp]: "lead_monom (\i\A. f i) = (\i\A. lead_monom (f i))" (is ?th1) and lead_coeff_prod' [simp]: "lead_coeff (\i\A. f i) = (\i\A. lead_coeff (f i))" (is ?th2) proof - from assms have "(\i\A. lead_coeff (f i)) \ 0" by (cases "finite A") auto thus ?th1 ?th2 by (simp_all add: lead_monom_prod lead_coeff_prod) qed lemma fixes p :: "'a :: comm_semiring_1 mpoly" assumes "lead_coeff p ^ n \ 0" shows lead_monom_power: "lead_monom (p ^ n) = of_nat n * lead_monom p" and lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" using assms lead_monom_prod[of "\_. p" "{.._. p" "{.. 0" shows lead_monom_power' [simp]: "lead_monom (p ^ n) = of_nat n * lead_monom p" and lead_coeff_power' [simp]: "lead_coeff (p ^ n) = lead_coeff p ^ n" using assms lead_monom_prod'[of "{.._. p"] lead_coeff_prod'[of "{.._. p"] by simp_all subsection \Turning a set of variables into a monomial\ text \ Given a finite set $\{X_1,\ldots,X_n\}$ of variables, the following is the monomial $X_1\ldots X_n$: \ lift_definition monom_of_set :: "nat set \ (nat \\<^sub>0 nat)" is "\X x. if finite X \ x \ X then 1 else 0" by auto lemma lookup_monom_of_set: "Poly_Mapping.lookup (monom_of_set X) i = (if finite X \ i \ X then 1 else 0)" by transfer auto lemma lookup_monom_of_set_1 [simp]: "finite X \ i \ X \ Poly_Mapping.lookup (monom_of_set X) i = 1" and lookup_monom_of_set_0 [simp]: "i \ X \ Poly_Mapping.lookup (monom_of_set X) i = 0" by (simp_all add: lookup_monom_of_set) lemma keys_monom_of_set: "keys (monom_of_set X) = (if finite X then X else {})" by transfer auto lemma keys_monom_of_set_finite [simp]: "finite X \ keys (monom_of_set X) = X" by (simp add: keys_monom_of_set) lemma monom_of_set_eq_iff [simp]: "finite X \ finite Y \ monom_of_set X = monom_of_set Y \ X = Y" by transfer (auto simp: fun_eq_iff) lemma monom_of_set_empty [simp]: "monom_of_set {} = 0" by transfer auto lemma monom_of_set_eq_zero_iff [simp]: "monom_of_set X = 0 \ infinite X \ X = {}" by transfer (auto simp: fun_eq_iff) lemma zero_eq_monom_of_set_iff [simp]: "0 = monom_of_set X \ infinite X \ X = {}" by transfer (auto simp: fun_eq_iff) subsection \Permuting the variables of a polynomial\ text \ Next, we define the operation of permuting the variables of a monomial and polynomial. \ lift_definition permutep :: "('a \ 'a) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b :: zero)" is "\f p. if bij f then p \ f else p" proof - fix f :: "'a \ 'a" and g :: "'a \ 'b" assume *: "finite {x. g x \ 0}" show "finite {x. (if bij f then g \ f else g) x \ 0}" proof (cases "bij f") case True with * have "finite (f -` {x. g x \ 0})" by (intro finite_vimageI) (auto dest: bij_is_inj) with True show ?thesis by auto qed (use * in auto) qed lift_definition mpoly_map_vars :: "(nat \ nat) \ 'a :: zero mpoly \ 'a mpoly" is "\f p. permutep (permutep f) p" . lemma keys_permutep: "bij f \ keys (permutep f m) = f -` keys m" by transfer auto lemma permutep_id'' [simp]: "permutep id = id" by transfer' (auto simp: fun_eq_iff) lemma permutep_id''' [simp]: "permutep (\x. x) = id" by transfer' (auto simp: fun_eq_iff) lemma permutep_0 [simp]: "permutep f 0 = 0" by transfer auto lemma permutep_single: "bij f \ permutep f (Poly_Mapping.single a b) = Poly_Mapping.single (inv_into UNIV f a) b" by transfer (auto simp: fun_eq_iff when_def inv_f_f surj_f_inv_f bij_is_inj bij_is_surj) lemma mpoly_map_vars_id [simp]: "mpoly_map_vars id = id" by transfer auto lemma mpoly_map_vars_id' [simp]: "mpoly_map_vars (\x. x) = id" by transfer auto lemma lookup_permutep: "Poly_Mapping.lookup (permutep f m) x = (if bij f then Poly_Mapping.lookup m (f x) else Poly_Mapping.lookup m x)" by transfer auto lemma inj_permutep [intro]: "inj (permutep (f :: 'a \ 'a) :: _ \ 'a \\<^sub>0 'b :: zero)" unfolding inj_def proof (transfer, safe) fix f :: "'a \ 'a" and x y :: "'a \ 'b" assume eq: "(if bij f then x \ f else x) = (if bij f then y \ f else y)" show "x = y" proof (cases "bij f") case True show ?thesis proof fix t :: 'a from \bij f\ obtain s where "t = f s" by (auto dest!: bij_is_surj) with eq and True show "x t = y t" by (auto simp: fun_eq_iff) qed qed (use eq in auto) qed lemma surj_permutep [intro]: "surj (permutep (f :: 'a \ 'a) :: _ \ 'a \\<^sub>0 'b :: zero)" unfolding surj_def proof (transfer, safe) fix f :: "'a \ 'a" and y :: "'a \ 'b" assume fin: "finite {t. y t \ 0}" show "\x\{f. finite {x. f x \ 0}}. y = (if bij f then x \ f else x)" proof (cases "bij f") case True with fin have "finite (the_inv f -` {t. y t \ 0})" by (intro finite_vimageI) (auto simp: bij_is_inj bij_betw_the_inv_into) moreover have "y \ the_inv f \ f = y" using True by (simp add: fun_eq_iff the_inv_f_f bij_is_inj) ultimately show ?thesis by (intro bexI[of _ "y \ the_inv f"]) (auto simp: True) qed (use fin in auto) qed lemma bij_permutep [intro]: "bij (permutep f)" using inj_permutep[of f] surj_permutep[of f] by (simp add: bij_def) lemma mpoly_map_vars_map_mpoly: "mpoly_map_vars f (map_mpoly g p) = map_mpoly g (mpoly_map_vars f p)" by (transfer', transfer') (auto simp: fun_eq_iff) lemma coeff_mpoly_map_vars: fixes f :: "nat \ nat" and p :: "'a :: zero mpoly" assumes "bij f" shows "MPoly_Type.coeff (mpoly_map_vars f p) mon = MPoly_Type.coeff p (permutep f mon)" using assms by transfer' (simp add: lookup_permutep bij_permutep) lemma permutep_monom_of_set: assumes "bij f" shows "permutep f (monom_of_set A) = monom_of_set (f -` A)" using assms by transfer (auto simp: fun_eq_iff bij_is_inj finite_vimage_iff) lemma permutep_comp: "bij f \ bij g \ permutep (f \ g) = permutep g \ permutep f" by transfer' (auto simp: fun_eq_iff bij_comp) lemma permutep_comp': "bij f \ bij g \ permutep (f \ g) mon = permutep g (permutep f mon)" by transfer (auto simp: fun_eq_iff bij_comp) lemma mpoly_map_vars_comp: "bij f \ bij g \ mpoly_map_vars f (mpoly_map_vars g p) = mpoly_map_vars (f \ g) p" by transfer (auto simp: bij_permutep permutep_comp) lemma permutep_id [simp]: "permutep id mon = mon" by transfer auto lemma permutep_id' [simp]: "permutep (\x. x) mon = mon" by transfer auto lemma inv_permutep [simp]: fixes f :: "'a \ 'a" assumes "bij f" shows "inv_into UNIV (permutep f) = permutep (inv_into UNIV f)" proof fix m :: "'a \\<^sub>0 'b" show "inv_into UNIV (permutep f) m = permutep (inv_into UNIV f) m" using permutep_comp'[of "inv_into UNIV f" f m] assms inj_iff[of f] by (intro inv_f_eq) (auto simp: bij_imp_bij_inv bij_is_inj) qed lemma mpoly_map_vars_monom: "bij f \ mpoly_map_vars f (monom m c) = monom (permutep (inv_into UNIV f) m) c" by transfer' (simp add: permutep_single bij_permutep) lemma vars_mpoly_map_vars: fixes f :: "nat \ nat" and p :: "'a :: zero mpoly" assumes "bij f" shows "vars (mpoly_map_vars f p) = f ` vars p" using assms unfolding vars_def proof transfer' fix f :: "nat \ nat" and p :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" assume f: "bij f" have eq: "f (inv_into UNIV f x) = x" for x using f by (subst surj_f_inv_f[of f]) (auto simp: bij_is_surj) show "\ (keys ` keys (permutep (permutep f) p)) = f ` \ (keys ` keys p)" proof safe fix m x assume mx: "m \ keys (permutep (permutep f) p)" "x \ keys m" from mx have "permutep f m \ keys p" by (auto simp: keys_permutep bij_permutep f) with mx have "f (inv_into UNIV f x) \ f ` (\m\keys p. keys m)" by (intro imageI) (auto intro!: bexI[of _ "permutep f m"] simp: keys_permutep f eq) with eq show "x \ f ` (\m\keys p. keys m)" by simp next fix m x assume mx: "m \ keys p" "x \ keys m" from mx have "permutep id m \ keys p" by simp also have "id = inv_into UNIV f \ f" using f by (intro ext) (auto simp: bij_is_inj inv_f_f) also have "permutep \ m = permutep f (permutep (inv_into UNIV f) m)" by (simp add: permutep_comp f bij_imp_bij_inv) finally have **: "permutep f (permutep (inv_into UNIV f) m) \ keys p" . moreover from f mx have "f x \ keys (permutep (inv_into UNIV f) m)" by (auto simp: keys_permutep bij_imp_bij_inv inv_f_f bij_is_inj) ultimately show "f x \ \ (keys ` keys (permutep (permutep f) p))" using f by (auto simp: keys_permutep bij_permutep) qed qed lemma permutep_eq_monom_of_set_iff [simp]: assumes "bij f" shows "permutep f mon = monom_of_set A \ mon = monom_of_set (f ` A)" proof assume eq: "permutep f mon = monom_of_set A" have "permutep (inv_into UNIV f) (permutep f mon) = monom_of_set (inv_into UNIV f -` A)" using assms by (simp add: eq bij_imp_bij_inv assms permutep_monom_of_set) also have "inv_into UNIV f -` A = f ` A" using assms by (force simp: bij_is_surj image_iff inv_f_f bij_is_inj surj_f_inv_f) also have "permutep (inv_into UNIV f) (permutep f mon) = permutep (f \ inv_into UNIV f) mon" using assms by (simp add: permutep_comp bij_imp_bij_inv) also have "f \ inv_into UNIV f = id" by (subst surj_iff [symmetric]) (insert assms, auto simp: bij_is_surj) finally show "mon = monom_of_set (f ` A)" by simp qed (insert assms, auto simp: permutep_monom_of_set inj_vimage_image_eq bij_is_inj) lemma permutep_monom_of_set_permutes [simp]: assumes "\ permutes A" shows "permutep \ (monom_of_set A) = monom_of_set A" using assms by transfer (auto simp: if_splits fun_eq_iff permutes_in_image) lemma mpoly_map_vars_0 [simp]: "mpoly_map_vars f 0 = 0" by (transfer, transfer') (simp add: o_def) lemma permutep_eq_0_iff [simp]: "permutep f m = 0 \ m = 0" proof transfer fix f :: "'a \ 'a" and m :: "'a \ 'b" assume "finite {x. m x \ 0}" show "((if bij f then m \ f else m) = (\k. 0)) = (m = (\k. 0))" proof (cases "bij f") case True hence "(\x. m (f x) = 0) \ (\x. m x = 0)" using bij_iff[of f] by metis with True show ?thesis by (auto simp: fun_eq_iff) qed (auto simp: fun_eq_iff) qed lemma mpoly_map_vars_1 [simp]: "mpoly_map_vars f 1 = 1" by (transfer, transfer') (auto simp: o_def fun_eq_iff when_def) lemma permutep_Const\<^sub>0 [simp]: "(\x. f x = 0 \ x = 0) \ permutep f (Const\<^sub>0 c) = Const\<^sub>0 c" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_add [simp]: "permutep f (m1 + m2) = permutep f m1 + permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_diff [simp]: "permutep f (m1 - m2) = permutep f m1 - permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_uminus [simp]: "permutep f (-m) = -permutep f m" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_mult [simp]: "(\x y. f (x + y) = f x + f y) \ permutep f (m1 * m2) = permutep f m1 * permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff prod_fun_compose_bij) lemma mpoly_map_vars_Const [simp]: "mpoly_map_vars f (Const c) = Const c" by transfer (auto simp: o_def fun_eq_iff when_def) lemma mpoly_map_vars_add [simp]: "mpoly_map_vars f (p + q) = mpoly_map_vars f p + mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_diff [simp]: "mpoly_map_vars f (p - q) = mpoly_map_vars f p - mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_uminus [simp]: "mpoly_map_vars f (-p) = -mpoly_map_vars f p" by transfer simp lemma permutep_smult: "permutep (permutep f) (Poly_Mapping.map ((*) c) p) = Poly_Mapping.map ((*) c) (permutep (permutep f) p)" by transfer' (auto split: if_splits simp: fun_eq_iff) lemma mpoly_map_vars_smult [simp]: "mpoly_map_vars f (smult c p) = smult c (mpoly_map_vars f p)" by transfer (simp add: permutep_smult) lemma mpoly_map_vars_mult [simp]: "mpoly_map_vars f (p * q) = mpoly_map_vars f p * mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_sum [simp]: "mpoly_map_vars f (sum g A) = (\x\A. mpoly_map_vars f (g x))" by (induction A rule: infinite_finite_induct) auto lemma mpoly_map_vars_prod [simp]: "mpoly_map_vars f (prod g A) = (\x\A. mpoly_map_vars f (g x))" by (induction A rule: infinite_finite_induct) auto lemma mpoly_map_vars_eq_0_iff [simp]: "mpoly_map_vars f p = 0 \ p = 0" by transfer auto lemma permutep_eq_iff [simp]: "permutep f p = permutep f q \ p = q" by transfer (auto simp: comp_bij_eq_iff) lemma mpoly_map_vars_Sum_any [simp]: "mpoly_map_vars f (Sum_any g) = Sum_any (\x. mpoly_map_vars f (g x))" by (simp add: Sum_any.expand_set) lemma mpoly_map_vars_power [simp]: "mpoly_map_vars f (p ^ n) = mpoly_map_vars f p ^ n" by (induction n) auto lemma mpoly_map_vars_monom_single [simp]: assumes "bij f" shows "mpoly_map_vars f (monom (Poly_Mapping.single i n) c) = monom (Poly_Mapping.single (f i) n) c" using assms by (simp add: mpoly_map_vars_monom permutep_single bij_imp_bij_inv inv_inv_eq) lemma insertion_mpoly_map_vars: assumes "bij f" shows "insertion g (mpoly_map_vars f p) = insertion (g \ f) p" proof - have "insertion g (mpoly_map_vars f p) = (\m. coeff p (permutep f m) * (\i. g i ^ lookup m i))" using assms by (simp add: insertion_altdef coeff_mpoly_map_vars) also have "\ = Sum_any (\m. coeff p (permutep f m) * Prod_any (\i. g (f i) ^ lookup m (f i)))" by (intro Sum_any.cong arg_cong[where ?f = "\y. x * y" for x] Prod_any.reindex_cong[OF assms]) (auto simp: o_def) also have "\ = Sum_any (\m. coeff p m * (\i. g (f i) ^ lookup m i))" by (intro Sum_any.reindex_cong [OF bij_permutep[of f], symmetric]) (auto simp: o_def lookup_permutep assms) also have "\ = insertion (g \ f) p" by (simp add: insertion_altdef) finally show ?thesis . qed lemma permutep_cong: assumes "f permutes (-keys p)" "g permutes (-keys p)" "p = q" shows "permutep f p = permutep g q" proof (intro poly_mapping_eqI) fix k :: 'a show "lookup (permutep f p) k = lookup (permutep g q) k" proof (cases "k \ keys p") case False with assms have "f k \ keys p" "g k \ keys p" using permutes_in_image[of _ "-keys p" k] by auto thus ?thesis using assms by (auto simp: lookup_permutep permutes_bij in_keys_iff) qed (insert assms, auto simp: lookup_permutep permutes_bij permutes_not_in) qed lemma mpoly_map_vars_cong: assumes "f permutes (-vars p)" "g permutes (-vars q)" "p = q" shows "mpoly_map_vars f p = mpoly_map_vars g (q :: 'a :: zero mpoly)" proof (intro mpoly_eqI) fix mon :: "nat \\<^sub>0 nat" show "coeff (mpoly_map_vars f p) mon = coeff (mpoly_map_vars g q) mon" proof (cases "keys mon \ vars p") case True with assms have "permutep f mon = permutep g mon" by (intro permutep_cong assms(1,2)[THEN permutes_subset]) auto thus ?thesis using assms by (simp add: coeff_mpoly_map_vars permutes_bij) next case False hence "\(keys mon \ f ` vars q)" "\(keys mon \ g ` vars q)" using assms by (auto simp: subset_iff permutes_not_in) thus ?thesis using assms by (subst (1 2) coeff_notin_vars) (auto simp: coeff_notin_vars vars_mpoly_map_vars permutes_bij) qed qed subsection \Symmetric polynomials\ text \ A polynomial is symmetric on a set of variables if it is invariant under any permutation of that set. \ definition symmetric_mpoly :: "nat set \ 'a :: zero mpoly \ bool" where "symmetric_mpoly A p = (\\. \ permutes A \ mpoly_map_vars \ p = p)" lemma symmetric_mpoly_empty [simp, intro]: "symmetric_mpoly {} p" by (simp add: symmetric_mpoly_def) text \ A polynomial is trivially symmetric on any set of variables that do not occur in it. \ lemma symmetric_mpoly_orthogonal: assumes "vars p \ A = {}" shows "symmetric_mpoly A p" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" with assms have "\ x = x" if "x \ vars p" for x using that permutes_not_in[of \ A x] by auto from assms have "mpoly_map_vars \ p = mpoly_map_vars id p" by (intro mpoly_map_vars_cong permutes_subset[OF \] permutes_id) auto also have "\ = p" by simp finally show "mpoly_map_vars \ p = p" . qed lemma symmetric_mpoly_monom [intro]: assumes "keys m \ A = {}" shows "symmetric_mpoly A (monom m c)" using assms vars_monom_subset[of m c] by (intro symmetric_mpoly_orthogonal) auto lemma symmetric_mpoly_subset: assumes "symmetric_mpoly A p" "B \ A" shows "symmetric_mpoly B p" unfolding symmetric_mpoly_def proof safe fix \ assume "\ permutes B" with assms have "\ permutes A" using permutes_subset by blast with assms show "mpoly_map_vars \ p = p" by (auto simp: symmetric_mpoly_def) qed text \ If a polynomial is symmetric over some set of variables, that set must either be a subset of the variables occurring in the polynomial or disjoint from it. \ lemma symmetric_mpoly_imp_orthogonal_or_subset: assumes "symmetric_mpoly A p" shows "vars p \ A = {} \ A \ vars p" proof (rule ccontr) assume "\(vars p \ A = {} \ A \ vars p)" then obtain x y where xy: "x \ vars p \ A" "y \ A - vars p" by auto - define \ where "\ = Fun.swap x y id" + define \ where "\ = transpose x y" from xy have \: "\ permutes A" unfolding \_def by (intro permutes_swap_id) auto - from xy have "y \ \ ` vars p" by (auto simp: \_def Fun.swap_def) + from xy have "y \ \ ` vars p" by (auto simp: \_def transpose_def) also from \ have "\ ` vars p = vars (mpoly_map_vars \ p)" by (auto simp: vars_mpoly_map_vars permutes_bij) also have "mpoly_map_vars \ p = p" using assms \ by (simp add: symmetric_mpoly_def) finally show False using xy by auto qed text \ Symmetric polynomials are closed under ring operations. \ lemma symmetric_mpoly_add [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p + q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_diff [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p - q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_uminus [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (-p)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_uminus_iff [simp]: "symmetric_mpoly A (-p) \ symmetric_mpoly A p" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_smult [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (smult c p)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_mult [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p * q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_0 [simp, intro]: "symmetric_mpoly A 0" and symmetric_mpoly_1 [simp, intro]: "symmetric_mpoly A 1" and symmetric_mpoly_Const [simp, intro]: "symmetric_mpoly A (Const c)" by (simp_all add: symmetric_mpoly_def) lemma symmetric_mpoly_power [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (p ^ n)" by (induction n) (auto intro!: symmetric_mpoly_mult) lemma symmetric_mpoly_sum [intro]: "(\i. i \ B \ symmetric_mpoly A (f i)) \ symmetric_mpoly A (sum f B)" by (induction B rule: infinite_finite_induct) (auto intro!: symmetric_mpoly_add) lemma symmetric_mpoly_prod [intro]: "(\i. i \ B \ symmetric_mpoly A (f i)) \ symmetric_mpoly A (prod f B)" by (induction B rule: infinite_finite_induct) (auto intro!: symmetric_mpoly_mult) text \ An symmetric sum or product over polynomials yields a symmetric polynomial: \ lemma symmetric_mpoly_symmetric_sum: assumes "g permutes X" assumes "\x \. x \ X \ \ permutes A \ mpoly_map_vars \ (f x) = f (g x)" shows "symmetric_mpoly A (\x\X. f x)" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" have "mpoly_map_vars \ (sum f X) = (\x\X. mpoly_map_vars \ (f x))" by simp also have "\ = (\x\X. f (g x))" by (intro sum.cong assms \ refl) also have "\ = (\x\g`X. f x)" using assms by (subst sum.reindex) (auto simp: permutes_inj_on) also have "g ` X = X" using assms by (simp add: permutes_image) finally show "mpoly_map_vars \ (sum f X) = sum f X" . qed lemma symmetric_mpoly_symmetric_prod: assumes "g permutes X" assumes "\x \. x \ X \ \ permutes A \ mpoly_map_vars \ (f x) = f (g x)" shows "symmetric_mpoly A (\x\X. f x)" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" have "mpoly_map_vars \ (prod f X) = (\x\X. mpoly_map_vars \ (f x))" by simp also have "\ = (\x\X. f (g x))" by (intro prod.cong assms \ refl) also have "\ = (\x\g`X. f x)" using assms by (subst prod.reindex) (auto simp: permutes_inj_on) also have "g ` X = X" using assms by (simp add: permutes_image) finally show "mpoly_map_vars \ (prod f X) = prod f X" . qed text \ If $p$ is a polynomial that is symmetric on some subset of variables $A$, then for the leading monomial of $p$, the exponents of these variables are decreasing w.\,r.\,t.\ the variable ordering. \ theorem lookup_lead_monom_decreasing: assumes "symmetric_mpoly A p" defines "m \ lead_monom p" assumes "i \ A" "j \ A" "i \ j" shows "lookup m i \ lookup m j" proof (cases "p = 0") case [simp]: False show ?thesis proof (intro leI notI) assume less: "lookup m i < lookup m j" - define \ where "\ = Fun.swap i j id" + define \ where "\ = transpose i j" from assms have \: "\ permutes A" unfolding \_def by (intro permutes_swap_id) auto have [simp]: "\ \ \ = id" "\ i = j" "\ j = i" "\k. k \ i \ k \ j \ \ k = k" by (auto simp: \_def Fun.swap_def fun_eq_iff) have "0 \ lead_coeff p" by simp also have "lead_coeff p = MPoly_Type.coeff (mpoly_map_vars \ p) (permutep \ m)" using \ by (simp add: lead_coeff_def m_def coeff_mpoly_map_vars permutes_bij permutep_comp' [symmetric]) also have "mpoly_map_vars \ p = p" using \ assms by (simp add: symmetric_mpoly_def) finally have "permutep \ m \ m" by (auto simp: m_def) moreover have "lookup m i < lookup (permutep \ m) i" and "(\k m) k)" using assms \ less by (auto simp: lookup_permutep permutes_bij) hence "m < permutep \ m" by (auto simp: less_poly_mapping_def less_fun_def) ultimately show False by simp qed qed (auto simp: m_def) subsection \The elementary symmetric polynomials\ text \ The $k$-th elementary symmetric polynomial for a finite set of variables $A$, with $k$ ranging between 1 and $|A|$, is the sum of the product of all subsets of $A$ with cardinality $k$: \ lift_definition sym_mpoly_aux :: "nat set \ nat \ (nat \\<^sub>0 nat) \\<^sub>0 'a :: {zero_neq_one}" is "\X k mon. if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0" proof - fix k :: nat and X :: "nat set" show "finite {x. (if finite X \ (\Y\X. card Y = k \ x = monom_of_set Y) then 1 else 0) \ (0::'a)}" (is "finite ?A") proof (cases "finite X") case True have "?A \ monom_of_set ` Pow X" by auto moreover from True have "finite (monom_of_set ` Pow X)" by simp ultimately show ?thesis by (rule finite_subset) qed auto qed lemma lookup_sym_mpoly_aux: "Poly_Mapping.lookup (sym_mpoly_aux X k) mon = (if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0)" by transfer' simp lemma lookup_sym_mpoly_aux_monom_of_set [simp]: assumes "finite X" "Y \ X" "card Y = k" shows "Poly_Mapping.lookup (sym_mpoly_aux X k) (monom_of_set Y) = 1" using assms by (auto simp: lookup_sym_mpoly_aux) lemma keys_sym_mpoly_aux: "m \ keys (sym_mpoly_aux A k) \ keys m \ A" by transfer' (auto split: if_splits simp: keys_monom_of_set) lift_definition sym_mpoly :: "nat set \ nat \ 'a :: {zero_neq_one} mpoly" is "sym_mpoly_aux" . lemma vars_sym_mpoly_subset: "vars (sym_mpoly A k) \ A" using keys_sym_mpoly_aux by (auto simp: vars_def sym_mpoly.rep_eq) lemma coeff_sym_mpoly: "MPoly_Type.coeff (sym_mpoly X k) mon = (if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0)" by transfer' (simp add: lookup_sym_mpoly_aux) lemma sym_mpoly_infinite: "\finite A \ sym_mpoly A k = 0" by (transfer, transfer) auto lemma sym_mpoly_altdef: "sym_mpoly A k = (\X | X \ A \ card X = k. monom (monom_of_set X) 1)" proof (cases "finite A") case False hence *: "infinite {X. X \ A \ infinite X}" by (rule infinite_infinite_subsets) have "infinite {X. X \ A \ card X = 0}" by (rule infinite_super[OF _ *]) auto moreover have **: "infinite {X. X \ A \ finite X \ card X = k}" if "k \ 0" using that infinite_card_subsets[of A k] False by auto have "infinite {X. X \ A \ card X = k}" if "k \ 0" by (rule infinite_super[OF _ **[OF that]]) auto ultimately show ?thesis using False by (cases "k = 0") (simp_all add: sym_mpoly_infinite) next case True show ?thesis proof (intro mpoly_eqI, goal_cases) case (1 m) show ?case proof (cases "\X. X \ A \ card X = k \ m = monom_of_set X") case False thus ?thesis by (auto simp: coeff_sym_mpoly coeff_sum coeff_monom) next case True then obtain X where X: "X \ A" "card X = k" "m = monom_of_set X" by blast have "coeff (\X | X \ A \ card X = k. monom (monom_of_set X) 1) m = (\X\{X}. 1)" unfolding coeff_sum proof (intro sum.mono_neutral_cong_right ballI) fix Y assume Y: "Y \ {X. X \ A \ card X = k} - {X}" hence "X = Y" if "monom_of_set X = monom_of_set Y" using that finite_subset[OF X(1)] finite_subset[of Y A] \finite A\ by auto thus "coeff (monom (monom_of_set Y) 1) m = 0" using X Y by (auto simp: coeff_monom when_def ) qed (insert X \finite A\, auto simp: coeff_monom) thus ?thesis using \finite A\ by (auto simp: coeff_sym_mpoly coeff_sum coeff_monom) qed qed qed lemma coeff_sym_mpoly_monom_of_set [simp]: assumes "finite X" "Y \ X" "card Y = k" shows "MPoly_Type.coeff (sym_mpoly X k) (monom_of_set Y) = 1" using assms by (auto simp: coeff_sym_mpoly) lemma coeff_sym_mpoly_0: "coeff (sym_mpoly X k) 0 = (if finite X \ k = 0 then 1 else 0)" proof - consider "finite X" "k = 0" | "finite X" "k \ 0" | "infinite X" by blast thus ?thesis proof cases assume "finite X" "k = 0" hence "coeff (sym_mpoly X k) (monom_of_set {}) = 1" by (subst coeff_sym_mpoly_monom_of_set) auto thus ?thesis unfolding monom_of_set_empty using \finite X\ \k = 0\ by simp next assume "finite X" "k \ 0" hence "\(\Y. finite Y \ Y \ X \ card Y = k \ monom_of_set Y = 0)" by auto thus ?thesis using \k \ 0\ by (auto simp: coeff_sym_mpoly) next assume "infinite X" thus ?thesis by (simp add: coeff_sym_mpoly) qed qed lemma symmetric_sym_mpoly [intro]: assumes "A \ B" shows "symmetric_mpoly A (sym_mpoly B k :: 'a :: zero_neq_one mpoly)" unfolding symmetric_mpoly_def proof (safe intro!: mpoly_eqI) fix \ and mon :: "nat \\<^sub>0 nat" assume \: "\ permutes A" from \ have \': "\ permutes B" by (rule permutes_subset) fact from \ have "MPoly_Type.coeff (mpoly_map_vars \ (sym_mpoly B k :: 'a mpoly)) mon = MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) (permutep \ mon)" by (simp add: coeff_mpoly_map_vars permutes_bij) also have "\ = 1 \ MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) mon = 1" (is "?lhs = 1 \ ?rhs = 1") proof assume "?rhs = 1" then obtain Y where "finite B" and Y: "Y \ B" "card Y = k" "mon = monom_of_set Y" by (auto simp: coeff_sym_mpoly split: if_splits) with \' have "\ -` Y \ B" "card (\ -` Y) = k" "permutep \ mon = monom_of_set (\ -` Y)" by (auto simp: permutes_in_image card_vimage_inj permutep_monom_of_set permutes_bij permutes_inj permutes_surj) thus "?lhs = 1" using \finite B\ by (auto simp: coeff_sym_mpoly) next assume "?lhs = 1" then obtain Y where "finite B" and Y: "Y \ B" "card Y = k" "permutep \ mon = monom_of_set Y" by (auto simp: coeff_sym_mpoly split: if_splits) from Y(1) have "inj_on \ Y" using inj_on_subset[of \ UNIV Y] \' by (auto simp: permutes_inj) with Y \' have "\ ` Y \ B" "card (\ ` Y) = k" "mon = monom_of_set (\ ` Y)" by (auto simp: permutes_in_image card_image permutep_monom_of_set permutes_bij permutes_inj permutes_surj) thus "?rhs = 1" using \finite B\ by (auto simp: coeff_sym_mpoly) qed hence "?lhs = ?rhs" by (auto simp: coeff_sym_mpoly split: if_splits) finally show "MPoly_Type.coeff (mpoly_map_vars \ (sym_mpoly B k :: 'a mpoly)) mon = MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) mon" . qed lemma insertion_sym_mpoly: assumes "finite X" shows "insertion f (sym_mpoly X k) = (\Y | Y \ X \ card Y = k. prod f Y)" using assms proof (transfer, transfer) fix f :: "nat \ 'a" and k :: nat and X :: "nat set" assume X: "finite X" have "insertion_fun f (\mon. if finite X \ (\Y\X. card Y = k \ mon = monom_of_set Y) then 1 else 0) = (\m. (\v. f v ^ poly_mapping.lookup m v) when (\Y\X. card Y = k \ m = monom_of_set Y))" by (auto simp add: insertion_fun_def X when_def intro!: Sum_any.cong) also have "\ = (\m | \Y\Pow X. card Y = k \ m = monom_of_set Y. (\v. f v ^ poly_mapping.lookup m v) when (\Y\X. card Y = k \ m = monom_of_set Y))" by (rule Sum_any.expand_superset) (use X in auto) also have "\ = (\m | \Y\Pow X. card Y = k \ m = monom_of_set Y. (\v. f v ^ poly_mapping.lookup m v))" by (intro sum.cong) (auto simp: when_def) also have "\ = (\Y | Y \ X \ card Y = k. (\v. f v ^ poly_mapping.lookup (monom_of_set Y) v))" by (rule sum.reindex_bij_witness[of _ monom_of_set keys]) (auto simp: finite_subset[OF _ X]) also have "\ = (\Y | Y \ X \ card Y = k. \v\Y. f v)" proof (intro sum.cong when_cong refl, goal_cases) case (1 Y) hence "finite Y" by (auto dest: finite_subset[OF _ X]) with 1 have "(\v. f v ^ poly_mapping.lookup (monom_of_set Y) v) = (\v::nat. if v \ Y then f v else 1)" by (intro Prod_any.cong) (auto simp: lookup_monom_of_set) also have "\ = (\v\Y. f v)" by (rule Prod_any.conditionalize [symmetric]) fact+ finally show ?case . qed finally show "insertion_fun f (\mon. if finite X \ (\Y\X. card Y = k \ mon = monom_of_set Y) then 1 else 0) = (\Y | Y \ X \ card Y = k. prod f Y)" . qed lemma sym_mpoly_nz [simp]: assumes "finite A" "k \ card A" shows "sym_mpoly A k \ (0 :: 'a :: zero_neq_one mpoly)" proof - from assms obtain B where B: "B \ A" "card B = k" using ex_subset_of_card by blast with assms have "coeff (sym_mpoly A k :: 'a mpoly) (monom_of_set B) = 1" by (intro coeff_sym_mpoly_monom_of_set) thus ?thesis by auto qed lemma coeff_sym_mpoly_0_or_1: "coeff (sym_mpoly A k) m \ {0, 1}" by (transfer, transfer) auto lemma lead_coeff_sym_mpoly [simp]: assumes "finite A" "k \ card A" shows "lead_coeff (sym_mpoly A k) = 1" proof - from assms have "lead_coeff (sym_mpoly A k) \ 0" by simp thus ?thesis using coeff_sym_mpoly_0_or_1[of A k "lead_monom (sym_mpoly A k)"] unfolding lead_coeff_def by blast qed lemma lead_monom_sym_mpoly: assumes "sorted xs" "distinct xs" "k \ length xs" shows "lead_monom (sym_mpoly (set xs) k :: 'a :: zero_neq_one mpoly) = monom_of_set (set (take k xs))" (is "lead_monom ?p = _") proof - let ?m = "lead_monom ?p" have sym: "symmetric_mpoly (set xs) (sym_mpoly (set xs) k)" by (intro symmetric_sym_mpoly) auto from assms have [simp]: "card (set xs) = length xs" by (subst distinct_card) auto from assms have "lead_coeff ?p = 1" by (subst lead_coeff_sym_mpoly) auto then obtain X where X: "X \ set xs" "card X = k" "?m = monom_of_set X" unfolding lead_coeff_def by (subst (asm) coeff_sym_mpoly) (auto split: if_splits) define ys where "ys = map (\x. if x \ X then 1 else 0 :: nat) xs" have [simp]: "length ys = length xs" by (simp add: ys_def) have ys_altdef: "ys = map (lookup ?m) xs" unfolding ys_def using X finite_subset[OF X(1)] by (intro map_cong) (auto simp: lookup_monom_of_set) define i where "i = Min (insert (length xs) {i. i < length xs \ ys ! i = 0})" have "i \ length xs" by (auto simp: i_def) have in_X: "xs ! j \ X" if "j < i" for j using that unfolding i_def by (auto simp: ys_def) have not_in_X: "xs ! j \ X" if "i \ j" "j < length xs" for j proof - have ne: "{i. i < length xs \ ys ! i = 0} \ {}" proof assume [simp]: "{i. i < length xs \ ys ! i = 0} = {}" from that show False by (simp add: i_def) qed hence "Min {i. i < length xs \ ys ! i = 0} \ {i. i < length xs \ ys ! i = 0}" using that by (intro Min_in) auto also have "Min {i. i < length xs \ ys ! i = 0} = i" unfolding i_def using ne by (subst Min_insert) (auto simp: min_def) finally have i: "ys ! i = 0" "i < length xs" by simp_all have "lookup ?m (xs ! j) \ lookup ?m (xs ! i)" using that assms by (intro lookup_lead_monom_decreasing[OF sym]) (auto intro!: sorted_nth_mono simp: set_conv_nth) also have "\ = 0" using i by (simp add: ys_altdef) finally show ?thesis using that X finite_subset[OF X(1)] by (auto simp: lookup_monom_of_set) qed from X have "k = card X" by simp also have "X = (\i. xs ! i) ` {i. i < length xs \ xs ! i \ X}" using X by (auto simp: set_conv_nth) also have "card \ = (\i | i < length xs \ xs ! i \ X. 1)" using assms by (subst card_image) (auto intro!: inj_on_nth) also have "\ = (\i | i < length xs. if xs ! i \ X then 1 else 0)" by (intro sum.mono_neutral_cong_left) auto also have "\ = sum_list ys" by (auto simp: sum_list_sum_nth ys_def intro!: sum.cong) also have "ys = take i ys @ drop i ys" by simp also have "sum_list \ = sum_list (take i ys) + sum_list (drop i ys)" by (subst sum_list_append) auto also have "take i ys = replicate i 1" using \i \ length xs\ in_X by (intro replicate_eqI) (auto simp: ys_def set_conv_nth) also have "sum_list \ = i" by simp also have "drop i ys = replicate (length ys - i) 0" using \i \ length xs\ not_in_X by (intro replicate_eqI) (auto simp: ys_def set_conv_nth) also have "sum_list \ = 0" by simp finally have "i = k" by simp have "X = set (filter (\x. x \ X) xs)" using X by auto also have "xs = take i xs @ drop i xs" by simp also note filter_append also have "filter (\x. x \ X) (take i xs) = take i xs" using in_X by (intro filter_True) (auto simp: set_conv_nth) also have "filter (\x. x \ X) (drop i xs) = []" using not_in_X by (intro filter_False) (auto simp: set_conv_nth) finally have "X = set (take i xs)" by simp with \i = k\ and X show ?thesis by simp qed subsection \Induction on the leading monomial\ text \ We show that the monomial ordering for a fixed set of variables is well-founded, so we can perform induction on the leading monomial of a polynomial. \ definition monom_less_on where "monom_less_on A = {(m1, m2). m1 < m2 \ keys m1 \ A \ keys m2 \ A}" lemma wf_monom_less_on: assumes "finite A" shows "wf (monom_less_on A :: ((nat \\<^sub>0 'b :: {zero, wellorder}) \ _) set)" proof (rule wf_subset) define n where "n = Suc (Max (insert 0 A))" have less_n: "k < n" if "k \ A" for k using that assms by (auto simp: n_def less_Suc_eq_le Max_ge_iff) define f :: "(nat \\<^sub>0 'b) \ 'b list" where "f = (\m. map (lookup m) [0.. inv_image (lexn {(x, y). x < y} n) f" proof safe fix m1 m2 :: "nat \\<^sub>0 'b" assume "(m1, m2) \ monom_less_on A" hence m12: "m1 < m2" "keys m1 \ A" "keys m2 \ A" by (auto simp: monom_less_on_def) then obtain k where k: "lookup m1 k < lookup m2 k" "\i(lookup m1 k = 0 \ lookup m2 k = 0)" proof (intro notI) assume "lookup m1 k = 0 \ lookup m2 k = 0" hence [simp]: "lookup m1 k = 0" "lookup m2 k = 0" by blast+ from k(1) show False by simp qed hence "k \ A" using m12 by (auto simp: in_keys_iff) hence "k < n" by (simp add: less_n) define as where "as = map (lookup m1) [0..k < n\ by (simp flip: upt_conv_Cons upt_add_eq_append') have [simp]: "length as = k" "length bs1 = n - Suc k" "length bs2 = n - Suc k" by (simp_all add: as_def bs1_def bs2_def) have "f m1 = as @ [lookup m1 k] @ bs1" unfolding f_def by (subst decomp) (simp add: as_def bs1_def) moreover have "f m2 = as @ [lookup m2 k] @ bs2" unfolding f_def using k by (subst decomp) (simp add: as_def bs2_def) ultimately show "(m1, m2) \ inv_image (lexn {(x,y). x < y} n) f" using k(1) \k < n\ unfolding lexn_conv by fastforce qed qed lemma lead_monom_induct [consumes 2, case_names less]: fixes p :: "'a :: zero mpoly" assumes fin: "finite A" and vars: "vars p \ A" assumes IH: "\p. vars p \ A \ (\p'. vars p' \ A \ lead_monom p' < lead_monom p \ P p') \ P p" shows "P p" using assms(2) proof (induct m \ "lead_monom p" arbitrary: p rule: wf_induct_rule[OF wf_monom_less_on[OF fin]]) case (1 p) show ?case proof (rule IH) fix p' :: "'a mpoly" assume *: "vars p' \ A" "lead_monom p' < lead_monom p" show "P p'" by (rule 1) (insert * "1.prems" keys_lead_monom_subset, auto simp: monom_less_on_def) qed (insert 1, auto) qed lemma lead_monom_induct' [case_names less]: fixes p :: "'a :: zero mpoly" assumes IH: "\p. (\p'. vars p' \ vars p \ lead_monom p' < lead_monom p \ P p') \ P p" shows "P p" proof - have "finite (vars p)" "vars p \ vars p" by (auto simp: vars_finite) thus ?thesis by (induction rule: lead_monom_induct) (use IH in blast) qed subsection \The fundamental theorem of symmetric polynomials\ lemma lead_coeff_sym_mpoly_powerprod: assumes "finite A" "\x. x \ X \ f x \ {1..card A}" shows "lead_coeff (\x\X. sym_mpoly A (f (x::'a)) ^ g x) = 1" proof - have eq: "lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly) = 1" if "x \ X" for x using that assms by (subst lead_coeff_power) (auto simp: lead_coeff_sym_mpoly assms) hence "(\x\X. lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly)) = (\x\X. 1)" by (intro prod.cong eq refl) also have "\ = 1" by simp finally have eq': "(\x\X. lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly)) = 1" . show ?thesis by (subst lead_coeff_prod) (auto simp: eq eq') qed context fixes A :: "nat set" and xs n f and decr :: "'a :: comm_ring_1 mpoly \ bool" defines "xs \ sorted_list_of_set A" defines "n \ card A" defines "f \ (\i. if i < n then xs ! i else 0)" defines "decr \ (\p. \i\A. \j\A. i \ j \ lookup (lead_monom p) i \ lookup (lead_monom p) j)" begin text \ The computation of the witness for the fundamental theorem works like this: Given some polynomial $p$ (that is assumed to be symmetric in the variables in $A$), we inspect its leading monomial, which is of the form $c X_1^{i_1}\ldots X_n{i_n}$ where the $A = \{X_1,\ldots, X_n\}$, $c$ contains only variables not in $A$, and the sequence $i_j$ is decreasing. The latter holds because $p$ is symmetric. Now, we form the polynomial $q := c e_1^{i_1 - i_2} e_2^{i_2 - i_3} \ldots e_n^{i_n}$, which has the same leading term as $p$. Then $p - q$ has a smaller leading monomial, so by induction, we can assume it to be of the required form and obtain a witness for $p - q$. Now, we only need to add $c Y_1^{i_1 - i_2} \ldots Y_n^{i_n}$ to that witness and we obtain a witness for $p$. \ definition fund_sym_step_coeff :: "'a mpoly \ 'a mpoly" where "fund_sym_step_coeff p = monom (restrictpm (-A) (lead_monom p)) (lead_coeff p)" definition fund_sym_step_monom :: "'a mpoly \ (nat \\<^sub>0 nat)" where "fund_sym_step_monom p = ( let g = (\i. if i < n then lookup (lead_monom p) (f i) else 0) in (\i 'a mpoly" where "fund_sym_step_poly p = ( let g = (\i. if i < n then lookup (lead_monom p) (f i) else 0) in fund_sym_step_coeff p * (\i The following function computes the witness, with the convention that it returns a constant polynomial if the input was not symmetric: \ function (domintros) fund_sym_poly_wit :: "'a :: comm_ring_1 mpoly \ 'a mpoly mpoly" where "fund_sym_poly_wit p = (if \symmetric_mpoly A p \ lead_monom p = 0 \ vars p \ A = {} then Const p else fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (fund_sym_step_coeff p))" by auto lemma coeff_fund_sym_step_coeff: "coeff (fund_sym_step_coeff p) m \ {lead_coeff p, 0}" by (auto simp: fund_sym_step_coeff_def coeff_monom when_def) lemma vars_fund_sym_step_coeff: "vars (fund_sym_step_coeff p) \ vars p - A" unfolding fund_sym_step_coeff_def using keys_lead_monom_subset[of p] by (intro order.trans[OF vars_monom_subset]) auto lemma keys_fund_sym_step_monom: "keys (fund_sym_step_monom p) \ {1..n}" unfolding fund_sym_step_monom_def Let_def by (intro order.trans[OF keys_sum] UN_least, subst keys_single) auto lemma coeff_fund_sym_step_poly: assumes C: "\m. coeff p m \ C" and "ring_closed C" shows "coeff (fund_sym_step_poly p) m \ C" proof - interpret ring_closed C by fact have *: "\m. coeff (p ^ x) m \ C" if "\m. coeff p m \ C" for p x using that by (induction x) (auto simp: coeff_mpoly_times mpoly_coeff_1 intro!: prod_fun_closed) have **: "\m. coeff (prod f X) m \ C" if "\i m. i \ X \ coeff (f i) m \ C" for X and f :: "nat \ _" using that by (induction X rule: infinite_finite_induct) (auto simp: coeff_mpoly_times mpoly_coeff_1 intro!: prod_fun_closed) show ?thesis using C unfolding fund_sym_step_poly_def Let_def fund_sym_step_coeff_def coeff_mpoly_times by (intro prod_fun_closed) (auto simp: coeff_monom when_def lead_coeff_def coeff_sym_mpoly intro!: * **) qed text \ We now show various relevant properties of the subtracted polynomial: \<^enum> Its leading term is the same as that of the input polynomial. \<^enum> It contains now new variables. \<^enum> It is symmetric in the variables in \A\. \ lemma fund_sym_step_poly: shows "finite A \ p \ 0 \ decr p \ lead_monom (fund_sym_step_poly p) = lead_monom p" and "finite A \ p \ 0 \ decr p \ lead_coeff (fund_sym_step_poly p) = lead_coeff p" and "finite A \ p \ 0 \ decr p \ fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" and "vars (fund_sym_step_poly p) \ vars p \ A" and "symmetric_mpoly A (fund_sym_step_poly p)" proof - define g where "g = (\i. if i < n then lookup (lead_monom p) (f i) else 0)" define q where "q = (\i vars p \ A" using keys_lead_monom_subset[of p] vars_monom_subset[of "restrictpm (-A) (lead_monom p)" "lead_coeff p"] unfolding c_def q_def by (intro order.trans[OF vars_mult] order.trans[OF vars_prod] order.trans[OF vars_power] Un_least UN_least order.trans[OF vars_sym_mpoly_subset]) auto thus "vars (fund_sym_step_poly p) \ vars p \ A" by simp have "symmetric_mpoly A (c * q)" unfolding c_def q_def by (intro symmetric_mpoly_mult symmetric_mpoly_monom symmetric_mpoly_prod symmetric_mpoly_power symmetric_sym_mpoly) auto thus "symmetric_mpoly A (fund_sym_step_poly p)" by simp assume finite: "finite A" and [simp]: "p \ 0" and "decr p" have "set xs = A" "distinct xs" and [simp]: "length xs = n" using finite by (auto simp: xs_def n_def) have [simp]: "lead_coeff c = lead_coeff p" "lead_monom c = restrictpm (- A) (lead_monom p)" by (simp_all add: c_def lead_monom_monom) hence f_range [simp]: "f i \ A" if "i < n" for i using that \set xs = A\ by (auto simp: f_def set_conv_nth) have "sorted xs" by (simp add: xs_def) hence f_mono: "f i \ f j" if "i \ j" "j < n" for i j using that by (auto simp: f_def n_def intro: sorted_nth_mono) hence g_mono: "g i \ g j" if "i \ j" for i j unfolding g_def using that using \decr p\ by (auto simp: decr_def) have *: "(\iifinite A\ by (intro prod.cong) (auto simp: n_def lead_coeff_power) hence "lead_coeff q = (\i = (\ifinite A\ by (intro prod.cong) (auto simp: lead_coeff_power n_def) finally have [simp]: "lead_coeff q = 1" by simp have "lead_monom q = (\i = (\ifinite A\ by (intro sum.cong) (auto simp: lead_monom_power n_def) also have "\ = (\iset xs = A\) also from 1 have "\ = monom_of_set (set (take (Suc i) xs))" by (subst lead_monom_sym_mpoly) (auto simp: xs_def n_def) finally show ?case by simp qed finally have lead_monom_q: "lead_monom q = (\i = lead_monom p" (is "?S = _") proof (intro poly_mapping_eqI) fix i :: nat show "lookup (lead_monom c + lead_monom q) i = lookup (lead_monom p) i" proof (cases "i \ A") case False hence "lookup (lead_monom c + lead_monom q) i = lookup (lead_monom p) i + (\jset xs = A\ dest!: in_set_takeD) finally show ?thesis by simp next case True with \set xs = A\ obtain m where m: "i = xs ! m" "m < n" by (auto simp: set_conv_nth) have "lookup (lead_monom c + lead_monom q) i = (\j = (\j | j < n \ i \ set (take (Suc j) xs). g j - g (Suc j))" by (intro sum.mono_neutral_cong_right) auto also have "{j. j < n \ i \ set (take (Suc j) xs)} = {m..distinct xs\ by (force simp: set_conv_nth nth_eq_iff_index_eq) also have "(\j\\. g j - g (Suc j)) = (\j\\. g j) - (\j\\. g (Suc j))" by (subst sum_subtractf_nat) (auto intro!: g_mono) also have "(\j\{m..j\{m<..n}. g j)" by (intro sum.reindex_bij_witness[of _ "\j. j - 1" Suc]) auto also have "\ = (\j\{m<..j\{m.. = (\j\{m..j\\. g j) = lookup (lead_monom p) i" using m by (auto simp: g_def not_less le_Suc_eq f_def) finally show ?thesis . qed qed finally show "lead_monom (fund_sym_step_poly p) = lead_monom p" by simp show "lead_coeff (fund_sym_step_poly p) = lead_coeff p" by (simp add: lead_coeff_mult) have *: "lookup (fund_sym_step_monom p) k = (if k \ {1..n} then g (k - 1) - g k else 0)" for k proof - have "lookup (fund_sym_step_monom p) k = (\x\(if k \ {1..n} then {k - 1} else {}). g (k - 1) - g k)" unfolding fund_sym_step_monom_def lookup_sum Let_def by (intro sum.mono_neutral_cong_right) (auto simp: g_def lookup_single when_def split: if_splits) thus ?thesis by simp qed hence "(\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x :: 'a mpoly) = (\x\{1..n}. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" by (intro Prod_any.expand_superset) auto also have "\ = (\xi. i - 1"]) auto also have "\ = q" unfolding q_def by (intro prod.cong) (auto simp: *) finally show "fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" by (simp add: c_def q_def f_def g_def fund_sym_step_monom_def fund_sym_step_coeff_def) qed text \ If the input is well-formed, a single step of the procedure always decreases the leading monomial. \ lemma lead_monom_fund_sym_step_poly_less: assumes "finite A" and "lead_monom p \ 0" and "decr p" shows "lead_monom (p - fund_sym_step_poly p) < lead_monom p" proof (cases "p = fund_sym_step_poly p") case True thus ?thesis using assms by (auto simp: order.strict_iff_order) next case False from assms have [simp]: "p \ 0" by auto let ?q = "fund_sym_step_poly p" and ?m = "lead_monom p" have "coeff (p - ?q) ?m = 0" using fund_sym_step_poly[of p] assms by (simp add: lead_coeff_def) moreover have "lead_coeff (p - ?q) \ 0" using False by auto ultimately have "lead_monom (p - ?q) \ ?m" unfolding lead_coeff_def by auto moreover have "lead_monom (p - ?q) \ ?m" using fund_sym_step_poly[of p] assms by (intro order.trans[OF lead_monom_diff] max.boundedI) auto ultimately show ?thesis by (auto simp: order.strict_iff_order) qed text \ Finally, we prove that the witness is indeed well-defined for all inputs. \ lemma fund_sym_poly_wit_dom_aux: assumes "finite B" "vars p \ B" "A \ B" shows "fund_sym_poly_wit_dom p" using assms(1-3) proof (induction p rule: lead_monom_induct) case (less p) have [simp]: "finite A" by (rule finite_subset[of _ B]) fact+ show ?case proof (cases "lead_monom p = 0 \ \symmetric_mpoly A p") case False hence [simp]: "p \ 0" by auto note decr = lookup_lead_monom_decreasing[of A p] have "vars (p - fund_sym_step_poly p) \ B" using fund_sym_step_poly[of p] decr False less.prems less.hyps \A \ B\ by (intro order.trans[OF vars_diff]) auto hence "fund_sym_poly_wit_dom (p - local.fund_sym_step_poly p)" using False less.prems less.hyps decr by (intro less.IH fund_sym_step_poly symmetric_mpoly_diff lead_monom_fund_sym_step_poly_less) (auto simp: decr_def) thus ?thesis using fund_sym_poly_wit.domintros by blast qed (auto intro: fund_sym_poly_wit.domintros) qed lemma fund_sym_poly_wit_dom [intro]: "fund_sym_poly_wit_dom p" proof - consider "\symmetric_mpoly A p" | "vars p \ A = {}" | "symmetric_mpoly A p" "A \ vars p" using symmetric_mpoly_imp_orthogonal_or_subset[of A p] by blast thus ?thesis proof cases assume "symmetric_mpoly A p" "A \ vars p" thus ?thesis using fund_sym_poly_wit_dom_aux[of "vars p" p] by (auto simp: vars_finite) qed (auto intro: fund_sym_poly_wit.domintros) qed termination fund_sym_poly_wit by (intro allI fund_sym_poly_wit_dom) (*<*) lemmas [simp del] = fund_sym_poly_wit.simps (*>*) text \ Next, we prove that our witness indeed fulfils all the properties stated by the fundamental theorem: \<^enum> If the original polynomial was in $R[X_1,\ldots,X_n,\ldots, X_m]$ where the $X_1$ to $X_n$ are the symmetric variables, then the witness is a polynomial in $R[X_{n+1},\ldots,X_m][Y_1,\ldots,Y_n]$. This means that its coefficients are polynomials in the variables of the original polynomial, minus the symmetric ones, and the (new and independent) variables of the witness polynomial range from $1$ to $n$. \<^enum> Substituting the \i\-th symmetric polynomial $e_i(X_1,\ldots,X_n)$ for the $Y_i$ variable for every \i\ yields the original polynomial. \<^enum> The coefficient ring $R$ need not be the entire type; if the coefficients of the original polynomial are in some subring, then the coefficients of the coefficients of the witness also do. \ lemma fund_sym_poly_wit_coeffs_aux: assumes "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" shows "vars (coeff (fund_sym_poly_wit p) m) \ B - A" using assms proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto with 1 False have "vars (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p)) m) \ B - A" by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) auto hence "vars (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (fund_sym_step_coeff p)) m) \ B - A" unfolding coeff_add coeff_monom using vars_fund_sym_step_coeff[of p] "1.prems" by (intro order.trans[OF vars_add] Un_least order.trans[OF vars_monom_subset]) (auto simp: when_def) thus ?thesis using "1.prems" False unfolding fund_sym_poly_wit.simps[of p] by simp qed (insert "1.prems", auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const lead_monom_eq_0_iff) qed lemma fund_sym_poly_wit_coeffs: assumes "symmetric_mpoly A p" shows "vars (coeff (fund_sym_poly_wit p) m) \ vars p - A" proof (cases "A \ vars p") case True with fund_sym_poly_wit_coeffs_aux[of "vars p" p m] assms show ?thesis by (auto simp: vars_finite) next case False hence "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[OF assms] by auto thus ?thesis by (auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const) qed lemma fund_sym_poly_wit_vars: "vars (fund_sym_poly_wit p) \ {1..n}" proof (cases "symmetric_mpoly A p \ A \ vars p") case True define B where "B = vars p" have "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" using True unfolding B_def by (auto simp: vars_finite) thus ?thesis proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto hence "vars (local.fund_sym_poly_wit (p - local.fund_sym_step_poly p)) \ {1..n}" using False "1.prems" by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) (auto simp: lead_monom_eq_0_iff) hence "vars (fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (local.fund_sym_step_coeff p)) \ {1..n}" by (intro order.trans[OF vars_add] Un_least order.trans[OF vars_monom_subset] keys_fund_sym_step_monom) auto thus ?thesis using "1.prems" False unfolding fund_sym_poly_wit.simps[of p] by simp qed (insert "1.prems", auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const lead_monom_eq_0_iff) qed next case False then consider "\symmetric_mpoly A p" | "symmetric_mpoly A p" "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[of A p] by auto thus ?thesis by cases (auto simp: fund_sym_poly_wit.simps[of p]) qed lemma fund_sym_poly_wit_insertion_aux: assumes "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" shows "insertion (sym_mpoly A) (fund_sym_poly_wit p) = p" using assms proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) from "1.prems" have "decr p" using lookup_lead_monom_decreasing[of A p] by (auto simp: decr_def) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto hence "insertion (sym_mpoly A) (fund_sym_poly_wit (p - fund_sym_step_poly p)) = p - fund_sym_step_poly p" using 1 False by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) auto moreover have "fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" using "1.prems" finite_subset[of A B] False \decr p\ by (intro fund_sym_step_poly) auto ultimately show ?thesis unfolding fund_sym_poly_wit.simps[of p] by (auto simp: insertion_add) qed (auto simp: fund_sym_poly_wit.simps[of p]) qed lemma fund_sym_poly_wit_insertion: assumes "symmetric_mpoly A p" shows "insertion (sym_mpoly A) (fund_sym_poly_wit p) = p" proof (cases "A \ vars p") case False hence "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[OF assms] by auto thus ?thesis by (auto simp: fund_sym_poly_wit.simps[of p]) next case True with fund_sym_poly_wit_insertion_aux[of "vars p" p] assms show ?thesis by (auto simp: vars_finite) qed lemma fund_sym_poly_wit_coeff: assumes "\m. coeff p m \ C" "ring_closed C" shows "\m m'. coeff (coeff (fund_sym_poly_wit p) m) m' \ C" using assms(1) proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) interpret ring_closed C by fact show ?case proof (cases "\symmetric_mpoly A p \ lead_monom p = 0 \ vars p \ A = {}") case True thus ?thesis using "1.prems" by (auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const) next case False have *: "\m m'. coeff (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p)) m) m' \ C" using False "1.prems" assms coeff_fund_sym_step_poly [of p] by (intro 1) auto show ?thesis proof (intro allI, goal_cases) case (1 m m') thus ?case using * False coeff_fund_sym_step_coeff[of p m'] "1.prems" by (auto simp: fund_sym_poly_wit.simps[of p] coeff_monom lead_coeff_def when_def) qed qed qed subsection \Uniqueness\ text \ Next, we show that the polynomial representation of a symmetric polynomial in terms of the elementary symmetric polynomials not only exists, but is unique. The key property here is that products of powers of elementary symmetric polynomials uniquely determine the exponent vectors, i.\,e.\ if $e_1, \ldots, e_n$ are the elementary symmetric polynomials, $a = (a_1,\ldots, a_n)$ and $b = (b_1,\ldots,b_n)$ are vectors of natural numbers, then: \[e_1^{a_1}\ldots e_n^{a_n} = e_1^{b_1}\ldots e_n^{b_n} \longleftrightarrow a = b\] We show this now. \ lemma lead_monom_sym_mpoly_prod: assumes "finite A" shows "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly A i :: 'a mpoly))" proof - have "(\i=1..n. lead_coeff (sym_mpoly A i ^ h i :: 'a mpoly)) = 1" using assms unfolding n_def by (intro prod.neutral allI) (auto simp: lead_coeff_power) hence "lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i=1..n. lead_monom (sym_mpoly A i ^ h i :: 'a mpoly))" by (subst lead_monom_prod) auto also have "\ = (\i=1..n. of_nat (h i) * lead_monom (sym_mpoly A i :: 'a mpoly))" by (intro sum.cong refl, subst lead_monom_power) (auto simp: lead_coeff_power assms n_def) finally show ?thesis . qed lemma lead_monom_sym_mpoly_prod_notin: assumes "finite A" "k \ A" shows "lookup (lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly)) k = 0" proof - have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) have "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly (set xs) i :: 'a mpoly))" by (subst lead_monom_sym_mpoly_prod) (use xs assms in auto) also have "lookup \ k = 0" unfolding lookup_sum by (intro sum.neutral ballI, subst lead_monom_sym_mpoly) (insert xs assms, auto simp: xs lead_monom_sym_mpoly lookup_monom_of_set set_conv_nth) finally show ?thesis . qed lemma lead_monom_sym_mpoly_prod_in: assumes "finite A" "k < n" shows "lookup (lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly)) (xs ! k) = (\i=k+1..n. h i)" proof - have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) have "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly (set xs) i :: 'a mpoly))" by (subst lead_monom_sym_mpoly_prod) (use xs assms in simp_all) also have "\ = (\i=1..n. of_nat (h i) * monom_of_set (set (take i xs)))" using xs by (intro sum.cong refl, subst lead_monom_sym_mpoly) auto also have "lookup \ (xs ! k) = (\i | i \ {1..n} \ xs ! k \ set (take i xs). h i)" unfolding lookup_sum lookup_monom_of_set by (intro sum.mono_neutral_cong_right) auto also have "{i. i \ {1..n} \ xs ! k \ set (take i xs)} = {k+1..n}" proof (intro equalityI subsetI) fix i assume i: "i \ {k+1..n}" hence "take i xs ! k = xs ! k" "k < n" "k < i" using assms by auto with i show "i \ {i. i \ {1..n} \ xs ! k \ set (take i xs)}" by (force simp: set_conv_nth) qed (insert assms xs, auto simp: set_conv_nth Suc_le_eq nth_eq_iff_index_eq) finally show ?thesis . qed lemma lead_monom_sym_poly_powerprod_inj: assumes "lead_monom (\i. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" assumes "finite A" "keys m1 \ {1..n}" "keys m2 \ {1..n}" shows "m1 = m2" proof (rule poly_mapping_eqI) fix k :: nat have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) from assms(3,4) have *: "i \ {1..n}" if "lookup m1 i \ 0 \ lookup m2 i \ 0" for i using that by (auto simp: subset_iff in_keys_iff) have **: "(\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = (\i=1..n. sym_mpoly A i ^ lookup m i :: 'a mpoly)" if "m \ {m1, m2}" for m using that * by (intro Prod_any.expand_superset subsetI * ) (auto intro!: Nat.gr0I) have ***: "lead_monom (\i=1..n. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i=1..n. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" using assms by (simp add: ** ) have sum_eq: "sum (lookup m1) {Suc k..n} = sum (lookup m2) {Suc k..n}" if "k < n" for k using arg_cong[OF ***, of "\m. lookup m (xs ! k)"] \finite A\ that by (subst (asm) (1 2) lead_monom_sym_mpoly_prod_in) auto show "lookup m1 k = lookup m2 k" proof (cases "k \ {1..n}") case False hence "lookup m1 k = 0" "lookup m2 k = 0" using assms by (auto simp: in_keys_iff) thus ?thesis by simp next case True thus ?thesis proof (induction "n - k" arbitrary: k rule: less_induct) case (less l) have "sum (lookup m1) {Suc (l - 1)..n} = sum (lookup m2) {Suc (l - 1)..n}" using less by (intro sum_eq) auto also have "{Suc (l - 1)..n} = insert l {Suc l..n}" using less by auto also have "sum (lookup m1) \ = lookup m1 l + (\i=Suc l..n. lookup m1 i)" by (subst sum.insert) auto also have "(\i=Suc l..n. lookup m1 i) = (\i=Suc l..n. lookup m2 i)" by (intro sum.cong less) auto also have "sum (lookup m2) (insert l {Suc l..n}) = lookup m2 l + (\i=Suc l..n. lookup m2 i)" by (subst sum.insert) auto finally show "lookup m1 l = lookup m2 l" by simp qed qed qed text \ We now show uniqueness by first showing that the zero polynomial has a unique representation. We fix some polynomial $p$ with $p(e_1,\ldots, e_n) = 0$ and then show, by contradiction, that $p = 0$. We have \[p(e_1,\ldots,e_n) = \sum c_{a_1,\ldots,a_n} e_1^{a_1}\ldots e_n^{a_n}\] and due to the injectivity of products of powers of elementary symmetric polynomials, the leading term of that sum is precisely the leading term of the summand with the biggest leading monomial, since summands cannot cancel each other. However, we also know that $p(e_1,\ldots,e_n) = 0$, so it follows that all summands must have leading term 0, and it is then easy to see that they must all be identically 0. \ lemma sym_mpoly_representation_unique_aux: fixes p :: "'a mpoly mpoly" assumes "finite A" "insertion (sym_mpoly A) p = 0" "\m. vars (coeff p m) \ A = {}" "vars p \ {1..n}" shows "p = 0" proof (rule ccontr) assume p: "p \ 0" have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) define h where "h = (\m. coeff p m * (\i. sym_mpoly A i ^ lookup m i))" define M where "M = {m. coeff p m \ 0}" define maxm where "maxm = Max ((lead_monom \ h) ` M)" have "finite M" by (auto intro!: finite_subset[OF _ finite_coeff_support[of p]] simp: h_def M_def) have keys_subset: "keys m \ {1..n}" if "coeff p m \ 0" for m using that assms coeff_notin_vars[of m p] by blast have lead_coeff: "lead_coeff (h m) = lead_coeff (coeff p m)" (is ?th1) and lead_monom: "lead_monom (h m) = lead_monom (coeff p m) + lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" (is ?th2) if [simp]: "coeff p m \ 0" for m proof - have "(\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = (\i | lookup m i \ 0. sym_mpoly A i ^ lookup m i :: 'a mpoly)" by (intro Prod_any.expand_superset) (auto intro!: Nat.gr0I) also have "lead_coeff \ = 1" using assms keys_subset[of m] by (intro lead_coeff_sym_mpoly_powerprod) (auto simp: in_keys_iff subset_iff n_def) finally have eq: "lead_coeff (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = 1" . thus ?th1 unfolding h_def using \coeff p m \ 0\ by (subst lead_coeff_mult) auto show ?th2 unfolding h_def by (subst lead_monom_mult) (auto simp: eq) qed have "insertion (sym_mpoly A) p = (\m\M. h m)" unfolding insertion_altdef h_def M_def by (intro Sum_any.expand_superset) auto also have "lead_monom \ = maxm" unfolding maxm_def proof (rule lead_monom_sum) from p obtain m where "coeff p m \ 0" using mpoly_eqI[of p 0] by auto hence "m \ M" using \coeff p m \ 0\ lead_coeff[of m] by (auto simp: M_def) thus "M \ {}" by auto next have restrict_lead_monom: "restrictpm A (lead_monom (h m)) = lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" if [simp]: "coeff p m \ 0" for m proof - have "restrictpm A (lead_monom (h m)) = restrictpm A (lead_monom (coeff p m)) + restrictpm A (lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly))" by (auto simp: lead_monom restrictpm_add) also have "restrictpm A (lead_monom (coeff p m)) = 0" using assms by (intro restrictpm_orthogonal order.trans[OF keys_lead_monom_subset]) auto also have "restrictpm A (lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)) = lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" by (intro restrictpm_id order.trans[OF keys_lead_monom_subset] order.trans[OF vars_Prod_any] UN_least order.trans[OF vars_power] vars_sym_mpoly_subset) finally show ?thesis by simp qed show "inj_on (lead_monom \ h) M" proof fix m1 m2 assume m12: "m1 \ M" "m2 \ M" "(lead_monom \ h) m1 = (lead_monom \ h) m2" hence [simp]: "coeff p m1 \ 0" "coeff p m2 \ 0" by (auto simp: M_def h_def) have "restrictpm A (lead_monom (h m1)) = restrictpm A (lead_monom (h m2))" using m12 by simp hence "lead_monom (\i. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" by (simp add: restrict_lead_monom) thus "m1 = m2" by (rule lead_monom_sym_poly_powerprod_inj) (use \finite A\ keys_subset[of m1] keys_subset[of m2] in auto) qed next fix m assume "m \ M" hence "lead_coeff (h m) = lead_coeff (coeff p m)" by (simp add: lead_coeff M_def) with \m \ M\ show "h m \ 0" by (auto simp: M_def) qed fact+ finally have "maxm = 0" by (simp add: assms) have only_zero: "m = 0" if "m \ M" for m proof - from that have nz [simp]: "coeff p m \ 0" by (auto simp: M_def h_def) from that have "(lead_monom \ h) m \ maxm" using \finite M\ unfolding maxm_def by (intro Max_ge imageI finite_imageI) with \maxm = 0\ have [simp]: "lead_monom (h m) = 0" by simp have lookup_nzD: "k \ {1..n}" if "lookup m k \ 0" for k using keys_subset[of m] that by (auto simp: in_keys_iff subset_iff) have "lead_monom (coeff p m) + 0 \ lead_monom (h m)" unfolding lead_monom[OF nz] by (intro add_left_mono) auto also have "\ = 0" by simp finally have lead_monom_0: "lead_monom (coeff p m) = 0" by simp have "sum (lookup m) {1..n} = 0" proof (rule ccontr) assume "sum (lookup m) {1..n} \ 0" hence "sum (lookup m) {1..n} > 0" by presburger have "0 \ lead_coeff (MPoly_Type.coeff p m)" by auto also have "lead_coeff (MPoly_Type.coeff p m) = lead_coeff (h m)" by (simp add: lead_coeff) also have "lead_coeff (h m) = coeff (h m) 0" by (simp add: lead_coeff_def) also have "\ = coeff (coeff p m) 0 * coeff (\i. sym_mpoly A i ^ lookup m i) 0" by (simp add: h_def mpoly_coeff_times_0) also have "(\i. sym_mpoly A i ^ lookup m i) = (\i=1..n. sym_mpoly A i ^ lookup m i)" by (intro Prod_any.expand_superset subsetI lookup_nzD) (auto intro!: Nat.gr0I) also have "coeff \ 0 = (\i=1..n. 0 ^ lookup m i)" unfolding mpoly_coeff_prod_0 mpoly_coeff_power_0 by (intro prod.cong) (auto simp: coeff_sym_mpoly_0) also have "\ = 0 ^ (\i=1..n. lookup m i)" by (simp add: power_sum) also have "\ = 0" using zero_power[OF \sum (lookup m) {1..n} > 0\] by simp finally show False by auto qed hence "lookup m k = 0" for k using keys_subset[of m] by (cases "k \ {1..n}") (auto simp: in_keys_iff) thus "m = 0" by (intro poly_mapping_eqI) auto qed have "0 = insertion (sym_mpoly A) p" using assms by simp also have "insertion (sym_mpoly A) p = (\m\M. h m)" by fact also have "\ = (\m\{0}. h m)" using only_zero by (intro sum.mono_neutral_left) (auto simp: h_def M_def) also have "\ = coeff p 0" by (simp add: h_def) finally have "0 \ M" by (auto simp: M_def) with only_zero have "M = {}" by auto hence "p = 0" by (intro mpoly_eqI) (auto simp: M_def) with \p \ 0\ show False by contradiction qed text \ The general uniqueness theorem now follows easily. This essentially shows that the substitution $Y_i \mapsto e_i(X_1,\ldots,X_n)$ is an isomorphism between the ring $R[Y_1,\ldots, Y_n]$ and the ring $R[X_1,\ldots,X_n]^{S_n}$ of symmetric polynomials. \ theorem sym_mpoly_representation_unique: fixes p :: "'a mpoly mpoly" assumes "finite A" "insertion (sym_mpoly A) p = insertion (sym_mpoly A) q" "\m. vars (coeff p m) \ A = {}" "\m. vars (coeff q m) \ A = {}" "vars p \ {1..n}" "vars q \ {1..n}" shows "p = q" proof - have "p - q = 0" proof (rule sym_mpoly_representation_unique_aux) fix m show "vars (coeff (p - q) m) \ A = {}" using vars_diff[of "coeff p m" "coeff q m"] assms(3,4)[of m] by auto qed (insert assms vars_diff[of p q], auto simp: insertion_diff) thus ?thesis by simp qed theorem eq_fund_sym_poly_witI: fixes p :: "'a mpoly" and q :: "'a mpoly mpoly" assumes "finite A" "symmetric_mpoly A p" "insertion (sym_mpoly A) q = p" "\m. vars (coeff q m) \ A = {}" "vars q \ {1..n}" shows "q = fund_sym_poly_wit p" using fund_sym_poly_wit_insertion[of p] fund_sym_poly_wit_vars[of p] fund_sym_poly_wit_coeffs[of p] by (intro sym_mpoly_representation_unique) (insert assms, auto simp: fund_sym_poly_wit_insertion) subsection \A recursive characterisation of symmetry\ text \ In a similar spirit to the proof of the fundamental theorem, we obtain a nice recursive and executable characterisation of symmetry. \ (*<*) lemmas [fundef_cong] = disj_cong conj_cong (*>*) function (domintros) check_symmetric_mpoly where "check_symmetric_mpoly p \ (vars p \ A = {} \ A \ vars p \ decr p \ check_symmetric_mpoly (p - fund_sym_step_poly p))" by auto lemma check_symmetric_mpoly_dom_aux: assumes "finite B" "vars p \ B" "A \ B" shows "check_symmetric_mpoly_dom p" using assms(1-3) proof (induction p rule: lead_monom_induct) case (less p) have [simp]: "finite A" by (rule finite_subset[of _ B]) fact+ show ?case proof (cases "lead_monom p = 0 \ \decr p") case False hence [simp]: "p \ 0" by auto have "vars (p - fund_sym_step_poly p) \ B" using fund_sym_step_poly[of p] False less.prems less.hyps \A \ B\ by (intro order.trans[OF vars_diff]) auto hence "check_symmetric_mpoly_dom (p - local.fund_sym_step_poly p)" using False less.prems less.hyps by (intro less.IH fund_sym_step_poly symmetric_mpoly_diff lead_monom_fund_sym_step_poly_less) (auto simp: decr_def) thus ?thesis using check_symmetric_mpoly.domintros by blast qed (auto intro: check_symmetric_mpoly.domintros simp: lead_monom_eq_0_iff) qed lemma check_symmetric_mpoly_dom [intro]: "check_symmetric_mpoly_dom p" proof - show ?thesis proof (cases "A \ vars p") assume "A \ vars p" thus ?thesis using check_symmetric_mpoly_dom_aux[of "vars p" p] by (auto simp: vars_finite) qed (auto intro: check_symmetric_mpoly.domintros) qed termination check_symmetric_mpoly by (intro allI check_symmetric_mpoly_dom) lemmas [simp del] = check_symmetric_mpoly.simps lemma check_symmetric_mpoly_correct: "check_symmetric_mpoly p \ symmetric_mpoly A p" proof (induction p rule: check_symmetric_mpoly.induct) case (1 p) have "symmetric_mpoly A (p - fund_sym_step_poly p) \ symmetric_mpoly A p" (is "?lhs = ?rhs") proof assume ?rhs thus ?lhs by (intro symmetric_mpoly_diff fund_sym_step_poly) next assume ?lhs hence "symmetric_mpoly A (p - fund_sym_step_poly p + fund_sym_step_poly p)" by (intro symmetric_mpoly_add fund_sym_step_poly) thus ?rhs by simp qed moreover have "decr p" if "symmetric_mpoly A p" using lookup_lead_monom_decreasing[of A p] that by (auto simp: decr_def) ultimately show "check_symmetric_mpoly p \ symmetric_mpoly A p" using 1 symmetric_mpoly_imp_orthogonal_or_subset[of A p] by (auto simp: Let_def check_symmetric_mpoly.simps[of p] intro: symmetric_mpoly_orthogonal) qed end subsection \Symmetric functions of roots of a univariate polynomial\ text \ Consider a factored polynomial \[p(X) = c_n X^n + c_{n-1} X^{n-1} + \ldots + c_1X + c_0 = (X - x_1)\ldots(X - x_n)\ .\] where $c_n$ is a unit. Then any symmetric polynomial expression $q(x_1, \ldots, x_n)$ in the roots $x_i$ can be written as a polynomial expression $q'(c_0,\ldots, c_{n-1})$ in the $c_i$. Moreover, if the coefficients of $q$ and the inverse of $c_n$ all lie in some subring, the coefficients of $q'$ do as well. \ context fixes C :: "'b :: comm_ring_1 set" and A :: "nat set" and root :: "nat \ 'a :: comm_ring_1" and l :: "'a \ 'b" and q :: "'b mpoly" and n :: nat defines "n \ card A" assumes C: "ring_closed C" "\m. coeff q m \ C" assumes l: "ring_homomorphism l" assumes finite: "finite A" assumes sym: "symmetric_mpoly A q" and vars: "vars q \ A" begin interpretation ring_closed C by fact interpretation ring_homomorphism l by fact theorem symmetric_poly_of_roots_conv_poly_of_coeffs: assumes c: "cinv * l c = 1" "cinv \ C" assumes "p = Polynomial.smult c (\i\A. [:-root i, 1:])" obtains q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" proof - define q' where "q' = fund_sym_poly_wit A q" define q'' where "q'' = mapm_mpoly (\m x. (\i. (cinv * l (- 1) ^ i) ^ lookup m i) * insertion (\_. 0) x) q'" define reindex where "reindex = (\i. if i \ n then n - i else i)" have "bij reindex" by (intro bij_betwI[of reindex _ _ reindex]) (auto simp: reindex_def) have "vars q' \ {1..n}" unfolding q'_def n_def by (intro fund_sym_poly_wit_vars) hence "vars q'' \ {1..n}" unfolding q''_def using vars_mapm_mpoly_subset by auto have "insertion (l \ root) (insertion (sym_mpoly A) q') = insertion (\n. insertion (l \ root) (sym_mpoly A n)) (map_mpoly (insertion (l \ root)) q')" by (rule insertion_insertion) also have "insertion (sym_mpoly A) q' = q" unfolding q'_def by (intro fund_sym_poly_wit_insertion sym) also have "insertion (\i. insertion (l \ root) (sym_mpoly A i)) (map_mpoly (insertion (l \ root)) q') = insertion (\i. cinv * l ((- 1) ^ i) * l (poly.coeff p (n - i))) (map_mpoly (insertion (l \ root)) q')" proof (intro insertion_irrelevant_vars, goal_cases) case (1 i) hence "i \ vars q'" using vars_map_mpoly_subset by auto also have "\ \ {1..n}" unfolding q'_def n_def by (intro fund_sym_poly_wit_vars) finally have i: "i \ {1..n}" . have "insertion (l \ root) (sym_mpoly A i) = l (\Y | Y \ A \ card Y = i. prod root Y)" using \finite A\ by (simp add: insertion_sym_mpoly) also have "\ = cinv * l (c * (\Y | Y \ A \ card Y = i. prod root Y))" unfolding mult mult.assoc[symmetric] \cinv * l c = 1\ by simp also have "c * (\Y | Y \ A \ card Y = i. prod root Y) = ((-1) ^ i * poly.coeff p (n - i))" using coeff_poly_from_roots[of A "n - i" root] i assms finite by (auto simp: n_def minus_one_power_iff) finally show ?case by (simp add: o_def) qed also have "map_mpoly (insertion (l \ root)) q' = map_mpoly (insertion (\_. 0)) q'" using fund_sym_poly_wit_coeffs[OF sym] vars by (intro map_mpoly_cong insertion_irrelevant_vars) (auto simp: q'_def) also have "insertion (\i. cinv * l ((- 1) ^ i) * l (poly.coeff p (n - i))) \ = insertion (\i. l (poly.coeff p (n - i))) q''" unfolding insertion_substitute_linear map_mpoly_conv_mapm_mpoly q''_def by (subst mapm_mpoly_comp) auto also have "\ = insertion (l \ poly.coeff p) (mpoly_map_vars reindex q'')" using \bij reindex\ and \vars q'' \ {1..n}\ by (subst insertion_mpoly_map_vars) (auto simp: o_def reindex_def intro!: insertion_irrelevant_vars) finally have "insertion (l \ root) q = insertion (l \ poly.coeff p) (mpoly_map_vars reindex q'')" . moreover have "coeff (mpoly_map_vars reindex q'') m \ C" for m unfolding q''_def q'_def using \bij reindex\ fund_sym_poly_wit_coeff[of q C A] C \cinv \ C\ by (auto simp: coeff_mpoly_map_vars intro!: mult_closed Prod_any_closed power_closed Sum_any_closed) moreover have "vars (mpoly_map_vars reindex q'') \ {0..bij reindex\ and \vars q'' \ {1..n}\ by (subst vars_mpoly_map_vars) (auto simp: reindex_def subset_iff)+ ultimately show ?thesis using that[of "mpoly_map_vars reindex q''"] by auto qed corollary symmetric_poly_of_roots_conv_poly_of_coeffs_monic: assumes "p = (\i\A. [:-root i, 1:])" obtains q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" proof - obtain q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by (rule symmetric_poly_of_roots_conv_poly_of_coeffs[of 1 1 p]) (use assms in auto) thus ?thesis by (intro that[of q']) auto qed text \ As a corollary, we obtain the following: Let $R, S$ be rings with $R\subseteq S$. Consider a polynomial $p\in R[X]$ whose leading coefficient $c$ is a unit in $R$ and that has a full set of roots $x_1,\ldots, x_n \in S$, i.\,e.\ $p(X) = c(X - x_1) \ldots (X - x_n)$. Let $q \in R[X_1,\ldots, X_n]$ be some symmetric polynomial expression in the roots. Then $q(x_1, \ldots, x_n) \in R$. A typical use case is $R = \mathbb{Q}$ and $S = \mathbb{C}$, i.\,e.\ any symmetric polynomial expression with rational coefficients in the roots of a rational polynomial is again rational. Similarly, any symmetric polynomial expression with integer coefficients in the roots of a monic integer polynomial is agan an integer. This is remarkable, since the roots themselves are usually not rational (possibly not even real). This particular fact is a key ingredient used in the standard proof that $\pi$ is transcendental. \ corollary symmetric_poly_of_roots_in_subring: assumes "cinv * l c = 1" "cinv \ C" assumes "p = Polynomial.smult c (\i\A. [:-root i, 1:])" assumes "\i. l (poly.coeff p i) \ C" shows "insertion (\x. l (root x)) q \ C" proof - obtain q' where q': "vars q' \ {0..m. coeff q' m \ C" "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by (rule symmetric_poly_of_roots_conv_poly_of_coeffs[of cinv c p]) (use assms in simp_all) have "insertion (l \ poly.coeff p) q' \ C" using C assms unfolding insertion_altdef by (intro Sum_any_closed mult_closed q' Prod_any_closed power_closed) auto also have "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by fact finally show ?thesis by (simp add: o_def) qed corollary symmetric_poly_of_roots_in_subring_monic: assumes "p = (\i\A. [:-root i, 1:])" assumes "\i. l (poly.coeff p i) \ C" shows "insertion (\x. l (root x)) q \ C" proof - interpret ring_closed C by fact interpret ring_homomorphism l by fact show ?thesis by (rule symmetric_poly_of_roots_in_subring[of 1 1 p]) (use assms in auto) qed end end \ No newline at end of file