diff --git a/thys/ADS_Functor/Generic_ADS_Construction.thy b/thys/ADS_Functor/Generic_ADS_Construction.thy --- a/thys/ADS_Functor/Generic_ADS_Construction.thy +++ b/thys/ADS_Functor/Generic_ADS_Construction.thy @@ -1,466 +1,469 @@ (* Author: Andreas Lochbihler, Digital Asset Author: Ognjen Maric, Digital Asset *) theory Generic_ADS_Construction imports "Merkle_Interface" "HOL-Library.BNF_Axiomatization" begin section \Generic construction of authenticated data structures\ subsection \Functors\ subsubsection \Source functor\ text \ We want to allow ADSs of arbitrary ADTs, which we call "source trees". The ADTs we are interested in can in general be represented as the least fixpoints of some bounded natural (bi-)functor (BNF) \('a, 'b) F\, where @{typ 'a} is the type of "source" data, and @{typ 'b} is a recursion "handle". However, Isabelle's type system does not support higher kinds, necessary to parameterize our definitions over functors. Instead, we first develop a general theory of ADSs over an arbitrary, but fixed functor, and its least fixpoint. We show that the theory is compositional, in that the functor's least fixed point can then be reused as the "source" data of another functor. We start by defining the arbitrary fixed functor, its fixpoints, and showing how composition can be done. A higher-level explanation is found in the paper. \ bnf_axiomatization ('a, 'b) F [wits: "'a \ ('a, 'b) F"] context notes [[typedef_overloaded]] begin datatype 'a T = T "('a, 'a T) F" end subsubsection \Base Merkle functor\ text \ This type captures the ADS hashes. \ bnf_axiomatization ('a, 'b) F\<^sub>h [wits: "'a \ ('a, 'b) F\<^sub>h"] text \ It intuitively contains mixed garbage and source values. The functor's recursive handle @{typ 'b} might contain partial garbage. \ text \ This type captures the ADS inclusion proofs. The functor \('a, 'a', 'b, 'b') F\<^sub>m\ has all type variables doubled. This type represents all values including the information which parts are blinded. The original type variable @{typ 'a} now represents the source data, which for compositionality can contain blindable positions. The type @{typ 'b} is a recursive handle to inclusion sub-proofs (which can be partialy blinded). The type @{typ 'a'} represent "hashes" of the source data in @{typ 'a}, i.e., a mix of source values and garbage. The type @{typ 'b'} is a recursive handle to ADS hashes of subtrees. The corresponding type of recursive authenticated trees is then a fixpoint of this functor. \ bnf_axiomatization ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) F\<^sub>m [wits: "'a\<^sub>m \ 'a\<^sub>h \ 'b\<^sub>h \ ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) F\<^sub>m"] subsubsection \Least fixpoint\ context notes [[typedef_overloaded]] begin datatype 'a\<^sub>h T\<^sub>h = T\<^sub>h "('a\<^sub>h, 'a\<^sub>h T\<^sub>h) F\<^sub>h" end context notes [[typedef_overloaded]] begin datatype ('a\<^sub>m, 'a\<^sub>h) T\<^sub>m = T\<^sub>m (the_T\<^sub>m: "('a\<^sub>m, 'a\<^sub>h, ('a\<^sub>m, 'a\<^sub>h) T\<^sub>m, 'a\<^sub>h T\<^sub>h) F\<^sub>m") end subsubsection \ Composition \ text \ Finally, we show how to compose two Merkle functors. For simplicity, we reuse @{typ \('a, 'b) F\} and @{typ \'a T\}. \ context notes [[typedef_overloaded]] begin datatype ('a, 'b) G = G "('a T, 'b) F" datatype ('a\<^sub>h, 'b\<^sub>h) G\<^sub>h = G\<^sub>h (the_G\<^sub>h: "('a\<^sub>h T\<^sub>h, 'b\<^sub>h) F\<^sub>h") datatype ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) G\<^sub>m = G\<^sub>m (the_G\<^sub>m: "(('a\<^sub>m, 'a\<^sub>h) T\<^sub>m, 'a\<^sub>h T\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) F\<^sub>m") end subsection \Root hash\ subsubsection \Base functor\ text \ The root hash of an authenticated value is modelled as a blindable value of type @{typ "('a', 'b') F\<^sub>h"}. (Actually, we want to use an abstract datatype for root hashes, but we omit this distinction here for simplicity.) \ consts root_hash_F' :: "(('a\<^sub>h, 'a\<^sub>h, 'b\<^sub>h, 'b\<^sub>h) F\<^sub>m, ('a\<^sub>h, 'b\<^sub>h) F\<^sub>h) hash" \ \Root hash operation where we assume that all atoms have already been replaced by root hashes. This assumption is reflected in the equality of the type parameters of @{type F\<^sub>m} \ type_synonym ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) hash_F = "('a\<^sub>m, 'a\<^sub>h) hash \ ('b\<^sub>m, 'b\<^sub>h) hash \ (('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) F\<^sub>m, ('a\<^sub>h, 'b\<^sub>h) F\<^sub>h) hash" definition root_hash_F :: "('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) hash_F" where "root_hash_F rha rhb = root_hash_F' \ map_F\<^sub>m rha id rhb id" subsubsection \Least fixpoint\ primrec root_hash_T' :: "(('a\<^sub>h, 'a\<^sub>h) T\<^sub>m, 'a\<^sub>h T\<^sub>h) hash" where "root_hash_T' (T\<^sub>m x) = T\<^sub>h (root_hash_F' (map_F\<^sub>m id id root_hash_T' id x))" definition root_hash_T :: "('a\<^sub>m, 'a\<^sub>h) hash \ (('a\<^sub>m, 'a\<^sub>h) T\<^sub>m, 'a\<^sub>h T\<^sub>h) hash" where "root_hash_T rha = root_hash_T' \ map_T\<^sub>m rha id" lemma root_hash_T_simps [simp]: "root_hash_T rha (T\<^sub>m x) = T\<^sub>h (root_hash_F rha (root_hash_T rha) x)" by(simp add: root_hash_T_def F\<^sub>m.map_comp root_hash_F_def T\<^sub>h.map_id0) subsubsection \Composition\ primrec root_hash_G' :: "(('a\<^sub>h, 'a\<^sub>h, 'b\<^sub>h, 'b\<^sub>h) G\<^sub>m, ('a\<^sub>h, 'b\<^sub>h) G\<^sub>h) hash" where "root_hash_G' (G\<^sub>m x) = G\<^sub>h (root_hash_F' (map_F\<^sub>m root_hash_T' id id id x))" definition root_hash_G :: "('a\<^sub>m, 'a\<^sub>h) hash \ ('b\<^sub>m, 'b\<^sub>h) hash \ (('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) G\<^sub>m, ('a\<^sub>h, 'b\<^sub>h) G\<^sub>h) hash" where "root_hash_G rha rhb = root_hash_G' \ map_G\<^sub>m rha id rhb id" lemma root_hash_G_unfold: "root_hash_G rha rhb = G\<^sub>h \ root_hash_F (root_hash_T rha) rhb \ the_G\<^sub>m" apply(rule ext) subgoal for x by(cases x)(simp add: root_hash_G_def fun_eq_iff root_hash_F_def root_hash_T_def F\<^sub>m.map_comp T\<^sub>m.map_comp o_def T\<^sub>h.map_id id_def[symmetric]) done lemma root_hash_G_simps [simp]: "root_hash_G rha rhb (G\<^sub>m x) = G\<^sub>h (root_hash_F (root_hash_T rha) rhb x)" by(simp add: root_hash_G_def root_hash_T_def F\<^sub>m.map_comp root_hash_F_def T\<^sub>h.map_id0) subsection \Blinding relation\ text \ The blinding relation determines whether one ADS value is a blinding of another. \ subsubsection \ Blinding on the base functor (@{type F\<^sub>m}) \ type_synonym ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) blinding_of_F = "('a\<^sub>m, 'a\<^sub>h) hash \ 'a\<^sub>m blinding_of \ ('b\<^sub>m, 'b\<^sub>h) hash \ 'b\<^sub>m blinding_of \ ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) F\<^sub>m blinding_of" \ \ Computes whether a partially blinded ADS is a blinding of another one \ axiomatization blinding_of_F :: "('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) blinding_of_F" where blinding_of_F_mono: "\ boa \ boa'; bob \ bob' \ \ blinding_of_F rha boa rhb bob \ blinding_of_F rha boa' rhb bob'" \ \ Monotonicity must be unconditional (without the assumption @{text "blinding_of_on"}) such that we can justify the recursive definition for the least fixpoint. \ and blinding_respects_hashes_F [locale_witness]: "\ blinding_respects_hashes rha boa; blinding_respects_hashes rhb bob \ \ blinding_respects_hashes (root_hash_F rha rhb) (blinding_of_F rha boa rhb bob)" and blinding_of_on_F [locale_witness]: "\ blinding_of_on A rha boa; blinding_of_on B rhb bob \ \ blinding_of_on {x. set1_F\<^sub>m x \ A \ set3_F\<^sub>m x \ B} (root_hash_F rha rhb) (blinding_of_F rha boa rhb bob)" lemma blinding_of_F_mono_inductive: assumes a: "\x y. boa x y \ boa' x y" and b: "\x y. bob x y \ bob' x y" shows "blinding_of_F rha boa rhb bob x y \ blinding_of_F rha boa' rhb bob' x y" using assms by(blast intro: blinding_of_F_mono[THEN predicate2D, OF predicate2I predicate2I]) subsubsection \ Blinding on least fixpoints \ context fixes rh :: "('a\<^sub>m, 'a\<^sub>h) hash" and bo :: "'a\<^sub>m blinding_of" begin inductive blinding_of_T :: "('a\<^sub>m, 'a\<^sub>h) T\<^sub>m blinding_of" where "blinding_of_T (T\<^sub>m x) (T\<^sub>m y)" if "blinding_of_F rh bo (root_hash_T rh) blinding_of_T x y" monos blinding_of_F_mono_inductive end lemma blinding_of_T_mono: assumes "bo \ bo'" shows "blinding_of_T rh bo \ blinding_of_T rh bo'" by(rule predicate2I; erule blinding_of_T.induct) (blast intro: blinding_of_T.intros blinding_of_F_mono[THEN predicate2D, OF assms, rotated -1]) lemma blinding_of_T_root_hash: assumes "bo \ vimage2p rh rh (=)" shows "blinding_of_T rh bo \ vimage2p (root_hash_T rh) (root_hash_T rh) (=)" apply(rule predicate2I vimage2pI)+ apply(erule blinding_of_T.induct) apply simp apply(drule blinding_respects_hashes_F[unfolded blinding_respects_hashes_def, THEN predicate2D, rotated -1]) apply(rule assms) apply(blast intro: vimage2pI) apply(simp add: vimage2p_def) done lemma blinding_respects_hashes_T [locale_witness]: "blinding_respects_hashes rh bo \ blinding_respects_hashes (root_hash_T rh) (blinding_of_T rh bo)" unfolding blinding_respects_hashes_def by(rule blinding_of_T_root_hash) lemma blinding_of_on_T [locale_witness]: assumes "blinding_of_on A rh bo" shows "blinding_of_on {x. set1_T\<^sub>m x \ A} (root_hash_T rh) (blinding_of_T rh bo)" (is "blinding_of_on ?A ?h ?bo") proof - interpret a: blinding_of_on A rh bo by fact show ?thesis proof have "?bo x x \ (?bo x y \ ?bo y z \ ?bo x z) \ (?bo x y \ ?bo y x \ x = y)" if "x \ ?A" for x y z using that proof(induction x arbitrary: y z) case (T\<^sub>m x) interpret blinding_of_on "{a. set1_F\<^sub>m a \ A \ set3_F\<^sub>m a \ set3_F\<^sub>m x}" "root_hash_F rh ?h" "blinding_of_F rh bo ?h ?bo" apply(rule blinding_of_on_F[OF assms]) apply unfold_locales subgoal using T\<^sub>m.IH T\<^sub>m.prems by(force simp add: eq_onp_def) subgoal for a b c using T\<^sub>m.IH[of a b c] T\<^sub>m.prems by auto subgoal for a b using T\<^sub>m.IH[of a b] T\<^sub>m.prems by auto done show ?case using T\<^sub>m.prems apply(intro conjI) subgoal by(auto intro: blinding_of_T.intros refl) subgoal by(auto elim!: blinding_of_T.cases trans intro!: blinding_of_T.intros) subgoal by(auto elim!: blinding_of_T.cases dest: antisym) done qed then show "x \ ?A \ ?bo x x" and "\ ?bo x y; ?bo y z; x \ ?A \ \ ?bo x z" and "\ ?bo x y; ?bo y x; x \ ?A \ \ x = y" for x y z by blast+ qed qed lemmas blinding_of_T [locale_witness] = blinding_of_on_T[where A=UNIV, simplified] subsubsection \ Blinding on composition \ context fixes rha :: "('a\<^sub>m, 'a\<^sub>h) hash" and boa :: "'a\<^sub>m blinding_of" and rhb :: "('b\<^sub>m, 'b\<^sub>h) hash" and bob :: "'b\<^sub>m blinding_of" begin inductive blinding_of_G :: "('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) G\<^sub>m blinding_of" where "blinding_of_G (G\<^sub>m x) (G\<^sub>m y)" if "blinding_of_F (root_hash_T rha) (blinding_of_T rha boa) rhb bob x y" lemma blinding_of_G_unfold: "blinding_of_G = vimage2p the_G\<^sub>m the_G\<^sub>m (blinding_of_F (root_hash_T rha) (blinding_of_T rha boa) rhb bob)" apply(rule ext)+ subgoal for x y by(cases x; cases y)(simp_all add: blinding_of_G.simps fun_eq_iff vimage2p_def) done end lemma blinding_of_G_mono: assumes "boa \ boa'" "bob \ bob'" shows "blinding_of_G rha boa rhb bob \ blinding_of_G rha boa' rhb bob'" unfolding blinding_of_G_unfold by(rule vimage2p_mono' blinding_of_F_mono blinding_of_T_mono assms)+ lemma blinding_of_G_root_hash: assumes "boa \ vimage2p rha rha (=)" and "bob \ vimage2p rhb rhb (=)" shows "blinding_of_G rha boa rhb bob \ vimage2p (root_hash_G rha rhb) (root_hash_G rha rhb) (=)" unfolding blinding_of_G_unfold root_hash_G_unfold vimage2p_comp o_apply apply(rule vimage2p_mono') apply(rule order_trans) apply(rule blinding_respects_hashes_F[unfolded blinding_respects_hashes_def]) apply(rule blinding_of_T_root_hash) apply(rule assms)+ apply(rule vimage2p_mono') apply(simp add: vimage2p_def) done lemma blinding_of_on_G [locale_witness]: assumes "blinding_of_on A rha boa" "blinding_of_on B rhb bob" shows "blinding_of_on {x. set1_G\<^sub>m x \ A \ set3_G\<^sub>m x \ B} (root_hash_G rha rhb) (blinding_of_G rha boa rhb bob)" (is "blinding_of_on ?A ?h ?bo") proof - interpret a: blinding_of_on A rha boa by fact interpret b: blinding_of_on B rhb bob by fact interpret FT: blinding_of_on "{x. set1_F\<^sub>m x \ {x. set1_T\<^sub>m x \ A} \ set3_F\<^sub>m x \ B}" "root_hash_F (root_hash_T rha) rhb" "blinding_of_F (root_hash_T rha) (blinding_of_T rha boa) rhb bob" .. show ?thesis proof show "?bo \ vimage2p ?h ?h (=)" using a.hash b.hash by(rule blinding_of_G_root_hash) show "?bo x x" if "x \ ?A" for x using that by(cases x; hypsubst)(rule blinding_of_G.intros; rule FT.refl; auto) show "?bo x z" if "?bo x y" "?bo y z" "x \ ?A" for x y z using that by(fastforce elim!: blinding_of_G.cases intro!: blinding_of_G.intros elim!: FT.trans) show "x = y" if "?bo x y" "?bo y x" "x \ ?A" for x y using that by(clarsimp elim!: blinding_of_G.cases)(erule (1) FT.antisym; auto) qed qed lemmas blinding_of_G [locale_witness] = blinding_of_on_G[where A=UNIV and B=UNIV, simplified] subsection \Merging\ text \Two Merkle values with the same root hash can be merged into a less blinded Merkle value. The operation is unspecified for trees with different root hashes. \ subsubsection \ Merging on the base functor \ axiomatization merge_F :: "('a\<^sub>m, 'a\<^sub>h) hash \ 'a\<^sub>m merge \ ('b\<^sub>m, 'b\<^sub>h) hash \ 'b\<^sub>m merge \ ('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) F\<^sub>m merge" where merge_F_cong [fundef_cong]: "\ \a b. a \ set1_F\<^sub>m x \ ma a b = ma' a b; \a b. a \ set3_F\<^sub>m x \ mb a b = mb' a b \ \ merge_F rha ma rhb mb x y = merge_F rha ma' rhb mb' x y" and merge_on_F [locale_witness]: "\ merge_on A rha boa ma; merge_on B rhb bob mb \ \ merge_on {x. set1_F\<^sub>m x \ A \ set3_F\<^sub>m x \ B} (root_hash_F rha rhb) (blinding_of_F rha boa rhb bob) (merge_F rha ma rhb mb)" lemmas merge_F [locale_witness] = merge_on_F[where A=UNIV and B=UNIV, simplified] subsubsection \ Merging on the least fixpoint \ lemma wfP_subterm_T: "wfP (\x y. x \ set3_F\<^sub>m (the_T\<^sub>m y))" apply(rule wfPUNIVI) subgoal premises IH[rule_format] for P x by(induct x)(auto intro: IH) done +lemma irrefl_subterm_T: "x \ set3_F\<^sub>m y \ y \ the_T\<^sub>m x" + using wfP_subterm_T by (auto simp: wfP_def elim!: wf_irrefl) + context fixes rh :: "('a\<^sub>m, 'a\<^sub>h) hash" fixes m :: "'a\<^sub>m merge" begin function merge_T :: "('a\<^sub>m, 'a\<^sub>h) T\<^sub>m merge" where "merge_T (T\<^sub>m x) (T\<^sub>m y) = map_option T\<^sub>m (merge_F rh m (root_hash_T rh) merge_T x y)" by pat_completeness auto termination apply(relation "{(x, y). x \ set3_F\<^sub>m (the_T\<^sub>m y)} <*lex*> {(x, y). x \ set3_F\<^sub>m (the_T\<^sub>m y)}") - apply(auto simp add: wfP_def[symmetric] wfP_subterm_T) + apply(auto simp add: wfP_def[symmetric] wfP_subterm_T dest!: irrefl_subterm_T) done lemma merge_on_T [locale_witness]: assumes "merge_on A rh bo m" shows "merge_on {x. set1_T\<^sub>m x \ A} (root_hash_T rh) (blinding_of_T rh bo) merge_T" (is "merge_on ?A ?h ?bo ?m") proof - interpret a: merge_on A rh bo m by fact show ?thesis proof have "(?h a = ?h b \ (\ab. ?m a b = Some ab \ ?bo a ab \ ?bo b ab \ (\u. ?bo a u \ ?bo b u \ ?bo ab u))) \ (?h a \ ?h b \ ?m a b = None)" if "a \ ?A" for a b using that unfolding mem_Collect_eq proof(induction a arbitrary: b) case (T\<^sub>m x y) interpret merge_on "{y. set1_F\<^sub>m y \ A \ set3_F\<^sub>m y \ set3_F\<^sub>m x}" "root_hash_F rh ?h" "blinding_of_F rh bo ?h ?bo" "merge_F rh m ?h ?m" proof fix a assume a: "a \ set3_F\<^sub>m x" with T\<^sub>m.prems have a': "set1_T\<^sub>m a \ A" by auto fix b from T\<^sub>m.IH[OF a a', rule_format, of b] show "root_hash_T rh a = root_hash_T rh b \ \ab. merge_T a b = Some ab \ blinding_of_T rh bo a ab \ blinding_of_T rh bo b ab \ (\u. blinding_of_T rh bo a u \ blinding_of_T rh bo b u \ blinding_of_T rh bo ab u)" and "root_hash_T rh a \ root_hash_T rh b \ merge_T a b = None" by(auto dest: sym) qed show ?case using T\<^sub>m.prems apply(intro conjI strip) subgoal by(cases y)(auto dest!: join simp add: blinding_of_T.simps) subgoal by(cases y)(auto dest!: undefined) done qed then show "?h a = ?h b \ \ab. ?m a b = Some ab \ ?bo a ab \ ?bo b ab \ (\u. ?bo a u \ ?bo b u \ ?bo ab u)" "?h a \ ?h b \ ?m a b = None" if "a \ ?A" for a b using that by blast+ qed qed lemmas merge_T [locale_witness] = merge_on_T[where A=UNIV, simplified] end lemma merge_T_cong [fundef_cong]: assumes "\a b. a \ set1_T\<^sub>m x \ m a b = m' a b" shows "merge_T rh m x y = merge_T rh m' x y" using assms apply(induction x y rule: merge_T.induct) apply simp apply(rule arg_cong[where f="map_option _"]) apply(blast intro: merge_F_cong) done subsubsection \ Merging and composition \ context fixes rha :: "('a\<^sub>m, 'a\<^sub>h) hash" fixes ma :: "'a\<^sub>m merge" fixes rhb :: "('b\<^sub>m, 'b\<^sub>h) hash" fixes mb :: "'b\<^sub>m merge" begin primrec merge_G :: "('a\<^sub>m, 'a\<^sub>h, 'b\<^sub>m, 'b\<^sub>h) G\<^sub>m merge" where "merge_G (G\<^sub>m x) y' = (case y' of G\<^sub>m y \ map_option G\<^sub>m (merge_F (root_hash_T rha) (merge_T rha ma) rhb mb x y))" lemma merge_G_simps [simp]: "merge_G (G\<^sub>m x) (G\<^sub>m y) = map_option G\<^sub>m (merge_F (root_hash_T rha) (merge_T rha ma) rhb mb x y)" by(simp) declare merge_G.simps [simp del] lemma merge_on_G: assumes a: "merge_on A rha boa ma" and b: "merge_on B rhb bob mb" shows "merge_on {x. set1_G\<^sub>m x \ A \ set3_G\<^sub>m x \ B} (root_hash_G rha rhb) (blinding_of_G rha boa rhb bob) merge_G" (is "merge_on ?A ?h ?bo ?m") proof - interpret a: merge_on A rha boa ma by fact interpret b: merge_on B rhb bob mb by fact interpret F: merge_on "{x. set1_F\<^sub>m x \ {x. set1_T\<^sub>m x \ A} \ set3_F\<^sub>m x \ B}" "root_hash_F (root_hash_T rha) rhb" "blinding_of_F (root_hash_T rha) (blinding_of_T rha boa) rhb bob" "merge_F (root_hash_T rha) (merge_T rha ma) rhb mb" .. show ?thesis proof show "\ab. ?m a b = Some ab \ ?bo a ab \ ?bo b ab \ (\u. ?bo a u \ ?bo b u \ ?bo ab u)" if "?h a = ?h b" "a \ ?A" for a b using that by(cases a; cases b)(auto dest!: F.join simp add: blinding_of_G.simps) show "?m a b = None" if "?h a \ ?h b" "a \ ?A" for a b using that by(cases a; cases b)(auto dest!: F.undefined) qed qed lemmas merge_G [locale_witness] = merge_on_G[where A=UNIV and B=UNIV, simplified] end lemma merge_G_cong [fundef_cong]: "\ \a b. a \ set1_G\<^sub>m x \ ma a b = ma' a b; \a b. a \ set3_G\<^sub>m x \ mb a b = mb' a b \ \ merge_G rha ma rhb mb x y = merge_G rha ma' rhb mb' x y" apply(cases x; cases y; simp) apply(rule arg_cong[where f="map_option _"]) apply(blast intro: merge_F_cong merge_T_cong) done end diff --git a/thys/Collections/Examples/ICF/Exploration.thy b/thys/Collections/Examples/ICF/Exploration.thy --- a/thys/Collections/Examples/ICF/Exploration.thy +++ b/thys/Collections/Examples/ICF/Exploration.thy @@ -1,295 +1,295 @@ (* Title: Isabelle Collections Library Author: Peter Lammich Maintainer: Peter Lammich *) section \\isaheader{State Space Exploration}\ theory Exploration imports Main Collections.Collections begin text_raw \\label{thy:Exploration}\ text \ In this theory, a workset algorithm for state-space exploration is defined. It explores the set of states that are reachable by a given relation, starting at a given set of initial states. The specification makes use of the data refinement framework for while-algorithms (cf. Section~\ref{thy:DatRef}), and is thus suited for being refined to an executable algorithm, using the Isabelle Collections Framework to provide the necessary data structures. \ subsection "Generic Search Algorithm" \ \The algorithm contains a set of discovered states and a workset\ type_synonym '\ sse_state = "'\ set \ '\ set" \ \Loop body\ inductive_set sse_step :: "('\\'\) set \ ('\ sse_state \ '\ sse_state) set" for R where "\ \\W; \' = \ \ (R``{\}); W' = (W-{\}) \ ((R``{\}) - \) \ \ ((\,W),(\',W'))\sse_step R" \ \Loop condition\ definition sse_cond :: "'\ sse_state set" where "sse_cond = {(\,W). W\{}}" \ \Initial state\ definition sse_initial :: "'\ set \ '\ sse_state" where "sse_initial \i == (\i,\i)" \ \Invariants: \begin{itemize} \item The workset contains only states that are already discovered. \item All discovered states are target states \item If there is a target state that is not discovered yet, then there is an element in the workset from that this target state is reachable without using discovered states as intermediate states. This supports the intuition of the workset as a frontier between the sets of discovered and undiscovered states. \end{itemize}\ definition sse_invar :: "'\ set \ ('\\'\) set \ '\ sse_state set" where "sse_invar \i R = {(\,W). W\\ \ (\ \ R\<^sup>*``\i) \ (\\\(R\<^sup>*``\i)-\. \\h\W. (\h,\)\(R - (UNIV \ \))\<^sup>*) }" definition "sse_algo \i R == \ wa_cond=sse_cond, wa_step=sse_step R, wa_initial = {sse_initial \i}, wa_invar = sse_invar \i R \" definition "sse_term_rel \i R == { (\',\). \\sse_invar \i R \ (\,\')\sse_step R }" \ \Termination: Either a new state is discovered, or the workset shrinks\ theorem sse_term: assumes finite[simp, intro!]: "finite (R\<^sup>*``\i)" shows "wf (sse_term_rel \i R)" proof - have "wf (({(\',\). \ \ \' \ \' \ (R\<^sup>*``\i)}) <*lex*> finite_psubset)" by (auto intro: wf_bounded_supset) moreover have "sse_term_rel \i R \ \" (is "_ \ ?R") proof fix S assume A: "S\sse_term_rel \i R" obtain \ W \' W' \ where [simp]: "S=((\',W'),(\,W))" and S: "(\,W) \ sse_invar \i R" "\\W" "\' = \ \ R``{\}" "W' = (W-{\}) \ (R``{\} - \)" proof - obtain \ W \' W' where SF[simp]: "S=((\',W'),(\,W))" by (cases S) force from A have R: "(\,W) \ sse_invar \i R" "((\,W),(\',W'))\sse_step R" by (auto simp add: sse_term_rel_def) from sse_step.cases[OF R(2)] obtain \ where S: "\\W" "\' = \ \ R``{\}" "W' = (W-{\}) \ (R``{\} - \)" by metis thus ?thesis by (rule_tac that[OF SF R(1) S]) qed from S(1) have [simp, intro!]: "finite \" "finite W" and WSS: "W\\" and SSS: "\\R\<^sup>*``\i" by (auto simp add: sse_invar_def intro: finite_subset) show "S\?R" proof (cases "R``{\} \ \") case True with S have "\'=\" "W' \ W" by auto thus ?thesis by (simp) next case False with S have "\' \ \" by auto moreover from S(2) WSS SSS have "\\R\<^sup>*``\i" by auto hence "R``{\} \ R\<^sup>*``\i" by (auto intro: rtrancl_into_rtrancl) with S(3) SSS have "\' \ R\<^sup>*``\i" by auto - ultimately show ?thesis by simp + ultimately show ?thesis by force qed qed ultimately show ?thesis by (auto intro: wf_subset) qed lemma sse_invar_initial: "(sse_initial \i) \ sse_invar \i R" by (unfold sse_invar_def sse_initial_def) (auto elim: rtrancl_last_touch) \ \Correctness theorem: If the loop terminates, the discovered states are exactly the reachable states\ theorem sse_invar_final: "\S. S\wa_invar (sse_algo \i R) \ S\wa_cond (sse_algo \i R) \ fst S = R\<^sup>*``\i" by (intro allI, case_tac S) (auto simp add: sse_invar_def sse_cond_def sse_algo_def) lemma sse_invar_step: "\S\sse_invar \i R; (S,S')\sse_step R\ \ S'\sse_invar \i R" \ \Split the goal by the invariant:\ apply (cases S, cases S') apply clarsimp apply (erule sse_step.cases) apply clarsimp apply (subst sse_invar_def) apply (simp add: Let_def split_conv) apply (intro conjI) \ \Solve the easy parts automatically\ apply (auto simp add: sse_invar_def) [3] apply (force simp add: sse_invar_def dest: rtrancl_into_rtrancl) [1] \ \Tackle the complex part (last part of the invariant) in Isar\ proof (intro ballI) fix \ W \ \' assume A: "(\,W)\sse_invar \i R" "\\W" "\'\R\<^sup>* `` \i - (\ \ R `` {\})" \ \Using the invariant of the original state, we obtain a state in the original workset and a path not touching the originally discovered states\ from A(3) have "\' \ R\<^sup>* `` \i - \" by auto with A(1) obtain \h where IP: "\h\W" "(\h,\')\(R - (UNIV \ \))\<^sup>*" and SS: "W\\" "\\R\<^sup>* `` \i" by (unfold sse_invar_def) force \ \We now make a case distinction, whether the obtained path contains states from @{term "post \"} or not:\ from IP(2) show "\\h\W - {\} \ (R `` {\} - \). (\h, \') \ (R - UNIV \ (\ \ R `` {\}))\<^sup>*" proof (cases rule: rtrancl_last_visit[where S="R `` {\}"]) case no_visit \ \In the case that the obtained path contains no states from @{term "post \"}, we can take it.\ hence G1: "(\h,\')\(R- (UNIV \ (\\R `` {\})))\<^sup>*" by (simp add: set_diff_diff_left Sigma_Un_distrib2) moreover have "\h \ \" \ \We may exclude the case that our obtained path started at @{text \}, as all successors of @{text \} are in @{term "R `` {\}"}\ proof assume [simp]: "\h=\" from A SS have "\\\'" by auto with G1 show False by (auto simp add: elim: converse_rtranclE) qed ultimately show ?thesis using IP(1) by auto next case (last_visit_point \t) \ \If the obtained path contains a state from @{text "R `` {\}"}, we simply pick the last one:\ hence "(\t,\')\(R- (UNIV \ (\\R `` {\})))\<^sup>*" by (simp add: set_diff_diff_left Sigma_Un_distrib2) moreover from last_visit_point(2) have "\t\\" by (auto elim: trancl.cases) ultimately show ?thesis using last_visit_point(1) by auto qed qed \ \The sse-algorithm is a well-defined while-algorithm\ theorem sse_while_algo: "finite (R\<^sup>*``\i) \ while_algo (sse_algo \i R)" apply unfold_locales apply (auto simp add: sse_algo_def intro: sse_invar_step sse_invar_initial) apply (drule sse_term) apply (erule_tac wf_subset) apply (unfold sse_term_rel_def) apply auto done subsection "Depth First Search" text_raw \\label{sec:sse:dfs}\ text \ In this section, the generic state space exploration algorithm is refined to a DFS-algorithm, that uses a stack to implement the workset. \ type_synonym '\ dfs_state = "'\ set \ '\ list" definition dfs_\ :: "'\ dfs_state \ '\ sse_state" where "dfs_\ S == let (\,W)=S in (\,set W)" definition dfs_invar_add :: "'\ dfs_state set" where "dfs_invar_add == {(\,W). distinct W}" definition "dfs_invar \i R == dfs_invar_add \ { s. dfs_\ s \ sse_invar \i R }" inductive_set dfs_initial :: "'\ set \ '\ dfs_state set" for \i where "\ distinct W; set W = \i\ \ (\i,W)\dfs_initial \i" inductive_set dfs_step :: "('\\'\) set \ ('\ dfs_state \'\ dfs_state) set" for R where "\ W=\#Wtl; distinct Wn; set Wn = R``{\} - \; W' = Wn@Wtl; \' = R``{\} \ \ \ \ ((\,W),(\',W'))\dfs_step R" definition dfs_cond :: "'\ dfs_state set" where "dfs_cond == { (\,W). W\[]}" definition "dfs_algo \i R == \ wa_cond = dfs_cond, wa_step = dfs_step R, wa_initial = dfs_initial \i, wa_invar = dfs_invar \i R \" \ \The DFS-algorithm refines the state-space exploration algorithm\ theorem dfs_pref_sse: "wa_precise_refine (dfs_algo \i R) (sse_algo \i R) dfs_\" apply (unfold_locales) apply (auto simp add: dfs_algo_def sse_algo_def dfs_cond_def sse_cond_def dfs_\_def) apply (erule dfs_step.cases) apply (rule_tac \=\ in sse_step.intros) apply (auto simp add: dfs_invar_def sse_invar_def dfs_invar_add_def sse_initial_def dfs_\_def elim: dfs_initial.cases) done \ \The DFS-algorithm is a well-defined while-algorithm\ theorem dfs_while_algo: assumes finite[simp, intro!]: "finite (R\<^sup>*``\i)" shows "while_algo (dfs_algo \i R)" proof - interpret wa_precise_refine "(dfs_algo \i R)" "(sse_algo \i R)" dfs_\ using dfs_pref_sse . have [simp]: "wa_invar (sse_algo \i R) = sse_invar \i R" by (simp add: sse_algo_def) show ?thesis apply (rule wa_intro) apply (simp add: sse_while_algo) apply (simp add: dfs_invar_def dfs_algo_def) apply (auto simp add: dfs_invar_add_def dfs_algo_def dfs_\_def dfs_cond_def sse_invar_def elim!: dfs_step.cases) [1] apply (auto simp add: dfs_invar_add_def dfs_algo_def elim: dfs_initial.cases) [1] done qed \ \The result of the DFS-algorithm is correct\ lemmas dfs_invar_final = wa_precise_refine.transfer_correctness[OF dfs_pref_sse sse_invar_final] end diff --git a/thys/Higher_Order_Terms/Name.thy b/thys/Higher_Order_Terms/Name.thy --- a/thys/Higher_Order_Terms/Name.thy +++ b/thys/Higher_Order_Terms/Name.thy @@ -1,126 +1,126 @@ chapter \Names as a unique datatype\ theory Name imports Main begin text \ I would like to model names as @{typ string}s. Unfortunately, there is no default order on lists, as there could be multiple reasonable implementations: e.g.\ lexicographic and point-wise. For both choices, users can import the corresponding instantiation. In Isabelle, only at most one implementation of a given type class for a given type may be present in the same theory. Consequently, I avoided importing a list ordering from the library, because it may cause conflicts with users who use another ordering. The general approach for these situations is to introduce a type copy. The full flexibility of strings (i.e.\ string manipulations) is only required where fresh names are being produced. Otherwise, only a linear order on terms is needed. Conveniently, Sternagel and Thiemann @{cite sternagel2015deriving} provide tooling to automatically generate such a lexicographic order. \ datatype name = Name (as_string: string) \ \Mostly copied from \List_Lexorder\\ instantiation name :: ord begin definition less_name where "xs < ys \ (as_string xs, as_string ys) \ lexord {(u, v). (of_char u :: nat) < of_char v}" definition less_eq_name where "(xs :: name) \ ys \ xs < ys \ xs = ys" instance .. end instance name :: order proof fix xs :: "name" show "xs \ xs" by (simp add: less_eq_name_def) next fix xs ys zs :: "name" assume "xs \ ys" and "ys \ zs" then show "xs \ zs" apply (auto simp add: less_eq_name_def less_name_def) apply (rule lexord_trans) - apply (auto intro: transI) + apply (auto intro: transI simp: antisym_def) done next fix xs ys :: "name" assume "xs \ ys" and "ys \ xs" then show "xs = ys" apply (auto simp add: less_eq_name_def less_name_def) apply (rule lexord_irreflexive [THEN notE]) defer apply (rule lexord_trans) - apply (auto intro: transI) + apply (auto intro: transI simp: antisym_def) done next fix xs ys :: "name" show "xs < ys \ xs \ ys \ \ ys \ xs" apply (auto simp add: less_name_def less_eq_name_def) defer apply (rule lexord_irreflexive [THEN notE]) apply auto apply (rule lexord_irreflexive [THEN notE]) defer apply (rule lexord_trans) - apply (auto intro: transI) + apply (auto intro: transI simp: antisym_def) done qed instance name :: linorder proof fix xs ys :: "name" have "(as_string xs, as_string ys) \ lexord {(u, v). (of_char u::nat) < of_char v} \ xs = ys \ (as_string ys, as_string xs) \ lexord {(u, v). (of_char u::nat) < of_char v}" by (metis (no_types, lifting) case_prodI lexord_linear linorder_neqE_nat mem_Collect_eq name.expand of_char_eq_iff) then show "xs \ ys \ ys \ xs" by (auto simp add: less_eq_name_def less_name_def) qed lemma less_name_code[code]: "Name xs < Name [] \ False" "Name [] < Name (x # xs) \ True" "Name (x # xs) < Name (y # ys) \ (of_char x::nat) < of_char y \ x = y \ Name xs < Name ys" unfolding less_name_def by auto lemma le_name_code[code]: "Name (x # xs) \ Name [] \ False" "Name [] \ Name (x # xs) \ True" "Name (x # xs) \ Name (y # ys) \ (of_char x::nat) < of_char y \ x = y \ Name xs \ Name ys" unfolding less_eq_name_def less_name_def by auto context begin qualified definition append :: "name \ name \ name" where "append v1 v2 = Name (as_string v1 @ as_string v2)" lemma name_append_less: assumes "xs \ Name []" shows "append ys xs > ys" proof - have "Name (ys @ xs) > Name ys" if "xs \ []" for xs ys using that proof (induction ys) case Nil thus ?case unfolding less_name_def by (cases xs) auto next case (Cons y ys) thus ?case unfolding less_name_def by auto qed with assms show ?thesis unfolding append_def by (cases xs, cases ys) auto qed end end \ No newline at end of file diff --git a/thys/IMP2/doc/Examples.thy b/thys/IMP2/doc/Examples.thy --- a/thys/IMP2/doc/Examples.thy +++ b/thys/IMP2/doc/Examples.thy @@ -1,1872 +1,1873 @@ theory Examples imports "../IMP2" "../lib/IMP2_Aux_Lemmas" begin section \Examples\ lemmas nat_distribs = nat_add_distrib nat_diff_distrib Suc_diff_le nat_mult_distrib nat_div_distrib subsection \Common Loop Patterns\ subsubsection \Count Up\ text \ Counter \c\ counts from \0\ to \n\, such that loop is executed \n\ times. The result is computed in an accumulator \a\. \ text \The invariant states that we have computed the function for the counter value \c\\ text \The variant is the difference between \n\ and \c\, i.e., the number of loop iterations that we still have to do\ program_spec exp_count_up assumes "0\n" ensures "a = 2^nat n\<^sub>0" defines \ a = 1; c = 0; while (cn-c\ @invariant \0\c \ c\n \ a=2^nat c\ { G_par = a; scope { a = G_par; a=2*a; G_ret = a }; a = G_ret; c=c+1 } \ apply vcg by (auto simp: algebra_simps nat_distribs) program_spec sum_prog assumes "n \ 0" ensures "s = \{0..n\<^sub>0}" defines \ s = 0; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n\<^sub>0 = n \ 0 \ i \ i \ n \ s = \{0..i}\ { i = i + 1; s = s + i } \ apply vcg_cs by (simp_all add: intvs_incdec) program_spec sq_prog assumes "n \ 0" ensures "a = n\<^sub>0 * n\<^sub>0" defines \ a = 0; z = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n\<^sub>0 = n \ 0 \ i \ i \ n \ a = i * i \ z = 2 * i + 1\ { a = a + z; z = z + 2; i = i + 1 } \ by vcg_cs (simp add: algebra_simps) fun factorial :: "int \ int" where "factorial i = (if i \ 0 then 1 else i * factorial (i - 1))" program_spec factorial_prog assumes "n \ 0" ensures "a = factorial n\<^sub>0" defines \ a = 1; i = 1; while (i \ n) @variant \n\<^sub>0 + 1 - i\ @invariant \n\<^sub>0 = n \ 1 \ i \ i \ n + 1 \ a = factorial (i - 1)\ { a = a * i; i = i + 1 } \ by vcg (simp add: antisym_conv)+ fun fib :: "int \ int" where "fib i = (if i \ 0 then 0 else if i = 1 then 1 else fib (i - 2) + fib (i - 1))" lemma fib_simps[simp]: "i \ 0 \ fib i = 0" "i = 1 \ fib i = 1" "i > 1 \ fib i = fib (i - 2) + fib (i - 1)" by simp+ lemmas [simp del] = fib.simps text \With precondition\ program_spec fib_prog assumes "n \ 0" ensures "a = fib n" defines \ a = 0; b = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n = n\<^sub>0 \ 0 \ i \ i \ n \ a = fib i \ b = fib (i + 1)\ { c = b; b = a + b; a = c; i = i + 1 } \ by vcg_cs (simp add: algebra_simps) text \Without precondition, returning \0\ for negative numbers\ program_spec fib_prog' assumes True ensures "a = fib n\<^sub>0" defines \ a = 0; b = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n = n\<^sub>0 \ (0 \ i \ i \ n \ n\<^sub>0 < 0 \ i = 0) \ a = fib i \ b = fib (i+ 1)\ { c = b; b = a + b; a = c; i = i + 1 } \ by vcg_cs (auto simp: algebra_simps) subsubsection \Count down\ text \Essentially the same as count up, but we (ab)use the input variable as a counter.\ text \The invariant is the same as for count-up. Only that we have to compute the actual number of loop iterations by \n\<^sub>0 - n\. We locally introduce the name \c\ for that. \ program_spec exp_count_down assumes "0\n" ensures "a = 2^nat n\<^sub>0" defines \ a = 1; while (n>0) @variant \n\ @invariant \let c = n\<^sub>0-n in 0\n \ n\n\<^sub>0 \ a=2^nat c\ { a=2*a; n=n-1 } \ apply vcg_cs by (auto simp: algebra_simps nat_distribs) subsubsection \Approximate from Below\ text \Used to invert a monotonic function. We count up, until we overshoot the desired result, then we subtract one. \ text \The invariant states that the \r-1\ is not too big. When the loop terminates, \r-1\ is not too big, but \r\ is already too big, so \r-1\ is the desired value (rounding down). \ text \The variant measures the gap that we have to the correct result. Note that the loop will do a final iteration, when the result has been reached exactly. We account for that by adding one, such that the measure also decreases in this case. \ program_spec sqr_approx_below assumes "0\n" ensures "0\r \ r\<^sup>2 \ n\<^sub>0 \ n\<^sub>0 < (r+1)\<^sup>2" defines \ r=1; while (r*r \ n) @variant \n + 1 - r*r\ @invariant \0\r \ (r-1)\<^sup>2 \ n\<^sub>0\ { r = r + 1 }; r = r - 1 \ apply vcg by (auto simp: algebra_simps power2_eq_square) subsubsection \Bisection\ text \A more efficient way of inverting monotonic functions is by bisection, that is, one keeps track of a possible interval for the solution, and halfs the interval in each step. The program will need $O(\log n)$ iterations, and is thus very efficient in practice. Although the final algorithm looks quite simple, getting it right can be quite tricky. \ text \The invariant is surprisingly easy, just stating that the solution is in the interval \l... \ lemma "\h l n\<^sub>0 :: int. \\''invar-final''; 0 \ n\<^sub>0; \ 1 + l < h; 0 \ l; l < h; l * l \ n\<^sub>0; n\<^sub>0 < h * h\ \ n\<^sub>0 < 1 + (l * l + l * 2)" by (smt mult.commute semiring_normalization_rules(3)) program_spec sqr_bisect assumes "0\n" ensures "r\<^sup>2\n\<^sub>0 \ n\<^sub>0<(r+1)\<^sup>2" defines \ l=0; h=n+1; while (l+1 < h) @variant \h-l\ @invariant \0\l \ l l\<^sup>2\n \ n < h\<^sup>2\ { m = (l + h) / 2; if (m*m \ n) l=m else h=m }; r=l \ apply vcg text \We use quick-and-dirty apply style proof to discharge the VCs\ apply (auto simp: power2_eq_square algebra_simps add_pos_pos) apply (smt not_sum_squares_lt_zero) by (smt mult.commute semiring_normalization_rules(3)) subsection \Debugging\ subsubsection \Testing Programs\ text \Stepwise\ schematic_goal "Map.empty: (sqr_approx_below,<''n'':=\_. 4>) \ ?s" unfolding sqr_approx_below_def apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' done text \Or all steps at once\ schematic_goal "Map.empty: (sqr_bisect,<''n'':=\_. 4900000001>) \ ?s" unfolding sqr_bisect_def by big_step subsection \More Numeric Algorithms\ subsubsection \Euclid's Algorithm (with subtraction)\ (* Crucial Lemmas *) thm gcd.commute gcd_diff1 program_spec euclid1 assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0" defines \ while (a\b) @invariant \gcd a b = gcd a\<^sub>0 b\<^sub>0 \ (a>0 \ b>0)\ @variant \a+b\ { if (a apply vcg_cs apply (metis gcd.commute gcd_diff1) apply (metis gcd_diff1) done subsubsection \Euclid's Algorithm (with mod)\ (* Crucial Lemmas *) thm gcd_red_int[symmetric] program_spec euclid2 assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0" defines \ while (b\0) @invariant \gcd a b = gcd a\<^sub>0 b\<^sub>0 \ b\0 \ a>0\ @variant \b\ { t = a; a = b; b = t mod b } \ apply vcg_cs by (simp add: gcd_red_int[symmetric]) subsubsection \Extended Euclid's Algorithm\ (* Provided by Simon Wimmer. *) locale extended_euclid_aux_lemmas begin lemma aux2: fixes a b :: int assumes "b = t * b\<^sub>0 + s * a\<^sub>0" "q = a div b" "gcd a b = gcd a\<^sub>0 b\<^sub>0" shows "gcd b (a - (a\<^sub>0 * (s * q) + b\<^sub>0 * (t * q))) = gcd a\<^sub>0 b\<^sub>0" proof - have "a - (a\<^sub>0 * (s * q) + b\<^sub>0 * (t * q)) = a - b * q" unfolding \b = _\ by (simp add: algebra_simps) also have "a - b * q = a mod b" unfolding \q = _\ by (simp add: algebra_simps) finally show ?thesis unfolding \gcd a b = _\[symmetric] by (simp add: gcd_red_int[symmetric]) qed lemma aux3: fixes a b :: int assumes "b = t * b\<^sub>0 + s * a\<^sub>0" "q = a div b" "b > 0" shows "t * (b\<^sub>0 * q) + s * (a\<^sub>0 * q) \ a" proof - have "t * (b\<^sub>0 * q) + s * (a\<^sub>0 * q) = q * b" unfolding \b = _\ by (simp add: algebra_simps) then show ?thesis using \b > 0\ by (simp add: algebra_simps \q = _\) (smt Euclidean_Division.pos_mod_sign cancel_div_mod_rules(1) mult.commute) qed end text \The following is a direct translation of the pseudocode for the Extended Euclidean algorithm as described by the English version of Wikipedia (\<^url>\https://en.wikipedia.org/wiki/Extended_Euclidean_algorithm\): \ program_spec euclid_extended assumes "a>0 \ b>0" ensures "old_r = gcd a\<^sub>0 b\<^sub>0 \ gcd a\<^sub>0 b\<^sub>0 = a\<^sub>0 * old_s + b\<^sub>0 * old_t" defines \ s = 0; old_s = 1; t = 1; old_t = 0; r = b; old_r = a; while (r\0) @invariant \ gcd old_r r = gcd a\<^sub>0 b\<^sub>0 \ r\0 \ old_r>0 \ a\<^sub>0 * old_s + b\<^sub>0 * old_t = old_r \ a\<^sub>0 * s + b\<^sub>0 * t = r \ @variant \r\ { quotient = old_r / r; temp = old_r; old_r = r; r = temp - quotient * r; temp = old_s; old_s = s; s = temp - quotient * s; temp = old_t; old_t = t; t = temp - quotient * t } \ proof - interpret extended_euclid_aux_lemmas . show ?thesis apply vcg_cs apply (simp add: algebra_simps) apply (simp add: aux2 aux3 minus_div_mult_eq_mod)+ done qed text \Non-Wikipedia version\ context extended_euclid_aux_lemmas begin lemma aux: fixes a b q x y:: int assumes "a = old_y * b\<^sub>0 + old_x * a\<^sub>0" "b = y * b\<^sub>0 + x * a\<^sub>0" "q = a div b" shows "a mod b + (a\<^sub>0 * (x * q) + b\<^sub>0 * (y * q)) = a" proof - have *: "a\<^sub>0 * (x * q) + b\<^sub>0 * (y * q) = q * b" unfolding \b = _\ by (simp add: algebra_simps) show ?thesis unfolding * unfolding \q = _\ by simp qed end program_spec euclid_extended' assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0 \ gcd a\<^sub>0 b\<^sub>0 = a\<^sub>0 * x + b\<^sub>0 * y" defines \ x = 0; y = 1; old_x = 1; old_y = 0; while (b\0) @invariant \ gcd a b = gcd a\<^sub>0 b\<^sub>0 \ b\0 \ a>0 \ a = a\<^sub>0 * old_x + b\<^sub>0 * old_y \ b = a\<^sub>0 * x + b\<^sub>0 * y \ @variant \b\ { q = a / b; t = a; a = b; b = t mod b; t = x; x = old_x - q * x; old_x = t; t = y; y = old_y - q * y; old_y = t }; x = old_x; y = old_y \ proof - interpret extended_euclid_aux_lemmas . show ?thesis apply vcg_cs apply (simp add: gcd_red_int[symmetric]) apply (simp add: algebra_simps) apply (rule aux; simp add: algebra_simps) done qed subsubsection \Exponentiation by Squaring\ (* Contributed by Simon Wimmer *) lemma ex_by_sq_aux: fixes x :: int and n :: nat assumes "n mod 2 = 1" shows "x * (x * x) ^ (n div 2) = x ^ n" proof - have "n > 0" using assms by presburger have "2 * (n div 2) = n - 1" using assms by presburger then have "(x * x) ^ (n div 2) = x ^ (n - 1)" by (simp add: semiring_normalization_rules) with \0 < n\ show ?thesis by simp (metis Suc_pred power.simps(2)) qed text \A classic algorithm for computing \x\<^sup>n\ works by repeated squaring, using the following recurrence: \<^item> \x\<^sup>n = x * x\<^bsup>(n-1)/2\<^esup>\ if \n\ is odd \<^item> \x\<^sup>n = x\<^bsup>n/2\<^esup>\ if \n\ is even \ program_spec ex_by_sq assumes "n\0" ensures "r = x\<^sub>0 ^ nat n\<^sub>0" defines \ r = 1; while (n\0) @invariant \ n \ 0 \ r * x ^ nat n = x\<^sub>0 ^ nat n\<^sub>0 \ @variant \n\ { if (n mod 2 == 1) { r = r * x }; x = x * x; n = n / 2 } \ apply vcg_cs subgoal apply (subst (asm) ex_by_sq_aux[symmetric]) apply (smt Suc_1 nat_1 nat_2 nat_mod_distrib nat_one_as_int) by (simp add: div_int_pos_iff int_div_less_self nat_div_distrib) apply (simp add: algebra_simps semiring_normalization_rules nat_mult_distrib; fail) apply (simp add: algebra_simps semiring_normalization_rules nat_mult_distrib; fail) done subsubsection \Power-Tower of 2s\ fun tower2 where "tower2 0 = 1" | "tower2 (Suc n) = 2 ^ tower2 n" definition "tower2' n = int (tower2 (nat n))" program_spec tower2_imp assumes \m>0\ ensures \a = tower2' m\<^sub>0\ defines \ a=1; while (m>0) @variant \m\ @invariant \0\m \ m\m\<^sub>0 \ a = tower2' (m\<^sub>0-m)\ { n=a; a = 1; while (n>0) @variant \n\ @invariant \True\ \ \This will get ugly, there is no \n\<^sub>0\ that we could use!\ { a=2*a; n=n-1 }; m=m-1 } \ oops text \We prove the inner loop separately instead! (It happens to be exactly our \<^const>\exp_count_down\ program.) \ program_spec tower2_imp assumes \m>0\ ensures \a = tower2' m\<^sub>0\ defines \ a=1; while (m>0) @variant \m\ @invariant \0\m \ m\m\<^sub>0 \ a = tower2' (m\<^sub>0-m)\ { n=a; inline exp_count_down; m=m-1 } \ apply vcg_cs by (auto simp: algebra_simps tower2'_def nat_distribs) subsection \Array Algorithms\ (* Presentation in lecture: introduce array syntax and semantics CONVENTIONS: * Every variable is an array. * Arrays are modeled as int\int. Negative indexes allowed. * Syntactic sugar to use it at index 0 by default V vname \ Vidx vname aexp Syntax: x[i] abbreviation "V x = Vidx x 0" Assign vname aexp \ AssignIdx vname aexp aexp Syntax: x[i] = a abbreviation: Assign x a = AssignIdx x 0 a NEW COMMAND: ArrayCpy vname vname Syntax: x[] = y copy two arrays as a whole. Assign x (V y) only copies index 0 semantics: Obvious aval (Vidx x i) s = s x (aval i s) (x[i] = a, s) \ s(x := (s x)(aval i s := aval a s)) (x[] = y, s) \ s(x := s y) weakest preconditions: Also obvious \ Reasoning with arrays: Usually: Model as int\int is appropriate. Common idioms: * {l..i Warning: Sledgehammer does not perform well on these! Data Refinement Sometimes, abstraction to list may be useful. Idea: Model an algorithm on lists first. Prove it correct. Then implement algorithm. Show that it implements the model. By transitivity, get: Algorithm is correct! Sometimes, abstraction to list may be useful *) subsubsection \Summation\ program_spec array_sum assumes "l\h" ensures "r = (\i=l\<^sub>0..0. a\<^sub>0 i)" defines \ r = 0; while (ll\<^sub>0\l \ l\h \ r = (\i=l\<^sub>0..0 i)\ @variant \h-l\ { r = r + a[l]; l=l+1 } \ apply vcg_cs by (auto simp: intvs_incdec) (* Exercise: Remove l\h precondition! *) subsubsection \Finding Least Index of Element\ program_spec find_least_idx assumes \l\h\ ensures \if l=h\<^sub>0 then x\<^sub>0 \ a\<^sub>0`{l\<^sub>0..0} else l\{l\<^sub>0..0} \ a\<^sub>0 l = x\<^sub>0 \ x\<^sub>0\a\<^sub>0`{l\<^sub>0.. defines \ while (l a[l] \ x) @variant \h-l\ @invariant \l\<^sub>0\l \ l\h \ x\a`{l\<^sub>0.. l=l+1 \ apply vcg_cs by (smt atLeastLessThan_iff imageI) subsubsection \Check for Sortedness\ term ran_sorted program_spec check_sorted assumes \l\h\ ensures \r\0 \ ran_sorted a\<^sub>0 l\<^sub>0 h\<^sub>0\ defines \ if (l==h) r=1 else { l=l+1; while (l a[l-1] \ a[l]) @variant \h-l\ @invariant \l\<^sub>0 l\h \ ran_sorted a l\<^sub>0 l\ l=l+1; if (l==h) r=1 else r=0 } \ apply vcg_cs apply (auto simp: ran_sorted_def) by (smt atLeastLessThan_iff) subsubsection \Find Equilibrium Index\ definition "is_equil a l h i \ l\i \ i (\j=l..j=i..l\h\ ensures \is_equil a l h i \ i=h \ (\i. is_equil a l h i)\ defines \ usum=0; i=l; while (ih-i\ @invariant \l\i \ i\h \ usum = (\j=l.. { usum = usum + a[i]; i=i+1 }; i=l; lsum=0; while (usum \ lsum \ ih-i\ @invariant \l\i \ i\h \ lsum=(\j=l.. usum=(\j=i.. (\jis_equil a l h j) \ { lsum = lsum + a[i]; usum = usum - a[i]; i=i+1 } \ apply vcg_cs apply (auto simp: intvs_incdec is_equil_def) apply (metis atLeastLessThan_iff eq_iff finite_atLeastLessThan_int sum_diff1) apply force by force subsubsection \Rotate Right\ program_spec rotate_right assumes "0i\{0..0 ((i-1) mod n)" defines \ i = 0; prev = a[n - 1]; while (i < n) @invariant \ 0 \ i \ i \ n \ (\j\{0..0 ((j-1) mod n)) \ (\j\{i..0 j) \ prev = a\<^sub>0 ((i-1) mod n) \ @variant \n - i\ { temp = a[i]; a[i] = prev; prev = temp; i = i + 1 } \ apply vcg_cs by (simp add: zmod_minus1) subsubsection \Binary Search, Leftmost Element\ text \We first specify the pre- and postcondition\ definition "bin_search_pre a l h \ l\h \ ran_sorted a l h" definition "bin_search_post a l h x i \ l\i \ i\h \ (\i\{l.. (\i\{i.. a i)" text \Then we prove that the program is correct\ program_spec binsearch assumes \bin_search_pre a l h\ ensures \bin_search_post a\<^sub>0 l\<^sub>0 h\<^sub>0 x\<^sub>0 l\ defines \ while (l < h) @variant \h-l\ @invariant \l\<^sub>0\l \ l\h \ h\h\<^sub>0 \ (\i\{l\<^sub>0.. (\i\{h..0}. x \ a i)\ { m = (l + h) / 2; if (a[m] < x) l = m + 1 else h = m } \ apply vcg_cs apply (auto simp: algebra_simps bin_search_pre_def bin_search_post_def) txt \Driving sledgehammer to its limits ...\ apply (smt div_add_self1 even_succ_div_two odd_two_times_div_two_succ ran_sorted_alt) by (smt div_add_self1 even_succ_div_two odd_two_times_div_two_succ ran_sorted_alt) text \Next, we show that our postcondition (which was easy to prove) implies the expected properties of the algorithm. \ lemma assumes "bin_search_pre a l h" "bin_search_post a l h x i" shows bin_search_decide_membership: "x\a`{l.. (i x = a i)" and bin_search_leftmost: "x\a`{l..Naive String Search\ (* By Simon Wimmer *) program_spec match_string assumes "l1 \ h1" ensures "(\j \ {0.. (i < h1 - l1 \ a (l + i) \ b (l1 + i)) \ 0 \ i \ i \ h1 - l1" defines \ i = 0; while (l1 + i < h1 \ a[l + i] == b[l1 + i]) @invariant \(\j \ {0.. 0 \ i \ i \ h1 - l1\ @variant \(h1 - (l1 + i))\ { i = i + 1 } \ by vcg_cs auto lemma lran_eq_iff': "lran a l1 (l1 + (h - l)) = lran a' l h \ (\i. 0 \ i \ i < h - l \ a (l1 + i) = a' (l + i))" if "l \ h" using that proof (induction "nat (h - l)" arbitrary: h) case 0 then show ?case by auto next case (Suc x) then have *: "x = nat (h - 1 - l)" "l \ h - 1" by auto note IH = Suc.hyps(1)[OF *] from * have 1: "lran a l1 (l1 + (h - l)) = lran a l1 (l1 + (h - 1 - l)) @ [a (l1 + (h - 1 - l))]" "lran a' l h = lran a' l (h - 1) @ [a' (h - 1)]" by (auto simp: lran_bwd_simp algebra_simps lran_butlast[symmetric]) from IH * Suc.hyps(2) Suc.prems show ?case unfolding 1 apply auto subgoal for i by (cases "i = h - 1 - l") auto done qed program_spec match_string' assumes "l1 \ h1" ensures "i = h1 - l1 \ lran a l (l + (h1 - l1)) = lran b l1 h1" for i h1 l1 l a[] b[] defines \inline match_string\ by vcg_cs (auto simp: lran_eq_iff') program_spec substring assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..0}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; while (l < h \ match == 0) @invariant\l\<^sub>0 \ l \ l \ h \ match \ {0,1} \ (if match = 1 then lran a l (l + (h1 - l1)) = lran b l1 h1 \ l < h else (\j \ {l\<^sub>0.. lran b l1 h1))\ @variant\(h - l) * (1 - match)\ { inline match_string'; if (i == h1 - l1) {match = 1} else {l = l + 1} } \ by vcg_cs auto program_spec substring' assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..h\<^sub>0-(h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; if (l + (h1 - l1) \ h) { h = h - (h1 - l1) + 1; inline substring } \ by vcg_cs auto (* Or combined: *) program_spec substring'' assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..0-(h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; if (l + (h1 - l1) \ h) { while (l + (h1 - l1) < h \ match == 0) @invariant\l\<^sub>0 \ l \ l \ h - (h1 - l1) \ match \ {0,1} \ (if match = 1 then lran a l (l + (h1 - l1)) = lran b l1 h1 \ l < h - (h1 - l1) else (\j \ {l\<^sub>0.. lran b l1 h1))\ @variant\(h - l) * (1 - match)\ { inline match_string'; if (i == h1 - l1) {match = 1} else {l = l + 1} } } \ by vcg_cs auto lemma lran_split: "lran a l h = lran a l p @ lran a p h" if "l \ p" "p \ h" using that by (induction p; simp; simp add: lran.simps) lemma lran_eq_append_iff: "lran a l h = as @ bs \ (\ i. l \ i \ i \ h \ as = lran a l i \ bs = lran a i h)" if "l \ h" apply safe subgoal using that proof (induction as arbitrary: l) case Nil then show ?case by auto next case (Cons x as) from this(2-) have "lran a (l + 1) h = as @ bs" "a l = x" "l + 1 \ h" apply - subgoal by simp subgoal by (smt append_Cons list.inject lran.simps lran_append1) subgoal using add1_zle_eq by fastforce done from Cons.IH[OF this(1,3)] guess i by safe note IH = this with \a l = x\ show ?case apply (intro exI[where x = i]) apply auto by (smt IH(3) lran_prepend1) qed apply (subst lran_split; simp) done lemma lran_split': "(\j\{l..h - (h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1) = (\as bs. lran a l h = as @ lran b l1 h1 @ bs)" if "l \ h" "l1 \ h1" proof safe fix j assume j: "j \ {l..h - (h1 - l1)}" and match: "lran a j (j + (h1 - l1)) = lran b l1 h1" with \l1 \ h1\ have "lran a l h = lran a l j @ lran a j (j + (h1 - l1)) @ lran a (j + (h1 - l1)) h" apply (subst lran_split[where p = j], simp, simp) apply (subst (2) lran_split[where p = "j + (h1 - l1)"]; simp) done then show "\as bs. lran a l h = as @ lran b l1 h1 @ bs" by (auto simp: match) next fix as bs assume "lran a l h = as @ lran b l1 h1 @ bs" with that lran_eq_append_iff[of l h a as "lran b l1 h1 @ bs"] obtain i where "as = lran a l i" "lran a i h = lran b l1 h1 @ bs" "l \ i" "i \ h" by auto with lran_eq_append_iff[of i h a "lran b l1 h1" bs] obtain j where j: "lran b l1 h1 = lran a i j" "bs = lran a j h" "i \ j" "j \ h" by auto moreover have "j = i + (h1 - l1)" proof - have "length (lran b l1 h1) = nat (h1 - l1)" "length (lran a i j) = nat (j - i)" by auto with j(1,3) that show ?thesis by auto qed ultimately show "\j\{l..h - (h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1" using \l \ i\ by auto qed program_spec substring_final assumes "l \ h \ 0 \ len" ensures "match = 1 \ (\as bs. lran a l\<^sub>0 h\<^sub>0 = as @ lran b 0 len @ bs)" for l h len match a[] b[] defines \l1 = 0; h1 = len; inline substring'\ supply [simp] = lran_split'[symmetric] apply vcg_cs done subsubsection \Insertion Sort\ text \ We first prove the inner loop. The specification here specifies what the algorithm does as closely as possible, such that it becomes easier to prove. In this case, sortedness is not a precondition for the inner loop to move the key element backwards over all greater elements. \ definition "insort_insert_post l j a\<^sub>0 a i \ (let key = a\<^sub>0 j in i\{l-1.. \\i\ is in range\ \ \Content of new array\ \ (\k\{l..i}. a k = a\<^sub>0 k) \ a (i+1) = key \ (\k\{i+2..j}. a k = a\<^sub>0 (k-1)) \ a = a\<^sub>0 on -{l..j} \ \Placement of \key\\ \ (i\l \ a i \ key) \ \Element at \i\ smaller than \key\, if it exists \ \ (\k\{i+2..j}. a k > key) \ \Elements \\i+2\ greater than \key\\ ) " for l j i :: int and a\<^sub>0 a :: "int \ int" program_spec insort_insert assumes "l0 a i" defines \ key = a[j]; i = j-1; while (i\l \ a[i]>key) @variant \i-l+1\ @invariant \l-1\i \ i (\k\{l..i}. a k = a\<^sub>0 k) \ (\k\{i+2..j}. a k > key \ a k = a\<^sub>0 (k-1)) \ a = a\<^sub>0 on -{l..j} \ { a[i+1] = a[i]; i=i-1 }; a[i+1] = key \ unfolding insort_insert_post_def Let_def apply vcg apply auto by (smt atLeastAtMost_iff) text \Next, we show that our specification that was easy to prove implies the specification that the outer loop expects:\ text \Invoking \insort_insert\ will sort in the element\ lemma insort_insert_sorted: assumes "lInvoking \insort_insert\ will only mutate the elements\ lemma insort_insert_ran1: assumes "li+1" "i+1\j" unfolding insort_insert_post_def Let_def by auto have ranshift: "mset_ran (a o (+)(-1)) {i+2..j} = mset_ran a {i+1..j-1}" by (simp add: mset_ran_shift algebra_simps) have "mset_ran a' {l..j} = mset_ran a' {l..i} + {# a' (i+1) #} + mset_ran a' {i+2..j}" using \l \l\i+1\ \i+1\j\ apply (simp add: mset_ran_combine) by (auto intro: arg_cong[where f="mset_ran a'"]) also have "\ = mset_ran a {l..i} + {# a j #} + mset_ran (a o (+)(-1)) {i+2..j}" using EQS(1,3)[THEN mset_ran_cong] EQS(2) by simp also have "\ = mset_ran a {l..i} + mset_ran a {i+1..j-1} + {# a j #}" by (simp add: mset_ran_shift algebra_simps) also have "\ = mset_ran a {l..j}" using \l \l\i+1\ \i+1\j\ apply (simp add: mset_ran_combine) by (auto intro: arg_cong[where f="mset_ran a"]) finally show ?thesis . qed text \The property @{thm insort_insert_ran1} extends to the whole array to be sorted\ lemma insort_insert_ran2: assumes "linsort_insert_post l j a a' i\ have "a' = a on {j<..Finally, we specify and prove correct the outer loop\ program_spec insort assumes "l mset_ran a {l..0 {l.. a=a\<^sub>0 on -{l.. j = l + 1; while (jh-j\ @invariant \ l j\h \ \\j\ in range\ \ ran_sorted a l j \ \Array is sorted up to \j\\ \ mset_ran a {l..0 {l.. \Elements in range only permuted\ \ a=a\<^sub>0 on -{l.. { inline insort_insert; j=j+1 } \ apply vcg_cs apply (intro conjI) subgoal by (rule insort_insert_sorted) subgoal using insort_insert_ran2(1) by auto subgoal apply (frule (2) insort_insert_ran2(2)) by (auto simp: eq_on_def) done subsubsection \Quicksort\ procedure_spec partition_aux(a,l,h,p) returns (a,i) assumes "l\h" ensures "mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\j\{l\<^sub>0..0) \ (\j\{i..0}. a j \ p\<^sub>0) \ l\<^sub>0\i \ i\h\<^sub>0 \ a\<^sub>0 = a on -{l\<^sub>0..0} " defines \ i=l; j=l; while (j l\i \ i\j \ j\h \ mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\k\{l.. (\k\{i.. p) \ (\k\{j..0 k = a k) \ a\<^sub>0 = a on -{l\<^sub>0..0} \ @variant \(h-j)\ { if (a[j] supply lran_eq_iff[simp] lran_tail[simp del] apply vcg_cs subgoal by (simp add: mset_ran_swap[unfolded swap_def]) subgoal by auto done procedure_spec partition(a,l,h,p) returns (a,i) assumes "l0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\j\{l\<^sub>0.. (\j\{i..0}. a j \ a i) \ l\<^sub>0\i \ i0 \ a\<^sub>0 (h\<^sub>0-1) = a i \ a\<^sub>0 = a on -{l\<^sub>0..0} " defines \ p = a[h-1]; (a,i) = partition_aux(a,l,h-1,p); a[h-1] = a[i]; a[i] = p \ apply vcg_cs apply (auto simp: eq_on_def mset_ran_swap[unfolded swap_def] intvs_incdec intro: mset_ran_combine_eq_diff) done lemma quicksort_sorted_aux: assumes BOUNDS: "l \ i" "i < h" assumes LESS: "\j\{l..1 j < a\<^sub>1 i" assumes GEQ: "\j\{i..1 i \ a\<^sub>1 j" assumes R1: "mset_ran a\<^sub>1 {l..2 {l..1 = a\<^sub>2 on - {l..2 l i" assumes R2: "mset_ran a\<^sub>2 {i + 1..3 {i + 1..2 = a\<^sub>3 on - {i + 1..3 (i + 1) h" shows "ran_sorted a\<^sub>3 l h" proof - have [simp]: "{l.. - {i + 1..1 i = a\<^sub>3 i" using E1 E2 by (auto simp: eq_on_def) note X1 = mset_ran_xfer_pointwise[where P = \\x. x < p\ for p, OF R1, simplified] note X2 = eq_on_xfer_pointwise[where P = \\x. x < p\ for p, OF E2, of "{l..j\{l..3 j < a\<^sub>3 i" by (simp add: X1 X2) from GEQ have GEQ1: "\j\{i+1..1 i \ a\<^sub>1 j" by auto have [simp]: "{i + 1.. - {l..\x. x \ p\ for p, OF E1, of "{i+1..\x. x \ p\ for p, OF R2, simplified] from GEQ1 have GEQ': "\j\{i+1..3 i \ a\<^sub>3 j" by (simp add: X3 X4) from SL eq_on_xfer_ran_sorted[OF E2, of l i] have SL': "ran_sorted a\<^sub>3 l i" by simp show ?thesis using combine_sorted_pivot[OF BOUNDS SL' SH LESS' GEQ'] . qed lemma quicksort_mset_aux: assumes B: "l\<^sub>0\i" "i0" assumes R1: "mset_ran a {l\<^sub>0..0..0..0} = mset_ran ab {i + 1..0}" assumes E2: "aa = ab on - {i + 1..0}" shows "mset_ran a {l\<^sub>0..0} = mset_ran ab {l\<^sub>0..0}" apply (rule trans) apply (rule mset_ran_eq_extend[OF R1 E1]) using B apply auto [2] apply (rule mset_ran_eq_extend[OF R2 E2]) using B apply auto [2] done recursive_spec quicksort(a,l,h) returns a assumes "True" ensures "ran_sorted a l\<^sub>0 h\<^sub>0 \ mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ a\<^sub>0=a on -{l\<^sub>0..0}" variant "h-l" defines \ if (l apply (vcg_cs; (intro conjI)?) subgoal using quicksort_sorted_aux by metis subgoal using quicksort_mset_aux by metis subgoal by (smt ComplD ComplI atLeastLessThan_iff eq_on_def) subgoal by (auto simp: ran_sorted_def) done subsection \Data Refinement\ subsubsection \Filtering\ program_spec array_filter_negative assumes "l\h" ensures "lran a l\<^sub>0 i = filter (\x. x\0) (lran a\<^sub>0 l\<^sub>0 h\<^sub>0)" defines \ i=l; j=l; while (j l\i \ i\j \ j\h \ lran a l i = filter (\x. x\0) (lran a\<^sub>0 l j) \ lran a j h = lran a\<^sub>0 j h \ @variant \h-j\ { if (a[j]\0) {a[i] = a[j]; i=i+1}; j=j+1 } \ supply lran_eq_iff[simp] lran_tail[simp del] apply vcg_cs done (* Exercise: Use swap to modify this algorithm to partitioning! *) subsubsection \Merge Two Sorted Lists\ text \ We define the merge function abstractly first, as a functional program on lists. \ fun merge where "merge [] ys = ys" | "merge xs [] = xs" | "merge (x#xs) (y#ys) = (if xIt's straightforward to show that this produces a sorted list with the same elements.\ lemma merge_sorted: assumes "sorted xs" "sorted ys" shows "sorted (merge xs ys) \ set (merge xs ys) = set xs \ set ys" using assms apply (induction xs ys rule: merge.induct) apply auto done lemma merge_mset: "mset (merge xs ys) = mset xs + mset ys" by (induction xs ys rule: merge.induct) auto text \Next, we prove an equation that characterizes one step of the while loop, on the list level.\ lemma merge_eq: "xs\[] \ ys\[] \ merge xs ys = ( if ys=[] \ (xs\[] \ hd xs < hd ys) then hd xs # merge (tl xs) ys else hd ys # merge xs (tl ys) )" by (cases xs; cases ys; simp) text \ We do a first proof that our merge implementation on the arrays and indexes behaves like the functional merge on the corresponding lists. The annotations use the @{const lran} function to map from the implementation level to the list level. Moreover, the invariant of the implementation, \l\h\, is carried through explicitly. \ program_spec merge_imp' assumes "l1\h1 \ l2\h2" ensures "let ms = lran m 0 j; xs\<^sub>0 = lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0; ys\<^sub>0 = lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0 in j\0 \ ms = merge xs\<^sub>0 ys\<^sub>0" defines \ j=0; while (l1\h1 \ l2\h2) @variant \h1 + h2 - l1 - l2\ @invariant \let xs=lran a1 l1 h1; ys = lran a2 l2 h2; ms = lran m 0 j; xs\<^sub>0 = lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0; ys\<^sub>0 = lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0 in l1\h1 \ l2\h2 \ 0\j \ merge xs\<^sub>0 ys\<^sub>0 = ms@merge xs ys \ { if (l2==h2 \ (l1\h1 \ a1[l1] < a2[l2])) { m[j] = a1[l1]; l1=l1+1 } else { m[j] = a2[l2]; l2=l2+1 }; j=j+1 } \ text \Given the @{thm [source] merge_eq} theorem, which captures the essence of a loop step, and the theorems @{thm lran_append1}, @{thm lran_tail}, and @{thm hd_lran}, which convert from the operations on arrays and indexes to operations on lists, the proof is straightforward \ apply vcg_cs subgoal apply (subst merge_eq) by auto subgoal by linarith subgoal apply (subst merge_eq) by auto done text \ In a next step, we refine our proof to combine it with the abstract properties we have proved about merge. The program does not change (we simply inline the original one here). \ procedure_spec merge_imp (a1,l1,h1,a2,l2,h2) returns (m,j) assumes "l1\h1 \ l2\h2 \ sorted (lran a1 l1 h1) \ sorted (lran a2 l2 h2)" ensures "let ms = lran m 0 j in j\0 \ sorted ms \ mset ms = mset (lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0) + mset (lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0)" for l1 h1 l2 h2 a1[] a2[] m[] j defines \inline merge_imp'\ apply vcg_cs apply (auto simp: Let_def merge_mset dest: merge_sorted) done thm merge_imp_spec thm merge_imp_def lemma [named_ss vcg_bb]: "UNIV \ a = UNIV" "a \ UNIV = UNIV" by auto lemma merge_msets_aux: "\l\m; m\h\ \ mset (lran a l m) + mset (lran a m h) = mset (lran a l h)" by (auto simp: mset_lran mset_ran_combine ivl_disj_un_two) recursive_spec mergesort (a,l,h) returns (b,j) assumes "l\h" ensures \0\j \ sorted (lran b 0 j) \ mset (lran b 0 j) = mset (lran a\<^sub>0 l\<^sub>0 h\<^sub>0)\ variant \h-l\ for a[] b[] defines \ if (l==h) j=0 else if (l+1==h) { b[0] = a[l]; j=1 } else { m = (h+l) / 2; (a1,h1) = rec mergesort (a,l,m); (a2,h2) = rec mergesort (a,m,h); (b,j) = merge_imp (a1,0,h1,a2,0,h2) } \ apply vcg apply auto [] apply (auto simp: lran.simps) [] apply auto [] apply auto [] apply auto [] apply (auto simp: Let_def merge_msets_aux) [] done print_theorems subsubsection \Remove Duplicates from Array, using Bitvector Set\ text \We use an array to represent a set of integers. If we only insert elements in range \{0.., this representation is called bit-vector (storing a single bit per index is enough). \ definition set_of :: "(int \ int) \ int set" where "set_of a \ {i. a i \ 0}" context notes [simp] = set_of_def begin lemma set_of_empty[simp]: "set_of (\_. 0) = {}" by auto lemma set_of_insert[simp]: "x\0 \ set_of (a(i:=x)) = insert i (set_of a)" by auto lemma set_of_remove[simp]: "set_of (a(i:=0)) = set_of a - {i}" by auto lemma set_of_mem[simp]: "i\set_of a \ a i \ 0" by auto end program_spec dedup assumes \l\h\ ensures \set (lran a l i) = set (lran a\<^sub>0 l h) \ distinct (lran a l i)\ defines \ i=l; j=l; clear b[]; while (jh-j\ @invariant \l\i \ i\j \ j\h \ set (lran a l i) = set (lran a\<^sub>0 l j) \ distinct (lran a l i) \ lran a j h = lran a\<^sub>0 j h \ set_of b = set (lran a l i) \ { if (b[a[j]] == 0) { a[i] = a[j]; i=i+1; b[a[j]] = 1 }; j=j+1 } \ apply vcg_cs apply (auto simp: lran_eq_iff lran_upd_inside intro: arg_cong[where f=tl]) [] apply (auto simp: lran_eq_iff) [] done procedure_spec bv_init () returns b assumes True ensures \set_of b = {}\ defines \clear b[]\ by vcg_cs procedure_spec bv_insert (x, b) returns b assumes True ensures \set_of b = insert x\<^sub>0 (set_of b\<^sub>0)\ defines \b[x] = 1\ by vcg_cs procedure_spec bv_remove (x, b) returns b assumes True ensures \set_of b = set_of b\<^sub>0 - {x\<^sub>0}\ defines \b[x] = 0\ by vcg_cs procedure_spec bv_elem (x,b) returns r assumes True ensures \r\0 \ x\<^sub>0\set_of b\<^sub>0\ defines \r = b[x]\ by vcg_cs procedure_spec dedup' (a,l,h) returns (a,l,i) assumes \l\h\ ensures \set (lran a l i) = set (lran a\<^sub>0 l\<^sub>0 h\<^sub>0) \ distinct (lran a l i) \ for b[] defines \ b = bv_init(); i=l; j=l; while (jh-j\ @invariant \l\i \ i\j \ j\h \ set (lran a l i) = set (lran a\<^sub>0 l j) \ distinct (lran a l i) \ lran a j h = lran a\<^sub>0 j h \ set_of b = set (lran a l i) \ { mem = bv_elem (a[j],b); if (mem == 0) { a[i] = a[j]; i=i+1; b = bv_insert(a[j],b) }; j=j+1 } \ apply vcg_cs apply (auto simp: lran_eq_iff lran_upd_inside intro: arg_cong[where f=tl]) done subsection \Recursion\ subsubsection \Recursive Fibonacci\ recursive_spec fib_imp (i) returns r assumes True ensures \r = fib i\<^sub>0\ variant \i\ defines \ if (i\0) r=0 else if (i==1) r=1 else { r1=rec fib_imp (i-2); r2=rec fib_imp (i-1); r = r1+r2 } \ by vcg_cs subsubsection \Homeier's Cycling Termination\ text \A contrived example from Homeier's thesis. Only the termination proof is done.\ (* TODO: Also show correctness: pedal (a,b) will multiply its arguments! correctness proof may be a good test how to handle specs with logical vars! *) recursive_spec pedal (n,m) returns () assumes \n\0 \ m\0\ ensures True variant \n+m\ defines \ if (n\0 \ m\0) { G = G + m; if (n and coast (n,m) returns () assumes \n\0 \ m\0\ ensures True variant \n+m+1\ defines \ G = G + n; if (n by vcg_cs subsubsection \Ackermann\ fun ack :: "nat \ nat \ nat" where "ack 0 n = n+1" | "ack m 0 = ack (m-1) 1" | "ack m n = ack (m-1) (ack m (n-1))" lemma ack_add_simps[simp]: "m\0 \ ack m 0 = ack (m-1) 1" "\m\0; n\0\ \ ack m n = ack (m-1) (ack m (n-1))" subgoal by (cases m) auto subgoal by (cases "(m,n)" rule: ack.cases) (auto) done recursive_spec relation "less_than <*lex*> less_than" ack_imp (m,n) returns r assumes "m\0 \ n\0" ensures "r=int (ack (nat m\<^sub>0) (nat n\<^sub>0))" variant "(nat m, nat n)" defines \ if (m==0) r = n+1 else if (n==0) r = rec ack_imp (m-1,1) else { t = rec ack_imp (m,n-1); r = rec ack_imp (m-1,t) } \ supply nat_distribs[simp] by vcg_cs subsubsection \McCarthy's 91 Function\ text \A standard benchmark for verification of recursive functions. We use Homeier's version with a global variable.\ recursive_spec p91(y) assumes True ensures "if 1000 then G = y\<^sub>0-10 else G = 91" variant "101-y" for G defines \ if (100 apply vcg_cs apply (auto split: if_splits) done subsubsection \Odd/Even\ recursive_spec odd_imp (a) returns b assumes True ensures "b\0 \ odd a\<^sub>0" variant "\a\" defines \ if (a==0) b=0 else if (a<0) b = rec even_imp (a+1) else b = rec even_imp (a-1) \ and even_imp (a) returns b assumes True ensures "b\0 \ even a\<^sub>0" variant "\a\" defines \ if (a==0) b=1 else if (a<0) b = rec odd_imp (a+1) else b = rec odd_imp (a-1) \ apply vcg apply auto done thm even_imp_spec subsubsection \Pandya and Joseph's Product Producers\ text \Again, taking the version from Homeier's thesis, but with a modification to also terminate for negative \y\. \ recursive_spec relation \measure nat <*lex*> less_than\ product () assumes True ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,1::nat)" for GX GY GZ defines \ e = even_imp (GY); if (e\0) rec evenproduct() else rec oddproduct() \ and oddproduct() assumes \odd GY\ ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,0::nat)" for GX GY GZ defines \ if (GY<0) { GY = GY + 1; GZ = GZ - GX } else { GY = GY - 1; GZ = GZ + GX }; rec evenproduct() \ and evenproduct() assumes \even GY\ ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,0::nat)" for GX GY GZ defines \ if (GY\0) { GX = 2*GX; GY = GY / 2; rec product() } \ apply vcg_cs apply (auto simp: algebra_simps) apply presburger+ done (* TODO: * Infrastructure to prove different specification for same program (DONE) We can use inline, and just give program a new name. * ghost variables. Keep them existentially qualified, with naming tag. ABS_GHOST ''name'' (\x. P x) \ \x. P x GHOST ''name'' value = True assumes ABS_GHOST name (\x. P x) obtains x where "GHOST name x" "P x" assumes "GHOST name x" and "P x" shows "ABS_GHOST name (\x. P x)" Let_Ghost name (\x s. C x s) = skip assumes "\x. C x s" assumes "\x. \ GHOST name x; C x s \ \ Q s" shows "wp (Let_Ghost name (\x s. C x s)) Q s" Let specification command also abstract over ghost variables in program. *) (* IDEAS: Extend arrays to multiple dimensions: int list \ int Vidx vname "aexp list" AssignIdx vname "aexp list" aexp ArrayCpy vname vname -- copy all dimensions (it's simpler) ArrayInit vname -- Init all dimensions plain values on dimension []! then we can easily encode matrices and adjacency lists! *) subsection \Graph Algorithms\ subsubsection \DFS\ (* By Simon Wimmer *) text \A graph is stored as an array of integers. Each node is an index into this array, pointing to a size-prefixed list of successors. Example for node \i\, which has successors \s1... sn\: \<^verbatim>\ Indexes: ... | i | i+1 | ... | i+n | ... Data: ... | n | s1 | ... | sn | ... \ \ definition succs where "succs a i \ a ` {i+1.. int" definition Edges where "Edges a \ {(i, j). j \ succs a i}" procedure_spec push' (x, stack, ptr) returns (stack, ptr) assumes "ptr \ 0" ensures \lran stack 0 ptr = lran stack\<^sub>0 0 ptr\<^sub>0 @ [x\<^sub>0] \ ptr = ptr\<^sub>0 + 1\ defines \stack[ptr] = x; ptr = ptr + 1\ by vcg_cs procedure_spec push (x, stack, ptr) returns (stack, ptr) assumes "ptr \ 0" ensures \stack ` {0..0} \ stack\<^sub>0 ` {0..0} \ ptr = ptr\<^sub>0 + 1\ for stack[] defines \stack[ptr] = x; ptr = ptr + 1\ by vcg_cs (auto simp: fun_upd_image) program_spec get_succs assumes "j \ stop \ stop = a (j - 1) \ 0 \ i" ensures " stack ` {0..0 - 1, x) \ Edges a \ x \ set_of visited} \ stack\<^sub>0 ` {0..0} \ i \ i\<^sub>0" for i j stop stack[] a[] visited[] defines \ while (j < stop) @invariant \stack ` {0.. a ` {j\<^sub>0.. x \ set_of visited} \ stack\<^sub>0 ` {0..0} \ j \ stop \ i\<^sub>0 \ i \ j\<^sub>0 \ j \ @variant \(stop - j)\ { succ = a[j]; is_elem = bv_elem(succ,visited); if (is_elem == 0) { (stack, i) = push (succ, stack, i) }; j = j + 1 } \ by vcg_cs (auto simp: intvs_incr_h Edges_def succs_def) procedure_spec pop (stack, ptr) returns (x, ptr) assumes "ptr \ 1" ensures \stack\<^sub>0 ` {0..0} = stack\<^sub>0 ` {0.. {x} \ ptr\<^sub>0 = ptr + 1\ for stack[] defines \ptr = ptr - 1; x = stack[ptr]\ by vcg_cs (simp add: intvs_upper_decr) procedure_spec stack_init () returns i assumes True ensures \i = 0\ defines \i = 0\ by vcg_cs lemma Edges_empty: "Edges a `` {i} = {}" if "i + 1 \ a i" using that unfolding Edges_def succs_def by auto text \This is one of the main insights of the algorithm: if a set of visited states is closed w.r.t.\ to the edge relation, then it is guaranteed to contain all the states that are reachable from any state within the set.\ lemma reachability_invariant: assumes reachable: "(s, x) \ (Edges a)\<^sup>*" and closed: "\v\visited. Edges a `` {v} \ visited" and start: "s \ visited" shows "x \ visited" using reachable start closed by induction auto program_spec (partial) dfs assumes "0 \ x \ 0 \ s" ensures "b = 1 \ x \ (Edges a)\<^sup>* `` {s}" defines \ b = 0; clear stack[]; i = stack_init(); (stack, i) = push (s, stack, i); clear visited[]; while (b == 0 \ i \ 0) @invariant \0 \ i \ (s \ stack ` {0.. s \ set_of visited) \ (b = 0 \ b = 1) \ ( if b = 0 then stack ` {0.. (Edges a)\<^sup>* `` {s} \ (\v \ set_of visited. (Edges a) `` {v} \ set_of visited \ stack ` {0.. (x \ set_of visited) else x \ (Edges a)\<^sup>* `` {s}) \ { (next, i) = pop(stack, i); \\Take the top most element from the stack.\ visited = bv_insert(next, visited); \\Mark it as visited,\ if (next == x) { b = 1 \\If it is the target, we are done.\ } else { \\Else, put its successors on the stack if they are not yet visited.\ stop = a[next]; j = next + 1; if (j \ stop) { inline get_succs } } } \ apply vcg_cs subgoal by (auto simp: set_of_def) subgoal using intvs_lower_incr by (auto simp: Edges_empty) subgoal by auto (fastforce simp: set_of_def dest!: reachability_invariant) done text \Assuming that the input graph is finite, we can also prove that the algorithm terminates. We will thus use an \Isabelle context\ to fix a certain finite graph and a start state: \ context fixes start :: int and edges assumes finite_graph[intro!]: "finite ((Edges edges)\<^sup>* `` {start})" begin lemma sub_insert_same_iff: "s \ insert x s \ x\s" by auto program_spec dfs1 assumes "0 \ x \ 0 \ s \ start = s \ edges = a" ensures "b = 1 \ x \ (Edges a)\<^sup>* `` {s}" for visited[] defines \ b = 0; \\\i\ will point to the next free space in the stack (i.e. it is the size of the stack)\ i = 1; \\Initially, we put \s\ on the stack.\ stack[0] = s; visited = bv_init(); while (b == 0 \ i \ 0) @invariant \ 0 \ i \ (s \ stack ` {0.. s \ set_of visited) \ (b = 0 \ b = 1) \ set_of visited \ (Edges edges)\<^sup>* `` {start} \ ( if b = 0 then stack ` {0.. (Edges a)\<^sup>* `` {s} \ (\v \ set_of visited. (Edges a) `` {v} \ set_of visited \ stack ` {0.. (x \ set_of visited) else x \ (Edges a)\<^sup>* `` {s}) \ @relation \finite_psupset ((Edges edges)\<^sup>* `` {start}) <*lex*> less_than\ @variant \ (set_of visited, nat i)\ { \\Take the top most element from the stack.\ (next, i) = pop(stack, i); if (next == x) { \\If it is the target, we are done.\ visited = bv_insert(next, visited); b = 1 } else { is_elem = bv_elem(next, visited); if (is_elem == 0) { visited = bv_insert(next, visited); \\Else, put its successors on the stack if they are not yet visited.\ stop = a[next]; j = next + 1; if (j \ stop) { inline get_succs } } } } \ apply vcg_cs subgoal by auto subgoal by (auto simp add: image_constant_conv) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: set_of_def) - subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) + subgoal apply (clarsimp simp: finite_psupset_def sub_insert_same_iff) + by (metis insertI1 set_of_mem) subgoal by (auto simp: Edges_empty) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: set_of_def) subgoal by auto (fastforce simp: set_of_def dest!: reachability_invariant) done end end \ No newline at end of file diff --git a/thys/Simpl/Hoare.thy b/thys/Simpl/Hoare.thy --- a/thys/Simpl/Hoare.thy +++ b/thys/Simpl/Hoare.thy @@ -1,424 +1,423 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de License: LGPL *) (* Title: Hoare.thy Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer Some rights reserved, TU Muenchen This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) section \Auxiliary Definitions/Lemmas to Facilitate Hoare Logic\ theory Hoare imports HoarePartial HoareTotal begin syntax "_hoarep_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set, 'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/\ (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\\<^bsub>'/_\<^esub> (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\ (_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoarep_noAbr":: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoarep_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\ (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\ (_/ (_)/ _))" [61,1000,20,1000]60) "_hoaret_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/\\<^sub>t (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\\<^sub>t (_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoaret_noAbr":: "[('s,'p,'f) body,'f set, ('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoaret_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\\<^sub>t (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\\<^sub>t (_/ (_)/ _))" [61,1000,20,1000]60) syntax (ASCII) "_hoarep_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn,'s assn] \ bool" ("(3_,_/|- (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-'/_ (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-(_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoarep_noAbr":: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-'/_ (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoarep_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-(_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-'/_ (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-(_/ (_)/ _))" [61,1000,20,1000]60) "_hoaret_emptyFault":: "[('s,'p,'f) body,('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/|-t (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-t'/_ (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-t(_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoaret_noAbr":: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-t'/_ (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoaret_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-t(_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-t'/_ (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-t(_/ (_)/ _))" [61,1000,20,1000]60) translations "\\ P c Q,A" == "\\\<^bsub>/{}\<^esub> P c Q,A" "\\\<^bsub>/F\<^esub> P c Q,A" == "\,{}\\<^bsub>/F\<^esub> P c Q,A" "\,\\ P c Q" == "\,\\\<^bsub>/{}\<^esub> P c Q" "\,\\\<^bsub>/F\<^esub> P c Q" == "\,\\\<^bsub>/F\<^esub> P c Q,{}" "\,\\ P c Q,A" == "\,\\\<^bsub>/{}\<^esub> P c Q,A" "\\ P c Q" == "\\\<^bsub>/{}\<^esub> P c Q" "\\\<^bsub>/F\<^esub> P c Q" == "\,{}\\<^bsub>/F\<^esub> P c Q" "\\\<^bsub>/F\<^esub> P c Q" <= "\\\<^bsub>/F\<^esub> P c Q,{}" "\\ P c Q" <= "\\ P c Q,{}" "\\\<^sub>t P c Q,A" == "\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" == "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" "\,\\\<^sub>t P c Q" == "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q" "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q" == "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,{}" "\,\\\<^sub>t P c Q,A" == "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" "\\\<^sub>t P c Q" == "\\\<^sub>t\<^bsub>/{}\<^esub> P c Q" "\\\<^sub>t\<^bsub>/F\<^esub> P c Q" == "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q" "\\\<^sub>t\<^bsub>/F\<^esub> P c Q" <= "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,{}" "\\\<^sub>t P c Q" <= "\\\<^sub>t P c Q,{}" term "\\ P c Q" term "\\ P c Q,A" term "\\\<^bsub>/F\<^esub> P c Q" term "\\\<^bsub>/F\<^esub> P c Q,A" term "\,\\ P c Q" term "\,\\\<^bsub>/F\<^esub> P c Q" term "\,\\ P c Q,A" term "\,\\\<^bsub>/F\<^esub> P c Q,A" term "\\\<^sub>t P c Q" term "\\\<^sub>t P c Q,A" term "\\\<^sub>t\<^bsub>/F\<^esub> P c Q" term "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" term "\,\\ P c Q" term "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q" term "\,\\ P c Q,A" term "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" locale hoare = fixes \::"('s,'p,'f) body" primrec assoc:: "('a \'b) list \ 'a \ 'b " where "assoc [] x = undefined" | "assoc (p#ps) x = (if fst p = x then (snd p) else assoc ps x)" lemma conjE_simp: "(P \ Q \ PROP R) \ (P \ Q \ PROP R)" by rule simp_all lemma CollectInt_iff: "{s. P s} \ {s. Q s} = {s. P s \ Q s}" by auto lemma Compl_Collect:"-(Collect b) = {x. \(b x)}" by fastforce lemma Collect_False: "{s. False} = {}" by simp lemma Collect_True: "{s. True} = UNIV" by simp lemma triv_All_eq: "\x. P \ P" by simp lemma triv_Ex_eq: "\x. P \ P" by simp lemma Ex_True: "\b. b" by blast lemma Ex_False: "\b. \b" by blast definition mex::"('a \ bool) \ bool" where "mex P = Ex P" definition meq::"'a \ 'a \ bool" where "meq s Z = (s = Z)" lemma subset_unI1: "A \ B \ A \ B \ C" by blast lemma subset_unI2: "A \ C \ A \ B \ C" by blast lemma split_paired_UN: "(\p. (P p)) = (\a b. (P (a,b)))" by auto lemma in_insert_hd: "f \ insert f X" by simp lemma lookup_Some_in_dom: "\ p = Some bdy \ p \ dom \" by auto lemma unit_object: "(\u::unit. P u) = P ()" by auto lemma unit_ex: "(\u::unit. P u) = P ()" by auto lemma unit_meta: "(\(u::unit). PROP P u) \ PROP P ()" by auto lemma unit_UN: "(\z::unit. P z) = P ()" by auto lemma subset_singleton_insert1: "y = x \ {y} \ insert x A" by auto lemma subset_singleton_insert2: "{y} \ A \ {y} \ insert x A" by auto lemma in_Specs_simp: "(\x\\Z. {(P Z, p, Q Z, A Z)}. Prop x) = (\Z. Prop (P Z,p,Q Z,A Z))" by auto lemma in_set_Un_simp: "(\x\A \ B. P x) = ((\x \ A. P x) \ (\x \ B. P x))" by auto lemma split_all_conj: "(\x. P x \ Q x) = ((\x. P x) \ (\x. Q x))" by blast lemma image_Un_single_simp: "f ` (\Z. {P Z}) = (\Z. {f (P Z)}) " by auto lemma measure_lex_prod_def': "f <*mlex*> r \ ({(x,y). (x,y) \ measure f \ f x=f y \ (x,y) \ r})" by (auto simp add: mlex_prod_def inv_image_def) lemma in_measure_iff: "(x,y) \ measure f = (f x < f y)" by (simp add: measure_def inv_image_def) lemma in_lex_iff: - "((a,b),(x,y)) \ r <*lex*> s = ((a,x) \ r \ (a=x \ (b,y)\s))" + "((a,b),(x,y)) \ r <*lex*> s = (a\x \ (a,x) \ r \ (a=x \ (b,y)\s))" by (simp add: lex_prod_def) lemma in_mlex_iff: "(x,y) \ f <*mlex*> r = (f x < f y \ (f x=f y \ (x,y) \ r))" by (simp add: measure_lex_prod_def' in_measure_iff) lemma in_inv_image_iff: "(x,y) \ inv_image r f = ((f x, f y) \ r)" by (simp add: inv_image_def) text \This is actually the same as @{thm [source] wf_mlex}. However, this basic proof took me so long that I'm not willing to delete it. \ lemma wf_measure_lex_prod [simp,intro]: assumes wf_r: "wf r" shows "wf (f <*mlex*> r)" proof (rule ccontr) assume " \ wf (f <*mlex*> r)" then obtain g where "\i. (g (Suc i), g i) \ f <*mlex*> r" by (auto simp add: wf_iff_no_infinite_down_chain) hence g: "\i. (g (Suc i), g i) \ measure f \ f (g (Suc i)) = f (g i) \ (g (Suc i), g i) \ r" by (simp add: measure_lex_prod_def') hence le_g: "\i. f (g (Suc i)) \ f (g i)" - by (auto simp add: in_measure_iff order_le_less) + by (auto simp add: order_le_less) have "wf (measure f)" by simp hence " \Q. (\x. x \ Q) \ (\z\Q. \y. (y, z) \ measure f \ y \ Q)" by (simp add: wf_eq_minimal) from this [rule_format, of "g ` UNIV"] have "\z. z \ range g \ (\y. (y, z) \ measure f \ y \ range g)" by auto then obtain z where - z: "z \ range g" and - min_z: "\y. f y < f z \ y \ range g" - by (auto simp add: in_measure_iff) + z: "z \ range g" and min_z: "\y. f y < f z \ y \ range g" + by auto from z obtain k where k: "z = g k" by auto have "\i. k \ i \ f (g i) = f (g k)" proof (intro allI impI) fix i assume "k \ i" then show "f (g i) = f (g k)" proof (induct i) case 0 have "k \ 0" by fact hence "k = 0" by simp thus "f (g 0) = f (g k)" by simp next case (Suc n) have k_Suc_n: "k \ Suc n" by fact then show "f (g (Suc n)) = f (g k)" proof (cases "k = Suc n") case True thus ?thesis by simp next case False with k_Suc_n have "k \ n" by simp with Suc.hyps have n_k: "f (g n) = f (g k)" by simp from le_g have le: "f (g (Suc n)) \ f (g n)" by simp show ?thesis proof (cases "f (g (Suc n)) = f (g n)") case True with n_k show ?thesis by simp next case False with le have "f (g (Suc n)) < f (g n)" by simp with n_k k have "f (g (Suc n)) < f z" by simp with min_z have "g (Suc n) \ range g" by blast hence False by simp thus ?thesis by simp qed qed qed qed with k [symmetric] have "\i. k \ i \ f (g i) = f z" by simp hence "\i. k \ i \ f (g (Suc i)) = f (g i)" by simp with g have "\i. k \ i \ (g (Suc i),(g i)) \ r" - by (auto simp add: in_measure_iff order_less_le ) + by (auto simp add: order_less_le) hence "\i. (g (Suc (i+k)),(g (i+k))) \ r" by simp then have "\f. \i. (f (Suc i), f i) \ r" by - (rule exI [where x="\i. g (i+k)"],simp) with wf_r show False by (simp add: wf_iff_no_infinite_down_chain) qed lemmas all_imp_to_ex = all_simps (5) (*"\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" Avoid introduction of existential quantification of states on negative position. *) lemma all_imp_eq_triv: "(\x. x = k \ Q) = Q" "(\x. k = x \ Q) = Q" by auto end diff --git a/thys/Simpl/UserGuide.thy b/thys/Simpl/UserGuide.thy --- a/thys/Simpl/UserGuide.thy +++ b/thys/Simpl/UserGuide.thy @@ -1,1584 +1,1583 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de License: LGPL *) (* Title: UserGuide.thy Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer Some rights reserved, TU Muenchen This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) section \User Guide \label{sec:UserGuide}\ (*<*) theory UserGuide imports HeapList Vcg "HOL-Statespace.StateSpaceSyntax" "HOL-Library.LaTeXsugar" begin (*>*) (*<*) syntax "_statespace_updates" :: "('a \ 'b) \ updbinds \ ('a \ 'b)" ("_\_\" [900,0] 900) (*>*) text \ We introduce the verification environment with a couple of examples that illustrate how to use the different bits and pieces to verify programs. \ subsection \Basics\ text \ First of all we have to decide how to represent the state space. There are currently two implementations. One is based on records the other one on the concept called `statespace' that was introduced with Isabelle 2007 (see \texttt{HOL/Statespace}) . In contrast to records a 'satespace' does not define a new type, but provides a notion of state, based on locales. Logically the state is modelled as a function from (abstract) names to (abstract) values and the statespace infrastructure organises distinctness of names an projection/injection of concrete values into the abstract one. Towards the user the interface of records and statespaces is quite similar. However, statespaces offer more flexibility, inherited from the locale infrastructure, in particular multiple inheritance and renaming of components. In this user guide we prefer statespaces, but give some comments on the usage of records in Section \ref{sec:records}. \ hoarestate vars = A :: nat I :: nat M :: nat N :: nat R :: nat S :: nat text (in vars) \The command \isacommand{hoarestate} is a simple preprocessor for the command \isacommand{statespaces} which decorates the state components with the suffix \_'\, to avoid cluttering the namespace. Also note that underscores are printed as hyphens in this documentation. So what you see as @{term "A_'"} in this document is actually \texttt{A\_'}. Every component name becomes a fixed variable in the locale \vars\ and can no longer be used for logical variables. Lookup of a component @{term "A_'"} in a state @{term "s"} is written as @{term "s\A_'"}, and update with a value @{term "term v"} as @{term "s\A_' := v\"}. To deal with local and global variables in the context of procedures the program state is organised as a record containing the two componets @{const "locals"} and @{const "globals"}. The variables defined in hoarestate \vars\ reside in the @{const "locals"} part. \ text \ Here is a first example. \ lemma (in vars) "\\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply vcg txt \@{subgoals}\ apply simp txt \@{subgoals}\ done text \We enable the locale of statespace \vars\ by the \texttt{in vars} directive. The verification condition generator is invoked via the \vcg\ method and leaves us with the expected subgoal that can be proved by simplification.\ text (in vars) \ If we refer to components (variables) of the state-space of the program we always mark these with \\\ (in assertions and also in the program itself). It is the acute-symbol and is present on most keyboards. The assertions of the Hoare tuple are ordinary Isabelle sets. As we usually want to refer to the state space in the assertions, we provide special brackets for them. They can be written as {\verb+{| |}+} in ASCII or \\ \\ with symbols. Internally, marking variables has two effects. First of all we refer to the implicit state and secondary we get rid of the suffix \_'\. So the assertion @{term "{|\N = 5|}"} internally gets expanded to \{s. locals s \N_' = 5}\ written in ordinary set comprehension notation of Isabelle. It describes the set of states where the \N_'\ component is equal to \5\. An empty context and an empty postcondition for abrupt termination can be omitted. The lemma above is a shorthand for \\,{}\ \\N = 5\ \N :== 2 * \N \\N = 10\,{}\. \ text \We can step through verification condition generation by the method \vcg_step\. \ lemma (in vars) "\,{}\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply vcg_step txt \@{subgoals}\ txt \The last step of verification condition generation, transforms the inclusion of state sets to the corresponding predicate on components of the state space. \ apply vcg_step txt \@{subgoals}\ by simp text \ Although our assertions work semantically on the state space, stepping through verification condition generation ``feels'' like the expected syntactic substitutions of traditional Hoare logic. This is achieved by light simplification on the assertions calculated by the Hoare rules. \ lemma (in vars) "\\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply (rule HoarePartial.Basic) txt \@{subgoals}\ apply (simp only: mem_Collect_eq) txt \@{subgoals}\ apply (tactic \Hoare.BasicSimpTac @{context} Hoare.Function false [] (K all_tac) 1\) txt \@{subgoals}\ by simp text \The next example shows how we deal with the while loop. Note the invariant annotation. \ lemma (in vars) "\,{}\ \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b\ DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg txt \@{subgoals [display]}\ txt \The verification condition generator gives us three proof obligations, stemming from the path from the precondition to the invariant, from the invariant together with loop condition through the loop body to the invariant, and finally from the invariant together with the negated loop condition to the postcondition.\ apply auto done subsection \Procedures\ subsubsection \Declaration\ text \ Our first procedure is a simple square procedure. We provide the command \isacommand{procedures}, to declare and define a procedure. \ procedures Square (N::nat|R::nat) where I::nat in "\R :== \N * \N" text \A procedure is given by the signature of the procedure followed by the procedure body. The signature consists of the name of the procedure and a list of parameters together with their types. The parameters in front of the pipe \|\ are value parameters and behind the pipe are the result parameters. Value parameters model call by value semantics. The value of a result parameter at the end of the procedure is passed back to the caller. Local variables follow the \where\. If there are no local variables the \where \ in\ can be omitted. The variable @{term "I"} is actually unused in the body, but is used in the examples below.\ text \ The procedures command provides convenient syntax for procedure calls (that creates the proper @{term init}, @{term return} and @{term result} functions on the fly) and creates locales and statespaces to reason about the procedure. The purpose of locales is to set up logical contexts to support modular reasoning. Locales can be seen as freeze-dried proof contexts that get alive as you setup a new lemma or theorem (\cite{Ballarin-04-locales}). The locale the user deals with is named \Square_impl\. It defines the procedure name (internally @{term "Square_'proc"}), the procedure body (named \Square_body\) and the statespaces for parameters and local and global variables. Moreover it contains the assumption @{term "\ Square_'proc = Some Square_body"}, which states that the procedure is properly defined in the procedure context. The purpose of the locale is to give us easy means to setup the context in which we prove programs correct. In this locale the procedure context @{term "\"} is fixed. So we always use this letter for the procedure specification. This is crucial, if we prove programs under the assumption of some procedure specifications. \ (*<*) context Square_impl begin (*>*) text \The procedures command generates syntax, so that we can either write \CALL Square(\I,\R)\ or @{term "\I :== CALL Square(\R)"} for the procedure call. The internal term is the following: \ (*<*) declare [[hoare_use_call_tr' = false]] (*>*) text \\small @{term [display] "CALL Square(\I,\R)"}\ (*<*) declare [[hoare_use_call_tr' = true]] (*>*) text \Note the additional decoration (with the procedure name) of the parameter and local variable names.\ (*<*) end (*>*) text \The abstract syntax for the procedure call is @{term "call init p return result"}. The @{term "init"} function copies the values of the actual parameters to the formal parameters, the @{term return} function copies the global variables back (in our case there are no global variables), and the @{term "result"} function additionally copies the values of the formal result parameters to the actual locations. Actual value parameters can be all kind of expressions, since we only need their value. But result parameters must be proper ``lvalues'': variables (including dereferenced pointers) or array locations, since we have to assign values to them. \ subsubsection \Verification\ text (in Square_impl) \ A procedure specification is an ordinary Hoare tuple. We use the parameterless call for the specification; \\R :== PROC Square(\N)\ is syntactic sugar for \Call Square_'proc\. This emphasises that the specification describes the internal behaviour of the procedure, whereas parameter passing corresponds to the procedure call. The following precondition fixes the current value \\N\ to the logical variable @{term n}. Universal quantification of @{term "n"} enables us to adapt the specification to an actual parameter. The specification is used in the rule for procedure call when we come upon a call to @{term Square}. Thus @{term "n"} plays the role of the auxiliary variable @{term "Z"}. \ text \To verify the procedure we need to verify the body. We use a derived variant of the general recursion rule, tailored for non recursive procedures: @{thm [source] HoarePartial.ProcNoRec1}: \begin{center} @{thm [mode=Rule,mode=ParenStmt] HoarePartial.ProcNoRec1 [no_vars]} \end{center} The naming convention for the rule is the following: The \1\ expresses that we look at one procedure, and \NoRec\ that the procedure is non recursive. \ lemma (in Square_impl) shows "\n. \\\\N = n\ \R :== PROC Square(\N) \\R = n * n\" txt \The directive \in\ has the effect that the context of the locale @{term "Square_impl"} is included to the current lemma, and that the lemma is added as a fact to the locale, after it is proven. The next time locale @{term "Square_impl"} is invoked this lemma is immediately available as fact, which the verification condition generator can use. \ apply (hoare_rule HoarePartial.ProcNoRec1) txt "@{subgoals[display]}" txt \The method \hoare_rule\, like \rule\ applies a single rule, but additionally does some ``obvious'' steps: It solves the canonical side-conditions of various Hoare-rules and it automatically expands the procedure body: With @{thm [source] Square_impl}: @{thm [names_short] Square_impl [no_vars]} we get the procedure body out of the procedure context @{term "\"}; with @{thm [source] Square_body_def}: @{thm [names_short] Square_body_def [no_vars]} we can unfold the definition of the body. The proof is finished by the vcg and simp. \ txt "@{subgoals[display]}" by vcg simp text \If the procedure is non recursive and there is no specification given, the verification condition generator automatically expands the body.\ lemma (in Square_impl) Square_spec: shows "\n. \\\\N = n\ \R :== PROC Square(\N) \\R = n * n\" by vcg simp text \An important naming convention is to name the specification as \_spec\. The verification condition generator refers to this name in order to search for a specification in the theorem database. \ subsubsection \Usage\ text\Let us see how we can use procedure specifications.\ (* FIXME: maybe don't show this at all *) lemma (in Square_impl) shows "\\\\I = 2\ \R :== CALL Square(\I) \\R = 4\" txt \Remember that we have already proven @{thm [source] "Square_spec"} in the locale \Square_impl\. This is crucial for verification condition generation. When reaching a procedure call, it looks for the specification (by its name) and applies the rule @{thm [source,mode=ParenStmt] HoarePartial.ProcSpec} instantiated with the specification (as last premise). Before we apply the verification condition generator, let us take some time to think of what we can expect. Let's look at the specification @{thm [source] Square_spec} again: @{thm [display] Square_spec [no_vars]} The specification talks about the formal parameters @{term "N"} and @{term R}. The precondition @{term "\\N = n\"} just fixes the initial value of \N\. The actual parameters are @{term "I"} and @{term "R"}. We have to adapt the specification to this calling context. @{term "\n. \\ \\I = n\ \R :== CALL Square(\I) \\R = n * n\"}. From the postcondition @{term "\\R = n * n\"} we have to derive the actual postcondition @{term "\\R = 4\"}. So we gain something like: @{term "\n * n = (4::nat)\"}. The precondition is @{term "\\I = 2\"} and the specification tells us @{term "\\I = n\"} for the pre-state. So the value of @{term n} is the value of @{term I} in the pre-state. So we arrive at @{term "\\I = 2\ \ \\I * \I = 4\"}. \ apply vcg_step txt "@{subgoals[display]}" txt \ The second set looks slightly more involved: @{term "\\t. \<^bsup>t\<^esup>R = \I * \I \ \I * \I = 4\"}, this is an artefact from the procedure call rule. Originally \\I * \I = 4\ was \\<^bsup>t\<^esup>R = 4\. Where @{term "t"} denotes the final state of the procedure and the superscript notation allows to select a component from a particular state. \ apply vcg_step txt "@{subgoals[display]}" by simp text \ The adaption of the procedure specification to the actual calling context is done due to the @{term init}, @{term return} and @{term result} functions in the rule @{thm [source] HoarePartial.ProcSpec} (or in the variant @{thm [source] HoarePartial.ProcSpecNoAbrupt} which already incorporates the fact that the postcondition for abrupt termination is the empty set). For the readers interested in the internals, here a version without vcg. \ lemma (in Square_impl) shows "\\\\I = 2\ \R :== CALL Square(\I) \\R = 4\" apply (rule HoarePartial.ProcSpecNoAbrupt [OF _ _ Square_spec]) txt "@{subgoals[display]}" txt \This is the raw verification condition, It is interesting to see how the auxiliary variable @{term "Z"} is actually used. It is unified with @{term n} of the specification and fixes the state after parameter passing. \ apply simp txt "@{subgoals[display]}" prefer 2 apply vcg_step txt "@{subgoals[display]}" apply (auto intro: ext) done subsubsection \Recursion\ text \We want to define a procedure for the factorial. We first define a HOL function that calculates it, to specify the procedure later on. \ primrec fac:: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = (Suc n) * fac n" (*<*) lemma fac_simp [simp]: "0 < i \ fac i = i * fac (i - 1)" by (cases i) simp_all (*>*) text \Now we define the procedure.\ procedures Fac (N::nat | R::nat) "IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI" text \ Now let us prove that our implementation of @{term "Fac"} meets its specification. \ lemma (in Fac_impl) shows "\n. \\ \\N = n\ \R :== PROC Fac(\N) \\R = fac n\" apply (hoare_rule HoarePartial.ProcRec1) txt "@{subgoals[display]}" apply vcg txt "@{subgoals[display]}" apply simp done text \ Since the factorial is implemented recursively, the main ingredient of this proof is, to assume that the specification holds for the recursive call of @{term Fac} and prove the body correct. The assumption for recursive calls is added to the context by the rule @{thm [source] HoarePartial.ProcRec1} (also derived from the general rule for mutually recursive procedures): \begin{center} @{thm [mode=Rule,mode=ParenStmt] HoarePartial.ProcRec1 [no_vars]} \end{center} The verification condition generator infers the specification out of the context @{term "\"} when it encounters a recursive call of the factorial. \ subsection \Global Variables and Heap \label{sec:VcgHeap}\ text \ Now we define and verify some procedures on heap-lists. We consider list structures consisting of two fields, a content element @{term "cont"} and a reference to the next list element @{term "next"}. We model this by the following state space where every field has its own heap. \ hoarestate globals_heap = "next" :: "ref \ ref" cont :: "ref \ nat" text \It is mandatory to start the state name with `globals'. This is exploited by the syntax translations to store the components in the @{const globals} part of the state. \ text \Updates to global components inside a procedure are always propagated to the caller. This is implicitly done by the parameter passing syntax translations. \ text \We first define an append function on lists. It takes two references as parameters. It appends the list referred to by the first parameter with the list referred to by the second parameter. The statespace of the global variables has to be imported. \ procedures (imports globals_heap) append(p :: ref, q::ref | p::ref) "IF \p=Null THEN \p :== \q ELSE \p\\next :== CALL append(\p\\next,\q) FI" (*<*) context append_impl begin (*>*) text \ The difference of a global and a local variable is that global variables are automatically copied back to the procedure caller. We can study this effect on the translation of @{term "\p :== CALL append(\p,\q)"}: \ (*<*) declare [[hoare_use_call_tr' = false]] (*>*) text \ @{term [display] "\p :== CALL append(\p,\q)"} \ (*<*) declare [[hoare_use_call_tr' = true]] end (*>*) text \Below we give two specifications this time. One captures the functional behaviour and focuses on the entities that are potentially modified by the procedure, the second one is a pure frame condition. \ text \ The functional specification below introduces two logical variables besides the state space variable @{term "\"}, namely @{term "Ps"} and @{term "Qs"}. They are universally quantified and range over both the pre-and the postcondition, so that we are able to properly instantiate the specification during the proofs. The syntax \\\. \\\ is a shorthand to fix the current state: \{s. \ = s \}\. Moreover \\<^bsup>\\<^esup>x\ abbreviates the lookup of variable \x\ in the state \\\. The approach to specify procedures on lists basically follows \cite{MehtaN-CADE03}. From the pointer structure in the heap we (relationally) abstract to HOL lists of references. Then we can specify further properties on the level of HOL lists, rather then on the heap. The basic abstractions are: @{thm [display] Path.simps [no_vars]} @{term [show_types] "Path x h y ps"}: @{term ps} is a list of references that we can obtain out of the heap @{term h} by starting with the reference @{term x}, following the references in @{term h} up to the reference @{term y}. @{thm [display] List_def [no_vars]} A list @{term "List p h ps"} is a path starting in @{term p} and ending up in @{term Null}. \ lemma (in append_impl) append_spec1: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) txt \@{subgoals [margin=80,display]} Note that @{term "hoare_rule"} takes care of multiple auxiliary variables! @{thm [source] HoarePartial.ProcRec1} has only one auxiliary variable, namely @{term Z}. But the type of @{term Z} can be instantiated arbitrarily. So \hoare_rule\ instantiates @{term Z} with the tuple @{term "(\,Ps,Qs)"} and derives a proper variant of the rule. Therefore \hoare_rule\ depends on the proper quantification of auxiliary variables! \ apply vcg txt \@{subgoals [display]} For each branch of the \IF\ statement we have one conjunct to prove. The \THEN\ branch starts with \p = Null \ \\ and the \ELSE\ branch with \p \ Null \ \\. Let us focus on the \ELSE\ branch, were the recursive call to append occurs. First of all we have to prove that the precondition for the recursive call is fulfilled. That means we have to provide some witnesses for the lists @{term Psa} and @{term Qsa} which are referenced by \p\next\ (now written as @{term "next p"}) and @{term q}. Then we have to show that we can derive the overall postcondition from the postcondition of the recursive call. The state components that have changed by the recursive call are the ones with the suffix \a\, like \nexta\ and \pa\. \ apply fastforce done text \If the verification condition generator works on a procedure call it checks whether it can find a modifies clause in the context. If one is present the procedure call is simplified before the Hoare rule @{thm [source] HoarePartial.ProcSpec} is applied. Simplification of the procedure call means that the ``copy back'' of the global components is simplified. Only those components that occur in the modifies clause are actually copied back. This simplification is justified by the rule @{thm [source] HoarePartial.ProcModifyReturn}. So after this simplification all global components that do not appear in the modifies clause are treated as local variables.\ text \We study the effect of the modifies clause on the following examples, where we want to prove that @{term "append"} does not change the @{term "cont"} part of the heap. \ lemma (in append_impl) shows "\\ \\cont=c\ \p :== CALL append(Null,Null) \\cont=c\" proof - note append_spec = append_spec1 show ?thesis apply vcg txt \@{subgoals [display]}\ txt \Only focus on the very last line: @{term conta} is the heap component after the procedure call, and @{term cont} the heap component before the procedure call. Since we have not added the modified clause we do not know that they have to be equal. \ oops text \ We now add the frame condition. The list in the modifies clause names all global state components that may be changed by the procedure. Note that we know from the modifies clause that the @{term cont} parts are not changed. Also a small side note on the syntax. We use ordinary brackets in the postcondition of the modifies clause, and also the state components do not carry the acute, because we explicitly note the state @{term t} here. \ lemma (in append_impl) append_modifies: shows "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC append(\p,\q) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done text \We tell the verification condition generator to use only the modifies clauses and not to search for functional specifications by the parameter \spec=modifies\. It also tries to solve the verification conditions automatically. Again it is crucial to name the lemma with this naming scheme, since the verfication condition generator searches for these names. \ text \The modifies clause is equal to a state update specification of the following form. \ lemma (in append_impl) shows "{t. t may_only_modify_globals Z in [next]} = {t. \next. globals t=update id id next_' (K_statefun next) (globals Z)}" apply (unfold mex_def meq_def) apply simp done text \Now that we have proven the frame-condition, it is available within the locale \append_impl\ and the \vcg\ exploits it.\ lemma (in append_impl) shows "\\ \\cont=c\ \p :== CALL append(Null,Null) \\cont=c\" proof - note append_spec = append_spec1 show ?thesis apply vcg txt \@{subgoals [display]}\ txt \With a modifies clause present we know that no change to @{term cont} has occurred. \ by simp qed text \ Of course we could add the modifies clause to the functional specification as well. But separating both has the advantage that we split up the verification work. We can make use of the modifies clause before we apply the functional specification in a fully automatic fashion. \ text \ To prove that a procedure respects the modifies clause, we only need the modifies clauses of the procedures called in the body. We do not need the functional specifications. So we can always prove the modifies clause without functional specifications, but we may need the modifies clause to prove the functional specifications. So usually the modifies clause is proved before the proof of the functional specification, so that it can already be used by the verification condition generator. \ subsection \Total Correctness\ text \When proving total correctness the additional proof burden to the user is to come up with a well-founded relation and to prove that certain states get smaller according to this relation. Proving that a relation is well-founded can be quite hard. But fortunately there are ways to construct and stick together relations so that they are well-founded by construction. This infrastructure is already present in Isabelle/HOL. For example, @{term "measure f"} is always well-founded; the lexicographic product of two well-founded relations is again well-founded and the inverse image construction @{term "inv_image"} of a well-founded relation is again well-founded. The constructions are best explained by some equations: @{thm in_measure_iff [no_vars]}\\ @{thm in_lex_iff [no_vars]}\\ @{thm in_inv_image_iff [no_vars]} Another useful construction is \<*mlex*>\ which is a combination of a measure and a lexicographic product: @{thm in_mlex_iff [no_vars]}\\ In contrast to the lexicographic product it does not construct a product type. The state may either decrease according to the measure function @{term f} or the measure stays the same and the state decreases because of the relation @{term r}. Lets look at a loop: \ lemma (in vars) "\\\<^sub>t \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b \ \M \ a\ VAR MEASURE a - \M DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg txt \@{subgoals [display]} The first conjunct of the second subgoal is the proof obligation that the variant decreases in the loop body. \ by auto text \The variant annotation is preceded by \VAR\. The capital \MEASURE\ is a shorthand for \measure (\s. a - \<^bsup>s\<^esup>M)\. Analogous there is a capital \<*MLEX*>\. \ lemma (in Fac_impl) Fac_spec': shows "\\. \\\<^sub>t {\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). \<^bsup>s\<^esup>N)"]) txt \In case of the factorial the parameter @{term N} decreases in every call. This is easily expressed by the measure function. Note that the well-founded relation for recursive procedures is formally defined on tuples containing the state space and the procedure name. \ txt \@{subgoals [display]} The initial call to the factorial is in state @{term "\"}. Note that in the precondition @{term "{\} \ {\'}"}, @{term "\'"} stems from the lemma we want to prove and @{term "\"} stems from the recursion rule for total correctness. Both are synonym for the initial state. To use the assumption in the Hoare context we have to show that the call to the factorial is invoked on a smaller @{term N} compared to the initial \\<^bsup>\\<^esup>N\. \ apply vcg txt \@{subgoals [display]} The tribute to termination is that we have to show \N - 1 < N\ in case of the recursive call. \ by simp lemma (in append_impl) append_spec2: shows "\\ Ps Qs. \\\<^sub>t \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) txt \In case of the append function the length of the list referenced by @{term p} decreases in every recursive call. \ txt \@{subgoals [margin=80,display]}\ apply vcg apply (fastforce simp add: List_list) done text \ In case of the lists above, we have used a relational list abstraction @{term List} to construct the HOL lists @{term Ps} and @{term Qs} for the pre- and postcondition. To supply a proper measure function we use a functional abstraction @{term list}. The functional abstraction can be defined by means of the relational list abstraction, since the lists are already uniquely determined by the relational abstraction: @{thm islist_def [no_vars]}\\ @{thm list_def [no_vars]} \isacommand{lemma} @{thm List_conv_islist_list [no_vars]} \ text \ The next contrived example is taken from \cite{Homeier-95-vcg}, to illustrate a more complex termination criterion for mutually recursive procedures. The procedures do not calculate anything useful. \ procedures pedal(N::nat,M::nat) "IF 0 < \N THEN IF 0 < \M THEN CALL coast(\N- 1,\M- 1) FI;; CALL pedal(\N- 1,\M) FI" and coast(N::nat,M::nat) "CALL pedal(\N,\M);; IF 0 < \M THEN CALL coast(\N,\M- 1) FI" text \ In the recursive calls in procedure \pedal\ the first argument always decreases. In the body of \coast\ in the recursive call of \coast\ the second argument decreases, but in the call to \pedal\ no argument decreases. Therefore an relation only on the state space is insufficient. We have to take the procedure names into account, too. We consider the procedure \coast\ to be ``bigger'' than \pedal\ when we construct a well-founded relation on the product of state space and procedure names. \ ML \ML_Thms.bind_thm ("HoareTotal_ProcRec2", Hoare.gen_proc_rec @{context} Hoare.Total 2)\ text \ We provide the ML function {\tt gen\_proc\_rec} to automatically derive a convenient rule for recursion for a given number of mutually recursive procedures. \ lemma (in pedal_coast_clique) shows "(\\. \\\<^sub>t {\} PROC pedal(\N,\M) UNIV) \ (\\. \\\<^sub>t {\} PROC coast(\N,\M) UNIV)" apply (hoare_rule HoareTotal_ProcRec2 [where r= "((\(s,p). \<^bsup>s\<^esup>N) <*mlex*> (\(s,p). \<^bsup>s\<^esup>M) <*mlex*> measure (\(s,p). if p = coast_'proc then 1 else 0))"]) txt \We can directly express the termination condition described above with the \<*mlex*>\ construction. Either state component \N\ decreases, or it stays the same and \M\ decreases or this also stays the same, but then the procedure name has to decrease.\ txt \@{subgoals [margin=80,display]}\ apply simp_all txt \@{subgoals [margin=75,display]}\ by (vcg,simp)+ text \We can achieve the same effect without \<*mlex*>\ by using the ordinary lexicographic product \<*lex*>\, \inv_image\ and \measure\ \ lemma (in pedal_coast_clique) shows "(\\. \\\<^sub>t {\} PROC pedal(\N,\M) UNIV) \ (\\. \\\<^sub>t {\} PROC coast(\N,\M) UNIV)" apply (hoare_rule HoareTotal_ProcRec2 [where r= "inv_image (measure (\m. m) <*lex*> measure (\m. m) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)) (\(s,p). (\<^bsup>s\<^esup>N,\<^bsup>s\<^esup>M,p))"]) txt \With the lexicographic product we construct a well-founded relation on triples of type @{typ "(nat\nat\string)"}. With @{term inv_image} we project the components out of the state-space and the procedure names to this triple. \ txt \@{subgoals [margin=75,display]}\ apply simp_all -by (vcg,simp)+ - + by (vcg,force)+ text \By doing some arithmetic we can express the termination condition with a single measure function. \ lemma (in pedal_coast_clique) shows "(\\. \\\<^sub>t {\} PROC pedal(\N,\M) UNIV) \ (\\. \\\<^sub>t {\} PROC coast(\N,\M) UNIV)" apply(hoare_rule HoareTotal_ProcRec2 [where r= "measure (\(s,p). \<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M + (if p = coast_'proc then 1 else 0))"]) apply simp_all txt \@{subgoals [margin=75,display]}\ by (vcg,simp,arith?)+ subsection \Guards\ text (in vars) \The purpose of a guard is to guard the {\bf (sub-) expressions} of a statement against runtime faults. Typical runtime faults are array bound violations, dereferencing null pointers or arithmetical overflow. Guards make the potential runtime faults explicit, since the expressions themselves never ``fail'' because they are ordinary HOL expressions. To relieve the user from typing in lots of standard guards for every subexpression, we supply some input syntax for the common language constructs that automatically generate the guards. For example the guarded assignment \\M :==\<^sub>g (\M + 1) div \N\ gets expanded to guarded command @{term "\M :==\<^sub>g (\M + 1) div \N"}. Here @{term "in_range"} is uninterpreted by now. \ lemma (in vars) "\\\True\ \M :==\<^sub>g (\M + 1) div \N \True\" apply vcg txt \@{subgoals}\ oops text \ The user can supply on (overloaded) definition of \in_range\ to fit to his needs. Currently guards are generated for: \begin{itemize} \item overflow and underflow of numbers (\in_range\). For subtraction of natural numbers \a - b\ the guard \b \ a\ is generated instead of \in_range\ to guard against underflows. \item division by \0\ \item dereferencing of @{term Null} pointers \item array bound violations \end{itemize} Following (input) variants of guarded statements are available: \begin{itemize} \item Assignment: \\ :==\<^sub>g \\ \item If: \IF\<^sub>g \\ \item While: \WHILE\<^sub>g \\ \item Call: \CALL\<^sub>g \\ or \\ :== CALL\<^sub>g \\ \end{itemize} \ subsection \Miscellaneous Techniques\ subsubsection \Modifies Clause\ text \We look at some issues regarding the modifies clause with the example of insertion sort for heap lists. \ primrec sorted:: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted le [] = True" | "sorted le (x#xs) = ((\y\set xs. le x y) \ sorted le xs)" procedures (imports globals_heap) insert(r::ref,p::ref | p::ref) "IF \r=Null THEN SKIP ELSE IF \p=Null THEN \p :== \r;; \p\\next :== Null ELSE IF \r\\cont \ \p\\cont THEN \r\\next :== \p;; \p:==\r ELSE \p\\next :== CALL insert(\r,\p\\next) FI FI FI" lemma (in insert_impl) insert_modifies: "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC insert(\r,\p) {t. t may_only_modify_globals \ in [next]}" by (hoare_rule HoarePartial.ProcRec1) (vcg spec=modifies) lemma (in insert_impl) insert_spec: "\\ Ps . \\ \\. List \p \next Ps \ sorted (\) (map \cont Ps) \ \r \ Null \ \r \ set Ps\ \p :== PROC insert(\r,\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = insert \<^bsup>\\<^esup>r (set Ps) \ (\x. x \ set Qs \ \next x = \<^bsup>\\<^esup>next x)\" (*<*) apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply (intro conjI impI) apply fastforce apply fastforce apply fastforce apply (clarsimp) apply force done (*>*) text \ In the postcondition of the functional specification there is a small but important subtlety. Whenever we talk about the @{term "cont"} part we refer to the one of the pre-state. The reason is that we have separated out the information that @{term "cont"} is not modified by the procedure, to the modifies clause. So whenever we talk about unmodified parts in the postcondition we have to use the pre-state part, or explicitly state an equality in the postcondition. The reason is simple. If the postcondition would talk about \\cont\ instead of \mbox{\\<^bsup>\\<^esup>cont\}, we get a new instance of \cont\ during verification and the postcondition would only state something about this new instance. But as the verification condition generator uses the modifies clause the caller of @{term "insert"} instead still has the old \cont\ after the call. Thats the sense of the modifies clause. So the caller and the specification simply talk about two different things, without being able to relate them (unless an explicit equality is added to the specification). \ subsubsection \Annotations\ text \ Annotations (like loop invariants) are mere syntactic sugar of statements that are used by the \vcg\. Logically a statement with an annotation is equal to the statement without it. Hence annotations can be introduced by the user while building a proof: @{thm [source] HoarePartial.annotateI}: @{thm [mode=Rule] HoarePartial.annotateI [no_vars]} When introducing annotations it can easily happen that these mess around with the nesting of sequential composition. Then after stripping the annotations the resulting statement is no longer syntactically identical to original one, only equivalent modulo associativity of sequential composition. The following rule also deals with this case: @{thm [source] HoarePartial.annotate_normI}: @{thm [mode=Rule] HoarePartial.annotate_normI [no_vars]} \ text_raw \\paragraph{Loop Annotations} \mbox{} \medskip \mbox{} \ procedures (imports globals_heap) insertSort(p::ref| p::ref) where r::ref q::ref in "\r:==Null;; WHILE (\p \ Null) DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p:==\r" lemma (in insertSort_impl) insertSort_modifies: shows "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC insertSort(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done text \Insertion sort is not implemented recursively here, but with a loop. Note that the while loop is not annotated with an invariant in the procedure definition. The invariant only comes into play during verification. Therefore we annotate the loop first, before we run the \vcg\. \ lemma (in insertSort_impl) insertSort_spec: shows "\\ Ps. \\ \\. List \p \next Ps \ \p :== PROC insertSort(\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = set Ps\" apply (hoare_rule HoarePartial.ProcRec1) apply (hoare_rule anno= "\r :== Null;; WHILE \p \ Null INV \\Qs Rs. List \p \next Qs \ List \r \next Rs \ set Qs \ set Rs = {} \ sorted (\) (map \cont Rs) \ set Qs \ set Rs = set Ps \ \cont = \<^bsup>\\<^esup>cont \ DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p :== \r" in HoarePartial.annotateI) apply vcg txt \\\\\ (*<*) apply fastforce prefer 2 apply fastforce apply (clarsimp) apply (rule_tac x=ps in exI) apply (intro conjI) apply (rule heap_eq_ListI1) apply assumption apply clarsimp apply (subgoal_tac "x\p \ x \ set Rs") apply auto done (*>*) text \The method \hoare_rule\ automatically solves the side-condition that the annotated program is the same as the original one after stripping the annotations.\ text_raw \\paragraph{Specification Annotations} \mbox{} \medskip \mbox{} \ text \ When verifying a larger block of program text, it might be useful to split up the block and to prove the parts in isolation. This is especially useful to isolate loops. On the level of the Hoare calculus the parts can then be combined with the consequence rule. To automate this process we introduce the derived command @{term specAnno}, which allows to introduce a Hoare tuple (inclusive auxiliary variables) in the program text: @{thm specAnno_def [no_vars]} The whole annotation reduces to the body @{term "c undefined"}. The type of the assertions @{term "P"}, @{term "Q"} and @{term "A"} is @{typ "'a \ 's set"} and the type of command @{term c} is @{typ "'a \ ('s,'p,'f) com"}. All entities formally depend on an auxiliary (logical) variable of type @{typ "'a"}. The body @{term "c"} formally also depends on this variable, since a nested annotation or loop invariant may also depend on this logical variable. But the raw body without annotations does not depend on the logical variable. The logical variable is only used by the verification condition generator. We express this by defining the whole @{term specAnno} to be equivalent with the body applied to an arbitrary variable. The Hoare rule for \specAnno\ is mainly an instance of the consequence rule: @{thm [mode=Rule,mode=ParenStmt] HoarePartial.SpecAnno [no_vars]} The side-condition @{term "\Z. c Z = c undefined"} expresses the intention of body @{term c} explained above: The raw body is independent of the auxiliary variable. This side-condition is solved automatically by the \vcg\. The concrete syntax for this specification annotation is shown in the following example: \ lemma (in vars) "\\ {\} \I :== \M;; ANNO \. \\. \I = \<^bsup>\\<^esup>M\ \M :== \N;; \N :== \I \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>I\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" txt \With the annotation we can name an intermediate state @{term \}. Since the postcondition refers to @{term "\"} we have to link the information about the equivalence of \\<^bsup>\\<^esup>I\ and \\<^bsup>\\<^esup>M\ in the specification in order to be able to derive the postcondition. \ apply vcg_step apply vcg_step txt \@{subgoals [display]}\ txt \The first subgoal is the isolated Hoare tuple. The second one is the side-condition of the consequence rule that allows us to derive the outermost pre/post condition from our inserted specification. \\I = \<^bsup>\\<^esup>M\ is the precondition of the specification, The second conjunct is a simplified version of \\t. \<^bsup>t\<^esup>M = \N \ \<^bsup>t\<^esup>N = \I \ \<^bsup>t\<^esup>M = \<^bsup>\\<^esup>N \ \<^bsup>t\<^esup>N = \<^bsup>\\<^esup>M\ expressing that the postcondition of the specification implies the outermost postcondition. \ apply vcg txt \@{subgoals [display]}\ apply simp apply vcg txt \@{subgoals [display]}\ by simp lemma (in vars) "\\ {\} \I :== \M;; ANNO \. \\. \I = \<^bsup>\\<^esup>M\ \M :== \N;; \N :== \I \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>I\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" apply vcg txt \@{subgoals [display]}\ by simp_all text \Note that \vcg_step\ changes the order of sequential composition, to allow the user to decompose sequences by repeated calls to \vcg_step\, whereas \vcg\ preserves the order. The above example illustrates how we can introduce a new logical state variable @{term "\"}. You can introduce multiple variables by using a tuple: \ lemma (in vars) "\\ {\} \I :== \M;; ANNO (n,i,m). \\I = \<^bsup>\\<^esup>M \ \N=n \ \I=i \ \M=m\ \M :== \N;; \N :== \I \\M = n \ \N = i\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" apply vcg txt \@{subgoals [display]}\ by simp_all text_raw \\paragraph{Lemma Annotations} \mbox{} \medskip \mbox{} \ text \ The specification annotations described before split the verification into several Hoare triples which result in several subgoals. If we instead want to proof the Hoare triples independently as separate lemmas we can use the \LEMMA\ annotation to plug together the lemmas. It inserts the lemma in the same fashion as the specification annotation. \ lemma (in vars) foo_lemma: "\n m. \\ \\N = n \ \M = m\ \N :== \N + 1;; \M :== \M + 1 \\N = n + 1 \ \M = m + 1\" apply vcg apply simp done lemma (in vars) "\\ \\N = n \ \M = m\ LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; \N :== \N + 1 \\N = n + 2 \ \M = m + 1\" apply vcg apply simp done lemma (in vars) "\\ \\N = n \ \M = m\ LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END \\N = n + 2 \ \M = m + 2\" apply vcg apply simp done lemma (in vars) "\\ \\N = n \ \M = m\ \N :== \N + 1;; \M :== \M + 1;; \N :== \N + 1;; \M :== \M + 1 \\N = n + 2 \ \M = m + 2\" apply (hoare_rule anno= "LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END" in HoarePartial.annotate_normI) apply vcg apply simp done subsubsection \Total Correctness of Nested Loops\ text \ When proving termination of nested loops it is sometimes necessary to express that the loop variable of the outer loop is not modified in the inner loop. To express this one has to fix the value of the outer loop variable before the inner loop and use this value in the invariant of the inner loop. This can be achieved by surrounding the inner while loop with an \ANNO\ specification as explained previously. However, this leads to repeating the invariant of the inner loop three times: in the invariant itself and in the the pre- and postcondition of the \ANNO\ specification. Moreover one has to deal with the additional subgoal introduced by \ANNO\ that expresses how the pre- and postcondition is connected to the invariant. To avoid this extra specification and verification work, we introduce an variant of the annotated while-loop, where one can introduce logical variables by \FIX\. As for the \ANNO\ specification multiple logical variables can be introduced via a tuple (\FIX (a,b,c).\). The Hoare logic rule for the augmented while-loop is a mixture of the invariant rule for loops and the consequence rule for \ANNO\: \begin{center} @{thm [mode=Rule,mode=ParenStmt] HoareTotal.WhileAnnoFix' [no_vars]} \end{center} The first premise expresses that the precondition implies the invariant and that the invariant together with the negated loop condition implies the postcondition. Since both implications may depend on the choice of the auxiliary variable @{term "Z"} these two implications are expressed in a single premise and not in two of them as for the usual while rule. The second premise is the preservation of the invariant by the loop body. And the third premise is the side-condition that the computational part of the body does not depend on the auxiliary variable. Finally the last premise is the well-foundedness of the variant. The last two premises are usually discharged automatically by the verification condition generator. Hence usually two subgoals remain for the user, stemming from the first two premises. The following example illustrates the usage of this rule. The outer loop increments the loop variable @{term "M"} while the inner loop increments @{term "N"}. To discharge the proof obligation for the termination of the outer loop, we need to know that the inner loop does not mess around with @{term "M"}. This is expressed by introducing the logical variable @{term "m"} and fixing the value of @{term "M"} to it. \ lemma (in vars) "\\\<^sub>t \\M=0 \ \N=0\ WHILE (\M < i) INV \\M \ i \ (\M \ 0 \ \N = j) \ \N \ j\ VAR MEASURE (i - \M) DO \N :== 0;; WHILE (\N < j) FIX m. INV \\M=m \ \N \ j\ VAR MEASURE (j - \N) DO \N :== \N + 1 OD;; \M :== \M + 1 OD \\M=i \ (\M\0 \ \N=j)\" apply vcg txt \@{subgoals [display]} The first subgoal is from the precondition to the invariant of the outer loop. The fourth subgoal is from the invariant together with the negated loop condition of the outer loop to the postcondition. The subgoals two and three are from the body of the outer while loop which is mainly the inner while loop. Because we introduce the logical variable @{term "m"} here, the while Rule described above is used instead of the ordinary while Rule. That is why we end up with two subgoals for the inner loop. Subgoal two is from the invariant and the loop condition of the outer loop to the invariant of the inner loop. And at the same time from the invariant of the inner loop to the invariant of the outer loop (together with the proof obligation that the measure of the outer loop decreases). The universal quantified variables @{term "Ma"} and @{term "N"} are the ``fresh'' state variables introduced for the final state of the inner loop. The equality @{term "Ma=M"} is the result of the equality \\M=m\ in the inner invariant. Subgoal three is the preservation of the invariant by the inner loop body (together with the proof obligation that the measure of the inner loop decreases). \ (*<*) apply (simp) apply (simp,arith) apply (simp,arith) done (*>*) subsection \Functional Correctness, Termination and Runtime Faults\ text \ Total correctness of a program with guards conceptually leads to three verification tasks. \begin{itemize} \item functional (partial) correctness \item absence of runtime faults \item termination \end{itemize} In case of a modifies specification the functional correctness part can be solved automatically. But the absence of runtime faults and termination may be non trivial. Fortunately the modifies clause is usually just a helpful companion of another specification that expresses the ``real'' functional behaviour. Therefor the task to prove the absence of runtime faults and termination can be dealt with during the proof of this functional specification. In most cases the absence of runtime faults and termination heavily build on the functional specification parts. So after all there is no reason why we should again prove the absence of runtime faults and termination for the modifies clause. Therefor it suffices to have partial correctness of the modifies clause for a program were all guards are ignored. This leads to the following pattern:\ procedures foo (N::nat|M::nat) "\M :== \M \ \think of body with guards instead\" foo_spec: "\\. \\\<^sub>t (P \) \M :== PROC foo(\N) (Q \)" foo_modifies: "\\. \\\<^bsub>/UNIV\<^esub> {\} \M :== PROC foo(\N) {t. t may_only_modify_globals \ in []}" text \ The verification condition generator can solve those modifies clauses automatically and can use them to simplify calls to \foo\ even in the context of total correctness. \ subsection \Procedures and Locales \label{sec:Locales}\ text \ Verification of a larger program is organised on the granularity of procedures. We proof the procedures in a bottom up fashion. Of course you can also always use Isabelle's dummy proof \sorry\ to prototype your formalisation. So you can write the theory in a bottom up fashion but actually prove the lemmas in any other order. Here are some explanations of handling of locales. In the examples below, consider \proc\<^sub>1\ and \proc\<^sub>2\ to be ``leaf'' procedures, which do not call any other procedure. Procedure \proc\ directly calls \proc\<^sub>1\ and \proc\<^sub>2\. \isacommand{lemma} (\isacommand{in} \proc\<^sub>1_impl\) \proc\<^sub>1_modifies\:\\ \isacommand{shows} \\\ After the proof of \proc\<^sub>1_modifies\, the \isacommand{in} directive stores the lemma in the locale \proc\<^sub>1_impl\. When we later on include \proc\<^sub>1_impl\ or prove another theorem in locale \proc\<^sub>1_impl\ the lemma \proc\<^sub>1_modifies\ will already be available as fact. \isacommand{lemma} (\isacommand{in} \proc\<^sub>1_impl\) \proc\<^sub>1_spec\:\\ \isacommand{shows} \\\ \isacommand{lemma} (\isacommand{in} \proc\<^sub>2_impl\) \proc\<^sub>2_modifies\:\\ \isacommand{shows} \\\ \isacommand{lemma} (\isacommand{in} \proc\<^sub>2_impl\) \proc\<^sub>2_spec\:\\ \isacommand{shows} \\\ \isacommand{lemma} (\isacommand{in} \proc_impl\) \proc_modifies\:\\ \isacommand{shows} \\\ Note that we do not explicitly include anything about \proc\<^sub>1\ or \proc\<^sub>2\ here. This is handled automatically. When defining an \impl\-locale it imports all \impl\-locales of procedures that are called in the body. In case of \proc_impl\ this means, that \proc\<^sub>1_impl\ and \proc\<^sub>2_impl\ are imported. This has the neat effect that all theorems that are proven in \proc\<^sub>1_impl\ and \proc\<^sub>2_impl\ are also present in \proc_impl\. \isacommand{lemma} (\isacommand{in} \proc_impl\) \proc_spec\:\\ \isacommand{shows} \\\ As we have seen in this example you only have to prove a procedure in its own \impl\ locale. You do not have to include any other locale. \ subsection \Records \label{sec:records}\ text \ Before @{term "statespaces"} where introduced the state was represented as a @{term "record"}. This is still supported. Compared to the flexibility of statespaces there are some drawbacks in particular with respect to modularity. Even names of local variables and parameters are globally visible and records can only be extended in a linear fashion, whereas statespaces also allow multiple inheritance. The usage of records is quite similar to the usage of statespaces. We repeat the example of an append function for heap lists. First we define the global components. Again the appearance of the prefix `globals' is mandatory. This is the way the syntax layer distinguishes local and global variables. \ record globals_list = next_' :: "ref \ ref" cont_' :: "ref \ nat" text \The local variables also have to be defined as a record before the actual definition of the procedure. The parent record \state\ defines a generic @{term "globals"} field as a place-holder for the record of global components. In contrast to the statespace approach there is no single @{term "locals"} slot. The local components are just added to the record. \ record 'g list_vars = "'g state" + p_' :: "ref" q_' :: "ref" r_' :: "ref" root_' :: "ref" tmp_' :: "ref" text \Since the parameters and local variables are determined by the record, there are no type annotations or definitions of local variables while defining a procedure. \ procedures append'(p,q|p) = "IF \p=Null THEN \p :== \q ELSE \p \\next:== CALL append'(\p\\next,\q) FI" text \As in the statespace approach, a locale called \append'_impl\ is created. Note that we do not give any explicit information which global or local state-record to use. Since the records are already defined we rely on Isabelle's type inference. Dealing with the locale is analogous to the case with statespaces. \ lemma (in append'_impl) append'_modifies: shows "\\. \\ {\} \p :== PROC append'(\p,\q) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in append'_impl) append'_spec: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append'(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply fastforce done text \ However, in some corner cases the inferred state type in a procedure definition can be too general which raises problems when attempting to proof a suitable specifications in the locale. Consider for example the simple procedure body @{term "\p :== NULL"} for a procedure \init\. \ procedures init (|p) = "\p:== Null" text \ Here Isabelle can only infer the local variable record. Since no reference to any global variable is made the type fixed for the global variables (in the locale \init'_impl\) is a type variable say @{typ "'g"} and not a @{term "globals_list"} record. Any specification mentioning @{term "next"} or @{term "cont"} restricts the state type and cannot be added to the locale \init_impl\. Hence we have to restrict the body @{term "\p :== NULL"} in the first place by adding a typing annotation: \ procedures init' (|p) = "\p:== Null::(('a globals_list_scheme, 'b) list_vars_scheme, char list, 'c) com" subsubsection \Extending State Spaces\ text \ The records in Isabelle are extensible \cite{Nipkow-02-hol,NaraschewskiW-TPHOLs98}. In principle this can be exploited during verification. The state space can be extended while we we add procedures. But there is one major drawback: \begin{itemize} \item records can only be extended in a linear fashion (there is no multiple inheritance) \end{itemize} You can extend both the main state record as well as the record for the global variables. \ subsubsection \Mapping Variables to Record Fields\ text \ Generally the state space (global and local variables) is flat and all components are accessible from everywhere. Locality or globality of variables is achieved by the proper \init\ and \return\/\result\ functions in procedure calls. What is the best way to map programming language variables to the state records? One way is to disambiguate all names, by using the procedure names as prefix or the structure names for heap components. This leads to long names and lots of record components. But for local variables this is not necessary, since variable @{term i} of procedure @{term A} and variable @{term "i"} of procedure @{term B} can be mapped to the same record component, without any harm, provided they have the same logical type. Therefor for local variables it is preferable to map them per type. You only have to distinguish a variable with the same name if they have a different type. Note that all pointers just have logical type \ref\. So you even do not have to distinguish between a pointer \p\ to a integer and a pointer \p\ to a list. For global components (global variables and heap structures) you have to disambiguate the name. But hopefully the field names of structures have different names anyway. Also note that there is no notion of hiding of a global component by a local one in the logic. You have to disambiguate global and local names! As the names of the components show up in the specifications and the proof obligations, names are even more important as for programming. Try to find meaningful and short names, to avoid cluttering up your reasoning. \ (*<*) text \ in locales, includes, spec or impl? Names: per type not per procedure\ downgrading total to partial\ \ (*>*) text \\ (*<*) end (*>*) diff --git a/thys/Simpl/ex/VcgExTotal.thy b/thys/Simpl/ex/VcgExTotal.thy --- a/thys/Simpl/ex/VcgExTotal.thy +++ b/thys/Simpl/ex/VcgExTotal.thy @@ -1,386 +1,370 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de License: LGPL *) (* Title: VcgExTotal.thy Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer Some rights reserved, TU Muenchen This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) section \Examples for Total Correctness\ theory VcgExTotal imports "../HeapList" "../Vcg" begin record 'g vars = "'g state" + A_' :: nat I_' :: nat M_' :: nat N_' :: nat R_' :: nat S_' :: nat Abr_':: string lemma "\\\<^sub>t \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b \ \M \ a\ VAR MEASURE a - \M DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg apply (auto) done lemma "\\\<^sub>t \\I \ 3\ WHILE \I < 10 INV \\I\ 10\ VAR MEASURE 10 - \I DO \I :== \I + 1 OD \\I = 10\" apply vcg apply auto done text \Total correctness of a nested loop. In the inner loop we have to express that the loop variable of the outer loop is not changed. We use \FIX\ to introduce a new logical variable \ lemma "\\\<^sub>t \\M=0 \ \N=0\ WHILE (\M < i) INV \\M \ i \ (\M \ 0 \ \N = j) \ \N \ j\ VAR MEASURE (i - \M) DO \N :== 0;; WHILE (\N < j) FIX m. INV \\M=m \ \N \ j\ VAR MEASURE (j - \N) DO \N :== \N + 1 OD;; \M :== \M + 1 OD \\M=i \ (\M\0 \ \N=j)\" apply vcg apply simp_all apply arith done primrec fac:: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = (Suc n) * fac n" lemma fac_simp [simp]: "0 < i \ fac i = i * fac (i - 1)" by (cases i) simp_all procedures Fac (N | R) = "IF \N = 0 THEN \R :== 1 ELSE CALL Fac(\N - 1,\R);; \R :== \N * \R FI" lemma (in Fac_impl) Fac_spec: shows "\n. \\\<^sub>t \\N=n\ \R :== PROC Fac(\N) \\R = fac n\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). \<^bsup>s\<^esup>N)"]) apply vcg apply simp done procedures p91(R,N | R) = "IF 100 < \N THEN \R :== \N - 10 ELSE \R :== CALL p91(\R,\N+11);; \R :== CALL p91(\R,\R) FI" p91_spec: "\n. \\\<^sub>t \\N=n\ \R :== PROC p91(\R,\N) \if 100 < n then \R = n - 10 else \R = 91\,{}" lemma (in p91_impl) p91_spec: shows "\\. \\\<^sub>t {\} \R :== PROC p91(\R,\N) \if 100 < \<^bsup>\\<^esup>N then \R = \<^bsup>\\<^esup>N - 10 else \R = 91\,{}" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). 101 - \<^bsup>s\<^esup>N)"]) apply vcg apply clarsimp apply arith done record globals_list = next_' :: "ref \ ref" cont_' :: "ref \ nat" record 'g list_vars = "'g state" + p_' :: "ref" q_' :: "ref" r_' :: "ref" root_' :: "ref" tmp_' :: "ref" procedures append(p,q|p) = "IF \p=Null THEN \p :== \q ELSE \p\\next :== CALL append(\p\\next,\q) FI" lemma (in append_impl) shows "\\ Ps Qs. \\\<^sub>t \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) apply vcg apply (fastforce simp add: List_list) done lemma (in append_impl) shows "\\ Ps Qs. \\\<^sub>t \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) apply vcg apply (fastforce simp add: List_list) done lemma (in append_impl) shows append_spec: "\\. \\\<^sub>t ({\} \ \islist \p \next\) \p :== PROC append(\p,\q) \\Ps Qs. List \<^bsup>\\<^esup>p \<^bsup>\\<^esup>next Ps \ List \<^bsup>\\<^esup>q \<^bsup>\\<^esup>next Qs \ set Ps \ set Qs = {} \ List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) apply vcg apply fastforce done lemma "\\\List \p \next Ps\ \q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ DO \r :== \p;; \p :== \p\\next;; \r\\next :== \q;; \q :== \r OD;; \p :==\q \List \p \next (rev Ps)\ " apply vcg apply clarsimp apply clarsimp apply force apply simp done lemma conjI2: "\Q; Q \ P\ \ P \ Q" by blast procedures Rev(p|p) = "\q :== Null;; WHILE \p \ Null DO \r :== \p;; \\p \ Null\\ \p :== \p\\next;; \\r \ Null\\ \r\\next :== \q;; \q :== \r OD;; \p :==\q" Rev_spec: "\Ps. \\\<^sub>t \List \p \next Ps\ \p :== PROC Rev(\p) \List \p \next (rev Ps)\" Rev_modifies: "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC Rev(\p) {t. t may_only_modify_globals \ in [next]}" text \We only need partial correctness of modifies clause!\ lemma upd_hd_next: assumes p_ps: "List p next (p#ps)" shows "List (next p) (next(p := q)) ps" proof - from p_ps have "p \ set ps" by auto with p_ps show ?thesis by simp qed lemma (in Rev_impl) shows Rev_spec: "\Ps. \\\<^sub>t \List \p \next Ps\ \p :== PROC Rev(\p) \List \p \next (rev Ps)\" apply (hoare_rule HoareTotal.ProcNoRec1) apply (hoare_rule anno = "\q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ VAR MEASURE (length (list \p \next) ) DO \r :== \p;; \\p \ Null\\\p :== \p\\next;; \\r \ Null\\\r\\next :== \q;; \q :== \r OD;; \p :==\q" in HoareTotal.annotateI) apply vcg apply clarsimp apply clarsimp apply (rule conjI2) apply force apply clarsimp apply (subgoal_tac "List p next (p#ps)") prefer 2 apply simp apply (frule_tac q=q in upd_hd_next) apply (simp only: List_list) apply simp apply simp done lemma (in Rev_impl) shows Rev_modifies: "\\. \\\<^bsub>/UNIV \<^esub>{\} \p :== PROC Rev(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcNoRec1) apply (vcg spec=modifies) done lemma "\\\<^sub>t \List \p \next Ps\ \q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ VAR MEASURE (length (list \p \next) ) DO \r :== \p;; \p :== \p\\next;; \r\\next :== \q;; \q :== \r OD;; \p :==\q \List \p \next (rev Ps)\ " apply vcg apply clarsimp apply clarsimp apply (rule conjI2) apply force apply clarsimp apply (subgoal_tac "List p next (p#ps)") prefer 2 apply simp apply (frule_tac q=q in upd_hd_next) apply (simp only: List_list) apply simp apply simp done procedures pedal(N,M) = "IF 0 < \N THEN IF 0 < \M THEN CALL coast(\N- 1,\M- 1) FI;; CALL pedal(\N- 1,\M) FI" and coast(N,M) = "CALL pedal(\N,\M);; IF 0 < \M THEN CALL coast(\N,\M- 1) FI" ML \ML_Thms.bind_thm ("HoareTotal_ProcRec2", Hoare.gen_proc_rec @{context} Hoare.Total 2)\ lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply (hoare_rule HoareTotal_ProcRec2 [where ?r= "inv_image (measure (\m. m) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)) (\(s,p). (\<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M,p))"]) apply simp_all - apply vcg - apply simp - apply vcg - apply simp + apply (vcg,force)+ done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply (hoare_rule HoareTotal_ProcRec2 [where ?r= "inv_image (measure (\m. m) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)) (\(s,p). (\<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M,p))"]) apply simp_all - apply vcg - apply simp - apply vcg - apply simp + apply (vcg,force)+ done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply(hoare_rule HoareTotal_ProcRec2 [where ?r= "measure (\(s,p). \<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M + (if p = coast_'proc then 1 else 0))"]) apply simp_all - apply vcg - apply simp - apply arith - apply vcg - apply simp + apply (vcg,force)+ done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply(hoare_rule HoareTotal_ProcRec2 [where ?r= "(\(s,p). \<^bsup>s\<^esup>N) <*mlex*> (\(s,p). \<^bsup>s\<^esup>M) <*mlex*> measure (\(s,p). if p = coast_'proc then 1 else 0)"]) apply simp_all - apply vcg - apply simp - apply vcg - apply simp - done + apply (vcg,force)+ + done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply(hoare_rule HoareTotal_ProcRec2 [where ?r= "measure (\s. \<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)"]) apply simp_all - apply vcg - apply simp - apply vcg - apply simp - done + apply (vcg,force)+ + done end diff --git a/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy b/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy --- a/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy +++ b/thys/Symmetric_Polynomials/Symmetric_Polynomials.thy @@ -1,2890 +1,2890 @@ (* File: Symmetric_Polynomials.thy Author: Manuel Eberl (TU München) The definition of symmetric polynomials and the elementary symmetric polynomials. Proof of the fundamental theorem of symmetric polynomials and its corollaries. *) section \Symmetric Polynomials\ theory Symmetric_Polynomials imports Vieta "Polynomials.More_MPoly_Type" "HOL-Library.Permutations" begin subsection \Auxiliary facts\ (* TODO: Many of these facts and definitions should be moved elsewhere, especially the ones on polynomials (leading monomial, mapping, insertion etc.) *) text \ An infinite set has infinitely many infinite subsets. \ lemma infinite_infinite_subsets: assumes "infinite A" shows "infinite {X. X \ A \ infinite X}" proof - have "\k. \X. X \ A \ infinite X \ card (A - X) = k" for k :: nat proof fix k :: nat obtain Y where "finite Y" "card Y = k" "Y \ A" using infinite_arbitrarily_large[of A k] assms by auto moreover from this have "A - (A - Y) = Y" by auto ultimately show "\X. X \ A \ infinite X \ card (A - X) = k" using assms by (intro exI[of _ "A - Y"]) auto qed from choice[OF this] obtain f where f: "\k. f k \ A \ infinite (f k) \ card (A - f k) = k" by blast have "k = l" if "f k = f l" for k l proof (rule ccontr) assume "k \ l" hence "card (A - f k) \ card (A - f l)" using f[of k] f[of l] by auto with \f k = f l\ show False by simp qed hence "inj f" by (auto intro: injI) moreover have "range f \ {X. X \ A \ infinite X}" using f by auto ultimately show ?thesis by (subst infinite_iff_countable_subset) auto qed text \ An infinite set contains infinitely many finite subsets of any fixed nonzero cardinality. \ lemma infinite_card_subsets: assumes "infinite A" "k > 0" shows "infinite {X. X \ A \ finite X \ card X = k}" proof - obtain B where B: "B \ A" "finite B" "card B = k - 1" using infinite_arbitrarily_large[OF assms(1), of "k - 1"] by blast define f where "f = (\x. insert x B)" have "f ` (A - B) \ {X. X \ A \ finite X \ card X = k}" using assms B by (auto simp: f_def) moreover have "inj_on f (A - B)" by (auto intro!: inj_onI simp: f_def) hence "infinite (f ` (A - B))" using assms B by (subst finite_image_iff) auto ultimately show ?thesis by (rule infinite_super) qed lemma comp_bij_eq_iff: assumes "bij f" shows "g \ f = h \ f \ g = h" proof assume *: "g \ f = h \ f" show "g = h" proof fix x obtain y where [simp]: "x = f y" using bij_is_surj[OF assms] by auto have "(g \ f) y = (h \ f) y" by (simp only: *) thus "g x = h x" by simp qed qed auto lemma sum_list_replicate [simp]: "sum_list (replicate n x) = of_nat n * (x :: 'a :: semiring_1)" by (induction n) (auto simp: algebra_simps) lemma ex_subset_of_card: assumes "finite A" "card A \ k" shows "\B. B \ A \ card B = k" using assms proof (induction arbitrary: k rule: finite_induct) case empty thus ?case by auto next case (insert x A k) show ?case proof (cases "k = 0") case True thus ?thesis by (intro exI[of _ "{}"]) auto next case False from insert have "\B\A. card B = k - 1" by (intro insert.IH) auto then obtain B where B: "B \ A" "card B = k - 1" by auto with insert have [simp]: "x \ B" by auto have "insert x B \ insert x A" using B insert by auto moreover have "card (insert x B) = k" using insert B finite_subset[of B A] False by (subst card_insert) auto ultimately show ?thesis by blast qed qed lemma length_sorted_list_of_set [simp]: "length (sorted_list_of_set A) = card A" using distinct_card[of "sorted_list_of_set A"] by (cases "finite A") simp_all lemma upt_add_eq_append': "i \ j \ j \ k \ [i..Subrings and ring homomorphisms\ locale ring_closed = fixes A :: "'a :: comm_ring_1 set" assumes zero_closed [simp]: "0 \ A" assumes one_closed [simp]: "1 \ A" assumes add_closed [simp]: "x \ A \ y \ A \ (x + y) \ A" assumes mult_closed [simp]: "x \ A \ y \ A \ (x * y) \ A" assumes uminus_closed [simp]: "x \ A \ -x \ A" begin lemma minus_closed [simp]: "x \ A \ y \ A \ x - y \ A" using add_closed[of x "-y"] uminus_closed[of y] by auto lemma sum_closed [intro]: "(\x. x \ X \ f x \ A) \ sum f X \ A" by (induction X rule: infinite_finite_induct) auto lemma power_closed [intro]: "x \ A \ x ^ n \ A" by (induction n) auto lemma Sum_any_closed [intro]: "(\x. f x \ A) \ Sum_any f \ A" unfolding Sum_any.expand_set by (rule sum_closed) lemma prod_closed [intro]: "(\x. x \ X \ f x \ A) \ prod f X \ A" by (induction X rule: infinite_finite_induct) auto lemma Prod_any_closed [intro]: "(\x. f x \ A) \ Prod_any f \ A" unfolding Prod_any.expand_set by (rule prod_closed) lemma prod_fun_closed [intro]: "(\x. f x \ A) \ (\x. g x \ A) \ prod_fun f g x \ A" by (auto simp: prod_fun_def when_def intro!: Sum_any_closed mult_closed) lemma of_nat_closed [simp, intro]: "of_nat n \ A" by (induction n) auto lemma of_int_closed [simp, intro]: "of_int n \ A" by (induction n) auto end locale ring_homomorphism = fixes f :: "'a :: comm_ring_1 \ 'b :: comm_ring_1" assumes add[simp]: "f (x + y) = f x + f y" assumes uminus[simp]: "f (-x) = -f x" assumes mult[simp]: "f (x * y) = f x * f y" assumes zero[simp]: "f 0 = 0" assumes one [simp]: "f 1 = 1" begin lemma diff [simp]: "f (x - y) = f x - f y" using add[of x "-y"] by (simp del: add) lemma power [simp]: "f (x ^ n) = f x ^ n" by (induction n) auto lemma sum [simp]: "f (sum g A) = (\x\A. f (g x))" by (induction A rule: infinite_finite_induct) auto lemma prod [simp]: "f (prod g A) = (\x\A. f (g x))" by (induction A rule: infinite_finite_induct) auto end lemma ring_homomorphism_id [intro]: "ring_homomorphism id" by standard auto lemma ring_homomorphism_id' [intro]: "ring_homomorphism (\x. x)" by standard auto lemma ring_homomorphism_of_int [intro]: "ring_homomorphism of_int" by standard auto subsection \Various facts about multivariate polynomials\ lemma poly_mapping_nat_ge_0 [simp]: "(m :: nat \\<^sub>0 nat) \ 0" proof (cases "m = 0") case False hence "Poly_Mapping.lookup m \ Poly_Mapping.lookup 0" by transfer auto hence "\k. Poly_Mapping.lookup m k \ 0" by (auto simp: fun_eq_iff) from LeastI_ex[OF this] Least_le[of "\k. Poly_Mapping.lookup m k \ 0"] show ?thesis by (force simp: less_eq_poly_mapping_def less_fun_def) qed auto lemma poly_mapping_nat_le_0 [simp]: "(m :: nat \\<^sub>0 nat) \ 0 \ m = 0" unfolding less_eq_poly_mapping_def poly_mapping_eq_iff less_fun_def by auto lemma of_nat_diff_poly_mapping_nat: assumes "m \ n" shows "of_nat (m - n) = (of_nat m - of_nat n :: 'a :: monoid_add \\<^sub>0 nat)" by (auto intro!: poly_mapping_eqI simp: lookup_of_nat lookup_minus when_def) lemma mpoly_coeff_transfer [transfer_rule]: "rel_fun cr_mpoly (=) poly_mapping.lookup MPoly_Type.coeff" unfolding MPoly_Type.coeff_def by transfer_prover lemma mapping_of_sum: "(\x\A. mapping_of (f x)) = mapping_of (sum f A)" by (induction A rule: infinite_finite_induct) (auto simp: plus_mpoly.rep_eq zero_mpoly.rep_eq) lemma mapping_of_eq_0_iff [simp]: "mapping_of p = 0 \ p = 0" by transfer auto lemma Sum_any_mapping_of: "Sum_any (\x. mapping_of (f x)) = mapping_of (Sum_any f)" by (simp add: Sum_any.expand_set mapping_of_sum) lemma Sum_any_parametric_cr_mpoly [transfer_rule]: "(rel_fun (rel_fun (=) cr_mpoly) cr_mpoly) Sum_any Sum_any" by (auto simp: rel_fun_def cr_mpoly_def Sum_any_mapping_of) lemma lookup_mult_of_nat [simp]: "lookup (of_nat n * m) k = n * lookup m k" proof - have "of_nat n * m = (\i k = (\i = n * lookup m k" by simp finally show ?thesis . qed lemma mpoly_eqI: assumes "\mon. MPoly_Type.coeff p mon = MPoly_Type.coeff q mon" shows "p = q" using assms by (transfer, transfer) (auto simp: fun_eq_iff) lemma coeff_mpoly_times: "MPoly_Type.coeff (p * q) mon = prod_fun (MPoly_Type.coeff p) (MPoly_Type.coeff q) mon" by (transfer', transfer') auto lemma (in ring_closed) coeff_mult_closed [intro]: "(\x. coeff p x \ A) \ (\x. coeff q x \ A) \ coeff (p * q) x \ A" by (auto simp: coeff_mpoly_times prod_fun_closed) lemma coeff_notin_vars: assumes "\(keys m \ vars p)" shows "coeff p m = 0" using assms unfolding vars_def by transfer' (auto simp: in_keys_iff) lemma finite_coeff_support [intro]: "finite {m. coeff p m \ 0}" by transfer simp lemma insertion_altdef: "insertion f p = Sum_any (\m. coeff p m * Prod_any (\i. f i ^ lookup m i))" by (transfer', transfer') (simp add: insertion_fun_def) lemma mpoly_coeff_uminus [simp]: "coeff (-p) m = -coeff p m" by transfer auto lemma Sum_any_uminus: "Sum_any (\x. -f x :: 'a :: ab_group_add) = -Sum_any f" by (simp add: Sum_any.expand_set sum_negf) lemma insertion_uminus [simp]: "insertion f (-p :: 'a :: comm_ring_1 mpoly) = -insertion f p" by (simp add: insertion_altdef Sum_any_uminus) lemma Sum_any_lookup: "finite {x. g x \ 0} \ Sum_any (\x. lookup (g x) y) = lookup (Sum_any g) y" by (auto simp: Sum_any.expand_set lookup_sum intro!: sum.mono_neutral_left) lemma Sum_any_diff: assumes "finite {x. f x \ 0}" assumes "finite {x. g x \ 0}" shows "Sum_any (\x. f x - g x :: 'a :: ab_group_add) = Sum_any f - Sum_any g" proof - have "{x. f x - g x \ 0} \ {x. f x \ 0} \ {x. g x \ 0}" by auto moreover have "finite ({x. f x \ 0} \ {x. g x \ 0})" by (subst finite_Un) (insert assms, auto) ultimately have "finite {x. f x - g x \ 0}" by (rule finite_subset) with assms show ?thesis by (simp add: algebra_simps Sum_any.distrib [symmetric]) qed lemma insertion_diff: "insertion f (p - q :: 'a :: comm_ring_1 mpoly) = insertion f p - insertion f q" proof (transfer, transfer) fix f :: "nat \ 'a" and p q :: "(nat \\<^sub>0 nat) \ 'a" assume fin: "finite {x. p x \ 0}" "finite {x. q x \ 0}" have "insertion_fun f (\x. p x - q x) = (\m. p m * (\v. f v ^ lookup m v) - q m * (\v. f v ^ lookup m v))" by (simp add: insertion_fun_def algebra_simps Sum_any_diff) also have "\ = (\m. p m * (\v. f v ^ lookup m v)) - (\m. q m * (\v. f v ^ lookup m v))" by (subst Sum_any_diff) (auto intro: finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)]) also have "\ = insertion_fun f p - insertion_fun f q" by (simp add: insertion_fun_def) finally show "insertion_fun f (\x. p x - q x) = \" . qed lemma insertion_power: "insertion f (p ^ n) = insertion f p ^ n" by (induction n) (simp_all add: insertion_mult) lemma insertion_sum: "insertion f (sum g A) = (\x\A. insertion f (g x))" by (induction A rule: infinite_finite_induct) (auto simp: insertion_add) lemma insertion_prod: "insertion f (prod g A) = (\x\A. insertion f (g x))" by (induction A rule: infinite_finite_induct) (auto simp: insertion_mult) lemma coeff_Var: "coeff (Var i) m = (1 when m = Poly_Mapping.single i 1)" by transfer' (auto simp: Var\<^sub>0_def lookup_single when_def) lemma vars_Var: "vars (Var i :: 'a :: {one,zero} mpoly) = (if (0::'a) = 1 then {} else {i})" unfolding vars_def by (auto simp: Var.rep_eq Var\<^sub>0_def) lemma insertion_Var [simp]: "insertion f (Var i) = f i" proof - have "insertion f (Var i) = (\m. (1 when m = Poly_Mapping.single i 1) * (\i. f i ^ lookup m i))" by (simp add: insertion_altdef coeff_Var) also have "\ = (\j. f j ^ lookup (Poly_Mapping.single i 1) j)" by (subst Sum_any.expand_superset[of "{Poly_Mapping.single i 1}"]) (auto simp: when_def) also have "\ = f i" by (subst Prod_any.expand_superset[of "{i}"]) (auto simp: when_def lookup_single) finally show ?thesis . qed lemma insertion_Sum_any: assumes "finite {x. g x \ 0}" shows "insertion f (Sum_any g) = Sum_any (\x. insertion f (g x))" unfolding Sum_any.expand_set insertion_sum by (intro sum.mono_neutral_right) (auto intro!: finite_subset[OF _ assms]) lemma keys_diff_subset: "keys (f - g) \ keys f \ keys g" by transfer auto lemma keys_empty_iff [simp]: "keys p = {} \ p = 0" by transfer auto lemma mpoly_coeff_0 [simp]: "MPoly_Type.coeff 0 m = 0" by transfer auto lemma lookup_1: "lookup 1 m = (if m = 0 then 1 else 0)" by transfer (simp add: when_def) lemma mpoly_coeff_1: "MPoly_Type.coeff 1 m = (if m = 0 then 1 else 0)" by (simp add: MPoly_Type.coeff_def one_mpoly.rep_eq lookup_1) lemma lookup_Const\<^sub>0: "lookup (Const\<^sub>0 c) m = (if m = 0 then c else 0)" unfolding Const\<^sub>0_def by (simp add: lookup_single when_def) lemma mpoly_coeff_Const: "MPoly_Type.coeff (Const c) m = (if m = 0 then c else 0)" by (simp add: MPoly_Type.coeff_def Const.rep_eq lookup_Const\<^sub>0) lemma coeff_smult [simp]: "coeff (smult c p) m = (c :: 'a :: mult_zero) * coeff p m" by transfer (auto simp: map_lookup) lemma in_keys_mapI: "x \ keys m \ f (lookup m x) \ 0 \ x \ keys (Poly_Mapping.map f m)" by transfer auto lemma keys_uminus [simp]: "keys (-m) = keys m" by transfer auto lemma vars_uminus [simp]: "vars (-p) = vars p" unfolding vars_def by transfer' auto lemma vars_smult: "vars (smult c p) \ vars p" unfolding vars_def by (transfer', transfer') auto lemma vars_0 [simp]: "vars 0 = {}" unfolding vars_def by transfer' simp lemma vars_1 [simp]: "vars 1 = {}" unfolding vars_def by transfer' simp lemma vars_sum: "vars (sum f A) \ (\x\A. vars (f x))" using vars_add by (induction A rule: infinite_finite_induct) auto lemma vars_prod: "vars (prod f A) \ (\x\A. vars (f x))" using vars_mult by (induction A rule: infinite_finite_induct) auto lemma vars_Sum_any: "vars (Sum_any h) \ (\i. vars (h i))" unfolding Sum_any.expand_set by (intro order.trans[OF vars_sum]) auto lemma vars_Prod_any: "vars (Prod_any h) \ (\i. vars (h i))" unfolding Prod_any.expand_set by (intro order.trans[OF vars_prod]) auto lemma vars_power: "vars (p ^ n) \ vars p" using vars_mult by (induction n) auto lemma vars_diff: "vars (p1 - p2) \ vars p1 \ vars p2" unfolding vars_def proof transfer' fix p1 p2 :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "\ (keys ` keys (p1 - p2)) \ \(keys ` (keys p1)) \ \(keys ` (keys p2))" using keys_diff_subset[of p1 p2] by (auto simp flip: not_in_keys_iff_lookup_eq_zero) qed lemma insertion_smult [simp]: "insertion f (smult c p) = c * insertion f p" unfolding insertion_altdef by (subst Sum_any_right_distrib) (auto intro: finite_subset[OF _ finite_coeff_support[of p]] simp: mult.assoc) lemma coeff_add [simp]: "coeff (p + q) m = coeff p m + coeff q m" by transfer' (simp add: lookup_add) lemma coeff_diff [simp]: "coeff (p - q) m = coeff p m - coeff q m" by transfer' (simp add: lookup_minus) lemma insertion_monom [simp]: "insertion f (monom m c) = c * Prod_any (\x. f x ^ lookup m x)" proof - have "insertion f (monom m c) = (\m'. (c when m = m') * (\v. f v ^ lookup m' v))" by (simp add: insertion_def insertion_aux_def insertion_fun_def lookup_single) also have "\ = c * (\v. f v ^ lookup m v)" by (subst Sum_any.expand_superset[of "{m}"]) (auto simp: when_def) finally show ?thesis . qed lemma insertion_aux_Const\<^sub>0 [simp]: "insertion_aux f (Const\<^sub>0 c) = c" proof - have "insertion_aux f (Const\<^sub>0 c) = (\m. (c when m = 0) * (\v. f v ^ lookup m v))" by (simp add: Const\<^sub>0_def insertion_aux_def insertion_fun_def lookup_single) also have "\ = (\m\{0}. (c when m = 0) * (\v. f v ^ lookup m v))" by (intro Sum_any.expand_superset) (auto simp: when_def) also have "\ = c" by simp finally show ?thesis . qed lemma insertion_Const [simp]: "insertion f (Const c) = c" by (simp add: insertion_def Const.rep_eq) lemma coeffs_0 [simp]: "coeffs 0 = {}" by transfer auto lemma coeffs_1 [simp]: "coeffs 1 = {1}" by transfer auto lemma coeffs_Const: "coeffs (Const c) = (if c = 0 then {} else {c})" unfolding Const_def Const\<^sub>0_def by transfer' auto lemma coeffs_subset: "coeffs (Const c) \ {c}" by (auto simp: coeffs_Const) lemma keys_Const\<^sub>0: "keys (Const\<^sub>0 c) = (if c = 0 then {} else {0})" unfolding Const\<^sub>0_def by transfer' auto lemma vars_Const [simp]: "vars (Const c) = {}" unfolding vars_def by transfer' (auto simp: keys_Const\<^sub>0) lemma prod_fun_compose_bij: assumes "bij f" and f: "\x y. f (x + y) = f x + f y" shows "prod_fun m1 m2 (f x) = prod_fun (m1 \ f) (m2 \ f) x" proof - have [simp]: "f x = f y \ x = y" for x y using \bij f\ by (auto dest!: bij_is_inj inj_onD) have "prod_fun (m1 \ f) (m2 \ f) x = Sum_any ((\l. m1 l * Sum_any ((\q. m2 q when f x = l + q) \ f)) \ f)" by (simp add: prod_fun_def f(1) [symmetric] o_def) also have "\ = Sum_any ((\l. m1 l * Sum_any ((\q. m2 q when f x = l + q))))" by (simp only: Sum_any.reindex_cong[OF assms(1) refl, symmetric]) also have "\ = prod_fun m1 m2 (f x)" by (simp add: prod_fun_def) finally show ?thesis .. qed lemma add_nat_poly_mapping_zero_iff [simp]: "(a + b :: 'a \\<^sub>0 nat) = 0 \ a = 0 \ b = 0" by transfer (auto simp: fun_eq_iff) lemma prod_fun_nat_0: fixes f g :: "('a \\<^sub>0 nat) \ 'b::semiring_0" shows "prod_fun f g 0 = f 0 * g 0" proof - have "prod_fun f g 0 = (\l. f l * (\q. g q when 0 = l + q))" unfolding prod_fun_def .. also have "(\l. \q. g q when 0 = l + q) = (\l. \q\{0}. g q when 0 = l + q)" by (intro ext Sum_any.expand_superset) (auto simp: when_def) also have "(\l. f l * \ l) = (\l\{0}. f l * \ l)" by (intro ext Sum_any.expand_superset) (auto simp: when_def) finally show ?thesis by simp qed lemma mpoly_coeff_times_0: "coeff (p * q) 0 = coeff p 0 * coeff q 0" by (simp add: coeff_mpoly_times prod_fun_nat_0) lemma mpoly_coeff_prod_0: "coeff (\x\A. f x) 0 = (\x\A. coeff (f x) 0)" by (induction A rule: infinite_finite_induct) (auto simp: mpoly_coeff_times_0 mpoly_coeff_1) lemma mpoly_coeff_power_0: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induction n) (auto simp: mpoly_coeff_times_0 mpoly_coeff_1) lemma prod_fun_max: fixes f g :: "'a::{linorder, ordered_cancel_comm_monoid_add} \ 'b::semiring_0" assumes zero: "\m. m > a \ f m = 0" "\m. m > b \ g m = 0" assumes fin: "finite {m. f m \ 0}" "finite {m. g m \ 0}" shows "prod_fun f g (a + b) = f a * g b" proof - note fin' = finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)] have "prod_fun f g (a + b) = (\l. f l * (\q. g q when a + b = l + q))" by (simp add: prod_fun_def Sum_any_right_distrib) also have "\ = (\l. \q. f l * g q when a + b = l + q)" by (subst Sum_any_right_distrib) (auto intro!: Sum_any.cong fin'(2) simp: when_def) also { fix l q assume lq: "a + b = l + q" "(a, b) \ (l, q)" and nz: "f l * g q \ 0" from nz and zero have "l \ a" "q \ b" by (auto intro: leI) moreover from this and lq(2) have "l < a \ q < b" by auto ultimately have "l + q < a + b" by (auto intro: add_less_le_mono add_le_less_mono) with lq(1) have False by simp } hence "(\l. \q. f l * g q when a + b = l + q) = (\l. \q. f l * g q when (a, b) = (l, q))" by (intro Sum_any.cong refl) (auto simp: when_def) also have "\ = (\(l,q). f l * g q when (a, b) = (l, q))" by (intro Sum_any.cartesian_product[of "{(a, b)}"]) auto also have "\ = (\(l,q)\{(a,b)}. f l * g q when (a, b) = (l, q))" by (intro Sum_any.expand_superset) auto also have "\ = f a * g b" by simp finally show ?thesis . qed lemma prod_fun_gt_max_eq_zero: fixes f g :: "'a::{linorder, ordered_cancel_comm_monoid_add} \ 'b::semiring_0" assumes "m > a + b" assumes zero: "\m. m > a \ f m = 0" "\m. m > b \ g m = 0" assumes fin: "finite {m. f m \ 0}" "finite {m. g m \ 0}" shows "prod_fun f g m = 0" proof - note fin' = finite_subset[OF _ fin(1)] finite_subset[OF _ fin(2)] have "prod_fun f g m = (\l. f l * (\q. g q when m = l + q))" by (simp add: prod_fun_def Sum_any_right_distrib) also have "\ = (\l. \q. f l * g q when m = l + q)" by (subst Sum_any_right_distrib) (auto intro!: Sum_any.cong fin'(2) simp: when_def) also { fix l q assume lq: "m = l + q" and nz: "f l * g q \ 0" from nz and zero have "l \ a" "q \ b" by (auto intro: leI) hence "l + q \ a + b" by (intro add_mono) also have "\ < m" by fact finally have "l + q < m" . } hence "(\l. \q. f l * g q when m = l + q) = (\l. \q. f l * g q when False)" by (intro Sum_any.cong refl) (auto simp: when_def) also have "\ = 0" by simp finally show ?thesis . qed subsection \Restricting a monomial to a subset of variables\ lift_definition restrictpm :: "'a set \ ('a \\<^sub>0 'b :: zero) \ ('a \\<^sub>0 'b)" is "\A f x. if x \ A then f x else 0" by (erule finite_subset[rotated]) auto lemma lookup_restrictpm: "lookup (restrictpm A m) x = (if x \ A then lookup m x else 0)" by transfer auto lemma lookup_restrictpm_in [simp]: "x \ A \ lookup (restrictpm A m) x = lookup m x" and lookup_restrict_pm_not_in [simp]: "x \ A \ lookup (restrictpm A m) x = 0" by (simp_all add: lookup_restrictpm) lemma keys_restrictpm [simp]: "keys (restrictpm A m) = keys m \ A" by transfer auto lemma restrictpm_add: "restrictpm X (m1 + m2) = restrictpm X m1 + restrictpm X m2" by transfer auto lemma restrictpm_id [simp]: "keys m \ X \ restrictpm X m = m" by transfer (auto simp: fun_eq_iff) lemma restrictpm_orthogonal [simp]: "keys m \ -X \ restrictpm X m = 0" by transfer (auto simp: fun_eq_iff) lemma restrictpm_add_disjoint: "X \ Y = {} \ restrictpm X m + restrictpm Y m = restrictpm (X \ Y) m" by transfer (auto simp: fun_eq_iff) lemma restrictpm_add_complements: "restrictpm X m + restrictpm (-X) m = m" "restrictpm (-X) m + restrictpm X m = m" by (subst restrictpm_add_disjoint; force)+ subsection \Mapping over a polynomial\ lift_definition map_mpoly :: "('a :: zero \ 'b :: zero) \ 'a mpoly \ 'b mpoly" is "\(f :: 'a \ 'b) (p :: (nat \\<^sub>0 nat) \\<^sub>0 'a). Poly_Mapping.map f p" . lift_definition mapm_mpoly :: "((nat \\<^sub>0 nat) \ 'a :: zero \ 'b :: zero) \ 'a mpoly \ 'b mpoly" is "\(f :: (nat \\<^sub>0 nat) \ 'a \ 'b) (p :: (nat \\<^sub>0 nat) \\<^sub>0 'a). Poly_Mapping.mapp f p" . lemma poly_mapping_map_conv_mapp: "Poly_Mapping.map f = Poly_Mapping.mapp (\_. f)" by (auto simp: Poly_Mapping.mapp_def Poly_Mapping.map_def map_fun_def o_def fun_eq_iff when_def in_keys_iff cong: if_cong) lemma map_mpoly_conv_mapm_mpoly: "map_mpoly f = mapm_mpoly (\_. f)" by transfer' (auto simp: poly_mapping_map_conv_mapp) lemma map_mpoly_comp: "f 0 = 0 \ map_mpoly f (map_mpoly g p) = map_mpoly (f \ g) p" by (transfer', transfer') (auto simp: when_def fun_eq_iff) lemma mapp_mapp: "(\x. f x 0 = 0) \ Poly_Mapping.mapp f (Poly_Mapping.mapp g m) = Poly_Mapping.mapp (\x y. f x (g x y)) m" by transfer' (auto simp: fun_eq_iff lookup_mapp in_keys_iff) lemma mapm_mpoly_comp: "(\x. f x 0 = 0) \ mapm_mpoly f (mapm_mpoly g p) = mapm_mpoly (\m c. f m (g m c)) p" by transfer' (simp add: mapp_mapp) lemma coeff_map_mpoly: "coeff (map_mpoly f p) m = (if coeff p m = 0 then 0 else f (coeff p m))" by (transfer, transfer) auto lemma coeff_map_mpoly' [simp]: "f 0 = 0 \ coeff (map_mpoly f p) m = f (coeff p m)" by (subst coeff_map_mpoly) auto lemma coeff_mapm_mpoly: "coeff (mapm_mpoly f p) m = (if coeff p m = 0 then 0 else f m (coeff p m))" by (transfer, transfer') (auto simp: in_keys_iff) lemma coeff_mapm_mpoly' [simp]: "(\m. f m 0 = 0) \ coeff (mapm_mpoly f p) m = f m (coeff p m)" by (subst coeff_mapm_mpoly) auto lemma vars_map_mpoly_subset: "vars (map_mpoly f p) \ vars p" unfolding vars_def by (transfer', transfer') (auto simp: map_mpoly.rep_eq) lemma coeff_sum [simp]: "coeff (sum f A) m = (\x\A. coeff (f x) m)" by (induction A rule: infinite_finite_induct) auto lemma coeff_Sum_any: "finite {x. f x \ 0} \ coeff (Sum_any f) m = Sum_any (\x. coeff (f x) m)" by (auto simp add: Sum_any.expand_set intro!: sum.mono_neutral_right) lemma Sum_any_zeroI: "(\x. f x = 0) \ Sum_any f = 0" by (auto simp: Sum_any.expand_set) lemma insertion_Prod_any: "finite {x. g x \ 1} \ insertion f (Prod_any g) = Prod_any (\x. insertion f (g x))" by (auto simp: Prod_any.expand_set insertion_prod intro!: prod.mono_neutral_right) lemma insertion_insertion: "insertion g (insertion k p) = insertion (\x. insertion g (k x)) (map_mpoly (insertion g) p)" (is "?lhs = ?rhs") proof - have "insertion g (insertion k p) = (\x. insertion g (coeff p x) * insertion g (\i. k i ^ lookup x i))" unfolding insertion_altdef[of k p] by (subst insertion_Sum_any) (auto intro: finite_subset[OF _ finite_coeff_support[of p]] simp: insertion_mult) also have "\ = (\x. insertion g (coeff p x) * (\i. insertion g (k i) ^ lookup x i))" proof (intro Sum_any.cong) fix x show "insertion g (coeff p x) * insertion g (\i. k i ^ lookup x i) = insertion g (coeff p x) * (\i. insertion g (k i) ^ lookup x i)" by (subst insertion_Prod_any) (auto simp: insertion_power intro!: finite_subset[OF _ finite_lookup[of x]] Nat.gr0I) qed also have "\ = insertion (\x. insertion g (k x)) (map_mpoly (insertion g) p)" unfolding insertion_altdef[of _ "map_mpoly f p" for f] by auto finally show ?thesis . qed lemma insertion_substitute_linear: "insertion (\i. c i * f i) p = insertion f (mapm_mpoly (\m d. Prod_any (\i. c i ^ lookup m i) * d) p)" unfolding insertion_altdef proof (intro Sum_any.cong, goal_cases) case (1 m) have "coeff (mapm_mpoly (\m. (*) (\i. c i ^ lookup m i)) p) m * (\i. f i ^ lookup m i) = MPoly_Type.coeff p m * ((\i. c i ^ lookup m i) * (\i. f i ^ lookup m i))" by (simp add: mult_ac) also have "(\i. c i ^ lookup m i) * (\i. f i ^ lookup m i) = (\i. (c i * f i) ^ lookup m i)" by (subst Prod_any.distrib [symmetric]) (auto simp: power_mult_distrib intro!: finite_subset[OF _ finite_lookup[of m]] Nat.gr0I) finally show ?case by simp qed lemma vars_mapm_mpoly_subset: "vars (mapm_mpoly f p) \ vars p" unfolding vars_def using keys_mapp_subset[of f] by (auto simp: mapm_mpoly.rep_eq) lemma map_mpoly_cong: assumes "\m. f (coeff p m) = g (coeff p m)" "p = q" shows "map_mpoly f p = map_mpoly g q" using assms by (intro mpoly_eqI) (auto simp: coeff_map_mpoly) subsection \The leading monomial and leading coefficient\ text \ The leading monomial of a multivariate polynomial is the one with the largest monomial w.\,r.\,t.\ the monomial ordering induced by the standard variable ordering. The leading coefficient is the coefficient of the leading monomial. As a convention, the leading monomial of the zero polynomial is defined to be the same as that of any non-constant zero polynomial, i.\,e.\ the monomial $X_1^0 \ldots X_n^0$. \ lift_definition lead_monom :: "'a :: zero mpoly \ (nat \\<^sub>0 nat)" is "\f :: (nat \\<^sub>0 nat) \\<^sub>0 'a. Max (insert 0 (keys f))" . lemma lead_monom_geI [intro]: assumes "coeff p m \ 0" shows "m \ lead_monom p" using assms by (auto simp: lead_monom_def coeff_def in_keys_iff) lemma coeff_gt_lead_monom_zero [simp]: assumes "m > lead_monom p" shows "coeff p m = 0" using lead_monom_geI[of p m] assms by force lemma lead_monom_nonzero_eq: assumes "p \ 0" shows "lead_monom p = Max (keys (mapping_of p))" using assms by transfer (simp add: max_def) lemma lead_monom_0 [simp]: "lead_monom 0 = 0" by (simp add: lead_monom_def zero_mpoly.rep_eq) lemma lead_monom_1 [simp]: "lead_monom 1 = 0" by (simp add: lead_monom_def one_mpoly.rep_eq) lemma lead_monom_Const [simp]: "lead_monom (Const c) = 0" by (simp add: lead_monom_def Const.rep_eq Const\<^sub>0_def) lemma lead_monom_uminus [simp]: "lead_monom (-p) = lead_monom p" by (simp add: lead_monom_def uminus_mpoly.rep_eq) lemma keys_mult_const [simp]: fixes c :: "'a :: {semiring_0, semiring_no_zero_divisors}" assumes "c \ 0" shows "keys (Poly_Mapping.map ((*) c) p) = keys p" using assms by transfer auto lemma lead_monom_eq_0_iff: "lead_monom p = 0 \ vars p = {}" unfolding vars_def by transfer' (auto simp: Max_eq_iff) lemma lead_monom_monom: "lead_monom (monom m c) = (if c = 0 then 0 else m)" by (auto simp add: lead_monom_def monom.rep_eq Const\<^sub>0_def max_def ) lemma lead_monom_monom' [simp]: "c \ 0 \ lead_monom (monom m c) = m" by (simp add: lead_monom_monom) lemma lead_monom_numeral [simp]: "lead_monom (numeral n) = 0" unfolding monom_numeral[symmetric] by (subst lead_monom_monom) auto lemma lead_monom_add: "lead_monom (p + q) \ max (lead_monom p) (lead_monom q)" proof transfer fix p q :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "Max (insert 0 (keys (p + q))) \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof (rule Max.boundedI) fix m assume m: "m \ insert 0 (keys (p + q))" thus "m \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof assume "m \ keys (p + q)" with keys_add[of p q] have "m \ keys p \ m \ keys q" by (auto simp: in_keys_iff plus_poly_mapping.rep_eq) thus ?thesis by (auto simp: le_max_iff_disj) qed auto qed auto qed lemma lead_monom_diff: "lead_monom (p - q) \ max (lead_monom p) (lead_monom q)" proof transfer fix p q :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" show "Max (insert 0 (keys (p - q))) \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof (rule Max.boundedI) fix m assume m: "m \ insert 0 (keys (p - q))" thus "m \ max (Max (insert 0 (keys p))) (Max (insert 0 (keys q)))" proof assume "m \ keys (p - q)" with keys_diff_subset[of p q] have "m \ keys p \ m \ keys q" by auto thus ?thesis by (auto simp: le_max_iff_disj) qed auto qed auto qed definition lead_coeff where "lead_coeff p = coeff p (lead_monom p)" lemma vars_empty_iff: "vars p = {} \ p = Const (lead_coeff p)" proof assume "vars p = {}" hence [simp]: "lead_monom p = 0" by (simp add: lead_monom_eq_0_iff) have [simp]: "mon \ 0 \ (mon > (0 :: nat \\<^sub>0 nat))" for mon by (auto simp: order.strict_iff_order) thus "p = Const (lead_coeff p)" by (intro mpoly_eqI) (auto simp: mpoly_coeff_Const lead_coeff_def) next assume "p = Const (lead_coeff p)" also have "vars \ = {}" by simp finally show "vars p = {}" . qed lemma lead_coeff_0 [simp]: "lead_coeff 0 = 0" by (simp add: lead_coeff_def) lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by (simp add: lead_coeff_def mpoly_coeff_1) lemma lead_coeff_Const [simp]: "lead_coeff (Const c) = c" by (simp add: lead_coeff_def mpoly_coeff_Const) lemma lead_coeff_monom [simp]: "lead_coeff (monom p c) = c" by (simp add: lead_coeff_def coeff_monom when_def lead_monom_monom) lemma lead_coeff_nonzero [simp]: "p \ 0 \ lead_coeff p \ 0" unfolding lead_coeff_def lead_monom_def by (cases "keys (mapping_of p) = {}") (auto simp: coeff_def max_def) lemma fixes c :: "'a :: semiring_0" assumes "c * lead_coeff p \ 0" shows lead_monom_smult [simp]: "lead_monom (smult c p) = lead_monom p" and lead_coeff_smult [simp]: "lead_coeff (smult c p) = c * lead_coeff p" proof - from assms have *: "keys (mapping_of p) \ {}" by auto from assms have "coeff (MPoly_Type.smult c p) (lead_monom p) \ 0" by (simp add: lead_coeff_def) hence smult_nz: "MPoly_Type.smult c p \ 0" by (auto simp del: coeff_smult) with assms have **: "keys (mapping_of (smult c p)) \ {}" by simp have "Max (keys (mapping_of (smult c p))) = Max (keys (mapping_of p))" proof (safe intro!: antisym Max.coboundedI) have "lookup (mapping_of p) (Max (keys (mapping_of p))) = lead_coeff p" using * by (simp add: lead_coeff_def lead_monom_def max_def coeff_def) with assms show "Max (keys (mapping_of p)) \ keys (mapping_of (smult c p))" using * by (auto simp: smult.rep_eq intro!: in_keys_mapI) from smult_nz have "lead_coeff (smult c p) \ 0" by (intro lead_coeff_nonzero) auto hence "coeff p (Max (keys (mapping_of (smult c p)))) \ 0" using assms * ** by (auto simp: lead_coeff_def lead_monom_def max_def) thus "Max (keys (mapping_of (smult c p))) \ keys (mapping_of p)" by (auto simp: smult.rep_eq coeff_def in_keys_iff) qed auto with * ** show "lead_monom (smult c p) = lead_monom p" by (simp add: lead_monom_def max_def) thus "lead_coeff (smult c p) = c * lead_coeff p" by (simp add: lead_coeff_def) qed lemma lead_coeff_mult_aux: "coeff (p * q) (lead_monom p + lead_monom q) = lead_coeff p * lead_coeff q" proof (cases "p = 0 \ q = 0") case False define a b where "a = lead_monom p" and "b = lead_monom q" have "coeff (p * q) (a + b) = coeff p a * coeff q b" unfolding coeff_mpoly_times by (rule prod_fun_max) (insert False, auto simp: a_def b_def) thus ?thesis by (simp add: a_def b_def lead_coeff_def) qed auto lemma lead_monom_mult_le: "lead_monom (p * q) \ lead_monom p + lead_monom q" proof (cases "p * q = 0") case False show ?thesis proof (intro leI notI) assume "lead_monom p + lead_monom q < lead_monom (p * q)" hence "lead_coeff (p * q) = 0" unfolding lead_coeff_def coeff_mpoly_times by (rule prod_fun_gt_max_eq_zero) auto with False show False by simp qed qed auto lemma lead_monom_mult: assumes "lead_coeff p * lead_coeff q \ 0" shows "lead_monom (p * q) = lead_monom p + lead_monom q" by (intro antisym lead_monom_mult_le lead_monom_geI) (insert assms, auto simp: lead_coeff_mult_aux) lemma lead_coeff_mult: assumes "lead_coeff p * lead_coeff q \ 0" shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q" using assms by (simp add: lead_monom_mult lead_coeff_mult_aux lead_coeff_def) lemma keys_lead_monom_subset: "keys (lead_monom p) \ vars p" proof (cases "p = 0") case False hence "lead_coeff p \ 0" by simp hence "coeff p (lead_monom p) \ 0" unfolding lead_coeff_def . thus ?thesis unfolding vars_def by transfer' (auto simp: max_def in_keys_iff) qed auto lemma assumes "(\i\A. lead_coeff (f i)) \ 0" shows lead_monom_prod: "lead_monom (\i\A. f i) = (\i\A. lead_monom (f i))" (is ?th1) and lead_coeff_prod: "lead_coeff (\i\A. f i) = (\i\A. lead_coeff (f i))" (is ?th2) proof - have "?th1 \ ?th2" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) from insert have nz: "lead_coeff (f x) \ 0" "(\i\A. lead_coeff (f i)) \ 0" by auto note IH = insert.IH[OF this(2)] from insert have nz': "lead_coeff (f x) * lead_coeff (\i\A. f i) \ 0" by (subst IH) auto from insert.prems insert.hyps nz nz' show ?case by (auto simp: lead_monom_mult lead_coeff_mult IH) qed auto thus ?th1 ?th2 by blast+ qed lemma lead_monom_sum_le: "(\x. x \ X \ lead_monom (h x) \ ub) \ lead_monom (sum h X) \ ub" by (induction X rule: infinite_finite_induct) (auto intro!: order.trans[OF lead_monom_add]) text \ The leading monomial of a sum where the leading monomial the summands are distinct is simply the maximum of the leading monomials. \ lemma lead_monom_sum: assumes "inj_on (lead_monom \ h) X" and "finite X" and "X \ {}" and "\x. x \ X \ h x \ 0" defines "m \ Max ((lead_monom \ h) ` X)" shows "lead_monom (\x\X. h x) = m" proof (rule antisym) show "lead_monom (sum h X) \ m" unfolding m_def using assms by (intro lead_monom_sum_le Max_ge finite_imageI) auto next from assms have "m \ (lead_monom \ h) ` X" unfolding m_def by (intro Max_in finite_imageI) auto then obtain x where x: "x \ X" "m = lead_monom (h x)" by auto have "coeff (\x\X. h x) m = (\x\X. coeff (h x) m)" by simp also have "\ = (\x\{x}. coeff (h x) m)" proof (intro sum.mono_neutral_right ballI) fix y assume y: "y \ X - {x}" hence "(lead_monom \ h) y \ m" using assms unfolding m_def by (intro Max_ge finite_imageI) auto moreover have "(lead_monom \ h) y \ (lead_monom \ h) x" using \x \ X\ y inj_onD[OF assms(1), of x y] by auto ultimately have "lead_monom (h y) < m" using x by auto thus "coeff (h y) m = 0" by simp qed (insert x assms, auto) also have "\ = coeff (h x) m" by simp also have "\ = lead_coeff (h x)" using x by (simp add: lead_coeff_def) also have "\ \ 0" using assms x by auto finally show "lead_monom (sum h X) \ m" by (intro lead_monom_geI) qed lemma lead_coeff_eq_0_iff [simp]: "lead_coeff p = 0 \ p = 0" by (cases "p = 0") auto lemma fixes f :: "_ \ 'a :: semidom mpoly" assumes "\i. i \ A \ f i \ 0" shows lead_monom_prod' [simp]: "lead_monom (\i\A. f i) = (\i\A. lead_monom (f i))" (is ?th1) and lead_coeff_prod' [simp]: "lead_coeff (\i\A. f i) = (\i\A. lead_coeff (f i))" (is ?th2) proof - from assms have "(\i\A. lead_coeff (f i)) \ 0" by (cases "finite A") auto thus ?th1 ?th2 by (simp_all add: lead_monom_prod lead_coeff_prod) qed lemma fixes p :: "'a :: comm_semiring_1 mpoly" assumes "lead_coeff p ^ n \ 0" shows lead_monom_power: "lead_monom (p ^ n) = of_nat n * lead_monom p" and lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" using assms lead_monom_prod[of "\_. p" "{.._. p" "{.. 0" shows lead_monom_power' [simp]: "lead_monom (p ^ n) = of_nat n * lead_monom p" and lead_coeff_power' [simp]: "lead_coeff (p ^ n) = lead_coeff p ^ n" using assms lead_monom_prod'[of "{.._. p"] lead_coeff_prod'[of "{.._. p"] by simp_all subsection \Turning a set of variables into a monomial\ text \ Given a finite set $\{X_1,\ldots,X_n\}$ of variables, the following is the monomial $X_1\ldots X_n$: \ lift_definition monom_of_set :: "nat set \ (nat \\<^sub>0 nat)" is "\X x. if finite X \ x \ X then 1 else 0" by auto lemma lookup_monom_of_set: "Poly_Mapping.lookup (monom_of_set X) i = (if finite X \ i \ X then 1 else 0)" by transfer auto lemma lookup_monom_of_set_1 [simp]: "finite X \ i \ X \ Poly_Mapping.lookup (monom_of_set X) i = 1" and lookup_monom_of_set_0 [simp]: "i \ X \ Poly_Mapping.lookup (monom_of_set X) i = 0" by (simp_all add: lookup_monom_of_set) lemma keys_monom_of_set: "keys (monom_of_set X) = (if finite X then X else {})" by transfer auto lemma keys_monom_of_set_finite [simp]: "finite X \ keys (monom_of_set X) = X" by (simp add: keys_monom_of_set) lemma monom_of_set_eq_iff [simp]: "finite X \ finite Y \ monom_of_set X = monom_of_set Y \ X = Y" by transfer (auto simp: fun_eq_iff) lemma monom_of_set_empty [simp]: "monom_of_set {} = 0" by transfer auto lemma monom_of_set_eq_zero_iff [simp]: "monom_of_set X = 0 \ infinite X \ X = {}" by transfer (auto simp: fun_eq_iff) lemma zero_eq_monom_of_set_iff [simp]: "0 = monom_of_set X \ infinite X \ X = {}" by transfer (auto simp: fun_eq_iff) subsection \Permuting the variables of a polynomial\ text \ Next, we define the operation of permuting the variables of a monomial and polynomial. \ lift_definition permutep :: "('a \ 'a) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b :: zero)" is "\f p. if bij f then p \ f else p" proof - fix f :: "'a \ 'a" and g :: "'a \ 'b" assume *: "finite {x. g x \ 0}" show "finite {x. (if bij f then g \ f else g) x \ 0}" proof (cases "bij f") case True with * have "finite (f -` {x. g x \ 0})" by (intro finite_vimageI) (auto dest: bij_is_inj) with True show ?thesis by auto qed (use * in auto) qed lift_definition mpoly_map_vars :: "(nat \ nat) \ 'a :: zero mpoly \ 'a mpoly" is "\f p. permutep (permutep f) p" . lemma keys_permutep: "bij f \ keys (permutep f m) = f -` keys m" by transfer auto lemma permutep_id'' [simp]: "permutep id = id" by transfer' (auto simp: fun_eq_iff) lemma permutep_id''' [simp]: "permutep (\x. x) = id" by transfer' (auto simp: fun_eq_iff) lemma permutep_0 [simp]: "permutep f 0 = 0" by transfer auto lemma permutep_single: "bij f \ permutep f (Poly_Mapping.single a b) = Poly_Mapping.single (inv_into UNIV f a) b" by transfer (auto simp: fun_eq_iff when_def inv_f_f surj_f_inv_f bij_is_inj bij_is_surj) lemma mpoly_map_vars_id [simp]: "mpoly_map_vars id = id" by transfer auto lemma mpoly_map_vars_id' [simp]: "mpoly_map_vars (\x. x) = id" by transfer auto lemma lookup_permutep: "Poly_Mapping.lookup (permutep f m) x = (if bij f then Poly_Mapping.lookup m (f x) else Poly_Mapping.lookup m x)" by transfer auto lemma inj_permutep [intro]: "inj (permutep (f :: 'a \ 'a) :: _ \ 'a \\<^sub>0 'b :: zero)" unfolding inj_def proof (transfer, safe) fix f :: "'a \ 'a" and x y :: "'a \ 'b" assume eq: "(if bij f then x \ f else x) = (if bij f then y \ f else y)" show "x = y" proof (cases "bij f") case True show ?thesis proof fix t :: 'a from \bij f\ obtain s where "t = f s" by (auto dest!: bij_is_surj) with eq and True show "x t = y t" by (auto simp: fun_eq_iff) qed qed (use eq in auto) qed lemma surj_permutep [intro]: "surj (permutep (f :: 'a \ 'a) :: _ \ 'a \\<^sub>0 'b :: zero)" unfolding surj_def proof (transfer, safe) fix f :: "'a \ 'a" and y :: "'a \ 'b" assume fin: "finite {t. y t \ 0}" show "\x\{f. finite {x. f x \ 0}}. y = (if bij f then x \ f else x)" proof (cases "bij f") case True with fin have "finite (the_inv f -` {t. y t \ 0})" by (intro finite_vimageI) (auto simp: bij_is_inj bij_betw_the_inv_into) moreover have "y \ the_inv f \ f = y" using True by (simp add: fun_eq_iff the_inv_f_f bij_is_inj) ultimately show ?thesis by (intro bexI[of _ "y \ the_inv f"]) (auto simp: True) qed (use fin in auto) qed lemma bij_permutep [intro]: "bij (permutep f)" using inj_permutep[of f] surj_permutep[of f] by (simp add: bij_def) lemma mpoly_map_vars_map_mpoly: "mpoly_map_vars f (map_mpoly g p) = map_mpoly g (mpoly_map_vars f p)" by (transfer', transfer') (auto simp: fun_eq_iff) lemma coeff_mpoly_map_vars: fixes f :: "nat \ nat" and p :: "'a :: zero mpoly" assumes "bij f" shows "MPoly_Type.coeff (mpoly_map_vars f p) mon = MPoly_Type.coeff p (permutep f mon)" using assms by transfer' (simp add: lookup_permutep bij_permutep) lemma permutep_monom_of_set: assumes "bij f" shows "permutep f (monom_of_set A) = monom_of_set (f -` A)" using assms by transfer (auto simp: fun_eq_iff bij_is_inj finite_vimage_iff) lemma permutep_comp: "bij f \ bij g \ permutep (f \ g) = permutep g \ permutep f" by transfer' (auto simp: fun_eq_iff bij_comp) lemma permutep_comp': "bij f \ bij g \ permutep (f \ g) mon = permutep g (permutep f mon)" by transfer (auto simp: fun_eq_iff bij_comp) lemma mpoly_map_vars_comp: "bij f \ bij g \ mpoly_map_vars f (mpoly_map_vars g p) = mpoly_map_vars (f \ g) p" by transfer (auto simp: bij_permutep permutep_comp) lemma permutep_id [simp]: "permutep id mon = mon" by transfer auto lemma permutep_id' [simp]: "permutep (\x. x) mon = mon" by transfer auto lemma inv_permutep [simp]: fixes f :: "'a \ 'a" assumes "bij f" shows "inv_into UNIV (permutep f) = permutep (inv_into UNIV f)" proof fix m :: "'a \\<^sub>0 'b" show "inv_into UNIV (permutep f) m = permutep (inv_into UNIV f) m" using permutep_comp'[of "inv_into UNIV f" f m] assms inj_iff[of f] by (intro inv_f_eq) (auto simp: bij_imp_bij_inv bij_is_inj) qed lemma mpoly_map_vars_monom: "bij f \ mpoly_map_vars f (monom m c) = monom (permutep (inv_into UNIV f) m) c" by transfer' (simp add: permutep_single bij_permutep) lemma vars_mpoly_map_vars: fixes f :: "nat \ nat" and p :: "'a :: zero mpoly" assumes "bij f" shows "vars (mpoly_map_vars f p) = f ` vars p" using assms unfolding vars_def proof transfer' fix f :: "nat \ nat" and p :: "(nat \\<^sub>0 nat) \\<^sub>0 'a" assume f: "bij f" have eq: "f (inv_into UNIV f x) = x" for x using f by (subst surj_f_inv_f[of f]) (auto simp: bij_is_surj) show "\ (keys ` keys (permutep (permutep f) p)) = f ` \ (keys ` keys p)" proof safe fix m x assume mx: "m \ keys (permutep (permutep f) p)" "x \ keys m" from mx have "permutep f m \ keys p" by (auto simp: keys_permutep bij_permutep f) with mx have "f (inv_into UNIV f x) \ f ` (\m\keys p. keys m)" by (intro imageI) (auto intro!: bexI[of _ "permutep f m"] simp: keys_permutep f eq) with eq show "x \ f ` (\m\keys p. keys m)" by simp next fix m x assume mx: "m \ keys p" "x \ keys m" from mx have "permutep id m \ keys p" by simp also have "id = inv_into UNIV f \ f" using f by (intro ext) (auto simp: bij_is_inj inv_f_f) also have "permutep \ m = permutep f (permutep (inv_into UNIV f) m)" by (simp add: permutep_comp f bij_imp_bij_inv) finally have **: "permutep f (permutep (inv_into UNIV f) m) \ keys p" . moreover from f mx have "f x \ keys (permutep (inv_into UNIV f) m)" by (auto simp: keys_permutep bij_imp_bij_inv inv_f_f bij_is_inj) ultimately show "f x \ \ (keys ` keys (permutep (permutep f) p))" using f by (auto simp: keys_permutep bij_permutep) qed qed lemma permutep_eq_monom_of_set_iff [simp]: assumes "bij f" shows "permutep f mon = monom_of_set A \ mon = monom_of_set (f ` A)" proof assume eq: "permutep f mon = monom_of_set A" have "permutep (inv_into UNIV f) (permutep f mon) = monom_of_set (inv_into UNIV f -` A)" using assms by (simp add: eq bij_imp_bij_inv assms permutep_monom_of_set) also have "inv_into UNIV f -` A = f ` A" using assms by (force simp: bij_is_surj image_iff inv_f_f bij_is_inj surj_f_inv_f) also have "permutep (inv_into UNIV f) (permutep f mon) = permutep (f \ inv_into UNIV f) mon" using assms by (simp add: permutep_comp bij_imp_bij_inv) also have "f \ inv_into UNIV f = id" by (subst surj_iff [symmetric]) (insert assms, auto simp: bij_is_surj) finally show "mon = monom_of_set (f ` A)" by simp qed (insert assms, auto simp: permutep_monom_of_set inj_vimage_image_eq bij_is_inj) lemma permutep_monom_of_set_permutes [simp]: assumes "\ permutes A" shows "permutep \ (monom_of_set A) = monom_of_set A" using assms by transfer (auto simp: if_splits fun_eq_iff permutes_in_image) lemma mpoly_map_vars_0 [simp]: "mpoly_map_vars f 0 = 0" by (transfer, transfer') (simp add: o_def) lemma permutep_eq_0_iff [simp]: "permutep f m = 0 \ m = 0" proof transfer fix f :: "'a \ 'a" and m :: "'a \ 'b" assume "finite {x. m x \ 0}" show "((if bij f then m \ f else m) = (\k. 0)) = (m = (\k. 0))" proof (cases "bij f") case True hence "(\x. m (f x) = 0) \ (\x. m x = 0)" using bij_iff[of f] by metis with True show ?thesis by (auto simp: fun_eq_iff) qed (auto simp: fun_eq_iff) qed lemma mpoly_map_vars_1 [simp]: "mpoly_map_vars f 1 = 1" by (transfer, transfer') (auto simp: o_def fun_eq_iff when_def) lemma permutep_Const\<^sub>0 [simp]: "(\x. f x = 0 \ x = 0) \ permutep f (Const\<^sub>0 c) = Const\<^sub>0 c" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_add [simp]: "permutep f (m1 + m2) = permutep f m1 + permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_diff [simp]: "permutep f (m1 - m2) = permutep f m1 - permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_uminus [simp]: "permutep f (-m) = -permutep f m" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff) lemma permutep_mult [simp]: "(\x y. f (x + y) = f x + f y) \ permutep f (m1 * m2) = permutep f m1 * permutep f m2" unfolding Const\<^sub>0_def by transfer' (auto simp: when_def fun_eq_iff prod_fun_compose_bij) lemma mpoly_map_vars_Const [simp]: "mpoly_map_vars f (Const c) = Const c" by transfer (auto simp: o_def fun_eq_iff when_def) lemma mpoly_map_vars_add [simp]: "mpoly_map_vars f (p + q) = mpoly_map_vars f p + mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_diff [simp]: "mpoly_map_vars f (p - q) = mpoly_map_vars f p - mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_uminus [simp]: "mpoly_map_vars f (-p) = -mpoly_map_vars f p" by transfer simp lemma permutep_smult: "permutep (permutep f) (Poly_Mapping.map ((*) c) p) = Poly_Mapping.map ((*) c) (permutep (permutep f) p)" by transfer' (auto split: if_splits simp: fun_eq_iff) lemma mpoly_map_vars_smult [simp]: "mpoly_map_vars f (smult c p) = smult c (mpoly_map_vars f p)" by transfer (simp add: permutep_smult) lemma mpoly_map_vars_mult [simp]: "mpoly_map_vars f (p * q) = mpoly_map_vars f p * mpoly_map_vars f q" by transfer simp lemma mpoly_map_vars_sum [simp]: "mpoly_map_vars f (sum g A) = (\x\A. mpoly_map_vars f (g x))" by (induction A rule: infinite_finite_induct) auto lemma mpoly_map_vars_prod [simp]: "mpoly_map_vars f (prod g A) = (\x\A. mpoly_map_vars f (g x))" by (induction A rule: infinite_finite_induct) auto lemma mpoly_map_vars_eq_0_iff [simp]: "mpoly_map_vars f p = 0 \ p = 0" by transfer auto lemma permutep_eq_iff [simp]: "permutep f p = permutep f q \ p = q" by transfer (auto simp: comp_bij_eq_iff) lemma mpoly_map_vars_Sum_any [simp]: "mpoly_map_vars f (Sum_any g) = Sum_any (\x. mpoly_map_vars f (g x))" by (simp add: Sum_any.expand_set) lemma mpoly_map_vars_power [simp]: "mpoly_map_vars f (p ^ n) = mpoly_map_vars f p ^ n" by (induction n) auto lemma mpoly_map_vars_monom_single [simp]: assumes "bij f" shows "mpoly_map_vars f (monom (Poly_Mapping.single i n) c) = monom (Poly_Mapping.single (f i) n) c" using assms by (simp add: mpoly_map_vars_monom permutep_single bij_imp_bij_inv inv_inv_eq) lemma insertion_mpoly_map_vars: assumes "bij f" shows "insertion g (mpoly_map_vars f p) = insertion (g \ f) p" proof - have "insertion g (mpoly_map_vars f p) = (\m. coeff p (permutep f m) * (\i. g i ^ lookup m i))" using assms by (simp add: insertion_altdef coeff_mpoly_map_vars) also have "\ = Sum_any (\m. coeff p (permutep f m) * Prod_any (\i. g (f i) ^ lookup m (f i)))" by (intro Sum_any.cong arg_cong[where ?f = "\y. x * y" for x] Prod_any.reindex_cong[OF assms]) (auto simp: o_def) also have "\ = Sum_any (\m. coeff p m * (\i. g (f i) ^ lookup m i))" by (intro Sum_any.reindex_cong [OF bij_permutep[of f], symmetric]) (auto simp: o_def lookup_permutep assms) also have "\ = insertion (g \ f) p" by (simp add: insertion_altdef) finally show ?thesis . qed lemma permutep_cong: assumes "f permutes (-keys p)" "g permutes (-keys p)" "p = q" shows "permutep f p = permutep g q" proof (intro poly_mapping_eqI) fix k :: 'a show "lookup (permutep f p) k = lookup (permutep g q) k" proof (cases "k \ keys p") case False with assms have "f k \ keys p" "g k \ keys p" using permutes_in_image[of _ "-keys p" k] by auto thus ?thesis using assms by (auto simp: lookup_permutep permutes_bij in_keys_iff) qed (insert assms, auto simp: lookup_permutep permutes_bij permutes_not_in) qed lemma mpoly_map_vars_cong: assumes "f permutes (-vars p)" "g permutes (-vars q)" "p = q" shows "mpoly_map_vars f p = mpoly_map_vars g (q :: 'a :: zero mpoly)" proof (intro mpoly_eqI) fix mon :: "nat \\<^sub>0 nat" show "coeff (mpoly_map_vars f p) mon = coeff (mpoly_map_vars g q) mon" proof (cases "keys mon \ vars p") case True with assms have "permutep f mon = permutep g mon" by (intro permutep_cong assms(1,2)[THEN permutes_subset]) auto thus ?thesis using assms by (simp add: coeff_mpoly_map_vars permutes_bij) next case False hence "\(keys mon \ f ` vars q)" "\(keys mon \ g ` vars q)" using assms by (auto simp: subset_iff permutes_not_in) thus ?thesis using assms by (subst (1 2) coeff_notin_vars) (auto simp: coeff_notin_vars vars_mpoly_map_vars permutes_bij) qed qed subsection \Symmetric polynomials\ text \ A polynomial is symmetric on a set of variables if it is invariant under any permutation of that set. \ definition symmetric_mpoly :: "nat set \ 'a :: zero mpoly \ bool" where "symmetric_mpoly A p = (\\. \ permutes A \ mpoly_map_vars \ p = p)" lemma symmetric_mpoly_empty [simp, intro]: "symmetric_mpoly {} p" by (simp add: symmetric_mpoly_def) text \ A polynomial is trivially symmetric on any set of variables that do not occur in it. \ lemma symmetric_mpoly_orthogonal: assumes "vars p \ A = {}" shows "symmetric_mpoly A p" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" with assms have "\ x = x" if "x \ vars p" for x using that permutes_not_in[of \ A x] by auto from assms have "mpoly_map_vars \ p = mpoly_map_vars id p" by (intro mpoly_map_vars_cong permutes_subset[OF \] permutes_id) auto also have "\ = p" by simp finally show "mpoly_map_vars \ p = p" . qed lemma symmetric_mpoly_monom [intro]: assumes "keys m \ A = {}" shows "symmetric_mpoly A (monom m c)" using assms vars_monom_subset[of m c] by (intro symmetric_mpoly_orthogonal) auto lemma symmetric_mpoly_subset: assumes "symmetric_mpoly A p" "B \ A" shows "symmetric_mpoly B p" unfolding symmetric_mpoly_def proof safe fix \ assume "\ permutes B" with assms have "\ permutes A" using permutes_subset by blast with assms show "mpoly_map_vars \ p = p" by (auto simp: symmetric_mpoly_def) qed text \ If a polynomial is symmetric over some set of variables, that set must either be a subset of the variables occurring in the polynomial or disjoint from it. \ lemma symmetric_mpoly_imp_orthogonal_or_subset: assumes "symmetric_mpoly A p" shows "vars p \ A = {} \ A \ vars p" proof (rule ccontr) assume "\(vars p \ A = {} \ A \ vars p)" then obtain x y where xy: "x \ vars p \ A" "y \ A - vars p" by auto define \ where "\ = Fun.swap x y id" from xy have \: "\ permutes A" unfolding \_def by (intro permutes_swap_id) auto from xy have "y \ \ ` vars p" by (auto simp: \_def Fun.swap_def) also from \ have "\ ` vars p = vars (mpoly_map_vars \ p)" by (auto simp: vars_mpoly_map_vars permutes_bij) also have "mpoly_map_vars \ p = p" using assms \ by (simp add: symmetric_mpoly_def) finally show False using xy by auto qed text \ Symmetric polynomials are closed under ring operations. \ lemma symmetric_mpoly_add [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p + q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_diff [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p - q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_uminus [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (-p)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_uminus_iff [simp]: "symmetric_mpoly A (-p) \ symmetric_mpoly A p" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_smult [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (smult c p)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_mult [intro]: "symmetric_mpoly A p \ symmetric_mpoly A q \ symmetric_mpoly A (p * q)" unfolding symmetric_mpoly_def by simp lemma symmetric_mpoly_0 [simp, intro]: "symmetric_mpoly A 0" and symmetric_mpoly_1 [simp, intro]: "symmetric_mpoly A 1" and symmetric_mpoly_Const [simp, intro]: "symmetric_mpoly A (Const c)" by (simp_all add: symmetric_mpoly_def) lemma symmetric_mpoly_power [intro]: "symmetric_mpoly A p \ symmetric_mpoly A (p ^ n)" by (induction n) (auto intro!: symmetric_mpoly_mult) lemma symmetric_mpoly_sum [intro]: "(\i. i \ B \ symmetric_mpoly A (f i)) \ symmetric_mpoly A (sum f B)" by (induction B rule: infinite_finite_induct) (auto intro!: symmetric_mpoly_add) lemma symmetric_mpoly_prod [intro]: "(\i. i \ B \ symmetric_mpoly A (f i)) \ symmetric_mpoly A (prod f B)" by (induction B rule: infinite_finite_induct) (auto intro!: symmetric_mpoly_mult) text \ An symmetric sum or product over polynomials yields a symmetric polynomial: \ lemma symmetric_mpoly_symmetric_sum: assumes "g permutes X" assumes "\x \. x \ X \ \ permutes A \ mpoly_map_vars \ (f x) = f (g x)" shows "symmetric_mpoly A (\x\X. f x)" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" have "mpoly_map_vars \ (sum f X) = (\x\X. mpoly_map_vars \ (f x))" by simp also have "\ = (\x\X. f (g x))" by (intro sum.cong assms \ refl) also have "\ = (\x\g`X. f x)" using assms by (subst sum.reindex) (auto simp: permutes_inj_on) also have "g ` X = X" using assms by (simp add: permutes_image) finally show "mpoly_map_vars \ (sum f X) = sum f X" . qed lemma symmetric_mpoly_symmetric_prod: assumes "g permutes X" assumes "\x \. x \ X \ \ permutes A \ mpoly_map_vars \ (f x) = f (g x)" shows "symmetric_mpoly A (\x\X. f x)" unfolding symmetric_mpoly_def proof safe fix \ assume \: "\ permutes A" have "mpoly_map_vars \ (prod f X) = (\x\X. mpoly_map_vars \ (f x))" by simp also have "\ = (\x\X. f (g x))" by (intro prod.cong assms \ refl) also have "\ = (\x\g`X. f x)" using assms by (subst prod.reindex) (auto simp: permutes_inj_on) also have "g ` X = X" using assms by (simp add: permutes_image) finally show "mpoly_map_vars \ (prod f X) = prod f X" . qed text \ If $p$ is a polynomial that is symmetric on some subset of variables $A$, then for the leading monomial of $p$, the exponents of these variables are decreasing w.\,r.\,t.\ the variable ordering. \ theorem lookup_lead_monom_decreasing: assumes "symmetric_mpoly A p" defines "m \ lead_monom p" assumes "i \ A" "j \ A" "i \ j" shows "lookup m i \ lookup m j" proof (cases "p = 0") case [simp]: False show ?thesis proof (intro leI notI) assume less: "lookup m i < lookup m j" define \ where "\ = Fun.swap i j id" from assms have \: "\ permutes A" unfolding \_def by (intro permutes_swap_id) auto have [simp]: "\ \ \ = id" "\ i = j" "\ j = i" "\k. k \ i \ k \ j \ \ k = k" by (auto simp: \_def Fun.swap_def fun_eq_iff) have "0 \ lead_coeff p" by simp also have "lead_coeff p = MPoly_Type.coeff (mpoly_map_vars \ p) (permutep \ m)" using \ by (simp add: lead_coeff_def m_def coeff_mpoly_map_vars permutes_bij permutep_comp' [symmetric]) also have "mpoly_map_vars \ p = p" using \ assms by (simp add: symmetric_mpoly_def) finally have "permutep \ m \ m" by (auto simp: m_def) moreover have "lookup m i < lookup (permutep \ m) i" and "(\k m) k)" using assms \ less by (auto simp: lookup_permutep permutes_bij) hence "m < permutep \ m" by (auto simp: less_poly_mapping_def less_fun_def) ultimately show False by simp qed qed (auto simp: m_def) subsection \The elementary symmetric polynomials\ text \ The $k$-th elementary symmetric polynomial for a finite set of variables $A$, with $k$ ranging between 1 and $|A|$, is the sum of the product of all subsets of $A$ with cardinality $k$: \ lift_definition sym_mpoly_aux :: "nat set \ nat \ (nat \\<^sub>0 nat) \\<^sub>0 'a :: {zero_neq_one}" is "\X k mon. if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0" proof - fix k :: nat and X :: "nat set" show "finite {x. (if finite X \ (\Y\X. card Y = k \ x = monom_of_set Y) then 1 else 0) \ (0::'a)}" (is "finite ?A") proof (cases "finite X") case True have "?A \ monom_of_set ` Pow X" by auto moreover from True have "finite (monom_of_set ` Pow X)" by simp ultimately show ?thesis by (rule finite_subset) qed auto qed lemma lookup_sym_mpoly_aux: "Poly_Mapping.lookup (sym_mpoly_aux X k) mon = (if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0)" by transfer' simp lemma lookup_sym_mpoly_aux_monom_of_set [simp]: assumes "finite X" "Y \ X" "card Y = k" shows "Poly_Mapping.lookup (sym_mpoly_aux X k) (monom_of_set Y) = 1" using assms by (auto simp: lookup_sym_mpoly_aux) lemma keys_sym_mpoly_aux: "m \ keys (sym_mpoly_aux A k) \ keys m \ A" by transfer' (auto split: if_splits simp: keys_monom_of_set) lift_definition sym_mpoly :: "nat set \ nat \ 'a :: {zero_neq_one} mpoly" is "sym_mpoly_aux" . lemma vars_sym_mpoly_subset: "vars (sym_mpoly A k) \ A" using keys_sym_mpoly_aux by (auto simp: vars_def sym_mpoly.rep_eq) lemma coeff_sym_mpoly: "MPoly_Type.coeff (sym_mpoly X k) mon = (if finite X \ (\Y. Y \ X \ card Y = k \ mon = monom_of_set Y) then 1 else 0)" by transfer' (simp add: lookup_sym_mpoly_aux) lemma sym_mpoly_infinite: "\finite A \ sym_mpoly A k = 0" by (transfer, transfer) auto lemma sym_mpoly_altdef: "sym_mpoly A k = (\X | X \ A \ card X = k. monom (monom_of_set X) 1)" proof (cases "finite A") case False hence *: "infinite {X. X \ A \ infinite X}" by (rule infinite_infinite_subsets) have "infinite {X. X \ A \ card X = 0}" by (rule infinite_super[OF _ *]) auto moreover have **: "infinite {X. X \ A \ finite X \ card X = k}" if "k \ 0" using that infinite_card_subsets[of A k] False by auto have "infinite {X. X \ A \ card X = k}" if "k \ 0" by (rule infinite_super[OF _ **[OF that]]) auto ultimately show ?thesis using False by (cases "k = 0") (simp_all add: sym_mpoly_infinite) next case True show ?thesis proof (intro mpoly_eqI, goal_cases) case (1 m) show ?case proof (cases "\X. X \ A \ card X = k \ m = monom_of_set X") case False thus ?thesis by (auto simp: coeff_sym_mpoly coeff_sum coeff_monom) next case True then obtain X where X: "X \ A" "card X = k" "m = monom_of_set X" by blast have "coeff (\X | X \ A \ card X = k. monom (monom_of_set X) 1) m = (\X\{X}. 1)" unfolding coeff_sum proof (intro sum.mono_neutral_cong_right ballI) fix Y assume Y: "Y \ {X. X \ A \ card X = k} - {X}" hence "X = Y" if "monom_of_set X = monom_of_set Y" using that finite_subset[OF X(1)] finite_subset[of Y A] \finite A\ by auto thus "coeff (monom (monom_of_set Y) 1) m = 0" using X Y by (auto simp: coeff_monom when_def ) qed (insert X \finite A\, auto simp: coeff_monom) thus ?thesis using \finite A\ by (auto simp: coeff_sym_mpoly coeff_sum coeff_monom) qed qed qed lemma coeff_sym_mpoly_monom_of_set [simp]: assumes "finite X" "Y \ X" "card Y = k" shows "MPoly_Type.coeff (sym_mpoly X k) (monom_of_set Y) = 1" using assms by (auto simp: coeff_sym_mpoly) lemma coeff_sym_mpoly_0: "coeff (sym_mpoly X k) 0 = (if finite X \ k = 0 then 1 else 0)" proof - consider "finite X" "k = 0" | "finite X" "k \ 0" | "infinite X" by blast thus ?thesis proof cases assume "finite X" "k = 0" hence "coeff (sym_mpoly X k) (monom_of_set {}) = 1" by (subst coeff_sym_mpoly_monom_of_set) auto thus ?thesis unfolding monom_of_set_empty using \finite X\ \k = 0\ by simp next assume "finite X" "k \ 0" hence "\(\Y. finite Y \ Y \ X \ card Y = k \ monom_of_set Y = 0)" by auto thus ?thesis using \k \ 0\ by (auto simp: coeff_sym_mpoly) next assume "infinite X" thus ?thesis by (simp add: coeff_sym_mpoly) qed qed lemma symmetric_sym_mpoly [intro]: assumes "A \ B" shows "symmetric_mpoly A (sym_mpoly B k :: 'a :: zero_neq_one mpoly)" unfolding symmetric_mpoly_def proof (safe intro!: mpoly_eqI) fix \ and mon :: "nat \\<^sub>0 nat" assume \: "\ permutes A" from \ have \': "\ permutes B" by (rule permutes_subset) fact from \ have "MPoly_Type.coeff (mpoly_map_vars \ (sym_mpoly B k :: 'a mpoly)) mon = MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) (permutep \ mon)" by (simp add: coeff_mpoly_map_vars permutes_bij) also have "\ = 1 \ MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) mon = 1" (is "?lhs = 1 \ ?rhs = 1") proof assume "?rhs = 1" then obtain Y where "finite B" and Y: "Y \ B" "card Y = k" "mon = monom_of_set Y" by (auto simp: coeff_sym_mpoly split: if_splits) with \' have "\ -` Y \ B" "card (\ -` Y) = k" "permutep \ mon = monom_of_set (\ -` Y)" by (auto simp: permutes_in_image card_vimage_inj permutep_monom_of_set permutes_bij permutes_inj permutes_surj) thus "?lhs = 1" using \finite B\ by (auto simp: coeff_sym_mpoly) next assume "?lhs = 1" then obtain Y where "finite B" and Y: "Y \ B" "card Y = k" "permutep \ mon = monom_of_set Y" by (auto simp: coeff_sym_mpoly split: if_splits) from Y(1) have "inj_on \ Y" using inj_on_subset[of \ UNIV Y] \' by (auto simp: permutes_inj) with Y \' have "\ ` Y \ B" "card (\ ` Y) = k" "mon = monom_of_set (\ ` Y)" by (auto simp: permutes_in_image card_image permutep_monom_of_set permutes_bij permutes_inj permutes_surj) thus "?rhs = 1" using \finite B\ by (auto simp: coeff_sym_mpoly) qed hence "?lhs = ?rhs" by (auto simp: coeff_sym_mpoly split: if_splits) finally show "MPoly_Type.coeff (mpoly_map_vars \ (sym_mpoly B k :: 'a mpoly)) mon = MPoly_Type.coeff (sym_mpoly B k :: 'a mpoly) mon" . qed lemma insertion_sym_mpoly: assumes "finite X" shows "insertion f (sym_mpoly X k) = (\Y | Y \ X \ card Y = k. prod f Y)" using assms proof (transfer, transfer) fix f :: "nat \ 'a" and k :: nat and X :: "nat set" assume X: "finite X" have "insertion_fun f (\mon. if finite X \ (\Y\X. card Y = k \ mon = monom_of_set Y) then 1 else 0) = (\m. (\v. f v ^ poly_mapping.lookup m v) when (\Y\X. card Y = k \ m = monom_of_set Y))" by (auto simp add: insertion_fun_def X when_def intro!: Sum_any.cong) also have "\ = (\m | \Y\Pow X. card Y = k \ m = monom_of_set Y. (\v. f v ^ poly_mapping.lookup m v) when (\Y\X. card Y = k \ m = monom_of_set Y))" by (rule Sum_any.expand_superset) (use X in auto) also have "\ = (\m | \Y\Pow X. card Y = k \ m = monom_of_set Y. (\v. f v ^ poly_mapping.lookup m v))" by (intro sum.cong) (auto simp: when_def) also have "\ = (\Y | Y \ X \ card Y = k. (\v. f v ^ poly_mapping.lookup (monom_of_set Y) v))" by (rule sum.reindex_bij_witness[of _ monom_of_set keys]) (auto simp: finite_subset[OF _ X]) also have "\ = (\Y | Y \ X \ card Y = k. \v\Y. f v)" proof (intro sum.cong when_cong refl, goal_cases) case (1 Y) hence "finite Y" by (auto dest: finite_subset[OF _ X]) with 1 have "(\v. f v ^ poly_mapping.lookup (monom_of_set Y) v) = (\v::nat. if v \ Y then f v else 1)" by (intro Prod_any.cong) (auto simp: lookup_monom_of_set) also have "\ = (\v\Y. f v)" by (rule Prod_any.conditionalize [symmetric]) fact+ finally show ?case . qed finally show "insertion_fun f (\mon. if finite X \ (\Y\X. card Y = k \ mon = monom_of_set Y) then 1 else 0) = (\Y | Y \ X \ card Y = k. prod f Y)" . qed lemma sym_mpoly_nz [simp]: assumes "finite A" "k \ card A" shows "sym_mpoly A k \ (0 :: 'a :: zero_neq_one mpoly)" proof - from assms obtain B where B: "B \ A" "card B = k" using ex_subset_of_card by blast with assms have "coeff (sym_mpoly A k :: 'a mpoly) (monom_of_set B) = 1" by (intro coeff_sym_mpoly_monom_of_set) thus ?thesis by auto qed lemma coeff_sym_mpoly_0_or_1: "coeff (sym_mpoly A k) m \ {0, 1}" by (transfer, transfer) auto lemma lead_coeff_sym_mpoly [simp]: assumes "finite A" "k \ card A" shows "lead_coeff (sym_mpoly A k) = 1" proof - from assms have "lead_coeff (sym_mpoly A k) \ 0" by simp thus ?thesis using coeff_sym_mpoly_0_or_1[of A k "lead_monom (sym_mpoly A k)"] unfolding lead_coeff_def by blast qed lemma lead_monom_sym_mpoly: assumes "sorted xs" "distinct xs" "k \ length xs" shows "lead_monom (sym_mpoly (set xs) k :: 'a :: zero_neq_one mpoly) = monom_of_set (set (take k xs))" (is "lead_monom ?p = _") proof - let ?m = "lead_monom ?p" have sym: "symmetric_mpoly (set xs) (sym_mpoly (set xs) k)" by (intro symmetric_sym_mpoly) auto from assms have [simp]: "card (set xs) = length xs" by (subst distinct_card) auto from assms have "lead_coeff ?p = 1" by (subst lead_coeff_sym_mpoly) auto then obtain X where X: "X \ set xs" "card X = k" "?m = monom_of_set X" unfolding lead_coeff_def by (subst (asm) coeff_sym_mpoly) (auto split: if_splits) define ys where "ys = map (\x. if x \ X then 1 else 0 :: nat) xs" have [simp]: "length ys = length xs" by (simp add: ys_def) have ys_altdef: "ys = map (lookup ?m) xs" unfolding ys_def using X finite_subset[OF X(1)] by (intro map_cong) (auto simp: lookup_monom_of_set) define i where "i = Min (insert (length xs) {i. i < length xs \ ys ! i = 0})" have "i \ length xs" by (auto simp: i_def) have in_X: "xs ! j \ X" if "j < i" for j using that unfolding i_def by (auto simp: ys_def) have not_in_X: "xs ! j \ X" if "i \ j" "j < length xs" for j proof - have ne: "{i. i < length xs \ ys ! i = 0} \ {}" proof assume [simp]: "{i. i < length xs \ ys ! i = 0} = {}" from that show False by (simp add: i_def) qed hence "Min {i. i < length xs \ ys ! i = 0} \ {i. i < length xs \ ys ! i = 0}" using that by (intro Min_in) auto also have "Min {i. i < length xs \ ys ! i = 0} = i" unfolding i_def using ne by (subst Min_insert) (auto simp: min_def) finally have i: "ys ! i = 0" "i < length xs" by simp_all have "lookup ?m (xs ! j) \ lookup ?m (xs ! i)" using that assms by (intro lookup_lead_monom_decreasing[OF sym]) (auto intro!: sorted_nth_mono simp: set_conv_nth) also have "\ = 0" using i by (simp add: ys_altdef) finally show ?thesis using that X finite_subset[OF X(1)] by (auto simp: lookup_monom_of_set) qed from X have "k = card X" by simp also have "X = (\i. xs ! i) ` {i. i < length xs \ xs ! i \ X}" using X by (auto simp: set_conv_nth) also have "card \ = (\i | i < length xs \ xs ! i \ X. 1)" using assms by (subst card_image) (auto intro!: inj_on_nth) also have "\ = (\i | i < length xs. if xs ! i \ X then 1 else 0)" by (intro sum.mono_neutral_cong_left) auto also have "\ = sum_list ys" by (auto simp: sum_list_sum_nth ys_def intro!: sum.cong) also have "ys = take i ys @ drop i ys" by simp also have "sum_list \ = sum_list (take i ys) + sum_list (drop i ys)" by (subst sum_list_append) auto also have "take i ys = replicate i 1" using \i \ length xs\ in_X by (intro replicate_eqI) (auto simp: ys_def set_conv_nth) also have "sum_list \ = i" by simp also have "drop i ys = replicate (length ys - i) 0" using \i \ length xs\ not_in_X by (intro replicate_eqI) (auto simp: ys_def set_conv_nth) also have "sum_list \ = 0" by simp finally have "i = k" by simp have "X = set (filter (\x. x \ X) xs)" using X by auto also have "xs = take i xs @ drop i xs" by simp also note filter_append also have "filter (\x. x \ X) (take i xs) = take i xs" using in_X by (intro filter_True) (auto simp: set_conv_nth) also have "filter (\x. x \ X) (drop i xs) = []" using not_in_X by (intro filter_False) (auto simp: set_conv_nth) finally have "X = set (take i xs)" by simp with \i = k\ and X show ?thesis by simp qed subsection \Induction on the leading monomial\ text \ We show that the monomial ordering for a fixed set of variables is well-founded, so we can perform induction on the leading monomial of a polynomial. \ definition monom_less_on where "monom_less_on A = {(m1, m2). m1 < m2 \ keys m1 \ A \ keys m2 \ A}" lemma wf_monom_less_on: assumes "finite A" shows "wf (monom_less_on A :: ((nat \\<^sub>0 'b :: {zero, wellorder}) \ _) set)" proof (rule wf_subset) define n where "n = Suc (Max (insert 0 A))" have less_n: "k < n" if "k \ A" for k using that assms by (auto simp: n_def less_Suc_eq_le Max_ge_iff) define f :: "(nat \\<^sub>0 'b) \ 'b list" where "f = (\m. map (lookup m) [0.. inv_image (lexn {(x, y). x < y} n) f" proof safe fix m1 m2 :: "nat \\<^sub>0 'b" assume "(m1, m2) \ monom_less_on A" hence m12: "m1 < m2" "keys m1 \ A" "keys m2 \ A" by (auto simp: monom_less_on_def) then obtain k where k: "lookup m1 k < lookup m2 k" "\i(lookup m1 k = 0 \ lookup m2 k = 0)" proof (intro notI) assume "lookup m1 k = 0 \ lookup m2 k = 0" hence [simp]: "lookup m1 k = 0" "lookup m2 k = 0" by blast+ from k(1) show False by simp qed hence "k \ A" using m12 by (auto simp: in_keys_iff) hence "k < n" by (simp add: less_n) define as where "as = map (lookup m1) [0..k < n\ by (simp flip: upt_conv_Cons upt_add_eq_append') have [simp]: "length as = k" "length bs1 = n - Suc k" "length bs2 = n - Suc k" by (simp_all add: as_def bs1_def bs2_def) have "f m1 = as @ [lookup m1 k] @ bs1" unfolding f_def by (subst decomp) (simp add: as_def bs1_def) moreover have "f m2 = as @ [lookup m2 k] @ bs2" unfolding f_def using k by (subst decomp) (simp add: as_def bs2_def) ultimately show "(m1, m2) \ inv_image (lexn {(x,y). x < y} n) f" - using k(1) \k < n\ unfolding lexn_conv by auto + using k(1) \k < n\ unfolding lexn_conv by fastforce qed qed lemma lead_monom_induct [consumes 2, case_names less]: fixes p :: "'a :: zero mpoly" assumes fin: "finite A" and vars: "vars p \ A" assumes IH: "\p. vars p \ A \ (\p'. vars p' \ A \ lead_monom p' < lead_monom p \ P p') \ P p" shows "P p" using assms(2) proof (induct m \ "lead_monom p" arbitrary: p rule: wf_induct_rule[OF wf_monom_less_on[OF fin]]) case (1 p) show ?case proof (rule IH) fix p' :: "'a mpoly" assume *: "vars p' \ A" "lead_monom p' < lead_monom p" show "P p'" by (rule 1) (insert * "1.prems" keys_lead_monom_subset, auto simp: monom_less_on_def) qed (insert 1, auto) qed lemma lead_monom_induct' [case_names less]: fixes p :: "'a :: zero mpoly" assumes IH: "\p. (\p'. vars p' \ vars p \ lead_monom p' < lead_monom p \ P p') \ P p" shows "P p" proof - have "finite (vars p)" "vars p \ vars p" by (auto simp: vars_finite) thus ?thesis by (induction rule: lead_monom_induct) (use IH in blast) qed subsection \The fundamental theorem of symmetric polynomials\ lemma lead_coeff_sym_mpoly_powerprod: assumes "finite A" "\x. x \ X \ f x \ {1..card A}" shows "lead_coeff (\x\X. sym_mpoly A (f (x::'a)) ^ g x) = 1" proof - have eq: "lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly) = 1" if "x \ X" for x using that assms by (subst lead_coeff_power) (auto simp: lead_coeff_sym_mpoly assms) hence "(\x\X. lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly)) = (\x\X. 1)" by (intro prod.cong eq refl) also have "\ = 1" by simp finally have eq': "(\x\X. lead_coeff (sym_mpoly A (f x) ^ g x :: 'b mpoly)) = 1" . show ?thesis by (subst lead_coeff_prod) (auto simp: eq eq') qed context fixes A :: "nat set" and xs n f and decr :: "'a :: comm_ring_1 mpoly \ bool" defines "xs \ sorted_list_of_set A" defines "n \ card A" defines "f \ (\i. if i < n then xs ! i else 0)" defines "decr \ (\p. \i\A. \j\A. i \ j \ lookup (lead_monom p) i \ lookup (lead_monom p) j)" begin text \ The computation of the witness for the fundamental theorem works like this: Given some polynomial $p$ (that is assumed to be symmetric in the variables in $A$), we inspect its leading monomial, which is of the form $c X_1^{i_1}\ldots X_n{i_n}$ where the $A = \{X_1,\ldots, X_n\}$, $c$ contains only variables not in $A$, and the sequence $i_j$ is decreasing. The latter holds because $p$ is symmetric. Now, we form the polynomial $q := c e_1^{i_1 - i_2} e_2^{i_2 - i_3} \ldots e_n^{i_n}$, which has the same leading term as $p$. Then $p - q$ has a smaller leading monomial, so by induction, we can assume it to be of the required form and obtain a witness for $p - q$. Now, we only need to add $c Y_1^{i_1 - i_2} \ldots Y_n^{i_n}$ to that witness and we obtain a witness for $p$. \ definition fund_sym_step_coeff :: "'a mpoly \ 'a mpoly" where "fund_sym_step_coeff p = monom (restrictpm (-A) (lead_monom p)) (lead_coeff p)" definition fund_sym_step_monom :: "'a mpoly \ (nat \\<^sub>0 nat)" where "fund_sym_step_monom p = ( let g = (\i. if i < n then lookup (lead_monom p) (f i) else 0) in (\i 'a mpoly" where "fund_sym_step_poly p = ( let g = (\i. if i < n then lookup (lead_monom p) (f i) else 0) in fund_sym_step_coeff p * (\i The following function computes the witness, with the convention that it returns a constant polynomial if the input was not symmetric: \ function (domintros) fund_sym_poly_wit :: "'a :: comm_ring_1 mpoly \ 'a mpoly mpoly" where "fund_sym_poly_wit p = (if \symmetric_mpoly A p \ lead_monom p = 0 \ vars p \ A = {} then Const p else fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (fund_sym_step_coeff p))" by auto lemma coeff_fund_sym_step_coeff: "coeff (fund_sym_step_coeff p) m \ {lead_coeff p, 0}" by (auto simp: fund_sym_step_coeff_def coeff_monom when_def) lemma vars_fund_sym_step_coeff: "vars (fund_sym_step_coeff p) \ vars p - A" unfolding fund_sym_step_coeff_def using keys_lead_monom_subset[of p] by (intro order.trans[OF vars_monom_subset]) auto lemma keys_fund_sym_step_monom: "keys (fund_sym_step_monom p) \ {1..n}" unfolding fund_sym_step_monom_def Let_def by (intro order.trans[OF keys_sum] UN_least, subst keys_single) auto lemma coeff_fund_sym_step_poly: assumes C: "\m. coeff p m \ C" and "ring_closed C" shows "coeff (fund_sym_step_poly p) m \ C" proof - interpret ring_closed C by fact have *: "\m. coeff (p ^ x) m \ C" if "\m. coeff p m \ C" for p x using that by (induction x) (auto simp: coeff_mpoly_times mpoly_coeff_1 intro!: prod_fun_closed) have **: "\m. coeff (prod f X) m \ C" if "\i m. i \ X \ coeff (f i) m \ C" for X and f :: "nat \ _" using that by (induction X rule: infinite_finite_induct) (auto simp: coeff_mpoly_times mpoly_coeff_1 intro!: prod_fun_closed) show ?thesis using C unfolding fund_sym_step_poly_def Let_def fund_sym_step_coeff_def coeff_mpoly_times by (intro prod_fun_closed) (auto simp: coeff_monom when_def lead_coeff_def coeff_sym_mpoly intro!: * **) qed text \ We now show various relevant properties of the subtracted polynomial: \<^enum> Its leading term is the same as that of the input polynomial. \<^enum> It contains now new variables. \<^enum> It is symmetric in the variables in \A\. \ lemma fund_sym_step_poly: shows "finite A \ p \ 0 \ decr p \ lead_monom (fund_sym_step_poly p) = lead_monom p" and "finite A \ p \ 0 \ decr p \ lead_coeff (fund_sym_step_poly p) = lead_coeff p" and "finite A \ p \ 0 \ decr p \ fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" and "vars (fund_sym_step_poly p) \ vars p \ A" and "symmetric_mpoly A (fund_sym_step_poly p)" proof - define g where "g = (\i. if i < n then lookup (lead_monom p) (f i) else 0)" define q where "q = (\i vars p \ A" using keys_lead_monom_subset[of p] vars_monom_subset[of "restrictpm (-A) (lead_monom p)" "lead_coeff p"] unfolding c_def q_def by (intro order.trans[OF vars_mult] order.trans[OF vars_prod] order.trans[OF vars_power] Un_least UN_least order.trans[OF vars_sym_mpoly_subset]) auto thus "vars (fund_sym_step_poly p) \ vars p \ A" by simp have "symmetric_mpoly A (c * q)" unfolding c_def q_def by (intro symmetric_mpoly_mult symmetric_mpoly_monom symmetric_mpoly_prod symmetric_mpoly_power symmetric_sym_mpoly) auto thus "symmetric_mpoly A (fund_sym_step_poly p)" by simp assume finite: "finite A" and [simp]: "p \ 0" and "decr p" have "set xs = A" "distinct xs" and [simp]: "length xs = n" using finite by (auto simp: xs_def n_def) have [simp]: "lead_coeff c = lead_coeff p" "lead_monom c = restrictpm (- A) (lead_monom p)" by (simp_all add: c_def lead_monom_monom) hence f_range [simp]: "f i \ A" if "i < n" for i using that \set xs = A\ by (auto simp: f_def set_conv_nth) have "sorted xs" by (simp add: xs_def) hence f_mono: "f i \ f j" if "i \ j" "j < n" for i j using that by (auto simp: f_def n_def intro: sorted_nth_mono) hence g_mono: "g i \ g j" if "i \ j" for i j unfolding g_def using that using \decr p\ by (auto simp: decr_def) have *: "(\iifinite A\ by (intro prod.cong) (auto simp: n_def lead_coeff_power) hence "lead_coeff q = (\i = (\ifinite A\ by (intro prod.cong) (auto simp: lead_coeff_power n_def) finally have [simp]: "lead_coeff q = 1" by simp have "lead_monom q = (\i = (\ifinite A\ by (intro sum.cong) (auto simp: lead_monom_power n_def) also have "\ = (\iset xs = A\) also from 1 have "\ = monom_of_set (set (take (Suc i) xs))" by (subst lead_monom_sym_mpoly) (auto simp: xs_def n_def) finally show ?case by simp qed finally have lead_monom_q: "lead_monom q = (\i = lead_monom p" (is "?S = _") proof (intro poly_mapping_eqI) fix i :: nat show "lookup (lead_monom c + lead_monom q) i = lookup (lead_monom p) i" proof (cases "i \ A") case False hence "lookup (lead_monom c + lead_monom q) i = lookup (lead_monom p) i + (\jset xs = A\ dest!: in_set_takeD) finally show ?thesis by simp next case True with \set xs = A\ obtain m where m: "i = xs ! m" "m < n" by (auto simp: set_conv_nth) have "lookup (lead_monom c + lead_monom q) i = (\j = (\j | j < n \ i \ set (take (Suc j) xs). g j - g (Suc j))" by (intro sum.mono_neutral_cong_right) auto also have "{j. j < n \ i \ set (take (Suc j) xs)} = {m..distinct xs\ by (force simp: set_conv_nth nth_eq_iff_index_eq) also have "(\j\\. g j - g (Suc j)) = (\j\\. g j) - (\j\\. g (Suc j))" by (subst sum_subtractf_nat) (auto intro!: g_mono) also have "(\j\{m..j\{m<..n}. g j)" by (intro sum.reindex_bij_witness[of _ "\j. j - 1" Suc]) auto also have "\ = (\j\{m<..j\{m.. = (\j\{m..j\\. g j) = lookup (lead_monom p) i" using m by (auto simp: g_def not_less le_Suc_eq f_def) finally show ?thesis . qed qed finally show "lead_monom (fund_sym_step_poly p) = lead_monom p" by simp show "lead_coeff (fund_sym_step_poly p) = lead_coeff p" by (simp add: lead_coeff_mult) have *: "lookup (fund_sym_step_monom p) k = (if k \ {1..n} then g (k - 1) - g k else 0)" for k proof - have "lookup (fund_sym_step_monom p) k = (\x\(if k \ {1..n} then {k - 1} else {}). g (k - 1) - g k)" unfolding fund_sym_step_monom_def lookup_sum Let_def by (intro sum.mono_neutral_cong_right) (auto simp: g_def lookup_single when_def split: if_splits) thus ?thesis by simp qed hence "(\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x :: 'a mpoly) = (\x\{1..n}. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" by (intro Prod_any.expand_superset) auto also have "\ = (\xi. i - 1"]) auto also have "\ = q" unfolding q_def by (intro prod.cong) (auto simp: *) finally show "fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" by (simp add: c_def q_def f_def g_def fund_sym_step_monom_def fund_sym_step_coeff_def) qed text \ If the input is well-formed, a single step of the procedure always decreases the leading monomial. \ lemma lead_monom_fund_sym_step_poly_less: assumes "finite A" and "lead_monom p \ 0" and "decr p" shows "lead_monom (p - fund_sym_step_poly p) < lead_monom p" proof (cases "p = fund_sym_step_poly p") case True thus ?thesis using assms by (auto simp: order.strict_iff_order) next case False from assms have [simp]: "p \ 0" by auto let ?q = "fund_sym_step_poly p" and ?m = "lead_monom p" have "coeff (p - ?q) ?m = 0" using fund_sym_step_poly[of p] assms by (simp add: lead_coeff_def) moreover have "lead_coeff (p - ?q) \ 0" using False by auto ultimately have "lead_monom (p - ?q) \ ?m" unfolding lead_coeff_def by auto moreover have "lead_monom (p - ?q) \ ?m" using fund_sym_step_poly[of p] assms by (intro order.trans[OF lead_monom_diff] max.boundedI) auto ultimately show ?thesis by (auto simp: order.strict_iff_order) qed text \ Finally, we prove that the witness is indeed well-defined for all inputs. \ lemma fund_sym_poly_wit_dom_aux: assumes "finite B" "vars p \ B" "A \ B" shows "fund_sym_poly_wit_dom p" using assms(1-3) proof (induction p rule: lead_monom_induct) case (less p) have [simp]: "finite A" by (rule finite_subset[of _ B]) fact+ show ?case proof (cases "lead_monom p = 0 \ \symmetric_mpoly A p") case False hence [simp]: "p \ 0" by auto note decr = lookup_lead_monom_decreasing[of A p] have "vars (p - fund_sym_step_poly p) \ B" using fund_sym_step_poly[of p] decr False less.prems less.hyps \A \ B\ by (intro order.trans[OF vars_diff]) auto hence "fund_sym_poly_wit_dom (p - local.fund_sym_step_poly p)" using False less.prems less.hyps decr by (intro less.IH fund_sym_step_poly symmetric_mpoly_diff lead_monom_fund_sym_step_poly_less) (auto simp: decr_def) thus ?thesis using fund_sym_poly_wit.domintros by blast qed (auto intro: fund_sym_poly_wit.domintros) qed lemma fund_sym_poly_wit_dom [intro]: "fund_sym_poly_wit_dom p" proof - consider "\symmetric_mpoly A p" | "vars p \ A = {}" | "symmetric_mpoly A p" "A \ vars p" using symmetric_mpoly_imp_orthogonal_or_subset[of A p] by blast thus ?thesis proof cases assume "symmetric_mpoly A p" "A \ vars p" thus ?thesis using fund_sym_poly_wit_dom_aux[of "vars p" p] by (auto simp: vars_finite) qed (auto intro: fund_sym_poly_wit.domintros) qed termination fund_sym_poly_wit by (intro allI fund_sym_poly_wit_dom) (*<*) lemmas [simp del] = fund_sym_poly_wit.simps (*>*) text \ Next, we prove that our witness indeed fulfils all the properties stated by the fundamental theorem: \<^enum> If the original polynomial was in $R[X_1,\ldots,X_n,\ldots, X_m]$ where the $X_1$ to $X_n$ are the symmetric variables, then the witness is a polynomial in $R[X_{n+1},\ldots,X_m][Y_1,\ldots,Y_n]$. This means that its coefficients are polynomials in the variables of the original polynomial, minus the symmetric ones, and the (new and independent) variables of the witness polynomial range from $1$ to $n$. \<^enum> Substituting the \i\-th symmetric polynomial $e_i(X_1,\ldots,X_n)$ for the $Y_i$ variable for every \i\ yields the original polynomial. \<^enum> The coefficient ring $R$ need not be the entire type; if the coefficients of the original polynomial are in some subring, then the coefficients of the coefficients of the witness also do. \ lemma fund_sym_poly_wit_coeffs_aux: assumes "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" shows "vars (coeff (fund_sym_poly_wit p) m) \ B - A" using assms proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto with 1 False have "vars (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p)) m) \ B - A" by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) auto hence "vars (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (fund_sym_step_coeff p)) m) \ B - A" unfolding coeff_add coeff_monom using vars_fund_sym_step_coeff[of p] "1.prems" by (intro order.trans[OF vars_add] Un_least order.trans[OF vars_monom_subset]) (auto simp: when_def) thus ?thesis using "1.prems" False unfolding fund_sym_poly_wit.simps[of p] by simp qed (insert "1.prems", auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const lead_monom_eq_0_iff) qed lemma fund_sym_poly_wit_coeffs: assumes "symmetric_mpoly A p" shows "vars (coeff (fund_sym_poly_wit p) m) \ vars p - A" proof (cases "A \ vars p") case True with fund_sym_poly_wit_coeffs_aux[of "vars p" p m] assms show ?thesis by (auto simp: vars_finite) next case False hence "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[OF assms] by auto thus ?thesis by (auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const) qed lemma fund_sym_poly_wit_vars: "vars (fund_sym_poly_wit p) \ {1..n}" proof (cases "symmetric_mpoly A p \ A \ vars p") case True define B where "B = vars p" have "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" using True unfolding B_def by (auto simp: vars_finite) thus ?thesis proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto hence "vars (local.fund_sym_poly_wit (p - local.fund_sym_step_poly p)) \ {1..n}" using False "1.prems" by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) (auto simp: lead_monom_eq_0_iff) hence "vars (fund_sym_poly_wit (p - fund_sym_step_poly p) + monom (fund_sym_step_monom p) (local.fund_sym_step_coeff p)) \ {1..n}" by (intro order.trans[OF vars_add] Un_least order.trans[OF vars_monom_subset] keys_fund_sym_step_monom) auto thus ?thesis using "1.prems" False unfolding fund_sym_poly_wit.simps[of p] by simp qed (insert "1.prems", auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const lead_monom_eq_0_iff) qed next case False then consider "\symmetric_mpoly A p" | "symmetric_mpoly A p" "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[of A p] by auto thus ?thesis by cases (auto simp: fund_sym_poly_wit.simps[of p]) qed lemma fund_sym_poly_wit_insertion_aux: assumes "finite B" "vars p \ B" "symmetric_mpoly A p" "A \ B" shows "insertion (sym_mpoly A) (fund_sym_poly_wit p) = p" using assms proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) from "1.prems" have "decr p" using lookup_lead_monom_decreasing[of A p] by (auto simp: decr_def) show ?case proof (cases "lead_monom p = 0 \ vars p \ A = {}") case False have "vars (p - fund_sym_step_poly p) \ B" using "1.prems" fund_sym_step_poly[of p] by (intro order.trans[OF vars_diff]) auto hence "insertion (sym_mpoly A) (fund_sym_poly_wit (p - fund_sym_step_poly p)) = p - fund_sym_step_poly p" using 1 False by (intro 1 symmetric_mpoly_diff fund_sym_step_poly) auto moreover have "fund_sym_step_poly p = fund_sym_step_coeff p * (\x. sym_mpoly A x ^ lookup (fund_sym_step_monom p) x)" using "1.prems" finite_subset[of A B] False \decr p\ by (intro fund_sym_step_poly) auto ultimately show ?thesis unfolding fund_sym_poly_wit.simps[of p] by (auto simp: insertion_add) qed (auto simp: fund_sym_poly_wit.simps[of p]) qed lemma fund_sym_poly_wit_insertion: assumes "symmetric_mpoly A p" shows "insertion (sym_mpoly A) (fund_sym_poly_wit p) = p" proof (cases "A \ vars p") case False hence "vars p \ A = {}" using symmetric_mpoly_imp_orthogonal_or_subset[OF assms] by auto thus ?thesis by (auto simp: fund_sym_poly_wit.simps[of p]) next case True with fund_sym_poly_wit_insertion_aux[of "vars p" p] assms show ?thesis by (auto simp: vars_finite) qed lemma fund_sym_poly_wit_coeff: assumes "\m. coeff p m \ C" "ring_closed C" shows "\m m'. coeff (coeff (fund_sym_poly_wit p) m) m' \ C" using assms(1) proof (induction p rule: fund_sym_poly_wit.induct) case (1 p) interpret ring_closed C by fact show ?case proof (cases "\symmetric_mpoly A p \ lead_monom p = 0 \ vars p \ A = {}") case True thus ?thesis using "1.prems" by (auto simp: fund_sym_poly_wit.simps[of p] mpoly_coeff_Const) next case False have *: "\m m'. coeff (coeff (fund_sym_poly_wit (p - fund_sym_step_poly p)) m) m' \ C" using False "1.prems" assms coeff_fund_sym_step_poly [of p] by (intro 1) auto show ?thesis proof (intro allI, goal_cases) case (1 m m') thus ?case using * False coeff_fund_sym_step_coeff[of p m'] "1.prems" by (auto simp: fund_sym_poly_wit.simps[of p] coeff_monom lead_coeff_def when_def) qed qed qed subsection \Uniqueness\ text \ Next, we show that the polynomial representation of a symmetric polynomial in terms of the elementary symmetric polynomials not only exists, but is unique. The key property here is that products of powers of elementary symmetric polynomials uniquely determine the exponent vectors, i.\,e.\ if $e_1, \ldots, e_n$ are the elementary symmetric polynomials, $a = (a_1,\ldots, a_n)$ and $b = (b_1,\ldots,b_n)$ are vectors of natural numbers, then: \[e_1^{a_1}\ldots e_n^{a_n} = e_1^{b_1}\ldots e_n^{b_n} \longleftrightarrow a = b\] We show this now. \ lemma lead_monom_sym_mpoly_prod: assumes "finite A" shows "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly A i :: 'a mpoly))" proof - have "(\i=1..n. lead_coeff (sym_mpoly A i ^ h i :: 'a mpoly)) = 1" using assms unfolding n_def by (intro prod.neutral allI) (auto simp: lead_coeff_power) hence "lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i=1..n. lead_monom (sym_mpoly A i ^ h i :: 'a mpoly))" by (subst lead_monom_prod) auto also have "\ = (\i=1..n. of_nat (h i) * lead_monom (sym_mpoly A i :: 'a mpoly))" by (intro sum.cong refl, subst lead_monom_power) (auto simp: lead_coeff_power assms n_def) finally show ?thesis . qed lemma lead_monom_sym_mpoly_prod_notin: assumes "finite A" "k \ A" shows "lookup (lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly)) k = 0" proof - have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) have "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly (set xs) i :: 'a mpoly))" by (subst lead_monom_sym_mpoly_prod) (use xs assms in auto) also have "lookup \ k = 0" unfolding lookup_sum by (intro sum.neutral ballI, subst lead_monom_sym_mpoly) (insert xs assms, auto simp: xs lead_monom_sym_mpoly lookup_monom_of_set set_conv_nth) finally show ?thesis . qed lemma lead_monom_sym_mpoly_prod_in: assumes "finite A" "k < n" shows "lookup (lead_monom (\i=1..n. sym_mpoly A i ^ h i :: 'a mpoly)) (xs ! k) = (\i=k+1..n. h i)" proof - have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) have "lead_monom (\i = 1..n. sym_mpoly A i ^ h i :: 'a mpoly) = (\i = 1..n. of_nat (h i) * lead_monom (sym_mpoly (set xs) i :: 'a mpoly))" by (subst lead_monom_sym_mpoly_prod) (use xs assms in simp_all) also have "\ = (\i=1..n. of_nat (h i) * monom_of_set (set (take i xs)))" using xs by (intro sum.cong refl, subst lead_monom_sym_mpoly) auto also have "lookup \ (xs ! k) = (\i | i \ {1..n} \ xs ! k \ set (take i xs). h i)" unfolding lookup_sum lookup_monom_of_set by (intro sum.mono_neutral_cong_right) auto also have "{i. i \ {1..n} \ xs ! k \ set (take i xs)} = {k+1..n}" proof (intro equalityI subsetI) fix i assume i: "i \ {k+1..n}" hence "take i xs ! k = xs ! k" "k < n" "k < i" using assms by auto with i show "i \ {i. i \ {1..n} \ xs ! k \ set (take i xs)}" by (force simp: set_conv_nth) qed (insert assms xs, auto simp: set_conv_nth Suc_le_eq nth_eq_iff_index_eq) finally show ?thesis . qed lemma lead_monom_sym_poly_powerprod_inj: assumes "lead_monom (\i. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" assumes "finite A" "keys m1 \ {1..n}" "keys m2 \ {1..n}" shows "m1 = m2" proof (rule poly_mapping_eqI) fix k :: nat have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) from assms(3,4) have *: "i \ {1..n}" if "lookup m1 i \ 0 \ lookup m2 i \ 0" for i using that by (auto simp: subset_iff in_keys_iff) have **: "(\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = (\i=1..n. sym_mpoly A i ^ lookup m i :: 'a mpoly)" if "m \ {m1, m2}" for m using that * by (intro Prod_any.expand_superset subsetI * ) (auto intro!: Nat.gr0I) have ***: "lead_monom (\i=1..n. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i=1..n. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" using assms by (simp add: ** ) have sum_eq: "sum (lookup m1) {Suc k..n} = sum (lookup m2) {Suc k..n}" if "k < n" for k using arg_cong[OF ***, of "\m. lookup m (xs ! k)"] \finite A\ that by (subst (asm) (1 2) lead_monom_sym_mpoly_prod_in) auto show "lookup m1 k = lookup m2 k" proof (cases "k \ {1..n}") case False hence "lookup m1 k = 0" "lookup m2 k = 0" using assms by (auto simp: in_keys_iff) thus ?thesis by simp next case True thus ?thesis proof (induction "n - k" arbitrary: k rule: less_induct) case (less l) have "sum (lookup m1) {Suc (l - 1)..n} = sum (lookup m2) {Suc (l - 1)..n}" using less by (intro sum_eq) auto also have "{Suc (l - 1)..n} = insert l {Suc l..n}" using less by auto also have "sum (lookup m1) \ = lookup m1 l + (\i=Suc l..n. lookup m1 i)" by (subst sum.insert) auto also have "(\i=Suc l..n. lookup m1 i) = (\i=Suc l..n. lookup m2 i)" by (intro sum.cong less) auto also have "sum (lookup m2) (insert l {Suc l..n}) = lookup m2 l + (\i=Suc l..n. lookup m2 i)" by (subst sum.insert) auto finally show "lookup m1 l = lookup m2 l" by simp qed qed qed text \ We now show uniqueness by first showing that the zero polynomial has a unique representation. We fix some polynomial $p$ with $p(e_1,\ldots, e_n) = 0$ and then show, by contradiction, that $p = 0$. We have \[p(e_1,\ldots,e_n) = \sum c_{a_1,\ldots,a_n} e_1^{a_1}\ldots e_n^{a_n}\] and due to the injectivity of products of powers of elementary symmetric polynomials, the leading term of that sum is precisely the leading term of the summand with the biggest leading monomial, since summands cannot cancel each other. However, we also know that $p(e_1,\ldots,e_n) = 0$, so it follows that all summands must have leading term 0, and it is then easy to see that they must all be identically 0. \ lemma sym_mpoly_representation_unique_aux: fixes p :: "'a mpoly mpoly" assumes "finite A" "insertion (sym_mpoly A) p = 0" "\m. vars (coeff p m) \ A = {}" "vars p \ {1..n}" shows "p = 0" proof (rule ccontr) assume p: "p \ 0" have xs: "set xs = A" "distinct xs" "sorted xs" and [simp]: "length xs = n" using assms by (auto simp: xs_def n_def) define h where "h = (\m. coeff p m * (\i. sym_mpoly A i ^ lookup m i))" define M where "M = {m. coeff p m \ 0}" define maxm where "maxm = Max ((lead_monom \ h) ` M)" have "finite M" by (auto intro!: finite_subset[OF _ finite_coeff_support[of p]] simp: h_def M_def) have keys_subset: "keys m \ {1..n}" if "coeff p m \ 0" for m using that assms coeff_notin_vars[of m p] by blast have lead_coeff: "lead_coeff (h m) = lead_coeff (coeff p m)" (is ?th1) and lead_monom: "lead_monom (h m) = lead_monom (coeff p m) + lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" (is ?th2) if [simp]: "coeff p m \ 0" for m proof - have "(\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = (\i | lookup m i \ 0. sym_mpoly A i ^ lookup m i :: 'a mpoly)" by (intro Prod_any.expand_superset) (auto intro!: Nat.gr0I) also have "lead_coeff \ = 1" using assms keys_subset[of m] by (intro lead_coeff_sym_mpoly_powerprod) (auto simp: in_keys_iff subset_iff n_def) finally have eq: "lead_coeff (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly) = 1" . thus ?th1 unfolding h_def using \coeff p m \ 0\ by (subst lead_coeff_mult) auto show ?th2 unfolding h_def by (subst lead_monom_mult) (auto simp: eq) qed have "insertion (sym_mpoly A) p = (\m\M. h m)" unfolding insertion_altdef h_def M_def by (intro Sum_any.expand_superset) auto also have "lead_monom \ = maxm" unfolding maxm_def proof (rule lead_monom_sum) from p obtain m where "coeff p m \ 0" using mpoly_eqI[of p 0] by auto hence "m \ M" using \coeff p m \ 0\ lead_coeff[of m] by (auto simp: M_def) thus "M \ {}" by auto next have restrict_lead_monom: "restrictpm A (lead_monom (h m)) = lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" if [simp]: "coeff p m \ 0" for m proof - have "restrictpm A (lead_monom (h m)) = restrictpm A (lead_monom (coeff p m)) + restrictpm A (lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly))" by (auto simp: lead_monom restrictpm_add) also have "restrictpm A (lead_monom (coeff p m)) = 0" using assms by (intro restrictpm_orthogonal order.trans[OF keys_lead_monom_subset]) auto also have "restrictpm A (lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)) = lead_monom (\i. sym_mpoly A i ^ lookup m i :: 'a mpoly)" by (intro restrictpm_id order.trans[OF keys_lead_monom_subset] order.trans[OF vars_Prod_any] UN_least order.trans[OF vars_power] vars_sym_mpoly_subset) finally show ?thesis by simp qed show "inj_on (lead_monom \ h) M" proof fix m1 m2 assume m12: "m1 \ M" "m2 \ M" "(lead_monom \ h) m1 = (lead_monom \ h) m2" hence [simp]: "coeff p m1 \ 0" "coeff p m2 \ 0" by (auto simp: M_def h_def) have "restrictpm A (lead_monom (h m1)) = restrictpm A (lead_monom (h m2))" using m12 by simp hence "lead_monom (\i. sym_mpoly A i ^ lookup m1 i :: 'a mpoly) = lead_monom (\i. sym_mpoly A i ^ lookup m2 i :: 'a mpoly)" by (simp add: restrict_lead_monom) thus "m1 = m2" by (rule lead_monom_sym_poly_powerprod_inj) (use \finite A\ keys_subset[of m1] keys_subset[of m2] in auto) qed next fix m assume "m \ M" hence "lead_coeff (h m) = lead_coeff (coeff p m)" by (simp add: lead_coeff M_def) with \m \ M\ show "h m \ 0" by (auto simp: M_def) qed fact+ finally have "maxm = 0" by (simp add: assms) have only_zero: "m = 0" if "m \ M" for m proof - from that have nz [simp]: "coeff p m \ 0" by (auto simp: M_def h_def) from that have "(lead_monom \ h) m \ maxm" using \finite M\ unfolding maxm_def by (intro Max_ge imageI finite_imageI) with \maxm = 0\ have [simp]: "lead_monom (h m) = 0" by simp have lookup_nzD: "k \ {1..n}" if "lookup m k \ 0" for k using keys_subset[of m] that by (auto simp: in_keys_iff subset_iff) have "lead_monom (coeff p m) + 0 \ lead_monom (h m)" unfolding lead_monom[OF nz] by (intro add_left_mono) auto also have "\ = 0" by simp finally have lead_monom_0: "lead_monom (coeff p m) = 0" by simp have "sum (lookup m) {1..n} = 0" proof (rule ccontr) assume "sum (lookup m) {1..n} \ 0" hence "sum (lookup m) {1..n} > 0" by presburger have "0 \ lead_coeff (MPoly_Type.coeff p m)" by auto also have "lead_coeff (MPoly_Type.coeff p m) = lead_coeff (h m)" by (simp add: lead_coeff) also have "lead_coeff (h m) = coeff (h m) 0" by (simp add: lead_coeff_def) also have "\ = coeff (coeff p m) 0 * coeff (\i. sym_mpoly A i ^ lookup m i) 0" by (simp add: h_def mpoly_coeff_times_0) also have "(\i. sym_mpoly A i ^ lookup m i) = (\i=1..n. sym_mpoly A i ^ lookup m i)" by (intro Prod_any.expand_superset subsetI lookup_nzD) (auto intro!: Nat.gr0I) also have "coeff \ 0 = (\i=1..n. 0 ^ lookup m i)" unfolding mpoly_coeff_prod_0 mpoly_coeff_power_0 by (intro prod.cong) (auto simp: coeff_sym_mpoly_0) also have "\ = 0 ^ (\i=1..n. lookup m i)" by (simp add: power_sum) also have "\ = 0" using zero_power[OF \sum (lookup m) {1..n} > 0\] by simp finally show False by auto qed hence "lookup m k = 0" for k using keys_subset[of m] by (cases "k \ {1..n}") (auto simp: in_keys_iff) thus "m = 0" by (intro poly_mapping_eqI) auto qed have "0 = insertion (sym_mpoly A) p" using assms by simp also have "insertion (sym_mpoly A) p = (\m\M. h m)" by fact also have "\ = (\m\{0}. h m)" using only_zero by (intro sum.mono_neutral_left) (auto simp: h_def M_def) also have "\ = coeff p 0" by (simp add: h_def) finally have "0 \ M" by (auto simp: M_def) with only_zero have "M = {}" by auto hence "p = 0" by (intro mpoly_eqI) (auto simp: M_def) with \p \ 0\ show False by contradiction qed text \ The general uniqueness theorem now follows easily. This essentially shows that the substitution $Y_i \mapsto e_i(X_1,\ldots,X_n)$ is an isomorphism between the ring $R[Y_1,\ldots, Y_n]$ and the ring $R[X_1,\ldots,X_n]^{S_n}$ of symmetric polynomials. \ theorem sym_mpoly_representation_unique: fixes p :: "'a mpoly mpoly" assumes "finite A" "insertion (sym_mpoly A) p = insertion (sym_mpoly A) q" "\m. vars (coeff p m) \ A = {}" "\m. vars (coeff q m) \ A = {}" "vars p \ {1..n}" "vars q \ {1..n}" shows "p = q" proof - have "p - q = 0" proof (rule sym_mpoly_representation_unique_aux) fix m show "vars (coeff (p - q) m) \ A = {}" using vars_diff[of "coeff p m" "coeff q m"] assms(3,4)[of m] by auto qed (insert assms vars_diff[of p q], auto simp: insertion_diff) thus ?thesis by simp qed theorem eq_fund_sym_poly_witI: fixes p :: "'a mpoly" and q :: "'a mpoly mpoly" assumes "finite A" "symmetric_mpoly A p" "insertion (sym_mpoly A) q = p" "\m. vars (coeff q m) \ A = {}" "vars q \ {1..n}" shows "q = fund_sym_poly_wit p" using fund_sym_poly_wit_insertion[of p] fund_sym_poly_wit_vars[of p] fund_sym_poly_wit_coeffs[of p] by (intro sym_mpoly_representation_unique) (insert assms, auto simp: fund_sym_poly_wit_insertion) subsection \A recursive characterisation of symmetry\ text \ In a similar spirit to the proof of the fundamental theorem, we obtain a nice recursive and executable characterisation of symmetry. \ (*<*) lemmas [fundef_cong] = disj_cong conj_cong (*>*) function (domintros) check_symmetric_mpoly where "check_symmetric_mpoly p \ (vars p \ A = {} \ A \ vars p \ decr p \ check_symmetric_mpoly (p - fund_sym_step_poly p))" by auto lemma check_symmetric_mpoly_dom_aux: assumes "finite B" "vars p \ B" "A \ B" shows "check_symmetric_mpoly_dom p" using assms(1-3) proof (induction p rule: lead_monom_induct) case (less p) have [simp]: "finite A" by (rule finite_subset[of _ B]) fact+ show ?case proof (cases "lead_monom p = 0 \ \decr p") case False hence [simp]: "p \ 0" by auto have "vars (p - fund_sym_step_poly p) \ B" using fund_sym_step_poly[of p] False less.prems less.hyps \A \ B\ by (intro order.trans[OF vars_diff]) auto hence "check_symmetric_mpoly_dom (p - local.fund_sym_step_poly p)" using False less.prems less.hyps by (intro less.IH fund_sym_step_poly symmetric_mpoly_diff lead_monom_fund_sym_step_poly_less) (auto simp: decr_def) thus ?thesis using check_symmetric_mpoly.domintros by blast qed (auto intro: check_symmetric_mpoly.domintros simp: lead_monom_eq_0_iff) qed lemma check_symmetric_mpoly_dom [intro]: "check_symmetric_mpoly_dom p" proof - show ?thesis proof (cases "A \ vars p") assume "A \ vars p" thus ?thesis using check_symmetric_mpoly_dom_aux[of "vars p" p] by (auto simp: vars_finite) qed (auto intro: check_symmetric_mpoly.domintros) qed termination check_symmetric_mpoly by (intro allI check_symmetric_mpoly_dom) lemmas [simp del] = check_symmetric_mpoly.simps lemma check_symmetric_mpoly_correct: "check_symmetric_mpoly p \ symmetric_mpoly A p" proof (induction p rule: check_symmetric_mpoly.induct) case (1 p) have "symmetric_mpoly A (p - fund_sym_step_poly p) \ symmetric_mpoly A p" (is "?lhs = ?rhs") proof assume ?rhs thus ?lhs by (intro symmetric_mpoly_diff fund_sym_step_poly) next assume ?lhs hence "symmetric_mpoly A (p - fund_sym_step_poly p + fund_sym_step_poly p)" by (intro symmetric_mpoly_add fund_sym_step_poly) thus ?rhs by simp qed moreover have "decr p" if "symmetric_mpoly A p" using lookup_lead_monom_decreasing[of A p] that by (auto simp: decr_def) ultimately show "check_symmetric_mpoly p \ symmetric_mpoly A p" using 1 symmetric_mpoly_imp_orthogonal_or_subset[of A p] by (auto simp: Let_def check_symmetric_mpoly.simps[of p] intro: symmetric_mpoly_orthogonal) qed end subsection \Symmetric functions of roots of a univariate polynomial\ text \ Consider a factored polynomial \[p(X) = c_n X^n + c_{n-1} X^{n-1} + \ldots + c_1X + c_0 = (X - x_1)\ldots(X - x_n)\ .\] where $c_n$ is a unit. Then any symmetric polynomial expression $q(x_1, \ldots, x_n)$ in the roots $x_i$ can be written as a polynomial expression $q'(c_0,\ldots, c_{n-1})$ in the $c_i$. Moreover, if the coefficients of $q$ and the inverse of $c_n$ all lie in some subring, the coefficients of $q'$ do as well. \ context fixes C :: "'b :: comm_ring_1 set" and A :: "nat set" and root :: "nat \ 'a :: comm_ring_1" and l :: "'a \ 'b" and q :: "'b mpoly" and n :: nat defines "n \ card A" assumes C: "ring_closed C" "\m. coeff q m \ C" assumes l: "ring_homomorphism l" assumes finite: "finite A" assumes sym: "symmetric_mpoly A q" and vars: "vars q \ A" begin interpretation ring_closed C by fact interpretation ring_homomorphism l by fact theorem symmetric_poly_of_roots_conv_poly_of_coeffs: assumes c: "cinv * l c = 1" "cinv \ C" assumes "p = Polynomial.smult c (\i\A. [:-root i, 1:])" obtains q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" proof - define q' where "q' = fund_sym_poly_wit A q" define q'' where "q'' = mapm_mpoly (\m x. (\i. (cinv * l (- 1) ^ i) ^ lookup m i) * insertion (\_. 0) x) q'" define reindex where "reindex = (\i. if i \ n then n - i else i)" have "bij reindex" by (intro bij_betwI[of reindex _ _ reindex]) (auto simp: reindex_def) have "vars q' \ {1..n}" unfolding q'_def n_def by (intro fund_sym_poly_wit_vars) hence "vars q'' \ {1..n}" unfolding q''_def using vars_mapm_mpoly_subset by auto have "insertion (l \ root) (insertion (sym_mpoly A) q') = insertion (\n. insertion (l \ root) (sym_mpoly A n)) (map_mpoly (insertion (l \ root)) q')" by (rule insertion_insertion) also have "insertion (sym_mpoly A) q' = q" unfolding q'_def by (intro fund_sym_poly_wit_insertion sym) also have "insertion (\i. insertion (l \ root) (sym_mpoly A i)) (map_mpoly (insertion (l \ root)) q') = insertion (\i. cinv * l ((- 1) ^ i) * l (poly.coeff p (n - i))) (map_mpoly (insertion (l \ root)) q')" proof (intro insertion_irrelevant_vars, goal_cases) case (1 i) hence "i \ vars q'" using vars_map_mpoly_subset by auto also have "\ \ {1..n}" unfolding q'_def n_def by (intro fund_sym_poly_wit_vars) finally have i: "i \ {1..n}" . have "insertion (l \ root) (sym_mpoly A i) = l (\Y | Y \ A \ card Y = i. prod root Y)" using \finite A\ by (simp add: insertion_sym_mpoly) also have "\ = cinv * l (c * (\Y | Y \ A \ card Y = i. prod root Y))" unfolding mult mult.assoc[symmetric] \cinv * l c = 1\ by simp also have "c * (\Y | Y \ A \ card Y = i. prod root Y) = ((-1) ^ i * poly.coeff p (n - i))" using coeff_poly_from_roots[of A "n - i" root] i assms finite by (auto simp: n_def minus_one_power_iff) finally show ?case by (simp add: o_def) qed also have "map_mpoly (insertion (l \ root)) q' = map_mpoly (insertion (\_. 0)) q'" using fund_sym_poly_wit_coeffs[OF sym] vars by (intro map_mpoly_cong insertion_irrelevant_vars) (auto simp: q'_def) also have "insertion (\i. cinv * l ((- 1) ^ i) * l (poly.coeff p (n - i))) \ = insertion (\i. l (poly.coeff p (n - i))) q''" unfolding insertion_substitute_linear map_mpoly_conv_mapm_mpoly q''_def by (subst mapm_mpoly_comp) auto also have "\ = insertion (l \ poly.coeff p) (mpoly_map_vars reindex q'')" using \bij reindex\ and \vars q'' \ {1..n}\ by (subst insertion_mpoly_map_vars) (auto simp: o_def reindex_def intro!: insertion_irrelevant_vars) finally have "insertion (l \ root) q = insertion (l \ poly.coeff p) (mpoly_map_vars reindex q'')" . moreover have "coeff (mpoly_map_vars reindex q'') m \ C" for m unfolding q''_def q'_def using \bij reindex\ fund_sym_poly_wit_coeff[of q C A] C \cinv \ C\ by (auto simp: coeff_mpoly_map_vars intro!: mult_closed Prod_any_closed power_closed Sum_any_closed) moreover have "vars (mpoly_map_vars reindex q'') \ {0..bij reindex\ and \vars q'' \ {1..n}\ by (subst vars_mpoly_map_vars) (auto simp: reindex_def subset_iff)+ ultimately show ?thesis using that[of "mpoly_map_vars reindex q''"] by auto qed corollary symmetric_poly_of_roots_conv_poly_of_coeffs_monic: assumes "p = (\i\A. [:-root i, 1:])" obtains q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" proof - obtain q' where "vars q' \ {0..m. coeff q' m \ C" and "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by (rule symmetric_poly_of_roots_conv_poly_of_coeffs[of 1 1 p]) (use assms in auto) thus ?thesis by (intro that[of q']) auto qed text \ As a corollary, we obtain the following: Let $R, S$ be rings with $R\subseteq S$. Consider a polynomial $p\in R[X]$ whose leading coefficient $c$ is a unit in $R$ and that has a full set of roots $x_1,\ldots, x_n \in S$, i.\,e.\ $p(X) = c(X - x_1) \ldots (X - x_n)$. Let $q \in R[X_1,\ldots, X_n]$ be some symmetric polynomial expression in the roots. Then $q(x_1, \ldots, x_n) \in R$. A typical use case is $R = \mathbb{Q}$ and $S = \mathbb{C}$, i.\,e.\ any symmetric polynomial expression with rational coefficients in the roots of a rational polynomial is again rational. Similarly, any symmetric polynomial expression with integer coefficients in the roots of a monic integer polynomial is agan an integer. This is remarkable, since the roots themselves are usually not rational (possibly not even real). This particular fact is a key ingredient used in the standard proof that $\pi$ is transcendental. \ corollary symmetric_poly_of_roots_in_subring: assumes "cinv * l c = 1" "cinv \ C" assumes "p = Polynomial.smult c (\i\A. [:-root i, 1:])" assumes "\i. l (poly.coeff p i) \ C" shows "insertion (\x. l (root x)) q \ C" proof - obtain q' where q': "vars q' \ {0..m. coeff q' m \ C" "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by (rule symmetric_poly_of_roots_conv_poly_of_coeffs[of cinv c p]) (use assms in simp_all) have "insertion (l \ poly.coeff p) q' \ C" using C assms unfolding insertion_altdef by (intro Sum_any_closed mult_closed q' Prod_any_closed power_closed) auto also have "insertion (l \ poly.coeff p) q' = insertion (l \ root) q" by fact finally show ?thesis by (simp add: o_def) qed corollary symmetric_poly_of_roots_in_subring_monic: assumes "p = (\i\A. [:-root i, 1:])" assumes "\i. l (poly.coeff p i) \ C" shows "insertion (\x. l (root x)) q \ C" proof - interpret ring_closed C by fact interpret ring_homomorphism l by fact show ?thesis by (rule symmetric_poly_of_roots_in_subring[of 1 1 p]) (use assms in auto) qed end end \ No newline at end of file diff --git a/thys/Tree-Automata/AbsAlgo.thy b/thys/Tree-Automata/AbsAlgo.thy --- a/thys/Tree-Automata/AbsAlgo.thy +++ b/thys/Tree-Automata/AbsAlgo.thy @@ -1,1382 +1,1382 @@ (* Title: Tree Automata Author: Peter Lammich Maintainer: Peter Lammich *) section "Abstract Tree Automata Algorithms" theory AbsAlgo imports Ta Collections_Examples.Exploration Collections.CollectionsV1 begin no_notation fun_rel_syn (infixr "\" 60) text_raw \\label{sec:absalgo}\ text \This theory defines tree automata algorithms on an abstract level, that is using non-executable datatypes and constructs like sets, set-collecting operations, etc. These algorithms are then refined to executable algorithms in Section~\ref{sec:taimpl}. \ subsection \Word Problem\ text \ First, a recursive version of the @{const accs}-predicate is defined. \ fun r_match :: "'a set list \ 'a list \ bool" where "r_match [] [] \ True" | "r_match (A#AS) (a#as) \ a\A \ r_match AS as" | "r_match _ _ \ False" \ \@{const r_match} accepts two lists, if they have the same length and the elements in the second list are contained in the respective elements of the first list:\ lemma r_match_alt: "r_match L l \ length L = length l \ (\i L!i)" apply (induct L l rule: r_match.induct) apply auto apply (case_tac i) apply auto done \ \Whether a rule matches the given state, label and list of sets of states\ fun r_matchc where "r_matchc q l Qs (qr \ lr qsr) \ q=qr \ l=lr \ r_match Qs qsr" \ \recursive version of @{const accs}-predicate\ fun faccs :: "('Q,'L) ta_rule set \ 'L tree \ 'Q set" where "faccs \ (NODE f ts) = ( let Qs = map (faccs \) (ts) in {q. \r\\. r_matchc q f Qs r } )" lemma faccs_correct_aux: "q\faccs \ n = accs \ n q" (is ?T1) "(map (faccs \) ts = map (\t. { q . accs \ t q}) ts)" (is ?T2) proof - have "(\q. q\faccs \ n = accs \ n q) \ (map (faccs \) ts = map (\t. { q . accs \ t q}) ts)" proof (induct rule: compat_tree_tree_list.induct) case (NODE f ts) thus ?case apply (intro allI iffI) apply simp apply (erule bexE) apply (case_tac x) apply simp apply (rule accs.intros) apply assumption apply (unfold r_match_alt) apply simp apply fastforce apply simp apply (erule accs.cases) apply auto apply (rule_tac x="qa \ f qs" in bexI) apply simp apply (unfold r_match_alt) apply auto done qed auto thus ?T1 ?T2 by auto qed theorem faccs_correct1: "q\faccs \ n \ accs \ n q" by (simp add: faccs_correct_aux) theorem faccs_correct2: "accs \ n q \ q\faccs \ n" by (simp add: faccs_correct_aux) lemmas faccs_correct = faccs_correct1 faccs_correct2 lemma faccs_alt: "faccs \ t = {q. accs \ t q}" by (auto intro: faccs_correct) subsection \Backward Reduction and Emptiness Check\ subsubsection "Auxiliary Definitions" \ \Step function, that maps a set of states to those states that are reachable via one backward step.\ inductive_set bacc_step :: "('Q,'L) ta_rule set \ 'Q set \ 'Q set" for \ Q where "\ r\\; set (rhsq r) \ Q \ \ lhs r \ bacc_step \ Q" \ \If a set is closed under adding all states that are reachable from the set by one backward step, then this set contains all backward accessible states.\ lemma b_accs_as_closed: assumes A: "bacc_step \ Q \ Q" shows "b_accessible \ \ Q" proof (rule subsetI) fix q assume "q\b_accessible \" thus "q\Q" proof (induct rule: b_accessible.induct) fix q f qs assume BC: "(q\f qs)\\" "!!x. x\set qs \ x\b_accessible \" "!!x. x\set qs \ x\Q" from bacc_step.intros[OF BC(1)] BC(3) have "q\bacc_step \ Q" by auto with A show "q\Q" by blast qed qed subsubsection "Algorithms" text \ First, the basic workset algorithm is specified. Then, it is refined to contain a counter for each rule, that counts the number of undiscovered states on the RHS. For both levels of abstraction, a version that computes the backwards reduction, and a version that checks for emptiness is specified. Additionally, a version of the algorithm that computes a witness for non-emptiness is provided. Levels of abstraction: \begin{itemize} \item[\\\] On this level, the state consists of a set of discovered states and a workset. \item[\\'\] On this level, the state consists of a set of discovered states, a workset and a map from rules to number of undiscovered rhs states. This map can be used to make the discovery of rules that have to be considered more efficient. \end{itemize} \ text_raw \\paragraph {\\\ - Level:}\ \ \A state contains the set of discovered states and a workset\ type_synonym ('Q,'L) br_state = "'Q set \ 'Q set" \ \Set of states that are non-empty (accept a tree) after adding the state $q$ to the set of discovered states\ definition br_dsq :: "('Q,'L) ta_rule set \ 'Q \ ('Q,'L) br_state \ 'Q set" where "br_dsq \ q == \(Q,W). { lhs r | r. r\\ \ set (rhsq r) \ (Q-(W-{q})) }" \ \Description of a step: One state is removed from the workset, and all new states that become non-empty due to this state are added to, both, the workset and the set of discovered states\ inductive_set br_step :: "('Q,'L) ta_rule set \ (('Q,'L) br_state \ ('Q,'L) br_state) set" for \ where "\ q\W; Q' = Q \ br_dsq \ q (Q,W); W' = W - {q} \ (br_dsq \ q (Q,W) - Q) \ \ ((Q,W),(Q',W'))\br_step \" \ \Termination condition for backwards reduction: The workset is empty\ definition br_cond :: "('Q,'L) br_state set" where "br_cond == {(Q,W). W\{}}" \ \Termination condition for emptiness check: The workset is empty or a non-empty initial state has been discovered\ definition bre_cond :: "'Q set \ ('Q,'L) br_state set" where "bre_cond Qi == {(Q,W). W\{} \ (Qi\Q={})}" \ \Set of all states that occur on the lhs of a constant-rule\ definition br_iq :: "('Q,'L) ta_rule set \ 'Q set" where "br_iq \ == { lhs r | r. r\\ \ rhsq r = [] }" \ \Initial state for the iteration\ definition br_initial :: "('Q,'L) ta_rule set \ ('Q,'L) br_state" where "br_initial \ == (br_iq \, br_iq \)" \ \Invariant for the iteration: \begin{itemize} \item States on the workset have been discovered \item Only accessible states have been discovered \item If a state is non-empty due to a rule whose rhs-states have been discovered and processed (i.e. are in $Q-W$), then the lhs state of the rule has also been discovered. \item The set of discovered states is finite \end{itemize}\ definition br_invar :: "('Q,'L) ta_rule set \ ('Q,'L) br_state set" where "br_invar \ == {(Q,W). W\Q \ Q \ b_accessible \ \ bacc_step \ (Q - W) \ Q \ finite Q}" definition "br_algo \ == \ wa_cond = br_cond, wa_step = br_step \, wa_initial = {br_initial \}, wa_invar = br_invar \ \" definition "bre_algo Qi \ == \ wa_cond = bre_cond Qi, wa_step = br_step \, wa_initial = {br_initial \}, wa_invar = br_invar \ \" \ \Termination: Either a new state is added, or the workset decreases.\ definition "br_termrel \ == ({(Q',Q). Q \ Q' \ Q' \ b_accessible \}) <*lex*> finite_psubset" lemma bre_cond_imp_br_cond[intro, simp]: "bre_cond Qi \ br_cond" by (auto simp add: br_cond_def bre_cond_def) lemma br_termrel_wf[simp, intro!]: "finite \ \ wf (br_termrel \)" apply (unfold br_termrel_def) apply (auto simp add: wf_bounded_supset) done \ \Only accessible states are discovered\ lemma br_dsq_ss: assumes A: "(Q,W)\br_invar \" "W \ {}" "q\W" shows "br_dsq \ q (Q,W) \ b_accessible \" proof (rule subsetI) fix q' assume B: "q'\br_dsq \ q (Q,W)" then obtain r where R: "q' = lhs r" "r\\" and S: "set (rhsq r) \ (Q-(W-{q}))" by (unfold br_dsq_def) auto note S also have "(Q-(W-{q})) \ b_accessible \" using A(1,3) by (auto simp add: br_invar_def) finally show "q'\b_accessible \" using R by (cases r) (auto intro: b_accessible.intros) qed lemma br_step_in_termrel: assumes A: "\\br_cond" "\\br_invar \" "(\,\')\br_step \" shows "(\', \)\br_termrel \" proof - obtain Q W Q' W' where [simp]: "\=(Q,W)" "\'=(Q',W')" by (cases \, cases \', auto) obtain q where QIW: "q\W" and ASSFMT[simp]: "Q' = Q \ br_dsq \ q (Q, W)" "W' = W - {q} \ (br_dsq \ q (Q, W) - Q)" by (auto intro: br_step.cases[OF A(3)[simplified]]) from A(2) have [simp]: "finite Q" by (auto simp add: br_invar_def) from A(2)[unfolded br_invar_def] have [simp]: "finite W" by (auto simp add: finite_subset) from A(1) have WNE: "W\{}" by (unfold br_cond_def) auto note DSQSS = br_dsq_ss[OF A(2)[simplified] WNE QIW] { assume "br_dsq \ q (Q,W) - Q = {}" hence ?thesis using QIW by (simp add: br_termrel_def set_simps) } moreover { assume "br_dsq \ q (Q,W) - Q \ {}" hence "Q \ Q'" by auto moreover from DSQSS A(2)[unfolded br_invar_def] have "Q' \ b_accessible \" by auto ultimately have ?thesis - by (simp add: br_termrel_def) + by (auto simp add: br_termrel_def) } ultimately show ?thesis by blast qed lemma br_invar_initial[simp]: "finite \ \ (br_initial \)\br_invar \" apply (auto simp add: br_initial_def br_invar_def br_iq_def) apply (case_tac r) apply (fastforce intro: b_accessible.intros) apply (fastforce elim!: bacc_step.cases) done lemma br_invar_step: assumes [simp]: "finite \" assumes A: "\\br_cond" "\\br_invar \" "(\,\')\br_step \" shows "\'\br_invar \" proof - obtain Q W Q' W' where SF[simp]: "\=(Q,W)" "\'=(Q',W')" by (cases \, cases \', auto) obtain q where QIW: "q\W" and ASSFMT[simp]: "Q' = Q \ br_dsq \ q (Q, W)" "W' = W - {q} \ (br_dsq \ q (Q, W) - Q)" by (auto intro: br_step.cases[OF A(3)[simplified]]) from A(1) have WNE: "W\{}" by (unfold br_cond_def) auto have DSQSS: "br_dsq \ q (Q,W) \ b_accessible \" using br_dsq_ss[OF A(2)[simplified] WNE QIW] . show ?thesis apply (simp add: br_invar_def del: ASSFMT) proof (intro conjI) from A(2) have "W \ Q" by (simp add: br_invar_def) thus "W' \ Q'" by auto next from A(2) have "Q \ b_accessible \" by (simp add: br_invar_def) with DSQSS show "Q' \ b_accessible \" by auto next show "bacc_step \ (Q' - W') \ Q'" apply (rule subsetI) apply (erule bacc_step.cases) apply (auto simp add: br_dsq_def) done next show "finite Q'" using A(2) by (simp add: br_invar_def br_dsq_def) qed qed lemma br_invar_final: "\\. \\wa_invar (br_algo \) \ \\wa_cond (br_algo \) \ fst \ = b_accessible \" apply (simp add: br_invar_def br_cond_def br_algo_def) apply (auto intro: rev_subsetD[OF _ b_accs_as_closed]) done (* shows "\(Q,W)\br_invar \; (Q,W)\br_cond\ \ Q = b_accessible \" apply (simp add: br_invar_def br_cond_def) apply (auto intro: rev_subsetD[OF _ b_accs_as_closed]) done*) theorem br_while_algo: assumes FIN[simp]: "finite \" shows "while_algo (br_algo \)" apply (unfold_locales) apply (simp_all add: br_algo_def br_invar_step br_invar_initial br_step_in_termrel) apply (rule_tac r="br_termrel \" in wf_subset) apply (auto intro: br_step_in_termrel) done lemma bre_invar_final: "\\. \\wa_invar (bre_algo Qi \) \ \\wa_cond (bre_algo Qi \) \ ((Qi\fst \={}) \ (Qi \ b_accessible \ = {}))" apply (simp add: br_invar_def bre_cond_def bre_algo_def) apply safe apply (auto dest!: b_accs_as_closed) done theorem bre_while_algo: assumes FIN[simp]: "finite \" shows "while_algo (bre_algo Qi \)" apply (unfold_locales) apply (unfold bre_algo_def) apply (auto simp add: br_invar_initial br_step_in_termrel intro: br_invar_step dest: rev_subsetD[OF _ bre_cond_imp_br_cond]) apply (rule_tac r="br_termrel \" in wf_subset) apply (auto intro: br_step_in_termrel dest: rev_subsetD[OF _ bre_cond_imp_br_cond]) done text_raw \\paragraph{\\'\ - Level}\ text \ Here, an optimization is added: For each rule, the algorithm now maintains a counter that counts the number of undiscovered states on the rules RHS. Whenever a new state is discovered, this counter is decremented for all rules where the state occurs on the RHS. The LHS states of rules where the counter falls to 0 are added to the worklist. The idea is that decrementing the counter is more efficient than checking whether all states on the rule's RHS have been discovered. A similar algorithm is sketched in \cite{tata2007}(Exercise~1.18). \ type_synonym ('Q,'L) br'_state = "'Q set \ 'Q set \ (('Q,'L) ta_rule \ nat)" \ \Abstraction to @{text \}-level\ definition br'_\ :: "('Q,'L) br'_state \ ('Q,'L) br_state" where "br'_\ = (\(Q,W,rcm). (Q,W))" definition br'_invar_add :: "('Q,'L) ta_rule set \ ('Q,'L) br'_state set" where "br'_invar_add \ == {(Q,W,rcm). (\r\\. rcm r = Some (card (set (rhsq r) - (Q - W)))) \ {lhs r | r. r\\ \ the (rcm r) = 0} \ Q }" definition br'_invar :: "('Q,'L) ta_rule set \ ('Q,'L) br'_state set" where "br'_invar \ == br'_invar_add \ \ {\. br'_\ \ \ br_invar \}" inductive_set br'_step :: "('Q,'L) ta_rule set \ (('Q,'L) br'_state \ ('Q,'L) br'_state) set" for \ where "\ q\W; Q' = Q \ { lhs r | r. r\\ \ q \ set (rhsq r) \ the (rcm r) \ 1 }; W' = (W-{q}) \ ({ lhs r | r. r\\ \ q \ set (rhsq r) \ the (rcm r) \ 1 } - Q); !!r. r\\ \ rcm' r = ( if q \ set (rhsq r) then Some (the (rcm r) - 1) else rcm r ) \ \ ((Q,W,rcm),(Q',W',rcm')) \ br'_step \" definition br'_cond :: "('Q,'L) br'_state set" where "br'_cond == {(Q,W,rcm). W\{}}" definition bre'_cond :: "'Q set \ ('Q,'L) br'_state set" where "bre'_cond Qi == {(Q,W,rcm). W\{} \ (Qi\Q={})}" inductive_set br'_initial :: "('Q,'L) ta_rule set \ ('Q,'L) br'_state set" for \ where "\ !!r. r\\ \ rcm r = Some (card (set (rhsq r))) \ \ (br_iq \, br_iq \, rcm)\br'_initial \" definition "br'_algo \ == \ wa_cond=br'_cond, wa_step = br'_step \, wa_initial = br'_initial \, wa_invar = br'_invar \ \" definition "bre'_algo Qi \ == \ wa_cond=bre'_cond Qi, wa_step = br'_step \, wa_initial = br'_initial \, wa_invar = br'_invar \ \" lemma br'_step_invar: assumes finite[simp]: "finite \" assumes INV: "\\br'_invar_add \" "br'_\ \ \ br_invar \" assumes STEP: "(\,\') \ br'_step \" shows "\'\br'_invar_add \" proof - obtain Q W rcm where [simp]: "\=(Q,W,rcm)" by (cases \) auto obtain Q' W' rcm' where [simp]: "\'=(Q',W',rcm')" by (cases \') auto from STEP obtain q where STEPF: "q\W" "Q' = Q \ { lhs r | r. r\\ \ q \ set (rhsq r) \ the (rcm r) \ 1 }" "W' = (W-{q}) \ ({ lhs r | r. r\\ \ q \ set (rhsq r) \ the (rcm r) \ 1 } - Q)" "!!r. r\\ \ rcm' r = ( if q \ set (rhsq r) then Some (the (rcm r) - 1) else rcm r )" by (auto elim: br'_step.cases) from INV[unfolded br'_invar_def br_invar_def br'_invar_add_def br'_\_def, simplified] have INV: "(\r\\. rcm r = Some (card (set (rhsq r) - (Q - W))))" "{lhs r |r. r \ \ \ the (rcm r) = 0} \ Q" "W \ Q" "Q \ b_accessible \" "bacc_step \ (Q - W) \ Q" "finite Q" by auto { fix r assume A: "r\\" with INV(1) have RCMR: "rcm r = Some (card (set (rhsq r) - (Q - W)))" by auto have "rcm' r = Some (card (set (rhsq r) - (Q' - W')))" proof (cases "q\set (rhsq r)") case False with A STEPF(4) have "rcm' r = rcm r" by auto moreover from STEPF INV(3) False have "set (rhsq r) - (Q-W) = set (rhsq r) - (Q'-W')" by auto ultimately show ?thesis by (simp add: RCMR) next case True with A STEPF(4) RCMR have "rcm' r = Some ((card (set (rhsq r) - (Q - W))) - 1)" by simp moreover from STEPF INV(3) True have "set (rhsq r) - (Q-W) = insert q (set (rhsq r) - (Q'-W'))" "q\(set (rhsq r) - (Q'-W'))" by auto ultimately show ?thesis by (simp add: RCMR card_insert_disjoint') qed } moreover { fix r assume A: "r\\" "the (rcm' r) = 0" have "lhs r \ Q'" proof (cases "q\set (rhsq r)") case True with A(1) STEPF(4) have "rcm' r = Some (the (rcm r) - 1)" by auto with A(2) have "the (rcm r) - 1 = 0" by auto hence "the (rcm r) \ 1" by auto with STEPF(2) A(1) True show ?thesis by auto next case False with A(1) STEPF(4) have "rcm' r = rcm r" by auto with A(2) have "the (rcm r) = 0" by auto with A(1) INV(2) have "lhs r \ Q" by auto with STEPF(2) show ?thesis by auto qed } ultimately show ?thesis by (auto simp add: br'_invar_add_def) qed lemma br'_invar_initial: "br'_initial \ \ br'_invar_add \" apply safe apply (erule br'_initial.cases) apply (unfold br'_invar_add_def) apply (auto simp add: br_iq_def) done lemma br'_rcm_aux': "\ (Q,W,rcm)\br'_invar \; q\W \ \ {r \ \. q \ set (rhsq r) \ the (rcm r) \ Suc 0} = {r\\. q\set (rhsq r) \ set (rhsq r) \ (Q - (W-{q}))}" proof (intro subsetI equalityI, goal_cases) case prems: (1 r) hence B: "r\\" "q\set (rhsq r)" "the (rcm r) \ Suc 0" by auto from B(1,3) prems(1)[unfolded br'_invar_def br'_invar_add_def] have CARD: "card (set (rhsq r) - (Q - W)) \ Suc 0" by auto from prems(1)[unfolded br'_invar_def br_invar_def br'_\_def] have WSQ: "W\Q" by auto have "set (rhsq r) - (Q - W) = {q}" proof - from B(2) prems(2) have R1: "q\set (rhsq r) - (Q - W)" by auto moreover { fix x assume A: "x\q" "x\set (rhsq r) - (Q - W)" with R1 have "{x,q} \ set (rhsq r) - (Q - W)" by auto hence "card {x,q} \ card (set (rhsq r) - (Q - W))" by (auto simp add: card_mono) with CARD A(1) have False by auto } ultimately show ?thesis by auto qed with prems(2) WSQ have "set (rhsq r) \ Q - (W - {q})" by auto thus ?case using B(1,2) by auto next case prems: (2 r) hence B: "r\\" "q\set (rhsq r)" "set (rhsq r) \ Q - (W - {q})" by auto with prems(1)[unfolded br'_invar_def br'_invar_add_def br'_\_def br_invar_def] have IC: "W\Q" "the (rcm r) = card (set (rhsq r) - (Q - W))" by auto have "set (rhsq r) - (Q - W) \ {q}" using B(2,3) IC(1) by auto from card_mono[OF _ this] have "the (rcm r) \ Suc 0" by (simp add: IC(2)) with B(1,2) show ?case by auto qed lemma br'_rcm_aux: assumes A: "(Q,W,rcm)\br'_invar \" "q\W" shows "{lhs r |r. r \ \ \ q \ set (rhsq r) \ the (rcm r) \ Suc 0} = {lhs r | r. r\\ \ q\set (rhsq r) \ set (rhsq r) \ (Q - (W-{q}))}" proof - have "{lhs r |r. r \ \ \ q \ set (rhsq r) \ the (rcm r) \ Suc 0} = lhs ` {r \ \. q \ set (rhsq r) \ the (rcm r) \ Suc 0}" by auto also from br'_rcm_aux'[OF A] have "\ = lhs ` {r \ \. q \ set (rhsq r) \ set (rhsq r) \ Q - (W - {q})}" by simp also have "\ = {lhs r | r. r\\ \ q\set (rhsq r) \ set (rhsq r) \ (Q - (W-{q}))}" by auto finally show ?thesis . qed lemma br'_invar_QcD: "(Q,W,rcm) \ br'_invar \ \ {lhs r | r. r\\ \ set (rhsq r) \ (Q-W)} \ Q" proof (safe) fix r assume A: "(Q,W,rcm)\br'_invar \" "r\\" "set (rhsq r) \ Q - W" from A(1)[unfolded br'_invar_def br'_invar_add_def br'_\_def br_invar_def, simplified] have IC: "W \ Q" "finite Q" "(\r\\. rcm r = Some (card (set (rhsq r) - (Q - W))))" "{lhs r |r. r \ \ \ the (rcm r) = 0} \ Q" by auto from IC(3) A(2,3) have "the (rcm r) = 0" by auto with IC(4) A(2) show "lhs r \ Q" by auto qed lemma br'_rcm_aux2: "\ (Q,W,rcm)\br'_invar \; q\W \ \ Q \ br_dsq \ q (Q,W) = Q \ {lhs r |r. r \ \ \ q \ set (rhsq r) \ the (rcm r) \ Suc 0}" apply (simp only: br'_rcm_aux) apply (unfold br_dsq_def) apply simp apply (frule br'_invar_QcD) apply auto done lemma br'_rcm_aux3: "\ (Q,W,rcm)\br'_invar \; q\W \ \ br_dsq \ q (Q,W) - Q = {lhs r |r. r \ \ \ q \ set (rhsq r) \ the (rcm r) \ Suc 0} - Q" apply (simp only: br'_rcm_aux) apply (unfold br_dsq_def) apply simp apply (frule br'_invar_QcD) apply auto done lemma br'_step_abs: "\ \\br'_invar \; (\,\') \ br'_step \ \ \ (br'_\ \, br'_\ \')\br_step \" apply (cases \, cases \', simp) apply (erule br'_step.cases) apply (simp add: br'_\_def) apply (rule_tac q=q in br_step.intros) apply simp apply (simp only: br'_rcm_aux2) apply (simp only: br'_rcm_aux3) done lemma br'_initial_abs: "br'_\`(br'_initial \) = {br_initial \}" apply (force simp add: br_initial_def br'_\_def elim: br'_initial.cases intro: br'_initial.intros) done lemma br'_cond_abs: "\\br'_cond \ (br'_\ \) \ br_cond" by (cases \) (simp add: br'_cond_def br_cond_def br'_\_def image_Collect br'_algo_def br_algo_def) lemma bre'_cond_abs: "\\bre'_cond Qi \ (br'_\ \)\bre_cond Qi" by (cases \) (simp add: bre'_cond_def bre_cond_def br'_\_def image_Collect bre'_algo_def bre_algo_def) lemma br'_invar_abs: "br'_\`br'_invar \ \ br_invar \" by (auto simp add: br'_invar_def) theorem br'_pref_br: "wa_precise_refine (br'_algo \) (br_algo \) br'_\" apply unfold_locales apply (simp_all add: br'_algo_def br_algo_def) apply (simp_all add: br'_cond_abs br'_step_abs br'_invar_abs br'_initial_abs) done interpretation br'_pref: wa_precise_refine "br'_algo \" "br_algo \" "br'_\" using br'_pref_br . theorem br'_while_algo: "finite \ \ while_algo (br'_algo \)" apply (rule br'_pref.wa_intro) apply (simp add: br_while_algo) apply (simp_all add: br'_algo_def br_algo_def) apply (simp add: br'_invar_def) apply (erule (3) br'_step_invar) apply (simp add: br'_invar_initial) done lemma fst_br'_\: "fst (br'_\ s) = fst s" by (cases s) (simp add: br'_\_def) lemmas br'_invar_final = br'_pref.transfer_correctness[OF br_invar_final, unfolded fst_br'_\] theorem bre'_pref_br: "wa_precise_refine (bre'_algo Qi \) (bre_algo Qi \) br'_\" apply unfold_locales apply (simp_all add: bre'_algo_def bre_algo_def) apply (simp_all add: bre'_cond_abs br'_step_abs br'_invar_abs br'_initial_abs) done interpretation bre'_pref: wa_precise_refine "bre'_algo Qi \" "bre_algo Qi \" "br'_\" using bre'_pref_br . theorem bre'_while_algo: "finite \ \ while_algo (bre'_algo Qi \)" apply (rule bre'_pref.wa_intro) apply (simp add: bre_while_algo) apply (simp_all add: bre'_algo_def bre_algo_def) apply (simp add: br'_invar_def) apply (erule (3) br'_step_invar) apply (simp add: br'_invar_initial) done lemmas bre'_invar_final = bre'_pref.transfer_correctness[OF bre_invar_final, unfolded fst_br'_\] text_raw \\paragraph{Implementing a Step}\ text \ In this paragraph, it is shown how to implement a step of the br'-algorithm by iteration over the rules that have the discovered state on their RHS. \ definition br'_inner_step :: "('Q,'L) ta_rule \ ('Q,'L) br'_state \ ('Q,'L) br'_state" where "br'_inner_step == \r (Q,W,rcm). let c=the (rcm r) in ( if c\1 then insert (lhs r) Q else Q, if c\1 \ (lhs r) \ Q then insert (lhs r) W else W, rcm ( r \ (c-(1::nat))) ) " definition br'_inner_invar :: "('Q,'L) ta_rule set \ 'Q \ ('Q,'L) br'_state \ ('Q,'L) ta_rule set \ ('Q,'L) br'_state \ bool" where "br'_inner_invar rules q == \(Q,W,rcm) it (Q',W',rcm'). Q' = Q \ { lhs r | r. r\rules-it \ the (rcm r) \ 1 } \ W' = (W-{q}) \ ({ lhs r | r. r\rules-it \ the (rcm r) \ 1 } - Q) \ (\r. rcm' r = (if r\rules-it then Some (the (rcm r) - 1) else rcm r)) " lemma br'_inner_invar_imp_final: "\ q\W; br'_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) {} \' \ \ ((Q,W,rcm),\') \ br'_step \" apply (unfold br'_inner_invar_def) apply auto apply (rule br'_step.intros) apply assumption apply auto done lemma br'_inner_invar_step: "\ q\W; br'_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) it \'; r\it; it\{r\\. q\set (rhsq r)} \ \ br'_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) (it-{r}) (br'_inner_step r \') " apply (cases \', simp) apply (unfold br'_inner_invar_def br'_inner_step_def Let_def) apply auto done lemma br'_inner_invar_initial: "\ q\W \ \ br'_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) {r\\. q\set (rhsq r)} (Q,W-{q},rcm)" apply (simp add: br'_inner_invar_def) apply auto done lemma br'_inner_step_proof: fixes \s :: "'\ \ ('Q,'L) br'_state" fixes cstep :: "('Q,'L) ta_rule \ '\ \ '\" fixes \h :: "'\" fixes cinvar :: "('Q,'L) ta_rule set \ '\ \ bool" assumes iterable_set: "set_iteratei \ invar iteratei" assumes invar_initial: "cinvar {r\\. q\set (rhsq r)} \h" assumes invar_step: "!!it r \. \ r\it; it \ {r\\. q\set (rhsq r)}; cinvar it \ \ \ cinvar (it-{r}) (cstep r \)" assumes step_desc: "!!it r \. \ r\it; it\{r\\. q\set (rhsq r)}; cinvar it \ \ \ \s (cstep r \) = br'_inner_step r (\s \)" assumes it_set_desc: "invar it_set" "\ it_set = {r\\. q\set (rhsq r)}" assumes QIW[simp]: "q\W" assumes \_desc[simp]: "\s \ = (Q,W,rcm)" assumes \h_desc[simp]: "\s \h = (Q,W-{q},rcm)" shows "(\s \, \s (iteratei it_set (\_. True) cstep \h))\br'_step \" proof - interpret set_iteratei \ invar iteratei by fact show ?thesis apply (rule_tac I="\it \. cinvar it \ \ br'_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) it (\s \)" in iterate_rule_P) apply (simp_all add: it_set_desc invar_initial br'_inner_invar_initial invar_step step_desc br'_inner_invar_step) apply (rule br'_inner_invar_imp_final) apply (rule QIW) apply simp done qed text_raw \\paragraph{Computing Witnesses}\ text \ The algorithm is now refined further, such that it stores, for each discovered state, a witness for non-emptiness, i.e. a tree that is accepted with the discovered state. \ \ \A map from states to trees has the witness-property, if it maps states to trees that are accepted with that state:\ definition "witness_prop \ m == \q t. m q = Some t \ accs \ t q" \ \Construct a witness for the LHS of a rule, provided that the map contains witnesses for all states on the RHS:\ definition construct_witness :: "('Q \ 'L tree) \ ('Q,'L) ta_rule \ 'L tree" where "construct_witness Q r == NODE (rhsl r) (List.map (\q. the (Q q)) (rhsq r))" lemma witness_propD: "\witness_prop \ m; m q = Some t\ \ accs \ t q" by (auto simp add: witness_prop_def) lemma construct_witness_correct: "\ witness_prop \ Q; r\\; set (rhsq r) \ dom Q \ \ accs \ (construct_witness Q r) (lhs r)" apply (unfold construct_witness_def witness_prop_def) apply (cases r) apply simp apply (erule accs.intros) apply (auto dest: nth_mem) done lemma construct_witness_eq: "\ Q |` set (rhsq r) = Q' |` set (rhsq r) \ \ construct_witness Q r = construct_witness Q' r" apply (unfold construct_witness_def) apply auto apply (subgoal_tac "Q x = Q' x") apply (simp) apply (drule_tac x=x in fun_cong) apply auto done text \ The set of discovered states is refined by a map from discovered states to their witnesses: \ type_synonym ('Q,'L) brw_state = "('Q\'L tree) \ 'Q set \ (('Q,'L) ta_rule \ nat)" definition brw_\ :: "('Q,'L) brw_state \ ('Q,'L) br'_state" where "brw_\ = (\(Q,W,rcm). (dom Q,W,rcm))" definition brw_invar_add :: "('Q,'L) ta_rule set \ ('Q,'L) brw_state set" where "brw_invar_add \ == {(Q,W,rcm). witness_prop \ Q}" definition "brw_invar \ == brw_invar_add \ \ {s. brw_\ s \ br'_invar \}" (* TODO: This step description does not allow full flexibility, because we may want to construct new witnesses from other witnesses constructed in the same step! However, if we say t = construct_witness Q' r, may we run into cyclicity problems, where a cycle of witnesses may witness itself?. Hmm? As these cyclic witnesses would have to be infinite, they cannot exist? But, if we use a BFS search strategy, the current step description will compute minimal depth witnesses. The argumentation is, that: Initially, all witnesses of depth 1 (definitely minimal) are discovered A witness of depth n has children of length < n The states that are initially on the workset are all those with witnesses of depth 1. Thus, after they have been processed, all states with witnesses of depth 2 have been discovered. This argument can be iterated inductively. *) inductive_set brw_step :: "('Q,'L) ta_rule set \ (('Q,'L) brw_state \ ('Q,'L) brw_state) set" for \ where "\ q\W; dsqr = { r\\. q \ set (rhsq r) \ the (rcm r) \ 1 }; dom Q' = dom Q \ lhs`dsqr; !!q t. Q' q = Some t \ Q q = Some t \ (\r\dsqr. q=lhs r \ t=construct_witness Q r); W' = (W-{q}) \ (lhs`dsqr - dom Q); !!r. r\\ \ rcm' r = ( if q \ set (rhsq r) then Some (the (rcm r) - 1) else rcm r ) \ \ ((Q,W,rcm),(Q',W',rcm')) \ brw_step \" definition brw_cond :: "'Q set \ ('Q,'L) brw_state set" where "brw_cond Qi == {(Q,W,rcm). W\{} \ (Qi\dom Q={})}" inductive_set brw_iq :: "('Q,'L) ta_rule set \ ('Q \ 'L tree) set" for \ where "\ \q t. Q q = Some t \ (\r\\. rhsq r = [] \ q = lhs r \ t = NODE (rhsl r) []); \r\\. rhsq r = [] \ Q (lhs r) \ None \ \ Q \ brw_iq \" inductive_set brw_initial :: "('Q,'L) ta_rule set \ ('Q,'L) brw_state set" for \ where "\ !!r. r\\ \ rcm r = Some (card (set (rhsq r))); Q\brw_iq \ \ \ (Q, br_iq \, rcm)\brw_initial \" definition "brw_algo Qi \ == \ wa_cond=brw_cond Qi, wa_step = brw_step \, wa_initial = brw_initial \, wa_invar = brw_invar \ \" lemma brw_cond_abs: "\\brw_cond Qi \ (brw_\ \)\bre'_cond Qi" apply (cases \) apply (simp add: brw_cond_def bre'_cond_def brw_\_def) done lemma brw_initial_abs: "\\brw_initial \ \ brw_\ \ \ br'_initial \" apply (cases \, simp) apply (erule brw_initial.cases) apply (erule brw_iq.cases) apply (auto simp add: brw_\_def) apply (subgoal_tac "dom Qa = br_iq \") apply simp apply (rule br'_initial.intros) apply auto [1] apply (force simp add: br_iq_def) done lemma brw_invar_initial: "brw_initial \ \ brw_invar_add \" apply safe apply (unfold brw_invar_add_def) apply (auto simp add: witness_prop_def) apply (erule brw_initial.cases) apply (erule brw_iq.cases) apply auto proof goal_cases case prems: (1 q t rcm Q) from prems(3)[rule_format, OF prems(1)] obtain r where [simp]: "r\\" "rhsq r = []" "q=lhs r" "t=NODE (rhsl r) []" by blast have RF[simplified]: "r=((lhs r) \ (rhsl r) (rhsq r))" by (cases r) simp show ?case apply (simp) apply (rule accs.intros) apply (subst RF[symmetric]) apply auto done qed lemma brw_step_abs: "\ (\,\')\brw_step \ \ \ (brw_\ \, brw_\ \')\br'_step \" apply (cases \, cases \', simp) apply (erule brw_step.cases) apply (simp add: brw_\_def) apply hypsubst apply (rule br'_step.intros) apply assumption apply auto done lemma brw_step_invar: assumes FIN[simp]: "finite \" assumes INV: "\\brw_invar_add \" and BR'INV: "brw_\ \ \ br'_invar \" assumes STEP: "(\,\') \ brw_step \" shows "\'\brw_invar_add \" proof - obtain Q W rcm Q' W' rcm' where [simp]: "\=(Q,W,rcm)" "\'=(Q',W',rcm')" by (cases \, cases \') force from INV have WP: "witness_prop \ Q" by (simp_all add: brw_invar_add_def) obtain qw dsqr where SPROPS: "dsqr = {r \ \. qw \ set (rhsq r) \ the (rcm r) \ 1}" "qw\W" "dom Q' = dom Q \ lhs ` dsqr" "!!q t. Q' q = Some t \ Q q = Some t \ (\r\dsqr. q=lhs r \ t=construct_witness Q r)" by (auto intro: brw_step.cases[OF STEP[simplified]]) from br'_rcm_aux'[OF BR'INV[unfolded brw_\_def, simplified] SPROPS(2)] have DSQR_ALT: "dsqr = {r \ \. qw \ set (rhsq r) \ set (rhsq r) \ dom Q - (W - {qw})}" by (simp add: SPROPS(1)) have "witness_prop \ Q'" proof (unfold witness_prop_def, safe) fix q t assume A: "Q' q = Some t" from SPROPS(4)[OF A] have "Q q = Some t \ (\r\dsqr. q = lhs r \ t = construct_witness Q r)" . moreover { assume C: "Q q = Some t" from witness_propD[OF WP, OF C] have "accs \ t q" . } moreover { fix r assume "r\dsqr" and [simp]: "q=lhs r" "t=construct_witness Q r" from \r\dsqr\ have 1: "r\\" "set (rhsq r) \ dom Q" by (auto simp add: DSQR_ALT) from construct_witness_correct[OF WP 1] have "accs \ t q" by simp } ultimately show "accs \ t q" by blast qed thus ?thesis by (simp add: brw_invar_add_def) qed theorem brw_pref_bre': "wa_precise_refine (brw_algo Qi \) (bre'_algo Qi \) brw_\" apply (unfold_locales) apply (simp_all add: brw_algo_def bre'_algo_def) apply (auto simp add: brw_cond_abs brw_step_abs brw_initial_abs brw_invar_def) done interpretation brw_pref: wa_precise_refine "brw_algo Qi \" "bre'_algo Qi \" "brw_\" using brw_pref_bre' . theorem brw_while_algo: "finite \ \ while_algo (brw_algo Qi \)" apply (rule brw_pref.wa_intro) apply (simp add: bre'_while_algo) apply (simp_all add: brw_algo_def bre'_algo_def) apply (simp add: brw_invar_def) apply (auto intro: brw_step_invar simp add: brw_invar_initial) done lemma fst_brw_\: "fst (brw_\ s) = dom (fst s)" by (cases s) (simp add: brw_\_def) theorem brw_invar_final: "\sc. sc \ wa_invar (brw_algo Qi \) \ sc \ wa_cond (brw_algo Qi \) \ (Qi \ dom (fst sc) = {}) = (Qi \ b_accessible \ = {}) \ (witness_prop \ (fst sc))" apply (intro conjI allI impI) using brw_pref.transfer_correctness[OF bre'_invar_final, unfolded fst_brw_\] apply blast apply (auto simp add: brw_algo_def brw_invar_def brw_invar_add_def) done text_raw \\paragraph{Implementing a Step}\ inductive_set brw_inner_step :: "('Q,'L) ta_rule \ (('Q,'L) brw_state \ ('Q,'L) brw_state) set" for r where "\ c = the (rcm r); \ = (Q,W,rcm); \'=(Q',W',rcm'); if c\1 \ (lhs r) \ dom Q then Q' = Q(lhs r \ construct_witness Q r) else Q' = Q; if c\1 \ (lhs r) \ dom Q then W' = insert (lhs r) W else W' = W; rcm' = rcm ( r \ (c-(1::nat))) \ \ (\,\')\brw_inner_step r" definition brw_inner_invar :: "('Q,'L) ta_rule set \ 'Q \ ('Q,'L) brw_state \ ('Q,'L) ta_rule set \ ('Q,'L) brw_state \ bool" where "brw_inner_invar rules q == \(Q,W,rcm) it (Q',W',rcm'). (br'_inner_invar rules q (brw_\ (Q,W,rcm)) it (brw_\ (Q',W',rcm')) \ (Q'|`dom Q = Q) \ (let dsqr = { r\rules - it. the (rcm r) \ 1 } in (\q t. Q' q = Some t \ (Q q = Some t \ (Q q = None \ (\r\dsqr. q=lhs r \ t=construct_witness Q r)) ) ))) " lemma brw_inner_step_abs: "(\,\')\brw_inner_step r \ br'_inner_step r (brw_\ \) = brw_\ \'" apply (erule brw_inner_step.cases) apply (unfold br'_inner_step_def brw_\_def Let_def) apply auto done lemma brw_inner_invar_imp_final: "\ q\W; brw_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) {} \' \ \ ((Q,W,rcm),\') \ brw_step \" apply (unfold brw_inner_invar_def br'_inner_invar_def brw_\_def) apply (auto simp add: Let_def) apply (rule brw_step.intros) apply assumption apply (rule refl) apply auto done lemma brw_inner_invar_step: assumes INVI: "(Q,W,rcm)\brw_invar \" assumes A: "q\W" "r\it" "it\{r\\. q\set (rhsq r)}" assumes INVH: "brw_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) it \h" assumes STEP: "(\h,\')\brw_inner_step r" shows "brw_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) (it-{r}) \'" proof - from INVI have BR'_INV: "(dom Q,W,rcm)\br'_invar \" by (simp add: brw_invar_def brw_\_def) obtain c Qh Wh rcmh Q' W' rcm' where SIGMAF[simp]: "\h=(Qh,Wh,rcmh)" "\'=(Q',W',rcm')" and CF[simp]: "c = the (rcmh r)" and SF: "if c\1 \ (lhs r) \ dom Qh then Q' = Qh(lhs r \ (construct_witness Qh r)) else Q' = Qh" "if c\1 \ (lhs r) \ dom Qh then W' = insert (lhs r) Wh else W' = Wh" "rcm' = rcmh ( r \ (c-(1::nat)))" by (blast intro: brw_inner_step.cases[OF STEP]) let ?rules = "{r\\. q\set (rhsq r)}" let ?dsqr = "\it. { r\?rules - it. the (rcm r) \ 1 }" from INVH have INVHF: "br'_inner_invar ?rules q (dom Q, W-{q}, rcm) (it) (dom Qh,Wh,rcmh)" "Qh|`dom Q = Q" "(\q t. Qh q = Some t \ (Q q = Some t \ (Q q = None \ (\r\?dsqr it. q=lhs r \ t=construct_witness Q r)) ) )" by (auto simp add: brw_inner_invar_def Let_def brw_\_def) from INVHF(1)[unfolded br'_inner_invar_def] have INV'HF: "dom Qh = dom Q \ lhs`?dsqr it" "(\r. rcmh r = (if r \ ?rules - it then Some (the (rcm r) - 1) else rcm r))" by auto from brw_inner_step_abs[OF STEP] br'_inner_invar_step[OF A(1) INVHF(1) A(2,3)] have G1: "br'_inner_invar ?rules q (dom Q, W-{q}, rcm) (it-{r}) (dom Q',W',rcm')" by (simp add: brw_\_def) moreover have "(\q t. Q' q = Some t \ (Q q = Some t \ ( Q q = None \ (\r\?dsqr (it-{r}). q=lhs r \ t=construct_witness Q r) ) ) )" (is ?G1) "Q'|`dom Q = Q" (is ?G2) proof - { assume C: "\ c\1 \ lhs r \ dom Qh" with SF have "Q'=Qh" by auto with INVHF(2,3) have ?G1 ?G2 by auto } moreover { assume C: "c\1" "lhs r\ dom Qh" with SF have Q'F: "Q'=Qh(lhs r \ (construct_witness Qh r))" by auto from C(2) INVHF(2) INV'HF(1) have G2: ?G2 by (auto simp add: Q'F) from C(1) INV'HF A have RI: "r\?dsqr (it-{r})" and DSS: "dom Q \ dom Qh" by (auto) from br'_rcm_aux'[OF BR'_INV A(1)] RI have RDQ: "set (rhsq r) \ dom Q" by auto with INVHF(2) have "Qh |` set (rhsq r) = Q |` set (rhsq r)" by (blast intro: restrict_map_subset_eq) hence [simp]: "construct_witness Qh r = construct_witness Q r" by (blast dest: construct_witness_eq) from DSS C(2) have [simp]: "Q (lhs r) = None" "Qh (lhs r) = None" by auto have G1: ?G1 proof (intro allI impI, goal_cases) case prems: (1 q t) { assume [simp]: "q=lhs r" from prems Q'F have [simp]: "t = (construct_witness Qh r)" by simp from RI have ?case by auto } moreover { assume "q\lhs r" with Q'F prems have "Qh q = Some t" by auto with INVHF(3) have ?case by auto } ultimately show ?case by blast qed note G1 G2 } ultimately show ?G1 ?G2 by blast+ qed ultimately show ?thesis by (unfold brw_inner_invar_def Let_def brw_\_def) auto qed lemma brw_inner_invar_initial: "\q\W\ \ brw_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) {r\\. q\set (rhsq r)} (Q,W-{q},rcm)" by (simp add: brw_inner_invar_def br'_inner_invar_initial brw_\_def) theorem brw_inner_step_proof: fixes \s :: "'\ \ ('Q,'L) brw_state" fixes cstep :: "('Q,'L) ta_rule \ '\ \ '\" fixes \h :: "'\" fixes cinvar :: "('Q,'L) ta_rule set \ '\ \ bool" assumes set_iterate: "set_iteratei \ invar iteratei" assumes invar_start: "(\s \)\brw_invar \" assumes invar_initial: "cinvar {r\\. q\set (rhsq r)} \h" assumes invar_step: "!!it r \. \ r\it; it \ {r\\. q\set (rhsq r)}; cinvar it \ \ \ cinvar (it-{r}) (cstep r \)" assumes step_desc: "!!it r \. \ r\it; it\{r\\. q\set (rhsq r)}; cinvar it \ \ \ (\s \, \s (cstep r \)) \ brw_inner_step r" assumes it_set_desc: "invar it_set" "\ it_set = {r\\. q\set (rhsq r)}" assumes QIW[simp]: "q\W" assumes \_desc[simp]: "\s \ = (Q,W,rcm)" assumes \h_desc[simp]: "\s \h = (Q,W-{q},rcm)" shows "(\s \, \s (iteratei it_set (\_. True) cstep \h))\brw_step \" proof - interpret set_iteratei \ invar iteratei by fact show ?thesis apply (rule_tac I="\it \. cinvar it \ \ brw_inner_invar {r\\. q\set (rhsq r)} q (Q,W-{q},rcm) it (\s \)" in iterate_rule_P) apply (auto simp add: it_set_desc invar_initial brw_inner_invar_initial invar_step step_desc brw_inner_invar_step[OF invar_start[simplified]] brw_inner_invar_imp_final[OF QIW]) done qed subsection \Product Automaton\ text \ The forward-reduced product automaton can be described as a state-space exploration problem. In this section, the DFS-algorithm for state-space exploration (cf. Theory~@{theory Collections_Examples.Exploration} in the Isabelle Collections Framework) is refined to compute the product automaton. \ type_synonym ('Q1,'Q2,'L) frp_state = "('Q1\'Q2) set \ ('Q1\'Q2) list \ (('Q1\'Q2),'L) ta_rule set" definition frp_\ :: "('Q1,'Q2,'L) frp_state \ ('Q1\'Q2) dfs_state" where "frp_\ S == let (Q,W,\)=S in (Q, W)" definition "frp_invar_add \1 \2 == { (Q,W,\d). \d = { r. r\\_prod \1 \2 \ lhs r \ Q - set W} }" definition frp_invar :: "('Q1, 'L) tree_automaton_rec \ ('Q2, 'L) tree_automaton_rec \ ('Q1,'Q2,'L) frp_state set" where "frp_invar T1 T2 == frp_invar_add (ta_rules T1) (ta_rules T2) \ { s. frp_\ s \ dfs_invar (ta_initial T1 \ ta_initial T2) (f_succ (\_prod (ta_rules T1) (ta_rules T2))) }" inductive_set frp_step :: "('Q1,'L) ta_rule set \ ('Q2,'L) ta_rule set \ (('Q1,'Q2,'L) frp_state \ ('Q1,'Q2,'L) frp_state) set" for \1 \2 where "\ W=(q1,q2)#Wtl; distinct Wn; set Wn = f_succ (\_prod \1 \2) `` {(q1,q2)} - Q; W'=Wn@Wtl; Q'=Q \ f_succ (\_prod \1 \2) `` {(q1,q2)}; \d'=\d \ {r\\_prod \1 \2. lhs r = (q1,q2) } \ \ ((Q,W,\d),(Q',W',\d'))\frp_step \1 \2" inductive_set frp_initial :: "'Q1 set \ 'Q2 set \ ('Q1,'Q2,'L) frp_state set" for Q10 Q20 where "\ distinct W; set W = Q10\Q20 \ \ (Q10\Q20,W,{}) \ frp_initial Q10 Q20" definition frp_cond :: "('Q1,'Q2,'L) frp_state set" where "frp_cond == {(Q,W,\d). W\[]}" definition "frp_algo T1 T2 == \ wa_cond = frp_cond, wa_step = frp_step (ta_rules T1) (ta_rules T2), wa_initial = frp_initial (ta_initial T1) (ta_initial T2), wa_invar = frp_invar T1 T2 \" \ \The algorithm refines the DFS-algorithm\ theorem frp_pref_dfs: "wa_precise_refine (frp_algo T1 T2) (dfs_algo (ta_initial T1 \ ta_initial T2) (f_succ (\_prod (ta_rules T1) (ta_rules T2)))) frp_\" apply unfold_locales apply (auto simp add: frp_algo_def frp_\_def frp_cond_def dfs_algo_def dfs_cond_def frp_invar_def elim!: frp_step.cases frp_initial.cases intro: dfs_step.intros dfs_initial.intros ) done interpretation frp_ref: wa_precise_refine "(frp_algo T1 T2)" "(dfs_algo (ta_initial T1 \ ta_initial T2) (f_succ (\_prod (ta_rules T1) (ta_rules T2))))" "frp_\" using frp_pref_dfs . \ \The algorithm is a well-defined while-algorithm\ theorem frp_while_algo: assumes TA: "tree_automaton T1" "tree_automaton T2" shows "while_algo (frp_algo T1 T2)" proof - interpret t1: tree_automaton T1 by fact interpret t2: tree_automaton T2 by fact have finite: "finite ((f_succ (\_prod (ta_rules T1) (ta_rules T2)))\<^sup>* `` (ta_initial T1 \ ta_initial T2))" proof - have "((f_succ (\_prod (ta_rules T1) (ta_rules T2)))\<^sup>* `` (ta_initial T1 \ ta_initial T2)) \ ((ta_initial T1 \ ta_initial T2) \ \_states (\_prod (ta_rules T1) (ta_rules T2)))" apply rule apply (drule f_accessible_subset[unfolded f_accessible_def]) apply auto done moreover have "finite \" by auto ultimately show ?thesis by (simp add: finite_subset) qed show ?thesis apply (rule frp_ref.wa_intro) apply (rule dfs_while_algo[OF finite]) apply (simp add: frp_algo_def dfs_algo_def frp_invar_def) apply (auto simp add: dfs_algo_def frp_algo_def frp_\_def dfs_\_def frp_invar_add_def dfs_invar_def dfs_invar_add_def sse_invar_def elim!: frp_step.cases) [1] apply (force simp add: frp_algo_def frp_invar_add_def elim!: frp_initial.cases) done qed (* unused lemma f_succ_adv: "\lhs r \ (f_succ \)\<^sup>* `` Q0; r\\\ \ set (rhsq r) \ (f_succ \)\<^sup>* `` Q0" by (case_tac r) (auto dest: rtrancl_into_rtrancl intro: f_succ.intros) *) \ \If the algorithm terminates, the forward reduced product automaton can be constructed from the result\ theorem frp_inv_final: "\s. s\wa_invar (frp_algo T1 T2) \ s\wa_cond (frp_algo T1 T2) \ (case s of (Q,W,\d) \ \ ta_initial = ta_initial T1 \ ta_initial T2, ta_rules = \d \ = ta_fwd_reduce (ta_prod T1 T2))" apply (intro allI impI) apply (case_tac s) apply simp apply (simp add: ta_reduce_def ta_prod_def frp_algo_def) proof - fix Q W \d assume A: "(Q,W,\d)\frp_invar T1 T2 \ (Q,W,\d)\frp_cond" from frp_ref.transfer_correctness[OF dfs_invar_final, unfolded frp_algo_def, simplified, rule_format, OF A] have [simp]: "Q = f_accessible (\_prod (ta_rules T1) (ta_rules T2)) (ta_initial T1 \ ta_initial T2)" by (simp add: f_accessible_def dfs_\_def frp_\_def) from A show "\d = reduce_rules (\_prod (ta_rules T1) (ta_rules T2)) (f_accessible (\_prod (ta_rules T1) (ta_rules T2)) (ta_initial T1 \ ta_initial T2))" apply (auto simp add: reduce_rules_def f_accessible_def frp_invar_def frp_invar_add_def frp_\_def frp_cond_def) apply (case_tac x) apply (auto dest: rtrancl_into_rtrancl intro: f_succ.intros) done qed end