diff --git a/thys/Approximation_Algorithms/Approx_MIS_Hoare.thy b/thys/Approximation_Algorithms/Approx_MIS_Hoare.thy --- a/thys/Approximation_Algorithms/Approx_MIS_Hoare.thy +++ b/thys/Approximation_Algorithms/Approx_MIS_Hoare.thy @@ -1,561 +1,564 @@ section "Independent Set" theory Approx_MIS_Hoare imports "HOL-Hoare.Hoare_Logic" "HOL-Library.Disjoint_Sets" begin text \The algorithm is classical, the proofs are inspired by the ones by Berghammer and M\"uller-Olm \cite{BerghammerM03}. In particular the approximation ratio is improved from \\+1\ to \\\.\ subsection "Graph" text \A set set is simply a set of edges, where an edge is a 2-element set.\ definition independent_vertices :: "'a set set \ 'a set \ bool" where "independent_vertices E S \ S \ \E \ (\v1 v2. v1 \ S \ v2 \ S \ {v1, v2} \ E)" locale Graph_E = fixes E :: "'a set set" assumes finite_E: "finite E" assumes edges2: "e \ E \ card e = 2" begin fun vertices :: "'a set set \ 'a set" where "vertices G = \G" abbreviation V :: "'a set" where "V \ vertices E" definition approximation_miv :: "nat \ 'a set \ bool" where "approximation_miv n S \ independent_vertices E S \ (\S'. independent_vertices E S' \ card S' \ card S * n)" fun neighbors :: "'a \ 'a set" where "neighbors v = {u. {u,v} \ E}" fun degree_vertex :: "'a \ nat" where "degree_vertex v = card (neighbors v)" abbreviation \ :: nat where "\ \ Max{degree_vertex u|u. u \ V}" lemma finite_edges: "e \ E \ finite e" using card_ge_0_finite and edges2 by force lemma finite_V: "finite V" using finite_edges and finite_E by auto +lemma neighbors_subset_V: "neighbors u \ V" +by auto + lemma finite_neighbors: "finite (neighbors u)" - using finite_V and rev_finite_subset [of V "neighbors u"] by auto + using finite_V neighbors_subset_V[of u] by (simp) lemma independent_vertices_finite: "independent_vertices E S \ finite S" by (metis rev_finite_subset independent_vertices_def vertices.simps finite_V) lemma edge_ex_vertices: "e \ E \ \u v. u \ v \ e = {u, v}" proof - assume "e \ E" then have "card e = Suc (Suc 0)" using edges2 by auto then show "\u v. u \ v \ e = {u, v}" by (metis card_eq_SucD insertI1) qed lemma \_pos [simp]: "E = {} \ 0 < \" proof cases assume "E = {}" then show "E = {} \ 0 < \" by auto next assume 1: "E \ {}" then have "V \ {}" using edges2 by fastforce moreover have "finite {degree_vertex u |u. u \ V}" by (metis finite_V finite_imageI Setcompr_eq_image) ultimately have 2: "\ \ {degree_vertex u |u. u \ V}" using Max_in by auto have "\ \ 0" proof assume "\ = 0" with 2 obtain u where 3: "u \ V" and 4: "degree_vertex u = 0" by auto from 3 obtain e where 5: "e \ E" and "u \ e" by auto moreover with 4 have "\v. {u, v} \ e" using finite_neighbors insert_absorb by fastforce ultimately show False using edge_ex_vertices by auto qed then show "E = {} \ 0 < \" by auto qed lemma \_max_degree: "u \ V \ degree_vertex u \ \" proof - assume H: "u \ V" have "finite {degree_vertex u |u. u \ V}" by (metis finite_V finite_imageI Setcompr_eq_image) with H show "degree_vertex u \ \" using Max_ge by auto qed subsection \Wei's algorithm: \(\+1)\-approximation\ text \The 'functional' part of the invariant, used to prove that the algorithm produces an independent set of vertices.\ definition inv_iv :: "'a set \ 'a set \ bool" where "inv_iv S X \ independent_vertices E S \ X \ V \ (\v1 \ (V - X). \v2 \ S. {v1, v2} \ E) \ S \ X" text \Strenghten the invariant with an approximation ratio \r\:\ definition inv_approx :: "'a set \ 'a set \ nat \ bool" where "inv_approx S X r \ inv_iv S X \ card X \ card S * r" text \Preservation of the functional invariant:\ lemma inv_preserv: fixes S :: "'a set" and X :: "'a set" and x :: "'a" assumes inv: "inv_iv S X" and x_def: "x \ V - X" shows "inv_iv (insert x S) (X \ neighbors x \ {x})" proof - have inv1: "independent_vertices E S" and inv2: "X \ V" and inv3: "S \ X" and inv4: "\v1 v2. v1 \ (V - X) \ v2 \ S \ {v1, v2} \ E" using inv unfolding inv_iv_def by auto have finite_S: "finite S" using inv1 and independent_vertices_finite by auto have S1: "\y \ S. {x, y} \ E" using inv4 and x_def by blast have S2: "\x \ S. \y \ S. {x, y} \ E" using inv1 unfolding independent_vertices_def by metis have S3: "v1 \ insert x S \ v2 \ insert x S \ {v1, v2} \ E" for v1 v2 proof - assume "v1 \ insert x S" and "v2 \ insert x S" then consider (a) "v1 = x" and "v2 = x" | (b) "v1 = x" and "v2 \ S" | (c) "v1 \ S" and "v2 = x" | (d) "v1 \ S" and "v2 \ S" by auto then show "{v1, v2} \ E" proof cases case a then show ?thesis using edges2 by force next case b then show ?thesis using S1 by auto next case c then show ?thesis using S1 by (metis doubleton_eq_iff) next case d then show ?thesis using S2 by auto qed qed (* invariant conjunct 1 *) have "independent_vertices E (insert x S)" using S3 and inv1 and x_def unfolding independent_vertices_def by auto (* invariant conjunct 2 *) moreover have "X \ neighbors x \ {x} \ V" proof fix xa assume "xa \ X \ neighbors x \ {x}" then consider (a) "xa \ X" | (b) "xa \ neighbors x" | (c) "xa = x" by auto then show "xa \ V" proof cases case a then show ?thesis using inv2 by blast next case b then show ?thesis by auto next case c then show ?thesis using x_def by blast qed qed (* invariant conjunct 3 *) moreover have "insert x S \ X \ neighbors x \ {x}" using inv3 by auto (* invariant conjunct 4 *) moreover have "v1 \ V - (X \ neighbors x \ {x}) \ v2 \ insert x S \ {v1, v2} \ E" for v1 v2 proof - assume H: "v1 \ V - (X \ neighbors x \ {x})" and "v2 \ insert x S" then consider (a) "v2 = x" | (b) "v2 \ S" by auto then show "{v1, v2} \ E" proof cases case a with H have "v1 \ neighbors v2" by blast then show ?thesis by auto next case b from H have "v1 \ V - X" by blast with b and inv4 show ?thesis by blast qed qed (* conclusion *) ultimately show "inv_iv (insert x S) (X \ neighbors x \ {x})" unfolding inv_iv_def by blast qed lemma inv_approx_preserv: assumes inv: "inv_approx S X (\ + 1)" and x_def: "x \ V - X" shows "inv_approx (insert x S) (X \ neighbors x \ {x}) (\ + 1)" proof - have finite_S: "finite S" using inv and independent_vertices_finite unfolding inv_approx_def inv_iv_def by auto have Sx: "x \ S" using inv and x_def unfolding inv_approx_def inv_iv_def by blast (* main invariant is preserved *) from inv have "inv_iv S X" unfolding inv_approx_def by auto with x_def have "inv_iv (insert x S) (X \ neighbors x \ {x})" proof (intro inv_preserv, auto) qed (* the approximation ratio is preserved (at most \+1 vertices are removed in any iteration) *) moreover have "card (X \ neighbors x \ {x}) \ card (insert x S) * (\ + 1)" proof - have "degree_vertex x \ \" using \_max_degree and x_def by auto then have "card (neighbors x \ {x}) \ \ + 1" using card_Un_le [of "neighbors x" "{x}"] by auto then have "card (X \ neighbors x \ {x}) \ card X + \ + 1" using card_Un_le [of X "neighbors x \ {x}"] by auto also have "... \ card S * (\ + 1) + \ + 1" using inv unfolding inv_approx_def by auto also have "... = card (insert x S) * (\ + 1)" using finite_S and Sx by auto finally show ?thesis . qed (* conclusion *) ultimately show "inv_approx (insert x S) (X \ neighbors x \ {x}) (\ + 1)" unfolding inv_approx_def by auto qed (* the antecedent combines inv_approx (for an arbitrary ratio r) and the negated post-condition *) lemma inv_approx: "independent_vertices E S \ card V \ card S * r \ approximation_miv r S" proof - assume 1: "independent_vertices E S" and 2: "card V \ card S * r" have "independent_vertices E S' \ card S' \ card S * r" for S' proof - assume "independent_vertices E S'" then have "S' \ V" unfolding independent_vertices_def by auto then have "card S' \ card V" using finite_V and card_mono by auto also have "... \ card S * r" using 2 by auto finally show "card S' \ card S * r" . qed with 1 show "approximation_miv r S" unfolding approximation_miv_def by auto qed theorem wei_approx_\_plus_1: "VARS (S :: 'a set) (X :: 'a set) (x :: 'a) { True } S := {}; X := {}; WHILE X \ V INV { inv_approx S X (\ + 1) } DO x := (SOME x. x \ V - X); S := insert x S; X := X \ neighbors x \ {x} OD { approximation_miv (\ + 1) S }" proof (vcg, goal_cases) case (1 S X x) (* invariant initially true *) then show ?case unfolding inv_approx_def inv_iv_def independent_vertices_def by auto next case (2 S X x) (* invariant preserved by loop *) (* definedness of assignment *) let ?x = "(SOME x. x \ V - X)" have "V - X \ {}" using 2 unfolding inv_approx_def inv_iv_def by blast then have "?x \ V - X" using some_in_eq by metis with 2 show ?case using inv_approx_preserv by auto next case (3 S X x) (* invariant implies post-condition *) then show ?case using inv_approx unfolding inv_approx_def inv_iv_def by auto qed subsection \Wei's algorithm: \\\-approximation\ text \The previous approximation uses very little information about the optimal solution (it has at most as many vertices as the set itself). With some extra effort we can improve the ratio to \\\ instead of \\+1\. In order to do that we must show that among the vertices removed in each iteration, at most \\\ could belong to an optimal solution. This requires carrying around a set \P\ (via a ghost variable) which records the vertices deleted in each iteration.\ definition inv_partition :: "'a set \ 'a set \ 'a set set \ bool" where "inv_partition S X P \ inv_iv S X \ \P = X \ (\p \ P. \s \ V. p = {s} \ neighbors s) \ card P = card S \ finite P" lemma inv_partition_preserv: assumes inv: "inv_partition S X P" and x_def: "x \ V - X" shows "inv_partition (insert x S) (X \ neighbors x \ {x}) (insert ({x} \ neighbors x) P)" proof - have finite_S: "finite S" using inv and independent_vertices_finite unfolding inv_partition_def inv_iv_def by auto have Sx: "x \ S" using inv and x_def unfolding inv_partition_def inv_iv_def by blast (* main invariant is preserved *) from inv have "inv_iv S X" unfolding inv_partition_def by auto with x_def have "inv_iv (insert x S) (X \ neighbors x \ {x})" proof (intro inv_preserv, auto) qed (* conjunct 1 *) moreover have "\(insert ({x} \ neighbors x) P) = X \ neighbors x \ {x}" using inv unfolding inv_partition_def by auto (* conjunct 2 *) moreover have "(\p\insert ({x} \ neighbors x) P. \s \ V. p = {s} \ neighbors s)" using inv and x_def unfolding inv_partition_def by auto (* conjunct 3 *) moreover have "card (insert ({x} \ neighbors x) P) = card (insert x S)" proof - from x_def and inv have "x \ \P" unfolding inv_partition_def by auto then have "{x} \ neighbors x \ P" by auto then have "card (insert ({x} \ neighbors x) P) = card P + 1" using inv unfolding inv_partition_def by auto moreover have "card (insert x S) = card S + 1" using Sx and finite_S by auto ultimately show ?thesis using inv unfolding inv_partition_def by auto qed (* conjunct 4 *) moreover have "finite (insert ({x} \ neighbors x) P)" using inv unfolding inv_partition_def by auto (* conclusion *) ultimately show "inv_partition (insert x S) (X \ neighbors x \ {x}) (insert ({x} \ neighbors x) P)" unfolding inv_partition_def by auto qed lemma card_Union_le_sum_card: fixes U :: "'a set set" assumes "\u \ U. finite u" shows "card (\U) \ sum card U" proof (cases "finite U") case False then show "card (\U) \ sum card U" using card_eq_0_iff finite_UnionD by auto next case True then show "card (\U) \ sum card U" proof (induct U rule: finite_induct) case empty then show ?case by auto next case (insert x F) then have "card(\(insert x F)) \ card(x) + card (\F)" using card_Un_le by auto also have "... \ card(x) + sum card F" using insert.hyps by auto also have "... = sum card (insert x F)" using sum.insert_if and insert.hyps by auto finally show ?case . qed qed (* this lemma could be more generally about U :: "nat set", but this makes its application more difficult later *) lemma sum_card: fixes U :: "'a set set" and n :: nat assumes "\S \ U. card S \ n" shows "sum card U \ card U * n" proof cases assume "infinite U \ U = {}" then have "sum card U = 0" using sum.infinite by auto then show "sum card U \ card U * n" by auto next assume "\(infinite U \ U = {})" with assms have "finite U" and "U \ {}"and "\S \ U. card S \ n" by auto then show "sum card U \ card U * n" proof (induct U rule: finite_ne_induct) case (singleton x) then show ?case by auto next case (insert x F) assume "\S\insert x F. card S \ n" then have 1:"card x \ n" and 2:"sum card F \ card F * n" using insert.hyps by auto then have "sum card (insert x F) = card x + sum card F" using sum.insert_if and insert.hyps by auto also have "... \ n + card F * n" using 1 and 2 by auto also have "... = card (insert x F) * n" using card_insert_if and insert.hyps by auto finally show ?case . qed qed (* among the vertices deleted in each iteration, at most \ can belong to an independent set of vertices: the chosen vertex or (some of) its neighbors *) lemma x_or_neighbors: fixes P :: "'a set set" and S :: "'a set" assumes inv: "\p\P. \s \ V. p = {s} \ neighbors s" and ivS: "independent_vertices E S" shows "\p \ P. card (S \ p) \ \" proof fix p assume "p \ P" then obtain s where 1: "s \ V \ p = {s} \ neighbors s" using inv by blast then show "card (S \ p) \ \" proof cases assume "s \ S" then have "S \ neighbors s = {}" using ivS unfolding independent_vertices_def by auto then have "S \ p \ {s}" using 1 by auto then have 2: "card (S \ p) \ 1" using subset_singletonD by fastforce consider (a) "E = {}" | (b) "0 < \" using \_pos by auto then show "card (S \ p) \ \" proof cases case a then have "S = {}" using ivS unfolding independent_vertices_def by auto then show ?thesis by auto next case b then show ?thesis using 2 by auto qed next assume "s \ S" with 1 have "S \ p \ neighbors s" by auto then have "card (S \ p) \ degree_vertex s" using card_mono and finite_neighbors by auto then show "card (S \ p) \ \" using 1 and \_max_degree [of s] by auto qed qed (* the premise combines the invariant and the negated post-condition *) lemma inv_partition_approx: "inv_partition S V P \ approximation_miv \ S" proof - assume H1: "inv_partition S V P" then have "independent_vertices E S" unfolding inv_partition_def inv_iv_def by auto moreover have "independent_vertices E S' \ card S' \ card S * \" for S' proof - let ?I = "{S' \ p | p. p \ P}" (* split the optimal solution among the sets of P, which cover V so no element is lost. We obtain a cover of S' and show the required bound on its cardinality *) assume H2: "independent_vertices E S'" then have "S' \ V" unfolding independent_vertices_def using vertices.simps by blast with H1 have "S' = S' \ \P" unfolding inv_partition_def by auto then have "S' = (\p \ P. S' \ p)" using Int_Union by auto then have "S' = \?I" by blast moreover have "finite S'" using H2 and independent_vertices_finite by auto then have "p \ P \ finite (S' \ p)" for p by auto ultimately have "card S' \ sum card ?I" using card_Union_le_sum_card [of ?I] by auto also have "... \ card ?I * \" using x_or_neighbors [of P S'] and sum_card [of ?I \] and H1 and H2 unfolding inv_partition_def by auto also have "... \ card P * \" proof - have "finite P" using H1 unfolding inv_partition_def by auto then have "card ?I \ card P" using Setcompr_eq_image [of "\p. S' \ p" P] and card_image_le unfolding inv_partition_def by auto then show ?thesis by auto qed also have "... = card S * \" using H1 unfolding inv_partition_def by auto ultimately show "card S' \ card S * \" by auto qed ultimately show "approximation_miv \ S" unfolding approximation_miv_def by auto qed theorem wei_approx_\: "VARS (S :: 'a set) (X :: 'a set) (x :: 'a) { True } S := {}; X := {}; WHILE X \ V INV { \P. inv_partition S X P } DO x := (SOME x. x \ V - X); S := insert x S; X := X \ neighbors x \ {x} OD { approximation_miv \ S }" proof (vcg, goal_cases) case (1 S X x) (* invariant initially true *) (* the invariant is initially true with the ghost variable P := {} *) have "inv_partition {} {} {}" unfolding inv_partition_def inv_iv_def independent_vertices_def by auto then show ?case by auto next case (2 S X x) (* invariant preserved by loop *) (* definedness of assignment *) let ?x = "(SOME x. x \ V - X)" from 2 obtain P where I: "inv_partition S X P" by auto then have "V - X \ {}" using 2 unfolding inv_partition_def by auto then have "?x \ V - X" using some_in_eq by metis (* show that the invariant is true with the ghost variable P := insert ({?x} \ neighbors ?x) P *) with I have "inv_partition (insert ?x S) (X \ neighbors ?x \ {?x}) (insert ({?x} \ neighbors ?x) P)" using inv_partition_preserv by blast then show ?case by auto next case (3 S X x) (* invariant implies post-condition *) then show ?case using inv_partition_approx unfolding inv_approx_def by auto qed subsection "Wei's algorithm with dynamically computed approximation ratio" text \In this subsection, we augment the algorithm with a variable used to compute the effective approximation ratio of the solution. In addition, the vertex of smallest degree is picked. With this heuristic, the algorithm achieves an approximation ratio of \(\+2)/3\, but this is not proved here.\ definition vertex_heuristic :: "'a set \ 'a \ bool" where "vertex_heuristic X v = (\u \ V - X. card (neighbors v - X) \ card (neighbors u - X))" (* this lemma is needed to show that there exist a vertex to be picked by the heuristic *) lemma ex_min_finite_set: fixes S :: "'a set" and f :: "'a \ nat" shows "finite S \ S \ {} \ \x. x \ S \ (\y \ S. f x \ f y)" (is "?P1 \ ?P2 \ \x. ?minf S x") proof (induct S rule: finite_ne_induct) case (singleton x) have "?minf {x} x" by auto then show ?case by auto next case (insert x F) from insert(4) obtain y where Py: "?minf F y" by auto show "\z. ?minf (insert x F) z" proof cases assume "f x < f y" then have "?minf (insert x F) x" using Py by auto then show ?case by auto next assume "\f x < f y" then have "?minf (insert x F) y" using Py by auto then show ?case by auto qed qed lemma inv_approx_preserv2: fixes S :: "'a set" and X :: "'a set" and s :: nat and x :: "'a" assumes inv: "inv_approx S X s" and x_def: "x \ V - X" shows "inv_approx (insert x S) (X \ neighbors x \ {x}) (max (card (neighbors x \ {x} - X)) s)" proof - have finite_S: "finite S" using inv and independent_vertices_finite unfolding inv_approx_def inv_iv_def by auto have Sx: "x \ S" using inv and x_def unfolding inv_approx_def inv_iv_def by blast (* main invariant is preserved *) from inv have "inv_iv S X" unfolding inv_approx_def by auto with x_def have "inv_iv (insert x S) (X \ neighbors x \ {x})" proof (intro inv_preserv, auto) qed (* the approximation ratio is preserved *) moreover have "card (X \ neighbors x \ {x}) \ card (insert x S) * max (card (neighbors x \ {x} - X)) s" proof - let ?N = "neighbors x \ {x} - X" have "card (X \ ?N) \ card X + card ?N" using card_Un_le [of X ?N] by auto also have "... \ card S * s + card ?N" using inv unfolding inv_approx_def by auto also have "... \ card S * max (card ?N) s + card ?N" by auto also have "... \ card S * max (card ?N) s + max (card ?N) s" by auto also have "... = card (insert x S) * max (card ?N) s" using Sx and finite_S by auto finally show ?thesis by auto qed (* conclusion *) ultimately show "inv_approx (insert x S) (X \ neighbors x \ {x}) (max (card (neighbors x \ {x} - X)) s)" unfolding inv_approx_def by auto qed theorem wei_approx_min_degree_heuristic: "VARS (S :: 'a set) (X :: 'a set) (x :: 'a) (r :: nat) { True } S := {}; X := {}; r := 0; WHILE X \ V INV { inv_approx S X r } DO x := (SOME x. x \ V - X \ vertex_heuristic X x); S := insert x S; r := max (card (neighbors x \ {x} - X)) r; X := X \ neighbors x \ {x} OD { approximation_miv r S }" proof (vcg, goal_cases) case (1 S X x r) (* invariant initially true *) then show ?case unfolding inv_approx_def inv_iv_def independent_vertices_def by auto next case (2 S X x r) (* invariant preserved by loop *) (* definedness of assignment *) let ?x = "(SOME x. x \ V - X \ vertex_heuristic X x)" have "V - X \ {}" using 2 unfolding inv_approx_def inv_iv_def by blast moreover have "finite (V - X)" using 2 and finite_V by auto ultimately have "\x. x \ V - X \ vertex_heuristic X x" using ex_min_finite_set [where ?f = "\x. card (neighbors x - X)"] unfolding vertex_heuristic_def by auto then have x_def: "?x \ V - X \ vertex_heuristic X ?x" using someI_ex [where ?P = "\x. x \ V - X \ vertex_heuristic X x"] by auto with 2 show ?case using inv_approx_preserv2 by auto next case (3 S X x r) then show ?case using inv_approx unfolding inv_approx_def inv_iv_def by auto qed end end \ No newline at end of file diff --git a/thys/Buildings/Simplicial.thy b/thys/Buildings/Simplicial.thy --- a/thys/Buildings/Simplicial.thy +++ b/thys/Buildings/Simplicial.thy @@ -1,749 +1,748 @@ section \Simplicial complexes\ text \ In this section we develop the basic theory of abstract simplicial complexes as a collection of finite sets, where the power set of each member set is contained in the collection. Note that in this development we allow the empty simplex, since allowing it or not seemed of no logical consequence, but of some small practical consequence. \ theory Simplicial imports Prelim begin subsection \Geometric notions\ text \ The geometric notions attached to a simplicial complex of main interest to us are those of facets (subsets of codimension one), adjacency (sharing a facet in common), and chains of adjacent simplices. \ subsubsection \Facets\ definition facetrel :: "'a set \ 'a set \ bool" (infix "\" 60) where "y \ x \ \v. v \ y \ x = insert v y" lemma facetrelI: "v \ y \ x = insert v y \ y \ x" using facetrel_def by fast lemma facetrelI_card: "y \ x \ card (x-y) = 1 \ y \ x" using card1[of "x-y"] by (blast intro: facetrelI) lemma facetrel_complement_vertex: "y\x \ x = insert v y \ v\y" using facetrel_def[of y x] by fastforce lemma facetrel_diff_vertex: "v\x \ x-{v} \ x" by (auto intro: facetrelI) lemma facetrel_conv_insert: "y \ x \ v \ x - y \ x = insert v y" unfolding facetrel_def by fast lemma facetrel_psubset: "y \ x \ y \ x" unfolding facetrel_def by fast lemma facetrel_subset: "y \ x \ y \ x" using facetrel_psubset by fast lemma facetrel_card: "y \ x \ card (x-y) = 1" using insert_Diff_if[of _ y y] unfolding facetrel_def by fastforce lemma finite_facetrel_card: "finite x \ y\x \ card x = Suc (card y)" using facetrel_def[of y x] card_insert_disjoint[of x] by auto lemma facetrelI_cardSuc: "z\x \ card x = Suc (card z) \ z\x" - using card_ge_0_finite finite_subset[of z] card_Diff_subset[of z x] - by (force intro: facetrelI_card) + using card_ge_0_finite[of x] card_Diff_subset[of z x] by (auto intro: facetrelI_card) lemma facet2_subset: "\ z\x; z\y; x\y - z \ {} \ \ x \ y" unfolding facetrel_def by force lemma inj_on_pullback_facet: assumes "inj_on f x" "z \ f`x" obtains y where "y \ x" "f`y = z" proof from assms(2) obtain v where v: "v\z" "f`x = insert v z" using facetrel_def[of z] by auto define u and y where "u \ the_inv_into x f v" and y: "y \ {v\x. f v \ z}" moreover with assms(2) v have "x = insert u y" using the_inv_into_f_eq[OF assms(1)] the_inv_into_into[OF assms(1)] by fastforce ultimately show "y \ x" using v f_the_inv_into_f[OF assms(1)] by (force intro: facetrelI) from y assms(2) show "f`y = z" using facetrel_subset by fast qed subsubsection \Adjacency\ definition adjacent :: "'a set \ 'a set \ bool" (infix "\" 70) where "x \ y \ \z. z\x \ z\y" lemma adjacentI: "z\x \ z\y \ x \ y" using adjacent_def by fast lemma empty_not_adjacent: "\ {} \ x" unfolding facetrel_def adjacent_def by fast lemma adjacent_sym: "x \ y \ y \ x" unfolding adjacent_def by fast lemma adjacent_refl: assumes "x \ {}" shows "x \ x" proof- from assms obtain v where v: "v\x" by fast thus "x \ x" using facetrelI[of v "x-{v}"] unfolding adjacent_def by fast qed lemma common_facet: "\ z\x; z\y; x \ y \ \ z = x \ y" using facetrel_subset facet2_subset by fast lemma adjacent_int_facet1: "x \ y \ x \ y \ (x \ y) \ x" using common_facet unfolding adjacent_def by fast lemma adjacent_int_facet2: "x \ y \ x \ y \ (x \ y) \ y" using adjacent_sym adjacent_int_facet1 by (fastforce simp add: Int_commute) lemma adjacent_conv_insert: "x \ y \ v \ x - y \ x = insert v (x\y)" using adjacent_int_facet1 facetrel_conv_insert by fast lemma adjacent_int_decomp: "x \ y \ x \ y \ \v. v \ y \ x = insert v (x\y)" using adjacent_int_facet1 unfolding facetrel_def by fast lemma adj_antivertex: assumes "x\y" "x\y" shows "\!v. v\x-y" proof (rule ex_ex1I) from assms obtain w where w: "w\y" "x = insert w (x\y)" using adjacent_int_decomp by fast thus "\v. v\x-y" by auto from w have "\v. v\x-y \ v=w" by fast thus "\v v'. v\x-y \ v'\x-y \ v=v'" by auto qed lemma adjacent_card: "x \ y \ card x = card y" unfolding adjacent_def facetrel_def by (cases "finite x" "x=y" rule: two_cases) auto lemma adjacent_to_adjacent_int_subset: assumes "C \ D" "f`C \ f`D" "f`C \ f`D" shows "f`C \ f`D \ f`(C\D)" proof from assms(1,3) obtain v where v: "v \ D" "C = insert v (C\D)" using adjacent_int_decomp by fast from assms(2,3) obtain w where w: "w \ f`D" "f`C = insert w (f`C\f`D)" using adjacent_int_decomp[of "f`C" "f`D"] by fast from w have w': "w \ f`C - f`D" by fast with v assms(1,2) have fv_w: "f v = w" using adjacent_conv_insert by fast fix b assume "b \ f`C \ f`D" from this obtain a1 a2 where a1: "a1 \ C" "b = f a1" and a2: "a2 \ D" "b = f a2" by fast from v a1 a2(2) have "a1 \ D \ f a2 = w" using fv_w by auto with a2(1) w' have "a1 \ D" by fast with a1 show "b \ f`(C\D)" by fast qed lemma adjacent_to_adjacent_int: "\ C \ D; f`C \ f`D; f`C \ f`D \ \ f`(C\D) = f`C \ f`D" using adjacent_to_adjacent_int_subset by fast subsubsection \Chains of adjacent sets\ abbreviation "adjacentchain \ binrelchain adjacent" abbreviation "padjacentchain \ proper_binrelchain adjacent" lemmas adjacentchain_Cons_reduce = binrelchain_Cons_reduce [of adjacent] lemmas adjacentchain_obtain_proper = binrelchain_obtain_proper [of _ _ adjacent] lemma adjacentchain_card: "adjacentchain (x#xs@[y]) \ card x = card y" using adjacent_card by (induct xs arbitrary: x) auto subsection \Locale and basic facts\ locale SimplicialComplex = fixes X :: "'a set set" assumes finite_simplices: "\x\X. finite x" and faces : "x\X \ y\x \ y\X" context SimplicialComplex begin abbreviation "Subcomplex Y \ Y \ X \ SimplicialComplex Y" definition "maxsimp x \ x\X \ (\z\X. x\z \ z=x)" definition adjacentset :: "'a set \ 'a set set" where "adjacentset x = {y\X. x\y}" lemma finite_simplex: "x\X \ finite x" using finite_simplices by simp lemma singleton_simplex: "v\\X \ {v} \ X" using faces by auto lemma maxsimpI: "x \ X \ (\z. z\X \ x\z \ z=x) \ maxsimp x" using maxsimp_def by auto lemma maxsimpD_simplex: "maxsimp x \ x\X" using maxsimp_def by fast lemma maxsimpD_maximal: "maxsimp x \ z\X \ x\z \ z=x" using maxsimp_def by auto lemmas finite_maxsimp = finite_simplex[OF maxsimpD_simplex] lemma maxsimp_nempty: "X \ {{}} \ maxsimp x \ x \ {}" unfolding maxsimp_def by fast lemma maxsimp_vertices: "maxsimp x \ x\\X" using maxsimpD_simplex by fast lemma adjacentsetD_adj: "y \ adjacentset x \ x\y" using adjacentset_def by fast lemma max_in_subcomplex: "\ Subcomplex Y; y \ Y; maxsimp y \ \ SimplicialComplex.maxsimp Y y" using maxsimpD_maximal by (fast intro: SimplicialComplex.maxsimpI) lemma face_im: assumes "w \ X" "y \ f`w" defines "u \ {a\w. f a \ y}" shows "y \ f\X" using assms faces[of w u] image_eqI[of y "(`) f" u X] by fast lemma im_faces: "x \ f \ X \ y \ x \ y \ f \ X" using faces face_im[of _ y] by (cases "y={}") auto lemma map_is_simplicial_morph: "SimplicialComplex (f\X)" proof show "\x\f\X. finite x" using finite_simplices by fast show "\x y. x \f\X \ y\x \ y\f\X" using im_faces by fast qed lemma vertex_set_int: assumes "SimplicialComplex Y" shows "\(X\Y) = \X \ \Y" proof have "\v. v \ \X \ \Y \ v\ \(X\Y)" using faces SimplicialComplex.faces[OF assms] by auto thus "\(X\Y) \ \X \ \Y" by fast qed auto end (* context SimplicialComplex *) subsection \Chains of maximal simplices\ text \ Chains of maximal simplices (with respect to adjacency) will allow us to walk through chamber complexes. But there is much we can say about them in simplicial complexes. We will call a chain of maximal simplices proper (using the prefix \p\ as a naming convention to denote proper) if no maximal simplex appears more than once in the chain. (Some sources elect to call improper chains prechains, and reserve the name chain to describe a proper chain. And usually a slightly weaker notion of proper is used, requiring only that no maximal simplex appear twice in succession. But it essentially makes no difference, and we found it easier to use @{const distinct} rather than @{term "binrelchain not_equal"}.) \ context SimplicialComplex begin definition "maxsimpchain xs \ (\x\set xs. maxsimp x) \ adjacentchain xs" definition "pmaxsimpchain xs \ (\x\set xs. maxsimp x) \ padjacentchain xs" function min_maxsimpchain :: "'a set list \ bool" where "min_maxsimpchain [] = True" | "min_maxsimpchain [x] = maxsimp x" | "min_maxsimpchain (x#xs@[y]) = (x\y \ is_arg_min length (\zs. maxsimpchain (x#zs@[y])) xs)" by (auto, rule list_cases_Cons_snoc) termination by (relation "measure length") auto lemma maxsimpchain_snocI: "\ maxsimpchain (xs@[x]); maxsimp y; x\y \ \ maxsimpchain (xs@[x,y])" using maxsimpchain_def binrelchain_snoc maxsimpchain_def by auto lemma maxsimpchainD_maxsimp: "maxsimpchain xs \ x \ set xs \ maxsimp x" using maxsimpchain_def by fast lemma maxsimpchainD_adj: "maxsimpchain xs \ adjacentchain xs" using maxsimpchain_def by fast lemma maxsimpchain_CConsI: "\ maxsimp w; maxsimpchain (x#xs); w\x \ \ maxsimpchain (w#x#xs)" using maxsimpchain_def by auto lemma maxsimpchain_Cons_reduce: "maxsimpchain (x#xs) \ maxsimpchain xs" using adjacentchain_Cons_reduce maxsimpchain_def by fastforce lemma maxsimpchain_append_reduce1: "maxsimpchain (xs@ys) \ maxsimpchain xs" using binrelchain_append_reduce1 maxsimpchain_def by auto lemma maxsimpchain_append_reduce2: "maxsimpchain (xs@ys) \ maxsimpchain ys" using binrelchain_append_reduce2 maxsimpchain_def by auto lemma maxsimpchain_remdup_adj: "maxsimpchain (xs@[x,x]@ys) \ maxsimpchain (xs@[x]@ys)" using maxsimpchain_def binrelchain_remdup_adj by auto lemma maxsimpchain_rev: "maxsimpchain xs \ maxsimpchain (rev xs)" using maxsimpchainD_maxsimp adjacent_sym binrelchain_sym_rev[of adjacent] unfolding maxsimpchain_def by fastforce lemma maxsimpchain_overlap_join: "maxsimpchain (xs@[w]) \ maxsimpchain (w#ys) \ maxsimpchain (xs@w#ys)" using binrelchain_overlap_join maxsimpchain_def by auto lemma pmaxsimpchain: "pmaxsimpchain xs \ maxsimpchain xs" using maxsimpchain_def pmaxsimpchain_def by fast lemma pmaxsimpchainI_maxsimpchain: "maxsimpchain xs \ distinct xs \ pmaxsimpchain xs" using maxsimpchain_def pmaxsimpchain_def by fast lemma pmaxsimpchain_CConsI: "\ maxsimp w; pmaxsimpchain (x#xs); w\x; w \ set (x#xs) \ \ pmaxsimpchain (w#x#xs)" using pmaxsimpchain_def by auto lemmas pmaxsimpchainD_maxsimp = maxsimpchainD_maxsimp[OF pmaxsimpchain] lemmas pmaxsimpchainD_adj = maxsimpchainD_adj [OF pmaxsimpchain] lemma pmaxsimpchainD_distinct: "pmaxsimpchain xs \ distinct xs" using pmaxsimpchain_def by fast lemma pmaxsimpchain_Cons_reduce: "pmaxsimpchain (x#xs) \ pmaxsimpchain xs" using maxsimpchain_Cons_reduce pmaxsimpchain pmaxsimpchainD_distinct by (fastforce intro: pmaxsimpchainI_maxsimpchain) lemma pmaxsimpchain_append_reduce1: "pmaxsimpchain (xs@ys) \ pmaxsimpchain xs" using maxsimpchain_append_reduce1 pmaxsimpchain pmaxsimpchainD_distinct by (fastforce intro: pmaxsimpchainI_maxsimpchain) lemma maxsimpchain_obtain_pmaxsimpchain: assumes "x\y" "maxsimpchain (x#xs@[y])" shows "\ys. set ys \ set xs \ length ys \ length xs \ pmaxsimpchain (x#ys@[y])" proof- obtain ys where ys: "set ys \ set xs" "length ys \ length xs" "padjacentchain (x#ys@[y])" using maxsimpchainD_adj[OF assms(2)] adjacentchain_obtain_proper[OF assms(1)] by auto from ys(1) assms(2) have "\a\set (x#ys@[y]). maxsimp a" using maxsimpchainD_maxsimp by auto with ys show ?thesis unfolding pmaxsimpchain_def by auto qed lemma min_maxsimpchainD_maxsimpchain: assumes "min_maxsimpchain xs" shows "maxsimpchain xs" proof (cases xs rule: list_cases_Cons_snoc) case Nil thus ?thesis using maxsimpchain_def by simp next case Single with assms show ?thesis using maxsimpchain_def by simp next case Cons_snoc with assms show ?thesis using is_arg_minD1 by fastforce qed lemma min_maxsimpchainD_min_betw: "min_maxsimpchain (x#xs@[y]) \ maxsimpchain (x#ys@[y]) \ length ys \ length xs" using is_arg_minD2 by fastforce lemma min_maxsimpchainI_betw: assumes "x\y" "maxsimpchain (x#xs@[y])" "\ys. maxsimpchain (x#ys@[y]) \ length xs \ length ys" shows "min_maxsimpchain (x#xs@[y])" using assms by (simp add: is_arg_min_linorderI) lemma min_maxsimpchainI_betw_compare: assumes "x\y" "maxsimpchain (x#xs@[y])" "min_maxsimpchain (x#ys@[y])" "length xs = length ys" shows "min_maxsimpchain (x#xs@[y])" using assms min_maxsimpchainD_min_betw min_maxsimpchainI_betw by auto lemma min_maxsimpchain_pmaxsimpchain: assumes "min_maxsimpchain xs" shows "pmaxsimpchain xs" proof ( rule pmaxsimpchainI_maxsimpchain, rule min_maxsimpchainD_maxsimpchain, rule assms, cases xs rule: list_cases_Cons_snoc ) case (Cons_snoc x ys y) have "\ distinct (x#ys@[y]) \ False" proof (cases "x\set ys" "y\set ys" rule: two_cases) case both from both(1) obtain as bs where "ys = as@x#bs" using in_set_conv_decomp[of x ys] by fast with assms Cons_snoc show False using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_append_reduce2[of "x#as"] min_maxsimpchainD_min_betw[of x ys y] by fastforce next case one from one(1) obtain as bs where "ys = as@x#bs" using in_set_conv_decomp[of x ys] by fast with assms Cons_snoc show False using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_append_reduce2[of "x#as"] min_maxsimpchainD_min_betw[of x ys y] by fastforce next case other from other(2) obtain as bs where "ys = as@y#bs" using in_set_conv_decomp[of y ys] by fast with assms Cons_snoc show False using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_append_reduce1[of "x#as@[y]"] min_maxsimpchainD_min_betw[of x ys y] by fastforce next case neither moreover assume "\ distinct (x # ys @ [y])" ultimately obtain as a bs cs where "ys = as@[a]@bs@[a]@cs" using assms Cons_snoc not_distinct_decomp[of ys] by auto with assms Cons_snoc show False using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_append_reduce1[of "x#as@[a]"] maxsimpchain_append_reduce2[of "x#as@[a]@bs" "a#cs@[y]"] maxsimpchain_overlap_join[of "x#as" a "cs@[y]"] min_maxsimpchainD_min_betw[of x ys y "as@a#cs"] by auto qed with Cons_snoc show "distinct xs" by fast qed auto lemma min_maxsimpchain_rev: assumes "min_maxsimpchain xs" shows "min_maxsimpchain (rev xs)" proof (cases xs rule: list_cases_Cons_snoc) case Single with assms show ?thesis using min_maxsimpchainD_maxsimpchain maxsimpchainD_maxsimp by simp next case (Cons_snoc x ys y) moreover have "min_maxsimpchain (y # rev ys @ [x])" proof (rule min_maxsimpchainI_betw) from Cons_snoc assms show "y\x" using min_maxsimpchain_pmaxsimpchain pmaxsimpchainD_distinct by auto from Cons_snoc show "maxsimpchain (y # rev ys @ [x])" using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_rev by fastforce from Cons_snoc assms show "\zs. maxsimpchain (y#zs@[x]) \ length (rev ys) \ length zs" using maxsimpchain_rev min_maxsimpchainD_min_betw[of x ys y] by fastforce qed ultimately show ?thesis by simp qed simp lemma min_maxsimpchain_adj: "\ maxsimp x; maxsimp y; x\y; x\y \ \ min_maxsimpchain [x,y]" using maxsimpchain_def min_maxsimpchainI_betw[of x y "[]"] by simp lemma min_maxsimpchain_betw_CCons_reduce: assumes "min_maxsimpchain (w#x#ys@[z])" shows "min_maxsimpchain (x#ys@[z])" proof (rule min_maxsimpchainI_betw) from assms show "x\z" using min_maxsimpchain_pmaxsimpchain pmaxsimpchainD_distinct by fastforce show "maxsimpchain (x#ys@[z])" using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_Cons_reduce by fast next fix zs assume "maxsimpchain (x#zs@[z])" hence "maxsimpchain (w#x#zs@[z])" using min_maxsimpchainD_maxsimpchain[OF assms] maxsimpchain_def by fastforce with assms show "length ys \ length zs" using min_maxsimpchainD_min_betw[of w "x#ys" z "x#zs"] by simp qed lemma min_maxsimpchain_betw_uniform_length: assumes "min_maxsimpchain (x#xs@[y])" "min_maxsimpchain (x#ys@[y])" shows "length xs = length ys" using min_maxsimpchainD_min_betw[OF assms(1)] min_maxsimpchainD_min_betw[OF assms(2)] min_maxsimpchainD_maxsimpchain[OF assms(1)] min_maxsimpchainD_maxsimpchain[OF assms(2)] by fastforce lemma not_min_maxsimpchainI_betw: "\ maxsimpchain (x#ys@[y]); length ys < length xs \ \ \ min_maxsimpchain (x#xs@[y])" using min_maxsimpchainD_min_betw not_less by blast lemma maxsimpchain_in_subcomplex: "\ Subcomplex Y; set ys \ Y; maxsimpchain ys \ \ SimplicialComplex.maxsimpchain Y ys" using maxsimpchain_def max_in_subcomplex SimplicialComplex.maxsimpchain_def by force end (* context SimplicialComplex *) subsection \Isomorphisms of simplicial complexes\ text \ Here we develop the concept of isomorphism of simplicial complexes. Note that we have not bothered to first develop the concept of morphism of simplicial complexes, since every function on the vertex set of a simplicial complex can be considered a morphism of complexes (see lemma \map_is_simplicial_morph\ above). \ locale SimplicialComplexIsomorphism = SimplicialComplex X for X :: "'a set set" + fixes f :: "'a \ 'b" assumes inj: "inj_on f (\X)" begin lemmas morph = map_is_simplicial_morph[of f] lemma iso_codim_map: "x \ X \ y \ X \ card (f`x - f`y) = card (x-y)" using inj inj_on_image_set_diff[of f _ x y] finite_simplex subset_inj_on[of f _ "x-y"] inj_on_iff_eq_card[of "x-y"] by fastforce lemma maxsimp_im_max: "maxsimp x \ w \ X \ f`x \ f`w \ f`w = f`x" using maxsimpD_simplex inj_onD[OF inj] maxsimpD_maximal[of x w] by blast lemma maxsimp_map: "maxsimp x \ SimplicialComplex.maxsimp (f\X) (f`x)" using maxsimpD_simplex maxsimp_im_max morph SimplicialComplex.maxsimpI[of "f\X" "f`x"] by fastforce lemma iso_adj_int_im: assumes "maxsimp x" "maxsimp y" "x\y" "x\y" shows "(f`x \ f`y) \ f`x" proof (rule facetrelI_card) from assms(1,2) have 1: "f ` x \ f ` y \ f ` y = f ` x" using maxsimp_map SimplicialComplex.maxsimpD_simplex[OF morph] SimplicialComplex.maxsimpD_maximal[OF morph] by simp thus "f`x \ f`y \ f`x" by fast from assms(1) have "card (f`x - f`x \ f`y) \ card (f`x - f`(x\y))" using finite_maxsimp card_mono[of "f`x - f`(x\y)" "f`x - f`x \ f`y"] by fast moreover from assms(1,3,4) have "card (f`x - f`(x\y)) = 1" using maxsimpD_simplex faces[of x] maxsimpD_simplex iso_codim_map adjacent_int_facet1[of x y] facetrel_card by fastforce ultimately have "card (f`x - f`x \ f`y) \ 1" by simp moreover from assms(1,2,4) have "card (f`x - f`x \ f`y) \ 0" using 1 maxsimpD_simplex finite_maxsimp inj_onD[OF induced_pow_fun_inj_on, OF inj, of x y] by auto ultimately show "card (f`x - f`x \ f`y) = 1" by simp qed lemma iso_adj_map: assumes "maxsimp x" "maxsimp y" "x\y" "x\y" shows "f`x \ f`y" using assms(3,4) iso_adj_int_im[OF assms] adjacent_sym iso_adj_int_im[OF assms(2) assms(1)] by (auto simp add: Int_commute intro: adjacentI) lemma pmaxsimpchain_map: "pmaxsimpchain xs \ SimplicialComplex.pmaxsimpchain (f\X) (f\xs)" proof (induct xs rule: list_induct_CCons) case Nil show ?case using map_is_simplicial_morph SimplicialComplex.pmaxsimpchain_def by fastforce next case (Single x) thus ?case using map_is_simplicial_morph pmaxsimpchainD_maxsimp maxsimp_map SimplicialComplex.pmaxsimpchain_def by fastforce next case (CCons x y xs) have "SimplicialComplex.pmaxsimpchain (f \ X) ( f`x # f`y # f\xs)" proof ( rule SimplicialComplex.pmaxsimpchain_CConsI, rule map_is_simplicial_morph ) from CCons(2) show "SimplicialComplex.maxsimp (f\X) (f`x)" using pmaxsimpchainD_maxsimp maxsimp_map by simp from CCons show "SimplicialComplex.pmaxsimpchain (f\X) (f`y # f\xs)" using pmaxsimpchain_Cons_reduce by simp from CCons(2) show "f`x \ f`y" using pmaxsimpchain_def iso_adj_map by simp from inj CCons(2) have "distinct (f\(x#y#xs))" using maxsimpD_simplex inj_on_distinct_setlistmapim unfolding pmaxsimpchain_def by blast thus "f`x \ set (f`y # f\xs)" by simp qed thus ?case by simp qed end (* context SimplicialComplexIsomorphism *) subsection \The complex associated to a poset\ text \ A simplicial complex is naturally a poset under the subset relation. The following develops the reverse direction: constructing a simplicial complex from a suitable poset. \ context ordering begin definition PosetComplex :: "'a set \ 'a set set" where "PosetComplex P \ (\x\P. { {y. pseudominimal_in (P.\<^bold>\x) y} })" lemma poset_is_SimplicialComplex: assumes "\x\P. simplex_like (P.\<^bold>\x)" shows "SimplicialComplex (PosetComplex P)" proof (rule SimplicialComplex.intro, rule ballI) fix a assume "a \ PosetComplex P" from this obtain x where "x\P" "a = {y. pseudominimal_in (P.\<^bold>\x) y}" unfolding PosetComplex_def by fast with assms show "finite a" using pseudominimal_inD1 simplex_likeD_finite finite_subset[of a "P.\<^bold>\x"] by fast next fix a b assume ab: "a \ PosetComplex P" "b\a" from ab(1) obtain x where x: "x\P" "a = {y. pseudominimal_in (P.\<^bold>\x) y}" unfolding PosetComplex_def by fast from assms x(1) obtain f and A::"nat set" where fA: "OrderingSetIso less_eq less (\) (\) (P.\<^bold>\x) f" "f`(P.\<^bold>\x) = Pow A" using simplex_likeD_iso[of "P.\<^bold>\x"] by auto define x' where x': "x' \ the_inv_into (P.\<^bold>\x) f (\(f`b))" from fA x(2) ab(2) x' have x'_P: "x'\P" using collect_pseudominimals_below_in_poset[of P x f] by simp moreover from x fA ab(2) x' have "b = {y. pseudominimal_in (P.\<^bold>\x') y}" using collect_pseudominimals_below_in_eq[of x P f] by simp ultimately show "b \ PosetComplex P" unfolding PosetComplex_def by fast qed definition poset_simplex_map :: "'a set \ 'a \ 'a set" where "poset_simplex_map P x = {y. pseudominimal_in (P.\<^bold>\x) y}" lemma poset_to_PosetComplex_OrderingSetMap: assumes "\x. x\P \ simplex_like (P.\<^bold>\x)" shows "OrderingSetMap (\<^bold>\) (\<^bold><) (\) (\) P (poset_simplex_map P)" proof from assms show "\a b. \ a\P; b\P; a\<^bold>\b \ \ poset_simplex_map P a \ poset_simplex_map P b" using simplex_like_has_bottom pseudominimal_in_below_in unfolding poset_simplex_map_def by fast qed end (* context ordering *) text \ When a poset affords a simplicial complex, there is a natural morphism of posets from the source poset into the poset of sets in the complex, as above. However, some further assumptions are necessary to ensure that this morphism is an isomorphism. These conditions are collected in the following locale. \ locale ComplexLikePoset = ordering less_eq less for less_eq :: "'a\'a\bool" (infix "\<^bold>\" 50) and less :: "'a\'a\bool" (infix "\<^bold><" 50) + fixes P :: "'a set" assumes below_in_P_simplex_like: "x\P \ simplex_like (P.\<^bold>\x)" and P_has_bottom : "has_bottom P" and P_has_glbs : "x\P \ y\P \ \b. glbound_in_of P x y b" begin abbreviation "smap \ poset_simplex_map P" lemma smap_onto_PosetComplex: "smap ` P = PosetComplex P" using poset_simplex_map_def PosetComplex_def by auto lemma ordsetmap_smap: "\ a\P; b\P; a\<^bold>\b \ \ smap a \ smap b" using OrderingSetMap.ordsetmap[ OF poset_to_PosetComplex_OrderingSetMap, OF below_in_P_simplex_like ] poset_simplex_map_def by simp lemma inj_on_smap: "inj_on smap P" proof (rule inj_onI) fix x y assume xy: "x\P" "y\P" "smap x = smap y" show "x = y" proof (cases "smap x = {}") case True with xy show ?thesis using poset_simplex_map_def below_in_P_simplex_like P_has_bottom simplex_like_no_pseudominimal_in_below_in_imp_singleton[of x P] simplex_like_no_pseudominimal_in_below_in_imp_singleton[of y P] below_in_singleton_is_bottom[of P x] below_in_singleton_is_bottom[of P y] by auto next case False from this obtain z where "z \ smap x" by fast with xy(3) have z1: "z \ P.\<^bold>\x" "z \ P.\<^bold>\y" using pseudominimal_inD1 poset_simplex_map_def by auto hence "lbound_of x y z" by (auto intro: lbound_ofI) with z1(1) obtain b where b: "glbound_in_of P x y b" using xy(1,2) P_has_glbs by fast moreover have "b \ P.\<^bold>\x" "b \ P.\<^bold>\y" using glbound_in_ofD_in[OF b] glbound_in_of_less_eq1[OF b] glbound_in_of_less_eq2[OF b] by auto ultimately show ?thesis using xy below_in_P_simplex_like pseudominimal_in_below_in_less_eq_glbound[of P x _ y b] simplex_like_below_in_above_pseudominimal_is_top[of x P] simplex_like_below_in_above_pseudominimal_is_top[of y P] unfolding poset_simplex_map_def by force qed qed lemma OrderingSetIso_smap: "OrderingSetIso (\<^bold>\) (\<^bold><) (\) (\) P smap" proof (rule OrderingSetMap.isoI) show "OrderingSetMap (\<^bold>\) (\<^bold><) (\) (\) P smap" using poset_simplex_map_def below_in_P_simplex_like poset_to_PosetComplex_OrderingSetMap by simp next fix x y assume xy: "x\P" "y\P" "smap x \ smap y" from xy(2) have "simplex_like (P.\<^bold>\y)" using below_in_P_simplex_like by fast from this obtain g and A::"nat set" where "OrderingSetIso (\<^bold>\) (\<^bold><) (\) (\) (P.\<^bold>\y) g" "g`(P.\<^bold>\y) = Pow A" using simplex_likeD_iso[of "P.\<^bold>\y"] by auto with xy show "x\<^bold>\y" using poset_simplex_map_def collect_pseudominimals_below_in_eq[of y P g] collect_pseudominimals_below_in_poset[of P y g] inj_onD[OF inj_on_smap, of "the_inv_into (P.\<^bold>\y) g (\(g ` smap x))" x] collect_pseudominimals_below_in_less_eq_top[of P y g A "smap x"] by simp qed (rule inj_on_smap) lemmas rev_ordsetmap_smap = OrderingSetIso.rev_ordsetmap[OF OrderingSetIso_smap] end (* context ComplexLikePoset *) end (* theory *) diff --git a/thys/Design_Theory/Block_Designs.thy b/thys/Design_Theory/Block_Designs.thy --- a/thys/Design_Theory/Block_Designs.thy +++ b/thys/Design_Theory/Block_Designs.thy @@ -1,551 +1,551 @@ (* Title: Block_Designs.thy Author: Chelsea Edmonds *) section \Block and Balanced Designs\ text \We define a selection of the many different types of block and balanced designs, building up to properties required for defining a BIBD, in addition to several base generalisations\ theory Block_Designs imports Design_Operations begin subsection \Block Designs\ text \A block design is a design where all blocks have the same size.\ subsubsection \K Block Designs\ text \An important generalisation of a typical block design is the $\mathcal{K}$ block design, where all blocks must have a size $x$ where $x \in \mathcal{K}$\ locale K_block_design = proper_design + fixes sizes :: "nat set" ("\") assumes block_sizes: "bl \# \ \ card bl \ \" assumes positive_ints: "x \ \ \ x > 0" begin lemma sys_block_size_subset: "sys_block_sizes \ \" using block_sizes sys_block_sizes_obtain_bl by blast end subsubsection\Uniform Block Design\ text \The typical uniform block design is defined below\ locale block_design = proper_design + fixes u_block_size :: nat ("\") assumes uniform [simp]: "bl \# \ \ card bl = \" begin lemma k_non_zero: "\ \ 1" proof - obtain bl where bl_in: "bl \# \" using design_blocks_nempty by auto then have "card bl \ 1" using block_size_gt_0 by (metis less_not_refl less_one not_le_imp_less) thus ?thesis by (simp add: bl_in) qed lemma uniform_alt_def_all: "\ bl \# \ .card bl = \" using uniform by auto lemma uniform_unfold_point_set: "bl \# \ \ card {p \ \. p \ bl} = \" using uniform wellformed by (simp add: Collect_conj_eq inf.absorb_iff2) lemma uniform_unfold_point_set_mset: "bl \# \ \ size {#p \# mset_set \. p \ bl #} = \" using uniform_unfold_point_set by (simp add: finite_sets) lemma sys_block_sizes_uniform [simp]: "sys_block_sizes = {\}" proof - have "sys_block_sizes = {bs . \ bl . bs = card bl \ bl\# \}" by (simp add: sys_block_sizes_def) then have "sys_block_sizes = {bs . bs = \}" using uniform uniform_unfold_point_set b_positive block_set_nempty_imp_block_ex by (smt (verit, best) Collect_cong design_blocks_nempty) thus ?thesis by auto qed lemma sys_block_sizes_uniform_single: "is_singleton (sys_block_sizes)" by simp lemma uniform_size_incomp: "\ \ \ - 1 \ bl \# \ \ incomplete_block bl" using uniform k_non_zero by (metis block_size_lt_v diff_diff_cancel diff_is_0_eq' less_numeral_extra(1) nat_less_le) lemma uniform_complement_block_size: assumes "bl \# \\<^sup>C" shows "card bl = \ - \" proof - obtain bl' where bl_assm: "bl = bl'\<^sup>c \ bl' \# \" using wellformed assms by (auto simp add: complement_blocks_def) then have "int (card bl') = \" by simp thus ?thesis using bl_assm block_complement_size wellformed by (simp add: block_size_lt_order of_nat_diff) qed lemma uniform_complement[intro]: assumes "\ \ \ - 1" shows "block_design \ \\<^sup>C (\ - \)" proof - interpret des: proper_design \ "\\<^sup>C" using uniform_size_incomp assms complement_proper_design by auto show ?thesis using assms uniform_complement_block_size by (unfold_locales) (simp) qed lemma block_size_lt_v: "\ \ \" using v_non_zero block_size_lt_v design_blocks_nempty uniform by auto end lemma (in proper_design) block_designI[intro]: "(\ bl . bl \# \ \ card bl = k) \ block_design \ \ k" by (unfold_locales) (auto) context block_design begin lemma block_design_multiple: "n > 0 \ block_design \ (multiple_blocks n) \" using elem_in_repeat_in_original multiple_proper_design proper_design.block_designI by (metis uniform_alt_def_all) end text \A uniform block design is clearly a type of $K$\_block\_design with a singleton $K$ set\ sublocale block_design \ K_block_design \ \ "{\}" using k_non_zero uniform by unfold_locales simp_all subsubsection \Incomplete Designs\ text \An incomplete design is a design where $k < v$, i.e. no block is equal to the point set\ locale incomplete_design = block_design + assumes incomplete: "\ < \" begin lemma incomplete_imp_incomp_block: "bl \# \ \ incomplete_block bl" using incomplete uniform uniform_size_incomp by fastforce lemma incomplete_imp_proper_subset: "bl \# \ \ bl \ \" using incomplete_block_proper_subset incomplete_imp_incomp_block by auto end lemma (in block_design) incomplete_designI[intro]: "\ < \ \ incomplete_design \ \ \" by unfold_locales auto context incomplete_design begin lemma multiple_incomplete: "n > 0 \ incomplete_design \ (multiple_blocks n) \" using block_design_multiple incomplete by (simp add: block_design.incomplete_designI) lemma complement_incomplete: "incomplete_design \ (\\<^sup>C) (\ - \)" proof - have "\ - \ < \" using v_non_zero k_non_zero by linarith thus ?thesis using uniform_complement incomplete incomplete_designI by (simp add: block_design.incomplete_designI) qed end subsection \Balanced Designs\ text \t-wise balance is a design with the property that all point subsets of size $t$ occur in $\lambda_t$ blocks\ locale t_wise_balance = proper_design + fixes grouping :: nat ("\") and index :: nat ("\\<^sub>t") assumes t_non_zero: "\ \ 1" assumes t_lt_order: "\ \ \" assumes balanced [simp]: "ps \ \ \ card ps = \ \ \ index ps = \\<^sub>t" begin lemma t_non_zero_suc: "\ \ Suc 0" using t_non_zero by auto lemma balanced_alt_def_all: "\ ps \ \ . card ps = \ \ \ index ps = \\<^sub>t" using balanced by auto end lemma (in proper_design) t_wise_balanceI[intro]: "\ \ \ \ \ \ 1 \ (\ ps . ps \ \ \ card ps = \ \ \ index ps = \\<^sub>t) \ t_wise_balance \ \ \ \\<^sub>t" by (unfold_locales) auto context t_wise_balance begin lemma obtain_t_subset_points: obtains T where "T \ \" "card T = \" "finite T" - using obtain_subset_with_card_n design_points_nempty t_lt_order t_non_zero finite_sets by auto + using obtain_subset_with_card_n t_lt_order by auto lemma multiple_t_wise_balance_index [simp]: assumes "ps \ \" assumes "card ps = \" shows "(multiple_blocks n) index ps = \\<^sub>t * n" using multiple_point_index balanced assms by fastforce lemma multiple_t_wise_balance: assumes "n > 0" shows "t_wise_balance \ (multiple_blocks n) \ (\\<^sub>t * n)" proof - interpret des: proper_design \ "(multiple_blocks n)" by (simp add: assms multiple_proper_design) show ?thesis using t_non_zero t_lt_order multiple_t_wise_balance_index by (unfold_locales) (simp_all) qed lemma twise_set_pair_index: "ps \ \ \ ps2 \ \ \ ps \ ps2 \ card ps = \ \ card ps2 = \ \ \ index ps = \ index ps2" using balanced by simp lemma t_wise_balance_alt: "ps \ \ \ card ps = \ \ \ index ps = l2 \ (\ ps . ps \ \ \ card ps = \ \ \ index ps = l2)" using twise_set_pair_index by blast lemma index_1_imp_mult_1 [simp]: assumes "\\<^sub>t = 1" assumes "bl \# \" assumes "card bl \ \" shows "multiplicity bl = 1" proof (rule ccontr) assume "\ (multiplicity bl = 1)" then have not: "multiplicity bl \ 1" by simp have "multiplicity bl \ 0" using assms by simp then have m: "multiplicity bl \ 2" using not by linarith obtain ps where ps: "ps \ bl \ card ps = \" using assms obtain_t_subset_points by (metis obtain_subset_with_card_n) then have "\ index ps \ 2" using m points_index_count_min ps by blast then show False using balanced ps antisym_conv2 not_numeral_less_zero numeral_le_one_iff points_index_ps_nin semiring_norm(69) zero_neq_numeral by (metis assms(1)) qed end subsubsection \Sub-types of t-wise balance\ text \Pairwise balance is when $t = 2$. These are commonly of interest\ locale pairwise_balance = t_wise_balance \ \ 2 \ for point_set ("\") and block_collection ("\") and index ("\") text \We can combine the balance properties with $K$\_block design to define tBD's (t-wise balanced designs), and PBD's (pairwise balanced designs)\ locale tBD = t_wise_balance + K_block_design + assumes block_size_gt_t: "k \ \ \ k \ \" locale \_PBD = pairwise_balance + K_block_design + assumes block_size_gt_t: "k \ \ \ k \ 2" sublocale \_PBD \ tBD \ \ 2 \ \ using t_lt_order block_size_gt_t by (unfold_locales) (simp_all) locale PBD = \_PBD \ \ 1 \ for point_set ("\") and block_collection ("\") and sizes ("\") begin lemma multiplicity_is_1: assumes "bl \# \" shows "multiplicity bl = 1" using block_size_gt_t index_1_imp_mult_1 by (simp add: assms block_sizes) end sublocale PBD \ simple_design using multiplicity_is_1 by (unfold_locales) text \PBD's are often only used in the case where $k$ is uniform, defined here.\ locale k_\_PBD = pairwise_balance + block_design + assumes block_size_t: "2 \ \" sublocale k_\_PBD \ \_PBD \ \ \ "{\}" using k_non_zero uniform block_size_t by(unfold_locales) (simp_all) locale k_PBD = k_\_PBD \ \ 1 \ for point_set ("\") and block_collection ("\") and u_block_size ("\") sublocale k_PBD \ PBD \ \ "{\}" using block_size_t by (unfold_locales, simp_all) subsubsection \Covering and Packing Designs\ text \Covering and packing designs involve a looser balance restriction. Upper/lower bounds are placed on the points index, instead of a strict equality\ text \A t-covering design is a relaxed version of a tBD, where, for all point subsets of size t, a lower bound is put on the points index\ locale t_covering_design = block_design + fixes grouping :: nat ("\") fixes min_index :: nat ("\\<^sub>t") assumes covering: "ps \ \ \ card ps = \ \ \ index ps \ \\<^sub>t" assumes block_size_t: "\ \ \" assumes t_non_zero: "\ \ 1" begin lemma covering_alt_def_all: "\ ps \ \ . card ps = \ \ \ index ps \ \\<^sub>t" using covering by auto end lemma (in block_design) t_covering_designI [intro]: "t \ \ \ t \ 1 \ (\ ps. ps \ \ \ card ps = t \ \ index ps \ \\<^sub>t) \ t_covering_design \ \ \ t \\<^sub>t" by (unfold_locales) simp_all text \A t-packing design is a relaxed version of a tBD, where, for all point subsets of size t, an upper bound is put on the points index\ locale t_packing_design = block_design + fixes grouping :: nat ("\") fixes min_index :: nat ("\\<^sub>t") assumes packing: "ps \ \ \ card ps = \ \ \ index ps \ \\<^sub>t" assumes block_size_t: "\ \ \" assumes t_non_zero: "\ \ 1" begin lemma packing_alt_def_all: "\ ps \ \ . card ps = \ \ \ index ps \ \\<^sub>t" using packing by auto end lemma (in block_design) t_packing_designI [intro]: "t \ \ \ t \ 1 \ (\ ps . ps \ \ \ card ps = t \ \ index ps \ \\<^sub>t) \ t_packing_design \ \ \ t \\<^sub>t" by (unfold_locales) simp_all lemma packing_covering_imp_balance: assumes "t_packing_design V B k t \\<^sub>t" assumes "t_covering_design V B k t \\<^sub>t" shows "t_wise_balance V B t \\<^sub>t" proof - from assms interpret des: proper_design V B using block_design.axioms(1) t_covering_design.axioms(1) by blast show ?thesis proof (unfold_locales) show "1 \ t" using assms t_packing_design.t_non_zero by auto show "t \ des.\" using block_design.block_size_lt_v t_packing_design.axioms(1) by (metis assms(1) dual_order.trans t_packing_design.block_size_t) show "\ps. ps \ V \ card ps = t \ B index ps = \\<^sub>t" using t_packing_design.packing t_covering_design.covering by (metis assms dual_order.antisym) qed qed subsection \Constant Replication Design\ text \When the replication number for all points in a design is constant, it is the design replication number.\ locale constant_rep_design = proper_design + fixes design_rep_number :: nat ("\") assumes rep_number [simp]: "x \ \ \ \ rep x = \" begin lemma rep_number_alt_def_all: "\ x \ \. \ rep x = \" by (simp) lemma rep_number_unfold_set: "x \ \ \ size {#bl \# \ . x \ bl#} = \" using rep_number by (simp add: point_replication_number_def) lemma rep_numbers_constant [simp]: "replication_numbers = {\}" unfolding replication_numbers_def using rep_number design_points_nempty Collect_cong finite.cases finite_sets insertCI singleton_conv by (smt (verit, ccfv_threshold) fst_conv snd_conv) lemma replication_number_single: "is_singleton (replication_numbers)" using is_singleton_the_elem by simp lemma constant_rep_point_pair: "x1 \ \ \ x2 \ \ \ x1 \ x2 \ \ rep x1 = \ rep x2" using rep_number by auto lemma constant_rep_alt: "x1 \ \ \ \ rep x1 = r2 \ (\ x . x \ \ \ \ rep x = r2)" by (simp) lemma constant_rep_point_not_0: assumes "x \ \" shows "\ rep x \ 0" proof (rule ccontr) assume "\ \ rep x \ 0" then have "\ x . x \ \ \ \ rep x = 0" using rep_number assms by auto then have "\ x . x \ \ \ size {#bl \# \ . x \ bl#} = 0" by (simp add: point_replication_number_def) then show False using design_blocks_nempty wf_design wf_design_iff wf_invalid_point by (metis ex_in_conv filter_mset_empty_conv multiset_nonemptyE size_eq_0_iff_empty) qed lemma rep_not_zero: "\ \ 0" using rep_number constant_rep_point_not_0 design_points_nempty by auto lemma r_gzero: "\ > 0" using rep_not_zero by auto lemma r_lt_eq_b: "\ \ \" using rep_number max_point_rep by (metis all_not_in_conv design_points_nempty) lemma complement_rep_number: assumes "\ bl . bl \# \ \ incomplete_block bl" shows "constant_rep_design \ \\<^sup>C (\ - \)" proof - interpret d: proper_design \ "(\\<^sup>C)" using complement_proper_design by (simp add: assms) show ?thesis using complement_rep_number rep_number by (unfold_locales) simp qed lemma multiple_rep_number: assumes "n > 0" shows "constant_rep_design \ (multiple_blocks n) (\ * n)" proof - interpret d: proper_design \ "(multiple_blocks n)" using multiple_proper_design by (simp add: assms) show ?thesis using multiple_point_rep_num by (unfold_locales) (simp_all) qed end lemma (in proper_design) constant_rep_designI [intro]: "(\ x . x \ \ \ \ rep x = \) \ constant_rep_design \ \ \" by unfold_locales auto subsection \T-designs\ text \All the before mentioned designs build up to the concept of a t-design, which has uniform block size and is t-wise balanced. We limit $t$ to be less than $k$, so the balance condition has relevance\ locale t_design = incomplete_design + t_wise_balance + assumes block_size_t: "\ \ \" begin lemma point_indices_balanced: "point_indices \ = {\\<^sub>t}" proof - have "point_indices \ = {i . \ ps . i = \ index ps \ card ps = \ \ ps \ \}" by (simp add: point_indices_def) then have "point_indices \ = {i . i = \\<^sub>t}" using balanced Collect_cong obtain_t_subset_points by (smt (verit, best)) thus ?thesis by auto qed lemma point_indices_singleton: "is_singleton (point_indices \)" using point_indices_balanced is_singleton_the_elem by simp end lemma t_designI [intro]: assumes "incomplete_design V B k" assumes "t_wise_balance V B t \\<^sub>t" assumes "t \ k" shows "t_design V B k t \\<^sub>t" by (simp add: assms(1) assms(2) assms(3) t_design.intro t_design_axioms.intro) sublocale t_design \ t_covering_design \ \ \ \ \\<^sub>t using t_non_zero by (unfold_locales) (auto simp add: block_size_t) sublocale t_design \ t_packing_design \ \ \ \ \\<^sub>t using t_non_zero by (unfold_locales) (auto simp add: block_size_t) lemma t_design_pack_cov [intro]: assumes "k < card V" assumes "t_covering_design V B k t \\<^sub>t" assumes "t_packing_design V B k t \\<^sub>t" shows "t_design V B k t \\<^sub>t" proof - from assms interpret id: incomplete_design V B k using block_design.incomplete_designI t_packing_design.axioms(1) by blast from assms interpret balance: t_wise_balance V B t \\<^sub>t using packing_covering_imp_balance by blast show ?thesis using assms(3) by (unfold_locales) (simp_all add: t_packing_design.block_size_t) qed sublocale t_design \ tBD \ \ \ \\<^sub>t "{\}" using uniform k_non_zero block_size_t by (unfold_locales) simp_all context t_design begin lemma multiple_t_design: "n > 0 \ t_design \ (multiple_blocks n) \ \ (\\<^sub>t * n)" using multiple_t_wise_balance multiple_incomplete block_size_t by (simp add: t_designI) lemma t_design_min_v: "\ > 1" using k_non_zero incomplete by simp end subsection \Steiner Systems\ text \Steiner systems are a special type of t-design where $\Lambda_t = 1$\ locale steiner_system = t_design \ \ \ \ 1 for point_set ("\") and block_collection ("\") and u_block_size ("\") and grouping ("\") begin lemma block_multiplicity [simp]: assumes "bl \# \" shows "multiplicity bl = 1" by (simp add: assms block_size_t) end sublocale steiner_system \ simple_design by unfold_locales (simp) lemma (in t_design) steiner_systemI[intro]: "\\<^sub>t = 1 \ steiner_system \ \ \ \" using t_non_zero t_lt_order block_size_t by unfold_locales auto subsection \Combining block designs\ text \We define some closure properties for various block designs under the combine operator. This is done using locales to reason on multiple instances of the same type of design, building on what was presented in the design operations theory\ locale two_t_wise_eq_points = two_designs_proper \ \ \ \' + des1: t_wise_balance \ \ \ \\<^sub>t + des2: t_wise_balance \ \' \ \\<^sub>t' for \ \ \ \\<^sub>t \' \\<^sub>t' begin lemma combine_t_wise_balance_index: "ps \ \ \ card ps = \ \ \\<^sup>+ index ps = (\\<^sub>t + \\<^sub>t')" using des1.balanced des2.balanced by (simp add: combine_points_index) lemma combine_t_wise_balance: "t_wise_balance \\<^sup>+ \\<^sup>+ \ (\\<^sub>t + \\<^sub>t')" proof (unfold_locales, simp add: des1.t_non_zero_suc) have "card \\<^sup>+ \ card \" by simp then show "\ \ card (\\<^sup>+)" using des1.t_lt_order by linarith show "\ps. ps \ \\<^sup>+ \ card ps = \ \ (\\<^sup>+ index ps) = \\<^sub>t + \\<^sub>t'" using combine_t_wise_balance_index by blast qed sublocale combine_t_wise_des: t_wise_balance "\\<^sup>+" "\\<^sup>+" "\" "(\\<^sub>t + \\<^sub>t')" using combine_t_wise_balance by auto end locale two_k_block_designs = two_designs_proper \ \ \' \' + des1: block_design \ \ \ + des2: block_design \' \' \ for \ \ \ \' \' begin lemma block_design_combine: "block_design \\<^sup>+ \\<^sup>+ \" using des1.uniform des2.uniform by (unfold_locales) (auto) sublocale combine_block_des: block_design "\\<^sup>+" "\\<^sup>+" "\" using block_design_combine by simp end locale two_rep_designs_eq_points = two_designs_proper \ \ \ \' + des1: constant_rep_design \ \ \ + des2: constant_rep_design \ \' \' for \ \ \ \' \' begin lemma combine_rep_number: "constant_rep_design \\<^sup>+ \\<^sup>+ (\ + \')" using combine_rep_number des1.rep_number des2.rep_number by (unfold_locales) (simp) sublocale combine_const_rep: constant_rep_design "\\<^sup>+" "\\<^sup>+" "(\ + \')" using combine_rep_number by simp end locale two_incomplete_designs = two_k_block_designs \ \ \ \' \' + des1: incomplete_design \ \ \ + des2: incomplete_design \' \' \ for \ \ \ \' \' begin lemma combine_is_incomplete: "incomplete_design \\<^sup>+ \\<^sup>+ \" using combine_order des1.incomplete des2.incomplete by (unfold_locales) (simp) sublocale combine_incomplete: incomplete_design "\\<^sup>+" "\\<^sup>+" "\" using combine_is_incomplete by simp end locale two_t_designs_eq_points = two_incomplete_designs \ \ \ \ \' + two_t_wise_eq_points \ \ \ \\<^sub>t \' \\<^sub>t' + des1: t_design \ \ \ \ \\<^sub>t + des2: t_design \ \' \ \ \\<^sub>t' for \ \ \ \' \ \\<^sub>t \\<^sub>t' begin lemma combine_is_t_des: "t_design \\<^sup>+ \\<^sup>+ \ \ (\\<^sub>t + \\<^sub>t')" using des1.block_size_t des2.block_size_t by (unfold_locales) sublocale combine_t_des: t_design "\\<^sup>+" "\\<^sup>+" "\" "\" "(\\<^sub>t + \\<^sub>t')" using combine_is_t_des by blast end end \ No newline at end of file diff --git a/thys/Khovanskii_Theorem/Khovanskii.thy b/thys/Khovanskii_Theorem/Khovanskii.thy --- a/thys/Khovanskii_Theorem/Khovanskii.thy +++ b/thys/Khovanskii_Theorem/Khovanskii.thy @@ -1,1258 +1,1258 @@ section\Khovanskii's Theorem\ text\We formalise the proof of an important theorem in additive combinatorics due to Khovanskii, attesting that the cardinality of the set of all sums of $n$ many elements of $A$, where $A$ is a finite subset of an abelian group, is a polynomial in $n$ for all sufficiently large $n$. We follow a proof due to Nathanson and Ruzsa as presented in the notes “Introduction to Additive Combinatorics” by Timothy Gowers for the University of Cambridge.\ theory Khovanskii imports Complex_Main FiniteProduct "HOL-Library.Equipollence" "Pluennecke_Ruzsa_Inequality.Pluennecke_Ruzsa_Inequality" "Bernoulli.Bernoulli" \ \sums of a fixed power are polynomials\ "HOL-Analysis.Weierstrass_Theorems" \ \needed for polynomial function\ "HOL-Library.List_Lenlexorder" \ \lexicographic ordering for the type @{typ \nat list\}\ begin lemma real_polynomial_function_sum_of_powers: "\p. real_polynomial_function p \ (\n. (\i\n. real i ^ j) = p (real n))" proof (intro exI conjI strip) let ?p = "\n. (bernpoly (Suc j) (1 + n) - bernpoly (Suc j) 0) / (Suc j)" show "real_polynomial_function ?p" by (force simp add: bernpoly_def) qed (simp add: add.commute sum_of_powers) text \The sum of the elements of a list\ abbreviation "\ \ sum_list" text \Related to the nsets of Ramsey.thy, but simpler\ definition finsets :: "['a set, nat] \ 'a set set" where "finsets A n \ {N. N \ A \ card N = n}" lemma card_finsets: "finite N \ card (finsets N k) = card N choose k" by (simp add: finsets_def n_subsets) lemma sorted_map_plus_iff [simp]: fixes a::"'a::linordered_cancel_ab_semigroup_add" shows "sorted (map ((+) a) xs) \ sorted xs" by (induction xs) auto lemma distinct_map_plus_iff [simp]: fixes a::"'a::linordered_cancel_ab_semigroup_add" shows "distinct (map ((+) a) xs) \ distinct xs" by (induction xs) auto subsection \Arithmetic operations on lists, pointwise on the elements\ text \Weak type class properties. Cancellation is difficult to arrange because of complications when lists differ in length.\ instantiation list :: (plus) plus begin definition "plus_list \ map2 (+)" instance.. end lemma length_plus_list [simp]: fixes xs :: "'a::plus list" shows "length (xs+ys) = min (length xs) (length ys)" by (simp add: plus_list_def) lemma plus_Nil [simp]: "[] + xs = []" by (simp add: plus_list_def) lemma plus_Cons: "(y # ys) + (x # xs) = (y+x) # (ys+xs)" by (simp add: plus_list_def) lemma nth_plus_list [simp]: fixes xs :: "'a::plus list" assumes "i < length xs" "i < length ys" shows "(xs+ys)!i = xs!i + ys!i" by (simp add: plus_list_def assms) instantiation list :: (minus) minus begin definition "minus_list \ map2 (-)" instance.. end lemma length_minus_list [simp]: fixes xs :: "'a::minus list" shows "length (xs-ys) = min (length xs) (length ys)" by (simp add: minus_list_def) lemma minus_Nil [simp]: "[] - xs = []" by (simp add: minus_list_def) lemma minus_Cons: "(y # ys) - (x # xs) = (y-x) # (ys-xs)" by (simp add: minus_list_def) lemma nth_minus_list [simp]: fixes xs :: "'a::minus list" assumes "i < length xs" "i < length ys" shows "(xs-ys)!i = xs!i - ys!i" by (simp add: minus_list_def assms) instance list :: (ab_semigroup_add) ab_semigroup_add proof have "map2 (+) (map2 (+) xs ys) zs = map2 (+) xs (map2 (+) ys zs)" for xs ys zs :: "'a list" proof (induction xs arbitrary: ys zs) case (Cons x xs) show ?case proof (cases "ys=[] \ zs=[]") case False then obtain y ys' z zs' where "ys = y#ys'" "zs = z # zs'" by (meson list.exhaust) then show ?thesis by (simp add: Cons add.assoc) qed auto qed auto then show "a + b + c = a + (b + c)" for a b c :: "'a list" by (auto simp: plus_list_def) next have "map2 (+) xs ys = map2 (+) ys xs" for xs ys :: "'a list" proof (induction xs arbitrary: ys) case (Cons x xs) show ?case proof (cases ys) case (Cons y ys') then show ?thesis by (simp add: Cons.IH add.commute) qed auto qed auto then show "a + b = b + a" for a b :: "'a list" by (auto simp: plus_list_def) qed subsection \The pointwise ordering on two equal-length lists of natural numbers\ text \Gowers uses the usual symbol ($\le$) for his pointwise ordering. In our development, $\le$ is the lexicographic ordering and $\unlhd$ is the pointwise ordering.\ definition pointwise_le :: "[nat list, nat list] \ bool" (infixr "\" 50 ) where "x \ y \ list_all2 (\) x y" definition pointwise_less :: "[nat list, nat list] \ bool" (infixr "\" 50 ) where "x \ y \ x \ y \ x \ y" lemma pointwise_le_iff_nth: "x \ y \ length x = length y \ (\i < length x. x!i \ y!i)" by (simp add: list_all2_conv_all_nth pointwise_le_def) lemma pointwise_le_iff: "x \ y \ length x = length y \ (\(i,j) \ set (zip x y). i\j)" by (simp add: list_all2_iff pointwise_le_def) lemma pointwise_append_le_iff [simp]: "u@x \ u@y \ x \ y" by (auto simp: pointwise_le_iff_nth nth_append) lemma pointwise_le_refl [iff]: "x \ x" by (simp add: list.rel_refl pointwise_le_def) lemma pointwise_le_antisym: "\x \ y; y \ x\ \ x=y" by (metis antisym list_all2_antisym pointwise_le_def) lemma pointwise_le_trans: "\x \ y; y \ z\ \ x \ z" by (smt (verit, del_insts) le_trans list_all2_trans pointwise_le_def) lemma pointwise_le_Nil [simp]: "Nil \ x \ x = Nil" by (simp add: pointwise_le_def) lemma pointwise_le_Nil2 [simp]: "x \ Nil \ x = Nil" by (simp add: pointwise_le_def) lemma pointwise_le_iff_less_equal: "x \ y \ x \ y \ x = y" using pointwise_less_def by blast lemma pointwise_less_iff: "x \ y \ x \ y \ (\(i,j) \ set (zip x y). i y \ x \ y \ (\k < length x. x!k Nil \ x" by (simp add: pointwise_less_def) lemma pointwise_less_Nil2 [simp]: "\ x \ Nil" by (simp add: pointwise_less_def) lemma zero_pointwise_le_iff [simp]: "replicate r 0 \ x \ length x = r" by (auto simp add: pointwise_le_iff_nth) lemma pointwise_le_imp_\: assumes "xs \ ys" shows "\ xs \ \ ys" using assms proof (induction ys arbitrary: xs) case Nil then show ?case by (simp add: pointwise_le_iff) next case (Cons y ys) then obtain x xs' where "x\y" "xs = x#xs'" "xs' \ ys" by (auto simp: pointwise_le_def list_all2_Cons2) then show ?case by (simp add: Cons.IH add_le_mono) qed lemma sum_list_plus: fixes xs :: "'a::comm_monoid_add list" assumes "length xs = length ys" shows "\ (xs + ys) = \ xs + \ ys" using assms by (simp add: plus_list_def case_prod_unfold sum_list_addf) lemma sum_list_minus: assumes "xs \ ys" shows "\ (ys - xs) = \ ys - \ xs" using assms proof (induction ys arbitrary: xs) case (Cons y ys) then obtain x xs' where "x\y" "xs = x#xs'" "xs' \ ys" by (auto simp: pointwise_le_def list_all2_Cons2) then show ?case using pointwise_le_imp_\ by (auto simp: Cons minus_Cons) qed (auto simp: in_set_conv_nth) subsection \Pointwise minimum and maximum of a set of lists\ definition min_pointwise :: "[nat, nat list set] \ nat list" where "min_pointwise \ \r U. map (\i. Min ((\u. u!i) ` U)) [0..u \ U; finite U\ \ min_pointwise (length u) U \ u" by (simp add: min_pointwise_def pointwise_le_iff_nth) lemma min_pointwise_ge_iff: assumes "finite U" "U \ {}" "\u. u \ U \ length u = r" "length x = r" shows "x \ min_pointwise r U \ (\u \ U. x \ u)" by (auto simp: min_pointwise_def pointwise_le_iff_nth assms) definition max_pointwise :: "[nat, nat list set] \ nat list" where "max_pointwise \ \r U. map (\i. Max ((\u. u!i) ` U)) [0..u \ U; finite U\ \ u \ max_pointwise (length u) U" by (simp add: max_pointwise_def pointwise_le_iff_nth) lemma max_pointwise_le_iff: assumes "finite U" "U \ {}" "\u. u \ U \ length u = r" "length x = r" shows "max_pointwise r U \ x \ (\u \ U. u \ x)" by (auto simp: max_pointwise_def pointwise_le_iff_nth assms) lemma max_pointwise_mono: assumes "X' \ X" "finite X" "X' \ {}" shows "max_pointwise r X' \ max_pointwise r X" - using assms by (simp add: max_pointwise_def pointwise_le_iff_nth Max_mono image_mono) + using assms by (auto simp add: max_pointwise_def pointwise_le_iff_nth Max_mono image_mono) lemma pointwise_le_plus: "\xs \ ys; length ys \ length zs\ \ xs \ ys+zs" proof (induction xs arbitrary: ys zs) case (Cons x xs) then obtain y ys' z zs' where "ys = y#ys'" "zs = z#zs'" unfolding pointwise_le_iff by (metis Suc_le_length_iff le_refl length_Cons) with Cons show ?case by (auto simp: plus_list_def pointwise_le_def) qed (simp add: pointwise_le_iff) lemma pairwise_minus_cancel: "\z \ x; z \ y; x - z = y - z\ \ x = y" unfolding pointwise_le_iff_nth by (metis eq_diff_iff nth_equalityI nth_minus_list) subsection \A locale to fix the finite subset @{term "A \ G"}\ locale Khovanskii = additive_abelian_group + fixes A :: "'a set" assumes AsubG: "A \ G" and finA: "finite A" and nonempty: "A \ {}" begin text \finite products of a group element\ definition Gmult :: "'a \ nat \ 'a" where "Gmult a n \ (((\)a) ^^ n) \" lemma Gmult_0 [simp]: "Gmult a 0 = \" by (simp add: Gmult_def) lemma Gmult_1 [simp]: "a \ G \ Gmult a (Suc 0) = a" by (simp add: Gmult_def) lemma Gmult_Suc [simp]: "Gmult a (Suc n) = a \ Gmult a n" by (simp add: Gmult_def) lemma Gmult_in_G [simp,intro]: "a \ G \ Gmult a n \ G" by (induction n) auto lemma Gmult_add_add: assumes "a \ G" shows "Gmult a (m+n) = Gmult a m \ Gmult a n" by (induction m) (use assms local.associative in fastforce)+ lemma Gmult_add_diff: assumes "a \ G" shows "Gmult a (n+k) \ Gmult a n = Gmult a k" by (metis Gmult_add_add Gmult_in_G assms commutative inverse_closed invertible invertible_left_inverse2) lemma Gmult_diff: assumes "a \ G" "n\m" shows "Gmult a m \ Gmult a n = Gmult a (m-n)" by (metis Gmult_add_diff assms le_add_diff_inverse) text \Mapping elements of @{term A} to their numeric subscript\ abbreviation "idx \ to_nat_on A" text \The elements of @{term A} in order\ definition aA :: "'a list" where "aA \ map (from_nat_into A) [0.. :: "nat list \ 'a" where "\ \ \x. fincomp (\i. Gmult (aA!i) (x!i)) {..The underlying assumption is @{term "length y = length x"}\ definition useless:: "nat list \ bool" where "useless x \ \y < x. \ y = \ x \ \ y = \ x \ length y = length x" abbreviation "useful x \ \ useless x" lemma alpha_replicate_0 [simp]: "\ (replicate (card A) 0) = \" by (auto simp: \_def intro: fincomp_unit_eqI) lemma idx_less_cardA: assumes "a \ A" shows "idx a < card A" by (metis assms bij_betw_def finA imageI lessThan_iff to_nat_on_finite) lemma aA_idx_eq [simp]: assumes "a \ A" shows "aA ! (idx a) = a" by (simp add: aA_def assms countable_finite finA idx_less_cardA) lemma set_aA: "set aA = A" using bij_betw_from_nat_into_finite [OF finA] by (simp add: aA_def atLeast0LessThan bij_betw_def) lemma nth_aA_in_G [simp]: "i < card A \ aA!i \ G" using AsubG aA_def set_aA by auto lemma alpha_in_G [iff]: "\ x \ G" using nth_aA_in_G fincomp_closed by (simp add: \_def) lemma Gmult_in_PiG [simp]: "(\i. Gmult (aA!i) (f i)) \ {.. G" by simp lemma alpha_plus: assumes "length x = card A" "length y = card A" shows "\ (x + y) = \ x \ \ y" proof - have "\ (x + y) = fincomp (\i. Gmult (aA!i) (map2 (+) x y!i)) {.._def plus_list_def) also have "\ = fincomp (\i. Gmult (aA!i) (x!i + y!i)) {.. = fincomp (\i. Gmult (aA!i) (x!i) \ Gmult (aA!i) (y!i)) {.. = \ x \ \ y" by (simp add: \_def fincomp_comp) finally show ?thesis . qed lemma alpha_minus: assumes "y \ x" "length y = card A" shows "\ (x - y) = \ x \ \ y" proof - have "\ (x - y) = fincomp (\i. Gmult (aA!i) (map2 (-) x y!i)) {.._def minus_list_def) also have "\ = fincomp (\i. Gmult (aA!i) (x!i - y!i)) {.. = fincomp (\i. Gmult (aA!i) (x!i) \ Gmult (aA!i) (y!i)) {.. = \ x \ \ y" by (simp add: \_def fincomp_comp fincomp_inverse) finally show ?thesis . qed subsection \Adding one to a list element\ definition list_incr :: "nat \ nat list \ nat list" where "list_incr i x \ x[i := Suc (x!i)]" lemma list_incr_Nil [simp]: "list_incr i [] = []" by (simp add: list_incr_def) lemma list_incr_Cons [simp]: "list_incr (Suc i) (k#ks) = k # list_incr i ks" by (simp add: list_incr_def) lemma sum_list_incr [simp]: "i < length x \ \ (list_incr i x) = Suc (\ x)" by (auto simp: list_incr_def sum_list_update) lemma length_list_incr [simp]: "length (list_incr i x) = length x" by (auto simp: list_incr_def) lemma nth_le_list_incr: "i < card A \ x!i \ list_incr (idx a) x!i" unfolding list_incr_def by (metis Suc_leD linorder_not_less list_update_beyond nth_list_update_eq nth_list_update_neq order_refl) lemma list_incr_nth_diff: "i < length x \ list_incr j x!i - x!i = (if i = j then 1 else 0)" by (simp add: list_incr_def) subsection \The set of all @{term r}-tuples that sum to @{term n}\ definition length_sum_set :: "nat \ nat \ nat list set" where "length_sum_set r n \ {x. length x = r \ \ x = n}" lemma length_sum_set_Nil [simp]: "length_sum_set 0 n = (if n=0 then {[]} else {})" by (auto simp: length_sum_set_def) lemma length_sum_set_Suc [simp]: "k#ks \ length_sum_set (Suc r) n \ (\m. ks \ length_sum_set r m \ n = m+k)" by (auto simp: length_sum_set_def) lemma length_sum_set_Suc_eqpoll: "length_sum_set (Suc r) n \ Sigma {..n} (\i. length_sum_set r (n-i))" (is "?L \ ?R") unfolding eqpoll_def proof let ?f = "(\l. (hd l, tl l))" show "bij_betw ?f ?L ?R" proof (intro bij_betw_imageI) show "inj_on ?f ?L" by (force simp: inj_on_def length_sum_set_def intro: list.expand) show "?f ` ?L = ?R" by (force simp: length_sum_set_def length_Suc_conv) qed qed lemma finite_length_sum_set: "finite (length_sum_set r n)" proof (induction r arbitrary: n) case 0 then show ?case by (auto simp: length_sum_set_def) next case (Suc r) then show ?case using length_sum_set_Suc_eqpoll eqpoll_finite_iff by blast qed lemma card_length_sum_set: "card (length_sum_set (Suc r) n) = (\i\n. card (length_sum_set r (n-i)))" proof - have "card (length_sum_set (Suc r) n) = card (Sigma {..n} (\i. length_sum_set r (n-i)))" by (metis eqpoll_finite_iff eqpoll_iff_card finite_length_sum_set length_sum_set_Suc_eqpoll) also have "\ = (\i\n. card (length_sum_set r (n-i)))" by (simp add: finite_length_sum_set) finally show ?thesis . qed lemma sum_up_index_split': assumes "N \ n" shows "(\i\n. f i) = (\i\n-N. f i) + (\i=Suc (n-N)..n. f i)" by (metis assms diff_add sum_up_index_split) lemma sum_invert: "N \ n \ (\i = Suc (n - N)..n. f (n - i)) = (\j n" shows "(\i\n - N. real (n - i) ^ j) = (\i\n. real i ^ j) - (\ii\n - N. real (n - i) ^ j) = (\i\(-) n ` {N..n}. real (n - i) ^ j)" proof (rule sum.cong) have "\x. x \ n - N \ \m\N. m \ n \ x = n - m" by (metis assms diff_diff_cancel diff_le_mono2 diff_le_self le_trans) then show "{..n - N} = (-) n ` {N..n}" by (auto simp add: image_iff Bex_def) qed auto also have "\ = (\i=N..n. real i ^ j)" using sum.reindex [OF inj, of "\i. real (n - i) ^ j", symmetric] by (simp add: ) also have "\ = (\i\n. real i ^ j) - (\ip. real_polynomial_function p \ (\n>0. real (card (length_sum_set r n)) = p (real n))" proof (induction r) case 0 have "\n>0. real (card (length_sum_set 0 n)) = 0" by auto then show ?case by blast next case (Suc r) then obtain p where poly: "real_polynomial_function p" and p: "\n. n>0 \ real (card (length_sum_set r n)) = p (real n)" by blast then obtain a n where p_eq: "p = (\x. \i\n. a i * x ^ i)" using real_polynomial_function_iff_sum by auto define q where "q \ \x. \j\n. a j * ((bernpoly (Suc j) (1 + x) - bernpoly (Suc j) 0) / (1 + real j) - 0 ^ j)" have rp_q: "real_polynomial_function q" by (fastforce simp: bernpoly_def p_eq q_def) have q_eq: "(\x\n - 1. p (real (n - x))) = q (real n)" if "n>0" for n using that by (simp add: p_eq q_def sum.swap sum_diff_split add.commute sum_of_powers flip: sum_distrib_left) define p' where "p' \ \x. q x + real (card (length_sum_set r 0))" have "real_polynomial_function p'" using rp_q by (force simp add: p'_def) moreover have "(\x\n - Suc 0. p (real (n - x))) + real (card (length_sum_set r 0)) = p' (real n)" if "n>0" for n using that q_eq by (auto simp: p'_def) ultimately show ?case unfolding card_length_sum_set by (force simp: sum_up_index_split' [of 1] p sum_invert) qed lemma all_zeroes_replicate: "length_sum_set r 0 = {replicate r 0}" by (auto simp: length_sum_set_def replicate_eqI) lemma length_sum_set_Suc_eq_UN: "length_sum_set r (Suc n) = (\ii list_incr i ` length_sum_set r n" if "\ x = Suc n" and "r = length x" for x proof - have "x \ replicate r 0" using that by (metis sum_list_replicate Zero_not_Suc mult_zero_right) then obtain i where i: "i < r" "x!i \ 0" by (metis \r = length x\ in_set_conv_nth replicate_eqI) with that have "x[i := x!i - 1] \ length_sum_set r n" by (simp add: sum_list_update length_sum_set_def) with i that show ?thesis unfolding list_incr_def by force qed then show ?thesis by (auto simp: length_sum_set_def Bex_def) qed lemma alpha_list_incr: assumes "a \ A" "x \ length_sum_set (card A) n" shows "\ (list_incr (idx a) x) = a \ \ x" proof - have lenx: "length x = card A" using assms length_sum_set_def by blast have "\ (list_incr (idx a) x) \ \ x = fincomp (\i. Gmult (aA!i) (list_incr (idx a) x!i) \ Gmult (aA!i) (x!i)) {.._def fincomp_comp fincomp_inverse) also have "\ = fincomp (\i. Gmult (aA!i) (list_incr (idx a) x!i - x!i)) {.. = fincomp (\i. if i = idx a then (aA!i) else \) {.. = a" using assms by (simp add: fincomp_singleton_swap idx_less_cardA) finally have "\ (list_incr (idx a) x) \ \ x = a" . then show ?thesis by (metis alpha_in_G associative inverse_closed invertible invertible_left_inverse right_unit) qed lemma sumset_iterated_enum: defines "r \ card A" shows "sumset_iterated A n = \ ` length_sum_set r n" proof (induction n) case 0 then show ?case by (simp add: all_zeroes_replicate r_def) next case (Suc n) have eq: "{..a\A. (\i. a \ \ i) ` length_sum_set r n)" using AsubG by (auto simp: Suc sumset) also have "\ = (\a\A. (\i. \ (list_incr (idx a) i)) ` length_sum_set r n)" by (simp add: alpha_list_incr r_def) also have "\ = \ ` length_sum_set r (Suc n)" by (simp add: image_UN image_comp length_sum_set_Suc_eq_UN eq) finally show ?case . qed subsection \Lemma 2.7 in Gowers's notes\ text\The following lemma corresponds to a key fact about the cardinality of the set of all sums of $n$ many elements of $A$, stated before Gowers's Lemma 2.7.\ lemma card_sumset_iterated_length_sum_set_useful: defines "r \ card A" shows "card(sumset_iterated A n) = card (length_sum_set r n \ {x. useful x})" (is "card ?L = card ?R") proof - have "\ x \ \ ` (length_sum_set r n \ {x. useful x})" if "x \ length_sum_set r n" for x proof - define y where "y \ LEAST y. y \ length_sum_set r n \ \ y = \ x" have y: "y \ length_sum_set (card A) n \ \ y = \ x" by (metis (mono_tags, lifting) LeastI r_def y_def that) moreover have "useful y" proof (clarsimp simp: useless_def) show False if "\ z = \ y" "length z = length y" and "z < y" "\ z = \ y" for z using that Least_le length_sum_set_def not_less_Least r_def y y_def by fastforce qed ultimately show ?thesis unfolding image_iff length_sum_set_def r_def by (smt (verit) Int_Collect) qed then have "sumset_iterated A n = \ ` (length_sum_set r n \ {x. useful x})" by (auto simp: sumset_iterated_enum length_sum_set_def r_def) moreover have "inj_on \ (length_sum_set r n \ {x. useful x})" apply (simp add: image_iff length_sum_set_def r_def inj_on_def useless_def Ball_def) by (metis linorder_less_linear) ultimately show ?thesis by (simp add: card_image length_sum_set_def) qed text\The following lemma corresponds to Lemma 2.7 in Gowers's notes.\ lemma useless_leq_useless: defines "r \ card A" assumes "useless x" and "x \ y" and "length x = r" shows "useless y " proof - have leny: "length y = r" using pointwise_le_iff assms by auto obtain x' where "x'< x" and \x': "\ x' = \ x" and \x': "\ x' = \ x" and lenx': "length x' = length x" using assms useless_def by blast obtain i where "i < card A" and xi: "x'!i < x!i" and takex': "take i x' = take i x" using \x' lenx' assms by (auto simp: list_less_def lenlex_def elim!: lex_take_index) define y' where "y' \ y+x'-x" have leny': "length y' = length y" using assms lenx' pointwise_le_iff by (simp add: y'_def) have "x!i \ y!i" using \x \ y\ \i < card A\ assms by (simp add: pointwise_le_iff_nth) then have "y'!i < y!i" using \i < card A\ assms lenx' xi pointwise_le_iff by (simp add: y'_def plus_list_def minus_list_def) moreover have "take i y' = take i y" proof (intro nth_equalityI) show "length (take i y') = length (take i y)" by (simp add: leny') show "\k. k < length (take i y') \ take i y' ! k = take i y!k" using takex' by (simp add: y'_def plus_list_def minus_list_def take_map take_zip) qed ultimately have "y' < y" using leny' \i < card A\ assms pointwise_le_iff by (auto simp: list_less_def lenlex_def lexord_lex lexord_take_index_conv) moreover have "\ y' = \ y" using assms by (simp add: \x' lenx' leny pointwise_le_plus sum_list_minus sum_list_plus y'_def) moreover have "\ y' = \ y" using assms lenx' \x' leny by (fastforce simp add: y'_def pointwise_le_plus alpha_minus alpha_plus local.associative) ultimately show ?thesis using leny' useless_def by blast qed inductive_set minimal_elements for U where "\x \ U; \y. y \ U \ \ y \ x\ \ x \ minimal_elements U" lemma pointwise_less_imp_\: assumes "xs \ ys" shows "\ xs < \ ys" proof - have eq: "length ys = length xs" and "xs \ ys" using assms by (auto simp: pointwise_le_iff pointwise_less_iff) have "\k ys!k" using \xs \ ys\ list_all2_nthD pointwise_le_def by auto moreover have "\k: "wf (inv_image less_than \)" by blast lemma WFP: "wfP (\)" by (auto simp: wfP_def pointwise_less_imp_\ intro: wf_subset [OF wf_measure_\]) text\The following is a direct corollary of the above lemma, i.e. a corollary of Lemma 2.7 in Gowers's notes.\ corollary useless_iff: assumes "length x = card A" shows "useless x \ (\x' \ minimal_elements (Collect useless). x' \ x)" (is "_=?R") proof assume "useless x" obtain z where z: "useless z" "z \ x" and zmin: "\y. y \ z \ y \ x \ useful y" using wfE_min [to_pred, where Q = "{z. useless z \ z \ x}", OF WFP] by (metis (no_types, lifting) \useless x\ mem_Collect_eq pointwise_le_refl) then show ?R by (smt (verit) mem_Collect_eq minimal_elements.intros pointwise_le_trans pointwise_less_def) next assume ?R with useless_leq_useless minimal_elements.cases show "useless x" by (metis assms mem_Collect_eq pointwise_le_iff) qed subsection \The set of minimal elements of a set of $r$-tuples is finite\ text\The following general finiteness claim corresponds to Lemma 2.8 in Gowers's notes and is key to the main proof.\ lemma minimal_elements_set_tuples_finite: assumes Ur: "\x. x \ U \ length x = r" shows "finite (minimal_elements U)" using assms proof (induction r arbitrary: U) case 0 then have "U \ {[]}" by auto then show ?case by (metis finite.simps minimal_elements.cases finite_subset subset_eq) next case (Suc r) show ?case proof (cases "U={}") case True with Suc.IH show ?thesis by blast next case False then obtain u where u: "u \ U" and zmin: "\y. y \ u \ y \ U" using wfE_min [to_pred, where Q = "U", OF WFP] by blast define V where "V = {v \ U. \ u \ v}" define VF where "VF \ \i t. {v \ V. v!i = t}" have [simp]: "length v = Suc r" if "v \ VF i t" for v i t using that by (simp add: Suc.prems VF_def V_def) have *: "\i\r. v!i < u!i" if "v \ V" for v using that u Suc.prems by (force simp add: V_def pointwise_le_iff_nth not_le less_Suc_eq_le) with u have "minimal_elements U \ insert u (\i\r. \t < u!i. minimal_elements (VF i t))" by (force simp: VF_def V_def minimal_elements.simps pointwise_less_def) moreover have "finite (minimal_elements (VF i t))" if "i\r" "t < u!i" for i t proof - define delete where "delete \ \v::nat list. take i v @ drop (Suc i) v" \ \deletion of @{term i}\ have len_delete[simp]: "length (delete u) = r" if "u \ VF i t" for u using Suc.prems VF_def V_def \i \ r\ delete_def that by auto have nth_delete: "delete u!k = (if k VF i t" "k delete v \ u \ v" if "u \ VF i t" "v \ VF i t" for u v proof assume "delete u \ delete v" then have "\j. (j < i \ u!j \ v!j) \ (j < r \ i \ j \ u!Suc j \ v!Suc j)" using that \i \ r\ by (force simp add: pointwise_le_iff_nth nth_delete split: if_split_asm cong: conj_cong) then show "u \ v" using that \i \ r\ apply (simp add: pointwise_le_iff_nth VF_def) by (metis eq_iff le_Suc_eq less_Suc_eq_0_disj linorder_not_less) next assume "u \ v" then show "delete u \ delete v" using that by (simp add: pointwise_le_iff_nth nth_delete) qed then have delete_eq_iff: "delete u = delete v \ u = v" if "u \ VF i t" "v \ VF i t" for u v by (metis that pointwise_le_antisym pointwise_le_refl) have delete_less_iff: "delete u \ delete v \ u \ v" if "u \ VF i t" "v \ VF i t" for u v by (metis delete_le_iff pointwise_le_antisym pointwise_less_def that) have "length (delete v) = r" if "v \ V" for v using id_take_nth_drop Suc.prems V_def \i \ r\ delete_def that by auto then have "finite (minimal_elements (delete ` V))" by (metis (mono_tags, lifting) Suc.IH image_iff) moreover have "inj_on delete (minimal_elements (VF i t))" by (simp add: delete_eq_iff inj_on_def minimal_elements.simps) moreover have "delete ` (minimal_elements (VF i t)) \ minimal_elements (delete ` (VF i t))" by (auto simp: delete_less_iff minimal_elements.simps) ultimately show ?thesis by (metis (mono_tags, lifting) Suc.IH image_iff inj_on_finite len_delete) qed ultimately show ?thesis by (force elim: finite_subset) qed qed subsection \Towards Lemma 2.9 in Gowers's notes\ text \Increasing sequences\ fun augmentum :: "nat list \ nat list" where "augmentum [] = []" | "augmentum (n#ns) = n # map ((+)n) (augmentum ns)" definition dementum:: "nat list \ nat list" where "dementum xs \ xs - (0#xs)" lemma dementum_Nil [simp]: "dementum [] = []" by (simp add: dementum_def) lemma zero_notin_augmentum [simp]: "0 \ set ns \ 0 \ set (augmentum ns)" by (induction ns) auto lemma length_augmentum [simp]:"length (augmentum xs) = length xs" by (induction xs) auto lemma sorted_augmentum [simp]: "0 \ set ns \ sorted (augmentum ns)" by (induction ns) auto lemma distinct_augmentum [simp]: "0 \ set ns \ distinct (augmentum ns)" by (induction ns) (simp_all add: image_iff) lemma augmentum_subset_sum_list: "set (augmentum ns) \ {..\ ns}" by (induction ns) auto lemma sum_list_augmentum: "\ ns \ set (augmentum ns) \ length ns > 0" by (induction ns) auto lemma length_dementum [simp]: "length (dementum xs) = length xs" by (simp add: dementum_def) lemma sorted_imp_pointwise: assumes "sorted (xs@[n])" shows "0 # xs \ xs @ [n]" using assms by (simp add: pointwise_le_iff_nth nth_Cons' nth_append sorted_append sorted_wrt_append sorted_wrt_nth_less) lemma sum_list_dementum: assumes "sorted (xs@[n])" shows "\ (dementum (xs@[n])) = n" proof - have "dementum (xs@[n]) = (xs@[n]) - (0 # xs)" by (rule nth_equalityI; simp add: nth_append dementum_def nth_Cons') then show ?thesis by (simp add: sum_list_minus sorted_imp_pointwise assms) qed lemma augmentum_cancel: "map ((+)k) (augmentum ns) - (k # map ((+)k) (augmentum ns)) = ns" proof (induction ns arbitrary: k) case Nil then show ?case by simp next case (Cons n ns) have "(+) k \ (+) n = (+) (k+n)" by auto then show ?case by (simp add: minus_Cons Cons) qed lemma dementum_augmentum [simp]: assumes "0 \ set ns" shows "(dementum \ sorted_list_of_set) ((set \ augmentum) ns) = ns" (is "?L ns = _") using assms augmentum_cancel [of 0] by (simp add: dementum_def map_idI sorted_list_of_set.idem_if_sorted_distinct) lemma dementum_nonzero: assumes ns: "sorted_wrt (<) ns" and 0: "0 \ set ns" shows "0 \ set (dementum ns)" unfolding dementum_def minus_list_def using sorted_wrt_nth_less [OF ns] 0 by (auto simp add: in_set_conv_nth image_iff set_zip nth_Cons' dest: leD) lemma nth_augmentum [simp]: "i < length ns \ augmentum ns!i = (\j\i. ns!j)" proof (induction ns arbitrary: i) case Nil then show ?case by simp next case (Cons a ns) show ?case proof (cases "i=0") case False then have "augmentum (a # ns)!i = a + sum ((!) ns) {..i-1}" using Cons.IH Cons.prems by auto also have "\ = a + (\j\{0<..i}. ns!(j-1))" using sum.reindex [of Suc "{..i - Suc 0}" "\j. ns!(j-1)", symmetric] False by (simp add: image_Suc_atMost atLeastSucAtMost_greaterThanAtMost del: sum.cl_ivl_Suc) also have "\ = (\j = 0..i. if j=0 then a else ns!(j-1))" by (simp add: sum.head) also have "\ = sum ((!) (a # ns)) {..i}" by (simp add: nth_Cons' atMost_atLeast0) finally show ?thesis . qed auto qed lemma augmentum_dementum [simp]: assumes "0 \ set ns" "sorted ns" shows "augmentum (dementum ns) = ns" proof (rule nth_equalityI) fix i assume "i < length (augmentum (dementum ns))" then have i: "i < length ns" by simp show "augmentum (dementum ns)!i = ns!i" proof (cases "i=0") case True then show ?thesis using nth_augmentum dementum_def i by auto next case False have ns_le: "\j. \0 < j; j \ i\ \ ns ! (j - Suc 0) \ ns ! j" using \sorted ns\ i by (simp add: sorted_iff_nth_mono) have "augmentum (dementum ns)!i = (\j\i. ns!j - (if j = 0 then 0 else ns!(j-1)))" using i by (simp add: dementum_def nth_Cons') also have "\ = (\j=0..i. if j = 0 then ns!0 else ns!j - ns!(j-1))" by (smt (verit, best) diff_zero sum.cong atMost_atLeast0) also have "\ = ns!0 + (\j\{0<..i}. ns!j - ns!(j-1))" by (simp add: sum.head) also have "\ = ns!0 + ((\j\{0<..i}. ns!j) - (\j\{0<..i}. ns!(j-1)))" by (auto simp: ns_le intro: sum_subtractf_nat) also have "\ = ns!0 + (\j\{0<..i}. ns!j) - (\j\{0<..i}. ns!(j-1))" proof - have "(\j\{0<..i}. ns ! (j - 1)) \ sum ((!) ns) {0<..i}" by (metis One_nat_def greaterThanAtMost_iff ns_le sum_mono) then show ?thesis by simp qed also have "\ = ns!0 + (\j\{0<..i}. ns!j) - (\j\i-Suc 0. ns!j)" using sum.reindex [of Suc "{..i - Suc 0}" "\j. ns!(j-1)", symmetric] False by (simp add: image_Suc_atMost atLeastSucAtMost_greaterThanAtMost) also have "\ = (\j=0..i. ns!j) - (\j\i-Suc 0. ns!j)" by (simp add: sum.head [of 0 i]) also have "\ = (\j=0..i-Suc 0. ns!j) + ns!i - (\j\i-Suc 0. ns!j)" by (metis False Suc_pred less_Suc0 not_less_eq sum.atLeast0_atMost_Suc) also have "\ = ns!i" by (simp add: atLeast0AtMost) finally show "augmentum (dementum ns)!i = ns!i" . qed qed auto text\The following lemma corresponds to Lemma 2.9 in Gowers's notes. The proof involves introducing bijective maps between r-tuples that fulfill certain properties/r-tuples and subsets of naturals, so as to show the cardinality claim. \ lemma bound_sum_list_card: assumes "r > 0" and n: "n \ \ x'" and lenx': "length x' = r" defines "S \ {x. x' \ x \ \ x = n}" shows "card S = (n - \ x' + r - 1) choose (r-1)" proof- define m where "m \ n - \ x'" define f where "f \ \x::nat list. x - x'" have f: "bij_betw f S (length_sum_set r m)" proof (intro bij_betw_imageI) show "inj_on f S" using pairwise_minus_cancel by (force simp: S_def f_def inj_on_def) have "\x. x \ S \ f x \ length_sum_set r m" by (simp add: S_def f_def length_sum_set_def lenx' m_def pointwise_le_iff sum_list_minus) moreover have "x \ f ` S" if "x \ length_sum_set r m" for x proof have x[simp]: "length x = r" "\ x = m" using that by (auto simp: length_sum_set_def) have "x = x' + x - x'" by (rule nth_equalityI; simp add: lenx') then show "x = f (x' + x)" unfolding f_def by fastforce have "x' \ x' + x" by (simp add: lenx' pointwise_le_plus) moreover have "\ (x' + x) = n" by (simp add: lenx' m_def n sum_list_plus) ultimately show "x' + x \ S" using S_def by blast qed ultimately show "f ` S = length_sum_set r m" by auto qed define g where "g \ \x::nat list. map Suc x" define g' where "g' \ \x::nat list. x - replicate (length x) 1" define T where "T \ length_sum_set r (m+r) \ lists(-{0})" have g: "bij_betw g (length_sum_set r m) T" proof (intro bij_betw_imageI) show "inj_on g (length_sum_set r m)" by (auto simp: g_def inj_on_def) have "\x. x \ length_sum_set r m \ g x \ T" by (auto simp: g_def length_sum_set_def sum_list_Suc T_def) moreover have "x \ g ` length_sum_set r m" if "x \ T" for x proof have [simp]: "length x = r" using length_sum_set_def that T_def by auto have r1_x: "replicate r (Suc 0) \ x" using that unfolding T_def pointwise_le_iff_nth by (simp add: lists_def in_listsp_conv_set Suc_leI) show "x = g (g' x)" using that by (intro nth_equalityI) (auto simp: g_def g'_def T_def) show "g' x \ length_sum_set r m" using that T_def by (simp add: g'_def r1_x sum_list_minus length_sum_set_def sum_list_replicate) qed ultimately show "g ` (length_sum_set r m) = T" by auto qed define U where "U \ (insert (m+r)) ` finsets {0<.. augmentum) T U" proof (intro bij_betw_imageI) show "inj_on ((set \ augmentum)) T" unfolding inj_on_def T_def by (metis ComplD IntE dementum_augmentum in_listsD insertI1) have "(set \ augmentum) t \ U" if "t \ T" for t proof - have t: "length t = r" "\ t = m+r" "0 \ set t" using that by (force simp add: T_def length_sum_set_def)+ then have mrt: "m + r \ set (augmentum t)" by (metis \r>0\ sum_list_augmentum) then have "set (augmentum t) = insert (m + r) (set (augmentum t) - {m + r})" by blast moreover have "set (augmentum t) - {m + r} \ finsets {0<.. (set \ augmentum) ` T" if "u \ U" for u proof from that obtain N where u: "u = insert (m + r) N" and Nsub: "N \ {0<.. N" "m+r \ N" "finite N" using finite_subset Nsub by auto have [simp]: "card u = r" using Nsub \r>0\ by (auto simp add: u card_insert_if) have ssN: "sorted (sorted_list_of_set N @ [m + r])" using Nsub by (simp add: less_imp_le_nat sorted_wrt_append subset_eq) have so_u_N: "sorted_list_of_set u = insort (m+r) (sorted_list_of_set N)" by (simp add: u) also have "\ = sorted_list_of_set N @ [m+r]" using Nsub by (force intro: sorted_insort_is_snoc) finally have so_u: "sorted_list_of_set u = sorted_list_of_set N @ [m+r]" . have 0: "0 \ set (sorted_list_of_set u)" by (simp add: \r>0\ set_insort_key so_u_N) show "u = (set \ augmentum) ((dementum \ sorted_list_of_set)u)" using 0 so_u ssN u by force have sortd_wrt_u: "sorted_wrt (<) (sorted_list_of_set u)" by simp show "(dementum \ sorted_list_of_set) u \ T" apply (simp add: T_def length_sum_set_def) using sum_list_dementum [OF ssN] sortd_wrt_u 0 by (force simp: so_u dementum_nonzero)+ qed ultimately show "(set \ augmentum) ` T = U" by auto qed obtain \ where "bij_betw \ S U" by (meson bij_betw_trans f g h) moreover have "card U = (n - \ x' + r-1) choose (r-1)" proof - have "inj_on (insert (m + r)) (finsets {0<.. = (n - \ x' + r-1) choose (r-1)" by (simp add: card_finsets m_def) finally show ?thesis . qed ultimately show ?thesis by (metis bij_betw_same_card) qed subsection \Towards the main theorem\ lemma extend_tuple: assumes "\ xs \ n" "length xs \ 0" obtains ys where "\ ys = n" "xs \ ys" proof - obtain x xs' where xs: "xs = x#xs'" using assms list.exhaust by auto define y where "y \ x + n - \ xs" show thesis proof show "\ (y#xs') = n" using assms xs y_def by auto show "xs \ y#xs'" using assms y_def pointwise_le_def xs by auto qed qed lemma extend_preserving: assumes "\ xs \ n" "length xs > 1" "i < length xs" obtains ys where "\ ys = n" "xs \ ys" "ys!i = xs!i" proof - define j where "j \ Suc i mod length xs" define xs1 where "xs1 = take j xs" define xs2 where "xs2 = drop (Suc j) xs" define x where "x = xs!j" have xs: "xs = xs1 @ [x] @ xs2" using assms apply (simp add: Cons_nth_drop_Suc assms x_def xs1_def xs2_def j_def) by (meson Suc_lessD id_take_nth_drop mod_less_divisor) define y where "y \ x + n - \ xs" define ys where "ys \ xs1 @ [y] @ xs2" have "x \ y" using assms y_def by linarith show thesis proof show "\ ys = n" using assms(1) xs y_def ys_def by auto show "xs \ ys" using xs ys_def \x \ y\ pointwise_append_le_iff pointwise_le_def by fastforce have "length xs1 \ i" using assms by (simp add: xs1_def j_def min_def mod_Suc) then show "ys!i = xs!i" by (auto simp: ys_def xs nth_append nth_Cons') qed qed text\The proof of the main theorem will make use of the inclusion-exclusion formula, in addition to the previously shown results.\ theorem Khovanskii: assumes "card A > 1" defines "f \ \n. card(sumset_iterated A n)" obtains N p where "real_polynomial_function p" "\n. n \ N \ real (f n) = p (real n)" proof - define r where "r \ card A" define C where "C \ \n x'. {x. x' \ x \ \ x = n}" define X where "X \ minimal_elements {x. useless x \ length x = r}" have "r > 1" "r \ 0" using assms r_def by auto have Csub: "C n x' \ length_sum_set (length x') n" for n x' by (auto simp add: C_def length_sum_set_def pointwise_le_iff) then have finC: "finite (C n x')" for n x' by (meson finite_length_sum_set finite_subset) have "finite X" using "minimal_elements_set_tuples_finite" X_def by force then have max_X: "\x'. x' \ X \ \ x' \ \ (max_pointwise r X)" using X_def max_pointwise_ge minimal_elements.simps pointwise_le_imp_\ by force let ?z0 = "replicate r 0" have Cn0: "C n ?z0 = length_sum_set r n" for n by (auto simp: C_def length_sum_set_def) then obtain p0 where pf_p0: "real_polynomial_function p0" and p0: "\n. n>0 \ p0 (real n) = real (card (C n ?z0))" by (metis real_polynomial_function_length_sum_set) obtain q where pf_q: "real_polynomial_function q" and q: "\x. q x = x gchoose (r-1)" using real_polynomial_function_gchoose by metis define p where "p \ \x::real. p0 x - (\Y | Y \ X \ Y \ {}. (- 1) ^ (card Y + 1) * q((x - real(\ (max_pointwise r Y)) + real r - 1)))" show thesis proof note pf_q' = real_polynomial_function_compose [OF _ pf_q, unfolded o_def] note pf_intros = real_polynomial_function_sum real_polynomial_function_diff real_polynomial_function.intros show "real_polynomial_function p" unfolding p_def using \finite X\ by (intro pf_p0 pf_q' pf_intros | force)+ next fix n assume "n \ max 1 (\ (max_pointwise r X))" then have nlarge: "n \ \ (max_pointwise r X)" and "n > 0" by auto define U where "U \ \n. length_sum_set r n \ {x. useful x}" have 2: "(length_sum_set r n \ {x. useless x}) = (\x'\X. C n x')" unfolding C_def X_def length_sum_set_def r_def using useless_leq_useless by (force simp: minimal_elements.simps pointwise_le_iff useless_iff) define SUM1 where "SUM1 \ \I | I \ C n ` X \ I \ {}. (- 1) ^ (card I + 1) * int (card (\I))" define SUM2 where "SUM2 \ \Y | Y \ X \ Y \ {}. (- 1) ^ (card Y + 1) * int (card (\(C n ` Y)))" have SUM1_card: "card(length_sum_set r n \ {x. useless x}) = nat SUM1" unfolding SUM1_def using \finite X\ by (simp add: finC 2 card_UNION) have "SUM1 \ 0" unfolding SUM1_def using card_UNION_nonneg finC \finite X\ by auto have C_empty_iff: "C n x' = {} \ \ x' > n" if "length x' \ 0" for x' by (simp add: set_eq_iff C_def) (meson extend_tuple linorder_not_le pointwise_le_imp_\ that) have C_eq_1: "C n x' = {[n]}" if "\ x' \ n" "length x' = 1" for x' using that by (auto simp: C_def length_Suc_conv pointwise_le_def elim!: list.rel_cases) have n_ge_X: "\ x \ n" if "x \ X" for x by (meson le_trans max_X nlarge that) have len_X_r: "x \ X \ length x = r" for x by (auto simp: X_def minimal_elements.simps) have "min_pointwise r (C n x') = x'" if "r > 1" "x' \ X" for x' proof (rule pointwise_le_antisym) have [simp]: "length x' = r" "\ x' \ n" using X_def minimal_elements.cases that(2) n_ge_X by auto have [simp]: "length (min_pointwise r (C n x')) = r" by (simp add: min_pointwise_def) show "min_pointwise r (C n x') \ x'" proof (clarsimp simp add: pointwise_le_iff_nth) fix i assume "i < r" then obtain y where "\ y = n \ x' \ y \ y!i \ x'!i" by (metis extend_preserving \1 < r\ \length x' = r\ \x' \ X\ order.refl n_ge_X) then have "\y\C n x'. y!i \ x'!i" using C_def by blast with \i < r\ show "min_pointwise r (C n x')!i \ x'!i" by (simp add: min_pointwise_def Min_le_iff finC C_empty_iff leD) qed have "x' \ min_pointwise r (C n x')" if "\ x' \ n" "length x' = r" for x' by (smt (verit, del_insts) C_def C_empty_iff \r \ 0\ finC leD mem_Collect_eq min_pointwise_ge_iff pointwise_le_iff that) then show "x' \ min_pointwise r (C n x')" using X_def minimal_elements.cases that by force qed then have inj_C: "inj_on (C n) X" by (smt (verit, best) inj_onI mem_Collect_eq \r>1\) have inj_on_imageC: "inj_on (image (C n)) (Pow X - {{}})" by (simp add: inj_C inj_on_diff inj_on_image_Pow) have "Pow (C n ` X) - {{}} \ (image (C n)) ` (Pow X - {{}})" by (metis Pow_empty image_Pow_surj image_diff_subset image_empty) then have "(image (C n)) ` (Pow X - {{}}) = Pow (C n ` X) - {{}}" by blast then have "SUM1 = sum (\I. (- 1) ^ (card I + 1) * int (card (\I))) ((image (C n)) ` (Pow X - {{}}))" unfolding SUM1_def by (auto intro: sum.cong) also have "\ = sum ((\I. (- 1) ^ (card I + 1) * int (card (\ I))) \ (image (C n))) (Pow X - {{}})" by (simp add: sum.reindex inj_on_imageC) also have "\ = SUM2" unfolding SUM2_def using subset_inj_on [OF inj_C] by (force simp: card_image intro: sum.cong) finally have "SUM1 = SUM2" . have "length_sum_set r n = (length_sum_set r n \ {x. useful x}) \ (length_sum_set r n \ {x. useless x})" by auto then have "card (length_sum_set r n) = card (length_sum_set r n \ {x. useful x}) + card (length_sum_set r n \ Collect useless)" by (simp add: finite_length_sum_set disjnt_iff flip: card_Un_disjnt) moreover have "C n ?z0 = length_sum_set r n" by (auto simp: C_def length_sum_set_def) ultimately have "card (C n ?z0) = card (U n) + nat SUM2" by (simp add: U_def flip: \SUM1 = SUM2\ SUM1_card) then have SUM2_le: "nat SUM2 \ card (C n ?z0)" by arith have \_max_pointwise_le: "\Y. \Y \ X; Y \ {}\ \ \ (max_pointwise r Y) \ n" by (meson \finite X\ le_trans max_pointwise_mono nlarge pointwise_le_imp_\) have card_C_max: "card (C n (max_pointwise r Y)) = (n - \ (max_pointwise r Y) + r - Suc 0 choose (r - Suc 0))" if "Y \ X" "Y \ {}" for Y proof - have [simp]: "length (max_pointwise r Y) = r" by (simp add: max_pointwise_def) then show ?thesis using \r \ 0\ that C_def by (simp add: bound_sum_list_card [of r] \_max_pointwise_le) qed define SUM3 where "SUM3 \ (\Y | Y \ X \ Y \ {}. - ((- 1) ^ (card Y) * ((n - \ (max_pointwise r Y) + r - 1 choose (r - 1)))))" have "\(C n ` Y) = C n (max_pointwise r Y)" if "Y \ X" "Y \ {}" for Y proof show "\ (C n ` Y) \ C n (max_pointwise r Y)" unfolding C_def proof clarsimp fix x assume "\y\Y. y \ x \ \ x = n" moreover have "finite Y" using \finite X\ infinite_super that by blast moreover have "\u. u \ Y \ length u = r" using len_X_r that by blast ultimately show "max_pointwise r Y \ x \ \ x = n" by (smt (verit, del_insts) all_not_in_conv max_pointwise_le_iff pointwise_le_iff_nth that(2)) qed next show "C n (max_pointwise r Y) \ \ (C n ` Y)" apply (clarsimp simp: C_def) by (metis \finite X\ finite_subset len_X_r max_pointwise_ge pointwise_le_trans subsetD that(1)) qed then have "SUM2 = SUM3" by (simp add: SUM2_def SUM3_def card_C_max) have "U n = C n ?z0 - (length_sum_set r n \ {x. useless x})" by (auto simp: U_def C_def length_sum_set_def) then have "card (U n) = card (C n ?z0) - card(length_sum_set r n \ {x. useless x})" using finite_length_sum_set by (simp add: C_def Collect_mono_iff inf.coboundedI1 length_sum_set_def flip: card_Diff_subset) then have card_U_eq_diff: "card (U n) = card (C n ?z0) - nat SUM1" using SUM1_card by presburger have "SUM3 \ 0" using \0 \ SUM1\ \SUM1 = SUM2\ \SUM2 = SUM3\ by blast have **: "\Y. \Y \ X; Y \ {}\ \ Suc (\ (max_pointwise r Y)) \ n + r" by (metis \1 < r\ \_max_pointwise_le add.commute add_le_mono less_or_eq_imp_le plus_1_eq_Suc) have "real (f n) = card (U n)" unfolding f_def r_def U_def length_sum_set_def using card_sumset_iterated_length_sum_set_useful length_sum_set_def by presburger also have "\ = card (C n ?z0) - nat SUM3" using card_U_eq_diff \SUM1 = SUM2\ \SUM2 = SUM3\ by presburger also have "\ = real (card (C n (replicate r 0))) - real (nat SUM3)" using SUM2_le \SUM2 = SUM3\ of_nat_diff by blast also have "\ = p (real n)" using \1 < r\ \n>0\ apply (simp add: p_def p0 q \SUM3 \ 0\) apply (simp add: SUM3_def binomial_gbinomial of_nat_diff \_max_pointwise_le algebra_simps **) done finally show "real (f n) = p (real n)" . qed qed end end diff --git a/thys/Matroids/Matroid.thy b/thys/Matroids/Matroid.thy --- a/thys/Matroids/Matroid.thy +++ b/thys/Matroids/Matroid.thy @@ -1,1338 +1,1317 @@ (* File: Matroid.thy Author: Jonas Keinholz *) section \Matroids\ theory Matroid imports Indep_System begin -lemma card_subset_ex: - assumes "finite A" "n \ card A" - shows "\B \ A. card B = n" -using assms -proof (induction A arbitrary: n rule: finite_induct) - case (insert x A) - show ?case - proof (cases n) - case 0 - then show ?thesis using card.empty by blast - next - case (Suc k) - then have "\B \ A. card B = k" using insert by auto - then obtain B where "B \ A" "card B = k" by auto - moreover from this have "finite B" using insert.hyps finite_subset by auto - ultimately have "card (insert x B) = n" - using Suc insert.hyps card_insert_disjoint by fastforce - then show ?thesis using \B \ A\ by blast - qed -qed auto - locale matroid = indep_system + assumes augment_aux: "indep X \ indep Y \ card X = Suc (card Y) \ \x \ X - Y. indep (insert x Y)" begin lemma augment: assumes "indep X" "indep Y" "card Y < card X" shows "\x \ X - Y. indep (insert x Y)" proof - obtain X' where "X' \ X" "card X' = Suc (card Y)" - using assms card_subset_ex[of X "Suc (card Y)"] indep_finite by auto + using assms obtain_subset_with_card_n[of "Suc (card Y)" X] indep_finite by (metis Suc_leI) then obtain x where "x \ X' - Y" "indep (insert x Y)" using assms augment_aux[of X' Y] indep_subset by auto then show ?thesis using \X' \ X\ by auto qed lemma augment_psubset: assumes "indep X" "indep Y" "Y \ X" shows "\x \ X - Y. indep (insert x Y)" using assms augment psubset_card_mono indep_finite by blast subsection \Minors\ text \ A subset of the ground set induces a matroid. \ lemma matroid_subset [simp, intro]: assumes "\ \ carrier" shows "matroid \ (indep_in \)" unfolding matroid_def matroid_axioms_def proof (safe, goal_cases indep_system augment) case indep_system then show ?case using indep_system_subset[OF assms] . next case (augment X Y) then show ?case using augment_aux[of X Y] unfolding indep_in_def by auto qed context fixes \ assumes "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using \\ \ carrier\ by auto lemmas augment_aux_indep_in = \.augment_aux lemmas augment_indep_in = \.augment lemmas augment_psubset_indep_in = \.augment_psubset end subsection \Bases\ lemma basis_card: assumes "basis B\<^sub>1" assumes "basis B\<^sub>2" shows "card B\<^sub>1 = card B\<^sub>2" proof (rule ccontr, goal_cases False) case False then have "card B\<^sub>1 < card B\<^sub>2 \ card B\<^sub>2 < card B\<^sub>1" by auto moreover { fix B\<^sub>1 B\<^sub>2 assume "basis B\<^sub>1" "basis B\<^sub>2" "card B\<^sub>1 < card B\<^sub>2" then obtain x where "x \ B\<^sub>2 - B\<^sub>1" "indep (insert x B\<^sub>1)" using augment basisD by blast then have "x \ carrier - B\<^sub>1" using \basis B\<^sub>1\ basisD indep_subset_carrier by blast then have "\ indep (insert x B\<^sub>1)" using \basis B\<^sub>1\ basisD by auto then have "False" using \indep (insert x B\<^sub>1)\ by auto } ultimately show ?case using assms by auto qed lemma basis_indep_card: assumes "indep X" assumes "basis B" shows "card X \ card B" proof - obtain B' where "basis B'" "X \ B'" using assms indep_imp_subset_basis by auto then show ?thesis using assms basis_finite basis_card[of B B'] by (auto intro: card_mono) qed lemma basis_augment: assumes "basis B\<^sub>1" "basis B\<^sub>2" "x \ B\<^sub>1 - B\<^sub>2" shows "\y \ B\<^sub>2 - B\<^sub>1. basis (insert y (B\<^sub>1 - {x}))" proof - let ?B\<^sub>1 = "B\<^sub>1 - {x}" have "card ?B\<^sub>1 < card B\<^sub>2" using assms basis_card[of B\<^sub>1 B\<^sub>2] card_Diff1_less[OF basis_finite, of B\<^sub>1] by auto moreover have "indep ?B\<^sub>1" using assms basis_indep[of B\<^sub>1] indep_subset[of B\<^sub>1 ?B\<^sub>1] by auto ultimately obtain y where y: "y \ B\<^sub>2 - ?B\<^sub>1" "indep (insert y ?B\<^sub>1)" using assms augment[of B\<^sub>2 ?B\<^sub>1] basis_indep by auto let ?B\<^sub>1' = "insert y ?B\<^sub>1" have "basis ?B\<^sub>1'" using \indep ?B\<^sub>1'\ proof (rule basisI, goal_cases "insert") case (insert x) have "card (insert x ?B\<^sub>1') > card B\<^sub>1" proof - have "card (insert x ?B\<^sub>1') = Suc (card ?B\<^sub>1')" using insert card.insert_remove[OF indep_finite, of ?B\<^sub>1'] y by auto also have "\ = Suc (Suc (card ?B\<^sub>1))" using card.insert_remove[OF indep_finite, of ?B\<^sub>1] \indep ?B\<^sub>1\ y by auto also have "\ = Suc (card B\<^sub>1)" using assms basis_finite[of B\<^sub>1] card.remove[of B\<^sub>1] by auto finally show ?thesis by auto qed then have "\indep (insert x (insert y ?B\<^sub>1))" using assms basis_indep_card[of "insert x (insert y ?B\<^sub>1)" B\<^sub>1] by auto moreover have "insert x (insert y ?B\<^sub>1) \ carrier" using assms insert y basis_finite indep_subset_carrier by auto ultimately show ?case by auto qed then show ?thesis using assms y by auto qed context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using \\ \ carrier\ by auto lemmas basis_in_card = \.basis_card[OF basis_inD_aux[OF *] basis_inD_aux[OF *]] lemmas basis_in_indep_in_card = \.basis_indep_card[OF _ basis_inD_aux[OF *]] lemma basis_in_augment: assumes "basis_in \ B\<^sub>1" "basis_in \ B\<^sub>2" "x \ B\<^sub>1 - B\<^sub>2" shows "\y \ B\<^sub>2 - B\<^sub>1. basis_in \ (insert y (B\<^sub>1 - {x}))" using assms \.basis_augment unfolding basis_in_def by auto end subsection \Circuits\ lemma circuit_elim: assumes "circuit C\<^sub>1" "circuit C\<^sub>2" "C\<^sub>1 \ C\<^sub>2" "x \ C\<^sub>1 \ C\<^sub>2" shows "\C\<^sub>3 \ (C\<^sub>1 \ C\<^sub>2) - {x}. circuit C\<^sub>3" proof - let ?C = "(C\<^sub>1 \ C\<^sub>2) - {x}" let ?carrier = "C\<^sub>1 \ C\<^sub>2" have assms': "circuit_in carrier C\<^sub>1" "circuit_in carrier C\<^sub>2" using assms circuit_imp_circuit_in by auto have "?C \ carrier" using assms circuit_subset_carrier by auto show ?thesis proof (cases "indep ?C") case False then show ?thesis using dep_iff_supset_circuit \?C \ carrier\ by auto next case True then have "indep_in ?carrier ?C" using \?C \ carrier\ by (auto intro: indep_inI) have *: "?carrier \ carrier" using assms circuit_subset_carrier by auto obtain y where y: "y \ C\<^sub>2" "y \ C\<^sub>1" using assms circuit_subset_eq by blast then have "indep_in ?carrier (C\<^sub>2 - {y})" using assms' circuit_in_min_dep_in[OF * circuit_in_supI[OF *, of C\<^sub>2]] by auto then obtain B where B: "basis_in ?carrier B" "C\<^sub>2 - {y} \ B" using * assms indep_in_imp_subset_basis_in[of ?carrier "C\<^sub>2 - {y}"] by auto have "y \ B" proof (rule ccontr, goal_cases False) case False then have "C\<^sub>2 \ B" using B by auto moreover have "circuit_in ?carrier C\<^sub>2" using * assms' circuit_in_supI by auto ultimately have "\ indep_in ?carrier B" using B basis_in_subset_carrier[OF *] supset_circuit_in_imp_dep_in[OF *] by auto then show False using assms B basis_in_indep_in[OF *] by auto qed have "C\<^sub>1 - B \ {}" proof (rule ccontr, goal_cases False) case False then have "C\<^sub>1 - (C\<^sub>1 \ B) = {}" by auto then have "C\<^sub>1 = C\<^sub>1 \ B" using assms circuit_subset_eq by auto moreover have "indep (C\<^sub>1 \ B)" using assms B basis_in_indep_in[OF *] indep_in_subset[OF *, of B "C\<^sub>1 \ B"] indep_in_indep by auto ultimately show ?case using assms circuitD by auto qed then obtain z where z: "z \ C\<^sub>1" "z \ B" by auto have "y \ z" using y z by auto have "x \ C\<^sub>1" "x \ C\<^sub>2" using assms by auto have "finite ?carrier" using assms carrier_finite finite_subset by auto have "card B \ card (?carrier - {y, z})" proof (rule card_mono) show "finite (C\<^sub>1 \ C\<^sub>2 - {y, z})" using \finite ?carrier\ by auto next show "B \ C\<^sub>1 \ C\<^sub>2 - {y, z}" using B basis_in_subset_carrier[OF *, of B] \y \ B\ \z \ B\ by auto qed also have "\ = card ?carrier - 2" using \finite ?carrier\ \y \ C\<^sub>2\ \z \ C\<^sub>1\ \y \ z\ card_Diff_subset_Int by auto also have "\ < card ?carrier - 1" proof - have "card ?carrier = card C\<^sub>1 + card C\<^sub>2 - card (C\<^sub>1 \ C\<^sub>2)" using assms \finite ?carrier\ card_Un_Int[of C\<^sub>1 C\<^sub>2] by auto also have "\ = card C\<^sub>1 + (card C\<^sub>2 - card (C\<^sub>1 \ C\<^sub>2))" using assms \finite ?carrier\ card_mono[of C\<^sub>2] by auto also have "\ = card C\<^sub>1 + card (C\<^sub>2 - C\<^sub>1)" proof - have "card (C\<^sub>2 - C\<^sub>1) = card C\<^sub>2 - card (C\<^sub>2 \ C\<^sub>1)" using assms \finite ?carrier\ card_Diff_subset_Int[of C\<^sub>2 C\<^sub>1] by auto also have "\ = card C\<^sub>2 - card (C\<^sub>1 \ C\<^sub>2)" by (simp add: inf_commute) finally show ?thesis by auto qed finally have "card (C\<^sub>1 \ C\<^sub>2) = card C\<^sub>1 + card (C\<^sub>2 - C\<^sub>1)" . moreover have "card C\<^sub>1 > 0" using assms circuit_nonempty \finite ?carrier\ by auto moreover have "card (C\<^sub>2 - C\<^sub>1) > 0" using assms \finite ?carrier\ \y \ C\<^sub>2\ \y \ C\<^sub>1\ by auto ultimately show ?thesis by auto qed also have "\ = card ?C" using \finite ?carrier\ card_Diff_singleton \x \ C\<^sub>1\ \x \ C\<^sub>2\ by auto finally have "card B < card ?C" . then have False using basis_in_indep_in_card[OF *, of ?C B] B \indep_in ?carrier ?C\ by auto then show ?thesis by auto qed qed lemma min_dep_imp_supset_circuit: assumes "indep X" assumes "circuit C" assumes "C \ insert x X" shows "x \ C" proof (rule ccontr) assume "x \ C" then have "C \ X" using assms by auto then have "indep C" using assms indep_subset by auto then show False using assms circuitD by auto qed lemma min_dep_imp_ex1_supset_circuit: assumes "x \ carrier" assumes "indep X" assumes "\ indep (insert x X)" shows "\!C. circuit C \ C \ insert x X" proof - obtain C where C: "circuit C" "C \ insert x X" using assms indep_subset_carrier dep_iff_supset_circuit by auto show ?thesis proof (rule ex1I, goal_cases ex unique) show "circuit C \ C \ insert x X" using C by auto next { fix C' assume C': "circuit C'" "C' \ insert x X" have "C' = C" proof (rule ccontr) assume "C' \ C" moreover have "x \ C' \ C" using C C' assms min_dep_imp_supset_circuit by auto ultimately have "\ indep (C' \ C - {x})" using circuit_elim[OF C(1) C'(1), of x] supset_circuit_imp_dep[of _ "C' \ C - {x}"] by auto moreover have "C' \ C - {x} \ X" using C C' by auto ultimately show False using assms indep_subset by auto qed } then show "\C'. circuit C' \ C' \ insert x X \ C' = C" by auto qed qed lemma basis_ex1_supset_circuit: assumes "basis B" assumes "x \ carrier - B" shows "\!C. circuit C \ C \ insert x B" using assms min_dep_imp_ex1_supset_circuit basisD by auto definition fund_circuit :: "'a \ 'a set \ 'a set" where "fund_circuit x B \ (THE C. circuit C \ C \ insert x B)" lemma circuit_iff_fund_circuit: "circuit C \ (\x B. x \ carrier - B \ basis B \ C = fund_circuit x B)" proof (safe, goal_cases LTR RTL) case LTR then obtain x where "x \ C" using circuit_nonempty by auto then have "indep (C - {x})" using LTR unfolding circuit_def by auto then obtain B where B: "basis B" "C - {x} \ B" using indep_imp_subset_basis by auto then have "x \ carrier" using LTR circuit_subset_carrier \x \ C\ by auto moreover have "x \ B" proof (rule ccontr, goal_cases False) case False then have "C \ B" using \C - {x} \ B\ by auto then have "\ indep B" using LTR B basis_subset_carrier supset_circuit_imp_dep by auto then show ?case using B basis_indep by auto qed ultimately show ?case unfolding fund_circuit_def using LTR B theI_unique[OF basis_ex1_supset_circuit[of B x], of C] by auto next case (RTL x B) then have "\!C. circuit C \ C \ insert x B" using min_dep_imp_ex1_supset_circuit basisD[of B] by auto then show ?case unfolding fund_circuit_def using theI[of "\C. circuit C \ C \ insert x B"] by fastforce qed lemma fund_circuitI: assumes "basis B" assumes "x \ carrier - B" assumes "circuit C" assumes "C \ insert x B" shows "fund_circuit x B = C" unfolding fund_circuit_def using assms theI_unique[OF basis_ex1_supset_circuit, of B x C] by auto definition fund_circuit_in where "fund_circuit_in \ x B \ matroid.fund_circuit \ (indep_in \) x B" context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using \\ \ carrier\ by auto lemma fund_circuit_inI_aux: "\.fund_circuit x B = fund_circuit_in \ x B" unfolding fund_circuit_in_def by auto lemma circuit_in_elim: assumes "circuit_in \ C\<^sub>1" "circuit_in \ C\<^sub>2" "C\<^sub>1 \ C\<^sub>2" "x \ C\<^sub>1 \ C\<^sub>2" shows "\C\<^sub>3 \ (C\<^sub>1 \ C\<^sub>2) - {x}. circuit_in \ C\<^sub>3" using assms \.circuit_elim unfolding circuit_in_def by auto lemmas min_dep_in_imp_supset_circuit_in = \.min_dep_imp_supset_circuit[OF _ circuit_inD_aux[OF *]] lemma min_dep_in_imp_ex1_supset_circuit_in: assumes "x \ \" assumes "indep_in \ X" assumes "\ indep_in \ (insert x X)" shows "\!C. circuit_in \ C \ C \ insert x X" using assms \.min_dep_imp_ex1_supset_circuit unfolding circuit_in_def by auto lemma basis_in_ex1_supset_circuit_in: assumes "basis_in \ B" assumes "x \ \ - B" shows "\!C. circuit_in \ C \ C \ insert x B" using assms \.basis_ex1_supset_circuit unfolding circuit_in_def basis_in_def by auto lemma fund_circuit_inI: assumes "basis_in \ B" assumes "x \ \ - B" assumes "circuit_in \ C" assumes "C \ insert x B" shows "fund_circuit_in \ x B = C" using assms \.fund_circuitI unfolding basis_in_def circuit_in_def fund_circuit_in_def by auto end context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using \\ \ carrier\ by auto lemma fund_circuit_in_sub_cong: assumes "\' \ \" assumes "x \ \' - B" assumes "basis_in \' B" shows "\.fund_circuit_in \' x B = fund_circuit_in \' x B" proof - obtain C where C: "circuit_in \' C" "C \ insert x B" using * assms basis_in_ex1_supset_circuit_in[of \' B x] by auto then have "fund_circuit_in \' x B = C" using * assms fund_circuit_inI by auto also have "\ = \.fund_circuit_in \' x B" using * assms C \.fund_circuit_inI basis_in_sub_cong[of \] circuit_in_sub_cong[of \] by auto finally show ?thesis by auto qed end subsection \Ranks\ abbreviation rank_of where "rank_of \ lower_rank_of" lemmas rank_of_def = lower_rank_of_def lemmas rank_of_sub_cong = lower_rank_of_sub_cong lemmas rank_of_le = lower_rank_of_le context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using * by auto lemma lower_rank_of_eq_upper_rank_of: "lower_rank_of \ = upper_rank_of \" proof - obtain B where "basis_in \ B" using basis_in_ex[OF *] by auto then have "{card B | B. basis_in \ B} = {card B}" by safe (auto dest: basis_in_card[OF *]) then show ?thesis unfolding lower_rank_of_def upper_rank_of_def by auto qed lemma rank_of_eq_card_basis_in: assumes "basis_in \ B" shows "rank_of \ = card B" proof - have "{card B | B. basis_in \ B} = {card B}" using assms by safe (auto dest: basis_in_card[OF *]) then show ?thesis unfolding rank_of_def by auto qed lemma rank_of_indep_in_le: assumes "indep_in \ X" shows "card X \ rank_of \" proof - { fix B assume "basis_in \ B" moreover obtain B' where "basis_in \ B'" "X \ B'" using assms indep_in_imp_subset_basis_in[OF *] by auto ultimately have "card X \ card B" using card_mono[OF basis_in_finite[OF *]] basis_in_card[OF *, of B B'] by auto } moreover have "finite {card B | B. basis_in \ B}" using collect_basis_in_finite[OF *] by auto ultimately show ?thesis unfolding rank_of_def using basis_in_ex[OF *] by auto qed end lemma rank_of_mono: assumes "X \ Y" assumes "Y \ carrier" shows "rank_of X \ rank_of Y" proof - obtain B\<^sub>X where B\<^sub>X: "basis_in X B\<^sub>X" using assms basis_in_ex[of X] by auto moreover obtain B\<^sub>Y where B\<^sub>Y: "basis_in Y B\<^sub>Y" using assms basis_in_ex[of Y] by auto moreover have "card B\<^sub>X \ card B\<^sub>Y" using assms basis_in_indep_in_card[OF _ _ B\<^sub>Y] basis_in_indep_in[OF _ B\<^sub>X] indep_in_subI_subset by auto ultimately show ?thesis using assms rank_of_eq_card_basis_in by auto qed lemma rank_of_insert_le: assumes "X \ carrier" assumes "x \ carrier" shows "rank_of (insert x X) \ Suc (rank_of X)" proof - obtain B where B: "basis_in X B" using assms basis_in_ex[of X] by auto have "basis_in (insert x X) B \ basis_in (insert x X) (insert x B)" proof - obtain B' where B': "B' \ insert x X - X" "basis_in (insert x X) (B \ B')" using assms B basis_in_subI[of "insert x X" X B] by auto then have "B' = {} \ B' = {x}" by auto then show ?thesis proof assume "B' = {}" then have "basis_in (insert x X) B" using B' by auto then show ?thesis by auto next assume "B' = {x}" then have "basis_in (insert x X) (insert x B)" using B' by auto then show ?thesis by auto qed qed then show ?thesis proof assume "basis_in (insert x X) B" then show ?thesis using assms B rank_of_eq_card_basis_in by auto next assume "basis_in (insert x X) (insert x B)" then have "rank_of (insert x X) = card (insert x B)" using assms rank_of_eq_card_basis_in by auto also have "\ = Suc (card (B - {x}))" using assms card.insert_remove[of B x] using B basis_in_finite by auto also have "\ \ Suc (card B)" using assms B basis_in_finite card_Diff1_le[of B] by auto also have "\ = Suc (rank_of X)" using assms B rank_of_eq_card_basis_in by auto finally show ?thesis . qed qed lemma rank_of_Un_Int_le: assumes "X \ carrier" assumes "Y \ carrier" shows "rank_of (X \ Y) + rank_of (X \ Y) \ rank_of X + rank_of Y" proof - obtain B_Int where B_Int: "basis_in (X \ Y) B_Int" using assms basis_in_ex[of "X \ Y"] by auto then have "indep_in (X \ Y) B_Int" using assms indep_in_subI_subset[OF _ basis_in_indep_in[of "X \ Y" B_Int], of "X \ Y"] by auto then obtain B_Un where B_Un: "basis_in (X \ Y) B_Un" "B_Int \ B_Un" using assms indep_in_imp_subset_basis_in[of "X \ Y" B_Int] by auto have "card (B_Un \ (X \ Y)) + card (B_Un \ (X \ Y)) = card ((B_Un \ X) \ (B_Un \ Y)) + card ((B_Un \ X) \ (B_Un \ Y))" by (simp add: inf_assoc inf_left_commute inf_sup_distrib1) also have "\ = card (B_Un \ X) + card (B_Un \ Y)" proof - have "finite (B_Un \ X)" "finite (B_Un \ Y)" using assms finite_subset[OF _ carrier_finite] by auto then show ?thesis using card_Un_Int[of "B_Un \ X" "B_Un \ Y"] by auto qed also have "\ \ rank_of X + rank_of Y" proof - have "card (B_Un \ X) \ rank_of X" proof - have "indep_in X (B_Un \ X)" using assms basis_in_indep_in[OF _ B_Un(1)] indep_in_subI by auto then show ?thesis using assms rank_of_indep_in_le by auto qed moreover have "card (B_Un \ Y) \ rank_of Y" proof - have "indep_in Y (B_Un \ Y)" using assms basis_in_indep_in[OF _ B_Un(1)] indep_in_subI by auto then show ?thesis using assms rank_of_indep_in_le by auto qed ultimately show ?thesis by auto qed finally have "rank_of X + rank_of Y \ card (B_Un \ (X \ Y)) + card (B_Un \ (X \ Y))" . moreover have "B_Un \ (X \ Y) = B_Un" using assms basis_in_subset_carrier[OF _ B_Un(1)] by auto moreover have "B_Un \ (X \ Y) = B_Int" proof - have "card (B_Un \ (X \ Y)) \ card B_Int" proof - have "indep_in (X \ Y) (B_Un \ (X \ Y))" using assms basis_in_indep_in[OF _ B_Un(1)] indep_in_subI by auto then show ?thesis using assms basis_in_indep_in_card[of "X \ Y" _ B_Int] B_Int by auto qed moreover have "finite (B_Un \ (X \ Y))" using assms carrier_finite finite_subset[of "B_Un \ (X \ Y)"] by auto moreover have "B_Int \ B_Un \ (X \ Y)" using assms B_Un B_Int basis_in_subset_carrier[of "X \ Y" B_Int] by auto ultimately show ?thesis using card_seteq by blast qed ultimately have "rank_of X + rank_of Y \ card B_Un + card B_Int" by auto moreover have "card B_Un = rank_of (X \ Y)" using assms rank_of_eq_card_basis_in[OF _ B_Un(1)] by auto moreover have "card B_Int = rank_of (X \ Y)" using assms rank_of_eq_card_basis_in[OF _ B_Int] by fastforce ultimately show "rank_of X + rank_of Y \ rank_of (X \ Y) + rank_of (X \ Y)" by auto qed lemma rank_of_Un_absorbI: assumes "X \ carrier" "Y \ carrier" assumes "\y. y \ Y - X \ rank_of (insert y X) = rank_of X" shows "rank_of (X \ Y) = rank_of X" proof - have "finite (Y - X)" using finite_subset[OF \Y \ carrier\] carrier_finite by auto then show ?thesis using assms proof (induction "Y - X" arbitrary: Y rule: finite_induct ) case empty then have "X \ Y = X" by auto then show ?case by auto next case (insert y F) have "rank_of (X \ Y) + rank_of X \ rank_of X + rank_of X" proof - have "rank_of (X \ Y) + rank_of X = rank_of ((X \ (Y - {y})) \ (insert y X)) + rank_of ((X \ (Y - {y})) \ (insert y X))" proof - have "X \ Y = (X \ (Y - {y})) \ (insert y X)" "X = (X \ (Y - {y})) \ (insert y X)" using insert by auto then show ?thesis by auto qed also have "\ \ rank_of (X \ (Y - {y})) + rank_of (insert y X)" proof (rule rank_of_Un_Int_le) show "X \ (Y - {y}) \ carrier" using insert by auto next show "insert y X \ carrier" using insert by auto qed also have "\ = rank_of (X \ (Y - {y})) + rank_of X" proof - have "y \ Y - X" using insert by auto then show ?thesis using insert by auto qed also have "\ = rank_of X + rank_of X" proof - have "F = (Y - {y}) - X" "Y - {y} \ carrier" using insert by auto then show ?thesis using insert insert(3)[of "Y - {y}"] by auto qed finally show ?thesis . qed moreover have "rank_of (X \ Y) + rank_of X \ rank_of X + rank_of X" using insert rank_of_mono by auto ultimately show ?case by auto qed qed lemma indep_iff_rank_of: assumes "X \ carrier" shows "indep X \ rank_of X = card X" proof (standard, goal_cases LTR RTL) case LTR then have "indep_in X X" by (auto intro: indep_inI) then have "basis_in X X" by (auto intro: basis_inI[OF assms]) then show ?case using rank_of_eq_card_basis_in[OF assms] by auto next case RTL obtain B where B: "basis_in X B" using basis_in_ex[OF assms] by auto then have "card B = card X" using RTL rank_of_eq_card_basis_in[OF assms] by auto then have "B = X" using basis_in_subset_carrier[OF assms B] card_seteq[OF finite_subset[OF assms carrier_finite]] by auto then show ?case using basis_in_indep_in[OF assms B] indep_in_indep by auto qed lemma basis_iff_rank_of: assumes "X \ carrier" shows "basis X \ rank_of X = card X \ rank_of X = rank_of carrier" proof (standard, goal_cases LTR RTL) case LTR then have "rank_of X = card X" using assms indep_iff_rank_of basis_indep by auto moreover have "\ = rank_of carrier" using LTR rank_of_eq_card_basis_in[of carrier X] basis_iff_basis_in by auto ultimately show ?case by auto next case RTL show ?case proof (rule basisI) show "indep X" using assms RTL indep_iff_rank_of by blast next fix x assume x: "x \ carrier - X" show "\ indep (insert x X)" proof (rule ccontr, goal_cases False) case False then have "card (insert x X) \ rank_of carrier" using assms x indep_inI rank_of_indep_in_le by auto also have "\ = card X" using RTL by auto finally show ?case using finite_subset[OF assms carrier_finite] x by auto qed qed qed lemma circuit_iff_rank_of: assumes "X \ carrier" shows "circuit X \ X \ {} \ (\x \ X. rank_of (X - {x}) = card (X - {x}) \ card (X - {x}) = rank_of X)" proof (standard, goal_cases LTR RTL) case LTR then have "X \ {}" using circuit_nonempty by auto moreover have indep_remove: "\x. x \ X \ rank_of (X - {x}) = card (X - {x})" proof - fix x assume "x \ X" then have "indep (X - {x})" using circuit_min_dep[OF LTR] by auto moreover have "X - {x} \ carrier" using assms by auto ultimately show "rank_of (X - {x}) = card (X - {x})" using indep_iff_rank_of by auto qed moreover have "\x. x \ X \ rank_of (X - {x}) = rank_of X" proof - fix x assume *: "x \ X" have "rank_of X \ card X" using assms rank_of_le by auto moreover have "rank_of X \ card X" using assms LTR circuitD indep_iff_rank_of[of X] by auto ultimately have "rank_of X < card X" by auto then have "rank_of X \ card (X - {x})" using * finite_subset[OF assms] carrier_finite by auto also have "\ = rank_of (X - {x})" using indep_remove \x \ X\ by auto finally show "rank_of (X - {x}) = rank_of X" using assms rank_of_mono[of "X - {x}" X] by auto qed ultimately show ?case by auto next case RTL then have "X \ {}" and indep_remove: "\x. x \ X \ rank_of (X - {x}) = card (X - {x})" and dep: "\x. x \ X \ rank_of (X - {x}) = rank_of X" by auto show ?case using assms proof (rule circuitI) obtain x where x: "x \ X" using \X \ {}\ by auto then have "rank_of X = card (X - {x})" using dep indep_remove by auto also have "\ < card X" using card_Diff1_less[OF finite_subset[OF assms carrier_finite] x] . finally show "\ indep X" using indep_iff_rank_of[OF assms] by auto next fix x assume "x \ X" then show "indep (X - {x})" using assms indep_remove[of x] indep_iff_rank_of[of "X - {x}"] by auto qed qed context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using * by auto lemma indep_in_iff_rank_of: assumes "X \ \" shows "indep_in \ X \ rank_of X = card X" using assms \.indep_iff_rank_of rank_of_sub_cong[OF * assms] by auto lemma basis_in_iff_rank_of: assumes "X \ \" shows "basis_in \ X \ rank_of X = card X \ rank_of X = rank_of \" using \.basis_iff_rank_of[OF assms] rank_of_sub_cong[OF *] assms unfolding basis_in_def by auto lemma circuit_in_iff_rank_of: assumes "X \ \" shows "circuit_in \ X \ X \ {} \ (\x \ X. rank_of (X - {x}) = card (X - {x}) \ card (X - {x}) = rank_of X)" proof - have "circuit_in \ X \ \.circuit X" unfolding circuit_in_def .. also have "\ \ X \ {} \ (\x \ X. \.rank_of (X - {x}) = card (X - {x}) \ card (X - {x}) = \.rank_of X)" using \.circuit_iff_rank_of[OF assms] . also have "\ \ X \ {} \ (\x \ X. rank_of (X - {x}) = card (X - {x}) \ card (X - {x}) = rank_of X)" proof - { fix x have "\.rank_of (X - {x}) = rank_of (X - {x})" "\.rank_of X = rank_of X" using assms rank_of_sub_cong[OF *, of "X - {x}"] rank_of_sub_cong[OF *, of X] by auto then have "\.rank_of (X - {x}) = card (X - {x}) \ card (X - {x}) = \.rank_of X \ rank_of (X - {x}) = card (X - {x}) \ card (X - {x}) = rank_of X" by auto } then show ?thesis by (auto simp: simp del: card_Diff_insert) qed finally show ?thesis . qed end subsection \Closure\ definition cl :: "'a set \ 'a set" where "cl X \ {x \ carrier. rank_of (insert x X) = rank_of X}" lemma clI: assumes "x \ carrier" assumes "rank_of (insert x X) = rank_of X" shows "x \ cl X" unfolding cl_def using assms by auto lemma cl_altdef: assumes "X \ carrier" shows "cl X = \{Y \ Pow carrier. X \ Y \ rank_of Y = rank_of X}" proof - { fix x assume *: "x \ cl X" have "x \ \{Y \ Pow carrier. X \ Y \ rank_of Y = rank_of X}" proof show "insert x X \ {Y \ Pow carrier. X \ Y \ rank_of Y = rank_of X}" using assms * unfolding cl_def by auto qed auto } moreover { fix x assume *: "x \ \{Y \ Pow carrier. X \ Y \ rank_of Y = rank_of X}" then obtain Y where Y: "x \ Y" "Y \ carrier" "X \ Y" "rank_of Y = rank_of X" by auto have "rank_of (insert x X) = rank_of X" proof - have "rank_of (insert x X) \ rank_of X" proof - have "insert x X \ Y" using Y by auto then show ?thesis using rank_of_mono[of "insert x X" Y] Y by auto qed moreover have "rank_of X \ rank_of (insert x X)" using Y by (auto intro: rank_of_mono) ultimately show ?thesis by auto qed then have "x \ cl X" using * unfolding cl_def by auto } ultimately show ?thesis by blast qed lemma cl_rank_of: "x \ cl X \ rank_of (insert x X) = rank_of X" unfolding cl_def by auto lemma cl_subset_carrier: "cl X \ carrier" unfolding cl_def by auto lemmas clD = cl_rank_of cl_subset_carrier lemma cl_subset: assumes "X \ carrier" shows "X \ cl X" using assms using insert_absorb[of _ X] by (auto intro!: clI) lemma cl_mono: assumes "X \ Y" assumes "Y \ carrier" shows "cl X \ cl Y" proof fix x assume "x \ cl X" then have "x \ carrier" using cl_subset_carrier by auto have "insert x X \ carrier" using assms \x \ cl X\ cl_subset_carrier[of X] by auto then interpret X_insert: matroid "insert x X" "indep_in (insert x X)" by auto have "insert x Y \ carrier" using assms \x \ cl X\ cl_subset_carrier[of X] by auto then interpret Y_insert: matroid "insert x Y" "indep_in (insert x Y)" by auto show "x \ cl Y" using \x \ carrier\ proof (rule clI, cases "x \ X") case True then show "rank_of (insert x Y) = rank_of Y" using assms insert_absorb[of x Y] by auto next case False obtain B\<^sub>X where B\<^sub>X: "basis_in X B\<^sub>X" using assms basis_in_ex[of X] by auto have "basis_in (insert x X) B\<^sub>X" proof - have "rank_of B\<^sub>X = card B\<^sub>X \ rank_of B\<^sub>X = rank_of (insert x X)" proof - have "rank_of B\<^sub>X = card B\<^sub>X \ rank_of B\<^sub>X = rank_of X" using assms B\<^sub>X basis_in_subset_carrier[where \ = X and X = B\<^sub>X] basis_in_iff_rank_of[where \ = X and X = B\<^sub>X] by blast then show ?thesis using cl_rank_of[OF \x \ cl X\] by auto qed then show ?thesis using assms basis_in_subset_carrier[OF _ B\<^sub>X] \x \ carrier\ basis_in_iff_rank_of[of "insert x X" B\<^sub>X] by auto qed have "indep_in (insert x Y) B\<^sub>X" using assms basis_in_indep_in[OF _ B\<^sub>X] indep_in_subI_subset[of X "insert x Y"] by auto then obtain B\<^sub>Y where B\<^sub>Y: "basis_in (insert x Y) B\<^sub>Y" "B\<^sub>X \ B\<^sub>Y" using assms \x \ carrier\ indep_in_iff_subset_basis_in[of "insert x Y" B\<^sub>X] by auto have "basis_in Y B\<^sub>Y" proof - have "x \ B\<^sub>Y" proof (rule ccontr, goal_cases False) case False then have "insert x B\<^sub>X \ B\<^sub>Y" using \B\<^sub>X \ B\<^sub>Y\ by auto then have "indep_in (insert x Y) (insert x B\<^sub>X)" using assms \x \ carrier\ B\<^sub>Y basis_in_indep_in[where \ = "insert x Y" and X = B\<^sub>Y] indep_in_subset[where \ = "insert x Y" and X = B\<^sub>Y and Y = "insert x B\<^sub>X"] by auto then have "indep_in (insert x X) (insert x B\<^sub>X)" using assms B\<^sub>X basis_in_subset_carrier[where \ = X and X = B\<^sub>X] indep_in_supI[where \ = "insert x Y" and \' = "insert x X" and X = "insert x B\<^sub>X"] by auto moreover have "x \ insert x X - B\<^sub>X" using assms \x \ X\ B\<^sub>X basis_in_subset_carrier[where \ = X and X = B\<^sub>X] by auto ultimately show ?case using assms \x \ carrier\ \basis_in (insert x X) B\<^sub>X\ basis_in_max_indep_in[where \ = "insert x X" and X = B\<^sub>X and x = x] by auto qed then show ?thesis using assms \x \ carrier\ B\<^sub>Y basis_in_subset_carrier[of "insert x Y" B\<^sub>Y] basis_in_supI[where \ = "insert x Y" and \' = Y and B = B\<^sub>Y] by auto qed show "rank_of (insert x Y) = rank_of Y" proof - have "rank_of (insert x Y) = card B\<^sub>Y" using assms \x \ carrier\ \basis_in (insert x Y) B\<^sub>Y\ basis_in_subset_carrier using basis_in_iff_rank_of[where \ = "insert x Y" and X = B\<^sub>Y] by auto also have "\ = rank_of Y" using assms \x \ carrier\ \basis_in Y B\<^sub>Y\ basis_in_subset_carrier using basis_in_iff_rank_of[where \ = Y and X = B\<^sub>Y] by auto finally show ?thesis . qed qed qed lemma cl_insert_absorb: assumes "X \ carrier" assumes "x \ cl X" shows "cl (insert x X) = cl X" proof show "cl (insert x X) \ cl X" proof (standard, goal_cases elem) case (elem y) then have *: "x \ carrier" "y \ carrier" using assms cl_subset_carrier by auto have "rank_of (insert y X) = rank_of (insert y (insert x X))" proof - have "rank_of (insert y X) \ rank_of (insert y (insert x X))" using assms * by (auto intro: rank_of_mono) moreover have "rank_of (insert y (insert x X)) = rank_of (insert y X)" proof - have "insert y (insert x X) = insert x (insert y X)" by auto then have "rank_of (insert y (insert x X)) = rank_of (insert x (insert y X))" by auto also have "\ = rank_of (insert y X)" proof - have "cl X \ cl (insert y X)" by (rule cl_mono) (auto simp add: assms \y \ carrier\) then have "x \ cl (insert y X)" using assms by auto then show ?thesis unfolding cl_def by auto qed finally show ?thesis . qed ultimately show ?thesis by auto qed also have "\ = rank_of (insert x X)" using elem using cl_rank_of by auto also have "\ = rank_of X" using assms cl_rank_of by auto finally show "y \ cl X" using * by (auto intro: clI) qed next have "insert x X \ carrier" using assms cl_subset_carrier by auto moreover have "X \ insert x X" using assms by auto ultimately show "cl X \ cl (insert x X)" using assms cl_subset_carrier cl_mono by auto qed lemma cl_cl_absorb: assumes "X \ carrier" shows "cl (cl X) = cl X" proof show "cl (cl X) \ cl X" proof (standard, goal_cases elem) case (elem x) then have "x \ carrier" using cl_subset_carrier by auto then show ?case proof (rule clI) have "rank_of (insert x X) \ rank_of X" using assms \x \ carrier\ rank_of_mono[of X "insert x X"] by auto moreover have "rank_of (insert x X) \ rank_of X" proof - have "rank_of (insert x X) \ rank_of (insert x (cl X))" using assms \x \ carrier\ cl_subset_carrier cl_subset[of X] rank_of_mono[of "insert x X" "insert x (cl X)"] by auto also have "\ = rank_of (cl X)" using elem cl_rank_of by auto also have "\ = rank_of (X \ (cl X - X))" using Diff_partition[OF cl_subset[OF assms]] by auto also have "\ = rank_of X" using \X \ carrier\ proof (rule rank_of_Un_absorbI) show "cl X - X \ carrier" using assms cl_subset_carrier by auto next fix y assume "y \ cl X - X - X" then show "rank_of (insert y X) = rank_of X" unfolding cl_def by auto qed finally show ?thesis . qed ultimately show "rank_of (insert x X) = rank_of X" by auto qed qed next show "cl X \ cl (cl X)" using cl_subset[OF cl_subset_carrier] by auto qed lemma cl_augment: assumes "X \ carrier" assumes "x \ carrier" assumes "y \ cl (insert x X) - cl X" shows "x \ cl (insert y X)" using \x \ carrier\ proof (rule clI) have "rank_of (insert y X) \ rank_of (insert x (insert y X))" using assms cl_subset_carrier by (auto intro: rank_of_mono) moreover have "rank_of (insert x (insert y X)) \ rank_of (insert y X)" proof - have "rank_of (insert x (insert y X)) = rank_of (insert y (insert x X))" proof - have "insert x (insert y X) = insert y (insert x X)" by auto then show ?thesis by auto qed also have "rank_of (insert y (insert x X)) = rank_of (insert x X)" using assms cl_def by auto also have "\ \ Suc (rank_of X)" using assms cl_subset_carrier by (auto intro: rank_of_insert_le) also have "\ = rank_of (insert y X)" proof - have "rank_of (insert y X) \ Suc (rank_of X)" using assms cl_subset_carrier by (auto intro: rank_of_insert_le) moreover have "rank_of (insert y X) \ rank_of X" using assms cl_def by auto moreover have "rank_of X \ rank_of (insert y X)" using assms cl_subset_carrier by (auto intro: rank_of_mono) ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show "rank_of (insert x (insert y X)) = rank_of (insert y X)" by auto qed lemma clI_insert: assumes "x \ carrier" assumes "indep X" assumes "\ indep (insert x X)" shows "x \ cl X" using \x \ carrier\ proof (rule clI) have *: "X \ carrier" using assms indep_subset_carrier by auto then have **: "insert x X \ carrier" using assms by auto have "indep_in (insert x X) X" using assms by (auto intro: indep_inI) then obtain B where B: "basis_in (insert x X) B" "X \ B" using assms indep_in_iff_subset_basis_in[OF **] by auto have "x \ B" proof (rule ccontr, goal_cases False) case False then have "indep_in (insert x X) (insert x X)" using B indep_in_subset[OF ** basis_in_indep_in[OF **]] by auto then show ?case using assms indep_in_indep by auto qed have "basis_in X B" using * proof (rule basis_inI, goal_cases indep max_indep) case indep show ?case proof (rule indep_in_supI[where \ = "insert x X"]) show "B \ X" using B basis_in_subset_carrier[OF **] \x \ B\ by auto next show "indep_in (insert x X) B" using basis_in_indep_in[OF ** B(1)] . qed auto next case (max_indep y) then have "\ indep_in (insert x X) (insert y B)" using B basis_in_max_indep_in[OF **] by auto then show ?case by (auto intro: indep_in_subI_subset) qed then show "rank_of (insert x X) = rank_of X" using B rank_of_eq_card_basis_in[OF *] rank_of_eq_card_basis_in[OF **] by auto qed lemma indep_in_carrier [simp]: "indep_in carrier = indep" using indep_subset_carrier by (auto simp: indep_in_def fun_eq_iff) context fixes I defines "I \ (\X. X \ carrier \ (\x\X. x \ cl (X - {x})))" begin lemma I_mono: "I Y" if "Y \ X" "I X" for X Y :: "'a set" proof - have "\x\Y. x \ cl (Y - {x})" proof (intro ballI) fix x assume x: "x \ Y" with that have "cl (Y - {x}) \ cl (X - {x})" by (intro cl_mono) (auto simp: I_def) with that and x show "x \ cl (Y - {x})" by (auto simp: I_def) qed with that show ?thesis by (auto simp: I_def) qed lemma clI': assumes "I X" "x \ carrier" "\I (insert x X)" shows "x \ cl X" proof - from assms have x: "x \ X" by (auto simp: insert_absorb) from assms obtain y where y: "y \ insert x X" "y \ cl (insert x X - {y})" by (force simp: I_def) show "x \ cl X" proof (cases "x = y") case True thus ?thesis using assms x y by (auto simp: I_def) next case False have "y \ cl (insert x X - {y})" by fact also from False have "insert x X - {y} = insert x (X - {y})" by auto finally have "y \ cl (insert x (X - {y})) - cl (X - {y})" using assms False y unfolding I_def by blast hence "x \ cl (insert y (X - {y}))" using cl_augment[of "X - {y}" x y] assms False y by (auto simp: I_def) also from y and False have "insert y (X - {y}) = X" by auto finally show ?thesis . qed qed lemma matroid_I: "matroid carrier I" proof (unfold_locales, goal_cases) show "finite carrier" by (rule carrier_finite) next case (4 X Y) have "\x\Y. x \ cl (Y - {x})" proof (intro ballI) fix x assume x: "x \ Y" with 4 have "cl (Y - {x}) \ cl (X - {x})" by (intro cl_mono) (auto simp: I_def) with 4 and x show "x \ cl (Y - {x})" by (auto simp: I_def) qed with 4 show ?case by (auto simp: I_def) next case (5 X Y) have "~(\X Y. I X \ I Y \ card X < card Y \ (\x\Y-X. \I (insert x X)))" proof assume *: "\X Y. I X \ I Y \ card X < card Y \ (\x\Y-X. \I (insert x X))" (is "\X Y. ?P X Y") define n where "n = Max ((\(X, Y). card (X \ Y)) ` {(X, Y). ?P X Y})" have "{(X, Y). ?P X Y} \ Pow carrier \ Pow carrier" by (auto simp: I_def) hence finite: "finite {(X, Y). ?P X Y}" by (rule finite_subset) (insert carrier_finite, auto) hence "n \ ((\(X, Y). card (X \ Y)) ` {(X, Y). ?P X Y})" unfolding n_def using * by (intro Max_in finite_imageI) auto then obtain X Y where XY: "?P X Y" "n = card (X \ Y)" by auto hence finite': "finite X" "finite Y" using finite_subset[OF _ carrier_finite] XY by (auto simp: I_def) from XY finite' have "~(Y \ X)" using card_mono[of X Y] by auto then obtain y where y: "y \ Y - X" by blast have False proof (cases "X \ cl (Y - {y})") case True from y XY have [simp]: "y \ carrier" by (auto simp: I_def) assume "X \ cl (Y - {y})" hence "cl X \ cl (cl (Y - {y}))" by (intro cl_mono cl_subset_carrier) also have "\ = cl (Y - {y})" using XY by (intro cl_cl_absorb) (auto simp: I_def) finally have "cl X \ cl (Y - {y})" . moreover have "y \ cl (Y - {y})" using y I_def XY(1) by blast ultimately have "y \ cl X" by blast thus False unfolding I_def using XY y clI' \y \ carrier\ by blast next case False with y XY have [simp]: "y \ carrier" by (auto simp: I_def) assume "\(X \ cl (Y - {y}))" then obtain t where t: "t \ X" "t \ cl (Y - {y})" by auto with XY have [simp]: "t \ carrier" by (auto simp: I_def) have "t \ X - Y" using t y clI[of t "Y - {y}"] by (cases "t = y") (auto simp: insert_absorb) moreover have "I (Y - {y})" using XY(1) I_mono[of "Y - {y}" Y] by blast ultimately have *: "I (insert t (Y - {y}))" using clI'[of "Y - {y}" t] t by auto from XY have "finite Y" by (intro finite_subset[OF _ carrier_finite]) (auto simp: I_def) moreover from y have "Y \ {}" by auto ultimately have [simp]: "card (insert t (Y - {y})) = card Y" using \t \ X - Y\ y by (simp add: Suc_diff_Suc card_gt_0_iff) have "\x\Y - X. I (insert x X)" proof (rule ccontr) assume "\?thesis" hence "?P X (insert t (Y - {y}))" using XY * \t \ X - Y\ by auto hence "card (X \ insert t (Y - {y})) \ n" unfolding n_def using finite by (intro Max_ge) auto also have "X \ insert t (Y - {y}) = insert t ((X \ Y) - {y})" using y \t \ X - Y\ by blast also have "card \ = Suc (card (X \ Y))" using y \t \ X - Y\ \finite Y\ by (simp add: card.insert_remove) finally show False using XY by simp qed with XY show False by blast qed thus False . qed with 5 show ?case by auto qed (auto simp: I_def) end definition cl_in where "cl_in \ X = matroid.cl \ (indep_in \) X" lemma cl_eq_cl_in: assumes "X \ carrier" shows "cl X = cl_in carrier X" proof - interpret \: matroid carrier "indep_in carrier" by (intro matroid_subset) auto have "cl X = {x \ carrier. rank_of (insert x X) = rank_of X}" unfolding cl_def by auto also have "\ = {x \ carrier. \.rank_of (insert x X) = \.rank_of X}" using rank_of_sub_cong[of carrier] assms by auto also have "\ = cl_in carrier X" unfolding cl_in_def \.cl_def by auto finally show ?thesis . qed context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using * by auto lemma cl_inI_aux: "x \ \.cl X \ x \ cl_in \ X" unfolding cl_in_def by auto lemma cl_inD_aux: "x \ cl_in \ X \ x \ \.cl X" unfolding cl_in_def by auto lemma cl_inI: assumes "X \ \" assumes "x \ \" assumes "rank_of (insert x X) = rank_of X" shows "x \ cl_in \ X" proof - have "\.rank_of (insert x X) = rank_of (insert x X)" "\.rank_of X = rank_of X" using assms rank_of_sub_cong[OF *] by auto then show ?thesis unfolding cl_in_def using assms by (auto intro: \.clI) qed lemma cl_in_altdef: assumes "X \ \" shows "cl_in \ X = \{Y \ Pow \. X \ Y \ rank_of Y = rank_of X}" unfolding cl_in_def proof (safe, goal_cases LTR RTL) case (LTR x) then have "x \ \{Y \ Pow \. X \ Y \ \.rank_of Y = \.rank_of X}" using \.cl_altdef[OF assms] by auto then obtain Y where Y: "x \ Y" "Y \ Pow \" "X \ Y" "\.rank_of Y = \.rank_of X" by auto then show ?case using rank_of_sub_cong[OF *] by auto next case (RTL x Y) then have "x \ \{Y \ Pow \. X \ Y \ \.rank_of Y = \.rank_of X}" using rank_of_sub_cong[OF *, of X] rank_of_sub_cong[OF *, of Y] by auto then show ?case using \.cl_altdef[OF assms] by auto qed lemma cl_in_subset_carrier: "cl_in \ X \ \" using \.cl_subset_carrier unfolding cl_in_def . lemma cl_in_rank_of: assumes "X \ \" assumes "x \ cl_in \ X" shows "rank_of (insert x X) = rank_of X" proof - have "\.rank_of (insert x X) = \.rank_of X" using assms \.cl_rank_of unfolding cl_in_def by auto moreover have "\.rank_of (insert x X) = rank_of (insert x X)" using assms rank_of_sub_cong[OF *, of "insert x X"] cl_in_subset_carrier by auto moreover have "\.rank_of X = rank_of X" using assms rank_of_sub_cong[OF *] by auto ultimately show ?thesis by auto qed lemmas cl_inD = cl_in_rank_of cl_in_subset_carrier lemma cl_in_subset: assumes "X \ \" shows "X \ cl_in \ X" using \.cl_subset[OF assms] unfolding cl_in_def . lemma cl_in_mono: assumes "X \ Y" assumes "Y \ \" shows "cl_in \ X \ cl_in \ Y" using \.cl_mono[OF assms] unfolding cl_in_def . lemma cl_in_insert_absorb: assumes "X \ \" assumes "x \ cl_in \ X" shows "cl_in \ (insert x X) = cl_in \ X" using assms \.cl_insert_absorb unfolding cl_in_def by auto lemma cl_in_augment: assumes "X \ \" assumes "x \ \" assumes "y \ cl_in \ (insert x X) - cl_in \ X" shows "x \ cl_in \ (insert y X)" using assms \.cl_augment unfolding cl_in_def by auto lemmas cl_inI_insert = cl_inI_aux[OF \.clI_insert] end lemma cl_in_subI: assumes "X \ \'" "\' \ \" "\ \ carrier" shows "cl_in \' X \ cl_in \ X" proof (safe, goal_cases elem) case (elem x) then have "x \ \'" "rank_of (insert x X) = rank_of X" using assms cl_inD[where \ = \' and X = X] by auto then show "x \ cl_in \ X" using assms by (auto intro: cl_inI) qed context fixes \ assumes *: "\ \ carrier" begin interpretation \: matroid \ "indep_in \" using * by auto lemma cl_in_sub_cong: assumes "X \ \'" "\' \ \" shows "\.cl_in \' X = cl_in \' X" proof (safe, goal_cases LTR RTL) case (LTR x) then have "x \ \'" "\.rank_of (insert x X) = \.rank_of X" using assms \.cl_in_rank_of[where \ = \' and X = X and x = x] \.cl_in_subset_carrier[where \ = \'] by auto moreover have "\.rank_of X = rank_of X" using assms rank_of_sub_cong[OF *] by auto moreover have "\.rank_of (insert x X) = rank_of (insert x X)" using assms rank_of_sub_cong[OF *, of "insert x X"] \x \ \'\ by auto ultimately show ?case using assms * by (auto intro: cl_inI) next case (RTL x) then have "x \ \'" "rank_of (insert x X) = rank_of X" using * assms cl_inD[where \ = \' and X = X] by auto moreover have "\.rank_of X = rank_of X" using assms rank_of_sub_cong[OF *] by auto moreover have "\.rank_of (insert x X) = rank_of (insert x X)" using assms rank_of_sub_cong[OF *, of "insert x X"] \x \ \'\ by auto ultimately show ?case using assms by (auto intro: \.cl_inI) qed end end end \ No newline at end of file diff --git a/thys/Network_Security_Policy_Verification/Lib/FiniteGraph.thy b/thys/Network_Security_Policy_Verification/Lib/FiniteGraph.thy --- a/thys/Network_Security_Policy_Verification/Lib/FiniteGraph.thy +++ b/thys/Network_Security_Policy_Verification/Lib/FiniteGraph.thy @@ -1,323 +1,323 @@ theory FiniteGraph imports Main begin (*Lots of this theory is based on a work by Benedikt Nordhoff and Peter Lammich*) section \Specification of a finite directed graph\ text\A graph \G=(V,E)\ consits of a set of vertices \V\, also called nodes, and a set of edges \E\. The edges are tuples of vertices. Both, the set of vertices and edges is finite.\ (* Inspired by Title: Dijkstra's Shortest Path Algorithm Author: Benedikt Nordhoff and Peter Lammich http://isa-afp.org/entries/Dijkstra_Shortest_Path.shtml *) section \Graph\ subsection\Definitions\ text \A graph is represented by a record.\ record 'v graph = nodes :: "'v set" edges :: "('v \'v) set" text \In a well-formed graph, edges only go from nodes to nodes.\ locale wf_graph = fixes G :: "'v graph" \ \Edges only reference to existing nodes\ assumes E_wf: "fst ` (edges G) \ (nodes G)" "snd ` (edges G) \ (nodes G)" and finiteE: "finite (edges G)" (*implied by finiteV*) and finiteV: "finite (nodes G)" begin abbreviation "V \ (nodes G)" abbreviation "E \ (edges G)" lemma E_wfD: assumes "(v,v') \ E" shows "v \ V" "v' \ V" apply - apply (rule subsetD[OF E_wf(1)]) using assms apply force apply (rule subsetD[OF E_wf(2)]) using assms apply force done lemma E_wfD2: "\e \ E. fst e \ V \ snd e \ V" by (auto simp add: E_wfD) end subsection \Basic operations on Graphs\ text \The empty graph.\ definition empty :: "'v graph" where "empty \ \ nodes = {}, edges = {} \" text \Adds a node to a graph.\ definition add_node :: "'v \ 'v graph \ 'v graph" where "add_node v G \ \ nodes = ({v} \ (nodes G)), edges=edges G \" text \Deletes a node from a graph. Also deletes all adjacent edges.\ definition delete_node where "delete_node v G \ \ nodes = (nodes G) - {v}, edges = {(e1, e2). (e1, e2) \ edges G \ e1 \ v \ e2 \ v} \" text \Adds an edge to a graph.\ definition add_edge where "add_edge v v' G = \nodes = nodes G \ {v,v'}, edges = {(v, v')} \ edges G \" text \Deletes an edge from a graph.\ definition delete_edge where "delete_edge v v' G \ \ nodes = nodes G, edges = {(e1,e2). (e1, e2) \ edges G \ (e1,e2) \ (v,v')} \" definition delete_edges::"'v graph \ ('v \ 'v) set \ 'v graph" where "delete_edges G es \ \ nodes = nodes G, edges = {(e1,e2). (e1, e2) \ edges G \ (e1,e2) \ es} \" fun delete_edges_list::"'v graph \ ('v \ 'v) list \ 'v graph" where "delete_edges_list G [] = G"| "delete_edges_list G ((v,v')#es) = delete_edges_list (delete_edge v v' G) es" definition fully_connected :: "'v graph \ 'v graph" where "fully_connected G \ \nodes = nodes G, edges = nodes G \ nodes G \" text \Extended graph operations\ text \Reflexive transitive successors of a node. Or: All reachable nodes for \v\ including \v\.\ definition succ_rtran :: "'v graph \ 'v \ 'v set" where "succ_rtran G v = {e2. (v,e2) \ (edges G)\<^sup>*}" text \Transitive successors of a node. Or: All reachable nodes for \v\.\ definition succ_tran :: "'v graph \ 'v \ 'v set" where "succ_tran G v = {e2. (v,e2) \ (edges G)\<^sup>+}" \ \succ_tran is always finite\ lemma succ_tran_finite: "wf_graph G \ finite (succ_tran G v)" proof - assume "wf_graph G" from wf_graph.finiteE[OF this] have "finite ((edges G)\<^sup>+)" using finite_trancl[symmetric, of "edges G"] by metis from this have "finite {(e1,e2). (e1, e2) \ (edges G)\<^sup>+}" by simp from this have finite: "finite (snd ` {(e1,e2). (e1, e2) \ (edges G)\<^sup>+})" by (metis finite_imageI) have "{(e1,e2). (e1, e2) \ (edges G)\<^sup>+ \ e1 = v} \ {(e1,e2). (e1, e2) \ (edges G)\<^sup>+}" by blast have 1: "snd ` {(e1,e2). (e1, e2) \ (edges G)\<^sup>+ \ e1 = v} \ snd ` {(e1,e2). (e1, e2) \ (edges G)\<^sup>+}" by blast have 2: "snd ` {(e1,e2). (e1, e2) \ (edges G)\<^sup>+ \ e1 = v} = {e2. (v,e2) \ (edges G)\<^sup>+}" by force from 1 2 have "{e2. (v,e2) \ (edges G)\<^sup>+} \ snd ` {(e1,e2). (e1, e2) \ (edges G)\<^sup>+}" by blast from this finite have "finite {e2. (v, e2) \ (edges G)\<^sup>+}" by (metis finite_subset) thus "finite (succ_tran G v)" using succ_tran_def by metis qed text\If there is no edge leaving from \v\, then \v\ has no successors\ lemma succ_tran_empty: "\ wf_graph G; v \ (fst ` edges G) \ \ succ_tran G v = {}" unfolding succ_tran_def using image_iff tranclD by fastforce text\@{const succ_tran} is subset of nodes\ lemma succ_tran_subseteq_nodes: "\ wf_graph G \ \ succ_tran G v \ nodes G" unfolding succ_tran_def using tranclD2 wf_graph.E_wfD(2) by fastforce text \The number of reachable nodes from \v\\ definition num_reachable :: "'v graph \ 'v \ nat" where "num_reachable G v = card (succ_tran G v)" definition num_reachable_norefl :: "'v graph \ 'v \ nat" where "num_reachable_norefl G v = card (succ_tran G v - {v})" text\@{const card} returns @{term 0} for infinite sets. Here, for a well-formed graph, if @{const num_reachable} is zero, there are actually no nodes reachable.\ lemma num_reachable_zero: "\wf_graph G; num_reachable G v = 0\ \ succ_tran G v = {}" unfolding num_reachable_def apply(subgoal_tac "finite (succ_tran G v)") apply(simp) apply(blast intro: succ_tran_finite) done lemma num_succtran_zero: "\succ_tran G v = {}\ \ num_reachable G v = 0" unfolding num_reachable_def by simp lemma num_reachable_zero_iff: "\wf_graph G\ \ (num_reachable G v = 0) \ (succ_tran G v = {})" by(metis num_succtran_zero num_reachable_zero) section\Undirected Graph\ subsection\undirected graph simulation\ text \Create undirected graph from directed graph by adding backward links\ definition backflows :: "('v \ 'v) set \ ('v \ 'v) set" where "backflows E \ {(r,s). (s,r) \ E}" definition undirected :: "'v graph \ 'v graph" where "undirected G = \ nodes = nodes G, edges = (edges G) \ {(b,a). (a,b) \ edges G} \" section \Graph Lemmas\ lemma graph_eq_intro: "(nodes (G::'a graph) = nodes G') \ (edges G = edges G') \ G = G'" by simp \ \finite\ lemma wf_graph_finite_filterE: "wf_graph G \ finite {(e1, e2). (e1, e2) \ edges G \ P e1 e2}" by(simp add: wf_graph.finiteE split_def) lemma wf_graph_finite_filterV: "wf_graph G \ finite {n. n \ nodes G \ P n}" by(simp add: wf_graph.finiteV) \ \empty\ lemma empty_wf[simp]: "wf_graph empty" unfolding empty_def by unfold_locales auto lemma nodes_empty[simp]: "nodes empty = {}" unfolding empty_def by simp lemma edges_empty[simp]: "edges empty = {}" unfolding empty_def by simp \ \add node\ lemma add_node_wf[simp]: "wf_graph g \ wf_graph (add_node v g)" unfolding add_node_def wf_graph_def by (auto) lemma delete_node_wf[simp]: "wf_graph G \ wf_graph (delete_node v G)" by(auto simp add: delete_node_def wf_graph_def wf_graph_finite_filterE) \ \add edgde\ lemma add_edge_wf[simp]: "wf_graph G \ wf_graph (add_edge v v' G)" by(auto simp add: add_edge_def add_node_def wf_graph_def) \ \delete edge\ lemma delete_edge_wf[simp]: "wf_graph G \ wf_graph (delete_edge v v' G)" by(auto simp add: delete_edge_def add_node_def wf_graph_def split_def) \ \delte edges\ lemma delete_edges_list_wf[simp]: "wf_graph G \ wf_graph (delete_edges_list G E)" by(induction E arbitrary: G, simp, force) lemma delete_edges_wf[simp]: "wf_graph G \ wf_graph (delete_edges G E)" by(auto simp add: delete_edges_def add_node_def wf_graph_def split_def) lemma delete_edges_list_set: "delete_edges_list G E = delete_edges G (set E)" proof(induction E arbitrary: G) case Nil thus ?case by (simp add: delete_edges_def) next case (Cons e E) thus ?case by(cases e)(simp add: delete_edge_def delete_edges_def) qed lemma delete_edges_list_union: "delete_edges_list G (ff @ keeps) = delete_edges G (set ff \ set keeps)" by(simp add: delete_edges_list_set) lemma add_edge_delete_edges_list: "(add_edge (fst a) (snd a) (delete_edges_list G (a # ff))) = (add_edge (fst a) (snd a) (delete_edges G (set ff)))" by(auto simp add: delete_edges_list_set delete_edges_def add_edge_def add_node_def) lemma delete_edges_empty[simp]: "delete_edges G {} = G" by(simp add: delete_edges_def) lemma delete_edges_simp2: "delete_edges G E = \ nodes = nodes G, edges = edges G - E\" by(auto simp add: delete_edges_def) lemma delete_edges_set_nodes: "nodes (delete_edges G E) = nodes G" by(simp add: delete_edges_simp2) lemma delete_edges_edges_mono: "E' \ E \ edges (delete_edges G E) \ edges (delete_edges G E')" by(simp add: delete_edges_def, fast) lemma delete_edges_edges_empty: "(delete_edges G (edges G)) = G\edges := {}\" by(simp add: delete_edges_simp2) \ \add delete\ lemma add_delete_edge: "wf_graph (G::'a graph) \ (a,b) \ edges G \ add_edge a b (delete_edge a b G) = G" apply(simp add: delete_edge_def add_edge_def wf_graph_def) apply(intro graph_eq_intro) by auto lemma add_delete_edges: "wf_graph (G::'v graph) \ (a,b) \ edges G \ (a,b) \ fs \ add_edge a b (delete_edges G (insert (a, b) fs)) = (delete_edges G fs)" by(auto simp add: delete_edges_simp2 add_edge_def wf_graph_def) \ \fully_connected\ lemma fully_connected_simp: "fully_connected \nodes = N, edges = ignore \\ \nodes = N, edges = N \ N \" by(simp add: fully_connected_def) lemma fully_connected_wf: "wf_graph G \ wf_graph (fully_connected G)" by(simp add: fully_connected_def wf_graph_def) \ \succ_tran\ lemma succ_tran_mono: "wf_graph \nodes=N, edges=E\ \ E' \ E \ succ_tran \nodes=N, edges=E'\ v \ succ_tran \nodes=N, edges=E\ v" apply(drule wf_graph.finiteE) apply(frule_tac A="E'" in rev_finite_subset, simp) apply(simp add: num_reachable_def) apply(simp add: succ_tran_def) apply(metis (lifting, full_types) Collect_mono trancl_mono) done \ \num_reachable\ lemma num_reachable_mono: "wf_graph \nodes=N, edges=E\ \ E' \ E \ num_reachable \nodes=N, edges=E'\ v \ num_reachable \nodes=N, edges=E\ v" apply(simp add: num_reachable_def) apply(frule_tac E'="E'" and v="v" in succ_tran_mono, simp) apply(frule_tac v="v" in succ_tran_finite) apply(simp add: card_mono) done \ \num_reachable_norefl\ lemma num_reachable_norefl_mono: "wf_graph \nodes=N, edges=E\ \ E' \ E \ num_reachable_norefl \nodes=N, edges=E'\ v \ num_reachable_norefl \nodes=N, edges=E\ v" apply(simp add: num_reachable_norefl_def) apply(frule_tac E'="E'" and v="v" in succ_tran_mono, simp) apply(frule_tac v="v" in succ_tran_finite) using card_mono by (metis Diff_mono finite_Diff subset_refl) \ \backflows\ lemma backflows_wf: "wf_graph \nodes=N, edges=E\ \ wf_graph \nodes=N, edges=backflows E\" using [[simproc add: finite_Collect]] by(auto simp add: wf_graph_def backflows_def) lemma undirected_backflows: "undirected G = \ nodes = nodes G, edges = (edges G) \ backflows (edges G) \" by(simp add: backflows_def undirected_def) lemma backflows_id: "backflows (backflows E) = E" by(simp add: backflows_def) lemma backflows_finite: "finite E \ finite (backflows E)" using [[simproc add: finite_Collect]] by(simp add: backflows_def) lemma backflows_minus_backflows: "backflows (X - backflows Y) = (backflows X) - Y" by(auto simp add: backflows_def) lemma backflows_subseteq: "X \ Y \ backflows X \ backflows Y" by(auto simp add: backflows_def) lemma backflows_un: "backflows (A \ B) = (backflows A) \ (backflows B)" by(auto simp add: backflows_def) lemma backflows_inter: "backflows (A \ B) = (backflows A) \ (backflows B)" by(auto simp add: backflows_def) lemma backflows_alt_fstsnd: "backflows E = (\e. (snd e, fst e)) ` E" by(auto simp add: backflows_def, force) lemmas graph_ops=add_node_def delete_node_def add_edge_def delete_edge_def delete_edges_simp2 \ \wf_graph\ lemma wf_graph_remove_edges: "wf_graph \ nodes = V, edges = E \ \ wf_graph \ nodes = V, edges=E - X\" by (metis delete_edges_simp2 delete_edges_wf select_convs(1) select_convs(2)) lemma wf_graph_remove_edges_union: "wf_graph \ nodes = V, edges = E \ E' \ \ wf_graph \ nodes = V, edges=E\" by(auto simp add: wf_graph_def) lemma wf_graph_union_edges: "\ wf_graph \ nodes = V, edges = E \; wf_graph \ nodes = V, edges=E'\ \ \ wf_graph \ nodes = V, edges=E \ E'\" by(auto simp add: wf_graph_def) lemma wf_graph_add_subset_edges: "\ wf_graph \ nodes = V, edges = E \; E' \ E \ \ wf_graph \ nodes = V, edges= E \ E'\" - by(auto simp add: wf_graph_def) (metis rev_finite_subset) + by(auto simp add: wf_graph_def) (*Inspired by Benedikt Nordhoff and Peter Lammich Dijkstra's Shortest Path Algorithm http://isa-afp.org/entries/Dijkstra_Shortest_Path.shtml*) (*more a literal copy of http://isa-afp.org/browser_info/current/AFP/Dijkstra_Shortest_Path/Graph.html*) text \Successors of a node.\ definition succ :: "'v graph \ 'v \ 'v set" where "succ G v \ {v'. (v,v')\edges G}" lemma succ_finite[simp, intro]: "finite (edges G) \ finite (succ G v)" unfolding succ_def by (rule finite_subset[where B="snd`edges G"]) force+ lemma succ_empty: "succ empty v = {}" unfolding empty_def succ_def by auto lemma (in wf_graph) succ_subset: "succ G v \ V" unfolding succ_def using E_wf by (force) end diff --git a/thys/Ordinal_Partitions/Partitions.thy b/thys/Ordinal_Partitions/Partitions.thy --- a/thys/Ordinal_Partitions/Partitions.thy +++ b/thys/Ordinal_Partitions/Partitions.thy @@ -1,926 +1,926 @@ section \Ordinal Partitions\ text \Material from Jean A. Larson, A short proof of a partition theorem for the ordinal $\omega^\omega$. \emph{Annals of Mathematical Logic}, 6:129--145, 1973. Also from ``Partition Relations'' by A. Hajnal and J. A. Larson, in \emph{Handbook of Set Theory}, edited by Matthew Foreman and Akihiro Kanamori (Springer, 2010).\ theory Partitions imports Library_Additions "ZFC_in_HOL.ZFC_Typeclasses" "ZFC_in_HOL.Cantor_NF" begin abbreviation tp :: "V set \ V" where "tp A \ ordertype A VWF" subsection \Ordinal Partitions: Definitions\ definition partn_lst :: "[('a \ 'a) set, 'a set, V list, nat] \ bool" where "partn_lst r B \ n \ \f \ [B]\<^bsup>n\<^esup> \ {..}. \i < length \. \H. H \ B \ ordertype H r = (\!i) \ f ` (nsets H n) \ {i}" abbreviation partn_lst_VWF :: "V \ V list \ nat \ bool" where "partn_lst_VWF \ \ partn_lst VWF (elts \)" lemma partn_lst_E: assumes "partn_lst r B \ n" "f \ nsets B n \ {.." obtains i H where "i < l" "H \ B" "ordertype H r = \!i" "f ` (nsets H n) \ {i}" using assms by (auto simp: partn_lst_def) lemma partn_lst_VWF_nontriv: assumes "partn_lst_VWF \ \ n" "l = length \" "Ord \" "l > 0" obtains i where "i < l" "\!i \ \" proof - have "{.. {}" by (simp add: \l > 0\ lessThan_empty_iff) then obtain f where "f \ nsets (elts \) n \ {.. elts \" and eq: "tp H = \!i" using assms by (metis partn_lst_E) then have "\!i \ \" by (metis \H \ elts \\ \Ord \\ eq ordertype_le_Ord) then show thesis using \i < l\ that by auto qed lemma partn_lst_triv0: assumes "\!i = 0" "i < length \" "n \ 0" shows "partn_lst r B \ n" by (metis partn_lst_def assms bot_least image_empty nsets_empty_iff ordertype_empty) lemma partn_lst_triv1: assumes "\!i \ 1" "i < length \" "n > 1" "B \ {}" "wf r" shows "partn_lst r B \ n" unfolding partn_lst_def proof clarsimp obtain \ where "\ \ B" "\ \ []" using assms mem_0_Ord by fastforce have 01: "\!i = 0 \ \!i = 1" using assms by (fastforce simp: one_V_def) fix f assume f: "f \ [B]\<^bsup>n\<^esup> \ {..}" with assms have "ordertype {\} r = 1 \ f ` [{\}]\<^bsup>n\<^esup> \ {i}" "ordertype {} r = 0 \ f ` [{}]\<^bsup>n\<^esup> \ {i}" by (auto simp: one_V_def ordertype_insert nsets_eq_empty) with assms 01 show "\i. \H\B. ordertype H r = \ ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" using \\ \ B\ by auto qed lemma partn_lst_two_swap: assumes "partn_lst r B [x,y] n" shows "partn_lst r B [y,x] n" proof - { fix f :: "'a set \ nat" assume f: "f \ [B]\<^bsup>n\<^esup> \ {..<2}" then have f': "(\i. 1 - i) \ f \ [B]\<^bsup>n\<^esup> \ {..<2}" by (auto simp: Pi_def) obtain i H where "i<2" "H \ B" "ordertype H r = ([x,y]!i)" "((\i. 1 - i) \ f) ` (nsets H n) \ {i}" by (auto intro: partn_lst_E [OF assms f']) moreover have "f x = Suc 0" if "Suc 0 \ f x" "x\[H]\<^bsup>n\<^esup>" for x using f that \H \ B\ nsets_mono by (fastforce simp: Pi_iff) ultimately have "ordertype H r = [y,x] ! (1-i) \ f ` [H]\<^bsup>n\<^esup> \ {1-i}" by (force simp: eval_nat_numeral less_Suc_eq) then have "\i H. i<2 \ H\B \ ordertype H r = [y,x] ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" by (metis Suc_1 \H \ B\ diff_less_Suc) } then show ?thesis by (auto simp: partn_lst_def eval_nat_numeral) qed lemma partn_lst_greater_resource: assumes M: "partn_lst r B \ n" and "B \ C" shows "partn_lst r C \ n" proof (clarsimp simp: partn_lst_def) fix f assume "f \ [C]\<^bsup>n\<^esup> \ {..}" then have "f \ [B]\<^bsup>n\<^esup> \ {..}" by (metis \B \ C\ part_fn_def part_fn_subset) then obtain i H where "i < length \" and "H \ B" "ordertype H r = (\!i)" and "f ` nsets H n \ {i}" using M partn_lst_def by metis then show "\i. \H\C. ordertype H r = \ ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" using \B \ C\ by blast qed lemma partn_lst_less: assumes M: "partn_lst r B \ n" and eq: "length \' = length \" and "List.set \' \ ON" and le: "\i. i < length \ \ \'!i \ \!i " and r: "wf r" "trans r" "total_on B r" and "small B" shows "partn_lst r B \' n" proof (clarsimp simp: partn_lst_def) fix f assume "f \ [B]\<^bsup>n\<^esup> \ {..'}" then obtain i H where "i < length \" and "H \ B" "small H" and H: "ordertype H r = (\!i)" and fi: "f ` nsets H n \ {i}" using assms by (auto simp: partn_lst_def smaller_than_small) then have bij: "bij_betw (ordermap H r) H (elts (\!i))" using ordermap_bij [of r H] by (smt assms(8) in_mono r(1) r(3) smaller_than_small total_on_def) define H' where "H' = inv_into H (ordermap H r) ` (elts (\'!i))" have "H' \ H" using bij \i < length \\ bij_betw_imp_surj_on le by (force simp: H'_def image_subset_iff intro: inv_into_into) moreover have ot: "ordertype H' r = (\'!i)" proof (subst ordertype_eq_iff) show "Ord (\' ! i)" using assms by (simp add: \i < length \\ subset_eq) show "small H'" by (simp add: H'_def) show "\f. bij_betw f H' (elts (\' ! i)) \ (\x\H'. \y\H'. (f x < f y) = ((x, y) \ r))" proof (intro exI conjI ballI) show "bij_betw (ordermap H r) H' (elts (\' ! i))" using \H' \ H\ by (metis H'_def \i < length \\ bij bij_betw_inv_into_RIGHT bij_betw_subset le less_eq_V_def) show "(ordermap H r x < ordermap H r y) = ((x, y) \ r)" if "x \ H'" "y \ H'" for x y proof (intro iffI ordermap_mono_less) assume "ordermap H r x < ordermap H r y" then show "(x, y) \ r" by (metis \H \ B\ assms(8) calculation in_mono leD ordermap_mono_le r smaller_than_small that total_on_def) qed (use assms that \H' \ H\ \small H\ in auto) qed show "total_on H' r" using r by (meson \H \ B\ \H' \ H\ subsetD total_on_def) qed (use r in auto) ultimately show "\i'. \H\B. ordertype H r = \' ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" using \H \ B\ \i < length \\ fi assms by (metis image_mono nsets_mono subset_trans) qed text \Holds because no $n$-sets exist!\ lemma partn_lst_VWF_degenerate: assumes "k < n" shows "partn_lst_VWF \ (ord_of_nat k # \s) n" proof (clarsimp simp: partn_lst_def) fix f :: "V set \ nat" have "[elts (ord_of_nat k)]\<^bsup>n\<^esup> = {}" by (simp add: nsets_eq_empty assms finite_Ord_omega) then have "f ` [elts (ord_of_nat k)]\<^bsup>n\<^esup> \ {0}" by auto then show "\i < Suc (length \s). \H\elts \. tp H = (ord_of_nat k # \s) ! i \ f ` [H]\<^bsup>n\<^esup> \ {i}" using assms ordertype_eq_Ord [of "ord_of_nat k"] elts_ord_of_nat less_Suc_eq_0_disj by fastforce qed lemma partn_lst_VWF_\_2: assumes "Ord \" shows "partn_lst_VWF (\ \ (1+\)) [2, \ \ (1+\)] 2" (is "partn_lst_VWF ?\ _ _") proof (clarsimp simp: partn_lst_def) fix f assume f: "f \ [elts ?\]\<^bsup>2\<^esup> \ {..iH\elts ?\. tp H = [2, ?\] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" proof (cases "\x \ elts ?\. \y \ elts ?\. x \ y \ f{x,y} = 0") case True then obtain x y where "x \ elts ?\" "y \ elts ?\" "x \ y" "f {x, y} = 0" by auto then have "{x,y} \ elts ?\" "tp {x,y} = 2" "f ` [{x, y}]\<^bsup>2\<^esup> \ {0}" by auto (simp add: eval_nat_numeral ordertype_VWF_finite_nat) with \x \ y\ show ?thesis by (metis nth_Cons_0 zero_less_Suc) next case False with f have "\x\elts ?\. \y\elts ?\. x \ y \ f {x, y} = 1" unfolding Pi_iff using lessThan_Suc by force then have "tp (elts ?\) = ?\" "f ` [elts ?\]\<^bsup>2\<^esup> \ {Suc 0}" by (auto simp: assms nsets_2_eq) then show ?thesis by (metis lessI nth_Cons_0 nth_Cons_Suc subsetI) qed qed subsection \Relating partition properties on @{term VWF} to the general case\ text \Two very similar proofs here!\ lemma partn_lst_imp_partn_lst_VWF_eq: assumes part: "partn_lst r U \ n" and \: "ordertype U r = \" "small U" and r: "wf r" "trans r" "total_on U r" shows "partn_lst_VWF \ \ n" unfolding partn_lst_def proof clarsimp fix f assume f: "f \ [elts \]\<^bsup>n\<^esup> \ {..}" define cv where "cv \ \X. ordermap U r ` X" have bij: "bij_betw (ordermap U r) U (elts \)" using ordermap_bij [of "r" U] assms by blast then have bij_cv: "bij_betw cv ([U]\<^bsup>n\<^esup>) ([elts \]\<^bsup>n\<^esup>)" using bij_betw_nsets cv_def by blast then have func: "f \ cv \ [U]\<^bsup>n\<^esup> \ {..}" and "inj_on (ordermap U r) U" using bij bij_betw_def bij_betw_apply f by fastforce+ then have cv_part: "\\x\[X]\<^bsup>n\<^esup>. f (cv x) = i; X \ U; a \ [cv X]\<^bsup>n\<^esup>\ \ f a = i" for a X i n by (force simp: cv_def nsets_def subset_image_iff inj_on_subset finite_image_iff card_image) have ot_eq [simp]: "tp (cv X) = ordertype X r" if "X \ U" for X unfolding cv_def proof (rule ordertype_inc_eq) fix u v assume "u \ X" "v \ X" and "(u,v) \ r" with that have "ordermap U r u < ordermap U r v" by (simp add: assms ordermap_mono_less subset_eq) then show "(ordermap U r u, ordermap U r v) \ VWF" by (simp add: r) next show "total_on X r" using that r by (auto simp: total_on_def) show "small X" by (meson \small U\ smaller_than_small that) qed (use assms in auto) obtain X i where "X \ U" and X: "ordertype X r = \!i" "(f \ cv) ` [X]\<^bsup>n\<^esup> \ {i}" and "i < length \" using part func by (auto simp: partn_lst_def) show "\i < length \. \H\elts \. tp H = \!i \ f ` [H]\<^bsup>n\<^esup> \ {i}" proof (intro exI conjI) show "i < length \" by (simp add: \i < length \\) show "cv X \ elts \" using \X \ U\ bij bij_betw_imp_surj_on cv_def by blast show "tp (cv X) = \ ! i" by (simp add: X(1) \X \ U\) show "f ` [cv X]\<^bsup>n\<^esup> \ {i}" using X \X \ U\ cv_part unfolding image_subset_iff cv_def by (metis comp_apply insertCI singletonD) qed qed lemma partn_lst_imp_partn_lst_VWF: assumes part: "partn_lst r U \ n" and \: "ordertype U r \ \" "small U" and r: "wf r" "trans r" "total_on U r" shows "partn_lst_VWF \ \ n" by (metis assms less_eq_V_def partn_lst_imp_partn_lst_VWF_eq partn_lst_greater_resource) lemma partn_lst_VWF_imp_partn_lst_eq: assumes part: "partn_lst_VWF \ \ n" and \: "ordertype U r = \" "small U" and r: "wf r" "trans r" "total_on U r" shows "partn_lst r U \ n" unfolding partn_lst_def proof clarsimp fix f assume f: "f \ [U]\<^bsup>n\<^esup> \ {..}" define cv where "cv \ \X. inv_into U (ordermap U r) ` X" have bij: "bij_betw (ordermap U r) U (elts \)" using ordermap_bij [of "r" U] assms by blast then have bij_cv: "bij_betw cv ([elts \]\<^bsup>n\<^esup>) ([U]\<^bsup>n\<^esup>)" using bij_betw_nsets bij_betw_inv_into unfolding cv_def by blast then have func: "f \ cv \ [elts \]\<^bsup>n\<^esup> \ {..}" using bij_betw_apply f by fastforce have "inj_on (ordermap U r) U" using bij bij_betw_def by blast then have cv_part: "\\x\[X]\<^bsup>n\<^esup>. f (cv x) = i; X \ elts \; a \ [cv X]\<^bsup>n\<^esup>\ \ f a = i" for a X i n apply ( simp add: cv_def nsets_def subset_image_iff inj_on_subset finite_image_iff card_image) by (metis bij bij_betw_def card_image inj_on_finite inj_on_inv_into subset_eq) have ot_eq [simp]: "ordertype (cv X) r = tp X" if "X \ elts \" for X unfolding cv_def proof (rule ordertype_inc_eq) show "small X" using down that by auto show "(inv_into U (ordermap U r) x, inv_into U (ordermap U r) y) \ r" if "x \ X" "y \ X" and "(x,y) \ VWF" for x y proof - have xy: "x \ ordermap U r ` U" "y \ ordermap U r ` U" using \X \ elts \\ \x \ X\ \y \ X\ bij bij_betw_imp_surj_on by blast+ then have eq: "ordermap U r (inv_into U (ordermap U r) x) = x" "ordermap U r (inv_into U (ordermap U r) y) = y" by (meson f_inv_into_f)+ then have "y \ elts x" by (metis (no_types) VWF_non_refl mem_imp_VWF that(3) trans_VWF trans_def) then show ?thesis by (metis (no_types) VWF_non_refl xy eq assms(3) inv_into_into ordermap_mono r(1) r(3) that(3) total_on_def) qed qed (use r in auto) obtain X i where "X \ elts \" and X: "tp X = \!i" "(f \ cv) ` [X]\<^bsup>n\<^esup> \ {i}" and "i < length \" using part func by (auto simp: partn_lst_def) show "\i < length \. \H\U. ordertype H r = \!i \ f ` [H]\<^bsup>n\<^esup> \ {i}" proof (intro exI conjI) show "i < length \" by (simp add: \i < length \\) show "cv X \ U" using \X \ elts \\ bij bij_betw_imp_surj_on bij_betw_inv_into cv_def by blast show "ordertype (cv X) r = \ ! i" by (simp add: X(1) \X \ elts \\) show "f ` [cv X]\<^bsup>n\<^esup> \ {i}" using X \X \ elts \\ cv_part unfolding image_subset_iff cv_def by (metis comp_apply insertCI singletonD) qed qed corollary partn_lst_VWF_imp_partn_lst: assumes "partn_lst_VWF \ \ n" and \: "ordertype U r \ \" "small U" "wf r" "trans r" "total_on U r" shows "partn_lst r U \ n" by (metis assms less_eq_V_def partn_lst_VWF_imp_partn_lst_eq partn_lst_greater_resource) subsection \Simple consequences of the definitions\ text \A restatement of the infinite Ramsey theorem using partition notation\ lemma Ramsey_partn: "partn_lst_VWF \ [\,\] 2" proof (clarsimp simp: partn_lst_def) fix f assume "f \ [elts \]\<^bsup>2\<^esup> \ {..x\elts \. \y\elts \. x \ y \ f {x, y} < 2" by (auto simp: nsets_def eval_nat_numeral) obtain H i where H: "H \ elts \" and "infinite H" and t: "i < Suc (Suc 0)" and teq: "\x\H. \y\H. x \ y \ f {x, y} = i" using Ramsey2 [OF infinite_\ *] by (auto simp: eval_nat_numeral) then have "tp H = [\, \] ! i" using less_2_cases eval_nat_numeral ordertype_infinite_\ by force moreover have "f ` {N. N \ H \ finite N \ card N = 2} \ {i}" by (force simp: teq card_2_iff) ultimately have "f ` [H]\<^bsup>2\<^esup> \ {i}" by (metis (no_types) nsets_def numeral_2_eq_2) then show "\iH\elts \. tp H = [\,\] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" using H \tp H = [\, \] ! i\ t by blast qed text \This is the counterexample sketched in Hajnal and Larson, section 9.1.\ proposition omega_basic_counterexample: assumes "Ord \" shows "\ partn_lst_VWF \ [succ (vcard \), \] 2" proof - obtain \ where fun\: "\ \ elts \ \ elts (vcard \)" and inj\: "inj_on \ (elts \)" using inj_into_vcard by auto have Ord\: "Ord (\ x)" if "x \ elts \" for x using Ord_in_Ord fun\ that by fastforce define f where "f A \ @i::nat. \x y. A = {x,y} \ x < y \ (\ x < \ y \ i=0 \ \ y < \ x \ i=1)" for A have f_Pi: "f \ [elts \]\<^bsup>2\<^esup> \ {.. [elts \]\<^bsup>2\<^esup>" then obtain x y where xy: "x \ elts \" "y \ elts \" "x < y" and A: "A = {x,y}" apply (clarsimp simp: nsets_2_eq) by (metis Ord_in_Ord Ord_linear_lt assms insert_commute) consider "\ x < \ y" | "\ y < \ x" by (metis Ord\ Ord_linear_lt inj\ inj_onD less_imp_not_eq2 xy) then show "f A \ {..x < y\ A exE_some [OF _ f_def]) qed have fiff: "\ x < \ y \ i=0 \ \ y < \ x \ i=1" if f: "f {x,y} = i" and xy: "x \ elts \" "y \ elts \" "x x < \ y" | "\ y < \ x" using xy by (metis Ord\ Ord_linear_lt inj\ inj_onD less_V_def) then show ?thesis proof cases case 1 then have "f{x,y} = 0" using \x by (force simp: f_def Set.doubleton_eq_iff) then show ?thesis using "1" f by auto next case 2 then have "f{x,y} = 1" using \x by (force simp: f_def Set.doubleton_eq_iff) then show ?thesis using "2" f by auto qed qed have False if eq: "tp H = succ (vcard \)" and H: "H \ elts \" and 0: "\A. A \ [H]\<^bsup>2\<^esup> \ f A = 0" for H proof - have [simp]: "small H" using H down by auto have OH: "Ord x" if "x \ H" for x using H Ord_in_Ord \Ord \\ that by blast have \: "\ x < \ y" if "x\H" "y\H" "x ` H \ elts (vcard \)" using H fun\ by auto have "tp H = tp (\ ` H)" proof (rule ordertype_VWF_inc_eq [symmetric]) show "\ ` H \ ON" using H Ord\ by blast qed (auto simp: \ OH subsetI) also have "\ \ vcard \" by (simp add: H sub_vcard assms ordertype_le_Ord) finally show False by (simp add: eq succ_le_iff) qed moreover have False if eq: "tp H = \" and H: "H \ elts \" and 1: "\A. A \ [H]\<^bsup>2\<^esup> \ f A = Suc 0" for H proof - have [simp]: "small H" using H down by auto define \ where "\ \ inv_into H (ordermap H VWF) \ ord_of_nat" have bij: "bij_betw (ordermap H VWF) H (elts \)" by (metis ordermap_bij \small H\ eq total_on_VWF wf_VWF) then have "bij_betw (inv_into H (ordermap H VWF)) (elts \) H" by (simp add: bij_betw_inv_into) then have \: "bij_betw \ UNIV H" unfolding \_def by (metis \_def bij_betw_comp_iff2 bij_betw_def elts_of_set inf inj_ord_of_nat order_refl) moreover have Ord\: "Ord (\ k)" for k by (meson H Ord_in_Ord UNIV_I \ assms bij_betw_apply subsetD) moreover obtain k where k: "(\ (\(Suc k)), \ (\ k)) \ VWF" by (meson wf_VWF wf_iff_no_infinite_down_chain) moreover have \: "\ y < \ x" if "x\H" "y\H" "x (Suc k) \ \ k" by (metis H Ord\ Ord_linear2 VWF_iff_Ord_less bij_betw_def rangeI subset_iff) then have "(\ (Suc k), \ k) \ VWF \ \ (Suc k) = \ k" using Ord\ Ord_mem_iff_lt by auto then have "ordermap H VWF (\ (Suc k)) \ ordermap H VWF (\ k)" by (metis \ \small H\ bij_betw_imp_surj_on ordermap_mono_le rangeI trans_VWF wf_VWF) moreover have "ordermap H VWF (\ (Suc k)) = succ (ord_of_nat k)" unfolding \_def using bij bij_betw_inv_into_right by force moreover have "ordermap H VWF (\ k) = ord_of_nat k" using \_def bij bij_betw_inv_into_right by force ultimately show False by (simp add: less_eq_V_def) qed ultimately show ?thesis apply (simp add: partn_lst_def image_subset_iff) by (metis f_Pi less_2_cases nth_Cons_0 nth_Cons_Suc numeral_2_eq_2) qed subsection \Specker's theorem\ definition form_split :: "[nat,nat,nat,nat,nat] \ bool" where "form_split a b c d i \ a \ c \ (i=0 \ a b c i=1 \ a c b i=2 \ a c d i=3 \ a=c \ b\d)" definition form :: "[(nat*nat)set, nat] \ bool" where "form u i \ \a b c d. u = {(a,b),(c,d)} \ form_split a b c d i" definition scheme :: "[(nat*nat)set] \ nat set" where "scheme u \ fst ` u \ snd ` u" definition UU :: "(nat*nat) set" where "UU \ {(a,b). a < b}" lemma ordertype_UNIV_\2: "ordertype UNIV pair_less = \\2" using ordertype_Times [of concl: UNIV UNIV less_than less_than] by (simp add: total_less_than pair_less_def ordertype_nat_\ numeral_2_eq_2) lemma ordertype_UU_ge_\2: "ordertype UNIV pair_less \ ordertype UU pair_less" proof (rule ordertype_inc_le) define \ where "\ \ \(m,n). (m, Suc (m+n))" show "(\ (x::nat \ nat), \ y) \ pair_less" if "(x, y) \ pair_less" for x y using that by (auto simp: \_def pair_less_def split: prod.split) show "range \ \ UU" by (auto simp: \_def UU_def) qed auto lemma ordertype_UU_\2: "ordertype UU pair_less = \\2" by (metis eq_iff ordertype_UNIV_\2 ordertype_UU_ge_\2 ordertype_mono small top_greatest trans_pair_less wf_pair_less) text \Lemma 2.3 of Jean A. Larson, A short proof of a partition theorem for the ordinal $\omega^\omega$. \emph{Annals of Mathematical Logic}, 6:129--145, 1973.\ lemma lemma_2_3: fixes f :: "(nat \ nat) set \ nat" assumes "f \ [UU]\<^bsup>2\<^esup> \ {..k u. \k < 4; u \ [UU]\<^bsup>2\<^esup>; form u k; scheme u \ N\ \ f u = js!k" proof - have f_less2: "f {p,q} < Suc (Suc 0)" if "p \ q" "p \ UU" "q \ UU" for p q proof - have "{p,q} \ [UU]\<^bsup>2\<^esup>" using that by (simp add: nsets_def) then show ?thesis using assms by (simp add: Pi_iff) qed define f0 where "f0 \ (\A::nat set. THE x. \a b c d. A = {a,b,c,d} \ a b c x = f {(a,b),(c,d)})" have f0: "f0 {a,b,c,d} = f {(a,b),(c,d)}" if "a [X]\<^bsup>4\<^esup>" using that by (auto simp: nsets_def) then obtain a b c d where "X = {a,b,c,d} \ a b cN t. infinite N \ t < Suc (Suc 0) \ (\X. X \ N \ finite X \ card X = 4 \ f0 X = t)" using Ramsey [of UNIV 4 f0 2] by (simp add: eval_nat_numeral) then obtain N0 j0 where "infinite N0" and j0: "j0 < Suc (Suc 0)" and N0: "\A. A \ [N0]\<^bsup>4\<^esup> \ f0 A = j0" by (auto simp: nsets_def) define f1 where "f1 \ (\A::nat set. THE x. \a b c d. A = {a,b,c,d} \ a b c x = f {(a,c),(b,d)})" have f1: "f1 {a,b,c,d} = f {(a,c),(b,d)}" if "a [X]\<^bsup>4\<^esup>" using that by (auto simp: nsets_def) then obtain a b c d where "X = {a,b,c,d} \ a b cN t. N \ N0 \ infinite N \ t < Suc (Suc 0) \ (\X. X \ N \ finite X \ card X = 4 \ f1 X = t)" using \infinite N0\ Ramsey [of N0 4 f1 2] by (simp add: eval_nat_numeral) then obtain N1 j1 where "N1 \ N0" "infinite N1" and j1: "j1 < Suc (Suc 0)" and N1: "\A. A \ [N1]\<^bsup>4\<^esup> \ f1 A = j1" by (auto simp: nsets_def) define f2 where "f2 \ (\A::nat set. THE x. \a b c d. A = {a,b,c,d} \ a b c x = f {(a,d),(b,c)})" have f2: "f2 {a,b,c,d} = f {(a,d),(b,c)}" if "a [X]\<^bsup>4\<^esup>" using that by (auto simp: nsets_def) then obtain a b c d where "X = {a,b,c,d} \ a b cN t. N \ N1 \ infinite N \ t < Suc (Suc 0) \ (\X. X \ N \ finite X \ card X = 4 \ f2 X = t)" using \infinite N1\ Ramsey [of N1 4 f2 2] by (simp add: eval_nat_numeral) then obtain N2 j2 where "N2 \ N1" "infinite N2" and j2: "j2 < Suc (Suc 0)" and N2: "\A. A \ [N2]\<^bsup>4\<^esup> \ f2 A = j2" by (auto simp: nsets_def) define f3 where "f3 \ (\A::nat set. THE x. \a b c. A = {a,b,c} \ a b x = f {(a,b),(a,c)})" have f3: "f3 {a,b,c} = f {(a,b),(a,c)}" if "a [X]\<^bsup>3\<^esup>" using that by (auto simp: nsets_def) then obtain a b c where "X = {a,b,c} \ a bN t. N \ N2 \ infinite N \ t < Suc (Suc 0) \ (\X. X \ N \ finite X \ card X = 3 \ f3 X = t)" using \infinite N2\ Ramsey [of N2 3 f3 2] by (simp add: eval_nat_numeral) then obtain N3 j3 where "N3 \ N2" "infinite N3" and j3: "j3 < Suc (Suc 0)" and N3: "\A. A \ [N3]\<^bsup>3\<^esup> \ f3 A = j3" by (auto simp: nsets_def) show thesis proof fix k u assume "k < 4" and u: "form u k" "scheme u \ N3" and UU: "u \ [UU]\<^bsup>2\<^esup>" then consider (0) "k=0" | (1) "k=1" | (2) "k=2" | (3) "k=3" by linarith then show "f u = [j0,j1,j2,j3] ! k" proof cases case 0 have "N3 \ N0" using \N1 \ N0\ \N2 \ N1\ \N3 \ N2\ by auto then show ?thesis using u 0 apply (auto simp: form_def form_split_def scheme_def simp flip: f0) apply (force simp: nsets_def intro: N0) done next case 1 have "N3 \ N1" using \N2 \ N1\ \N3 \ N2\ by auto then show ?thesis using u 1 apply (auto simp: form_def form_split_def scheme_def simp flip: f1) apply (force simp: nsets_def intro: N1) done next case 2 then show ?thesis using u \N3 \ N2\ apply (auto simp: form_def form_split_def scheme_def nsets_def simp flip: f2) apply (force simp: nsets_def intro: N2) done next case 3 { fix a b d assume "{(a, b), (a, d)} \ [UU]\<^bsup>2\<^esup>" and *: "a \ N3" "b \ N3" "d \ N3" "b \ d" then have "ainfinite N3\) qed text \Lemma 2.4 of Jean A. Larson, ibid.\ lemma lemma_2_4: assumes "infinite N" "k < 4" obtains M where "M \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M]\<^bsup>2\<^esup> \ form u k" "\u. u \ [M]\<^bsup>2\<^esup> \ scheme u \ N" proof - obtain f:: "nat \ nat" where "bij_betw f UNIV N" "strict_mono f" using assms by (meson bij_enumerate enumerate_mono strict_monoI) then have iff[simp]: "f x = f y \ x=y" "f x < f y \ x N" for x using bij_betw_apply [OF \bij_betw f UNIV N\] by blast define M0 where "M0 = (\i. (f(2*i), f(Suc(2*i)))) ` {..i. (f i, f(m+i))) ` {..i. (f i, f(2*m-i))) ` {..i. (f 0, f (Suc i))) ` {..i. (f (2 * i), f (Suc (2 * i)))) {.. [UU]\<^bsup>m\<^esup>" by (simp add: M0_def nsets_def card_image UU_def image_subset_iff) next fix u assume u: "(u::(nat \ nat) set) \ [M0]\<^bsup>2\<^esup>" then obtain x y where "u = {x,y}" "x \ y" "x \ M0" "y \ M0" by (auto simp: nsets_2_eq) then obtain i j where "i f (2 * j)" by (simp add: \i less_imp_le_nat) ultimately show "form u k" apply (simp add: 0 form_def form_split_def nsets_def) apply (rule_tac x="f (2 * i)" in exI) apply (rule_tac x="f (Suc (2 * i))" in exI) apply (rule_tac x="f (2 * j)" in exI) apply (rule_tac x="f (Suc (2 * j))" in exI) apply auto done show "scheme u \ N" using ueq by (auto simp: scheme_def) qed next case 1 show ?thesis proof have "inj_on (\i. (f i, f(m+i))) {.. [UU]\<^bsup>m\<^esup>" by (simp add: M1_def nsets_def card_image UU_def image_subset_iff) next fix u assume u: "(u::(nat \ nat) set) \ [M1]\<^bsup>2\<^esup>" then obtain x y where "u = {x,y}" "x \ y" "x \ M1" "y \ M1" by (auto simp: nsets_2_eq) then obtain i j where "i N" using ueq by (auto simp: scheme_def) qed next case 2 show ?thesis proof have "inj_on (\i. (f i, f(2*m-i))) {.. [UU]\<^bsup>m\<^esup>" by (auto simp: M2_def nsets_def card_image UU_def image_subset_iff) next fix u assume u: "(u::(nat \ nat) set) \ [M2]\<^bsup>2\<^esup>" then obtain x y where "u = {x,y}" "x \ y" "x \ M2" "y \ M2" by (auto simp: nsets_2_eq) then obtain i j where "i N" using ueq by (auto simp: scheme_def) qed next case 3 show ?thesis proof have "inj_on (\i. (f 0, f (Suc i))) {.. [UU]\<^bsup>m\<^esup>" by (auto simp: M3_def nsets_def card_image UU_def image_subset_iff) next fix u assume u: "(u::(nat \ nat) set) \ [M3]\<^bsup>2\<^esup>" then obtain x y where "u = {x,y}" "x \ y" "x \ M3" "y \ M3" by (auto simp: nsets_2_eq) then obtain i j where "i N" using ueq by (auto simp: scheme_def) qed qed qed text \Lemma 2.5 of Jean A. Larson, ibid.\ lemma lemma_2_5: assumes "infinite N" obtains X where "X \ UU" "ordertype X pair_less = \\2" "\u. u \ [X]\<^bsup>2\<^esup> \ (\k<4. form u k) \ scheme u \ N" proof - obtain C where dis: "pairwise (\i j. disjnt (C i) (C j)) UNIV" and N: "(\i. C i) \ N" and infC: "\i::nat. infinite (C i)" using assms infinite_infinite_partition by blast then have "\\::nat \ nat. inj \ \ range \ = C i \ strict_mono \" for i by (metis bij_betw_imp_inj_on bij_betw_imp_surj_on bij_enumerate enumerate_mono infC strict_mono_def) then obtain \:: "[nat,nat] \ nat" where \: "\i. inj (\ i) \ range (\ i) = C i \ strict_mono (\ i)" by metis then have \_in_C [simp]: "\ i j \ C i' \ i'=i" for i i' j using dis by (fastforce simp: pairwise_def disjnt_def) have less_iff [simp]: "\ i j' < \ i j \ j' < j" for i j' j by (simp add: \ strict_mono_less) let ?a = "\ 0" define X where "X \ {(?a i, b) | i b. ?a i < b \ b \ C (Suc i)}" show thesis proof show "X \ UU" by (auto simp: X_def UU_def) show "ordertype X pair_less = \\2" proof (rule antisym) have "ordertype X pair_less \ ordertype UU pair_less" by (simp add: \X \ UU\ ordertype_mono) then show "ordertype X pair_less \ \\2" using ordertype_UU_\2 by auto define \ where "\ \ \(i,j::nat). (?a i, \ (Suc i) (?a j))" have "\i j. i < j \ \ 0 i < \ (Suc i) (\ 0 j)" by (meson \ le_less_trans less_iff strict_mono_imp_increasing) then have subX: "\ ` UU \ X" by (auto simp: UU_def \_def X_def) then have "ordertype (\ ` UU) pair_less \ ordertype X pair_less" by (simp add: ordertype_mono) moreover have "ordertype (\ ` UU) pair_less = ordertype UU pair_less" proof (rule ordertype_inc_eq) show "(\ x, \ y) \ pair_less" if "x \ UU" "y \ UU" and "(x, y) \ pair_less" for x y using that by (auto simp: UU_def \_def pair_less_def) qed auto ultimately show "\\2 \ ordertype X pair_less" using ordertype_UU_\2 by simp qed next fix U assume "U \ [X]\<^bsup>2\<^esup>" then obtain a b c d where Ueq: "U = {(a,b),(c,d)}" and ne: "(a,b) \ (c,d)" and inX: "(a,b) \ X" "(c,d) \ X" and "a \ c" apply (auto simp: nsets_def subset_iff eval_nat_numeral card_Suc_eq Set.doubleton_eq_iff) apply (metis nat_le_linear)+ done show "(\k<4. form U k) \ scheme U \ N" proof show "scheme U \ N" using inX N \ by (fastforce simp: scheme_def Ueq X_def) next consider "a < c" | "a = c \ b \ d" using \a \ c\ ne nat_less_le by blast then show "\k<4. form U k" proof cases case 1 have *: "a < b" "b \ c" "c < d" using inX by (auto simp: X_def) moreover have "\a < c; c < b; \ d < b\ \ b < d" using inX apply (clarsimp simp: X_def not_less) by (metis \ \_in_C imageE nat.inject nat_less_le) ultimately consider (k0) "a b c c b c da \ c\ by blast then show ?thesis by force next case k1 then have "form U 1" unfolding form_def form_split_def using Ueq \a \ c\ by blast then show ?thesis by force next case k2 then have "form U 2" unfolding form_def form_split_def using Ueq \a \ c\ by blast then show ?thesis by force qed next case 2 then have "form_split a b c d 3" by (auto simp: form_split_def) then show ?thesis using Ueq form_def leI by force qed qed qed qed text \Theorem 2.1 of Jean A. Larson, ibid.\ lemma Specker_aux: assumes "\ \ elts \" shows "partn_lst pair_less UU [\\2,\] 2" unfolding partn_lst_def proof clarsimp fix f assume f: "f \ [UU]\<^bsup>2\<^esup> \ {..: "?P0 \ ?P1" proof (rule disjCI) assume "\ ?P1" then have not1: "\M. \M \ UU; ordertype M pair_less = \\ \ \x\[M]\<^bsup>2\<^esup>. f x \ Suc 0" by auto obtain m where m: "\ = ord_of_nat m" using assms elts_\ by auto then have f_eq_0: "M \ [UU]\<^bsup>m\<^esup> \ \x\[M]\<^bsup>2\<^esup>. f x = 0" for M using not1 [of M] finite_ordertype_eq_card [of M pair_less m] f apply (clarsimp simp: nsets_def eval_nat_numeral Pi_def) - by (meson less_Suc0 not_less_less_Suc_eq subset_trans) + by (metis less_Suc0 less_antisym order.trans finite_subset) obtain N js where "infinite N" and N: "\k u. \k < 4; u \ [UU]\<^bsup>2\<^esup>; form u k; scheme u \ N\ \ f u = js!k" using f lemma_2_3 by blast obtain M0 where M0: "M0 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M0]\<^bsup>2\<^esup> \ form u 0" "\u. u \ [M0]\<^bsup>2\<^esup> \ scheme u \ N" by (rule lemma_2_4 [OF \infinite N\]) auto obtain M1 where M1: "M1 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M1]\<^bsup>2\<^esup> \ form u 1" "\u. u \ [M1]\<^bsup>2\<^esup> \ scheme u \ N" by (rule lemma_2_4 [OF \infinite N\]) auto obtain M2 where M2: "M2 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M2]\<^bsup>2\<^esup> \ form u 2" "\u. u \ [M2]\<^bsup>2\<^esup> \ scheme u \ N" by (rule lemma_2_4 [OF \infinite N\]) auto obtain M3 where M3: "M3 \ [UU]\<^bsup>m\<^esup>" "\u. u \ [M3]\<^bsup>2\<^esup> \ form u 3" "\u. u \ [M3]\<^bsup>2\<^esup> \ scheme u \ N" by (rule lemma_2_4 [OF \infinite N\]) auto have "js!0 = 0" using N [of 0 ] M0 f_eq_0 [of M0] by (force simp: nsets_def eval_nat_numeral) moreover have "js!1 = 0" using N [of 1] M1 f_eq_0 [of M1] by (force simp: nsets_def eval_nat_numeral) moreover have "js!2 = 0" using N [of 2 ] M2 f_eq_0 [of M2] by (force simp: nsets_def eval_nat_numeral) moreover have "js!3 = 0" using N [of 3 ] M3 f_eq_0 [of M3] by (force simp: nsets_def eval_nat_numeral) ultimately have js0: "js!k = 0" if "k < 4" for k using that by (auto simp: eval_nat_numeral less_Suc_eq) obtain X where "X \ UU" and otX: "ordertype X pair_less = \\2" and X: "\u. u \ [X]\<^bsup>2\<^esup> \ (\k<4. form u k) \ scheme u \ N" using \infinite N\ lemma_2_5 by auto moreover have "f ` [X]\<^bsup>2\<^esup> \ {0}" proof (clarsimp simp: image_subset_iff) fix u assume u: "u \ [X]\<^bsup>2\<^esup>" then have u_UU2: "u \ [UU]\<^bsup>2\<^esup>" using \X \ UU\ nsets_mono by blast show "f u = 0" using X u N [OF _ u_UU2] js0 by auto qed ultimately show "\X \ UU. ordertype X pair_less = \\2 \ f ` [X]\<^bsup>2\<^esup> \ {0}" by blast qed then show "\iH\UU. ordertype H pair_less = [\\2, \] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" proof show "?P0 \ ?thesis" by (metis nth_Cons_0 numeral_2_eq_2 pos2) show "?P1 \ ?thesis" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc) qed qed theorem Specker: "\ \ elts \ \ partn_lst_VWF (\\2) [\\2,\] 2" using partn_lst_imp_partn_lst_VWF_eq [OF Specker_aux] ordertype_UU_\2 wf_pair_less by blast end diff --git a/thys/Rep_Fin_Groups/Rep_Fin_Groups.thy b/thys/Rep_Fin_Groups/Rep_Fin_Groups.thy --- a/thys/Rep_Fin_Groups/Rep_Fin_Groups.thy +++ b/thys/Rep_Fin_Groups/Rep_Fin_Groups.thy @@ -1,9513 +1,9513 @@ theory Rep_Fin_Groups imports "HOL-Library.Function_Algebras" "HOL-Library.Set_Algebras" "HOL-Computational_Algebra.Polynomial" begin section \Preliminaries\ text \ In this section, we establish some basic facts about logic, sets, and functions that are not available in the HOL library. As well, we develop some theory for almost-everywhere-zero functions in preparation of the definition of the group ring. \ subsection \Logic\ lemma conjcases [case_names BothTrue OneTrue OtherTrue BothFalse] : assumes BothTrue: "P \ Q \ R" and OneTrue: "P \ \Q \ R" and OtherTrue: "\P \ Q \ R" and BothFalse: "\P \ \Q \ R" shows "R" using assms by fast subsection \Sets\ lemma empty_set_diff_single : "A - {x} = {} \ A = {} \ A = {x}" by auto lemma seteqI : "(\a. a \ A \ a \ B) \ (\b. b \ B \ b \ A) \ A = B" using subset_antisym subsetI by fast lemma prod_ballI : "(\a b. (a,b) \ AxB \ P a b) \ \(a,b)\AxB. P a b" by fast lemma good_card_imp_finite : "of_nat (card A) \ (0::'a::semiring_1) \ finite A" using card_ge_0_finite[of A] by fastforce subsection \Lists\ subsubsection \\zip\\ lemma zip_truncate_left : "zip xs ys = zip (take (length ys) xs) ys" by (induct xs ys rule:list_induct2') auto lemma zip_truncate_right : "zip xs ys = zip xs (take (length xs) ys)" by (induct xs ys rule:list_induct2') auto text \ Lemmas \zip_append1\ and \zip_append2\ in theory @{theory HOL.List} have unnecessary \take (length _)\ in them. Here are replacements. \ lemma zip_append_left : "zip (xs@ys) zs = zip xs zs @ zip ys (drop (length xs) zs)" using zip_append1 zip_truncate_right[of xs zs] by simp lemma zip_append_right : "zip xs (ys@zs) = zip xs ys @ zip (drop (length ys) xs) zs" using zip_append2 zip_truncate_left[of xs ys] by simp lemma length_concat_map_split_zip : "length [f x y. (x,y)\zip xs ys] = min (length xs) (length ys)" by (induct xs ys rule: list_induct2') auto lemma concat_map_split_eq_map_split_zip : "[f x y. (x,y)\zip xs ys] = map (case_prod f) (zip xs ys)" by (induct xs ys rule: list_induct2') auto lemma set_zip_map2 : "(a,z) \ set (zip xs (map f ys)) \ \b. (a,b) \ set (zip xs ys) \ z = f b" by (induct xs ys rule: list_induct2') auto subsubsection \\concat\\ lemma concat_eq : "list_all2 (\xs ys. length xs = length ys) xss yss \ concat xss = concat yss \ xss = yss" by (induct xss yss rule: list_all2_induct) auto lemma match_concat : fixes bss :: "'b list list" defines eq_len: "eq_len \ \xs ys. length xs = length ys" shows "\as::'a list. length as = length (concat bss) \ (\css::'a list list. as = concat css \ list_all2 eq_len css bss)" proof (induct bss) from eq_len show "\as. length as = length (concat []) \ (\css. as = concat css \ list_all2 eq_len css [])" by simp next fix fs :: "'b list" and fss :: "'b list list" assume prevcase: "\as. length as = length (concat fss) \ (\css. as = concat css \ list_all2 eq_len css fss)" have "\as. length as = length (concat (fs # fss)) \ (\css. as = concat css \ list_all2 eq_len css (fs # fss))" proof fix as :: "'a list" assume as: "length as = length (concat (fs#fss))" define xs ys where "xs = take (length fs) as" and "ys = drop (length fs) as" define gss where "gss = (SOME css. ys = concat css \ list_all2 eq_len css fss)" define hss where "hss = xs # gss" with xs_def ys_def as gss_def eq_len prevcase show "as = concat hss \ list_all2 eq_len hss (fs#fss)" using someI_ex[of "\css. ys = concat css \ list_all2 eq_len css fss"] by auto qed thus "\as. length as = length (concat (fs # fss)) \ (\css. as = concat css \ list_all2 eq_len css (fs # fss))" by fast qed subsubsection \\strip_while\\ lemma strip_while_0_nnil : "as \ [] \ set as \ 0 \ strip_while ((=) 0) as \ []" by (induct as rule: rev_nonempty_induct) auto subsubsection \\sum_list\\ lemma const_sum_list : "\x \ set xs. f x = a \ sum_list (map f xs) = a * (length xs)" by (induct xs) auto lemma sum_list_prod_cong : "\(x,y) \ set xys. f x y = g x y \ (\(x,y)\xys. f x y) = (\(x,y)\xys. g x y)" using arg_cong[of "map (case_prod f) xys" "map (case_prod g) xys" sum_list] by fastforce lemma sum_list_prod_map2 : "(\(a,y)\zip as (map f bs). g a y) = (\(a,b)\zip as bs. g a (f b))" by (induct as bs rule: list_induct2') auto lemma sum_list_fun_apply : "(\x\xs. f x) y = (\x\xs. f x y)" by (induct xs) auto lemma sum_list_prod_fun_apply : "(\(x,y)\xys. f x y) z = (\(x,y)\xys. f x y z)" by (induct xys) auto lemma (in comm_monoid_add) sum_list_plus : "length xs = length ys \ sum_list xs + sum_list ys = sum_list [a+b. (a,b)\zip xs ys]" proof (induct xs ys rule: list_induct2) case Cons thus ?case by (simp add: algebra_simps) qed simp lemma sum_list_const_mult_prod : fixes f :: "'a \ 'b \ 'r::semiring_0" shows "r * (\(x,y)\xys. f x y) = (\(x,y)\xys. r * (f x y))" using sum_list_const_mult[of r "case_prod f"] prod.case_distrib[of "\x. r*x" f] by simp lemma sum_list_mult_const_prod : fixes f :: "'a \ 'b \ 'r::semiring_0" shows "(\(x,y)\xys. f x y) * r = (\(x,y)\xys. (f x y) * r)" using sum_list_mult_const[of "case_prod f" r] prod.case_distrib[of "\x. x*r" f] by simp lemma sum_list_update : fixes xs :: "'a::ab_group_add list" shows "n < length xs \ sum_list (xs[n := y]) = sum_list xs - xs!n + y" proof (induct xs arbitrary: n) case Cons thus ?case by (cases n) auto qed simp lemma sum_list_replicate0 : "sum_list (replicate n 0) = 0" by (induct n) auto subsubsection \\listset\\ lemma listset_ConsI : "x \ X \ xs \ listset Xs \ x#xs \ listset (X#Xs)" unfolding listset_def set_Cons_def by simp lemma listset_ConsD : "x#xs \ listset (A # As) \ x \ A \ xs \ listset As" unfolding listset_def set_Cons_def by auto lemma listset_Cons_conv : "xs \ listset (A # As) \ (\y ys. y \ A \ ys \ listset As \ xs = y#ys)" unfolding listset_def set_Cons_def by auto lemma listset_length : "xs \ listset Xs \ length xs = length Xs" using listset_ConsD unfolding listset_def set_Cons_def by (induct xs Xs rule: list_induct2') auto lemma set_sum_list_element : "x \ (\A\As. A) \ \as \ listset As. x = (\a\as. a)" proof (induct As arbitrary: x) case Nil hence "x = (\a\[]. a)" by simp moreover have "[] \ listset []" by simp ultimately show ?case by fast next case (Cons A As) from this obtain a as where a_as: "a \ A" "as \ listset As" "x = (\b\(a#as). b)" using set_plus_def[of A] by fastforce have "listset (A#As) = set_Cons A (listset As)" by simp with a_as(1,2) have "a#as \ listset (A#As)" unfolding set_Cons_def by fast with a_as(3) show "\bs\listset (A#As). x = (\b\bs. b)" by fast qed lemma set_sum_list_element_Cons : assumes "x \ (\X\(A#As). X)" shows "\a as. a\A \ as \ listset As \ x = a + (\b\as. b)" proof- from assms obtain xs where xs: "xs \ listset (A#As)" "x = (\b\xs. b)" using set_sum_list_element by fast from xs(1) obtain a as where "a \ A" "as \ listset As" "xs = a # as" using listset_Cons_conv by fast with xs(2) show ?thesis by auto qed lemma sum_list_listset : "as \ listset As \ sum_list as \ (\A\As. A)" proof- have "length as = length As \ as \ listset As \ sum_list as \ (\A\As. A)" proof (induct as As rule: list_induct2) case Nil show ?case by simp next case (Cons a as A As) thus ?case using listset_ConsD[of a] set_plus_def by auto qed thus "as \ listset As \ sum_list as \ (\A\As. A)" using listset_length by fast qed lemma listsetI_nth : "length xs = length Xs \ \n Xs!n \ xs \ listset Xs" proof (induct xs Xs rule: list_induct2) case Nil show ?case by simp next case (Cons x xs X Xs) thus "x#xs \ listset (X#Xs)" using listset_ConsI[of x X xs Xs] by fastforce qed lemma listsetD_nth : "xs \ listset Xs \ \n Xs!n" proof- have "length xs = length Xs \ xs \ listset Xs \ \n Xs!n" proof (induct xs Xs rule: list_induct2) case Nil show ?case by simp next case (Cons x xs X Xs) from Cons(3) have x_xs: "x \ X" "xs \ listset Xs" using listset_ConsD[of x] by auto with Cons(2) have 1: "(x#xs)!0 \ (X#Xs)!0" "\n Xs!n" by auto have "\n. n < length (x#xs) \ (x#xs)!n \ (X#Xs)!n" proof- fix n assume "n < length (x#xs)" with 1 show "(x#xs)!n \ (X#Xs)!n" by (cases n) auto qed thus "\n < length (x#xs). (x#xs)!n \ (X#Xs)!n" by fast qed thus "xs \ listset Xs \ \n Xs!n" using listset_length by fast qed lemma set_listset_el_subset : "xs \ listset Xs \ \X\set Xs. X \ A \ set xs \ A" proof- have "\ length xs = length Xs; xs \ listset Xs; \X\set Xs. X \ A \ \ set xs \ A" proof (induct xs Xs rule: list_induct2) case Cons thus ?case using listset_ConsD by force qed simp thus "xs \ listset Xs \ \X\set Xs. X \ A \ set xs \ A" using listset_length by fast qed subsection \Functions\ subsubsection \Miscellaneous facts\ lemma sum_fun_apply : "finite A \ (\a\A. f a) x = (\a\A. f a x)" by (induct set: finite) auto lemma sum_single_nonzero : "finite A \ (\x\A. \y\A. f x y = (if y = x then g x else 0)) \ (\x\A. sum (f x) A = g x)" proof (induct A rule: finite_induct) case (insert a A) show "\x\insert a A. sum (f x) (insert a A) = g x" proof fix x assume x: "x \ insert a A" show "sum (f x) (insert a A) = g x" proof (cases "x = a") case True moreover with insert(2,4) have "\y\A. f x y = 0" by simp ultimately show ?thesis using insert(1,2,4) by simp next case False with x insert show ?thesis by simp qed qed qed simp lemma distrib_comp_sum_right : "(T + T') \ S = (T \ S) + (T' \ S)" by auto subsubsection \Support of a function\ definition supp :: "('a \ 'b::zero) \ 'a set" where "supp f = {x. f x \ 0}" lemma suppI: "f x \ 0 \ x \ supp f" using supp_def by fast lemma suppI_contra: "x \ supp f \ f x = 0" using suppI by fast lemma suppD: "x \ supp f \ f x \ 0" using supp_def by fast lemma suppD_contra: "f x = 0 \ x \ supp f" using suppD by fast lemma zerofun_imp_empty_supp : "supp 0 = {}" unfolding supp_def by simp lemma supp_zerofun_subset_any : "supp 0 \ A" using zerofun_imp_empty_supp by fast lemma supp_sum_subset_union_supp : fixes f g :: "'a \ 'b::monoid_add" shows "supp (f + g) \ supp f \ supp g" unfolding supp_def by auto lemma supp_neg_eq_supp : fixes f :: "'a \ 'b::group_add" shows "supp (- f) = supp f" unfolding supp_def by auto lemma supp_diff_subset_union_supp : fixes f g :: "'a \ 'b::group_add" shows "supp (f - g) \ supp f \ supp g" unfolding supp_def by auto abbreviation restrict0 :: "('a\'b::zero) \ 'a set \ ('a\'b)" (infix "\" 70) where "restrict0 f A \ (\a. if a \ A then f a else 0)" lemma supp_restrict0 : "supp (f\A) \ A" proof- have "\a. a \ A \ a \ supp (f\A)" using suppD_contra[of "f\A"] by simp thus ?thesis by fast qed lemma bij_betw_restrict0 : "bij_betw f A B \ bij_betw (f \ A) A B" using bij_betw_imp_inj_on bij_betw_imp_surj_on unfolding bij_betw_def inj_on_def by auto subsubsection \Convolution\ definition convolution :: "('a::group_add \ 'b::{comm_monoid_add,times}) \ ('a\'b) \ ('a\'b)" where "convolution f g = (\x. \y|x - y \ supp f \ y \ supp g. (f (x - y)) * g y)" \ \More often than not, this definition will be used in the case that @{text "'b"} is of class @{text "mult_zero"}, in which case the conditions @{term "x - y \ supp f"} and @{term "y \ supp g"} are obviously mathematically unnecessary. However, they also serve to ensure that the sum is taken over a finite set in the case that at least one of @{term f} and @{term g} is almost everywhere zero.\ lemma convolution_zero : fixes f g :: "'a::group_add \ 'b::{comm_monoid_add,mult_zero}" shows "f = 0 \ g = 0 \ convolution f g = 0" unfolding convolution_def by auto lemma convolution_symm : fixes f g :: "'a::group_add \ 'b::{comm_monoid_add,times}" shows "convolution f g = (\x. \y|y \ supp f \ -y + x \ supp g. (f y) * g (-y + x))" proof fix x::'a define c1 c2 i S1 S2 where "c1 y = (f (x - y)) * g y" and "c2 y = (f y) * g (-y + x)" and "i y = -y + x" and "S1 = {y. x - y \ supp f \ y \ supp g}" and "S2 = {y. y \ supp f \ -y + x \ supp g}" for y have "inj_on i S2" unfolding inj_on_def using i_def by simp hence "(\y\(i ` S2). c1 y) = (\y\S2. (c1 \ i) y)" using sum.reindex by fast moreover have S1_iS2: "S1 = i ` S2" proof (rule seteqI) fix y assume y_S1: "y \ S1" define z where "z = x - y" hence y_eq: "-z + x = y" by (auto simp add: algebra_simps) hence "-z + x \ supp g" using y_S1 S1_def by fast moreover have "z \ supp f" using z_def y_S1 S1_def by fast ultimately have "z \ S2" using S2_def by fast moreover have "y = i z" using i_def [abs_def] y_eq by fast ultimately show "y \ i ` S2" by fast next fix y assume "y \ i ` S2" from this obtain z where z_S2: "z \ S2" and y_eq: "y = -z + x" using i_def by fast from y_eq have "x - y = z" by (auto simp add: algebra_simps) hence "x - y \ supp f \ y \ supp g" using y_eq z_S2 S2_def by fastforce thus "y \ S1" using S1_def by fast qed ultimately have "(\y\S1. c1 y) = (\y\S2. (c1 \ i) y)" by fast with i_def c1_def c2_def have "(\y\S1. c1 y) = (\y\S2. c2 y)" using diff_add_eq_diff_diff_swap[of x _ x] by simp thus "convolution f g x = (\y|y \ supp f \ -y + x \ supp g. (f y) * g (-y + x))" unfolding S1_def c1_def S2_def c2_def convolution_def by fast qed lemma supp_convolution_subset_sum_supp : fixes f g :: "'a::group_add \ 'b::{comm_monoid_add,times}" shows "supp (convolution f g) \ supp f + supp g" proof- define SS where "SS x = {y. x-y \ supp f \ y \ supp g}" for x have "convolution f g = (\x. sum (\y. (f (x - y)) * g y) (SS x))" unfolding SS_def convolution_def by fast moreover have "\x. x \ supp f + supp g \ SS x = {}" proof- have "\x. SS x \ {} \ x \ supp f + supp g" proof- fix x::'a assume "SS x \ {}" from this obtain y where "x - y \ supp f" and y_G: "y \ supp g" using SS_def by fast from this obtain z where z_F: "z \ supp f" and z_eq: "x - y = z" by fast from z_eq have "x = z + y" using diff_eq_eq by fast with z_F y_G show "x \ supp f + supp g" by fast qed thus "\x. x \ supp f + supp g \ SS x = {}" by fast qed ultimately have "\x. x \ supp f + supp g \ convolution f g x = sum (\y. (f (x - y)) * g y) {}" by simp hence "\x. x \ supp f + supp g \ convolution f g x = 0" using sum.empty by simp thus ?thesis unfolding supp_def by fast qed subsection \Almost-everywhere-zero functions\ subsubsection \Definition and basic properties\ definition "aezfun_set = {f::'a\'b::zero. finite (supp f)}" lemma aezfun_setD: "f \ aezfun_set \ finite (supp f)" unfolding aezfun_set_def by fast lemma aezfun_setI: "finite (supp f) \ f \ aezfun_set" unfolding aezfun_set_def by fast lemma zerofun_is_aezfun : "0 \ aezfun_set" unfolding supp_def aezfun_set_def by auto lemma sum_of_aezfun_is_aezfun : fixes f g :: "'a\'b::monoid_add" shows "f \ aezfun_set \ g \ aezfun_set \ f + g \ aezfun_set" using supp_sum_subset_union_supp[of f g] finite_subset[of _ "supp f \ supp g"] unfolding aezfun_set_def by fastforce lemma neg_of_aezfun_is_aezfun : fixes f :: "'a\'b::group_add" shows "f \ aezfun_set \ - f \ aezfun_set" using supp_neg_eq_supp[of f] unfolding aezfun_set_def by simp lemma diff_of_aezfun_is_aezfun : fixes f g :: "'a\'b::group_add" shows "f \ aezfun_set \ g \ aezfun_set \ f - g \ aezfun_set" using supp_diff_subset_union_supp[of f g] finite_subset[of _ "supp f \ supp g"] unfolding aezfun_set_def by fastforce lemma restrict_and_extend0_aezfun_is_aezfun : assumes "f \ aezfun_set" shows "f\A \ aezfun_set" proof (rule aezfun_setI) have "\a. a \ supp f \ A \ a \ supp (f\A)" proof- fix a assume "a \ supp f \ A" thus "a \ supp (f\A)" using suppI_contra[of a] suppD_contra[of "f\A" a] by (cases "a \ A") auto qed with assms show "finite (supp (f\A))" - using aezfun_setD finite_subset[of "supp (f\A)"] by auto + using aezfun_setD finite_subset[of "supp (f\A)"] by blast qed subsubsection \Delta (impulse) functions\ text \The notation is set up in the order output-input so that later when these are used to define the group ring RG, it will be in order ring-element-group-element.\ definition deltafun :: "'b::zero \ 'a \ ('a \ 'b)" (infix "\" 70) where "b \ a = (\x. if x = a then b else 0)" lemma deltafun_apply_eq : "(b \ a) a = b" unfolding deltafun_def by simp lemma deltafun_apply_neq : "x \ a \ (b \ a) x = 0" unfolding deltafun_def by simp lemma deltafun0 : "0 \ a = 0" unfolding deltafun_def by auto lemma deltafun_plus : fixes b c :: "'b::monoid_add" shows "(b+c) \ a = (b \ a) + (c \ a)" unfolding deltafun_def by auto lemma supp_delta0fun : "supp (0 \ a) = {}" unfolding supp_def deltafun_def by simp lemma supp_deltafun : "b \ 0 \ supp (b \ a) = {a}" unfolding supp_def deltafun_def by simp lemma deltafun_is_aezfun : "b \ a \ aezfun_set" proof (cases "b = 0") case True hence "supp (b \ a) = {}" using supp_delta0fun[of a] by fast thus ?thesis unfolding aezfun_set_def by simp next case False thus ?thesis using supp_deltafun[of b a] unfolding aezfun_set_def by simp qed lemma aezfun_common_supp_spanning_set' : "finite A \ \as. distinct as \ set as = A \ ( \f::'a \ 'b::semiring_1. supp f \ A \ (\bs. length bs = length as \ f = (\(b,a)\zip bs as. b \ a)) )" proof (induct rule: finite_induct) case empty show ?case unfolding supp_def by auto next case (insert a A) from insert(3) obtain as where as: "distinct as" "set as = A" "\f::'a \ 'b. supp f \ A \ \bs. length bs = length as \ f = (\(b,a)\zip bs as. b \ a)" by fast from as(1,2) insert(2) have "distinct (a#as)" "set (a#as) = insert a A" by auto moreover have "\f::'a \ 'b::semiring_1. supp f \ insert a A \ (\bs. length bs = length (a#as) \ f = (\(b,a)\zip bs (a#as). b \ a))" proof- fix f :: "'a \ 'b" assume supp_f : "supp f \ insert a A" define g where "g x = (if x = a then 0 else f x)" for x have "supp g \ A" proof fix x assume x: "x \ supp g" with x supp_f g_def have "x \ insert a A" unfolding supp_def by auto moreover from x g_def have "x \ a" unfolding supp_def by auto ultimately show "x \ A" by fast qed with as(3) obtain bs where bs: "length bs = length as" "g = (\(b,a)\zip bs as. b \ a)" by fast from bs(1) have "length ((f a) # bs) = length (a#as)" by auto moreover from g_def bs(2) have "f = (\(b,a)\zip ((f a) # bs) (a#as). b \ a)" using deltafun_apply_eq[of "f a" a] deltafun_apply_neq[of _ a "f a"] by (cases) auto ultimately show "\bs. length bs = length (a#as) \ f = (\(b,a)\zip bs (a#as). b \ a)" by fast qed ultimately show ?case by fast qed subsubsection \Convolution of almost-everywhere-zero functions\ lemma convolution_eq_sum_over_supp_right : fixes g f :: "'a::group_add \ 'b::{comm_monoid_add,mult_zero}" assumes "g \ aezfun_set" shows "convolution f g = (\x. \y\supp g. (f (x - y)) * g y )" proof fix x::'a define SS where "SS = {y. x - y \ supp f \ y \ supp g}" have "finite (supp g)" using assms unfolding aezfun_set_def by fast moreover have "SS \ supp g" unfolding SS_def by fast moreover have "\y. y \ supp g - SS \ (f (x - y)) * g y = 0" using SS_def unfolding supp_def by auto ultimately show "convolution f g x = (\y\supp g. (f (x - y)) * g y )" unfolding convolution_def using SS_def sum.mono_neutral_left[of "supp g" SS "\y. (f (x - y)) * g y"] by fast qed lemma convolution_symm_eq_sum_over_supp_left : fixes f g :: "'a::group_add \ 'b::{comm_monoid_add,mult_zero}" assumes "f \ aezfun_set" shows "convolution f g = (\x. \y\supp f. (f y) * g (-y + x))" proof fix x::'a define SS where "SS = {y. y \ supp f \ -y + x \ supp g}" have "finite (supp f)" using assms unfolding aezfun_set_def by fast moreover have "SS \ supp f" using SS_def by fast moreover have "\y. y \ supp f - SS \ (f y) * g (-y + x) = 0" using SS_def unfolding supp_def by auto ultimately have "(\y\SS. (f y) * g (-y + x)) = (\y\supp f. (f y) * g (-y + x) )" unfolding convolution_def using SS_def sum.mono_neutral_left[of "supp f" SS "\y. (f y) * g (-y + x)"] by fast thus "convolution f g x = (\y\supp f. (f y) * g (-y + x) )" using SS_def convolution_symm[of f g] by simp qed lemma convolution_delta_left : fixes b :: "'b::{comm_monoid_add,mult_zero}" and a :: "'a::group_add" and f :: "'a \ 'b" shows "convolution (b \ a) f = (\x. b * f (-a + x))" proof (cases "b = 0") case True moreover have "convolution (b \ a) f = 0" proof- from True have "convolution (b \ a) f = convolution 0 f" using deltafun0[of a] arg_cong[of "0 \ a" "0::'a\'b"] by (simp add: \0 \ a = 0\ \b = 0\) thus ?thesis using convolution_zero by auto qed ultimately show ?thesis by auto next case False thus ?thesis using deltafun_is_aezfun[of b a] convolution_symm_eq_sum_over_supp_left supp_deltafun[of b a] deltafun_apply_eq[of b a] by fastforce qed lemma convolution_delta_right : fixes b :: "'b::{comm_monoid_add,mult_zero}" and f :: "'a::group_add \ 'b" and a::'a shows "convolution f (b \ a) = (\x. f (x - a) * b)" proof (cases "b = 0") case True moreover have "convolution f (b \ a) = 0" proof- from True have "convolution f (b \ a) = convolution f 0" using deltafun0[of a] arg_cong[of "0 \ a" "0::'a\'b"] by (simp add: \0 \ a = 0\) thus ?thesis using convolution_zero by auto qed ultimately show ?thesis by auto next case False thus ?thesis using deltafun_is_aezfun[of b a] convolution_eq_sum_over_supp_right supp_deltafun[of b a] deltafun_apply_eq[of b a] by fastforce qed lemma convolution_delta_delta : fixes b1 b2 :: "'b::{comm_monoid_add,mult_zero}" and a1 a2 :: "'a::group_add" shows "convolution (b1 \ a1) (b2 \ a2) = (b1 * b2) \ (a1 + a2)" proof fix x::'a have 1: "convolution (b1 \ a1) (b2 \ a2) x = (b1 \ a1) (x - a2) * b2" using convolution_delta_right[of "b1 \ a1"] by simp show "convolution (b1 \ a1) (b2 \ a2) x = ((b1 * b2) \ (a1 + a2)) x" proof (cases "x = a1 + a2") case True hence "x - a2 = a1" by (simp add: algebra_simps) with 1 have "convolution (b1 \ a1) (b2 \ a2) x = b1 * b2" using deltafun_apply_eq[of b1 a1] by simp with True show ?thesis using deltafun_apply_eq[of "b1 * b2" "a1 + a2"] by simp next case False hence "x - a2 \ a1" by (simp add: algebra_simps) with 1 have "convolution (b1 \ a1) (b2 \ a2) x = 0" using deltafun_apply_neq[of "x - a2" a1 b1] by simp with False show ?thesis using deltafun_apply_neq by simp qed qed lemma convolution_of_aezfun_is_aezfun : fixes f g :: "'a::group_add \ 'b::{comm_monoid_add,times}" shows "f \ aezfun_set \ g \ aezfun_set \ convolution f g \ aezfun_set" using supp_convolution_subset_sum_supp[of f g] finite_set_plus[of "supp f" "supp g"] finite_subset unfolding aezfun_set_def by fastforce lemma convolution_assoc : fixes f h g :: "'a::group_add \ 'b::semiring_0" assumes f_aez: "f \ aezfun_set" and h_aez: "h \ aezfun_set" shows "convolution (convolution f g) h = convolution f (convolution g h)" proof define fg gh where "fg = convolution f g" and "gh = convolution g h" fix x::'a have "convolution fg h x = (\y\supp f. (\z\supp h. f y * g (-y + x - z) * h z) )" proof- have "convolution fg h x = (\z\supp h. fg (x - z) * h z )" using h_aez convolution_eq_sum_over_supp_right[of h fg] by simp moreover have "\z. fg (x - z) * h z = (\y\supp f. f y * g (-y + x - z) * h z)" proof- fix z::'a have "fg (x - z) = (\y\supp f. f y * g (-y + (x - z)) )" using fg_def f_aez convolution_symm_eq_sum_over_supp_left by fastforce hence "fg (x - z) * h z = (\y\supp f. f y * g (-y + (x - z)) * h z )" using sum_distrib_right by simp thus "fg (x - z) * h z = (\y\supp f. f y * g (-y + x - z) * h z )" by (simp add: algebra_simps) qed ultimately have "convolution fg h x = (\z\supp h. (\y\supp f. f y * g (-y + x - z) * h z) )" using sum.cong by simp thus ?thesis using sum.swap by simp qed moreover have "convolution f gh x = (\y\supp f. (\z\supp h. f y * g (-y + x - z) * h z) )" proof- have "convolution f gh x = (\y\supp f. f y * gh (-y + x) )" using f_aez convolution_symm_eq_sum_over_supp_left[of f gh] by simp moreover have "\y. f y * gh (-y + x) = (\z\supp h. f y * g (-y + x - z) * h z)" proof- fix y::'a have triple_cong: "\z. f y * (g (-y + x - z) * h z) = f y * g (-y + x - z) * h z" using mult.assoc[of "f y"] by simp have "gh (-y + x) = (\z\supp h. g (-y + x - z) * h z)" using gh_def h_aez convolution_eq_sum_over_supp_right by fastforce hence "f y * gh (-y + x) = (\z\supp h. f y * (g (-y + x - z) * h z))" using sum_distrib_left by simp also have "\ = (\z\supp h. f y * g (-y + x - z) * h z)" using triple_cong sum.cong by simp finally show "f y * gh (-y + x) = (\z\supp h. f y * g (-y + x - z) * h z)" by fast qed ultimately show ?thesis using sum.cong by simp qed ultimately show "convolution fg h x = convolution f gh x" by simp qed lemma convolution_distrib_left : fixes g h f :: "'a::group_add \ 'b::semiring_0" assumes "g \ aezfun_set" "h \ aezfun_set" shows "convolution f (g + h) = convolution f g + convolution f h" proof define gh GH where "gh = g + h" and "GH = supp g \ supp h" have fin_GH: "finite GH" using GH_def assms unfolding aezfun_set_def by fast have gh_aezfun: "gh \ aezfun_set" using gh_def assms sum_of_aezfun_is_aezfun by fast fix x::'a have zero_ext_g : "\y. y \ GH - supp g \ (f (x - y)) * g y = 0" and zero_ext_h : "\y. y \ GH - supp h \ (f (x - y)) * h y = 0" and zero_ext_gh: "\y. y \ GH - supp gh \ (f (x - y)) * gh y = 0" unfolding supp_def by auto have "convolution f gh x = (\y\supp gh. (f (x - y)) * gh y)" using assms gh_aezfun convolution_eq_sum_over_supp_right[of gh f] by simp also from gh_def GH_def have "\ = (\y\GH. (f (x - y)) * gh y)" using fin_GH supp_sum_subset_union_supp zero_ext_gh sum.mono_neutral_left[of GH "supp gh" "(\y. (f (x - y)) * gh y)"] by fast also from gh_def have "\ = (\y\GH. (f (x - y)) * g y) + (\y\GH. (f (x - y)) * h y)" using sum.distrib by (simp add: algebra_simps) finally show "convolution f gh x = (convolution f g + convolution f h) x" using assms GH_def fin_GH zero_ext_g zero_ext_h sum.mono_neutral_right[of GH "supp g" "(\y. (f (x - y)) * g y)"] sum.mono_neutral_right[of GH "supp h" "(\y. (f (x - y)) * h y)"] convolution_eq_sum_over_supp_right[of g f] convolution_eq_sum_over_supp_right[of h f] by fastforce qed lemma convolution_distrib_right : fixes f g h :: "'a::group_add \ 'b::semiring_0" assumes "f \ aezfun_set" "g \ aezfun_set" shows "convolution (f + g) h = convolution f h + convolution g h" proof define fg FG where "fg = f + g" and "FG = supp f \ supp g" have fin_FG: "finite FG" using FG_def assms unfolding aezfun_set_def by fast have fg_aezfun: "fg \ aezfun_set" using fg_def assms sum_of_aezfun_is_aezfun by fast fix x::'a have zero_ext_f : "\y. y \ FG - supp f \ (f y) * h (-y + x) = 0" and zero_ext_g : "\y. y \ FG - supp g \ (g y) * h (-y + x) = 0" and zero_ext_fg: "\y. y \ FG - supp fg \ (fg y) * h (-y + x) = 0" unfolding supp_def by auto from assms have "convolution fg h x = (\y\supp fg. (fg y) * h (-y + x))" using fg_aezfun convolution_symm_eq_sum_over_supp_left[of fg h] by simp also from fg_def FG_def have "\ = (\y\FG. (fg y) * h (-y + x))" using fin_FG supp_sum_subset_union_supp zero_ext_fg sum.mono_neutral_left[of FG "supp fg" "(\y. (fg y) * h (-y + x))"] by fast also from fg_def have "\ = (\y\FG. (f y) * h (-y + x)) + (\y\FG. (g y) * h (-y + x))" using sum.distrib by (simp add: algebra_simps) finally show "convolution fg h x = (convolution f h + convolution g h) x" using assms FG_def fin_FG zero_ext_f zero_ext_g sum.mono_neutral_right[of FG "supp f" "(\y. (f y) * h (-y + x))"] sum.mono_neutral_right[of FG "supp g" "(\y. (g y) * h (-y + x))"] convolution_symm_eq_sum_over_supp_left[of f h] convolution_symm_eq_sum_over_supp_left[of g h] by fastforce qed subsubsection \Type definition, instantiations, and instances\ typedef (overloaded) ('a::zero,'b) aezfun = "aezfun_set :: ('b\'a) set" morphisms aezfun Abs_aezfun using zerofun_is_aezfun by fast setup_lifting type_definition_aezfun lemma aezfun_finite_supp : "finite (supp (aezfun a))" using aezfun.aezfun unfolding aezfun_set_def by fast lemma aezfun_transfer : "aezfun a = aezfun b \ a = b" by transfer fast instantiation aezfun :: (zero, type) zero begin lift_definition zero_aezfun :: "('a,'b) aezfun" is "0::'b\'a" using zerofun_is_aezfun by fast instance .. end lemma zero_aezfun_transfer : "Abs_aezfun ((0::'b::zero) \ (0::'a::zero)) = 0" proof- define zb za where "zb = (0::'b)" and "za = (0::'a)" hence "zb \ za = 0" using deltafun0[of za] by fast moreover have "aezfun 0 = 0" using zero_aezfun.rep_eq by fast ultimately have "zb \ za = aezfun 0" by simp with zb_def za_def show ?thesis using aezfun_inverse by simp qed lemma zero_aezfun_apply [simp]: "aezfun 0 x = 0" by transfer simp instantiation aezfun :: (monoid_add, type) plus begin lift_definition plus_aezfun :: "('a, 'b) aezfun \ ('a, 'b) aezfun \ ('a, 'b) aezfun" is "\f g. f + g" using sum_of_aezfun_is_aezfun by auto instance .. end lemma plus_aezfun_apply [simp]: "aezfun (a+b) x = aezfun a x + aezfun b x" by transfer simp instance aezfun :: (monoid_add, type) semigroup_add proof fix a b c :: "('a, 'b) aezfun" have "aezfun (a + b + c) = aezfun (a + (b + c))" proof fix x::'b show "aezfun (a + b + c) x = aezfun (a + (b + c)) x" using add.assoc[of "aezfun a x"] by simp qed thus "a + b + c = a + (b + c)" by transfer fast qed instance aezfun :: (monoid_add, type) monoid_add proof fix a b c :: "('a, 'b) aezfun" show "0 + a = a" by transfer simp show "a + 0 = a" by transfer simp qed lemma sum_list_aezfun_apply [simp] : "aezfun (sum_list as) x = (\a\as. aezfun a x)" by (induct as) auto lemma sum_list_map_aezfun_apply [simp] : "aezfun (\a\as. f a) x = (\a\as. aezfun (f a) x)" by (induct as) auto lemma sum_list_map_aezfun [simp] : "aezfun (\a\as. f a) = (\a\as. aezfun (f a))" using sum_list_map_aezfun_apply[of f] sum_list_fun_apply[of "aezfun \ f"] by auto lemma sum_list_prod_map_aezfun_apply : "aezfun (\(x,y)\xys. f x y) a = (\(x,y)\xys. aezfun (f x y) a)" by (induct xys) auto lemma sum_list_prod_map_aezfun : "aezfun (\(x,y)\xys. f x y) = (\(x,y)\xys. aezfun (f x y))" using sum_list_prod_map_aezfun_apply[of f] sum_list_prod_fun_apply[of "\y z. aezfun (f y z)"] by auto instance aezfun :: (comm_monoid_add, type) comm_monoid_add proof fix a b :: "('a, 'b) aezfun" have "aezfun (a + b) = aezfun (b + a)" proof fix x::'b show "aezfun (a + b) x = aezfun (b + a) x" using add.commute[of "aezfun a x"] by simp qed thus "a + b = b + a" by transfer fast show "0 + a = a" by simp qed lemma sum_aezfun_apply [simp] : "finite A \ aezfun (\A) x = (\a\A. aezfun a x)" by (induct set: finite) auto instantiation aezfun :: (group_add, type) minus begin lift_definition minus_aezfun :: "('a, 'b) aezfun \ ('a, 'b) aezfun \ ('a, 'b) aezfun" is "\f g. f - g" using diff_of_aezfun_is_aezfun by fast instance .. end lemma minus_aezfun_apply [simp]: "aezfun (a-b) x = aezfun a x - aezfun b x" by transfer simp instantiation aezfun :: (group_add, type) uminus begin lift_definition uminus_aezfun :: "('a, 'b) aezfun \ ('a, 'b) aezfun" is "\f. - f" using neg_of_aezfun_is_aezfun by fast instance .. end lemma uminus_aezfun_apply [simp]: "aezfun (-a) x = - aezfun a x" by transfer simp lemma aezfun_left_minus [simp] : fixes a :: "('a::group_add, 'b) aezfun" shows "- a + a = 0" by transfer simp lemma aezfun_diff_minus [simp] : fixes a b :: "('a::group_add, 'b) aezfun" shows "a - b = a + - b" by transfer auto instance aezfun :: (group_add, type) group_add proof fix a b :: "('a::group_add, 'b) aezfun" show "- a + a = 0" "a + - b = a - b" by auto qed instance aezfun :: (ab_group_add, type) ab_group_add proof fix a b :: "('a::ab_group_add, 'b) aezfun" show "- a + a = 0" by simp show "a - b = a + - b" using aezfun_diff_minus by fast qed instantiation aezfun :: ("{one,zero}", zero) one begin lift_definition one_aezfun :: "('a,'b) aezfun" is "1 \ 0" using deltafun_is_aezfun by fast instance .. end lemma one_aezfun_transfer : "Abs_aezfun (1 \ 0) = 1" proof- define z n where "z = (0::'b::zero)" and "n = (1::'a::{one,zero})" hence "aezfun 1 = n \ z" using one_aezfun.rep_eq by fast hence "Abs_aezfun (n \ z) = Abs_aezfun (aezfun 1)" by simp with z_def n_def show ?thesis using aezfun_inverse by simp qed lemma one_aezfun_apply [simp]: "aezfun 1 x = (1 \ 0) x" by transfer rule lemma one_aezfun_apply_eq [simp]: "aezfun 1 0 = 1" using deltafun_apply_eq by simp lemma one_aezfun_apply_neq [simp]: "x \ 0 \ aezfun 1 x = 0" using deltafun_apply_neq by simp instance aezfun :: (zero_neq_one, zero) zero_neq_one proof have "(0::'a) \ 1" "aezfun 0 0 = 0" "aezfun (1::('a,'b) aezfun) 0 = 1" using zero_neq_one one_aezfun_apply_eq by auto thus "(0::('a,'b) aezfun) \ 1" using zero_neq_one one_aezfun_apply_eq fun_eq_iff[of "aezfun (0::('a,'b) aezfun)" "aezfun 1"] by auto qed instantiation aezfun :: ("{comm_monoid_add,times}", group_add) times begin lift_definition times_aezfun :: "('a, 'b) aezfun \ ('a, 'b) aezfun \ ('a, 'b) aezfun" is "\ f g. convolution f g" using convolution_of_aezfun_is_aezfun by fast instance .. end lemma convolution_transfer : assumes "f \ aezfun_set" "g \ aezfun_set" shows "Abs_aezfun (convolution f g) = Abs_aezfun f * Abs_aezfun g" proof (rule aezfun_transfer) from assms have "aezfun (Abs_aezfun (convolution f g)) = convolution f g" using convolution_of_aezfun_is_aezfun Abs_aezfun_inverse by fast moreover from assms have "aezfun (Abs_aezfun f * Abs_aezfun g) = convolution f g" using times_aezfun.rep_eq[of "Abs_aezfun f"] Abs_aezfun_inverse[of f] Abs_aezfun_inverse[of g] by simp ultimately show "aezfun (Abs_aezfun (convolution f g)) = aezfun (Abs_aezfun f * Abs_aezfun g)" by simp qed instance aezfun :: ("{comm_monoid_add,mult_zero}", group_add) mult_zero proof fix a :: "('a, 'b) aezfun" show "0 * a = 0" using convolution_zero[of _ "aezfun a"] by transfer fast show "a * 0 = 0" using convolution_zero[of "aezfun a"] by transfer fast qed instance aezfun :: (semiring_0, group_add) semiring_0 proof fix a b c :: "('a, 'b) aezfun" show "a * b * c = a * (b * c)" using convolution_assoc[of "aezfun a" "aezfun c" "aezfun b"] by transfer show "(a + b) * c = a * c + b * c" using convolution_distrib_right[of "aezfun a" "aezfun b" "aezfun c"] by transfer show "a * (b + c) = a * b + a * c" using convolution_distrib_left[of "aezfun b" "aezfun c" "aezfun a"] by transfer qed instance aezfun :: (ring, group_add) ring .. instance aezfun :: ("{semiring_0,monoid_mult,zero_neq_one}", group_add) monoid_mult proof fix a :: "('a, 'b) aezfun" show "1 * a = a" proof- have "aezfun (1 * a) = convolution (1 \ 0) (aezfun a)" by transfer fast hence "aezfun (1 * a) = (aezfun a)" using one_neq_zero convolution_delta_left[of 1 0 "aezfun a"] minus_zero by simp thus "1 * a = a" by transfer qed show "a * 1 = a" proof- have "aezfun (a * 1) = convolution (aezfun a) (1 \ 0)" by transfer fast hence "aezfun (a * 1) = (aezfun a)" using one_neq_zero convolution_delta_right[of "aezfun a"] by simp thus ?thesis by transfer qed qed instance aezfun :: (ring_1, group_add) ring_1 .. subsubsection \Transfer facts\ abbreviation aezdeltafun :: "'b::zero \ 'a \ ('b,'a) aezfun" (infix "\\" 70) where "b \\ a \ Abs_aezfun (b \ a)" lemma aezdeltafun : "aezfun (b \\ a) = b \ a" using deltafun_is_aezfun[of b a] Abs_aezfun_inverse by fast lemma aezdeltafun_plus : "(b+c) \\ a = (b \\ a) + (c \\ a)" using aezdeltafun[of "b+c" a] deltafun_plus aezdeltafun[of b a] aezdeltafun[of c a] plus_aezfun.rep_eq[of "b \\ a"] aezfun_transfer[of "(b+c) \\ a" "(b \\ a) + (c \\ a)"] by fastforce lemma times_aezdeltafun_aezdeltafun : fixes b1 b2 :: "'b::{comm_monoid_add,mult_zero}" shows "(b1 \\ a1) * (b2 \\ a2) = (b1 * b2) \\ (a1 + a2)" using deltafun_is_aezfun convolution_transfer[of "b1 \ a1", THEN sym] convolution_delta_delta[of b1 a1 b2 a2] by fastforce lemma aezfun_restrict_and_extend0 : "(aezfun x)\A \ aezfun_set" using aezfun.aezfun restrict_and_extend0_aezfun_is_aezfun[of "aezfun x"] by fast lemma aezdeltafun_decomp : fixes b :: "'b::semiring_1" shows "b \\ a = (b \\ 0) * (1 \\ a)" using convolution_delta_delta[of b 0 1 a] deltafun_is_aezfun[of b 0] deltafun_is_aezfun[of 1 a] convolution_transfer by fastforce lemma aezdeltafun_decomp' : fixes b :: "'b::semiring_1" shows "b \\ a = (1 \\ a) * (b \\ 0)" using convolution_delta_delta[of 1 a b 0] deltafun_is_aezfun[of b 0] deltafun_is_aezfun[of 1 a] convolution_transfer by fastforce lemma supp_aezfun1 : "supp ( aezfun ( 1 :: ('a::zero_neq_one,'b::zero) aezfun ) ) = 0" using supp_deltafun[of "1::'a" "0::'b"] by transfer simp lemma supp_aezfun_diff : "supp (aezfun (x - y)) \ supp (aezfun x) \ supp (aezfun y)" proof- have "supp (aezfun (x - y)) = supp ( (aezfun x) - (aezfun y) )" by transfer fast thus ?thesis using supp_diff_subset_union_supp by fast qed lemma supp_aezfun_times : "supp (aezfun (x * y)) \ supp (aezfun x) + supp (aezfun y)" proof- have "supp (aezfun (x * y)) = supp (convolution (aezfun x) (aezfun y))" by transfer fast thus ?thesis using supp_convolution_subset_sum_supp by fast qed subsubsection \Almost-everywhere-zero functions with constrained support\ text \The name of the next definition anticipates \aezfun_common_supp_spanning_set\ below.\ definition aezfun_setspan :: "'a set \ ('b::zero,'a) aezfun set" where "aezfun_setspan A = {x. supp (aezfun x) \ A}" lemma aezfun_setspanD : "x \ aezfun_setspan A \ supp (aezfun x) \ A" unfolding aezfun_setspan_def by fast lemma aezfun_setspanI : "supp (aezfun x) \ A \ x \ aezfun_setspan A" unfolding aezfun_setspan_def by fast lemma aezfun_common_supp_spanning_set : assumes "finite A" shows "\as. distinct as \ set as = A \ ( \x::('b::semiring_1,'a) aezfun \ aezfun_setspan A. \bs. length bs = length as \ x = (\(b,a)\zip bs as. b \\ a) )" proof- from assms aezfun_common_supp_spanning_set'[of A] obtain as where as: "distinct as" "set as = A" "\f::'a \ 'b. supp f \ A \ (\bs. length bs = length as \ f = (\(b,a)\zip bs as. b \ a))" by fast have "\x::('b,'a) aezfun. x \ aezfun_setspan A \ (\bs. length bs = length as \ x = (\(b,a)\zip bs as. b \\ a))" proof- fix x::"('b,'a) aezfun" assume "x \ aezfun_setspan A" with as(3) obtain bs where bs: "length bs = length as" "aezfun x = (\(b,a)\zip bs as. b \ a)" using aezfun_setspanD by fast have "\b a. (b,a) \ set (zip bs as) \ b \ a = aezfun (b \\ a)" proof- fix b a assume "(b,a) \ set (zip bs as)" show "b \ a = aezfun (b \\ a)" using aezdeltafun[of b a] by simp qed with bs show "\bs. length bs = length as \ x = (\(b,a)\zip bs as. b \\ a)" using sum_list_prod_cong[of "zip bs as" deltafun "\b a. aezfun (b \\ a)"] sum_list_prod_map_aezfun[of aezdeltafun "zip bs as"] aezfun_transfer[of x] by fastforce qed with as(1,2) show ?thesis by fast qed lemma aezfun_common_supp_spanning_set_decomp : fixes G :: "'g::group_add set" assumes "finite G" shows "\gs. distinct gs \ set gs = G \ ( \x::('r::semiring_1,'g) aezfun \ aezfun_setspan G. \rs. length rs = length gs \ x = (\(r,g)\zip rs gs. (r \\ 0) * (1 \\ g)) )" proof- from aezfun_common_supp_spanning_set[OF assms] obtain gs where gs: "distinct gs" "set gs = G" "\x::('r,'g) aezfun \ aezfun_setspan G. \rs. length rs = length gs \ x = (\(r,g)\zip rs gs. r \\ g)" by fast have "\x::('r,'g) aezfun. x \ aezfun_setspan G \ \rs. length rs = length gs \ x = (\(r,g)\zip rs gs. (r \\ 0) * (1 \\ g))" proof- fix x::"('r,'g) aezfun" assume "x \ aezfun_setspan G" with gs(3) obtain rs where "length rs = length gs" "x = (\(r,g)\zip rs gs. r \\ g)" using aezfun_setspanD by fast thus "\rs. length rs = length gs \ x = (\(r,g)\zip rs gs. (r \\ 0) * (1 \\ g))" using aezdeltafun_decomp sum_list_prod_cong[ of "zip rs gs" "\r g. r \\ g" "\r g. (r \\ 0) * (1 \\ g)" ] by auto qed with gs(1,2) show ?thesis by fast qed lemma aezfun_decomp_aezdeltafun : fixes c :: "('r::semiring_1,'a) aezfun" shows "\ras. set (map snd ras) = supp (aezfun c) \ c = (\(r,a)\ras. r \\ a)" proof- from aezfun_finite_supp[of c] obtain as where as: "set as = supp (aezfun c)" "\x::('r,'a) aezfun \ aezfun_setspan (supp (aezfun c)). \bs. length bs = length as \ x = (\(b,a)\zip bs as. b \\ a)" using aezfun_common_supp_spanning_set[of "supp (aezfun c)"] by fast from as(2) obtain bs where bs: "length bs = length as" "c = (\(b,a)\zip bs as. b \\ a)" using aezfun_setspanI[of c "supp (aezfun c)"] by fast from bs(1) as(1) have "set (map snd (zip bs as)) = supp (aezfun c)" by simp with bs(2) show ?thesis by fast qed lemma aezfun_setspan_el_decomp_aezdeltafun : fixes c :: "('r::semiring_1,'a) aezfun" shows "c \ aezfun_setspan A \ \ras. set (map snd ras) \ A \ c = (\(r,a)\ras. r \\ a)" using aezfun_setspanD aezfun_decomp_aezdeltafun by fast lemma aezdelta0fun_commutes' : fixes b1 b2 :: "'b::comm_semiring_1" shows "b1 \\ a * (b2 \\ 0) = b2 \\ 0 * (b1 \\ a)" using times_aezdeltafun_aezdeltafun[of b1 a] times_aezdeltafun_aezdeltafun[of b2 0 b1 a] by (simp add: algebra_simps) lemma aezdelta0fun_commutes : fixes b :: "'b::comm_semiring_1" shows "c * (b \\ 0) = b \\ 0 * c" proof- from aezfun_decomp_aezdeltafun obtain ras where c: "c = (\(r,a)\ras. r \\ a)" by fast thus ?thesis using sum_list_mult_const_prod[of "\r a. r \\ a" ras] aezdelta0fun_commutes' sum_list_prod_cong[of ras "\r a. r \\ a * (b \\ 0)" "\r a. b \\ 0 * (r \\ a)"] sum_list_const_mult_prod[of "b \\ 0" "\r a. r \\ a" ras] by auto qed text \ The following definition constrains the support of arbitrary almost-everywhere-zero functions, as a sort of projection onto a \aezfun_setspan\. \ definition aezfun_setspan_proj :: "'a set \ ('b::zero,'a) aezfun \ ('b::zero,'a) aezfun" where "aezfun_setspan_proj A x \ Abs_aezfun ((aezfun x)\A)" lemma aezfun_setspan_projD1 : "a \ A \ aezfun (aezfun_setspan_proj A x) a = aezfun x a" using aezfun_restrict_and_extend0[of A x] Abs_aezfun_inverse[of "(aezfun x)\A"] unfolding aezfun_setspan_proj_def by simp lemma aezfun_setspan_projD2 : "a \ A \ aezfun (aezfun_setspan_proj A x) a = 0" using aezfun_restrict_and_extend0[of A x] Abs_aezfun_inverse[of "(aezfun x)\A"] unfolding aezfun_setspan_proj_def by simp lemma aezfun_setspan_proj_in_setspan : "aezfun_setspan_proj A x \ aezfun_setspan A" using aezfun_setspan_projD2[of _ A] suppD_contra[of "aezfun (aezfun_setspan_proj A x)"] aezfun_setspanI[of "aezfun_setspan_proj A x" A] by auto lemma aezfun_setspan_proj_zero : "aezfun_setspan_proj A 0 = 0" proof- have "aezfun (aezfun_setspan_proj A 0) = aezfun 0" proof fix a show "aezfun (aezfun_setspan_proj A 0) a = aezfun 0 a" using aezfun_setspan_projD1[of a A 0] aezfun_setspan_projD2[of a A 0] by (cases "a\A") auto qed thus ?thesis using aezfun_transfer by fast qed lemma aezfun_setspan_proj_aezdeltafun : "aezfun_setspan_proj A (b \\ a) = (if a \ A then b \\ a else 0)" proof- have "aezfun (aezfun_setspan_proj A (b \\ a)) = aezfun (if a \ A then b \\ a else 0)" proof fix x show "aezfun (aezfun_setspan_proj A (b \\ a)) x = aezfun (if a \ A then b \\ a else 0) x" proof (cases "x \ A") case True thus ?thesis using aezfun_setspan_projD1[of x A "b \\ a"] aezdeltafun[of b a] deltafun_apply_neq[of x] by fastforce next case False hence "aezfun (aezfun_setspan_proj A (b \\ a)) x = 0" using aezfun_setspan_projD2[of x A] by simp moreover from False have "a \ A \ aezfun (if a \ A then b \\ a else 0) x = 0" using aezdeltafun[of b a] deltafun_apply_neq[of x a b] by auto ultimately show ?thesis by auto qed qed thus ?thesis using aezfun_transfer by fast qed lemma aezfun_setspan_proj_add : "aezfun_setspan_proj A (x+y) = aezfun_setspan_proj A x + aezfun_setspan_proj A y" proof- have "aezfun (aezfun_setspan_proj A (x+y)) = aezfun (aezfun_setspan_proj A x + aezfun_setspan_proj A y)" proof fix a show "aezfun (aezfun_setspan_proj A (x+y)) a = aezfun (aezfun_setspan_proj A x + aezfun_setspan_proj A y) a" using aezfun_setspan_projD1[of a A "x+y"] aezfun_setspan_projD2[of a A "x+y"] aezfun_setspan_projD1[of a A x] aezfun_setspan_projD1[of a A y] aezfun_setspan_projD2[of a A x] aezfun_setspan_projD2[of a A y] by (cases "a \ A") auto qed thus ?thesis using aezfun_transfer by auto qed lemma aezfun_setspan_proj_sum_list : "aezfun_setspan_proj A (\x\xs. f x) = (\x\xs. aezfun_setspan_proj A (f x))" proof (induct xs) case Nil show ?case using aezfun_setspan_proj_zero by simp next case (Cons x xs) thus ?case using aezfun_setspan_proj_add[of A "f x"] by simp qed lemma aezfun_setspan_proj_sum_list_prod : "aezfun_setspan_proj A (\(x,y)\xys. f x y) = (\(x,y)\xys. aezfun_setspan_proj A (f x y))" using aezfun_setspan_proj_sum_list[of A "\xy. case_prod f xy"] prod.case_distrib[of "aezfun_setspan_proj A" f] by simp subsection \Polynomials\ lemma nonzero_coeffs_nonzero_poly : "as \ [] \ set as \ 0 \ Poly as \ 0" using coeffs_Poly[of as] strip_while_0_nnil[of as] by fastforce lemma const_poly_nonzero_coeff : assumes "degree p = 0" "p \ 0" shows "coeff p 0 \ 0" proof assume z: "coeff p 0 = 0" have "\n. coeff p n = 0" proof- fix n from z assms show "coeff p n = 0" using coeff_eq_0[of p] by (cases "n = 0") auto qed with assms(2) show False using poly_eqI[of p 0] by simp qed lemma pCons_induct2 [case_names 00 lpCons rpCons pCons2]: assumes 00: "P 0 0" and lpCons: "\a p. a \ 0 \ p \ 0 \ P (pCons a p) 0" and rpCons: "\b q. b \ 0 \ q \ 0 \ P 0 (pCons b q)" and pCons2: "\a p b q. a \ 0 \ p \ 0 \ b \ 0 \ q \ 0 \ P p q \ P (pCons a p) (pCons b q)" shows "P p q" proof (induct p arbitrary: q) case 0 show ?case proof (cases q) fix b q' assume "q = pCons b q'" with 00 rpCons show ?thesis by (cases "b \ 0 \ q' \ 0") auto qed next case (pCons a p) show ?case proof (cases q) fix b q' assume "q = pCons b q'" with pCons lpCons pCons2 show ?thesis by (cases "b \ 0 \ q' \ 0") auto qed qed subsection \Algebra of sets\ subsubsection \General facts\ lemma zeroset_eqI: "0 \ A \ (\a. a \ A \ a = 0) \ A = 0" by auto lemma sum_list_sets_single : "(\X\[A]. X) = A" using add_0_right[of A] by simp lemma sum_list_sets_double : "(\X\[A,B]. X) = A + B" using add_0_right[of B] by simp subsubsection \Additive independence of sets\ primrec add_independentS :: "'a::monoid_add set list \ bool" where "add_independentS [] = True" | "add_independentS (A#As) = ( add_independentS As \ (\x\(\B\As. B). \a\A. a + x = 0 \ a = 0) )" lemma add_independentS_doubleI: assumes "\b a. b\B \ a\A \ a + b = 0 \ a = 0" shows "add_independentS [A,B]" using assms sum_list_sets_single[of B] by simp lemma add_independentS_doubleD: assumes "add_independentS [A,B]" shows "\b a. b\B \ a\A \ a + b = 0 \ a = 0" using assms sum_list_sets_single[of B] by simp lemma add_independentS_double_iff : "add_independentS [A,B] = (\b\B. \a\A. a + b = 0 \ a = 0 )" using add_independentS_doubleI add_independentS_doubleD by fast lemma add_independentS_Cons_conv_sum_right : "add_independentS (A#As) = (add_independentS [A,\B\As. B] \ add_independentS As)" using add_independentS_double_iff[of A] by auto lemma add_independentS_double_sum_conv_append : "\ \X\set As. 0 \ X; add_independentS As; add_independentS Bs; add_independentS [\X\As. X, \X\Bs. X] \ \ add_independentS (As@Bs)" proof (induct As) case (Cons A As) have "add_independentS [\X\As. X, \X\Bs. X]" proof (rule add_independentS_doubleI) fix b a assume ba: "b \ (\X\Bs. X)" "a \ (\X\As. X)" "a + b = 0" from Cons(2) ba(2) have "a \ (\X\A#As. X)" using set_plus_intro[of 0 A a] by simp with ba(1,3) Cons(5) show "a = 0" using add_independentS_doubleD[of "\X\A # As. X" "\X\Bs. X" b a] by simp qed moreover have "\x a. \ x \ (\X\As@Bs. X); a \ A; a + x = 0 \ \ a = 0" proof- fix x a assume x_a: "x \ (\X\As@Bs. X)" "a \ A" "a + x = 0" from x_a(1) obtain xa xb where xa_xb: "x = xa + xb" "xa \ (\X\As. X)" "xb \ (\X\Bs. X)" using set_plus_elim[of x "\X\As. X"] by auto from x_a(2) xa_xb(2) have "a + xa \ A + (\X\As. X)" using set_plus_intro by auto with Cons(3,5) xa_xb x_a(2,3) show "a = 0" using add_independentS_doubleD[ of "\X\A # As. X" "\X\Bs. X" xb "a+xa" ] add.assoc[of a] add_independentS_doubleD by simp qed ultimately show "add_independentS ((A#As)@Bs)" using Cons by simp qed simp lemma add_independentS_ConsI : assumes "add_independentS As" "\x a. \ x\(\X\As. X); a \ A; a+x = 0 \ \ a = 0" shows "add_independentS (A#As)" using assms by simp lemma add_independentS_append_reduce_right : "add_independentS (As@Bs) \ add_independentS Bs" by (induct As) auto lemma add_independentS_append_reduce_left : "add_independentS (As@Bs) \ 0 \ (\X\Bs. X) \ add_independentS As" proof (induct As) case (Cons A As) show "add_independentS (A#As)" proof (rule add_independentS_ConsI) from Cons show "add_independentS As" by simp next fix x a assume x: "x \ (\X\As. X)" and a: "a \ A" and sum: "a+x = 0" from x Cons(3) have "x + 0 \ (\X\As. X) + (\X\Bs. X)" by fast with a sum Cons(2) show "a = 0" by simp qed qed simp lemma add_independentS_append_conv_double_sum : "add_independentS (As@Bs) \ add_independentS [\X\As. X, \X\Bs. X]" proof (induct As) case (Cons A As) show "add_independentS [\X\(A#As). X, \X\Bs. X]" proof (rule add_independentS_doubleI) fix b x assume bx: "b \ (\X\Bs. X)" "x \ (\X\A # As. X)" "x + b = 0" from bx(2) obtain a as where a_as: "a \ A" "as \ listset As" "x = a + (\z\as. z)" using set_sum_list_element_Cons by fast from Cons(2) have "add_independentS [A,\X\As@Bs. X]" using add_independentS_Cons_conv_sum_right[of A "As@Bs"] by simp moreover from a_as(2) bx(1) have "(\z\as. z) + b \ (\X\(As@Bs). X)" using sum_list_listset set_plus_intro by auto ultimately have "a = 0" using a_as(1,3) bx(3) add_independentS_doubleD[of A _ _ a] add.assoc[of a] by auto with a_as(2,3) bx(1,3) Cons show "x = 0" using sum_list_listset add_independentS_doubleD[of "\X\As. X" "\X\Bs. X" b "\z\as. z"] by auto qed qed simp subsubsection \Inner direct sums\ definition inner_dirsum :: "'a::monoid_add set list \ 'a set" where "inner_dirsum As = (if add_independentS As then \A\As. A else 0)" text\Some syntactic sugar for \inner_dirsum\, borrowed from theory @{theory HOL.List}.\ syntax "_inner_dirsum" :: "pttrn => 'a list => 'b => 'b" ("(3\_\_. _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\M\Ms. b" == "CONST inner_dirsum (CONST map (%M. b) Ms)" abbreviation inner_dirsum_double :: "'a::monoid_add set \ 'a set \ 'a set" (infixr "\" 70) where "inner_dirsum_double A B \ inner_dirsum [A,B]" lemma inner_dirsumI : "M = (\N\Ns. N) \ add_independentS Ns \ M = (\N\Ns. N)" unfolding inner_dirsum_def by simp lemma inner_dirsum_doubleI : "M = A + B \ add_independentS [A,B] \ M = A \ B" using inner_dirsumI[of M "[A,B]"] sum_list_sets_double[of A] by simp lemma inner_dirsumD : "add_independentS Ms \ (\M\Ms. M) = (\M\Ms. M)" unfolding inner_dirsum_def by simp lemma inner_dirsumD2 : "\ add_independentS Ms \ (\M\Ms. M) = 0" unfolding inner_dirsum_def by simp lemma inner_dirsum_Nil : "(\A\[]. A) = 0" unfolding inner_dirsum_def by simp lemma inner_dirsum_singleD : "(\N\[M]. N) = M" using inner_dirsumD[of "[M]"] sum_list_sets_single[of M] by simp lemma inner_dirsum_doubleD : "add_independentS [M,N] \ M \ N = M + N" using inner_dirsumD[of "[M,N]"] sum_list_sets_double[of M N] by simp lemma inner_dirsum_Cons : "add_independentS (A # As) \ (\X\(A#As). X) = A \ (\X\As. X)" using inner_dirsumD[of "A#As"] add_independentS_Cons_conv_sum_right[of A] inner_dirsum_doubleD[of A] inner_dirsumD[of As] by simp lemma inner_dirsum_append : "add_independentS (As@Bs) \ 0 \ (\X\Bs. X) \ (\X\(As@Bs). X) = (\X\As. X) \ (\X\Bs. X)" using inner_dirsumD[of "As@Bs"] add_independentS_append_reduce_left[of As] inner_dirsumD[of As] inner_dirsumD[of Bs] add_independentS_append_reduce_right[of As Bs] add_independentS_append_conv_double_sum[of As] inner_dirsum_doubleD[of "\X\As. X"] by simp lemma inner_dirsum_double_left0: "0 \ A = A" using add_independentS_doubleD inner_dirsum_doubleI[of "0+A"] add_0_left[of A] by simp lemma add_independentS_Cons_conv_dirsum_right : "add_independentS (A#As) = (add_independentS [A,\B\As. B] \ add_independentS As)" using add_independentS_Cons_conv_sum_right[of A As] inner_dirsumD by auto lemma sum_list_listset_dirsum : "add_independentS As \ as \ listset As \ sum_list as \ (\A\As. A)" using inner_dirsumD sum_list_listset by fast section \Groups\ subsection \Locales and basic facts\ subsubsection \Locale \Group\ and finite variant \FinGroup\\ text \ Define a \Group\ to be a closed subset of @{term UNIV} for the \group_add\ class. \ locale Group = fixes G :: "'g::group_add set" assumes nonempty : "G \ {}" and diff_closed: "\g h. g \ G \ h \ G \ g - h \ G" lemma trivial_Group : "Group 0" by unfold_locales auto locale FinGroup = Group G for G :: "'g::group_add set" + assumes finite: "finite G" lemma (in FinGroup) Group : "Group G" by unfold_locales lemma (in Group) FinGroupI : "finite G \ FinGroup G" by unfold_locales context Group begin abbreviation Subgroup :: "'g set \ bool" where "Subgroup H \ Group H \ H \ G" lemma SubgroupD1 : "Subgroup H \ Group H" by fast lemma zero_closed : "0 \ G" proof- from nonempty obtain g where "g \ G" by fast hence "g - g \ G" using diff_closed by fast thus ?thesis by simp qed lemma obtain_nonzero: assumes "G \ 0" obtains g where "g \ G" and "g \ 0" using assms zero_closed by auto lemma zeroS_closed : "0 \ G" using zero_closed by simp lemma neg_closed : "g \ G \ -g \ G" using zero_closed diff_closed[of 0 g] by simp lemma add_closed : "g \ G \ h \ G \ g + h \ G" using neg_closed[of h] diff_closed[of g "-h"] by simp lemma neg_add_closed : "g \ G \ h \ G \ -g + h \ G" using neg_closed add_closed by fast lemma sum_list_closed : "set (map f as) \ G \ (\a\as. f a) \ G" using zero_closed add_closed by (induct as) auto lemma sum_list_closed_prod : "set (map (case_prod f) xys) \ G \ (\(x,y)\xys. f x y) \ G" using sum_list_closed by fast lemma set_plus_closed : "A \ G \ B \ G \ A + B \ G" using set_plus_def[of A B] add_closed by force lemma zip_add_closed : "set as \ G \ set bs \ G \ set [a + b. (a,b)\zip as bs] \ G" using add_closed by (induct as bs rule: list_induct2') auto lemma list_diff_closed : "set gs \ G \ set hs \ G \ set [x-y. (x,y)\zip gs hs] \ G" using diff_closed by (induct gs hs rule: list_induct2') auto lemma add_closed_converse_right : "g+x \ G \ g \ G \ x \ G" using neg_add_closed add.assoc[of "-g" g x] by fastforce lemma add_closed_inverse_right : "x \ G \ g \ G \ g+x \ G" using add_closed_converse_right by fast lemma add_closed_converse_left : "g+x \ G \ x \ G \ g \ G" using diff_closed add.assoc[of g] by fastforce lemma add_closed_inverse_left : "g \ G \ x \ G \ g+x \ G" using add_closed_converse_left by fast lemma right_translate_bij : assumes "g \ G" shows "bij_betw (\x. x + g) G G" unfolding bij_betw_def proof show "inj_on (\x. x + g) G" by (rule inj_onI) simp show "(\x. x + g) ` G = G" proof show "(\x. x + g) ` G \ G" using assms add_closed by fast show "(\x. x + g) ` G \ G" proof fix x assume "x \ G" with assms have "x - g \ G" "x = (\x. x + g) (x - g)" using diff_closed diff_add_cancel[of x] by auto thus "x \ (\x. x + g) ` G" by fast qed qed qed lemma right_translate_sum : "g \ G \ (\h\G. f h) = (\h\G. f (h + g))" using right_translate_bij[of g] bij_betw_def[of "\h. h + g"] sum.reindex[of "\h. h + g" G] by simp end (* context Group *) subsubsection \Abelian variant locale \AbGroup\\ locale AbGroup = Group G for G :: "'g::ab_group_add set" begin lemmas nonempty = nonempty lemmas zero_closed = zero_closed lemmas diff_closed = diff_closed lemmas add_closed = add_closed lemmas neg_closed = neg_closed lemma sum_closed : "finite A \ f ` A \ G \ (\a\A. f a) \ G" proof (induct set: finite) case empty show ?case using zero_closed by simp next case (insert a A) thus ?case using add_closed by simp qed lemma subset_plus_right : "A \ G + A" using zero_closed set_zero_plus2 by fast lemma subset_plus_left : "A \ A + G" using subset_plus_right add.commute by fast end (* context AbGroup *) subsection \Right cosets\ context Group begin definition rcoset_rel :: "'g set \ ('g\'g) set" where "rcoset_rel H \ {(g,g'). g \ G \ g' \ G \ g - g' \ H}" lemma (in Group) rcosets : assumes subgrp: "Subgroup H" and g: "g \ G" shows "(rcoset_rel H)``{g} = H + {g}" proof (rule seteqI) fix x assume "x \ (rcoset_rel H)``{g}" hence "x \ G" "g - x \ H" using rcoset_rel_def by auto with subgrp have "x - g \ H" using Group.neg_closed minus_diff_eq[of g x] by fastforce from this obtain h where h: "h \ H" "x - g = h" by fast from h(2) have "x = h + g" by (simp add: algebra_simps) with h(1) show "x \ H + {g}" using set_plus_def by fast next fix x assume "x \ H + {g}" from this obtain h where h: "h \ H" "x = h + g" unfolding set_plus_def by fast with subgrp g have 1: "x \ G" using add_closed by fast from h(2) have "x - g = h" by (simp add: algebra_simps) with subgrp h(1) have "g - x \ H" using Group.neg_closed by fastforce with g 1 show "x \ (rcoset_rel H)``{g}" using rcoset_rel_def by fast qed lemma rcoset_equiv : assumes "Subgroup H" shows "equiv G (rcoset_rel H)" proof (rule equivI) show "refl_on G (rcoset_rel H)" proof (rule refl_onI) show "(rcoset_rel H) \ G \ G" using rcoset_rel_def by auto next fix x assume "x \ G" with assms show "(x,x) \ (rcoset_rel H)" using rcoset_rel_def Group.zero_closed by auto qed show "sym (rcoset_rel H)" proof (rule symI) fix a b assume "(a,b) \ (rcoset_rel H)" with assms show "(b,a) \ (rcoset_rel H)" using rcoset_rel_def Group.neg_closed[of H "a - b"] minus_diff_eq by simp qed show "trans (rcoset_rel H)" proof (rule transI) fix x y z assume "(x,y) \ (rcoset_rel H)" "(y,z) \ (rcoset_rel H)" with assms show "(x,z) \ (rcoset_rel H)" using rcoset_rel_def Group.add_closed[of H "x - y" "y - z"] by (simp add: algebra_simps) qed qed lemma rcoset0 : "Subgroup H \ (rcoset_rel H)``{0} = H" using zero_closed rcosets unfolding set_plus_def by simp definition is_rcoset_replist :: "'g set \ 'g list \ bool" where "is_rcoset_replist H gs \ set gs \ G \ distinct (map (\g. (rcoset_rel H)``{g}) gs) \ G = (\g\set gs. (rcoset_rel H)``{g})" lemma is_rcoset_replistD_set : "is_rcoset_replist H gs \ set gs \ G" unfolding is_rcoset_replist_def by fast lemma is_rcoset_replistD_distinct : "is_rcoset_replist H gs \ distinct (map (\g. (rcoset_rel H)``{g}) gs)" unfolding is_rcoset_replist_def by fast lemma is_rcoset_replistD_cosets : "is_rcoset_replist H gs \ G = (\g\set gs. (rcoset_rel H)``{g})" unfolding is_rcoset_replist_def by fast lemma group_eq_subgrp_rcoset_un : "Subgroup H \ is_rcoset_replist H gs \ G = (\g\set gs. H + {g})" using is_rcoset_replistD_set is_rcoset_replistD_cosets rcosets by (auto, smt UN_E subsetCE, blast) lemma is_rcoset_replist_imp_nrelated_nth : assumes "Subgroup H" "is_rcoset_replist H gs" shows "\i j. i < length gs \ j < length gs \ i \ j \ gs!i - gs!j \ H" proof fix i j assume ij: "i < length gs" "j < length gs" "i\j" "gs!i - gs!j \ H" from assms(2) ij(1,2,4) have "(gs!i,gs!j) \ rcoset_rel H" using set_conv_nth is_rcoset_replistD_set rcoset_rel_def by fastforce with assms(1) ij(1,2) have "(map (\g. (rcoset_rel H)``{g}) gs)!i = (map (\g. (rcoset_rel H)``{g}) gs)!j" using rcoset_equiv equiv_class_eq by fastforce with assms(2) ij(1-3) show False using is_rcoset_replistD_distinct distinct_conv_nth[ of "map (\g. (rcoset_rel H)``{g}) gs" ] by auto qed lemma is_rcoset_replist_Cons : "is_rcoset_replist H (g#gs) \ g \ G \ set gs \ G \ (rcoset_rel H)``{g} \ set (map (\x. (rcoset_rel H)``{x}) gs) \ distinct (map (\x. (rcoset_rel H)``{x}) gs) \ G = (rcoset_rel H)``{g} \ (\x\set gs. (rcoset_rel H)``{x})" using is_rcoset_replist_def[of H "g#gs"] by auto lemma rcoset_replist_Hrep : assumes "Subgroup H" "is_rcoset_replist H gs" shows "\g\set gs. g \ H" proof- from assms(2) obtain g where g: "g \ set gs" "0 \ (rcoset_rel H)``{g}" using zero_closed is_rcoset_replistD_cosets by fast from assms(1) g(2) have "g \ (rcoset_rel H)``{0}" using rcoset_equiv equivE sym_def[of "rcoset_rel H"] by force with assms(1) g show "\g\set gs. g \ H" using rcoset0 by fast qed lemma rcoset_replist_reorder : "is_rcoset_replist H (gs @ g # gs') \ is_rcoset_replist H (g # gs @ gs')" using is_rcoset_replist_def by auto lemma rcoset_replist_replacehd : assumes "Subgroup H" "g' \ (rcoset_rel H)``{g}" "is_rcoset_replist H (g # gs)" shows "is_rcoset_replist H (g' # gs)" proof- from assms(2) have "g' \ G" using rcoset_rel_def by simp moreover from assms(3) have "set gs \ G" using is_rcoset_replistD_set by fastforce moreover from assms(1-3) have "(rcoset_rel H)``{g'} \ set (map (\x. (rcoset_rel H)``{x}) gs)" using set_conv_nth[of gs] rcoset_equiv equiv_class_eq_iff[of G] is_rcoset_replistD_distinct by fastforce moreover from assms(3) have "distinct (map (\g. (rcoset_rel H)``{g}) gs)" using is_rcoset_replistD_distinct by fastforce moreover from assms(1-3) have "G = (rcoset_rel H)``{g'} \ (\x\set gs. (rcoset_rel H)``{x})" using is_rcoset_replistD_cosets[of H "g#gs"] rcoset_equiv equiv_class_eq_iff[of G] by simp ultimately show ?thesis using is_rcoset_replist_Cons by auto qed end (* context Group *) lemma (in FinGroup) ex_rcoset_replist : assumes "Subgroup H" shows "\gs. is_rcoset_replist H gs" proof- define r where "r = rcoset_rel H" hence equiv_r: "equiv G r" using rcoset_equiv[OF assms] by fast have "\x\G//r. \g. g \ x" proof fix x assume "x \ G//r" from this obtain g where g: "g \ G" "x = r``{g}" using quotient_def[of G r] by auto hence "g \ x" using equiv_r equivE[of G r] refl_onD[of G r] by auto thus "\g. g \ x" by fast qed from this obtain f where f: "\x\G//r. f x \ x" using bchoice by force from r_def have "r \ G \ G" using rcoset_rel_def by auto with finite have "finite (f`(G//r))" using finite_quotient by auto from this obtain gs where gs: "distinct gs" "set gs = f`(G//r)" using finite_distinct_list by force have 1: "set gs \ G" proof fix g assume "g \ set gs" with gs(2) obtain x where x: "x \ G//r" "g = f x" by fast with f show "g \ G" using equiv_r Union_quotient by fast qed moreover have "distinct (map (\g. r``{g}) gs)" proof- have "\i j. \ i < length (map (\g. r``{g}) gs); j < length (map (\g. r``{g}) gs); i \ j \ \ (map (\g. r``{g}) gs)!i \ (map (\g. r``{g}) gs)!j" proof fix i j assume ij: "i < length (map (\g. r``{g}) gs)" "j< length (map (\g. r``{g}) gs)" "i \ j" "(map (\g. r``{g}) gs)!i = (map (\g. r``{g}) gs)!j" from ij(1,2) have "gs!i \ set gs" "gs!j \ set gs" using set_conv_nth by auto from this gs(2) obtain x y where x: "x \ G//r" "gs!i = f x" and y: "y \ G//r" "gs!j = f y" by auto have "x = y" using equiv_r x(1) y(1) proof (rule quotient_eqI[of G r]) from ij(1,2,4) have "r``{gs!i} = r``{gs!j}" by simp with ij(1,2) 1 show "(gs!i,gs!j) \ r" using eq_equiv_class_iff[OF equiv_r] by force from x y f show "gs!i \ x" "gs!j \ y" by auto qed with x(2) y(2) ij(1-3) gs(1) show False using distinct_conv_nth by fastforce qed thus ?thesis using distinct_conv_nth by fast qed moreover have "G = (\g\set gs. r``{g})" proof (rule subset_antisym, rule subsetI) fix g assume "g \ G" hence rg: "r``{g} \ G//r" using quotientI by fast with gs(2) obtain g' where g': "g' \ set gs" "g' = f (r``{g})" by fast from f g'(2) rg have "g \ r``{g'}" using equiv_r equivE sym_def[of r] by force with g'(1) show "g \ (\g\set gs. r``{g})" by fast next from r_def show "G \ (\g\set gs. r``{g})" using rcoset_rel_def by auto qed ultimately show ?thesis using r_def unfolding is_rcoset_replist_def by fastforce qed lemma (in FinGroup) ex_rcoset_replist_hd0 : assumes "Subgroup H" shows "\gs. is_rcoset_replist H (0#gs)" proof- from assms obtain xs where xs: "is_rcoset_replist H xs" using ex_rcoset_replist by fast with assms obtain x where x: "x \ set xs" "x \ H" using rcoset_replist_Hrep by fast from x(2) have "(0,x) \ rcoset_rel H" using rcoset0[OF assms] by auto moreover have "sym (rcoset_rel H)" using rcoset_equiv[OF assms] equivE[of G "rcoset_rel H"] by simp ultimately have "0 \ (rcoset_rel H)``{x}" using sym_def by fast with x(1) xs assms show "\gs. is_rcoset_replist H (0#gs)" using split_list rcoset_replist_reorder rcoset_replist_replacehd by fast qed subsection \Group homomorphisms\ subsubsection \Preliminaries\ definition ker :: "('a\'b::zero) \ 'a set" where "ker f = {a. f a = 0}" lemma kerI : "f a = 0 \ a \ ker f" unfolding ker_def by fast lemma kerD : "a \ ker f \ f a = 0" unfolding ker_def by fast lemma ker_im_iff : "(A \ {} \ A \ ker f) = (f ` A = 0)" proof assume A: "A \ {} \ A \ ker f" show "f ` A = 0" proof (rule zeroset_eqI) from A obtain a where a: "a \ A" by fast with A have "f a = 0" using kerD by fastforce from this[THEN sym] a show "0 \ f ` A" by fast next fix b assume "b \ f ` A" from this obtain a where "a \ A" "b = f a" by fast with A show "b = 0" using kerD by fast qed next assume fA: "f ` A = 0" have "A \ {}" proof- from fA obtain a where "a \ A" "f a = 0" by force thus ?thesis by fast qed moreover have "A \ ker f" proof fix a assume "a \ A" with fA have "f a = 0" by auto thus "a \ ker f" using kerI by fast qed ultimately show "A \ {} \ A \ ker f" by fast qed subsubsection \Locales\ text \The \supp\ condition is not strictly necessary, but helps with equality and uniqueness arguments.\ locale GroupHom = Group G for G :: "'g::group_add set" + fixes T :: "'g \ 'h::group_add" assumes hom : "\g g'. g \ G \ g' \ G \ T (g + g') = T g + T g'" and supp: "supp T \ G" abbreviation (in GroupHom) "Ker \ ker T \ G" abbreviation (in GroupHom) "ImG \ T ` G" locale GroupEnd = GroupHom G T for G :: "'g::group_add set" and T :: "'g \ 'g" + assumes endomorph: "ImG \ G" locale GroupIso = GroupHom G T for G :: "'g::group_add set" and T :: "'g \ 'h::group_add" + fixes H :: "'h set" assumes bijective: "bij_betw T G H" subsubsection \Basic facts\ lemma (in Group) trivial_GroupHom : "GroupHom G (0::('g\'h::group_add))" proof fix g g' define z z_map where "z = (0::'h)" and "z_map = (0::'g\'h)" thus "z_map (g + g') = z_map g + z_map g'" by simp qed (rule supp_zerofun_subset_any) lemma (in Group) GroupHom_idhom : "GroupHom G (id\G)" using add_closed supp_restrict0 by unfold_locales simp context GroupHom begin lemma im_zero : "T 0 = 0" using zero_closed hom[of 0 0] add_diff_cancel[of "T 0" "T 0"] by simp lemma zero_in_Ker : "0 \ Ker" using im_zero kerI zero_closed by fast lemma comp_zero : "T \ 0 = 0" using im_zero by auto lemma im_neg : "T (- g) = - T g" using im_zero hom[of g "- g"] neg_closed[of g] minus_unique[of "T g"] neg_closed[of "-g"] supp suppI_contra[of g T] suppI_contra[of "-g" T] by fastforce lemma im_diff : "g \ G \ g' \ G \ T (g - g') = T g - T g'" using hom neg_closed hom[of g "-g'"] im_neg by simp lemma eq_im_imp_diff_in_Ker : "\ g \ G; h \ G; T g = T h \ \ g - h \ Ker" using im_diff kerI diff_closed[of g h] by force lemma im_sum_list_prod : "set (map (case_prod f) xys) \ G \ T (\(x,y)\xys. f x y) = (\(x,y)\xys. T (f x y))" proof (induct xys) case Nil show ?case using im_zero by simp next case (Cons xy xys) define Tf where "Tf = T \ (case_prod f)" have "T (\(x,y)\(xy#xys). f x y) = T ( (case_prod f) xy + (\(x,y)\xys. f x y) )" by simp moreover from Cons(2) have "(case_prod f) xy \ G" "set (map (case_prod f) xys) \ G" by auto ultimately have "T (\(x,y)\(xy#xys). f x y) = Tf xy + (\(x,y)\xys. Tf (x,y))" using Tf_def sum_list_closed[of "case_prod f"] hom Cons by auto also have "\ = (\(x,y)\(xy#xys). Tf (x,y))" by simp finally show ?case using Tf_def by simp qed lemma distrib_comp_sum_left : "range S \ G \ range S' \ G \ T \ (S + S') = (T \ S) + (T \ S')" using hom by (auto simp add: fun_eq_iff) lemma Ker_Im_iff : "(Ker = G) = (ImG = 0)" using nonempty ker_im_iff[of G T] by fast lemma Ker0_imp_inj_on : assumes "Ker \ 0" shows "inj_on T G" proof (rule inj_onI) fix x y assume xy: "x \ G" "y \ G" "T x = T y" hence "T (x - y) = 0" using im_diff by simp with xy(1,2) have "x - y \ Ker" using diff_closed kerI by fast with assms show "x = y" using zero_in_Ker by auto qed lemma inj_on_imp_Ker0 : assumes "inj_on T G" shows "Ker = 0" using zero_in_Ker kerD zero_closed im_zero inj_onD[OF assms] by fastforce lemma nonzero_Ker_el_imp_n_inj : "g \ G \ g \ 0 \ T g = 0 \ \ inj_on T G" using inj_on_imp_Ker0 kerI[of T] by auto lemma Group_Ker : "Group Ker" proof show "Ker \ {}" using zero_in_Ker by fast next fix g h assume "g \ Ker" "h \ Ker" thus "g - h \ Ker" using im_diff kerD[of g T] kerD[of h T] diff_closed kerI[of T] by auto qed lemma Group_Im : "Group ImG" proof show "ImG \ {}" using nonempty by fast next fix g' h' assume "g' \ ImG" "h' \ ImG" from this obtain g h where gh: "g \ G" "g' = T g" "h \ G" "h' = T h" by fast thus "g' - h' \ ImG" using im_diff diff_closed by force qed lemma GroupHom_restrict0_subgroup : assumes "Subgroup H" shows "GroupHom H (T \ H)" proof (intro_locales, rule SubgroupD1[OF assms], unfold_locales) show "supp (T \ H) \ H" using supp_restrict0 by fast next fix h h' assume hh': "h \ H" "h' \ H" with assms show "(T \ H) (h + h') = (T \ H) h + (T \ H) h'" using Group.add_closed hom[of h h'] by auto qed lemma im_subgroup : assumes "Subgroup H" shows "Group.Subgroup ImG (T ` H)" proof from assms have "Group ((T \ H) ` H)" using GroupHom_restrict0_subgroup GroupHom.Group_Im by fast moreover have "(T \ H) ` H = T ` H" by auto ultimately show "Group (T ` H)" by simp from assms show "T ` H \ ImG" by fast qed lemma GroupHom_composite_left : assumes "ImG \ H" "GroupHom H S" shows "GroupHom G (S \ T)" proof fix g g' assume "g \ G" "g' \ G" with hom assms(1) show "(S \ T) (g + g') = (S \ T) g + (S \ T) g'" using GroupHom.hom[OF assms(2)] by fastforce next from supp have "\g. g \ G \ (S \ T) g = 0" using suppI_contra GroupHom.im_zero[OF assms(2)] by fastforce thus "supp (S \ T) \ G" using suppD_contra by fast qed lemma idhom_left : "T ` G \ H \ (id\H) \ T = T" using supp suppI_contra by fastforce end (* context GroupHom *) subsubsection \Basic facts about endomorphisms\ context GroupEnd begin lemmas hom = hom lemma range : "range T \ G" proof (rule image_subsetI) fix x show "T x \ G" proof (cases "x \ G") case True with endomorph show ?thesis by fast next case False with supp show ?thesis using suppI_contra zero_closed by fastforce qed qed lemma proj_decomp : assumes "\g. g \ G \ T (T g) = T g" shows "G = Ker \ ImG" proof (rule inner_dirsum_doubleI, rule subset_antisym, rule subsetI) fix g assume g: "g \ G" have "g = (g - T g) + T g" using diff_add_cancel[of g] by simp moreover have "g - T g \ Ker" proof from g endomorph assms have "T (g - T g) = 0" using im_diff by auto thus "g - T g \ ker T" using kerI by fast from g endomorph show "g - T g \ G" using diff_closed by fast qed moreover from g have "T g \ ImG" by fast ultimately show "g \ Ker + ImG" using set_plus_intro[of "g - T g" Ker "T g"] by simp next from endomorph show "G \ Ker + ImG" using set_plus_closed by simp show "add_independentS [Ker,ImG]" proof (rule add_independentS_doubleI) fix g h assume gh: "h \ ImG" "g \ Ker" "g + h = 0" from gh(1) obtain g' where "g' \ G" "h = T g'" by fast with gh(2,3) endomorph assms have "h = 0" using im_zero hom[of g "T g'"] kerD by fastforce with gh(3) show "g = 0" by simp qed qed end (* context GroupEnd *) subsubsection \Basic facts about isomorphisms\ context GroupIso begin abbreviation "invT \ (the_inv_into G T) \ H" lemma ImG : "ImG = H" using bijective bij_betw_imp_surj_on by fast lemma GroupH : "Group H" using ImG Group_Im by fast lemma invT_onto : "invT ` H = G" using bijective bij_betw_imp_inj_on[of T] ImG the_inv_into_onto[of T] by force lemma inj_on_invT : "inj_on invT H" using bijective bij_betw_imp_inj_on[of T G] ImG inj_on_the_inv_into[of T] unfolding inj_on_def by force lemma bijective_invT : "bij_betw invT H G" using inj_on_invT invT_onto unfolding bij_betw_def by fast lemma invT_into : "h \ H \ invT h \ G" using bijective bij_betw_imp_inj_on ImG the_inv_into_into[of T] by force lemma T_invT : "h \ H \ T (invT h) = h" using bijective bij_betw_imp_inj_on ImG f_the_inv_into_f[of T] by force lemma invT_eq: "g \ G \ T g = h \ invT h = g" using bijective bij_betw_imp_inj_on ImG the_inv_into_f_eq[of T] by force lemma inv : "GroupIso H invT G" proof (intro_locales, rule GroupH, unfold_locales) show "supp invT \ H" using supp_restrict0 by fast show "bij_betw invT H G" using bijective_invT by fast next fix h h' assume "h \ H" "h' \ H" thus "invT (h + h') = invT h + invT h'" using invT_into hom T_invT add_closed invT_eq by simp qed end (* context GroupIso *) subsubsection \Hom-sets\ definition GroupHomSet :: "'g::group_add set \ 'h::group_add set \ ('g \ 'h) set" where "GroupHomSet G H \ {T. GroupHom G T} \ {T. T ` G \ H}" lemma GroupHomSetI : "GroupHom G T \ T ` G \ H \ T \ GroupHomSet G H" unfolding GroupHomSet_def by fast lemma GroupHomSetD_GroupHom : "T \ GroupHomSet G H \ GroupHom G T" unfolding GroupHomSet_def by fast lemma GroupHomSetD_Im : "T \ GroupHomSet G H \ T ` G \ H" unfolding GroupHomSet_def by fast lemma (in Group) Group_GroupHomSet : fixes H :: "'h::ab_group_add set" assumes "AbGroup H" shows "Group (GroupHomSet G H)" proof show "GroupHomSet G H \ {}" using trivial_GroupHom AbGroup.zero_closed[OF assms] GroupHomSetI by fastforce next fix S T assume ST: "S \ GroupHomSet G H" "T \ GroupHomSet G H" show "S - T \ GroupHomSet G H" proof (rule GroupHomSetI, unfold_locales) from ST show "supp (S - T) \ G" using GroupHomSetD_GroupHom[of S] GroupHomSetD_GroupHom[of T] GroupHom.supp[of G S] GroupHom.supp[of G T] supp_diff_subset_union_supp[of S T] by auto show "(S - T) ` G \ H" proof (rule image_subsetI) fix g assume "g \ G" with ST have "S g \ H" "T g \ H" using GroupHomSetD_Im[of S G] GroupHomSetD_Im[of T G] by auto thus "(S - T) g \ H" using AbGroup.diff_closed[OF assms] by simp qed next fix g g' assume "g \ G" "g' \ G" with ST show "(S - T) (g + g') = (S - T) g + (S - T) g'" using GroupHomSetD_GroupHom[of S] GroupHom.hom[of G S] GroupHomSetD_GroupHom[of T] GroupHom.hom[of G T] by simp qed qed subsection \Facts about collections of groups\ lemma listset_Group_plus_closed : "\ \G\set Gs. Group G; as \ listset Gs; bs \ listset Gs\ \ [a+b. (a,b)\zip as bs] \ listset Gs" proof- have "\ length as = length bs; length bs = length Gs; as \ listset Gs; bs \ listset Gs; \G\set Gs. Group G\ \ [a+b. (a,b)\zip as bs] \ listset Gs" proof (induct as bs Gs rule: list_induct3) case (Cons a as b bs G Gs) thus "[x+y. (x,y)\zip (a#as) (b#bs)] \ listset (G#Gs)" using listset_ConsD[of a] listset_ConsD[of b] Group.add_closed listset_ConsI[of "a+b" G] by fastforce qed simp thus "\ \G\set Gs. Group G; as \ listset Gs; bs \ listset Gs\ \ [a+b. (a,b)\zip as bs] \ listset Gs" using listset_length[of as Gs] listset_length[of bs Gs, THEN sym] by fastforce qed lemma AbGroup_set_plus : assumes "AbGroup H" "AbGroup G" shows "AbGroup (H + G)" proof from assms show "H + G \ {}" using AbGroup.nonempty by blast next fix x y assume "x \ H + G" "y \ H + G" from this obtain xh xg yh yg where xy: "xh \ H" "xg \ G" "x = xh + xg" "yh \ H" "yg \ G" "y = yh + yg" unfolding set_plus_def by fast hence "x - y = (xh - yh) + (xg - yg)" by simp thus "x - y \ H + G" using assms xy(1,2,4,5) AbGroup.diff_closed by auto qed lemma AbGroup_sum_list : "(\G\set Gs. AbGroup G) \ AbGroup (\G\Gs. G)" using trivial_Group AbGroup.intro AbGroup_set_plus by (induct Gs) auto lemma AbGroup_subset_sum_list : "\G \ set Gs. AbGroup G \ H \ set Gs \ H \ (\G\Gs. G)" proof (induct Gs arbitrary: H) case (Cons G Gs) show "H \ (\X\(G#Gs). X)" proof (cases "H = G") case True with Cons(2) show ?thesis using AbGroup_sum_list AbGroup.subset_plus_left by auto next case False with Cons have "H \ (\G\Gs. G)" by simp with Cons(2) show ?thesis using AbGroup.subset_plus_right by auto qed qed simp lemma independent_AbGroups_pairwise_int0 : "\ \G\set Gs. AbGroup G; add_independentS Gs; G \ set Gs; G' \ set Gs; G \ G' \ \ G \ G' = 0" proof (induct Gs arbitrary: G G') case (Cons H Hs) from Cons(1-3) have "\A B. \ A \ set Hs; B \ set Hs; A \ B \ \ A \ B \ 0" by simp moreover have "\A. A \ set Hs \ A \ H \ A \ H \ 0" proof fix A g assume A: "A \ set Hs" "A \ H" and g: "g \ A \ H" from A(1) g Cons(2) have "-g \ (\X\Hs. X)" using AbGroup.neg_closed AbGroup_subset_sum_list by force moreover have "g + (-g) = 0" by simp ultimately show "g \ 0" using g Cons(3) by simp qed ultimately have "\A B. \ A \ set (H#Hs); B \ set (H#Hs); A \ B \ \ A \ B \ 0" by auto with Cons(4-6) have "G \ G' \ 0" by fast moreover from Cons(2,4,5) have "0 \ G \ G'" using AbGroup.zero_closed by auto ultimately show ?case by fast qed simp lemma independent_AbGroups_pairwise_int0_double : assumes "AbGroup G" "AbGroup G'" "add_independentS [G,G']" shows "G \ G' = 0" proof (cases "G = 0") case True with assms(2) show ?thesis using AbGroup.zero_closed by auto next case False show ?thesis proof (rule independent_AbGroups_pairwise_int0) from assms(1,2) show "\G\set [G,G']. AbGroup G" by simp from assms(3) show "add_independentS [G,G']" by fast show "G \ set [G,G']" "G' \ set [G,G']" by auto show "G \ G'" proof assume GG': "G = G'" from False assms(1) obtain g where g: "g \ G" "g \ 0" using AbGroup.nonempty by auto moreover from assms(2) GG' g(1) have "-g \ G'" using AbGroup.neg_closed by fast moreover have "g + (-g) = 0" by simp ultimately show False using assms(3) by force qed qed qed subsection \Inner direct sums of Abelian groups\ subsubsection \General facts\ lemma AbGroup_inner_dirsum : "\G\set Gs. AbGroup G \ AbGroup (\G\Gs. G)" using inner_dirsumD[of Gs] inner_dirsumD2[of Gs] AbGroup_sum_list AbGroup.intro trivial_Group by (cases "add_independentS Gs") auto lemma inner_dirsum_double_leftfull_imp_right0: assumes "Group A" "B \ {}" "A = A \ B" shows "B = 0" proof (cases "add_independentS [A,B]") case True with assms(3) have 1: "A = A + B" using inner_dirsum_doubleD by fast have "\b. b \ B \ b = 0" proof- fix b assume b: "b \ B" from assms(1) obtain a where a: "a \ A" using Group.nonempty by fast with b 1 have "a + b \ A" by fast from this obtain a' where a': "a' \ A" "a + b = a'" by fast hence "(-a'+a) + b = 0" by (simp add: add.assoc) with assms(1) True a a'(1) b show "b = 0" using Group.neg_add_closed[of A] add_independentS_doubleD[of A B b "-a'+a"] by simp qed with assms(2) show ?thesis by auto next case False hence 1: "A \ B = 0" unfolding inner_dirsum_def by auto moreover with assms(3) have "A = 0" by fast ultimately show ?thesis using inner_dirsum_double_left0 by auto qed lemma AbGroup_subset_inner_dirsum : "\ \G \ set Gs. AbGroup G; add_independentS Gs; H \ set Gs \ \ H \ (\G\Gs. G)" using AbGroup_subset_sum_list inner_dirsumD by fast lemma AbGroup_nth_subset_inner_dirsum : "\ \G \ set Gs. AbGroup G; add_independentS Gs; n < length Gs \ \ Gs!n \ (\G\Gs. G)" using AbGroup_subset_inner_dirsum by force lemma AbGroup_inner_dirsum_el_decomp_ex1_double : assumes "AbGroup G" "AbGroup H" "add_independentS [G,H]" "x \ G \ H" shows "\!gh. fst gh \ G \ snd gh \ H \ x = fst gh + snd gh" proof (rule ex_ex1I) from assms(3,4) obtain g h where "x = g + h" "g \ G" "h \ H" using inner_dirsum_doubleD set_plus_elim by blast from this have 1: "fst (g,h) \ G" "snd (g,h) \ H" "x = fst (g,h) + snd (g,h)" by auto thus "\gh. fst gh \ G \ snd gh \ H \ x = fst gh + snd gh" by fast next fix gh gh' assume A: "fst gh \ G \ snd gh \ H \ x = fst gh + snd gh " "fst gh' \ G \ snd gh' \ H \ x = fst gh' + snd gh'" show "gh = gh'" proof from A assms(1,2) have "fst gh - fst gh' \ G" "snd gh - snd gh' \ H" using AbGroup.diff_closed by auto moreover from A have z: "(fst gh - fst gh') + (snd gh - snd gh') = 0" by (simp add: algebra_simps) ultimately show "fst gh = fst gh'" using assms(3) add_independentS_doubleD[of G H "snd gh - snd gh'" "fst gh - fst gh'"] by simp with z show "snd gh = snd gh'" by simp qed qed lemma AbGroup_inner_dirsum_el_decomp_ex1 : "\ \G \ set Gs. AbGroup G; add_independentS Gs \ \ \x \ (\G\Gs. G). \!gs\listset Gs. x = sum_list gs" proof (induct Gs) case Nil have "\x::'a. x \ (\H\[]. H) \ \!gs\listset []. x = sum_list gs" proof fix x::'a assume "x \ (\G\[]. G)" moreover define f :: "'a \ 'a list" where "f x = []" for x ultimately show "f x \ listset [] \ x = sum_list (f x)" using inner_dirsum_Nil by auto next fix x::'a and gs assume x: "x \ (\G\[]. G)" and gs: "gs \ listset [] \ x = sum_list gs" thus "gs = []" by simp qed thus "\x::'a \ (\H\[]. H). \!gs\listset []. x = sum_list gs" by fast next case (Cons G Gs) hence prevcase: "\x\(\H\Gs. H). \!gs\listset Gs. x = sum_list gs" by auto from Cons(2) have grps: "AbGroup G" "AbGroup (\H\Gs. H)" using AbGroup_inner_dirsum by auto from Cons(3) have ind: "add_independentS [G, \H\Gs. H]" using add_independentS_Cons_conv_dirsum_right by fast have "\x. x \ (\H\(G#Gs). H) \ \!gs\listset (G#Gs). x = sum_list gs" proof (rule ex_ex1I) fix x assume "x \ (\H\(G#Gs). H)" with Cons(3) have "x \ G \ (\H\Gs. H)" using inner_dirsum_Cons by fast with grps ind obtain gh where gh: "fst gh \ G" "snd gh \ (\H\Gs. H)" "x = fst gh + snd gh" using AbGroup_inner_dirsum_el_decomp_ex1_double by blast from gh(2) prevcase obtain gs where gs: "gs \ listset Gs" "snd gh = sum_list gs" by fast with gh(1) gs(1) have "fst gh # gs \ listset (G#Gs)" using set_Cons_def by fastforce moreover from gh(3) gs(2) have "x = sum_list (fst gh # gs)" by simp ultimately show "\gs. gs \ listset (G#Gs) \ x = sum_list gs" by fast next fix x gs hs assume "x \ (\H\(G#Gs). H)" and gs: "gs \ listset (G#Gs) \ x = sum_list gs" and hs: "hs \ listset (G#Gs) \ x = sum_list hs" hence "gs \ set_Cons G (listset Gs)" "hs \ set_Cons G (listset Gs)" by auto from this obtain a as b bs where a_as: "gs = a#as" "a\G" "as \ listset Gs" and b_bs: "hs = b#bs" "b\G" "bs \ listset Gs" unfolding set_Cons_def by fast from a_as(3) b_bs(3) Cons(3) have as: "sum_list as \ (\H\Gs. H)" and bs: "sum_list bs \ (\H\Gs. H)" using sum_list_listset_dirsum by auto with a_as(2) b_bs(2) grps have "a - b \ G" "sum_list as - sum_list bs \ (\H\Gs. H)" using AbGroup.diff_closed by auto moreover from gs hs a_as(1) b_bs(1) have z: "(a - b) + (sum_list as - sum_list bs) = 0" by (simp add: algebra_simps) ultimately have "a - b = 0" using ind add_independentS_doubleD by blast with z have 1: "a = b" and z': "sum_list as = sum_list bs" by auto from z' prevcase as a_as(3) bs b_bs(3) have 2: "as = bs" by fast from 1 2 a_as(1) b_bs(1) show "gs = hs" by fast qed thus "\x\(\H\(G#Gs). H). \!gs. gs \ listset (G#Gs) \ x = sum_list gs" by fast qed lemma AbGroup_inner_dirsum_pairwise_int0 : "\ \G \ set Gs. AbGroup G; add_independentS Gs; G \ set Gs; G' \ set Gs; G \ G' \ \ G \ G' = 0" proof (induct Gs arbitrary: G G') case (Cons H Hs) from Cons(1-3) have "\A B. \ A \ set Hs; B \ set Hs; A \ B \ \ A \ B \ 0" by simp moreover have "\A. A \ set Hs \ A \ H \ A \ H \ 0" proof fix A g assume A: "A \ set Hs" "A \ H" and g: "g \ A \ H" from A(1) g Cons(2) have "-g \ (\X\Hs. X)" using AbGroup.neg_closed AbGroup_subset_sum_list by force moreover have "g + (-g) = 0" by simp ultimately show "g \ 0" using g Cons(3) by simp qed ultimately have "\A B. \ A \ set (H#Hs); B \ set (H#Hs); A \ B \ \ A \ B \ 0" by auto with Cons(4-6) have "G \ G' \ 0" by fast moreover from Cons(2,4,5) have "0 \ G \ G'" using AbGroup.zero_closed by auto ultimately show ?case by fast qed simp lemma AbGroup_inner_dirsum_pairwise_int0_double : assumes "AbGroup G" "AbGroup G'" "add_independentS [G,G']" shows "G \ G' = 0" proof (cases "G = 0") case True with assms(2) show ?thesis using AbGroup.zero_closed by auto next case False show ?thesis proof (rule AbGroup_inner_dirsum_pairwise_int0) from assms(1,2) show "\G\set [G,G']. AbGroup G" by simp from assms(3) show "add_independentS [G,G']" by fast show "G \ set [G,G']" "G' \ set [G,G']" by auto show "G \ G'" proof assume GG': "G = G'" from False assms(1) obtain g where g: "g \ G" "g \ 0" using AbGroup.nonempty by auto moreover from assms(2) GG' g(1) have "-g \ G'" using AbGroup.neg_closed by fast moreover have "g + (-g) = 0" by simp ultimately show False using assms(3) by force qed qed qed subsubsection \Element decomposition and projection\ definition inner_dirsum_el_decomp :: "'g::ab_group_add set list \ ('g \ 'g list)" ("\_\") where "\Gs\ = (\x. if x \ (\G\Gs. G) then THE gs. gs \ listset Gs \ x = sum_list gs else [])" abbreviation inner_dirsum_el_decomp_double :: "'g::ab_group_add set \ 'g set \ ('g \ 'g list)" ("_\_\") where "G\H\ \ \[G,H]\" abbreviation inner_dirsum_el_decomp_nth :: "'g::ab_group_add set list \ nat \ ('g \ 'g)" ("\_\_") where "\Gs\n \ restrict0 (\x. (\Gs\x)!n) (\G\Gs. G)" lemma AbGroup_inner_dirsum_el_decompI : "\ \G \ set Gs. AbGroup G; add_independentS Gs; x \ (\G\Gs. G) \ \ (\Gs\x) \ listset Gs \ x = sum_list (\Gs\x)" using AbGroup_inner_dirsum_el_decomp_ex1 theI'[ of "\gs. gs \ listset Gs \ x = sum_list gs" ] unfolding inner_dirsum_el_decomp_def by fastforce lemma (in AbGroup) abSubgroup_inner_dirsum_el_decomp_set : "\ \H \ set Hs. Subgroup H; add_independentS Hs; x \ (\H\Hs. H) \ \ set (\Hs\x) \ G" using AbGroup.intro AbGroup_inner_dirsum_el_decompI[of Hs x] set_listset_el_subset[of "(\Hs\x)" Hs G] by fast lemma AbGroup_inner_dirsum_el_decomp_eq : "\ \G \ set Gs. AbGroup G; add_independentS Gs; x \ (\G\Gs. G); gs \ listset Gs; x = sum_list gs \ \ (\Gs\x) = gs" using AbGroup_inner_dirsum_el_decomp_ex1[of Gs] inner_dirsum_el_decomp_def[of Gs] by force lemma AbGroup_inner_dirsum_el_decomp_plus : assumes "\G \ set Gs. AbGroup G" "add_independentS Gs" "x \ (\G\Gs. G)" "y \ (\G\Gs. G)" shows "(\Gs\(x+y)) = [a+b. (a,b)\zip (\Gs\x) (\Gs\y)]" proof- define xs ys where "xs = (\Gs\x)" and "ys = (\Gs\y)" with assms have xy: "xs \ listset Gs" "x = sum_list xs" "ys \ listset Gs" "y = sum_list ys" using AbGroup_inner_dirsum_el_decompI by auto from assms(1) xy(1,3) have "[a+b. (a,b)\zip xs ys] \ listset Gs" using AbGroup.axioms listset_Group_plus_closed by fast moreover from xy have "x + y = sum_list [a+b. (a,b)\zip xs ys]" using listset_length[of xs Gs] listset_length[of ys Gs, THEN sym] sum_list_plus by simp ultimately show "(\Gs\(x+y)) = [a+b. (a,b)\zip xs ys]" using assms AbGroup_inner_dirsum AbGroup.add_closed AbGroup_inner_dirsum_el_decomp_eq by fast qed lemma AbGroup_length_inner_dirsum_el_decomp : "\ \G \ set Gs. AbGroup G; add_independentS Gs; x \ (\G\Gs. G) \ \ length (\Gs\x) = length Gs" using AbGroup_inner_dirsum_el_decompI listset_length by fastforce lemma AbGroup_inner_dirsum_el_decomp_in_nth : assumes "\G \ set Gs. AbGroup G" "add_independentS Gs" "n < length Gs" "x \ Gs!n" shows "(\Gs\x) = (replicate (length Gs) 0)[n := x]" proof- from assms have x: "x \ (\G\Gs. G)" using AbGroup_nth_subset_inner_dirsum by fast define xgs where "xgs = (replicate (length Gs) (0::'a))[n := x]" hence "length xgs = length Gs" by simp moreover have "\k Gs!k" proof- have "\k. k < length xgs \ xgs!k \ Gs!k" proof- fix k assume "k < length xgs" with assms(1,4) xgs_def show "xgs!k \ Gs!k" using AbGroup.zero_closed[of "Gs!k"] by (cases "k = n") auto qed thus ?thesis by fast qed ultimately have "xgs \ listset Gs" using listsetI_nth by fast moreover from xgs_def assms(3) have "x = sum_list xgs" using sum_list_update[of n "replicate (length Gs) 0" x] nth_replicate sum_list_replicate0 by simp ultimately show "(\Gs\x) = xgs" using assms(1,2) x xgs_def AbGroup_inner_dirsum_el_decomp_eq by fast qed lemma AbGroup_inner_dirsum_el_decomp_nth_in_nth : "\ \G \ set Gs. AbGroup G; add_independentS Gs; k < length Gs; n < length Gs; x \ Gs!n \ \ (\Gs\k) x = (if k = n then x else 0)" using AbGroup_nth_subset_inner_dirsum AbGroup_inner_dirsum_el_decomp_in_nth[of Gs n x] by auto lemma AbGroup_inner_dirsum_el_decomp_nth_id_on_nth : "\ \G \ set Gs. AbGroup G; add_independentS Gs; n < length Gs; x \ Gs!n \ \ (\Gs\n) x = x" using AbGroup_inner_dirsum_el_decomp_nth_in_nth by fastforce lemma AbGroup_inner_dirsum_el_decomp_nth_onto_nth : assumes "\G \ set Gs. AbGroup G" "add_independentS Gs" "n < length Gs" shows "(\Gs\n) ` (\G\Gs. G) = Gs!n" proof from assms show "(\Gs\n) ` (\G\Gs. G) \ Gs!n" using AbGroup_nth_subset_inner_dirsum[of Gs n] AbGroup_inner_dirsum_el_decomp_nth_id_on_nth[of Gs n] by force from assms show "(\Gs\n) ` (\G\Gs. G) \ Gs!n" using AbGroup_inner_dirsum_el_decompI listset_length listsetD_nth by fastforce qed lemma AbGroup_inner_dirsum_subset_proj_eq_0 : assumes "Gs \ []" "\G \ set Gs. AbGroup G" "add_independentS Gs" "X \ (\G\Gs. G)" "\i < length Gs. (\Gs\i) ` X = 0" shows "X = 0" proof- have "X \ 0" proof fix x assume x: "x \ X" with assms(2-4) have "x = (\i=0..< length Gs. (\Gs\i) x)" using AbGroup_inner_dirsum_el_decompI sum_list_sum_nth[of "(\Gs\x)"] AbGroup_length_inner_dirsum_el_decomp by fastforce moreover from x assms(5) have "\iGs\i) x = 0" by auto ultimately show "x \ 0" by simp qed moreover from assms(1,5) have "X \ {}" by auto ultimately show ?thesis by auto qed lemma GroupEnd_inner_dirsum_el_decomp_nth : assumes "\G \ set Gs. AbGroup G" "add_independentS Gs" "n < length Gs" shows "GroupEnd (\G\Gs. G) (\Gs\n)" proof (intro_locales) from assms(1) show grp: "Group (\G\Gs. G)" using AbGroup_inner_dirsum AbGroup.axioms by fast show "GroupHom_axioms (\G\Gs. G) \Gs\n" proof show "supp (\Gs\n) \ (\G\Gs. G)" using supp_restrict0 by fast next fix x y assume xy: "x \ (\G\Gs. G)" "y \ (\G\Gs. G)" with assms(1,2) have "(\Gs\(x+y)) = [x+y. (x,y)\zip (\Gs\x) (\Gs\y)]" using AbGroup_inner_dirsum_el_decomp_plus by fast hence "(\Gs\(x+y)) = map (case_prod (+)) (zip (\Gs\x) (\Gs\y))" using concat_map_split_eq_map_split_zip by simp moreover from assms xy have "n < length (\Gs\x)" "n < length (\Gs\y)" "n < length (zip (\Gs\x) (\Gs\y))" using AbGroup_length_inner_dirsum_el_decomp[of Gs x] AbGroup_length_inner_dirsum_el_decomp[of Gs y] by auto ultimately show "(\Gs\n) (x + y) = (\Gs\n) x + (\Gs\n) y" using xy assms(1) AbGroup_inner_dirsum AbGroup.add_closed[of "\G\Gs. G"] by auto qed show "GroupEnd_axioms (\G\Gs. G) \Gs\n" using assms AbGroup_inner_dirsum_el_decomp_nth_onto_nth AbGroup_nth_subset_inner_dirsum by unfold_locales fast qed subsection \Rings\ subsubsection \Preliminaries\ lemma (in ring_1) map_times_neg1_eq_map_uminus : "[(-1)*r. r\rs] = [-r. r\rs]" using map_eq_conv by simp subsubsection \Locale and basic facts\ text \Define a \Ring1\ to be a multiplicatively closed additive subgroup of @{term UNIV} for the \ring_1\ class.\ (* Don't need to use AbGroup for R because ring_1 already assumes add_commute *) locale Ring1 = Group R for R :: "'r::ring_1 set" + assumes one_closed : "1 \ R" and mult_closed: "\r s. r \ R \ s \ R \ r * s \ R" begin lemma AbGroup : "AbGroup R" using Ring1_axioms Ring1.axioms(1) AbGroup.intro by fast lemmas zero_closed = zero_closed lemmas add_closed = add_closed lemmas neg_closed = neg_closed lemmas diff_closed = diff_closed lemmas zip_add_closed = zip_add_closed lemmas sum_closed = AbGroup.sum_closed[OF AbGroup] lemmas sum_list_closed = sum_list_closed lemmas sum_list_closed_prod = sum_list_closed_prod lemmas list_diff_closed = list_diff_closed abbreviation Subring1 :: "'r set \ bool" where "Subring1 S \ Ring1 S \ S \ R" lemma Subring1D1 : "Subring1 S \ Ring1 S" by fast end (* context Ring1 *) lemma (in ring_1) full_Ring1 : "Ring1 UNIV" by unfold_locales auto subsection \The group ring\ subsubsection \Definition and basic facts\ text \ Realize the group ring as the set of almost-every-zero functions from group to ring. One can recover the usual notion of group ring element by considering such a function to send group elements to their coefficients. Here the codomain of such functions is not restricted to some \Ring1\ subset since we will not be interested in having the ability to change the ring of scalars for a group ring. \ context Group begin abbreviation group_ring :: "('a::zero, 'g) aezfun set" where "group_ring \ aezfun_setspan G" lemmas group_ringD = aezfun_setspan_def[of G] lemma RG_one_closed : "(1::('r::zero_neq_one,'g) aezfun) \ group_ring" proof- have "supp (aezfun (1::('r,'g) aezfun)) \ G" using supp_aezfun1 zeroS_closed by fast thus ?thesis using group_ringD by fast qed lemma RG_zero_closed : "(0::('r::zero,'g) aezfun) \ group_ring" proof- have "aezfun (0::('r,'g) aezfun) = (0::'g\'r)" using zero_aezfun.rep_eq by fast hence "supp (aezfun (0::('r,'g) aezfun)) = supp (0::'g\'r)" by simp moreover have "supp (0::'g\'r) \ G" using supp_zerofun_subset_any by fast ultimately show ?thesis using group_ringD by fast qed lemma RG_n0 : "group_ring \ (0::('r::zero_neq_one, 'g) aezfun set)" using RG_one_closed zero_neq_one by force lemma RG_mult_closed : defines RG: "RG \ group_ring :: ('r::ring_1, 'g) aezfun set" shows "x \ RG \ y \ RG \ x * y \ RG" using RG supp_aezfun_times[of x y] set_plus_closed[of "supp (aezfun x)" "supp (aezfun y)"] group_ringD by blast lemma Ring1_RG : defines RG: "RG \ group_ring :: ('r::ring_1, 'g) aezfun set" shows "Ring1 RG" proof from RG show "RG \ {}" "1 \ RG" "\x y. x \ RG \ y \ RG \ x * y \ RG" using RG_one_closed RG_mult_closed by auto next fix x y assume "x \ RG" "y \ RG" with RG show "x - y \ RG" using supp_aezfun_diff[of x y] group_ringD by blast qed lemma RG_aezdeltafun_closed : defines RG: "RG \ group_ring :: ('r::ring_1, 'g) aezfun set" assumes "g \ G" shows "r \\ g \ RG" proof- have supp: "supp (aezfun (r \\ g)) = supp (r \ g)" using aezdeltafun arg_cong[of _ _ "supp"] by fast have "supp (aezfun (r \\ g)) \ G" proof (cases "r = 0") case True with supp show ?thesis using supp_delta0fun by fast next case False with assms supp show ?thesis using supp_deltafun[of r g] by fast qed with RG show ?thesis using group_ringD by fast qed lemma RG_aezdelta0fun_closed : "(r::'r::ring_1) \\ 0 \ group_ring" using zero_closed RG_aezdeltafun_closed[of 0] by fast lemma RG_sum_list_rddg_closed : defines RG: "RG \ group_ring :: ('r::ring_1, 'g) aezfun set" assumes "set (map snd rgs) \ G" shows "(\(r,g)\rgs. r \\ g) \ RG" proof (rule Ring1.sum_list_closed_prod) from RG show "Ring1 RG" using Ring1_RG by fast from assms show "set (map (case_prod aezdeltafun) rgs) \ RG" using RG_aezdeltafun_closed by fastforce qed lemmas RG_el_decomp_aezdeltafun = aezfun_setspan_el_decomp_aezdeltafun[of _ G] lemma Subgroup_imp_Subring : fixes H :: "'g set" and FH :: "('r::ring_1,'g) aezfun set" and FG :: "('r,'g) aezfun set" defines "FH \ Group.group_ring H" and "FG \ group_ring" shows "Subgroup H \ Ring1.Subring1 FG FH" using assms Group.Ring1_RG Group.RG_el_decomp_aezdeltafun RG_sum_list_rddg_closed by fast end (* context Group *) lemma (in FinGroup) group_ring_spanning_set : "\gs. distinct gs \ set gs = G \ (\f\ (group_ring :: ('b::semiring_1, 'g) aezfun set). \bs. length bs = length gs \ f = (\(b,g)\zip bs gs. (b \\ 0) * (1 \\ g)) )" using finite aezfun_common_supp_spanning_set_decomp[of G] group_ringD by fast subsubsection \Projecting almost-everywhere-zero functions onto a group ring\ context Group begin abbreviation "RG_proj \ aezfun_setspan_proj G" lemmas RG_proj_in_RG = aezfun_setspan_proj_in_setspan lemmas RG_proj_sum_list_prod = aezfun_setspan_proj_sum_list_prod[of G] lemma RG_proj_mult_leftdelta' : fixes r s :: "'r::{comm_monoid_add,mult_zero}" shows "g \ G \ RG_proj (r \\ g * (s \\ g')) = r \\ g * RG_proj (s \\ g')" using add_closed add_closed_inverse_right times_aezdeltafun_aezdeltafun[of r g] aezfun_setspan_proj_aezdeltafun[of G "r*s"] aezfun_setspan_proj_aezdeltafun[of G s] by simp lemma RG_proj_mult_leftdelta : fixes r :: "'r::semiring_1" assumes "g \ G" shows "RG_proj ((r \\ g) * x) = r \\ g * RG_proj x" proof- from aezfun_decomp_aezdeltafun obtain rgs where rgs: "x = (\(s,h)\rgs. s \\ h)" using RG_el_decomp_aezdeltafun by fast hence "RG_proj ((r \\ g) * x) = (\(s,h)\rgs. RG_proj ((r \\ g) * (s \\ h)))" using sum_list_const_mult_prod[of "r \\ g" "\s h. s \\ h"] RG_proj_sum_list_prod by simp also from assms rgs have "\ = (r \\ g) * RG_proj x" using RG_proj_mult_leftdelta'[of g r] sum_list_const_mult_prod[of "r \\ g" "\s h. RG_proj (s \\ h)"] RG_proj_sum_list_prod[of "\s h. s \\ h" rgs] by simp finally show ?thesis by fast qed lemma RG_proj_mult_rightdelta' : fixes r s :: "'r::{comm_monoid_add,mult_zero}" assumes "g' \ G" shows "RG_proj (r \\ g * (s \\ g')) = RG_proj (r \\ g) * (s \\ g')" using assms times_aezdeltafun_aezdeltafun[of r g] aezfun_setspan_proj_aezdeltafun[of G "r*s"] add_closed add_closed_inverse_left aezfun_setspan_proj_aezdeltafun[of G r] by simp lemma RG_proj_mult_rightdelta : fixes r :: "'r::semiring_1" assumes "g \ G" shows "RG_proj (x * (r \\ g)) = (RG_proj x) * (r \\ g)" proof- from aezfun_decomp_aezdeltafun obtain rgs where rgs: "x = (\(s,h)\rgs. s \\ h)" using RG_el_decomp_aezdeltafun by fast hence "RG_proj (x * (r \\ g)) = (\(s,h)\rgs. RG_proj ((s \\ h) * (r \\ g)))" using sum_list_mult_const_prod[of "\s h. s \\ h" rgs] RG_proj_sum_list_prod by simp with assms rgs show ?thesis using RG_proj_mult_rightdelta'[of g _ _ r] sum_list_prod_cong[of rgs "\s h. RG_proj ((s \\ h) * (r \\ g))" "\s h. RG_proj (s \\ h) * (r \\ g)" ] sum_list_mult_const_prod[of "\s h. RG_proj (s \\ h)" rgs] RG_proj_sum_list_prod[of "\s h. s \\ h" rgs] sum_list_mult_const_prod[of "\s h. RG_proj (s \\ h)" rgs "r \\ g"] RG_proj_sum_list_prod[of "\s h. s \\ h" rgs] by simp qed lemma RG_proj_mult_right : "x \ (group_ring :: ('r::ring_1, 'g) aezfun set) \ RG_proj (y * x) = RG_proj y * x" using RG_el_decomp_aezdeltafun sum_list_const_mult_prod[of y "\r g. r \\ g"] RG_proj_sum_list_prod[of "\r g. y * (r \\ g)"] RG_proj_mult_rightdelta[of _ y] sum_list_prod_cong[ of _ "\r g. RG_proj (y * (r \\ g))" "\r g. RG_proj y * (r \\ g)" ] sum_list_const_mult_prod[of "RG_proj y" "\r g. r \\ g"] by fastforce end (* context Group *) section \Modules\ subsection \Locales and basic facts\ subsubsection \Locales\ locale scalar_mult = fixes smult :: "'r::ring_1 \ 'm::ab_group_add \ 'm" (infixr "\" 70) locale R_scalar_mult = scalar_mult smult + Ring1 R for R :: "'r::ring_1 set" and smult :: "'r \ 'm::ab_group_add \ 'm" (infixr "\" 70) lemma (in scalar_mult) R_scalar_mult : "R_scalar_mult UNIV" using full_Ring1 R_scalar_mult.intro by fast lemma (in R_scalar_mult) Ring1 : "Ring1 R" .. locale RModule = R_scalars?: R_scalar_mult R smult + VecGroup?: Group M for R :: "'r::ring_1 set" and smult :: "'r \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" + assumes smult_closed : "\ r \ R; m \ M \ \ r \ m \ M" and smult_distrib_left [simp] : "\ r \ R; m \ M; n \ M \ \ r \ (m + n) = r \ m + r \ n" and smult_distrib_right [simp] : "\ r \ R; s \ R; m \ M \ \ (r + s) \ m = r \ m + s \ m" and smult_assoc [simp] : "\ r \ R; s \ R; m \ M \ \ r \ s \ m = (r * s) \ m" and one_smult [simp] : "m \ M \ 1 \ m = m" lemmas RModuleI = RModule.intro[OF R_scalar_mult.intro] locale Module = RModule UNIV smult M for smult :: "'r::ring_1 \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" lemmas ModuleI = RModuleI[of UNIV, OF full_Ring1, THEN Module.intro] subsubsection \Basic facts\ lemma trivial_RModule : fixes smult :: "'r::ring_1 \ 'm::ab_group_add \ 'm" (infixr "\" 70) assumes "Ring1 R" "\r\R. smult r (0::'m::ab_group_add) = 0" shows "RModule R smult (0::'m set)" proof (rule RModuleI, rule assms(1), rule trivial_Group, unfold_locales) define Z where "Z = (0::'m set)" fix r s m n assume rsmn: "r \ R" "s \ R" "m \ Z" "n \ Z" from rsmn(1,3) Z_def assms(2) show "r \ m \ Z" by simp from rsmn(1,3,4) Z_def assms(2) show "r \ (m+n) = r \ m + r \ n" by simp from rsmn(1-3) Z_def assms show "(r + s) \ m = r \ m + s \ m" using Ring1.add_closed by auto from rsmn(1-3) Z_def assms show "r \ (s \ m) = (r*s) \ m" using Ring1.mult_closed by auto next define Z where "Z = (0::'m set)" fix m assume "m \ Z" with Z_def assms show "1 \ m = m" using Ring1.one_closed by auto qed context RModule begin abbreviation RSubmodule :: "'m set \ bool" where "RSubmodule N \ RModule R smult N \ N \ M" lemma Group : "Group M" using RModule_axioms RModule.axioms(2) by fast lemma Subgroup_RSubmodule : "RSubmodule N \ Subgroup N" using RModule.Group by fast lemma AbGroup : "AbGroup M" using AbGroup.intro Group by fast lemmas zero_closed = zero_closed lemmas diff_closed = diff_closed lemmas set_plus_closed = set_plus_closed lemmas sum_closed = AbGroup.sum_closed[OF AbGroup] lemma map_smult_closed : "r \ R \ set ms \ M \ set (map ((\) r) ms) \ M" using smult_closed by (induct ms) auto lemma zero_smult : "m \ M \ 0 \ m = 0" using R_scalars.zero_closed smult_distrib_right[of 0] add_left_imp_eq by simp lemma smult_zero : "r \ R \ r \ 0 = 0" using zero_closed smult_distrib_left[of r 0] by simp lemma neg_smult : "r \ R \ m \ M \ (-r) \ m = - (r \ m)" using R_scalars.neg_closed smult_distrib_right[of r "-r" m] zero_smult minus_unique[of "r \ m"] by simp lemma neg_eq_neg1_smult : "m \ M \ (-1) \ m = - m" using one_closed neg_smult one_smult by fastforce lemma smult_neg : "r \ R \ m \ M \ r \ (- m) = - (r \ m)" using neg_eq_neg1_smult one_closed R_scalars.neg_closed smult_assoc[of r "- 1"] smult_closed by force lemma smult_distrib_left_diff : "\ r \ R; m \ M; n \ M \ \ r \ (m - n) = r \ m - r \ n" using neg_closed smult_distrib_left[of r m "-n"] smult_neg by (simp add: algebra_simps) lemma smult_distrib_right_diff : "\ r \ R; s \ R; m \ M \ \ (r - s) \ m = r \ m - s \ m" using R_scalars.neg_closed smult_distrib_right[of r "-s"] neg_smult by (simp add: algebra_simps) lemma smult_sum_distrib : assumes "r \ R" shows "finite A \ f ` A \ M \ r \ (\a\A. f a) = (\a\A. r \ f a)" proof (induct set: finite) case empty from assms show ?case using smult_zero by simp next case (insert a A) with assms show ?case using sum_closed[of A] by simp qed lemma sum_smult_distrib : assumes "m \ M" shows "finite A \ f ` A \ R \ (\a\A. f a) \ m = (\a\A. (f a) \ m)" proof (induct set: finite) case empty from assms show ?case using zero_smult by simp next case (insert a A) with assms show ?case using R_scalars.sum_closed[of A] by simp qed lemma smult_sum_list_distrib : "r \ R \ set ms \ M \ r \ (sum_list ms) = (\m\ms. r \ m)" using smult_zero sum_list_closed[of id] by (induct ms) auto lemma sum_list_prod_map_smult_distrib : "m \ M \ set (map (case_prod f) xys) \ R \ (\(x,y)\xys. f x y) \ m = (\(x,y)\xys. f x y \ m)" using zero_smult R_scalars.sum_list_closed_prod[of f] by (induct xys) auto lemma RSubmoduleI : assumes "Subgroup N" "\r n. r \ R \ n \ N \ r \ n \ N" shows "RSubmodule N" proof show "RModule R smult N" proof (intro_locales, rule SubgroupD1[OF assms(1)], unfold_locales) from assms(2) show "\r m. r \ R \ m \ N \ r \ m \ N" by fast from assms(1) show "\r m n. \ r \ R; m \ N; n \ N \ \ r \ (m + n) = r \ m + r \ n" using smult_distrib_left by blast from assms(1) show "\r s m. \ r \ R; s \ R; m \ N \ \ (r + s) \ m = r \ m + s \ m" using smult_distrib_right by blast from assms(1) show "\r s m. \ r \ R; s \ R; m \ N \ \ r \ s \ m = (r * s) \ m" using smult_assoc by blast from assms(1) show "\m. m \ N \ 1 \ m = m" using one_smult by blast qed from assms(1) show "N \ M" by fast qed end (* context RModule *) lemma (in R_scalar_mult) listset_RModule_Rsmult_closed : "\ \M\set Ms. RModule R smult M; r \ R; ms \ listset Ms \ \ [r \ m. m\ms] \ listset Ms" proof- have "\ length ms = length Ms; ms \ listset Ms; \M\set Ms. RModule R smult M; r \ R \ \ [r \ m. m\ms] \ listset Ms" proof (induct ms Ms rule: list_induct2) case (Cons m ms M Ms) thus ?case using listset_ConsD[of m] RModule.smult_closed listset_ConsI[of "r \ m" M] by fastforce qed simp thus "\ \M\set Ms. RModule R smult M; r \ R; ms \ listset Ms \ \ [r \ m. m\ms] \ listset Ms" using listset_length[of ms Ms] by fast qed context Module begin abbreviation Submodule :: "'m set \ bool" where "Submodule \ RModule.RSubmodule UNIV smult M" lemmas AbGroup = AbGroup lemmas SubmoduleI = RSubmoduleI end (* context Module *) subsubsection \Module and submodule instances\ lemma (in R_scalar_mult) trivial_RModule : "(\r. r \ R \ r \ 0 = 0) \ RModule R smult 0" using trivial_Group add_closed mult_closed one_closed by unfold_locales auto context RModule begin lemma trivial_RSubmodule : "RSubmodule 0" using zeroS_closed smult_zero trivial_RModule by fast lemma RSubmodule_set_plus : assumes "RSubmodule L" "RSubmodule N" shows "RSubmodule (L + N)" proof (rule RSubmoduleI) from assms have "Group (L + N)" using RModule.AbGroup AbGroup_set_plus[of L N] AbGroup.axioms by fast moreover from assms have "L + N \ M" using Group Group.set_plus_closed by auto ultimately show "Subgroup (L + N)" by fast next fix r x assume rx: "r \ R" "x \ L + N" from rx(2) obtain m n where mn: "m \ L" "n \ N" "x = m + n" using set_plus_def[of L N] by fast with assms rx(1) show "r \ x \ L + N" using RModule.smult_closed[of R smult L] RModule.smult_closed[of R smult N] smult_distrib_left set_plus_def by fast qed lemma RSubmodule_sum_list : "(\N\set Ns. RSubmodule N) \ RSubmodule (\N\Ns. N)" using trivial_RSubmodule RSubmodule_set_plus by (induct Ns) auto lemma RSubmodule_inner_dirsum : assumes "(\N\set Ns. RSubmodule N)" shows "RSubmodule (\N\Ns. N)" proof (cases "add_independentS Ns") case True with assms show ?thesis using RSubmodule_sum_list inner_dirsumD by fastforce next case False thus ?thesis using inner_dirsumD2[of Ns] trivial_RSubmodule by simp qed lemma RModule_inner_dirsum : "(\N\set Ns. RSubmodule N) \ RModule R smult (\N\Ns. N)" using RSubmodule_inner_dirsum by fast lemma SModule_restrict_scalars : assumes "Subring1 S" shows "RModule S smult M" proof (rule RModuleI, rule Subring1D1[OF assms], rule Group, unfold_locales) from assms show "\r m. r \ S \ m \ M \ r \ m \ M" "\r m n. r \ S \ m \ M \ n \ M \ r \ (m + n) = r \ m + r \ n" "\m. m \ M \ 1 \ m = m" using smult_closed smult_distrib_left by auto from assms show "\r s m. r \ S \ s \ S \ m \ M \ (r + s) \ m = r \ m + s \ m" using Ring1.add_closed smult_distrib_right by fast from assms show "\r s m. r \ S \ s \ S \ m \ M \ r \ s \ m = (r * s) \ m" using Ring1.mult_closed smult_assoc by fast qed end (* context RModule *) subsection \Linear algebra in modules\ subsubsection \Linear combinations: \lincomb\\ context scalar_mult begin definition lincomb :: "'r list \ 'm list \ 'm" (infix "\\" 70) where "rs \\ ms = (\(r,m)\zip rs ms. r \ m)" text \Note: \zip\ will truncate if lengths of coefficient and vector lists differ.\ lemma lincomb_Nil : "rs = [] \ ms = [] \ rs \\ ms = 0" unfolding lincomb_def by auto lemma lincomb_singles : "[a] \\ [m] = a \ m" using lincomb_def by simp lemma lincomb_Cons : "(r # rs) \\ (m # ms) = r \ m + rs \\ ms" unfolding lincomb_def by simp lemma lincomb_append : "length rs = length ms \ (rs@ss) \\ (ms@ns) = rs \\ ms + ss \\ ns" unfolding lincomb_def by simp lemma lincomb_append_left : "(rs @ ss) \\ ms = rs \\ ms + ss \\ drop (length rs) ms" using zip_append_left[of rs ss ms] unfolding lincomb_def by simp lemma lincomb_append_right : "rs \\ (ms@ns) = rs \\ ms + (drop (length ms) rs) \\ ns" using zip_append_right[of rs ms] unfolding lincomb_def by simp lemma lincomb_conv_take_right : "rs \\ ms = rs \\ take (length rs) ms" using lincomb_Nil lincomb_Cons by (induct rs ms rule: list_induct2') auto end (* context scalar_mult *) context RModule begin lemmas lincomb_Nil = lincomb_Nil lemmas lincomb_Cons = lincomb_Cons lemma lincomb_closed : "set rs \ R \ set ms \ M \ rs \\ ms \ M" proof (induct ms arbitrary: rs) case Nil show ?case using lincomb_Nil zero_closed by simp next case (Cons m ms) hence Cons1: "\rs. set rs \ R \ rs \\ ms \ M" "m \ M" "set rs \ R" by auto show "rs \\ (m#ms) \ M" proof (cases rs) case Nil thus ?thesis using lincomb_Nil zero_closed by simp next case Cons with Cons1 show ?thesis using lincomb_Cons smult_closed add_closed by fastforce qed qed lemma smult_lincomb : "\ set rs \ R; s \ R; set ms \ M \ \ s \ (rs \\ ms) = [s*r. r\rs] \\ ms" using lincomb_Nil smult_zero lincomb_Cons smult_closed lincomb_closed by (induct rs ms rule: list_induct2') auto lemma neg_lincomb : "set rs \ R \ set ms \ M \ - (rs \\ ms) = [-r. r\rs] \\ ms" using lincomb_closed neg_eq_neg1_smult one_closed R_scalars.neg_closed[of 1] smult_lincomb[of rs "- 1"] map_times_neg1_eq_map_uminus by auto lemma lincomb_sum_left : "\ set rs \ R; set ss \ R; set ms \ M; length rs \ length ss \ \ [r + s. (r,s)\zip rs ss] \\ ms = rs \\ ms + (take (length rs) ss) \\ ms" proof (induct rs ss arbitrary: ms rule: list_induct2') case 1 show ?case using lincomb_Nil by simp next case (2 r rs) show "\ms. length (r#rs) \ length [] \ [a + b. (a,b)\zip (r#rs) []] \\ ms = (r#rs) \\ ms + (take (length (r#rs)) []) \\ ms" by simp next case 3 show ?case using lincomb_Nil by simp next case (4 r rs s ss) thus "[a+b. (a,b)\zip (r#rs) (s#ss)] \\ ms = (r#rs) \\ ms + (take (length (r#rs)) (s#ss)) \\ ms" using lincomb_Nil lincomb_Cons by (cases ms) auto qed lemma lincomb_sum : assumes "set rs \ R" "set ss \ R" "set ms \ M" "length rs \ length ss" shows "rs \\ ms + ss \\ ms = ([a + b. (a,b)\zip rs ss] @ (drop (length rs) ss)) \\ ms" proof- define zs fss bss where "zs = [a + b. (a,b)\zip rs ss]" and "fss = take (length rs) ss" and "bss = drop (length rs) ss" from assms(4) zs_def fss_def have "length zs = length rs" "length fss = length rs" using length_concat_map_split_zip[of "\a b. a + b" rs] by auto hence "(zs @ bss) \\ ms = rs \\ ms + (fss @ bss) \\ ms" using assms(1,2,3) zs_def fss_def lincomb_sum_left lincomb_append_left by simp thus ?thesis using fss_def bss_def zs_def by simp qed lemma lincomb_diff_left : "\ set rs \ R; set ss \ R; set ms \ M; length rs \ length ss \ \ [r - s. (r,s)\zip rs ss] \\ ms = rs \\ ms - (take (length rs) ss) \\ ms" proof (induct rs ss arbitrary: ms rule: list_induct2') case 1 show ?case using lincomb_Nil by simp next case (2 r rs) show "\ms. length (r#rs) \ length [] \ [a - b. (a,b)\zip (r#rs) []] \\ ms = (r#rs) \\ ms - (take (length (r#rs)) []) \\ ms" by simp next case 3 show ?case using lincomb_Nil by simp next case (4 r rs s ss) thus "[a-b. (a,b)\zip (r#rs) (s#ss)] \\ ms = (r#rs) \\ ms - (take (length (r#rs)) (s#ss)) \\ ms" using lincomb_Nil lincomb_Cons smult_distrib_right_diff by (cases ms) auto qed lemma lincomb_replicate_left : "r \ R \ set ms \ M \ (replicate k r) \\ ms = r \ ( \m\(take k ms). m )" proof (induct k arbitrary: ms) case 0 thus ?case using lincomb_Nil smult_zero by simp next case (Suc k) show ?case proof (cases ms) case Nil with Suc(2) show ?thesis using lincomb_Nil smult_zero by simp next case (Cons m ms) with Suc show ?thesis using lincomb_Cons set_take_subset[of k ms] sum_list_closed[of id] by auto qed qed lemma lincomb_replicate0_left : "set ms \ M \ (replicate k 0) \\ ms = 0" proof- assume ms: "set ms \ M" hence "(replicate k 0) \\ ms = 0 \ (\m\(take k ms). m)" using R_scalars.zero_closed lincomb_replicate_left by fast moreover from ms have "(\m\(take k ms). m) \ M" using set_take_subset sum_list_closed by fastforce ultimately show "(replicate k 0) \\ ms = 0" using zero_smult by simp qed lemma lincomb_0coeffs : "set ms \ M \ \s\set rs. s = 0 \ rs \\ ms = 0" using lincomb_Nil lincomb_Cons zero_smult by (induct rs ms rule: list_induct2') auto lemma delta_scalars_lincomb_eq_nth : "set ms \ M \ n < length ms \ ((replicate (length ms) 0)[n := 1]) \\ ms = ms!n" proof (induct ms arbitrary: n) case (Cons m ms) thus ?case using lincomb_Cons lincomb_replicate0_left zero_smult by (cases n) auto qed simp lemma lincomb_obtain_same_length_Rcoeffs : "set rs \ R \ set ms \ M \ \ss. set ss \ R \ length ss = length ms \ take (length rs) ss = take (length ms) rs \ rs \\ ms = ss \\ ms" proof (induct rs ms rule: list_induct2') case 1 show ?case using lincomb_Nil by simp next case 2 thus ?case using lincomb_Nil by simp next case (3 m ms) define ss where "ss = replicate (Suc (length ms)) (0::'r)" from 3(2) ss_def have "set ss \ R" "length ss = length (m#ms)" "[] \\ (m#ms) = ss \\ (m#ms)" using R_scalars.zero_closed lincomb_Nil lincomb_replicate0_left[of "m#ms" "Suc (length ms)"] by auto thus ?case by auto next case (4 r rs m ms) from this obtain ss where ss: "set ss \ R" "length ss = length ms" "take (length rs) ss = take (length ms) rs" "rs \\ ms = ss \\ ms" by auto from 4(2) ss have "set (r#ss) \ R" "length (r#ss) = length (m#ms)" "take (length (r#rs)) (r#ss) = take (length (m#ms)) (r#rs)" "(r#rs) \\ (m#ms) = (r#ss) \\ (m#ms)" using lincomb_Cons by auto thus ?case by fast qed lemma lincomb_concat : "list_all2 (\rs ms. length rs = length ms) rss mss \ (concat rss) \\ (concat mss) = (\(rs,ms)\zip rss mss. rs \\ ms)" using lincomb_Nil lincomb_append by (induct rss mss rule: list_induct2') auto lemma lincomb_snoc0 : "set ms \ M \ (as@[0]) \\ ms = as \\ ms" using lincomb_append_left set_drop_subset lincomb_replicate0_left[of _ 1] by fastforce lemma lincomb_strip_while_0coeffs : assumes "set ms \ M" shows "(strip_while ((=) 0) as) \\ ms = as \\ ms" proof (induct as rule: rev_induct) case (snoc a as) hence caseassm: "strip_while ((=) 0) as \\ ms = as \\ ms" by fast show ?case proof (cases "a = 0") case True moreover with assms have "(as@[a]) \\ ms = as \\ ms" using lincomb_snoc0 by fast ultimately show "strip_while ((=) 0) (as @ [a]) \\ ms = (as@[a]) \\ ms" using caseassm by simp qed simp qed simp end (* context RModule *) lemmas (in Module) lincomb_obtain_same_length_coeffs = lincomb_obtain_same_length_Rcoeffs lemmas (in Module) lincomb_concat = lincomb_concat subsubsection \Spanning: \RSpan\ and \Span\\ context R_scalar_mult begin primrec RSpan :: "'m list \ 'm set" where "RSpan [] = 0" | "RSpan (m#ms) = { r \ m | r. r \ R } + RSpan ms" lemma RSpan_single : "RSpan [m] = { r \ m | r. r \ R }" using add_0_right[of "{ r \ m | r. r \ R }"] by simp lemma RSpan_Cons : "RSpan (m#ms) = RSpan [m] + RSpan ms" using RSpan_single by simp lemma in_RSpan_obtain_same_length_coeffs : "n \ RSpan ms \ \rs. set rs \ R \ length rs = length ms \ n = rs \\ ms" proof (induct ms arbitrary: n) case Nil hence "n = 0" by simp thus "\rs. set rs \ R \ length rs = length [] \ n = rs \\ []" using lincomb_Nil by simp next case (Cons m ms) from this obtain r rs where "set (r#rs) \ R" "length (r#rs) = length (m#ms)" "n = (r#rs) \\ (m#ms)" using set_plus_def[of _ "RSpan ms"] lincomb_Cons by fastforce thus "\rs. set rs \ R \ length rs = length (m#ms) \ n = rs \\ (m#ms)" by fast qed lemma in_RSpan_Cons_obtain_same_length_coeffs : "n \ RSpan (m#ms) \ \r rs. set (r#rs) \ R \ length rs = length ms \ n = r \ m + rs \\ ms" proof- assume "n \ RSpan (m#ms)" from this obtain x y where "x \ RSpan [m]" "y \ RSpan ms" "n = x + y" using RSpan_Cons set_plus_def[of "RSpan [m]"] by auto thus "\r rs. set (r # rs) \ R \ length rs = length ms \ n = r \ m + rs \\ ms" using RSpan_single in_RSpan_obtain_same_length_coeffs[of y ms] by auto qed lemma RSpanD_lincomb : "RSpan ms = { rs \\ ms | rs. set rs \ R \ length rs = length ms}" proof show "RSpan ms \ {rs \\ ms |rs. set rs \ R \ length rs = length ms}" using in_RSpan_obtain_same_length_coeffs by fast show "{rs \\ ms |rs. set rs \ R \ length rs = length ms} \ RSpan ms" proof fix x assume "x \ {rs \\ ms |rs. set rs \ R \ length rs = length ms}" from this obtain rs where rs: "set rs \ R" "length rs = length ms" "x = rs \\ ms" by fast from rs(2) have "set rs \ R \ rs \\ ms \ RSpan ms" using lincomb_Nil lincomb_Cons by (induct rs ms rule: list_induct2) auto with rs(1,3) show "x \ RSpan ms" by fast qed qed lemma RSpan_append : "RSpan (ms @ ns) = RSpan ms + RSpan ns" proof (induct ms) case Nil show ?case using add_0_left[of "RSpan ns"] by simp next case (Cons m ms) thus ?case using RSpan_Cons[of m "ms@ns"] add.assoc by fastforce qed end (* context R_scalar_mult *) context scalar_mult begin abbreviation "Span \ R_scalar_mult.RSpan UNIV smult" lemmas Span_append = R_scalar_mult.RSpan_append[OF R_scalar_mult, of smult] lemmas SpanD_lincomb = R_scalar_mult.RSpanD_lincomb [OF R_scalar_mult, of smult] lemmas in_Span_obtain_same_length_coeffs = R_scalar_mult.in_RSpan_obtain_same_length_coeffs[ OF R_scalar_mult, of _ smult ] end (* context scalar_mult *) context RModule begin lemma RSpan_contains_spanset_single : "m \ M \ m \ RSpan [m]" using one_closed RSpan_single by fastforce lemma RSpan_single_nonzero : "m \ M \ m \ 0 \ RSpan [m] \ 0" using RSpan_contains_spanset_single by auto lemma Group_RSpan_single : assumes "m \ M" shows "Group (RSpan [m])" proof from assms show "RSpan [m] \ {}" using RSpan_contains_spanset_single by fast next fix x y assume "x \ RSpan [m]" "y \ RSpan [m]" from this obtain r s where rs: "r \ R" "x = r \ m" "s \ R" "y = s \ m" using RSpan_single by auto with assms have "x - y = (r - s) \ m" using smult_distrib_right_diff by simp with rs(1,3) show "x - y \ RSpan [m]" using R_scalars.diff_closed[of r s] RSpan_single[of m] by auto qed lemma Group_RSpan : "set ms \ M \ Group (RSpan ms)" proof (induct ms) case Nil show ?case using trivial_Group by simp next case (Cons m ms) hence "Group (RSpan [m])" "Group (RSpan ms)" using Group_RSpan_single[of m] by auto thus ?case using RSpan_Cons[of m ms] AbGroup.intro AbGroup_set_plus AbGroup.axioms(1) by fastforce qed lemma RSpanD_lincomb_arb_len_coeffs : "set ms \ M \ RSpan ms = { rs \\ ms | rs. set rs \ R }" proof show "RSpan ms \ { rs \\ ms | rs. set rs \ R }" using RSpanD_lincomb by fast show "set ms \ M \ RSpan ms \ { rs \\ ms | rs. set rs \ R }" proof (induct ms) case Nil show ?case using lincomb_Nil by auto next case (Cons m ms) show ?case proof fix x assume "x \ { rs \\ (m#ms) | rs. set rs \ R }" from this obtain rs where rs: "set rs \ R" "x = rs \\ (m#ms)" by fast with Cons show "x \ RSpan (m#ms)" using lincomb_Nil Group_RSpan[of "m#ms"] Group.zero_closed lincomb_Cons by (cases rs) auto qed qed qed lemma RSpanI_lincomb_arb_len_coeffs : "set rs \ R \ set ms \ M \ rs \\ ms \ RSpan ms" using RSpanD_lincomb_arb_len_coeffs by fast lemma RSpan_contains_RSpans_Cons_left : "set ms \ M \ RSpan [m] \ RSpan (m#ms)" using RSpan_Cons Group_RSpan AbGroup.intro AbGroup.subset_plus_left by fast lemma RSpan_contains_RSpans_Cons_right : "m \ M \ RSpan ms \ RSpan (m#ms)" using RSpan_Cons Group_RSpan_single AbGroup.intro AbGroup.subset_plus_right by fast lemma RSpan_contains_RSpans_append_left : "set ns \ M \ RSpan ms \ RSpan (ms@ns)" using RSpan_append Group_RSpan AbGroup.intro AbGroup.subset_plus_left by fast lemma RSpan_contains_spanset : "set ms \ M \ set ms \ RSpan ms" proof (induct ms) case Nil show ?case by simp next case (Cons m ms) thus ?case using RSpan_contains_spanset_single RSpan_contains_RSpans_Cons_left[of ms m] RSpan_contains_RSpans_Cons_right[of m ms] by auto qed lemma RSpan_contains_spanset_append_left : "set ms \ M \ set ns \ M \ set ms \ RSpan (ms@ns)" using RSpan_contains_spanset[of "ms@ns"] by simp lemma RSpan_contains_spanset_append_right : "set ms \ M \ set ns \ M \ set ns \ RSpan (ms@ns)" using RSpan_contains_spanset[of "ms@ns"] by simp lemma RSpan_zero_closed : "set ms \ M \ 0 \ RSpan ms" using Group_RSpan Group.zero_closed by fast lemma RSpan_single_closed : "m \ M \ RSpan [m] \ M" using RSpan_single smult_closed by auto lemma RSpan_closed : "set ms \ M \ RSpan ms \ M" proof (induct ms) case Nil show ?case using zero_closed by simp next case (Cons m ms) thus ?case using RSpan_single_closed RSpan_Cons Group Group.set_plus_closed[of M] by simp qed lemma RSpan_smult_closed : assumes "r \ R" "set ms \ M" "n \ RSpan ms" shows "r \ n \ RSpan ms" proof- from assms(2,3) obtain rs where rs: "set rs \ R" "n = rs \\ ms" using RSpanD_lincomb_arb_len_coeffs by fast with assms(1,2) show ?thesis using smult_lincomb[OF rs(1) assms(1,2)] mult_closed RSpanI_lincomb_arb_len_coeffs[of "[r*a. a\rs]" ms] by auto qed lemma RSpan_add_closed : assumes "set ms \ M" "n \ RSpan ms" "n' \ RSpan ms" shows "n + n' \ RSpan ms" proof- from assms (2,3) obtain rs ss where rs: "set rs \ R" "length rs = length ms" "n = rs \\ ms" and ss: "set ss \ R" "length ss = length ms" "n' = ss \\ ms" using RSpanD_lincomb by auto with assms(1) have "n + n' = [r + s. (r,s)\zip rs ss] \\ ms" using lincomb_sum_left by simp moreover from rs(1) ss(1) have "set [r + s. (r,s)\zip rs ss] \ R" using set_zip_leftD[of _ _ rs ss] set_zip_rightD[of _ _ rs ss] R_scalars.add_closed R_scalars.zip_add_closed by blast ultimately show "n + n' \ RSpan ms" using assms(1) RSpanI_lincomb_arb_len_coeffs by simp qed lemma RSpan_lincomb_closed : "\ set rs \ R; set ms \ M; set ns \ RSpan ms \ \ rs \\ ns \ RSpan ms" using lincomb_Nil RSpan_zero_closed lincomb_Cons RSpan_smult_closed RSpan_add_closed by (induct rs ns rule: list_induct2') auto lemma RSpanI : "set ms \ M \ M \ RSpan ms \ M = RSpan ms" using RSpan_closed by fast lemma RSpan_contains_RSpan_take : "set ms \ M \ RSpan (take k ms) \ RSpan ms" using append_take_drop_id set_drop_subset RSpan_contains_RSpans_append_left[of "drop k ms"] by fastforce lemma RSubmodule_RSpan_single : assumes "m \ M" shows "RSubmodule (RSpan [m])" proof (rule RSubmoduleI) from assms show "Subgroup (RSpan [m])" using Group_RSpan_single RSpan_closed[of "[m]"] by simp next fix r n assume rn: "r \ R" "n \ RSpan [m]" from rn(2) obtain s where s: "s \ R" "n = s \ m" using RSpan_single by fast with assms rn(1) have "r * s \ R" "r \ n = (r * s) \ m" using mult_closed by auto thus "r \ n \ RSpan [m]" using RSpan_single by fast qed lemma RSubmodule_RSpan : "set ms \ M \ RSubmodule (RSpan ms)" proof (induct ms) case Nil show ?case using trivial_RSubmodule by simp next case (Cons m ms) hence "RSubmodule (RSpan [m])" "RSubmodule (RSpan ms)" using RSubmodule_RSpan_single by auto thus ?case using RSpan_Cons RSubmodule_set_plus by simp qed lemma RSpan_RSpan_closed : "set ms \ M \ set ns \ RSpan ms \ RSpan ns \ RSpan ms" using RSpanD_lincomb[of ns] RSpan_lincomb_closed by auto lemma spanset_reduce_Cons : "set ms \ M \ m \ RSpan ms \ RSpan (m#ms) = RSpan ms" using RSpan_Cons RSpan_RSpan_closed[of ms "[m]"] RSpan_contains_RSpans_Cons_right[of m ms] RSubmodule_RSpan[of ms] RModule.set_plus_closed[of R smult "RSpan ms" "RSpan [m]" "RSpan ms"] by auto lemma RSpan_replace_hd : assumes "n \ M" "set ms \ M" "m \ RSpan (n # ms)" shows "RSpan (m # ms) \ RSpan (n # ms)" proof fix x assume "x \ RSpan (m # ms)" from this assms(3) obtain r rs s ss where r_rs: "r \ R" "set rs \ R" "length rs = length ms" "x = r \ m + rs \\ ms" and s_ss: "s \ R" "set ss \ R" "length ss = length ms" "m = s \ n + ss \\ ms" using in_RSpan_Cons_obtain_same_length_coeffs[of x m ms] in_RSpan_Cons_obtain_same_length_coeffs[of m n ms] by fastforce from r_rs(1) s_ss(2) have set1: "set [r*a. a\ss] \ R" using mult_closed by auto have "x = ((r * s) # [a + b. (a,b)\zip [r*a. a\ss] rs]) \\ (n # ms)" proof- from r_rs(2,3) s_ss(3) assms(2) have "[r*a. a\ss] \\ ms + rs \\ ms = [a + b. (a,b)\zip [r*a. a\ss] rs] \\ ms" using set1 lincomb_sum by simp moreover from assms(1,2) r_rs(1,2,4) s_ss(1,2,4) have "x = (r * s) \ n + ([r*a. a\ss] \\ ms + rs \\ ms)" using smult_closed lincomb_closed smult_lincomb mult_closed lincomb_sum by simp ultimately show ?thesis using lincomb_Cons by simp qed moreover have "set ((r * s) # [a + b. (a,b)\zip [r*a. a\ss] rs]) \ R" proof- from r_rs(2) have "set [a + b. (a,b)\zip [r*a. a\ss] rs] \ R" using set1 R_scalars.zip_add_closed by fast with r_rs(1) s_ss(1) show ?thesis using mult_closed by simp qed ultimately show "x \ RSpan (n # ms)" using assms(1,2) RSpanI_lincomb_arb_len_coeffs[of _ "n#ms"] by fastforce qed end (* context RModule *) lemmas (in scalar_mult) Span_Cons = R_scalar_mult.RSpan_Cons[OF R_scalar_mult, of smult] context Module begin lemmas SpanD_lincomb_arb_len_coeffs = RSpanD_lincomb_arb_len_coeffs lemmas SpanI = RSpanI lemmas SpanI_lincomb_arb_len_coeffs = RSpanI_lincomb_arb_len_coeffs lemmas Span_contains_Spans_Cons_right = RSpan_contains_RSpans_Cons_right lemmas Span_contains_spanset = RSpan_contains_spanset lemmas Span_contains_spanset_append_left = RSpan_contains_spanset_append_left lemmas Span_contains_spanset_append_right = RSpan_contains_spanset_append_right lemmas Span_closed = RSpan_closed lemmas Span_smult_closed = RSpan_smult_closed lemmas Span_contains_Span_take = RSpan_contains_RSpan_take lemmas Span_replace_hd = RSpan_replace_hd lemmas Submodule_Span = RSubmodule_RSpan end (* context Module *) subsubsection \Finitely generated modules\ context R_scalar_mult begin abbreviation "R_fingen M \ (\ms. set ms \ M \ RSpan ms = M)" text \ Similar to definition of \card\ for finite sets, we default \dim\ to 0 if no finite spanning set exists. Note that @{term "RSpan [] = 0"} implies that @{term "dim_R {0} = 0"}. \ definition dim_R :: "'m set \ nat" where "dim_R M = (if R_fingen M then ( LEAST n. \ms. length ms = n \ set ms \ M \ RSpan ms = M ) else 0)" lemma dim_R_nonzero : assumes "dim_R M > 0" shows "M \ 0" proof assume M: "M = 0" hence "dim_R M = (LEAST n. \ms. length ms = n \ set ms \ M \ RSpan ms = M)" using dim_R_def by simp moreover from M have "length [] = 0 \ set [] \ M \ RSpan [] = M" by simp ultimately show False using assms by simp qed end (* context R_scalar_mult *) hide_const real_vector.dim hide_const (open) Real_Vector_Spaces.dim abbreviation (in scalar_mult) "fingen \ R_scalar_mult.R_fingen UNIV smult" abbreviation (in scalar_mult) "dim \ R_scalar_mult.dim_R UNIV smult" lemmas (in Module) dim_nonzero = dim_R_nonzero subsubsection \@{term R}-linear independence\ context R_scalar_mult begin primrec R_lin_independent :: "'m list \ bool" where R_lin_independent_Nil: "R_lin_independent [] = True" | R_lin_independent_Cons: "R_lin_independent (m#ms) = (R_lin_independent ms \ (\r rs. (set (r#rs) \ R \ (r#rs) \\ (m#ms) = 0) \ r = 0))" lemma R_lin_independent_ConsI : assumes "R_lin_independent ms" "\r rs. set (r#rs) \ R \ (r#rs) \\ (m#ms) = 0 \ r = 0" shows "R_lin_independent (m#ms)" using assms R_lin_independent_Cons by fast lemma R_lin_independent_ConsD1 : "R_lin_independent (m#ms) \ R_lin_independent ms" by simp lemma R_lin_independent_ConsD2 : "\ R_lin_independent (m#ms); set (r#rs) \ R; (r#rs) \\ (m#ms) = 0 \ \ r = 0" by auto end (* context R_scalar_mult *) context RModule begin lemma R_lin_independent_imp_same_scalars : "\ length rs = length ss; length rs \ length ms; set rs \ R; set ss \ R; set ms \ M; R_lin_independent ms; rs \\ ms = ss \\ ms \ \ rs = ss" proof (induct rs ss arbitrary: ms rule: list_induct2) case (Cons r rs s ss) from Cons(3) have "ms \ []" by auto from this obtain n ns where ms: "ms = n#ns" using neq_Nil_conv[of ms] by fast from Cons(4,5) have "set ([a-b. (a,b)\zip (r#rs) (s#ss)]) \ R" using Ring1 Ring1.list_diff_closed by fast hence "set ((r-s)#[a-b. (a,b)\zip rs ss]) \ R" by simp moreover from Cons(1,4-6,8) ms have 1: "((r-s)#[a-b. (a,b)\zip rs ss]) \\ (n#ns) = 0" using lincomb_diff_left[of "r#rs" "s#ss"] by simp ultimately have "r - s = 0" using Cons(7) ms R_lin_independent_Cons by fast hence 2: "r = s" by simp with 1 Cons(1,4-6) ms have "rs \\ ns = ss \\ ns" using lincomb_Cons zero_smult lincomb_diff_left by simp with Cons(2-7) ms have "rs = ss" by simp with 2 show ?case by fast qed fast lemma R_lin_independent_obtain_unique_scalars : "\ set ms \ M; R_lin_independent ms; n \ RSpan ms \ \ (\! rs. set rs \ R \ length rs = length ms \ n = rs \\ ms)" using in_RSpan_obtain_same_length_coeffs[of n ms] R_lin_independent_imp_same_scalars[of _ _ ms] by auto lemma R_lin_independentI_all_scalars : "set ms \ M \ (\rs. set rs \ R \ length rs = length ms \ rs \\ ms = 0 \ set rs \ 0) \ R_lin_independent ms" proof (induct ms) case (Cons m ms) show ?case proof (rule R_lin_independent_ConsI) have "\rs. \ set rs \ R; length rs = length ms; rs \\ ms = 0 \ \ set rs \ 0" proof- fix rs assume rs: "set rs \ R" "length rs = length ms" "rs \\ ms = 0" with Cons(2) have "set (0#rs) \ R" "length (0#rs) = length (m#ms)" "(0#rs) \\ (m#ms) = 0" using R_scalars.zero_closed lincomb_Cons zero_smult by auto with Cons(3) have "set (0#rs) \ 0" by fast thus "set rs \ 0" by simp qed with Cons(1,2) show "R_lin_independent ms" by simp next fix r rs assume r_rs: "set (r # rs) \ R" "(r # rs) \\ (m # ms) = 0" from r_rs(1) Cons(2) obtain ss where ss: "set ss \ R" "length ss = length ms" "rs \\ ms = ss \\ ms" using lincomb_obtain_same_length_Rcoeffs[of rs ms] by auto with r_rs have "(r#ss) \\ (m#ms) = 0" using lincomb_Cons by simp moreover from r_rs(1) ss(1) have "set (r#ss) \ R" by simp moreover from ss(2) have "length (r#ss) = length (m#ms)" by simp ultimately have "set (r#ss) \ 0" using Cons(3) by fast thus "r = 0" by simp qed qed simp lemma R_lin_independentI_concat_all_scalars : defines eq_len: "eq_len \ \xs ys. length xs = length ys" assumes "set (concat mss) \ M" "\rss. set (concat rss) \ R \ list_all2 eq_len rss mss \ (concat rss) \\ (concat mss) = 0 \ (\rs \ set rss. set rs \ 0)" shows "R_lin_independent (concat mss)" using assms(2) proof (rule R_lin_independentI_all_scalars) have "\rs. \ set rs \ R; length rs = length (concat mss); rs \\ concat mss = 0 \ \ set rs \ 0" proof- fix rs assume rs: "set rs \ R" "length rs = length (concat mss)" "rs \\ concat mss = 0" from rs(2) eq_len obtain rss where "rs = concat rss" "list_all2 eq_len rss mss" using match_concat by fast with rs(1,3) assms(3) show "set rs \ 0" by auto qed thus "\rs. set rs \ R \ length rs = length (concat mss) \ rs \\ concat mss = 0 \ set rs \ 0" by auto qed lemma R_lin_independentD_all_scalars : "\ set rs \ R; set ms \ M; length rs \ length ms; R_lin_independent ms; rs \\ ms = 0 \ \ set rs \ 0" proof (induct rs ms rule: list_induct2') case (4 r rs m ms) from 4(2,5,6) have "r = 0" by auto moreover with 4 have "set rs \ 0" using lincomb_Cons zero_smult by simp ultimately show ?case by simp qed auto lemma R_lin_independentD_all_scalars_nth : assumes "set rs \ R" "set ms \ M" "R_lin_independent ms" "rs \\ ms = 0" "k < min (length rs) (length ms)" shows "rs!k = 0" proof- from assms(1,2) obtain ss where ss: "set ss \ R" "length ss = length ms" "take (length rs) ss = take (length ms) rs" "rs \\ ms = ss \\ ms" using lincomb_obtain_same_length_Rcoeffs[of rs ms] by fast from ss(1,2,4) assms(2,3,4) have "set ss \ 0" using R_lin_independentD_all_scalars by auto moreover from assms(5) ss(3) have "rs!k = (take (length rs) ss)!k" by simp moreover from assms(5) ss(2) have "k < length (take (length rs) ss)" by simp ultimately show ?thesis using in_set_conv_nth by force qed lemma R_lin_dependent_dependence_relation : "set ms \ M \ \ R_lin_independent ms \ \rs. set rs \ R \ set rs \ 0 \ length rs = length ms \ rs \\ ms = 0" proof (induct ms) case (Cons m ms) show ?case proof (cases "R_lin_independent ms") case True with Cons(3) have "\ (\r rs. (set (r#rs) \ R \ (r#rs) \\ (m#ms) = 0) \ r = 0)" by simp from this obtain r rs where r_rs: "set (r#rs) \ R" "(r#rs) \\ (m#ms) = 0" "r \ 0" by fast from r_rs(1) Cons(2) obtain ss where ss: "set ss \ R" "length ss = length ms" "rs \\ ms = ss \\ ms" using lincomb_obtain_same_length_Rcoeffs[of rs ms] by auto from ss r_rs have "set (r#ss) \ R \ set (r#ss) \ 0 \ length (r#ss) = length (m#ms) \ (r#ss) \\ (m#ms) = 0" using lincomb_Cons by simp thus ?thesis by fast next case False with Cons(1,2) obtain rs where rs: "set rs \ R" "set rs \ 0" "length rs = length ms" "rs \\ ms = 0" by fastforce from False rs Cons(2) have "set (0#rs) \ R \ set (0#rs) \ 0 \ length (0#rs) = length (m#ms) \ (0#rs) \\ (m#ms) = 0" using Ring1.zero_closed[OF Ring1] lincomb_Cons[of 0 rs m ms] zero_smult[of m] empty_set_diff_single[of "set rs"] by fastforce thus ?thesis by fast qed qed simp lemma R_lin_independent_imp_distinct : "set ms \ M \ R_lin_independent ms \ distinct ms" proof (induct ms) case (Cons m ms) have "\n. n \ set ms \ m \ n" proof fix n assume n: "n \ set ms" "m = n" from n(1) obtain xs ys where "ms = xs @ n # ys" using split_list by fast with Cons(2) n(2) have "(1 # replicate (length xs) 0 @ [-1]) \\ (m # ms) = 0" using lincomb_Cons lincomb_append lincomb_replicate0_left lincomb_Nil neg_eq_neg1_smult by simp with Cons(3) have "1 = 0" using R_scalars.zero_closed one_closed R_scalars.neg_closed by force thus False using one_neq_zero by fast qed with Cons show ?case by auto qed simp lemma R_lin_independent_imp_independent_take : "set ms \ M \ R_lin_independent ms \ R_lin_independent (take n ms)" proof (induct ms arbitrary: n) case (Cons m ms) show ?case proof (cases n) case (Suc k) hence "take n (m#ms) = m # take k ms" by simp moreover have "R_lin_independent (m # take k ms)" proof (rule R_lin_independent_ConsI) from Cons show "R_lin_independent (take k ms)" by simp next fix r rs assume r_rs: "set (r#rs) \ R" "(r#rs)\\(m # take k ms) = 0" from r_rs(1) Cons(2) obtain ss where ss: "set ss \ R" "length ss = length (take k ms)" "rs \\ take k ms = ss \\ take k ms" using set_take_subset[of k ms] lincomb_obtain_same_length_Rcoeffs by force from r_rs(1) ss(1) have "set (r#ss) \ R" by simp moreover from r_rs(2) ss have "(r#ss) \\ (m#ms) = 0" using lincomb_Cons lincomb_Nil lincomb_append_right[of ss "take k ms" "drop k ms"] by simp ultimately show "r = 0" using Cons(3) by auto qed ultimately show ?thesis by simp qed simp qed simp lemma R_lin_independent_Cons_imp_independent_RSpans : assumes "m \ M" "R_lin_independent (m#ms)" shows "add_independentS [RSpan [m], RSpan ms]" proof (rule add_independentS_doubleI) fix x y assume xy: "x \ RSpan [m]" "y \ RSpan ms" "x + y = 0" from xy(1,2) obtain r rs where r_rs: "r \ R" "x = r \ m" "set rs \ R" "y = rs \\ ms" using RSpan_single RSpanD_lincomb by fast with xy(3) have "set (r#rs) \ R" "(r#rs) \\ (m#ms) = 0" using lincomb_Cons by auto with assms r_rs(2) show "x = 0" using zero_smult by auto qed lemma hd0_imp_R_lin_dependent : "\ R_lin_independent (0#ms)" using lincomb_Cons[of 1 "[]" 0 ms] lincomb_Nil[of "[]" ms] smult_zero one_closed R_lin_independent_Cons by fastforce lemma R_lin_independent_imp_hd_n0 : "R_lin_independent (m#ms) \ m \ 0" using hd0_imp_R_lin_dependent by fast lemma R_lin_independent_imp_hd_independent_from_RSpan : assumes "m \ M" "set ms \ M" "R_lin_independent (m#ms)" shows "m \ RSpan ms" proof assume m: "m \ RSpan ms" with assms(2) have "(-1) \ m \ RSpan ms" using RSubmodule_RSpan[of ms] RModule.smult_closed[of R smult "RSpan ms" "-1" m] one_closed R_scalars.neg_closed[of 1] by simp moreover from assms(1) have "m + (-1) \ m = 0" using neg_eq_neg1_smult by simp ultimately show False using RSpan_contains_spanset_single assms R_lin_independent_Cons_imp_independent_RSpans add_independentS_doubleD R_lin_independent_imp_hd_n0 by fast qed lemma R_lin_independent_reduce : assumes "n \ M" shows "set ms \ M \ R_lin_independent (ms @ n # ns) \ R_lin_independent (ms @ ns)" proof (induct ms) case (Cons m ms) moreover have "\r rs. set (r#rs) \ R \ (r#rs) \\ (m#ms@ns) = 0 \ r = 0" proof- fix r rs assume r_rs: "set (r#rs) \ R" "(r#rs) \\ (m # ms @ ns) = 0" from Cons(2) r_rs(1) obtain ss where ss: "set ss \ R" "length ss = length ms" "rs \\ ms = ss \\ ms" using lincomb_obtain_same_length_Rcoeffs[of rs ms] by auto from assms ss(2,3) r_rs(2) have "(r # ss @ 0 # drop (length ms) rs) \\ (m # ms @ n # ns) = 0" using lincomb_Cons lincomb_append_right add.assoc[of "r\m" "rs\\ms" "(drop (length ms) rs)\\ns"] zero_smult lincomb_append by simp moreover from r_rs(1) ss(1) have "set (r # ss @ 0 # drop (length ms) rs) \ R" using R_scalars.zero_closed set_drop_subset[of _ rs] by auto ultimately show "r = 0" using Cons(3) R_lin_independent_ConsD2[of m _ r "ss @ 0 # drop (length ms) rs"] by simp qed ultimately show "R_lin_independent ( (m#ms) @ ns)" by auto qed simp lemma R_lin_independent_vs_lincomb0 : assumes "set (ms@n#ns) \ M" "R_lin_independent (ms @ n # ns)" "set (rs@s#ss) \ R" "length rs = length ms" "(rs@s#ss) \\ (ms@n#ns) = 0" shows "s = 0" proof- define k where "k = length rs" hence "(rs@s#ss)!k = s" by simp moreover from k_def assms(4) have "k < min (length (rs@s#ss)) (length (ms@n#ns))" by simp ultimately show ?thesis using assms(1,2,3,5) R_lin_independentD_all_scalars_nth[of "rs@s#ss" "ms@n#ns"] by simp qed lemma R_lin_independent_append_imp_independent_RSpans : "set ms \ M \ R_lin_independent (ms@ns) \ add_independentS [RSpan ms, RSpan ns]" proof (induct ms) case (Cons m ms) show ?case proof (rule add_independentS_doubleI) fix x y assume xy: "y \ RSpan ns""x \ RSpan (m#ms)" "x + y = 0" from xy(2) obtain x1 x2 where x1_x2: "x1 \ RSpan [m]" "x2 \ RSpan ms" "x = x1 + x2" using RSpan_Cons set_plus_def[of "RSpan [m]"] by auto from x1_x2(1,2) xy(1) obtain r rs ss where r_rs_ss: "set (r#(rs@ss)) \ R" "length rs = length ms" "x1 = r \ m" "x2 = rs \\ ms" "y = ss \\ ns" using RSpan_single in_RSpan_obtain_same_length_coeffs[of x2 ms] RSpanD_lincomb[of ns] by auto have x1_0: "x1 = 0" proof- from xy(3) x1_x2(3) r_rs_ss(2-5) have "(r#(rs@ss)) \\ (m#(ms@ns)) = 0" using lincomb_append lincomb_Cons by (simp add: algebra_simps) with r_rs_ss(1,3) Cons(2,3) show ?thesis using R_lin_independent_ConsD2[of m "ms@ns" r "rs@ss"] zero_smult by simp qed moreover have "x2 = 0" proof- from x1_0 xy(3) x1_x2(3) have "x2 + y = 0" by simp with xy(1) x1_x2(2) Cons show ?thesis using add_independentS_doubleD by simp qed ultimately show "x = 0" using x1_x2(3) by simp qed qed simp end (* context RModule *) subsubsection \Linear independence over \UNIV\\ context scalar_mult begin abbreviation "lin_independent ms \ R_scalar_mult.R_lin_independent UNIV smult ms" lemmas lin_independent_ConsI = R_scalar_mult.R_lin_independent_ConsI [OF R_scalar_mult, of smult] lemmas lin_independent_ConsD1 = R_scalar_mult.R_lin_independent_ConsD1[OF R_scalar_mult, of smult] end (* context scalar_mult *) context Module begin lemmas lin_independent_imp_independent_take = R_lin_independent_imp_independent_take lemmas lin_independent_reduce = R_lin_independent_reduce lemmas lin_independent_vs_lincomb0 = R_lin_independent_vs_lincomb0 lemmas lin_dependent_dependence_relation = R_lin_dependent_dependence_relation lemmas lin_independent_imp_distinct = R_lin_independent_imp_distinct lemmas lin_independent_imp_hd_independent_from_Span = R_lin_independent_imp_hd_independent_from_RSpan lemmas lin_independent_append_imp_independent_Spans = R_lin_independent_append_imp_independent_RSpans end (* context Module *) subsubsection \Rank\ context R_scalar_mult begin definition R_finrank :: "'m set \ bool" where "R_finrank M = (\n. \ms. set ms \ M \ R_lin_independent ms \ length ms \ n)" lemma R_finrankI : "(\ms. set ms \ M \ R_lin_independent ms \ length ms \ n) \ R_finrank M" unfolding R_finrank_def by blast lemma R_finrankD : "R_finrank M \ \n. \ms. set ms \ M \ R_lin_independent ms \ length ms \ n" unfolding R_finrank_def by fast lemma submodule_R_finrank : "R_finrank M \ N \ M \ R_finrank N" unfolding R_finrank_def by blast end (* context R_scalar_mult *) context scalar_mult begin abbreviation finrank :: "'m set \ bool" where "finrank \ R_scalar_mult.R_finrank UNIV smult" lemmas finrankI = R_scalar_mult.R_finrankI[OF R_scalar_mult, of _ smult] lemmas finrankD = R_scalar_mult.R_finrankD[OF R_scalar_mult, of smult] lemmas submodule_finrank = R_scalar_mult.submodule_R_finrank [OF R_scalar_mult, of smult] end (* context scalar_mult *) subsection \Module homomorphisms\ subsubsection \Locales\ locale RModuleHom = Domain?: RModule R smult M + Codomain?: scalar_mult smult' + GroupHom?: GroupHom M T for R :: "'r::ring_1 set" and smult :: "'r \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" and smult' :: "'r \ 'n::ab_group_add \ 'n" (infixr "\" 70) and T :: "'m \ 'n" + assumes R_map: "\r m. r \ R \ m \ M \ T (r \ m) = r \ T m" abbreviation (in RModuleHom) lincomb' :: "'r list \ 'n list \ 'n" (infix "\\" 70) where "lincomb' \ Codomain.lincomb" lemma (in RModule) RModuleHomI : assumes "GroupHom M T" "\r m. r \ R \ m \ M \ T (r \ m) = smult' r (T m)" shows "RModuleHom R smult M smult' T" by ( rule RModuleHom.intro, rule RModule_axioms, rule assms(1), unfold_locales, rule assms(2) ) locale RModuleEnd = RModuleHom R smult M smult T for R :: "'r::ring_1 set" and smult :: "'r \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" and T :: "'m \ 'm" + assumes endomorph: "ImG \ M" locale ModuleHom = RModuleHom UNIV smult M smult' T for smult :: "'r::ring_1 \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" and smult' :: "'r \ 'n::ab_group_add \ 'n" (infixr "\" 70) and T :: "'m \ 'n" lemmas (in ModuleHom) hom = hom lemmas (in Module) ModuleHomI = RModuleHomI[THEN ModuleHom.intro] locale ModuleEnd = ModuleHom smult M smult T for smult :: "'r::ring_1 \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" and T :: "'m \ 'm" + assumes endomorph: "ImG \ M" locale RModuleIso = RModuleHom R smult M smult' T for R :: "'r::ring_1 set" and smult :: "'r \ 'm::ab_group_add \ 'm" (infixr "\" 70) and M :: "'m set" and smult' :: "'r \ 'n::ab_group_add \ 'n" (infixr "\" 70) and T :: "'m \ 'n" + fixes N :: "'n set" assumes bijective: "bij_betw T M N" lemma (in RModule) RModuleIsoI : assumes "GroupIso M T N" "\r m. r \ R \ m \ M \ T (r \ m) = smult' r (T m)" shows "RModuleIso R smult M smult' T N" proof (rule RModuleIso.intro) from assms show "RModuleHom R (\) M smult' T" using GroupIso.axioms(1) RModuleHomI by fastforce from assms(1) show "RModuleIso_axioms M T N" using GroupIso.bijective by unfold_locales qed subsubsection \Basic facts\ lemma (in RModule) trivial_RModuleHom : "\r\R. smult' r 0 = 0 \ RModuleHom R smult M smult' 0" using trivial_GroupHom RModuleHomI by fastforce lemma (in RModule) RModHom_idhom : "RModuleHom R smult M smult (id\M)" using RModule_axioms GroupHom_idhom proof (rule RModuleHom.intro) show "RModuleHom_axioms R (\) M (\) (id \ M)" using smult_closed by unfold_locales simp qed context RModuleHom begin lemmas additive = hom lemmas supp = supp lemmas im_zero = im_zero lemmas im_diff = im_diff lemmas Ker_Im_iff = Ker_Im_iff lemmas Ker0_imp_inj_on = Ker0_imp_inj_on lemma GroupHom : "GroupHom M T" .. lemma codomain_smult_zero : "r \ R \ r \ 0 = 0" using im_zero smult_zero zero_closed R_map[of r 0] by simp lemma RSubmodule_Ker : "Domain.RSubmodule Ker" proof (rule Domain.RSubmoduleI, rule conjI, rule Group_Ker) fix r m assume r: "r \ R" and m: "m \ Ker" thus "r \ m \ Ker" using R_map[of r m] kerD[of m T] codomain_smult_zero kerI Domain.smult_closed by simp qed fast lemma RModule_Im : "RModule R smult' ImG" using Ring1 Group_Im proof (rule RModuleI, unfold_locales) show "\n. n \ T ` M \ 1 \ n = n" using one_closed R_map[of 1] by auto next fix r s m n assume r: "r \ R" and s: "s \ R" and m: "m \ T ` M" and n: "n \ T ` M" from m n obtain m' n' where m': "m' \ M" "m = T m'" and n': "n' \ M" "n = T n'" by fast from m' r R_map have "r \ m = T (r \ m')" by simp with r m'(1) show "r \ m \ T ` M" using smult_closed by fast from r m' n' show "r \ (m + n) = r \ m + r \ n" using hom add_closed R_map[of r "m'+n'"] smult_closed R_map[of r] by simp from r s m' show "(r + s) \ m = r \ m + s \ m" using R_scalars.add_closed R_map[of "r+s"] smult_closed hom R_map by simp from r s m' show "r \ s \ m = (r * s) \ m" using smult_closed R_map[of s] R_map[of r "s \ m'"] mult_closed R_map[of "r*s"] by simp qed lemma im_submodule : assumes "RSubmodule N" shows "RModule.RSubmodule R smult' ImG (T ` N)" proof (rule RModule.RSubmoduleI, rule RModule_Im) from assms show "Group.Subgroup (T ` M) (T ` N)" using im_subgroup Subgroup_RSubmodule by fast from assms R_map show "\r n. r \ R \ n \ T ` N \ r \ n \ T ` N" using RModule.smult_closed by force qed lemma RModHom_composite_left : assumes "T ` M \ N" "RModuleHom R smult' N smult'' S" shows "RModuleHom R smult M smult'' (S \ T)" proof (rule RModule.RModuleHomI, rule RModule_axioms) from assms(1) show "GroupHom M (S \ T)" using RModuleHom.GroupHom[OF assms(2)] GroupHom_composite_left by auto from assms(1) show "\r m. r \ R \ m \ M \ (S \ T) (r \ m) = smult'' r ((S \ T) m)" using R_map RModuleHom.R_map[OF assms(2)] by auto qed lemma RModuleHom_restrict0_submodule : assumes "RSubmodule N" shows "RModuleHom R smult N smult' (T \ N)" proof (rule RModuleHom.intro) from assms show "RModule R (\) N" by fast from assms show "GroupHom N (T \ N)" using RModule.Group GroupHom_restrict0_subgroup by fast show "RModuleHom_axioms R (\) N (\) (T \ N)" proof fix r m assume "r \ R" "m \ N" with assms show "(T \ N) (r \ m) = r \ (T \ N) m" using RModule.smult_closed R_map by fastforce qed qed lemma distrib_lincomb : "set rs \ R \ set ms \ M \ T (rs \\ ms) = rs \\ map T ms" using Domain.lincomb_Nil im_zero Codomain.lincomb_Nil R_map Domain.lincomb_Cons Domain.smult_closed Domain.lincomb_closed additive Codomain.lincomb_Cons by (induct rs ms rule: list_induct2') auto lemma same_image_on_RSpanset_imp_same_hom : assumes "RModuleHom R smult M smult' S" "set ms \ M" "M = Domain.R_scalars.RSpan ms" "\m\set ms. S m = T m" shows "S = T" proof fix m show "S m = T m" proof (cases "m \ M") case True with assms(2,3) obtain rs where rs: "set rs \ R" "m = rs \\ ms" using Domain.RSpanD_lincomb_arb_len_coeffs by fast from rs(1) assms(2) have "S (rs \\ ms) = rs \\ (map S ms)" using RModuleHom.distrib_lincomb[OF assms(1)] by simp moreover from rs(1) assms(2) have "T (rs \\ ms) = rs \\ (map T ms)" using distrib_lincomb by simp ultimately show ?thesis using assms(4) map_ext[of ms S T] rs(2) by auto next case False with assms(1) supp show ?thesis using RModuleHom.supp suppI_contra[of _ S] suppI_contra[of _ T] by fastforce qed qed end (* context RModuleHom *) lemma RSubmodule_eigenspace : fixes smult :: "'r::ring_1 \ 'm::ab_group_add \ 'm" (infixr "\" 70) assumes RModHom: "RModuleHom R smult M smult T" and r: "r \ R" "\s m. s \ R \ m \ M \ s \ r \ m = r \ s \ m" defines E: "E \ {m \ M. T m = r \ m}" shows "RModule.RSubmodule R smult M E" proof (rule RModule.RSubmoduleI) from RModHom show rmod: "RModule R smult M" using RModuleHom.axioms(1) by fast have "Group E" proof from r(1) E show "E \ {}" using RModule.zero_closed[OF rmod] RModuleHom.im_zero[OF RModHom] RModule.smult_zero[OF rmod] by auto next fix m n assume "m \ E" "n \ E" with r(1) E show "m - n \ E" using RModule.diff_closed[OF rmod] RModuleHom.im_diff[OF RModHom] RModule.smult_distrib_left_diff[OF rmod] by simp qed with E show "Group.Subgroup M E" by fast show "\s m. s \ R \ m \ E \ s \ m \ E" proof- fix s m assume "s \ R" "m \ E" with E r RModuleHom.R_map[OF RModHom] show "s \ m \ E" using RModule.smult_closed[OF rmod] by simp qed qed subsubsection \Basic facts about endomorphisms\ lemma (in RModule) Rmap_endomorph_is_RModuleEnd : assumes grpend: "GroupEnd M T" and Rmap : "\r m. r \ R \ m \ M \ T (r \ m) = r \ (T m)" shows "RModuleEnd R smult M T" proof (rule RModuleEnd.intro, rule RModuleHomI) from grpend show "GroupHom M T" using GroupEnd.axioms(1) by fast from grpend show "RModuleEnd_axioms M T" using GroupEnd.endomorph by unfold_locales qed (rule Rmap) lemma (in RModuleEnd) GroupEnd : "GroupEnd M T" proof (rule GroupEnd.intro) from endomorph show "GroupEnd_axioms M T" by unfold_locales qed (unfold_locales) lemmas (in RModuleEnd) proj_decomp = GroupEnd.proj_decomp[OF GroupEnd] lemma (in ModuleEnd) RModuleEnd : "RModuleEnd UNIV smult M T" using endomorph RModuleEnd.intro by unfold_locales lemmas (in ModuleEnd) GroupEnd = RModuleEnd.GroupEnd[OF RModuleEnd] lemma RModuleEnd_over_UNIV_is_ModuleEnd : "RModuleEnd UNIV rsmult M T \ ModuleEnd rsmult M T" proof (rule ModuleEnd.intro, rule ModuleHom.intro) assume endo: "RModuleEnd UNIV rsmult M T" thus "RModuleHom UNIV rsmult M rsmult T" using RModuleEnd.axioms(1) by fast from endo show "ModuleEnd_axioms M T" using RModuleEnd.endomorph by unfold_locales qed subsubsection \Basic facts about isomorphisms\ context RModuleIso begin abbreviation "invT \ (the_inv_into M T) \ N" lemma GroupIso : "GroupIso M T N" proof (rule GroupIso.intro) show "GroupHom M T" .. from bijective show "GroupIso_axioms M T N" by unfold_locales qed lemmas ImG = GroupIso.ImG [OF GroupIso] lemmas GroupHom_inv = GroupIso.inv [OF GroupIso] lemmas invT_into = GroupIso.invT_into [OF GroupIso] lemmas T_invT = GroupIso.T_invT [OF GroupIso] lemmas invT_eq = GroupIso.invT_eq [OF GroupIso] lemma RModuleN : "RModule R smult' N" using RModule_Im ImG by fast lemma inv : "RModuleIso R smult' N smult invT M" using RModuleN GroupHom_inv proof (rule RModule.RModuleIsoI) fix r n assume rn: "r \ R" "n \ N" thus "invT (r \ n) = r \ invT n" using invT_into smult_closed R_map T_invT invT_eq by simp qed end (* context RModuleIso *) subsection \Inner direct sums of RModules\ lemma (in RModule) RModule_inner_dirsum_el_decomp_Rsmult : assumes "\N\set Ns. RSubmodule N" "add_independentS Ns" "r \ R" "x \ (\N\Ns. N)" shows "(\Ns\(r \ x)) = [r \ m. m\(\Ns\x)]" proof- define xs where "xs = (\Ns\x)" with assms have x: "xs \ listset Ns" "x = sum_list xs" using RModule.AbGroup[of R] AbGroup_inner_dirsum_el_decompI[of Ns x] by auto from assms(1,2,4) xs_def have xs_M: "set xs \ M" using Subgroup_RSubmodule AbGroup.abSubgroup_inner_dirsum_el_decomp_set[OF AbGroup] by fast from assms(1,3) x(1) have "[r \ m. m\xs] \ listset Ns" using listset_RModule_Rsmult_closed by fast moreover from x assms(3) xs_M have "r \ x = sum_list [r \ m. m\xs]" using smult_sum_list_distrib by fast moreover from assms(1,3,4) have "r \ x \ (\M\Ns. M)" using RModule_inner_dirsum RModule.smult_closed by fast ultimately show "(\Ns\(r \ x)) = [r \ m. m\xs]" using assms(1,2) RModule.AbGroup AbGroup_inner_dirsum_el_decomp_eq by fast qed lemma (in RModule) RModuleEnd_inner_dirsum_el_decomp_nth : assumes "\N \ set Ns. RSubmodule N" "add_independentS Ns" "n < length Ns" shows "RModuleEnd R smult (\N\Ns. N) (\Ns\n)" proof (rule RModule.Rmap_endomorph_is_RModuleEnd) from assms(1) show "RModule R smult (\N\Ns. N)" using RSubmodule_inner_dirsum by fast from assms show "GroupEnd (\N\Ns. N) \Ns\n" using RModule.AbGroup GroupEnd_inner_dirsum_el_decomp_nth[of Ns] by fast show "\r m. r \ R \ m \ (\N\Ns. N) \ (\Ns\n) (r \ m) = r \ ((\Ns\n) m)" proof- fix r m assume "r \ R" "m \ (\N\Ns. N)" moreover with assms(1) have "r \ m \ (\M\Ns. M)" using RModule_inner_dirsum RModule.smult_closed by fast ultimately show "(\Ns\n) (r \ m) = r \ (\Ns\n) m" using assms RModule.AbGroup[of R smult] AbGroup_length_inner_dirsum_el_decomp[of Ns] RModule_inner_dirsum_el_decomp_Rsmult by simp qed qed section \Vector Spaces\ subsection \Locales and basic facts\ text \Here we don't care about being able to switch scalars.\ locale fscalar_mult = scalar_mult smult for smult :: "'f::field \ 'v::ab_group_add \ 'v" (infixr "\" 70) abbreviation (in fscalar_mult) "findim \ fingen" locale VectorSpace = Module smult V for smult :: "'f::field \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" lemmas VectorSpaceI = ModuleI[THEN VectorSpace.intro] sublocale VectorSpace < fscalar_mult proof- qed locale FinDimVectorSpace = VectorSpace + assumes findim: "findim V" lemma (in VectorSpace) FinDimVectorSpaceI : "findim V \ FinDimVectorSpace (\) V" by unfold_locales fast context VectorSpace begin abbreviation Subspace :: "'v set \ bool" where "Subspace \ Submodule" lemma SubspaceD1 : "Subspace U \ VectorSpace smult U" using VectorSpace.intro Module.intro by fast lemmas AbGroup = AbGroup lemmas add_closed = add_closed lemmas smult_closed = smult_closed lemmas one_smult = one_smult lemmas smult_assoc = smult_assoc lemmas smult_distrib_left = smult_distrib_left lemmas smult_distrib_right = smult_distrib_right lemmas zero_closed = zero_closed lemmas zero_smult = zero_smult lemmas smult_zero = smult_zero lemmas smult_lincomb = smult_lincomb lemmas smult_distrib_left_diff = smult_distrib_left_diff lemmas smult_sum_distrib = smult_sum_distrib lemmas sum_smult_distrib = sum_smult_distrib lemmas lincomb_sum = lincomb_sum lemmas lincomb_closed = lincomb_closed lemmas lincomb_concat = lincomb_concat lemmas lincomb_replicate0_left = lincomb_replicate0_left lemmas delta_scalars_lincomb_eq_nth = delta_scalars_lincomb_eq_nth lemmas SpanI = SpanI lemmas Span_closed = Span_closed lemmas SpanD_lincomb_arb_len_coeffs = SpanD_lincomb_arb_len_coeffs lemmas SpanI_lincomb_arb_len_coeffs = SpanI_lincomb_arb_len_coeffs lemmas in_Span_obtain_same_length_coeffs = in_Span_obtain_same_length_coeffs lemmas SubspaceI = SubmoduleI lemmas subspace_finrank = submodule_finrank lemma cancel_scalar: "\ a \ 0; u \ V; v \ V; a \ u = a \ v \ \ u = v" using smult_assoc[of "1/a" a u] by simp end (* context VectorSpace *) subsection \Linear algebra in vector spaces\ subsubsection \Linear independence and spanning\ context VectorSpace begin lemmas Subspace_Span = Submodule_Span lemmas lin_independent_Nil = R_lin_independent_Nil lemmas lin_independentI_concat_all_scalars = R_lin_independentI_concat_all_scalars lemmas lin_independentD_all_scalars = R_lin_independentD_all_scalars lemmas lin_independent_obtain_unique_scalars = R_lin_independent_obtain_unique_scalars lemma lincomb_Cons_0_imp_in_Span : "\ v \ V; set vs \ V; a \ 0; (a#as) \\ (v#vs) = 0 \ \ v \ Span vs" using lincomb_Cons eq_neg_iff_add_eq_0[of "a \ v" "as \\ vs"] neg_lincomb smult_assoc[of "1/a" a v] smult_lincomb SpanD_lincomb_arb_len_coeffs by auto lemma lin_independent_Cons_conditions : "\ v \ V; set vs \ V; v \ Span vs; lin_independent vs \ \ lin_independent (v#vs)" using lincomb_Cons_0_imp_in_Span lin_independent_ConsI by fast lemma coeff_n0_imp_in_Span_others : assumes "v \ V" "set us \ V" "set vs \ V" "b \ 0" "length as = length us" "w = (as @ b # bs) \\ (us @ v # vs)" shows "v \ Span (w # us @ vs)" proof- define x where "x = (1 # [- c. c\as@bs]) \\ (w # us @ vs)" from assms(1,4-6) have "v = (1/b) \ (w + - ( (as@bs) \\ (us@vs) ))" using lincomb_append lincomb_Cons by simp moreover from assms(1,2,3,6) have w: "w \ V" using lincomb_closed by simp ultimately have "v = (1/b) \ x" using x_def assms(2,3) neg_lincomb[of _ "us@vs"] lincomb_Cons[of 1 _ w] by simp with x_def w assms(2,3) show ?thesis using SpanD_lincomb_arb_len_coeffs[of "w # us @ vs"] Span_smult_closed[of "1/b" "w # us @ vs" x] by auto qed lemma lin_independent_replace1_by_lincomb : assumes "set us \ V" "v \ V" "set vs \ V" "lin_independent (us @ v # vs)" "length as = length us" "b \ 0" shows "lin_independent ( ((as @ b # bs) \\ (us @ v # vs)) # us @ vs )" proof- define w where "w = (as @ b # bs) \\ (us @ v # vs)" from assms(1,2,4) have "lin_independent (us @ vs)" using lin_independent_reduce by fast hence "lin_independent (w # us @ vs)" proof (rule lin_independent_ConsI) fix c cs assume A: "(c#cs) \\ (w # us @ vs) = 0" from assms(1,3) obtain ds es fs where dsesfs: "length ds = length vs" "bs \\ vs = ds \\ vs" "length es = length vs" "(drop (length us) cs) \\ vs = es \\ vs" "length fs = length us" "cs \\ us = fs \\ us" using lincomb_obtain_same_length_coeffs[of bs vs] lincomb_obtain_same_length_coeffs[of "drop (length us) cs" vs] lincomb_obtain_same_length_coeffs[of cs us] by auto define xs ys where "xs = [x+y. (x,y)\zip [c*a. a\as] fs]" and "ys = [x+y. (x,y)\zip es [c*d. d\ds]]" with assms(5) dsesfs(5) have len_xs: "length xs = length us" using length_concat_map_split_zip[of _ "[c*a. a\as]" fs] by simp from A w_def assms(1-3,5) dsesfs(2,4,6) have "0 = c \ as \\ us + fs \\ us + (c * b) \ v + es \\ vs + c \ ds \\ vs" using lincomb_Cons lincomb_append_right lincomb_append add_closed smult_closed lincomb_closed by (simp add: algebra_simps) also from assms(1,3,5) dsesfs(1,3,5) xs_def ys_def len_xs have "\ = (xs @ (c * b) # ys) \\ (us @ v # vs)" using smult_lincomb lincomb_sum lincomb_Cons lincomb_append by simp finally have "(xs @ (c * b) # ys) \\ (us @ v # vs) = 0" by simp with assms(1-3,4,6) len_xs show "c = 0" using lin_independent_vs_lincomb0 by fastforce qed with w_def show ?thesis by fast qed lemma build_lin_independent_seq : assumes us_V: "set us \ V" and indep_us: "lin_independent us" shows "\ws. set ws \ V \ lin_independent (ws @ us) \ (Span (ws @ us) = V \ length ws = n)" proof (induct n) case 0 from indep_us show ?case by force next case (Suc m) from this obtain ws where ws: "set ws \ V" "lin_independent (ws @ us)" "Span (ws@us) = V \ length ws = m" by auto show ?case proof (cases "V = Span (ws@us)") case True with ws show ?thesis by fast next case False moreover from ws(1) us_V have ws_us_V: "set (ws @ us) \ V" by simp ultimately have "Span (ws@us) \ V" using Span_closed by fast from this obtain w where w: "w \ V" "w \ Span (ws@us)" by fast define vs where "vs = w # ws" with w ws_us_V ws(2,3) have "set (vs @ us) \ V" "lin_independent (vs @ us)" "length vs = Suc m" using lin_independent_Cons_conditions[of w "ws@us"] by auto thus ?thesis by auto qed qed end (* context VectorSpace *) subsubsection \Basis for a vector space: \basis_for\\ abbreviation (in fscalar_mult) basis_for :: "'v set \ 'v list \ bool" where "basis_for V vs \ (set vs \ V \ V = Span vs \ lin_independent vs)" context VectorSpace begin lemma spanset_contains_basis : "set vs \ V \ \us. set us \ set vs \ basis_for (Span vs) us" proof (induct vs) case Nil show ?case using lin_independent_Nil by simp next case (Cons v vs) from this obtain ws where ws: "set ws \ set vs" "basis_for (Span vs) ws" by auto show ?case proof (cases "v \ Span vs") case True with Cons(2) ws(2) have "basis_for (Span (v#vs)) ws" using spanset_reduce_Cons by force with ws(1) show ?thesis by auto next case False from Cons(2) ws have "set (v#ws) \ set (v#vs)" "set (v#ws) \ Span (v#vs)" "Span (v#vs) = Span (v#ws)" using Span_contains_spanset[of "v#vs"] Span_contains_Spans_Cons_right[of v vs] Span_Cons by auto moreover have "lin_independent (v#ws)" proof (rule lin_independent_Cons_conditions) from Cons(2) ws(1) show "v \ V" "set ws \ V" by auto from ws(2) False show "v \ Span ws" "lin_independent ws" by auto qed ultimately show ?thesis by blast qed qed lemma basis_for_Span_ex : "set vs \ V \ \us. basis_for (Span vs) us" using spanset_contains_basis by fastforce lemma replace_basis_one_step : assumes closed: "set vs \ V" "set us \ V" and indep: "lin_independent (us@vs)" and new_w: "w \ Span (us@vs) - Span us" shows "\xs y ys. vs = xs @ y # ys \ basis_for (Span (us@vs)) (w # us @ xs @ ys)" proof- from new_w obtain u v where uv: "u \ Span us" "v \ Span vs" "w = u + v" using Span_append set_plus_def[of "Span us"] by auto from uv(1,3) new_w have v_n0: "v \ 0" by auto from uv(1,2) obtain as bs where as_bs: "length as = length us" "u = as \\ us" "length bs = length vs" "v = bs \\ vs" using in_Span_obtain_same_length_coeffs by blast from v_n0 as_bs(4) closed(1) obtain b where b: "b \ set bs" "b \ 0" using lincomb_0coeffs[of vs] by auto from b(1) obtain cs ds where cs_ds: "bs = cs @ b # ds" using split_list by fast define n where "n = length cs" define fvs where "fvs = take n vs" define y where "y = vs!n" define bvs where "bvs = drop (Suc n) vs" define ufvs where "ufvs = us @ fvs" define acs where "acs = as @ cs" from as_bs(1,3) cs_ds n_def acs_def ufvs_def fvs_def have n_len_vs: "n < length vs" and len_acs: "length acs = length ufvs" by auto from n_len_vs fvs_def y_def bvs_def have vs_decomp: "vs = fvs @ y # bvs" using id_take_nth_drop by simp with uv(3) as_bs(1,2,4) cs_ds acs_def ufvs_def have w_decomp: "w = (acs @ b # ds) \\ (ufvs @ y # bvs)" using lincomb_append by simp from closed(1) vs_decomp have y_V: "y \ V" and fvs_V: "set fvs \ V" and bvs_V: "set bvs \ V" by auto from ufvs_def fvs_V closed(2) have ufvs_V: "set ufvs \ V" by simp from w_decomp ufvs_V y_V bvs_V have w_V: "w \ V" using lincomb_closed by simp have "Span (us@vs) = Span (w # ufvs @ bvs)" proof from vs_decomp ufvs_def have 1: "Span (us@vs) = Span (y # ufvs @ bvs)" using Span_append Span_Cons[of y bvs] Span_Cons[of y ufvs] Span_append[of "y#ufvs" bvs] by (simp add: algebra_simps) with new_w y_V ufvs_V bvs_V show "Span (w # ufvs @ bvs) \ Span (us@vs)" using Span_replace_hd by simp from len_acs w_decomp y_V ufvs_V bvs_V have "y \ Span (w # ufvs @ bvs)" using b(2) coeff_n0_imp_in_Span_others by simp with w_V ufvs_V bvs_V have "Span (y # ufvs @ bvs) \ Span (w # ufvs @ bvs)" using Span_replace_hd by simp with 1 show "Span (us@vs) \ Span (w # ufvs @ bvs)" by fast qed moreover from ufvs_V y_V bvs_V ufvs_def indep vs_decomp w_decomp len_acs b(2) have "lin_independent (w # ufvs @ bvs)" using lin_independent_replace1_by_lincomb[of ufvs y bvs acs b ds] by simp moreover have "set (w # (us@fvs) @ bvs) \ Span (us@vs)" proof- from new_w have "w \ Span (us@vs)" by fast moreover from closed have "set us \ Span (us@vs)" using Span_contains_spanset_append_left by fast moreover from closed fvs_def have "set fvs \ Span (us@vs)" using Span_contains_spanset_append_right[of us] set_take_subset by fastforce moreover from closed bvs_def have "set bvs \ Span (us@vs)" using Span_contains_spanset_append_right[of us] set_drop_subset by fastforce ultimately show ?thesis by simp qed ultimately show ?thesis using ufvs_def vs_decomp by auto qed lemma replace_basis : assumes closed: "set vs \ V" and indep_vs: "lin_independent vs" shows "\ length us \ length vs; set us \ Span vs; lin_independent us \ \ \pvs. length pvs = length vs \ set pvs = set vs \ basis_for (Span vs) (take (length vs) (us @ pvs))" proof (induct us) case Nil from closed indep_vs show ?case using Span_contains_spanset[of vs] by fastforce next case (Cons u us) from this obtain ppvs where ppvs: "length ppvs = length vs" "set ppvs = set vs" "basis_for (Span vs) (take (length vs) (us @ ppvs))" using lin_independent_ConsD1[of u us] by auto define fppvs bppvs where "fppvs = take (length vs - length us) ppvs" and "bppvs = drop (length vs - length us) ppvs" with ppvs(1) Cons(2) have ppvs_decomp: "ppvs = fppvs @ bppvs" and len_fppvs : "length fppvs = length vs - length us" by auto from closed Cons(3) have uus_V: "u \ V" "set us \ V" using Span_closed by auto from closed ppvs(2) have "set ppvs \ V" by fast with fppvs_def have fppvs_V: "set fppvs \ V" using set_take_subset[of _ ppvs] by fast from fppvs_def Cons(2) have prev_basis_decomp: "take (length vs) (us @ ppvs) = us @ fppvs" by auto with Cons(3,4) ppvs(3) fppvs_V uus_V obtain xs y ys where xs_y_ys: "fppvs = xs @ y # ys" "basis_for (Span vs) (u # us @ xs @ ys)" using lin_independent_imp_hd_independent_from_Span replace_basis_one_step[of fppvs us u] by auto define pvs where "pvs = xs @ ys @ y # bppvs" with xs_y_ys len_fppvs ppvs_decomp ppvs(1,2) have "length pvs = length vs" "set pvs = set vs" "basis_for (Span vs) (take (length vs) ((u # us) @ pvs))" using take_append[of "length vs" "u # us @ xs @ ys"] by auto thus ?case by fast qed lemma replace_basis_completely : "\ set vs \ V; lin_independent vs; length us = length vs; set us \ Span vs; lin_independent us \ \ basis_for (Span vs) us" using replace_basis[of vs us] by auto lemma basis_for_obtain_unique_scalars : "basis_for V vs \ v \ V \ \! as. length as = length vs \ v = as \\ vs" using lin_independent_obtain_unique_scalars by fast lemma add_unique_scalars : assumes vs: "basis_for V vs" and v: "v \ V" and v': "v' \ V" defines as: "as \ (THE ds. length ds = length vs \ v = ds \\ vs)" and bs: "bs \ (THE ds. length ds = length vs \ v' = ds \\ vs)" and cs: "cs \ (THE ds. length ds = length vs \ v+v' = ds \\ vs)" shows "cs = [a+b. (a,b)\zip as bs]" proof- from vs v v' as bs have as': "length as = length vs \ v = as \\ vs" and bs': "length bs = length vs \ v' = bs \\ vs" using basis_for_obtain_unique_scalars theI'[ of "\ds. length ds = length vs \ v = ds \\ vs" ] theI'[of "\ds. length ds = length vs \ v' = ds \\ vs"] by auto have "length [a+b. (a,b)\zip as bs] = length (zip as bs)" by (induct as bs rule: list_induct2') auto with vs as' bs' have "length [a+b. (a,b)\zip as bs] = length vs \ v + v' = [a + b. (a,b)\zip as bs] \\ vs" using lincomb_sum by auto moreover from vs v v' have "\! ds. length ds = length vs \ v+v' = ds \\ vs" using add_closed basis_for_obtain_unique_scalars by force ultimately show ?thesis using cs the1_equality by fast qed lemma smult_unique_scalars : fixes a::'f assumes vs: "basis_for V vs" and v: "v \ V" defines as: "as \ (THE cs. length cs = length vs \ v = cs \\ vs)" and bs: "bs \ (THE cs. length cs = length vs \ a \ v = cs \\ vs)" shows "bs = map ((*) a) as" proof- from vs v as have "length as = length vs \ v = as \\ vs" using basis_for_obtain_unique_scalars theI'[ of "\cs. length cs = length vs \ v = cs \\ vs" ] by auto with vs have "length (map ((*) a) as) = length vs \ a \ v = (map ((*) a) as) \\ vs" using smult_lincomb by auto moreover from vs v have "\! cs. length cs = length vs \ a \ v = cs \\ vs" using smult_closed basis_for_obtain_unique_scalars by fast ultimately show ?thesis using bs the1_equality by fast qed lemma max_lin_independent_set_in_Span : assumes "set vs \ V" "set us \ Span vs" "lin_independent us" shows "length us \ length vs" proof (cases us) case (Cons x xs) from assms(1) spanset_contains_basis[of vs] obtain bvs where bvs: "set bvs \ set vs" "basis_for (Span vs) bvs" by auto with assms(1) have len_bvs: "length bvs \ length vs" using lin_independent_imp_distinct[of bvs] distinct_card finite_set card_mono[of "set vs" "set bvs"] card_length[of vs] by fastforce moreover have "length (x#xs) > length bvs \ \ lin_independent (x#xs)" proof assume A: "length (x#xs) > length bvs" "lin_independent (x#xs)" define ws where "ws = take (length bvs) xs" from Cons assms(1,2) have xxs_V: "x \ V" "set xs \ V" using Span_closed by auto from ws_def A(1) have "length ws = length bvs" by simp moreover from Cons assms(2) bvs(2) ws_def have "set ws \ Span bvs" using set_take_subset by fastforce ultimately have "basis_for (Span vs) ws" using A(2) ws_def assms(1) bvs xxs_V lin_independent_ConsD1 lin_independent_imp_independent_take replace_basis_completely[of bvs ws] by force with Cons assms(2) ws_def A(2) xxs_V show False using Span_contains_Span_take[of xs] lin_independent_imp_hd_independent_from_Span[of x xs] by auto qed ultimately show ?thesis using Cons assms(3) by fastforce qed simp lemma finrank_Span : "set vs \ V \ finrank (Span vs)" using max_lin_independent_set_in_Span finrankI by blast end (* context VectorSpace *) subsection \Finite dimensional spaces\ context VectorSpace begin lemma dim_eq_size_basis : "basis_for V vs \ length vs = dim V" using max_lin_independent_set_in_Span Least_equality[ of "\n::nat. \us. length us = n \ set us \ V \ RSpan us = V" "length vs" ] unfolding dim_R_def by fastforce lemma finrank_imp_findim : assumes "finrank V" shows "findim V" proof- from assms obtain n where "\vs. set vs \ V \ lin_independent vs \ length vs \ n" using finrankD by fastforce moreover from build_lin_independent_seq[of "[]"] obtain ws where "set ws \ V" "lin_independent ws" "Span ws = V \ length ws = Suc n" by auto ultimately show ?thesis by auto qed lemma subspace_Span_is_findim : "\ set vs \ V; Subspace W; W \ Span vs \ \ findim W" using finrank_Span subspace_finrank[of "Span vs" W] SubspaceD1[of W] VectorSpace.finrank_imp_findim by auto end (* context VectorSpace *) context FinDimVectorSpace begin lemma Subspace_is_findim : "Subspace U \ findim U" using findim subspace_Span_is_findim by fast lemma basis_ex : "\vs. basis_for V vs" using findim basis_for_Span_ex by auto lemma lin_independent_length_le_dim : "set us \ V \ lin_independent us \ length us \ dim V" using basis_ex max_lin_independent_set_in_Span dim_eq_size_basis by force lemma too_long_lin_dependent : "set us \ V \ length us > dim V \ \ lin_independent us" using lin_independent_length_le_dim by fastforce lemma extend_lin_independent_to_basis : assumes "set us \ V" "lin_independent us" shows "\vs. basis_for V (vs @ us)" proof- define n where "n = Suc (dim V - length us)" from assms obtain vs where vs: "set vs \ V" "lin_independent (vs @ us)" "Span (vs @ us) = V \ length vs = n" using build_lin_independent_seq[of us n] by fast with assms n_def show ?thesis using set_append lin_independent_length_le_dim[of "vs @ us"] by auto qed lemma extend_Subspace_basis : "U \ V \ basis_for U us \ \vs. basis_for V (vs@us)" using Span_contains_spanset extend_lin_independent_to_basis by fast lemma Subspace_dim_le : assumes "Subspace U" shows "dim U \ dim V" using assms findim proof- from assms obtain us where "basis_for U us" using Subspace_is_findim SubspaceD1 VectorSpace.FinDimVectorSpaceI[of "(\)" U] FinDimVectorSpace.basis_ex[of "(\)" U] by auto with assms show ?thesis using RSpan_contains_spanset[of us] lin_independent_length_le_dim[of us] SubspaceD1 VectorSpace.dim_eq_size_basis[of "(\)" U us] by auto qed lemma Subspace_eqdim_imp_equal : assumes "Subspace U" "dim U = dim V" shows "U = V" proof- from assms(1) obtain us where us: "basis_for U us" using Subspace_is_findim SubspaceD1 VectorSpace.FinDimVectorSpaceI[of "(\)" U] FinDimVectorSpace.basis_ex[of "(\)" U] by auto with assms(1) obtain vs where vs: "basis_for V (vs@us)" using extend_Subspace_basis[of U us] by fast from assms us vs show ?thesis using SubspaceD1 VectorSpace.dim_eq_size_basis[of smult U] dim_eq_size_basis[of "vs@us"] by auto qed lemma Subspace_dim_lt : "Subspace U \ U \ V \ dim U < dim V" using Subspace_dim_le Subspace_eqdim_imp_equal by fastforce lemma semisimple : assumes "Subspace U" shows "\W. Subspace W \ (V = W \ U)" proof- from assms obtain us where us: "basis_for U us" using SubspaceD1 Subspace_is_findim VectorSpace.FinDimVectorSpaceI FinDimVectorSpace.basis_ex[of _ U] by fastforce with assms obtain ws where basis: "basis_for V (ws@us)" using extend_Subspace_basis by fastforce hence ws_V: "set ws \ V" and ind_ws_us: "lin_independent (ws@us)" and V_eq: "V = Span (ws@us)" by auto have "V = Span ws \ Span us" proof (rule inner_dirsum_doubleI) from V_eq show "V = Span ws + Span us" using Span_append by fast from ws_V ind_ws_us show "add_independentS [Span ws, Span us]" using lin_independent_append_imp_independent_Spans by auto qed with us ws_V have "Subspace (Span ws) \ V = (Span ws) \ U" using Subspace_Span by auto thus ?thesis by fast qed end (* context FinDimVectorSpace *) subsection \Vector space homomorphisms\ subsubsection \Locales\ locale VectorSpaceHom = ModuleHom smult V smult' T for smult :: "'f::field \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" and smult' :: "'f \ 'w::ab_group_add \ 'w" (infixr "\" 70) and T :: "'v \ 'w" sublocale VectorSpaceHom < VectorSpace .. lemmas (in VectorSpace) VectorSpaceHomI = ModuleHomI[THEN VectorSpaceHom.intro] lemma (in VectorSpace) VectorSpaceHomI_fromaxioms : assumes "\g g'. g \ V \ g' \ V \ T (g + g') = T g + T g'" "supp T \ V" "\r m. r \ UNIV \ m \ V \ T (r \ m) = smult' r (T m)" shows "VectorSpaceHom smult V smult' T" using assms by unfold_locales locale VectorSpaceEnd = VectorSpaceHom smult V smult T for smult :: "'f::field \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" and T :: "'v \ 'v" + assumes endomorph: "ImG \ V" abbreviation (in VectorSpace) "VEnd \ VectorSpaceEnd smult V" lemma VectorSpaceEndI : fixes smult :: "'f::field \ 'v::ab_group_add \ 'v" assumes "VectorSpaceHom smult V smult T" "T ` V \ V" shows "VectorSpaceEnd smult V T" by (rule VectorSpaceEnd.intro, rule assms(1), unfold_locales, rule assms(2)) lemma (in VectorSpaceEnd) VectorSpaceHom: "VectorSpaceHom smult V smult T" .. lemma (in VectorSpaceEnd) ModuleEnd : "ModuleEnd smult V T" using endomorph ModuleEnd.intro by unfold_locales locale VectorSpaceIso = VectorSpaceHom smult V smult' T for smult :: "'f::field \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" and smult' :: "'f \ 'w::ab_group_add \ 'w" (infixr "\" 70) and T :: "'v \ 'w" + fixes W :: "'w set" assumes bijective: "bij_betw T V W" abbreviation (in VectorSpace) isomorphic :: "('f \ 'w::ab_group_add \ 'w) \ 'w set \ bool" where "isomorphic smult' W \ (\ T. VectorSpaceIso smult V smult' T W)" subsubsection \Basic facts\ lemma (in VectorSpace) trivial_VectorSpaceHom : "(\a. smult' a 0 = 0) \ VectorSpaceHom smult V smult' 0" using trivial_RModuleHom[of smult'] ModuleHom.intro VectorSpaceHom.intro by fast lemma (in VectorSpace) VectorSpaceHom_idhom : "VectorSpaceHom smult V smult (id\V)" using smult_zero RModHom_idhom ModuleHom.intro VectorSpaceHom.intro by fast context VectorSpaceHom begin lemmas hom = hom lemmas supp = supp lemmas f_map = R_map lemmas im_zero = im_zero lemmas im_sum_list_prod = im_sum_list_prod lemmas additive = additive lemmas GroupHom = GroupHom lemmas distrib_lincomb = distrib_lincomb lemmas same_image_on_spanset_imp_same_hom = same_image_on_RSpanset_imp_same_hom[ OF ModuleHom.axioms(1), OF VectorSpaceHom.axioms(1) ] lemma VectorSpace_Im : "VectorSpace smult' ImG" using RModule_Im VectorSpace.intro Module.intro by fast lemma VectorSpaceHom_scalar_mul : "VectorSpaceHom smult V smult' (\v. a \ T v)" proof show "\v v'. v \ V \ v' \ V \ a \ T (v + v') = a \ T v + a \ T v'" using additive VectorSpace.smult_distrib_left[OF VectorSpace_Im] by simp have "\v. v \ V \ v \ supp (\v. a \ T v)" proof- fix v assume "v \ V" hence "a \ T v = 0" using supp suppI_contra[of _ T] codomain_smult_zero by fastforce thus "v \ supp (\v. a \ T v)" using suppD_contra by fast qed thus "supp (\v. a \ T v) \ V" by fast show "\c v. v \ V \ a \ T (c \ v) = c \ a \ T v" using f_map VectorSpace.smult_assoc[OF VectorSpace_Im] by (simp add: field_simps) qed lemma VectorSpaceHom_composite_left : assumes "ImG \ W" "VectorSpaceHom smult' W smult'' S" shows "VectorSpaceHom smult V smult'' (S \ T)" proof- have "RModuleHom UNIV smult' W smult'' S" using VectorSpaceHom.axioms(1)[OF assms(2)] ModuleHom.axioms(1) by fast with assms(1) have "RModuleHom UNIV smult V smult'' (S \ T)" using RModHom_composite_left[of W] by fast thus ?thesis using ModuleHom.intro VectorSpaceHom.intro by fast qed lemma findim_domain_findim_image : assumes "findim V" shows "fscalar_mult.findim smult' ImG" proof- from assms obtain vs where vs: "set vs \ V" "scalar_mult.Span smult vs = V" by fast define ws where "ws = map T vs" with vs(1) have 1: "set ws \ ImG" by auto moreover have "Span ws = ImG" proof show "Span ws \ ImG" using 1 VectorSpace.Span_closed[OF VectorSpace_Im] by fast from vs ws_def show "Span ws \ ImG" using 1 SpanD_lincomb_arb_len_coeffs distrib_lincomb VectorSpace.SpanD_lincomb_arb_len_coeffs[OF VectorSpace_Im] by auto qed ultimately show ?thesis by fast qed end (* context VectorSpaceHom *) lemma (in VectorSpace) basis_im_defines_hom : fixes smult' :: "'f \ 'w::ab_group_add \ 'w" (infixr "\" 70) and lincomb' :: "'f list \ 'w list \ 'w" (infixr "\\" 70) defines lincomb' : "lincomb' \ scalar_mult.lincomb smult'" assumes VSpW : "VectorSpace smult' W" and basisV : "basis_for V vs" and basisV_im : "set ws \ W" "length ws = length vs" shows "\! T. VectorSpaceHom smult V smult' T \ map T vs = ws" proof (rule ex_ex1I) define T where "T = restrict0 (\v. (THE as. length as = length vs \ v = as \\ vs) \\ ws) V" have "VectorSpaceHom (\) V smult' T" proof fix v v' assume vv': "v \ V" "v' \ V" with T_def lincomb' basisV basisV_im(1) show "T (v + v') = T v + T v'" using basis_for_obtain_unique_scalars theI'[ of "\ds. length ds = length vs \ v = ds \\ vs" ] theI'[of "\ds. length ds = length vs \ v' = ds \\ vs"] add_closed add_unique_scalars VectorSpace.lincomb_sum[OF VSpW] by auto next from T_def show "supp T \ V" using supp_restrict0 by fast next fix a v assume v: "v \ V" with basisV basisV_im(1) T_def lincomb' show "T (a \ v) = a \ T v" using smult_closed smult_unique_scalars VectorSpace.smult_lincomb[OF VSpW] by auto qed moreover have "map T vs = ws" proof (rule nth_equalityI) from basisV_im(2) show "length (map T vs) = length ws" by simp have "\i. i map T vs ! i = ws ! i" proof- fix i assume i: "i < length (map T vs)" define zs where "zs = (replicate (length vs) (0::'f))[i:=1]" with basisV i have "length zs = length vs \ vs!i = zs \\ vs" using delta_scalars_lincomb_eq_nth by auto moreover from basisV i have "vs!i \ V" by auto ultimately show "(map T vs)!i = ws!i" using basisV basisV_im T_def lincomb' zs_def i basis_for_obtain_unique_scalars[of vs "vs!i"] the1_equality[of "\zs. length zs = length vs \ vs!i = zs \\ vs"] VectorSpace.delta_scalars_lincomb_eq_nth[OF VSpW, of ws] by force qed thus "\i. i < length (map T vs) \ map T vs ! i = ws ! i" by fast qed ultimately have "VectorSpaceHom (\) V smult' T \ map T vs = ws" by fast thus "\T. VectorSpaceHom (\) V smult' T \ map T vs = ws" by fast next fix S T assume "VectorSpaceHom (\) V smult' S \ map S vs = ws" "VectorSpaceHom (\) V smult' T \ map T vs = ws" with basisV show "S = T" using VectorSpaceHom.same_image_on_spanset_imp_same_hom map_eq_conv by fastforce (* slow *) qed subsubsection \Hom-sets\ definition VectorSpaceHomSet :: "('f::field \ 'v::ab_group_add \ 'v) \ 'v set \ ('f \ 'w::ab_group_add \ 'w) \ 'w set \ ('v \ 'w) set" where "VectorSpaceHomSet fsmult V fsmult' W \ {T. VectorSpaceHom fsmult V fsmult' T} \ {T. T ` V \ W}" abbreviation (in VectorSpace) "VectorSpaceEndSet \ {S. VEnd S}" lemma VectorSpaceHomSetI : "VectorSpaceHom fsmult V fsmult' T \ T ` V \ W \ T \ VectorSpaceHomSet fsmult V fsmult' W" unfolding VectorSpaceHomSet_def by fast lemma VectorSpaceHomSetD_VectorSpaceHom : "T \ VectorSpaceHomSet fsmult V fsmult' N \ VectorSpaceHom fsmult V fsmult' T" unfolding VectorSpaceHomSet_def by fast lemma VectorSpaceHomSetD_Im : "T \ VectorSpaceHomSet fsmult V fsmult' W \ T ` V \ W" unfolding VectorSpaceHomSet_def by fast context VectorSpace begin lemma VectorSpaceHomSet_is_fmaps_in_GroupHomSet : fixes smult' :: "'f \ 'w::ab_group_add \ 'w" (infixr "\" 70) shows "VectorSpaceHomSet smult V smult' W = (GroupHomSet V W) \ {T. \a. \v\V. T (a \ v) = a \ (T v)}" proof show "VectorSpaceHomSet smult V smult' W \ (GroupHomSet V W) \ {T. \a. \v\V. T (a \ v) = a \ (T v)}" using VectorSpaceHomSetD_VectorSpaceHom[of _ smult] VectorSpaceHomSetD_Im[of _ smult] VectorSpaceHom.GroupHom[of smult] GroupHomSetI[of V _ W] VectorSpaceHom.f_map[of smult] by fastforce show "VectorSpaceHomSet smult V smult' W \ (GroupHomSet V W) \ {T. \a. \v\V. T (a \ v) = a \ (T v)}" proof fix T assume T: "T \ (GroupHomSet V W) \ {T. \a. \v\V. T (a \ v) = a \ (T v)}" have "VectorSpaceHom smult V smult' T" proof (rule VectorSpaceHom.intro, rule ModuleHom.intro, rule RModuleHom.intro) show "RModule UNIV (\) V" .. from T show "GroupHom V T" using GroupHomSetD_GroupHom by fast from T show "RModuleHom_axioms UNIV smult V smult' T" by unfold_locales fast qed with T show "T \ VectorSpaceHomSet smult V smult' W" using GroupHomSetD_Im[of T] VectorSpaceHomSetI by fastforce qed qed lemma Group_VectorSpaceHomSet : fixes smult' :: "'f \ 'w::ab_group_add \ 'w" (infixr "\" 70) assumes "VectorSpace smult' W" shows "Group (VectorSpaceHomSet smult V smult' W)" proof show "VectorSpaceHomSet smult V smult' W \ {}" using VectorSpace.smult_zero[OF assms] VectorSpace.zero_closed[OF assms] trivial_VectorSpaceHom[of smult'] VectorSpaceHomSetI by fastforce next fix S T assume S: "S \ VectorSpaceHomSet smult V smult' W" and T: "T \ VectorSpaceHomSet smult V smult' W" from S T have ST: "S \ (GroupHomSet V W) \ {T. \a. \v\V. T (a \ v) = a \ (T v)}" "T \ (GroupHomSet V W) \ {T. \a. \v\V. T (a \ v) = a \ (T v)}" using VectorSpaceHomSet_is_fmaps_in_GroupHomSet by auto hence "S - T \ GroupHomSet V W" using VectorSpace.AbGroup[OF assms] Group_GroupHomSet Group.diff_closed by fast moreover have "\a v. v \ V \ (S - T) (a \ v) = a \ ((S-T) v)" proof- fix a v assume "v \ V" with ST show "(S - T) (a \ v) = a \ ((S - T) v)" using GroupHomSetD_Im VectorSpace.smult_distrib_left_diff[OF assms, of a "S v" "T v"] by fastforce qed ultimately show "S - T \ VectorSpaceHomSet (\) V (\) W" using VectorSpaceHomSet_is_fmaps_in_GroupHomSet[of smult' W] by fast qed lemma VectorSpace_VectorSpaceHomSet : fixes smult' :: "'f \ 'w::ab_group_add \ 'w" (infixr "\" 70) and hom_smult :: "'f \ ('v \ 'w) \ ('v \ 'w)" (infixr "\\" 70) defines hom_smult: "hom_smult \ \a T v. a \ T v" assumes VSpW: "VectorSpace smult' W" shows "VectorSpace hom_smult (VectorSpaceHomSet smult V smult' W)" proof (rule VectorSpace.intro, rule Module.intro, rule RModule.intro, rule R_scalar_mult) from VSpW show "Group (VectorSpaceHomSet (\) V (\) W)" using Group_VectorSpaceHomSet by fast show "RModule_axioms UNIV hom_smult (VectorSpaceHomSet (\) V (\) W)" proof fix a b S T assume S: "S \ VectorSpaceHomSet (\) V (\) W" and T: "T \ VectorSpaceHomSet (\) V (\) W" show "a \\ T \ VectorSpaceHomSet (\) V (\) W" proof (rule VectorSpaceHomSetI) from assms T show "VectorSpaceHom (\) V (\) (a \\ T)" using VectorSpaceHomSetD_VectorSpaceHom VectorSpaceHomSetD_Im VectorSpaceHom.VectorSpaceHom_scalar_mul by fast from hom_smult show "(a \\ T) ` V \ W" using VectorSpaceHomSetD_Im[OF T] VectorSpace.smult_closed[OF VSpW] by auto qed show "a \\ (S + T) = a \\ S + a \\ T" proof fix v from assms show "(a \\ (S + T)) v = (a \\ S + a \\ T) v" using VectorSpaceHomSetD_Im[OF S] VectorSpaceHomSetD_Im[OF T] VectorSpace.smult_distrib_left[OF VSpW, of a "S v" "T v"] VectorSpaceHomSetD_VectorSpaceHom[OF S] VectorSpaceHomSetD_VectorSpaceHom[OF S] VectorSpaceHom.supp suppI_contra[of v S] suppI_contra[of v T] VectorSpace.smult_zero by fastforce qed show "(a + b) \\ T = a \\ T + b \\ T" proof fix v from assms show "((a + b) \\ T) v = (a \\ T + b \\ T) v" using VectorSpaceHomSetD_Im[OF T] VectorSpace.smult_distrib_right VectorSpaceHomSetD_VectorSpaceHom[OF T] VectorSpaceHom.supp suppI_contra[of v] VectorSpace.smult_zero by fastforce qed show "a \\ b \\ T = (a * b) \\ T" proof fix v from assms show "(a \\ b \\ T) v = ((a * b) \\ T) v" using VectorSpaceHomSetD_Im[OF T] VectorSpace.smult_assoc VectorSpaceHomSetD_VectorSpaceHom[OF T] VectorSpaceHom.supp suppI_contra[of v] VectorSpace.smult_zero[OF VSpW, of b] VectorSpace.smult_zero[OF VSpW, of a] VectorSpace.smult_zero[OF VSpW, of "a * b"] by fastforce qed show "1 \\ T = T" proof fix v from assms T show "(1 \\ T) v = T v" using VectorSpaceHomSetD_Im VectorSpace.one_smult VectorSpaceHomSetD_VectorSpaceHom VectorSpaceHom.supp suppI_contra[of v] VectorSpace.smult_zero by fastforce qed qed qed end (* context VectorSpace *) subsubsection \Basic facts about endomorphisms\ lemma ModuleEnd_over_field_is_VectorSpaceEnd : fixes smult :: "'f::field \ 'v::ab_group_add \ 'v" assumes "ModuleEnd smult V T" shows "VectorSpaceEnd smult V T" proof (rule VectorSpaceEndI, rule VectorSpaceHom.intro) from assms show "ModuleHom smult V smult T" using ModuleEnd.axioms(1) by fast from assms show "T ` V \ V" using ModuleEnd.endomorph by fast qed context VectorSpace begin lemmas VectorSpaceEnd_inner_dirsum_el_decomp_nth = RModuleEnd_inner_dirsum_el_decomp_nth[ THEN RModuleEnd_over_UNIV_is_ModuleEnd, THEN ModuleEnd_over_field_is_VectorSpaceEnd ] abbreviation end_smult :: "'f \ ('v \ 'v) \ ('v \ 'v)" (infixr "\\" 70) where "a \\ T \ (\v. a \ T v)" abbreviation end_lincomb :: "'f list \ (('v \ 'v) list) \ ('v \ 'v)" (infixr "\\\" 70) where "end_lincomb \ scalar_mult.lincomb end_smult" lemma end_smult0: "a \\ 0 = 0" using smult_zero by auto lemma end_0smult: "range T \ V \ 0 \\ T = 0" using zero_smult by fastforce lemma end_smult_distrib_left : assumes "range S \ V" "range T \ V" shows "a \\ (S + T) = a \\ S + a \\ T" proof fix v from assms show "(a \\ (S + T)) v = (a \\ S + a \\ T) v" using smult_distrib_left[of a "S v" "T v"] by fastforce qed lemma end_smult_distrib_right : assumes "range T \ V" shows "(a+b) \\ T = a \\ T + b \\ T" proof fix v from assms show "((a+b) \\ T) v = (a \\ T + b \\ T) v" using smult_distrib_right[of a b "T v"] by fastforce qed lemma end_smult_assoc : assumes "range T \ V" shows "a \\ b \\ T = (a * b) \\ T" proof fix v from assms show "(a \\ b \\ T) v = ((a * b) \\ T) v" using smult_assoc[of a b "T v"] by fastforce qed lemma end_smult_comp_comm_left : "(a \\ T) \ S = a \\ (T \ S)" by auto lemma end_idhom : "VEnd (id\V)" by (rule VectorSpaceEnd.intro, rule VectorSpaceHom_idhom, unfold_locales) auto lemma VectorSpaceEndSet_is_VectorSpaceHomSet : "VectorSpaceHomSet smult V smult V = {T. VEnd T}" proof show "VectorSpaceHomSet smult V smult V \ {T. VEnd T}" using VectorSpaceHomSetD_VectorSpaceHom VectorSpaceHomSetD_Im VectorSpaceEndI by fast show "VectorSpaceHomSet smult V smult V \ {T. VEnd T}" using VectorSpaceEnd.VectorSpaceHom[of smult V] VectorSpaceEnd.endomorph[of smult V] VectorSpaceHomSetI[of smult V smult _ V] by fast qed lemma VectorSpace_VectorSpaceEndSet : "VectorSpace end_smult VectorSpaceEndSet" using VectorSpace_axioms VectorSpace_VectorSpaceHomSet VectorSpaceEndSet_is_VectorSpaceHomSet by fastforce end (* context VectorSpace *) context VectorSpaceEnd begin lemmas f_map = R_map lemmas supp = supp lemmas GroupEnd = ModuleEnd.GroupEnd[OF ModuleEnd] lemmas idhom_left = idhom_left lemmas range = GroupEnd.range[OF GroupEnd] lemmas Ker0_imp_inj_on = Ker0_imp_inj_on lemmas inj_on_imp_Ker0 = inj_on_imp_Ker0 lemmas nonzero_Ker_el_imp_n_inj = nonzero_Ker_el_imp_n_inj lemmas VectorSpaceHom_composite_left = VectorSpaceHom_composite_left[OF endomorph] lemma in_VEndSet : "T \ VectorSpaceEndSet" using VectorSpaceEnd_axioms by fast lemma end_smult_comp_comm_right : "range S \ V \ T \ (a \\ S) = a \\ (T \ S)" using f_map by fastforce lemma VEnd_end_smult_VEnd : "VEnd (a \\ T)" using in_VEndSet VectorSpace.smult_closed[OF VectorSpace_VectorSpaceEndSet] by fast lemma VEnd_composite_left : assumes "VEnd S" shows "VEnd (S \ T)" using endomorph VectorSpaceEnd.axioms(1)[OF assms] VectorSpaceHom_composite_left VectorSpaceEnd.endomorph[OF assms] VectorSpaceEndI[of smult V "S \ T"] by fastforce lemma VEnd_composite_right : "VEnd S \ VEnd (T \ S)" using VectorSpaceEnd_axioms VectorSpaceEnd.VEnd_composite_left by fast end (* context VectorSpaceEnd *) lemma (in VectorSpace) inj_comp_end : assumes "VEnd S" "inj_on S V" "VEnd T" "inj_on T V" shows "inj_on (S \ T) V" proof- have "ker (S \ T) \ V \ 0" proof fix v assume "v \ ker (S \ T) \ V" moreover hence "T v = 0" using kerD[of v "S \ T"] using VectorSpaceEnd.endomorph[OF assms(3)] kerI[of S] VectorSpaceEnd.inj_on_imp_Ker0[OF assms(1,2)] by auto ultimately show "v \ 0" using kerI[of T] VectorSpaceEnd.inj_on_imp_Ker0[OF assms(3,4)] by auto qed with assms(1,3) show ?thesis using VectorSpaceEnd.VEnd_composite_right VectorSpaceEnd.Ker0_imp_inj_on by fast qed lemma (in VectorSpace) n_inj_comp_end : "\ VEnd S; VEnd T; \ inj_on (S \ T) V \ \ \ inj_on S V \ \ inj_on T V" using inj_comp_end by fast subsubsection \Polynomials of endomorphisms\ context VectorSpaceEnd begin primrec endpow :: "nat \ ('v\'v)" where endpow0: "endpow 0 = id\V" | endpowSuc: "endpow (Suc n) = T \ (endpow n)" definition polymap :: "'f poly \ ('v\'v)" where "polymap p \ (coeffs p) \\\ (map endpow [0.. (endpow k))" using VEnd_composite_right by fast moreover have "endpow (Suc k) = T \ (endpow k)" by simp ultimately show "VEnd (endpow (Suc k))" by simp qed lemma endpow_list_apply_closed : "v \ V \ set (map (\S. S v) (map endpow [0.. V" using VEnd_endpow VectorSpaceEnd.endomorph by fastforce lemma map_endpow_Suc : "map endpow [0..V) # map ((\) T) (map endpow [0.. V # map ((\) T) (map endpow [0..) T) [endpow k]" by auto also have "\ = id \ V # map ((\) T) (map endpow ([0..S. S v) (map endpow [0..S. S v) (map ((\) T) (map endpow [0.. {S. VEnd S}" using VEnd_endpow by auto thus ?thesis using polymap_def VectorSpace_VectorSpaceEndSet VectorSpace.lincomb_closed by fastforce qed lemma polymap_pCons : "polymap (pCons a p) = a \\ (id\V) + (T \ (polymap p))" proof cases assume p: "p = 0" show ?thesis proof cases assume "a = 0" with p show ?thesis using polymap0 VectorSpace_VectorSpaceEndSet VectorSpace.zero_smult end_idhom comp_zero by fastforce next assume a: "a \ 0" define zmap where "zmap = (0::'v\'v)" from a p have "polymap (pCons a p) = a \\ (endpow 0)" using polymap_def scalar_mult.lincomb_singles by simp moreover have "a \\ (id\V) + (T \ (polymap p)) = a \\ (id\V)" using p polymap0 comp_zero by simp ultimately show ?thesis by simp qed next assume "p \ 0" hence "polymap (pCons a p) = (a # (coeffs p)) \\\ (map endpow [0.. = (a # (coeffs p)) \\\ ((id\V) # map ((\) T) (map endpow [0.. = a \\ (id\V) + (coeffs p) \\\ (map ((\) T) (map endpow [0.. = a \\ (id\V) + (\(c,S) \zip (coeffs p) (map ((\) T) (map endpow [0..\ S)" using scalar_mult.lincomb_def by simp finally have calc: "polymap (pCons a p) = a \\ (id\V) + (\(c,k)\zip (coeffs p) [0..\ (T \ (endpow k)))" using sum_list_prod_map2[ of "\c S. c \\ S" "coeffs p" "(\) T" "map endpow [0..c S. c \\ (T \ S)" "coeffs p" endpow "[0..\ (id\V)) + (T \ (polymap p))) v" proof (cases "v \ V") case True with calc have "polymap (pCons a p) v = a \ v + (\(c,k) \zip (coeffs p) [0.. T (endpow k v))" using sum_list_prod_fun_apply[of "\c k. c \\ (T \ (endpow k))"] by simp hence "polymap (pCons a p) v = a \ v + (coeffs p) \\ (map T (map (\S. S v) (map endpow [0..c S. c \ T (S v)" "coeffs p" endpow "[0..c u. c \ T u" "coeffs p" "\S. S v" "map endpow [0..c u. c \ u" "coeffs p" T "map (\S. S v) (map endpow [0.. = a \ v + T ( (coeffs p) \\ (map (\S. S v) (map endpow [0..c u. c \ u" "coeffs p" "\S. S v" "map endpow [0..c S. c \\ S"] polymap_def scalar_mult.lincomb_def[of end_smult] by simp next case False hence "polymap (pCons a p) v = 0" using VEnd_polymap VectorSpaceEnd.supp suppI_contra by fast moreover from False have "((a \\ (id\V)) + (T \ (polymap p))) v = 0" using smult_zero VEnd_polymap[of p] VectorSpaceEnd.supp suppI_contra im_zero by fastforce ultimately show ?thesis by simp qed qed qed lemma polymap_plus : "polymap (p + q) = polymap p + polymap q" proof (induct p q rule: pCons_induct2) case 00 show ?case using polymap0 by simp case lpCons show ?case using polymap0 by simp case rpCons show ?case using polymap0 by simp next case (pCons2 a p b q) have "polymap (pCons a p + pCons b q) = a \\ (id\V) + b \\ (id\V) + (T \ (polymap (p+q)))" using polymap_pCons end_idhom end_smult_distrib_right[OF VectorSpaceEnd.range] by simp also from pCons2(3) have "\ = a \\ (id\V) + b \\ (id\V) + (T \ (polymap p + polymap q))" by auto finally show ?case using pCons2(3) distrib_comp_sum_left[of "polymap p" "polymap q"] VEnd_polymap VectorSpaceEnd.range polymap_pCons by fastforce qed lemma polymap_polysmult : "polymap (Polynomial.smult a p) = a \\ polymap p" proof (induct p) case 0 show "polymap (Polynomial.smult a 0) = a \\ polymap 0" using polymap0 end_smult0 by simp next case (pCons b p) hence "polymap (Polynomial.smult a (pCons b p)) = a \\ b \\ (id\V) + a \\ (T \ polymap p)" using polymap_pCons VectorSpaceEnd.range[OF VEnd_polymap] end_smult_comp_comm_right VectorSpaceEnd.range[OF end_idhom] end_smult_assoc by simp thus "polymap (Polynomial.smult a (pCons b p)) = a \\ (polymap (pCons b p))" using VectorSpaceEnd.VEnd_end_smult_VEnd[OF end_idhom, of b] VEnd_composite_right[OF VEnd_polymap, of p] end_smult_distrib_left[ OF VectorSpaceEnd.range VectorSpaceEnd.range, of smult _ smult "T \ polymap p" ] polymap_pCons by simp qed lemma polymap_times : "polymap (p * q) = (polymap p) \ (polymap q)" proof (induct p) case 0 show ?case using polymap0 by auto next case (pCons a p) have "polymap (pCons a p * q) = a \\ polymap q + (T \ (polymap (p*q)))" using polymap_plus polymap_polysmult polymap_pCons end_idhom end_0smult[OF VectorSpaceEnd.range] by simp also from pCons(2) have "\ = a \\ ((id\V) \ polymap q) + (T \ polymap p \ polymap q)" using VectorSpaceEnd.endomorph[OF VEnd_polymap] VectorSpaceEnd.idhom_left[OF VEnd_polymap] by auto finally show "polymap (pCons a p * q) = (polymap (pCons a p)) \ (polymap q)" using end_smult_comp_comm_left distrib_comp_sum_right[of "a \\ id \ V" _ "polymap q"] polymap_pCons by simp qed lemma polymap_apply : assumes "v \ V" shows "polymap p v = (coeffs p) \\ (map (\S. S v) (map endpow [0..\ id\V" using polymap_pCons polymap0 comp_zero by simp ultimately show ?thesis using assms pCons(1) lincomb_singles by simp next case False from assms pCons(2) have "polymap (pCons a p) v = a \ v + T (coeffs p \\ map (\S. S v) (map endpow [0..\ (v # map T (map (\S. S v) (map endpow [0..S. S v) (map endpow [0..S. S v) (map ((\) T) (map endpow [0..\ (v # map (\S. S v) (map ((\) T) (map endpow [0..x. polymap (pCons a p) v = (coeffs (pCons a p)) \\ (v # x)" ] by simp with assms have 3: "polymap (pCons a p) v = (coeffs (pCons a p)) \\ (map (\S. S v) (id\V # map ((\) T) (map endpow [0.. V # map ((\) T) (map endpow [0..x. polymap (pCons a p) v = (coeffs (pCons a p)) \\ (map (\S. S v) x)" ] by simp qed qed lemma polymap_apply_linear : "v \ V \ polymap [:-c, 1:] v = T v - c \ v" using polymap_apply lincomb_def neg_smult endomorph by auto lemma polymap_const_inj : assumes "degree p = 0" "p \ 0" shows "inj_on (polymap p) V" proof (rule inj_onI) fix u v assume uv: "u \ V" "v \ V" "polymap p u = polymap p v" from assms have p: "coeffs p = [coeff p 0]" unfolding coeffs_def by simp from uv assms have "(coeff p 0) \ u = (coeff p 0) \ v" using polymap_apply lincomb_singles unfolding coeffs_def by simp with assms uv(1,2) show "u = v" using const_poly_nonzero_coeff cancel_scalar by auto qed lemma n_inj_polymap_times : "\ inj_on (polymap (p * q)) V \ \ inj_on (polymap p) V \ \ inj_on (polymap q) V" using polymap_times VEnd_polymap n_inj_comp_end by fastforce text \In the following lemma, @{term "[:-c, 1:]"} is the linear polynomial @{term "x - c"}.\ lemma n_inj_polymap_findlinear : assumes alg_closed: "\p::'f poly. degree p > 0 \ \c. poly p c = 0" shows "p \ 0 \ \ inj_on (polymap p) V \ \c. \ inj_on (polymap [:-c, 1:]) V" proof (induct n \ "degree p" arbitrary: p) case (0 p) thus ?case using polymap_const_inj by simp next case (Suc n p) from Suc(2) alg_closed obtain c where c: "poly p c = 0" by fastforce define q where "q = synthetic_div p c" with c have p_decomp: "p = [:-c, 1:] * q" using synthetic_div_correct'[of c p] by simp show ?case proof (cases "inj_on (polymap q) V") case True with Suc(4) show ?thesis using p_decomp n_inj_polymap_times by fast next case False then have "n = degree q" using degree_synthetic_div [of p c] q_def \Suc n = degree p\ by auto moreover have "q \ 0" using \p \ 0\ p_decomp by auto ultimately show ?thesis using False by (rule Suc.hyps) qed qed end (* context VectorSpaceEnd *) subsubsection \Existence of eigenvectors of endomorphisms of finite-dimensional vector spaces\ lemma (in FinDimVectorSpace) endomorph_has_eigenvector : assumes alg_closed: "\p::'a poly. degree p > 0 \ \c. poly p c = 0" and dim : "dim V > 0" and endo : "VectorSpaceEnd smult V T" shows "\c u. u \ V \ u \ 0 \ T u = c \ u" proof- define Tpolymap where "Tpolymap = VectorSpaceEnd.polymap smult V T" from dim obtain v where v: "v \ V" "v \ 0" using dim_nonzero nonempty by auto define Tpows where "Tpows = map (VectorSpaceEnd.endpow V T) [0..S. S v) Tpows" with endo Tpows_def v(1) have Tpows_v_V: "set Tpows_v \ V" using VectorSpaceEnd.endpow_list_apply_closed by fast moreover from Tpows_v_def Tpows_def Tpows_v_V have "\ lin_independent Tpows_v" using too_long_lin_dependent by simp ultimately obtain as where as: "set as \ 0" "length as = length Tpows_v" "as \\ Tpows_v = 0" using lin_dependent_dependence_relation by fast define p where "p = Poly as" with dim Tpows_def Tpows_v_def as(1,2) have p_n0: "p \ 0" using nonzero_coeffs_nonzero_poly[of as] by fastforce define Tpows' where "Tpows' = map (VectorSpaceEnd.endpow V T) [0..S. S v) Tpows'" have "Tpows' = take (Suc (degree p)) Tpows" proof- from Tpows_def have 1: "take (Suc (degree p)) Tpows = map (VectorSpaceEnd.endpow V T) (take (Suc (degree p)) [0..S. S v" Tpows] by simp moreover from p_def Tpows_v_V as(3) Tpows_v'_def have "(coeffs p) \\ Tpows_v = 0" using lincomb_strip_while_0coeffs by simp ultimately have "(coeffs p) \\ Tpows_v' = 0" using p_n0 lincomb_conv_take_right[of "coeffs p"] length_coeffs_degree[of p] by simp with Tpolymap_def v(1) Tpows_v'_def Tpows'_def have "Tpolymap p v = 0" using VectorSpaceEnd.polymap_apply[OF endo] by simp with alg_closed Tpolymap_def v endo p_n0 obtain c where "\ inj_on (Tpolymap [:-c, 1:]) V" using VectorSpaceEnd.VEnd_polymap VectorSpaceEnd.nonzero_Ker_el_imp_n_inj VectorSpaceEnd.n_inj_polymap_findlinear[OF endo] by fastforce with Tpolymap_def have "(GroupHom.Ker V (Tpolymap [:-c, 1:])) - 0 \ {}" using VectorSpaceEnd.VEnd_polymap[OF endo] VectorSpaceEnd.Ker0_imp_inj_on by fast from this obtain u where "u \ V" "Tpolymap [:-c, 1:] u = 0" "u \ 0" using kerD by fastforce with Tpolymap_def show ?thesis using VectorSpaceEnd.polymap_apply_linear[OF endo] by auto qed section \Modules Over a Group Ring\ subsection \Almost-everywhere-zero functions as scalars\ locale aezfun_scalar_mult = scalar_mult smult for smult :: "('r::ring_1, 'g::group_add) aezfun \ 'v::ab_group_add \ 'v" (infixr "\" 70) begin definition fsmult :: "'r \ 'v \ 'v" (infixr "\\" 70) where "a \\ v \ (a \\ 0) \ v" abbreviation flincomb :: "'r list \ 'v list \ 'v" (infixr "\\\" 70) where "as \\\ vs \ scalar_mult.lincomb fsmult as vs" abbreviation f_lin_independent :: "'v list \ bool" where "f_lin_independent \ scalar_mult.lin_independent fsmult" abbreviation fSpan :: "'v list \ 'v set" where "fSpan \ scalar_mult.Span fsmult" definition Gmult :: "'g \ 'v \ 'v" (infixr "*\" 70) where "g *\ v \ (1 \\ g) \ v" lemmas R_scalar_mult = R_scalar_mult lemma fsmultD : "a \\ v = (a \\ 0) \ v" unfolding fsmult_def by fast lemma GmultD : "g *\ v = (1 \\ g) \ v" unfolding Gmult_def by fast definition negGorbit_list :: "'g list \ ('a \ 'v) \ 'a list \ 'v list list" where "negGorbit_list gs T as \ map (\g. map (Gmult (-g) \ T) as) gs" lemma negGorbit_Cons : "negGorbit_list (g#gs) T as = (map (Gmult (-g) \ T) as) # negGorbit_list gs T as" using negGorbit_list_def[of _ T as] by simp lemma length_negGorbit_list : "length (negGorbit_list gs T as) = length gs" using negGorbit_list_def[of gs T] by simp lemma length_negGorbit_list_sublist : "fs \ set (negGorbit_list gs T as) \ length fs = length as" using negGorbit_list_def[of gs T] by auto lemma length_concat_negGorbit_list : "length (concat (negGorbit_list gs T as)) = (length gs) * (length as)" using length_concat[of "negGorbit_list gs T as"] length_negGorbit_list_sublist[of _ gs T as] const_sum_list[of "negGorbit_list gs T as" length "length as"] length_negGorbit_list by auto lemma negGorbit_list_nth : "\i. i < length gs \ (negGorbit_list gs T as)!i = map (Gmult (-gs!i) \ T) as" proof (induct gs) case (Cons g gs) thus ?case using negGorbit_Cons[of _ _ T] by (cases i) auto qed simp end (* context aezfun_scalar_mult *) subsection \Locale and basic facts\ locale FGModule = ActingGroup?: Group G + FGMod?: RModule ActingGroup.group_ring smult V for G :: "'g::group_add set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" sublocale FGModule < aezfun_scalar_mult proof- qed lemma (in Group) trivial_FGModule : fixes smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" assumes smult_zero: "\a\group_ring. smult a (0::'v) = 0" shows "FGModule G smult (0::'v set)" proof (rule FGModule.intro) from assms show "RModule group_ring smult 0" using Ring1_RG trivial_RModule by fast qed (unfold_locales) context FGModule begin abbreviation FG :: "('f,'g) aezfun set" where "FG \ ActingGroup.group_ring" abbreviation "FGSubmodule \ RSubmodule" abbreviation "FG_proj \ ActingGroup.RG_proj" lemma GroupG: "Group G" .. lemmas zero_closed = zero_closed lemmas neg_closed = neg_closed lemmas diff_closed = diff_closed lemmas zero_smult = zero_smult lemmas smult_zero = smult_zero lemmas AbGroup = AbGroup lemmas sum_closed = AbGroup.sum_closed[OF AbGroup] lemmas FGSubmoduleI = RSubmoduleI lemmas FG_proj_mult_leftdelta = ActingGroup.RG_proj_mult_leftdelta lemmas FG_proj_mult_right = ActingGroup.RG_proj_mult_right lemmas FG_el_decomp = ActingGroup.RG_el_decomp_aezdeltafun lemma FG_n0: "FG \ 0" using ActingGroup.RG_n0 by fast lemma FG_proj_in_FG : "FG_proj x \ FG" using ActingGroup.RG_proj_in_RG by fast lemma FG_fddg_closed : "g \ G \ a \\ g \ FG" using ActingGroup.RG_aezdeltafun_closed by fast lemma FG_fdd0_closed : "a \\ 0 \ FG" using ActingGroup.RG_aezdelta0fun_closed by fast lemma Gmult_closed : "g \ G \ v \ V \ g *\ v \ V" using FG_fddg_closed smult_closed GmultD by simp lemma map_Gmult_closed : "g \ G \ set vs \ V \ set (map ((*\) g) vs) \ V" using Gmult_def FG_fddg_closed map_smult_closed[of "1 \\ g" vs] by auto lemma Gmult0 : assumes "v \ V" shows "0 *\ v = v" proof- have "0 *\ v = (1 \\ 0) \ v" using GmultD by fast moreover have "1 \\ 0 = (1::('f,'g) aezfun)" using one_aezfun_transfer by fast ultimately have "0 *\ v = (1::('f,'g) aezfun) \ v" by simp with assms show ?thesis using one_smult by simp qed lemma Gmult_assoc : assumes "g \ G" "h \ G" "v \ V" shows "g *\ h *\ v = (g + h) *\ v" proof- define n where "n = (1::'f)" with assms have "g *\ h *\ v = ((n \\ g) * (n \\ h)) \ v" using FG_fddg_closed GmultD by simp moreover from n_def have "n \\ g * (n \\ h) = n \\ (g + h)" using times_aezdeltafun_aezdeltafun[of n g n h] by simp ultimately show ?thesis using n_def GmultD by simp qed lemma Gmult_distrib_left : "\ g \ G; v \ V; v' \ V \ \ g *\ (v + v') = g *\ v + g *\ v'" using GmultD FG_fddg_closed by simp lemma neg_Gmult : "g \ G \ v \ V \ g *\ (- v) = - (g *\ v)" using GmultD FG_fddg_closed smult_neg by simp lemma Gmult_neg_left : "g \ G \ v \ V \ (- g) *\ g *\ v = v" using ActingGroup.neg_closed Gmult_assoc[of "- g" g] Gmult0 by simp lemma fddg_smult_decomp : "g \ G \ v \ V \ (f \\ g) \ v = f \\ g *\ v" using aezdeltafun_decomp[of f g] FG_fddg_closed FG_fdd0_closed GmultD fsmult_def by simp lemma sum_list_aezdeltafun_smult_distrib : assumes "v \ V" "set (map snd fgs) \ G" shows "(\(f,g)\fgs. f \\ g) \ v = (\(f,g)\fgs. f \\ g *\ v)" proof- from assms(2) have "set (map (case_prod aezdeltafun) fgs) \ FG" using FG_fddg_closed by auto with assms(1) have "(\(f,g)\fgs. f \\ g) \ v = (\(f,g)\fgs. (f \\ g) \ v)" using sum_list_prod_map_smult_distrib by auto also have "\ = (\(f,g)\fgs. f \\ g *\ v)" using assms fddg_smult_decomp sum_list_prod_cong[of fgs "\ f g. (f \\ g) \ v" "\ f g. f \\ g *\ v"] by fastforce finally show ?thesis by fast qed abbreviation "GSubspace \ RSubmodule" abbreviation "GSpan \ RSpan" abbreviation "G_fingen \ R_fingen" lemma GSubspaceI : "FGModule G smult U \ U \ V \ GSubspace U" using FGModule.axioms(2) by fast lemma GSubspace_is_FGModule : assumes "GSubspace U" shows "FGModule G smult U" proof (rule FGModule.intro, rule GroupG) from assms show "RModule FG (\) U" by fast qed (unfold_locales) lemma restriction_to_subgroup_is_module : fixes H :: "'g set" assumes subgrp: "Group.Subgroup G H" shows "FGModule H smult V" proof (rule FGModule.intro) from subgrp show "Group H" by fast from assms show "RModule (Group.group_ring H) (\) V" using ActingGroup.Subgroup_imp_Subring SModule_restrict_scalars by fast qed lemma negGorbit_list_V : assumes "set gs \ G" "T ` (set as) \ V" shows "set (concat (negGorbit_list gs T as)) \ V" proof- from assms(2) have "set (concat (negGorbit_list gs T as)) \ (\g\set gs. (Gmult (-g)) ` V)" using set_concat negGorbit_list_def[of gs T as] by force moreover from assms(1) have "\g. g \ set gs \ (Gmult (-g)) ` V \ V" using ActingGroup.neg_closed Gmult_closed by fast ultimately show ?thesis by fast qed lemma negGorbit_list_Cons0 : "T ` (set as) \ V \ negGorbit_list (0#gs) T as = (map T as) # (negGorbit_list gs T as)" using negGorbit_Cons[of 0 gs T as] Gmult0 by auto end (* context FGModule *) subsection \Modules over a group ring as a vector spaces\ context FGModule begin lemma fVectorSpace : "VectorSpace fsmult V" proof (rule VectorSpaceI, unfold_locales) fix a::'f show "\v. v \ V \ a \\ v \ V" using fsmult_def smult_closed FG_fdd0_closed by simp next fix a::'f show "\u v. u \ V \ v \ V \ a \\ (u + v) = a \\ u + a \\ v" using fsmult_def FG_fdd0_closed by simp next fix a b :: 'f and v :: 'v assume v: "v \ V" have "(a+b) \\ v = (a \\ 0 + b \\ 0) \ v" using aezdeltafun_plus[of a b 0] arg_cong[of _ _ "\r. r \ v"] fsmult_def by fastforce with v show "(a+b) \\ v = a \\ v + b \\ v" using fsmult_def FG_fdd0_closed by simp next fix a b :: 'f show "\v. v \ V \ a \\ (b \\ v) = (a * b) \\ v" using times_aezdeltafun_aezdeltafun[of a 0 b 0] arg_cong fsmult_def FG_fdd0_closed by simp next fix v :: 'v assume "v \ V" thus "1 \\ v = v" using one_aezfun_transfer arg_cong[of "1 \\ 0" 1 "\a. a \ v"] fsmult_def by fastforce qed abbreviation "fSubspace \ VectorSpace.Subspace fsmult V" abbreviation "fbasis_for \ fscalar_mult.basis_for fsmult" abbreviation "fdim \ scalar_mult.dim fsmult V" lemma VectorSpace_fSubspace : "fSubspace W \ VectorSpace fsmult W" using Module.intro VectorSpace.intro by fast lemma fsmult_closed : "v \ V \ a \\ v \ V" using FG_fdd0_closed smult_closed fsmult_def by simp lemmas one_fsmult [simp] = VectorSpace.one_smult [OF fVectorSpace] lemmas fsmult_assoc [simp] = VectorSpace.smult_assoc [OF fVectorSpace] lemmas fsmult_zero [simp] = VectorSpace.smult_zero [OF fVectorSpace] lemmas fsmult_distrib_left [simp] = VectorSpace.smult_distrib_left [OF fVectorSpace] lemmas flincomb_closed = VectorSpace.lincomb_closed [OF fVectorSpace] lemmas fsmult_sum_distrib = VectorSpace.smult_sum_distrib [OF fVectorSpace] lemmas sum_fsmult_distrib = VectorSpace.sum_smult_distrib [OF fVectorSpace] lemmas flincomb_concat = VectorSpace.lincomb_concat [OF fVectorSpace] lemmas fSpan_closed = VectorSpace.Span_closed [OF fVectorSpace] lemmas flin_independentD_all_scalars = VectorSpace.lin_independentD_all_scalars[OF fVectorSpace] lemmas in_fSpan_obtain_same_length_coeffs = VectorSpace.in_Span_obtain_same_length_coeffs [OF fVectorSpace] lemma fsmult_smult_comm : "r \ FG \ v \ V \ a \\ r \ v = r \ a \\ v" using fsmultD FG_fdd0_closed smult_assoc aezdelta0fun_commutes[of r] by simp lemma fsmult_Gmult_comm : "g \ G \ v \ V \ a \\ g *\ v = g *\ a \\ v" using aezdeltafun_decomp[of a g] aezdeltafun_decomp'[of a g] FG_fddg_closed FG_fdd0_closed fsmult_def GmultD by simp lemma Gmult_flincomb_comm : assumes "g \ G" "set vs \ V" shows "g *\ as \\\ vs = as \\\ (map (Gmult g) vs)" proof- have "g *\ as \\\ vs = (1 \\ g) \ (\(a,v)\zip as vs. a \\ v)" using Gmult_def scalar_mult.lincomb_def[of fsmult] by simp with assms have "g *\ as \\\ vs = sum_list (map ((\) (1 \\ g) \ (\(x, y). x \\ y)) (zip as vs))" using set_zip_rightD fsmult_closed FG_fddg_closed[of g "1::'f"] smult_sum_list_distrib[of "1 \\ g" "map (case_prod (\\)) (zip as vs)"] map_map[of "(\) (1 \\ g)" "case_prod (\\)" "zip as vs"] by fastforce moreover have "(\) (1 \\ g) \ (\(x, y). x \\ y) = (\(x,y). (1 \\ g) \ (x \\ y))" by auto ultimately have "g *\ as \\\ vs = sum_list (map (\(x,y). g *\ x \\ y) (zip as vs))" using Gmult_def by simp moreover from assms have "\(x,y) \ set (zip as vs). g *\ x \\ y = x \\ g *\ y" using set_zip_rightD fsmult_Gmult_comm by fastforce ultimately have "g *\ as \\\ vs = sum_list (map (\(x,y). x \\ y) (zip as (map (Gmult g) vs)))" using sum_list_prod_cong sum_list_prod_map2[of "\x y. x \\ y" as "Gmult g"] by force thus ?thesis using scalar_mult.lincomb_def[of fsmult] by simp qed lemma GSubspace_is_Subspace : "GSubspace U \ VectorSpace.Subspace fsmult V U" using GSubspace_is_FGModule FGModule.fVectorSpace VectorSpace.axioms Module.axioms by fast end (* context FGModule *) subsection \Homomorphisms of modules over a group ring\ subsubsection \Locales\ locale FGModuleHom = ActingGroup?: Group G + RModHom?: RModuleHom ActingGroup.group_ring smult V smult' T for G :: "'g::group_add set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" and smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and T :: "'v \ 'w" sublocale FGModuleHom < FGModule .. lemma (in FGModule) FGModuleHomI_fromaxioms : assumes "\v v'. v \ V \ v' \ V \ T (v + v') = T v + T v'" "supp T \ V" "\r m. r \ FG \ m \ V \ T (r \ m) = smult' r (T m)" shows "FGModuleHom G smult V smult' T" using assms by unfold_locales locale FGModuleEnd = FGModuleHom G smult V smult T for G :: "'g::group_add set" and FG :: "('f::field,'g) aezfun set" and smult :: "('f, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" and T :: "'v \ 'v" + assumes endomorph: "ImG \ V" locale FGModuleIso = FGModuleHom G smult V smult' T for G :: "'g::group_add set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixr "\" 70) and V :: "'v set" and smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and T :: "'v \ 'w" + fixes W :: "'w set" assumes bijective: "bij_betw T V W" abbreviation (in FGModule) isomorphic :: "(('f,'g) aezfun \ 'w::ab_group_add \ 'w) \ 'w set \ bool" where "isomorphic smult' W \ (\ T. FGModuleIso G smult V smult' T W)" subsubsection \Basic facts\ context FGModule begin lemma trivial_FGModuleHom : assumes "\r. r \ FG \ smult' r 0 = 0" shows "FGModuleHom G smult V smult' 0" proof (rule FGModuleHom.intro) from assms show "RModuleHom FG (\) V smult' 0" using trivial_RModuleHom by auto qed (unfold_locales) lemma FGModHom_idhom : "FGModuleHom G smult V smult (id\V)" proof (rule FGModuleHom.intro) show "RModuleHom FG smult V smult (id\V)" using RModHom_idhom by fast qed (unfold_locales) lemma VecHom_GMap_is_FGModuleHom : fixes smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) and Gmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) defines fsmult': "fsmult' \ aezfun_scalar_mult.fsmult smult'" and Gmult' : "Gmult' \ aezfun_scalar_mult.Gmult smult'" assumes hom : "VectorSpaceHom fsmult V fsmult' T" and Im_W : "FGModule G smult' W" "T ` V \ W" and G_map : "\g v. g \ G \ v \ V \ T (g *\ v) = g *\ (T v)" shows "FGModuleHom G smult V smult' T" proof show "\v v'. v \ V \ v' \ V \ T (v + v') = T v + T v'" using VectorSpaceHom.GroupHom[OF hom] GroupHom.hom by auto from hom show "supp T \ V" using VectorSpaceHom.supp by fast show "\r v. r \ FG \ v \ V \ T (r \ v) = r \ T v" proof- fix r v assume r: "r \ FG" and v: "v \ V" from r obtain fgs where fgs: "set (map snd fgs) \ G" "r = (\(f,g)\fgs. f \\ g)" using FG_el_decomp by fast from fgs v have "r \ v = (\(f,g)\fgs. f \\ g *\ v)" using sum_list_aezdeltafun_smult_distrib by simp moreover from v fgs(1) have "set (map (\(f,g). f \\ g *\ v) fgs) \ V" using Gmult_closed fsmult_closed by auto ultimately have "T (r \ v) = (\(f,g)\fgs. T (f \\ g *\ v))" using hom VectorSpaceHom.im_sum_list_prod by auto moreover from hom G_map fgs(1) v have "\(f,g) \ set fgs. T (f \\ g *\ v) = f \\ g *\ T v" using Gmult_closed VectorSpaceHom.f_map[of fsmult V fsmult' T] by auto ultimately have "T (r \ v) = (\(f,g)\fgs. f \\ g *\ T v)" using sum_list_prod_cong by simp with v fgs fsmult' Gmult' Im_W(2) show "T (r \ v) = r \ (T v)" using FGModule.sum_list_aezdeltafun_smult_distrib[OF Im_W(1)] by auto qed qed lemma VecHom_GMap_on_fbasis_is_FGModuleHom : fixes smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) and Gmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) and flincomb' :: "'f list \ 'w list \ 'w" (infixr "\\\" 70) defines fsmult' : "fsmult' \ aezfun_scalar_mult.fsmult smult'" and Gmult' : "Gmult' \ aezfun_scalar_mult.Gmult smult'" and flincomb' : "flincomb' \ aezfun_scalar_mult.flincomb smult'" assumes fbasis : "fbasis_for V vs" and hom : "VectorSpaceHom fsmult V fsmult' T" and Im_W : "FGModule G smult' W" "T ` V \ W" and G_map : "\g v. g \ G \ v \ set vs \ T (g *\ v) = g *\ (T v)" shows "FGModuleHom G smult V smult' T" proof (rule VecHom_GMap_is_FGModuleHom) from fsmult' hom show "VectorSpaceHom (\\) V (aezfun_scalar_mult.fsmult (\)) T" by fast next fix g v assume g: "g \ G" and v: "v \ V" from v fbasis obtain cs where cs: "v = cs \\\ vs" using VectorSpace.in_Span_obtain_same_length_coeffs[OF fVectorSpace] by fast with g(1) fbasis fsmult' flincomb' have "T (g *\ v) = cs \\\ (map (T \ (Gmult g)) vs)" using Gmult_flincomb_comm map_Gmult_closed VectorSpaceHom.distrib_lincomb[OF hom] by auto moreover have "T \ (Gmult g) = (\v. T (g *\ v))" by auto ultimately have "T (g *\ v) = cs \\\ (map (\v. g *\ (T v)) vs)" using fbasis g(1) G_map map_cong[of vs vs "\v. T (g *\ v)"] by simp moreover have "(\v. g *\ (T v)) = (Gmult' g) \ T" by auto ultimately have "T (g *\ v) = g *\ cs \\\ (map T vs)" using g(1) fbasis Im_W(2) Gmult' flincomb' FGModule.Gmult_flincomb_comm[OF Im_W(1), of g "map T vs"] by fastforce thus "T (g *\ v) = aezfun_scalar_mult.Gmult (\) g (T v)" using fbasis fsmult' Gmult' flincomb' cs VectorSpaceHom.distrib_lincomb[OF hom] by auto qed (rule Im_W(1), rule Im_W(2)) end (* context FGModule *) context FGModuleHom begin abbreviation fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) where "fsmult' \ aezfun_scalar_mult.fsmult smult'" abbreviation Gmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) where "Gmult' \ aezfun_scalar_mult.Gmult smult'" lemmas supp = supp lemmas additive = additive lemmas FG_map = R_map lemmas FG_fdd0_closed = FG_fdd0_closed lemmas fsmult_smult_domain_comm = fsmult_smult_comm lemmas GSubspace_Ker = RSubmodule_Ker lemmas Ker_Im_iff = Ker_Im_iff lemmas Ker0_imp_inj_on = Ker0_imp_inj_on lemmas eq_im_imp_diff_in_Ker = eq_im_imp_diff_in_Ker lemmas im_submodule = im_submodule lemmas fsmultD' = aezfun_scalar_mult.fsmultD[of smult'] lemmas GmultD' = aezfun_scalar_mult.GmultD[of smult'] lemma f_map : "v \ V \ T (a \\ v) = a \\ T v" using fsmultD ActingGroup.RG_aezdelta0fun_closed[of a] FG_map fsmultD' by simp lemma G_map : "g \ G \ v \ V \ T (g *\ v) = g *\ T v" using GmultD ActingGroup.RG_aezdeltafun_closed[of g 1] FG_map GmultD' by simp lemma VectorSpaceHom : "VectorSpaceHom fsmult V fsmult' T" by ( rule VectorSpace.VectorSpaceHomI, rule fVectorSpace, unfold_locales, rule f_map ) lemmas distrib_flincomb = VectorSpaceHom.distrib_lincomb[OF VectorSpaceHom] lemma FGModule_Im : "FGModule G smult' ImG" by (rule FGModule.intro, rule GroupG, rule RModule_Im, unfold_locales) lemma FGModHom_composite_left : assumes "FGModuleHom G smult' W smult'' S" "T ` V \ W" shows "FGModuleHom G smult V smult'' (S \ T)" proof (rule FGModuleHom.intro) from assms(2) show "RModuleHom FG smult V smult'' (S \ T)" using FGModuleHom.axioms(2)[OF assms(1)] RModHom_composite_left[of W] by fast qed (rule GroupG, unfold_locales) lemma restriction_to_subgroup_is_hom : fixes H :: "'g set" assumes subgrp: "Group.Subgroup G H" shows "FGModuleHom H smult V smult' T" proof (rule FGModule.FGModuleHomI_fromaxioms) have "FGModule G smult V" .. with assms show "FGModule H (\) V" using FGModule.restriction_to_subgroup_is_module by fast from supp show "supp T \ V" by fast from assms show "\r m. \ r \ (Group.group_ring H); m \ V \ \ T (r \ m) = r \ T m" using FG_map ActingGroup.Subgroup_imp_Subring by fast qed (rule hom) lemma FGModuleHom_restrict0_GSubspace : assumes "GSubspace U" shows "FGModuleHom G smult U smult' (T \ U)" proof (rule FGModuleHom.intro) from assms show "RModuleHom FG (\) U (\) (T \ U)" using RModuleHom_restrict0_submodule by fast qed (unfold_locales) lemma FGModuleHom_fscalar_mul : "FGModuleHom G smult V smult' (\v. a \\ T v)" proof have vsphom: "VectorSpaceHom fsmult V fsmult' (\v. a \\ T v)" using VectorSpaceHom.VectorSpaceHom_scalar_mul[OF VectorSpaceHom] by fast thus "\v v'. v \ V \ v' \ V \ a \\ T (v + v') = a \\ T v + a \\ T v'" using VectorSpaceHom.additive[of fsmult V] by auto from vsphom show "supp (\v. a \\ T v) \ V" using VectorSpaceHom.supp by fast next fix r v assume rv: "r \ FG" "v \ V" thus "a \\ T (r \ v) = r \ a \\ T v" using FG_map FGModule.fsmult_smult_comm[OF FGModule_Im] by auto qed end (* context FGModuleHom *) lemma GSubspace_eigenspace : fixes e :: "'f::field" and E :: "'v::ab_group_add set" and smult :: "('f::field, 'g::group_add) aezfun \ 'v \ 'v" (infixr "\" 70) assumes FGModHom: "FGModuleHom G smult V smult T" defines E : "E \ {v \ V. T v = aezfun_scalar_mult.fsmult smult e v}" shows "FGModule.GSubspace G smult V E" proof- have "FGModule.GSubspace G smult V {v \ V. T v = (e \\ 0) \ v}" using FGModuleHom.axioms(2)[OF FGModHom] proof (rule RSubmodule_eigenspace) show "e \\ 0 \ FGModule.FG G" using FGModuleHom.FG_fdd0_closed[OF FGModHom] by fast show "\s v. s \ FGModule.FG G \ v \ V \ s \ (e \\ 0) \ v = (e \\ 0) \ s \ v" using FGModuleHom.fsmult_smult_domain_comm[OF FGModHom] aezfun_scalar_mult.fsmultD[of smult] by simp qed with E show ?thesis using aezfun_scalar_mult.fsmultD[of smult] by simp qed subsubsection \Basic facts about endomorphisms\ lemma RModuleEnd_over_group_ring_is_FGModuleEnd : fixes G :: "'g::group_add set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" assumes G : "Group G" and endo: "RModuleEnd (Group.group_ring G) smult V T" shows "FGModuleEnd G smult V T" proof (rule FGModuleEnd.intro, rule FGModuleHom.intro, rule G) from endo show "RModuleHom (Group.group_ring G) smult V smult T" using RModuleEnd.axioms(1) by fast from endo show "FGModuleEnd_axioms V T" using RModuleEnd.endomorph by unfold_locales qed lemma (in FGModule) VecEnd_GMap_is_FGModuleEnd : assumes endo : "VectorSpaceEnd fsmult V T" and G_map: "\g v. g \ G \ v \ V \ T (g *\ v) = g *\ (T v)" shows "FGModuleEnd G smult V T" proof (rule FGModuleEnd.intro, rule VecHom_GMap_is_FGModuleHom) from endo show "VectorSpaceHom (\\) V (\\) T" using VectorSpaceEnd.axioms(1) by fast from endo show "T ` V \ V" using VectorSpaceEnd.endomorph by fast from endo show "FGModuleEnd_axioms V T" using VectorSpaceEnd.endomorph by unfold_locales qed (unfold_locales, rule G_map) lemma (in FGModule) GEnd_inner_dirsum_el_decomp_nth : "\ \U\set Us. GSubspace U; add_independentS Us; n < length Us \ \ FGModuleEnd G smult (\U\Us. U) (\Us\n)" using GroupG RModuleEnd_inner_dirsum_el_decomp_nth RModuleEnd_over_group_ring_is_FGModuleEnd by fast context FGModuleEnd begin lemma RModuleEnd : "RModuleEnd ActingGroup.group_ring smult V T" using endomorph by unfold_locales lemma VectorSpaceEnd : "VectorSpaceEnd fsmult V T" by ( rule VectorSpaceEnd.intro, rule VectorSpaceHom, unfold_locales, rule endomorph ) lemmas proj_decomp = RModuleEnd.proj_decomp[OF RModuleEnd] lemmas GSubspace_Ker = GSubspace_Ker lemmas FGModuleHom_restrict0_GSubspace = FGModuleHom_restrict0_GSubspace end (* context FGModuleEnd *) subsubsection \Basic facts about isomorphisms\ context FGModuleIso begin lemmas VectorSpaceHom = VectorSpaceHom abbreviation "invT \ (the_inv_into V T) \ W" lemma RModuleIso : "RModuleIso FG smult V smult' T W" proof (rule RModuleIso.intro) show "RModuleHom FG (\) V (\) T" using FGModuleIso_axioms FGModuleIso.axioms(1) FGModuleHom.axioms(2) by fast qed (unfold_locales, rule bijective) lemmas ImG = RModuleIso.ImG[OF RModuleIso] lemma FGModuleIso_restrict0_GSubspace : assumes "GSubspace U" shows "FGModuleIso G smult U smult' (T \ U) (T ` U)" proof (rule FGModuleIso.intro) from assms show "FGModuleHom G (\) U (\) (T \ U)" using FGModuleHom_restrict0_GSubspace by fast show "FGModuleIso_axioms U (T \ U) (T ` U)" proof from assms bijective have "bij_betw T U (T ` U)" using subset_inj_on unfolding bij_betw_def by auto thus "bij_betw (T \ U) U (T ` U)" unfolding bij_betw_def inj_on_def by auto qed qed lemma inv : "FGModuleIso G smult' W smult invT V" proof (rule FGModuleIso.intro, rule FGModuleHom.intro) show "RModuleHom FG (\) W (\) invT" using RModuleIso.inv[OF RModuleIso] RModuleIso.axioms(1) by fast show "FGModuleIso_axioms W invT V" using RModuleIso.inv[OF RModuleIso] RModuleIso.bijective by unfold_locales qed (unfold_locales) lemma FGModIso_composite_left : assumes "FGModuleIso G smult' W smult'' S X" shows "FGModuleIso G smult V smult'' (S \ T) X" proof (rule FGModuleIso.intro) from assms show "FGModuleHom G (\) V smult'' (S \ T)" using FGModuleIso.axioms(1) ImG FGModHom_composite_left by fast show "FGModuleIso_axioms V (S \ T) X" using bijective FGModuleIso.bijective[OF assms] bij_betw_trans by unfold_locales qed lemma isomorphic_sym : "FGModule.isomorphic G smult' W smult V" using inv by fast lemma isomorphic_trans : "FGModule.isomorphic G smult' W smult'' X \ FGModule.isomorphic G smult V smult'' X" using FGModIso_composite_left by fast lemma isomorphic_to_zero_left : "V = 0 \ W = 0" using bijective bij_betw_imp_surj_on im_zero by fastforce lemma isomorphic_to_zero_right : "W = 0 \ V = 0" using isomorphic_sym FGModuleIso.isomorphic_to_zero_left by fast lemma isomorphic_to_irr_right' : assumes "\U. FGModule.GSubspace G smult' W U \ U = 0 \ U = W" shows "\U. GSubspace U \ U = 0 \ U = V" proof- fix U assume U: "GSubspace U" have "U \ V \ U = 0" proof- assume UV: "U \ V" from U bijective have "T ` U = T ` V \ U = V" using bij_betw_imp_inj_on[of T V W] inj_onD[of T V] by fast with UV bijective have "T ` U \ W" using bij_betw_imp_surj_on by fast moreover from U have "FGModule.GSubspace G smult' W (T ` U)" using ImG im_submodule by fast ultimately show "U = 0" using assms U FGModuleIso_restrict0_GSubspace FGModuleIso.isomorphic_to_zero_right by fast qed thus "U = 0 \ U = V" by fast qed end (* context FGModuleIso *) context FGModule begin lemma isomorphic_sym : "isomorphic smult' W \ FGModule.isomorphic G smult' W smult V" using FGModuleIso.inv by fast lemma isomorphic_trans : "isomorphic smult' W \ FGModule.isomorphic G smult' W smult'' X \ isomorphic smult'' X" using FGModuleIso.FGModIso_composite_left by fast lemma isomorphic_to_zero_left : "V = 0 \ isomorphic smult' W \ W = 0" using FGModuleIso.isomorphic_to_zero_left by fast lemma isomorphic_to_zero_right : "isomorphic smult' 0 \ V = 0" using FGModuleIso.isomorphic_to_zero_right by fast lemma FGModIso_idhom : "FGModuleIso G smult V smult (id\V) V" using FGModHom_idhom proof (rule FGModuleIso.intro) show "FGModuleIso_axioms V (id\V) V" using bij_betw_id bij_betw_restrict0 by unfold_locales fast qed lemma isomorphic_refl : "isomorphic smult V" using FGModIso_idhom by fast end (* context FGModule *) subsubsection \Hom-sets\ definition FGModuleHomSet :: "'g::group_add set \ (('f::field,'g) aezfun \ 'v::ab_group_add \ 'v) \ 'v set \ (('f,'g) aezfun \ 'w::ab_group_add \ 'w) \ 'w set \ ('v \ 'w) set" where "FGModuleHomSet G fgsmult V fgsmult' W \ {T. FGModuleHom G fgsmult V fgsmult' T} \ {T. T ` V \ W}" lemma FGModuleHomSetI : "FGModuleHom G fgsmult V fgsmult' T \ T ` V \ W \ T \ FGModuleHomSet G fgsmult V fgsmult' W" unfolding FGModuleHomSet_def by fast lemma FGModuleHomSetD_FGModuleHom : "T \ FGModuleHomSet G fgsmult V fgsmult' W \ FGModuleHom G fgsmult V fgsmult' T" unfolding FGModuleHomSet_def by fast lemma FGModuleHomSetD_Im : "T \ FGModuleHomSet G fgsmult V fgsmult' W \ T ` V \ W" unfolding FGModuleHomSet_def by fast context FGModule begin lemma FGModuleHomSet_is_Gmaps_in_VectorSpaceHomSet : fixes smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) and Gmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) defines fsmult' : "fsmult' \ aezfun_scalar_mult.fsmult smult'" and Gmult' : "Gmult' \ aezfun_scalar_mult.Gmult smult'" assumes FGModW : "FGModule G smult' W" shows "FGModuleHomSet G smult V smult' W = (VectorSpaceHomSet fsmult V fsmult' W) \ {T. \g\G. \v\V. T (g *\ v) = g *\ (T v)}" proof from fsmult' Gmult' show "FGModuleHomSet G smult V smult' W \ (VectorSpaceHomSet fsmult V fsmult' W) \ {T. \g\G. \v\V. T (g *\ v) = g *\ T v}" using FGModuleHomSetD_FGModuleHom[of _ G smult V smult'] FGModuleHom.VectorSpaceHom[of G smult V smult'] FGModuleHomSetD_Im[of _ G smult V smult'] VectorSpaceHomSetI[of fsmult V fsmult'] FGModuleHom.G_map[of G smult V smult'] by auto show "FGModuleHomSet G smult V smult' W \ (VectorSpaceHomSet fsmult V fsmult' W) \ {T. \g\G. \v\V. T (g *\ v) = g *\ T v}" proof fix T assume T: "T \ (VectorSpaceHomSet fsmult V fsmult' W) \ {T. \g\G. \v\V. T (g *\ v) = g *\ T v}" show "T \ FGModuleHomSet G smult V smult' W" proof (rule FGModuleHomSetI, rule VecHom_GMap_is_FGModuleHom) from T fsmult' show "VectorSpaceHom (\\) V (aezfun_scalar_mult.fsmult smult') T" using VectorSpaceHomSetD_VectorSpaceHom by fast from T show "T ` V \ W" using VectorSpaceHomSetD_Im by fast from T Gmult' show "\g v. g \ G \ v \ V \ T (g *\ v) = aezfun_scalar_mult.Gmult (\) g (T v)" by fast from T show "T ` V \ W" using VectorSpaceHomSetD_Im by fast qed (rule FGModW) qed qed lemma Group_FGModuleHomSet : fixes smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) and Gmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) defines fsmult' : "fsmult' \ aezfun_scalar_mult.fsmult smult'" and Gmult' : "Gmult' \ aezfun_scalar_mult.Gmult smult'" assumes FGModW : "FGModule G smult' W" shows "Group (FGModuleHomSet G smult V smult' W)" proof from FGModW show "FGModuleHomSet G (\) V smult' W \ {}" using FGModule.smult_zero trivial_FGModuleHom[of smult'] FGModule.zero_closed FGModuleHomSetI by fastforce next fix S T assume S: "S \ FGModuleHomSet G (\) V smult' W" and T: "T \ FGModuleHomSet G (\) V smult' W" with assms have ST: "S \ (VectorSpaceHomSet fsmult V fsmult' W) \ {T. \g\G. \v\V. T (g *\ v) = g *\ T v}" "T \ (VectorSpaceHomSet fsmult V fsmult' W) \ {T. \g\G. \v\V. T (g *\ v) = g *\ T v}" using FGModuleHomSet_is_Gmaps_in_VectorSpaceHomSet by auto with fsmult' have "S - T \ VectorSpaceHomSet fsmult V fsmult' W" using FGModule.fVectorSpace[OF FGModW] VectorSpace.Group_VectorSpaceHomSet[OF fVectorSpace] Group.diff_closed by fast moreover have "\g v. g\G \ v\V \ (S-T) (g *\ v) = g *\ ((S-T) v)" proof- fix g v assume "g \ G" "v \ V" moreover with ST have "S v \ W" "T v \ W" "- T v \ W" using VectorSpaceHomSetD_Im[of S fsmult V fsmult'] VectorSpaceHomSetD_Im[of T fsmult V fsmult'] FGModule.neg_closed[OF FGModW] by auto ultimately show "(S-T) (g *\ v) = g *\ ((S-T) v)" using ST Gmult' FGModule.neg_Gmult[OF FGModW] FGModule.Gmult_distrib_left[OF FGModW, of g "S v" "- T v"] by auto qed ultimately show "S - T \ FGModuleHomSet G (\) V smult' W" using fsmult' Gmult' FGModuleHomSet_is_Gmaps_in_VectorSpaceHomSet[OF FGModW] by fast qed lemma Subspace_FGModuleHomSet : fixes smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) and Gmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) and hom_fsmult :: "'f \ ('v \ 'w) \ ('v \ 'w)" (infixr "\\\" 70) defines fsmult' : "fsmult' \ aezfun_scalar_mult.fsmult smult'" and Gmult' : "Gmult' \ aezfun_scalar_mult.Gmult smult'" defines hom_fsmult : "hom_fsmult \ \a T v. a \\ T v" assumes FGModW : "FGModule G smult' W" shows "VectorSpace.Subspace hom_fsmult (VectorSpaceHomSet fsmult V fsmult' W) (FGModuleHomSet G smult V smult' W)" proof (rule VectorSpace.SubspaceI) from hom_fsmult fsmult' show "VectorSpace (\\\) (VectorSpaceHomSet (\\) V (\\) W)" using FGModule.fVectorSpace[OF FGModW] VectorSpace.VectorSpace_VectorSpaceHomSet[OF fVectorSpace] by fast from fsmult' Gmult' FGModW show "Group (FGModuleHomSet G (\) V (\) W) \ FGModuleHomSet G (\) V (\) W \ VectorSpaceHomSet (\\) V (\\) W" using Group_FGModuleHomSet FGModuleHomSet_is_Gmaps_in_VectorSpaceHomSet by fast next fix a T assume T: "T \ FGModuleHomSet G (\) V (\) W" from hom_fsmult fsmult' have "FGModuleHom G smult V smult' (a \\\ T)" using FGModuleHomSetD_FGModuleHom[OF T] FGModuleHomSetD_Im[OF T] FGModuleHom.FGModuleHom_fscalar_mul by simp moreover from hom_fsmult fsmult' have "(a \\\ T) ` V \ W" using FGModuleHomSetD_Im[OF T] FGModule.fsmult_closed[OF FGModW] by auto ultimately show "a \\\ T \ FGModuleHomSet G (\) V (\) W" using FGModuleHomSetI by fastforce qed lemma VectorSpace_FGModuleHomSet : fixes smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) and hom_fsmult :: "'f \ ('v \ 'w) \ ('v \ 'w)" (infixr "\\\" 70) defines "fsmult' \ aezfun_scalar_mult.fsmult smult'" defines "hom_fsmult \ \a T v. a \\ T v" assumes "FGModule G smult' W" shows "VectorSpace hom_fsmult (FGModuleHomSet G smult V smult' W)" using assms Subspace_FGModuleHomSet Module.intro VectorSpace.intro by fast end (* context FGModule *) subsection \Induced modules\ subsubsection \Additive function spaces\ definition addfunset :: "'a::monoid_add set \ 'm::monoid_add set \ ('a \ 'm) set" where "addfunset A M \ {f. supp f \ A \ range f \ M \ (\x\A. \y\A. f (x+y) = f x + f y) }" lemma addfunsetI : "\ supp f \ A; range f \ M; \x\A. \y\A. f (x+y) = f x + f y \ \ f \ addfunset A M" unfolding addfunset_def by fast lemma addfunsetD_supp : "f \ addfunset A M \ supp f \ A" unfolding addfunset_def by fast lemma addfunsetD_range : "f \ addfunset A M \ range f \ M" unfolding addfunset_def by fast lemma addfunsetD_range' : "f \ addfunset A M \ f x \ M" using addfunsetD_range by fast lemma addfunsetD_add : "\ f \ addfunset A M; x\A; y\A \ \ f (x+y) = f x + f y" unfolding addfunset_def by fast lemma addfunset0 : "addfunset A (0::'m::monoid_add set) = 0" proof show "addfunset A 0 \ 0" using addfunsetD_range' by fastforce have "(0::'a\'m) \ addfunset A 0" using supp_zerofun_subset_any by (rule addfunsetI) auto thus "addfunset A (0::'m::monoid_add set) \ 0" by simp qed lemma Group_addfunset : fixes M::"'g::ab_group_add set" assumes "Group M" shows "Group (addfunset R M)" proof from assms show "addfunset R M \ {}" using addfunsetI[of 0 R M] supp_zerofun_subset_any Group.zero_closed by fastforce next fix g h assume gh: "g \ addfunset R M" "h \ addfunset R M" show "g - h \ addfunset R M" proof (rule addfunsetI) from gh show "supp (g - h) \ R" using addfunsetD_supp supp_diff_subset_union_supp by fast from gh show "range (g - h) \ M" using addfunsetD_range Group.diff_closed [OF assms] by (simp add: addfunsetD_range' image_subsetI) show "\x\R. \y\R. (g - h) (x + y) = (g - h) x + (g - h) y" using addfunsetD_add[OF gh(1)] addfunsetD_add[OF gh(2)] by simp qed qed subsubsection \Spaces of functions which transform under scalar multiplication by almost-everywhere-zero functions\ context aezfun_scalar_mult begin definition smultfunset :: "'g set \ ('r,'g) aezfun set \ (('r,'g) aezfun \ 'v) set" where "smultfunset G FH \ {f. (\a::'r. \g\G. \x\FH. f ( a \\ g * x ) = (a \\ g) \ (f x))}" lemma smultfunsetD : "\f \ smultfunset G FH; g\G; x\FH \ \ f ( a \\ g * x ) = (a \\ g) \ (f x)" unfolding smultfunset_def by fast lemma smultfunsetI : "\a::'r. \g\G. \x\FH. f ( a \\ g * x ) = (a \\ g) \ (f x) \ f \ smultfunset G FH" unfolding smultfunset_def by fast end (* context aezfun_scalar_mult *) subsubsection \General induced spaces of functions on a group ring\ context aezfun_scalar_mult begin definition indspace :: "'g set \ ('r,'g) aezfun set \ 'v set \ (('r,'g) aezfun \ 'v) set" where "indspace G FH V = addfunset FH V \ smultfunset G FH" lemma indspaceD : "f \ indspace G FH V \ f \ addfunset FH V \ smultfunset G FH" using indspace_def by fast lemma indspaceD_supp : "f \ indspace G FH V \ supp f \ FH" using indspace_def addfunsetD_supp by fast lemma indspaceD_supp' : "f \ indspace G FH V \ x \ FH \ f x = 0" using indspaceD_supp suppI_contra by fast lemma indspaceD_range : "f \ indspace G FH V \ range f \ V" using indspace_def addfunsetD_range by fast lemma indspaceD_range': "f \ indspace G FH V \ f x \ V" using indspaceD_range by fast lemma indspaceD_add : "\ f \ indspace G FH V; x\FH; y\FH \ \ f (x+y) = f x + f y" using indspace_def addfunsetD_add by auto lemma indspaceD_transform : "\f \ indspace G FH V; g\G; x\FH \ \ f ( a \\ g * x ) = (a \\ g) \ (f x)" using indspace_def smultfunsetD by auto lemma indspaceI : "f \ addfunset FH V \ f \ smultfunset G FH \ f \ indspace G FH V" using indspace_def by fast lemma indspaceI' : "\ supp f \ FH; range f \ V; \x\FH. \y\FH. f (x+y) = f x + f y; \a::'r. \g\G. \x\FH. f ( a \\ g * x ) = (a \\ g) \ (f x) \ \ f \ indspace G FH V" using smultfunsetI addfunsetI[of f] indspaceI by simp lemma mono_indspace : "mono (indspace G FH)" proof (rule monoI) fix U V :: "'v set" assume U_V: "U \ V" show "indspace G FH U \ indspace G FH V" proof fix f assume f: "f \ indspace G FH U" show "f \ indspace G FH V" using indspaceD_supp[OF f] proof (rule indspaceI') from f U_V show "range f \ V" using indspaceD_range[of f G FH] by auto from f show "\x\FH. \y\FH. f (x+y) = f x + f y" using indspaceD_add by auto from f show "\a::'r. \g\G. \x\FH. f ( a \\ g * x ) = (a \\ g) \ (f x)" using indspaceD_transform by auto qed qed qed end (* context aezfun_scalar_mult *) context FGModule begin lemma zero_transforms : "0 \ smultfunset G FH" using smultfunsetI FG_fddg_closed smult_zero by simp lemma indspace0 : "indspace G FH 0 = 0" using zero_transforms addfunset0 indspace_def by auto lemma Group_indspace : assumes "Ring1 FH" shows "Group (indspace G FH V)" proof from zero_closed have "0 \ V" by simp with mono_indspace [of G FH] have "indspace G FH 0 \ indspace G FH V" by (auto dest!: monoD [of _ 0 V]) then show "indspace G FH V \ {}" using indspace0 [of FH] by auto next fix f1 f2 assume ff: "f1 \ indspace G FH V" "f2 \ indspace G FH V" hence "f1 - f2 \ addfunset FH V" using assms indspaceD indspaceD Group Group_addfunset Group.diff_closed by fast moreover from ff have "f1 - f2 \ smultfunset G FH" using indspaceD_transform FG_fddg_closed indspaceD_range' smult_distrib_left_diff smultfunsetI by simp ultimately show "f1 - f2 \ indspace G FH V" using indspaceI by fast qed end (* context FGModule *) subsubsection \The right regular action\ context Ring1 begin definition rightreg_scalar_mult :: "'r::ring_1 \ ('r \ 'm::ab_group_add) \ ('r \ 'm)" (infixr "\" 70) where "rightreg_scalar_mult r f = (\x. if x \ R then f (x*r) else 0)" lemma rightreg_scalar_multD1 : "x \ R \ (r \ f) x = f (x*r)" unfolding rightreg_scalar_mult_def by simp lemma rightreg_scalar_multD2 : "x \ R \ (r \ f) x = 0" unfolding rightreg_scalar_mult_def by simp lemma rrsmult_supp : "supp (r \ f) \ R" using rightreg_scalar_multD2 suppD_contra by force lemma rrsmult_range : "range (r \ f) \ {0} \ range f" proof (rule image_subsetI) fix x show "(r \ f) x \ {0} \ range f" using rightreg_scalar_multD1[of x r f] image_eqI rightreg_scalar_multD2[of x r f] by (cases "x \ R") auto qed lemma rrsmult_distrib_left : "r \ (f + g) = r \ f + r \ g" proof fix x show "(r \ (f + g)) x = (r \ f + r \ g) x" unfolding rightreg_scalar_mult_def by (cases "x \ R") auto qed lemma rrsmult_distrib_right : assumes "\x y. x \ R \ y \ R \ f (x+y) = f x + f y" "r \ R" "s \ R" shows "(r + s) \ f = r \ f + s \ f" proof fix x show "((r + s) \ f) x = (r \ f + s \ f) x" using assms mult_closed unfolding rightreg_scalar_mult_def by (cases "x \ R") (auto simp add: distrib_left) qed lemma RModule_addfunset : fixes M::"'g::ab_group_add set" assumes "Group M" shows "RModule R rightreg_scalar_mult (addfunset R M)" proof (rule RModuleI) from assms show "Group (addfunset R M)" using Group_addfunset by fast show "RModule_axioms R (\) (addfunset R M)" proof fix r f assume r: "r \ R" and f: "f \ addfunset R M" show "r \ f \ addfunset R M" proof (rule addfunsetI) show "supp (r \ f) \ R" using rightreg_scalar_multD2 suppD_contra by force show "range (r \ f) \ M" using addfunsetD_range[OF f] Group.zero_closed[OF assms] unfolding rightreg_scalar_mult_def by fastforce from r show "\x\R. \y\R. (r \ f) (x + y) = (r \ f) x + (r \ f) y" using mult_closed add_closed addfunsetD_add[OF f] unfolding rightreg_scalar_mult_def by (simp add: distrib_right) qed next show "\r f g. r \ (f + g) = r \ f + r \ g" using rrsmult_distrib_left by fast next fix r s f assume "r \ R" "s \ R" "f \ addfunset R M" thus "(r + s) \ f = r \ f + s \ f" using addfunsetD_add[of f] rrsmult_distrib_right[of f] by simp next fix r s f assume r: "r \ R" and s: "s \ R" and f: "f \ addfunset R M" show "r \ s \ f = (r * s) \ f" proof fix x from r show "(r \ s \ f) x = ((r * s) \ f) x" using mult_closed unfolding rightreg_scalar_mult_def by (cases "x \ R") (auto simp add: mult.assoc) qed next fix f assume f: "f \ addfunset R M" show "1 \ f = f" proof fix x show "(1 \ f) x = f x" unfolding rightreg_scalar_mult_def using addfunsetD_supp[OF f] suppI_contra[of x f] contra_subsetD[of "supp f"] by (cases "x \ R") auto qed qed qed (unfold_locales) end (* context Ring1 *) subsubsection \Locale and basic facts\ text \ In the following locale, @{term G} is a subgroup of @{term H}, @{term V} is a module over the group ring for @{term G}, and the induced space @{term indV} will be shown to be a module over the group ring for @{term H} under the right regular scalar multiplication @{term rrsmult}. \ locale InducedFHModule = Supgroup?: Group H + BaseFGMod? : FGModule G smult V + induced_smult?: aezfun_scalar_mult rrsmult for H :: "'g::group_add set" and G :: "'g set" and FG :: "('f::field, 'g) aezfun set" and smult :: "('f, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixl "\" 70) and V :: "'v set" and rrsmult :: "('f,'g) aezfun \ (('f,'g) aezfun \ 'v) \ (('f,'g) aezfun \ 'v)" (infixl "\" 70) + fixes FH :: "('f, 'g) aezfun set" and indV :: "(('f, 'g) aezfun \ 'v) set" defines FH : "FH \ Supgroup.group_ring" and indV : "indV \ BaseFGMod.indspace G FH V" assumes rrsmult : "rrsmult = Ring1.rightreg_scalar_mult FH" and Subgroup: "Supgroup.Subgroup G" (* G is a subgroup of H *) begin abbreviation indfsmult :: "'f \ (('f, 'g) aezfun \ 'v) \ (('f, 'g) aezfun \ 'v)" (infixl "\\" 70) where "indfsmult \ induced_smult.fsmult" abbreviation indflincomb :: "'f list \ (('f, 'g) aezfun \ 'v) list \ (('f, 'g) aezfun \ 'v)" (infixl "\\\" 70) where "indflincomb \ induced_smult.flincomb" abbreviation Hmult :: "'g \ (('f, 'g) aezfun \ 'v) \ (('f, 'g) aezfun \ 'v)" (infixl "*\" 70) where "Hmult \ induced_smult.Gmult" lemma Ring1_FH : "Ring1 FH" using FH Supgroup.Ring1_RG by fast lemma FG_subring_FH : "Ring1.Subring1 FH BaseFGMod.FG" using FH Supgroup.Subgroup_imp_Subring[OF Subgroup] by fast lemma rrsmultD1 : "x \ FH \ (r \ f) x = f (x*r)" using Ring1.rightreg_scalar_multD1[OF Ring1_FH] rrsmult by simp lemma rrsmultD2 : "x \ FH \ (r \ f) x = 0" using Ring1.rightreg_scalar_multD2[OF Ring1_FH] rrsmult by fast lemma rrsmult_supp : "supp (r \ f) \ FH" using Ring1.rrsmult_supp[OF Ring1_FH] rrsmult by auto lemma rrsmult_range : "range (r \ f) \ {0} \ range f" using Ring1.rrsmult_range[OF Ring1_FH] rrsmult by auto lemma FHModule_addfunset : "FGModule H rrsmult (addfunset FH V)" proof (rule FGModule.intro) from FH rrsmult show "RModule Supgroup.group_ring (\) (addfunset FH V)" using Group Supgroup.Ring1_RG Ring1.RModule_addfunset by fast qed (unfold_locales) lemma FHSubmodule_indspace : "FGModule.FGSubmodule H rrsmult (addfunset FH V) indV" proof (rule FGModule.FGSubmoduleI[of H], rule FHModule_addfunset, rule conjI) from Ring1_FH indV show "Group indV" using Group_indspace by fast from indV show "indV \ addfunset FH V" using BaseFGMod.indspaceD by fast next fix r f assume rf: "r \ (Supgroup.group_ring :: ('f,'g) aezfun set)" "f \ indV" from rf(2) indV have rf2': "f \ BaseFGMod.indspace G FH V" by fast show "r \ f \ indV" unfolding indV proof (rule BaseFGMod.indspaceI', rule rrsmult_supp) show "range (r \ f) \ V" using rrsmult_range BaseFGMod.indspaceD_range[OF rf2'] zero_closed by force from FH rf(1) rf2' show "\x\FH. \y\FH. (r \ f) (x + y) = (r \ f) x + (r \ f) y" using Ring1.add_closed[OF Ring1_FH] rrsmultD1[of _ r f] Ring1.mult_closed[OF Ring1_FH] BaseFGMod.indspaceD_add by (simp add: distrib_right) { fix a g x assume gx: "g\G" "x\FH" with FH have "a \\ g * x \ FH" using FG_fddg_closed FG_subring_FH Ring1.mult_closed[OF Ring1_FH] by fast with FH rf(1) gx(2) have "(r \ f) (a \\ g * x) = a \\ g \ ((r \ f) x)" using rrsmultD1[of _ r f] Ring1.mult_closed[OF Ring1_FH] BaseFGMod.indspaceD_transform[OF rf2' gx(1)] by (simp add: mult.assoc) } thus "\a. \g\G. \x\FH. (r \ f) (a \\ g * x) = a \\ g \ (r \ f) x" by fast qed qed lemma FHModule_indspace : "FGModule H rrsmult indV" proof (rule FGModule.intro) show "RModule Supgroup.group_ring (\) indV" using FHSubmodule_indspace by fast qed (unfold_locales) lemmas fVectorSpace_indspace = FGModule.fVectorSpace[OF FHModule_indspace] lemmas restriction_is_FGModule = FGModule.restriction_to_subgroup_is_module[OF FHModule_indspace] definition induced_vector :: "'v \ (('f, 'g) aezfun \ 'v)" where "induced_vector v \ (if v \ V then (\y. if y \ FH then (FG_proj y) \ v else 0) else 0)" lemma induced_vector_apply1 : "v \ V \ x \ FH \ induced_vector v x = (FG_proj x) \ v" using induced_vector_def by simp lemma induced_vector_apply2 : "v \ V \ x \ FH \ induced_vector v x = 0" using induced_vector_def by simp lemma induced_vector_indV : assumes v: "v \ V" shows "induced_vector v \ indV" unfolding indV proof (rule BaseFGMod.indspaceI') from assms show "supp (induced_vector v) \ FH" using induced_vector_def supp_restrict0[of FH "\y. (FG_proj y) \ v"] by simp show "range (induced_vector v) \ V" proof (rule image_subsetI) fix y from v show "(induced_vector v) y \ V" using induced_vector_def zero_closed aezfun_setspan_proj_in_setspan[of G y] smult_closed ActingGroup.group_ringD by auto qed { fix x y assume xy: "x \ FH" "y \ FH" with v have "(induced_vector v) (x + y) = (induced_vector v) x + (induced_vector v) y" using Ring1_FH Ring1.add_closed aezfun_setspan_proj_add[of G x y] FG_proj_in_FG smult_distrib_left induced_vector_def by auto } thus "\x\FH. \y\FH. induced_vector v (x + y) = induced_vector v x + induced_vector v y" by fast { fix a g x assume g: "g \ G" and x: "x \ FH" with v FH have "(induced_vector v) (a \\ g * x) = a \\ g \ (induced_vector v) x" using FG_subring_FH FG_fddg_closed Ring1_FH Ring1.mult_closed[of FH "a \\ g"] FG_proj_mult_leftdelta[of g a] FG_fddg_closed FG_proj_in_FG smult_assoc induced_vector_def by fastforce } thus "\a. \g\G. \x\FH. induced_vector v (a \\ g * x) = a \\ g \ induced_vector v x" by fast qed lemma induced_vector_additive : "v \ V \ v' \ V \ induced_vector (v+v') = induced_vector v + induced_vector v'" using add_closed induced_vector_def FG_proj_in_FG smult_distrib_left by auto lemma hom_induced_vector : "FGModuleHom G smult V rrsmult induced_vector" proof show "\v v'. v \ V \ v' \ V \ induced_vector (v + v') = induced_vector v + induced_vector v'" using induced_vector_additive by fast have "induced_vector = (\v. if v\V then \y. if y \ FH then (FG_proj y) \ v else 0 else 0)" using induced_vector_def by fast thus "supp induced_vector \ V" using supp_restrict0[of V] by fastforce show "\x v. x \ BaseFGMod.FG \ v \ V \ induced_vector (x \ v) = x \ induced_vector v" proof- fix x v assume xv: "x \ BaseFGMod.FG" "v \ V" show "induced_vector (x \ v) = x \ induced_vector v" proof fix a from xv FH show "induced_vector (x \ v) a = (x \ induced_vector v) a" using smult_closed induced_vector_def FG_proj_in_FG smult_assoc rrsmultD1 FG_subring_FH Ring1.mult_closed[OF Ring1_FH] induced_vector_apply1 FG_proj_mult_right[of x] smult_closed induced_vector_apply2 rrsmultD2 by auto qed qed qed lemma indspace_sum_list_fddh: "\ fhs \ []; set (map snd fhs) \ H; f \ indV \ \ f (\(a,h)\fhs. a \\ h) = (\(a,h)\fhs. f (a \\ h))" proof (induct fhs rule: list_nonempty_induct) case (single fh) show ?case using split_beta[of "\a h. a \\ h" fh] split_beta[of "\a h. f (a \\ h)" fh] by simp next case (cons fh fhs) hence prevcase: "snd fh \ H" "set (map snd fhs) \ H" "f \ indV" "f (\(a,h)\fhs. a \\ h) = (\(a,h)\fhs. f (a \\ h))" by auto have "f (\(a,h)\fh#fhs. a \\ h) = f ((fst fh) \\ (snd fh) + (\ah\fhs. case_prod (\a h. a \\ h) ah))" using split_beta[of "\a h. a \\ h" fh] by simp moreover from prevcase(1) FH have "(fst fh) \\ (snd fh) \ FH" using Supgroup.RG_aezdeltafun_closed by fast moreover from prevcase(2) FH have "(\ah\fhs. case_prod (\a h. a \\ h) ah) \ FH" using Supgroup.RG_aezdeltafun_closed Ring1.sum_list_closed[OF Ring1_FH, of "\ah. case_prod (\a h. a \\ h) ah" fhs] by fastforce ultimately have "f (\(a,h)\fh#fhs. a \\ h) = f ((fst fh) \\ (snd fh)) + f (\(a,h)\fhs. a \\ h)" using indV prevcase(3) BaseFGMod.indspaceD_add by simp with prevcase(4) show ?case using split_beta[of "\a h. f (a \\ h)" fh] by simp qed lemma induced_fsmult_conv_fsmult_1ddh : "f \ indV \ h \ H \ (r \\ f) (1 \\ h) = r \\ (f (1 \\ h))" using FH indV induced_smult.fsmultD Supgroup.RG_aezdeltafun_closed[of h "1::'f"] rrsmultD1 aezdeltafun_decomp'[of r h] aezdeltafun_decomp[of r h] Supgroup.RG_aezdeltafun_closed[of h "1::'f"] Group.zero_closed[OF GroupG] BaseFGMod.indspaceD_transform[of f G FH V 0 "(1::'f) \\ h" r] BaseFGMod.fsmultD by simp lemma indspace_el_eq_on_1ddh_imp_eq_on_rddh : assumes "HmodG \ H" "H = (\h\HmodG. G + {h})" "f \ indV" "f' \ indV" "\h\HmodG. f (1 \\ h) = f' (1 \\ h)" "h \ H" shows "f (r \\ h) = f' (r \\ h)" proof- from assms(2,6) obtain h' where h': "h' \ HmodG" "h \ G + {h'}" by fast from h'(2) obtain g where g: "g \ G" "h = g + h'" using set_plus_def[of G] by auto from g(2) have "r \\ h = r \\ 0 * (1 \\ (g+h'))" using aezdeltafun_decomp by simp moreover have "(1::'f) \\ (g+h') = 1 \\ g * (1 \\ h')" using times_aezdeltafun_aezdeltafun[of "1::'f", THEN sym] by simp ultimately have "r \\ h = r \\ g * (1 \\ h')" using aezdeltafun_decomp[of r g] by (simp add: algebra_simps) with indV FH assms(1,3,4) g(1) h'(1) have "f (r \\ h) = r \\ g \ f (1 \\ h')" "f' (r \\ h) = r \\ g \ f' (1 \\ h')" using Supgroup.RG_aezdeltafun_closed[of h' 1] BaseFGMod.indspaceD_transform[of f G FH V g "1 \\ h'" r] BaseFGMod.indspaceD_transform[of f' G FH V g "1 \\ h'" r] by auto thus "f (r \\ h) = f' (r \\ h)" using h'(1) assms(5) by simp qed lemma indspace_el_eq : assumes "HmodG \ H" "H = (\h\HmodG. G + {h})" "f \ indV" "f' \ indV" "\h\HmodG. f (1 \\ h) = f' (1 \\ h)" shows "f = f'" proof fix x show "f x = f' x" proof (cases "x = 0" "x \ FH" rule: conjcases) case BothTrue hence "x = 0 \\ 0" using zero_aezfun_transfer by simp with assms show ?thesis using indspace_el_eq_on_1ddh_imp_eq_on_rddh[of HmodG f f'] Supgroup.zero_closed by auto next case OneTrue with FH show ?thesis using Supgroup.RG_zero_closed by fast next case OtherTrue with FH obtain rhs where rhs: "set (map snd rhs) \ H" "x = (\(r,h)\rhs. r \\ h)" using Supgroup.RG_el_decomp_aezdeltafun by fast from OtherTrue rhs(2) have rhs_nnil: "rhs \ []" by auto with assms(3,4) rhs have "f x = (\(r,h)\rhs. f (r \\ h))" "f' x = (\(r,h)\rhs. f' (r \\ h))" using indspace_sum_list_fddh by auto moreover from rhs(1) assms have "\(r,h) \ set rhs. f (r \\ h) = f' (r \\ h)" using indspace_el_eq_on_1ddh_imp_eq_on_rddh[of HmodG f f'] by fastforce ultimately show ?thesis using sum_list_prod_cong[of rhs "\r h. f (r \\ h)"] by simp next case BothFalse with indV assms(3,4) show ?thesis using BaseFGMod.indspaceD_supp'[of f] BaseFGMod.indspaceD_supp'[of f'] by simp qed qed lemma indspace_el_eq' : assumes "set hs \ H" "H = (\h\set hs. G + {h})" "f \ indV" "f' \ indV" "\i\ (hs!i)) = f' (1 \\ (hs!i))" shows "f = f'" using assms(1-4) proof (rule indspace_el_eq[of "set hs"]) have "\h. h\set hs \ f (1 \\ h) = f' (1 \\ h)" proof- fix h assume "h \ set hs" from this obtain i where "i < length hs" "h = hs!i" using in_set_conv_nth[of h] by fast with assms(5) show "f (1 \\ h) = f' (1 \\ h)" by simp qed thus "\h\set hs. f (1 \\ h) = f' (1 \\ h)" by fast qed end (* context InducedFHModule *) section \Representations of Finite Groups\ subsection \Locale and basic facts\ text \ Define a group representation to be a module over the group ring that is finite-dimensional as a vector space. The only restriction on the characteristic of the field is that it does not divide the order of the group. Also, we don't explicitly assume that the group is finite; instead, the \good_char\ assumption implies that the cardinality of G is not zero, which implies G is finite. (See lemma \good_card_imp_finite\.) \ locale FinGroupRepresentation = FGModule G smult V for G :: "'g::group_add set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixl "\" 70) and V :: "'v set" + assumes good_char: "of_nat (card G) \ (0::'f)" and findim : "fscalar_mult.findim fsmult V" lemma (in Group) trivial_FinGroupRep : fixes smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" assumes good_char : "of_nat (card G) \ (0::'f)" and smult_zero : "\a\group_ring. smult a (0::'v) = 0" shows "FinGroupRepresentation G smult (0::'v set)" proof (rule FinGroupRepresentation.intro) from smult_zero show "FGModule G smult (0::'v set)" using trivial_FGModule by fast have "fscalar_mult.findim (aezfun_scalar_mult.fsmult smult) 0" by auto (metis R_scalar_mult.RSpan.simps(1) aezfun_scalar_mult.R_scalar_mult empty_set empty_subsetI set_zero) with good_char show "FinGroupRepresentation_axioms G smult 0" by unfold_locales qed context FinGroupRepresentation begin abbreviation ordG :: 'f where "ordG \ of_nat (card G)" abbreviation "GRepHom \ FGModuleHom G smult V" abbreviation "GRepIso \ FGModuleIso G smult V" abbreviation "GRepEnd \ FGModuleEnd G smult V" lemmas zero_closed = zero_closed lemmas Group = Group lemmas GSubmodule_GSpan_single = RSubmodule_RSpan_single lemmas GSpan_single_nonzero = RSpan_single_nonzero lemma finiteG: "finite G" using good_char good_card_imp_finite by fast lemma FinDimVectorSpace: "FinDimVectorSpace fsmult V" using findim fVectorSpace VectorSpace.FinDimVectorSpaceI by fast lemma GSubspace_is_FinGroupRep : assumes "GSubspace U" shows "FinGroupRepresentation G smult U" proof ( rule FinGroupRepresentation.intro, rule GSubspace_is_FGModule[OF assms], unfold_locales ) from assms show "fscalar_mult.findim fsmult U" using FinDimVectorSpace GSubspace_is_Subspace FinDimVectorSpace.Subspace_is_findim by fast qed (rule good_char) lemma isomorphic_imp_GRep : assumes "isomorphic smult' W" shows "FinGroupRepresentation G smult' W" proof (rule FinGroupRepresentation.intro) from assms show "FGModule G smult' W" using FGModuleIso.ImG FGModuleHom.FGModule_Im[OF FGModuleIso.axioms(1)] by fast from assms have "fscalar_mult.findim (aezfun_scalar_mult.fsmult smult') W" using FGModuleIso.ImG findim FGModuleIso.VectorSpaceHom VectorSpaceHom.findim_domain_findim_image by fastforce with good_char show "FinGroupRepresentation_axioms G smult' W" by unfold_locales qed end (* context FinGroupRepresentation *) subsection \Irreducible representations\ locale IrrFinGroupRepresentation = FinGroupRepresentation + assumes irr: "GSubspace U \ U = 0 \ U = V" begin lemmas AbGroup = AbGroup lemma zero_isomorphic_to_FG_zero : assumes "V = 0" shows "isomorphic (*) (0::('b,'a) aezfun set)" proof show "GRepIso (*) 0 0" proof (rule FGModuleIso.intro) show "GRepHom (*) 0" using trivial_FGModuleHom[of "(*)"] by simp show "FGModuleIso_axioms V 0 0" proof from assms show "bij_betw 0 V 0" unfolding bij_betw_def inj_on_def by simp qed qed qed lemma eq_GSpan_single : "v \ V \ v \ 0 \ V = GSpan [v]" using irr RSubmodule_RSpan_single RSpan_single_nonzero by fast end (* context IrrFinGroupRepresentation *) lemma (in Group) trivial_IrrFinGroupRepI : fixes smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" assumes "of_nat (card G) \ (0::'f)" and "\a\group_ring. smult a (0::'v) = 0" shows "IrrFinGroupRepresentation G smult (0::'v set)" proof (rule IrrFinGroupRepresentation.intro) from assms show "FinGroupRepresentation G smult 0" using trivial_FinGroupRep by fast show "IrrFinGroupRepresentation_axioms G smult 0" using RModule.zero_closed by unfold_locales auto qed lemma (in Group) trivial_IrrFinGroupRepresentation_in_FG : "of_nat (card G) \ (0::'f::field) \ IrrFinGroupRepresentation G (*) (0::('f,'g) aezfun set)" using trivial_IrrFinGroupRepI[of "(*)"] by simp context FinGroupRepresentation begin lemma IrrFinGroupRep_trivialGSubspace : "IrrFinGroupRepresentation G smult (0::'v set)" proof- have "ordG \ (0::'f)" using good_char by fast moreover have "\a\FG. a \ 0 = 0" using smult_zero by simp ultimately show ?thesis using ActingGroup.trivial_IrrFinGroupRepI[of smult] by fast qed lemma IrrI : assumes "\U. FGModule.GSubspace G smult V U \ U = 0 \ U = V" shows "IrrFinGroupRepresentation G smult V" using assms IrrFinGroupRepresentation.intro by unfold_locales lemma notIrr : "\ IrrFinGroupRepresentation G smult V \ \U. GSubspace U \ U \ 0 \ U \ V" using IrrI by fast end (* context FinGroupRepresentation *) subsection \Maschke's theorem\ subsubsection \Averaged projection onto a G-subspace\ context FinGroupRepresentation begin lemma avg_proj_eq_id_on_right : assumes "VectorSpace fsmult W" "add_independentS [W,V]" "v \ V" defines P : "P \ (\[W,V]\1)" defines CP: "CP \ (\g. Gmult (- g) \ P \ Gmult g)" defines T : "T \ fsmult (1/ordG) \ (\g\G. CP g)" shows "T v = v" proof- from P assms(2,3) have "\g. g \ G \ P (g *\ v) = g *\ v" using Gmult_closed VectorSpace.AbGroup[OF assms(1)] AbGroup AbGroup_inner_dirsum_el_decomp_nth_id_on_nth[of "[W,V]"] by simp with CP assms(3) have "\g. g \ G \ CP g v = v" using Gmult_neg_left by simp with assms(3) good_char T show "T v = v" using finiteG sum_fun_apply[of G CP] sum_fsmult_distrib[of v G "\x. 1"] fsmult_assoc[of _ ordG v] by simp qed lemma avg_proj_onto_right : assumes "VectorSpace fsmult W" "GSubspace U" "add_independentS [W,U]" "V = W \ U" defines P : "P \ (\[W,U]\1)" defines CP: "CP \ (\g. Gmult (- g) \ P \ Gmult g)" defines T : "T \ fsmult (1/ordG) \ (\g\G. CP g)" shows "T ` V = U" proof from assms(2) have U: "FGModule G smult U" using GSubspace_is_FGModule by fast show "T ` V \ U" proof (rule image_subsetI) fix v assume v: "v \ V" with assms(1,3,4) P U have "\g. g \ G \ P (g *\ v) \ U" using Gmult_closed VectorSpace.AbGroup FGModule.AbGroup AbGroup_inner_dirsum_el_decomp_nth_onto_nth[of "[W,U]" 1] by fastforce with U CP have "\g. g \ G \ CP g v \ U" using FGModule.Gmult_closed GroupG Group.neg_closed by fastforce with assms(2) U T show "T v \ U" using finiteG FGModule.sum_closed[of G smult U G "\g. CP g v"] sum_fun_apply[of G CP] FGModule.fsmult_closed[of G smult U] by fastforce qed show "T ` V \ U" proof fix u assume u: "u \ U" with u T CP P assms(1,2,3) have "T u = u" using GSubspace_is_FinGroupRep FinGroupRepresentation.avg_proj_eq_id_on_right by fast from this[THEN sym] assms(1-4) u show "u \ T ` V" using Module.AbGroup RModule.AbGroup AbGroup_subset_inner_dirsum by fast qed qed lemma FGModuleEnd_avg_proj_right : assumes "fSubspace W" "GSubspace U" "add_independentS [W,U]" "V = W \ U" defines P : "P \ (\[W,U]\1)" defines CP: "CP \ (\g. Gmult (- g) \ P \ Gmult g)" defines T : "T \ (fsmult (1/ordG) \ (\g\G. CP g))\V" shows "FGModuleEnd G smult V T" proof (rule VecEnd_GMap_is_FGModuleEnd) from T have T_apply: "\v. v\V \ T v = (1/ordG) \\ (\g\G. CP g v)" using finiteG sum_fun_apply[of G CP] by simp from assms(1-4) P have Pgv: "\g v. g \ G \ v \ V \ P ( g *\ v ) \ V" using Gmult_closed VectorSpace_fSubspace VectorSpace.AbGroup[of fsmult W] RModule.AbGroup[of FG smult U] GroupEnd_inner_dirsum_el_decomp_nth[of "[W,U]" 1] GroupEnd.endomorph[of V] by force have im_CP_V: "\v. v \ V \ (\g. CP g v) ` G \ V" proof- fix v assume "v \ V" thus "(\g. CP g v) ` G \ V" using CP Pgv[of _ v] ActingGroup.neg_closed Gmult_closed finiteG by auto qed have sumCP_V: "\v. v \ V \ (\g\G. CP g v) \ V" using finiteG im_CP_V sum_closed by force show "VectorSpaceEnd (\\) V T" proof ( rule VectorSpaceEndI, rule VectorSpace.VectorSpaceHomI, rule fVectorSpace ) show "GroupHom V T" proof fix v v' assume vv': "v \ V" "v' \ V" with T_apply have "T (v + v') = (1/ordG) \\ (\g\G. CP g (v + v'))" using add_closed by fast moreover have "\g. g \ G \ CP g (v + v') = CP g v + CP g v'" proof- fix g assume g: "g \ G" with CP P vv' assms(1-4) have "CP g (v + v') = (- g) *\ ( P ( g *\ v ) + P ( g *\ v' ) )" using Gmult_distrib_left Gmult_closed VectorSpace_fSubspace VectorSpace.AbGroup[of fsmult W] RModule.AbGroup[of FG smult U] GroupEnd_inner_dirsum_el_decomp_nth[of "[W,U]" 1] GroupEnd.hom[of V P] by simp with g vv' have "CP g (v + v') = (- g) *\ P ( g *\ v ) + (- g) *\ P ( g *\ v' )" using Pgv ActingGroup.neg_closed Gmult_distrib_left by simp thus "CP g (v + v') = CP g v + CP g v'" using CP by simp qed ultimately show "T (v + v') = T v + T v'" using vv' sumCP_V[of v] sumCP_V[of v'] sum.distrib[of "\g. CP g v"] T_apply by simp next from T show "supp T \ V" using supp_restrict0 by fast qed show "\a v. v \ V \ T (a \\ v) = a \\ T v" proof- fix a::'f and v assume v: "v \ V" with T_apply have "T (a \\ v) = (1/ordG) \\ (\g\G. CP g (a \\ v))" using fsmult_closed by fast moreover have "\g. g \ G \ CP g (a \\ v) = a \\ CP g v" proof- fix g assume "g \ G" with assms(1-4) CP P v show "CP g (a \\ v) = a \\ CP g v" using fsmult_Gmult_comm GSubspace_is_Subspace Gmult_closed fVectorSpace VectorSpace.VectorSpaceEnd_inner_dirsum_el_decomp_nth[of fsmult] VectorSpaceEnd.f_map[of fsmult "(\N\[W, U]. N)" P a "g *\ v"] ActingGroup.neg_closed Pgv by simp qed ultimately show "T (a \\ v) = a \\ T v" using v im_CP_V sumCP_V T_apply finiteG fsmult_sum_distrib[of a G "\g. CP g v"] fsmult_assoc[of "1/ordG" a "(\g\G. CP g v)"] by simp qed show "T ` V \ V" using sumCP_V fsmult_closed T_apply image_subsetI by auto qed show "\g v. g \ G \ v \ V \ T (g *\ v) = g *\ T v" proof- fix g v assume g: "g \ G" and v: "v \ V" with T_apply have "T (g *\ v) = (1/ordG) \\ (\h\G. CP h (g *\ v))" using Gmult_closed by fast with g have "T (g *\ v) = (1/ordG) \\ (\h\G. CP (h + - g) (g *\ v))" using GroupG Group.neg_closed Group.right_translate_sum[of G "- g" "\h. CP h (g *\ v)"] by fastforce moreover from CP have "\h. h \ G \ CP (h + - g) (g *\ v) = g *\ CP h v" using g v Gmult_closed[of g v] ActingGroup.neg_closed Gmult_assoc[of _ "- g" "g *\ v", THEN sym] Gmult_neg_left minus_add[of _ "- g"] Pgv Gmult_assoc by simp ultimately show "T (g *\ v) = g *\ T v" using g v GmultD finiteG FG_fddg_closed im_CP_V sumCP_V smult_sum_distrib[of "1 \\ g" G] fsmult_Gmult_comm[of g "\h\G. (CP h v)"] T_apply by simp qed qed lemma avg_proj_is_proj_right : assumes "VectorSpace fsmult W" "GSubspace U" "add_independentS [W,U]" "V = W \ U" "v \ V" defines P : "P \ (\[W,U]\1)" defines CP: "CP \ (\g. Gmult (- g) \ P \ Gmult g)" defines T : "T \ fsmult (1/ordG) \ (\g\G. CP g)" shows "T (T v) = T v" using assms avg_proj_onto_right GSubspace_is_FinGroupRep FinGroupRepresentation.avg_proj_eq_id_on_right by fast end (* context FinGroupRepresentation *) subsubsection \The theorem\ context FinGroupRepresentation begin theorem Maschke : assumes "GSubspace U" shows "\W. GSubspace W \ V = W \ U" proof (cases "V = 0") case True moreover from assms True have "U = 0" using RModule.zero_closed by auto ultimately have "V = 0 + U" using set_plus_def by fastforce moreover have "add_independentS [0,U]" by simp ultimately have "V = 0 \ U" using inner_dirsum_doubleI by fast moreover have "GSubspace 0" using trivial_RSubmodule zero_closed by auto ultimately show "\W. GSubspace W \ V = W \ U" by fast next case False from assms obtain W' where W': "VectorSpace.Subspace fsmult V W' \ V = W' \ U" using GSubspace_is_Subspace FinDimVectorSpace FinDimVectorSpace.semisimple by force hence vsp_W': "VectorSpace fsmult W'" and dirsum: "V = W' \ U" using VectorSpace.SubspaceD1[OF fVectorSpace] by auto from False dirsum have indS: "add_independentS [W',U]" using inner_dirsumD2 by fastforce define P where "P = (\[W',U]\1)" define CP where "CP = (\g. Gmult (- g) \ P \ Gmult g)" define S where "S = fsmult (1/ordG) \ (\g\G. CP g)" define W where "W = GroupHom.Ker V (S\V)" from assms W' indS S_def CP_def P_def have endo: "GRepEnd (S\V)" using FGModuleEnd_avg_proj_right by fast moreover from S_def CP_def P_def have "\v. v \ V \ (S\V) ((S\V) v) = (S\V) v" using endo FGModuleEnd.endomorph avg_proj_is_proj_right[OF vsp_W' assms indS dirsum] by fastforce moreover have "(S\V) ` V = U" proof- from assms indS P_def CP_def S_def dirsum vsp_W' have "S ` V = U" using avg_proj_onto_right by fast moreover have "(S\V) ` V = S ` V" by auto ultimately show ?thesis by fast qed ultimately have "V = W \ U" "GSubspace W" using W_def FGModuleEnd.proj_decomp[of G smult V "S\V"] FGModuleEnd.GSubspace_Ker by auto thus ?thesis by fast qed corollary Maschke_proper : assumes "GSubspace U" "U \ 0" "U \ V" shows "\W. GSubspace W \ W \ 0 \ W \ V \ V = W \ U" proof- from assms(1) obtain W where W: "GSubspace W" "V = W \ U" using Maschke by fast from assms(3) W(2) have "W \ 0" using inner_dirsum_double_left0 by fast moreover from assms(1,2) W have "W \ V" using Subgroup_RSubmodule Group.nonempty inner_dirsum_double_leftfull_imp_right0[of W U] by fastforce ultimately show ?thesis using W by fast qed end (* context FinGroupRepresentation *) subsubsection \Consequence: complete reducibility\ lemma (in FinGroupRepresentation) notIrr_decompose : "\ IrrFinGroupRepresentation G smult V \ \U W. GSubspace U \ U \ 0 \ U \ V \ GSubspace W \ W \ 0 \ W \ V \ V = U \ W" using notIrr Maschke_proper by blast text \ In the following decomposition lemma, we do not need to explicitly include the condition that all @{term U} in @{term "set Us"} are subsets of @{term V}. (See lemma \AbGroup_subset_inner_dirsum\.) \ lemma FinGroupRepresentation_reducible' : fixes n::nat shows "\V. FinGroupRepresentation G fgsmult V \ n = FGModule.fdim fgsmult V \ (\Us. Ball (set Us) (IrrFinGroupRepresentation G fgsmult) \ (0 \ set Us) \ V = (\U\Us. U) )" proof (induct n rule: full_nat_induct) fix n V define GRep IGRep GSubspace GSpan fdim where "GRep = FinGroupRepresentation G fgsmult" and "IGRep = IrrFinGroupRepresentation G fgsmult" and "GSubspace = FGModule.GSubspace G fgsmult V" and "GSpan = FGModule.GSpan G fgsmult" and "fdim = FGModule.fdim fgsmult" assume "\m. Suc m \ n \ (\x. GRep x \ m = fdim x \ ( \Us. Ball (set Us) IGRep \ (0 \ set Us) \ x = (\U\Us. U)) )" hence prev_case: "\m. m < n \ (\W. GRep W \ m = fdim W \ ( \Us. Ball (set Us) IGRep \ (0 \ set Us) \ W = (\U\Us. U)))" using Suc_le_eq by auto show "GRep V \ n = fdim V \ ( \Us. Ball (set Us) IGRep \ (0 \ set Us) \ V = (\U\Us. U) )" proof- assume V: "GRep V \ n = fdim V" show "(\Us. Ball (set Us) IGRep \ (0 \ set Us) \ V = (\U\Us. U))" proof (cases "IGRep V" "V = 0" rule: conjcases) case BothTrue moreover have "Ball (set []) IGRep" "\U\set []. U \ 0" by auto ultimately show ?thesis using inner_dirsum_Nil by fast next case OneTrue with V GRep_def obtain v where v: "v \ V" "v \ 0" using FinGroupRepresentation.Group[of G fgsmult] Group.obtain_nonzero by auto from v(1) V GRep_def GSpan_def GSubspace_def have GSub: "GSubspace (GSpan [v])" using FinGroupRepresentation.GSubmodule_GSpan_single by fast moreover from v V GRep_def GSpan_def have nzero: "GSpan [v] \ 0" using FinGroupRepresentation.GSpan_single_nonzero by fast ultimately have "V = GSpan [v]" using OneTrue GSpan_def GSubspace_def IGRep_def IrrFinGroupRepresentation.irr by fast with OneTrue have "Ball (set [GSpan [v]]) IGRep" "0 \ set [GSpan [v]]" "V = (\U\[GSpan [v]]. U)" using nzero GSub inner_dirsum_singleD by auto thus ?thesis by fast next case OtherTrue with V GRep_def IGRep_def show ?thesis using FinGroupRepresentation.IrrFinGroupRep_trivialGSubspace by fast next case BothFalse with V GRep_def IGRep_def GSubspace_def obtain U W where U: "GSubspace U" "U \ 0" "U \ V" and W: "GSubspace W" "W \ 0" "W \ V" and Vdecompose: "V = U \ W" using FinGroupRepresentation.notIrr_decompose[of G fgsmult V] by auto from U(1,3) W(1,3) V GRep_def GSubspace_def fdim_def have "fdim U < fdim V" "fdim W < fdim V" using FinGroupRepresentation.axioms(1) FGModule.GSubspace_is_Subspace[of G fgsmult V U] FGModule.GSubspace_is_Subspace[of G fgsmult V W] FinGroupRepresentation.FinDimVectorSpace[of G fgsmult V] FinDimVectorSpace.Subspace_dim_lt[ of "aezfun_scalar_mult.fsmult fgsmult" V U ] FinDimVectorSpace.Subspace_dim_lt[ of "aezfun_scalar_mult.fsmult fgsmult" V W ] by auto from this U(1) W(1) V GSubspace_def obtain Us Ws where "Ball (set Us) IGRep \ (0 \ set Us) \ U = (\X\Us. X)" and "Ball (set Ws) IGRep \ (0 \ set Ws) \ W = (\X\Ws. X)" using prev_case[of "fdim U"] prev_case[of "fdim W"] GRep_def FinGroupRepresentation.GSubspace_is_FinGroupRep[ of G fgsmult V U ] FinGroupRepresentation.GSubspace_is_FinGroupRep[ of G fgsmult V W ] by fastforce hence Us: "Ball (set Us) IGRep" "0 \ set Us" "U = (\X\Us. X)" and Ws: "Ball (set Ws) IGRep" "0 \ set Ws" "W = (\X\Ws. X)" by auto from Us(1) Ws(1) have "Ball (set (Us@Ws)) IGRep" by auto moreover from Us(2) Ws(2) have "0 \ set (Us@Ws)" by auto moreover have "V = (\X\(Us@Ws). X)" proof- from U(2) Us(3) W(2) Ws(3) have indUs: "add_independentS Us" and indWs: "add_independentS Ws" using inner_dirsumD2 by auto moreover from IGRep_def Us(1) have "Ball (set Us) ((\) 0)" using IrrFinGroupRepresentation.axioms(1)[of G fgsmult] FinGroupRepresentation.zero_closed[of G fgsmult] by fast moreover from Us(3) Ws(3) BothFalse Vdecompose indUs indWs have "add_independentS [(\X\Us. X),(\X\Ws. X)]" using inner_dirsumD2[of "[U,W]"] inner_dirsumD[of Us] inner_dirsumD[of Ws] by auto ultimately have "add_independentS (Us@Ws)" by (metis add_independentS_double_sum_conv_append) moreover from W(1) Ws(3) indWs have "0 \ (\X\Ws. X)" using inner_dirsumD GSubspace_def RModule.zero_closed by fast ultimately show ?thesis using Vdecompose Us(3) Ws(3) inner_dirsum_append by fast qed ultimately show ?thesis by fast qed qed qed theorem (in FinGroupRepresentation) reducible : "\Us. (\U\set Us. IrrFinGroupRepresentation G smult U) \ (0 \ set Us) \ V = (\U\Us. U)" using FinGroupRepresentation_axioms FinGroupRepresentation_reducible' by fast subsubsection \Consequence: decomposition relative to a homomomorphism\ lemma (in FinGroupRepresentation) GRepHom_decomp : fixes T :: "'v \ 'w::ab_group_add" defines KerT : "KerT \ (ker T \ V)" assumes hom : "GRepHom smult' T" and nonzero: "V \ 0" shows "\U. GSubspace U \ V = U \ KerT \ FGModule.isomorphic G smult U smult' (T ` V)" proof- from KerT have KerT': "GSubspace KerT" using FGModuleHom.GSubspace_Ker[OF hom] by fast from this obtain U where U: "GSubspace U" "V = U \ KerT" using Maschke[of KerT] by fast from nonzero U(2) have indep: "add_independentS [U,KerT]" using inner_dirsumD2 by fastforce have "FGModuleIso G smult U smult' (T \ U) (T ` V)" proof (rule FGModuleIso.intro) from U(1) show "FGModuleHom G (\) U smult' (T \ U)" using FGModuleHom.FGModuleHom_restrict0_GSubspace[OF hom] by fast show "FGModuleIso_axioms U (T \ U) (T ` V)" unfolding FGModuleIso_axioms_def bij_betw_def proof (rule conjI, rule inj_onI) show "(T \ U) ` U = T ` V" proof from U(1) show "(T \ U) ` U \ T ` V" by auto show "(T \ U) ` U \ T ` V" proof (rule image_subsetI) fix v assume "v \ V" with U(2) obtain u k where uk: "u \ U" "k \ KerT" "v = u + k" using inner_dirsum_doubleD[OF indep] set_plus_def[of U KerT] by fast with KerT U(1) have "T v = (T \ U) u" using kerD FGModuleHom.additive[OF hom] by force with uk(1) show "T v \ (T \ U) ` U" by fast qed qed next fix x y assume xy: "x \ U" "y \ U" "(T \ U) x = (T \ U) y" with U(1) KerT have "x - y \ U \ KerT" using FGModuleHom.eq_im_imp_diff_in_Ker[OF hom] GSubspace_is_FGModule FGModule.diff_closed[of G smult U x y] by auto moreover from U(1) have "AbGroup U" using RModule.AbGroup by fast moreover from KerT' have "AbGroup KerT" using RModule.AbGroup by fast ultimately show "x = y" using indep AbGroup_inner_dirsum_pairwise_int0_double[of U KerT] by force qed qed with U show ?thesis by fast qed subsection \Schur's lemma\ lemma (in IrrFinGroupRepresentation) Schur_Ker : "GRepHom smult' T \ T ` V \ 0 \ inj_on T V" using irr FGModuleHom.GSubspace_Ker[of G smult V smult' T] FGModuleHom.Ker_Im_iff[of G smult V smult' T] FGModuleHom.Ker0_imp_inj_on[of G smult V smult' T] by auto lemma (in FinGroupRepresentation) Schur_Im : assumes "IrrFinGroupRepresentation G smult' W" "GRepHom smult' T" "T ` V \ W" "T ` V \ 0" shows "T ` V = W" proof- have "FGModule.GSubspace G smult' W (T ` V)" proof from assms(2) show "RModule FG smult' (T ` V)" using FGModuleHom.axioms(2)[of G] RModuleHom.RModule_Im[of FG smult V smult' T] by fast from assms(3) show "T ` V \ W" by fast qed with assms(1,4) show ?thesis using IrrFinGroupRepresentation.irr by fast qed theorem (in IrrFinGroupRepresentation) Schur1 : assumes "IrrFinGroupRepresentation G smult' W" "GRepHom smult' T" "T ` V \ W" "T ` V \ 0" shows "GRepIso smult' T W" proof (rule FGModuleIso.intro, rule assms(2), unfold_locales) show "bij_betw T V W" using IrrFinGroupRepresentation_axioms assms IrrFinGroupRepresentation.axioms(1)[of G] FinGroupRepresentation.Schur_Im[of G] IrrFinGroupRepresentation.Schur_Ker[of G smult V smult' T] unfolding bij_betw_def by fast qed theorem (in IrrFinGroupRepresentation) Schur2 : assumes GRep : "GRepEnd T" and fdim : "fdim > 0" and alg_closed: "\p::'b poly. degree p > 0 \ \c. poly p c = 0" shows "\c. \v \ V. T v = c \\ v" proof- from fdim alg_closed obtain e u where eu: "u \ V" "u \ 0" "T u = e \\ u" using FGModuleEnd.VectorSpaceEnd[OF GRep] FinDimVectorSpace.endomorph_has_eigenvector[ OF FinDimVectorSpace, of T ] by fast define U where "U = {v \ V. T v = e \\ v}" moreover from eu U_def have "U \ 0" by auto ultimately have "\v \ V. T v = e \\ v" using GRep irr FGModuleEnd.axioms(1)[of G smult V T] GSubspace_eigenspace[of G smult] by fast thus ?thesis by fast qed subsection \The group ring as a representation space\ subsubsection \The group ring is a representation space\ lemma (in Group) FGModule_FG : defines FG: "FG \ group_ring :: ('f::field, 'g) aezfun set" shows "FGModule G (*) FG" proof (rule FGModule.intro, rule Group_axioms, rule RModuleI) show 1: "Ring1 group_ring" using Ring1_RG by fast from 1 FG show "Group FG" using Ring1.axioms(1) by fast from 1 FG show "RModule_axioms group_ring (*) FG" using Ring1.mult_closed by unfold_locales (auto simp add: algebra_simps) qed theorem (in Group) FinGroupRepresentation_FG : defines FG: "FG \ group_ring :: ('f::field, 'g) aezfun set" assumes good_char: "of_nat (card G) \ (0::'f)" shows "FinGroupRepresentation G (*) FG" proof (rule FinGroupRepresentation.intro) from FG show "FGModule G (*) FG" using FGModule_FG by fast show "FinGroupRepresentation_axioms G (*) FG" proof from FG good_char obtain gs where gs: "set gs = G" "\f\ FG. \bs. length bs = length gs \ f = (\(b,g)\zip bs gs. (b \\ 0) * (1 \\ g))" using good_card_imp_finite FinGroupI FinGroup.group_ring_spanning_set by fast define xs where "xs = map ((\\) (1::'f)) gs" with FG gs(1) have 1: "set xs \ FG" using RG_aezdeltafun_closed by auto moreover have "aezfun_scalar_mult.fSpan (*) xs = FG" proof from 1 FG show "aezfun_scalar_mult.fSpan (*) xs \ FG" using FGModule_FG FGModule.fSpan_closed by fast show "aezfun_scalar_mult.fSpan (*) xs \ FG" proof fix x assume "x \ FG" from this gs(2) obtain bs where bs: "length bs = length gs" "x = (\(b,g)\zip bs gs. (b \\ 0) * (1 \\ g))" by fast from bs(2) xs_def have "x = (\(b,a)\zip bs xs. (b \\ 0) * a)" using sum_list_prod_map2[THEN sym] by fast with bs(1) xs_def show "x \ aezfun_scalar_mult.fSpan (*) xs" using aezfun_scalar_mult.fsmultD[of "(*)", THEN sym] sum_list_prod_cong[ of "zip bs xs" "\b a. (b \\ 0) * a" "\b a. aezfun_scalar_mult.fsmult (*) b a" ] scalar_mult.lincomb_def[of "aezfun_scalar_mult.fsmult (*)" bs xs] scalar_mult.SpanD_lincomb[of "aezfun_scalar_mult.fsmult (*)"] by force qed qed ultimately show "\xs. set xs \ FG \ aezfun_scalar_mult.fSpan (*) xs = FG" by fast qed (rule good_char) qed lemma (in FinGroupRepresentation) FinGroupRepresentation_FG : "FinGroupRepresentation G (*) FG" using good_char ActingGroup.FinGroupRepresentation_FG by fast lemma (in Group) FG_reducible : assumes "of_nat (card G) \ (0::'f::field)" shows "\Us::('f,'g) aezfun set list. (\U\set Us. IrrFinGroupRepresentation G (*) U) \ 0 \ set Us \ group_ring = (\U\Us. U)" using assms FinGroupRepresentation_FG FinGroupRepresentation.reducible by fast subsubsection \Irreducible representations are constituents of the group ring\ lemma (in FGModuleIso) isomorphic_to_irr_right : assumes "IrrFinGroupRepresentation G smult' W" shows "IrrFinGroupRepresentation G smult V" proof (rule FinGroupRepresentation.IrrI) from assms show "FinGroupRepresentation G (\) V" using IrrFinGroupRepresentation.axioms(1) isomorphic_sym FinGroupRepresentation.isomorphic_imp_GRep by fast from assms show "\U. GSubspace U \ U = 0 \ U = V" using IrrFinGroupRepresentation.irr isomorphic_to_irr_right' by fast qed lemma (in FinGroupRepresentation) IrrGSubspace_iso_constituent : assumes nonzero : "V \ 0" and subsp : "W \ V" "W \ 0" "IrrFinGroupRepresentation G smult W" and V_decomp: "\U\set Us. IrrFinGroupRepresentation G smult U" "0 \ set Us" "V = (\U\Us. U)" shows "\U\set Us. FGModule.isomorphic G smult W smult U" proof- from V_decomp(1) have abGrp: "\U\set Us. AbGroup U" using IrrFinGroupRepresentation.AbGroup by fast from nonzero V_decomp(3) have indep: "add_independentS Us" using inner_dirsumD2 by fast with V_decomp (3) have "\U\set Us. U \ V" using abGrp AbGroup_subset_inner_dirsum by fast with subsp(1,3) V_decomp(1) have GSubspaces: "GSubspace W" "\U\ set Us. GSubspace U" using IrrFinGroupRepresentation.axioms(1)[of G smult] FinGroupRepresentation.axioms(1)[of G smult] GSubspaceI by auto have "\i. i < length Us \ (\Us\i) ` W \ 0 \ FGModuleIso G smult W smult ( (\Us\i) \ W ) (Us!i)" proof- fix i assume i: "i < length Us" "(\Us\i) ` W \ 0" from i(1) V_decomp(3) have "GRepEnd (\Us\i)" using GSubspaces(2) indep GEnd_inner_dirsum_el_decomp_nth by fast hence "FGModuleHom G smult W smult ( (\Us\i) \ W )" using GSubspaces(1) FGModuleEnd.FGModuleHom_restrict0_GSubspace by fast moreover have "( (\Us\i) \ W ) ` W \ Us!i" proof- from V_decomp(3) i(1) subsp(1,3) have "(\Us\i) ` W \ Us!i" using indep abGrp AbGroup_inner_dirsum_el_decomp_nth_onto_nth by fast thus ?thesis by auto qed moreover from i(1) V_decomp(1) have "IrrFinGroupRepresentation G smult (Us!i)" by simp ultimately show "FGModuleIso G smult W smult ( (\Us\i) \ W ) (Us!i)" using i(2) IrrFinGroupRepresentation.Schur1[ OF subsp(3), of smult "Us!i" "(\Us\i) \ W " ] by auto qed moreover from nonzero V_decomp(3) have "\iUs\i) ` W = 0 \ W = 0" using inner_dirsum_Nil abGrp subsp(1) indep AbGroup_inner_dirsum_subset_proj_eq_0[of Us W] by fastforce ultimately have "\iUs\i) \ W ) (Us!i)" using subsp(2) by auto thus ?thesis using set_conv_nth[of Us] by auto qed theorem (in IrrFinGroupRepresentation) iso_FG_constituent : assumes nonzero : "V \ 0" and FG_decomp: "\U\set Us. IrrFinGroupRepresentation G (*) U" "0 \ set Us" "FG = (\U\Us. U)" shows "\U\set Us. isomorphic (*) U" proof- from nonzero obtain v where v: "v \ V" "v \ 0" using nonempty by auto define T where "T = (\x. x \ v)\FG" have "FGModuleHom G (*) FG smult T" proof (rule FGModule.FGModuleHomI_fromaxioms) show "FGModule G (*) FG" using ActingGroup.FGModule_FG by fast from T_def v(1) show "\v v'. v \ FG \ v' \ FG \ T (v + v') = T v + T v'" using Ring1.add_closed[OF Ring1] smult_distrib_right by auto from T_def show "supp T \ FG" using supp_restrict0 by fast from T_def v(1) show "\r m. r \ FG \ m \ FG \ T (r * m) = r \ T m" using ActingGroup.RG_mult_closed by auto qed then obtain W where W: "FGModule.GSubspace G (*) FG W" "FG = W \ (ker T \ FG)" "FGModule.isomorphic G (*) W smult (T ` FG)" using FG_n0 FinGroupRepresentation.GRepHom_decomp[ OF FinGroupRepresentation_FG ] by fast from T_def v have "T ` FG = V" using eq_GSpan_single RSpan_single by auto with W(3) have W': "FGModule.isomorphic G (*) W smult V" by fast with W(1) nonzero have "W \ 0" using FGModule.GSubspace_is_FGModule[OF ActingGroup.FGModule_FG] FGModule.isomorphic_to_zero_left by fastforce moreover from W' have "IrrFinGroupRepresentation G (*) W" using IrrFinGroupRepresentation_axioms FGModuleIso.isomorphic_to_irr_right by fast ultimately have "\U\set Us. FGModule.isomorphic G (*) W (*) U" using FG_decomp W(1) good_char FG_n0 FinGroupRepresentation.IrrGSubspace_iso_constituent[ OF ActingGroup.FinGroupRepresentation_FG, of W ] by simp with W(1) W' show ?thesis using FGModule.GSubspace_is_FGModule[OF ActingGroup.FGModule_FG] FGModule.isomorphic_sym[of G "(*)" W smult] isomorphic_trans by fast qed subsection \Isomorphism classes of irreducible representations\ text \ We have already demonstrated that the relation \FGModule.isomorphic\ is reflexive (lemma \FGModule.isomorphic_refl\), symmetric (lemma \FGModule.isomorphic_sym\), and transitive (lemma \FGModule.isomorphic_trans\). In this section, we provide a finite set of representatives for the resulting isomorphism classes of irreducible representations. \ context Group begin primrec remisodups :: "('f::field,'g) aezfun set list \ ('f,'g) aezfun set list" where "remisodups [] = []" | "remisodups (U # Us) = (if (\W\set Us. FGModule.isomorphic G (*) U (*) W) then remisodups Us else U # remisodups Us)" lemma set_remisodups : "set (remisodups Us) \ set Us" by (induct Us) auto lemma isodistinct_remisodups : "\ \U\set Us. FGModule G (*) U; V \ set (remisodups Us); W \ set (remisodups Us); V \ W \ \ \ (FGModule.isomorphic G (*) V (*) W)" proof (induct Us arbitrary: V W) case (Cons U Us) show ?case proof (cases "\X\set Us. FGModule.isomorphic G (*) U (*) X") case True with Cons show ?thesis by simp next case False show ?thesis proof (cases "V=U" "W=U" rule: conjcases) case BothTrue with Cons(5) show ?thesis by fast next case OneTrue with False Cons(4,5) show ?thesis using set_remisodups by auto next case OtherTrue with False Cons(2,3) show ?thesis using set_remisodups FGModule.isomorphic_sym[of G "(*)" V "(*)" W] by fastforce next case BothFalse with Cons False show ?thesis by simp qed qed qed simp definition "FG_constituents \ SOME Us. (\U\set Us. IrrFinGroupRepresentation G (*) U) \ 0 \ set Us \ group_ring = (\U\Us. U)" lemma FG_constituents_irr : "of_nat (card G) \ (0::'f::field) \ \U\set (FG_constituents::('f,'g) aezfun set list). IrrFinGroupRepresentation G (*) U" using someI_ex[OF FG_reducible] unfolding FG_constituents_def by fast lemma FG_consitutents_n0: "of_nat (card G) \ (0::'f::field) \ 0 \ set (FG_constituents::('f,'g) aezfun set list)" using someI_ex[OF FG_reducible] unfolding FG_constituents_def by fast lemma FG_constituents_constituents : "of_nat (card G) \ (0::'f::field) \ (group_ring::('f,'g) aezfun set) = (\U\FG_constituents. U)" using someI_ex[OF FG_reducible] unfolding FG_constituents_def by fast definition "GIrrRep_repset \ 0 \ set (remisodups FG_constituents)" lemma finite_GIrrRep_repset : "finite GIrrRep_repset" unfolding GIrrRep_repset_def by simp lemma all_irr_GIrrRep_repset : assumes "of_nat (card G) \ (0::'f::field)" shows "\U\(GIrrRep_repset::('f,'g) aezfun set set). IrrFinGroupRepresentation G (*) U" proof fix U :: "('f,'g) aezfun set" assume "U \ GIrrRep_repset" with assms show "IrrFinGroupRepresentation G (*) U" using trivial_IrrFinGroupRepresentation_in_FG GIrrRep_repset_def set_remisodups FG_constituents_irr by (cases "U = 0") auto qed lemma isodistinct_GIrrRep_repset : defines "GIRRS \ GIrrRep_repset :: ('f::field,'g) aezfun set set" assumes "of_nat (card G) \ (0::'f)" "V \ GIRRS" "W \ GIRRS" "V \ W" shows "\ (FGModule.isomorphic G (*) V (*) W)" proof (cases "V = 0" "W = 0" rule: conjcases) case BothTrue with assms(5) show ?thesis by fast next case OneTrue with assms(1,2,4,5) show ?thesis using GIrrRep_repset_def set_remisodups FG_consitutents_n0 trivial_FGModule[of "(*)"] FGModule.isomorphic_to_zero_left[of G "(*)"] by fastforce next case OtherTrue moreover with assms(1,3) have "V \ set FG_constituents" using GIrrRep_repset_def set_remisodups by auto ultimately show ?thesis using assms(2) FG_consitutents_n0 FG_constituents_irr IrrFinGroupRepresentation.axioms(1) FinGroupRepresentation.axioms(1) FGModule.isomorphic_to_zero_right[of G "(*)" V "(*)"] by fastforce next case BothFalse with assms(1,3,4) have "V \ set (remisodups FG_constituents)" "W \ set (remisodups FG_constituents)" using GIrrRep_repset_def by auto with assms(2,5) show ?thesis using FG_constituents_irr IrrFinGroupRepresentation.axioms(1) FinGroupRepresentation.axioms(1) isodistinct_remisodups by fastforce qed end (* context Group *) lemma (in FGModule) iso_in_list_imp_iso_in_remisodups : "\U\set Us. isomorphic (*) U \ \U\set (ActingGroup.remisodups Us). isomorphic (*) U" proof (induct Us) case (Cons U Us) from Cons(2) obtain W where W: "W \ set (U#Us)" "isomorphic (*) W" by fast show ?case proof ( cases "W = U" "\X\set Us. FGModule.isomorphic G (*) U (*) X" rule: conjcases ) case BothTrue with W(2) Cons(1) show ?thesis using isomorphic_trans[of "(*)" W] by force next case OneTrue with W(2) show ?thesis by simp next case OtherTrue with Cons(1) W show ?thesis by auto next case BothFalse with Cons(1) W show ?thesis by auto qed qed simp lemma (in IrrFinGroupRepresentation) iso_to_GIrrRep_rep : "\U\ActingGroup.GIrrRep_repset. isomorphic (*) U" using zero_isomorphic_to_FG_zero ActingGroup.GIrrRep_repset_def good_char ActingGroup.FG_constituents_irr ActingGroup.FG_consitutents_n0 ActingGroup.FG_constituents_constituents iso_FG_constituent iso_in_list_imp_iso_in_remisodups ActingGroup.GIrrRep_repset_def by (cases "V = 0") auto theorem (in Group) iso_class_reps : defines "GIRRS \ GIrrRep_repset :: ('f::field,'g) aezfun set set" assumes "of_nat (card G) \ (0::'f)" shows "finite GIRRS" "\U\GIRRS. IrrFinGroupRepresentation G (*) U" "\U W. \ U \ GIRRS; W \ GIRRS; U \ W \ \ \ (FGModule.isomorphic G (*) U (*) W)" "\fgsmult V. IrrFinGroupRepresentation G fgsmult V \ \U\GIRRS. FGModule.isomorphic G fgsmult V (*) U" using assms finite_GIrrRep_repset all_irr_GIrrRep_repset isodistinct_GIrrRep_repset IrrFinGroupRepresentation.iso_to_GIrrRep_rep by auto subsection \Induced representations\ subsubsection \Locale and basic facts\ locale InducedFinGroupRepresentation = Supgroup?: Group H + BaseRep?: FinGroupRepresentation G smult V + induced_smult?: aezfun_scalar_mult rrsmult for H :: "'g::group_add set" and G :: "'g set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixl "\" 70) and V :: "'v set" and rrsmult :: "('f,'g) aezfun \ (('f,'g) aezfun \'v) \ (('f,'g) aezfun \'v)" (infixl "\" 70) + fixes FH :: "('f, 'g) aezfun set" and indV :: "(('f, 'g) aezfun \ 'v) set" defines FH : "FH \ Supgroup.group_ring" and indV : "indV \ BaseRep.indspace G FH V" assumes rrsmult : "rrsmult = Ring1.rightreg_scalar_mult FH" and good_ordSupgrp: "of_nat (card H) \ (0::'f)" (* this implies H is finite *) and Subgroup : "Supgroup.Subgroup G" (* G is a subgroup of H *) sublocale InducedFinGroupRepresentation < InducedFHModule using FH indV rrsmult Subgroup by unfold_locales fast context InducedFinGroupRepresentation begin abbreviation ordH :: 'f where "ordH \ of_nat (card H)" abbreviation "is_Vfbasis \ fbasis_for V" abbreviation "GRepHomSet \ FGModuleHomSet G smult V" abbreviation "HRepHom \ FGModuleHom H rrsmult indV" abbreviation "HRepHomSet \ FGModuleHomSet H rrsmult indV" lemma finiteSupgroup: "finite H" using good_ordSupgrp good_card_imp_finite by fast lemma FinGroupSupgroup : "FinGroup H" using finiteSupgroup Supgroup.FinGroupI by fast lemmas fVectorSpace = fVectorSpace lemmas FinDimVectorSpace = FinDimVectorSpace lemmas ex_rcoset_replist_hd0 = FinGroup.ex_rcoset_replist_hd0[OF FinGroupSupgroup Subgroup] end (* context InducedFinGroupRepresentation *) subsubsection \A basis for the induced space\ context InducedFinGroupRepresentation begin abbreviation "negHorbit_list \ induced_smult.negGorbit_list" lemmas ex_rcoset_replist = FinGroup.ex_rcoset_replist[OF FinGroupSupgroup Subgroup] lemmas length_negHorbit_list = induced_smult.length_negGorbit_list lemmas length_negHorbit_list_sublist = induced_smult.length_negGorbit_list_sublist lemmas negHorbit_list_indV = FGModule.negGorbit_list_V[OF FHModule_indspace] lemma flincomb_Horbit_induced_veclist_reduce : fixes vs :: "'v list" and hs :: "'g list" defines hfvss : "hfvss \ negHorbit_list hs induced_vector vs" assumes vs : "set vs \ V" and i: "i < length hs" and scalars : "list_all2 (\rs ms. length rs = length ms) css hfvss" and rcoset_reps : "Supgroup.is_rcoset_replist G hs" shows "((concat css) \\\ (concat hfvss)) (1 \\ (hs!i)) = (css!i) \\\ vs" proof- have mostly_zero: "\k j. k < length hs \ j < length hs \ ((css!j) \\\ (hfvss!j)) (1 \\ hs!k) = (if j = k then (css!k) \\\ vs else 0)" proof- fix k j assume k: "k < length hs" and j: "j < length hs" hence hsk_H: "hs!k \ H" and hsj_H: "hs!j \ H" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] by auto define LHS where "LHS = ((css!j) \\\ (hfvss!j)) (1 \\ hs!k)" with hfvss have "LHS = (\(r,m)\zip (css!j) (hfvss!j). (r \\ m) (1 \\ hs!k))" using length_negHorbit_list scalar_mult.lincomb_def[of induced_smult.fsmult] sum_list_prod_fun_apply by simp moreover have "\(r,m) \ set (zip (css!j) (hfvss!j)). (induced_smult.fsmult r m) (1 \\ hs!k) = r \\ m (1 \\ hs!k)" proof (rule prod_ballI) fix r m assume "(r,m) \ set (zip (css!j) (hfvss!j))" with k j vs hfvss show "(induced_smult.fsmult r m) (1 \\ hs!k) = r \\ m (1 \\ hs!k)" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] set_zip_rightD set_concat length_negHorbit_list[of hs induced_vector vs] nth_mem[of j hfvss] hsk_H induced_fsmult_conv_fsmult_1ddh[of m "hs!k" r] induced_vector_indV negHorbit_list_indV[of hs induced_vector vs] by force qed ultimately have "LHS = (\(r,v)\zip (css!j) vs. r \\ (induced_vector v) (1 \\ hs!k * (1 \\ - hs!j)))" using FH j hfvss induced_smult.negGorbit_list_def[of hs induced_vector vs] sum_list_prod_cong[of _ "\r m. (induced_smult.fsmult r m) (1 \\ hs!k)"] sum_list_prod_map2[of "\r m. r \\ m (1 \\ hs!k)" _ "Hmult (- hs!j)" "map induced_vector vs" ] sum_list_prod_map2[of "\r v. r \\ (Hmult (-hs!j) v) (1 \\ hs!k)"] induced_smult.GmultD hsk_H Supgroup.RG_aezdeltafun_closed[of "hs!k" "1::'f"] rrsmultD1[of "1 \\ (hs!k)"] by force moreover have "(1::'f) \\ hs!k * (1 \\ - hs!j) = 1 \\ (hs!k - hs!j)" using times_aezdeltafun_aezdeltafun[of "1::'f" "hs!k" 1 "-(hs!j)"] by (simp add: algebra_simps) ultimately have "LHS = (\(r,v)\zip (css!j) vs. r \\ (induced_vector v) (1 \\ (hs!k - hs!j)))" using sum_list_prod_map2 by simp moreover from FH vs have "\(r,v) \ set (zip (css!j) vs). r \\ (induced_vector v) (1 \\ (hs!k - hs!j)) = r \\ (FG_proj (1 \\ (hs!k - hs!j)) \ v)" using set_zip_rightD induced_vector_def hsk_H hsj_H Supgroup.diff_closed Supgroup.RG_aezdeltafun_closed[of _ "1::'f"] by fastforce ultimately have calc: "LHS = (\(r,v)\zip (css!j) vs. r \\ (FG_proj (1 \\ (hs!k - hs!j)) \ v) )" using sum_list_prod_cong by force show "LHS = (if j = k then (css!k) \\\ vs else 0)" proof (cases "j = k") case True with calc have "LHS = (\(r,v)\zip (css!k) vs. r \\ 1 \\ v)" using Group.zero_closed[OF GroupG] aezfun_setspan_proj_aezdeltafun[of G "1::'f"] BaseRep.fsmult_def by simp moreover from vs have "\(r,v) \ set (zip (css!k) vs). r \\ 1 \\ v = r \\ v" using set_zip_rightD BaseRep.fsmult_assoc by fastforce ultimately show ?thesis using True sum_list_prod_cong[of _ "\r v. r \\ 1 \\ v"] scalar_mult.lincomb_def[of BaseRep.fsmult] by simp next case False with k j calc have "LHS = (\(r,v)\zip (css!j) vs. r \\ (0 \ v))" using Supgroup.is_rcoset_replist_imp_nrelated_nth[OF Subgroup rcoset_reps] aezfun_setspan_proj_aezdeltafun[of G "1::'f"] by simp moreover from vs have "\(r,v) \ set (zip (css!j) vs). r \\ (0 \ v) = 0" using set_zip_rightD BaseRep.zero_smult by fastforce ultimately have "LHS = (\(r,v)\zip (css!j) vs. (0::'v))" using sum_list_prod_cong[of _ "\r v. r \\ (0 \ v)"] by simp hence "LHS = (\rv\zip (css!j) vs. case_prod (\r v. (0::'v)) rv)" by fastforce with False show ?thesis by simp qed qed define terms LHS where "terms = map (\a. case_prod (\cs hfvs. (cs \\\ hfvs) (1 \\ hs!i)) a) (zip css hfvss)" and "LHS = ((concat css) \\\ (concat hfvss)) (1 \\ (hs!i))" hence "LHS = sum_list terms" using scalars VectorSpace.lincomb_concat[OF fVectorSpace_indspace, of css hfvss] sum_list_prod_fun_apply by simp hence "LHS = (\j\{0..j\{0..\\ (hfvss!j)) (1 \\ hs!i)" by simp ultimately show "LHS = (css!i) \\\ vs" using terms_def sum.cong scalars list_all2_lengthD[of _ css hfvss] hfvss length_negHorbit_list[of hs induced_vector vs] i mostly_zero sum_single_nonzero[ of "{0..i j. ((css!j) \\\ (hfvss!j)) (1 \\ (hs!i))" "\i. (css!i) \\\ vs" ] by simp qed lemma indspace_fspanning_set : fixes vs :: "'v list" and hs :: "'g list" defines hfvss : "hfvss \ negHorbit_list hs induced_vector vs" assumes base_spset : "set vs \ V" "V = BaseRep.fSpan vs" and rcoset_reps : "Supgroup.is_rcoset_replist G hs" shows "indV = induced_smult.fSpan (concat hfvss)" proof (rule VectorSpace.SpanI[OF fVectorSpace_indspace]) from base_spset(1) hfvss show "set (concat hfvss) \ indV" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] induced_vector_indV negHorbit_list_indV by fast show "indV \ R_scalar_mult.RSpan UNIV (aezfun_scalar_mult.fsmult (\)) (concat hfvss)" proof fix f assume f: "f \ indV" hence "\h \ set hs. f (1 \\ h) \ V" using indV BaseRep.indspaceD_range by fast with base_spset(2) have coeffs_exist: "\h \ set hs. \ahs. length ahs = length vs \ f (1 \\ h) = ahs \\\ vs" using BaseRep.in_fSpan_obtain_same_length_coeffs by fast define f_coeffs where "f_coeffs h = (SOME ahs. length ahs = length vs \ f (1 \\ h) = ahs \\\ vs)" for h define ahss where "ahss = map f_coeffs hs" hence len_ahss: "length ahss = length hs" by simp with hfvss have len_zip_ahss_hfvss: "length (zip ahss hfvss) = length hs" using length_negHorbit_list[of hs induced_vector vs] by simp have len_ahss_el: "\ahs\set ahss. length ahs = length vs" proof fix ahs assume "ahs \ set ahss" from this ahss_def obtain h where h: "h \ set hs" "ahs = f_coeffs h" using set_map by auto from h(1) have "\ahs. length ahs = length vs \ f (1 \\ h) = ahs \\\ vs" using coeffs_exist by fast with h(2) show "length ahs = length vs" unfolding f_coeffs_def using someI_ex[of "\ahs. length ahs = length vs \ f (1 \\ h) = ahs \\\ vs"] by fast qed have "\(ahs,hfvs)\set (zip ahss hfvss). length ahs = length hfvs" proof fix x assume x: "x \ set (zip ahss hfvss)" show "case x of (ahs, hfvs) \ length ahs = length hfvs" proof fix ahs hfvs assume "x = (ahs,hfvs)" with x hfvss have "length ahs = length vs" "length hfvs = length vs" using set_zip_leftD[of ahs hfvs] len_ahss_el set_zip_rightD[of ahs hfvs] length_negHorbit_list_sublist[of _ hs induced_vector] by auto thus "length ahs = length hfvs" by simp qed qed with hfvss have list_all2_len_ahss_hfvss: "list_all2 (\rs ms. length rs = length ms) ahss hfvss" using len_ahss length_negHorbit_list[of hs induced_vector vs] list_all2I[of ahss hfvss] by auto define f' where "f' = (concat ahss) \\\ (concat hfvss)" have "f = f'" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] Supgroup.group_eq_subgrp_rcoset_un[OF Subgroup rcoset_reps] f proof (rule indspace_el_eq'[of hs]) from f'_def hfvss base_spset(1) show "f' \ indV" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] induced_vector_indV negHorbit_list_indV[of hs induced_vector vs] FGModule.flincomb_closed[OF FHModule_indspace] by fast have "\i. i f (1 \\ (hs!i)) = f' (1 \\ (hs!i))" proof- fix i assume i: "i < length hs" with f_coeffs_def have "f (1 \\ (hs!i)) = (f_coeffs (hs!i)) \\\ vs" using coeffs_exist someI_ex[of "\ahs. length ahs = length vs \ f (1 \\ hs!i) = ahs \\\ vs"] by auto moreover from i hfvss f'_def base_spset(1) rcoset_reps ahss_def have "f' (1 \\ (hs!i)) = (f_coeffs (hs!i)) \\\ vs" using list_all2_len_ahss_hfvss flincomb_Horbit_induced_veclist_reduce by simp ultimately show "f (1 \\ (hs!i)) = f' (1 \\ (hs!i))" by simp qed thus "\i\ (hs!i)) = f' (1 \\ (hs!i))" by fast qed with f'_def hfvss base_spset(1) show "f \ induced_smult.fSpan (concat hfvss)" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] induced_vector_indV negHorbit_list_indV[of hs induced_vector vs] VectorSpace.SpanI_lincomb_arb_len_coeffs[OF fVectorSpace_indspace] by fast qed qed lemma indspace_basis : fixes vs :: "'v list" and hs :: "'g list" defines hfvss : "hfvss \ negHorbit_list hs induced_vector vs" assumes base_spset : "BaseRep.fbasis_for V vs" and rcoset_reps : "Supgroup.is_rcoset_replist G hs" shows "fscalar_mult.basis_for induced_smult.fsmult indV (concat hfvss)" proof- from assms have 1: "set (concat hfvss) \ indV" and "indV = induced_smult.fSpan (concat hfvss)" using Supgroup.is_rcoset_replistD_set[OF rcoset_reps] induced_vector_indV negHorbit_list_indV[of hs induced_vector vs] indspace_fspanning_set[of vs hs] by auto moreover have "induced_smult.f_lin_independent (concat hfvss)" proof ( rule VectorSpace.lin_independentI_concat_all_scalars[OF fVectorSpace_indspace], rule 1 ) fix rss assume rss: "list_all2 (\xs ys. length xs = length ys) rss hfvss" "(concat rss) \\\ (concat hfvss) = 0" from rss(1) have len_rss_hfvsss: "length rss = length hfvss" using list_all2_lengthD by fast with hfvss have len_rss_hs: "length rss = length hs" using length_negHorbit_list by fastforce show "\rs\set rss. set rs \ 0" proof fix rs assume "rs \ set rss" from this obtain i where i: "i < length rss" "rs = rss!i" using in_set_conv_nth[of rs] by fast with hfvss rss(1) have "length rs = length vs" using list_all2_nthD len_rss_hfvsss in_set_conv_nth[of _ hfvss] length_negHorbit_list_sublist by fastforce moreover from hfvss rss i base_spset(1) rcoset_reps have "rs \\\ vs = 0" using len_rss_hs flincomb_Horbit_induced_veclist_reduce by force ultimately show "set rs \ 0" using base_spset flin_independentD_all_scalars by force qed qed ultimately show ?thesis by fast qed lemma induced_vector_decomp : fixes vs :: "'v list" and hs :: "'g list" and cs :: "'f list" defines hfvss : "hfvss \ negHorbit_list (0#hs) induced_vector vs" and extrazeros : "extrazeros \ replicate ((length hs)*(length vs)) 0" assumes base_spset : "BaseRep.fbasis_for V vs" and rcoset_reps : "Supgroup.is_rcoset_replist G (0#hs)" and cs : "length cs = length vs" and v : "v = cs \\\ vs" shows "induced_vector v = (cs@extrazeros) \\\ (concat hfvss)" proof- from hfvss base_spset have "hfvss = (map induced_vector vs) # (negHorbit_list hs induced_vector vs)" using induced_vector_indV FGModule.negGorbit_list_Cons0[OF FHModule_indspace] by fastforce with cs extrazeros base_spset rcoset_reps v show "induced_vector v = (cs@extrazeros) \\\ (concat hfvss)" using scalar_mult.lincomb_append[of cs _ induced_smult.fsmult] Supgroup.is_rcoset_replistD_set induced_vector_indV negHorbit_list_indV[of hs induced_vector vs] VectorSpace.lincomb_replicate0_left[OF fVectorSpace_indspace] FGModuleHom.VectorSpaceHom[OF hom_induced_vector] VectorSpaceHom.distrib_lincomb by fastforce qed end (* context InducedFinGroupRepresentation *) subsubsection \The induced space is a representation space\ context InducedFinGroupRepresentation begin lemma indspace_findim : "fscalar_mult.findim induced_smult.fsmult indV" proof- from BaseRep.findim obtain vs where vs: "set vs \ V" "V = BaseRep.fSpan vs" by fast obtain hs where hs: "Supgroup.is_rcoset_replist G hs" using ex_rcoset_replist by fast define hfvss where "hfvss = negHorbit_list hs induced_vector vs" with vs hs have "set (concat hfvss) \ indV" "indV = induced_smult.fSpan (concat hfvss)" using Supgroup.is_rcoset_replistD_set[OF hs] induced_vector_indV negHorbit_list_indV[of hs induced_vector vs] indspace_fspanning_set by auto thus ?thesis by fast qed theorem FinGroupRepresentation_indspace : "FinGroupRepresentation H rrsmult indV" using FHModule_indspace proof (rule FinGroupRepresentation.intro) from good_ordSupgrp show "FinGroupRepresentation_axioms H (\) indV" using indspace_findim by unfold_locales fast qed end (* context InducedFinGroupRepresentation *) subsection \Frobenius reciprocity\ subsubsection \Locale and basic facts\ text \There are a number of defined objects and lemmas concerning those objects leading up to the theorem of Frobenius reciprocity, so we create a locale to contain it all.\ locale FrobeniusReciprocity = GRep?: InducedFinGroupRepresentation H G smult V rrsmult + HRep?: FinGroupRepresentation H smult' W for H :: "'g::group_add set" and G :: "'g set" and smult :: "('f::field, 'g) aezfun \ 'v::ab_group_add \ 'v" (infixl "\" 70) and V :: "'v set" and rrsmult :: "('f,'g) aezfun \ (('f,'g) aezfun \ 'v) \ (('f,'g) aezfun \ 'v)" (infixl "\" 70) and smult' :: "('f, 'g) aezfun \ 'w::ab_group_add \ 'w" (infixr "\" 70) and W :: "'w set" begin abbreviation fsmult' :: "'f \ 'w \ 'w" (infixr "\\" 70) where "fsmult' \ HRep.fsmult" abbreviation flincomb' :: "'f list \ 'w list \ 'w" (infixr "\\\" 70) where "flincomb' \ HRep.flincomb" abbreviation Hmult' :: "'g \ 'w \ 'w" (infixr "*\" 70) where "Hmult' \ HRep.Gmult" definition Tsmult1 :: "'f \ ((('f, 'g) aezfun \ 'v)\'w) \ ((('f, 'g) aezfun \ 'v)\'w)" (infixr "\\" 70) where "Tsmult1 \ \a T. \f. a \\ (T f)" definition Tsmult2 :: "'f \ ('v\'w) \ ('v\'w)" (infixr "\\" 70) where "Tsmult2 \ \a T. \v. a \\ (T v)" lemma FHModuleW : "FGModule H (\) W" .. lemma FGModuleW: "FGModule G smult' W" using FHModuleW Subgroup HRep.restriction_to_subgroup_is_module by fast text \ In order to build an inverse for the required isomorphism of Hom-sets, we will need a basis for the induced @{term H}-space. \ definition Vfbasis :: "'v list" where "Vfbasis \ (SOME vs. is_Vfbasis vs)" lemma Vfbasis : "is_Vfbasis Vfbasis" using Vfbasis_def FinDimVectorSpace.basis_ex[OF GRep.FinDimVectorSpace] someI_ex by simp lemma Vfbasis_V : "set Vfbasis \ V" using Vfbasis by fast definition nonzero_H_rcoset_reps :: "'g list" where "nonzero_H_rcoset_reps \ (SOME hs. Group.is_rcoset_replist H G (0#hs))" definition H_rcoset_reps :: "'g list" where "H_rcoset_reps \ 0 # nonzero_H_rcoset_reps" lemma H_rcoset_reps : "Group.is_rcoset_replist H G H_rcoset_reps" using H_rcoset_reps_def nonzero_H_rcoset_reps_def GRep.ex_rcoset_replist_hd0 someI_ex by simp lemma H_rcoset_reps_H : "set H_rcoset_reps \ H" using H_rcoset_reps Group.is_rcoset_replistD_set[OF HRep.GroupG] by fast lemma nonzero_H_rcoset_reps_H : "set nonzero_H_rcoset_reps \ H" using H_rcoset_reps_H H_rcoset_reps_def by simp abbreviation negHorbit_homVfbasis :: "('v \ 'w) \ 'w list list" where "negHorbit_homVfbasis T \ HRep.negGorbit_list H_rcoset_reps T Vfbasis" lemma negHorbit_Hom_indVfbasis_W : "T ` V \ W \ set (concat (negHorbit_homVfbasis T)) \ W" using H_rcoset_reps_H Vfbasis_V HRep.negGorbit_list_V[of H_rcoset_reps T Vfbasis] by fast lemma negHorbit_HomSet_indVfbasis_W : "T \ GRepHomSet smult' W \ set (concat (negHorbit_homVfbasis T)) \ W" using FGModuleHomSetD_Im negHorbit_Hom_indVfbasis_W by fast definition indVfbasis :: "(('f, 'g) aezfun \ 'v) list list" where "indVfbasis \ GRep.negHorbit_list H_rcoset_reps induced_vector Vfbasis" lemma indVfbasis : "fscalar_mult.basis_for induced_smult.fsmult indV (concat indVfbasis)" using Vfbasis H_rcoset_reps indVfbasis_def indspace_basis[of Vfbasis H_rcoset_reps] by auto lemma indVfbasis_indV : "hfvs \ set indVfbasis \ set hfvs \ indV" using indVfbasis by auto end (* context FrobeniusReciprocity *) subsubsection \The required isomorphism of Hom-sets\ context FrobeniusReciprocity begin text \The following function will demonstrate the required isomorphism of Hom-sets (as vector spaces).\ definition \ :: "((('f, 'g) aezfun \ 'v) \ 'w) \ ('v \ 'w)" where "\ \ restrict0 (\T. T \ GRep.induced_vector) (HRepHomSet smult' W)" lemma \_im : "\ ` HRepHomSet (\) W \ GRepHomSet (\) W" proof (rule image_subsetI) fix T assume T: "T \ HRepHomSet (\) W" show "\ T \ GRepHomSet (\) W" proof (rule FGModuleHomSetI) from T have "FGModuleHom G rrsmult indV smult' T" using FGModuleHomSetD_FGModuleHom GRep.Subgroup FGModuleHom.restriction_to_subgroup_is_hom by fast thus "BaseRep.GRepHom (\) (\ T)" using T \_def GRep.hom_induced_vector GRep.induced_vector_indV FGModuleHom.FGModHom_composite_left by fastforce show "\ T ` V \ W" using T \_def GRep.induced_vector_indV FGModuleHomSetD_Im by fastforce qed qed end (* context FrobeniusReciprocity *) subsubsection \The inverse map of Hom-sets\ text \In this section we build an inverse for the required isomorphism, @{term "\"}.\ context FrobeniusReciprocity begin definition \_condition :: "('v \ 'w) \ ((('f, 'g) aezfun \ 'v) \ 'w) \ bool" where "\_condition T S \ VectorSpaceHom induced_smult.fsmult indV fsmult' S \ map (map S) indVfbasis = negHorbit_homVfbasis T" lemma inverse_im_exists' : assumes "T \ GRepHomSet (\) W" shows "\! S. VectorSpaceHom induced_smult.fsmult indV fsmult' S \ map S (concat indVfbasis) = concat (negHorbit_homVfbasis T)" proof ( rule VectorSpace.basis_im_defines_hom, rule fVectorSpace_indspace, rule HRep.fVectorSpace, rule indVfbasis ) from assms show "set (concat (negHorbit_homVfbasis T)) \ W" using negHorbit_HomSet_indVfbasis_W by fast show "length (concat (negHorbit_homVfbasis T)) = length (concat indVfbasis)" using length_concat_negGorbit_list indVfbasis_def induced_smult.length_concat_negGorbit_list[of H_rcoset_reps induced_vector] by simp qed lemma inverse_im_exists : assumes "T \ GRepHomSet (\) W" shows "\! S. \_condition T S" proof- have "\ S. \_condition T S" proof- from assms obtain S where S: "VectorSpaceHom induced_smult.fsmult indV fsmult' S" "map S (concat indVfbasis) = concat (negHorbit_homVfbasis T)" using inverse_im_exists' by fast from S(2) have "concat (map (map S) indVfbasis) = concat (negHorbit_homVfbasis T)" using map_concat[of S] by simp moreover have "list_all2 (\xs ys. length xs = length ys) (map (map S) indVfbasis) (negHorbit_homVfbasis T)" proof (rule iffD2[OF list_all2_iff], rule conjI) show "length (map (map S) indVfbasis) = length (negHorbit_homVfbasis T)" using indVfbasis_def induced_smult.length_negGorbit_list HRep.length_negGorbit_list[of H_rcoset_reps T] by auto show "\(xs,ys)\set (zip (map (map S) indVfbasis) (negHorbit_homVfbasis T)). length xs = length ys" proof (rule prod_ballI) fix xs ys assume xsys: "(xs,ys) \ set (zip (map (map S) indVfbasis) (negHorbit_homVfbasis T))" from this obtain zs where zs: "zs \ set indVfbasis" "xs = map S zs" using set_zip_leftD by fastforce with xsys show "length xs = length ys" using indVfbasis_def set_zip_rightD[of xs ys] HRep.length_negGorbit_list_sublist[of ys H_rcoset_reps T Vfbasis] induced_smult.length_negGorbit_list_sublist by simp qed qed ultimately have "map (map S) indVfbasis = negHorbit_homVfbasis T" using concat_eq[of "map (map S) indVfbasis"] by fast with S(1) show ?thesis using \_condition_def by fast qed moreover have "\S U. \_condition T S \ \_condition T U \ S = U" proof- fix S U assume "\_condition T S" "\_condition T U" hence "VectorSpaceHom induced_smult.fsmult indV fsmult' S" "map S (concat indVfbasis) = concat (negHorbit_homVfbasis T)" "VectorSpaceHom induced_smult.fsmult indV fsmult' U" "map U (concat indVfbasis) = concat (negHorbit_homVfbasis T)" using \_condition_def map_concat[of S] map_concat[of U] by auto with assms show "S = U" using inverse_im_exists' by fast qed ultimately show ?thesis by fast qed definition \ :: "('v \ 'w) \ ((('f, 'g) aezfun \ 'v) \ 'w)" where "\ \ restrict0 (\T. THE S. \_condition T S) (GRepHomSet (\) W)" lemma \D : "T \ GRepHomSet (\) W \ \_condition T (\ T)" using \_def inverse_im_exists[of T] theI'[of "\S. \_condition T S"] by simp lemma \D_VectorSpaceHom : "T \ GRepHomSet (\) W \ VectorSpaceHom induced_smult.fsmult indV fsmult' (\ T)" using \D \_condition_def by fast lemma \D_im : "T \ GRepHomSet (\) W \ map (map (\ T)) indVfbasis = aezfun_scalar_mult.negGorbit_list (\) H_rcoset_reps T Vfbasis" using \D \_condition_def by fast lemma \D_im_single : assumes "T \ GRepHomSet (\) W" "h \ set H_rcoset_reps" "v \ set Vfbasis" shows "\ T ((- h) *\ (induced_vector v)) = (-h) *\ (T v)" proof- from assms(2,3) obtain i j where i: "i < length H_rcoset_reps" "h = H_rcoset_reps!i" and j: "j < length Vfbasis" "v = Vfbasis!j" using set_conv_nth[of H_rcoset_reps] set_conv_nth[of Vfbasis] by auto moreover hence "map (map (\ T)) indVfbasis !i !j = \ T ((-h) *\ (induced_vector v))" using indVfbasis_def aezfun_scalar_mult.length_negGorbit_list[ of rrsmult H_rcoset_reps induced_vector ] aezfun_scalar_mult.negGorbit_list_nth[ of i H_rcoset_reps rrsmult induced_vector ] by auto ultimately show ?thesis using assms(1) HRep.negGorbit_list_nth[of i H_rcoset_reps T] \D_im by simp qed lemma \T_W : assumes "T \ GRepHomSet (\) W" shows "\ T ` indV \ W" proof (rule image_subsetI) from assms have T: "VectorSpaceHom induced_smult.fsmult indV fsmult' (\ T)" using \D_VectorSpaceHom by fast fix f assume "f \ indV" from this obtain cs where cs:"length cs = length (concat indVfbasis)" "f = cs \\\ (concat indVfbasis)" using indVfbasis scalar_mult.in_Span_obtain_same_length_coeffs by fast from cs(1) obtain css where css: "cs = concat css" "list_all2 (\xs ys. length xs = length ys) css indVfbasis" using match_concat by fast from assms cs(2) css have "\ T f = \ T (\(cs,hfvs)\zip css indVfbasis. cs \\\ hfvs)" using VectorSpace.lincomb_concat[OF fVectorSpace_indspace] by simp also have "\ = (\(cs,hfvs)\zip css indVfbasis. \ T (cs \\\ hfvs))" using set_zip_rightD[of _ _ css indVfbasis] indVfbasis_indV VectorSpace.lincomb_closed[OF GRep.fVectorSpace_indspace] VectorSpaceHom.im_sum_list_prod[OF T] by force finally have "\ T f = (\(cs,\Thfvs)\zip css (map (map (\ T)) indVfbasis). cs \\\ \Thfvs)" using set_zip_rightD[of _ _ css indVfbasis] indVfbasis_indV VectorSpaceHom.distrib_lincomb[OF T] sum_list_prod_cong[of "zip css indVfbasis" "\cs hfvs. \ T (cs \\\ hfvs)" "\cs hfvs. cs \\\ (map (\ T) hfvs)" ] sum_list_prod_map2[of "\cs \Thfvs. cs \\\ \Thfvs" css "map (\ T)"] by fastforce moreover from css(2) have "list_all2 (\xs ys. length xs = length ys) css (map (map (\ T)) indVfbasis)" using list_all2_iff[of _ css indVfbasis] set_zip_map2 list_all2_iff[of _ css "map (map (\ T)) indVfbasis"] by force ultimately have "\ T f = (concat css) \\\ (concat (negHorbit_homVfbasis T))" using HRep.flincomb_concat map_concat[of "\ T"] \D_im[OF assms] by simp thus "\ T f \ W" using assms negHorbit_HomSet_indVfbasis_W HRep.flincomb_closed by simp qed lemma \T_Hmap_on_indVfbasis : assumes "T \ GRepHomSet (\) W" shows "\x f. x \ H \ f \ set (concat indVfbasis) \ \ T (x *\ f) = x *\ (\ T f)" proof- fix x f assume x: "x \ H" and f: "f \ set (concat indVfbasis)" from f obtain i where i: "i < length indVfbasis" "f \ set (indVfbasis!i)" using set_concat set_conv_nth[of indVfbasis] by auto from i(1) have i': "i < length H_rcoset_reps" using indVfbasis_def aezfun_scalar_mult.length_negGorbit_list[ of rrsmult H_rcoset_reps induced_vector ] by simp define hi where "hi = H_rcoset_reps!i" with i' have hi_H: "hi \ H" using set_conv_nth H_rcoset_reps_H by fast from hi_def i(2) have "f \ set (map (Hmult (-hi) \ induced_vector) Vfbasis)" using indVfbasis_def i' aezfun_scalar_mult.negGorbit_list_nth[ of i H_rcoset_reps rrsmult induced_vector ] by simp from this obtain v where v: "v \ set Vfbasis" "f = (-hi) *\ (induced_vector v)" by auto from v(1) have v_V: "v \ V" and Tv_W: "T v \ W" using Vfbasis_V FGModuleHomSetD_Im[OF assms] by auto from x have "hi - x \ H" using hi_H Supgroup.diff_closed by fast from this obtain j where j: "j < length H_rcoset_reps" "hi - x \ G + {H_rcoset_reps!j}" using set_conv_nth[of H_rcoset_reps] H_rcoset_reps Group.group_eq_subgrp_rcoset_un[OF HRep.GroupG Subgroup H_rcoset_reps] by force from j(1) have j': "j < length indVfbasis" using indVfbasis_def aezfun_scalar_mult.length_negGorbit_list[ of rrsmult H_rcoset_reps induced_vector ] by simp define hj where "hj = H_rcoset_reps!j" with j(1) have hj_H: "hj \ H" using set_conv_nth H_rcoset_reps_H by fast from hj_def j(2) obtain g where g: "g \ G" "hi - x = g + hj" unfolding set_plus_def by fast from g(2) have x_hi: "x - hi = - hj + - g" using minus_diff_eq[of hi x] minus_add[of g] by simp from g(1) have "-g *\ v \ V" using v_V ActingGroup.neg_closed BaseRep.Gmult_closed by fast from this obtain cs where cs: "length cs = length Vfbasis" "-g *\ v = cs \\\ Vfbasis" using Vfbasis VectorSpace.in_Span_obtain_same_length_coeffs[OF GRep.fVectorSpace] by fast from v(2) x have "\ T (x *\ f) = \ T ((x-hi) *\ (induced_vector v))" using hi_H Supgroup.neg_closed v_V induced_vector_indV FGModule.Gmult_assoc[OF GRep.FHModule_indspace] by (simp add: algebra_simps) also from g(1) have "\ = \ T ((-hj) *\ (induced_vector (-g *\ v)))" using x_hi hj_H Subgroup Supgroup.neg_closed v_V induced_vector_indV FGModule.Gmult_assoc[OF GRep.FHModule_indspace] ActingGroup.neg_closed FGModuleHom.G_map[OF hom_induced_vector] by auto also from cs(2) hj_def j(1) have "\ = \ T (cs \\\ (indVfbasis!j))" using hj_H Vfbasis_V FGModuleHom.distrib_flincomb[OF hom_induced_vector] indVfbasis_def Supgroup.neg_closed[of hj] induced_vector_indV FGModule.Gmult_flincomb_comm[ OF GRep.FHModule_indspace, of "-hj" "map induced_vector Vfbasis" ] aezfun_scalar_mult.negGorbit_list_nth[ of j H_rcoset_reps rrsmult induced_vector ] by fastforce also have "\ = cs \\\ ((map (map (\ T)) indVfbasis)!j)" using \D_VectorSpaceHom[OF assms] indVfbasis_indV j' set_conv_nth VectorSpaceHom.distrib_lincomb[of induced_smult.fsmult indV fsmult'] by simp also from j(1) hj_def have "\ = (- hj) *\ cs \\\ (map T Vfbasis)" using \D_im[OF assms] aezfun_scalar_mult.negGorbit_list_nth[of j H_rcoset_reps smult' T] hj_H Group.neg_closed[OF HRep.GroupG] Vfbasis_V FGModuleHomSetD_Im[OF assms] HRep.Gmult_flincomb_comm[of "- hj" "map T Vfbasis"] by fastforce also from cs(2) g(1) have "\ = (- hj) *\ (-g) *\ (T v)" using v_V FGModuleHomSetD_FGModuleHom[OF assms] Vfbasis_V FGModuleHom.distrib_flincomb[of G smult V smult'] ActingGroup.neg_closed FGModuleHom.G_map[of G smult V smult' T "-g" v] by auto also from g(1) v(1) have "\ = (x - hi) *\ (T v)" using FGModuleHomSetD_FGModuleHom[OF assms] Vfbasis_V Supgroup.neg_closed hj_H Subgroup FGModuleHomSetD_Im[OF assms] HRep.Gmult_assoc[of "-hj" "-g" "T v"] x_hi by auto also from x(1) have "\ = x *\ (- hi) *\ (T v)" using hi_H Supgroup.neg_closed Tv_W HRep.Gmult_assoc by (simp add: algebra_simps) finally show "\ T (x *\ f) = x *\ (\ T f)" using assms(1) v hi_def i' set_conv_nth[of H_rcoset_reps] \D_im_single by fastforce qed lemma \T_hom : assumes "T \ GRepHomSet (\) W" shows "HRepHom (\) (\ T)" using indVfbasis \D_VectorSpaceHom[OF assms] FHModuleW proof ( rule FGModule.VecHom_GMap_on_fbasis_is_FGModuleHom[ OF GRep.FHModule_indspace ] ) show "\ T ` indV \ W" using indVfbasis \T_W[OF assms] by fast show "\g v. g \ H \ v \ set (concat indVfbasis) \ \ T (g *\ v) = g *\ \ T v" using \T_Hmap_on_indVfbasis[OF assms] by fast qed lemma \_im : "\ ` GRepHomSet (\) W \ HRepHomSet (\) W" using \T_W \T_hom FGModuleHomSetI by fastforce end (* context FrobeniusReciprocity *) subsubsection \Demonstration of bijectivity\ text \Now we demonstrate that @{term "\"} is bijective via the inverse @{term "\"}.\ context FrobeniusReciprocity begin lemma \\ : assumes "T \ GRepHomSet smult' W" shows "(\ \ \) T = T" proof fix v show "(\ \ \) T v = T v" proof (cases "v \ V") case True from this obtain cs where cs: "length cs = length Vfbasis" "v = cs \\\ Vfbasis" using Vfbasis VectorSpace.in_Span_obtain_same_length_coeffs[OF GRep.fVectorSpace] by fast define extrazeros where "extrazeros = replicate ((length nonzero_H_rcoset_reps)*(length Vfbasis)) (0::'f)" with cs have "GRep.induced_vector v = (cs@extrazeros) \\\ (concat indVfbasis)" using H_rcoset_reps induced_vector_decomp[OF Vfbasis] unfolding H_rcoset_reps_def indVfbasis_def by auto with assms have "(\ \ \) T v = (cs@extrazeros) \\\ (map (\ T) (concat indVfbasis))" using \_im \_def indVfbasis VectorSpaceHom.distrib_lincomb[OF \D_VectorSpaceHom] by auto also have "\ = (cs@extrazeros) \\\ (map T Vfbasis @ concat (HRep.negGorbit_list nonzero_H_rcoset_reps T Vfbasis))" using map_concat[of "\ T"] \D_im[OF assms] H_rcoset_reps_def FGModuleHomSetD_Im[OF assms] Vfbasis_V HRep.negGorbit_list_Cons0 by fastforce also from cs(1) have "\ = cs \\\ (map T Vfbasis) + extrazeros \\\ (concat (HRep.negGorbit_list nonzero_H_rcoset_reps T Vfbasis))" using scalar_mult.lincomb_append[of cs _ fsmult'] by simp also have "\ = cs \\\ (map T Vfbasis)" using nonzero_H_rcoset_reps_H Vfbasis FGModuleHomSetD_Im[OF assms] HRep.negGorbit_list_V VectorSpace.lincomb_replicate0_left[OF HRep.fVectorSpace] unfolding extrazeros_def by force also from cs(2) have "\ = T v" using FGModuleHomSetD_FGModuleHom[OF assms] FGModuleHom.VectorSpaceHom Vfbasis VectorSpaceHom.distrib_lincomb[of "aezfun_scalar_mult.fsmult smult"] by fastforce finally show ?thesis by fast next case False with assms show ?thesis using \_im \_def GRep.induced_vector_def \D_VectorSpaceHom VectorSpaceHom.im_zero FGModuleHomSetD_FGModuleHom[of T G smult V] FGModuleHom.supp suppI_contra by fastforce qed qed lemma \_inverse_im : "\ ` HRepHomSet (\) W \ GRepHomSet (\) W" using \\ \_im by force lemma bij_\ : "bij_betw \ (HRepHomSet (\) W) (GRepHomSet (\) W)" unfolding bij_betw_def proof have "\ S T. \ S \ HRepHomSet (\) W; T \ HRepHomSet (\) W; \ S = \ T \ \ S = T" proof (rule VectorSpaceHom.same_image_on_spanset_imp_same_hom) fix S T assume ST: "S \ HRepHomSet (\) W" "T \ HRepHomSet (\) W" "\ S = \ T" from ST(1,2) have ST': "HRepHom smult' S" "HRepHom smult' T" using FGModuleHomSetD_FGModuleHom[of _ H rrsmult] by auto from ST' show "VectorSpaceHom induced_smult.fsmult indV fsmult' S" "VectorSpaceHom induced_smult.fsmult indV fsmult' T" using FGModuleHom.VectorSpaceHom[of H rrsmult indV smult'] by auto show "indV = induced_smult.fSpan (concat indVfbasis)" "set (concat indVfbasis) \ indV" using indVfbasis by auto show "\f\set (concat indVfbasis). S f = T f" proof fix f assume "f \ set (concat indVfbasis)" from this obtain hfvs where hfvs: "hfvs \ set indVfbasis" "f \ set hfvs" using set_concat by fast from hfvs(1) obtain h where h: "h \ set H_rcoset_reps" "hfvs = map (Hmult (-h) \ induced_vector) Vfbasis" using indVfbasis_def induced_smult.negGorbit_list_def[of H_rcoset_reps induced_vector] by auto from hfvs(2) h(2) obtain v where v: "v \ set Vfbasis" "f = (-h) *\ (induced_vector v)" by auto from v h(1) ST(1) have "S f = (-h) *\ (\ S v)" using ST'(1) H_rcoset_reps_H Group.neg_closed[OF HRep.GroupG] GRep.induced_vector_indV Vfbasis_V \_def FGModuleHom.G_map by fastforce moreover from v h(1) ST(2) have "T f = (-h) *\ (\ T v)" using ST'(2) H_rcoset_reps_H Group.neg_closed[OF HRep.GroupG] GRep.induced_vector_indV Vfbasis_V \_def FGModuleHom.G_map by fastforce ultimately show "S f = T f" using ST(3) by simp qed qed thus "inj_on \ (HRepHomSet (\) W)" unfolding inj_on_def by fast show "\ ` HRepHomSet (\) W = GRepHomSet (\) W" using \_im \_inverse_im by fast qed end (* context FrobeniusReciprocity *) subsubsection \The theorem\ text \ Finally we demonstrate that @{term "\"} is an isomorphism of vector spaces between the two hom-sets, leading to Frobenius reciprocity. \ context FrobeniusReciprocity begin lemma VectorSpaceIso_\ : "VectorSpaceIso Tsmult1 (HRepHomSet (\) W) Tsmult2 \ (GRepHomSet (\) W)" proof (rule VectorSpaceIso.intro, rule VectorSpace.VectorSpaceHomI_fromaxioms) from Tsmult1_def show "VectorSpace Tsmult1 (HRepHomSet (\) W)" using FHModule_indspace FHModuleW FGModule.VectorSpace_FGModuleHomSet by simp from \_def show "supp \ \ HRepHomSet (\) W" using suppD_contra[of \] by fastforce have "bij_betw \ (HRepHomSet (\) W) (GRepHomSet (\) W)" using bij_\ by fast thus "VectorSpaceIso_axioms (HRepHomSet (\) W) \ (GRepHomSet (\) W)" by unfold_locales next fix S T assume "S \ HRepHomSet (\) W" "T \ HRepHomSet (\) W" thus "\ (S + T) = \ S + \ T" using \_def Group.add_closed FGModule.Group_FGModuleHomSet[OF FHModule_indspace FHModuleW] by auto next fix a T assume T: "T \ HRepHomSet (\) W" moreover with Tsmult1_def have aT: "a \\ T \ HRepHomSet (\) W" using FGModule.VectorSpace_FGModuleHomSet[ OF FHModule_indspace FHModuleW ] VectorSpace.smult_closed by simp ultimately show "\ (a \\ T) = a \\ (\ T)" using \_def Tsmult1_def Tsmult2_def by auto qed theorem FrobeniusReciprocity : "VectorSpace.isomorphic Tsmult1 (HRepHomSet smult' W) Tsmult2 (GRepHomSet smult' W)" using VectorSpaceIso_\ by fast end (* context FrobeniusReciprocity *) end (* theory *) diff --git a/thys/Solidity/Reentrancy.thy b/thys/Solidity/Reentrancy.thy --- a/thys/Solidity/Reentrancy.thy +++ b/thys/Solidity/Reentrancy.thy @@ -1,3308 +1,3310 @@ section\Reentrancy\ text \ In the following we use our semantics to verify a contract implementing a simple token. The contract is defined by definition \<^emph>\victim\ and consist of one state variable and two methods: \<^item> The state variable "balance" is a mapping which assigns a balance to each address. \<^item> Method "deposit" allows to send money to the contract which is then added to the sender's balance. \<^item> Method "withdraw" allows to withdraw the callers balance. \ text \ We then verify that the following invariant (defined by \<^emph>\INV\) is preserved by both methods: The difference between \<^item> the contracts own account-balance and \<^item> the sum of all the balances kept in the contracts state variable is larger than a certain threshold. \ text \ There are two things to note here: First, Solidity implicitly triggers the call of a so-called fallback method whenever we transfer money to a contract. In particular if another contract calls "withdraw", this triggers an implict call to the callee's fallback method. This functionality was exploited in the infamous DAO attack which we demonstrate it in terms of an example later on. Since we do not know all potential contracts which call "withdraw", we need to verify our invariant for all possible Solidity programs. Thus, the core result here is a lemma which shows that the invariant is preserved by every Solidity program which is not executed in the context of our own contract. For our own methods we show that the invariant holds after executing it. Since our own program as well as the unknown program may depend on each other both properties are encoded in a single lemma (\<^emph>\secure\) which is then proved by induction over all statements. The final result is then given in terms of two corollaries for the corresponding methods of our contract. \ text \ The second thing to note is that we were not able to verify that the difference is indeed constant. During verification it turned out that this is not the case since in the fallback method a contract could just send us additional money withouth calling "deposit". In such a case the difference would change. In particular it would grow. However, we were able to verify that the difference does never shrink which is what we actually want to ensure. \ theory Reentrancy imports Solidity_Evaluator begin subsection\Example of Re-entrancy\ (* The following value can take several minutes to process. *) value "eval 1000 stmt (COMP (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''deposit'') [] (UINT 256 10)) (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] (UINT 256 0))) (STR ''Attacker'') (STR '''') (STR ''0'') [(STR ''Victim'', STR ''100''), (STR ''Attacker'', STR ''100'')] [ (STR ''Attacker'', [], ITE (LESS (BALANCE THIS) (UINT 256 125)) (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] (UINT 256 0)) SKIP), (STR ''Victim'', [ (STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], ASSIGN (Ref (STR ''balance'') [SENDER]) VALUE, None)), (STR ''withdraw'', Method ([], ITE (LESS (UINT 256 0) (LVAL (Ref (STR ''balance'') [SENDER]))) (COMP (TRANSFER SENDER (LVAL (Ref (STR ''balance'') [SENDER]))) (ASSIGN (Ref (STR ''balance'') [SENDER]) (UINT 256 0))) SKIP , None))], SKIP) ] []" subsection\Definition of Contract\ abbreviation myrexp::L where "myrexp \ Ref (STR ''balance'') [SENDER]" abbreviation mylval::E where "mylval \ LVAL myrexp" abbreviation assign::S where "assign \ ASSIGN (Ref (STR ''balance'') [SENDER]) (UINT 256 0)" abbreviation transfer::S where "transfer \ TRANSFER SENDER (LVAL (Id (STR ''bal'')))" abbreviation comp::S where "comp \ COMP assign transfer" abbreviation keep::S where "keep \ BLOCK ((STR ''bal'', Value (TUInt 256)), Some mylval) comp" abbreviation deposit::S where "deposit \ ASSIGN (Ref (STR ''balance'') [SENDER]) (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE)" definition victim::"(Identifier, Member) fmap" where "victim \ fmap_of_list [ (STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], deposit, None)), (STR ''withdraw'', Method ([], keep, None))]" subsection\Definition of Invariant\ abbreviation "SUMM s \ \(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x. ReadL\<^sub>i\<^sub>n\<^sub>t x" abbreviation "POS s \ \ad x. fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" abbreviation "INV st s val bal \ fmlookup (storage st) (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - val \ bal \ bal \ 0" definition frame_def: "frame bal st \ (\s. INV st s (SUMM s) bal \ POS s)" subsection\Verification\ lemma conj3: "P \ Q \ R \ P \ (Q \ R)" by simp lemma fmfinite: "finite ({(ad, x). fmlookup y ad = Some x})" proof - have "{(ad, x). fmlookup y ad = Some x} \ dom (fmlookup y) \ ran (fmlookup y)" proof fix x assume "x \ {(ad, x). fmlookup y ad = Some x}" then have "fmlookup y (fst x) = Some (snd x)" by auto then have "fst x \ dom (fmlookup y)" and "snd x \ ran (fmlookup y)" using Map.ranI by (blast,metis) then show "x \ dom (fmlookup y) \ ran (fmlookup y)" by (simp add: mem_Times_iff) qed thus ?thesis by (simp add: finite_ran finite_subset) qed lemma fmlookup_finite: fixes f :: "'a \ 'a" and y :: "('a, 'b) fmap" assumes "inj_on (\(ad, x). (f ad, x)) {(ad, x). (fmlookup y \ f) ad = Some x}" shows "finite {(ad, x). (fmlookup y \ f) ad = Some x}" proof (cases rule: inj_on_finite[OF assms(1), of "{(ad, x)|ad x. (fmlookup y) ad = Some x}"]) case 1 then show ?case by auto next case 2 then have *: "finite {(ad, x) |ad x. fmlookup y ad = Some x}" using fmfinite[of y] by simp show ?case using finite_image_set[OF *, of "\(ad, x). (ad, x)"] by simp qed lemma balance_inj: "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''),x)) {(ad, x). (fmlookup y \ f) ad = Some x}" proof fix v v' assume asm1: "v \ {(ad, x). (fmlookup y \ f) ad = Some x}" and asm2: "v' \ {(ad, x). (fmlookup y \ f) ad = Some x}" and *:"(case v of (ad, x) \ (ad + (STR ''.'' + STR ''balance''),x)) = (case v' of (ad, x) \ (ad + (STR ''.'' + STR ''balance''),x))" then obtain ad ad' x x' where **: "v = (ad,x)" and ***: "v'=(ad',x')" by auto moreover from * ** *** have "ad + (STR ''.'' + STR ''balance'') = ad' + (STR ''.'' + STR ''balance'')" by simp with * ** have "ad = ad'" using String_Cancel[of ad "STR ''.'' + STR ''balance''" ad' ] by simp moreover from asm1 asm2 ** *** have "(fmlookup y \ f) ad = Some x" and "(fmlookup y \ f) ad' = Some x'" by auto with calculation(3) have "x=x'" by simp ultimately show "v=v'" by simp qed lemma transfer_frame: assumes "Accounts.transfer ad adv v (accounts st) = Some acc" and "frame bal st" and "ad \ STR ''Victim''" shows "frame bal (st\accounts := acc\)" proof (cases "adv = STR ''Victim''") case True define st' where "st' = st\accounts := acc, stack := emptyStore, memory := emptyStore\" from True assms(2) transfer_mono[OF assms(1)] have "(\s. fmlookup (storage st) (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add:frame_def) then have "(\s. fmlookup (storage st') (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add: st'_def) then show ?thesis using assms(2) by (auto simp add:st'_def frame_def) next case False define st' where "st' = st\accounts := acc, stack := emptyStore, memory := emptyStore\" from False assms(2) assms(3) transfer_eq[OF assms(1)] have "(\s. fmlookup (storage st) (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add:frame_def) then have "(\s. fmlookup (storage st') (STR ''Victim'') = Some s \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0)" by (auto simp add: st'_def) then show ?thesis using assms(2) by (auto simp add:st'_def frame_def) qed lemma decl_frame: assumes "frame bal st" and "decl a1 a2 a3 cp cd mem c env st = Normal (rv, st')" shows "frame bal st'" proof (cases a2) case (Value t) then show ?thesis proof (cases a3) case None with Value show ?thesis using assms by (auto simp add:frame_def) next case (Some a) show ?thesis proof (cases a) case (Pair a b) then show ?thesis proof (cases a) case (KValue v) then show ?thesis proof (cases b) case v2: (Value t') show ?thesis proof (cases "Valuetypes.convert t' t v") case None with Some Pair KValue Value v2 show ?thesis using assms by simp next case s2: (Some a) show ?thesis proof (cases a) case p2: (Pair a b) with Some Pair KValue Value v2 s2 show ?thesis using assms by (auto simp add:frame_def) qed qed next case (Calldata x2) with Some Pair KValue Value show ?thesis using assms by simp next case (Memory x3) with Some Pair KValue Value show ?thesis using assms by simp next case (Storage x4) with Some Pair KValue Value show ?thesis using assms by simp qed next case (KCDptr x2) with Some Pair Value show ?thesis using assms by simp next case (KMemptr x3) with Some Pair Value show ?thesis using assms by simp next case (KStoptr x4) with Some Pair Value show ?thesis using assms by simp qed qed qed next case (Calldata x2) then show ?thesis proof (cases cp) case True then show ?thesis proof (cases x2) case (MTArray x t) then show ?thesis proof (cases a3) case None with Calldata show ?thesis using assms by simp next case (Some a) show ?thesis proof (cases a) case (Pair a b) then show ?thesis proof (cases a) case (KValue x1) with Calldata Some Pair show ?thesis using assms by simp next case (KCDptr p) define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" obtain c' where c_def: "\x. (x, c') = allocate c" by simp show ?thesis proof (cases "cpm2m p l x t cd c'") case None with Calldata MTArray Some Pair KCDptr l_def c_def True show ?thesis using assms by simp next case s2: (Some a) with Calldata MTArray Some Pair KCDptr l_def c_def True show ?thesis using assms by (auto simp add:frame_def) qed next case (KMemptr p) define l where "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc c)" obtain c' where c_def: "\x. (x, c') = allocate c" by simp show ?thesis proof (cases "cpm2m p l x t mem c'") case None with Calldata MTArray Some Pair KMemptr l_def c_def True show ?thesis using assms by simp next case s2: (Some a) with Calldata MTArray Some Pair KMemptr l_def c_def True show ?thesis using assms by (auto simp add:frame_def) qed next case (KStoptr x4) with Calldata Some Pair show ?thesis using assms by simp qed qed qed next case (MTValue x2) with Calldata show ?thesis using assms by simp qed next case False with Calldata show ?thesis using assms by simp qed next case (Memory x3) then show ?thesis proof (cases x3) case (MTArray x t) then show ?thesis proof (cases a3) case None with Memory MTArray None show ?thesis using assms by (auto simp add:frame_def simp add:Let_def) next case (Some a) then show ?thesis proof (cases a) case (Pair a b) then show ?thesis proof (cases a) case (KValue x1) with Memory MTArray Some Pair show ?thesis using assms by simp next case (KCDptr p) define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" obtain m' where m'_def: "\x. (x, m') = allocate m" by simp then show ?thesis proof (cases "cpm2m p l x t cd m'") case None with Memory MTArray Some Pair KCDptr m_def l_def m'_def show ?thesis using assms by simp next case s2: (Some a) with Memory MTArray Some Pair KCDptr m_def l_def m'_def show ?thesis using assms by (auto simp add:frame_def) qed next case (KMemptr p) then show ?thesis proof (cases cp) case True define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" obtain m' where m'_def: "\x. (x, m') = allocate m" by simp then show ?thesis proof (cases "cpm2m p l x t mem m'") case None with Memory MTArray Some Pair KMemptr True m_def l_def m'_def show ?thesis using assms by simp next case s2: (Some a) with Memory MTArray Some Pair KMemptr True m_def l_def m'_def show ?thesis using assms by (auto simp add:frame_def) qed next case False with Memory MTArray Some Pair KMemptr show ?thesis using assms by (auto simp add:frame_def) qed next case (KStoptr p) then show ?thesis proof (cases b) case (Value x1) with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp next case (Calldata x2) with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp next case m2: (Memory x3) with Memory MTArray Some Pair KStoptr show ?thesis using assms by simp next case (Storage x4) then show ?thesis proof (cases x4) case (STArray x' t') define m l where "m = memory st" and "l = ShowL\<^sub>n\<^sub>a\<^sub>t (toploc m)" obtain m' where m'_def: "\x. (x, m') = allocate m" by simp from assms(2) Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def obtain s where *: "fmlookup (storage st) (address env) = Some s" using Let_def by (auto simp add: Let_def split:option.split_asm) then show ?thesis proof (cases "cps2m p l x' t' s m'") case None with Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def * show ?thesis using assms by simp next case s2: (Some a) with Memory MTArray Some Pair KStoptr Storage STArray m_def l_def m'_def * show ?thesis using assms by (auto simp add:frame_def) qed next case (STMap x21 x22) with Memory MTArray Some Pair KStoptr Storage show ?thesis using assms by simp next case (STValue x3) with Memory MTArray Some Pair KStoptr Storage show ?thesis using assms by simp qed qed qed qed qed next case (MTValue x2) with Memory show ?thesis using assms by simp qed next case (Storage x4) then show ?thesis proof (cases x4) case (STArray x t) then show ?thesis proof (cases a3) case None with Storage STArray show ?thesis using assms by simp next case (Some a) then show ?thesis proof (cases a) case (Pair a b) then show ?thesis proof (cases a) case (KValue x1) with Storage STArray Some Pair show ?thesis using assms by simp next case (KCDptr x2) with Storage STArray Some Pair show ?thesis using assms by simp next case (KMemptr x3) with Storage STArray Some Pair show ?thesis using assms by simp next case (KStoptr x4) with Storage STArray Some Pair show ?thesis using assms by (auto simp add:frame_def) qed qed qed next case (STMap t t') then show ?thesis proof (cases a3) case None with Storage STMap show ?thesis using assms by simp next case (Some a) then show ?thesis proof (cases a) case (Pair a b) then show ?thesis proof (cases a) case (KValue x1) with Storage STMap Some Pair show ?thesis using assms by simp next case (KCDptr x2) with Storage STMap Some Pair show ?thesis using assms by simp next case (KMemptr x3) with Storage STMap Some Pair show ?thesis using assms by simp next case (KStoptr x4) with Storage STMap Some Pair show ?thesis using assms by (auto simp add:frame_def) qed qed qed next case (STValue x3) with Storage show ?thesis using assms by simp qed qed context statement_with_gas begin lemma secureassign: assumes "stmt assign ep env cd st = Normal((), st')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "address env = (STR ''Victim'')" and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" and "accessStore x (stack st) = Some (KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s))" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - (SUMM s) \ bal" and "POS s" obtains s' where "fmlookup (storage st') (STR ''Victim'') = Some s'" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - (SUMM s' + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)) \ bal" and "accessStore x (stack st') = Some (KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s))" and "POS s'" proof - define st'' where "st'' = st\gas := gas st - costs assign ep env cd st\" define st''' where "st''' = st''\gas := gas st'' - costs\<^sub>e (UINT 256 0) ep env cd st''\" define st'''' where "st'''' = st'''\gas := gas st''' - costs\<^sub>e SENDER ep env cd st'''\" from assms(1) have c1: "gas st > costs assign ep env cd st" by (auto split:if_split_asm) have c2: "gas st'' > costs\<^sub>e (UINT 256 0) ep env cd st''" proof (rule ccontr) assume "\ costs\<^sub>e (UINT 256 0) ep env cd st'' < gas st''" with c1 show False using assms(1) st''_def st'''_def by auto qed hence *:"expr (UINT 256 0) ep env cd st'' = Normal ((KValue (createUInt 256 0),Value (TUInt 256)), st''')" using expr.simps(2)[of 256 0 ep env cd "st\gas := gas st - costs assign ep env cd st\"] st''_def st'''_def by simp moreover have "gas st''' > costs\<^sub>e SENDER ep env cd st'''" proof (rule ccontr) assume "\ costs\<^sub>e SENDER ep env cd st''' < gas st'''" with c1 c2 show False using assms(1,4) st''_def st'''_def by auto qed with assms(4) have **:"lexp (Ref (STR ''balance'') [SENDER]) ep env cd st''' = Normal ((LStoreloc ((sender env) + (STR ''.'' + STR ''balance'')), Storage (STValue (TUInt 256))), st'''')" using st''''_def by simp moreover have "Valuetypes.convert (TUInt 256) (TUInt 256) (ShowL\<^sub>i\<^sub>n\<^sub>t 0) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t 0, TUInt 256)" by simp moreover from * ** st''_def assms(1) obtain s'' where ***: "fmlookup (storage st'''') (address env) = Some s''" by (auto split:if_split_asm option.split_asm) ultimately have ****:"st' = st''''\storage := fmupd (STR ''Victim'') (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t 0) s'') (storage st)\" using c1 st''_def st'''_def st''''_def assms(1,3) by auto moreover define s' where s'_def: "s' = (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t 0) s'')" ultimately have "fmlookup (storage st') (STR ''Victim'') = Some s'" and *****:"fmlookup s' ((sender env) + (STR ''.'' + STR ''balance'')) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t 0)" by simp_all moreover have "SUMM s' + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) = SUMM s" proof - have s1: "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)" proof (cases "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = None") case True then have "accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s = ShowL\<^sub>i\<^sub>n\<^sub>t 0" by simp moreover have "{(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof fix x assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using True by auto qed next show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x }" proof fix x assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" using True by auto qed qed then have "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp ultimately show ?thesis using Read_ShowL_id by simp next case False then obtain val where val_def: "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = Some val" by auto have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp then have "finite {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s] by simp - then have sum1: "finite ({(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + then have sum1: "finite ({(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}"] + by fastforce moreover have sum2: "(sender env,val) \ {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp moreover from sum1 x1 val_def have "insert (sender env,val) {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] val_def by simp qed moreover have s2: "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" proof - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp then have "finite {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'] by simp - then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] + by fastforce moreover have sum2: "(sender env,ShowL\<^sub>i\<^sub>n\<^sub>t 0) \ {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp moreover from ***** have "insert (sender env,ShowL\<^sub>i\<^sub>n\<^sub>t 0) {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] Read_ShowL_id by simp qed moreover have s3: "(\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) =(\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" proof - have "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof fix xx assume "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" then obtain ad x where "xx = (ad,x)" and "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto moreover have "s''=s" using assms(2,3) s'_def *** st''''_def st'''_def st''_def by simp moreover from `ad \ sender env` have "ad + (STR ''.'' + STR ''balance'') \ (sender env) + (STR ''.'' + STR ''balance'')" using String_Cancel[where c="(STR ''.'' + STR ''balance'')"] by auto ultimately show "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using s'_def by simp qed next show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof fix xx assume "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" then obtain ad x where "xx = (ad,x)" and "fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto moreover have "s''=s" using assms(2,3) s'_def *** st''''_def st'''_def st''_def by simp moreover from `ad \ sender env` have "ad + (STR ''.'' + STR ''balance'') \ (sender env) + (STR ''.'' + STR ''balance'')" using String_Cancel[where c="(STR ''.'' + STR ''balance'')"] by auto ultimately show "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using s'_def by simp qed qed thus ?thesis by simp qed ultimately have "SUMM s' = SUMM s - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) " proof - from s2 have "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp also from s3 have "\ = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp also from s1 have "\ = SUMM s - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) " by simp finally show ?thesis . qed then show ?thesis by simp qed moreover have "POS s'" proof (rule allI[OF allI]) fix x xa show "fmlookup s' (x + (STR ''.'' + STR ''balance'')) = Some xa \ 0 \ (\xa\::int)" proof assume a1: "fmlookup s' (x + (STR ''.'' + STR ''balance'')) = Some xa" show "0 \ (\xa\::int)" proof (cases "x = sender env") case True then show ?thesis using s'_def a1 using Read_ShowL_id by auto next case False moreover have "s''=s" using assms(2,3) s'_def *** st''''_def st'''_def st''_def by simp ultimately have "fmlookup s (x + (STR ''.'' + STR ''balance'')) = Some xa" using s'_def a1 String_Cancel by (auto split:if_split_asm) then show ?thesis using assms(7) by simp qed qed qed moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" using **** st''_def st'''_def st''''_def by simp moreover from assms(5) have "accessStore x (stack st') = Some (KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) )"using **** st''_def st'''_def st''''_def by simp ultimately show ?thesis using assms(6) that by simp qed lemma securesender: assumes "expr SENDER ep env cd st = Normal((KValue v,t), st')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" obtains s' where "v = sender env" and "t = Value TAddr" and "fmlookup (storage st') (STR ''Victim'') = Some s'" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" using assms by (auto split:if_split_asm) lemma securessel: assumes "ssel type loc [] ep env cd st = Normal (x, st')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" obtains s' where "x = (loc, type)" and "fmlookup (storage st') (STR ''Victim'') = Some s'" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" using assms by auto lemma securessel2: assumes "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st = Normal ((loc, type), st')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" obtains s' where "loc = sender env + (STR ''.'' + STR ''balance'')" and "type = STValue (TUInt 256)" and "fmlookup (storage st') (STR ''Victim'') = Some s'" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" proof - from assms(1) obtain v t st'' st''' x where *: "expr SENDER ep env cd st = Normal ((KValue v, t), st'')" and **: "ssel (STValue (TUInt 256)) (hash (STR ''balance'') v) [] ep env cd st'' = Normal (x,st''')" and "st' = st'''" by (auto split:if_split_asm) moreover obtain s'' where "v =sender env" and "t = Value TAddr" and ***:"fmlookup (storage st'') (STR ''Victim'') = Some s''" and ****: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - SUMM s'' \ bal \ POS s''" using securesender[OF * assms(2,3)] by auto moreover obtain s''' where "x = (hash (STR ''balance'') v, STValue (TUInt 256))" and "fmlookup (storage st''') (STR ''Victim'') = Some s'''" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st''') (STR ''Victim'')) - SUMM s''' \ bal \ POS s'''" using securessel[OF ** *** ****] by auto ultimately show ?thesis using assms(1) that by simp qed lemma securerexp: assumes "rexp myrexp e\<^sub>p env cd st = Normal ((v, t), st')" and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ POS s" and "address env = STR ''Victim''" obtains s' where "fmlookup (storage st') (address env) = Some s'" and "v = KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s')" and "t = Value (TUInt 256)" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ POS s'" proof - from assms(1,2) obtain l' t'' st'' s where *: "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] e\<^sub>p env cd st = Normal ((l', STValue t''), st'')" and "fmlookup (storage st'') (address env) = Some s" and "v = KValue (accessStorage t'' l' s)" and "t = Value t''" and "st'=st''" by (simp split:if_split_asm option.split_asm) moreover obtain s'' where "fmlookup (storage st'') (STR ''Victim'') = Some s''" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - SUMM s'' \ bal \ POS s''" and "l'=sender env + (STR ''.'' + STR ''balance'')" and "t'' = TUInt 256" using securessel2[OF * assms(3,4)] by blast ultimately show ?thesis using assms(1,5) that by auto qed lemma securelval: assumes "expr mylval ep env cd st = Normal((v,t), st')" and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal \ bal \ 0 \ POS s" and "address env = STR ''Victim''" obtains s' where "fmlookup (storage st') (STR ''Victim'') = Some s'" and "v = KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s')" and "t = Value (TUInt 256)" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal \ bal \ 0 \ POS s'" proof - define st'' where "st'' = st\gas := gas st - costs\<^sub>e mylval ep env cd st\" with assms(3,4) have *: "fmlookup (storage st'') (STR ''Victim'') = Some s" and **: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - SUMM s \ bal \ POS s" by simp+ from assms(1) st''_def obtain v' t' st''' where ***: "rexp myrexp ep env cd st'' = Normal ((v, t), st''')" and "v' = v" and "t' = t" and "st''' = st'" by (simp split:if_split_asm) with securerexp[OF *** assms(2) * **] show ?thesis using assms(1,4,5) that by auto qed lemma plus_frame: assumes "expr (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st = Normal (kv, st')" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) < 2^256" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) \ 0" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal" and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" and "address env = (STR ''Victim'')" shows "kv = (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env))), Value (TUInt 256))" and "fmlookup (storage st') (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" proof - define st0 where "st0 = st\gas := gas st - costs\<^sub>e (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st\" define st1 where "st1 = st0\gas := gas st0 - costs\<^sub>e (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0\" define st2 where "st2 = st1\gas := gas st1 - costs\<^sub>e SENDER ep env cd st1\" define st3 where "st3 = st2\gas := gas st2 - costs\<^sub>e VALUE ep env cd st2\" from assms(1) obtain v1 t1 v2 t2 st'' st''' st'''' v t where *: "expr (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0 = Normal ((KValue v1, Value t1), st'')" and **: "expr VALUE ep env cd st'' = Normal ((KValue v2, Value t2), st''')" and ***: "add t1 t2 v1 v2 = Some (v,t)" and ****: "expr (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st = Normal ((KValue v, Value t), st'''')" using st0_def by (auto simp del: expr.simps simp add:expr.simps(11) split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm option.split_asm) moreover have "expr (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0 = Normal ((KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s), Value (TUInt 256)), st'')" and "st'' = st2" proof - from * obtain l' t' s'' where *****: "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st1 = Normal ((l', STValue t'), st'')" and ******: "fmlookup (storage st'') (address env) = Some s''" and "v1 = (accessStorage t' l' s'')" and "t' = t1" using st0_def st1_def assms(4,6) by (auto simp del: accessStorage.simps ssel.simps split:if_split_asm option.split_asm STypes.split_asm result.split_asm) moreover from ***** have "expr SENDER ep env cd st1 = Normal ((KValue (sender env), Value TAddr), st2)" using st2_def by (simp split:if_split_asm) with ***** have "st'' = st2" and "l' = sender env + (STR ''.'' + STR ''balance'')" and "t' = TUInt 256" by auto moreover from ****** `st'' = st2` have "s''=s" using st0_def st1_def st2_def assms(4,7) by auto ultimately show "expr (LVAL (Ref (STR ''balance'') [SENDER])) ep env cd st0 = Normal ((KValue (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s), Value (TUInt 256)), st'')" and "st'' = st2" using * by (auto split:if_split_asm) qed moreover from ** `st'' = st2` have "expr VALUE ep env cd st2 = Normal ((KValue (svalue env), Value (TUInt 256)), st3)" and "st''' = st3" using st1_def st3_def by (auto split:if_split_asm) moreover have "add (TUInt 256) (TUInt 256) (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) (svalue env) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)), TUInt 256)" (is "?LHS = ?RHS") proof - have "?LHS = Some (createUInt 256 (\(accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)\+ \(svalue env)\), TUInt 256)" using add_def olift.simps(2)[of "(+)" 256 256] by simp with assms(2,3) show "?LHS = ?RHS" by simp qed ultimately have "v= (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)))" and "t = TUInt 256" and "st' = st3" using st0_def assms(1) by (auto split:if_split_asm) with assms(1) **** have "kv = (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env))), Value (TUInt 256))" using st0_def by simp moreover from assms(4) st0_def st1_def st2_def st3_def `st' = st3` have "fmlookup (storage st') (STR ''Victim'') = Some s" by simp moreover from assms(5) st0_def st1_def st2_def st3_def `st' = st3` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s \ bal" by simp moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" using st0_def st1_def st2_def st3_def `st' = st3` by simp ultimately show "kv = (KValue (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env))), Value (TUInt 256))" and "fmlookup (storage st') (STR ''Victim'') = Some s" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" by auto qed lemma deposit_frame: assumes "stmt deposit ep env cd st = Normal ((), st')" and "fmlookup (storage st) (STR ''Victim'') = Some s" and "address env = (STR ''Victim'')" and "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim'')) - SUMM s \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) < 2^256" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env) \ 0" and "POS s" obtains s' where "fmlookup (storage st') (STR ''Victim'') = Some s'" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) - SUMM s' \ bal" and "POS s'" proof - define st0 where "st0 = st\gas := gas st - costs deposit ep env cd st\" from assms(1) st0_def obtain kv st'' where *: "expr (PLUS (LVAL (Ref (STR ''balance'') [SENDER])) VALUE) ep env cd st0 = Normal (kv, st'')" by (auto simp del: expr.simps split:if_split_asm result.split_asm) moreover have "fmlookup (storage st0) (STR ''Victim'') = Some s" using st0_def assms(2) by simp moreover from assms(5) have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st0) (STR ''Victim'')) - SUMM s \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue env)" using st0_def by simp ultimately have **: "kv = (KValue \(\accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s\::int) + \svalue env\\, Value (TUInt 256))" and st''_s:"fmlookup (storage st'') STR ''Victim'' = Some s" and ac: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st0) (STR ''Victim''))" using plus_frame[OF _ assms(6,7) _ _ assms(4,3), of ep cd st0 kv st''] by auto define v where "v= (\accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s\::int) + \svalue env\" moreover from * ** assms(1) st0_def obtain rl st''' where ***: "lexp (Ref (STR ''balance'') [SENDER]) ep env cd st'' = Normal (rl, st''')" by (auto simp del:expr.simps lexp.simps accessStorage.simps split:if_split_asm result.split_asm) moreover from *** assms have "rl = (LStoreloc ((sender env) + (STR ''.'' + STR ''balance'')), Storage (STValue (TUInt 256)))" and st'''_def: "st''' = st''\gas := gas st'' - costs\<^sub>e SENDER ep env cd st''\" proof - from *** assms(4) obtain l' t' where "fmlookup (denvalue env) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc (STR ''balance''))" and *:"ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st'' = Normal ((l',t'), st''')" and "rl = (LStoreloc l', Storage t')" by (auto simp del: ssel.simps split:if_split_asm option.split_asm result.split_asm) moreover from * have "ssel (STMap TAddr (STValue (TUInt 256))) (STR ''balance'') [SENDER] ep env cd st'' = Normal ((((sender env) + (STR ''.'' + STR ''balance'')), STValue (TUInt 256)), st''\gas := gas st'' - costs\<^sub>e SENDER ep env cd st''\)" by (simp split:if_split_asm) ultimately show "rl = (LStoreloc ((sender env) + (STR ''.'' + STR ''balance'')), Storage (STValue (TUInt 256)))" and st'''_def: "st''' = st''\gas := gas st'' - costs\<^sub>e SENDER ep env cd st''\" by auto qed moreover have "Valuetypes.convert (TUInt 256) (TUInt 256) (ShowL\<^sub>i\<^sub>n\<^sub>t v) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t v, TUInt 256)" by simp moreover from st''_s st'''_def have s'''_s: "fmlookup (storage st''') (STR ''Victim'') = Some s" by simp ultimately have ****:"st' = st'''\storage := fmupd (STR ''Victim'') (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t v) s) (storage st''')\" using assms(1) * ** st0_def assms(3) by (auto simp del:expr.simps lexp.simps accessStorage.simps split:if_split_asm) moreover define s' where "s' = (fmupd ((sender env) + (STR ''.'' + STR ''balance'')) (ShowL\<^sub>i\<^sub>n\<^sub>t v) s)" ultimately have "fmlookup (storage st') (STR ''Victim'') = Some s'" and *****:"fmlookup s' ((sender env) + (STR ''.'' + STR ''balance'')) = Some (ShowL\<^sub>i\<^sub>n\<^sub>t v)" by simp_all moreover have "SUMM s' = SUMM s + \svalue env\" proof - have s1: "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s)" proof (cases "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = None") case True then have "accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s = ShowL\<^sub>i\<^sub>n\<^sub>t 0" by simp moreover have "{(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof fix x assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" using True by auto qed next show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x }" proof fix x assume "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" then show "x \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" using True by auto qed qed then have "SUMM s = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" by simp ultimately show ?thesis using Read_ShowL_id by simp next case False then obtain val where val_def: "fmlookup s (sender env + (STR ''.'' + STR ''balance'')) = Some val" by auto have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp then have "finite {(ad, x). (fmlookup s \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s] by simp - then have sum1: "finite ({(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + then have sum1: "finite ({(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by fastforce moreover have sum2: "(sender env,val) \ {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp moreover from sum1 x1 val_def have "insert (sender env,val) {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] val_def by simp qed moreover have s2: "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + v" proof - have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp then have "finite {(ad, x). (fmlookup s' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'] by simp - then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto + then have sum1: "finite ({(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env})" using finite_subset[of "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by fastforce moreover have sum2: "(sender env,ShowL\<^sub>i\<^sub>n\<^sub>t v) \ {(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp moreover from ***** have "insert (sender env,ShowL\<^sub>i\<^sub>n\<^sub>t v) {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x}" by auto ultimately show ?thesis using sum.insert[OF sum1 sum2, of "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] Read_ShowL_id by simp qed moreover have s3: "(\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) =(\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x)" proof - have "{(ad,x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} = {(ad,x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof show "{(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof fix xx assume "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" then obtain ad x where "xx = (ad,x)" and "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto then have "fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x" using s'_def String_Cancel[of ad "(STR ''.'' + STR ''balance'')" "sender env"] by (simp split:if_split_asm) with `ad \ sender env` `xx = (ad,x)` show "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp qed next show "{(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env} \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" proof fix xx assume "xx \ {(ad, x). fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" then obtain ad x where "xx = (ad,x)" and "fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x" and "ad \ sender env" by auto then have "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x" using s'_def String_Cancel[of ad "(STR ''.'' + STR ''balance'')" "sender env"] by (auto split:if_split_asm) with `ad \ sender env` `xx = (ad,x)` show "xx \ {(ad, x). fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env}" by simp qed qed thus ?thesis by simp qed moreover from s'_def v_def have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s') = ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + \svalue env\" using Read_ShowL_id by (simp split:option.split_asm) ultimately have "SUMM s' = SUMM s + \svalue env\" proof - from s2 have "SUMM s' = (\(ad,x)|fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + v" by simp also from s3 have "\ = (\(ad,x)|fmlookup s (ad + (STR ''.'' + STR ''balance'')) = Some x \ ad \ sender env. ReadL\<^sub>i\<^sub>n\<^sub>t x) + v" by simp also from s1 have "\ = SUMM s - ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender env + (STR ''.'' + STR ''balance'')) s) + v" by simp finally show ?thesis using v_def by simp qed then show ?thesis by simp qed moreover have "POS s'" proof (rule allI[OF allI]) fix ad xa show "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some xa \ 0 \ (\xa\::int)" proof assume a1: "fmlookup s' (ad + (STR ''.'' + STR ''balance'')) = Some xa" show "0 \ (\xa\::int)" proof (cases "ad = sender env") case True then show ?thesis using s'_def assms(7) Read_ShowL_id a1 v_def by auto next case False then show ?thesis using s'_def assms(7,8) using Read_ShowL_id a1 v_def by (auto split:if_split_asm) qed qed qed moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st') (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st) (STR ''Victim''))" using **** ac st0_def st'''_def by simp ultimately show ?thesis using that assms(5) by simp qed lemma secure: "address ev1 \ (STR ''Victim'') \ fmlookup ep1 (STR ''Victim'') = Some (victim, SKIP) \ (\rv1 st1' bal. frame bal st1 \ msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal (rv1, st1') \ frame bal st1')" "address ev2 \ (STR ''Victim'') \ fmlookup ep2 (STR ''Victim'') = Some (victim, SKIP) \ (\rv2 st2' bal. frame bal st2 \ ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal (rv2, st2') \ frame bal st2')" "address ev5 \ (STR ''Victim'') \ fmlookup ep5 (STR ''Victim'') = Some (victim, SKIP) \ (\rv3 st5' bal. frame bal st5 \ lexp l5 ep5 ev5 cd5 st5 = Normal (rv3, st5') \ frame bal st5')" "address ev4 \ (STR ''Victim'') \ fmlookup ep4 (STR ''Victim'') = Some (victim, SKIP) \ (\rv4 st4' bal. frame bal st4 \ expr e4 ep4 ev4 cd4 st4 = Normal (rv4, st4') \ frame bal st4')" "address lev \ (STR ''Victim'') \ fmlookup lep (STR ''Victim'') = Some (victim, SKIP) \ (\ev cd st st' bal. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \ (frame bal lst0 \ frame bal st) \ (frame bal lst \ frame bal st') \ address lev0 = address ev \ sender lev0 = sender ev \ svalue lev0 = svalue ev)" "address ev3 \ (STR ''Victim'') \ fmlookup ep3 (STR ''Victim'') = Some (victim, SKIP) \ (\rv3 st3' bal. frame bal st3 \ rexp l3 ep3 ev3 cd3 st3 = Normal (rv3, st3') \ frame bal st3')" "(fmlookup ep6 (STR ''Victim'') = Some (victim, SKIP) \ (\st6'. stmt s6 ep6 ev6 cd6 st6 = Normal((), st6') \ ((address ev6 \ (STR ''Victim'') \ (\bal. frame bal st6 \ frame bal st6')) \ (address ev6 = (STR ''Victim'') \ (\s val bal x. s6 = transfer \ INV st6 s (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) bal \ POS s \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) \ accessStore x (stack st6) = Some (KValue val) \ sender ev6 \ address ev6 \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ (\s bal x. s6 = comp \ INV st6 s (SUMM s) bal \ POS s \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') \ accessStore x (stack st6) = Some (KValue (accessStorage (TUInt 256) (sender ev6 + (STR ''.'' + STR ''balance'')) s)) \ sender ev6 \ address ev6 \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ (\s bal. s6 = keep \ INV st6 s (SUMM s) bal \ POS s \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') \ sender ev6 \ address ev6 \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) ))))" proof (induct rule: msel_ssel_lexp_expr_load_rexp_stmt.induct [where ?P1.0="\c1 t1 l1 xe1 ep1 ev1 cd1 st1. address ev1 \ (STR ''Victim'') \ fmlookup ep1 (STR ''Victim'') = Some (victim, SKIP) \ (\rv1 st1' bal. frame bal st1 \ msel c1 t1 l1 xe1 ep1 ev1 cd1 st1 = Normal (rv1, st1') \ frame bal st1')" and ?P2.0="\t2 l2 xe2 ep2 ev2 cd2 st2. address ev2 \ (STR ''Victim'') \ fmlookup ep2 (STR ''Victim'') = Some (victim, SKIP) \ (\rv2 st2' bal. frame bal st2 \ ssel t2 l2 xe2 ep2 ev2 cd2 st2 = Normal (rv2, st2') \ frame bal st2')" and ?P3.0="\l5 ep5 ev5 cd5 st5. address ev5 \ (STR ''Victim'') \ fmlookup ep5 (STR ''Victim'') = Some (victim, SKIP) \ (\rv5 st5' bal. frame bal st5 \ lexp l5 ep5 ev5 cd5 st5 = Normal (rv5, st5') \ frame bal st5')" and ?P4.0="\e4 ep4 ev4 cd4 st4. address ev4 \ (STR ''Victim'') \ fmlookup ep4 (STR ''Victim'') = Some (victim, SKIP) \ (\rv4 st4' bal. frame bal st4 \ expr e4 ep4 ev4 cd4 st4 = Normal (rv4, st4') \ frame bal st4')" and ?P5.0="\lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst. address lev \ (STR ''Victim'') \ fmlookup lep (STR ''Victim'') = Some (victim, SKIP) \ (\ev cd st st' bal. load lcp lis lxs lep lev0 lcd0 lst0 lev lcd lst = Normal ((ev, cd, st), st') \ (frame bal lst0 \ frame bal st) \ (frame bal lst \ frame bal st') \ address lev0 = address ev \ sender lev0 = sender ev \ svalue lev0 = svalue ev)" and ?P6.0="\l3 ep3 ev3 cd3 st3. address ev3 \ (STR ''Victim'') \ fmlookup ep3 (STR ''Victim'') = Some (victim, SKIP) \ (\rv3 st3' bal. frame bal st3 \ rexp l3 ep3 ev3 cd3 st3 = Normal (rv3, st3') \ frame bal st3')" and ?P7.0="\s6 ep6 ev6 cd6 st6. (fmlookup ep6 (STR ''Victim'') = Some (victim, SKIP) \ (\st6'. stmt s6 ep6 ev6 cd6 st6 = Normal((), st6') \ ((address ev6 \ (STR ''Victim'') \ (\bal. frame bal st6 \ frame bal st6')) \ (address ev6 = (STR ''Victim'') \ (\s val bal x. s6 = transfer \ INV st6 s (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) bal \ POS s \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) \ accessStore x (stack st6) = Some (KValue val) \ sender ev6 \ address ev6 \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ (\s bal x. s6 = comp \ INV st6 s (SUMM s) bal \ POS s \ fmlookup (denvalue ev6) (STR ''bal'') = Some (Value (TUInt 256), Stackloc x) \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') \ accessStore x (stack st6) = Some (KValue (accessStorage (TUInt 256) (sender ev6 + (STR ''.'' + STR ''balance'')) s)) \ sender ev6 \ address ev6 \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) \ (\s bal. s6 = keep \ INV st6 s (SUMM s) bal \ POS s \ fmlookup (denvalue ev6) (STR ''balance'') = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'') \ sender ev6 \ address ev6 \ (\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')) ))))" ]) case (1 uu uv uw ux uy uz st) then show ?case by simp next case (2 va vb vc vd ve vf vg st) then show ?case by simp next case (3 vh al t loc x e\<^sub>p env cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ msel vh (MTArray al t) loc [x] e\<^sub>p env cd st = Normal (rv1, st')" moreover from * obtain v4 t4 st4' where **: "expr x e\<^sub>p env cd st = Normal ((v4, t4), st4')" by (auto split: result.split_asm) moreover from * ** have "frame bal st4'" using 3(1) asm by (auto split:if_split_asm) ultimately show "frame bal st'" by (simp split:Stackvalue.split_asm Type.split_asm if_split_asm) qed qed next case (4 mm al t loc x y ys e\<^sub>p env cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ msel mm (MTArray al t) loc (x # y # ys) e\<^sub>p env cd st = Normal (rv1, st')" moreover from * obtain v4 t4 st'' where **: "expr x e\<^sub>p env cd st = Normal ((KValue v4, Value t4), st'')" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have f1: "frame bal st''" using 4(1) asm by (auto split:if_split_asm) moreover from * ** have ***: "Valuetypes.less t4 (TUInt 256) v4 \al\ = Some (\True\, TBool)" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm if_split_asm) moreover from * ** *** obtain vb st''' where ****: "(applyf (\st. if mm then memory st else cd) st'') = Normal (vb, st''')" and f2: "frame bal st'''" using f1 by (simp split:Stackvalue.split_asm Type.split_asm if_split_asm) moreover from * ** *** **** obtain ll where *****: "accessStore (hash loc v4) vb = Some (MPointer ll)" by (simp split: Type.split_asm if_split_asm option.split_asm Memoryvalue.split_asm) moreover from * ** *** **** ***** obtain l1' st'''' where ******: "msel mm t ll (y # ys) e\<^sub>p env cd st''' = Normal (l1', st'''')" by (simp split: Type.split_asm if_split_asm option.split_asm Memoryvalue.split_asm) ultimately have "st' = st''''" by simp moreover have x1: "\rv1' st1' bal. (frame bal st''') \ local.msel mm t ll (y # ys) e\<^sub>p env cd st''' = Normal (rv1', st1') \ frame bal st1'" using 4(2)[OF ** _ _ _ *** _ *****] **** asm apply safe by auto ultimately show "frame bal st'" using f2 ****** by blast qed qed next case (5 tp loc vi vj vk st) then show ?case by simp next case (6 vl vm vn vo vp vq vr st) then show ?case by simp next case (7 al t loc x xs e\<^sub>p env cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ ssel (STArray al t) loc (x # xs) e\<^sub>p env cd st = Normal (rv1, st')" moreover from * obtain v4 t4 st4' where **: "expr x e\<^sub>p env cd st = Normal ((KValue v4, Value t4), st4')" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have f1: "frame bal st4'" using 7(1) asm by (auto split:if_split_asm) moreover from * ** have ***: "Valuetypes.less t4 (TUInt 256) v4 \al\ = Some (\True\, TBool)" by (auto split: result.split_asm Stackvalue.split_asm Type.split_asm if_split_asm) moreover from * ** *** obtain l1' st''' where ****: "ssel t (hash loc v4) xs e\<^sub>p env cd st4' = Normal (l1', st''')" by (simp split: Type.split_asm if_split_asm option.split_asm Memoryvalue.split_asm) ultimately have "st' = st'''" by simp moreover have "\rv' st2' bal. (frame bal st4') \ ssel t (hash loc v4) xs e\<^sub>p env cd st4' = Normal (rv', st2') \ frame bal st2'" using 7(2)[OF ** _ _ _ ***] asm apply safe by auto ultimately show "frame bal st'" using f1 **** by blast qed qed next case (8 vs t loc x xs e\<^sub>p env cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address env \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ ssel (STMap vs t) loc (x # xs) e\<^sub>p env cd st = Normal (rv1, st')" moreover from * obtain v4 t4 st4' where **: "expr x e\<^sub>p env cd st = Normal ((KValue v4, t4), st4')" by (auto split: result.split_asm Stackvalue.split_asm) moreover from * ** have ***: "frame bal st4'" using 8(1) asm by (auto split:if_split_asm) moreover from * ** *** obtain l1' st''' where ****:"ssel t (hash loc v4) xs e\<^sub>p env cd st4' = Normal (l1', st''')" by simp moreover from *** **** have "frame bal st'''" using 8(2)[OF **,of "KValue v4" t4 v4] asm by blast ultimately show "frame bal st'" by (simp split:Stackvalue.split_asm) qed qed next case (9 i vt e vu st) then show ?case by (auto split:option.split_asm result.split_asm Denvalue.split_asm) next case (10 i r e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ lexp (Ref i r) e\<^sub>p e cd st = Normal (rv1, st')" show "frame bal st'" proof (cases "fmlookup (denvalue e) i") case None with * show ?thesis by simp next case (Some a) then show ?thesis proof (cases a) case (Pair tp b) then show ?thesis proof (cases b) case (Stackloc l') then show ?thesis proof (cases "accessStore l' (stack st)") case None with * show ?thesis using Some Pair Stackloc by simp next case s2: (Some k) then show ?thesis proof (cases k) case (KValue x1) with * show ?thesis using Some Pair Stackloc s2 by simp next case (KCDptr x2) with * show ?thesis using Some Pair Stackloc s2 by simp next case (KMemptr l'') then show ?thesis proof (cases tp) case (Value x1) with * show ?thesis using Some Pair Stackloc s2 KMemptr by simp next case (Calldata x2) with * show ?thesis using Some Pair Stackloc s2 KMemptr by simp next case (Memory x3) with * Some Pair Stackloc KMemptr s2 obtain l1' t1' where "msel True x3 l'' r e\<^sub>p e cd st = Normal ((l1', t1'), st')" by (auto split: result.split_asm) with * 10(1)[OF Some Pair Stackloc _ _ KMemptr, of "Some k" st x3] show ?thesis using s2 Memory asm by auto next case (Storage x4) with * show ?thesis using Some Pair Stackloc s2 KMemptr by simp qed next case (KStoptr l'') then show ?thesis proof (cases tp) case (Value x1) with * show ?thesis using Some Pair Stackloc s2 KStoptr by simp next case (Calldata x2) with * show ?thesis using Some Pair Stackloc s2 KStoptr by simp next case (Memory x3) with * show ?thesis using Some Pair Stackloc s2 KStoptr by simp next case (Storage x4) with * Some Pair Stackloc KStoptr s2 obtain l1' t1' where "ssel x4 l'' r e\<^sub>p e cd st = Normal ((l1', t1'), st')" by (auto split: result.split_asm) with * 10(2)[OF Some Pair Stackloc _ _ KStoptr, of "Some k" st x4] show ?thesis using s2 Storage asm by auto qed qed qed next case (Storeloc l'') then show ?thesis proof (cases tp) case (Value x1) with * show ?thesis using Some Pair Storeloc by simp next case (Calldata x2) with * show ?thesis using Some Pair Storeloc by simp next case (Memory x3) with * show ?thesis using Some Pair Storeloc by simp next case (Storage x4) with * Some Pair Storeloc obtain l1' t1' where "ssel x4 l'' r e\<^sub>p e cd st = Normal ((l1', t1'), st')" by (auto split: result.split_asm) with * 10(3)[OF Some Pair Storeloc Storage] asm show ?thesis by auto qed qed qed qed qed qed next case (11 b x e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (12 b x e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (13 ad e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (14 ad e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (BALANCE ad) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain adv st'' where **:"expr ad e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (BALANCE ad) e\<^sub>p e cd st)\) = Normal ((KValue adv, Value TAddr),st'')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Types.split_asm Type.split_asm) with * ** have "frame bal st''" using 14(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** have "st' = st''" by (simp split:if_split_asm) ultimately show "frame bal st'" by simp qed qed next case (15 e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (16 e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (17 e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (18 e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (19 e\<^sub>p e cd st) then show ?case by (simp add:frame_def) next case (20 x e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (NOT x) e\<^sub>p e cd st = Normal (rv1, st')" then have f1: "frame bal (st\gas:=gas st - (costs\<^sub>e (NOT x) e\<^sub>p e cd st)\)" by (simp add:frame_def) moreover from * obtain v t st'' where **: "expr x e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (NOT x) e\<^sub>p e cd st)\) = Normal ((KValue v, Value t), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 20(1) asm by (auto simp add:frame_def split:if_split_asm) show "frame bal st'" proof (cases "v = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True") case True with * ** *** obtain x st''' where "expr FALSE e\<^sub>p e cd st'' = Normal (x, st''')" and "frame bal st'''" by (auto simp add:frame_def split:if_split_asm) with * ** *** True show ?thesis by (auto split: if_split_asm) next case False with * ** *** obtain x st''' where "expr TRUE e\<^sub>p e cd st'' = Normal (x, st''')" and "frame bal st'''" by (auto simp add:frame_def split:if_split_asm) with * ** *** False show ?thesis by (auto split: if_split_asm) qed qed qed next case (21 e1 e2 e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (PLUS e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (PLUS e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 21(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** *** **** have "frame bal st'''" using 21(2)[OF _ **] asm by (auto split:if_split_asm) moreover from * ** **** obtain v t where "add t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) ultimately show "frame bal st'" by (auto split:if_split_asm) qed qed next case (22 e1 e2 e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (MINUS e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (MINUS e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 22(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** *** **** have "frame bal st'''" using 22(2)[OF _ **] asm by (auto split:if_split_asm) moreover from * ** **** obtain v t where "sub t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) ultimately show "frame bal st'" by (auto split:if_split_asm) qed qed next case (23 e1 e2 e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (LESS e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (LESS e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 23(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** *** **** have "frame bal st'''" using 23(2)[OF _ **] asm by (auto split:if_split_asm) moreover from * ** **** obtain v t where "Valuetypes.less t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) ultimately show "frame bal st'" by (auto split:if_split_asm) qed qed next case (24 e1 e2 e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (EQUAL e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (EQUAL e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 24(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** *** **** have "frame bal st'''" using 24(2)[OF _ **] asm by (auto split:if_split_asm) moreover from * ** **** obtain v t where "Valuetypes.equal t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) ultimately show "frame bal st'" by (auto split:if_split_asm) qed qed next case (25 e1 e2 e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (AND e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (AND e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 25(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** *** **** have "frame bal st'''" using 25(2)[OF _ **] asm by (auto split:if_split_asm) moreover from * ** **** obtain v t where "Valuetypes.vtand t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) ultimately show "frame bal st'" by (auto split:if_split_asm) qed qed next case (26 e1 e2 e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (OR e1 e2) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * obtain v1 t1 st'' where **: "expr e1 e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (OR e1 e2) e\<^sub>p e cd st)\) = Normal ((KValue v1, Value t1), st'')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** have ***: "frame bal st''" using 26(1) asm by (auto simp add:frame_def split:if_split_asm) moreover from * ** *** obtain v2 t2 st''' where ****: "expr e2 e\<^sub>p e cd st'' = Normal ((KValue v2, Value t2), st''')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm) moreover from * ** *** **** have "frame bal st'''" using 26(2)[OF _ **] asm by (auto split:if_split_asm) moreover from * ** **** obtain v t where "Valuetypes.vtor t1 t2 v1 v2 = Some (v, t)" by (auto split:if_split_asm option.split_asm) ultimately show "frame bal st'" by (auto split:if_split_asm) qed qed next case (27 i e\<^sub>p e cd st) show ?case using 27(1)[of "()" "st\gas:=gas st - (costs\<^sub>e (LVAL i) e\<^sub>p e cd st)\"] apply safe by (auto simp add:frame_def split:if_split_asm) next case (28 i xe e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (CALL i xe) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * have a1: "(applyf (costs\<^sub>e (CALL i xe) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st\)" by auto moreover from * obtain ct bla where **: "fmlookup e\<^sub>p (address e) = Some (ct, bla)" by (auto split:if_split_asm option.split_asm) moreover from * ** obtain fp f x where ***: "fmlookup ct i = Some (Method (fp, f, Some x))" by (auto split:if_split_asm option.split_asm Member.split_asm) moreover define e' where "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" moreover from * ** *** obtain e'' cd' st'' st''' where ****: "load False fp xe e\<^sub>p e' emptyStore (st\gas:=gas st - (costs\<^sub>e (CALL i xe) e\<^sub>p e cd st), stack:=emptyStore\) e cd (st\gas:=gas st - (costs\<^sub>e (CALL i xe) e\<^sub>p e cd st)\) = Normal ((e'', cd', st''), st''')" using e'_def by (auto split:if_split_asm result.split_asm) moreover from * **** have f1: "frame bal st''" and ad: "address e' = address e''" using asm 28(1)[OF _ ** _ *** _ _ _ _ e'_def, of _ "st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st\" bla "(fp, f, Some x)" fp "(f, Some x)" f "Some x" x "st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st, stack := emptyStore\" "st\gas := gas st - costs\<^sub>e (CALL i xe) e\<^sub>p e cd st\"] by (auto simp add:frame_def split:if_split_asm result.split_asm) moreover from e'_def have ad2: "address e = address e'" using ffold_init_ad_same[of ct "(emptyEnv (address e) (sender e) (svalue e))" "(fmdom ct)" e'] by simp moreover from * ** *** **** e'_def obtain st'''' where *****: "stmt f e\<^sub>p e'' cd' st'' = Normal ((), st'''')" by (auto split:if_split_asm result.split_asm) moreover from f1 ad ad2 asm ***** have f2:"frame bal st''''" using 28(2)[OF a1 ** _ *** _ _ _ _ e'_def _ ****, of bla "(fp, f, Some x)" "(f, Some x)" f "Some x" x e'' "(cd', st'')" "cd'" "st''" st''' st''' "()" st''] by (simp add:frame_def) moreover from * ** *** **** ***** f1 f2 e'_def obtain rv st''''' where ******: "expr x e\<^sub>p e'' cd' st'''' = Normal (rv, st''''')" by (auto split:if_split_asm result.split_asm) ultimately have "st' = st'''''\stack:=stack st''', memory := memory st'''\" apply safe by auto moreover from f1 f2 ad ad2 asm a1 ***** ****** have "\rv4 st4' bal. frame bal st'''' \ local.expr x e\<^sub>p e'' cd' st'''' = Normal (rv4, st4') \ frame bal st4'" using e'_def asm 28(3)[OF a1 ** _ *** _ _ _ _ e'_def _ **** _ _ _ _ *****, of bla "(fp, f, Some x)" " (f, Some x)" "Some x" x "(cd', st'')" st'' st''' st''' "()"] apply safe by auto with ****** f2 have "frame bal st'''''" by blast ultimately show "frame bal st'" by (simp add:frame_def) qed qed next case (29 ad i xe val e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv1 and st' and bal assume *: "frame bal st \ expr (ECALL ad i xe val) e\<^sub>p e cd st = Normal (rv1, st')" moreover from * have a1: "(applyf (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd st\)" by auto moreover from * obtain adv st'' where **: "expr ad e\<^sub>p e cd (st\gas:=gas st - (costs\<^sub>e (ECALL ad i xe val) e\<^sub>p e cd st)\) = Normal ((KValue adv, Value TAddr), st'')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) moreover from * ** have f1: "frame bal st''"using asm 29(1) by (auto simp add:frame_def split:if_split_asm) moreover from * ** obtain ct bla where ***: "fmlookup e\<^sub>p adv = Some (ct, bla)" by (auto split:if_split_asm option.split_asm) moreover from * ** *** obtain fp f x where ****: "fmlookup ct i = Some (Method (fp, f, Some x))" by (auto split:if_split_asm option.split_asm Member.split_asm) moreover from * ** *** **** obtain v t st''' where *****: "expr val e\<^sub>p e cd st'' = Normal ((KValue v, Value t), st''')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm) moreover from f1 ***** asm have f2: "frame bal st'''" and f3: "frame bal (st'''\stack := emptyStore, memory := emptyStore\)" using asm 29(2)[OF a1 ** _ _ _ _ *** _ ****] by (auto simp add:frame_def) moreover define e' where "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" moreover from * ** *** **** ***** obtain e'' cd' st'''' st''''' where *******: "load True fp xe e\<^sub>p e' emptyStore (st'''\stack:=emptyStore, memory:=emptyStore\) e cd st''' = Normal ((e'', cd', st''''), st''''')" using e'_def by (auto split:if_split_asm result.split_asm option.split_asm) moreover have "(\ev cda st st' bal. local.load True fp xe e\<^sub>p e' emptyStore (st'''\stack := emptyStore, memory := emptyStore\) e cd st''' = Normal ((ev, cda, st), st') \ (frame bal (st'''\stack := emptyStore, memory := emptyStore\) \ frame bal st) \ (frame bal st''' \ frame bal st') \ address e' = address ev \ sender e' = sender ev \ svalue e' = svalue ev)" using 29(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ _ ***** _ _ _ e'_def, of "KValue adv" "Value TAddr" TAddr bla "(fp, f, Some x)" fp "(f, Some x)" f "Some x" x "KValue v" "Value t" t "st'''\stack := emptyStore, memory := emptyStore\" st'''] asm ******* by simp then have "frame bal st'''' \ frame bal st''''' \ address e' = address e''" using ******* f2 f3 by blast then have f4: "frame bal st''''" and ad1: "address e' = address e''" by auto moreover from * ** *** **** ***** ******* e'_def obtain acc where ******: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) then have ******: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) moreover from f4 have f5: "frame bal (st''''\accounts := acc\)" using transfer_frame[OF ******] asm by simp moreover from e'_def have ad2: "adv = address e'" using ffold_init_ad_same[of ct "(emptyEnv adv (address e) v)" "(fmdom ct)" e'] by simp moreover from * ** *** **** ***** ****** ******* ****** obtain st'''''' where ********: "stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" using e'_def by (auto simp del: transfer.simps split:if_split_asm result.split_asm) moreover have "adv \ STR ''Victim''" proof (rule ccontr) assume "\ adv \ STR ''Victim''" with asm ** *** **** show False using victim_def fmap_of_list_SomeD[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], deposit, None)), (STR ''withdraw'', Method ([], keep, None))]"] by auto qed with ad1 ad2 have ad: "address e'' \ STR ''Victim'' \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" using asm by simp then have "(\bal. frame bal (st''''\accounts := acc\) \ frame bal st'''''')" using 29(4)[OF a1 ** _ _ _ _ *** _ **** _ _ _ _ ***** _ _ _ _ _ ******* _ _ _ ******, of "KValue adv" "Value TAddr" TAddr bla "(fp, f, Some x)" "(f, Some x)" f "Some x" x "KValue v" "Value t" t e'' "(cd', st'''')" cd' st''''' st''''' "()" "st''''\accounts := acc\"] ******** e'_def by auto then have f4: "frame bal st''''''" using f5 ******** by auto moreover from * ** *** **** ***** ****** ******* ******** obtain rv st''''''' where *********: "expr x e\<^sub>p e'' cd' st'''''' = Normal (rv, st''''''')" using e'_def by (auto split:if_split_asm result.split_asm) ultimately have "st' = st'''''''\stack:=stack st''''', memory := memory st'''''\" apply safe by auto moreover from ad have "\rv4 st4' bal. frame bal st'''''' \ local.expr x e\<^sub>p e'' cd' st'''''' = Normal (rv4, st4') \ frame bal st4'" using e'_def 29(5)[OF a1 ** _ _ _ _ *** _ **** _ _ _ _ ***** _ _ _ _ _ ******* _ _ _ ****** _ ********, of "KValue adv" "Value TAddr" TAddr bla "(fp, f, Some x)"] by auto then have"frame bal st'''''''" using f4 ********* by blast ultimately show "frame bal st'" by (simp add:frame_def) qed qed next case (30 cp i\<^sub>p t\<^sub>p pl e el e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e\<^sub>v \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF allI[OF allI[OF impI]]]]]) fix ev and cda and sta and st'a and bal assume *: "local.load cp ((i\<^sub>p, t\<^sub>p) # pl) (e # el) e\<^sub>p e\<^sub>v' cd' st' e\<^sub>v cd st = Normal ((ev, cda, sta), st'a)" moreover from * obtain v t st'' where **: "expr e e\<^sub>p e\<^sub>v cd st = Normal ((v,t),st'')" by (auto split: result.split_asm) moreover from * ** obtain cd'' e\<^sub>v'' st''' where ***: "decl i\<^sub>p t\<^sub>p (Some (v,t)) cp cd (memory st'') cd' e\<^sub>v' st' = Normal ((cd'', e\<^sub>v''),st''')" by (auto split: result.split_asm) moreover from *** have ad: "address e\<^sub>v' = address e\<^sub>v'' \ sender e\<^sub>v' = sender e\<^sub>v'' \ svalue e\<^sub>v' = svalue e\<^sub>v''" using decl_gas_address by simp moreover from * ** *** obtain ev' cda' sta' st'a' where ****: "local.load cp pl el e\<^sub>p e\<^sub>v'' cd'' st''' e\<^sub>v cd st''= Normal ((ev', cda', sta'), st'a')" by (auto split: result.split_asm) ultimately have "ev = ev'" and "sta = sta'" and "st'a = st'a'" by simp+ from **** asm have IH: "(frame bal st''' \ frame bal sta') \ (frame bal st'' \ frame bal st'a') \ address e\<^sub>v'' = address ev' \ sender e\<^sub>v'' = sender ev' \ svalue e\<^sub>v'' = svalue ev'" using 30(2)[OF ** _ _ _ ***, of st'' "()" cd'' e\<^sub>v'' st''' st''' "()" st''] apply safe by (auto simp add:frame_def) show "(frame bal st' \ frame bal sta) \ (frame bal st \ frame bal st'a) \ address e\<^sub>v' = address ev \ sender e\<^sub>v' = sender ev \ svalue e\<^sub>v' = svalue ev" proof (rule conj3) show "frame bal st' \ frame bal sta" proof assume "frame bal st'" with * ** *** have "frame bal st'''" using decl_frame by simp with IH have "frame bal sta'" by simp with `sta = sta'` show "frame bal sta" by simp qed next show "frame bal st \ frame bal st'a" proof assume "frame bal st" with ** have "frame bal st''" using 30(1) asm by simp with IH have "frame bal st'a'" by simp with `st'a = st'a'` show "frame bal st'a" by simp qed next from ad IH show "address e\<^sub>v' = address ev \ sender e\<^sub>v' = sender ev \ svalue e\<^sub>v' = svalue ev" using `ev = ev'` by simp qed qed qed next case (31 vv vw vx vy vz wa wb wc wd st) then show ?case by simp next case (32 we wf wg wh wi wj wk wl wm st) then show ?case by simp next case (33 wn wo e\<^sub>v' cd' st' e\<^sub>v cd st) then show ?case by simp next case (34 i e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv3 and st3' and bal assume *: "frame bal st \ local.rexp (L.Id i) e\<^sub>p e cd st = Normal (rv3, st3')" show "frame bal st3'" proof (cases "fmlookup (denvalue e) i") case None with * show ?thesis by simp next case (Some a) then show ?thesis proof (cases a) case (Pair tp b) then show ?thesis proof (cases b) case (Stackloc l) then show ?thesis proof (cases "accessStore l (stack st)") case None with * Some Pair Stackloc show ?thesis by (auto split: Type.split_asm STypes.split_asm) next case s2: (Some a) with * Some Pair Stackloc s2 show ?thesis by (auto split: Type.split_asm STypes.split_asm Stackvalue.split_asm) qed next case (Storeloc x2) with * Some Pair Storeloc show ?thesis by (auto split: Type.split_asm STypes.split_asm option.split_asm) qed qed qed qed qed next case (35 i r e\<^sub>p e cd st) show ?case (is "_ \ ?RHS") proof assume asm: "address e \ (STR ''Victim'') \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" show ?RHS proof (rule allI[OF allI[OF allI[OF impI]]]) fix rv3 and st3' and bal assume *: "frame bal st \ local.rexp (L.Ref i r) e\<^sub>p e cd st = Normal (rv3, st3')" show "frame bal st3'" proof (cases "fmlookup (denvalue e) i") case None with * show ?thesis by simp next case (Some a) then show ?thesis proof (cases a) case (Pair tp b) then show ?thesis proof (cases b) case (Stackloc l') then show ?thesis proof (cases "accessStore l' (stack st)") case None with * Some Pair Stackloc show ?thesis by simp next case s2: (Some a) then show ?thesis proof (cases a) case (KValue x1) with * Some Pair Stackloc s2 show ?thesis by simp next case (KCDptr l'') then show ?thesis proof (cases tp) case (Value x1) with * Some Pair Stackloc s2 KCDptr show ?thesis by simp next case (Calldata t) with * Some Pair Stackloc s2 KCDptr obtain l''' t' st' where **: "msel False t l'' r e\<^sub>p e cd st = Normal ((l''',t'), st')" by (auto split: Type.split_asm STypes.split_asm result.split_asm) then have "\rv1 st1' bal. frame bal st \ local.msel False t l'' r e\<^sub>p e cd st = Normal (rv1, st1') \ frame bal st1'" using asm 35(1)[OF Some Pair Stackloc _ s2 KCDptr Calldata] by auto with * ** have f2: "frame bal st'" by blast then show ?thesis proof (cases t') case (MTArray x t'') then show ?thesis proof (cases "accessStore l''' cd") case None with * ** Some Pair Stackloc s2 KCDptr Calldata MTArray show ?thesis by simp next case s3: (Some a) then show ?thesis proof (cases a) case (MValue x1) with * ** Some Pair Stackloc s2 KCDptr Calldata MTArray s3 show ?thesis by simp next case (MPointer x2) with * ** f2 Some Pair Stackloc s2 KCDptr Calldata MTArray s3 show ?thesis by simp qed qed next case (MTValue t''') then show ?thesis proof (cases "accessStore l''' cd") case None with * ** Some Pair Stackloc s2 KCDptr Calldata MTValue show ?thesis by simp next case s3: (Some a) then show ?thesis proof (cases a) case (MValue x1) with * ** f2 Some Pair Stackloc s2 KCDptr Calldata MTValue s3 show ?thesis by simp next case (MPointer x2) with * ** Some Pair Stackloc s2 KCDptr Calldata MTValue s3 show ?thesis by simp qed qed qed next case (Memory x3) with * Some Pair Stackloc s2 KCDptr show ?thesis by simp next case (Storage x4) with * Some Pair Stackloc s2 KCDptr show ?thesis by simp qed next case (KMemptr l'') then show ?thesis proof (cases tp) case (Value x1) with * Some Pair Stackloc s2 KMemptr show ?thesis by simp next case (Calldata x2) with * Some Pair Stackloc s2 KMemptr show ?thesis by simp next case (Memory t) with * Some Pair Stackloc s2 KMemptr obtain l''' t' st' where **: "msel True t l'' r e\<^sub>p e cd st = Normal ((l''',t'), st')" by (auto split: Type.split_asm STypes.split_asm result.split_asm) then have "\rv1 st1' bal. frame bal st \ local.msel True t l'' r e\<^sub>p e cd st = Normal (rv1, st1') \ frame bal st1'" using asm 35(2)[OF Some Pair Stackloc _ s2 KMemptr Memory, of st] by auto with * ** have f2: "frame bal st'" by blast then show ?thesis proof (cases t') case (MTArray x11 x12) then show ?thesis proof (cases "accessStore l''' (memory st')") case None with * ** Some Pair Stackloc s2 KMemptr Memory MTArray show ?thesis by simp next case s3: (Some a) then show ?thesis proof (cases a) case (MValue x1) with * ** Some Pair Stackloc s2 KMemptr Memory MTArray s3 show ?thesis by simp next case (MPointer x2) with * ** f2 Some Pair Stackloc s2 KMemptr Memory MTArray s3 show ?thesis by simp qed qed next case (MTValue x2) then show ?thesis proof (cases "accessStore l''' (memory st')") case None with * ** Some Pair Stackloc s2 KMemptr Memory MTValue show ?thesis by simp next case s3: (Some a) then show ?thesis proof (cases a) case (MValue x1) with * ** f2 Some Pair Stackloc s2 KMemptr Memory MTValue s3 show ?thesis by simp next case (MPointer x2) with * ** Some Pair Stackloc s2 KMemptr Memory MTValue s3 show ?thesis by simp qed qed qed next case (Storage x4) with * Some Pair Stackloc s2 KMemptr show ?thesis by simp qed next case (KStoptr l'') then show ?thesis proof (cases tp) case (Value x1) with * Some Pair Stackloc s2 KStoptr show ?thesis by simp next case (Calldata x2) with * Some Pair Stackloc s2 KStoptr show ?thesis by simp next case (Memory x3) with * Some Pair Stackloc s2 KStoptr show ?thesis by simp next case (Storage t) with * Some Pair Stackloc s2 KStoptr obtain l''' t' st' where **: "ssel t l'' r e\<^sub>p e cd st = Normal ((l''',t'), st')" by (auto split: Type.split_asm STypes.split_asm result.split_asm) then have "\rv2 st2' bal. frame bal st \ local.ssel t l'' r e\<^sub>p e cd st = Normal (rv2, st2') \ frame bal st2'" using asm 35(3)[OF Some Pair Stackloc _ s2 KStoptr Storage, of st] by auto with * ** have "frame bal st'" by blast with * ** Some Pair Stackloc s2 KStoptr Storage show ?thesis by (simp split: STypes.split_asm option.split_asm) qed qed qed next case (Storeloc l') then show ?thesis proof (cases tp) case (Value x1) with * Some Pair Storeloc show ?thesis by simp next case (Calldata x2) with * Some Pair Storeloc show ?thesis by simp next case (Memory x3) with * Some Pair Storeloc show ?thesis by simp next case (Storage t) with * Some Pair Storeloc obtain l'' t' st' where **: "ssel t l' r e\<^sub>p e cd st = Normal ((l'',t'), st')" by (auto split: result.split_asm) then have "\rv2 st2' bal. frame bal st \ local.ssel t l' r e\<^sub>p e cd st = Normal (rv2, st2') \ frame bal st2'" using asm 35(4)[OF Some Pair Storeloc Storage] by auto with * ** have "frame bal st'" by blast with * ** Some Pair Storeloc Storage show ?thesis by (simp split: STypes.split_asm option.split_asm) qed qed qed qed qed qed next case (36 e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume *: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st6' show "?RHS st6'" (is "?LHS \ ?RHS") proof assume t0: "stmt SKIP e\<^sub>p e cd st = Normal ((), st6')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume ad: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st6'" proof fix bal show "frame bal st \ frame bal st6'" proof assume "frame bal st" with t0 * show "frame bal st6'" by (auto simp add: frame_def split:if_split_asm) qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume "address e = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (37 lv ex e\<^sub>p env cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st6' show "?RHS st6'" (is "?LHS \ ?RHS") proof assume *: "stmt (ASSIGN lv ex) e\<^sub>p env cd st = Normal ((), st6')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address env \ STR ''Victim''" show "\bal. frame bal st \ frame bal st6'" proof fix bal show "frame bal st \ frame bal st6'" proof assume "frame bal st" with * have a1: "(applyf (costs (ASSIGN lv ex) e\<^sub>p env cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas:=gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\)" and f1: "frame bal (st\gas:=gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\)" by (auto simp add:frame_def) moreover from * obtain kv kt st' where **: "expr ex e\<^sub>p env cd (st\gas:=gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\) = Normal ((kv, kt), st')" by (auto split:if_split_asm result.split_asm) ultimately have "\rv4 st4' (ev4'::Environment) bal. frame bal (st\gas := gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\) \ local.expr ex e\<^sub>p env cd (st\gas := gas st - costs (ASSIGN lv ex) e\<^sub>p env cd st\) = Normal (rv4, st4') \ frame bal st4'" using asm 0 37(1) by simp with f1 ** have f2: "frame bal st'" by blast show "frame bal st6'" proof (cases kv) case (KValue v) then show ?thesis proof (cases kt) case (Value t) with * ** KValue obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) with KValue Value have "\rv5 st5' (ev5'::Environment) bal. frame bal st' \ local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ frame bal st5'" using asm 0 37(2)[OF a1 **] by simp with f2 *** have f3: "frame bal st''" by blast then show ?thesis proof (cases rv) case (LStackloc l') then show ?thesis proof (cases rt) case v2: (Value t') then show ?thesis proof (cases "Valuetypes.convert t t' v") case None with * ** *** KValue Value LStackloc v2 show ?thesis by (auto split:if_split_asm) next case (Some a) then show ?thesis proof (cases a) case (Pair v' b) with * ** *** KValue Value LStackloc v2 Some have "st6' = st'' \stack := updateStore l' (KValue v') (stack st'')\" by (auto split:if_split_asm) with f3 show ?thesis by (simp add:frame_def) qed qed next case (Calldata x2) with * ** *** KValue Value LStackloc show ?thesis by (auto split:if_split_asm) next case (Memory x3) with * ** *** KValue Value LStackloc show ?thesis by (auto split:if_split_asm) next case (Storage x4) with * ** *** KValue Value LStackloc show ?thesis by (auto split:if_split_asm) qed next case (LMemloc l') then show ?thesis proof (cases rt) case v2: (Value t') with * ** *** KValue Value LMemloc show ?thesis by (auto split:if_split_asm) next case (Calldata x2) with * ** *** KValue Value LMemloc show ?thesis by (auto split:if_split_asm) next case (Memory x3) then show ?thesis proof (cases x3) case (MTArray x11 x12) with * ** *** KValue Value LMemloc Memory show ?thesis by (auto split:if_split_asm) next case (MTValue t') then show ?thesis proof (cases "Valuetypes.convert t t' v") case None with * ** *** KValue Value LMemloc Memory MTValue show ?thesis by (auto split:if_split_asm) next case (Some a) then show ?thesis proof (cases a) case (Pair v' b) with * ** *** KValue Value LMemloc Memory MTValue Some have "st6' = st'' \memory := updateStore l' (MValue v') (memory st'')\" by (auto split:if_split_asm) with f3 show ?thesis by (simp add:frame_def) qed qed qed next case (Storage x4) with * ** *** KValue Value LMemloc Storage show ?thesis by (auto split:if_split_asm) qed next case (LStoreloc l') then show ?thesis proof (cases rt) case v2: (Value x1) with * ** *** KValue Value LStoreloc show ?thesis by (auto split:if_split_asm) next case (Calldata x2) with * ** *** KValue Value LStoreloc show ?thesis by (auto split:if_split_asm) next case (Memory x3) with * ** *** KValue Value LStoreloc show ?thesis by (auto split:if_split_asm) next case (Storage x4) then show ?thesis proof (cases x4) case (STArray x11 x12) with * ** *** KValue Value LStoreloc Storage show ?thesis by (auto split:if_split_asm) next case (STMap x21 x22) with * ** *** KValue Value LStoreloc Storage show ?thesis by (auto split:if_split_asm) next case (STValue t') then show ?thesis proof (cases "Valuetypes.convert t t' v") case None with * ** *** KValue Value LStoreloc Storage STValue show ?thesis by (auto split:if_split_asm) next case (Some a) then show ?thesis proof (cases a) case (Pair v' b) then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KValue Value LStoreloc Storage STValue Some Pair show ?thesis by (auto split:if_split_asm) next case s2: (Some s) with * ** *** KValue Value LStoreloc Storage STValue Some Pair have "st6' = st''\storage := fmupd (address env) (fmupd l' v' s) (storage st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed qed qed qed qed next case (Calldata x2) with * ** KValue show ?thesis by (auto split:if_split_asm) next case (Memory x3) with * ** KValue show ?thesis by (auto split:if_split_asm) next case (Storage x4) with * ** KValue show ?thesis by (auto split:if_split_asm) qed next case (KCDptr p) then show ?thesis proof (cases kt) case (Value t) with * ** KCDptr show ?thesis by (auto split:if_split_asm) next case (Calldata x2) then show ?thesis proof (cases x2) case (MTArray x t) with * ** KCDptr Calldata obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) with KCDptr Calldata MTArray have "\rv5 st5' (ev5'::Environment) bal. frame bal st' \ local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ frame bal st5'" using asm 0 37(3)[OF a1 **] by auto with f2 *** have f3: "frame bal st''" by blast then show ?thesis proof (cases rv) case (LStackloc l') then show ?thesis proof (cases rt) case (Value x1) with * ** *** KCDptr Calldata MTArray LStackloc show ?thesis by (auto split:if_split_asm) next case c2: (Calldata x2) with * ** *** KCDptr Calldata MTArray LStackloc show ?thesis by (auto split:if_split_asm) next case (Memory x3) with f3 * ** *** KCDptr Calldata MTArray LStackloc show ?thesis by (auto simp add:frame_def split:if_split_asm) next case (Storage x4) then show ?thesis proof (cases "accessStore l' (stack st'')") case None with * ** *** KCDptr Calldata MTArray LStackloc Storage show ?thesis by (simp split:if_split_asm) next case (Some sv) then show ?thesis proof (cases sv) case (KValue x1) with * ** *** KCDptr Calldata MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) next case c2: (KCDptr x2) with * ** *** KCDptr Calldata MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) next case (KMemptr x3) with * ** *** KCDptr Calldata MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) next case (KStoptr p') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KCDptr Calldata MTArray LStackloc Storage Some KStoptr show ?thesis by (simp split:if_split_asm) next case s2: (Some s) then show ?thesis proof (cases "cpm2s p p' x t cd s") case None with * ** *** KCDptr Calldata MTArray LStackloc Storage Some KStoptr s2 show ?thesis by (simp split:if_split_asm) next case s3: (Some s') with * ** *** KCDptr Calldata MTArray LStackloc Storage Some KStoptr s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed qed qed qed next case (LMemloc l') then show ?thesis proof (cases "cpm2m p l' x t cd (memory st'')") case None with * ** *** KCDptr Calldata MTArray LMemloc show ?thesis by (auto split:if_split_asm) next case (Some m) with * ** *** KCDptr Calldata MTArray LMemloc have "st6' = st'' \memory := m\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed next case (LStoreloc l') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KCDptr Calldata MTArray LStoreloc show ?thesis by (auto split:if_split_asm) next case (Some s) then show ?thesis proof (cases "cpm2s p l' x t cd s") case None with * ** *** KCDptr Calldata MTArray LStoreloc Some show ?thesis by (auto split:if_split_asm) next case s2: (Some s') with * ** *** KCDptr Calldata MTArray LStoreloc Some s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed qed next case (MTValue x2) with * ** KCDptr Calldata show ?thesis by (simp split:if_split_asm) qed next case (Memory x3) with * ** KCDptr show ?thesis by (simp split:if_split_asm) next case (Storage x4) with * ** KCDptr show ?thesis by (simp split:if_split_asm) qed next case (KMemptr p) then show ?thesis proof (cases kt) case (Value t) with * ** KMemptr show ?thesis by (auto split:if_split_asm) next case (Calldata x2) with * ** KMemptr show ?thesis by (simp split:if_split_asm) next case (Memory x3) then show ?thesis proof (cases x3) case (MTArray x t) with * ** KMemptr Memory obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) with KMemptr Memory MTArray have "\rv5 st5' (ev5'::Environment) bal. frame bal st' \ local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ frame bal st5'" using asm 0 37(4)[OF a1 **] by auto with f2 *** have f3: "frame bal st''" by blast then show ?thesis proof (cases rv) case (LStackloc l') then show ?thesis proof (cases rt) case (Value x1) with * ** *** KMemptr Memory MTArray LStackloc show ?thesis by (auto split:if_split_asm) next case (Calldata x2) with * ** *** KMemptr Memory MTArray LStackloc show ?thesis by (auto split:if_split_asm) next case m3: (Memory x3) with f3 * ** *** KMemptr Memory MTArray LStackloc show ?thesis by (auto simp add:frame_def split:if_split_asm) next case (Storage x4) then show ?thesis proof (cases "accessStore l' (stack st'')") case None with * ** *** KMemptr Memory MTArray LStackloc Storage show ?thesis by (simp split:if_split_asm) next case (Some sv) then show ?thesis proof (cases sv) case (KValue x1) with * ** *** KMemptr Memory MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) next case (KCDptr x2) with * ** *** KMemptr Memory MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) next case m2: (KMemptr x3) with * ** *** KMemptr Memory MTArray LStackloc Storage Some show ?thesis by (simp split:if_split_asm) next case (KStoptr p') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KMemptr Memory MTArray LStackloc Storage Some KStoptr show ?thesis by (simp split:if_split_asm) next case s2: (Some s) then show ?thesis proof (cases "cpm2s p p' x t (memory st'') s") case None with * ** *** KMemptr Memory MTArray LStackloc Storage Some KStoptr s2 show ?thesis by (simp split:if_split_asm) next case s3: (Some s') with * ** *** KMemptr Memory MTArray LStackloc Storage Some KStoptr s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed qed qed qed next case (LMemloc l') with * ** *** KMemptr Memory MTArray LMemloc have "st6' = st'' \memory := updateStore l' (MPointer p) (memory st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) next case (LStoreloc l') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KMemptr Memory MTArray LStoreloc show ?thesis by (auto split:if_split_asm) next case (Some s) then show ?thesis proof (cases "cpm2s p l' x t (memory st'') s") case None with * ** *** KMemptr Memory MTArray LStoreloc Some show ?thesis by (auto split:if_split_asm) next case s2: (Some s') with * ** *** KMemptr Memory MTArray LStoreloc Some s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed qed next case (MTValue x2) with * ** KMemptr Memory show ?thesis by (simp split:if_split_asm) qed next case (Storage x4) with * ** KMemptr show ?thesis by (simp split:if_split_asm) qed next case (KStoptr p) then show ?thesis proof (cases kt) case (Value t) with * ** KStoptr show ?thesis by (auto split:if_split_asm) next case (Calldata x2) with * ** KStoptr show ?thesis by (simp split:if_split_asm) next case (Storage x3) then show ?thesis proof (cases x3) case (STArray x t) with * ** KStoptr Storage obtain rv rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((rv,rt), st'')" by (auto split:if_split_asm result.split_asm) with KStoptr Storage STArray have "\rv5 st5' bal. frame bal st' \ local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ frame bal st5'" using asm 0 37(5)[OF a1 **] by auto with f2 *** have f3: "frame bal st''" by blast then show ?thesis proof (cases rv) case (LStackloc l') then show ?thesis proof (cases rt) case (Value x1) with * ** *** KStoptr Storage STArray LStackloc show ?thesis by (auto split:if_split_asm) next case (Calldata x2) with * ** *** KStoptr Storage STArray LStackloc show ?thesis by (auto split:if_split_asm) next case (Memory x3) then show ?thesis proof (cases "accessStore l' (stack st'')") case None with * ** *** KStoptr Storage STArray LStackloc Memory show ?thesis by (simp split:if_split_asm) next case (Some sv) then show ?thesis proof (cases sv) case (KValue x1) with * ** *** KStoptr Storage STArray LStackloc Memory Some show ?thesis by (simp split:if_split_asm) next case (KCDptr x2) with * ** *** KStoptr Storage STArray LStackloc Memory Some show ?thesis by (simp split:if_split_asm) next case (KMemptr p') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KStoptr Storage STArray LStackloc Memory Some KMemptr show ?thesis by (simp split:if_split_asm) next case s2: (Some s) then show ?thesis proof (cases "cps2m p p' x t s (memory st'')") case None with * ** *** KStoptr Storage STArray LStackloc Memory Some KMemptr s2 show ?thesis by (simp split:if_split_asm) next case s3: (Some m) with * ** *** KStoptr Storage STArray LStackloc Memory Some KMemptr s2 have "st6' = st'' \memory := m\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed next case m2: (KStoptr x3) with * ** *** KStoptr Storage STArray LStackloc Memory Some show ?thesis by (simp split:if_split_asm) qed qed next case st2: (Storage x4) with f3 * ** *** KStoptr Storage STArray LStackloc show ?thesis by (auto simp add:frame_def split:if_split_asm) qed next case (LStoreloc l') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KStoptr Storage STArray LStoreloc show ?thesis by (auto split:if_split_asm) next case (Some s) then show ?thesis proof (cases "copy p l' x t s") case None with * ** *** KStoptr Storage STArray LStoreloc Some show ?thesis by (auto split:if_split_asm) next case s2: (Some s') with * ** *** KStoptr Storage STArray LStoreloc Some s2 have "st6' = st'' \storage := fmupd (address env) s' (storage st'')\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed next case (LMemloc l') then show ?thesis proof (cases "fmlookup (storage st'') (address env)") case None with * ** *** KStoptr Storage STArray LMemloc show ?thesis by (auto split:if_split_asm) next case (Some s) then show ?thesis proof (cases "cps2m p l' x t s (memory st'')") case None with * ** *** KStoptr Storage STArray LMemloc Some show ?thesis by (auto split:if_split_asm) next case s2: (Some m) with * ** *** KStoptr Storage STArray LMemloc Some s2 have "st6' = st'' \memory := m\" by (auto split:if_split_asm) with f3 show ?thesis using asm by (simp add:frame_def) qed qed qed next case (STMap t t') with * ** KStoptr Storage obtain l' rt st'' where ***: "lexp lv e\<^sub>p env cd st' = Normal ((LStackloc l',rt), st'')" by (auto split:if_split_asm result.split_asm LType.split_asm) with KStoptr Storage STMap have "\rv5 st5' (ev5'::Environment) bal. frame bal st' \ local.lexp lv e\<^sub>p env cd st' = Normal (rv5, st5') \ frame bal st5'" using asm 0 37(6)[OF a1 **] by auto with f2 *** have f3: "frame bal st''" by blast moreover from * ** *** KStoptr Storage STMap have "st6' = st'' \stack := updateStore l' (KStoptr p) (stack st'')\" by (auto split:if_split_asm) ultimately show ?thesis using asm f3 by (simp add:frame_def) next case (STValue x2) with * ** KStoptr Storage show ?thesis by (simp split:if_split_asm) qed next case (Memory x4) with * ** KStoptr show ?thesis by (simp split:if_split_asm) qed qed qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume "address env = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (38 s1 s2 e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st6' show "?RHS st6'" (is "?LHS \ ?RHS") proof assume *: "stmt (COMP s1 s2) e\<^sub>p e cd st = Normal ((), st6')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st6'" proof fix bal show "frame bal st \ frame bal st6'" proof assume "frame bal st" with * have a1: "(applyf (costs (COMP s1 s2) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\)" and f1: "frame bal (st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\)" by (auto simp add:frame_def) then have "\rv4 st4' bal. frame bal (st\gas := gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) \ stmt s1 e\<^sub>p e cd (st\gas := gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) = Normal (rv4, st4') \ frame bal st4'" using asm 0 38(1) by (simp add:frame_def) moreover from * obtain st' where **: "stmt s1 e\<^sub>p e cd (st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) = Normal ((), st')" by (auto split:if_split_asm result.split_asm) ultimately have f2: "frame bal st'" using f1 by blast have "\rv4 st4' bal. frame bal st' \ stmt s2 e\<^sub>p e cd st' = Normal (rv4, st4') \ frame bal st4'" using asm 0 38(2)[OF a1 **] by (simp add:frame_def) moreover from * ** obtain st'' where ***: "stmt s2 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) ultimately have f3: "frame bal st''" using f2 by blast from a1 * ** *** have "st6' = st''" by (simp split:if_split_asm) with f3 asm show "frame bal st6'" by simp qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume ad: "address e = STR ''Victim''" show ?RHS (is "?A \ ?B \ ?C") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS(is "?A1 \ ?A2 \ ?A3 \ ?A4 \ ?A5 \ ?A6") then have ?A1 and ?A2 and ?A3 and ?A4 and ?A5 and ?A6 by auto with * have c1: "gas st > costs comp e\<^sub>p e cd st" by (auto split:if_split_asm) with `?A1` * obtain st'' where 00: "stmt assign e\<^sub>p e cd (st\gas := gas st - costs comp e\<^sub>p e cd st\) = Normal((), st'')" by (auto split:result.split_asm) moreover from `?A2` have "fmlookup (storage (st\gas := gas st - costs comp e\<^sub>p e cd st\)) (STR ''Victim'') = Some s" by simp moreover from `?A2` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st\gas := gas st - costs comp e\<^sub>p e cd st\)) (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0" by simp moreover from `?A3` have "POS s" by simp moreover from `?A6` have "accessStore x (stack (st\gas := gas st - costs comp e\<^sub>p e cd st\)) = Some (KValue (accessStorage (TUInt 256) (sender e + (STR ''.'' + STR ''balance'')) s))" by simp ultimately obtain s'' where "fmlookup (storage st'') (STR ''Victim'') = Some s''" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - (SUMM s'' + ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e + (STR ''.'' + STR ''balance'')) s)) \ bal \ bal \ 0" and **: "accessStore x (stack st'') = Some (KValue (accessStorage (TUInt 256) (sender e + (STR ''.'' + STR ''balance'')) s))" and "POS s''" using secureassign[OF 00 _ ad `?A5`] that by blast moreover from c1 `?A1` * 00 obtain st''' where ***: "stmt transfer e\<^sub>p e cd st'' = Normal((), st''')" and "st6' = st'''" by auto moreover from `?A1` 00 have x1: "stmt s1 e\<^sub>p e cd (st\gas := gas st - costs (COMP s1 s2) e\<^sub>p e cd st\) = Normal((), st'')" by simp moreover from * have x2: "(applyf (costs (COMP s1 s2) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas:=gas st - costs (COMP s1 s2) e\<^sub>p e cd st\)" by (simp split: if_split_asm) ultimately show "\s'. fmlookup (storage st6') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s'" using 38(2)[OF x2 x1] `?A1` `?A4` ad 0 ** `?A6` by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (39 ex s1 s2 e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st6' show "?RHS st6'" (is "?LHS \ ?RHS") proof assume *: "stmt (ITE ex s1 s2) e\<^sub>p e cd st = Normal ((), st6')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st6'" proof fix bal show "frame bal st \ frame bal st6'" proof assume "frame bal st" with * have a1: "(applyf (costs (ITE ex s1 s2) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas:=gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\)" and f1: "frame bal (st\gas:=gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\)" by (auto simp add:frame_def) from * obtain b st' where **: "expr ex e\<^sub>p e cd (st\gas:=gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\) = Normal ((KValue b, Value TBool), st')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) moreover from asm have "\rv4 st4' bal. frame bal (st\gas := gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\) \ expr ex e\<^sub>p e cd (st\gas := gas st - costs (ITE ex s1 s2) e\<^sub>p e cd st\) = Normal (rv4, st4') \ frame bal st4'" using 0 39(1)[OF a1] by (simp add:frame_def) ultimately have f2: "frame bal st'" using f1 by blast show "frame bal st6'" proof (cases "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True") case True then have "\st6' bal. frame bal st' \ local.stmt s1 e\<^sub>p e cd st' = Normal ((), st6') \ frame bal st6'" using asm 0 39(2)[OF a1 **, of "KValue b" "Value TBool" b TBool] by (simp add:frame_def) moreover from * ** True obtain st'' where ***: "stmt s1 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) ultimately have "frame bal st''" using f2 by blast moreover from a1 * ** True *** have "st6' = st''" by (simp split:if_split_asm) ultimately show "frame bal st6'" using asm by simp next case False then have "\st6' bal. frame bal st' \ local.stmt s2 e\<^sub>p e cd st' = Normal ((), st6') \ frame bal st6'" using 0 asm 39(3)[OF a1 **, of "KValue b" "Value TBool" b TBool] by (simp add:frame_def) moreover from * ** False obtain st'' where ***: "stmt s2 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) ultimately have "frame bal st''" using f2 by blast moreover from a1 * ** False *** have "st6' = st''" by (simp split:if_split_asm) ultimately show "frame bal st6'" using asm by simp qed qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume "address e = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (40 ex s0 e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st6' show "?RHS st6'" (is "?LHS \ ?RHS") proof assume *: "stmt (WHILE ex s0) e\<^sub>p e cd st = Normal ((), st6')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st6'" proof fix bal show "frame bal st \ frame bal st6'" proof assume "frame bal st" with * have a1: "(applyf (costs (WHILE ex s0) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas:=gas st - costs (WHILE ex s0) e\<^sub>p e cd st\)" and f1: "frame bal (st\gas:=gas st - costs (WHILE ex s0) e\<^sub>p e cd st\)" by (auto simp add:frame_def) from * obtain b st' where **: "expr ex e\<^sub>p e cd (st\gas:=gas st - costs (WHILE ex s0) e\<^sub>p e cd st\) = Normal ((KValue b, Value TBool), st')" by (auto split:if_split_asm result.split_asm prod.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) moreover from asm have "\rv4 st4' bal. frame bal (st\gas := gas st - costs (WHILE ex s0) e\<^sub>p e cd st\) \ expr ex e\<^sub>p e cd (st\gas := gas st - costs (WHILE ex s0) e\<^sub>p e cd st\) = Normal (rv4, st4') \ frame bal st4'" using 0 40(1)[OF a1] by (simp add:frame_def) ultimately have f2: "frame bal st'" using f1 by blast show "frame bal st6'" proof (cases "b = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True") case True then have "\st6' bal. frame bal st' \ local.stmt s0 e\<^sub>p e cd st' = Normal ((), st6') \ frame bal st6'" using 0 asm 40(2)[OF a1 **, of "KValue b" "Value TBool" b TBool] by (simp add:frame_def) moreover from * ** True obtain st'' where ***: "stmt s0 e\<^sub>p e cd st' = Normal ((), st'')" by (auto split:if_split_asm result.split_asm) ultimately have f3: "frame bal st''" using f2 by blast have "\st6' bal. frame bal st'' \ local.stmt (WHILE ex s0) e\<^sub>p e cd st'' = Normal ((), st6') \ frame bal st6'" using 0 asm 40(3)[OF a1 ** _ _ _ _ True ***] by (simp add:frame_def) moreover from * ** True *** obtain st''' where ****: "stmt (WHILE ex s0) e\<^sub>p e cd st'' = Normal ((), st''')" by (auto split:if_split_asm result.split_asm) ultimately have "frame bal st'''" using f3 by blast moreover from a1 * ** True *** **** have "st6' = st'''" by (simp split:if_split_asm) ultimately show "frame bal st6'" using asm by simp next case False then show "frame bal st6'" using * ** f1 f2 asm by (simp split:if_split_asm) qed qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume "address e = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (41 i xe e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st' show "?RHS st'" (is "?LHS \ ?RHS") proof assume *: "stmt (INVOKE i xe) e\<^sub>p e cd st = Normal ((), st')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st'" proof fix bal show "frame bal st \ frame bal st'" proof assume ff: "frame bal st" moreover from * have a1: "(applyf (costs (INVOKE i xe) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (INVOKE i xe) e\<^sub>p e cd st\)" by auto moreover from * obtain ct bla where **: "fmlookup e\<^sub>p (address e) = Some (ct, bla)" by (auto split:if_split_asm option.split_asm) moreover from * ** obtain fp f where ***: "fmlookup ct i = Some (Method (fp, f, None))" by (auto split:if_split_asm option.split_asm Member.split_asm) moreover define e' where "e' = ffold_init ct (emptyEnv (address e) (sender e) (svalue e)) (fmdom ct)" moreover from * ** *** obtain e'' cd' st'' st''' where ****: "load False fp xe e\<^sub>p e' emptyStore (st\gas:=gas st - (costs (INVOKE i xe) e\<^sub>p e cd st), stack:=emptyStore\) e cd (st\gas:=gas st - (costs (INVOKE i xe) e\<^sub>p e cd st)\) = Normal ((e'', cd', st''), st''')" using e'_def by (auto split:if_split_asm result.split_asm) moreover from * **** have f1: "frame bal st''" and ad: "address e' = address e''" using asm ff 0 41(1)[OF a1 ** _ *** _ _ _ _ e'_def, of bla "(fp, f, None)" fp "(f, None)" f None] by (auto simp add:frame_def) moreover from e'_def have ad2: "address e = address e'" using ffold_init_ad_same[of ct "(emptyEnv (address e) (sender e) (svalue e))" "(fmdom ct)" e'] by simp moreover from * ** *** **** e'_def obtain st'''' where *****: "stmt f e\<^sub>p e'' cd' st'' = Normal ((), st'''')" by (auto split:if_split_asm result.split_asm) ultimately have "st' = st''''\stack:=stack st''', memory := memory st'''\" using * apply safe by simp moreover from f1 ad ad2 asm ***** have f2:"frame bal st''''" using 41(2)[OF a1 ** _ *** _ _ _ _ e'_def _ ****] using 0 * by (auto simp add:frame_def) ultimately show "frame bal st'" by (simp add:frame_def) qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume "address e = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (42 ad i xe val e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st' show "?RHS st'" (is "?LHS \ ?RHS") proof assume *: "stmt (EXTERNAL ad i xe val) e\<^sub>p e cd st = Normal ((), st')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st'" proof fix bal show "frame bal st \ frame bal st'" proof assume ff: "frame bal st" moreover from * have a1: "(applyf (costs (EXTERNAL ad i xe val) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (EXTERNAL ad i xe val) e\<^sub>p e cd st\)" by auto moreover from * obtain adv st'' where **: "expr ad e\<^sub>p e cd (st\gas:=gas st - (costs (EXTERNAL ad i xe val) e\<^sub>p e cd st)\) = Normal ((KValue adv, Value TAddr), st'')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) moreover from * ** ff have f1: "frame bal st''" using asm 0 42(1) by (simp add:frame_def split:if_split_asm) moreover from * ** obtain ct fb where ***: "fmlookup e\<^sub>p adv = Some (ct, fb)" by (auto split:if_split_asm option.split_asm) moreover from * ** *** f1 obtain v t st''' where ****: "expr val e\<^sub>p e cd st'' = Normal ((KValue v, Value t), st''')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm) moreover from **** f1 have "frame bal st'''" using asm 42(2)[OF a1 ** _ _ _ _ ***] 0 by (simp split:if_split_asm) then have f2: "frame bal (st'''\stack := emptyStore, memory := emptyStore\)" by (simp add:frame_def) show "frame bal st'" proof (cases "fmlookup ct i") case None with * ** *** **** obtain acc where trans: "Accounts.transfer (address e) adv v (accounts st''') = Some acc" by (auto split:if_split_asm option.split_asm) with * ** *** **** None obtain st'''' where *****: "stmt fb e\<^sub>p (emptyEnv adv (address e) v) cd (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'''')" by (auto split:if_split_asm result.split_asm) moreover have f4: "frame bal (st''''\stack:=stack st'', memory := memory st''\)" proof (cases "adv = STR ''Victim''") case True with 0 *** have "fb = SKIP" by simp moreover from f2 have "frame bal (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF trans] asm by (simp add:frame_def) ultimately show ?thesis using ***** by (auto split:if_split_asm simp add:frame_def) next case False moreover from f2 have "frame bal (st'''\accounts := acc,stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF trans] asm by (simp add:frame_def) then have "frame bal st''''" using f2 0 42(5)[OF a1 ** _ _ _ _ *** _ **** _ _ _ None _ trans, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t" t st''' st''' st'''] asm ***** False by (auto simp add:frame_def) then show ?thesis by (simp add:frame_def) qed ultimately show "frame bal st'" using a1 * ** *** **** None trans by (auto simp add:frame_def) next case (Some a) with * ** *** **** obtain fp f where *****: "fmlookup ct i = Some (Method (fp, f, None))" by (auto split:if_split_asm option.split_asm Member.split_asm) moreover define e' where e'_def: "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" moreover from * ** *** ***** **** obtain e'' cd' st'''' st''''' where *******: "load True fp xe e\<^sub>p e' emptyStore (st'''\stack:=emptyStore, memory:=emptyStore\) e cd st''' = Normal ((e'', cd', st''''), st''''')" using e'_def by (auto split:if_split_asm result.split_asm option.split_asm) moreover from e'_def have ad2: "address e' = adv" and send2: "sender e' = address e" and sval2: "svalue e' = v" using ffold_init_ad_same[of ct "(emptyEnv adv (address e) v)" "(fmdom ct)" e'] by auto moreover from * ** *** ***** **** ******* e'_def obtain acc where trans: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) then have ******: "Accounts.transfer (address e) adv v (accounts st'''') = Some acc" by (auto split:if_split_asm option.split_asm) moreover from * ** *** ***** **** ****** ******* obtain st'''''' where ********: "stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" using e'_def by (auto split:if_split_asm result.split_asm) moreover have f4: "frame bal st''''''" proof (cases "adv = STR ''Victim''") case True with 0 *** have ct_def: "ct = victim" by simp moreover have "(\(ev::Environment) cda st st' bal. local.load True fp xe e\<^sub>p e' emptyStore (st'''\stack := emptyStore, memory := emptyStore\) e cd st''' = Normal ((ev, cda, st), st') \ (frame bal (st'''\stack := emptyStore, memory := emptyStore\) \ frame bal st) \ (frame bal st''' \ frame bal st') \ address e' = address ev \ sender e' = sender ev \ svalue e' = svalue ev)" using 0 42(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def] asm by simp with f2 ******* have f3: "frame bal st''''" and ad1: "address e' = address e''" and send1: "sender e' = sender e''" and sval1: "svalue e' = svalue e''" by auto from ct_def ***** consider (withdraw) "i = STR ''withdraw''" | (deposit) "i = STR ''deposit''" using victim_def fmap_of_list_SomeD[of "[(STR ''balance'', Var (STMap TAddr (STValue (TUInt 256)))), (STR ''deposit'', Method ([], deposit, None)), (STR ''withdraw'', Method ([], keep, None))]"] by auto then show ?thesis proof cases case withdraw moreover have "fmlookup victim (STR ''withdraw'') = Some (Method ([], keep, None))" using victim_def by eval ultimately have "f=keep" and "fp = []" using *** ***** True 0 by auto with ******** have *********: "stmt keep e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" by simp have "fmlookup (denvalue e'') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" proof - from victim_def have some: "fmlookup victim (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by eval with ct_def have "fmlookup ct (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by simp moreover have "STR ''balance'' |\| fmdom (denvalue (emptyEnv adv (address e) v))" by simp moreover from ct_def some have "STR ''balance'' |\| fmdom ct" using fmdomI by simp ultimately have "fmlookup (denvalue e') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" using e'_def ffold_init_fmap[of ct "STR ''balance''" "(STMap TAddr (STValue (TUInt 256)))" "(emptyEnv adv (address e) v)" "(fmdom ct)"] by simp moreover have "e'' = e'" proof (cases "xe=[]") case True with ******* `fp = []` show ?thesis by simp next case False then obtain xx xe' where "xe = xx # xe'" using list.exhaust by auto with ******* `fp = []` show ?thesis by simp qed ultimately show ?thesis by simp qed moreover from ad1 ad2 True have ad: "address e'' = STR ''Victim''" by simp moreover from ad send1 send2 asm have "sender e'' \ address e''" by simp moreover from f3 have f4: "frame bal (st''''\accounts := acc\)" using transfer_frame[OF ******] asm by simp then obtain s'''' where "fmlookup (storage (st''''\accounts := acc\)) (STR ''Victim'') = Some s'''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) - (SUMM s'''') \ bal \ bal \ 0 \ POS s''''" by (auto simp add:frame_def) ultimately have "(\s'. fmlookup (storage st'''''') (STR ''Victim'') = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''''') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')" using 0 ******** `f=keep` 42(4)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def _ ******* _ _ _ trans, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t" t "(fp, f, None)" "(f, None)" f None e'' "(cd', st'''')" cd' st''''' st''''' "()" "st''''\accounts := acc\"] apply safe by auto then show ?thesis by (simp add:frame_def) next case deposit moreover from f2 have "frame bal (st'''\stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF ******] asm by simp moreover have "fmlookup victim (STR ''deposit'') = Some (Method ([], deposit, None))" using victim_def by eval ultimately have "f=deposit" and "fp = []" using *** ***** True 0 by auto with ******** have *: "stmt deposit e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st'''''')" by simp have f4: "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t v) (st''''\accounts := acc\)" and ad1: "address e' = address e''" and send1: "sender e' = sender e''" and sval1: "svalue e' = svalue e''" proof - have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) + ReadL\<^sub>i\<^sub>n\<^sub>t v" using transfer_add[OF ******] asm True by simp moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF ******] by simp ultimately have "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t v) (st''''\accounts := acc\)" using f3 by (auto simp add:frame_def) then show "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t v) (st''''\accounts := acc\)" and "address e' = address e''" and "sender e' = sender e''" and "svalue e' = svalue e''" using f2 0 42(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t"] asm ******* by (auto simp add:frame_def) qed moreover from sval1 sval2 have "v = svalue e''" by simp ultimately have "frame (bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e'')) (st''''\accounts := acc\)" by simp then obtain s''''' where II: "INV (st''''\accounts := acc\) s''''' (SUMM s''''') (bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e''))" and III:"POS s'''''" by (auto simp add:frame_def) then have s'''''_def: "fmlookup (storage (st''''\accounts := acc\)) STR ''Victim'' = Some s'''''" by simp have yyy: "fmlookup (denvalue e'') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" proof - from victim_def have some: "fmlookup victim (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by eval with ct_def have "fmlookup ct (STR ''balance'') = Some (Var (STMap TAddr (STValue (TUInt 256))))" by simp moreover have "STR ''balance'' |\| fmdom (denvalue (emptyEnv adv (address e) v))" by simp moreover from ct_def some have "STR ''balance'' |\| fmdom ct" using fmdomI by simp ultimately have "fmlookup (denvalue e') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" using e'_def ffold_init_fmap[of ct "STR ''balance''" "(STMap TAddr (STValue (TUInt 256)))" "(emptyEnv adv (address e) v)" "(fmdom ct)"] by simp moreover have "e'' = e'" proof (cases "xe=[]") case True with ******* `fp = []` show ?thesis by simp next case False then obtain xx xe' where "xe = xx # xe'" using list.exhaust by auto with ******* `fp = []` show ?thesis by simp qed ultimately show ?thesis by simp qed from asm True have "address e \ (STR ''Victim'')" by simp then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) + ReadL\<^sub>i\<^sub>n\<^sub>t v < 2^256" using transfer_val2[OF ******] True by simp moreover from `address e \ (STR ''Victim'')` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) + ReadL\<^sub>i\<^sub>n\<^sub>t v" using transfer_add[OF ******] (******) True by simp ultimately have abc: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) < 2^256" by simp from II have "fmlookup (storage (st''''\accounts := acc\)) (STR ''Victim'') = Some s''''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) - (SUMM s''''') \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e'') \ bal + ReadL\<^sub>i\<^sub>n\<^sub>t (svalue e'') \ 0" by (auto) moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') + \svalue e''\ < 2 ^ 256 \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') + \svalue e''\ \ 0" proof (cases "fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance'')) = None") case True hence "(accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') = ShowL\<^sub>i\<^sub>n\<^sub>t 0" by simp moreover have "(\svalue e''\::int) < 2 ^ 256" proof - from II have "bal + \svalue e''\ + SUMM s''''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp moreover have "0 \ SUMM s'''''" using III sum_nonneg[of "{(ad,x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"] by auto ultimately have "bal + \svalue e''\ \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp moreover from ff have "bal\0" by (auto simp add:frame_def) ultimately show "(\svalue e''\::int) < 2 ^ 256" using abc by simp qed moreover have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF ******] by simp with `svalue e' = v` sval1 have "(\svalue e''\::int) \ 0" by simp ultimately show ?thesis using Read_ShowL_id by simp next case False then obtain x where x_def: "fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance'')) = Some x" by auto moreover from II have "bal + \svalue e''\ + SUMM s''''' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp moreover have "(case (sender e'', x) of (ad, x) \ \x\) \ (\(ad, x)\{(ad, x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}. ReadL\<^sub>i\<^sub>n\<^sub>t x)" proof (cases rule: member_le_sum[of "(sender e'',x)" "{(ad,x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "\(ad,x). ReadL\<^sub>i\<^sub>n\<^sub>t x"]) case 1 then show ?case using x_def by simp next case (2 x) with III show ?case by auto next case 3 have "inj_on (\(ad, x). (ad + (STR ''.'' + STR ''balance''), x)) {(ad, x). (fmlookup s''''' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using balance_inj by simp then have "finite {(ad, x). (fmlookup s''''' \ (\ad. ad + (STR ''.'' + STR ''balance''))) ad = Some x}" using fmlookup_finite[of "\ad. (ad + (STR ''.'' + STR ''balance''))" s'''''] by simp then show ?case using finite_subset[of "{(ad, x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}" "{(ad, x). fmlookup s''''' (ad + (STR ''.'' + STR ''balance'')) = Some x}"] by auto qed then have "ReadL\<^sub>i\<^sub>n\<^sub>t x \ SUMM s'''''" by simp ultimately have "bal + \svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp moreover from ff have "bal\0" by (auto simp add:frame_def) ultimately have "\svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim''))" by simp with abc have "\svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x < 2^256" by simp moreover have "fmlookup s''''' (sender e' + (STR ''.'' + STR ''balance'')) = fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance''))" using send1 by simp ultimately have "bal + \svalue e''\ \ \accessBalance (accounts (st''''\accounts := acc\)) STR ''Victim''\ - SUMM s'''''" and lck: "fmlookup s''''' (sender e'' + (STR ''.'' + STR ''balance'')) = Some x" and "ReadL\<^sub>i\<^sub>n\<^sub>t x + \svalue e''\ < 2 ^ 256" using ad1 ad2 True II x_def by simp+ moreover from lck have "(accessStorage (TUInt 256) (sender e'' + (STR ''.'' + STR ''balance'')) s''''') = x" by simp moreover have "\svalue e''\ + ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" proof - have "ReadL\<^sub>i\<^sub>n\<^sub>t v \ 0" using transfer_val1[OF ******] by simp with `svalue e' = v` sval1 have "(\svalue e''\::int) \ 0" by simp moreover from III have "ReadL\<^sub>i\<^sub>n\<^sub>t x \ 0" using x_def by simp ultimately show ?thesis by simp qed ultimately show ?thesis using Read_ShowL_id by simp qed moreover have "address e'' = STR ''Victim''" using ad1 ad2 True by simp ultimately obtain s'''''' where "fmlookup (storage st'''''') (STR ''Victim'') = Some s''''''" and "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''''') (STR ''Victim'')) - SUMM s'''''' \ bal" and "POS s''''''" using deposit_frame[OF * s'''''_def _ yyy] III by auto then show ?thesis using ff by (auto simp add:frame_def) qed next case False moreover from f2 have "frame bal (st'''\stack:=emptyStore, memory:=emptyStore\)" using transfer_frame[OF ******] asm by simp then have "frame bal st''''" and ad1: "address e' = address e''" using f2 0 42(3)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def, of "KValue adv" "Value TAddr" TAddr fb "KValue v" "Value t"] asm ******* by (auto simp add:frame_def) then have f4: "frame bal (st''''\accounts := acc\)" using transfer_frame[OF ******] asm by simp moreover from ad1 ad2 have "address e'' \ STR ''Victim'' \ fmlookup e\<^sub>p (STR ''Victim'') = Some (victim, SKIP)" using 0 False by simp then have "\st6' bal. frame bal (st''''\accounts := acc\) \ local.stmt f e\<^sub>p e'' cd' (st''''\accounts := acc\) = Normal ((), st6') \ frame bal st6'" using 42(4)[OF a1 ** _ _ _ _ *** _ **** _ _ _ ***** _ _ _ _ e'_def _ ******* _ _ _ ******] False asm by (auto simp add:frame_def) ultimately show ?thesis using ******** by blast qed ultimately show "frame bal st'" using a1 * ** *** **** by (auto simp add:frame_def) qed qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume "address e = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (43 ad ex e\<^sub>p e cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st' show "?RHS st'" (is "?LHS \ ?RHS") proof assume *: "stmt (TRANSFER ad ex) e\<^sub>p e cd st = Normal ((), st')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e \ STR ''Victim''" show "\bal. frame bal st \ frame bal st'" proof fix bal show "frame bal st \ frame bal st'" proof assume ff: "frame bal st" from * have a1: "(applyf (costs (TRANSFER ad ex) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (TRANSFER ad ex) e\<^sub>p e cd st\)" by auto from * obtain v t st'' where **: "expr ex e\<^sub>p e cd (st\gas:=gas st - (costs (TRANSFER ad ex) e\<^sub>p e cd st)\) = Normal ((KValue v, Value t), st'')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm) from asm ff * ** have f1: "frame bal st''" using 43(1)[OF a1] 0 by (simp add:frame_def) from * ** obtain adv st''' where ***: "expr ad e\<^sub>p e cd st'' = Normal ((KValue adv, Value TAddr), st''')" by (auto split:if_split_asm result.split_asm Stackvalue.split_asm Type.split_asm Types.split_asm) from asm * *** f1 have f2: "frame bal st'''" using asm 43(2)[OF a1 **] 0 by (simp add:frame_def) from * ** *** obtain acc where *****: "Accounts.transfer (address e) adv v (accounts st''') = Some acc" by (auto split:if_split_asm option.split_asm) from f2 have f3: "frame bal (st'''\accounts := acc\)" using transfer_frame[OF *****] asm by simp show "frame bal st'" proof (cases "fmlookup e\<^sub>p adv") case None with a1 * ** *** ***** show ?thesis using f3 by auto next case (Some a) then show ?thesis proof (cases a) case (Pair ct f) moreover define e' where "e' = ffold_init ct (emptyEnv adv (address e) v) (fmdom ct)" moreover from * ** *** Some Pair ***** e'_def obtain st'''' where ******: "stmt f e\<^sub>p e' emptyStore (st'''\accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st'''')" by (auto split:if_split_asm option.split_asm result.split_asm) moreover from e'_def have ad: "adv = address e'" using ffold_init_ad_same[of ct "(emptyEnv adv (address e) v)" "(fmdom ct)" e'] by simp moreover have f4: "frame bal st''''" proof (cases "adv = STR ''Victim''") case True with 0 ** *** Some Pair have "f = SKIP" using victim_def by simp with ****** have "st''''= st''' \accounts := acc, stack := emptyStore, memory := emptyStore, gas := gas st''' - costs SKIP e\<^sub>p e' emptyStore (st'''\accounts := acc, stack := emptyStore, memory := emptyStore\)\" by (simp split:if_split_asm) with f3 show ?thesis by (simp add:frame_def) next case False with asm ad have "\st6' bal. frame bal (st'''\accounts := acc\) \ local.stmt f e\<^sub>p e' emptyStore (st'''\accounts := acc, stack := emptyStore, memory := emptyStore\) = Normal ((), st6') \ frame bal st6'" using asm Some Pair 43(3)[OF a1 ** _ _ _ *** _ _ _ _ _ ***** _ _ e'_def, where s'e = "st'''\accounts := acc, stack:=emptyStore, memory:=emptyStore\", of "KValue v" "Value t" t "KValue adv" "Value TAddr" TAddr st''' _ f st''' st''' "()"] using 0 by (simp add:frame_def) with f3 ****** show ?thesis by blast qed moreover from * ** *** Some Pair ***** ****** e'_def have st'_def: "st' = st''''\stack:=stack st''', memory := memory st'''\" by (simp split:if_split_asm) ultimately show "frame bal st'" apply safe by (simp_all add:frame_def) qed qed qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume ad: "address e = STR ''Victim''" show ?RHS (is "?A \ ?B \ ?C") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS (is "?A1 \ ?A2 \ ?A3 \ ?A4 \ ?A5") then have ?A1 and ?A2 and ?A3 and ?A4 and ?A5 by auto define st'' where "st'' = st\gas := gas st - costs transfer e\<^sub>p e cd st\" define st''' where "st''' = st''\gas := gas st'' - costs\<^sub>e (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st''\" define st'''' where "st'''' = st'''\gas := gas st''' - costs\<^sub>e SENDER e\<^sub>p e cd st'''\" from `?A1` * have c1: "gas st > costs transfer e\<^sub>p e cd st" by (auto split:if_split_asm) have c2: "gas st'' > costs\<^sub>e (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st''" proof (rule ccontr) assume "\ costs\<^sub>e (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st'' < gas st''" with c1 show False using `?A1` * st''_def st'''_def by auto qed with `?A4` `?A5` have 00:"expr (LVAL (Id (STR ''bal''))) e\<^sub>p e cd st'' = Normal ((KValue val, Value (TUInt 256)), st''')" using st'''_def st''_def by simp moreover have "gas st'''>costs\<^sub>e SENDER e\<^sub>p e cd st'''" proof (rule ccontr) assume "\ costs\<^sub>e SENDER e\<^sub>p e cd st''' < gas st'''" with c1 c2 show False using `?A1` `?A4` `?A5` * st''_def st'''_def by auto qed then have **:"expr SENDER e\<^sub>p e cd st''' = Normal ((KValue (sender e), Value TAddr), st'''')" using st''''_def by simp then obtain acc where ***:"Accounts.transfer (address e) (sender e) val (accounts st'''') = Some acc" and ****: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (ReadL\<^sub>i\<^sub>n\<^sub>t val)" proof - from `?A1` * c1 00 ** obtain acc where acc_def: "Accounts.transfer (address e) (sender e) val (accounts st'''') = Some acc" using st''''_def st'''_def st''_def by (auto split: option.split_asm) with ad obtain acc' where *: "subBalance (STR ''Victim'') val (accounts st'''') = Some acc'" and "addBalance (sender e) val acc' = Some acc" by (simp split: option.split_asm) moreover from * have "acc' = updateBalance(STR ''Victim'') (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - ReadL\<^sub>i\<^sub>n\<^sub>t val)) (accounts st'''')" by (simp split: if_split_asm) then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc' (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (ReadL\<^sub>i\<^sub>n\<^sub>t val)" using Read_ShowL_id by simp moreover from `?A5` ad have "sender e \ (STR ''Victim'')" by simp ultimately have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) = ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (ReadL\<^sub>i\<^sub>n\<^sub>t val)" using addBalance_eq[of "sender e" val acc' acc " STR ''Victim''"] by simp with acc_def show ?thesis using that by simp qed show ?RHS proof (cases "fmlookup e\<^sub>p (sender e)") case None with `?A1` 00 * ** *** have "st' = st''''\accounts := acc\" using c1 st''_def by auto moreover from `?A2` have "fmlookup (storage st'''') (STR ''Victim'') = Some s" using st''_def st'''_def st''''_def by simp moreover from `?A2` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) \ bal \ bal \ 0" using st''_def st'''_def st''''_def by simp with **** have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc\)) (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp ultimately show ?thesis using `?A3` by (simp add:frame_def) next case (Some a) then show ?thesis proof (cases a) case (Pair ct f) moreover define e' where e'_def: "e'=ffold_init ct (emptyEnv (sender e) (address e) val) (fmdom ct)" ultimately obtain st''''' where *****: "stmt f e\<^sub>p e' emptyStore (st''''\accounts := acc, stack:=emptyStore, memory:=emptyStore\) = Normal ((), st''''')" and ******: "st' = st'''''\stack:=stack st'''', memory := memory st''''\" using `?A1` 00 ** *** Some * stmt.simps(8)[of SENDER "(LVAL (Id (STR ''bal'')))" e\<^sub>p e cd st] c1 st''_def st'''_def st''''_def by (auto split: result.split_asm unit.split_asm) from `?A1` * have x1: "(applyf (costs (TRANSFER ad ex) e\<^sub>p e cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st'')" using st''_def by (simp split:if_split_asm) from 00 `?A1` have x2: "expr ex e\<^sub>p e cd st'' = Normal ((KValue val, Value (TUInt 256)), st''')" by simp have x3: "(KValue val, Value (TUInt 256)) = (KValue val, Value (TUInt 256))" by simp have x4: "KValue val = KValue val" by simp have x5: "Value (TUInt 256) = Value (TUInt 256)" by simp from ** `?A1` have x6: "expr ad e\<^sub>p e cd st''' = Normal ((KValue (sender e), Value TAddr), st'''')" by simp have x7: "(KValue (sender e), Value TAddr) = (KValue (sender e), Value TAddr)" by simp have x8: "KValue (sender e) = KValue (sender e)" by simp have x9: "Value TAddr = Value TAddr" by simp have x10: "TAddr = TAddr" by simp have x11: "applyf accounts st'''' = Normal (accounts st'''', st'''')" by simp from *** have x12: "Accounts.transfer (address e) (sender e) val (accounts st'''') = Some acc" by simp from Some Pair have x13: "fmlookup e\<^sub>p (sender e) = Some (ct,f)" by simp have x14: "(ct, f) = (ct, f)" by simp from e'_def have x15: "e' = ffold_init ct (emptyEnv (sender e) (address e) val) (fmdom ct)" by simp have x16: "get st'''' = Normal (st'''', st'''')" by simp have x17: "modify (\st. st\accounts := acc, stack := emptyStore, memory := emptyStore\) st'''' = Normal ((), st''''\accounts := acc, stack := emptyStore, memory := emptyStore\)" by simp from `?A2` have "fmlookup (storage st'''') (STR ''Victim'') = Some s" using st''_def st'''_def st''''_def by simp moreover from `?A2` have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (SUMM s + ReadL\<^sub>i\<^sub>n\<^sub>t val) \ bal \ bal \ 0" using st''_def st'''_def st''''_def by simp with **** have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance acc (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp then have "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts (st''''\accounts := acc, stack:=emptyStore, memory:=emptyStore\)) (STR ''Victim'')) - SUMM s \ bal \ bal \ 0" by simp moreover from `?A5` ad have "sender e \ (STR ''Victim'')" by simp with e'_def have "address e' \ STR ''Victim''" using ffold_init_ad_same[of ct "(emptyEnv (sender e) (address e) val)" "(fmdom ct)" e'] by simp ultimately have "frame bal st'''''" using 0 ***** 43(3)[OF x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17] `?A3` apply safe by (auto simp add:frame_def) with "******" show ?RHS by (auto simp add:frame_def) qed qed qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed qed qed qed qed qed qed next case (44 id0 tp ex smt e\<^sub>p e\<^sub>v cd st) show ?case (is "?LHS \ ?RHS") proof assume 0: "fmlookup e\<^sub>p STR ''Victim'' = Some (victim, SKIP)" show ?RHS (is "\st6'. ?RHS st6'") proof fix st6' show "?RHS st6'" (is "?LHS \ ?RHS") proof assume *: "stmt (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st = Normal ((), st6')" show ?RHS (is "?LHS \ ?RHS") proof show "?LHS" proof assume asm: "address e\<^sub>v \ STR ''Victim''" show "\bal. frame bal st \ frame bal st6'" proof fix bal show "frame bal st \ frame bal st6'" proof assume ff: "frame bal st" with * have a1: "(applyf (costs(BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)" by auto from * ff have f1: "frame bal (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)" by (simp add:frame_def) show "frame bal st6'" proof (cases ex) case None with * obtain cd' e' st' where **:"decl id0 tp None False cd (memory (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)) cd e\<^sub>v (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\) = Normal ((cd', e'), st')" by (auto split:result.split_asm if_split_asm) with * have f2: "frame bal st'" using decl_frame[OF f1 **] by simp moreover from * None ** obtain st'' where ***: "stmt smt e\<^sub>p e' cd' st' = Normal ((), st'')" by (simp split:if_split_asm) moreover from ** have ad: "address e' = address e\<^sub>v" using decl_gas_address by simp moreover from *** asm f2 ad 0 have "frame bal st''" using 44(3)[OF a1 None _ **, of cd' e'] by (simp add:frame_def) moreover from * None ** *** have "st6' = st''" by (auto split:if_split_asm) ultimately show ?thesis using f1 asm by auto next case (Some ex') with * obtain v t st' where **:"expr ex' e\<^sub>p e\<^sub>v cd (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\) = Normal ((v, t), st')" by (auto split:result.split_asm if_split_asm) with * f1 Some have f2: "frame bal st'" using 44(1)[OF a1 Some] asm 0 by simp moreover from Some * ** obtain cd' e' st'' where ***:"decl id0 tp (Some (v,t)) False cd (memory st') cd e\<^sub>v st' = Normal ((cd', e'), st'')" by (auto split:result.split_asm if_split_asm) with * have f3: "frame bal st''" using decl_frame[OF f2 ***] by simp moreover from ** *** have ad: "address e' = address e\<^sub>v" using decl_gas_address by simp moreover from * Some ** *** obtain st''' where ****: "stmt smt e\<^sub>p e' cd' st'' = Normal ((), st''')" by (simp split:if_split_asm) moreover from **** asm f3 ad have "frame bal st'''" using 44(2)[OF a1 Some ** _ _ ***] 0 by (simp add:frame_def) moreover from * Some ** *** **** have "st6' = st'''" by (auto split:if_split_asm) ultimately show ?thesis using f1 asm by auto qed qed qed qed next show "?RHS" (is "?LHS \ ?RHS") proof assume ad: "address e\<^sub>v = STR ''Victim''" show ?RHS (is "?A \ (?B \ ?C)") proof (rule conj3) show ?A (is "\s val bal x. ?LHS s val bal x") proof (rule allI[OF allI[OF allI[OF allI]]]) fix s val bal x show "?LHS s val bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?B (is "\s bal x. ?LHS s bal x") proof (rule allI[OF allI[OF allI]]) fix s bal x show "?LHS s bal x" (is "?LHS \ ?RHS") proof assume ?LHS then show ?RHS by simp qed qed next show ?C (is "\s bal. ?LHS s bal") proof (rule allI[OF allI]) fix s bal show "?LHS s bal" (is "?LHS \ ?RHS") proof assume ?LHS (is "?A1 \ ?A2 \ ?A3 \ ?A4 \ ?A5") then have ?A1 and ?A2 and ?A3 and ?A4 and ?A5 by auto define st'' where "st'' = st\gas := gas st - costs keep e\<^sub>p e\<^sub>v cd st\" with `?A2` `?A3` have 00: "fmlookup (storage st'') (STR ''Victim'') = Some s" and **: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'') (STR ''Victim'')) - (SUMM s) \ bal \ bal \ 0 \ POS s" by simp+ from `?A1` * st''_def obtain v t st''' cd' e' st'''' st''''' where ***: "expr mylval e\<^sub>p e\<^sub>v cd st'' = Normal ((v,t), st''')" and ****: "decl (STR ''bal'') (Value (TUInt 256)) (Some (v, t)) False cd (memory st''') cd e\<^sub>v st''' = Normal ((cd', e'),st'''')" and *****: "stmt comp e\<^sub>p e' cd' st'''' = Normal ((), st''''')" and "st6' = st'''''" by (auto split:if_split_asm result.split_asm) obtain s''' where f1: "fmlookup (storage st''') (STR ''Victim'') = Some s'''" and v_def: "v = KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''')" and t_def: "t = Value (TUInt 256)" and f2: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st''') (STR ''Victim'')) - (SUMM s''') \ bal \ bal \ 0 \ POS s'''" using securelval[OF *** `?A4` 00 ** ad] by auto with **** obtain s'''' where *******: "fmlookup (storage st'''') (STR ''Victim'') = Some s''''" and bbal: "ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st'''') (STR ''Victim'')) - (SUMM s'''') \ bal \ bal \ 0 \ POS s''''" using decl_frame frame_def by auto from ad `?A5` have ad2: "address e' = STR ''Victim''" and ss: "sender e'\address e'" using decl_gas_address[OF ****] by auto then obtain x where ******: "fmlookup (denvalue e') (STR ''bal'') = Some (Value (TUInt 256), (Stackloc x))" and lkup: "fmlookup (denvalue e') STR ''balance'' = Some (Storage (STMap TAddr (STValue (TUInt 256))), Storeloc STR ''balance'')" and "accessStore x (stack st'''') = Some (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''''))" proof - have "Valuetypes.convert (TUInt 256) (TUInt 256) (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''') = Some (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''', TUInt 256)" by simp with **** v_def t_def have "append (STR ''bal'') (Value (TUInt 256)) (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''')) cd e\<^sub>v st''' = Normal ((cd', e'),st'''')" by simp with f1 v_def t_def have st''''_def: "st'''' = st'''\stack := push v (stack st''')\" and "e' = updateEnv (STR ''bal'') t (Stackloc (ShowL\<^sub>n\<^sub>a\<^sub>t (toploc (stack st''')))) e\<^sub>v" by auto moreover from ******* f1 st''''_def have "Some (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''')) = Some (KValue (accessStorage (TUInt 256) (sender e\<^sub>v + (STR ''.'' + STR ''balance'')) s''''))" by simp ultimately show ?thesis using t_def v_def `?A4` that by simp qed with decl_gas_address **** have sck: "accessStore x (stack st'''') = Some (KValue (accessStorage (TUInt 256) (sender e' + (STR ''.'' + STR ''balance'')) s''''))" by simp from * have a1: "(applyf (costs(BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd) \ (\g. assert Gas (\st. gas st \ g) (modify (\st. st\gas := gas st - g\)))) st = Normal ((), st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\)" by auto from `?A1` have a2: "ex = Some (LVAL (Ref STR ''balance'' [SENDER]))" by simp from `?A1` *** have a3: "local.expr (LVAL (Ref STR ''balance'' [SENDER])) e\<^sub>p e\<^sub>v cd (st\gas := gas st - costs (BLOCK ((id0, tp), ex) smt) e\<^sub>p e\<^sub>v cd st\) = Normal ((v, t), st''')" using st''_def by simp from `?A1` **** have a4: "decl id0 tp (Some (v, t)) False cd (memory st''') cd e\<^sub>v st''' = Normal ((cd', e'), st'''')" by simp from `?A1` ***** `st6' = st'''''` have a5: "local.stmt smt e\<^sub>p e' cd' st'''' = Normal ((), st6')" by simp show "(\s'. fmlookup (storage st6') STR ''Victim'' = Some s' \ ReadL\<^sub>i\<^sub>n\<^sub>t (accessBalance (accounts st6') (STR ''Victim'')) - (SUMM s') \ bal \ bal \ 0 \ POS s')" using 44(2)[OF a1 a2 a3 _ _ a4, of cd' e'] 0 a5 ad2 `?A1` ******* bbal ****** lkup sck `?A4` ss apply safe by auto qed qed qed qed qed qed qed qed qed corollary final1: assumes "fmlookup ep (STR ''Victim'') = Some (victim, SKIP)" and "stmt (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] val) ep env cd st = Normal((), st')" and "address env \ (STR ''Victim'')" and "frame bal st" shows "frame bal st'" using assms secure(7)[of ep "(EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''withdraw'') [] val)" env cd st] by simp corollary final2: assumes "fmlookup ep (STR ''Victim'') = Some (victim, SKIP)" and "stmt (EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''deposit'') [] val) ep env cd st = Normal((), st')" and "address env \ (STR ''Victim'')" and "frame bal st" shows "frame bal st'" using assms secure(7)[of ep "(EXTERNAL (ADDRESS (STR ''Victim'')) (STR ''deposit'') [] val)" env cd st] by simp end end diff --git a/thys/VectorSpace/MonoidSums.thy b/thys/VectorSpace/MonoidSums.thy --- a/thys/VectorSpace/MonoidSums.thy +++ b/thys/VectorSpace/MonoidSums.thy @@ -1,118 +1,116 @@ section \Sums in monoids\ theory MonoidSums imports Main "HOL-Algebra.Module" RingModuleFacts FunctionLemmas begin text \We build on the finite product simplifications in FiniteProduct.thy and the analogous ones for finite sums (see "lemmas" in Ring.thy).\ text \Use as an intro rule\ lemma (in comm_monoid) factors_equal: "\a=b; c=d\ \ a\\<^bsub>G\<^esub>c = b\\<^bsub>G\<^esub>d" by simp lemma (in comm_monoid) extend_prod: fixes a A S assumes fin: "finite S" and subset: "A\S" and a: "a\A\carrier G" shows "(\\<^bsub>G\<^esub> x\S. (if x\A then a x else \\<^bsub>G\<^esub>)) = (\\<^bsub>G\<^esub> x\A. a x)" (is "(\\<^bsub>G\<^esub> x\S. ?b x) = (\\<^bsub>G\<^esub> x\A. a x)") proof - from subset have uni:"S = A \ (S-A)" by auto from assms subset show ?thesis apply (subst uni) apply (subst finprod_Un_disjoint, auto) by (auto cong: finprod_cong if_cong elim: finite_subset simp add:Pi_def finite_subset) (*Pi_def is key for automation here*) qed text \Scalar multiplication distributes over scalar multiplication (on left).\(*Add to module.*) lemma (in module) finsum_smult: "[| c\ carrier R; g \ A \ carrier M |] ==> (c \\<^bsub>M\<^esub> finsum M g A) = finsum M (%x. c \\<^bsub>M\<^esub> g x) A " proof (induct A rule: infinite_finite_induct) case (insert a A) from insert.hyps insert.prems have 1: "finsum M g (insert a A) = g a \\<^bsub>M\<^esub> finsum M g A" by (intro finsum_insert, auto) from insert.hyps insert.prems have 2: "(\\<^bsub>M\<^esub>x\insert a A. c \\<^bsub>M\<^esub> g x) = c \\<^bsub>M\<^esub> g a \\<^bsub>M\<^esub>(\\<^bsub>M\<^esub>x\A. c \\<^bsub>M\<^esub> g x)" by (intro finsum_insert, auto) from insert.hyps insert.prems show ?case by (auto simp add:1 2 smult_r_distr) qed auto text \Scalar multiplication distributes over scalar multiplication (on right).\(*Add to module.*) lemma (in module) finsum_smult_r: "[| v\ carrier M; f \ A \ carrier R |] ==> (finsum R f A \\<^bsub>M\<^esub> v) = finsum M (%x. f x \\<^bsub>M\<^esub> v) A " proof (induct A rule: infinite_finite_induct) case (insert a A) from insert.hyps insert.prems have 1: "finsum R f (insert a A) = f a \\<^bsub>R\<^esub> finsum R f A" by (intro R.finsum_insert, auto) from insert.hyps insert.prems have 2: "(\\<^bsub>M\<^esub>x\insert a A. f x \\<^bsub>M\<^esub> v) = f a \\<^bsub>M\<^esub> v \\<^bsub>M\<^esub>(\\<^bsub>M\<^esub>x\A. f x \\<^bsub>M\<^esub> v)" by (intro finsum_insert, auto) from insert.hyps insert.prems show ?case by (auto simp add:1 2 smult_l_distr) qed auto text \A sequence of lemmas that shows that the product does not depend on the ambient group. Note I had to dig back into the definitions of foldSet to show this.\ (*Add the following 2 lemmas to Group.*) lemma foldSet_not_depend: fixes A E assumes h1: "D\E" shows "foldSetD D f e \foldSetD E f e" proof - from h1 have 1: "\x1 x2. (x1,x2) \ foldSetD D f e \ (x1, x2) \ foldSetD E f e" proof - fix x1 x2 assume 2: "(x1,x2) \ foldSetD D f e" from h1 2 show "?thesis x1 x2" apply (intro foldSetD.induct[where ?D="D" and ?f="f" and ?e="e" and ?x1.0="x1" and ?x2.0="x2" and ?P = "\x1 x2. ((x1, x2)\ foldSetD E f e)"]) apply auto apply (intro emptyI, auto) by (intro insertI, auto) qed from 1 show ?thesis by auto qed lemma foldD_not_depend: fixes D E B f e A assumes h1: "LCD B D f" and h2: "LCD B E f" and h3: "D\E" and h4: "e\D" and h5: "A\B" and h6: "finite B" shows "foldD D f e A = foldD E f e A" proof - from assms have 1: "\y. (A,y)\foldSetD D f e" - apply (intro finite_imp_foldSetD, auto) - apply (metis finite_subset) - by (unfold LCD_def, auto) + by (intro finite_imp_foldSetD, auto simp: LCD_def) from 1 obtain y where 2: "(A,y)\foldSetD D f e" by auto from assms 2 have 3: "foldD D f e A = y" by (intro LCD.foldD_equality[of B], auto) from h3 have 4: "foldSetD D f e \ foldSetD E f e" by (rule foldSet_not_depend) from 2 4 have 5: "(A,y)\foldSetD E f e" by auto from assms 5 have 6: "foldD E f e A = y" by (intro LCD.foldD_equality[of B], auto) (*(A,y)\f*) from 3 6 show ?thesis by auto qed lemma (in comm_monoid) finprod_all1[simp]: assumes all1:" \a. a\A\f a = \\<^bsub>G\<^esub>" shows "(\\<^bsub>G\<^esub> a\A. f a) = \\<^bsub>G\<^esub>" (* "[| \a. a\A\f a = \\<^bsub>G\<^esub> |] ==> (\\<^bsub>G\<^esub> a\A. f a) = \\<^bsub>G\<^esub>" won't work with proof - *) proof - from assms show ?thesis by (simp cong: finprod_cong) qed context abelian_monoid begin lemmas summands_equal = add.factors_equal lemmas extend_sum = add.extend_prod lemmas finsum_all0 = add.finprod_all1 end end diff --git a/thys/VectorSpace/VectorSpace.thy b/thys/VectorSpace/VectorSpace.thy --- a/thys/VectorSpace/VectorSpace.thy +++ b/thys/VectorSpace/VectorSpace.thy @@ -1,1210 +1,1207 @@ section \Basic theory of vector spaces, using locales\ theory VectorSpace imports Main "HOL-Algebra.Module" "HOL-Algebra.Coset" RingModuleFacts MonoidSums LinearCombinations SumSpaces begin subsection \Basic definitions and facts carried over from modules\ text \A \vectorspace\ is a module where the ring is a field. Note that we switch notation from $(R, M)$ to $(K, V)$.\ locale vectorspace = module?: module K V + field?: field K for K and V (* Use sets for bases, and functions from the sets to carrier K represent the coefficients. *) text \A \subspace\ of a vectorspace is a nonempty subset that is closed under addition and scalar multiplication. These properties have already been defined in submodule. Caution: W is a set, while V is a module record. To get W as a vectorspace, write vs W.\ locale subspace = fixes K and W and V (structure) assumes vs: "vectorspace K V" and submod: "submodule K W V" lemma (in vectorspace) is_module[simp]: "subspace K W V\submodule K W V" by (unfold subspace_def, auto) text \We introduce some basic facts and definitions copied from module. We introduce some abbreviations, to match convention.\ abbreviation (in vectorspace) vs::"'c set \ ('a, 'c, 'd) module_scheme" where "vs W \ V\carrier :=W\" lemma (in vectorspace) carrier_vs_is_self [simp]: "carrier (vs W) = W" by auto lemma (in vectorspace) subspace_is_vs: fixes W::"'c set" assumes 0: "subspace K W V" shows "vectorspace K (vs W)" proof - from 0 show ?thesis apply (unfold vectorspace_def subspace_def, auto) by (intro submodule_is_module, auto) qed abbreviation (in module) subspace_sum:: "['c set, 'c set] \ 'c set" where "subspace_sum W1 W2 \submodule_sum W1 W2" lemma (in vectorspace) vs_zero_lin_dep: assumes h2: "S\carrier V" and h3: "lin_indpt S" shows "\\<^bsub>V\<^esub> \ S" proof - have vs: "vectorspace K V".. from vs have nonzero: "carrier K \{\\<^bsub>K\<^esub>}" by (metis one_zeroI zero_not_one) from h2 h3 nonzero show ?thesis by (rule zero_nin_lin_indpt) qed text \A \linear_map\ is a module homomorphism between 2 vectorspaces over the same field.\ locale linear_map = V?: vectorspace K V + W?: vectorspace K W + mod_hom?: mod_hom K V W T for K and V and W and T context linear_map begin lemmas T_hom = f_hom lemmas T_add = f_add lemmas T_smult = f_smult lemmas T_im = f_im lemmas T_neg = f_neg lemmas T_minus = f_minus lemmas T_ker = f_ker abbreviation imT:: "'e set" where "imT \ mod_hom.im" abbreviation kerT:: "'c set" where "kerT \ mod_hom.ker" lemmas T0_is_0[simp] = f0_is_0 lemma kerT_is_subspace: "subspace K ker V" proof - have vs: "vectorspace K V".. from vs show ?thesis apply (unfold subspace_def, auto) by (rule ker_is_submodule) qed lemma imT_is_subspace: "subspace K imT W" proof - have vs: "vectorspace K W".. from vs show ?thesis apply (unfold subspace_def, auto) by (rule im_is_submodule) qed end lemma vs_criteria: fixes K and V assumes field: "field K" and zero: "\\<^bsub>V\<^esub>\ carrier V" and add: "\v w. v\carrier V \ w\carrier V\ v\\<^bsub>V\<^esub> w\ carrier V" and neg: "\v\carrier V. (\neg_v\carrier V. v\\<^bsub>V\<^esub>neg_v=\\<^bsub>V\<^esub>)" and smult: "\c v. c\ carrier K \ v\carrier V\ c\\<^bsub>V\<^esub> v \ carrier V" and comm: "\v w. v\carrier V \ w\carrier V\ v\\<^bsub>V\<^esub> w=w\\<^bsub>V\<^esub> v" and assoc: "\v w x. v\carrier V \ w\carrier V \ x\carrier V\ (v\\<^bsub>V\<^esub> w)\\<^bsub>V\<^esub> x= v\\<^bsub>V\<^esub> (w\\<^bsub>V\<^esub> x)" and add_id: "\v\carrier V. (v\\<^bsub>V\<^esub> \\<^bsub>V\<^esub> =v)" and compat: "\a b v. a\ carrier K \ b\ carrier K \ v\carrier V\ (a\\<^bsub>K\<^esub> b)\\<^bsub>V\<^esub> v =a\\<^bsub>V\<^esub> (b\\<^bsub>V\<^esub> v)" and smult_id: "\v\carrier V. (\\<^bsub>K\<^esub> \\<^bsub>V\<^esub> v =v)" and dist_f: "\a b v. a\ carrier K \ b\ carrier K \ v\carrier V\ (a\\<^bsub>K\<^esub> b)\\<^bsub>V\<^esub> v =(a\\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> (b\\<^bsub>V\<^esub> v)" and dist_add: "\a v w. a\ carrier K \ v\carrier V \ w\carrier V\ a\\<^bsub>V\<^esub> (v\\<^bsub>V\<^esub> w) =(a\\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> (a\\<^bsub>V\<^esub> w)" shows "vectorspace K V" proof - from field have 1: "cring K" by (unfold field_def domain_def, auto) from assms 1 have 2: "module K V" by (intro module_criteria, auto) from field 2 show ?thesis by (unfold vectorspace_def module_def, auto) qed text \For any set $S$, the space of functions $S\to K$ forms a vector space.\ lemma (in vectorspace) func_space_is_vs: fixes S shows "vectorspace K (func_space S)" proof - have 0: "field K".. have 1: "module K (func_space S)" by (rule func_space_is_module) from 0 1 show ?thesis by (unfold vectorspace_def module_def, auto) qed lemma direct_sum_is_vs: fixes K V1 V2 assumes h1: "vectorspace K V1" and h2: "vectorspace K V2" shows "vectorspace K (direct_sum V1 V2)" proof - from h1 h2 have mod: "module K (direct_sum V1 V2)" by (unfold vectorspace_def, intro direct_sum_is_module, auto) from mod h1 show ?thesis by (unfold vectorspace_def, auto) qed lemma inj1_linear: fixes K V1 V2 assumes h1: "vectorspace K V1" and h2: "vectorspace K V2" shows "linear_map K V1 (direct_sum V1 V2) (inj1 V1 V2)" proof - from h1 h2 have mod: "mod_hom K V1 (direct_sum V1 V2) (inj1 V1 V2)" by (unfold vectorspace_def, intro inj1_hom, auto) from mod h1 h2 show ?thesis by (unfold linear_map_def vectorspace_def , auto, intro direct_sum_is_module, auto) qed lemma inj2_linear: fixes K V1 V2 assumes h1: "vectorspace K V1" and h2: "vectorspace K V2" shows "linear_map K V2 (direct_sum V1 V2) (inj2 V1 V2)" proof - from h1 h2 have mod: "mod_hom K V2 (direct_sum V1 V2) (inj2 V1 V2)" by (unfold vectorspace_def, intro inj2_hom, auto) from mod h1 h2 show ?thesis by (unfold linear_map_def vectorspace_def , auto, intro direct_sum_is_module, auto) qed text \For subspaces $V_1,V_2\subseteq V$, the map $V_1\oplus V_2\to V$ given by $(v_1,v_2)\mapsto v_1+v_2$ is linear.\ lemma (in vectorspace) sum_map_linear: fixes V1 V2 assumes h1: "subspace K V1 V" and h2: "subspace K V2 V" shows "linear_map K (direct_sum (vs V1) (vs V2)) V (\ v. (fst v) \\<^bsub>V\<^esub> (snd v))" proof - from h1 h2 have mod: "mod_hom K (direct_sum (vs V1) (vs V2)) V (\ v. (fst v) \\<^bsub>V\<^esub> (snd v))" by ( intro sum_map_hom, unfold subspace_def, auto) from mod h1 h2 show ?thesis apply (unfold linear_map_def, auto) apply (intro direct_sum_is_vs subspace_is_vs, auto).. qed lemma (in vectorspace) sum_is_subspace: fixes W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" shows "subspace K (subspace_sum W1 W2) V" proof - from h1 h2 have mod: "submodule K (submodule_sum W1 W2) V" by ( intro sum_is_submodule, unfold subspace_def, auto) from mod h1 h2 show ?thesis by (unfold subspace_def, auto) qed text \If $W_1,W_2\subseteq V$ are subspaces, $W_1\subseteq W_1 + W_2$\ lemma (in vectorspace) in_sum_vs: fixes W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" shows "W1 \ subspace_sum W1 W2" proof - from h1 h2 show ?thesis by (intro in_sum, unfold subspace_def, auto) qed lemma (in vectorspace) vsum_comm: fixes W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" shows "(subspace_sum W1 W2) = (subspace_sum W2 W1)" proof - from h1 h2 show ?thesis by (intro msum_comm, unfold subspace_def, auto) qed text \If $W_1,W_2\subseteq V$ are subspaces, then $W_1+W_2$ is the minimal subspace such that both $W_1\subseteq W$ and $W_2\subseteq W$.\ lemma (in vectorspace) vsum_is_minimal: fixes W W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" and h3: "subspace K W V" shows "(subspace_sum W1 W2) \ W \ W1 \ W \ W2 \ W" proof - from h1 h2 h3 show ?thesis by (intro sum_is_minimal, unfold subspace_def, auto) qed lemma (in vectorspace) span_is_subspace: fixes S assumes h2: "S\carrier V" shows "subspace K (span S) V" proof - have 0: "vectorspace K V".. from h2 have 1: "submodule K (span S) V" by (rule span_is_submodule) from 0 1 show ?thesis by (unfold subspace_def mod_hom_def linear_map_def, auto) qed subsubsection \Facts specific to vector spaces\ text \If $av = w$ and $a\neq 0$, $v=a^{-1}w$.\ lemma (in vectorspace) mult_inverse: assumes h1: "a\carrier K" and h2: "v\carrier V" and h3: "a \\<^bsub>V\<^esub> v = w" and h4: "a\\\<^bsub>K\<^esub>" shows "v = (inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub> w" proof - from h1 h2 h3 have 1: "w\carrier V" by auto from h3 1 have 2: "(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>(a \\<^bsub>V\<^esub> v) =(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>w" by auto from h1 h4 have 3: "inv\<^bsub>K\<^esub> a\carrier K" by auto interpret g: group "(units_group K)" by (rule units_form_group) have f: "field K".. from f h1 h4 have 4: "a\Units K" by (unfold field_def field_axioms_def, simp) from 4 h1 h4 have 5: "((inv\<^bsub>K\<^esub> a) \\<^bsub>K\<^esub>a) = \\<^bsub>K\<^esub>" by (intro Units_l_inv, auto) from 5 have 6: "(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>(a \\<^bsub>V\<^esub> v) = v" proof - from h1 h2 h4 have 7: "(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>(a \\<^bsub>V\<^esub> v) =(inv\<^bsub>K\<^esub> a \\<^bsub>K\<^esub>a) \\<^bsub>V\<^esub> v" by (auto simp add: smult_assoc1) from 5 h2 have 8: "(inv\<^bsub>K\<^esub> a \\<^bsub>K\<^esub>a) \\<^bsub>V\<^esub> v = v" by auto from 7 8 show ?thesis by auto qed from 2 6 show ?thesis by auto qed text \If $w\in S$ and $\sum_{w\in S} a_ww=0$, we have $v=\sum_{w\not\in S}a_v^{-1}a_ww$\ lemma (in vectorspace) lincomb_isolate: fixes A v assumes h1: "finite A" and h2: "A\carrier V" and h3: "a\A\carrier K" and h4: "v\A" and h5: "a v \ \\<^bsub>K\<^esub>" and h6: "lincomb a A=\\<^bsub>V\<^esub>" shows "v=lincomb (\w. \\<^bsub>K\<^esub>(inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> a w) (A-{v})" and "v\ span (A-{v})" proof - from h1 h2 h3 h4 have 1: "lincomb a A = ((a v) \\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (rule lincomb_del2) from 1 have 2: "\\<^bsub>V\<^esub>= ((a v) \\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (simp add: h6) from h1 h2 h3 have 5: "lincomb a (A-{v}) \carrier V" by auto (*intro lincomb_closed*) from 2 h1 h2 h3 h4 have 3: " \\<^bsub>V\<^esub> lincomb a (A-{v}) = ((a v) \\<^bsub>V\<^esub> v)" by (auto intro!: M.minus_equality) have 6: "v = (\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v))) \\<^bsub>V\<^esub> lincomb a (A-{v})" proof - from h2 h3 h4 h5 3 have 7: "v = inv\<^bsub>K\<^esub> (a v) \\<^bsub>V\<^esub> (\\<^bsub>V\<^esub> lincomb a (A-{v}))" by (intro mult_inverse, auto) from assms have 8: "inv\<^bsub>K\<^esub> (a v)\carrier K" by auto from assms 5 8 have 9: "inv\<^bsub>K\<^esub> (a v) \\<^bsub>V\<^esub> (\\<^bsub>V\<^esub> lincomb a (A-{v})) = (\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v))) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (simp add: smult_assoc_simp smult_minus_1_back r_minus) from 7 9 show ?thesis by auto qed from h1 have 10: "finite (A-{v})" by auto from assms have 11 : "(\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v)))\ carrier K" by auto from assms have 12: "lincomb (\w. \\<^bsub>K\<^esub>(inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> a w) (A-{v}) = (\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v))) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (intro lincomb_smult, auto) from 6 12 show "v=lincomb (\w. \\<^bsub>K\<^esub>(inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> a w) (A-{v})" by auto with assms show "v\ span (A-{v})" unfolding span_def by (force simp add: 11 ring_subset_carrier) qed text \The map $(S\to K)\mapsto V$ given by $(a_v)_{v\in S}\mapsto \sum_{v\in S} a_v v$ is linear.\ lemma (in vectorspace) lincomb_is_linear: fixes S assumes h: "finite S" and h2: "S\carrier V" shows "linear_map K (func_space S) V (\a. lincomb a S)" proof - have 0: "vectorspace K V".. from h h2 have 1: "mod_hom K (func_space S) V (\a. lincomb a S)" by (rule lincomb_is_mod_hom) from 0 1 show ?thesis by (unfold vectorspace_def mod_hom_def linear_map_def, auto) qed subsection \Basic facts about span and linear independence\ text \If $S$ is linearly independent, then $v\in \text{span}S$ iff $S\cup \{v\}$ is linearly dependent.\ theorem (in vectorspace) lin_dep_iff_in_span: fixes A v S assumes h1: "S \ carrier V" and h2: "lin_indpt S" and h3: "v\ carrier V" and h4: "v\S" shows "v\ span S \ lin_dep (S \ {v})" proof - let ?T = "S \ {v}" have 0: "v\?T " by auto from h1 h3 have h1_1: "?T \ carrier V" by auto have a1:"lin_dep ?T \ v\ span S" proof - assume a11: "lin_dep ?T" from a11 obtain a w A where a: "(finite A \ A\?T \ (a\ (A\carrier K)) \ (lincomb a A = \\<^bsub>V\<^esub>) \ (w\A) \ (a w\ \\<^bsub>K\<^esub>))" by (metis lin_dep_def) from assms a have nz2: "\v\A-S. a v\\\<^bsub>K\<^esub>" by (intro lincomb_must_include[where ?v="w" and ?T="S\{v}"], auto) from a nz2 have singleton: "{v}=A-S" by auto from singleton nz2 have nz3: "a v\\\<^bsub>K\<^esub>" by auto (*Can modularize this whole section out. "solving for one variable"*) let ?b="(\w. \\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> (a w))" from singleton have Ains: "(A\S) = A-{v}" by auto from assms a singleton nz3 have a31: "v= lincomb ?b (A\S)" apply (subst Ains) by (intro lincomb_isolate(1), auto) from a a31 nz3 singleton show ?thesis apply (unfold span_def, auto) apply (rule_tac x="?b" in exI) apply (rule_tac x="A\S" in exI) by (auto intro!: m_closed) qed have a2: "v\ (span S) \ lin_dep ?T" proof - assume inspan: "v\ (span S)" from inspan obtain a A where a: "A\S \ finite A \ (v = lincomb a A)\ a\A\carrier K" by (simp add: span_def, auto) let ?b = "\ w. if (w=v) then (\\<^bsub>K\<^esub> \\<^bsub>K\<^esub>) else a w" (*consider -v + \sum a_w w*) have lc0: " lincomb ?b (A\{v})=\\<^bsub>V\<^esub>" proof - from assms a have lc_ins: "lincomb ?b (A\{v}) = ((?b v) \\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> lincomb ?b A" by (intro lincomb_insert, auto) from assms a have lc_elim: "lincomb ?b A=lincomb a A" by (intro lincomb_elim_if, auto) from assms lc_ins lc_elim a show ?thesis by (simp add: M.l_neg smult_minus_1) qed from a lc0 show ?thesis apply (unfold lin_dep_def) apply (rule_tac x="A\{v}" in exI) apply (rule_tac x="?b" in exI) apply (rule_tac x="v" in exI) by auto qed from a1 a2 show ?thesis by auto qed text \If $v\in \text{span} A$ then $\text{span}A =\text{span}(A\cup \{v\})$\ lemma (in vectorspace) already_in_span: fixes v A assumes inC: "A\carrier V" and inspan: "v\span A" shows "span A= span (A\{v})" proof - from inC inspan have dir1: "span A \ span (A\{v})" by (intro span_is_monotone, auto) from inC have inown: "A\span A" by (rule in_own_span) from inC have subm: "submodule K (span A) V" by (rule span_is_submodule) from inown inspan subm have dir2: "span (A \ {v}) \ span A" by (intro span_is_subset, auto) from dir1 dir2 show ?thesis by auto qed subsection \The Replacement Theorem\ text \If $A,B\subseteq V$ are finite, $A$ is linearly independent, $B$ generates $W$, and $A\subseteq W$, then there exists $C\subseteq V$ disjoint from $A$ such that $\text{span}(A\cup C) = W$ and $|C|\le |B|-|A|$. In other words, we can complete any linearly independent set to a generating set of $W$ by adding at most $|B|-|A|$ more elements.\ theorem (in vectorspace) replacement: fixes A B (*A B are lists of vectors (colloquially we refer to them as sets)*) assumes h1: "finite A" and h2: "finite B" and h3: "B\carrier V" and h4: "lin_indpt A" (*A is linearly independent*) and h5: "A\span B" (*All entries of A are in K*) shows "\C. finite C \ C\carrier V \ C\span B \ C\A={} \ int (card C) \ (int (card B)) - (int (card A)) \ (span (A \ C) = span B)" (is "\C. ?P A B C") (*There is a set C of cardinality |B| - |A| such that A\C generates V*) using h1 h2 h3 h4 h5 proof (induct "card A" arbitrary: A B) case 0 from "0.prems"(1) "0.hyps" have a0: "A={}" by auto from "0.prems"(3) have a3: "B\span B" by (rule in_own_span) from a0 a3 "0.prems" show ?case by (rule_tac x="B" in exI, auto) next case (Suc m) let ?W="span B" from Suc.prems(3) have BinC: "span B\carrier V" by (rule span_is_subset2) (*everything you want to know about A*) from Suc.prems Suc.hyps BinC have A: "finite A" "lin_indpt A" "A\span B" "Suc m = card A" "A\carrier V" by auto (*everything you want to know about B*) from Suc.prems have B: "finite B" "B\carrier V" by auto (*A B are lists of vectors (colloquially we refer to them as sets)*) from Suc.hyps(2) obtain v where v: "v\A" by fastforce let ?A'="A-{v}" (*?A' is linearly independent because it is the subset of a linearly independent set, A.*) from A(2) have liA': "lin_indpt ?A'" apply (intro subset_li_is_li[of "A" "?A'"]) by auto from v liA' Suc.prems Suc.hyps(2) have "\C'. ?P ?A' B C'" apply (intro Suc.hyps(1)) by auto from this obtain C' where C': "?P ?A' B C'" by auto show ?case proof (cases "v\ C'") case True have vinC': "v\C'" by fact from vinC' v have seteq: "A - {v} \ C' = A \ (C' - {v})" by auto from C' seteq have spaneq: "span (A \ (C' - {v})) = span (B)" by algebra from Suc.prems Suc.hyps C' vinC' v spaneq show ?thesis apply (rule_tac x="C'-{v}" in exI) apply (subgoal_tac "card C' >0") by auto next case False have f: "v\C'" by fact from A v C' have "\a. a\(?A'\C')\carrier K \ lincomb a (?A' \ C') =v" by (intro finite_in_span, auto) from this obtain a where a: "a\(?A'\C')\carrier K \ v= lincomb a (?A' \ C')" by metis let ?b="(\ w. if (w=v) then \\<^bsub>K\<^esub>\\<^bsub>K\<^esub> else a w)" from a have b: "?b\A\C'\carrier K" by auto from v have rewrite_ins: "A\C'=(?A'\C')\{v}" by auto from f have "v\?A'\C'" by auto from this A C' v a f have lcb: "lincomb ?b (A \ C') = \\<^bsub>V\<^esub>" apply (subst rewrite_ins) apply (subst lincomb_insert) apply (simp_all add: ring_subset_carrier coeff_in_ring) apply (auto split: if_split_asm) apply (subst lincomb_elim_if) by (auto simp add: smult_minus_1 l_neg ring_subset_carrier) (*NOTE: l_neg got deleted from the simp rules, but it is very useful.*) from C' f have rewrite_minus: "C'=(A\C')-A" by auto from A C' b lcb v have exw: "\w\ C'. ?b w\\\<^bsub>K\<^esub>" apply (subst rewrite_minus) apply (intro lincomb_must_include[where ?T="A\C'" and ?v="v"]) by auto from exw obtain w where w: "w\ C'" "?b w\\\<^bsub>K\<^esub>" by auto from A C' w f b lcb have w_in: "w\span ((A\ C') -{w})" apply (intro lincomb_isolate[where a="?b"]) by auto have spaneq2: "span (A\(C'-{w})) = span B" proof - have 1: "span (?A'\C') = span (A\ C')" (*adding v doesn't change the span*) proof - from A C' v have m1: "span (?A'\C') = span ((?A'\ C')\{v})" apply (intro already_in_span) by auto from f m1 show ?thesis by (metis rewrite_ins) qed have 2: "span (A\ (C'-{w})) = span (A\ C')" (*removing w doesn't change the span*) proof - from C' w(1) f have b60: "A\ (C'-{w}) = (A\ C') -{w}" by auto from w(1) have b61: "A\ C'= (A\ C' -{w})\{w}" by auto from A C' w_in show ?thesis apply (subst b61) apply (subst b60) apply (intro already_in_span) by auto qed from C' 1 2 show ?thesis by auto qed(* "span (A\(C'-{w})) = span B"*) from A C' w f v spaneq2 show ?thesis apply (rule_tac x="C'-{w}" in exI) apply (subgoal_tac "card C' >0") by auto qed qed subsection \Defining dimension and bases.\ text \Finite dimensional is defined as having a finite generating set.\ definition (in vectorspace) fin_dim:: "bool" where "fin_dim = (\ A. ((finite A) \ (A \ carrier V) \ (gen_set A)))" text \The dimension is the size of the smallest generating set. For equivalent characterizations see below.\ definition (in vectorspace) dim:: "nat" where "dim = (LEAST n. (\ A. ((finite A) \ (card A = n) \ (A \ carrier V) \ (gen_set A))))" text \A \basis\ is a linearly independent generating set.\ definition (in vectorspace) basis:: "'c set \ bool" where "basis A = ((lin_indpt A) \ (gen_set A)\ (A\carrier V))" text \From the replacement theorem, any linearly independent set is smaller than any generating set.\ lemma (in vectorspace) li_smaller_than_gen: fixes A B assumes h1: "finite A" and h2: "finite B" and h3: "A\carrier V" and h4: "B\carrier V" and h5: "lin_indpt A" and h6: "gen_set B" shows "card A \ card B" proof - from h3 h6 have 1: "A\span B" by auto from h1 h2 h4 h5 1 obtain C where 2: "finite C \ C\carrier V \ C\span B \ C\A={} \ int (card C) \ int (card B) - int (card A) \ (span (A \ C) = span B)" by (metis replacement) from 2 show ?thesis by arith qed text \The dimension is the cardinality of any basis. (In particular, all bases are the same size.)\ lemma (in vectorspace) dim_basis: fixes A assumes fin: "finite A" and h2: "basis A" shows "dim = card A" proof - have 0: "\B m. ((finite B) \ (card B = m) \ (B \ carrier V) \ (gen_set B)) \ card A \ m" proof - fix B m assume 1: "((finite B) \ (card B = m) \ (B \ carrier V) \ (gen_set B))" from 1 fin h2 have 2: "card A \ card B" apply (unfold basis_def) apply (intro li_smaller_than_gen) by auto from 1 2 show "?thesis B m" by auto qed from fin h2 0 show ?thesis apply (unfold dim_def basis_def) apply (intro Least_equality) apply (rule_tac x="A" in exI) by auto qed (*can define more generally with posets*) text \A \maximal\ set with respect to $P$ is such that if $B\supseteq A$ and $P$ is also satisfied for $B$, then $B=A$.\ definition maximal::"'a set \ ('a set \ bool) \ bool" where "maximal A P = ((P A) \ (\B. B\A \ P B \ B = A))" text \A \minimal\ set with respect to $P$ is such that if $B\subseteq A$ and $P$ is also satisfied for $B$, then $B=A$.\ definition minimal::"'a set \ ('a set \ bool) \ bool" where "minimal A P = ((P A) \ (\B. B\A \ P B \ B = A))" text \A maximal linearly independent set is a generating set.\ lemma (in vectorspace) max_li_is_gen: fixes A assumes h1: "maximal A (\S. S\carrier V \ lin_indpt S)" shows "gen_set A" proof (rule ccontr) assume 0: "\(gen_set A)" from h1 have 1: " A\ carrier V \ lin_indpt A" by (unfold maximal_def, auto) from 1 have 2: "span A \ carrier V" by (intro span_is_subset2, auto) from 0 1 2 have 3: "\v. v\carrier V \ v \ (span A)" by auto from 3 obtain v where 4: "v\carrier V \ v \ (span A)" by auto have 5: "v\A" proof - from h1 1 have 51: "A\span A" apply (intro in_own_span) by auto from 4 51 show ?thesis by auto qed from lin_dep_iff_in_span have 6: "\S v. S \ carrier V\ lin_indpt S \ v\ carrier V \ v\S \ v\ span S \ (lin_indpt (S \ {v}))" by auto from 1 4 5 have 7: "lin_indpt (A \ {v})" apply (intro 6) by auto (* assumes h0:"finite S" and h1: "S \ carrier V" and h2: "lin_indpt S" and h3: "v\ carrier V" and h4: "v\S" shows "v\ span S \ \ (lin_indpt (S \ {v}))"*) have 9: "\(maximal A (\S. S\carrier V \ lin_indpt S))" proof - from 1 4 5 7 have 8: "(\B. A \ B \ B \ carrier V \ lin_indpt B \ B \ A)" apply (rule_tac x="A\{v}" in exI) by auto from 8 show ?thesis apply (unfold maximal_def) by simp qed from h1 9 show False by auto qed text \A minimal generating set is linearly independent.\ lemma (in vectorspace) min_gen_is_li: fixes A assumes h1: "minimal A (\S. S\carrier V \ gen_set S)" shows "lin_indpt A" proof (rule ccontr) assume 0: "\lin_indpt A" from h1 have 1: " A\ carrier V \ gen_set A" by (unfold minimal_def, auto) from 1 have 2: "span A = carrier V" by auto from 0 1 obtain a v A' where 3: "finite A' \ A'\A \ a \ A' \ carrier K \ LinearCombinations.module.lincomb V a A' = \\<^bsub>V\<^esub> \ v \ A' \ a v \ \\<^bsub>K\<^esub>" by (unfold lin_dep_def, auto) have 4: "gen_set (A-{v})" proof - from 1 3 have 5: "v\span (A'-{v})" apply (intro lincomb_isolate[where a="a" and v="v"]) by auto from 3 5 have 51: "v\span (A-{v})" apply (intro subsetD[where ?A="span (A'-{v})" and ?B="span (A-{v})" and ?c="v"]) by (intro span_is_monotone, auto) from 1 have 6: "A\span A" apply (intro in_own_span) by auto from 1 51 have 7: "span (A-{v}) = span ((A-{v})\{v})" apply (intro already_in_span) by auto from 3 have 8: "A = ((A-{v})\{v})" by auto from 2 7 8 have 9:"span (A-{v}) = carrier V" by auto (*can't use 3 directly :( *) from 9 show ?thesis by auto qed have 10: "\(minimal A (\S. S\carrier V \ gen_set S))" proof - from 1 3 4 have 11: "(\B. A \ B \ B \ carrier V \ gen_set B \ B \ A)" apply (rule_tac x="A-{v}" in exI) by auto from 11 show ?thesis apply (unfold minimal_def) by auto qed from h1 10 show False by auto qed text \Given that some finite set satisfies $P$, there is a minimal set that satisfies $P$.\ lemma minimal_exists: fixes A P assumes h1: "finite A" and h2: "P A" shows "\B. B\A \ minimal B P" using h1 h2 proof (induct "card A" arbitrary: A rule: less_induct) case (less A) show ?case proof (cases "card A = 0") case True from True less.hyps less.prems show ?thesis apply (rule_tac x="{}" in exI) apply (unfold minimal_def) by auto next case False show ?thesis proof (cases "minimal A P") case True then show ?thesis apply (rule_tac x="A" in exI) by auto next case False have 2: "\minimal A P" by fact from less.prems 2 have 3: "\B. P B \ B \ A \ B\A" apply (unfold minimal_def) by auto from 3 obtain B where 4: "P B \ B \ A \ B\A" by auto from 4 have 5: "card B < card A" by (metis less.prems(1) psubset_card_mono) from less.hyps less.prems 3 4 5 have 6: "\C\B. minimal C P" - apply (intro less.hyps) - apply auto - by (metis rev_finite_subset) + by (intro less.hyps, auto) from 6 obtain C where 7: "C\B \ minimal C P" by auto from 4 7 show ?thesis apply (rule_tac x="C" in exI) apply (unfold minimal_def) by auto qed qed qed text \If $V$ is finite-dimensional, then any linearly independent set is finite.\ lemma (in vectorspace) fin_dim_li_fin: assumes fd: "fin_dim" and li: "lin_indpt A" and inC: "A\carrier V" shows fin: "finite A" proof (rule ccontr) assume A: "\finite A" from fd obtain C where C: "finite C \ C\carrier V \ gen_set C" by (unfold fin_dim_def, auto) from A obtain B where B: "B\A\ finite B \ card B = card C + 1" by (metis infinite_arbitrarily_large) from B li have liB: "lin_indpt B" by (intro subset_li_is_li[where ?A="A" and ?B="B"], auto) from B C liB inC have "card B \ card C" by (intro li_smaller_than_gen, auto) from this B show False by auto qed text \If $V$ is finite-dimensional (has a finite generating set), then a finite basis exists.\ lemma (in vectorspace) finite_basis_exists: assumes h1: "fin_dim" shows "\\. finite \ \ basis \" proof - from h1 obtain A where 1: "finite A \ A\carrier V \ gen_set A" by (metis fin_dim_def) hence 2: "\\. \\A \ minimal \ (\S. S\carrier V \ gen_set S)" apply (intro minimal_exists) by auto then obtain \ where 3: "\\A \ minimal \ (\S. S\carrier V \ gen_set S)" by auto hence 4: "lin_indpt \" apply (intro min_gen_is_li) by auto moreover from 3 have 5: "gen_set \ \ \\carrier V" apply (unfold minimal_def) by auto moreover from 1 3 have 6: "finite \" by (auto simp add: finite_subset) ultimately show ?thesis apply (unfold basis_def) by auto qed text\ The proof is as follows. \begin{enumerate} \item Because $V$ is finite-dimensional, there is a finite generating set (we took this as our definition of finite-dimensional). \item Hence, there is a minimal $\beta \subseteq A$ such that $\beta$ generates $V$. \item $\beta$ is linearly independent because a minimal generating set is linearly independent. \end{enumerate} Finally, $\beta$ is a basis because it is both generating and linearly independent. \ text \Any linearly independent set has cardinality at most equal to the dimension.\ lemma (in vectorspace) li_le_dim: fixes A assumes fd: "fin_dim" and c: "A\carrier V" and l: "lin_indpt A" shows "finite A" "card A \ dim" proof - from fd c l show fa: "finite A" by (intro fin_dim_li_fin, auto) from fd obtain \ where 1: "finite \ \ basis \" by (metis finite_basis_exists) from assms fa 1 have 2: "card A \ card \" apply (intro li_smaller_than_gen, auto) by (unfold basis_def, auto) from assms 1 have 3: "dim = card \" by (intro dim_basis, auto) from 2 3 show "card A \ dim" by auto qed text \Any generating set has cardinality at least equal to the dimension.\ lemma (in vectorspace) gen_ge_dim: fixes A assumes fa: "finite A" and c: "A\carrier V" and l: "gen_set A" shows "card A \ dim" proof - from assms have fd: "fin_dim" by (unfold fin_dim_def, auto) from fd obtain \ where 1: "finite \ \ basis \" by (metis finite_basis_exists) from assms 1 have 2: "card A \ card \" apply (intro li_smaller_than_gen, auto) by (unfold basis_def, auto) from assms 1 have 3: "dim = card \" by (intro dim_basis, auto) from 2 3 show ?thesis by auto qed text \If there is an upper bound on the cardinality of sets satisfying $P$, then there is a maximal set satisfying $P$.\ lemma maximal_exists: fixes P B N assumes maxc: "\A. P A \ finite A \ (card A \N)" and b: "P B" shows "\A. finite A \ maximal A P" proof - (*take the maximal*) let ?S="{card A| A. P A}" let ?n="Max ?S" from maxc have 1:"finite ?S" apply (simp add: finite_nat_set_iff_bounded_le) by auto from 1 have 2: "?n\?S" by (metis (mono_tags, lifting) Collect_empty_eq Max_in b) from assms 2 have 3: "\A. P A \ finite A \ card A = ?n" by auto from 3 obtain A where 4: "P A \ finite A \ card A = ?n" by auto from 1 maxc have 5: "\A. P A \ finite A \ (card A \?n)" by (metis (mono_tags, lifting) Max.coboundedI mem_Collect_eq) from 4 5 have 6: "maximal A P" apply (unfold maximal_def) by (metis card_seteq) from 4 6 show ?thesis by auto qed text \Any maximal linearly independent set is a basis.\ lemma (in vectorspace) max_li_is_basis: fixes A assumes h1: "maximal A (\S. S\carrier V \ lin_indpt S)" shows "basis A" proof - from h1 have 1: "gen_set A" by (rule max_li_is_gen) from assms 1 show ?thesis by (unfold basis_def maximal_def, auto) qed text \Any minimal linearly independent set is a generating set.\ lemma (in vectorspace) min_gen_is_basis: fixes A assumes h1: "minimal A (\S. S\carrier V \ gen_set S)" shows "basis A" proof - from h1 have 1: "lin_indpt A" by (rule min_gen_is_li) from assms 1 show ?thesis by (unfold basis_def minimal_def, auto) qed text \Any linearly independent set with cardinality at least the dimension is a basis.\ lemma (in vectorspace) dim_li_is_basis: fixes A assumes fd: "fin_dim" and fa: "finite A" and ca: "A\carrier V" and li: "lin_indpt A" and d: "card A \ dim" (*\*) shows "basis A" proof - from fd have 0: "\S. S\carrier V \ lin_indpt S \ finite S \ card S \dim" by (auto intro: li_le_dim) (*fin_dim_li_fin*) from 0 assms have h1: "finite A \ maximal A (\S. S\carrier V \ lin_indpt S)" apply (unfold maximal_def) apply auto by (metis card_seteq eq_iff) from h1 show ?thesis by (auto intro: max_li_is_basis) qed text \Any generating set with cardinality at most the dimension is a basis.\ lemma (in vectorspace) dim_gen_is_basis: fixes A assumes fa: "finite A" and ca: "A\carrier V" and li: "gen_set A" and d: "card A \ dim" shows "basis A" proof - have 0: "\S. finite S\ S\carrier V \ gen_set S \ card S \dim" by (intro gen_ge_dim, auto) (*li_le_dim*) from 0 assms have h1: "minimal A (\S. finite S \ S\carrier V \ gen_set S)" - apply (unfold minimal_def) - apply auto - by (metis card_seteq eq_iff) + unfolding minimal_def + by (metis card_seteq le_antisym) (*slightly annoying: we have to get rid of "finite S" inside.*) from h1 have h: "\B. B \ A \ B \ carrier V \ LinearCombinations.module.gen_set K V B \ B = A" proof - fix B assume asm: "B \ A \ B \ carrier V \ LinearCombinations.module.gen_set K V B" from asm h1 have "finite B" apply (unfold minimal_def) apply (intro finite_subset[where ?A="B" and ?B="A"]) by auto from h1 asm this show "?thesis B" apply (unfold minimal_def) by simp qed from h1 h have h2: "minimal A (\S. S\carrier V \ gen_set S)" apply (unfold minimal_def) by presburger from h2 show ?thesis by (rule min_gen_is_basis) qed text \$\beta$ is a basis iff for all $v\in V$, there exists a unique $(a_v)_{v\in S}$ such that $\sum_{v\in S} a_v v=v$.\ lemma (in vectorspace) basis_criterion: assumes A_fin: "finite A" and AinC: "A\carrier V" shows "basis A \ (\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v))" proof - have 1: "\(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v)) \ \basis A" proof - assume "\(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v))" then obtain v where v: "v\ carrier V \ \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v)" by metis (*either there is more than 1 rep, or no reps*) from v have vinC: "v\carrier V" by auto from v have "\(\ a. a\A \\<^sub>E carrier K \ lincomb a A = v) \ (\ a b. a\A \\<^sub>E carrier K \ lincomb a A = v \ b\A \\<^sub>E carrier K \ lincomb b A = v \ a\b)" by metis then show ?thesis proof assume a: "\(\ a. a\A \\<^sub>E carrier K \ lincomb a A = v)" from A_fin AinC have "\a. a\A \ carrier K \ lincomb a A = lincomb (restrict a A) A" unfolding lincomb_def restrict_def by (simp cong: finsum_cong add: ring_subset_carrier coeff_in_ring) with a have "\(\ a. a\A \ carrier K \ lincomb a A = v)" by auto with A_fin AinC have "v\span A" using finite_in_span by blast with AinC v show "\(basis A)" by (unfold basis_def, auto) next assume a2: "(\ a b. a\A \\<^sub>E carrier K \ lincomb a A = v \ b\A \\<^sub>E carrier K \ lincomb b A = v \ a\b)" then obtain a b where ab: "a\A \\<^sub>E carrier K \ lincomb a A = v \ b\A \\<^sub>E carrier K \ lincomb b A = v \ a\b" by metis from ab obtain w where w: "w\A \ a w \ b w" apply (unfold PiE_def, auto) by (metis extensionalityI) let ?c="\ x. (if x\A then ((a x) \\<^bsub>K\<^esub> (b x)) else undefined)" from ab have a_fun: "a\A \ carrier K" and b_fun: "b\A \ carrier K" by (unfold PiE_def, auto) from w a_fun b_fun have abinC: "a w \carrier K" "b w \carrier K" by auto from abinC w have nz: "a w \\<^bsub>K\<^esub> b w \ \\<^bsub>K\<^esub>" by auto (*uses M.minus_other_side*) from A_fin AinC a_fun b_fun ab vinC have a_b: "LinearCombinations.module.lincomb V (\x. if x \ A then a x \\<^bsub>K\<^esub> b x else undefined) A = \\<^bsub>V\<^esub>" by (simp cong: lincomb_cong add: coeff_in_ring lincomb_diff) from A_fin AinC ab w v nz a_b have "lin_dep A" apply (intro lin_dep_crit[where ?A="A" and ?a="?c" and ?v="w"]) apply (auto simp add: PiE_def) by auto thus "\basis A" by (unfold basis_def, auto) qed qed have 2: "(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v)) \ basis A" proof - assume b1: "(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v))" (is "(\v. v\ carrier V \(\! a. ?Q a v))") from b1 have b2: "(\v. v\ carrier V \(\ a. a\A \ carrier K \ lincomb a A = v))" apply (unfold PiE_def) by blast from A_fin AinC b2 have "gen_set A" apply (unfold span_def) by blast from b1 have A_li: "lin_indpt A" proof - let ?z="\ x. (if (x\A) then \\<^bsub>K\<^esub> else undefined)" from A_fin AinC have zero: "?Q ?z \\<^bsub>V\<^esub>" by (unfold PiE_def extensional_def lincomb_def, auto simp add: ring_subset_carrier) (*uses finsum_all0*) from A_fin AinC show ?thesis proof (rule finite_lin_indpt2) fix a assume a_fun: "a \ A \ carrier K" and lc_a: "LinearCombinations.module.lincomb V a A = \\<^bsub>V\<^esub>" from a_fun have a_res: "restrict a A \ A \\<^sub>E carrier K" by auto from a_fun A_fin AinC lc_a have lc_a_res: "LinearCombinations.module.lincomb V (restrict a A) A = \\<^bsub>V\<^esub>" apply (unfold lincomb_def restrict_def) by (simp cong: finsum_cong2 add: coeff_in_ring ring_subset_carrier) from a_fun a_res lc_a_res zero b1 have "restrict a A = ?z" by auto from this show "\v\A. a v = \\<^bsub>K\<^esub>" apply (unfold restrict_def) by meson qed qed have A_gen: "gen_set A" proof - from AinC have dir1: "span A\carrier V" by (rule span_is_subset2) have dir2: "carrier V\span A" proof (auto) fix v assume v: "v\carrier V" from v b2 obtain a where "a\A \ carrier K \ lincomb a A = v" by auto from this A_fin AinC show "v\span A" by (subst finite_span, auto) qed from dir1 dir2 show ?thesis by auto qed from A_li A_gen AinC show "basis A" by (unfold basis_def, auto) qed from 1 2 show ?thesis by satx qed lemma (in linear_map) surj_imp_imT_carrier: assumes surj: "T` (carrier V) = carrier W" shows "(imT) = carrier W" by (simp add: surj im_def) subsection \The rank-nullity (dimension) theorem\ text \If $V$ is finite-dimensional and $T:V\to W$ is a linear map, then $\text{dim}(\text{im}(T))+ \text{dim}(\text{ker}(T)) = \text{dim} V$. Moreover, we prove that if $T$ is surjective linear-map between $V$ and $W$, where $V$ is finite-dimensional, then also $W$ is finite-dimensional.\ theorem (in linear_map) rank_nullity_main: assumes fd: "V.fin_dim" shows "(vectorspace.dim K (W.vs imT)) + (vectorspace.dim K (V.vs kerT)) = V.dim" "T ` (carrier V) = carrier W \ W.fin_dim" proof - \ \First interpret kerT, imT as vectorspaces\ have subs_ker: "subspace K kerT V" by (intro kerT_is_subspace) from subs_ker have vs_ker: "vectorspace K (V.vs kerT)" by (rule V.subspace_is_vs) from vs_ker interpret ker: vectorspace K "(V.vs kerT)" by auto have kerInC: "kerT\carrier V" by (unfold ker_def, auto) have subs_im: "subspace K imT W" by (intro imT_is_subspace) from subs_im have vs_im: "vectorspace K (W.vs imT)" by (rule W.subspace_is_vs) from vs_im interpret im: vectorspace K "(W.vs imT)" by auto have imInC: "imT\carrier W" by (unfold im_def, auto) (* obvious fact *) have zero_same[simp]: "\\<^bsub>V.vs kerT\<^esub> = \\<^bsub>V\<^esub>" apply (unfold ker_def) by auto \ \Show ker T has a finite basis. This is not obvious. Show that any linearly independent set has size at most that of V. There exists a maximal linearly independent set, which is the basis.\ have every_li_small: "\A. (A \ kerT)\ ker.lin_indpt A \ finite A \ card A \ V.dim" proof - fix A assume eli_asm: "(A \ kerT)\ ker.lin_indpt A" (*annoying: I can't just use subst V.module.span_li_not_depend(2) in the show ?thesis statement because it doesn't appear in the conclusion.*) note V.module.span_li_not_depend(2)[where ?N="kerT" and ?S="A"] from this subs_ker fd eli_asm kerInC show "?thesis A" apply (intro conjI) by (auto intro!: V.li_le_dim) qed from every_li_small have exA: "\A. finite A \ maximal A (\S. S\carrier (V.vs kerT) \ ker.lin_indpt S)" apply (intro maximal_exists[where ?N="V.dim" and ?B="{}"]) apply auto by (unfold ker.lin_dep_def, auto) from exA obtain A where A:" finite A \ maximal A (\S. S\carrier (V.vs kerT) \ ker.lin_indpt S)" by blast hence finA: "finite A" and Ainker: "A\carrier (V.vs kerT)" and AinC: "A\carrier V" by (unfold maximal_def ker_def, auto) \ \We obtain the basis A of kerT. It is also linearly independent when considered in V rather than kerT\ from A have Abasis: "ker.basis A" by (intro ker.max_li_is_basis, auto) from subs_ker Abasis have spanA: "V.module.span A = kerT" apply (unfold ker.basis_def) by (subst sym[OF V.module.span_li_not_depend(1)[where ?N="kerT"]], auto) from Abasis have Akerli: "ker.lin_indpt A" apply (unfold ker.basis_def) by auto from subs_ker Ainker Akerli have Ali: "V.module.lin_indpt A" by (auto simp add: V.module.span_li_not_depend(2)) txt\Use the replacement theorem to find C such that $A\cup C$ is a basis of V.\ from fd obtain B where B: "finite B\ V.basis B" by (metis V.finite_basis_exists) from B have Bfin: "finite B" and Bbasis:"V.basis B" by auto from B have Bcard: "V.dim = card B" by (intro V.dim_basis, auto) from Bbasis have 62: "V.module.span B = carrier V" by (unfold V.basis_def, auto) from A Abasis Ali B vs_ker have "\C. finite C \ C\carrier V \ C\ V.module.span B \ C\A={} \ int (card C) \ (int (card B)) - (int (card A)) \ (V.module.span (A \ C) = V.module.span B)" apply (intro V.replacement) apply (unfold vectorspace.basis_def V.basis_def) by (unfold ker_def, auto) txt \From replacement we got $|C|\leq |B|-|A|$. Equality must actually hold, because no generating set can be smaller than $B$. Now $A\cup C$ is a maximal generating set, hence a basis; its cardinality equals the dimension.\ txt \We claim that $T(C)$ is basis for $\text{im}(T)$.\ then obtain C where C: "finite C \ C\carrier V \ C\ V.module.span B \ C\A={} \ int (card C) \ (int (card B)) - (int (card A)) \ (V.module.span (A \ C) = V.module.span B)" by auto hence Cfin: "finite C" and CinC: "C\carrier V" and CinspanB: " C\V.module.span B" and CAdis: "C\A={}" and Ccard: "int (card C) \ (int (card B)) - (int (card A))" and ACspanB: "(V.module.span (A \ C) = V.module.span B)" by auto from C have cardLe: "card A + card C \ card B" by auto from B C have ACgen: "V.module.gen_set (A\C)" apply (unfold V.basis_def) by auto from finA C ACgen AinC B have cardGe: "card (A\C) \ card B" by (intro V.li_smaller_than_gen, unfold V.basis_def, auto) from finA C have cardUn: "card (A\C)\ card A + card C" by (metis Int_commute card_Un_disjoint le_refl) from cardLe cardUn cardGe Bcard have cardEq: "card (A\C) = card A + card C" "card (A\C) = card B" "card (A\C) = V.dim" by auto from Abasis C cardEq have disj: "A\C={}" by auto from finA AinC C cardEq 62 have ACfin: "finite (A\C)" and ACbasis: "V.basis (A\C)" by (auto intro!: V.dim_gen_is_basis) have lm: "linear_map K V W T".. txt \Let $C'$ be the image of $C$ under $T$. We will show $C'$ is a basis for $\text{im}(T)$.\ let ?C' = "T`C" from Cfin have C'fin: "finite ?C'" by auto from AinC C have cim: "?C'\imT" by (unfold im_def, auto) txt \"There is a subtle detail: we first have to show $T$ is injective on $C$.\ txt \We establish that no nontrivial linear combination of $C$ can have image 0 under $T$, because that would mean it is a linear combination of $A$, giving that $A\cup C$ is linearly dependent, contradiction. We use this result in 2 ways: (1) if $T$ is not injective on $C$, then we obtain $v$, $w\in C$ such that $v-w$ is in the kernel, contradiction, (2) if $T(C)$ is linearly dependent, taking the inverse image of that linear combination gives a linear combination of $C$ in the kernel, contradiction. Hence $T$ is injective on $C$ and $T(C)$ is linearly independent.\ have lc_in_ker: "\d D v. \D\C; d\D\carrier K; T (V.module.lincomb d D) = \\<^bsub>W\<^esub>; v\D; d v \\\<^bsub>K\<^esub>\\False" proof - fix d D v assume D: "D\C" and d: "d\D\carrier K" and T0: "T (V.module.lincomb d D) = \\<^bsub>W\<^esub>" and v: "v\D" and dvnz: "d v \\\<^bsub>K\<^esub>" from D Cfin have Dfin: "finite D" by (auto intro: finite_subset) from D CinC have DinC: "D\carrier V" by auto from T0 d Dfin DinC have lc_d: "V.module.lincomb d D\kerT" by (unfold ker_def, auto) from lc_d spanA AinC have "\a' A'. A'\A \ a'\A'\carrier K \ V.module.lincomb a' A'= V.module.lincomb d D" by (intro V.module.in_span, auto) then obtain a' A' where a': "A'\A \ a'\A'\carrier K \ V.module.lincomb d D = V.module.lincomb a' A'" by metis hence A'sub: "A'\A" and a'fun: "a'\A'\carrier K" and a'_lc:"V.module.lincomb d D = V.module.lincomb a' A'" by auto from a' finA Dfin have A'fin: "finite (A')" by (auto intro: finite_subset) from AinC A'sub have A'inC: "A'\carrier V" by auto let ?e = "(\v. if v \ A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v)" from a'fun d have e_fun: "?e \ A' \ D \ carrier K" apply (unfold Pi_def) by auto from A'fin Dfin (*finiteness*) A'inC DinC (*in carrier*) a'fun d e_fun (*coefficients valid*) disj D A'sub (*A and C disjoint*) have lccomp1: "V.module.lincomb a' A' \\<^bsub>V\<^esub> \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>V\<^esub> V.module.lincomb d D = V.module.lincomb (\v. if v\A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v) (A'\D)" apply (subst sym[OF V.module.lincomb_smult]) apply (simp_all) apply (subst V.module.lincomb_union2) by (auto) from A'fin (*finiteness*) A'inC (*in carrier*) a'fun (*coefficients valid*) have lccomp2: "V.module.lincomb a' A' \\<^bsub>V\<^esub> \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>V\<^esub> V.module.lincomb d D = \\<^bsub>V\<^esub>" by (simp add: a'_lc V.module.smult_minus_1 V.module.M.r_neg) from lccomp1 lccomp2 have lc0: "V.module.lincomb (\v. if v\A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v) (A'\D) =\\<^bsub>V\<^esub>" by auto from disj a' v D have v_nin: "v\A'" by auto from A'fin Dfin (*finiteness*) A'inC DinC (*in carrier*) e_fun d (*coefficients valid*) A'sub D disj (*A' D are disjoint subsets*) v dvnz (*d v is nonzero coefficient*) lc0 have AC_ld: "V.module.lin_dep (A\C)" apply (intro V.module.lin_dep_crit[where ?A="A'\D" and ?S="A\C" and ?a="\v. if v\A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v" and ?v="v"]) by (auto dest: integral) from AC_ld ACbasis show False by (unfold V.basis_def, auto) qed have C'_card: "inj_on T C" "card C = card ?C'" proof - show "inj_on T C" proof (rule ccontr) assume "\inj_on T C" then obtain v w where "v\C" "w\C" "v\w" "T v = T w" by (unfold inj_on_def, auto) from this CinC show False apply (intro lc_in_ker[where ?D1="{v,w}" and ?d1="\x. if x=v then \\<^bsub>K\<^esub> else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>" and ?v1="v"]) by (auto simp add: V.module.lincomb_def hom_sum ring_subset_carrier W.module.smult_minus_1 r_neg T_im) qed from this Cfin show "card C = card ?C'" by (metis card_image) qed let ?f="the_inv_into C T" have f: "\x. x\C \ ?f (T x) = x" "\y. y\?C' \ T (?f y) = y" apply (insert C'_card(1)) apply (metis the_inv_into_f_f) by (metis f_the_inv_into_f) (*We show C' is a basis for the image. First we show it is linearly independent.*) have C'_li: "im.lin_indpt ?C'" proof (rule ccontr) assume Cld: "\im.lin_indpt ?C'" from Cld cim subs_im have CldW: "W.module.lin_dep ?C'" apply (subst sym[OF W.module.span_li_not_depend(2)[where ?S="T`C" and ?N="imT"]]) by auto from C CldW have "\c' v'. (c'\ (?C'\carrier K)) \ (W.module.lincomb c' ?C' = \\<^bsub>W\<^esub>) \ (v'\?C') \ (c' v'\ \\<^bsub>K\<^esub>)" by (intro W.module.finite_lin_dep, auto) then obtain c' v' where c': "(c'\ (?C'\carrier K)) \ (W.module.lincomb c' ?C' = \\<^bsub>W\<^esub>) \ (v'\?C') \ (c' v'\ \\<^bsub>K\<^esub>)" by auto hence c'fun: "(c'\ (?C'\carrier K))" and c'lc: "(W.module.lincomb c' ?C' = \\<^bsub>W\<^esub>)" and v':"(v'\?C')" and cvnz: "(c' v'\ \\<^bsub>K\<^esub>)" by auto (*can't get c' directly with metis/auto with W.module.finite_lin_dep*) txt \We take the inverse image of $C'$ under $T$ to get a linear combination of $C$ that is in the kernel and hence a linear combination of $A$. This contradicts $A\cup C$ being linearly independent.\ let ?c="\v. c' (T v)" from c'fun have c_fun: "?c\ C\carrier K" by auto from Cfin (*C finite*) c_fun c'fun (*coefficients valid*) C'_card (*bijective*) CinC (*C in carrier*) f (*inverse to T*) c'lc (*lc c' = 0*) have "T (V.module.lincomb ?c C) = \\<^bsub>W\<^esub>" apply (unfold V.module.lincomb_def W.module.lincomb_def) apply (subst hom_sum, auto) apply (simp cong: finsum_cong add: ring_subset_carrier coeff_in_ring) apply (subst finsum_reindex[where ?f="\w. c' w \\<^bsub>W\<^esub> w" and ?h="T" and ?A="C", THEN sym]) by auto with f c'fun cvnz v' show False by (intro lc_in_ker[where ?D1="C" and ?d1="?c" and ?v1="?f v'"], auto) qed have C'_gen: "im.gen_set ?C'" proof - have C'_span: "span ?C' = imT" proof (rule equalityI) from cim subs_im show "W.module.span ?C' \ imT" by (intro span_is_subset, unfold subspace_def, auto) next show "imT\W.module.span ?C'" proof (auto) fix w assume w: "w\imT" from this finA Cfin AinC CinC obtain v where v_inC: "v\carrier V" and w_eq_T_v: "w= T v" by (unfold im_def image_def, auto) from finA Cfin AinC CinC v_inC ACgen have "\a. a \ A\C \ carrier K\ V.module.lincomb a (A\C) = v" by (intro V.module.finite_in_span, auto) then obtain a where a_fun: "a \ A\C \ carrier K" and lc_a_v: "v= V.module.lincomb a (A\C)" by auto let ?a'="\v. a (?f v)" from finA Cfin AinC CinC a_fun disj Ainker f C'_card have Tv: "T v = W.module.lincomb ?a' ?C'" apply (subst lc_a_v) apply (subst V.module.lincomb_union, simp_all) (*Break up the union A\C*) apply (unfold lincomb_def V.module.lincomb_def) apply (subst hom_sum, auto) (*Take T inside the sum over A*) apply (simp add: subsetD coeff_in_ring hom_sum (*Take T inside the sum over C*) T_ker (*all terms become 0 because the vectors are in the kernel.*) ) apply (subst finsum_reindex[where ?h="T" and ?f="\v. ?a' v\\<^bsub>W\<^esub> v"], auto) by (auto cong: finsum_cong simp add: coeff_in_ring ring_subset_carrier) from a_fun f have a'_fun: "?a'\?C' \ carrier K" by auto from C'fin CinC this w_eq_T_v a'_fun Tv show "w \ LinearCombinations.module.span K W (T ` C)" by (subst finite_span, auto) qed qed from this subs_im CinC show ?thesis apply (subst span_li_not_depend(1)) by (unfold im_def subspace_def, auto) qed from C'_li C'_gen C cim have C'_basis: "im.basis (T`C)" by (unfold im.basis_def, auto) have C_card_im: "card C = (vectorspace.dim K (W.vs imT))" using C'_basis C'_card(2) C'fin im.dim_basis by auto from finA Abasis have "ker.dim = card A" by (rule ker.dim_basis) note * = this C_card_im cardEq show "(vectorspace.dim K (W.vs imT)) + (vectorspace.dim K (V.vs kerT)) = V.dim" using * by auto assume "T` (carrier V) = carrier W" from * surj_imp_imT_carrier[OF this] show "W.fin_dim" using C'_basis C'fin unfolding W.fin_dim_def im.basis_def by auto qed theorem (in linear_map) rank_nullity: assumes fd: "V.fin_dim" shows "(vectorspace.dim K (W.vs imT)) + (vectorspace.dim K (V.vs kerT)) = V.dim" by (rule rank_nullity_main[OF fd]) end diff --git a/thys/Verified_SAT_Based_AI_Planning/SAS_Plus_STRIPS.thy b/thys/Verified_SAT_Based_AI_Planning/SAS_Plus_STRIPS.thy --- a/thys/Verified_SAT_Based_AI_Planning/SAS_Plus_STRIPS.thy +++ b/thys/Verified_SAT_Based_AI_Planning/SAS_Plus_STRIPS.thy @@ -1,4403 +1,4402 @@ (* Author: Mohammad Abdulaziz, Fred Kurz *) theory SAS_Plus_STRIPS imports "STRIPS_Semantics" "SAS_Plus_Semantics" "Map_Supplement" begin section "SAS+/STRIPS Equivalence" text \ The following part is concerned with showing the equivalent expressiveness of SAS+ and STRIPS as discussed in \autoref{sub:equivalence-sas-plus-strips}. \ subsection "Translation of SAS+ Problems to STRIPS Problems" definition possible_assignments_for :: "('variable, 'domain) sas_plus_problem \ 'variable \ ('variable \ 'domain) list" where "possible_assignments_for \ v \ [(v, a). a \ the (range_of \ v)]" definition all_possible_assignments_for :: "('variable, 'domain) sas_plus_problem \ ('variable \ 'domain) list" where "all_possible_assignments_for \ \ concat [possible_assignments_for \ v. v \ variables_of \]" definition state_to_strips_state :: "('variable, 'domain) sas_plus_problem \ ('variable, 'domain) state \ ('variable, 'domain) assignment strips_state" ("\\<^sub>S _ _" 99) where "state_to_strips_state \ s \ let defined = filter (\v. s v \ None) (variables_of \) in map_of (map (\(v, a). ((v, a), the (s v) = a)) (concat [possible_assignments_for \ v. v \ defined]))" definition sasp_op_to_strips :: "('variable, 'domain) sas_plus_problem \ ('variable, 'domain) sas_plus_operator \ ('variable, 'domain) assignment strips_operator" ("\\<^sub>O _ _" 99) where "sasp_op_to_strips \ op \ let pre = precondition_of op ; add = effect_of op ; delete = [(v, a'). (v, a) \ effect_of op, a' \ filter ((\) a) (the (range_of \ v))] in STRIPS_Representation.operator_for pre add delete" definition sas_plus_problem_to_strips_problem :: "('variable, 'domain) sas_plus_problem \ ('variable, 'domain) assignment strips_problem" ("\ _ " 99) where "sas_plus_problem_to_strips_problem \ \ let vs = [as. v \ variables_of \, as \ (possible_assignments_for \) v] ; ops = map (sasp_op_to_strips \) (operators_of \) ; I = state_to_strips_state \ (initial_of \) ; G = state_to_strips_state \ (goal_of \) in STRIPS_Representation.problem_for vs ops I G" definition sas_plus_parallel_plan_to_strips_parallel_plan :: "('variable, 'domain) sas_plus_problem \ ('variable, 'domain) sas_plus_parallel_plan \ ('variable \ 'domain) strips_parallel_plan" ("\\<^sub>P _ _" 99) where "sas_plus_parallel_plan_to_strips_parallel_plan \ \ \ [[sasp_op_to_strips \ op. op \ ops]. ops \ \]" (* TODO first argument should be ('variable, 'domain) strips_problem *) definition strips_state_to_state :: "('variable, 'domain) sas_plus_problem \ ('variable, 'domain) assignment strips_state \ ('variable, 'domain) state" ("\\<^sub>S\ _ _" 99) where "strips_state_to_state \ s \ map_of (filter (\(v, a). s (v, a) = Some True) (all_possible_assignments_for \))" (* TODO remove problem argument *) definition strips_op_to_sasp :: "('variable, 'domain) sas_plus_problem \ ('variable \ 'domain) strips_operator \ ('variable, 'domain) sas_plus_operator" ("\\<^sub>O\ _ _" 99) where "strips_op_to_sasp \ op \ let precondition = strips_operator.precondition_of op ; effect = strips_operator.add_effects_of op in \ precondition_of = precondition, effect_of = effect \" (* TODO \strips_parallel_plan_to_sas_plus_parallel_plan \ \_P\\ and \strips_op_to_sasp \ \_O\\ *) definition strips_parallel_plan_to_sas_plus_parallel_plan :: "('variable, 'domain) sas_plus_problem \ ('variable \ 'domain) strips_parallel_plan \ ('variable, 'domain) sas_plus_parallel_plan" ("\\<^sub>P\ _ _" 99) where "strips_parallel_plan_to_sas_plus_parallel_plan \ \ \ [[strips_op_to_sasp \ op. op \ ops]. ops \ \]" text \ To set up the equivalence proof context, we declare a common locale \isaname{sas_plus_strips_equivalence} for both the STRIPS and SAS+ formalisms and make it a sublocale of both locale \isaname{strips} as well as \isaname{sas_plus}. The declaration itself is omitted for brevity since it basically just joins locales \isaname{sas_plus} and \isaname{strips} while renaming the locale parameter to avoid name clashes. The sublocale proofs are shown below. \footnote{We append a suffix identifying the respective formalism to the the parameter names passed to the parameter names in the locale. This is necessary to avoid ambiguous names in the sublocale declarations. For example, without addition of suffixes the type for \initial_of\ is ambiguous and will therefore not be bound to either \strips_problem.initial_of\ or \sas_plus_problem.initial_of\. Isabelle in fact considers it to be a a free variable in this case. We also qualify the parent locales in the sublocale declarations by adding \texttt{strips:} and \texttt{sas\_plus:} before the respective parent locale identifiers. } \ definition "range_of_strips \ x \ { True, False }" context begin \ \ Set-up simp rules. \ lemma[simp]: "(\ \) = (let vs = [as. v \ variables_of \, as \ (possible_assignments_for \) v] ; ops = map (sasp_op_to_strips \) (operators_of \) ; I = state_to_strips_state \ (initial_of \) ; G = state_to_strips_state \ (goal_of \) in STRIPS_Representation.problem_for vs ops I G)" and "(\\<^sub>S \ s) = (let defined = filter (\v. s v \ None) (variables_of \) in map_of (map (\(v, a). ((v, a), the (s v) = a)) (concat [possible_assignments_for \ v. v \ defined])))" and "(\\<^sub>O \ op) = (let pre = precondition_of op ; add = effect_of op ; delete = [(v, a'). (v, a) \ effect_of op, a' \ filter ((\) a) (the (range_of \ v))] in STRIPS_Representation.operator_for pre add delete)" and "(\\<^sub>P \ \) = [[\\<^sub>O \ op. op \ ops]. ops \ \]" and "(\\<^sub>S\ \ s')= map_of (filter (\(v, a). s' (v, a) = Some True) (all_possible_assignments_for \))" and "(\\<^sub>O\ \ op') = (let precondition = strips_operator.precondition_of op' ; effect = strips_operator.add_effects_of op' in \ precondition_of = precondition, effect_of = effect \)" and "(\\<^sub>P\ \ \) = [[\\<^sub>O\ \ op. op \ ops]. ops \ \]" unfolding SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.state_to_strips_state_def state_to_strips_state_def SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sas_plus_parallel_plan_to_strips_parallel_plan_def SAS_Plus_STRIPS.strips_state_to_state_def strips_state_to_state_def SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_parallel_plan_to_sas_plus_parallel_plan_def by blast+ lemmas [simp] = range_of'_def lemma is_valid_problem_sas_plus_dom_sas_plus_problem_range_of: assumes "is_valid_problem_sas_plus \" shows "\v \ set ((\)\<^sub>\\<^sub>+). v \ dom (sas_plus_problem.range_of \)" using assms(1) is_valid_problem_sas_plus_then(1) unfolding is_valid_problem_sas_plus_def by (meson domIff list.pred_set) lemma possible_assignments_for_set_is: assumes "v \ dom (sas_plus_problem.range_of \)" shows "set (possible_assignments_for \ v) = { (v, a) | a. a \ \\<^sub>+ \ v }" proof - have "sas_plus_problem.range_of \ v \ None" using assms(1) by auto thus ?thesis unfolding possible_assignments_for_def by fastforce qed lemma all_possible_assignments_for_set_is: assumes "\v \ set ((\)\<^sub>\\<^sub>+). range_of \ v \ None" shows "set (all_possible_assignments_for \) = (\v \ set ((\)\<^sub>\\<^sub>+). { (v, a) | a. a \ \\<^sub>+ \ v })" proof - let ?vs = "variables_of \" have "set (all_possible_assignments_for \) = (\(set ` (\v. map (\(v, a). (v, a)) (possible_assignments_for \ v)) ` set ?vs))" unfolding all_possible_assignments_for_def set_concat using set_map by auto also have "\ = (\((\v. set (possible_assignments_for \ v)) ` set ?vs))" using image_comp set_map by simp (* TODO slow *) also have "\ = (\((\v. { (v, a) | a. a \ \\<^sub>+ \ v }) ` set ?vs))" using possible_assignments_for_set_is assms by fastforce finally show ?thesis by force qed lemma state_to_strips_state_dom_is_i[simp]: assumes "\v \ set ((\)\<^sub>\\<^sub>+). v \ dom (sas_plus_problem.range_of \)" shows "set (concat [possible_assignments_for \ v. v \ filter (\v. s v \ None) (variables_of \)]) = (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }. { (v, a) | a. a \ \\<^sub>+ \ v })" proof - let ?vs = "variables_of \" let ?defined = "filter (\v. s v \ None) ?vs" let ?l = "concat [possible_assignments_for \ v. v \ ?defined]" have nb: "set ?defined = { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }" unfolding set_filter by force have "set ?l = \(set ` set (map (possible_assignments_for \) ?defined ))" unfolding set_concat image_Union by blast also have "\ = \(set ` (possible_assignments_for \) ` set ?defined)" unfolding set_map by blast also have "\ = (\v \ set ?defined. set (possible_assignments_for \ v))" by blast also have "\ = (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }. set (possible_assignments_for \ v))" using nb by argo finally show ?thesis using possible_assignments_for_set_is is_valid_problem_sas_plus_dom_sas_plus_problem_range_of assms(1) by fastforce qed lemma state_to_strips_state_dom_is: \ \ NOTE A transformed state is defined on all possible assignments for all variables defined in the original state. \ assumes "is_valid_problem_sas_plus \" shows "dom (\\<^sub>S \ s) = (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }. { (v, a) | a. a \ \\<^sub>+ \ v })" proof - let ?vs = "variables_of \" let ?l = "concat [possible_assignments_for \ v. v \ filter (\v. s v \ None) ?vs]" have nb: "\v \ set ((\)\<^sub>\\<^sub>+). v \ dom (sas_plus_problem.range_of \)" using is_valid_problem_sas_plus_dom_sas_plus_problem_range_of assms(1) by fastforce have "dom (\\<^sub>S \ s) = fst ` set (map (\(v, a). ((v, a), the (s v) = a)) ?l)" unfolding state_to_strips_state_def SAS_Plus_STRIPS.state_to_strips_state_def using dom_map_of_conv_image_fst[of "map (\(v, a). ((v, a), the (s v) = a)) ?l"] by presburger also have "\ = fst ` (\(v, a). ((v, a), the (s v) = a)) ` set ?l" unfolding set_map by blast also have "\ = (\(v, a). fst ((v, a), the (s v) = a)) ` set ?l" unfolding image_comp[of fst "\(v, a). ((v, a), the (s v) = a)"] comp_apply[of fst "\(v, a). ((v, a), the (s v) = a)"] prod.case_distrib by blast finally show ?thesis unfolding state_to_strips_state_dom_is_i[OF nb] by force qed corollary state_to_strips_state_dom_element_iff: assumes "is_valid_problem_sas_plus \" shows "(v, a) \ dom (\\<^sub>S \ s) \ v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None \ a \ \\<^sub>+ \ v" proof - let ?vs = "variables_of \" and ?s' = "\\<^sub>S \ s" show ?thesis proof (rule iffI) assume "(v, a) \ dom (\\<^sub>S \ s)" then have "v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }" and "a \ \\<^sub>+ \ v" unfolding state_to_strips_state_dom_is[OF assms(1)] by force+ moreover have "v \ set ?vs" and "s v \ None" using calculation(1) by fastforce+ ultimately show "v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None \ a \ \\<^sub>+ \ v" by force next assume "v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None \ a \ \\<^sub>+ \ v" then have "v \ set ((\)\<^sub>\\<^sub>+)" and "s v \ None" and a_in_range_of_v: "a \ \\<^sub>+ \ v" by simp+ then have "v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }" by force thus "(v, a) \ dom (\\<^sub>S \ s)" unfolding state_to_strips_state_dom_is[OF assms(1)] using a_in_range_of_v by blast qed qed lemma state_to_strips_state_range_is: assumes "is_valid_problem_sas_plus \" and "(v, a) \ dom (\\<^sub>S \ s)" shows "(\\<^sub>S \ s) (v, a) = Some (the (s v) = a)" proof - let ?vs = "variables_of \" let ?s' = "\\<^sub>S \ s" and ?defined = "filter (\v. s v \ None) ?vs" let ?l = "concat [possible_assignments_for \ v. v \ ?defined]" have v_in_set_vs: "v \ set ?vs" and s_of_v_is_not_None: "s v \ None" and a_in_range_of_v: "a \ \\<^sub>+ \ v" using assms(2) unfolding state_to_strips_state_dom_is[OF assms(1)] by fastforce+ moreover { have "\v \ set ((\)\<^sub>\\<^sub>+). v \ dom (sas_plus_problem.range_of \)" using assms(1) is_valid_problem_sas_plus_then(1) unfolding is_valid_problem_sas_plus_def by fastforce moreover have "(v, a) \ set ?l" unfolding state_to_strips_state_dom_is_i[OF calculation(1)] using s_of_v_is_not_None a_in_range_of_v v_in_set_vs by fastforce moreover have "set ?l \ {}" using calculation by fastforce \ \ TODO slow. \ ultimately have "(\\<^sub>S \ s) (v, a) = Some (the (s v) = a)" using map_of_from_function_graph_is_some_if[of ?l "(v, a)" "\(v, a). the (s v) = a"] unfolding SAS_Plus_STRIPS.state_to_strips_state_def state_to_strips_state_def Let_def case_prod_beta' by fastforce } thus ?thesis. qed \ \ Show that a STRIPS state corresponding to a SAS+ state via transformation is consistent w.r.t. to the variable subset with same left component (i.e. the original SAS+ variable). This is the consistency notion corresponding to SAS+ consistency: i.e. if no two assignments with different values for the same variable exist in the SAS+ state, then assigning the corresponding assignment both to @{text "True"} is impossible. Vice versa, if both are assigned to @{text "True"} then the assignment variables must be the same SAS+ variable/SAS+ value pair. \ lemma state_to_strips_state_effect_consistent: assumes "is_valid_problem_sas_plus \" and "(v, a) \ dom (\\<^sub>S \ s)" and "(v, a') \ dom (\\<^sub>S \ s)" and "(\\<^sub>S \ s) (v, a) = Some True" and "(\\<^sub>S \ s) (v, a') = Some True" shows "(v, a) = (v, a')" proof - have "the (s v) = a" and "the (s v) = a'" using state_to_strips_state_range_is[OF assms(1)] assms(2, 3, 4, 5) by fastforce+ thus ?thesis by argo qed lemma sasp_op_to_strips_set_delete_effects_is: assumes "is_valid_operator_sas_plus \ op" shows "set (strips_operator.delete_effects_of (\\<^sub>O \ op)) = (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" proof - let ?D = "range_of \" and ?effect = "effect_of op" let ?delete = "[(v, a'). (v, a) \ ?effect, a' \ filter ((\) a) (the (?D v))]" { fix v a assume "(v, a) \ set ?effect" then have "(\\<^sub>+ \ v) = set (the (?D v))" using assms using is_valid_operator_sas_plus_then_range_of_sas_plus_op_is_set_range_of_op by fastforce hence "set (filter ((\) a) (the (?D v))) = { a' \ \\<^sub>+ \ v. a' \ a }" unfolding set_filter by blast } note nb = this { \ \ TODO slow. \ have "set ?delete = \(set ` (\(v, a). map (Pair v) (filter ((\) a) (the (?D v)))) ` (set ?effect))" using set_concat by simp also have "\ = \((\(v, a). Pair v ` set (filter ((\) a) (the (?D v)))) ` (set ?effect))" unfolding image_comp[of set] set_map by auto \ \ TODO slow. \ also have "\ = (\(v, a) \ set ?effect. Pair v ` { a' \ \\<^sub>+ \ v. a' \ a })" using nb by fast finally have "set ?delete = (\(v, a) \ set ?effect. { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" by blast } thus ?thesis unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def Let_def by force qed lemma sas_plus_problem_to_strips_problem_variable_set_is: \ \ The variable set of \\\ is the set of all possible assignments that are possible using the variables of \\\ and the corresponding domains. \ assumes "is_valid_problem_sas_plus \" shows "set ((\ \)\<^sub>\) = (\v \ set ((\)\<^sub>\\<^sub>+). { (v, a) | a. a \ \\<^sub>+ \ v })" proof - let ?\ = "\ \" and ?vs = "variables_of \" { have "set (strips_problem.variables_of ?\) = set [as. v \ ?vs, as \ possible_assignments_for \ v]" unfolding sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def by force also have "\ = (\(set ` (\v. possible_assignments_for \ v) ` set ?vs))" using set_concat by auto also have "\ = (\((set \ possible_assignments_for \) ` set ?vs))" using image_comp[of set "\v. possible_assignments_for \ v" "set ?vs"] by argo finally have "set (strips_problem.variables_of ?\) = (\v \ set ?vs. set (possible_assignments_for \ v))" unfolding o_apply by blast } moreover have "\v \ set ?vs. v \ dom (sas_plus_problem.range_of \)" using is_valid_problem_sas_plus_dom_sas_plus_problem_range_of assms by force ultimately show ?thesis using possible_assignments_for_set_is by force qed corollary sas_plus_problem_to_strips_problem_variable_set_element_iff: assumes "is_valid_problem_sas_plus \" shows "(v, a) \ set ((\ \)\<^sub>\) \ v \ set ((\)\<^sub>\\<^sub>+) \ a \ \\<^sub>+ \ v" unfolding sas_plus_problem_to_strips_problem_variable_set_is[OF assms] by fast lemma sasp_op_to_strips_effect_consistent: assumes "op = \\<^sub>O \ op'" and "op' \ set ((\)\<^sub>\\<^sub>+)" and "is_valid_operator_sas_plus \ op'" shows "(v, a) \ set (add_effects_of op) \ (v, a) \ set (delete_effects_of op)" and "(v, a) \ set (delete_effects_of op) \ (v, a) \ set (add_effects_of op)" proof - have nb: "(\(v, a) \ set (effect_of op'). \(v', a') \ set (effect_of op'). v \ v' \ a = a')" using assms(3) unfolding is_valid_operator_sas_plus_def SAS_Plus_Representation.is_valid_operator_sas_plus_def list_all_iff ListMem_iff Let_def by argo { fix v a assume v_a_in_add_effects_of_op: "(v, a) \ set (add_effects_of op)" have "(v, a) \ set (delete_effects_of op)" proof (rule ccontr) assume "\(v, a) \ set (delete_effects_of op)" moreover have "(v, a) \ (\(v, a') \ set (effect_of op'). { (v, a'') | a''. a'' \ (\\<^sub>+ \ v) \ a'' \ a' })" using calculation sasp_op_to_strips_set_delete_effects_is assms by blast moreover obtain a' where "(v, a') \ set (effect_of op')" and "a \ a'" using calculation by blast moreover have "(v, a') \ set (add_effects_of op)" using assms(1) calculation(3) unfolding sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by fastforce moreover have "(v, a) \ set (effect_of op')" and "(v, a') \ set (effect_of op')" using assms(1) v_a_in_add_effects_of_op calculation(5) unfolding sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by force+ ultimately show False using nb by fast qed } moreover { fix v a assume v_a_in_delete_effects_of_op: "(v, a) \ set (delete_effects_of op)" have "(v, a) \ set (add_effects_of op)" proof (rule ccontr) assume "\(v, a) \ set (add_effects_of op)" moreover have "(v, a) \ set (add_effects_of op)" using calculation by blast moreover have "(v, a) \ (\(v, a') \ set (effect_of op'). { (v, a'') | a''. a'' \ (\\<^sub>+ \ v) \ a'' \ a' })" using sasp_op_to_strips_set_delete_effects_is nb assms(1, 3) v_a_in_delete_effects_of_op by force moreover obtain a' where "(v, a') \ set (effect_of op')" and "a \ a'" using calculation by blast moreover have "(v, a') \ set (add_effects_of op)" using assms(1) calculation(4) unfolding sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by fastforce moreover have "(v, a) \ set (effect_of op')" and "(v, a') \ set (effect_of op')" using assms(1) calculation(2, 6) unfolding sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by force+ ultimately show False using nb by fast qed } ultimately show "(v, a) \ set (add_effects_of op) \ (v, a) \ set (delete_effects_of op)" and "(v, a) \ set (delete_effects_of op) \ (v, a) \ set (add_effects_of op)" by blast+ qed lemma is_valid_problem_sas_plus_then_strips_transformation_too_iii: assumes "is_valid_problem_sas_plus \" shows "list_all (is_valid_operator_strips (\ \)) (strips_problem.operators_of (\ \))" proof - let ?\ = "\ \" let ?vs = "strips_problem.variables_of ?\" { fix op assume "op \ set (strips_problem.operators_of ?\)" \ \ TODO slow. \ then obtain op' where op_is: "op = \\<^sub>O \ op'" and op'_in_operators: "op' \ set ((\)\<^sub>\\<^sub>+)" unfolding SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def sas_plus_problem_to_strips_problem_def sasp_op_to_strips_def by auto then have is_valid_op': "is_valid_operator_sas_plus \ op'" using sublocale_sas_plus_finite_domain_representation_ii(2)[OF assms] by blast moreover { fix v a assume "(v, a) \ set (strips_operator.precondition_of op)" \ \ TODO slow. \ then have "(v, a) \ set (sas_plus_operator.precondition_of op')" using op_is unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by force moreover have "v \ set ((\)\<^sub>\\<^sub>+)" using is_valid_op' calculation using is_valid_operator_sas_plus_then(1) by fastforce moreover have "a \ \\<^sub>+ \ v" using is_valid_op' calculation(1) using is_valid_operator_sas_plus_then(2) by fast ultimately have "(v, a) \ set ?vs" using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] by force } moreover { fix v a assume "(v, a) \ set (strips_operator.add_effects_of op)" then have "(v, a) \ set (effect_of op')" using op_is unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by force then have "v \ set ((\)\<^sub>\\<^sub>+)" and "a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then is_valid_op' by fastforce+ hence "(v, a) \ set ?vs" using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] by force } moreover { fix v a' assume v_a'_in_delete_effects: "(v, a') \ set (strips_operator.delete_effects_of op)" moreover have "set (strips_operator.delete_effects_of op) = (\(v, a) \ set (effect_of op'). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" using sasp_op_to_strips_set_delete_effects_is[OF is_valid_op'] op_is by simp \ \ TODO slow. \ ultimately obtain a where "(v, a) \ set (effect_of op')" and a'_in: "a' \ { a' \ \\<^sub>+ \ v. a' \ a }" by blast moreover have "is_valid_operator_sas_plus \ op'" using op'_in_operators assms(1) is_valid_problem_sas_plus_then(2) by blast moreover have "v \ set ((\)\<^sub>\\<^sub>+)" using is_valid_operator_sas_plus_then calculation(1, 3) by fast moreover have "a' \ \\<^sub>+ \ v" using a'_in by blast ultimately have "(v, a') \ set ?vs" using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] by force } ultimately have "set (strips_operator.precondition_of op) \ set ?vs \ set (strips_operator.add_effects_of op) \ set ?vs \ set (strips_operator.delete_effects_of op) \ set ?vs \ (\v\set (add_effects_of op). v \ set (delete_effects_of op)) \ (\v\set (delete_effects_of op). v \ set (add_effects_of op))" using sasp_op_to_strips_effect_consistent[OF op_is op'_in_operators is_valid_op'] by fast+ } thus ?thesis unfolding is_valid_operator_strips_def STRIPS_Representation.is_valid_operator_strips_def list_all_iff ListMem_iff Let_def by blast qed lemma is_valid_problem_sas_plus_then_strips_transformation_too_iv: assumes "is_valid_problem_sas_plus \" shows "\x. ((\ \)\<^sub>I) x \ None \ ListMem x (strips_problem.variables_of (\ \))" proof - let ?vs = "variables_of \" and ?I = "initial_of \" and ?\ = "\ \" let ?vs' = "strips_problem.variables_of ?\" and ?I' = "strips_problem.initial_of ?\" { fix x have "?I' x \ None \ ListMem x ?vs'" proof (rule iffI) assume I'_of_x_is_not_None: "?I' x \ None" then have "x \ dom ?I'" by blast moreover obtain v a where x_is: "x = (v, a)" by fastforce ultimately have "(v, a) \ dom ?I'" by blast then have "v \ set ?vs" and "?I v \ None" and "a \ \\<^sub>+ \ v" using state_to_strips_state_dom_element_iff[OF assms(1), of v a ?I] unfolding sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def state_to_strips_state_def SAS_Plus_STRIPS.state_to_strips_state_def by simp+ thus "ListMem x ?vs'" unfolding ListMem_iff using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] x_is by auto next assume list_mem_x_vs': "ListMem x ?vs'" then obtain v a where x_is: "x = (v, a)" by fastforce then have "(v, a) \ set ?vs'" using list_mem_x_vs' unfolding ListMem_iff by blast then have "v \ set ?vs" and "a \ \\<^sub>+ \ v" using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] by force+ moreover have "?I v \ None" using is_valid_problem_sas_plus_then(3) assms(1) calculation(1) by auto ultimately have "(v, a) \ dom ?I'" using state_to_strips_state_dom_element_iff[OF assms(1), of v a ?I] unfolding SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.state_to_strips_state_def state_to_strips_state_def by force thus "?I' x \ None" using x_is by fastforce qed } thus ?thesis by simp qed private lemma is_valid_problem_sas_plus_then_strips_transformation_too_v: assumes "is_valid_problem_sas_plus \" shows "\x. ((\ \)\<^sub>G) x \ None \ ListMem x (strips_problem.variables_of (\ \))" proof - let ?vs = "variables_of \" and ?D = "range_of \" and ?G = "goal_of \" let ?\ = "\ \" let ?vs' = "strips_problem.variables_of ?\" and ?G' = "strips_problem.goal_of ?\" have nb: "?G' = \\<^sub>S \ ?G" by simp { fix x assume "?G' x \ None" moreover obtain v a where "x = (v, a)" by fastforce moreover have "(v, a) \ dom ?G'" using domIff calculation(1, 2) by blast moreover have "v \ set ?vs" and "a \ \\<^sub>+ \ v" using state_to_strips_state_dom_is[OF assms(1), of ?G] nb calculation(3) by auto+ ultimately have "x \ set ?vs'" using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] by auto } thus ?thesis unfolding ListMem_iff by simp qed text \ We now show that given \<^term>\\\ is a valid SASPlus problem, then \<^term>\\ \ \ \\ is a valid STRIPS problem as well. The proof unfolds the definition of \<^term>\is_valid_problem_strips\ and then shows each of the conjuncts for \<^term>\\\. These are: \begin{itemize} \item \<^term>\\\ has at least one variable; \item \<^term>\\\ has at least one operator; \item all operators are valid STRIPS operators; \item \<^term>\(\::'a strips_problem)\<^sub>I\ is defined for all variables in \<^term>\(\::'a strips_problem)\<^sub>\\; and finally, \item if \<^term>\((\::'a strips_problem)\<^sub>G) x\ is defined, then \<^term>\x\ is in \<^term>\(\::'a strips_problem)\<^sub>\\. \end{itemize} \ theorem is_valid_problem_sas_plus_then_strips_transformation_too: assumes "is_valid_problem_sas_plus \" shows "is_valid_problem_strips (\ \)" proof - let ?\ = "\ \" have "list_all (is_valid_operator_strips (\ \)) (strips_problem.operators_of (\ \))" using is_valid_problem_sas_plus_then_strips_transformation_too_iii[OF assms]. moreover have "\x. (((\ \)\<^sub>I) x \ None) = ListMem x (strips_problem.variables_of (\ \))" using is_valid_problem_sas_plus_then_strips_transformation_too_iv[OF assms]. moreover have "\x. ((\ \)\<^sub>G) x \ None \ ListMem x (strips_problem.variables_of (\ \))" using is_valid_problem_sas_plus_then_strips_transformation_too_v[OF assms]. ultimately show ?thesis using is_valid_problem_strips_def unfolding STRIPS_Representation.is_valid_problem_strips_def by fastforce qed lemma set_filter_all_possible_assignments_true_is: assumes "is_valid_problem_sas_plus \" shows "set (filter (\(v, a). s (v, a) = Some True) (all_possible_assignments_for \)) = (\v \ set ((\)\<^sub>\\<^sub>+). Pair v ` { a \ \\<^sub>+ \ v. s (v, a) = Some True })" proof - let ?vs = "sas_plus_problem.variables_of \" and ?P = "(\(v, a). s (v, a) = Some True)" let ?l = "filter ?P (all_possible_assignments_for \)" have "set ?l = set (concat (map (filter ?P) (map (possible_assignments_for \) ?vs)))" unfolding all_possible_assignments_for_def filter_concat[of ?P "map (possible_assignments_for \) (sas_plus_problem.variables_of \)"] by simp also have "\ = set (concat (map (\v. filter ?P (possible_assignments_for \ v)) ?vs))" unfolding map_map comp_apply by blast also have "\ = set (concat (map (\v. map (Pair v) (filter (?P \ Pair v) (the (range_of \ v)))) ?vs))" unfolding possible_assignments_for_def filter_map by blast also have "\ = set (concat (map (\v. map (Pair v) (filter (\a. s (v, a) = Some True) (the (range_of \ v)))) ?vs))" unfolding comp_apply by fast also have "\ = \(set ` ((\v. map (Pair v) (filter (\a. s (v, a) = Some True) (the (range_of \ v)))) ` set ?vs))" unfolding set_concat set_map.. also have "\ = (\v \ set ?vs. Pair v ` set (filter (\a. s (v, a) = Some True) (the (range_of \ v))))" unfolding image_comp[of set] comp_apply set_map.. also have "\ = (\v \ set ?vs. Pair v ` { a \ set (the (range_of \ v)). s (v, a) = Some True })" unfolding set_filter.. finally show ?thesis using set_the_range_of_is_range_of_sas_plus_if[OF assms(1)] by auto qed lemma strips_state_to_state_dom_is: assumes "is_valid_problem_sas_plus \" shows "dom (\\<^sub>S\ \ s) = (\v \ set ((\)\<^sub>\\<^sub>+). { v | a. a \ (\\<^sub>+ \ v) \ s (v, a) = Some True })" proof - let ?vs = "variables_of \" and ?s' = "\\<^sub>S\ \ s" and ?P = "(\(v, a). s (v, a) = Some True)" let ?l = "filter ?P (all_possible_assignments_for \)" { have "fst ` set ?l = fst ` (\v \ set ?vs. Pair v ` { a \ \\<^sub>+ \ v. s (v, a) = Some True })" unfolding set_filter_all_possible_assignments_true_is[OF assms] by auto also have "\ = (\v \ set ?vs. fst ` Pair v ` { a \ \\<^sub>+ \ v. s (v, a) = Some True })" by blast also have "\ = (\v \ set ?vs. (\a. fst (Pair v a)) ` { a \ \\<^sub>+ \ v. s (v, a) = Some True })" unfolding image_comp[of fst] comp_apply by blast finally have "fst ` set ?l = (\v \ set ((\)\<^sub>\\<^sub>+). { v | a. a \ (\\<^sub>+ \ v) \ s (v, a) = Some True })" unfolding setcompr_eq_image fst_conv by simp } thus ?thesis unfolding SAS_Plus_STRIPS.strips_state_to_state_def strips_state_to_state_def dom_map_of_conv_image_fst by blast qed lemma strips_state_to_state_range_is: assumes "is_valid_problem_sas_plus \" and "v \ set ((\)\<^sub>\\<^sub>+)" and "a \ \\<^sub>+ \ v" and "(v, a) \ dom s'" and "\(v, a) \ dom s'. \(v, a') \ dom s'. s' (v, a) = Some True \ s' (v, a') = Some True \ (v, a) = (v, a')" shows "(\\<^sub>S\ \ s') v = Some a \ the (s' (v, a))" proof - let ?vs = "variables_of \" and ?D = "range_of \" and ?s = "\\<^sub>S\ \ s'" let ?as = "all_possible_assignments_for \" let ?l = "filter (\(v, a). s' (v, a) = Some True) ?as" show ?thesis proof (rule iffI) assume s_of_v_is_Some_a: "?s v = Some a" { have "(v, a) \ set ?l" using s_of_v_is_Some_a unfolding SAS_Plus_STRIPS.strips_state_to_state_def strips_state_to_state_def using map_of_SomeD by fast hence "s' (v, a) = Some True" unfolding all_possible_assignments_for_set_is set_filter by blast } thus "the (s' (v, a))" by simp next assume the_of_s'_of_v_a_is: "the (s' (v, a))" then have s'_of_v_a_is_Some_true: "s' (v, a) = Some True" using assms(4) domIff by force \ \ TODO slow. \ moreover { fix v v' a a' assume "(v, a) \ set ?l" and "(v', a') \ set ?l" then have "v \ v' \ a = a'" using assms(5) by fastforce } moreover { have "\v \ set ((\)\<^sub>\\<^sub>+). sas_plus_problem.range_of \ v \ None" using is_valid_problem_sas_plus_then(1) assms(1) range_of_not_empty by force (* TODO slow. *) moreover have "set ?l = Set.filter (\(v, a). s' (v, a) = Some True) (\v \ set ((\)\<^sub>\\<^sub>+). { (v, a) | a. a \ \\<^sub>+ \ v })" using all_possible_assignments_for_set_is calculation by force ultimately have "(v, a) \ set ?l" using assms(2, 3) s'_of_v_a_is_Some_true by simp } ultimately show "?s v = Some a" using map_of_constant_assignments_defined_if[of ?l v a] unfolding SAS_Plus_STRIPS.strips_state_to_state_def strips_state_to_state_def by blast qed qed \ \ NOTE A technical lemma which characterizes the return values for possible assignments @{text "(v, a)"} when used as variables on a state @{text "s"} which was transformed from. \ lemma strips_state_to_state_inverse_is_i: assumes "is_valid_problem_sas_plus \" and "v \ set ((\)\<^sub>\\<^sub>+)" and "s v \ None" and "a \ \\<^sub>+ \ v" shows "(\\<^sub>S \ s) (v, a) = Some (the (s v) = a)" proof - let ?vs = "sas_plus_problem.variables_of \" let ?s' = "\\<^sub>S \ s" and ?f = "\(v, a). the (s v) = a" and ?l = "concat (map (possible_assignments_for \) (filter (\v. s v \ None) ?vs))" have "(v, a) \ dom ?s'" using state_to_strips_state_dom_element_iff[ OF assms(1)] assms(2, 3, 4) by presburger { have "v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }" using assms(2, 3) by blast moreover have "\v \ set ((\)\<^sub>\\<^sub>+). v \ dom (sas_plus_problem.range_of \)" using is_valid_problem_sas_plus_dom_sas_plus_problem_range_of[OF assms(1)]. moreover have "set ?l = (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }. { (v, a) |a. a \ \\<^sub>+ \ v })" unfolding state_to_strips_state_dom_is_i[OF calculation(2)] by blast ultimately have "(v, a) \ set ?l" using assms(4) by blast } moreover have "set ?l \ {}" using calculation by force \ \ TODO slow.\ ultimately show ?thesis unfolding SAS_Plus_STRIPS.state_to_strips_state_def state_to_strips_state_def using map_of_from_function_graph_is_some_if[of ?l "(v, a)" ?f] unfolding split_def by fastforce qed \ \ NOTE Show that the transformed strips state is consistent for pairs of assignments @{text "(v, a)"} and @{text "(v, a')"} in the same variable domain. \ (* TODO make private. *) corollary strips_state_to_state_inverse_is_ii: assumes "is_valid_problem_sas_plus \" and "v \ set ((\)\<^sub>\\<^sub>+)" and "s v = Some a" and "a \ \\<^sub>+ \ v" and "a' \ \\<^sub>+ \ v" and "a' \ a" shows "(\\<^sub>S \ s) (v, a') = Some False" proof - have "s v \ None" using assms(3) by simp moreover have "the (s v) \ a'" using assms(3, 6) by simp ultimately show ?thesis using strips_state_to_state_inverse_is_i[OF assms(1, 2) _ assms(5)] by force qed \ \ NOTE Follows from the corollary above by contraposition. \ (* TODO make private. *) corollary strips_state_to_state_inverse_is_iii: assumes "is_valid_problem_sas_plus \" and "v \ set ((\)\<^sub>\\<^sub>+)" and "s v = Some a" and "a \ \\<^sub>+ \ v" and "a' \ \\<^sub>+ \ v" and "(\\<^sub>S \ s) (v, a) = Some True" and "(\\<^sub>S \ s) (v, a') = Some True" shows "a = a'" proof - have "s v \ None" using assms(3) by blast thus ?thesis using strips_state_to_state_inverse_is_i[OF assms(1, 2)] assms(4, 5, 6, 7) by auto qed (* TODO make private. *) lemma strips_state_to_state_inverse_is_iv: assumes "is_valid_problem_sas_plus \" and "dom s \ set ((\)\<^sub>\\<^sub>+)" and "v \ set ((\)\<^sub>\\<^sub>+)" and "s v = Some a" and "a \ \\<^sub>+ \ v" shows "(\\<^sub>S\ \ (\\<^sub>S \ s)) v = Some a" proof - let ?vs = "variables_of \" and ?s' = "\\<^sub>S \ s" let ?s'' = "\\<^sub>S\ \ ?s'" let ?P = "\(v, a). ?s' (v, a) = Some True" let ?as = "filter ?P (all_possible_assignments_for \)" and ?As = "Set.filter ?P (\v \ set ((\)\<^sub>\\<^sub>+). { (v, a) | a. a \ \\<^sub>+ \ v })" { have "\v \ set ((\)\<^sub>\\<^sub>+). range_of \ v \ None" using sublocale_sas_plus_finite_domain_representation_ii(1)[OF assms(1)] range_of_not_empty by force (* TODO slow. *) hence "set ?as = ?As" unfolding set_filter using all_possible_assignments_for_set_is by force } note nb = this moreover { { fix v v' a a' assume "(v, a) \ set ?as" and "(v', a') \ set ?as" then have "(v, a) \ ?As" and "(v', a') \ ?As" using nb by blast+ then have v_in_set_vs: "v \ set ?vs" and v'_in_set_vs: "v' \ set ?vs" and a_in_range_of_v: "a \ \\<^sub>+ \ v" and a'_in_range_of_v: "a' \ \\<^sub>+ \ v'" and s'_of_v_a_is: "?s' (v, a) = Some True" and s'_of_v'_a'_is: "?s' (v', a') = Some True" by fastforce+ then have "(v, a) \ dom ?s'" by blast then have s_of_v_is_Some_a: "s v = Some a" using state_to_strips_state_dom_element_iff[OF assms(1)] state_to_strips_state_range_is[OF assms(1)] s'_of_v_a_is by auto have "v \ v' \ a = a'" proof (rule ccontr) assume "\(v \ v' \ a = a')" then have "v = v'" and "a \ a'" by simp+ thus False using a'_in_range_of_v a_in_range_of_v assms(1) v'_in_set_vs s'_of_v'_a'_is s'_of_v_a_is s_of_v_is_Some_a strips_state_to_state_inverse_is_iii by force qed } moreover { have "s v \ None" using assms(4) by simp then have "?s' (v, a) = Some True" using strips_state_to_state_inverse_is_i[OF assms(1, 3) _ assms(5)] assms(4) by simp (* TODO slow *) hence "(v, a) \ set ?as" using all_possible_assignments_for_set_is assms(3, 5) nb by simp } ultimately have "map_of ?as v = Some a" using map_of_constant_assignments_defined_if[of ?as v a] by blast } \ \ TODO slow. \ thus ?thesis unfolding SAS_Plus_STRIPS.strips_state_to_state_def strips_state_to_state_def all_possible_assignments_for_def by simp qed (* TODO the constraints on the state @{text "s"} could be refactored into a definition of valid states for a problem description. *) (* TODO The proof is not very elegant. Should be simplified. *) \ \ Show that that \\\<^sub>S\ \\ is the inverse of \\\<^sub>S \\. The additional constraints \<^term>\dom s = set ((\)\<^sub>\\<^sub>+)\ and \<^term>\\v \ dom s. the (s v) \ \\<^sub>+ \ v\ are needed because the transformation functions only take into account variables and domains declared in the problem description. They also sufficiently characterize a state that was transformed from SAS+ to STRIPS. \ lemma strips_state_to_state_inverse_is: assumes "is_valid_problem_sas_plus \" and "dom s \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom s. the (s v) \ \\<^sub>+ \ v" shows "s = (\\<^sub>S\ \ (\\<^sub>S \ s))" proof - let ?vs = "variables_of \" and ?D = "range_of \" let ?s' = "\\<^sub>S \ s" let ?s'' = "\\<^sub>S\ \ ?s'" \ \ NOTE Show the thesis by proving that @{text "s"} and @{text "?s'"} are mutual submaps. \ { fix v assume v_in_dom_s: "v \ dom s" then have v_in_set_vs: "v \ set ?vs" using assms(2) by auto then obtain a where the_s_v_is_a: "s v = Some a" and a_in_dom_v: "a \ \\<^sub>+ \ v" using assms(2, 3) v_in_dom_s by force moreover have "?s'' v = Some a" using strips_state_to_state_inverse_is_iv[OF assms(1, 2)] v_in_set_vs the_s_v_is_a a_in_dom_v by force ultimately have "s v = ?s'' v" by argo } note nb = this moreover { fix v assume "v \ dom ?s''" then obtain a where "a \ \\<^sub>+ \ v" and "?s' (v, a) = Some True" using strips_state_to_state_dom_is[OF assms(1)] by blast then have "(v, a) \ dom ?s'" by blast then have "s v \ None" using state_to_strips_state_dom_is[OF assms(1)] by simp then obtain a where "s v = Some a" by blast hence "?s'' v = s v" using nb by fastforce } \ \ TODO slow.\ ultimately show ?thesis using map_le_antisym[of s ?s''] map_le_def unfolding strips_state_to_state_def state_to_strips_state_def by blast qed \ \ An important lemma which shows that the submap relation does not change if we transform the states on either side from SAS+ to STRIPS. % TODO what is this called generally? Predicate monotony?? \ lemma state_to_strips_state_map_le_iff: assumes "is_valid_problem_sas_plus \" and "dom s \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom s. the (s v) \ \\<^sub>+ \ v" shows "s \\<^sub>m t \ (\\<^sub>S \ s) \\<^sub>m (\\<^sub>S \ t)" proof - let ?vs = "variables_of \" and ?D = "range_of \" and ?s' = "\\<^sub>S \ s" and ?t' = "\\<^sub>S \ t" show ?thesis proof (rule iffI) assume s_map_le_t: "s \\<^sub>m t" { fix v a assume "(v, a) \ dom ?s'" moreover have "v \ set ((\)\<^sub>\\<^sub>+)" and "s v \ None" and "a \ \\<^sub>+ \ v" using state_to_strips_state_dom_is[OF assms(1)] calculation by blast+ moreover have "?s' (v, a) = Some (the (s v) = a)" using state_to_strips_state_range_is[OF assms(1)] calculation(1) by meson moreover have "v \ dom s" using calculation(3) by auto moreover have "s v = t v" using s_map_le_t calculation(6) unfolding map_le_def by blast moreover have "t v \ None" using calculation(3, 7) by argo moreover have "(v, a) \ dom ?t'" using state_to_strips_state_dom_is[OF assms(1)] calculation(2, 4, 8) by blast moreover have "?t' (v, a) = Some (the (t v) = a)" using state_to_strips_state_range_is[OF assms(1)] calculation(9) by simp ultimately have "?s' (v, a) = ?t' (v, a)" by presburger } thus "?s' \\<^sub>m ?t'" unfolding map_le_def by fast next assume s'_map_le_t': "?s' \\<^sub>m ?t'" { fix v assume v_in_dom_s: "v \ dom s" moreover obtain a where the_of_s_of_v_is_a: "the (s v) = a" by blast moreover have v_in_vs: "v \ set ((\)\<^sub>\\<^sub>+)" and s_of_v_is_not_None: "s v \ None" and a_in_range_of_v: "a \ \\<^sub>+ \ v" using assms(2, 3) v_in_dom_s calculation by blast+ moreover have "(v, a) \ dom ?s'" using state_to_strips_state_dom_is[OF assms(1)] calculation(3, 4, 5) by simp moreover have "?s' (v, a) = ?t' (v, a)" using s'_map_le_t' calculation unfolding map_le_def by blast moreover have "(v, a) \ dom ?t'" using calculation unfolding domIff by argo moreover have "?s' (v, a) = Some (the (s v) = a)" and "?t' (v, a) = Some (the (t v) = a)" using state_to_strips_state_range_is[OF assms(1)] calculation by fast+ moreover have "s v = Some a" using calculation(2, 4) by force moreover have "?s' (v, a) = Some True" using calculation(9, 11) by fastforce moreover have "?t' (v, a) = Some True" using calculation(7, 12) by argo moreover have "the (t v) = a" using calculation(10, 13) try0 by force moreover { have "v \ dom t" using state_to_strips_state_dom_element_iff[OF assms(1)] calculation(8) by auto hence "t v = Some a" using calculation(14) by force } ultimately have "s v = t v" by argo } thus "s \\<^sub>m t" unfolding map_le_def by simp qed qed \ \ We also show that \\\<^sub>O\ \\ is the inverse of \\\<^sub>O \\. Note that this proof is completely mechanical since both the precondition and effect lists are simply being copied when transforming from SAS+ to STRIPS and when transforming back from STRIPS to SAS+. \ (* TODO rename \sasp_op_to_strips_inverse_is\ *) (* TODO prune assumptions (not required) *) lemma sas_plus_operator_inverse_is: assumes "is_valid_problem_sas_plus \" and "op \ set ((\)\<^sub>\\<^sub>+)" shows "(\\<^sub>O\ \ (\\<^sub>O \ op)) = op" proof - let ?op = "\\<^sub>O\ \ (\\<^sub>O \ op)" have "precondition_of ?op = precondition_of op" unfolding SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by fastforce moreover have "effect_of ?op = effect_of op" unfolding SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by force ultimately show ?thesis by simp qed \ \ Note that we have to make the assumption that \op'\ is a member of the operator set of the induced STRIPS problem \\ \\. This implies that \op'\ was transformed from an \op \ operators_of \\. If we don't make this assumption, then multiple STRIPS operators of the form \\ precondition_of = [], add_effects_of = [], delete_effects_of = [(v, a), ...] \\ correspond to one SAS+ operator (since the delete effects are being discarded in the transformation function). \ lemma strips_operator_inverse_is: assumes "is_valid_problem_sas_plus \" and "op' \ set ((\ \)\<^sub>\)" shows "(\\<^sub>O \ (\\<^sub>O\ \ op')) = op'" proof - let ?\ = "\ \" obtain op where "op \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op" using assms by auto moreover have "\\<^sub>O\ \ op' = op" using sas_plus_operator_inverse_is[OF assms(1) calculation(1)] calculation(2) by blast ultimately show ?thesis by argo qed (* \<^item> TODO Simplify | refactor proof. \<^item> TODO make private. *) lemma sas_plus_equivalent_to_strips_i_a_I: assumes "is_valid_problem_sas_plus \" and "set ops' \ set ((\ \)\<^sub>\)" and "STRIPS_Semantics.are_all_operators_applicable (\\<^sub>S \ s) ops'" and "op \ set [\\<^sub>O\ \ op'. op' \ ops']" shows "map_of (precondition_of op) \\<^sub>m (\\<^sub>S\ \ (\\<^sub>S \ s))" proof - let ?\ = "\ \" and ?s' = "\\<^sub>S \ s" let ?s = "\\<^sub>S\ \ ?s'" and ?D = "range_of \" and ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" and ?pre = "precondition_of op" have nb\<^sub>1: "\(v, a) \ dom ?s'. \(v, a') \ dom ?s'. ?s' (v, a) = Some True \ ?s' (v, a') = Some True \ (v, a) = (v, a')" using state_to_strips_state_effect_consistent[OF assms(1)] by blast { fix op' assume "op' \ set ops'" moreover have "op' \ set ((?\)\<^sub>\)" using assms(2) calculation by blast ultimately have "\op \ set ((\)\<^sub>\\<^sub>+). op' = (\\<^sub>O \ op)" by auto } note nb\<^sub>2 = this { fix op assume "op \ set ?ops" then obtain op' where "op' \ set ops'" and "op = \\<^sub>O\ \ op'" using assms(4) by auto moreover obtain op'' where "op'' \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op''" using nb\<^sub>2 calculation(1) by blast moreover have "op = op''" using sas_plus_operator_inverse_is[OF assms(1) calculation(3)] calculation(2, 4) by blast ultimately have "op \ set ((\)\<^sub>\\<^sub>+)" by blast } note nb\<^sub>3 = this { fix op v a assume "op \ set ?ops" and v_a_in_precondition_of_op': "(v, a) \ set (precondition_of op)" moreover obtain op' where "op' \ set ops'" and "op = \\<^sub>O\ \ op'" using calculation(1) by auto moreover have "strips_operator.precondition_of op' = precondition_of op" using calculation(4) unfolding SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def by simp ultimately have "\op' \ set ops'. op = (\\<^sub>O\ \ op') \ (v, a) \ set (strips_operator.precondition_of op')" by metis } note nb\<^sub>4 = this { fix op' v a assume "op' \ set ops'" and v_a_in_precondition_of_op': "(v, a) \ set (strips_operator.precondition_of op')" moreover have s'_of_v_a_is_Some_True: "?s' (v, a) = Some True" using assms(3) calculation(1, 2) unfolding are_all_operators_applicable_set by blast moreover { obtain op where "op \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op" using nb\<^sub>2 calculation(1) by blast moreover have "strips_operator.precondition_of op' = precondition_of op" using calculation(2) unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by simp moreover have "(v, a) \ set (precondition_of op)" using v_a_in_precondition_of_op' calculation(3) by argo moreover have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) assms(1) calculation(1) unfolding is_valid_operator_sas_plus_def by auto moreover have "v \ set ((\)\<^sub>\\<^sub>+)" and "a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then(1,2) calculation(4, 5) unfolding is_valid_operator_sas_plus_def by fastforce+ moreover have "v \ dom ?s" using strips_state_to_state_dom_is[OF assms(1), of ?s'] s'_of_v_a_is_Some_True calculation(6, 7) by blast moreover have "(v, a) \ dom ?s'" using s'_of_v_a_is_Some_True domIff by blast ultimately have "?s v = Some a" using strips_state_to_state_range_is[OF assms(1) _ _ _ nb\<^sub>1] s'_of_v_a_is_Some_True by simp } hence "?s v = Some a". } note nb\<^sub>5 = this { fix v assume "v \ dom (map_of ?pre)" then obtain a where "map_of ?pre v = Some a" by fast moreover have "(v, a) \ set ?pre" using map_of_SomeD calculation by fast moreover { have "op \ set ((\)\<^sub>\\<^sub>+)" using assms(4) nb\<^sub>3 by blast then have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) assms(1) unfolding is_valid_operator_sas_plus_def by auto hence "\(v, a) \ set ?pre. \(v', a') \ set ?pre. v \ v' \ a = a'" using is_valid_operator_sas_plus_then(5) unfolding is_valid_operator_sas_plus_def by fast } moreover have "map_of ?pre v = Some a" using map_of_constant_assignments_defined_if[of ?pre] calculation(2, 3) by blast moreover obtain op' where "op' \ set ops'" and "(v, a) \ set (strips_operator.precondition_of op')" using nb\<^sub>4[OF assms(4) calculation(2)] by blast moreover have "?s v = Some a" using nb\<^sub>5 calculation(5, 6) by fast ultimately have "map_of ?pre v = ?s v" by argo } thus ?thesis unfolding map_le_def by blast qed lemma to_sas_plus_list_of_transformed_sas_plus_problem_operators_structure: assumes "is_valid_problem_sas_plus \" and "set ops' \ set ((\ \)\<^sub>\)" and "op \ set [\\<^sub>O\ \ op'. op' \ ops']" shows "op \ set ((\)\<^sub>\\<^sub>+) \ (\op' \ set ops'. op' = \\<^sub>O \ op)" proof - let ?\ = "\ \" obtain op' where "op' \ set ops'" and "op = \\<^sub>O\ \ op'" using assms(3) by auto moreover have "op' \ set ((?\)\<^sub>\)" using assms(2) calculation(1) by blast moreover obtain op'' where "op'' \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op''" using calculation(3) by auto moreover have "op = op''" using sas_plus_operator_inverse_is[OF assms(1) calculation(4)] calculation(2, 5) by presburger ultimately show ?thesis by blast qed (* \<^item> TODO Prune premises (2nd premise and \are_all_operators_applicable s' ops'\ can be removed?). \<^item> TODO make private. \<^item> TODO adjust nb indexes *) lemma sas_plus_equivalent_to_strips_i_a_II: fixes \ :: "('variable, 'domain) sas_plus_problem" fixes s :: "('variable, 'domain) state" assumes "is_valid_problem_sas_plus \" and "set ops' \ set ((\ \)\<^sub>\)" and "STRIPS_Semantics.are_all_operators_applicable (\\<^sub>s \ s) ops' \ STRIPS_Semantics.are_all_operator_effects_consistent ops'" shows "are_all_operator_effects_consistent [\\<^sub>O\ \ op'. op' \ ops']" proof - let ?s' = "\\<^sub>S \ s" let ?s = "\\<^sub>S\ \ ?s'" and ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" and ?\ = "\ \" have nb: "\(v, a) \ dom ?s'. \(v, a') \ dom ?s'. ?s' (v, a) = Some True \ ?s' (v, a') = Some True \ (v, a) = (v, a')" using state_to_strips_state_effect_consistent[OF assms(1)] by blast { fix op\<^sub>1' op\<^sub>2' assume "op\<^sub>1' \ set ops'" and "op\<^sub>2' \ set ops'" hence "STRIPS_Semantics.are_operator_effects_consistent op\<^sub>1' op\<^sub>2'" using assms(3) unfolding STRIPS_Semantics.are_all_operator_effects_consistent_def list_all_iff by blast } note nb\<^sub>1 = this { fix op\<^sub>1 op\<^sub>1' op\<^sub>2 op\<^sub>2' assume op\<^sub>1_in_ops: "op\<^sub>1 \ set ?ops" and op\<^sub>1'_in_ops': "op\<^sub>1' \ set ops'" and op\<^sub>1'_is: "op\<^sub>1' = \\<^sub>O \ op\<^sub>1" and is_valid_op\<^sub>1: "is_valid_operator_sas_plus \ op\<^sub>1" and op\<^sub>2_in_ops: "op\<^sub>2 \ set ?ops" and op\<^sub>2'_in_ops': "op\<^sub>2' \ set ops'" and op\<^sub>2'_is: "op\<^sub>2' = \\<^sub>O \ op\<^sub>2" and is_valid_op\<^sub>2: "is_valid_operator_sas_plus \ op\<^sub>2" have "\(v, a) \ set (add_effects_of op\<^sub>1'). \(v', a') \ set (add_effects_of op\<^sub>2'). v \ v' \ a = a'" proof (rule ccontr) assume "\(\(v, a) \ set (add_effects_of op\<^sub>1'). \(v', a') \ set (add_effects_of op\<^sub>2'). v \ v' \ a = a')" then obtain v v' a a' where "(v, a) \ set (add_effects_of op\<^sub>1')" and "(v', a') \ set (add_effects_of op\<^sub>2')" and "v = v'" and "a \ a'" by blast \ \ TODO slow. \ moreover have "(v, a) \ set (effect_of op\<^sub>1)" using op\<^sub>1'_is op\<^sub>2'_is calculation(1, 2) unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by force moreover { have "(v', a') \ set (effect_of op\<^sub>2)" using op\<^sub>2'_is calculation(2) unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by force hence "a' \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then is_valid_op\<^sub>2 calculation(3) by fastforce } moreover have "(v, a') \ set (delete_effects_of op\<^sub>1')" using sasp_op_to_strips_set_delete_effects_is op\<^sub>1'_is is_valid_op\<^sub>1 calculation(3, 4, 5, 6) by blast moreover have "\STRIPS_Semantics.are_operator_effects_consistent op\<^sub>1' op\<^sub>2'" unfolding STRIPS_Semantics.are_operator_effects_consistent_def list_ex_iff using calculation(2, 3, 7) by meson ultimately show False using assms(3) op\<^sub>1'_in_ops' op\<^sub>2'_in_ops' unfolding STRIPS_Semantics.are_all_operator_effects_consistent_def list_all_iff by blast qed } note nb\<^sub>3 = this { fix op\<^sub>1 op\<^sub>2 assume op\<^sub>1_in_ops: "op\<^sub>1 \ set ?ops" and op\<^sub>2_in_ops: "op\<^sub>2 \ set ?ops" moreover have op\<^sub>1_in_operators_of_\: "op\<^sub>1 \ set ((\)\<^sub>\\<^sub>+)" and op\<^sub>2_in_operators_of_\: "op\<^sub>2 \ set ((\)\<^sub>\\<^sub>+)" using to_sas_plus_list_of_transformed_sas_plus_problem_operators_structure[OF assms(1, 2)] calculation by blast+ moreover have is_valid_operator_op\<^sub>1: "is_valid_operator_sas_plus \ op\<^sub>1" and is_valid_operator_op\<^sub>2: "is_valid_operator_sas_plus \ op\<^sub>2" using is_valid_problem_sas_plus_then(2) op\<^sub>1_in_operators_of_\ op\<^sub>2_in_operators_of_\ assms(1) unfolding is_valid_operator_sas_plus_def by auto+ moreover obtain op\<^sub>1' op\<^sub>2' where op\<^sub>1_in_ops': "op\<^sub>1' \ set ops'" and op\<^sub>1_is: "op\<^sub>1' = \\<^sub>O \ op\<^sub>1" and op\<^sub>2_in_ops': "op\<^sub>2' \ set ops'" and op\<^sub>2_is: "op\<^sub>2' = \\<^sub>O \ op\<^sub>2" using to_sas_plus_list_of_transformed_sas_plus_problem_operators_structure[OF assms(1, 2)] op\<^sub>1_in_ops op\<^sub>2_in_ops by blast \ \ TODO slow.\ ultimately have "\(v, a) \ set (add_effects_of op\<^sub>1'). \(v', a') \ set (add_effects_of op\<^sub>2'). v \ v' \ a = a'" using nb\<^sub>3 by auto hence "are_operator_effects_consistent op\<^sub>1 op\<^sub>2" using op\<^sub>1_is op\<^sub>2_is unfolding are_operator_effects_consistent_def sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def list_all_iff Let_def by simp } thus ?thesis unfolding are_all_operator_effects_consistent_def list_all_iff by fast qed \ \ A technical lemmas used in \sas_plus_equivalent_to_strips_i_a\ showing that the execution precondition is linear w.r.t. to STRIPS transformation to SAS+. The second premise states that the given STRIPS state corresponds to a consistent SAS+ state (i.e. no two assignments of the same variable to different values exist). \ (* \<^item> TODO make private. \<^item> TODO decrement suffix *) lemma sas_plus_equivalent_to_strips_i_a_IV: assumes "is_valid_problem_sas_plus \" and "set ops' \ set ((\ \)\<^sub>\)" and "STRIPS_Semantics.are_all_operators_applicable (\\<^sub>S \ s) ops' \ STRIPS_Semantics.are_all_operator_effects_consistent ops'" shows "are_all_operators_applicable_in (\\<^sub>S\ \ (\\<^sub>S \ s)) [\\<^sub>O\ \ op'. op' \ ops'] \ are_all_operator_effects_consistent [\\<^sub>O\ \ op'. op' \ ops']" proof - let ?\ = "\ \" and ?s' = "\\<^sub>S \ s" let ?vs' = "strips_problem.variables_of ?\" and ?ops' = "strips_problem.operators_of ?\" and ?vs = "variables_of \" and ?D = "range_of \" and ?s = "\\<^sub>S\ \ ?s'" and ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" have nb: "\(v, a) \ dom ?s'. \(v, a') \ dom (\\<^sub>S \ s). ?s' (v, a) = Some True \ ?s' (v, a') = Some True \ (v, a) = (v, a')" using state_to_strips_state_effect_consistent[OF assms(1)] by blast { have "STRIPS_Semantics.are_all_operators_applicable ?s' ops'" using assms(3) by simp moreover have "list_all (\op. map_of (precondition_of op) \\<^sub>m ?s) ?ops" using sas_plus_equivalent_to_strips_i_a_I[OF assms(1) assms(2)] calculation unfolding list_all_iff by blast moreover have "list_all (\op. list_all (are_operator_effects_consistent op) ?ops) ?ops" using sas_plus_equivalent_to_strips_i_a_II assms nb unfolding are_all_operator_effects_consistent_def is_valid_operator_sas_plus_def list_all_iff by blast ultimately have "are_all_operators_applicable_in ?s ?ops" unfolding are_all_operators_applicable_in_def is_operator_applicable_in_def list_all_iff by argo } moreover have "are_all_operator_effects_consistent ?ops" using sas_plus_equivalent_to_strips_i_a_II assms nb by simp ultimately show ?thesis by simp qed (* TODO: \<^item> prune premises + make private. \<^item> decrement suffixes *) lemma sas_plus_equivalent_to_strips_i_a_VI: assumes "is_valid_problem_sas_plus \" and "dom s \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom s. the (s v) \ \\<^sub>+ \ v" and "set ops' \ set ((\ \)\<^sub>\)" and "are_all_operators_applicable_in s [\\<^sub>O\ \ op'. op' \ ops'] \ are_all_operator_effects_consistent [\\<^sub>O\ \ op'. op' \ ops']" shows "STRIPS_Semantics.are_all_operators_applicable (\\<^sub>S \ s) ops'" proof - let ?vs = "variables_of \" and ?D = "range_of \" and ?\ = "\ \" and ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" and ?s' = "\\<^sub>S \ s" \ \ TODO refactor. \ { fix op' assume "op' \ set ops'" moreover obtain op where "op \ set ?ops" and "op = \\<^sub>O\ \ op'" using calculation by force moreover obtain op'' where "op'' \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op''" using assms(4) calculation(1) by auto moreover have "is_valid_operator_sas_plus \ op''" using is_valid_problem_sas_plus_then(2) assms(1) calculation(4) unfolding is_valid_operator_sas_plus_def by auto moreover have "op = op''" using sas_plus_operator_inverse_is[OF assms(1)] calculation(3, 4, 5) by blast ultimately have "\op \ set ?ops. op \ set ?ops \ op = (\\<^sub>O\ \ op') \ is_valid_operator_sas_plus \ op" by blast } note nb\<^sub>1 = this have nb\<^sub>2: "\(v, a) \ dom ?s'. \(v, a') \ dom ?s'. ?s' (v, a) = Some True \ ?s' (v, a') = Some True \ (v, a) = (v, a')" using state_to_strips_state_effect_consistent[OF assms(1), of _ _ s] by blast { fix op assume "op \ set ?ops" hence "map_of (precondition_of op) \\<^sub>m s" using assms(5) unfolding are_all_operators_applicable_in_def is_operator_applicable_in_def list_all_iff by blast } note nb\<^sub>3 = this { fix op' assume "op' \ set ops'" then obtain op where op_in_ops: "op \ set ?ops" and op_is: "op = (\\<^sub>O\ \ op')" and is_valid_operator_op: "is_valid_operator_sas_plus \ op" using nb\<^sub>1 by force moreover have preconditions_are_consistent: "\(v, a) \ set (precondition_of op). \(v', a') \ set (precondition_of op). v \ v' \ a = a'" using is_valid_operator_sas_plus_then(5) calculation(3) unfolding is_valid_operator_sas_plus_def by fast moreover { fix v a assume "(v, a) \ set (strips_operator.precondition_of op')" moreover have v_a_in_precondition_of_op: "(v, a) \ set (precondition_of op)" using op_is calculation unfolding SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def by auto moreover have "map_of (precondition_of op) v = Some a" using map_of_constant_assignments_defined_if[OF preconditions_are_consistent calculation(2)] by blast moreover have s_of_v_is: "s v = Some a" using nb\<^sub>3[OF op_in_ops] calculation(3) unfolding map_le_def by force moreover have "v \ set ((\)\<^sub>\\<^sub>+)" and "a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then(1, 2) is_valid_operator_op v_a_in_precondition_of_op unfolding is_valid_operator_sas_plus_def SAS_Plus_Representation.is_valid_operator_sas_plus_def Let_def list_all_iff ListMem_iff by auto+ moreover have "(v, a) \ dom ?s'" using state_to_strips_state_dom_is[OF assms(1)] s_of_v_is calculation by simp moreover have "(\\<^sub>S\ \ ?s') v = Some a" using strips_state_to_state_inverse_is[OF assms(1, 2, 3)] s_of_v_is by argo \ \ TODO slow. \ ultimately have "?s' (v, a) = Some True" using strips_state_to_state_range_is[OF assms(1)] nb\<^sub>2 by auto } ultimately have "\(v, a) \ set (strips_operator.precondition_of op'). ?s' (v, a) = Some True" by fast } thus ?thesis unfolding are_all_operators_applicable_def is_operator_applicable_in_def STRIPS_Representation.is_operator_applicable_in_def list_all_iff by simp qed (* TODO Prune premises. *) lemma sas_plus_equivalent_to_strips_i_a_VII: assumes "is_valid_problem_sas_plus \" and "dom s \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom s. the (s v) \ \\<^sub>+ \ v" and "set ops' \ set ((\ \)\<^sub>\)" and "are_all_operators_applicable_in s [\\<^sub>O\ \ op'. op' \ ops'] \ are_all_operator_effects_consistent [\\<^sub>O\ \ op'. op' \ ops']" shows "STRIPS_Semantics.are_all_operator_effects_consistent ops'" proof - let ?s' = "\\<^sub>S \ s" and ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" and ?D = "range_of \" and ?\ = "\ \" \ \ TODO refactor. \ { fix op' assume "op' \ set ops'" moreover obtain op where "op \ set ?ops" and "op = \\<^sub>O\ \ op'" using calculation by force moreover obtain op'' where "op'' \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op''" using assms(4) calculation(1) by auto moreover have "is_valid_operator_sas_plus \ op''" using is_valid_problem_sas_plus_then(2) assms(1) calculation(4) unfolding is_valid_operator_sas_plus_def by auto moreover have "op = op''" using sas_plus_operator_inverse_is[OF assms(1)] calculation(3, 4, 5) by blast ultimately have "\op \ set ?ops. op \ set ?ops \ op' = (\\<^sub>O \ op) \ is_valid_operator_sas_plus \ op" by blast } note nb\<^sub>1 = this { fix op\<^sub>1' op\<^sub>2' assume "op\<^sub>1' \ set ops'" and "op\<^sub>2' \ set ops'" and "\(v, a) \ set (add_effects_of op\<^sub>1'). \(v', a') \ set (delete_effects_of op\<^sub>2'). (v, a) = (v', a')" moreover obtain op\<^sub>1 op\<^sub>2 where "op\<^sub>1 \ set ?ops" and "op\<^sub>1' = \\<^sub>O \ op\<^sub>1" and "is_valid_operator_sas_plus \ op\<^sub>1" and "op\<^sub>2 \ set ?ops" and "op\<^sub>2' = \\<^sub>O \ op\<^sub>2" and is_valid_op\<^sub>2: "is_valid_operator_sas_plus \ op\<^sub>2" using nb\<^sub>1 calculation(1, 2) by meson moreover obtain v v' a a' where "(v, a) \ set (add_effects_of op\<^sub>1')" and "(v', a') \ set (delete_effects_of op\<^sub>2')" and "(v, a) = (v', a')" using calculation by blast moreover have "(v, a) \ set (effect_of op\<^sub>1)" using calculation(5, 10) unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by fastforce moreover have "v = v'" and "a = a'" using calculation(12) by simp+ \ \ The next proof block shows that \(v', a')\ is constructed from an effect \(v'', a'')\ s.t. \a' \ a''\. \ moreover { (* TODO slow. *) have "(v', a') \ (\(v'', a'') \ set (effect_of op\<^sub>2). { (v'', a''') | a'''. a''' \ (\\<^sub>+ \ v'') \ a''' \ a'' })" using sasp_op_to_strips_set_delete_effects_is calculation(8, 11) is_valid_op\<^sub>2 by blast then obtain v'' a'' where "(v'', a'') \ set (effect_of op\<^sub>2)" and "(v', a') \ { (v'', a''') | a'''. a''' \ (\\<^sub>+ \ v'') \ a''' \ a'' }" by blast moreover have "(v', a'') \ set (effect_of op\<^sub>2)" using calculation by blast moreover have "a' \ \\<^sub>+ \ v''" and "a' \ a''" using calculation(1, 2) by fast+ ultimately have "\a''. (v', a'') \ set (effect_of op\<^sub>2) \ a' \ (\\<^sub>+ \ v') \ a' \ a''" by blast } moreover obtain a'' where "(v', a'') \ set (effect_of op\<^sub>2)" and "a' \ \\<^sub>+ \ v'" and "a' \ a''" using calculation(16) by blast moreover have "\(v, a) \ set (effect_of op\<^sub>1). (\(v', a') \ set (effect_of op\<^sub>2). v = v' \ a \ a')" using calculation(13, 14, 15, 17, 19) by blast moreover have "\are_operator_effects_consistent op\<^sub>1 op\<^sub>2" unfolding are_operator_effects_consistent_def list_all_iff using calculation(20) by fastforce ultimately have "\are_all_operator_effects_consistent ?ops" unfolding are_all_operator_effects_consistent_def list_all_iff by meson } note nb\<^sub>2 = this { fix op\<^sub>1' op\<^sub>2' assume op\<^sub>1'_in_ops: "op\<^sub>1' \ set ops'" and op\<^sub>2'_in_ops: "op\<^sub>2' \ set ops'" have "STRIPS_Semantics.are_operator_effects_consistent op\<^sub>1' op\<^sub>2'" proof (rule ccontr) assume "\STRIPS_Semantics.are_operator_effects_consistent op\<^sub>1' op\<^sub>2'" then consider (A) "\(v, a) \ set (add_effects_of op\<^sub>1'). \(v', a') \ set (delete_effects_of op\<^sub>2'). (v, a) = (v', a')" | (B) "\(v, a) \ set (add_effects_of op\<^sub>2'). \(v', a') \ set (delete_effects_of op\<^sub>1'). (v, a) = (v', a')" unfolding STRIPS_Semantics.are_operator_effects_consistent_def list_ex_iff by fastforce thus False using nb\<^sub>2[OF op\<^sub>1'_in_ops op\<^sub>2'_in_ops] nb\<^sub>2[OF op\<^sub>2'_in_ops op\<^sub>1'_in_ops] assms(5) by (cases, argo, force) qed } thus ?thesis unfolding STRIPS_Semantics.are_all_operator_effects_consistent_def STRIPS_Semantics.are_operator_effects_consistent_def list_all_iff by blast qed lemma sas_plus_equivalent_to_strips_i_a_VIII: assumes "is_valid_problem_sas_plus \" and "dom s \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom s. the (s v) \ \\<^sub>+ \ v" and "set ops' \ set ((\ \)\<^sub>\)" and "are_all_operators_applicable_in s [\\<^sub>O\ \ op'. op' \ ops'] \ are_all_operator_effects_consistent [\\<^sub>O\ \ op'. op' \ ops']" shows "STRIPS_Semantics.are_all_operators_applicable (\\<^sub>S \ s) ops' \ STRIPS_Semantics.are_all_operator_effects_consistent ops'" using sas_plus_equivalent_to_strips_i_a_VI sas_plus_equivalent_to_strips_i_a_VII assms by fastforce (* TODO refactor. *) lemma sas_plus_equivalent_to_strips_i_a_IX: assumes "dom s \ V" and "\op \ set ops. \(v, a) \ set (effect_of op). v \ V" shows "dom (execute_parallel_operator_sas_plus s ops) \ V" proof - show ?thesis using assms proof (induction ops arbitrary: s) case Nil then show ?case unfolding execute_parallel_operator_sas_plus_def by simp next case (Cons op ops) let ?s' = "s ++ map_of (effect_of op)" \ \ TODO Wrap IH instantiation in block. \ { have "\(v, a) \ set (effect_of op). v \ V" using Cons.prems(2) by fastforce moreover have "fst ` set (effect_of op) \ V" using calculation by fastforce ultimately have "dom ?s' \ V" unfolding dom_map_add dom_map_of_conv_image_fst using Cons.prems(1) by blast } moreover have "\op \ set ops. \(v, a) \ set (effect_of op). v \ V" using Cons.prems(2) by fastforce ultimately have "dom (execute_parallel_operator_sas_plus ?s' ops) \ V" using Cons.IH[of ?s'] by fast thus ?case unfolding execute_parallel_operator_sas_plus_cons. qed qed \ \ NOTE Show that the domain value constraint on states is monotonous w.r.t. to valid operator execution. I.e. if a parallel operator is executed on a state for which the domain value constraint holds, the domain value constraint will also hold on the resultant state. \ (* TODO refactor. TODO Rewrite lemma without domain function, i.e. \set (the (D v)) \ D\ *) lemma sas_plus_equivalent_to_strips_i_a_X: assumes "dom s \ V" and "V \ dom D" and "\v \ dom s. the (s v) \ set (the (D v))" and "\op \ set ops. \(v, a) \ set (effect_of op). v \ V \ a \ set (the (D v))" shows "\v \ dom (execute_parallel_operator_sas_plus s ops). the (execute_parallel_operator_sas_plus s ops v) \ set (the (D v))" proof - show ?thesis using assms proof (induction ops arbitrary: s) case Nil then show ?case unfolding execute_parallel_operator_sas_plus_def by simp next case (Cons op ops) let ?s' = "s ++ map_of (effect_of op)" { { have "\(v, a) \ set (effect_of op). v \ V" using Cons.prems(4) by fastforce moreover have "fst ` set (effect_of op) \ V" using calculation by fastforce ultimately have "dom ?s' \ V" unfolding dom_map_add dom_map_of_conv_image_fst using Cons.prems(1) by blast } moreover { fix v assume v_in_dom_s': "v \ dom ?s'" hence "the (?s' v) \ set (the (D v))" proof (cases "v \ dom (map_of (effect_of op))") case True moreover have "?s' v = (map_of (effect_of op)) v" unfolding map_add_dom_app_simps(1)[OF True] by blast moreover obtain a where "(map_of (effect_of op)) v = Some a" using calculation(1) by fast moreover have "(v, a) \ set (effect_of op)" using map_of_SomeD calculation(3) by fast moreover have "a \ set (the (D v))" using Cons.prems(4) calculation(4) by fastforce ultimately show ?thesis by force next case False then show ?thesis unfolding map_add_dom_app_simps(3)[OF False] using Cons.prems(3) v_in_dom_s' by fast qed } moreover have "\op \ set ops. \(v, a) \ set (effect_of op). v \ V \ a \ set (the (D v))" using Cons.prems(4) by auto ultimately have "\v \ dom (execute_parallel_operator_sas_plus ?s' ops). the (execute_parallel_operator_sas_plus ?s' ops v) \ set (the (D v))" using Cons.IH[of "s ++ map_of (effect_of op)", OF _ Cons.prems(2)] by meson } thus ?case unfolding execute_parallel_operator_sas_plus_cons by blast qed qed lemma transfom_sas_plus_problem_to_strips_problem_operators_valid: assumes "is_valid_problem_sas_plus \" and "op' \ set ((\ \)\<^sub>\)" obtains op where "op \ set ((\)\<^sub>\\<^sub>+)" and "op' = (\\<^sub>O \ op)" "is_valid_operator_sas_plus \ op" proof - { obtain op where "op \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op" using assms by auto moreover have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) assms(1) calculation(1) by auto ultimately have "\op \ set ((\)\<^sub>\\<^sub>+). op' = (\\<^sub>O \ op) \ is_valid_operator_sas_plus \ op" by blast } thus ?thesis using that by blast qed lemma sas_plus_equivalent_to_strips_i_a_XI: assumes "is_valid_problem_sas_plus \" and "op' \ set ((\ \)\<^sub>\)" shows "(\\<^sub>S \ s) ++ map_of (effect_to_assignments op') = \\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op')))" proof - let ?\ = "\ \" let ?vs = "variables_of \" and?ops = "operators_of \" and ?ops' = "strips_problem.operators_of ?\" let ?s' = "\\<^sub>S \ s" let ?t = "?s' ++ map_of (effect_to_assignments op')" and ?t' = "\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op')))" obtain op where op'_is: "op' = (\\<^sub>O \ op)" and op_in_ops: "op \ set ((\)\<^sub>\\<^sub>+)" and is_valid_operator_op: "is_valid_operator_sas_plus \ op" using transfom_sas_plus_problem_to_strips_problem_operators_valid[OF assms] by auto have nb\<^sub>1: "(\\<^sub>O\ \ op') = op" using sas_plus_operator_inverse_is[OF assms(1)] op'_is op_in_ops by blast \ \ TODO refactor. \ { (*have "fst ` set (effect_to_assignments op') \ fst ` ((\v. (v, True)) ` set (add_effects_of op') \ (\v. (v, False)) ` set (delete_effects_of op'))" by auto then*) have "dom (map_of (effect_to_assignments op')) = set (strips_operator.add_effects_of op') \ set (strips_operator.delete_effects_of op')" unfolding dom_map_of_conv_image_fst by force \ \ TODO slow.\ also have "\ = set (effect_of op) \ set (strips_operator.delete_effects_of op')" using op'_is unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by auto \ \ TODO slow.\ finally have "dom (map_of (effect_to_assignments op')) = set (effect_of op) \ (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" using sasp_op_to_strips_set_delete_effects_is[OF is_valid_operator_op] op'_is by argo } note nb\<^sub>2 = this have nb\<^sub>3: "dom ?t = dom ?s' \ set (effect_of op) \ (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" unfolding nb\<^sub>2 dom_map_add by blast \ \ TODO refactor. \ have nb\<^sub>4: "dom (s ++ map_of (effect_of (\\<^sub>O\ \ op'))) = dom s \ fst ` set (effect_of op)" unfolding dom_map_add dom_map_of_conv_image_fst nb\<^sub>1 by fast { let ?u = "s ++ map_of (effect_of (\\<^sub>O\ \ op'))" have "dom ?t' = (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ ?u v \ None }. { (v, a) | a. a \ \\<^sub>+ \ v })" using state_to_strips_state_dom_is[OF assms(1)] by presburger } note nb\<^sub>5 = this \ \ TODO refactor. \ have nb\<^sub>6: "set (add_effects_of op') = set (effect_of op)" using op'_is unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def by auto \ \ TODO refactor. \ have nb\<^sub>7: "set (delete_effects_of op') = (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" using sasp_op_to_strips_set_delete_effects_is[OF is_valid_operator_op] op'_is by argo \ \ TODO refactor. \ { let ?Add = "set (effect_of op)" let ?Delete = "(\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" have dom_add: "dom (map_of (map (\v. (v, True)) (add_effects_of op'))) = ?Add" unfolding dom_map_of_conv_image_fst set_map image_comp comp_apply using nb\<^sub>6 by simp have dom_delete: "dom (map_of (map (\v. (v, False)) (delete_effects_of op'))) = ?Delete" unfolding dom_map_of_conv_image_fst set_map image_comp comp_apply using nb\<^sub>7 by auto { { fix v a assume v_a_in_dom_add: "(v, a) \ dom (map_of (map (\v. (v, True)) (add_effects_of op')))" have "(v, a) \ dom (map_of (map (\v. (v, False)) (delete_effects_of op')))" proof (rule ccontr) assume "\((v, a) \ dom (map_of (map (\v. (v, False)) (delete_effects_of op'))))" then have "(v, a) \ ?Delete" and "(v, a) \ ?Add" using dom_add dom_delete v_a_in_dom_add by argo+ moreover have "\(v', a') \ ?Add. v \ v' \ a = a'" using is_valid_operator_sas_plus_then(6) is_valid_operator_op calculation(2) unfolding is_valid_operator_sas_plus_def by fast ultimately show False by fast qed } hence "disjnt (dom (map_of (map (\v. (v, True)) (add_effects_of op')))) (dom (map_of (map (\v. (v, False)) (delete_effects_of op'))))" unfolding disjnt_def Int_def using nb\<^sub>7 by simp } hence "dom (map_of (map (\v. (v, True)) (add_effects_of op'))) = ?Add" and "dom (map_of (map (\v. (v, False)) (delete_effects_of op'))) = ?Delete" and "disjnt (dom (map_of (map (\v. (v, True)) (add_effects_of op')))) (dom (map_of (map (\v. (v, False)) (delete_effects_of op'))))" using dom_add dom_delete by blast+ } note nb\<^sub>8 = this \ \ TODO refactor. \ { let ?Add = "set (effect_of op)" let ?Delete = "(\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" \ \ TODO slow.\ have "\(v, a) \ ?Add. map_of (effect_to_assignments op') (v, a) = Some True" and "\(v, a) \ ?Delete. map_of (effect_to_assignments op') (v, a) = Some False" proof - { fix v a assume "(v, a) \ ?Add" hence "map_of (effect_to_assignments op') (v, a) = Some True" unfolding effect_to_assignments_simp using nb\<^sub>6 map_of_defined_if_constructed_from_list_of_constant_assignments[of "map (\v. (v, True)) (add_effects_of op')" True "add_effects_of op'"] by force } moreover { fix v a assume "(v, a) \ ?Delete" moreover have "(v, a) \ dom (map_of (map (\v. (v, False)) (delete_effects_of op')))" using nb\<^sub>8(2) calculation(1) by argo moreover have "(v, a) \ dom (map_of (map (\v. (v, True)) (add_effects_of op')))" using nb\<^sub>8 unfolding disjnt_def using calculation(1) by blast moreover have "map_of (effect_to_assignments op') (v, a) = map_of (map (\v. (v, False)) (delete_effects_of op')) (v, a)" unfolding effect_to_assignments_simp map_of_append using map_add_dom_app_simps(3)[OF calculation(3)] by presburger \ \ TODO slow. \ ultimately have "map_of (effect_to_assignments op') (v, a) = Some False" using map_of_defined_if_constructed_from_list_of_constant_assignments[ of "map (\v. (v, False)) (delete_effects_of op')" False "delete_effects_of op'"] nb\<^sub>7 by auto } ultimately show "\(v, a) \ ?Add. map_of (effect_to_assignments op') (v, a) = Some True" and "\(v, a) \ ?Delete. map_of (effect_to_assignments op') (v, a) = Some False" by blast+ qed } note nb\<^sub>9 = this { fix v a assume "(v, a) \ set (effect_of op)" moreover have "\(v, a) \ set (effect_of op). \(v', a') \ set (effect_of op). v \ v' \ a = a'" using is_valid_operator_sas_plus_then is_valid_operator_op unfolding is_valid_operator_sas_plus_def by fast ultimately have "map_of (effect_of op) v = Some a" using map_of_constant_assignments_defined_if[of "effect_of op"] by presburger } note nb\<^sub>1\<^sub>0 = this { fix v a assume v_a_in_effect_of_op: "(v, a) \ set (effect_of op)" and "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v \ None" moreover have "v \ set ?vs" using is_valid_operator_op is_valid_operator_sas_plus_then(3) calculation(1) by fastforce moreover { have "is_valid_problem_strips ?\" using is_valid_problem_sas_plus_then_strips_transformation_too assms(1) by blast thm calculation(1) nb\<^sub>6 assms(2) moreover have "set (add_effects_of op') \ set ((?\)\<^sub>\)" using assms(2) is_valid_problem_strips_operator_variable_sets(2) calculation by blast moreover have "(v, a) \ set ((?\)\<^sub>\)" using v_a_in_effect_of_op nb\<^sub>6 calculation(2) by blast ultimately have "a \ \\<^sub>+ \ v" using sas_plus_problem_to_strips_problem_variable_set_element_iff[OF assms(1)] by fast } \ \ TODO slow. \ ultimately have "(v, a) \ dom (\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op'))))" using state_to_strips_state_dom_is[OF assms(1), of "s ++ map_of (effect_of (\\<^sub>O\ \ op'))"] by simp } note nb\<^sub>1\<^sub>1 = this { fix v a assume "(v, a) \ set (effect_of op)" moreover have "v \ dom (map_of (effect_of op))" unfolding dom_map_of_conv_image_fst using calculation by force moreover have "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v = Some a" unfolding map_add_dom_app_simps(1)[OF calculation(2)] nb\<^sub>1 using nb\<^sub>1\<^sub>0 calculation(1) by blast moreover have "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v \ None" using calculation(3) by auto moreover have "(v, a) \ dom (\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op'))))" using nb\<^sub>1\<^sub>1 calculation(1, 4) by presburger ultimately have "(\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op')))) (v, a) = Some True" using state_to_strips_state_range_is[OF assms(1)] by simp } note nb\<^sub>1\<^sub>2 = this { fix v a' assume "(v, a') \ dom (map_of (effect_to_assignments op'))" and "(v, a') \ (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" moreover have "v \ dom (map_of (effect_of op))" unfolding dom_map_of_conv_image_fst using calculation(2) by force moreover have "v \ set ?vs" using calculation(3) is_valid_operator_sas_plus_then(3) is_valid_operator_op unfolding dom_map_of_conv_image_fst is_valid_operator_sas_plus_def by fastforce moreover obtain a where "(v, a) \ set (effect_of op)" and "a' \ \\<^sub>+ \ v" and "a' \ a" using calculation(2) by blast moreover have "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v = Some a" unfolding map_add_dom_app_simps(1)[OF calculation(3)] nb\<^sub>1 using nb\<^sub>1\<^sub>0 calculation(5) by blast moreover have "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v \ None" using calculation(8) by auto \ \ TODO slow. \ moreover have "(v, a') \ dom (\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op'))))" using state_to_strips_state_dom_is[OF assms(1), of "s ++ map_of (effect_of (\\<^sub>O\ \ op'))"] calculation(4, 6, 9) by simp \ \ TODO slow. \ ultimately have "(\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ op')))) (v, a') = Some False" using state_to_strips_state_range_is[OF assms(1), of v a' "s ++ map_of (effect_of (\\<^sub>O\ \ op'))"] by simp } note nb\<^sub>1\<^sub>3 = this { fix v a assume "(v, a) \ dom ?t" and "(v, a) \ dom (map_of (effect_to_assignments op'))" moreover have "(v, a) \ dom ?s'" using calculation(1, 2) unfolding dom_map_add by blast moreover have "?t (v, a) = ?s' (v, a)" unfolding map_add_dom_app_simps(3)[OF calculation(2)].. ultimately have "?t (v, a) = Some (the (s v) = a)" using state_to_strips_state_range_is[OF assms(1)] by presburger } note nb\<^sub>1\<^sub>4 = this { fix v a assume "(v, a) \ dom ?t" and v_a_not_in: "(v, a) \ dom (map_of (effect_to_assignments op'))" moreover have "(v, a) \ dom ?s'" using calculation(1, 2) unfolding dom_map_add by blast moreover have "(v, a) \ (\ v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }. { (v, a) | a. a \ \\<^sub>+ \ v })" using state_to_strips_state_dom_is[OF assms(1)] calculation(3) by presburger moreover have "v \ set ((\)\<^sub>\\<^sub>+)" and "s v \ None" and "a \ \\<^sub>+ \ v" using calculation(4) by blast+ \ \ NOTE Hasn't this been proved before? \ moreover { have "dom (map_of (effect_to_assignments op')) = (\(v, a) \ set (effect_of op). { (v, a) }) \ (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" unfolding nb\<^sub>2 by blast also have "\ = (\(v, a) \ set (effect_of op). { (v, a) } \ { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" by blast finally have "dom (map_of (effect_to_assignments op')) = (\(v, a) \ set (effect_of op). { (v, a) } \ { (v, a) | a. a \ \\<^sub>+ \ v })" by auto then have "(v, a) \ (\(v, a) \ set (effect_of op). { (v, a) | a. a \ \\<^sub>+ \ v })" using v_a_not_in by blast } \ \ TODO slow. \ moreover have "v \ dom (map_of (effect_of op))" using dom_map_of_conv_image_fst calculation by fastforce moreover have "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v = s v" unfolding nb\<^sub>1 map_add_dom_app_simps(3)[OF calculation(9)] by simp \ \ TODO slow. \ moreover have "(v, a) \ dom ?t'" using state_to_strips_state_dom_is[OF assms(1), of "s ++ map_of (effect_of (\\<^sub>O\ \ op'))"] calculation(5, 6, 7, 8, 10) by simp ultimately have "?t' (v, a) = Some (the (s v) = a)" using state_to_strips_state_range_is[OF assms(1)] by presburger } note nb\<^sub>1\<^sub>5 = this \ \ TODO refactor. \ have nb\<^sub>1\<^sub>6: "dom ?t = (\ v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None }. { (v, a) | a. a \ (\\<^sub>+ \ v) }) \ set (effect_of op) \ (\(v, a)\set (effect_of op). {(v, a') |a'. a' \ (\\<^sub>+ \ v) \ a' \ a})" unfolding dom_map_add nb\<^sub>2 using state_to_strips_state_dom_is[OF assms(1), of s] by auto { { fix v a assume "(v, a) \ dom ?t" then consider (A) "(v, a) \ dom (\\<^sub>S \ s)" | (B) "(v, a) \ dom (map_of (effect_to_assignments op'))" by fast hence "(v, a) \ dom ?t'" proof (cases) case A then have "v \ set ((\)\<^sub>\\<^sub>+)" and "s v \ None" and "a \ \\<^sub>+ \ v" unfolding state_to_strips_state_dom_element_iff[OF assms(1)] by blast+ thm map_add_None state_to_strips_state_dom_element_iff[OF assms(1)] moreover have "(s ++ map_of (effect_of (\\<^sub>O\ \ op'))) v \ None" using calculation(2) by simp ultimately show ?thesis unfolding state_to_strips_state_dom_element_iff[OF assms(1)] by blast next case B then have "(v, a) \ set (effect_of op) \ (\(v, a)\set (effect_of op). { (v, a') | a'. a' \ \\<^sub>+ \ v \ a' \ a })" unfolding nb\<^sub>2 by blast then consider (B\<^sub>1) "(v, a) \ set (effect_of op)" | (B\<^sub>2) "(v, a) \ (\(v, a)\set (effect_of op). { (v, a') | a'. a' \ \\<^sub>+ \ v \ a' \ a })" by blast thm nb\<^sub>1\<^sub>2 nb\<^sub>1\<^sub>3 nb\<^sub>2 thus ?thesis proof (cases) case B\<^sub>1 then show ?thesis using nb\<^sub>1\<^sub>2 by fast next case B\<^sub>2 then show ?thesis using nb\<^sub>1\<^sub>3 B by blast qed qed } moreover { let ?u = "s ++ map_of (effect_of (\\<^sub>O\ \ op'))" fix v a assume v_a_in_dom_t': "(v, a) \ dom ?t'" thm nb\<^sub>5 then have v_in_vs: "v \ set ((\)\<^sub>\\<^sub>+)" and u_of_v_is_not_None: "?u v \ None" and a_in_range_of_v: "a \ \\<^sub>+ \ v" using state_to_strips_state_dom_element_iff[OF assms(1)] v_a_in_dom_t' by meson+ { assume "(v, a) \ dom ?t" then have contradiction: "(v, a) \ (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None}. { (v, a) |a. a \ \\<^sub>+ \ v }) \ set (effect_of op) \ (\(v, a)\set (effect_of op). {(v, a') |a'. a' \ \\<^sub>+ \ v \ a' \ a})" unfolding nb\<^sub>1\<^sub>6 by fast hence False proof (cases "map_of (effect_of (\\<^sub>O\ \ op')) v = None") case True then have "s v \ None" using u_of_v_is_not_None by simp then have "(v, a) \ (\v \ { v | v. v \ set ((\)\<^sub>\\<^sub>+) \ s v \ None}. { (v, a) |a. a \ \\<^sub>+ \ v })" using v_in_vs a_in_range_of_v by blast thus ?thesis using contradiction by blast next case False then have "v \ dom (map_of (effect_of op))" using u_of_v_is_not_None nb\<^sub>1 by blast then obtain a' where map_of_effect_of_op_v_is: "map_of (effect_of op) v = Some a'" by blast then have v_a'_in: "(v, a') \ set (effect_of op)" using map_of_SomeD by fast then show ?thesis proof (cases "a = a'") case True then have "(v, a) \ set (effect_of op)" using v_a'_in by blast then show ?thesis using contradiction by blast next case False then have "(v, a) \ (\(v, a)\set (effect_of op). {(v, a') |a'. a' \ \\<^sub>+ \ v \ a' \ a})" using v_a'_in calculation a_in_range_of_v by blast thus ?thesis using contradiction by fast qed qed } hence "(v, a) \ dom ?t" by argo } moreover have "dom ?t \ dom ?t'" and "dom ?t' \ dom ?t" subgoal using calculation(1) subrelI[of "dom ?t" "dom ?t'"] by fast subgoal using calculation(2) subrelI[of "dom ?t'" "dom ?t"] by argo done ultimately have "dom ?t = dom ?t'" by force } note nb\<^sub>1\<^sub>7 = this { fix v a assume v_a_in_dom_t: "(v, a) \ dom ?t" hence "?t (v, a) = ?t' (v, a)" proof (cases "(v, a) \ dom (map_of (effect_to_assignments op'))") case True \ \ TODO slow. \ \ \ NOTE Split on the (disjunct) domain variable sets of @{text "map_of (effect_to_assignments op')"}. \ then consider (A1) "(v, a) \ set (effect_of op)" | (A2) "(v, a) \ (\(v, a) \ set (effect_of op). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" using nb\<^sub>2 by fastforce then show ?thesis proof (cases) case A1 then have "?t (v, a) = Some True" unfolding map_add_dom_app_simps(1)[OF True] using nb\<^sub>9(1) by fast moreover have "?t' (v, a) = Some True" using nb\<^sub>1\<^sub>2[OF A1]. ultimately show ?thesis.. next case A2 then have "?t (v, a) = Some False" unfolding map_add_dom_app_simps(1)[OF True] using nb\<^sub>9(2) by blast moreover have "?t' (v, a) = Some False" using nb\<^sub>1\<^sub>3[OF True A2]. ultimately show ?thesis.. qed next case False moreover have "?t (v, a) = Some (the (s v) = a)" using nb\<^sub>1\<^sub>4[OF v_a_in_dom_t False]. moreover have "?t' (v, a) = Some (the (s v) = a)" using nb\<^sub>1\<^sub>5[OF v_a_in_dom_t False]. ultimately show ?thesis by argo qed } note nb\<^sub>1\<^sub>8 = this moreover { fix v a assume "(v, a) \ dom ?t'" hence "?t (v, a) = ?t' (v, a)" using nb\<^sub>1\<^sub>7 nb\<^sub>1\<^sub>8 by presburger } \ \ TODO slow.\ ultimately have "?t \\<^sub>m ?t'" and "?t' \\<^sub>m ?t" unfolding map_le_def by fastforce+ thus ?thesis using map_le_antisym[of ?t ?t'] by fast qed \ \ NOTE This is the essential step in the SAS+/STRIPS equivalence theorem. We show that executing a given parallel STRIPS operator @{text "ops'"} on the corresponding STRIPS state @{text "s' = \\<^sub>S \ s"} yields the same state as executing the transformed SAS+ parallel operator @{text "ops = [\\<^sub>O\ (\ \) op'. op' \ ops']"} on the original SAS+ state @{text "s"} and the transforming the resultant SAS+ state to its corresponding STRIPS state. \ (* TODO refactor. *) lemma sas_plus_equivalent_to_strips_i_a_XII: assumes "is_valid_problem_sas_plus \" and "\op' \ set ops'. op' \ set ((\ \)\<^sub>\)" shows "execute_parallel_operator (\\<^sub>S \ s) ops' = \\<^sub>S \ (execute_parallel_operator_sas_plus s [\\<^sub>O\ \ op'. op' \ ops'])" using assms proof (induction ops' arbitrary: s) case Nil then show ?case unfolding execute_parallel_operator_def execute_parallel_operator_sas_plus_def by simp next case (Cons op' ops') let ?\ = "\ \" let ?t' = "(\\<^sub>S \ s) ++ map_of (effect_to_assignments op')" and ?t = "s ++ map_of (effect_of (\\<^sub>O\ \ op'))" have nb\<^sub>1: "?t' = \\<^sub>S \ ?t" using sas_plus_equivalent_to_strips_i_a_XI[OF assms(1)] Cons.prems(2) by force { have "\op' \ set ops'. op' \ set (strips_problem.operators_of ?\)" using Cons.prems(2) by simp then have "execute_parallel_operator (\\<^sub>S \ ?t) ops' = \\<^sub>S \ (execute_parallel_operator_sas_plus ?t [\\<^sub>O\ \ x. x \ ops'])" using Cons.IH[OF Cons.prems(1), of ?t] by fastforce hence "execute_parallel_operator ?t' ops' = \\<^sub>S \ (execute_parallel_operator_sas_plus ?t [\\<^sub>O\ \ x. x \ ops'])" using nb\<^sub>1 by argo } thus ?case by simp qed lemma sas_plus_equivalent_to_strips_i_a_XIII: assumes "is_valid_problem_sas_plus \" and "\op' \ set ops'. op' \ set ((\ \)\<^sub>\)" and "(\\<^sub>S \ G) \\<^sub>m execute_parallel_plan (execute_parallel_operator (\\<^sub>S \ I) ops') \" shows "(\\<^sub>S \ G) \\<^sub>m execute_parallel_plan (\\<^sub>S \ (execute_parallel_operator_sas_plus I [\\<^sub>O\ \ op'. op' \ ops'])) \" proof - let ?I' = "(\\<^sub>S \ I)" and ?G' = "\\<^sub>S \ G" and ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" and ?\ = "\ \" let ?J = "execute_parallel_operator_sas_plus I ?ops" { fix v a assume "(v, a) \ dom ?G'" then have "?G' (v, a) = execute_parallel_plan (execute_parallel_operator ?I' ops') \ (v, a)" using assms(3) unfolding map_le_def by auto hence "?G' (v, a) = execute_parallel_plan (\\<^sub>S \ ?J) \ (v, a)" using sas_plus_equivalent_to_strips_i_a_XII[OF assms(1, 2)] by simp } thus ?thesis unfolding map_le_def by fast qed \ \ NOTE This is a more abstract formulation of the proposition in \sas_plus_equivalent_to_strips_i\ which is better suited for induction proofs. We essentially claim that given a plan the execution in STRIPS semantics of which solves the problem of reaching a transformed goal state \\\<^sub>S \ G\ from a transformed initial state \\\<^sub>S \ I\---such as the goal and initial state of an induced STRIPS problem for a SAS+ problem---is equivalent to an execution in SAS+ semantics of the transformed plan \\\<^sub>P\ (\ \) \\ w.r.t to the original initial state \I\ and original goal state \G\. \ lemma sas_plus_equivalent_to_strips_i_a: assumes "is_valid_problem_sas_plus \" and "dom I \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom I. the (I v) \ \\<^sub>+ \ v" and "dom G \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom G. the (G v) \ \\<^sub>+ \ v" and "\ops' \ set \. \op' \ set ops'. op' \ set ((\ \)\<^sub>\)" and "(\\<^sub>S \ G) \\<^sub>m execute_parallel_plan (\\<^sub>S \ I) \" shows "G \\<^sub>m execute_parallel_plan_sas_plus I (\\<^sub>P\ \ \)" proof - let ?vs = "variables_of \" and ?\ = "\\<^sub>P\ \ \" show ?thesis using assms proof (induction \ arbitrary: I) case Nil then have "(\\<^sub>S \ G) \\<^sub>m (\\<^sub>S \ I)" by fastforce then have "G \\<^sub>m I" using state_to_strips_state_map_le_iff[OF assms(1, 4, 5)] by blast thus ?case unfolding SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_parallel_plan_to_sas_plus_parallel_plan_def by fastforce next case (Cons ops' \) let ?D = "range_of \" and ?\ = "\ \" and ?I' = "\\<^sub>S \ I" and ?G' = "\\<^sub>S \ G" let ?ops = "[\\<^sub>O\ \ op'. op' \ ops']" let ?J = "execute_parallel_operator_sas_plus I ?ops" and ?J' = "execute_parallel_operator ?I' ops'" have nb\<^sub>1: "set ops' \ set ((?\)\<^sub>\)" using Cons.prems(6) unfolding STRIPS_Semantics.is_parallel_solution_for_problem_def list_all_iff ListMem_iff by fastforce { fix op assume "op \ set ?ops" moreover obtain op' where "op' \ set ops'" and "op = \\<^sub>O\ \ op'" using calculation by auto moreover have "op' \ set ((?\)\<^sub>\)" using nb\<^sub>1 calculation(2) by blast moreover obtain op'' where "op'' \ set ((\)\<^sub>\\<^sub>+)" and "op' = \\<^sub>O \ op''" using calculation(4) by auto moreover have "op = op''" using sas_plus_operator_inverse_is[OF assms(1) calculation(5)] calculation(3, 6) by presburger ultimately have "op \ set ((\)\<^sub>\\<^sub>+) \ (\op' \ set ops'. op' = \\<^sub>O \ op)" by blast } note nb\<^sub>2 = this { fix op v a assume "op \ set ((\)\<^sub>\\<^sub>+)" and "(v, a) \ set (effect_of op)" moreover have "op \ set ((\)\<^sub>\\<^sub>+)" using nb\<^sub>2 calculation(1) by blast moreover have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) Cons.prems(1) calculation(3) by blast ultimately have "v \ set ((\)\<^sub>\\<^sub>+)" using is_valid_operator_sas_plus_then(3) by fastforce } note nb\<^sub>3 = this { fix op assume "op \ set ?ops" then have "op \ set ((\)\<^sub>\\<^sub>+)" using nb\<^sub>2 by blast then have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) Cons.prems(1) by blast hence "\(v, a) \ set (effect_of op). v \ set ((\)\<^sub>\\<^sub>+) \ a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then(3,4) by fast } note nb\<^sub>4 = this show ?case proof (cases "STRIPS_Semantics.are_all_operators_applicable ?I' ops' \ STRIPS_Semantics.are_all_operator_effects_consistent ops'") case True { { have "dom I \ set ((\)\<^sub>\\<^sub>+)" using Cons.prems(2) by blast hence "(\\<^sub>S\ \ ?I') = I" using strips_state_to_state_inverse_is[OF Cons.prems(1) _ Cons.prems(3)] by argo } then have "are_all_operators_applicable_in I ?ops \ are_all_operator_effects_consistent ?ops" using sas_plus_equivalent_to_strips_i_a_IV[OF assms(1) nb\<^sub>1, of I] True by simp moreover have "(\\<^sub>P\ \ (ops' # \)) = ?ops # (\\<^sub>P\ \ \)" unfolding SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def by simp ultimately have "execute_parallel_plan_sas_plus I (\\<^sub>P\ \ (ops' # \)) = execute_parallel_plan_sas_plus ?J (\\<^sub>P\ \ \)" by force } note nb\<^sub>5 = this \ \ Show the goal using the IH. \ { have dom_J_subset_eq_vs: "dom ?J \ set ((\)\<^sub>\\<^sub>+)" using sas_plus_equivalent_to_strips_i_a_IX[OF Cons.prems(2)] nb\<^sub>2 nb\<^sub>4 by blast moreover { have "set ((\)\<^sub>\\<^sub>+) \ dom (range_of \)" using is_valid_problem_sas_plus_then(1)[OF assms(1)] by fastforce moreover have "\v \ dom I. the (I v) \ set (the (range_of \ v))" using Cons.prems(2, 3) assms(1) set_the_range_of_is_range_of_sas_plus_if by force moreover have "\op \ set ?ops. \(v, a) \ set (effect_of op). v \ set ((\)\<^sub>\\<^sub>+) \ a \ set (the (?D v))" using set_the_range_of_is_range_of_sas_plus_if assms(1) nb\<^sub>4 by fastforce moreover have v_in_dom_J_range: "\v \ dom ?J. the (?J v) \ set (the (?D v))" using sas_plus_equivalent_to_strips_i_a_X[of I "set ((\)\<^sub>\\<^sub>+)" ?D ?ops, OF Cons.prems(2)] calculation(1, 2, 3) by fastforce { fix v assume "v \ dom ?J" moreover have "v \ set ((\)\<^sub>\\<^sub>+)" using nb\<^sub>2 calculation dom_J_subset_eq_vs by blast moreover have "set (the (range_of \ v)) = \\<^sub>+ \ v" using set_the_range_of_is_range_of_sas_plus_if[OF assms(1)] calculation(2) by presburger ultimately have "the (?J v) \ \\<^sub>+ \ v" using nb\<^sub>3 v_in_dom_J_range by blast } ultimately have "\v \ dom ?J. the (?J v) \ \\<^sub>+ \ v" by fast } moreover have "\ops' \ set \. \op'\set ops'. op' \ set ((\ \)\<^sub>\)" using Cons.prems(6) by simp moreover { have "?G' \\<^sub>m execute_parallel_plan ?J' \" using Cons.prems(7) True by auto hence "(\\<^sub>S \ G) \\<^sub>m execute_parallel_plan (\\<^sub>S \ ?J) \" using sas_plus_equivalent_to_strips_i_a_XIII[OF Cons.prems(1)] nb\<^sub>1 by blast } ultimately have "G \\<^sub>m execute_parallel_plan_sas_plus I (\\<^sub>P\ \ (ops' # \))" using Cons.IH[of ?J, OF Cons.prems(1) _ _ Cons.prems(4, 5)] Cons.prems(6) nb\<^sub>5 by presburger } thus ?thesis. next case False then have "?G' \\<^sub>m ?I'" using Cons.prems(7) by force moreover { have "dom I \ set ?vs" using Cons.prems(2) by simp hence "\(are_all_operators_applicable_in I ?ops \ are_all_operator_effects_consistent ?ops)" using sas_plus_equivalent_to_strips_i_a_VIII[OF Cons.prems(1) _ Cons.prems(3) nb\<^sub>1] False by force } moreover { have "(\\<^sub>P\ \ (ops' # \)) = ?ops # (\\<^sub>P\ \ \)" unfolding SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def by simp hence "G \\<^sub>m execute_parallel_plan_sas_plus I (?ops # (\\<^sub>P\ \ \)) \ G \\<^sub>m I" using calculation(2) by force } ultimately show ?thesis using state_to_strips_state_map_le_iff[OF Cons.prems(1, 4, 5)] unfolding SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def by force qed qed qed \ \ NOTE Show that a solution for the induced STRIPS problem for the given valid SAS+ problem, corresponds to a solution for the given SAS+ problem. Note that in the context of the SAS+ problem solving pipeline, we \begin{enumerate} \item convert the given valid SAS+ @{text "\"} problem to the corresponding STRIPS problem @{text "\"} (this is implicitely also valid by lemma @{text "is_valid_problem_sas_plus_then_strips_transformation_too"}); then, \item get a solution @{text "\"}---if it exists---for the induced STRIPS problem by executing SATPlan; and finally, \item convert @{text "\"} back to a solution @{text "\"} for the SAS+ problem. \end{enumerate} \ lemma sas_plus_equivalent_to_strips_i: assumes "is_valid_problem_sas_plus \" and "STRIPS_Semantics.is_parallel_solution_for_problem (\ \) \" shows "goal_of \ \\<^sub>m execute_parallel_plan_sas_plus (sas_plus_problem.initial_of \) (\\<^sub>P\ \ \)" proof - let ?vs = "variables_of \" and ?I = "initial_of \" and ?G = "goal_of \" let ?\ = "\ \" let ?G' = "strips_problem.goal_of ?\" and ?I' = "strips_problem.initial_of ?\" let ?\ = "\\<^sub>P\ \ \" have "dom ?I \ set ?vs" using is_valid_problem_sas_plus_then(3) assms(1) by auto moreover have "\v \ dom ?I. the (?I v) \ \\<^sub>+ \ v" using is_valid_problem_sas_plus_then(4) assms(1) calculation by auto moreover have "dom ?G \ set ?vs" and "\v \ dom ?G. the (?G v) \ \\<^sub>+ \ v" using is_valid_problem_sas_plus_then(5, 6) assms(1) by blast+ moreover have "\ops'\set \. \op'\set ops'. op' \ set ((?\)\<^sub>\)" using is_parallel_solution_for_problem_operator_set[OF assms(2)] by simp moreover { have "?G' \\<^sub>m execute_parallel_plan ?I' \" using assms(2) unfolding STRIPS_Semantics.is_parallel_solution_for_problem_def.. moreover have "?G' = \\<^sub>S \ ?G" and "?I' = \\<^sub>S \ ?I" by simp+ ultimately have "(\\<^sub>S \ ?G) \\<^sub>m execute_parallel_plan (\\<^sub>S \ ?I) \" by simp } ultimately show ?thesis using sas_plus_equivalent_to_strips_i_a[OF assms(1)] by simp qed \ \ NOTE Show that the operators for a given solution @{text "\"} to the induced STRIPS problem for a given SAS+ problem correspond to operators of the SAS+ problem. \ lemma sas_plus_equivalent_to_strips_ii: assumes "is_valid_problem_sas_plus \" and "STRIPS_Semantics.is_parallel_solution_for_problem (\ \) \" shows "list_all (list_all (\op. ListMem op (operators_of \))) (\\<^sub>P\ \ \)" proof - let ?\ = "\ \" let ?ops = "operators_of \" and ?\ = "\\<^sub>P\ \ \" have "is_valid_problem_strips ?\" using is_valid_problem_sas_plus_then_strips_transformation_too[OF assms(1)] by simp have nb\<^sub>1: "\op' \ set ((?\)\<^sub>\). (\op \ set ?ops. op' = (\\<^sub>O \ op))" by auto { fix ops' op' op assume "ops' \ set \" and "op' \ set ops'" then have "op' \ set (strips_problem.operators_of ?\)" using is_parallel_solution_for_problem_operator_set[OF assms(2)] by simp then obtain op where "op \ set ((\)\<^sub>\\<^sub>+)" and "op' = (\\<^sub>O \ op)" by auto then have "(\\<^sub>O\ \ op') \ set ((\)\<^sub>\\<^sub>+)" using sas_plus_operator_inverse_is[OF assms(1)] by presburger } thus ?thesis unfolding list_all_iff ListMem_iff strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_op_to_sasp_def strips_op_to_sasp_def by auto qed text \ We now show that for a parallel solution \<^term>\\\ of \<^term>\\\ the SAS+ plan \<^term>\\ \ \\<^sub>P\ \ \\ yielded by the STRIPS to SAS+ plan transformation is a solution for \<^term>\\\. The proof uses the definition of parallel STRIPS solutions and shows that the execution of \<^term>\\\ on the initial state of the SAS+ problem yields a state satisfying the problem's goal state, i.e. @{text[display, indent=4]"G \\<^sub>m execute_parallel_plan_sas_plus I \"} and by showing that all operators in all parallel operators of \<^term>\\\ are operators of the problem. \ theorem sas_plus_equivalent_to_strips: assumes "is_valid_problem_sas_plus \" and "STRIPS_Semantics.is_parallel_solution_for_problem (\ \) \" shows "is_parallel_solution_for_problem \ (\\<^sub>P\ \ \)" proof - let ?I = "initial_of \" and ?G = "goal_of \" and ?ops = "operators_of \" and ?\ = "\\<^sub>P\ \ \" show ?thesis unfolding is_parallel_solution_for_problem_def Let_def proof (rule conjI) show "?G \\<^sub>m execute_parallel_plan_sas_plus ?I ?\" using sas_plus_equivalent_to_strips_i[OF assms]. next show "list_all (list_all (\op. ListMem op ?ops)) ?\" using sas_plus_equivalent_to_strips_ii[OF assms]. qed qed private lemma strips_equivalent_to_sas_plus_i_a_I: assumes "is_valid_problem_sas_plus \" and "\op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" and "op' \ set [\\<^sub>O \ op. op \ ops]" obtains op where "op \ set ops" and "op' = \\<^sub>O \ op" proof - let ?\ = "\ \" let ?ops = "operators_of \" obtain op where "op \ set ops" and "op' = \\<^sub>O \ op" using assms(3) by auto thus ?thesis using that by blast qed private corollary strips_equivalent_to_sas_plus_i_a_II: assumes"is_valid_problem_sas_plus \" and "\op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" and "op' \ set [\\<^sub>O \ op. op \ ops]" shows "op' \ set ((\ \)\<^sub>\)" and "is_valid_operator_strips (\ \) op'" proof - let ?\ = "\ \" let ?ops = "operators_of \" and ?ops' = "strips_problem.operators_of ?\" obtain op where op_in: "op \ set ops" and op'_is: "op' = \\<^sub>O \ op" using strips_equivalent_to_sas_plus_i_a_I[OF assms]. then have nb: "op' \ set ((\ \)\<^sub>\)" using assms(2) op_in op'_is by fastforce thus "op' \ set ((\ \)\<^sub>\)" and "is_valid_operator_strips ?\ op'" proof - have "\op' \ set ?ops'. is_valid_operator_strips ?\ op'" using is_valid_problem_sas_plus_then_strips_transformation_too_iii[OF assms(1)] unfolding list_all_iff. thus "is_valid_operator_strips ?\ op'" using nb by fastforce qed fastforce qed (* TODO make private *) lemma strips_equivalent_to_sas_plus_i_a_III: assumes "is_valid_problem_sas_plus \" and "\op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" shows "execute_parallel_operator (\\<^sub>S \ s) [\\<^sub>O \ op. op \ ops] = (\\<^sub>S \ (execute_parallel_operator_sas_plus s ops))" proof - { fix op s assume "op \ set ((\)\<^sub>\\<^sub>+)" moreover have "(\\<^sub>O \ op) \ set ((\ \)\<^sub>\)" using calculation by simp moreover have "(\\<^sub>S \ s) ++ map_of (effect_to_assignments (\\<^sub>O \ op)) = (\\<^sub>S \ (s ++ map_of (effect_of (\\<^sub>O\ \ (\\<^sub>O \ op)))))" using sas_plus_equivalent_to_strips_i_a_XI[OF assms(1) calculation(2)] by blast moreover have "(\\<^sub>O\ \ (\\<^sub>O \ op)) = op" using sas_plus_operator_inverse_is[OF assms(1) calculation(1)]. ultimately have "(\\<^sub>S \ s) \ (\\<^sub>O \ op) = (\\<^sub>S \ (s \\<^sub>+ op))" unfolding execute_operator_def execute_operator_sas_plus_def by simp } note nb\<^sub>1 = this show ?thesis using assms proof (induction ops arbitrary: s) case Nil then show ?case unfolding execute_parallel_operator_def execute_parallel_operator_sas_plus_def by simp next case (Cons op ops) let ?t = "s \\<^sub>+ op" let ?s' = "\\<^sub>S \ s" and ?ops' = "[\\<^sub>O \ op. op \ op # ops]" let ?t' = "?s' \ (\\<^sub>O \ op)" have "execute_parallel_operator ?s' ?ops' = execute_parallel_operator ?t' [\\<^sub>O \ x. x \ ops]" unfolding execute_operator_def by simp moreover have "(\\<^sub>S \ (execute_parallel_operator_sas_plus s (op # ops))) = (\\<^sub>S \ (execute_parallel_operator_sas_plus ?t ops))" unfolding execute_operator_sas_plus_def by simp moreover { have "?t' = (\\<^sub>S \ ?t)" using nb\<^sub>1 Cons.prems(2) by simp hence "execute_parallel_operator ?t'[\\<^sub>O \ x. x \ ops] = (\\<^sub>S \ (execute_parallel_operator_sas_plus ?t ops))" using Cons.IH[of ?t] Cons.prems by simp } ultimately show ?case by argo qed qed private lemma strips_equivalent_to_sas_plus_i_a_IV: assumes "is_valid_problem_sas_plus \" and "\op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" and "are_all_operators_applicable_in I ops \ are_all_operator_effects_consistent ops" shows "STRIPS_Semantics.are_all_operators_applicable (\\<^sub>S \ I) [\\<^sub>O \ op. op \ ops] \ STRIPS_Semantics.are_all_operator_effects_consistent [\\<^sub>O \ op. op \ ops]" proof - let ?vs = "variables_of \" and ?ops = "operators_of \" let ?I' = "\\<^sub>S \ I" and ?ops' = "[\\<^sub>O \ op. op \ ops]" have nb\<^sub>1: "\op \ set ops. is_operator_applicable_in I op" using assms(3) unfolding are_all_operators_applicable_in_def list_all_iff by blast have nb\<^sub>2: "\op \ set ops. is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) assms(1, 2) unfolding is_valid_operator_sas_plus_def by auto have nb\<^sub>3: "\op \ set ops. map_of (precondition_of op) \\<^sub>m I" using nb\<^sub>1 unfolding is_operator_applicable_in_def list_all_iff by blast { fix op\<^sub>1 op\<^sub>2 assume "op\<^sub>1 \ set ops" and "op\<^sub>2 \ set ops" hence "are_operator_effects_consistent op\<^sub>1 op\<^sub>2" using assms(3) unfolding are_all_operator_effects_consistent_def list_all_iff by blast } note nb\<^sub>4 = this { fix op\<^sub>1 op\<^sub>2 assume "op\<^sub>1 \ set ops" and "op\<^sub>2 \ set ops" hence "\(v, a) \ set (effect_of op\<^sub>1). \(v', a') \ set (effect_of op\<^sub>2). v \ v' \ a = a'" using nb\<^sub>4 unfolding are_operator_effects_consistent_def Let_def list_all_iff by presburger } note nb\<^sub>5 = this { fix op\<^sub>1' op\<^sub>2' I assume "op\<^sub>1' \ set ?ops'" and "op\<^sub>2' \ set ?ops'" and "\(v, a) \ set (add_effects_of op\<^sub>1'). \(v', a') \ set (delete_effects_of op\<^sub>2'). (v, a) = (v', a')" moreover obtain op\<^sub>1 op\<^sub>2 where "op\<^sub>1 \ set ops" and "op\<^sub>1' = \\<^sub>O \ op\<^sub>1" and "op\<^sub>2 \ set ops" and "op\<^sub>2' = \\<^sub>O \ op\<^sub>2" using strips_equivalent_to_sas_plus_i_a_I[OF assms(1, 2)] calculation(1, 2) by auto moreover have "is_valid_operator_sas_plus \ op\<^sub>1" and is_valid_operator_op\<^sub>2: "is_valid_operator_sas_plus \ op\<^sub>2" using calculation(4, 6) nb\<^sub>2 by blast+ moreover obtain v v' a a' where "(v, a) \ set (add_effects_of op\<^sub>1')" and "(v', a') \ set (delete_effects_of op\<^sub>2')" and "(v, a) = (v', a')" using calculation by blast moreover have "(v, a) \ set (effect_of op\<^sub>1)" using calculation(5, 10) unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def Let_def by fastforce moreover have "v = v'" and "a = a'" using calculation(12) by simp+ moreover { have "(v', a') \ (\(v, a) \ set (effect_of op\<^sub>2). { (v, a') | a'. a' \ (\\<^sub>+ \ v) \ a' \ a })" using sasp_op_to_strips_set_delete_effects_is calculation(7, 9, 11) by blast then obtain v'' a'' where "(v'', a'') \ set (effect_of op\<^sub>2)" and "(v', a') \ { (v'', a''') | a'''. a''' \ (\\<^sub>+ \ v'') \ a''' \ a'' }" by blast moreover have "(v', a'') \ set (effect_of op\<^sub>2)" using calculation by blast moreover have "a' \ \\<^sub>+ \ v''" and "a' \ a''" using calculation(1, 2) by fast+ ultimately have "\a''. (v', a'') \ set (effect_of op\<^sub>2) \ a' \ (\\<^sub>+ \ v') \ a' \ a''" by blast } moreover obtain a'' where "a' \ \\<^sub>+ \ v'" and "(v', a'') \ set (effect_of op\<^sub>2)" and "a' \ a''" using calculation(16) by blast moreover have "\(v, a) \ set (effect_of op\<^sub>1). (\(v', a') \ set (effect_of op\<^sub>2). v = v' \ a \ a')" using calculation(13, 14, 15, 17, 18, 19) by blast \ \ TODO slow. \ ultimately have "\op\<^sub>1 \ set ops. \op\<^sub>2 \ set ops. \are_operator_effects_consistent op\<^sub>1 op\<^sub>2" unfolding are_operator_effects_consistent_def list_all_iff by fastforce } note nb\<^sub>6 = this show ?thesis proof (rule conjI) { fix op' assume "op' \ set ?ops'" moreover obtain op where op_in: "op \ set ops" and op'_is: "op' = \\<^sub>O \ op" and op'_in: "op' \ set ((\ \)\<^sub>\)" and is_valid_op': "is_valid_operator_strips (\ \) op'" using strips_equivalent_to_sas_plus_i_a_I[OF assms(1, 2)] strips_equivalent_to_sas_plus_i_a_II[OF assms(1, 2)] calculation by metis moreover have is_valid_op: "is_valid_operator_sas_plus \ op" using nb\<^sub>2 calculation(2).. { fix v a assume v_a_in_preconditions': "(v, a) \ set (strips_operator.precondition_of op')" have v_a_in_preconditions: "(v, a) \ set (precondition_of op)" using op'_is unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def Let_def using v_a_in_preconditions' by force moreover have "v \ set ?vs" and "a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then(1,2) is_valid_op calculation(1) by fastforce+ moreover { have "\(v, a) \ set (precondition_of op). \(v', a') \ set (precondition_of op). v \ v' \ a = a'" using is_valid_operator_sas_plus_then(5) is_valid_op by fast hence "map_of (precondition_of op) v = Some a" using map_of_constant_assignments_defined_if[OF _ v_a_in_preconditions] by blast } moreover have "v \ dom (map_of (precondition_of op))" using calculation(4) by blast moreover have "I v = Some a" using nb\<^sub>3 unfolding map_le_def using op_in calculation(4, 5) by metis moreover have "(v, a) \ dom ?I'" using state_to_strips_state_dom_element_iff[OF assms(1)] calculation(2, 3, 6) by simp ultimately have "?I' (v, a) = Some True" using state_to_strips_state_range_is[OF assms(1)] by simp } hence "STRIPS_Representation.is_operator_applicable_in ?I' op'" unfolding STRIPS_Representation.is_operator_applicable_in_def Let_def list_all_iff by fast } thus "are_all_operators_applicable ?I' ?ops'" unfolding are_all_operators_applicable_def list_all_iff by blast next { fix op\<^sub>1' op\<^sub>2' assume op\<^sub>1'_in_ops': "op\<^sub>1' \ set ?ops'" and op\<^sub>2'_in_ops': "op\<^sub>2' \ set ?ops'" have "STRIPS_Semantics.are_operator_effects_consistent op\<^sub>1' op\<^sub>2'" unfolding STRIPS_Semantics.are_operator_effects_consistent_def Let_def \ \ TODO proof is symmetrical... refactor into nb. \ proof (rule conjI) show "\list_ex (\x. list_ex ((=) x) (delete_effects_of op\<^sub>2')) (add_effects_of op\<^sub>1')" proof (rule ccontr) assume "\\list_ex (\v. list_ex ((=) v) (delete_effects_of op\<^sub>2')) (add_effects_of op\<^sub>1')" then have "\(v, a) \ set (delete_effects_of op\<^sub>2'). \(v', a') \ set (add_effects_of op\<^sub>1'). (v, a) = (v', a')" unfolding list_ex_iff by fastforce then obtain op\<^sub>1 op\<^sub>2 where "op\<^sub>1 \ set ops" and "op\<^sub>2 \ set ops" and "\are_operator_effects_consistent op\<^sub>1 op\<^sub>2" using nb\<^sub>6[OF op\<^sub>1'_in_ops' op\<^sub>2'_in_ops'] by blast thus False using nb\<^sub>4 by blast qed next show "\list_ex (\v. list_ex ((=) v) (add_effects_of op\<^sub>2')) (delete_effects_of op\<^sub>1')" proof (rule ccontr) assume "\\list_ex (\v. list_ex ((=) v) (add_effects_of op\<^sub>2')) (delete_effects_of op\<^sub>1')" then have "\(v, a) \ set (delete_effects_of op\<^sub>1'). \(v', a') \ set (add_effects_of op\<^sub>2'). (v, a) = (v', a')" unfolding list_ex_iff by fastforce then obtain op\<^sub>1 op\<^sub>2 where "op\<^sub>1 \ set ops" and "op\<^sub>2 \ set ops" and "\are_operator_effects_consistent op\<^sub>1 op\<^sub>2" using nb\<^sub>6[OF op\<^sub>2'_in_ops' op\<^sub>1'_in_ops'] by blast thus False using nb\<^sub>4 by blast qed qed } thus "STRIPS_Semantics.are_all_operator_effects_consistent ?ops'" unfolding STRIPS_Semantics.are_all_operator_effects_consistent_def list_all_iff by blast qed qed private lemma strips_equivalent_to_sas_plus_i_a_V: assumes "is_valid_problem_sas_plus \" and "\op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" and "\(are_all_operators_applicable_in s ops \ are_all_operator_effects_consistent ops)" shows "\(STRIPS_Semantics.are_all_operators_applicable (\\<^sub>S \ s) [\\<^sub>O \ op. op \ ops] \ STRIPS_Semantics.are_all_operator_effects_consistent [\\<^sub>O \ op. op \ ops])" proof - let ?vs = "variables_of \" and ?ops = "operators_of \" let ?s' = "\\<^sub>S \ s" and ?ops' = "[\\<^sub>O \ op. op \ ops]" { fix op assume "op \ set ops" hence "\op' \ set ?ops'. op' = \\<^sub>O \ op" by simp } note nb\<^sub>1 = this { fix op assume "op \ set ops" then have "op \ set ((\)\<^sub>\\<^sub>+)" using assms(2) by blast then have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) assms(1) unfolding is_valid_operator_sas_plus_def by auto hence "\(v, a) \ set (precondition_of op). \(v', a') \ set (precondition_of op). v \ v' \ a = a'" using is_valid_operator_sas_plus_then(5) unfolding is_valid_operator_sas_plus_def by fast } note nb\<^sub>2 = this { consider (A) "\are_all_operators_applicable_in s ops" | (B) "\are_all_operator_effects_consistent ops" using assms(3) by blast hence "\STRIPS_Semantics.are_all_operators_applicable ?s' ?ops' \ \STRIPS_Semantics.are_all_operator_effects_consistent ?ops'" proof (cases) case A then obtain op where op_in: "op \ set ops" and not_precondition_map_le_s: "\(map_of (precondition_of op) \\<^sub>m s)" using A unfolding are_all_operators_applicable_in_def list_all_iff is_operator_applicable_in_def by blast then obtain op' where op'_in: "op' \ set ?ops'" and op'_is: "op' = \\<^sub>O \ op" using nb\<^sub>1 by blast have "\are_all_operators_applicable ?s' ?ops'" proof (rule ccontr) assume "\\are_all_operators_applicable ?s' ?ops'" then have all_operators_applicable: "are_all_operators_applicable ?s' ?ops'" by simp moreover { fix v assume "v \ dom (map_of (precondition_of op))" moreover obtain a where "map_of (precondition_of op) v = Some a" using calculation by blast moreover have "(v, a) \ set (precondition_of op)" using map_of_SomeD[OF calculation(2)]. moreover have "(v, a) \ set (strips_operator.precondition_of op')" using op'_is unfolding sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def using calculation(3) by auto moreover have "?s' (v, a) = Some True" using all_operators_applicable calculation unfolding are_all_operators_applicable_def STRIPS_Representation.is_operator_applicable_in_def is_operator_applicable_in_def Let_def list_all_iff using op'_in by fast moreover have "(v, a) \ dom ?s'" using calculation(5) by blast moreover have "(v, a) \ set (precondition_of op)" using op'_is calculation(3) unfolding sasp_op_to_strips_def Let_def by fastforce moreover have "v \ set ?vs" and "a \ \\<^sub>+ \ v" and "s v \ None" using state_to_strips_state_dom_element_iff[OF assms(1)] calculation(6) by simp+ moreover have "?s' (v, a) = Some (the (s v) = a)" using state_to_strips_state_range_is[OF assms(1) calculation(6)]. moreover have "the (s v) = a" using calculation(5, 11) by fastforce moreover have "s v = Some a" using calculation(12) option.collapse[OF calculation(10)] by argo moreover have "map_of (precondition_of op) v = Some a" using map_of_constant_assignments_defined_if[OF nb\<^sub>2[OF op_in] calculation(7)]. ultimately have "map_of (precondition_of op) v = s v" by argo } then have "map_of (precondition_of op) \\<^sub>m s" unfolding map_le_def by blast thus False using not_precondition_map_le_s by simp qed thus ?thesis by simp next case B { obtain op\<^sub>1 op\<^sub>2 v v' a a' where "op\<^sub>1 \ set ops" and op\<^sub>2_in: "op\<^sub>2 \ set ops" and v_a_in: "(v, a) \ set (effect_of op\<^sub>1)" and v'_a'_in: "(v', a') \ set (effect_of op\<^sub>2)" and v_is: "v = v'" and a_is: "a \ a'" using B unfolding are_all_operator_effects_consistent_def are_operator_effects_consistent_def list_all_iff Let_def by blast moreover obtain op\<^sub>1' op\<^sub>2' where "op\<^sub>1' \ set ?ops'" and "op\<^sub>1' = \\<^sub>O \ op\<^sub>1" and "op\<^sub>1' \ set ?ops'" and op\<^sub>2'_is: "op\<^sub>2' = \\<^sub>O \ op\<^sub>2" using nb\<^sub>1[OF calculation(1)] nb\<^sub>1[OF calculation(2)] by blast moreover have "(v, a) \ set (add_effects_of op\<^sub>1')" using calculation(3, 8) unfolding SAS_Plus_STRIPS.sasp_op_to_strips_def sasp_op_to_strips_def Let_def by force moreover { have "is_valid_operator_sas_plus \ op\<^sub>1" using assms(2) calculation(1) is_valid_problem_sas_plus_then(2) assms(1) unfolding is_valid_operator_sas_plus_def by auto moreover have "is_valid_operator_sas_plus \ op\<^sub>2" using sublocale_sas_plus_finite_domain_representation_ii(2)[ OF assms(1)] assms(2) op\<^sub>2_in by blast moreover have "a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then(4) calculation v_a_in unfolding is_valid_operator_sas_plus_def by fastforce ultimately have "(v, a) \ set (delete_effects_of op\<^sub>2')" using sasp_op_to_strips_set_delete_effects_is[of \ op\<^sub>2] v'_a'_in v_is a_is using op\<^sub>2'_is by blast } \ \ TODO slow. \ ultimately have "\op\<^sub>1' \ set ?ops'. \op\<^sub>2' \ set ?ops'. \(v, a) \ set (delete_effects_of op\<^sub>2'). \(v', a') \ set (add_effects_of op\<^sub>1'). (v, a) = (v', a')" by fastforce } then have "\STRIPS_Semantics.are_all_operator_effects_consistent ?ops'" unfolding STRIPS_Semantics.are_all_operator_effects_consistent_def STRIPS_Semantics.are_operator_effects_consistent_def list_all_iff list_ex_iff Let_def by blast thus ?thesis by simp qed } thus ?thesis by blast qed (* TODO make private *) lemma strips_equivalent_to_sas_plus_i_a: assumes "is_valid_problem_sas_plus \" and "dom I \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom I. the (I v) \ \\<^sub>+ \ v" and "dom G \ set ((\)\<^sub>\\<^sub>+)" and "\v \ dom G. the (G v) \ \\<^sub>+ \ v" and "\ops \ set \. \op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" and "G \\<^sub>m execute_parallel_plan_sas_plus I \" shows "(\\<^sub>S \ G) \\<^sub>m execute_parallel_plan (\\<^sub>S \ I) (\\<^sub>P \ \)" proof - let ?\ = "\ \" and ?G' = "\\<^sub>S \ G" show ?thesis using assms proof (induction \ arbitrary: I) case Nil let ?I' = "\\<^sub>S \ I" have "G \\<^sub>m I" using Nil by simp moreover have "?G' \\<^sub>m ?I'" using state_to_strips_state_map_le_iff[OF Nil.prems(1, 4, 5)] calculation.. ultimately show ?case unfolding SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sas_plus_parallel_plan_to_strips_parallel_plan_def by simp next case (Cons ops \) let ?vs = "variables_of \" and ?ops = "operators_of \" and ?J = "execute_parallel_operator_sas_plus I ops" and ?\ = "\\<^sub>P \ (ops # \)" let ?I' = "\\<^sub>S \ I" and ?J' = "\\<^sub>S \ ?J" and ?ops' = "[\\<^sub>O \ op. op \ ops]" { fix op v a assume "op \ set ops" and "(v, a) \ set (effect_of op)" moreover have "op \ set ?ops" using Cons.prems(6) calculation(1) by simp moreover have "is_valid_operator_sas_plus \ op" using is_valid_problem_sas_plus_then(2) Cons.prems(1) calculation(3) unfolding is_valid_operator_sas_plus_def by auto ultimately have "v \ set ((\)\<^sub>\\<^sub>+)" and "a \ \\<^sub>+ \ v" using is_valid_operator_sas_plus_then(3,4) by fastforce+ } note nb\<^sub>1 = this show ?case proof (cases "are_all_operators_applicable_in I ops \ are_all_operator_effects_consistent ops") case True { have "(\\<^sub>P \ (ops # \)) = ?ops' # (\\<^sub>P \ \)" unfolding sas_plus_parallel_plan_to_strips_parallel_plan_def SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def by simp moreover have "\op \ set ops. op \ set ((\)\<^sub>\\<^sub>+)" using Cons.prems(6) by simp moreover have "STRIPS_Semantics.are_all_operators_applicable ?I' ?ops'" and "STRIPS_Semantics.are_all_operator_effects_consistent ?ops'" using strips_equivalent_to_sas_plus_i_a_IV[OF Cons.prems(1) _ True] calculation by blast+ ultimately have "execute_parallel_plan ?I' ?\ = execute_parallel_plan (execute_parallel_operator ?I' ?ops') (\\<^sub>P \ \)" by fastforce } \ \ NOTE Instantiate the IH on the next state of the SAS+ execution \execute_parallel_operator_sas_plus I ops\. \ moreover { { have "dom I \ set (sas_plus_problem.variables_of \)" using Cons.prems(2) by blast moreover have "\op \ set ops. \(v, a) \ set (effect_of op). v \ set ((\)\<^sub>\\<^sub>+)" using nb\<^sub>1(1) by blast ultimately have "dom ?J \ set ((\)\<^sub>\\<^sub>+)" using sas_plus_equivalent_to_strips_i_a_IX[of I "set ?vs"] by simp } note nb\<^sub>2 = this moreover { have "dom I \ set (sas_plus_problem.variables_of \)" using Cons.prems(2) by blast moreover have "set (sas_plus_problem.variables_of \) \ dom (range_of \)" using is_valid_problem_sas_plus_dom_sas_plus_problem_range_of assms(1) by auto moreover { fix v assume "v \ dom I" moreover have "v \ set ((\)\<^sub>\\<^sub>+)" using Cons.prems(2) calculation by blast ultimately have "the (I v) \ set (the (range_of \ v))" using Cons.prems(3) using set_the_range_of_is_range_of_sas_plus_if[OF assms(1)] by blast } moreover have "\op\set ops. \(v, a)\set (effect_of op). v \ set (sas_plus_problem.variables_of \) \ a \ set (the (range_of \ v))" using set_the_range_of_is_range_of_sas_plus_if[OF assms(1)] nb\<^sub>1(1) nb\<^sub>1(2) by force moreover have nb\<^sub>3: "\v \ dom ?J. the (?J v) \ set (the (range_of \ v))" using sas_plus_equivalent_to_strips_i_a_X[of I "set ?vs" "range_of \" ops] calculation by fast moreover { fix v assume "v \ dom ?J" moreover have "v \ set ((\)\<^sub>\\<^sub>+)" using nb\<^sub>2 calculation by blast moreover have "set (the (range_of \ v)) = \\<^sub>+ \ v" using set_the_range_of_is_range_of_sas_plus_if[OF assms(1)] calculation(2) by presburger ultimately have "the (?J v) \ \\<^sub>+ \ v" using nb\<^sub>3 by blast } ultimately have "\v \ dom ?J. the (?J v) \ \\<^sub>+ \ v" by fast } moreover have "\ops\set \. \op\set ops. op \ set ?ops" using Cons.prems(6) by auto moreover have "G \\<^sub>m execute_parallel_plan_sas_plus ?J \" using Cons.prems(7) True by simp ultimately have "(\\<^sub>S \ G) \\<^sub>m execute_parallel_plan ?J' (\\<^sub>P \ \)" using Cons.IH[of ?J, OF Cons.prems(1) _ _ Cons.prems(4, 5)] by fastforce } moreover have "execute_parallel_operator ?I' ?ops' = ?J'" using assms(1) strips_equivalent_to_sas_plus_i_a_III[OF assms(1)] Cons.prems(6) by auto ultimately show ?thesis by argo next case False then have nb: "G \\<^sub>m I" using Cons.prems(7) by force moreover { have "?\ = ?ops' # (\\<^sub>P \ \)" unfolding sas_plus_parallel_plan_to_strips_parallel_plan_def SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by auto moreover have "set ?ops' \ set (strips_problem.operators_of ?\)" using strips_equivalent_to_sas_plus_i_a_II(1)[OF assms(1)] Cons.prems(6) by auto moreover have "\(STRIPS_Semantics.are_all_operators_applicable ?I' ?ops' \ STRIPS_Semantics.are_all_operator_effects_consistent ?ops')" using strips_equivalent_to_sas_plus_i_a_V[OF assms(1) _ False] Cons.prems(6) by force ultimately have "execute_parallel_plan ?I' ?\ = ?I'" by auto } moreover have "?G' \\<^sub>m ?I'" using state_to_strips_state_map_le_iff[OF Cons.prems(1, 4, 5)] nb by blast ultimately show ?thesis by presburger qed qed qed (* TODO make private *) lemma strips_equivalent_to_sas_plus_i: assumes "is_valid_problem_sas_plus \" and "is_parallel_solution_for_problem \ \" shows "(strips_problem.goal_of (\ \)) \\<^sub>m execute_parallel_plan (strips_problem.initial_of (\ \)) (\\<^sub>P \ \)" proof - let ?vs = "variables_of \" and ?ops = "operators_of \" and ?I = "initial_of \" and ?G = "goal_of \" let ?\ = "\ \" let ?I' = "strips_problem.initial_of ?\" and ?G' = "strips_problem.goal_of ?\" have "dom ?I \ set ?vs" using is_valid_problem_sas_plus_then(3) assms(1) by auto moreover have "\v\dom ?I. the (?I v) \ \\<^sub>+ \ v" using is_valid_problem_sas_plus_then(4) assms(1) calculation by auto moreover have "dom ?G \ set ((\)\<^sub>\\<^sub>+)" using is_valid_problem_sas_plus_then(5) assms(1) by auto moreover have "\v \ dom ?G. the (?G v) \ \\<^sub>+ \ v" using is_valid_problem_sas_plus_then(6) assms(1) by auto moreover have "\ops \ set \. \op \ set ops. op \ set ?ops" using is_parallel_solution_for_problem_plan_operator_set[OF assms(2)] by fastforce moreover have "?G \\<^sub>m execute_parallel_plan_sas_plus ?I \" using assms(2) unfolding is_parallel_solution_for_problem_def by simp (* TODO slow *) ultimately show ?thesis using strips_equivalent_to_sas_plus_i_a[OF assms(1), of ?I ?G \] unfolding sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def state_to_strips_state_def SAS_Plus_STRIPS.state_to_strips_state_def by force qed (* TODO make private *) lemma strips_equivalent_to_sas_plus_ii: assumes "is_valid_problem_sas_plus \" and "is_parallel_solution_for_problem \ \" shows "list_all (list_all (\op. ListMem op (strips_problem.operators_of (\ \)))) (\\<^sub>P \ \)" proof - let ?ops = "operators_of \" let ?\ = "\ \" let ?ops' = "strips_problem.operators_of ?\" and ?\ = "\\<^sub>P \ \" have "is_valid_problem_strips ?\" using is_valid_problem_sas_plus_then_strips_transformation_too[OF assms(1)] by simp have nb\<^sub>1: "\op \ set ?ops. (\op' \ set ?ops'. op' = (\\<^sub>O \ op))" unfolding sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def Let_def sasp_op_to_strips_def by force { fix ops op op' assume "ops \ set \" and "op \ set ops" moreover have "op \ set ((\)\<^sub>\\<^sub>+)" using is_parallel_solution_for_problem_plan_operator_set[OF assms(2)] calculation by blast moreover obtain op' where "op' \ set ?ops'" and "op' = (\\<^sub>O \ op)" using nb\<^sub>1 calculation(3) by auto ultimately have "(\\<^sub>O \ op) \ set ?ops'" by blast } thus ?thesis unfolding list_all_iff ListMem_iff Let_def sas_plus_problem_to_strips_problem_def SAS_Plus_STRIPS.sas_plus_problem_to_strips_problem_def sas_plus_parallel_plan_to_strips_parallel_plan_def SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by auto qed text \ The following lemma proves the complementary proposition to theorem \ref{isathm:equivalence-parallel-strips-parallel-sas-plus}. Namely, given a parallel solution \<^term>\\\ for a SAS+ problem, the transformation to a STRIPS plan \<^term>\\\<^sub>P \ \\ also is a solution to the corresponding STRIPS problem \<^term>\\ \ (\ \)\. In this direction, we have to show that the execution of the transformed plan reaches the goal state \<^term>\G' \ strips_problem.goal_of \\ of the corresponding STRIPS problem, i.e. @{text[display, indent=4] "G' \\<^sub>m execute_parallel_plan I' \"} and that all operators in the transformed plan \<^term>\\\ are operators of \<^term>\\\. \ theorem strips_equivalent_to_sas_plus: assumes "is_valid_problem_sas_plus \" and "is_parallel_solution_for_problem \ \" shows "STRIPS_Semantics.is_parallel_solution_for_problem (\ \) (\\<^sub>P \ \)" proof - let ?\ = "\ \" let ?I' = "strips_problem.initial_of ?\" and ?G' = "strips_problem.goal_of ?\" and ?ops' = "strips_problem.operators_of ?\" and ?\ = "\\<^sub>P \ \" show ?thesis unfolding STRIPS_Semantics.is_parallel_solution_for_problem_def proof (rule conjI) show "?G' \\<^sub>m execute_parallel_plan ?I' ?\" using strips_equivalent_to_sas_plus_i[OF assms] by simp next show "list_all (list_all (\op. ListMem op ?ops')) ?\" using strips_equivalent_to_sas_plus_ii[OF assms]. qed qed lemma embedded_serial_sas_plus_plan_operator_structure: assumes "ops \ set (embed \)" obtains op where "op \ set \" and "[\\<^sub>O \ op. op \ ops] = [\\<^sub>O \ op]" proof - let ?\' = "embed \" { have "?\' = [[op]. op \ \]" by (induction \; force) moreover obtain op where "ops = [op]" and "op \ set \" using assms calculation by fastforce ultimately have "\op \ set \. [\\<^sub>O \ op. op \ ops] = [\\<^sub>O \ op]" by auto } thus ?thesis using that by meson qed private lemma serial_sas_plus_equivalent_to_serial_strips_i: assumes "ops \ set (\\<^sub>P \ (embed \))" obtains op where "op \ set \" and "ops = [\\<^sub>O \ op]" proof - let ?\' = "embed \" { have "set (\\<^sub>P \ (embed \)) = { [\\<^sub>O \ op. op \ ops] | ops. ops \ set ?\' }" unfolding sas_plus_parallel_plan_to_strips_parallel_plan_def SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sasp_op_to_strips_def set_map using setcompr_eq_image by blast moreover obtain ops' where "ops' \ set ?\'" and "ops = [\\<^sub>O \ op. op \ ops']" using assms(1) calculation by blast moreover obtain op where "op \ set \" and "ops = [\\<^sub>O \ op]" using embedded_serial_sas_plus_plan_operator_structure calculation(2, 3) by blast ultimately have "\op \ set \. ops = [\\<^sub>O \ op]" by meson } thus ?thesis using that.. qed private lemma serial_sas_plus_equivalent_to_serial_strips_ii[simp]: "concat (\\<^sub>P \ (embed \)) = [\\<^sub>O \ op. op \ \]" proof - let ?\' = "List_Supplement.embed \" have "concat (\\<^sub>P \ ?\') = map (\op. \\<^sub>O \ op) (concat ?\')" unfolding sas_plus_parallel_plan_to_strips_parallel_plan_def SAS_Plus_STRIPS.sas_plus_parallel_plan_to_strips_parallel_plan_def sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def map_concat by blast also have "\ = map (\op. \\<^sub>O \ op) \" unfolding concat_is_inverse_of_embed[of \].. finally show "concat (\\<^sub>P \ (embed \)) = [\\<^sub>O \ op. op \ \]". qed text \ Having established the equivalence of parallel STRIPS and SAS+, we can now show the equivalence in the serial case. The proof combines the embedding theorem for serial SAS+ solutions (\ref{isathm:serial-sas-plus-embedding}), the parallel plan equivalence theorem \ref{isathm:equivalence-parallel-sas-plus-parallel-strips}, and the flattening theorem for parallel STRIPS plans (\ref{isathm:embedded-serial-plan-flattening-strips}). More precisely, given a serial SAS+ solution \<^term>\\\ for a SAS+ problem \<^term>\\\, the embedding theorem confirms that the embedded plan \<^term>\embed \\ is an equivalent parallel solution to \<^term>\\\. By parallel plan equivalence, \<^term>\\ \ \\<^sub>P \ (embed \)\ is a parallel solution for the corresponding STRIPS problem \<^term>\\ \\. Moreover, since \<^term>\embed \\ is a plan consisting of singleton parallel operators, the same is true for \<^term>\\\. Hence, the flattening lemma applies and \<^term>\concat \\ is a serial solution for \<^term>\\ \\. Since \<^term>\concat\ moreover can be shown to be the inverse of \<^term>\embed\, the term @{text[display, indent=4] "concat \ = concat (\\<^sub>P \ (embed \))"} can be reduced to the intuitive form @{text[display, indent=4] "\ = [\\<^sub>O \ op. op \ \]"} which concludes the proof. \ theorem serial_sas_plus_equivalent_to_serial_strips: assumes "is_valid_problem_sas_plus \" and "SAS_Plus_Semantics.is_serial_solution_for_problem \ \" shows "STRIPS_Semantics.is_serial_solution_for_problem (\ \) [\\<^sub>O \ op. op \ \]" proof - let ?\' = "embed \" and ?\ = "\ \" let ?\' = "\\<^sub>P \ ?\'" let ?\ = "concat ?\'" { have "SAS_Plus_Semantics.is_parallel_solution_for_problem \ ?\'" using execute_serial_plan_sas_plus_is_execute_parallel_plan_sas_plus[OF assms] by simp hence "STRIPS_Semantics.is_parallel_solution_for_problem ?\ ?\'" using strips_equivalent_to_sas_plus[OF assms(1)] by simp } moreover have "?\ = [\\<^sub>O \ op. op \ \]" by simp moreover have "is_valid_problem_strips ?\" using is_valid_problem_sas_plus_then_strips_transformation_too[OF assms(1)]. moreover have "\ops \ set ?\'. \op \ set \. ops = [\\<^sub>O \ op]" using serial_sas_plus_equivalent_to_serial_strips_i[of _ \ \] by metis ultimately show ?thesis using STRIPS_Semantics.flattening_lemma[of ?\] by metis qed lemma embedded_serial_strips_plan_operator_structure: assumes "ops' \ set (embed \)" obtains op where "op \ set \" and "[\\<^sub>O\ \ op. op \ ops'] = [\\<^sub>O\ \ op]" proof - let ?\' = "embed \" { have "?\' = [[op]. op \ \]" by (induction \; force) moreover obtain op where "ops' = [op]" and "op \ set \" using calculation assms by fastforce ultimately have "\op \ set \. [\\<^sub>O\ \ op. op \ ops'] = [\\<^sub>O\ \ op]" by auto } thus ?thesis using that by meson qed private lemma serial_strips_equivalent_to_serial_sas_plus_i: assumes "ops \ set (\\<^sub>P\ \ (embed \))" obtains op where "op \ set \" and "ops = [\\<^sub>O\ \ op]" proof - let ?\' = "embed \" { have "set (\\<^sub>P\ \ (embed \)) = { [\\<^sub>O\ \ op. op \ ops] | ops. ops \ set ?\' }" unfolding strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_op_to_sasp_def set_map using setcompr_eq_image by blast moreover obtain ops' where "ops' \ set ?\'" and "ops = [\\<^sub>O\ \ op. op \ ops']" using assms(1) calculation by blast moreover obtain op where "op \ set \" and "ops = [\\<^sub>O\ \ op]" using embedded_serial_strips_plan_operator_structure calculation(2, 3) by blast ultimately have "\op \ set \. ops = [\\<^sub>O\ \ op]" by meson } thus ?thesis using that.. qed private lemma serial_strips_equivalent_to_serial_sas_plus_ii[simp]: "concat (\\<^sub>P\ \ (embed \)) = [\\<^sub>O\ \ op. op \ \]" proof - let ?\' = "List_Supplement.embed \" have "concat (\\<^sub>P\ \ ?\') = map (\op. \\<^sub>O\ \ op) (concat ?\')" unfolding strips_parallel_plan_to_sas_plus_parallel_plan_def SAS_Plus_STRIPS.strips_parallel_plan_to_sas_plus_parallel_plan_def strips_op_to_sasp_def SAS_Plus_STRIPS.strips_op_to_sasp_def Let_def map_concat by simp also have "\ = map (\op. \\<^sub>O\ \ op) \" unfolding concat_is_inverse_of_embed[of \].. finally show "concat (\\<^sub>P\ \ (embed \)) = [\\<^sub>O\ \ op. op \ \]". qed text \ Using the analogous lemmas for the opposite direction, we can show the counterpart to theorem \ref{isathm:equivalence-serial-sas-plus-serial-strips} which shows that serial solutions to STRIPS solutions can be transformed to serial SAS+ solutions via composition of embedding, transformation and flattening. \ theorem serial_strips_equivalent_to_serial_sas_plus: assumes "is_valid_problem_sas_plus \" and "STRIPS_Semantics.is_serial_solution_for_problem (\ \) \" shows "SAS_Plus_Semantics.is_serial_solution_for_problem \ [\\<^sub>O\ \ op. op \ \]" proof - let ?\' = "embed \" and ?\ = "\ \" let ?\' = "\\<^sub>P\ \ ?\'" let ?\ = "concat ?\'" { have "STRIPS_Semantics.is_parallel_solution_for_problem ?\ ?\'" using embedding_lemma[OF is_valid_problem_sas_plus_then_strips_transformation_too[OF assms(1)] assms(2)]. hence "SAS_Plus_Semantics.is_parallel_solution_for_problem \ ?\'" using sas_plus_equivalent_to_strips[OF assms(1)] by simp } moreover have "?\ = [\\<^sub>O\ \ op. op \ \]" by simp moreover have "is_valid_problem_strips ?\" using is_valid_problem_sas_plus_then_strips_transformation_too[OF assms(1)]. moreover have "\ops \ set ?\'. \op \ set \. ops = [\\<^sub>O\ \ op]" using serial_strips_equivalent_to_serial_sas_plus_i by metis ultimately show ?thesis using flattening_lemma[OF assms(1)] by metis qed subsection "Equivalence of SAS+ and STRIPS" \ \ Define the sets of plans with upper length bound as well as the sets of solutions with upper length bound for SAS problems and induced STRIPS problems. We keep this polymorphic by not specifying concrete types so it applies to both STRIPS and SAS+ plans. \ abbreviation bounded_plan_set where "bounded_plan_set ops k \ { \. set \ \ set ops \ length \ = k }" definition bounded_solution_set_sas_plus' :: "('variable, 'domain) sas_plus_problem \ nat \ ('variable, 'domain) sas_plus_plan set" where "bounded_solution_set_sas_plus' \ k \ { \. is_serial_solution_for_problem \ \ \ length \ = k}" abbreviation bounded_solution_set_sas_plus :: "('variable, 'domain) sas_plus_problem \ nat \ ('variable, 'domain) sas_plus_plan set" where "bounded_solution_set_sas_plus \ N \ (\k \ {0..N}. bounded_solution_set_sas_plus' \ k)" definition bounded_solution_set_strips' :: "('variable \ 'domain) strips_problem \ nat \ ('variable \ 'domain) strips_plan set" where "bounded_solution_set_strips' \ k \ { \. STRIPS_Semantics.is_serial_solution_for_problem \ \ \ length \ = k }" abbreviation bounded_solution_set_strips :: "('variable \ 'domain) strips_problem \ nat \ ('variable \ 'domain) strips_plan set" where "bounded_solution_set_strips \ N \ (\k \ {0..N}. bounded_solution_set_strips' \ k)" \ \ Show that plan transformation for all SAS Plus solutions yields a STRIPS solution for the induced STRIPS problem with same length. We first show injectiveness of plan transformation \\\. [\\<^sub>O \ op. op \ \]\ on the set of plans \P\<^sub>k \ bounded_plan_set (operators_of \) k\ with length bound \k\. The injectiveness of \Sol\<^sub>k \ bounded_solution_set_sas_plus \ k\---the set of solutions with length bound \k\--then follows from the subset relation \Sol\<^sub>k \ P\<^sub>k\. \ lemma sasp_op_to_strips_injective: assumes "(\\<^sub>O \ op\<^sub>1) = (\\<^sub>O \ op\<^sub>2)" shows "op\<^sub>1 = op\<^sub>2" proof - let ?op\<^sub>1' = "\\<^sub>O \ op\<^sub>1" and ?op\<^sub>2' = "\\<^sub>O \ op\<^sub>2" { have "strips_operator.precondition_of ?op\<^sub>1' = strips_operator.precondition_of ?op\<^sub>2'" using assms by argo hence "sas_plus_operator.precondition_of op\<^sub>1 = sas_plus_operator.precondition_of op\<^sub>2" unfolding sasp_op_to_strips_def SAS_Plus_STRIPS.sasp_op_to_strips_def Let_def by simp } moreover { have "strips_operator.add_effects_of ?op\<^sub>1' = strips_operator.add_effects_of ?op\<^sub>2'" using assms unfolding sasp_op_to_strips_def Let_def by argo hence "sas_plus_operator.effect_of op\<^sub>1 = sas_plus_operator.effect_of op\<^sub>2" unfolding sasp_op_to_strips_def Let_def SAS_Plus_STRIPS.sasp_op_to_strips_def by simp } ultimately show ?thesis by simp qed lemma sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_a: assumes "is_valid_problem_sas_plus \" shows "inj_on (\\. [\\<^sub>O \ op. op \ \]) (bounded_plan_set (sas_plus_problem.operators_of \) k)" proof - let ?ops = "sas_plus_problem.operators_of \" (* TODO refactor transformation definitions *) and ?\\<^sub>P = "\\. [\\<^sub>O \ op. op \ \]" let ?P = "bounded_plan_set ?ops" { fix \\<^sub>1 \\<^sub>2 assume \\<^sub>1_in: "\\<^sub>1 \ ?P k" and \\<^sub>2_in: "\\<^sub>2 \ ?P k" and \\<^sub>P_of_\\<^sub>1_is_\\<^sub>P_of_\\<^sub>2: "(?\\<^sub>P \\<^sub>1) = (?\\<^sub>P \\<^sub>2)" hence "\\<^sub>1 = \\<^sub>2" proof (induction k arbitrary: \\<^sub>1 \\<^sub>2) case 0 then have "length \\<^sub>1 = 0" and "length \\<^sub>2 = 0" using \\<^sub>1_in \\<^sub>2_in unfolding bounded_solution_set_sas_plus'_def by blast+ then show ?case by blast next case (Suc k) moreover have "length \\<^sub>1 = Suc k" and "length \\<^sub>2 = Suc k" using length_Suc_conv Suc(2, 3) unfolding bounded_solution_set_sas_plus'_def by blast+ moreover obtain op\<^sub>1 \\<^sub>1' where "\\<^sub>1 = op\<^sub>1 # \\<^sub>1'" and "set (op\<^sub>1 # \\<^sub>1') \ set ?ops" and "length \\<^sub>1' = k" using calculation(5) Suc(2) unfolding length_Suc_conv by blast moreover obtain op\<^sub>2 \\<^sub>2' where "\\<^sub>2 = op\<^sub>2 # \\<^sub>2'" and "set (op\<^sub>2 # \\<^sub>2') \ set ?ops" and "length \\<^sub>2' = k" using calculation(6) Suc(3) unfolding length_Suc_conv by blast moreover have "set \\<^sub>1' \ set ?ops" and "set \\<^sub>2' \ set ?ops" using calculation(8, 11) by auto+ moreover have "\\<^sub>1' \ ?P k" and "\\<^sub>2' \ ?P k" using calculation(9, 12, 13, 14) by fast+ moreover have "?\\<^sub>P \\<^sub>1' = ?\\<^sub>P \\<^sub>2'" using Suc.prems(3) calculation(7, 10) by fastforce moreover have "\\<^sub>1' = \\<^sub>2'" using Suc.IH[of \\<^sub>1' \\<^sub>2', OF calculation(15, 16, 17)] by simp moreover have "?\\<^sub>P \\<^sub>1 = (\\<^sub>O \ op\<^sub>1) # ?\\<^sub>P \\<^sub>1'" and "?\\<^sub>P \\<^sub>2 = (\\<^sub>O \ op\<^sub>2) # ?\\<^sub>P \\<^sub>2'" using Suc.prems(3) calculation(7, 10) by fastforce+ moreover have "(\\<^sub>O \ op\<^sub>1) = (\\<^sub>O \ op\<^sub>2)" using Suc.prems(3) calculation(17, 19, 20) by simp moreover have "op\<^sub>1 = op\<^sub>2" using sasp_op_to_strips_injective[OF calculation(21)]. ultimately show ?case by argo qed } thus ?thesis unfolding inj_on_def by blast qed private corollary sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_b: assumes "is_valid_problem_sas_plus \" shows "inj_on (\\. [\\<^sub>O \ op. op \ \]) (bounded_solution_set_sas_plus' \ k)" proof - let ?ops = "sas_plus_problem.operators_of \" and ?\\<^sub>P = "\\. [\\<^sub>O \ op. op \ \]" { fix \ assume "\ \ bounded_solution_set_sas_plus' \ k" then have "set \ \ set ?ops" and "length \ = k" unfolding bounded_solution_set_sas_plus'_def is_serial_solution_for_problem_def Let_def list_all_iff ListMem_iff by fast+ hence "\ \ bounded_plan_set ?ops k" by blast } hence "bounded_solution_set_sas_plus' \ k \ bounded_plan_set ?ops k" by blast moreover have "inj_on ?\\<^sub>P (bounded_plan_set ?ops k)" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_a[OF assms(1)]. ultimately show ?thesis using inj_on_subset[of ?\\<^sub>P "bounded_plan_set ?ops k" "bounded_solution_set_sas_plus' \ k"] by fast qed (* lemma "card ((\\. [\\<^sub>O \ op. op \ \]) ` (bounded_solution_set_sas_plus' \ k)) = card (bounded_solution_set_strips' (\ \) k)" sorry *) \ \ Show that mapping plan transformation \\\. [\\<^sub>O \ op. op \ \]\ over the solution set for a given SAS+ problem yields the solution set for the induced STRIPS problem. \ private lemma sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_c: assumes "is_valid_problem_sas_plus \" shows "(\\. [\\<^sub>O \ op. op \ \]) ` (bounded_solution_set_sas_plus' \ k) = bounded_solution_set_strips' (\ \) k" proof - let ?\ = "\ \" and ?\\<^sub>P = "\\. [\\<^sub>O \ op. op \ \]" let ?Sol\<^sub>k = "bounded_solution_set_sas_plus' \ k" and ?Sol\<^sub>k' = "bounded_solution_set_strips' ?\ k" { assume "?\\<^sub>P ` ?Sol\<^sub>k \ ?Sol\<^sub>k'" then consider (A) "\\ \ ?\\<^sub>P ` ?Sol\<^sub>k. \ \ ?Sol\<^sub>k'" | (B) "\\ \ ?Sol\<^sub>k'. \ \ ?\\<^sub>P ` ?Sol\<^sub>k" by blast hence False proof (cases) case A moreover obtain \ where "\ \ ?\\<^sub>P ` ?Sol\<^sub>k" and "\ \ ?Sol\<^sub>k'" using calculation by blast moreover obtain \ where "length \ = k" and "SAS_Plus_Semantics.is_serial_solution_for_problem \ \" and "\ = ?\\<^sub>P \" using calculation(2) unfolding bounded_solution_set_sas_plus'_def by blast moreover have "length \ = k" and "STRIPS_Semantics.is_serial_solution_for_problem ?\ \" subgoal using calculation(4, 6) by auto subgoal using serial_sas_plus_equivalent_to_serial_strips assms(1) calculation(5) calculation(6) by blast done moreover have "\ \ ?Sol\<^sub>k'" unfolding bounded_solution_set_strips'_def using calculation(7, 8) by simp ultimately show ?thesis by fast next case B moreover obtain \ where "\ \ ?Sol\<^sub>k'" and "\ \ ?\\<^sub>P ` ?Sol\<^sub>k" using calculation by blast moreover have "STRIPS_Semantics.is_serial_solution_for_problem ?\ \" and "length \ = k" using calculation(2) unfolding bounded_solution_set_strips'_def by simp+ \ \ Construct the counter example \\ \ [\\<^sub>O\ ?\ op. op \ \]\ and show that \\ \ ?Sol\<^sub>k\ as well as \?\\<^sub>P \ = \\ hence \\ \ ?\\<^sub>P ` ?Sol\<^sub>k\. \ moreover have "length [\\<^sub>O\ \ op. op \ \] = k" and "SAS_Plus_Semantics.is_serial_solution_for_problem \ [\\<^sub>O\ \ op. op \ \]" subgoal using calculation(5) by simp subgoal using serial_strips_equivalent_to_serial_sas_plus[OF assms(1)] calculation(4) by simp done moreover have "[\\<^sub>O\ \ op. op \ \] \ ?Sol\<^sub>k" unfolding bounded_solution_set_sas_plus'_def using calculation(6, 7) by blast (* TODO refactor transformation lemmas *) moreover { have "\op \ set \. op \ set ((?\)\<^sub>\)" using calculation(4) unfolding STRIPS_Semantics.is_serial_solution_for_problem_def list_all_iff ListMem_iff by simp hence "?\\<^sub>P [\\<^sub>O\ \ op. op \ \] = \" proof (induction \) case (Cons op \) moreover have "?\\<^sub>P [\\<^sub>O\ \ op. op \ op # \] = (\\<^sub>O \ (\\<^sub>O\ \ op)) # ?\\<^sub>P [\\<^sub>O\ \ op. op \ \]" by simp moreover have "op \ set ((?\)\<^sub>\)" using Cons.prems by simp moreover have "(\\<^sub>O \ (\\<^sub>O\ \ op)) = op" using strips_operator_inverse_is[OF assms(1) calculation(4)]. moreover have "?\\<^sub>P [\\<^sub>O\ \ op. op \ \] = \" using Cons.IH Cons.prems by auto ultimately show ?case by argo qed simp } moreover have "\ \ ?\\<^sub>P ` ?Sol\<^sub>k" using calculation(8, 9) by force ultimately show ?thesis by blast qed } thus ?thesis by blast qed private lemma sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_d: assumes "is_valid_problem_sas_plus \" shows "card (bounded_solution_set_sas_plus' \ k) \ card (bounded_solution_set_strips' (\ \) k)" proof - let ?\ = "\ \" and ?\\<^sub>P = "\\. [\\<^sub>O \ op. op \ \]" let ?Sol\<^sub>k = "bounded_solution_set_sas_plus' \ k" and ?Sol\<^sub>k' = "bounded_solution_set_strips' ?\ k" have "card (?\\<^sub>P ` ?Sol\<^sub>k) = card (?Sol\<^sub>k)" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_b[OF assms(1)] card_image by blast moreover have "?\\<^sub>P ` ?Sol\<^sub>k = ?Sol\<^sub>k'" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_c[OF assms(1)]. ultimately show ?thesis by simp qed \ \ The set of fixed length plans with operators in a given operator set is finite. \ lemma bounded_plan_set_finite: shows "finite { \. set \ \ set ops \ length \ = k }" proof (induction k) case (Suc k) let ?P = "{ \. set \ \ set ops \ length \ = k }" and ?P' = "{ \. set \ \ set ops \ length \ = Suc k }" let ?P'' = "(\op \ set ops. (\\ \ ?P. { op # \ }))" { have "\op \. finite { op # \ }" by simp then have "\op. finite (\\ \ ?P. { op # \ })" using finite_UN[of ?P] Suc by blast hence "finite ?P''" using finite_UN[of "set ops"] by blast } moreover { { fix \ assume "\ \ ?P'" moreover have "set \ \ set ops" and "length \ = Suc k" using calculation by simp+ moreover obtain op \' where "\ = op # \'" using calculation (3) unfolding length_Suc_conv by fast moreover have "set \' \ set ops" and "op \ set ops" using calculation(2, 4) by simp+ moreover have "length \' = k" using calculation(3, 4) by auto moreover have "\' \ ?P" using calculation(5, 7) by blast ultimately have "\ \ ?P''" by blast } hence "?P' \ ?P''" by blast } ultimately show ?case using rev_finite_subset[of ?P'' ?P'] by blast qed force \ \ The set of fixed length SAS+ solutions are subsets of the set of plans with fixed length and therefore also finite. \ private lemma sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_ii_a: assumes "is_valid_problem_sas_plus \" shows "finite (bounded_solution_set_sas_plus' \ k)" proof - let ?Ops = "set ((\)\<^sub>\\<^sub>+)" let ?Sol\<^sub>k = "bounded_solution_set_sas_plus' \ k" and ?P\<^sub>k = "{ \. set \ \ ?Ops \ length \ = k }" { fix \ assume "\ \ ?Sol\<^sub>k" then have "length \ = k" and "set \ \ ?Ops" unfolding bounded_solution_set_sas_plus'_def SAS_Plus_Semantics.is_serial_solution_for_problem_def Let_def list_all_iff ListMem_iff by fastforce+ hence "\ \ ?P\<^sub>k" by blast } then have "?Sol\<^sub>k \ ?P\<^sub>k" by force thus ?thesis - using bounded_plan_set_finite rev_finite_subset[of ?P\<^sub>k ?Sol\<^sub>k] - by auto + using bounded_plan_set_finite rev_finite_subset[of ?P\<^sub>k ?Sol\<^sub>k] by blast qed \ \ The set of fixed length STRIPS solutions are subsets of the set of plans with fixed length and therefore also finite. \ private lemma sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_ii_b: assumes "is_valid_problem_sas_plus \" shows "finite (bounded_solution_set_strips' (\ \) k)" proof - let ?\ = "\ \" let ?Ops = "set ((?\)\<^sub>\)" let ?Sol\<^sub>k = "bounded_solution_set_strips' ?\ k" and ?P\<^sub>k = "{ \. set \ \ ?Ops \ length \ = k }" { fix \ assume "\ \ ?Sol\<^sub>k" then have "length \ = k" and "set \ \ ?Ops" unfolding bounded_solution_set_strips'_def STRIPS_Semantics.is_serial_solution_for_problem_def Let_def list_all_iff ListMem_iff by fastforce+ hence "\ \ ?P\<^sub>k" by blast } then have "?Sol\<^sub>k \ ?P\<^sub>k" by force thus ?thesis using bounded_plan_set_finite rev_finite_subset[of ?P\<^sub>k ?Sol\<^sub>k] unfolding state_to_strips_state_def SAS_Plus_STRIPS.state_to_strips_state_def operators_of_def by blast qed text \ With the results on the equivalence of SAS+ and STRIPS solutions, we can now show that given problems in both formalisms, the solution sets have the same size. This is the property required by the definition of planning formalism equivalence presented earlier in theorem \ref{thm:solution-sets-sas-plus-strips-f} (\autoref{sub:equivalence-sas-plus-strips}) and thus end up with the desired equivalence result. The proof uses the finiteness and disjunctiveness of the solution sets for either problem to be able to equivalently transform the set cardinality over the union of sets of solutions with bounded lengths into a sum over the cardinality of the sets of solutions with bounded length. Moreover, since we know that for each SAS+ solution with a given length an equivalent STRIPS solution exists in the solution set of the transformed problem with the same length, both sets must have the same cardinality. Hence the cardinality of the SAS+ solution set over all lengths up to a given upper bound \<^term>\N\ has the same size as the solution set of the corresponding STRIPS problem over all length up to a given upper bound \<^term>\N\. \ theorem assumes "is_valid_problem_sas_plus \" shows "card (bounded_solution_set_sas_plus \ N) = card (bounded_solution_set_strips (\ \) N)" proof - let ?\ = "\ \" and ?R = "{0..N}" \ \ Due to the disjoint nature of the bounded solution sets for fixed plan length for different lengths, we can sum the individual set cardinality to obtain the cardinality of the overall SAS+ resp. STRIPS solution sets. \ have finite_R: "finite ?R" by simp moreover { have "\k \ ?R. finite (bounded_solution_set_sas_plus' \ k)" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_ii_a[OF assms(1)].. moreover have "\j \ ?R. \k \ ?R. j \ k \ bounded_solution_set_sas_plus' \ j \ bounded_solution_set_sas_plus' \ k = {}" unfolding bounded_solution_set_sas_plus'_def by blast (* TODO slow. *) ultimately have "card (bounded_solution_set_sas_plus \ N) = (\k \ ?R. card (bounded_solution_set_sas_plus' \ k))" using card_UN_disjoint by blast } moreover { have "\k \ ?R. finite (bounded_solution_set_strips' ?\ k)" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_ii_b[OF assms(1)].. moreover have "\j \ ?R. \k \ ?R. j \ k \ bounded_solution_set_strips' ?\ j \ bounded_solution_set_strips' ?\ k = {}" unfolding bounded_solution_set_strips'_def by blast (* TODO slow. *) ultimately have "card (bounded_solution_set_strips ?\ N) = (\k \ ?R. card (bounded_solution_set_strips' ?\ k))" using card_UN_disjoint by blast } moreover { fix k have "card (bounded_solution_set_sas_plus' \ k) = card ((\\. [\\<^sub>O \ op. op \ \]) ` bounded_solution_set_sas_plus' \ k)" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_b[OF assms] card_image[symmetric] by blast hence "card (bounded_solution_set_sas_plus' \ k) = card (bounded_solution_set_strips' ?\ k)" using sas_plus_formalism_and_induced_strips_formalism_are_equally_expressive_i_c[OF assms] by presburger } ultimately show ?thesis by presburger qed end end \ No newline at end of file