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,469 +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 dest!: irrefl_subterm_T) + apply(auto simp add: wfP_def[symmetric] wfP_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/CAVA_Automata/CAVA_Base/Lexord_List.thy b/thys/CAVA_Automata/CAVA_Base/Lexord_List.thy --- a/thys/CAVA_Automata/CAVA_Base/Lexord_List.thy +++ b/thys/CAVA_Automata/CAVA_Base/Lexord_List.thy @@ -1,113 +1,119 @@ theory Lexord_List imports Main begin typedef 'a lexlist = "{xs::'a list. True}" morphisms unlex Lex by auto definition "lexlist \ Lex" lemma lexlist_ext: "Lex xs = Lex ys \ xs = ys" by (auto simp: Lex_inject) lemma Lex_unlex [simp, code abstype]: "Lex (unlex lxs) = lxs" by (fact unlex_inverse) lemma unlex_lexlist [simp, code abstract]: "unlex (lexlist xs) = xs" by (metis lexlist_ext unlex_inverse lexlist_def) definition list_less :: "'a :: ord list \ 'a list \ bool" where "list_less xs ys \ (xs, ys) \ lexord {(u, v). u < v}" definition list_le where "list_le xs ys \ list_less xs ys \ xs = ys" lemma not_less_Nil [simp]: "\ list_less x []" by (simp add: list_less_def) lemma Nil_less_Cons [simp]: "list_less [] (a # x)" by (simp add: list_less_def) -lemma Cons_less_Cons [simp]: "list_less (a # x) (b # y) \ (if a = b then list_less x y else a < b)" +lemma Cons_less_Cons [simp]: "list_less (a # x) (b # y) \ a < b \ a = b \ list_less x y" by (simp add: list_less_def) lemma le_Nil [simp]: "list_le x [] \ x = []" unfolding list_le_def by (cases x) auto lemma Nil_le_Cons [simp]: "list_le [] x" unfolding list_le_def by (cases x) auto -lemma Cons_le_Cons [simp]: "list_le (a # x) (b # y) \ (if a = b then list_le x y else a < b)" +lemma Cons_le_Cons [simp]: "list_le (a # x) (b # y) \ a < b \ a = b \ list_le x y" unfolding list_le_def by auto lemma less_list_code [code]: "list_less xs [] \ False" "list_less [] (x # xs) \ True" - "list_less (x # xs) (y # ys) \ (if x = y then list_less xs ys else x < y)" + "list_less (x # xs) (y # ys) \ x < y \ x = y \ list_less xs ys" by simp_all lemma less_eq_list_code [code]: "list_le (x # xs) [] \ False" "list_le [] xs \ True" - "list_le (x # xs) (y # ys) \ (if x = y then list_le xs ys else x < y)" + "list_le (x # xs) (y # ys) \ x < y \ x = y \ list_le xs ys" by simp_all instantiation lexlist :: (ord) ord begin definition lexlist_less_def: "xs < ys \ list_less (unlex xs) (unlex ys)" definition lexlist_le_def: "(xs :: _ lexlist) \ ys \ list_le (unlex xs) (unlex ys)" instance .. lemmas lexlist_ord_defs = lexlist_le_def lexlist_less_def list_le_def list_less_def end instance lexlist :: (order) order proof fix xs :: "'a lexlist" show "xs \ xs" by (simp add: lexlist_le_def list_le_def) next fix xs ys zs :: "'a lexlist" assume "xs \ ys" and "ys \ zs" then show "xs \ zs" apply (auto simp add: lexlist_ord_defs) apply (rule lexord_trans) apply (auto intro: transI) - using antisym_def order.asym by auto + done next fix xs ys :: "'a lexlist" assume "xs \ ys" and "ys \ xs" then show "xs = ys" apply (auto simp add: lexlist_ord_defs) apply (rule lexord_irreflexive [THEN notE]) defer apply (rule lexord_trans) apply (auto intro: transI simp: unlex_inject) - using antisym_def by fastforce + done next fix xs ys :: "'a lexlist" show "xs < ys \ xs \ ys \ \ ys \ xs" apply (auto simp add: lexlist_ord_defs) - apply (metis (no_types, lifting) antisym_def case_prodD case_prodI less_trans lexord_irreflexive lexord_trans mem_Collect_eq order.asym trans_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) done qed instance lexlist :: (linorder) linorder proof fix xs ys :: "'a lexlist" have "(unlex xs, unlex ys) \ lexord {(u, v). u < v} \ unlex xs = unlex ys \ (unlex ys, unlex xs) \ lexord {(u, v). u < v}" by (rule lexord_linear) auto then show "xs \ ys \ ys \ xs" by (auto simp add: lexlist_ord_defs unlex_inject) qed end diff --git a/thys/CYK/CYK.thy b/thys/CYK/CYK.thy --- a/thys/CYK/CYK.thy +++ b/thys/CYK/CYK.thy @@ -1,1072 +1,1064 @@ (* Title: A formalisation of the Cocke-Younger-Kasami algorithm Author: Maksym Bortin *) theory CYK imports Main begin text \The theory is structured as follows. First section deals with modelling of grammars, derivations, and the language semantics. Then the basic properties are proved. Further, CYK is abstractly specified and its underlying recursive relationship proved. The final section contains a prototypical implementation accompanied by a proof of its correctness.\ section "Basic modelling" subsection "Grammars in Chomsky normal form" text "A grammar in Chomsky normal form is here simply modelled by a list of production rules (the type CNG below), each having a non-terminal symbol on the lhs and either two non-terminals or one terminal symbol on the rhs." datatype ('n, 't) RHS = Branch 'n 'n | Leaf 't type_synonym ('n, 't) CNG = "('n \ ('n, 't) RHS) list" text "Abbreviating the list append symbol for better readability" abbreviation list_append :: "'a list \ 'a list \ 'a list" (infixr "\" 65) where "xs \ ys \ xs @ ys" subsection "Derivation by grammars" text\A \emph{word form} (or sentential form) may be built of both non-terminal and terminal symbols, as opposed to a \emph{word} that contains only terminals. By the usage of disjoint union, non-terminals are injected into a word form by @{term "Inl"} whereas terminals -- by @{term "Inr"}.\ type_synonym ('n, 't) word_form = "('n + 't) list" type_synonym 't word = "'t list" text "A single step derivation relation on word forms is induced by a grammar in the standard way, replacing a non-terminal within a word form in accordance to the production rules." definition DSTEP :: "('n, 't) CNG \ (('n, 't) word_form \ ('n, 't) word_form) set" where "DSTEP G = {(l \ [Inl N] \ r, x) | l N r rhs x. (N, rhs) \ set G \ (case rhs of Branch A B \ x = l \ [Inl A, Inl B] \ r | Leaf t \ x = l \ [Inr t] \ r)}" abbreviation DSTEP' :: "('n, 't) word_form \ ('n, 't) CNG \ ('n, 't) word_form \ bool" ("_ -_\ _" [60, 61, 60] 61) where "w -G\ w' \ (w, w') \ DSTEP G" abbreviation DSTEP_reflc :: "('n, 't) word_form \ ('n, 't) CNG \ ('n, 't) word_form \ bool" ("_ -_\\<^sup>= _" [60, 61, 60] 61) where "w -G\\<^sup>= w' \ (w, w') \ (DSTEP G)\<^sup>=" abbreviation DSTEP_transc :: "('n, 't) word_form \ ('n, 't) CNG \ ('n, 't) word_form \ bool" ("_ -_\\<^sup>+ _" [60, 61, 60] 61) where "w -G\\<^sup>+ w' \ (w, w') \ (DSTEP G)\<^sup>+" abbreviation DSTEP_rtransc :: "('n, 't) word_form \ ('n, 't) CNG \ ('n, 't) word_form \ bool" ("_ -_\\<^sup>* _" [60, 61, 60] 61) where "w -G\\<^sup>* w' \ (w, w') \ (DSTEP G)\<^sup>*" subsection "The generated language semantics" text "The language generated by a grammar from a non-terminal symbol comprises all words that can be derived from the non-terminal in one or more steps. Notice that by the presented grammar modelling, languages containing the empty word cannot be generated. Hence in rare situations when such languages are required, the empty word case should be treated separately." definition Lang :: "('n, 't) CNG \ 'n \ 't word set" where "Lang G S = {w. [Inl S] -G\\<^sup>+ map Inr w }" text\So, for instance, a grammar generating the language $a^nb^n$ from the non-terminal @{term "''S''"} might look as follows.\ definition "G_anbn = [(''S'', Branch ''A'' ''T''), (''S'', Branch ''A'' ''B''), (''T'', Branch ''S'' ''B''), (''A'', Leaf ''a''), (''B'', Leaf ''b'')]" text\Now the term @{term "Lang G_anbn ''S''"} denotes the set of words of the form $a^nb^n$ with $n > 0$. This is intuitively clear, but not straight forward to show, and a lengthy proof for that is out of scope.\ section "Basic properties" lemma prod_into_DSTEP1 : "(S, Branch A B) \ set G \ L \ [Inl S] \ R -G\ L \ [Inl A, Inl B] \ R" by(simp add: DSTEP_def, rule_tac x="L" in exI, force) lemma prod_into_DSTEP2 : "(S, Leaf a) \ set G \ L \ [Inl S] \ R -G\ L \ [Inr a] \ R" by(simp add: DSTEP_def, rule_tac x="L" in exI, force) lemma DSTEP_D : "s -G\ t \ \L N R rhs. s = L \ [Inl N] \ R \ (N, rhs) \ set G \ (\A B. rhs = Branch A B \ t = L \ [Inl A, Inl B] \ R) \ (\x. rhs = Leaf x \ t = L \ [Inr x] \ R)" by(unfold DSTEP_def, clarsimp, simp split: RHS.split_asm, blast+) lemma DSTEP_append : assumes a: "s -G\ t" shows "L \ s \ R -G\ L \ t \ R" proof - from a have "\l N r rhs. s = l \ [Inl N] \ r \ (N, rhs) \ set G \ (\A B. rhs = Branch A B \ t = l \ [Inl A, Inl B] \ r) \ (\x. rhs = Leaf x \ t = l \ [Inr x] \ r)" (is "\l N r rhs. ?P l N r rhs") by(rule DSTEP_D) then obtain l N r rhs where "?P l N r rhs" by blast thus ?thesis by(simp add: DSTEP_def, rule_tac x="L \ l" in exI, rule_tac x=N in exI, rule_tac x="r \ R" in exI, simp, rule_tac x=rhs in exI, simp split: RHS.split) qed lemma DSTEP_star_mono : "s -G\\<^sup>* t \ length s \ length t" proof(erule rtrancl_induct, simp) fix t u assume "s -G\\<^sup>* t" assume a: "t -G\ u" assume b: "length s \ length t" show "length s \ length u" proof - from a have "\L N R rhs. t = L \ [Inl N] \ R \ (N, rhs) \ set G \ (\A B. rhs = Branch A B \ u = L \ [Inl A, Inl B] \ R) \ (\x. rhs = Leaf x \ u = L \ [Inr x] \ R)" (is "\L N R rhs. ?P L N R rhs") by(rule DSTEP_D) then obtain L N R rhs where "?P L N R rhs" by blast with b show ?thesis by(case_tac rhs, clarsimp+) qed qed lemma DSTEP_comp : assumes a: "l \ r -G\ t" shows "\l' r'. l -G\\<^sup>= l' \ r -G\\<^sup>= r' \ t = l' \ r'" proof - from a have "\L N R rhs. l \ r = L \ [Inl N] \ R \ (N, rhs) \ set G \ (\A B. rhs = Branch A B \ t = L \ [Inl A, Inl B] \ R) \ (\x. rhs = Leaf x \ t = L \ [Inr x] \ R)" (is "\L N R rhs. ?T L N R rhs") by(rule DSTEP_D) then obtain L N R rhs where b: "?T L N R rhs" by blast hence "l \ r = L \ Inl N # R" by simp hence "\u. (l = L \ u \ u \ r = Inl N # R) \ (l \ u = L \ r = u \ Inl N # R)" by(rule append_eq_append_conv2[THEN iffD1]) then obtain xs where c: "l = L \ xs \ xs \ r = Inl N # R \ l \ xs = L \ r = xs \ Inl N # R" (is "?C1 \ ?C2") by blast show ?thesis proof(cases rhs) case (Leaf x) with b have d: "t = L \ [Inr x] \ R \ (N, Leaf x) \ set G" by simp from c show ?thesis proof assume e: "?C1" show ?thesis proof(cases xs) case Nil with d and e show ?thesis by(clarsimp, rule_tac x=L in exI, simp add: DSTEP_def, simp split: RHS.split, blast) next case (Cons z zs) with d and e show ?thesis by(rule_tac x="L \ Inr x # zs" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast) qed next assume e: "?C2" show ?thesis proof(cases xs) case Nil with d and e show ?thesis by(rule_tac x=L in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast) next case (Cons z zs) with d and e show ?thesis by(rule_tac x="l" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, rule_tac x="z#zs" in exI, rule_tac x=N in exI, rule_tac x=R in exI, simp, rule_tac x="Leaf x" in exI, simp) qed qed next case (Branch A B) with b have d: "t = L \ [Inl A, Inl B] \ R \ (N, Branch A B) \ set G" by simp from c show ?thesis proof assume e: "?C1" show ?thesis proof(cases xs) case Nil with d and e show ?thesis by(clarsimp, rule_tac x=L in exI, simp add: DSTEP_def, simp split: RHS.split, blast) next case (Cons z zs) with d and e show ?thesis by(rule_tac x="L \ [Inl A, Inl B] \ zs" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast) qed next assume e: "?C2" show ?thesis proof(cases xs) case Nil with d and e show ?thesis by(rule_tac x=L in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast) next case (Cons z zs) with d and e show ?thesis by(rule_tac x="l" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, rule_tac x="z#zs" in exI, rule_tac x=N in exI, rule_tac x=R in exI, simp, rule_tac x="Branch A B" in exI, simp) qed qed qed qed theorem DSTEP_star_comp1 : assumes A: "l \ r -G\\<^sup>* t" shows "\l' r'. l -G\\<^sup>* l' \ r -G\\<^sup>* r' \ t = l' \ r'" proof - have "\s. s -G\\<^sup>* t \ \l r. s = l \ r \ (\l' r'. l -G\\<^sup>* l' \ r -G\\<^sup>* r' \ t = l' \ r')" (is "\s. ?P s t \ ?Q s t") proof(erule rtrancl_induct, force) fix s t u assume "?P s t" assume a: "t -G\ u" assume b: "?Q s t" show "?Q s u" proof(clarify) fix l r assume "s = l \ r" with b have "\l' r'. l -G\\<^sup>* l' \ r -G\\<^sup>* r' \ t = l' \ r'" by simp then obtain l' r' where c: "l -G\\<^sup>* l' \ r -G\\<^sup>* r' \ t = l' \ r'" by blast with a have "l' \ r' -G\ u" by simp hence "\l'' r''. l' -G\\<^sup>= l'' \ r' -G\\<^sup>= r'' \ u = l'' \ r''" by(rule DSTEP_comp) then obtain l'' r'' where "l' -G\\<^sup>= l'' \ r' -G\\<^sup>= r'' \ u = l'' \ r''" by blast hence "l' -G\\<^sup>* l'' \ r' -G\\<^sup>* r'' \ u = l'' \ r''" by blast with c show "\l' r'. l -G\\<^sup>* l' \ r -G\\<^sup>* r' \ u = l' \ r'" by(rule_tac x=l'' in exI, rule_tac x=r'' in exI, force) qed qed with A show ?thesis by force qed theorem DSTEP_star_comp2 : assumes A: "l -G\\<^sup>* l'" and B: "r -G\\<^sup>* r'" shows "l \ r -G\\<^sup>* l' \ r'" proof - have "l -G\\<^sup>* l' \ \r r'. r -G\\<^sup>* r' \ l \ r -G\\<^sup>* l' \ r'" (is "?P l l' \ ?Q l l'") proof(erule rtrancl_induct) show "?Q l l" proof(clarify, erule rtrancl_induct, simp) fix r s t assume a: "s -G\ t" assume b: "l \ r -G\\<^sup>* l \ s" show "l \ r -G\\<^sup>* l \ t" proof - from a have "l \ s -G\ l \ t" by(drule_tac L=l and R="[]" in DSTEP_append, simp) with b show ?thesis by simp qed qed next fix s t assume a: "s -G\ t" assume b: "?Q l s" show "?Q l t" proof(clarsimp) fix r r' assume "r -G\\<^sup>* r'" with b have c: "l \ r -G\\<^sup>* s \ r'" by simp from a have "s \ r' -G\ t \ r'" by(drule_tac L="[]" and R=r' in DSTEP_append, simp) with c show "l \ r -G\\<^sup>* t \ r'" by simp qed qed with A and B show ?thesis by simp qed lemma DSTEP_trancl_term : assumes A: "[Inl S] -G\\<^sup>+ t" and B: "Inr x \ set t" shows "\N. (N, Leaf x) \ set G" proof - have "[Inl S] -G\\<^sup>+ t \ \x. Inr x \ set t \ (\N. (N, Leaf x) \ set G)" (is "?P t \ ?Q t") proof(erule trancl_induct) fix t assume a: "[Inl S] -G\ t" show "?Q t" proof - from a have "\rhs. (S, rhs) \ set G \ (\A B. rhs = Branch A B \ t = [Inl A, Inl B]) \ (\x. rhs = Leaf x \ t = [Inr x])" (is "\rhs. ?P rhs") by(simp add: DSTEP_def, clarsimp, simp split: RHS.split_asm, case_tac l, force, simp, clarsimp, simp split: RHS.split_asm, case_tac l, force, simp) then obtain rhs where "?P rhs" by blast thus ?thesis by(case_tac rhs, clarsimp, force) qed next fix s t assume a: "s -G\ t" assume b: "?Q s" show "?Q t" proof - from a have "\L N R rhs. s = L \ [Inl N] \ R \ (N, rhs) \ set G \ (\A B. rhs = Branch A B \ t = L \ [Inl A, Inl B] \ R) \ (\x. rhs = Leaf x \ t = L \ [Inr x] \ R)" (is "\L N R rhs. ?P L N R rhs") by(rule DSTEP_D) then obtain L N R rhs where "?P L N R rhs" by blast with b show ?thesis by(case_tac rhs, clarsimp, force) qed qed with A and B show ?thesis by simp qed subsection "Properties of generated languages" lemma Lang_no_Nil : "w \ Lang G S \ w \ []" by(simp add: Lang_def, drule trancl_into_rtrancl, drule DSTEP_star_mono, force) lemma Lang_rtrancl_eq : "(w \ Lang G S) = [Inl S] -G\\<^sup>* map Inr w" (is "?L = (?p \ ?R\<^sup>*)") proof(simp add: Lang_def, rule iffI, erule trancl_into_rtrancl) assume "?p \ ?R\<^sup>*" hence "?p \ (?R\<^sup>+)\<^sup>=" by(subst rtrancl_trancl_reflcl[THEN sym], assumption) hence "[Inl S] = map Inr w \ ?p \ ?R\<^sup>+" by force thus "?p \ ?R\<^sup>+" by(case_tac w, simp_all) qed lemma Lang_term : "w \ Lang G S \ \x \ set w. \N. (N, Leaf x) \ set G" by(clarsimp simp add: Lang_def, drule DSTEP_trancl_term, simp, erule imageI, assumption) lemma Lang_eq1 : "([x] \ Lang G S) = ((S, Leaf x) \ set G)" proof(simp add: Lang_def, rule iffI, subst (asm) trancl_unfold_left, clarsimp) fix t assume a: "[Inl S] -G\ t" assume b: "t -G\\<^sup>* [Inr x]" note DSTEP_star_mono[OF b, simplified] hence c: "length t \ 1" by simp have "\z. t = [z]" proof(cases t) assume "t = []" with b have d: "[] -G\\<^sup>* [Inr x]" by simp have "\s. ([], s) \ (DSTEP G)\<^sup>* \ s = []" by(erule rtrancl_induct, simp_all, drule DSTEP_D, clarsimp) note this[OF d] thus ?thesis by simp next fix z zs assume "t = z#zs" with c show ?thesis by force qed with a have "\z. (S, Leaf z) \ set G \ t = [Inr z]" by(clarsimp simp add: DSTEP_def, simp split: RHS.split_asm, case_tac l, simp_all) with b show "(S, Leaf x) \ set G" proof(clarsimp) fix z assume c: "(S, Leaf z) \ set G" assume "[Inr z] -G\\<^sup>* [Inr x]" hence "([Inr z], [Inr x]) \ ((DSTEP G)\<^sup>+)\<^sup>=" by simp hence "[Inr z] = [Inr x] \ [Inr z] -G\\<^sup>+ [Inr x]" by force hence "x = z" proof assume "[Inr z] = [Inr x]" thus ?thesis by simp next assume "[Inr z] -G\\<^sup>+ [Inr x]" hence "\u. [Inr z] -G\ u \ u -G\\<^sup>* [Inr x]" by(subst (asm) trancl_unfold_left, force) then obtain u where "[Inr z] -G\ u" by blast thus ?thesis by(clarsimp simp add: DSTEP_def, case_tac l, simp_all) qed with c show ?thesis by simp qed next assume a: "(S, Leaf x) \ set G" show "[Inl S] -G\\<^sup>+ [Inr x]" by(rule r_into_trancl, simp add: DSTEP_def, rule_tac x="[]" in exI, rule_tac x="S" in exI, rule_tac x="[]" in exI, simp, rule_tac x="Leaf x" in exI, simp add: a) qed theorem Lang_eq2 : "(w \ Lang G S \ 1 < length w) = (\A B. (S, Branch A B) \ set G \ (\l r. w = l \ r \ l \ Lang G A \ r \ Lang G B))" (is "?L = ?R") proof(rule iffI, clarify, subst (asm) Lang_def, simp, subst (asm) trancl_unfold_left, clarsimp) have map_Inr_split : "\xs. \zs w. map Inr w = xs \ zs \ (\u v. w = u \ v \ xs = map Inr u \ zs = map Inr v)" by(induct_tac xs, simp, force) fix t assume a: "Suc 0 < length w" assume b: "[Inl S] -G\ t" assume c: "t -G\\<^sup>* map Inr w" from b have "\A B. (S, Branch A B) \ set G \ t = [Inl A, Inl B]" proof(simp add: DSTEP_def, clarify, case_tac l, simp_all, simp split: RHS.split_asm, clarify) fix x assume "t = [Inr x]" with c have d: "[Inr x] -G\\<^sup>* map Inr w"by simp have "\x s. [Inr x] -G\\<^sup>* s \ s = [Inr x]" by(erule rtrancl_induct, simp_all, drule DSTEP_D, clarsimp, case_tac L, simp_all) note this[OF d] hence "w = [x]" by(case_tac w, simp_all) with a show "False" by simp qed then obtain A B where d: "(S, Branch A B) \ set G \ t = [Inl A, Inl B]" by blast with c have e: "[Inl A] \ [Inl B] -G\\<^sup>* map Inr w" by simp note DSTEP_star_comp1[OF e] then obtain l' r' where e: "[Inl A] -G\\<^sup>* l' \ [Inl B] -G\\<^sup>* r' \ map Inr w = l' \ r'" by blast note map_Inr_split[rule_format, OF e[THEN conjunct2, THEN conjunct2]] then obtain u v where f: "w = u \ v \ l' = map Inr u \ r' = map Inr v" by blast with e have g: "[Inl A] -G\\<^sup>* map Inr u \ [Inl B] -G\\<^sup>* map Inr v" by simp show "?R" by(rule_tac x=A in exI, rule_tac x=B in exI, simp add: d, rule_tac x=u in exI, rule_tac x=v in exI, simp add: f, (subst Lang_rtrancl_eq)+, rule g) next assume "?R" then obtain A B l r where a: "(S, Branch A B) \ set G \ w = l \ r \ l \ Lang G A \ r \ Lang G B" by blast have "[Inl A] \ [Inl B] -G\\<^sup>* map Inr l \ map Inr r" by(rule DSTEP_star_comp2, subst Lang_rtrancl_eq[THEN sym], simp add: a, subst Lang_rtrancl_eq[THEN sym], simp add: a) hence b: "[Inl A] \ [Inl B] -G\\<^sup>* map Inr w" by(simp add: a) have c: "w \ Lang G S" by(simp add: Lang_def, subst trancl_unfold_left, rule_tac b="[Inl A] \ [Inl B]" in relcompI, simp add: DSTEP_def, rule_tac x="[]" in exI, rule_tac x="S" in exI, rule_tac x="[]" in exI, simp, rule_tac x="Branch A B" in exI, simp add: a[THEN conjunct1], rule b) thus "?L" proof show "1 < length w" proof(simp add: a, rule ccontr, drule leI) assume "length l + length r \ Suc 0" hence "l = [] \ r = []" by(case_tac l, simp_all) thus "False" proof assume "l = []" with a have "[] \ Lang G A" by simp note Lang_no_Nil[OF this] thus ?thesis by simp next assume "r = []" with a have "[] \ Lang G B" by simp note Lang_no_Nil[OF this] thus ?thesis by simp qed qed qed qed section "Abstract specification of CYK" text "A subword of a word $w$, starting at the position $i$ (first element is at the position $0$) and having the length $j$, is defined as follows." definition "subword w i j = take j (drop i w)" text "Thus, to any subword of the given word $w$ CYK assigns all non-terminals from which this subword is derivable by the grammar $G$." definition "CYK G w i j = {S. subword w i j \ Lang G S}" subsection \Properties of @{term "subword"}\ lemma subword_length : "i + j \ length w \ length(subword w i j) = j" by(simp add: subword_def) lemma subword_nth1 : "i + j \ length w \ k < j \ (subword w i j)!k = w!(i + k)" by(simp add: subword_def) lemma subword_nth2 : assumes A: "i + 1 \ length w" shows "subword w i 1 = [w!i]" proof - note subword_length[OF A] hence "\x. subword w i 1 = [x]" by(case_tac "subword w i 1", simp_all) then obtain x where a:"subword w i 1 = [x]" by blast note subword_nth1[OF A, where k="(0 :: nat)", simplified] with a have "x = w!i" by simp with a show ?thesis by simp qed lemma subword_self : "subword w 0 (length w) = w" by(simp add: subword_def) lemma take_split[rule_format] : "\n m. n \ length xs \ n \ m \ take n xs \ take (m - n) (drop n xs) = take m xs" by(induct_tac xs, clarsimp+, case_tac n, simp_all, case_tac m, simp_all) lemma subword_split : "i + j \ length w \ 0 < k \ k < j \ subword w i j = subword w i k \ subword w (i + k) (j - k)" by(simp add: subword_def, subst take_split[where n=k, THEN sym], simp_all, rule_tac f="\x. take (j - k) (drop x w)" in arg_cong, simp) lemma subword_split2 : assumes A: "subword w i j = l \ r" and B: "i + j \ length w" and C: "0 < length l" and D: "0 < length r" shows "l = subword w i (length l) \ r = subword w (i + length l) (j - length l)" proof - have a: "length(subword w i j) = j" by(rule subword_length, rule B) note arg_cong[where f=length, OF A] with a and D have b: "length l < j" by force with B have c: "i + length l \ length w" by force have "subword w i j = subword w i (length l) \ subword w (i + length l) (j - length l)" by(rule subword_split, rule B, rule C, rule b) with A have d: "l \ r = subword w i (length l) \ subword w (i + length l) (j - length l)" by simp show ?thesis by(rule append_eq_append_conv[THEN iffD1], subst subword_length, rule c, simp, rule d) qed subsection \Properties of @{term "CYK"}\ lemma CYK_Lang : "(S \ CYK G w 0 (length w)) = (w \ Lang G S)" by(simp add: CYK_def subword_self) lemma CYK_eq1 : "i + 1 \ length w \ CYK G w i 1 = {S. (S, Leaf (w!i)) \ set G}" by(simp add: CYK_def, subst subword_nth2[simplified], assumption, subst Lang_eq1, rule refl) theorem CYK_eq2 : assumes A: "i + j \ length w" and B: "1 < j" shows "CYK G w i j = {X | X A B k. (X, Branch A B) \ set G \ A \ CYK G w i k \ B \ CYK G w (i + k) (j - k) \ 1 \ k \ k < j}" proof(rule set_eqI, rule iffI, simp_all add: CYK_def) fix X assume a: "subword w i j \ Lang G X" show "\A B. (X, Branch A B) \ set G \ (\k. subword w i k \ Lang G A \ subword w (i + k) (j - k) \ Lang G B \ Suc 0 \ k \ k < j)" proof - have b: "1 < length(subword w i j)" by(subst subword_length, rule A, rule B) note Lang_eq2[THEN iffD1, OF conjI, OF a b] then obtain A B l r where c: "(X, Branch A B) \ set G \ subword w i j = l \ r \ l \ Lang G A \ r \ Lang G B" by blast note Lang_no_Nil[OF c[THEN conjunct2, THEN conjunct2, THEN conjunct1]] hence d: "0 < length l" by(case_tac l, simp_all) note Lang_no_Nil[OF c[THEN conjunct2, THEN conjunct2, THEN conjunct2]] hence e: "0 < length r" by(case_tac r, simp_all) note subword_split2[OF c[THEN conjunct2, THEN conjunct1], OF A, OF d, OF e] with c show ?thesis proof(rule_tac x=A in exI, rule_tac x=B in exI, simp, rule_tac x="length l" in exI, simp) show "Suc 0 \ length l \ length l < j" (is "?A \ ?B") proof from d show "?A" by(case_tac l, simp_all) next note arg_cong[where f=length, OF c[THEN conjunct2, THEN conjunct1], THEN sym] also have "length(subword w i j) = j" by(rule subword_length, rule A) finally have "length l + length r = j" by simp with e show ?B by force qed qed qed next fix X assume "\A B. (X, Branch A B) \ set G \ (\k. subword w i k \ Lang G A \ subword w (i + k) (j - k) \ Lang G B \ Suc 0 \ k \ k < j)" then obtain A B k where a: "(X, Branch A B) \ set G \ subword w i k \ Lang G A \ subword w (i + k) (j - k) \ Lang G B \ Suc 0 \ k \ k < j" by blast show "subword w i j \ Lang G X" proof(rule Lang_eq2[THEN iffD2, THEN conjunct1], rule_tac x=A in exI, rule_tac x=B in exI, simp add: a, rule_tac x="subword w i k" in exI, rule_tac x="subword w (i + k) (j - k)" in exI, simp add: a, rule subword_split, rule A) from a show "0 < k" by force next from a show "k < j" by simp qed qed section "Implementation" text "One of the particularly interesting features of CYK implementation is that it follows the principles of dynamic programming, constructing a table of solutions for sub-problems in the bottom-up style reusing already stored results." subsection "Main cycle" text "This is an auxiliary implementation of the membership test on lists." fun mem :: "'a \ 'a list \ bool" where "mem a [] = False" | "mem a (x#xs) = (x = a \ mem a xs)" lemma mem[simp] : "mem x xs = (x \ set xs)" by(induct_tac xs, simp, force) text "The purpose of the following is to collect non-terminals that appear on the lhs of a production such that the first non-terminal on its rhs appears in the first of two given lists and the second non-terminal -- in the second list." fun match_prods :: "('n, 't) CNG \ 'n list \ 'n list \ 'n list" where "match_prods [] ls rs = []" | "match_prods ((X, Branch A B)#ps) ls rs = (if mem A ls \ mem B rs then X # match_prods ps ls rs else match_prods ps ls rs)" | "match_prods ((X, Leaf a)#ps) ls rs = match_prods ps ls rs" lemma match_prods : "(X \ set(match_prods G ls rs)) = (\A \ set ls. \B \ set rs. (X, Branch A B) \ set G)" by(induct_tac G, clarsimp+, rename_tac l r ps, case_tac r, force+) text "The following function is the inner cycle of the algorithm. The parameters $i$ and $j$ identify a subword starting at $i$ with the length $j$, whereas $k$ is used to iterate through its splits (which are of course subwords as well) all having the length greater $0$ but less than $j$. The parameter $T$ represents a table containing CYK solutions for those splits." function inner :: "('n, 't) CNG \ (nat \ nat \ 'n list) \ nat \ nat \ nat \ 'n list" where "inner G T i k j = (if k < j then match_prods G (T(i, k)) (T(i + k, j - k)) @ inner G T i (k + 1) j else [])" by pat_completeness auto termination by(relation "measure(\(a, b, c, d, e). e - d)", rule wf_measure, simp) declare inner.simps[simp del] lemma inner : "(X \ set(inner G T i k j)) = (\l. k \ l \ l < j \ X \ set(match_prods G (T(i, l)) (T(i + l, j - l))))" (is "?L G T i k j = ?R G T i k j") proof(induct_tac G T i k j rule: inner.induct) fix G T i k j assume a: "k < j \ ?L G T i (k + 1) j = ?R G T i (k + 1) j" show "?L G T i k j = ?R G T i k j" proof(case_tac "k < j") assume b: "k < j" with a have c: "?L G T i (k + 1) j = ?R G T i (k + 1) j" by simp show ?thesis proof(subst inner.simps, simp add: b, rule iffI, erule disjE, rule_tac x=k in exI, simp add: b) assume "X \ set(inner G T i (Suc k) j)" with c have "?R G T i (k + 1) j" by simp thus "?R G T i k j" by(clarsimp, rule_tac x=l in exI, simp) next assume "?R G T i k j" then obtain l where d: "k \ l \ l < j \ X \ set(match_prods G (T(i, l)) (T(i + l, j - l)))" by blast show "X \ set(match_prods G (T(i, k)) (T(i + k, j - k))) \ ?L G T i (Suc k) j" proof(case_tac "Suc k \ l", rule disjI2, subst c[simplified], rule_tac x=l in exI, simp add: d, rule disjI1) assume "\ Suc k \ l" with d have "l = k" by force with d show "X \ set(match_prods G (T(i, k)) (T(i + k, j - k)))" by simp qed qed next assume "\ k < j" thus ?thesis by(subst inner.simps, simp) qed qed - -lemma in_lex_prod[simp]: "NO_MATCH less_than r \ ((a, b), (a', b')) \ r <*lex*> s \ a \ a' \ (a, a') \ r \ a = a' \ (b, b') \ s" - by (auto simp:lex_prod_def) - -text\compared with @{thm[source]in_lex_prod} this yields simpler results\ -lemma in_lex_prod_less_than[simp]: "((a, b), (a', b')) \ less_than <*lex*> s \a a = a' \ (b, b') \ s" - by auto - - + text\Now the main part of the algorithm just iterates through all subwords up to the given length $len$, calls @{term "inner"} on these, and stores the results in the table $T$. The length $j$ is supposed to be greater than $1$ -- the subwords of length $1$ will be handled in the initialisation phase below.\ function main :: "('n, 't) CNG \ (nat \ nat \ 'n list) \ nat \ nat \ nat \ (nat \ nat \ 'n list)" where "main G T len i j = (let T' = T((i, j) := inner G T i 1 j) in if i + j < len then main G T' len (i + 1) j else if j < len then main G T' len 0 (j + 1) else T')" by pat_completeness auto termination by(relation "inv_image (less_than <*lex*> less_than) (\(a, b, c, d, e). (c - e, c - (d + e)))", rule wf_inv_image, rule wf_lex_prod, (rule wf_less_than)+, simp_all) declare main.simps[simp del] lemma main : assumes "1 < j" and "i + j \ length w" and "\i' j'. j' < j \ 1 \ j' \ i' + j' \ length w \ set(T(i', j')) = CYK G w i' j'" and "\i'. i' < i \ i' + j \ length w \ set(T(i', j)) = CYK G w i' j" and "1 \ j'" and "i' + j' \ length w" shows "set((main G T (length w) i j)(i', j')) = CYK G w i' j'" proof - have "\len T' w. main G T len i j = T' \ length w = len \ 1 < j \ i + j \ len \ (\j' < j. \i'. 1 \ j' \ i' + j' \ len \ set(T(i', j')) = CYK G w i' j') \ (\i' < i. i' + j \ len \ set(T(i', j)) = CYK G w i' j) \ (\j' \ 1. \i'. i' + j' \ len \ set(T'(i', j')) = CYK G w i' j')" (is "\len. ?P G T len i j") proof(rule allI, induct_tac G T len i j rule: main.induct, (drule meta_spec, drule meta_mp, rule refl)+, clarify) fix G T i j i' j' fix w :: "'a list" assume a: "i + j < length w \ ?P G (T((i, j) := inner G T i 1 j)) (length w) (i + 1) j" assume b: "\ i + j < length w \ j < length w \ ?P G (T((i, j) := inner G T i 1 j)) (length w) 0 (j + 1)" assume c: "1 < j" assume d: "i + j \ length w" assume e: "(1::nat) \ j'" assume f: "i' + j' \ length w" assume g: "\j' < j. \i'. 1 \ j' \ i' + j' \ length w \ set(T(i', j')) = CYK G w i' j'" assume h: "\i' < i. i' + j \ length w \ set(T(i', j)) = CYK G w i' j" have inner: "set (inner G T i (Suc 0) j) = CYK G w i j" proof(rule set_eqI, subst inner, subst match_prods, subst CYK_eq2, rule d, rule c, simp) fix X show "(\l\Suc 0. l < j \ (\A \ set(T(i, l)). \B \ set(T(i + l, j - l)). (X, Branch A B) \ set G)) = (\A B. (X, Branch A B) \ set G \ (\k. A \ CYK G w i k \ B \ CYK G w (i + k) (j - k) \ Suc 0 \ k \ k < j))" (is "?L = ?R") proof assume "?L" thus "?R" proof(clarsimp, rule_tac x=A in exI, rule_tac x=B in exI, simp, rule_tac x=l in exI, simp) fix l A B assume i: "Suc 0 \ l" assume j: "l < j" assume k: "A \ set(T(i, l))" assume l: "B \ set(T(i + l, j - l))" note g[rule_format, where i'=i and j'=l] with d i j have A: "set(T(i, l)) = CYK G w i l" by force note g[rule_format, where i'="i + l" and j'="j - l"] with d i j have "set(T(i + l, j - l)) = CYK G w (i + l) (j - l)" by force with k l A show "A \ CYK G w i l \ B \ CYK G w (i + l) (j - l)" by simp qed next assume "?R" thus "?L" proof(clarsimp, rule_tac x=k in exI, simp) fix A B k assume i: "Suc 0 \ k" assume j: "k < j" assume k: "A \ CYK G w i k" assume l: "B \ CYK G w (i + k) (j - k)" assume m: "(X, Branch A B) \ set G" note g[rule_format, where i'=i and j'=k] with d i j have A: "CYK G w i k = set(T(i, k))" by force note g[rule_format, where i'="i + k" and j'="j - k"] with d i j have "CYK G w (i + k) (j - k) = set(T(i + k, j - k))" by force with k l A have "A \ set(T(i, k)) \ B \ set(T(i + k, j - k))" by simp with m show "\A \ set(T(i, k)). \B \ set(T(i + k, j - k)). (X, Branch A B) \ set G" by force qed qed qed (* inner *) show "set((main G T (length w) i j)(i', j')) = CYK G w i' j'" proof(case_tac "i + j = length w") assume i: "i + j = length w" show ?thesis proof(case_tac "j < length w") assume j: "j < length w" show ?thesis proof(subst main.simps, simp add: Let_def i j, rule b[rule_format, where w=w and i'=i' and j'=j', OF _ _ refl, simplified], simp_all add: inner) from i show "\ i + j < length w" by simp next from c show "0 < j" by simp next from j show "Suc j \ length w" by simp next from e show "Suc 0 \ j'" by simp next from f show "i' + j' \ length w" by assumption next fix i'' j'' assume k: "j'' < Suc j" assume l: "Suc 0 \ j''" assume m: "i'' + j'' \ length w" show "(i'' = i \ j'' \ j) \ set(T(i'',j'')) = CYK G w i'' j''" proof(case_tac "j'' = j", simp_all, clarify) assume n: "j'' = j" assume "i'' \ i" with i m n have "i'' < i" by simp with n m h show "set(T(i'', j)) = CYK G w i'' j" by simp next assume "j'' \ j" with k have "j'' < j" by simp with l m g show "set(T(i'', j'')) = CYK G w i'' j''" by simp qed qed next assume "\ j < length w" with i have j: "i = 0 \ j = length w" by simp show ?thesis proof(subst main.simps, simp add: Let_def j, intro conjI, clarify) from j and inner show "set (inner G T 0 (Suc 0) (length w)) = CYK G w 0 (length w)" by simp next show "0 < i' \ set(T(i', j')) = CYK G w i' j'" proof assume "0 < i'" with j and f have "j' < j" by simp with e g f show "set(T(i', j')) = CYK G w i' j'" by simp qed next show "j' \ length w \ set(T(i', j')) = CYK G w i' j'" proof assume "j' \ length w " with j and f have "j' < j" by simp with e g f show "set(T(i', j')) = CYK G w i' j'" by simp qed qed qed next assume "i + j \ length w" with d have i: "i + j < length w" by simp show ?thesis proof(subst main.simps, simp add: Let_def i, rule a[rule_format, where w=w and i'=i' and j'=j', OF i, OF refl, simplified]) from c show "Suc 0 < j" by simp next from i show "Suc(i + j) \ length w" by simp next from e show "Suc 0 \ j'" by simp next from f show "i' + j' \ length w" by assumption next fix i'' j'' assume "j'' < j" and "Suc 0 \ j''" and "i'' + j'' \ length w" with g show "set(T(i'', j'')) = CYK G w i'' j''" by simp next fix i'' assume j: "i'' < Suc i" show "set(if i'' = i then inner G T i (Suc 0) j else T(i'', j)) = CYK G w i'' j" proof(simp split: if_split, rule conjI, clarify, rule inner, clarify) assume "i'' \ i" with j have "i'' < i" by simp with d h show "set(T(i'', j)) = CYK G w i'' j" by simp qed qed qed qed with assms show ?thesis by force qed subsection "Initialisation phase" text\Similarly to @{term "match_prods"} above, here we collect non-terminals from which the given terminal symbol can be derived.\ fun init_match :: "('n, 't) CNG \ 't \ 'n list" where "init_match [] t = []" | "init_match ((X, Branch A B)#ps) t = init_match ps t" | "init_match ((X, Leaf a)#ps) t = (if a = t then X # init_match ps t else init_match ps t)" lemma init_match : "(X \ set(init_match G a)) = ((X, Leaf a) \ set G)" by(induct_tac G a rule: init_match.induct, simp_all) text "Representing the empty table." definition "emptyT = (\(i, j). [])" text "The following function initialises the empty table for subwords of length $1$, i.e. each symbol occurring in the given word." fun init' :: "('n, 't) CNG \ 't list \ nat \ nat \ nat \ 'n list" where "init' G [] k = emptyT" | "init' G (t#ts) k = (init' G ts (k + 1))((k, 1) := init_match G t)" lemma init' : assumes "i + 1 \ length w" shows "set(init' G w 0 (i, 1)) = CYK G w i 1" proof - have "\i. Suc i \ length w \ (\k. set(init' G w k (k + i, Suc 0)) = CYK G w i (Suc 0))" (is "\i. ?P i w \ (\k. ?Q i k w)") proof(induct_tac w, clarsimp+, rule conjI, clarsimp, rule set_eqI, subst init_match) fix x w S show "((S, Leaf x) \ set G) = (S \ CYK G (x#w) 0 (Suc 0))" by(subst CYK_eq1[simplified], simp_all) next fix x w i assume a: "\i. ?P i w \ (\k. ?Q i k w)" assume b: "i \ length w" show "0 < i \ (\k. set(init' G w (Suc k) (k + i, Suc 0)) = CYK G (x#w) i (Suc 0))" proof(clarify, case_tac i, simp_all, subst CYK_eq1[simplified], simp, erule subst, rule b, simp) fix k j assume c: "i = Suc j" note a[rule_format, where i=j and k="Suc k"] with b and c have "set(init' G w (Suc k) (Suc k + j, Suc 0)) = CYK G w j (Suc 0)" by simp also with b and c have "... = {S. (S, Leaf (w ! j)) \ set G}" by(subst CYK_eq1[simplified], simp_all) finally show "set(init' G w (Suc k) (Suc (k + j), Suc 0)) = {S. (S, Leaf (w ! j)) \ set G}" by simp qed qed with assms have "\k. ?Q i k w" by simp note this[rule_format, where k=0] thus ?thesis by simp qed text\The next version of initialization refines @{term "init'"} in that it takes additional account of the cases when the given word is empty or contains a terminal symbol that does not have any matching production (that is, @{term "init_match"} is an empty list). No initial table is then needed as such words can immediately be rejected.\ fun init :: "('n, 't) CNG \ 't list \ nat \ (nat \ nat \ 'n list) option" where "init G [] k = None" | "init G [t] k = (case (init_match G t) of [] \ None | xs \ Some(emptyT((k, 1) := xs)))" | "init G (t#ts) k = (case (init_match G t) of [] \ None | xs \ (case (init G ts (k + 1)) of None \ None | Some T \ Some(T((k, 1) := xs))))" lemma init1: \init' G w k = T\ if \init G w k = Some T\ using that by (induction G w k arbitrary: T rule: init.induct) (simp_all split: list.splits option.splits) lemma init2 : "(init G w k = None) = (w = [] \ (\a \ set w. init_match G a = []))" by(induct_tac G w k rule: init.induct, simp, simp split: list.split, simp split: list.split option.split, force) subsection \The overall procedure\ definition "cyk G S w = (case init G w 0 of None \ False | Some T \ let len = length w in if len = 1 then mem S (T(0, 1)) else let T' = main G T len 0 2 in mem S (T'(0, len)))" theorem cyk : "cyk G S w = (w \ Lang G S)" proof(simp add: cyk_def split: option.split, simp_all add: Let_def, rule conjI, subst init2, simp, rule conjI) show "w = [] \ [] \ Lang G S" by(clarify, drule Lang_no_Nil, clarify) next show "(\x\set w. init_match G x = []) \ w \ Lang G S" by(clarify, drule Lang_term, subst (asm) init_match[THEN sym], force) next show "\T. init G w 0 = Some T \ ((length w = Suc 0 \ S \ set(T(0, Suc 0))) \ (length w \ Suc 0 \ S \ set(main G T (length w) 0 2 (0, length w)))) = (w \ Lang G S)" (is "\T. ?P T \ ?L T = ?R") proof clarify fix T assume a: "?P T" hence b: "init' G w 0 = T" by(rule init1) note init2[THEN iffD2, OF disjI1] have c: "w \ []" by(clarify, drule init2[where G=G and k=0, THEN iffD2, OF disjI1], simp add: a) have "?L (init' G w 0) = ?R" proof(case_tac "length w = 1", simp_all) assume d: "length w = Suc 0" show "S \ set(init' G w 0 (0, Suc 0)) = ?R" by(subst init'[simplified], simp add: d, subst CYK_Lang[THEN sym], simp add: d) next assume "length w \ Suc 0" with c have "1 < length w" by(case_tac w, simp_all) hence d: "Suc(Suc 0) \ length w" by simp show "(S \ set(main G (init' G w 0) (length w) 0 2 (0, length w))) = (w \ Lang G S)" proof(subst main, simp_all, rule d) fix i' j' assume "j' < 2" and "Suc 0 \ j'" hence e: "j' = 1" by simp assume "i' + j' \ length w" with e have f: "i' + 1 \ length w" by simp have "set(init' G w 0 (i', 1)) = CYK G w i' 1" by(rule init', rule f) with e show "set(init' G w 0 (i', j')) = CYK G w i' j'" by simp next from d show "Suc 0 \ length w" by simp next show "(S \ CYK G w 0 (length w)) = (w \ Lang G S)" by(rule CYK_Lang) qed qed with b show "?L T = ?R" by simp qed qed value [code] "let G = [(0::int, Branch 1 2), (0, Branch 2 3), (1, Branch 2 1), (1, Leaf ''a''), (2, Branch 3 3), (2, Leaf ''b''), (3, Branch 1 2), (3, Leaf ''a'')] in map (cyk G 0) [[''b'',''a'',''a'',''b'',''a''], [''b'',''a'',''b'',''a'']]" end diff --git a/thys/Coinductive/Coinductive_List.thy b/thys/Coinductive/Coinductive_List.thy --- a/thys/Coinductive/Coinductive_List.thy +++ b/thys/Coinductive/Coinductive_List.thy @@ -1,5286 +1,5283 @@ (* Title: Coinductive_List.thy Author: Andreas Lochbihler *) section \Coinductive lists and their operations\ theory Coinductive_List imports "HOL-Library.Infinite_Set" "HOL-Library.Sublist" "HOL-Library.Simps_Case_Conv" Coinductive_Nat begin subsection \Auxiliary lemmata\ lemma funpow_Suc_conv [simp]: "(Suc ^^ n) m = m + n" by(induct n arbitrary: m) simp_all lemma wlog_linorder_le [consumes 0, case_names le symmetry]: assumes le: "\a b :: 'a :: linorder. a \ b \ P a b" and sym: "P b a \ P a b" shows "P a b" proof(cases "a \ b") case True thus ?thesis by(rule le) next case False hence "b \ a" by simp hence "P b a" by(rule le) thus ?thesis by(rule sym) qed subsection \Type definition\ codatatype (lset: 'a) llist = lnull: LNil | LCons (lhd: 'a) (ltl: "'a llist") for map: lmap rel: llist_all2 where "lhd LNil = undefined" | "ltl LNil = LNil" text \ Coiterator setup. \ primcorec unfold_llist :: "('a \ bool) \ ('a \ 'b) \ ('a \ 'a) \ 'a \ 'b llist" where "p a \ unfold_llist p g21 g22 a = LNil" | "_ \ unfold_llist p g21 g22 a = LCons (g21 a) (unfold_llist p g21 g22 (g22 a))" declare unfold_llist.ctr(1) [simp] llist.corec(1) [simp] text \ The following setup should be done by the BNF package. \ text \congruence rule\ declare llist.map_cong [cong] text \Code generator setup\ lemma corec_llist_never_stop: "corec_llist IS_LNIL LHD (\_. False) MORE LTL x = unfold_llist IS_LNIL LHD LTL x" by(coinduction arbitrary: x) auto text \lemmas about generated constants\ lemma eq_LConsD: "xs = LCons y ys \ xs \ LNil \ lhd xs = y \ ltl xs = ys" by auto lemma shows LNil_eq_lmap: "LNil = lmap f xs \ xs = LNil" and lmap_eq_LNil: "lmap f xs = LNil \ xs = LNil" by(cases xs,simp_all)+ declare llist.map_sel(1)[simp] lemma ltl_lmap[simp]: "ltl (lmap f xs) = lmap f (ltl xs)" by(cases xs, simp_all) declare llist.map_ident[simp] lemma lmap_eq_LCons_conv: "lmap f xs = LCons y ys \ (\x xs'. xs = LCons x xs' \ y = f x \ ys = lmap f xs')" by(cases xs)(auto) lemma lmap_conv_unfold_llist: "lmap f = unfold_llist (\xs. xs = LNil) (f \ lhd) ltl" (is "?lhs = ?rhs") proof fix x show "?lhs x = ?rhs x" by(coinduction arbitrary: x) auto qed lemma lmap_unfold_llist: "lmap f (unfold_llist IS_LNIL LHD LTL b) = unfold_llist IS_LNIL (f \ LHD) LTL b" by(coinduction arbitrary: b) auto lemma lmap_corec_llist: "lmap f (corec_llist IS_LNIL LHD endORmore TTL_end TTL_more b) = corec_llist IS_LNIL (f \ LHD) endORmore (lmap f \ TTL_end) TTL_more b" by(coinduction arbitrary: b rule: llist.coinduct_strong) auto lemma unfold_llist_ltl_unroll: "unfold_llist IS_LNIL LHD LTL (LTL b) = unfold_llist (IS_LNIL \ LTL) (LHD \ LTL) LTL b" by(coinduction arbitrary: b) auto lemma ltl_unfold_llist: "ltl (unfold_llist IS_LNIL LHD LTL a) = (if IS_LNIL a then LNil else unfold_llist IS_LNIL LHD LTL (LTL a))" by(simp) lemma unfold_llist_eq_LCons [simp]: "unfold_llist IS_LNIL LHD LTL b = LCons x xs \ \ IS_LNIL b \ x = LHD b \ xs = unfold_llist IS_LNIL LHD LTL (LTL b)" by(subst unfold_llist.code) auto lemma unfold_llist_id [simp]: "unfold_llist lnull lhd ltl xs = xs" by(coinduction arbitrary: xs) simp_all lemma lset_eq_empty [simp]: "lset xs = {} \ lnull xs" by(cases xs) simp_all declare llist.set_sel(1)[simp] lemma lset_ltl: "lset (ltl xs) \ lset xs" by(cases xs) auto lemma in_lset_ltlD: "x \ lset (ltl xs) \ x \ lset xs" using lset_ltl[of xs] by auto text \induction rules\ theorem llist_set_induct[consumes 1, case_names find step]: assumes "x \ lset xs" and "\xs. \ lnull xs \ P (lhd xs) xs" and "\xs y. \\ lnull xs; y \ lset (ltl xs); P y (ltl xs)\ \ P y xs" shows "P x xs" using assms by(induct)(fastforce simp del: llist.disc(2) iff: llist.disc(2), auto) text \Test quickcheck setup\ lemma "\xs. xs = LNil" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "LCons x xs = LCons x xs" quickcheck[narrowing, expect=no_counterexample] oops subsection \Properties of predefined functions\ lemmas lhd_LCons = llist.sel(1) lemmas ltl_simps = llist.sel(2,3) lemmas lhd_LCons_ltl = llist.collapse(2) lemma lnull_ltlI [simp]: "lnull xs \ lnull (ltl xs)" unfolding lnull_def by simp lemma neq_LNil_conv: "xs \ LNil \ (\x xs'. xs = LCons x xs')" by(cases xs) auto lemma not_lnull_conv: "\ lnull xs \ (\x xs'. xs = LCons x xs')" unfolding lnull_def by(simp add: neq_LNil_conv) lemma lset_LCons: "lset (LCons x xs) = insert x (lset xs)" by simp lemma lset_intros: "x \ lset (LCons x xs)" "x \ lset xs \ x \ lset (LCons x' xs)" by simp_all lemma lset_cases [elim?]: assumes "x \ lset xs" obtains xs' where "xs = LCons x xs'" | x' xs' where "xs = LCons x' xs'" "x \ lset xs'" using assms by(cases xs) auto lemma lset_induct' [consumes 1, case_names find step]: assumes major: "x \ lset xs" and 1: "\xs. P (LCons x xs)" and 2: "\x' xs. \ x \ lset xs; P xs \ \ P (LCons x' xs)" shows "P xs" using major apply(induct y\"x" xs rule: llist_set_induct) using 1 2 by(auto simp add: not_lnull_conv) lemma lset_induct [consumes 1, case_names find step, induct set: lset]: assumes major: "x \ lset xs" and find: "\xs. P (LCons x xs)" and step: "\x' xs. \ x \ lset xs; x \ x'; P xs \ \ P (LCons x' xs)" shows "P xs" using major apply(induct rule: lset_induct') apply(rule find) apply(case_tac "x'=x") apply(simp_all add: find step) done lemmas lset_LNil = llist.set(1) lemma lset_lnull: "lnull xs \ lset xs = {}" by(auto dest: llist.collapse) text \Alternative definition of @{term lset} for nitpick\ inductive lsetp :: "'a llist \ 'a \ bool" where "lsetp (LCons x xs) x" | "lsetp xs x \ lsetp (LCons x' xs) x" lemma lset_into_lsetp: "x \ lset xs \ lsetp xs x" by(induct rule: lset_induct)(blast intro: lsetp.intros)+ lemma lsetp_into_lset: "lsetp xs x \ x \ lset xs" by(induct rule: lsetp.induct)(blast intro: lset_intros)+ lemma lset_eq_lsetp [nitpick_unfold]: "lset xs = {x. lsetp xs x}" by(auto intro: lset_into_lsetp dest: lsetp_into_lset) hide_const (open) lsetp hide_fact (open) lsetp.intros lsetp.cases lsetp.induct lset_into_lsetp lset_eq_lsetp text \code setup for @{term lset}\ definition gen_lset :: "'a set \ 'a llist \ 'a set" where "gen_lset A xs = A \ lset xs" lemma gen_lset_code [code]: "gen_lset A LNil = A" "gen_lset A (LCons x xs) = gen_lset (insert x A) xs" by(auto simp add: gen_lset_def) lemma lset_code [code]: "lset = gen_lset {}" by(simp add: gen_lset_def fun_eq_iff) definition lmember :: "'a \ 'a llist \ bool" where "lmember x xs \ x \ lset xs" lemma lmember_code [code]: "lmember x LNil \ False" "lmember x (LCons y ys) \ x = y \ lmember x ys" by(simp_all add: lmember_def) lemma lset_lmember [code_unfold]: "x \ lset xs \ lmember x xs" by(simp add: lmember_def) lemmas lset_lmap [simp] = llist.set_map subsection \The subset of finite lazy lists @{term "lfinite"}\ inductive lfinite :: "'a llist \ bool" where lfinite_LNil: "lfinite LNil" | lfinite_LConsI: "lfinite xs \ lfinite (LCons x xs)" declare lfinite_LNil [iff] lemma lnull_imp_lfinite [simp]: "lnull xs \ lfinite xs" by(auto dest: llist.collapse) lemma lfinite_LCons [simp]: "lfinite (LCons x xs) = lfinite xs" by(auto elim: lfinite.cases intro: lfinite_LConsI) lemma lfinite_ltl [simp]: "lfinite (ltl xs) = lfinite xs" by(cases xs) simp_all lemma lfinite_code [code]: "lfinite LNil = True" "lfinite (LCons x xs) = lfinite xs" by simp_all lemma lfinite_induct [consumes 1, case_names LNil LCons]: assumes lfinite: "lfinite xs" and LNil: "\xs. lnull xs \ P xs" and LCons: "\xs. \ lfinite xs; \ lnull xs; P (ltl xs) \ \ P xs" shows "P xs" using lfinite by(induct)(auto intro: LCons LNil) lemma lfinite_imp_finite_lset: assumes "lfinite xs" shows "finite (lset xs)" using assms by induct auto subsection \Concatenating two lists: @{term "lappend"}\ primcorec lappend :: "'a llist \ 'a llist \ 'a llist" where "lappend xs ys = (case xs of LNil \ ys | LCons x xs' \ LCons x (lappend xs' ys))" simps_of_case lappend_code [code, simp, nitpick_simp]: lappend.code lemmas lappend_LNil_LNil = lappend_code(1)[where ys = LNil] lemma lappend_simps [simp]: shows lhd_lappend: "lhd (lappend xs ys) = (if lnull xs then lhd ys else lhd xs)" and ltl_lappend: "ltl (lappend xs ys) = (if lnull xs then ltl ys else lappend (ltl xs) ys)" by(case_tac [!] xs) simp_all lemma lnull_lappend [simp]: "lnull (lappend xs ys) \ lnull xs \ lnull ys" by(auto simp add: lappend_def) lemma lappend_eq_LNil_iff: "lappend xs ys = LNil \ xs = LNil \ ys = LNil" using lnull_lappend unfolding lnull_def . lemma LNil_eq_lappend_iff: "LNil = lappend xs ys \ xs = LNil \ ys = LNil" by(auto dest: sym simp add: lappend_eq_LNil_iff) lemma lappend_LNil2 [simp]: "lappend xs LNil = xs" by(coinduction arbitrary: xs) simp_all lemma shows lappend_lnull1: "lnull xs \ lappend xs ys = ys" and lappend_lnull2: "lnull ys \ lappend xs ys = xs" unfolding lnull_def by simp_all lemma lappend_assoc: "lappend (lappend xs ys) zs = lappend xs (lappend ys zs)" by(coinduction arbitrary: xs rule: llist.coinduct_strong) auto lemma lmap_lappend_distrib: "lmap f (lappend xs ys) = lappend (lmap f xs) (lmap f ys)" by(coinduction arbitrary: xs rule: llist.coinduct_strong) auto lemma lappend_snocL1_conv_LCons2: "lappend (lappend xs (LCons y LNil)) ys = lappend xs (LCons y ys)" by(simp add: lappend_assoc) lemma lappend_ltl: "\ lnull xs \ lappend (ltl xs) ys = ltl (lappend xs ys)" by simp lemma lfinite_lappend [simp]: "lfinite (lappend xs ys) \ lfinite xs \ lfinite ys" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs proof(induct zs\"lappend xs ys" arbitrary: xs ys) case lfinite_LNil thus ?case by(simp add: LNil_eq_lappend_iff) next case (lfinite_LConsI zs z) thus ?case by(cases xs)(cases ys, simp_all) qed next assume ?rhs (is "?xs \ ?ys") then obtain ?xs ?ys .. thus ?lhs by induct simp_all qed lemma lappend_inf: "\ lfinite xs \ lappend xs ys = xs" by(coinduction arbitrary: xs) auto lemma lfinite_lmap [simp]: "lfinite (lmap f xs) = lfinite xs" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by(induct zs\"lmap f xs" arbitrary: xs rule: lfinite_induct) fastforce+ next assume ?rhs thus ?lhs by(induct) simp_all qed lemma lset_lappend_lfinite [simp]: "lfinite xs \ lset (lappend xs ys) = lset xs \ lset ys" by(induct rule: lfinite.induct) auto lemma lset_lappend: "lset (lappend xs ys) \ lset xs \ lset ys" proof(cases "lfinite xs") case True thus ?thesis by simp next case False thus ?thesis by(auto simp add: lappend_inf) qed lemma lset_lappend1: "lset xs \ lset (lappend xs ys)" by(rule subsetI)(erule lset_induct, simp_all) lemma lset_lappend_conv: "lset (lappend xs ys) = (if lfinite xs then lset xs \ lset ys else lset xs)" by(simp add: lappend_inf) lemma in_lset_lappend_iff: "x \ lset (lappend xs ys) \ x \ lset xs \ lfinite xs \ x \ lset ys" by(simp add: lset_lappend_conv) lemma split_llist_first: assumes "x \ lset xs" shows "\ys zs. xs = lappend ys (LCons x zs) \ lfinite ys \ x \ lset ys" using assms proof(induct) case find thus ?case by(auto intro: exI[where x=LNil]) next case step thus ?case by(fastforce intro: exI[where x="LCons a b" for a b]) qed lemma split_llist: "x \ lset xs \ \ys zs. xs = lappend ys (LCons x zs) \ lfinite ys" by(blast dest: split_llist_first) subsection \The prefix ordering on lazy lists: @{term "lprefix"}\ coinductive lprefix :: "'a llist \ 'a llist \ bool" (infix "\" 65) where LNil_lprefix [simp, intro!]: "LNil \ xs" | Le_LCons: "xs \ ys \ LCons x xs \ LCons x ys" lemma lprefixI [consumes 1, case_names lprefix, case_conclusion lprefix LeLNil LeLCons]: assumes major: "(xs, ys) \ X" and step: "\xs ys. (xs, ys) \ X \ lnull xs \ (\x xs' ys'. xs = LCons x xs' \ ys = LCons x ys' \ ((xs', ys') \ X \ xs' \ ys'))" shows "xs \ ys" using major by(rule lprefix.coinduct)(auto dest: step) lemma lprefix_coinduct [consumes 1, case_names lprefix, case_conclusion lprefix LNil LCons, coinduct pred: lprefix]: assumes major: "P xs ys" and step: "\xs ys. P xs ys \ (lnull ys \ lnull xs) \ (\ lnull xs \ \ lnull ys \ lhd xs = lhd ys \ (P (ltl xs) (ltl ys) \ ltl xs \ ltl ys))" shows "xs \ ys" proof - from major have "(xs, ys) \ {(xs, ys). P xs ys}" by simp thus ?thesis proof(coinduct rule: lprefixI) case (lprefix xs ys) hence "P xs ys" by simp from step[OF this] show ?case by(cases xs)(auto simp add: not_lnull_conv) qed qed lemma lprefix_refl [intro, simp]: "xs \ xs" by(coinduction arbitrary: xs) simp_all lemma lprefix_LNil [simp]: "xs \ LNil \ lnull xs" unfolding lnull_def by(subst lprefix.simps)simp lemma lprefix_lnull: "lnull ys \ xs \ ys \ lnull xs" unfolding lnull_def by auto lemma lnull_lprefix: "lnull xs \ lprefix xs ys" by(simp add: lnull_def) lemma lprefix_LCons_conv: "xs \ LCons y ys \ xs = LNil \ (\xs'. xs = LCons y xs' \ xs' \ ys)" by(subst lprefix.simps) simp lemma LCons_lprefix_LCons [simp]: "LCons x xs \ LCons y ys \ x = y \ xs \ ys" by(simp add: lprefix_LCons_conv) lemma LCons_lprefix_conv: "LCons x xs \ ys \ (\ys'. ys = LCons x ys' \ xs \ ys')" by(cases ys)(auto) lemma lprefix_ltlI: "xs \ ys \ ltl xs \ ltl ys" by(cases ys)(auto simp add: lprefix_LCons_conv) lemma lprefix_code [code]: "LNil \ ys \ True" "LCons x xs \ LNil \ False" "LCons x xs \ LCons y ys \ x = y \ xs \ ys" by simp_all lemma lprefix_lhdD: "\ xs \ ys; \ lnull xs \ \ lhd xs = lhd ys" by(clarsimp simp add: not_lnull_conv LCons_lprefix_conv) lemma lprefix_lnullD: "\ xs \ ys; lnull ys \ \ lnull xs" by(auto elim: lprefix.cases) lemma lprefix_not_lnullD: "\ xs \ ys; \ lnull xs \ \ \ lnull ys" by(auto elim: lprefix.cases) lemma lprefix_expand: "(\ lnull xs \ \ lnull ys \ lhd xs = lhd ys \ ltl xs \ ltl ys) \ xs \ ys" by(cases xs ys rule: llist.exhaust[case_product llist.exhaust])(simp_all) lemma lprefix_antisym: "\ xs \ ys; ys \ xs \ \ xs = ys" by(coinduction arbitrary: xs ys)(auto simp add: not_lnull_conv lprefix_lnull) lemma lprefix_trans [trans]: "\ xs \ ys; ys \ zs \ \ xs \ zs" by(coinduction arbitrary: xs ys zs)(auto 4 3 dest: lprefix_lnullD lprefix_lhdD intro: lprefix_ltlI) lemma preorder_lprefix [cont_intro]: "class.preorder (\) (mk_less (\))" by(unfold_locales)(auto simp add: mk_less_def intro: lprefix_trans) lemma lprefix_lsetD: assumes "xs \ ys" shows "lset xs \ lset ys" proof fix x assume "x \ lset xs" thus "x \ lset ys" using assms by(induct arbitrary: ys)(auto simp add: LCons_lprefix_conv) qed lemma lprefix_lappend_sameI: assumes "xs \ ys" shows "lappend zs xs \ lappend zs ys" proof(cases "lfinite zs") case True thus ?thesis using assms by induct auto qed(simp add: lappend_inf) lemma not_lfinite_lprefix_conv_eq: assumes nfin: "\ lfinite xs" shows "xs \ ys \ xs = ys" proof assume "xs \ ys" with nfin show "xs = ys" by(coinduction arbitrary: xs ys)(auto dest: lprefix_lnullD lprefix_lhdD intro: lprefix_ltlI) qed simp lemma lprefix_lappend: "xs \ lappend xs ys" by(coinduction arbitrary: xs) auto lemma lprefix_down_linear: assumes "xs \ zs" "ys \ zs" shows "xs \ ys \ ys \ xs" proof(rule disjCI) assume "\ ys \ xs" with assms show "xs \ ys" by(coinduction arbitrary: xs ys zs)(auto simp add: not_lnull_conv LCons_lprefix_conv, simp add: lnull_def) qed lemma lprefix_lappend_same [simp]: "lappend xs ys \ lappend xs zs \ (lfinite xs \ ys \ zs)" (is "?lhs \ ?rhs") proof assume "?lhs" show "?rhs" proof assume "lfinite xs" thus "ys \ zs" using \?lhs\ by(induct) simp_all qed next assume "?rhs" show ?lhs proof(cases "lfinite xs") case True hence yszs: "ys \ zs" by(rule \?rhs\[rule_format]) from True show ?thesis by induct (simp_all add: yszs) next case False thus ?thesis by(simp add: lappend_inf) qed qed subsection \Setup for partial\_function\ primcorec lSup :: "'a llist set \ 'a llist" where "lSup A = (if \x\A. lnull x then LNil else LCons (THE x. x \ lhd ` (A \ {xs. \ lnull xs})) (lSup (ltl ` (A \ {xs. \ lnull xs}))))" declare lSup.simps[simp del] lemma lnull_lSup [simp]: "lnull (lSup A) \ (\x\A. lnull x)" by(simp add: lSup_def) lemma lhd_lSup [simp]: "\x\A. \ lnull x \ lhd (lSup A) = (THE x. x \ lhd ` (A \ {xs. \ lnull xs}))" by(auto simp add: lSup_def) lemma ltl_lSup [simp]: "ltl (lSup A) = lSup (ltl ` (A \ {xs. \ lnull xs}))" by(cases "\xs\A. lnull xs")(auto 4 3 simp add: lSup_def intro: llist.expand) lemma lhd_lSup_eq: assumes chain: "Complete_Partial_Order.chain (\) Y" shows "\ xs \ Y; \ lnull xs \ \ lhd (lSup Y) = lhd xs" by(subst lhd_lSup)(auto 4 3 dest: chainD[OF chain] lprefix_lhdD intro!: the_equality) lemma lSup_empty [simp]: "lSup {} = LNil" by(simp add: lSup_def) lemma lSup_singleton [simp]: "lSup {xs} = xs" by(coinduction arbitrary: xs) simp_all lemma LCons_image_Int_not_lnull: "(LCons x ` A \ {xs. \ lnull xs}) = LCons x ` A" by auto lemma lSup_LCons: "A \ {} \ lSup (LCons x ` A) = LCons x (lSup A)" by(rule llist.expand)(auto simp add: image_image lhd_lSup exI LCons_image_Int_not_lnull intro!: the_equality) lemma lSup_eq_LCons_iff: "lSup Y = LCons x xs \ (\x\Y. \ lnull x) \ x = (THE x. x \ lhd ` (Y \ {xs. \ lnull xs})) \ xs = lSup (ltl ` (Y \ {xs. \ lnull xs}))" by(auto dest: eq_LConsD simp add: lnull_def[symmetric] intro: llist.expand) lemma lSup_insert_LNil: "lSup (insert LNil Y) = lSup Y" by(rule llist.expand) simp_all lemma lSup_minus_LNil: "lSup (Y - {LNil}) = lSup Y" using lSup_insert_LNil[where Y="Y - {LNil}", symmetric] by(cases "LNil \ Y")(simp_all add: insert_absorb) lemma chain_lprefix_ltl: assumes chain: "Complete_Partial_Order.chain (\) A" shows "Complete_Partial_Order.chain (\) (ltl ` (A \ {xs. \ lnull xs}))" by(auto intro!: chainI dest: chainD[OF chain] intro: lprefix_ltlI) lemma lSup_finite_prefixes: "lSup {ys. ys \ xs \ lfinite ys} = xs" (is "lSup (?C xs) = _") proof(coinduction arbitrary: xs) case (Eq_llist xs) have ?lnull by(cases xs)(auto simp add: lprefix_LCons_conv) moreover have "\ lnull xs \ ltl ` ({ys. ys \ xs \ lfinite ys} \ {xs. \ lnull xs}) = {ys. ys \ ltl xs \ lfinite ys}" by(auto 4 4 intro!: rev_image_eqI[where x="LCons (lhd xs) ys" for ys] intro: llist.expand lprefix_ltlI simp add: LCons_lprefix_conv) hence ?LCons by(auto 4 3 intro!: the_equality dest: lprefix_lhdD intro: rev_image_eqI) ultimately show ?case .. qed lemma lSup_finite_gen_prefixes: assumes "zs \ xs" "lfinite zs" shows "lSup {ys. ys \ xs \ zs \ ys \ lfinite ys} = xs" using \lfinite zs\ \zs \ xs\ proof(induction arbitrary: xs) case lfinite_LNil thus ?case by(simp add: lSup_finite_prefixes) next case (lfinite_LConsI zs z) from \LCons z zs \ xs\ obtain xs' where xs: "xs = LCons z xs'" and "zs \ xs'" by(auto simp add: LCons_lprefix_conv) show ?case (is "?lhs = ?rhs") proof(rule llist.expand) show "lnull ?lhs = lnull ?rhs" using xs lfinite_LConsI by(auto 4 3 simp add: lprefix_LCons_conv del: disjCI intro: disjI2) next assume lnull: "\ lnull ?lhs" "\ lnull ?rhs" have "lhd ?lhs = lhd ?rhs" using lnull xs by(auto intro!: rev_image_eqI simp add: LCons_lprefix_conv) moreover have "ltl ` ({ys. ys \ xs \ LCons z zs \ ys \ lfinite ys} \ {xs. \ lnull xs}) = {ys. ys \ xs' \ zs \ ys \ lfinite ys}" using xs \\ lnull ?rhs\ by(auto 4 3 simp add: lprefix_LCons_conv intro: rev_image_eqI del: disjCI intro: disjI2) hence "ltl ?lhs = ltl ?rhs" using lfinite_LConsI.IH[OF \zs \ xs'\] xs by simp ultimately show "lhd ?lhs = lhd ?rhs \ ltl ?lhs = ltl ?rhs" .. qed qed lemma lSup_strict_prefixes: "\ lfinite xs \ lSup {ys. ys \ xs \ ys \ xs} = xs" (is "_ \ lSup (?C xs) = _") proof(coinduction arbitrary: xs) case (Eq_llist xs) then obtain x x' xs' where xs: "xs = LCons x (LCons x' xs')" "\ lfinite xs'" by(cases xs)(simp, rename_tac xs', case_tac xs', auto) hence ?lnull by(auto intro: exI[where x="LCons x (LCons x' LNil)"]) moreover hence "\ lnull (lSup {ys. ys \ xs \ ys \ xs})" using xs by auto hence "lhd (lSup {ys. ys \ xs \ ys \ xs}) = lhd xs" using xs by(auto 4 3 intro!: the_equality intro: rev_image_eqI dest: lprefix_lhdD) moreover from xs have "ltl ` ({ys. ys \ xs \ ys \ xs} \ {xs. \ lnull xs}) = {ys. ys \ ltl xs \ ys \ ltl xs}" by(auto simp add: lprefix_LCons_conv intro: image_eqI[where x="LCons x (LCons x' ys)" for ys] image_eqI[where x="LCons x LNil"]) ultimately show ?case using Eq_llist by clarsimp qed lemma chain_lprefix_lSup: "\ Complete_Partial_Order.chain (\) A; xs \ A \ \ xs \ lSup A" proof(coinduction arbitrary: xs A) case (lprefix xs A) note chain = \Complete_Partial_Order.chain (\) A\ from \xs \ A\ show ?case by(auto 4 3 dest: chainD[OF chain] lprefix_lhdD intro: chain_lprefix_ltl[OF chain] intro!: the_equality[symmetric]) qed lemma chain_lSup_lprefix: "\ Complete_Partial_Order.chain (\) A; \xs. xs \ A \ xs \ zs \ \ lSup A \ zs" proof(coinduction arbitrary: A zs) case (lprefix A zs) note chain = \Complete_Partial_Order.chain (\) A\ from \\xs. xs \ A \ xs \ zs\ show ?case by(auto 4 4 dest: lprefix_lnullD lprefix_lhdD intro: chain_lprefix_ltl[OF chain] lprefix_ltlI rev_image_eqI intro!: the_equality) qed lemma llist_ccpo [simp, cont_intro]: "class.ccpo lSup (\) (mk_less (\))" by(unfold_locales)(auto dest: lprefix_antisym intro: lprefix_trans chain_lprefix_lSup chain_lSup_lprefix simp add: mk_less_def) lemmas [cont_intro] = ccpo.admissible_leI[OF llist_ccpo] lemma llist_partial_function_definitions: "partial_function_definitions (\) lSup" by(unfold_locales)(auto dest: lprefix_antisym intro: lprefix_trans chain_lprefix_lSup chain_lSup_lprefix) interpretation llist: partial_function_definitions "(\)" lSup rewrites "lSup {} \ LNil" by(rule llist_partial_function_definitions)(simp) abbreviation "mono_llist \ monotone (fun_ord (\)) (\)" interpretation llist_lift: partial_function_definitions "fun_ord lprefix" "fun_lub lSup" rewrites "fun_lub lSup {} \ \_. LNil" by(rule llist_partial_function_definitions[THEN partial_function_lift])(simp) abbreviation "mono_llist_lift \ monotone (fun_ord (fun_ord lprefix)) (fun_ord lprefix)" lemma lprefixes_chain: "Complete_Partial_Order.chain (\) {ys. lprefix ys xs}" by(rule chainI)(auto dest: lprefix_down_linear) lemma llist_gen_induct: assumes adm: "ccpo.admissible lSup (\) P" and step: "\zs. zs \ xs \ lfinite zs \ (\ys. zs \ ys \ ys \ xs \ lfinite ys \ P ys)" shows "P xs" proof - from step obtain zs where zs: "zs \ xs" "lfinite zs" and ys: "\ys. \ zs \ ys; ys \ xs; lfinite ys \ \ P ys" by blast let ?C = "{ys. ys \ xs \ zs \ ys \ lfinite ys}" from lprefixes_chain[of xs] have "Complete_Partial_Order.chain (\) ?C" by(auto dest: chain_compr) with adm have "P (lSup ?C)" by(rule ccpo.admissibleD)(auto intro: ys zs) also have "lSup ?C = xs" using lSup_finite_gen_prefixes[OF zs] by simp finally show ?thesis . qed lemma llist_induct [case_names adm LNil LCons, induct type: llist]: assumes adm: "ccpo.admissible lSup (\) P" and LNil: "P LNil" and LCons: "\x xs. \ lfinite xs; P xs \ \ P (LCons x xs)" shows "P xs" proof - { fix ys :: "'a llist" assume "lfinite ys" hence "P ys" by(induct)(simp_all add: LNil LCons) } note [intro] = this show ?thesis using adm by(rule llist_gen_induct)(auto intro: exI[where x=LNil]) qed lemma LCons_mono [partial_function_mono, cont_intro]: "mono_llist A \ mono_llist (\f. LCons x (A f))" by(rule monotoneI)(auto dest: monotoneD) lemma mono2mono_LCons [THEN llist.mono2mono, simp, cont_intro]: shows monotone_LCons: "monotone (\) (\) (LCons x)" by(auto intro: monotoneI) lemma mcont2mcont_LCons [THEN llist.mcont2mcont, simp, cont_intro]: shows mcont_LCons: "mcont lSup (\) lSup (\) (LCons x)" by(simp add: mcont_def monotone_LCons lSup_LCons[symmetric] contI) lemma mono2mono_ltl[THEN llist.mono2mono, simp, cont_intro]: shows monotone_ltl: "monotone (\) (\) ltl" by(auto intro: monotoneI lprefix_ltlI) lemma cont_ltl: "cont lSup (\) lSup (\) ltl" proof(rule contI) fix Y :: "'a llist set" assume "Y \ {}" have "ltl (lSup Y) = lSup (insert LNil (ltl ` (Y \ {xs. \ lnull xs})))" by(simp add: lSup_insert_LNil) also have "insert LNil (ltl ` (Y \ {xs. \ lnull xs})) = insert LNil (ltl ` Y)" by auto also have "lSup \ = lSup (ltl ` Y)" by(simp add: lSup_insert_LNil) finally show "ltl (lSup Y) = lSup (ltl ` Y)" . qed lemma mcont2mcont_ltl [THEN llist.mcont2mcont, simp, cont_intro]: shows mcont_ltl: "mcont lSup (\) lSup (\) ltl" by(simp add: mcont_def monotone_ltl cont_ltl) lemma llist_case_mono [partial_function_mono, cont_intro]: assumes lnil: "monotone orda ordb lnil" and lcons: "\x xs. monotone orda ordb (\f. lcons f x xs)" shows "monotone orda ordb (\f. case_llist (lnil f) (lcons f) x)" by(rule monotoneI)(auto split: llist.split dest: monotoneD[OF lnil] monotoneD[OF lcons]) lemma mcont_llist_case [cont_intro, simp]: "\ mcont luba orda lubb ordb (\x. f x); \x xs. mcont luba orda lubb ordb (\y. g x xs y) \ \ mcont luba orda lubb ordb (\y. case xs of LNil \ f y | LCons x xs' \ g x xs' y)" by(cases xs) simp_all lemma monotone_lprefix_case [cont_intro, simp]: assumes mono: "\x. monotone (\) (\) (\xs. f x xs (LCons x xs))" shows "monotone (\) (\) (\xs. case xs of LNil \ LNil | LCons x xs' \ f x xs' xs)" by(rule llist.monotone_if_bot[where f="\xs. f (lhd xs) (ltl xs) xs" and bound=LNil])(auto 4 3 split: llist.split simp add: not_lnull_conv LCons_lprefix_conv dest: monotoneD[OF mono]) lemma mcont_lprefix_case_aux: fixes f bot defines "g \ \xs. f (lhd xs) (ltl xs) (LCons (lhd xs) (ltl xs))" assumes mcont: "\x. mcont lSup (\) lub ord (\xs. f x xs (LCons x xs))" and ccpo: "class.ccpo lub ord (mk_less ord)" and bot: "\x. ord bot x" shows "mcont lSup (\) lub ord (\xs. case xs of LNil \ bot | LCons x xs' \ f x xs' xs)" proof(rule llist.mcont_if_bot[where f=g and bound=LNil and bot=bot, OF ccpo bot]) fix Y :: "'a llist set" assume chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" "\x. x \ Y \ \ x \ LNil" from Y have Y': "Y \ {xs. \ lnull xs} \ {}" by auto from Y obtain x xs where "LCons x xs \ Y" by(fastforce simp add: not_lnull_conv) with Y(2) have eq: "Y = LCons x ` (ltl ` (Y \ {xs. \ lnull xs}))" by(force dest: chainD[OF chain] simp add: LCons_lprefix_conv lprefix_LCons_conv intro: imageI rev_image_eqI) show "g (lSup Y) = lub (g ` Y)" by(subst (1 2) eq)(simp add: lSup_LCons Y' g_def mcont_contD[OF mcont] chain chain_lprefix_ltl, simp add: image_image) qed(auto 4 3 split: llist.split simp add: not_lnull_conv LCons_lprefix_conv g_def intro: mcont_monoD[OF mcont]) lemma mcont_lprefix_case [cont_intro, simp]: assumes "\x. mcont lSup (\) lSup (\) (\xs. f x xs (LCons x xs))" shows "mcont lSup (\) lSup (\) (\xs. case xs of LNil \ LNil | LCons x xs' \ f x xs' xs)" using assms by(rule mcont_lprefix_case_aux)(simp_all add: llist_ccpo) lemma monotone_lprefix_case_lfp [cont_intro, simp]: fixes f :: "_ \ _ :: order_bot" assumes mono: "\x. monotone (\) (\) (\xs. f x xs (LCons x xs))" shows "monotone (\) (\) (\xs. case xs of LNil \ \ | LCons x xs \ f x xs (LCons x xs))" by(rule llist.monotone_if_bot[where bound=LNil and bot=\ and f="\xs. f (lhd xs) (ltl xs) xs"])(auto 4 3 simp add: not_lnull_conv LCons_lprefix_conv dest: monotoneD[OF mono] split: llist.split) lemma mcont_lprefix_case_lfp [cont_intro, simp]: fixes f :: "_ => _ :: complete_lattice" assumes "\x. mcont lSup (\) Sup (\) (\xs. f x xs (LCons x xs))" shows "mcont lSup (\) Sup (\) (\xs. case xs of LNil \ \ | LCons x xs \ f x xs (LCons x xs))" using assms by(rule mcont_lprefix_case_aux)(rule complete_lattice_ccpo', simp) declaration \Partial_Function.init "llist" @{term llist.fixp_fun} @{term llist.mono_body} @{thm llist.fixp_rule_uc} @{thm llist.fixp_strong_induct_uc} NONE\ subsection \Monotonicity and continuity of already defined functions\ lemma fixes f F defines "F \ \lmap xs. case xs of LNil \ LNil | LCons x xs \ LCons (f x) (lmap xs)" shows lmap_conv_fixp: "lmap f \ ccpo.fixp (fun_lub lSup) (fun_ord (\)) F" (is "?lhs \ ?rhs") and lmap_mono: "\xs. mono_llist (\lmap. F lmap xs)" (is "PROP ?mono") proof(intro eq_reflection ext) show mono: "PROP ?mono" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix xs show "?lhs xs = ?rhs xs" proof(coinduction arbitrary: xs) case Eq_llist show ?case by(subst (1 3 4) llist.mono_body_fixp[OF mono])(auto simp add: F_def split: llist.split) qed qed lemma mono2mono_lmap[THEN llist.mono2mono, simp, cont_intro]: shows monotone_lmap: "monotone (\) (\) (lmap f)" by(rule llist.fixp_preserves_mono1[OF lmap_mono lmap_conv_fixp]) simp lemma mcont2mcont_lmap[THEN llist.mcont2mcont, simp, cont_intro]: shows mcont_lmap: "mcont lSup (\) lSup (\) (lmap f)" by(rule llist.fixp_preserves_mcont1[OF lmap_mono lmap_conv_fixp]) simp lemma [partial_function_mono]: "mono_llist F \ mono_llist (\f. lmap g (F f))" by(rule mono2mono_lmap) lemma mono_llist_lappend2 [partial_function_mono]: "mono_llist A \ mono_llist (\f. lappend xs (A f))" by(auto intro!: monotoneI lprefix_lappend_sameI simp add: fun_ord_def dest: monotoneD) lemma mono2mono_lappend2 [THEN llist.mono2mono, cont_intro, simp]: shows monotone_lappend2: "monotone (\) (\) (lappend xs)" by(rule monotoneI)(rule lprefix_lappend_sameI) lemma mcont2mcont_lappend2 [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_lappend2: "mcont lSup (\) lSup (\) (lappend xs)" proof(cases "lfinite xs") case True thus ?thesis by induct(simp_all add: monotone_lappend2) next case False hence "lappend xs = (\_. xs)" by(simp add: fun_eq_iff lappend_inf) thus ?thesis by(simp add: ccpo.cont_const[OF llist_ccpo]) qed lemma fixes f F defines "F \ \lset xs. case xs of LNil \ {} | LCons x xs \ insert x (lset xs)" shows lset_conv_fixp: "lset \ ccpo.fixp (fun_lub Union) (fun_ord (\)) F" (is "_ \ ?fixp") and lset_mono: "\x. monotone (fun_ord (\)) (\) (\f. F f x)" (is "PROP ?mono") proof(rule eq_reflection ext antisym subsetI)+ show mono: "PROP ?mono" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix x xs show "?fixp xs \ lset xs" by(induct arbitrary: xs rule: lfp.fixp_induct_uc[of "\x. x" F "\x. x", OF mono reflexive refl])(auto simp add: F_def split: llist.split) assume "x \ lset xs" thus "x \ ?fixp xs" by induct(subst lfp.mono_body_fixp[OF mono], simp add: F_def)+ qed lemma mono2mono_lset [THEN lfp.mono2mono, cont_intro, simp]: shows monotone_lset: "monotone (\) (\) lset" by(rule lfp.fixp_preserves_mono1[OF lset_mono lset_conv_fixp]) simp lemma mcont2mcont_lset[THEN mcont2mcont, cont_intro, simp]: shows mcont_lset: "mcont lSup (\) Union (\) lset" by(rule lfp.fixp_preserves_mcont1[OF lset_mono lset_conv_fixp]) simp lemma lset_lSup: "Complete_Partial_Order.chain (\) Y \ lset (lSup Y) = \(lset ` Y)" by(cases "Y = {}")(simp_all add: mcont_lset[THEN mcont_contD]) lemma lfinite_lSupD: "lfinite (lSup A) \ \xs \ A. lfinite xs" by(induct ys\"lSup A" arbitrary: A rule: lfinite_induct) fastforce+ lemma monotone_enat_le_lprefix_case [cont_intro, simp]: "monotone (\) (\) (\x. f x (eSuc x)) \ monotone (\) (\) (\x. case x of 0 \ LNil | eSuc x' \ f x' x)" by(erule monotone_enat_le_case) simp_all lemma mcont_enat_le_lprefix_case [cont_intro, simp]: assumes "mcont Sup (\) lSup (\) (\x. f x (eSuc x))" shows "mcont Sup (\) lSup (\) (\x. case x of 0 \ LNil | eSuc x' \ f x' x)" using llist_ccpo assms by(rule mcont_enat_le_case) simp lemma compact_LConsI: assumes "ccpo.compact lSup (\) xs" shows "ccpo.compact lSup (\) (LCons x xs)" using llist_ccpo proof(rule ccpo.compactI) from assms have "ccpo.admissible lSup (\) (\ys. \ xs \ ys)" by cases hence [cont_intro]: "ccpo.admissible lSup (\) (\ys. \ xs \ ltl ys)" using mcont_ltl by(rule admissible_subst) have [cont_intro]: "ccpo.admissible lSup (\) (\ys. \ lnull ys \ lhd ys \ x)" proof(rule ccpo.admissibleI) fix Y assume chain: "Complete_Partial_Order.chain (\) Y" and *: "Y \ {}" "\ys\Y. \ lnull ys \ lhd ys \ x" from * show "\ lnull (lSup Y) \ lhd (lSup Y) \ x" by(subst lhd_lSup)(auto 4 4 dest: chainD[OF chain] intro!: the_equality[symmetric] dest: lprefix_lhdD) qed have eq: "(\ys. \ LCons x xs \ ys) = (\ys. \ xs \ ltl ys \ ys = LNil \ \ lnull ys \ lhd ys \ x)" by(auto simp add: fun_eq_iff LCons_lprefix_conv neq_LNil_conv) show "ccpo.admissible lSup (\) (\ys. \ LCons x xs \ ys)" by(simp add: eq) qed lemma compact_LConsD: assumes "ccpo.compact lSup (\) (LCons x xs)" shows "ccpo.compact lSup (\) xs" using llist_ccpo proof(rule ccpo.compactI) from assms have "ccpo.admissible lSup (\) (\ys. \ LCons x xs \ ys)" by cases hence "ccpo.admissible lSup (\) (\ys. \ LCons x xs \ LCons x ys)" by(rule admissible_subst)(rule mcont_LCons) thus "ccpo.admissible lSup (\) (\ys. \ xs \ ys)" by simp qed lemma compact_LCons_iff [simp]: "ccpo.compact lSup (\) (LCons x xs) \ ccpo.compact lSup (\) xs" by(blast intro: compact_LConsI compact_LConsD) lemma compact_lfiniteI: "lfinite xs \ ccpo.compact lSup (\) xs" by(induction rule: lfinite.induct) simp_all lemma compact_lfiniteD: assumes "ccpo.compact lSup (\) xs" shows "lfinite xs" proof(rule ccontr) assume inf: "\ lfinite xs" from assms have "ccpo.admissible lSup (\) (\ys. \ xs \ ys)" by cases moreover let ?C = "{ys. ys \ xs \ ys \ xs}" have "Complete_Partial_Order.chain (\) ?C" using lprefixes_chain[of xs] by(auto dest: chain_compr) moreover have "?C \ {}" using inf by(cases xs) auto ultimately have "\ xs \ lSup ?C" by(rule ccpo.admissibleD)(auto dest: lprefix_antisym) also have "lSup ?C = xs" using inf by(rule lSup_strict_prefixes) finally show False by simp qed lemma compact_eq_lfinite [simp]: "ccpo.compact lSup (\) = lfinite" by(blast intro: compact_lfiniteI compact_lfiniteD) subsection \More function definitions\ primcorec iterates :: "('a \ 'a) \ 'a \ 'a llist" where "iterates f x = LCons x (iterates f (f x))" primrec llist_of :: "'a list \ 'a llist" where "llist_of [] = LNil" | "llist_of (x#xs) = LCons x (llist_of xs)" definition list_of :: "'a llist \ 'a list" where [code del]: "list_of xs = (if lfinite xs then inv llist_of xs else undefined)" definition llength :: "'a llist \ enat" where [code del]: "llength = enat_unfold lnull ltl" primcorec ltake :: "enat \ 'a llist \ 'a llist" where "n = 0 \ lnull xs \ lnull (ltake n xs)" | "lhd (ltake n xs) = lhd xs" | "ltl (ltake n xs) = ltake (epred n) (ltl xs)" definition ldropn :: "nat \ 'a llist \ 'a llist" where "ldropn n xs = (ltl ^^ n) xs" context notes [[function_internals]] begin partial_function (llist) ldrop :: "enat \ 'a llist \ 'a llist" where "ldrop n xs = (case n of 0 \ xs | eSuc n' \ case xs of LNil \ LNil | LCons x xs' \ ldrop n' xs')" end primcorec ltakeWhile :: "('a \ bool) \ 'a llist \ 'a llist" where "lnull xs \ \ P (lhd xs) \ lnull (ltakeWhile P xs)" | "lhd (ltakeWhile P xs) = lhd xs" | "ltl (ltakeWhile P xs) = ltakeWhile P (ltl xs)" context fixes P :: "'a \ bool" notes [[function_internals]] begin partial_function (llist) ldropWhile :: "'a llist \ 'a llist" where "ldropWhile xs = (case xs of LNil \ LNil | LCons x xs' \ if P x then ldropWhile xs' else xs)" partial_function (llist) lfilter :: "'a llist \ 'a llist" where "lfilter xs = (case xs of LNil \ LNil | LCons x xs' \ if P x then LCons x (lfilter xs') else lfilter xs')" end primrec lnth :: "'a llist \ nat \ 'a" where "lnth xs 0 = (case xs of LNil \ undefined (0 :: nat) | LCons x xs' \ x)" | "lnth xs (Suc n) = (case xs of LNil \ undefined (Suc n) | LCons x xs' \ lnth xs' n)" declare lnth.simps [simp del] primcorec lzip :: "'a llist \ 'b llist \ ('a \ 'b) llist" where "lnull xs \ lnull ys \ lnull (lzip xs ys)" | "lhd (lzip xs ys) = (lhd xs, lhd ys)" | "ltl (lzip xs ys) = lzip (ltl xs) (ltl ys)" definition llast :: "'a llist \ 'a" where [nitpick_simp]: "llast xs = (case llength xs of enat n \ (case n of 0 \ undefined | Suc n' \ lnth xs n') | \ \ undefined)" coinductive ldistinct :: "'a llist \ bool" where LNil [simp]: "ldistinct LNil" | LCons: "\ x \ lset xs; ldistinct xs \ \ ldistinct (LCons x xs)" hide_fact (open) LNil LCons definition inf_llist :: "(nat \ 'a) \ 'a llist" where [code del]: "inf_llist f = lmap f (iterates Suc 0)" abbreviation repeat :: "'a \ 'a llist" where "repeat \ iterates (\x. x)" definition lstrict_prefix :: "'a llist \ 'a llist \ bool" where [code del]: "lstrict_prefix xs ys \ xs \ ys \ xs \ ys" text \longest common prefix\ definition llcp :: "'a llist \ 'a llist \ enat" where [code del]: "llcp xs ys = enat_unfold (\(xs, ys). lnull xs \ lnull ys \ lhd xs \ lhd ys) (map_prod ltl ltl) (xs, ys)" coinductive llexord :: "('a \ 'a \ bool) \ 'a llist \ 'a llist \ bool" for r :: "'a \ 'a \ bool" where llexord_LCons_eq: "llexord r xs ys \ llexord r (LCons x xs) (LCons x ys)" | llexord_LCons_less: "r x y \ llexord r (LCons x xs) (LCons y ys)" | llexord_LNil [simp, intro!]: "llexord r LNil ys" context notes [[function_internals]] begin partial_function (llist) lconcat :: "'a llist llist \ 'a llist" where "lconcat xss = (case xss of LNil \ LNil | LCons xs xss' \ lappend xs (lconcat xss'))" end definition lhd' :: "'a llist \ 'a option" where "lhd' xs = (if lnull xs then None else Some (lhd xs))" lemma lhd'_simps[simp]: "lhd' LNil = None" "lhd' (LCons x xs) = Some x" unfolding lhd'_def by auto definition ltl' :: "'a llist \ 'a llist option" where "ltl' xs = (if lnull xs then None else Some (ltl xs))" lemma ltl'_simps[simp]: "ltl' LNil = None" "ltl' (LCons x xs) = Some xs" unfolding ltl'_def by auto definition lnths :: "'a llist \ nat set \ 'a llist" where "lnths xs A = lmap fst (lfilter (\(x, y). y \ A) (lzip xs (iterates Suc 0)))" definition (in monoid_add) lsum_list :: "'a llist \ 'a" where "lsum_list xs = (if lfinite xs then sum_list (list_of xs) else 0)" subsection \Converting ordinary lists to lazy lists: @{term "llist_of"}\ lemma lhd_llist_of [simp]: "lhd (llist_of xs) = hd xs" by(cases xs)(simp_all add: hd_def lhd_def) lemma ltl_llist_of [simp]: "ltl (llist_of xs) = llist_of (tl xs)" by(cases xs) simp_all lemma lfinite_llist_of [simp]: "lfinite (llist_of xs)" by(induct xs) auto lemma lfinite_eq_range_llist_of: "lfinite xs \ xs \ range llist_of" proof assume "lfinite xs" thus "xs \ range llist_of" by(induct rule: lfinite.induct)(auto intro: llist_of.simps[symmetric]) next assume "xs \ range llist_of" thus "lfinite xs" by(auto intro: lfinite_llist_of) qed lemma lnull_llist_of [simp]: "lnull (llist_of xs) \ xs = []" by(cases xs) simp_all lemma llist_of_eq_LNil_conv: "llist_of xs = LNil \ xs = []" by(cases xs) simp_all lemma llist_of_eq_LCons_conv: "llist_of xs = LCons y ys \ (\xs'. xs = y # xs' \ ys = llist_of xs')" by(cases xs) auto lemma lappend_llist_of_llist_of: "lappend (llist_of xs) (llist_of ys) = llist_of (xs @ ys)" by(induct xs) simp_all lemma lfinite_rev_induct [consumes 1, case_names Nil snoc]: assumes fin: "lfinite xs" and Nil: "P LNil" and snoc: "\x xs. \ lfinite xs; P xs \ \ P (lappend xs (LCons x LNil))" shows "P xs" proof - from fin obtain xs' where xs: "xs = llist_of xs'" unfolding lfinite_eq_range_llist_of by blast show ?thesis unfolding xs by(induct xs' rule: rev_induct)(auto simp add: Nil lappend_llist_of_llist_of[symmetric] dest: snoc[rotated]) qed lemma lappend_llist_of_LCons: "lappend (llist_of xs) (LCons y ys) = lappend (llist_of (xs @ [y])) ys" by(induct xs) simp_all lemma lmap_llist_of [simp]: "lmap f (llist_of xs) = llist_of (map f xs)" by(induct xs) simp_all lemma lset_llist_of [simp]: "lset (llist_of xs) = set xs" by(induct xs) simp_all lemma llist_of_inject [simp]: "llist_of xs = llist_of ys \ xs = ys" proof assume "llist_of xs = llist_of ys" thus "xs = ys" proof(induct xs arbitrary: ys) case Nil thus ?case by(cases ys) auto next case Cons thus ?case by(cases ys) auto qed qed simp lemma inj_llist_of [simp]: "inj llist_of" by(rule inj_onI) simp subsection \Converting finite lazy lists to ordinary lists: @{term "list_of"}\ lemma list_of_llist_of [simp]: "list_of (llist_of xs) = xs" by(fastforce simp add: list_of_def intro: inv_f_f inj_onI) lemma llist_of_list_of [simp]: "lfinite xs \ llist_of (list_of xs) = xs" unfolding lfinite_eq_range_llist_of by auto lemma list_of_LNil [simp, nitpick_simp]: "list_of LNil = []" using list_of_llist_of[of "[]"] by simp lemma list_of_LCons [simp]: "lfinite xs \ list_of (LCons x xs) = x # list_of xs" proof(induct arbitrary: x rule: lfinite.induct) case lfinite_LNil show ?case using list_of_llist_of[of "[x]"] by simp next case (lfinite_LConsI xs' x') from \list_of (LCons x' xs') = x' # list_of xs'\ show ?case using list_of_llist_of[of "x # x' # list_of xs'"] llist_of_list_of[OF \lfinite xs'\] by simp qed lemma list_of_LCons_conv [nitpick_simp]: "list_of (LCons x xs) = (if lfinite xs then x # list_of xs else undefined)" by(auto)(auto simp add: list_of_def) lemma list_of_lappend: assumes "lfinite xs" "lfinite ys" shows "list_of (lappend xs ys) = list_of xs @ list_of ys" using \lfinite xs\ by induct(simp_all add: \lfinite ys\) lemma list_of_lmap [simp]: assumes "lfinite xs" shows "list_of (lmap f xs) = map f (list_of xs)" using assms by induct simp_all lemma set_list_of [simp]: assumes "lfinite xs" shows "set (list_of xs) = lset xs" using assms by(induct)(simp_all) lemma hd_list_of [simp]: "lfinite xs \ hd (list_of xs) = lhd xs" by(clarsimp simp add: lfinite_eq_range_llist_of) lemma tl_list_of: "lfinite xs \ tl (list_of xs) = list_of (ltl xs)" by(auto simp add: lfinite_eq_range_llist_of) text \Efficient implementation via tail recursion suggested by Brian Huffman\ definition list_of_aux :: "'a list \ 'a llist \ 'a list" where "list_of_aux xs ys = (if lfinite ys then rev xs @ list_of ys else undefined)" lemma list_of_code [code]: "list_of = list_of_aux []" by(simp add: fun_eq_iff list_of_def list_of_aux_def) lemma list_of_aux_code [code]: "list_of_aux xs LNil = rev xs" "list_of_aux xs (LCons y ys) = list_of_aux (y # xs) ys" by(simp_all add: list_of_aux_def) subsection \The length of a lazy list: @{term "llength"}\ lemma [simp, nitpick_simp]: shows llength_LNil: "llength LNil = 0" and llength_LCons: "llength (LCons x xs) = eSuc (llength xs)" by(simp_all add: llength_def enat_unfold) lemma llength_eq_0 [simp]: "llength xs = 0 \ lnull xs" by(cases xs) simp_all lemma llength_lnull [simp]: "lnull xs \ llength xs = 0" by simp lemma epred_llength: "epred (llength xs) = llength (ltl xs)" by(cases xs) simp_all lemmas llength_ltl = epred_llength[symmetric] lemma llength_lmap [simp]: "llength (lmap f xs) = llength xs" by(coinduction arbitrary: xs rule: enat_coinduct)(auto simp add: epred_llength) lemma llength_lappend [simp]: "llength (lappend xs ys) = llength xs + llength ys" by(coinduction arbitrary: xs ys rule: enat_coinduct)(simp_all add: iadd_is_0 epred_iadd1 split_paired_all epred_llength, auto) lemma llength_llist_of [simp]: "llength (llist_of xs) = enat (length xs)" by(induct xs)(simp_all add: zero_enat_def eSuc_def) lemma length_list_of: "lfinite xs \ enat (length (list_of xs)) = llength xs" apply(rule sym) by(induct rule: lfinite.induct)(auto simp add: eSuc_enat zero_enat_def) lemma length_list_of_conv_the_enat: "lfinite xs \ length (list_of xs) = the_enat (llength xs)" unfolding lfinite_eq_range_llist_of by auto lemma llength_eq_enat_lfiniteD: "llength xs = enat n \ lfinite xs" proof(induct n arbitrary: xs) case [folded zero_enat_def]: 0 thus ?case by simp next case (Suc n) note len = \llength xs = enat (Suc n)\ then obtain x xs' where "xs = LCons x xs'" by(cases xs)(auto simp add: zero_enat_def) moreover with len have "llength xs' = enat n" by(simp add: eSuc_def split: enat.split_asm) hence "lfinite xs'" by(rule Suc) ultimately show ?case by simp qed lemma lfinite_llength_enat: assumes "lfinite xs" shows "\n. llength xs = enat n" using assms by induct(auto simp add: eSuc_def zero_enat_def) lemma lfinite_conv_llength_enat: "lfinite xs \ (\n. llength xs = enat n)" by(blast dest: llength_eq_enat_lfiniteD lfinite_llength_enat) lemma not_lfinite_llength: "\ lfinite xs \ llength xs = \" by(simp add: lfinite_conv_llength_enat) lemma llength_eq_infty_conv_lfinite: "llength xs = \ \ \ lfinite xs" by(simp add: lfinite_conv_llength_enat) lemma lfinite_finite_index: "lfinite xs \ finite {n. enat n < llength xs}" by(auto dest: lfinite_llength_enat) text \tail-recursive implementation for @{term llength}\ definition gen_llength :: "nat \ 'a llist \ enat" where "gen_llength n xs = enat n + llength xs" lemma gen_llength_code [code]: "gen_llength n LNil = enat n" "gen_llength n (LCons x xs) = gen_llength (n + 1) xs" by(simp_all add: gen_llength_def iadd_Suc eSuc_enat[symmetric] iadd_Suc_right) lemma llength_code [code]: "llength = gen_llength 0" by(simp add: gen_llength_def fun_eq_iff zero_enat_def) lemma fixes F defines "F \ \llength xs. case xs of LNil \ 0 | LCons x xs \ eSuc (llength xs)" shows llength_conv_fixp: "llength \ ccpo.fixp (fun_lub Sup) (fun_ord (\)) F" (is "_ \ ?fixp") and llength_mono: "\xs. monotone (fun_ord (\)) (\) (\llength. F llength xs)" (is "PROP ?mono") proof(intro eq_reflection ext) show mono: "PROP ?mono" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix xs show "llength xs = ?fixp xs" by(coinduction arbitrary: xs rule: enat_coinduct)(subst (1 2 3) lfp.mono_body_fixp[OF mono], fastforce simp add: F_def split: llist.split del: iffI)+ qed lemma mono2mono_llength[THEN lfp.mono2mono, simp, cont_intro]: shows monotone_llength: "monotone (\) (\) llength" by(rule lfp.fixp_preserves_mono1[OF llength_mono llength_conv_fixp])(fold bot_enat_def, simp) lemma mcont2mcont_llength[THEN lfp.mcont2mcont, simp, cont_intro]: shows mcont_llength: "mcont lSup (\) Sup (\) llength" by(rule lfp.fixp_preserves_mcont1[OF llength_mono llength_conv_fixp])(fold bot_enat_def, simp) subsection \Taking and dropping from lazy lists: @{term "ltake"}, @{term "ldropn"}, and @{term "ldrop"}\ lemma ltake_LNil [simp, code, nitpick_simp]: "ltake n LNil = LNil" by simp lemma ltake_0 [simp]: "ltake 0 xs = LNil" by(simp add: ltake_def) lemma ltake_eSuc_LCons [simp]: "ltake (eSuc n) (LCons x xs) = LCons x (ltake n xs)" by(rule llist.expand)(simp_all) lemma ltake_eSuc: "ltake (eSuc n) xs = (case xs of LNil \ LNil | LCons x xs' \ LCons x (ltake n xs'))" by(cases xs) simp_all lemma lnull_ltake [simp]: "lnull (ltake n xs) \ lnull xs \ n = 0" by(cases n xs rule: enat_coexhaust[case_product llist.exhaust]) simp_all lemma ltake_eq_LNil_iff: "ltake n xs = LNil \ xs = LNil \ n = 0" by(cases n xs rule: enat_coexhaust[case_product llist.exhaust]) simp_all lemma LNil_eq_ltake_iff [simp]: "LNil = ltake n xs \ xs = LNil \ n = 0" by(cases n xs rule: enat_coexhaust[case_product llist.exhaust]) simp_all lemma ltake_LCons [code, nitpick_simp]: "ltake n (LCons x xs) = (case n of 0 \ LNil | eSuc n' \ LCons x (ltake n' xs))" by(rule llist.expand)(simp_all split: co.enat.split) lemma lhd_ltake [simp]: "n \ 0 \ lhd (ltake n xs) = lhd xs" by(cases n xs rule: enat_coexhaust[case_product llist.exhaust]) simp_all lemma ltl_ltake: "ltl (ltake n xs) = ltake (epred n) (ltl xs)" by(cases n xs rule: enat_coexhaust[case_product llist.exhaust]) simp_all lemmas ltake_epred_ltl = ltl_ltake [symmetric] declare ltake.sel(2) [simp del] lemma ltake_ltl: "ltake n (ltl xs) = ltl (ltake (eSuc n) xs)" by(cases xs) simp_all lemma llength_ltake [simp]: "llength (ltake n xs) = min n (llength xs)" by(coinduction arbitrary: n xs rule: enat_coinduct)(auto 4 3 simp add: enat_min_eq_0_iff epred_llength ltl_ltake) lemma ltake_lmap [simp]: "ltake n (lmap f xs) = lmap f (ltake n xs)" by(coinduction arbitrary: n xs)(auto 4 3 simp add: ltl_ltake) lemma ltake_ltake [simp]: "ltake n (ltake m xs) = ltake (min n m) xs" by(coinduction arbitrary: n m xs)(auto 4 4 simp add: enat_min_eq_0_iff ltl_ltake) lemma lset_ltake: "lset (ltake n xs) \ lset xs" proof(rule subsetI) fix x assume "x \ lset (ltake n xs)" thus "x \ lset xs" proof(induct "ltake n xs" arbitrary: xs n rule: lset_induct) case find thus ?case by(cases xs)(simp, cases n rule: enat_coexhaust, simp_all) next case step thus ?case by(cases xs)(simp, cases n rule: enat_coexhaust, simp_all) qed qed lemma ltake_all: "llength xs \ m \ ltake m xs = xs" by(coinduction arbitrary: m xs)(auto simp add: epred_llength[symmetric] ltl_ltake intro: epred_le_epredI) lemma ltake_llist_of [simp]: "ltake (enat n) (llist_of xs) = llist_of (take n xs)" proof(induct n arbitrary: xs) case 0 thus ?case unfolding zero_enat_def[symmetric] by(cases xs) simp_all next case (Suc n) thus ?case unfolding eSuc_enat[symmetric] by(cases xs) simp_all qed lemma lfinite_ltake [simp]: "lfinite (ltake n xs) \ lfinite xs \ n < \" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by(induct ys\"ltake n xs" arbitrary: n xs rule: lfinite_induct)(fastforce simp add: zero_enat_def ltl_ltake)+ next assume ?rhs (is "?xs \ ?n") thus ?lhs proof assume ?xs thus ?thesis by(induct xs arbitrary: n)(simp_all add: ltake_LCons split: enat_cosplit) next assume ?n then obtain n' where "n = enat n'" by auto moreover have "lfinite (ltake (enat n') xs)" by(induct n' arbitrary: xs) (auto simp add: zero_enat_def[symmetric] eSuc_enat[symmetric] ltake_eSuc split: llist.split) ultimately show ?thesis by simp qed qed lemma ltake_lappend1: "n \ llength xs \ ltake n (lappend xs ys) = ltake n xs" by(coinduction arbitrary: n xs)(auto intro!: exI simp add: llength_ltl epred_le_epredI ltl_ltake) lemma ltake_lappend2: "llength xs \ n \ ltake n (lappend xs ys) = lappend xs (ltake (n - llength xs) ys)" by(coinduction arbitrary: n xs rule: llist.coinduct_strong)(auto intro!: exI simp add: llength_ltl epred_le_epredI ltl_ltake) lemma ltake_lappend: "ltake n (lappend xs ys) = lappend (ltake n xs) (ltake (n - llength xs) ys)" by(coinduction arbitrary: n xs ys rule: llist.coinduct_strong)(auto intro!: exI simp add: llength_ltl ltl_ltake) lemma take_list_of: assumes "lfinite xs" shows "take n (list_of xs) = list_of (ltake (enat n) xs)" using assms by(induct arbitrary: n) (simp_all add: take_Cons zero_enat_def[symmetric] eSuc_enat[symmetric] split: nat.split) lemma ltake_eq_ltake_antimono: "\ ltake n xs = ltake n ys; m \ n \ \ ltake m xs = ltake m ys" by (metis ltake_ltake min_def) lemma ltake_is_lprefix [simp, intro]: "ltake n xs \ xs" by(coinduction arbitrary: xs n)(auto simp add: ltl_ltake) lemma lprefix_ltake_same [simp]: "ltake n xs \ ltake m xs \ n \ m \ llength xs \ m" (is "?lhs \ ?rhs") proof(rule iffI disjCI)+ assume ?lhs "\ llength xs \ m" thus "n \ m" proof(coinduction arbitrary: n m xs rule: enat_le_coinduct) case (le n m xs) thus ?case by(cases xs)(auto 4 3 simp add: ltake_LCons split: co.enat.splits) qed next assume ?rhs thus ?lhs proof assume "n \ m" thus ?thesis by(coinduction arbitrary: n m xs)(auto 4 4 simp add: ltl_ltake epred_le_epredI) qed(simp add: ltake_all) qed lemma fixes f F defines "F \ \ltake n xs. case xs of LNil \ LNil | LCons x xs \ case n of 0 \ LNil | eSuc n \ LCons x (ltake n xs)" shows ltake_conv_fixp: "ltake \ curry (ccpo.fixp (fun_lub lSup) (fun_ord (\)) (\ltake. case_prod (F (curry ltake))))" (is "?lhs \ ?rhs") and ltake_mono: "\nxs. mono_llist (\ltake. case nxs of (n, xs) \ F (curry ltake) n xs)" (is "PROP ?mono") proof(intro eq_reflection ext) show mono: "PROP ?mono" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix n xs show "?lhs n xs = ?rhs n xs" proof(coinduction arbitrary: n xs) case Eq_llist show ?case by(subst (1 3 4) llist.mono_body_fixp[OF mono])(auto simp add: F_def split: llist.split prod.split co.enat.split) qed qed lemma monotone_ltake: "monotone (rel_prod (\) (\)) (\) (case_prod ltake)" by(rule llist.fixp_preserves_mono2[OF ltake_mono ltake_conv_fixp]) simp lemma mono2mono_ltake1[THEN llist.mono2mono, cont_intro, simp]: shows monotone_ltake1: "monotone (\) (\) (\n. ltake n xs)" using monotone_ltake by auto lemma mono2mono_ltake2[THEN llist.mono2mono, cont_intro, simp]: shows monotone_ltake2: "monotone (\) (\) (ltake n)" using monotone_ltake by auto lemma mcont_ltake: "mcont (prod_lub Sup lSup) (rel_prod (\) (\)) lSup (\) (case_prod ltake)" by(rule llist.fixp_preserves_mcont2[OF ltake_mono ltake_conv_fixp]) simp lemma mcont2mcont_ltake1 [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_ltake1: "mcont Sup (\) lSup (\) (\n. ltake n xs)" using mcont_ltake by auto lemma mcont2mcont_ltake2 [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_ltake2: "mcont lSup (\) lSup (\) (ltake n)" using mcont_ltake by auto lemma [partial_function_mono]: "mono_llist F \ mono_llist (\f. ltake n (F f))" by(rule mono2mono_ltake2) lemma llist_induct2: assumes adm: "ccpo.admissible (prod_lub lSup lSup) (rel_prod (\) (\)) (\x. P (fst x) (snd x))" and LNil: "P LNil LNil" and LCons1: "\x xs. \ lfinite xs; P xs LNil \ \ P (LCons x xs) LNil" and LCons2: "\y ys. \ lfinite ys; P LNil ys \ \ P LNil (LCons y ys)" and LCons: "\x xs y ys. \ lfinite xs; lfinite ys; P xs ys \ \ P (LCons x xs) (LCons y ys)" shows "P xs ys" proof - let ?C = "(\n. (ltake n xs, ltake n ys)) ` range enat" have "Complete_Partial_Order.chain (rel_prod (\) (\)) ?C" by(rule chainI) auto with adm have "P (fst (prod_lub lSup lSup ?C)) (snd (prod_lub lSup lSup ?C))" proof(rule ccpo.admissibleD) fix xsys assume "xsys \ ?C" then obtain n where "xsys = (ltake (enat n) xs, ltake (enat n) ys)" by auto moreover { fix xs :: "'a llist" assume "lfinite xs" hence "P xs LNil" by induct(auto intro: LNil LCons1) } note 1 = this { fix ys :: "'b llist" assume "lfinite ys" hence "P LNil ys" by induct(auto intro: LNil LCons2) } note 2 = this have "P (ltake (enat n) xs) (ltake (enat n) ys)" by(induct n arbitrary: xs ys)(auto simp add: zero_enat_def[symmetric] LNil eSuc_enat[symmetric] ltake_eSuc split: llist.split intro: LNil LCons 1 2) ultimately show "P (fst xsys) (snd xsys)" by simp qed simp also have "fst (prod_lub lSup lSup ?C) = xs" unfolding prod_lub_def fst_conv by(subst image_image)(simp add: mcont_contD[OF mcont_ltake1, symmetric] ltake_all) also have "snd (prod_lub lSup lSup ?C) = ys" unfolding prod_lub_def snd_conv by(subst image_image)(simp add: mcont_contD[OF mcont_ltake1, symmetric] ltake_all) finally show ?thesis . qed lemma ldropn_0 [simp]: "ldropn 0 xs = xs" by(simp add: ldropn_def) lemma ldropn_LNil [code, simp]: "ldropn n LNil = LNil" by(induct n)(simp_all add: ldropn_def) lemma ldropn_lnull: "lnull xs \ ldropn n xs = LNil" by(simp add: lnull_def) lemma ldropn_LCons [code]: "ldropn n (LCons x xs) = (case n of 0 \ LCons x xs | Suc n' \ ldropn n' xs)" by(cases n)(simp_all add: ldropn_def funpow_swap1) lemma ldropn_Suc: "ldropn (Suc n) xs = (case xs of LNil \ LNil | LCons x xs' \ ldropn n xs')" by(simp split: llist.split)(simp add: ldropn_def funpow_swap1) lemma ldropn_Suc_LCons [simp]: "ldropn (Suc n) (LCons x xs) = ldropn n xs" by(simp add: ldropn_LCons) lemma ltl_ldropn: "ltl (ldropn n xs) = ldropn n (ltl xs)" proof(induct n arbitrary: xs) case 0 thus ?case by simp next case (Suc n) thus ?case by(cases xs)(simp_all add: ldropn_Suc split: llist.split) qed lemma ldrop_simps [simp]: shows ldrop_LNil: "ldrop n LNil = LNil" and ldrop_0: "ldrop 0 xs = xs" and ldrop_eSuc_LCons: "ldrop (eSuc n) (LCons x xs) = ldrop n xs" by(simp_all add: ldrop.simps split: co.enat.split) lemma ldrop_lnull: "lnull xs \ ldrop n xs = LNil" by(simp add: lnull_def) lemma fixes f F defines "F \ \ldropn xs. case xs of LNil \ \_. LNil | LCons x xs \ \n. if n = 0 then LCons x xs else ldropn xs (n - 1)" shows ldrop_conv_fixp: "(\xs n. ldropn n xs) \ ccpo.fixp (fun_lub (fun_lub lSup)) (fun_ord (fun_ord lprefix)) (\ldrop. F ldrop)" (is "?lhs \ ?rhs") and ldrop_mono: "\xs. mono_llist_lift (\ldrop. F ldrop xs)" (is "PROP ?mono") proof(intro eq_reflection ext) show mono: "PROP ?mono" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix n xs show "?lhs xs n = ?rhs xs n" by(induction n arbitrary: xs) (subst llist_lift.mono_body_fixp[OF mono], simp add: F_def split: llist.split)+ qed lemma ldropn_fixp_case_conv: "(\xs. case xs of LNil \ \_. LNil | LCons x xs \ \n. if n = 0 then LCons x xs else f xs (n - 1)) = (\xs n. case xs of LNil \ LNil | LCons x xs \ if n = 0 then LCons x xs else f xs (n - 1))" by(auto simp add: fun_eq_iff split: llist.split) lemma monotone_ldropn_aux: "monotone lprefix (fun_ord lprefix) (\xs n. ldropn n xs)" by(rule llist_lift.fixp_preserves_mono1[OF ldrop_mono ldrop_conv_fixp]) (simp add: ldropn_fixp_case_conv monotone_fun_ord_apply) lemma mono2mono_ldropn[THEN llist.mono2mono, cont_intro, simp]: shows monotone_ldropn': "monotone lprefix lprefix (\xs. ldropn n xs)" using monotone_ldropn_aux by(auto simp add: monotone_def fun_ord_def) lemma mcont_ldropn_aux: "mcont lSup lprefix (fun_lub lSup) (fun_ord lprefix) (\xs n. ldropn n xs)" by(rule llist_lift.fixp_preserves_mcont1[OF ldrop_mono ldrop_conv_fixp]) (simp add: ldropn_fixp_case_conv mcont_fun_lub_apply) lemma mcont2mcont_ldropn [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_ldropn: "mcont lSup lprefix lSup lprefix (ldropn n)" using mcont_ldropn_aux by(auto simp add: mcont_fun_lub_apply) lemma monotone_enat_cocase [cont_intro, simp]: "\ \n. monotone (\) ord (\n. f n (eSuc n)); \n. ord a (f n (eSuc n)); ord a a \ \ monotone (\) ord (\n. case n of 0 \ a | eSuc n' \ f n' n)" by(rule monotone_enat_le_case) lemma monotone_ldrop: "monotone (rel_prod (=) (\)) (\) (case_prod ldrop)" by(rule llist.fixp_preserves_mono2[OF ldrop.mono ldrop_def]) simp lemma mono2mono_ldrop2 [THEN llist.mono2mono, cont_intro, simp]: shows monotone_ldrop2: "monotone (\) (\) (ldrop n)" by(simp add: monotone_ldrop[simplified]) lemma mcont_ldrop: "mcont (prod_lub the_Sup lSup) (rel_prod (=) (\)) lSup (\) (case_prod ldrop)" by(rule llist.fixp_preserves_mcont2[OF ldrop.mono ldrop_def]) simp lemma mcont2monct_ldrop2 [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_ldrop2: "mcont lSup (\) lSup (\) (ldrop n)" by(simp add: mcont_ldrop[simplified]) lemma ldrop_eSuc_conv_ltl: "ldrop (eSuc n) xs = ltl (ldrop n xs)" proof(induct xs arbitrary: n) case LCons thus ?case by(cases n rule: co.enat.exhaust) simp_all qed simp_all lemma ldrop_ltl: "ldrop n (ltl xs) = ldrop (eSuc n) xs" proof(induct xs arbitrary: n) case LCons thus ?case by(cases n rule: co.enat.exhaust) simp_all qed simp_all lemma lnull_ldropn [simp]: "lnull (ldropn n xs) \ llength xs \ enat n" proof(induction n arbitrary: xs) case (Suc n) from Suc.IH[of "ltl xs"] show ?case by(cases xs)(simp_all add: eSuc_enat[symmetric]) qed(simp add: zero_enat_def[symmetric]) lemma ldrop_eq_LNil [simp]: "ldrop n xs = LNil \ llength xs \ n" proof(induction xs arbitrary: n) case (LCons x xs n) thus ?case by(cases n rule: co.enat.exhaust) simp_all qed simp_all lemma lnull_ldrop [simp]: "lnull (ldrop n xs) \ llength xs \ n" unfolding lnull_def by(fact ldrop_eq_LNil) lemma ldropn_eq_LNil: "(ldropn n xs = LNil) = (llength xs \ enat n)" using lnull_ldropn unfolding lnull_def . lemma ldropn_all: "llength xs \ enat m \ ldropn m xs = LNil" by(simp add: ldropn_eq_LNil) lemma ldrop_all: "llength xs \ m \ ldrop m xs = LNil" by(simp) lemma ltl_ldrop: "ltl (ldrop n xs) = ldrop n (ltl xs)" by(simp add: ldrop_eSuc_conv_ltl ldrop_ltl) lemma ldrop_eSuc: "ldrop (eSuc n) xs = (case xs of LNil \ LNil | LCons x xs' \ ldrop n xs')" by(cases xs) simp_all lemma ldrop_LCons: "ldrop n (LCons x xs) = (case n of 0 \ LCons x xs | eSuc n' \ ldrop n' xs)" by(cases n rule: enat_coexhaust) simp_all lemma ldrop_inf [code, simp]: "ldrop \ xs = LNil" by(induction xs)(simp_all add: ldrop_LCons enat_cocase_inf) lemma ldrop_enat [code]: "ldrop (enat n) xs = ldropn n xs" proof(induct n arbitrary: xs) case Suc thus ?case by(cases xs)(simp_all add: eSuc_enat[symmetric]) qed(simp add: zero_enat_def[symmetric]) lemma lfinite_ldropn [simp]: "lfinite (ldropn n xs) = lfinite xs" by(induct n arbitrary: xs)(simp_all add: ldropn_Suc split: llist.split) lemma lfinite_ldrop [simp]: "lfinite (ldrop n xs) \ lfinite xs \ n = \" by(cases n)(simp_all add: ldrop_enat) lemma ldropn_ltl: "ldropn n (ltl xs) = ldropn (Suc n) xs" by(simp add: ldropn_def funpow_swap1) lemmas ldrop_eSuc_ltl = ldropn_ltl[symmetric] lemma lset_ldropn_subset: "lset (ldropn n xs) \ lset xs" by(induct n arbitrary: xs)(fastforce simp add: ldropn_Suc split: llist.split_asm)+ lemma in_lset_ldropnD: "x \ lset (ldropn n xs) \ x \ lset xs" using lset_ldropn_subset[of n xs] by auto lemma lset_ldrop_subset: "lset (ldrop n xs) \ lset xs" proof(induct xs arbitrary: n) case LCons thus ?case by(cases n rule: co.enat.exhaust) auto qed simp_all lemma in_lset_ldropD: "x \ lset (ldrop n xs) \ x \ lset xs" using lset_ldrop_subset[of n xs] by(auto) lemma lappend_ltake_ldrop: "lappend (ltake n xs) (ldrop n xs) = xs" by(coinduction arbitrary: n xs rule: llist.coinduct_strong) (auto simp add: ldrop_ltl ltl_ltake intro!: arg_cong2[where f=lappend]) lemma ldropn_lappend: "ldropn n (lappend xs ys) = (if enat n < llength xs then lappend (ldropn n xs) ys else ldropn (n - the_enat (llength xs)) ys)" proof(induct n arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric] lappend_lnull1) next case (Suc n) { fix zs assume "llength zs \ enat n" hence "the_enat (eSuc (llength zs)) = Suc (the_enat (llength zs))" by(auto intro!: the_enat_eSuc iff del: not_infinity_eq) } note eq = this from Suc show ?case by(cases xs)(auto simp add: not_less not_le eSuc_enat[symmetric] eq) qed lemma ldropn_lappend2: "llength xs \ enat n \ ldropn n (lappend xs ys) = ldropn (n - the_enat (llength xs)) ys" by(auto simp add: ldropn_lappend) lemma lappend_ltake_enat_ldropn [simp]: "lappend (ltake (enat n) xs) (ldropn n xs) = xs" by(fold ldrop_enat)(rule lappend_ltake_ldrop) lemma ldrop_lappend: "ldrop n (lappend xs ys) = (if n < llength xs then lappend (ldrop n xs) ys else ldrop (n - llength xs) ys)" \ \cannot prove this directly using fixpoint induction, because @{term "(-) :: enat \ enat \ enat"} is not a least fixpoint\ by(cases n)(cases "llength xs", simp_all add: ldropn_lappend not_less ldrop_enat) lemma ltake_plus_conv_lappend: "ltake (n + m) xs = lappend (ltake n xs) (ltake m (ldrop n xs))" by(coinduction arbitrary: n m xs rule: llist.coinduct_strong)(auto intro!: exI simp add: iadd_is_0 ltl_ltake epred_iadd1 ldrop_ltl) lemma ldropn_eq_LConsD: "ldropn n xs = LCons y ys \ enat n < llength xs" proof(induct n arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric]) next case (Suc n) thus ?case by(cases xs)(simp_all add: Suc_ile_eq) qed lemma ldrop_eq_LConsD: "ldrop n xs = LCons y ys \ n < llength xs" by(rule ccontr)(simp add: not_less ldrop_all) lemma ldropn_lmap [simp]: "ldropn n (lmap f xs) = lmap f (ldropn n xs)" by(induct n arbitrary: xs)(simp_all add: ldropn_Suc split: llist.split) lemma ldrop_lmap [simp]: "ldrop n (lmap f xs) = lmap f (ldrop n xs)" proof(induct xs arbitrary: n) case LCons thus ?case by(cases n rule: co.enat.exhaust) simp_all qed simp_all lemma ldropn_ldropn [simp]: "ldropn n (ldropn m xs) = ldropn (n + m) xs" by(induct m arbitrary: xs)(auto simp add: ldropn_Suc split: llist.split) lemma ldrop_ldrop [simp]: "ldrop n (ldrop m xs) = ldrop (n + m) xs" proof(induct xs arbitrary: m) case LCons thus ?case by(cases m rule: co.enat.exhaust)(simp_all add: iadd_Suc_right) qed simp_all lemma llength_ldropn [simp]: "llength (ldropn n xs) = llength xs - enat n" proof(induct n arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric]) next case (Suc n) thus ?case by(cases xs)(simp_all add: eSuc_enat[symmetric]) qed lemma enat_llength_ldropn: "enat n \ llength xs \ enat (n - m) \ llength (ldropn m xs)" by(cases "llength xs") simp_all lemma ldropn_llist_of [simp]: "ldropn n (llist_of xs) = llist_of (drop n xs)" proof(induct n arbitrary: xs) case Suc thus ?case by(cases xs) simp_all qed simp lemma ldrop_llist_of: "ldrop (enat n) (llist_of xs) = llist_of (drop n xs)" proof(induct xs arbitrary: n) case Cons thus ?case by(cases n)(simp_all add: zero_enat_def[symmetric] eSuc_enat[symmetric]) qed simp lemma drop_list_of: "lfinite xs \ drop n (list_of xs) = list_of (ldropn n xs)" by (metis ldropn_llist_of list_of_llist_of llist_of_list_of) lemma llength_ldrop: "llength (ldrop n xs) = (if n = \ then 0 else llength xs - n)" proof(induct xs arbitrary: n) case (LCons x xs) thus ?case by simp(cases n rule: co.enat.exhaust, simp_all) qed simp_all lemma ltake_ldropn: "ltake n (ldropn m xs) = ldropn m (ltake (n + enat m) xs)" by(induct m arbitrary: n xs)(auto simp add: zero_enat_def[symmetric] ldropn_Suc eSuc_enat[symmetric] iadd_Suc_right split: llist.split) lemma ldropn_ltake: "ldropn n (ltake m xs) = ltake (m - enat n) (ldropn n xs)" by(induct n arbitrary: m xs)(auto simp add: zero_enat_def[symmetric] ltake_LCons ldropn_Suc eSuc_enat[symmetric] iadd_Suc_right split: llist.split co.enat.split_asm) lemma ltake_ldrop: "ltake n (ldrop m xs) = ldrop m (ltake (n + m) xs)" by(induct xs arbitrary: n m)(simp_all add: ldrop_LCons iadd_Suc_right split: co.enat.split) lemma ldrop_ltake: "ldrop n (ltake m xs) = ltake (m - n) (ldrop n xs)" by(induct xs arbitrary: n m)(simp_all add: ltake_LCons ldrop_LCons split: co.enat.split) subsection \Taking the $n$-th element of a lazy list: @{term "lnth" }\ lemma lnth_LNil: "lnth LNil n = undefined n" by(cases n)(simp_all add: lnth.simps) lemma lnth_0 [simp]: "lnth (LCons x xs) 0 = x" by(simp add: lnth.simps) lemma lnth_Suc_LCons [simp]: "lnth (LCons x xs) (Suc n) = lnth xs n" by(simp add: lnth.simps) lemma lnth_LCons: "lnth (LCons x xs) n = (case n of 0 \ x | Suc n' \ lnth xs n')" by(cases n) simp_all lemma lnth_LCons': "lnth (LCons x xs) n = (if n = 0 then x else lnth xs (n - 1))" by(simp add: lnth_LCons split: nat.split) lemma lhd_conv_lnth: "\ lnull xs \ lhd xs = lnth xs 0" by(auto simp add: lhd_def not_lnull_conv) lemmas lnth_0_conv_lhd = lhd_conv_lnth[symmetric] lemma lnth_ltl: "\ lnull xs \ lnth (ltl xs) n = lnth xs (Suc n)" by(auto simp add: not_lnull_conv) lemma lhd_ldropn: "enat n < llength xs \ lhd (ldropn n xs) = lnth xs n" proof(induct n arbitrary: xs) case 0 thus ?case by(cases xs) auto next case (Suc n) from \enat (Suc n) < llength xs\ obtain x xs' where [simp]: "xs = LCons x xs'" by(cases xs) auto from \enat (Suc n) < llength xs\ have "enat n < llength xs'" by(simp add: Suc_ile_eq) hence "lhd (ldropn n xs') = lnth xs' n" by(rule Suc) thus ?case by simp qed lemma lhd_ldrop: assumes "n < llength xs" shows "lhd (ldrop n xs) = lnth xs (the_enat n)" proof - from assms obtain n' where n: "n = enat n'" by(cases n) auto from assms show ?thesis unfolding n proof(induction n' arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric] lhd_conv_lnth) next case (Suc n') thus ?case by(cases xs)(simp_all add: eSuc_enat[symmetric], simp add: eSuc_enat) qed qed lemma lnth_beyond: "llength xs \ enat n \ lnth xs n = undefined (n - (case llength xs of enat m \ m))" proof(induct n arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric] lnth_def lnull_def) next case Suc thus ?case by(cases xs)(simp_all add: zero_enat_def lnth_def eSuc_enat[symmetric] split: enat.split, auto simp add: eSuc_enat) qed lemma lnth_lmap [simp]: "enat n < llength xs \ lnth (lmap f xs) n = f (lnth xs n)" proof(induct n arbitrary: xs) case 0 thus ?case by(cases xs) simp_all next case (Suc n) from \enat (Suc n) < llength xs\ obtain x xs' where xs: "xs = LCons x xs'" and len: "enat n < llength xs'" by(cases xs)(auto simp add: Suc_ile_eq) from len have "lnth (lmap f xs') n = f (lnth xs' n)" by(rule Suc) with xs show ?case by simp qed lemma lnth_ldropn [simp]: "enat (n + m) < llength xs \ lnth (ldropn n xs) m = lnth xs (m + n)" proof(induct n arbitrary: xs) case (Suc n) from \enat (Suc n + m) < llength xs\ obtain x xs' where "xs = LCons x xs'" by(cases xs) auto moreover with \enat (Suc n + m) < llength xs\ have "enat (n + m) < llength xs'" by(simp add: Suc_ile_eq) hence "lnth (ldropn n xs') m = lnth xs' (m + n)" by(rule Suc) ultimately show ?case by simp qed simp lemma lnth_ldrop [simp]: "n + enat m < llength xs \ lnth (ldrop n xs) m = lnth xs (m + the_enat n)" by(cases n)(simp_all add: ldrop_enat) lemma in_lset_conv_lnth: "x \ lset xs \ (\n. enat n < llength xs \ lnth xs n = x)" (is "?lhs \ ?rhs") proof assume ?rhs then obtain n where "enat n < llength xs" "lnth xs n = x" by blast thus ?lhs proof(induct n arbitrary: xs) case 0 thus ?case by(auto simp add: zero_enat_def[symmetric] not_lnull_conv) next case (Suc n) thus ?case by(cases xs)(auto simp add: eSuc_enat[symmetric]) qed next assume ?lhs thus ?rhs proof(induct) case (find xs) show ?case by(auto intro: exI[where x=0] simp add: zero_enat_def[symmetric]) next case (step x' xs) thus ?case by(auto 4 4 intro: exI[where x="Suc n" for n] ileI1 simp add: eSuc_enat[symmetric]) qed qed lemma lset_conv_lnth: "lset xs = {lnth xs n|n. enat n < llength xs}" by(auto simp add: in_lset_conv_lnth) lemma lnth_llist_of [simp]: "lnth (llist_of xs) = nth xs" proof(rule ext) fix n show "lnth (llist_of xs) n = xs ! n" proof(induct xs arbitrary: n) case Nil thus ?case by(cases n)(simp_all add: nth_def lnth_def) next case Cons thus ?case by(simp add: lnth_LCons split: nat.split) qed qed lemma nth_list_of [simp]: assumes "lfinite xs" shows "nth (list_of xs) = lnth xs" using assms by induct(auto intro: simp add: nth_def lnth_LNil nth_Cons split: nat.split) lemma lnth_lappend1: "enat n < llength xs \ lnth (lappend xs ys) n = lnth xs n" proof(induct n arbitrary: xs) case 0 thus ?case by(auto simp add: zero_enat_def[symmetric] not_lnull_conv) next case (Suc n) from \enat (Suc n) < llength xs\ obtain x xs' where [simp]: "xs = LCons x xs'" and len: "enat n < llength xs'" by(cases xs)(auto simp add: Suc_ile_eq) from len have "lnth (lappend xs' ys) n = lnth xs' n" by(rule Suc) thus ?case by simp qed lemma lnth_lappend_llist_of: "lnth (lappend (llist_of xs) ys) n = (if n < length xs then xs ! n else lnth ys (n - length xs))" proof(induct xs arbitrary: n) case (Cons x xs) thus ?case by(cases n) auto qed simp lemma lnth_lappend2: "\ llength xs = enat k; k \ n \ \ lnth (lappend xs ys) n = lnth ys (n - k)" proof(induct n arbitrary: xs k) case 0 thus ?case by(simp add: zero_enat_def[symmetric] lappend_lnull1) next case (Suc n) thus ?case by(cases xs)(auto simp add: eSuc_def zero_enat_def split: enat.split_asm) qed lemma lnth_lappend: "lnth (lappend xs ys) n = (if enat n < llength xs then lnth xs n else lnth ys (n - the_enat (llength xs)))" by(cases "llength xs")(auto simp add: lnth_lappend1 lnth_lappend2) lemma lnth_ltake: "enat m < n \ lnth (ltake n xs) m = lnth xs m" proof(induct m arbitrary: xs n) case 0 thus ?case by(cases n rule: enat_coexhaust)(auto, cases xs, auto) next case (Suc m) from \enat (Suc m) < n\ obtain n' where "n = eSuc n'" by(cases n rule: enat_coexhaust) auto with \enat (Suc m) < n\ have "enat m < n'" by(simp add: eSuc_enat[symmetric]) with Suc \n = eSuc n'\ show ?case by(cases xs) auto qed lemma ldropn_Suc_conv_ldropn: "enat n < llength xs \ LCons (lnth xs n) (ldropn (Suc n) xs) = ldropn n xs" proof(induct n arbitrary: xs) case 0 thus ?case by(cases xs) auto next case (Suc n) from \enat (Suc n) < llength xs\ obtain x xs' where [simp]: "xs = LCons x xs'" by(cases xs) auto from \enat (Suc n) < llength xs\ have "enat n < llength xs'" by(simp add: Suc_ile_eq) hence "LCons (lnth xs' n) (ldropn (Suc n) xs') = ldropn n xs'" by(rule Suc) thus ?case by simp qed lemma ltake_Suc_conv_snoc_lnth: "enat m < llength xs \ ltake (enat (Suc m)) xs = lappend (ltake (enat m) xs) (LCons (lnth xs m) LNil)" by(metis eSuc_enat eSuc_plus_1 ltake_plus_conv_lappend ldrop_enat ldropn_Suc_conv_ldropn ltake_0 ltake_eSuc_LCons one_eSuc) lemma lappend_eq_lappend_conv: assumes len: "llength xs = llength us" shows "lappend xs ys = lappend us vs \ xs = us \ (lfinite xs \ ys = vs)" (is "?lhs \ ?rhs") proof assume ?rhs thus ?lhs by(auto simp add: lappend_inf) next assume eq: ?lhs show ?rhs proof(intro conjI impI) show "xs = us" using len eq proof(coinduction arbitrary: xs us) case (Eq_llist xs us) thus ?case by(cases xs us rule: llist.exhaust[case_product llist.exhaust]) auto qed assume "lfinite xs" then obtain xs' where "xs = llist_of xs'" by(auto simp add: lfinite_eq_range_llist_of) hence "lappend (llist_of xs') ys = lappend (llist_of xs') vs" using eq \xs = us\ by blast thus "ys = vs" by (induct xs') simp_all qed qed subsection\iterates\ lemmas iterates [code, nitpick_simp] = iterates.ctr and lnull_iterates = iterates.simps(1) and lhd_iterates = iterates.simps(2) and ltl_iterates = iterates.simps(3) lemma lfinite_iterates [iff]: "\ lfinite (iterates f x)" proof assume "lfinite (iterates f x)" thus False by(induct zs\"iterates f x" arbitrary: x rule: lfinite_induct) auto qed lemma lmap_iterates: "lmap f (iterates f x) = iterates f (f x)" by(coinduction arbitrary: x) auto lemma iterates_lmap: "iterates f x = LCons x (lmap f (iterates f x))" by(simp add: lmap_iterates)(rule iterates) lemma lappend_iterates: "lappend (iterates f x) xs = iterates f x" by(coinduction arbitrary: x) auto lemma [simp]: fixes f :: "'a \ 'a" shows lnull_funpow_lmap: "lnull ((lmap f ^^ n) xs) \ lnull xs" and lhd_funpow_lmap: "\ lnull xs \ lhd ((lmap f ^^ n) xs) = (f ^^ n) (lhd xs)" and ltl_funpow_lmap: "\ lnull xs \ ltl ((lmap f ^^ n) xs) = (lmap f ^^ n) (ltl xs)" by(induct n) simp_all lemma iterates_equality: assumes h: "\x. h x = LCons x (lmap f (h x))" shows "h = iterates f" proof - { fix x have "\ lnull (h x)" "lhd (h x) = x" "ltl (h x) = lmap f (h x)" by(subst h, simp)+ } note [simp] = this { fix x define n :: nat where "n = 0" have "(lmap f ^^ n) (h x) = (lmap f ^^ n) (iterates f x)" by(coinduction arbitrary: n)(auto simp add: funpow_swap1 lmap_iterates intro: exI[where x="Suc n" for n]) } thus ?thesis by auto qed lemma llength_iterates [simp]: "llength (iterates f x) = \" by(coinduction arbitrary: x rule: enat_coinduct)(auto simp add: epred_llength) lemma ldropn_iterates: "ldropn n (iterates f x) = iterates f ((f ^^ n) x)" proof(induct n arbitrary: x) case 0 thus ?case by simp next case (Suc n) have "ldropn (Suc n) (iterates f x) = ldropn n (iterates f (f x))" by(subst iterates)simp also have "\ = iterates f ((f ^^ n) (f x))" by(rule Suc) finally show ?case by(simp add: funpow_swap1) qed lemma ldrop_iterates: "ldrop (enat n) (iterates f x) = iterates f ((f ^^ n) x)" proof(induct n arbitrary: x) case Suc thus ?case by(subst iterates)(simp add: eSuc_enat[symmetric] funpow_swap1) qed(simp add: zero_enat_def[symmetric]) lemma lnth_iterates [simp]: "lnth (iterates f x) n = (f ^^ n) x" proof(induct n arbitrary: x) case 0 thus ?case by(subst iterates) simp next case (Suc n) hence "lnth (iterates f (f x)) n = (f ^^ n) (f x)" . thus ?case by(subst iterates)(simp add: funpow_swap1) qed lemma lset_iterates: "lset (iterates f x) = {(f ^^ n) x|n. True}" by(auto simp add: lset_conv_lnth) lemma lset_repeat [simp]: "lset (repeat x) = {x}" by(simp add: lset_iterates id_def[symmetric]) subsection \ More on the prefix ordering on lazy lists: @{term "lprefix"} and @{term "lstrict_prefix"} \ lemma lstrict_prefix_code [code, simp]: "lstrict_prefix LNil LNil \ False" "lstrict_prefix LNil (LCons y ys) \ True" "lstrict_prefix (LCons x xs) LNil \ False" "lstrict_prefix (LCons x xs) (LCons y ys) \ x = y \ lstrict_prefix xs ys" by(auto simp add: lstrict_prefix_def) lemma lmap_lprefix: "xs \ ys \ lmap f xs \ lmap f ys" by(rule monotoneD[OF monotone_lmap]) lemma lprefix_llength_eq_imp_eq: "\ xs \ ys; llength xs = llength ys \ \ xs = ys" by(coinduction arbitrary: xs ys)(auto simp add: not_lnull_conv) lemma lprefix_llength_le: "xs \ ys \ llength xs \ llength ys" using monotone_llength by(rule monotoneD) lemma lstrict_prefix_llength_less: assumes "lstrict_prefix xs ys" shows "llength xs < llength ys" proof(rule ccontr) assume "\ llength xs < llength ys" moreover from \lstrict_prefix xs ys\ have "xs \ ys" "xs \ ys" unfolding lstrict_prefix_def by simp_all from \xs \ ys\ have "llength xs \ llength ys" by(rule lprefix_llength_le) ultimately have "llength xs = llength ys" by auto with \xs \ ys\ have "xs = ys" by(rule lprefix_llength_eq_imp_eq) with \xs \ ys\ show False by contradiction qed lemma lstrict_prefix_lfinite1: "lstrict_prefix xs ys \ lfinite xs" by (metis lstrict_prefix_def not_lfinite_lprefix_conv_eq) lemma wfP_lstrict_prefix: "wfP lstrict_prefix" proof(unfold wfP_def) have "wf {(x :: enat, y). x < y}" unfolding wf_def by(blast intro: less_induct) hence "wf (inv_image {(x, y). x < y} llength)" by(rule wf_inv_image) moreover have "{(xs, ys). lstrict_prefix xs ys} \ inv_image {(x, y). x < y} llength" by(auto intro: lstrict_prefix_llength_less) ultimately show "wf {(xs, ys). lstrict_prefix xs ys}" by(rule wf_subset) qed lemma llist_less_induct[case_names less]: "(\xs. (\ys. lstrict_prefix ys xs \ P ys) \ P xs) \ P xs" by(rule wfP_induct[OF wfP_lstrict_prefix]) blast lemma ltake_enat_eq_imp_eq: "(\n. ltake (enat n) xs = ltake (enat n) ys) \ xs = ys" by(coinduction arbitrary: xs ys)(auto simp add: zero_enat_def lnull_def neq_LNil_conv ltake_eq_LNil_iff eSuc_enat[symmetric] elim: allE[where x="Suc n" for n]) lemma ltake_enat_lprefix_imp_lprefix: assumes "\n. lprefix (ltake (enat n) xs) (ltake (enat n) ys)" shows "lprefix xs ys" proof - have "ccpo.admissible Sup (\) (\n. ltake n xs \ ltake n ys)" by simp hence "ltake (Sup (range enat)) xs \ ltake (Sup (range enat)) ys" by(rule ccpo.admissibleD)(auto intro: assms) thus ?thesis by(simp add: ltake_all) qed lemma lprefix_conv_lappend: "xs \ ys \ (\zs. ys = lappend xs zs)" (is "?lhs \ ?rhs") proof(rule iffI) assume ?lhs hence "ys = lappend xs (ldrop (llength xs) ys)" by(coinduction arbitrary: xs ys)(auto dest: lprefix_lnullD lprefix_lhdD intro: lprefix_ltlI simp add: not_lnull_conv lprefix_LCons_conv intro: exI[where x=LNil]) thus ?rhs .. next assume ?rhs thus ?lhs by(coinduct rule: lprefix_coinduct) auto qed lemma lappend_lprefixE: assumes "lappend xs ys \ zs" obtains zs' where "zs = lappend xs zs'" using assms unfolding lprefix_conv_lappend by(auto simp add: lappend_assoc) lemma lprefix_lfiniteD: "\ xs \ ys; lfinite ys \ \ lfinite xs" unfolding lprefix_conv_lappend by auto lemma lprefix_lappendD: assumes "xs \ lappend ys zs" shows "xs \ ys \ ys \ xs" proof(rule ccontr) assume "\ (xs \ ys \ ys \ xs)" hence "\ xs \ ys" "\ ys \ xs" by simp_all from \xs \ lappend ys zs\ obtain xs' where "lappend xs xs' = lappend ys zs" unfolding lprefix_conv_lappend by auto hence eq: "lappend (ltake (llength ys) xs) (lappend (ldrop (llength ys) xs) xs') = lappend ys zs" unfolding lappend_assoc[symmetric] by(simp only: lappend_ltake_ldrop) moreover have "llength xs \ llength ys" proof(rule ccontr) assume "\ ?thesis" hence "llength xs < llength ys" by simp hence "ltake (llength ys) xs = xs" by(simp add: ltake_all) hence "lappend xs (lappend (ldrop (llength ys) xs) xs') = lappend (ltake (llength xs) ys) (lappend (ldrop (llength xs) ys) zs)" unfolding lappend_assoc[symmetric] lappend_ltake_ldrop using eq by(simp add: lappend_assoc) hence xs: "xs = ltake (llength xs) ys" using \llength xs < llength ys\ by(subst (asm) lappend_eq_lappend_conv)(auto simp add: min_def) have "xs \ ys" by(subst xs) auto with \\ xs \ ys\ show False by contradiction qed ultimately have ys: "ys = ltake (llength ys) xs" by(subst (asm) lappend_eq_lappend_conv)(simp_all add: min_def) have "ys \ xs" by(subst ys) auto with \\ ys \ xs\ show False by contradiction qed lemma lstrict_prefix_lappend_conv: "lstrict_prefix xs (lappend xs ys) \ lfinite xs \ \ lnull ys" proof - { assume "lfinite xs" "xs = lappend xs ys" hence "lnull ys" by induct auto } thus ?thesis by(auto simp add: lstrict_prefix_def lprefix_lappend lappend_inf lappend_lnull2 elim: contrapos_np) qed lemma lprefix_llist_ofI: "\zs. ys = xs @ zs \ llist_of xs \ llist_of ys" by(clarsimp simp add: lappend_llist_of_llist_of[symmetric] lprefix_lappend) lemma lprefix_llist_of [simp]: "llist_of xs \ llist_of ys \ prefix xs ys" by(auto simp add: prefix_def lprefix_conv_lappend)(metis lfinite_lappend lfinite_llist_of list_of_lappend list_of_llist_of lappend_llist_of_llist_of)+ lemma llimit_induct [case_names LNil LCons limit]: \ \The limit case is just an instance of admissibility\ assumes LNil: "P LNil" and LCons: "\x xs. \ lfinite xs; P xs \ \ P (LCons x xs)" and limit: "(\ys. lstrict_prefix ys xs \ P ys) \ P xs" shows "P xs" proof(rule limit) fix ys assume "lstrict_prefix ys xs" hence "lfinite ys" by(rule lstrict_prefix_lfinite1) thus "P ys" by(induct)(blast intro: LNil LCons)+ qed lemma lmap_lstrict_prefix: "lstrict_prefix xs ys \ lstrict_prefix (lmap f xs) (lmap f ys)" by (metis llength_lmap lmap_lprefix lprefix_llength_eq_imp_eq lstrict_prefix_def) lemma lprefix_lnthD: assumes "xs \ ys" and "enat n < llength xs" shows "lnth xs n = lnth ys n" using assms by (metis lnth_lappend1 lprefix_conv_lappend) lemma lfinite_lSup_chain: assumes chain: "Complete_Partial_Order.chain (\) A" shows "lfinite (lSup A) \ finite A \ (\xs \ A. lfinite xs)" (is "?lhs \ ?rhs") proof(intro iffI conjI) assume ?lhs then obtain n where n: "llength (lSup A) = enat n" unfolding lfinite_conv_llength_enat .. have "llength ` A \ {..xs\A. lfinite xs" by simp_all show ?lhs proof(cases "A = {}") case False with chain \finite A\ have "lSup A \ A" by(rule ccpo.in_chain_finite[OF llist_ccpo]) with \\xs\A. lfinite xs\ show ?thesis .. qed simp qed(rule lfinite_lSupD) text \Setup for @{term "lprefix"} for Nitpick\ definition finite_lprefix :: "'a llist \ 'a llist \ bool" where "finite_lprefix = (\)" lemma finite_lprefix_nitpick_simps [nitpick_simp]: "finite_lprefix xs LNil \ xs = LNil" "finite_lprefix LNil xs \ True" "finite_lprefix xs (LCons y ys) \ xs = LNil \ (\xs'. xs = LCons y xs' \ finite_lprefix xs' ys)" by(simp_all add: lprefix_LCons_conv finite_lprefix_def lnull_def) lemma lprefix_nitpick_simps [nitpick_simp]: "xs \ ys = (if lfinite xs then finite_lprefix xs ys else xs = ys)" by(simp add: finite_lprefix_def not_lfinite_lprefix_conv_eq) hide_const (open) finite_lprefix hide_fact (open) finite_lprefix_def finite_lprefix_nitpick_simps lprefix_nitpick_simps subsection \Length of the longest common prefix\ lemma llcp_simps [simp, code, nitpick_simp]: shows llcp_LNil1: "llcp LNil ys = 0" and llcp_LNil2: "llcp xs LNil = 0" and llcp_LCons: "llcp (LCons x xs) (LCons y ys) = (if x = y then eSuc (llcp xs ys) else 0)" by(simp_all add: llcp_def enat_unfold split: llist.split) lemma llcp_eq_0_iff: "llcp xs ys = 0 \ lnull xs \ lnull ys \ lhd xs \ lhd ys" by(simp add: llcp_def) lemma epred_llcp: "\ \ lnull xs; \ lnull ys; lhd xs = lhd ys \ \ epred (llcp xs ys) = llcp (ltl xs) (ltl ys)" by(simp add: llcp_def) lemma llcp_commute: "llcp xs ys = llcp ys xs" by(coinduction arbitrary: xs ys rule: enat_coinduct)(auto simp add: llcp_eq_0_iff epred_llcp) lemma llcp_same_conv_length [simp]: "llcp xs xs = llength xs" by(coinduction arbitrary: xs rule: enat_coinduct)(auto simp add: llcp_eq_0_iff epred_llcp epred_llength) lemma llcp_lappend_same [simp]: "llcp (lappend xs ys) (lappend xs zs) = llength xs + llcp ys zs" by(coinduction arbitrary: xs rule: enat_coinduct)(auto simp add: iadd_is_0 llcp_eq_0_iff epred_iadd1 epred_llcp epred_llength) lemma llcp_lprefix1 [simp]: "xs \ ys \ llcp xs ys = llength xs" by (metis add_0_right lappend_LNil2 llcp_LNil1 llcp_lappend_same lprefix_conv_lappend) lemma llcp_lprefix2 [simp]: "ys \ xs \ llcp xs ys = llength ys" by (metis llcp_commute llcp_lprefix1) lemma llcp_le_length: "llcp xs ys \ min (llength xs) (llength ys)" proof - define m n where "m = llcp xs ys" and "n = min (llength xs) (llength ys)" hence "(m, n) \ {(llcp xs ys, min (llength xs) (llength ys)) |xs ys :: 'a llist. True}" by blast thus "m \ n" proof(coinduct rule: enat_leI) case (Leenat m n) then obtain xs ys :: "'a llist" where "m = llcp xs ys" "n = min (llength xs) (llength ys)" by blast thus ?case by(cases xs ys rule: llist.exhaust[case_product llist.exhaust])(auto 4 3 intro!: exI[where x="Suc 0"] simp add: eSuc_enat[symmetric] iadd_Suc_right zero_enat_def[symmetric]) qed qed lemma llcp_ltake1: "llcp (ltake n xs) ys = min n (llcp xs ys)" by(coinduction arbitrary: n xs ys rule: enat_coinduct)(auto simp add: llcp_eq_0_iff enat_min_eq_0_iff epred_llcp ltl_ltake) lemma llcp_ltake2: "llcp xs (ltake n ys) = min n (llcp xs ys)" by(metis llcp_commute llcp_ltake1) lemma llcp_ltake [simp]: "llcp (ltake n xs) (ltake m ys) = min (min n m) (llcp xs ys)" by(metis llcp_ltake1 llcp_ltake2 min.assoc) subsection \Zipping two lazy lists to a lazy list of pairs @{term "lzip" }\ lemma lzip_simps [simp, code, nitpick_simp]: "lzip LNil ys = LNil" "lzip xs LNil = LNil" "lzip (LCons x xs) (LCons y ys) = LCons (x, y) (lzip xs ys)" by(auto intro: llist.expand) lemma lnull_lzip [simp]: "lnull (lzip xs ys) \ lnull xs \ lnull ys" by(simp add: lzip_def) lemma lzip_eq_LNil_conv: "lzip xs ys = LNil \ xs = LNil \ ys = LNil" using lnull_lzip unfolding lnull_def . lemmas lhd_lzip = lzip.sel(1) and ltl_lzip = lzip.sel(2) lemma lzip_eq_LCons_conv: "lzip xs ys = LCons z zs \ (\x xs' y ys'. xs = LCons x xs' \ ys = LCons y ys' \ z = (x, y) \ zs = lzip xs' ys')" by(cases xs ys rule: llist.exhaust[case_product llist.exhaust]) auto lemma lzip_lappend: "llength xs = llength us \ lzip (lappend xs ys) (lappend us vs) = lappend (lzip xs us) (lzip ys vs)" by(coinduction arbitrary: xs ys us vs rule: llist.coinduct_strong)(auto 4 6 simp add: llength_ltl) lemma llength_lzip [simp]: "llength (lzip xs ys) = min (llength xs) (llength ys)" by(coinduction arbitrary: xs ys rule: enat_coinduct)(auto simp add: enat_min_eq_0_iff epred_llength) lemma ltake_lzip: "ltake n (lzip xs ys) = lzip (ltake n xs) (ltake n ys)" by(coinduction arbitrary: xs ys n)(auto simp add: ltl_ltake) lemma ldropn_lzip [simp]: "ldropn n (lzip xs ys) = lzip (ldropn n xs) (ldropn n ys)" by(induct n arbitrary: xs ys)(simp_all add: ldropn_Suc split: llist.split) lemma fixes F defines "F \ \lzip (xs, ys). case xs of LNil \ LNil | LCons x xs' \ case ys of LNil \ LNil | LCons y ys' \ LCons (x, y) (curry lzip xs' ys')" shows lzip_conv_fixp: "lzip \ curry (ccpo.fixp (fun_lub lSup) (fun_ord (\)) F)" (is "?lhs \ ?rhs") and lzip_mono: "mono_llist (\lzip. F lzip xs)" (is "?mono xs") proof(intro eq_reflection ext) show mono: "\xs. ?mono xs" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix xs ys show "lzip xs ys = ?rhs xs ys" proof(coinduction arbitrary: xs ys) case Eq_llist show ?case by(subst (1 3 4) llist.mono_body_fixp[OF mono])(auto simp add: F_def split: llist.split) qed qed lemma monotone_lzip: "monotone (rel_prod (\) (\)) (\) (case_prod lzip)" by(rule llist.fixp_preserves_mono2[OF lzip_mono lzip_conv_fixp]) simp lemma mono2mono_lzip1 [THEN llist.mono2mono, cont_intro, simp]: shows monotone_lzip1: "monotone (\) (\) (\xs. lzip xs ys)" by(simp add: monotone_lzip[simplified]) lemma mono2mono_lzip2 [THEN llist.mono2mono, cont_intro, simp]: shows monotone_lzip2: "monotone (\) (\) (\ys. lzip xs ys)" by(simp add: monotone_lzip[simplified]) lemma mcont_lzip: "mcont (prod_lub lSup lSup) (rel_prod (\) (\)) lSup (\) (case_prod lzip)" by(rule llist.fixp_preserves_mcont2[OF lzip_mono lzip_conv_fixp]) simp lemma mcont2mcont_lzip1 [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_lzip1: "mcont lSup (\) lSup (\) (\xs. lzip xs ys)" by(simp add: mcont_lzip[simplified]) lemma mcont2mcont_lzip2 [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_lzip2: "mcont lSup (\) lSup (\) (\ys. lzip xs ys)" by(simp add: mcont_lzip[simplified]) lemma ldrop_lzip [simp]: "ldrop n (lzip xs ys) = lzip (ldrop n xs) (ldrop n ys)" proof(induct xs arbitrary: ys n) case LCons thus ?case by(cases ys n rule: llist.exhaust[case_product co.enat.exhaust]) simp_all qed simp_all lemma lzip_iterates: "lzip (iterates f x) (iterates g y) = iterates (\(x, y). (f x, g y)) (x, y)" by(coinduction arbitrary: x y) auto lemma lzip_llist_of [simp]: "lzip (llist_of xs) (llist_of ys) = llist_of (zip xs ys)" proof(induct xs arbitrary: ys) case (Cons x xs') thus ?case by(cases ys) simp_all qed simp lemma lnth_lzip: "\ enat n < llength xs; enat n < llength ys \ \ lnth (lzip xs ys) n = (lnth xs n, lnth ys n)" proof(induct n arbitrary: xs ys) case 0 thus ?case by(simp add: zero_enat_def[symmetric] lnth_0_conv_lhd) next case (Suc n) thus ?case by(cases xs ys rule: llist.exhaust[case_product llist.exhaust])(auto simp add: eSuc_enat[symmetric]) qed lemma lset_lzip: "lset (lzip xs ys) = {(lnth xs n, lnth ys n)|n. enat n < min (llength xs) (llength ys)}" by(auto simp add: lset_conv_lnth lnth_lzip)(auto intro!: exI simp add: lnth_lzip) lemma lset_lzipD1: "(x, y) \ lset (lzip xs ys) \ x \ lset xs" proof(induct "lzip xs ys" arbitrary: xs ys rule: lset_induct) case [symmetric]: find thus ?case by(auto simp add: lzip_eq_LCons_conv) next case (step z zs) thus ?case by -(drule sym, auto simp add: lzip_eq_LCons_conv) qed lemma lset_lzipD2: "(x, y) \ lset (lzip xs ys) \ y \ lset ys" proof(induct "lzip xs ys" arbitrary: xs ys rule: lset_induct) case [symmetric]: find thus ?case by(auto simp add: lzip_eq_LCons_conv) next case (step z zs) thus ?case by -(drule sym, auto simp add: lzip_eq_LCons_conv) qed lemma lset_lzip_same [simp]: "lset (lzip xs xs) = (\x. (x, x)) ` lset xs" by(auto 4 3 simp add: lset_lzip in_lset_conv_lnth) lemma lfinite_lzip [simp]: "lfinite (lzip xs ys) \ lfinite xs \ lfinite ys" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by(induct zs\"lzip xs ys" arbitrary: xs ys rule: lfinite_induct) fastforce+ next assume ?rhs (is "?xs \ ?ys") thus ?lhs proof assume ?xs thus ?thesis proof(induct arbitrary: ys) case (lfinite_LConsI xs x) thus ?case by(cases ys) simp_all qed simp next assume ?ys thus ?thesis proof(induct arbitrary: xs) case (lfinite_LConsI ys y) thus ?case by(cases xs) simp_all qed simp qed qed lemma lzip_eq_lappend_conv: assumes eq: "lzip xs ys = lappend us vs" shows "\xs' xs'' ys' ys''. xs = lappend xs' xs'' \ ys = lappend ys' ys'' \ llength xs' = llength ys' \ us = lzip xs' ys' \ vs = lzip xs'' ys''" proof - let ?xs' = "ltake (llength us) xs" let ?xs'' = "ldrop (llength us) xs" let ?ys' = "ltake (llength us) ys" let ?ys'' = "ldrop (llength us) ys" from eq have "llength (lzip xs ys) = llength (lappend us vs)" by simp hence "min (llength xs) (llength ys) \ llength us" by(auto simp add: enat_le_plus_same) hence len: "llength xs \ llength us" "llength ys \ llength us" by(auto) hence leneq: "llength ?xs' = llength ?ys'" by(simp add: min_def) have xs: "xs = lappend ?xs' ?xs''" and ys: "ys = lappend ?ys' ?ys''" by(simp_all add: lappend_ltake_ldrop) hence "lappend us vs = lzip (lappend ?xs' ?xs'') (lappend ?ys' ?ys'')" using eq by simp with len have "lappend us vs = lappend (lzip ?xs' ?ys') (lzip ?xs'' ?ys'')" by(simp add: lzip_lappend min_def) hence us: "us = lzip ?xs' ?ys'" and vs: "lfinite us \ vs = lzip ?xs'' ?ys''" using len by(simp_all add: min_def lappend_eq_lappend_conv) show ?thesis proof(cases "lfinite us") case True with leneq xs ys us vs len show ?thesis by fastforce next case False let ?xs'' = "lmap fst vs" let ?ys'' = "lmap snd vs" from False have "lappend us vs = us" by(simp add: lappend_inf) moreover from False have "llength us = \" by(rule not_lfinite_llength) moreover with len have "llength xs = \" "llength ys = \" by auto moreover with \llength us = \\ have "xs = ?xs'" "ys = ?ys'" by(simp_all add: ltake_all) from \llength us = \\ len have "\ lfinite ?xs'" "\ lfinite ?ys'" by(auto simp del: llength_ltake lfinite_ltake simp add: ltake_all dest: lfinite_llength_enat) with \xs = ?xs'\ \ys = ?ys'\ have "xs = lappend ?xs' ?xs''" "ys = lappend ?ys' ?ys''" by(simp_all add: lappend_inf) moreover have "vs = lzip ?xs'' ?ys''" by(coinduction arbitrary: vs) auto ultimately show ?thesis using eq by(fastforce simp add: ltake_all) qed qed lemma lzip_lmap [simp]: "lzip (lmap f xs) (lmap g ys) = lmap (\(x, y). (f x, g y)) (lzip xs ys)" by(coinduction arbitrary: xs ys) auto lemma lzip_lmap1: "lzip (lmap f xs) ys = lmap (\(x, y). (f x, y)) (lzip xs ys)" by(subst (4) llist.map_ident[symmetric])(simp only: lzip_lmap) lemma lzip_lmap2: "lzip xs (lmap f ys) = lmap (\(x, y). (x, f y)) (lzip xs ys)" by(subst (1) llist.map_ident[symmetric])(simp only: lzip_lmap) lemma lmap_fst_lzip_conv_ltake: "lmap fst (lzip xs ys) = ltake (min (llength xs) (llength ys)) xs" by(coinduction arbitrary: xs ys)(auto simp add: enat_min_eq_0_iff ltl_ltake epred_llength) lemma lmap_snd_lzip_conv_ltake: "lmap snd (lzip xs ys) = ltake (min (llength xs) (llength ys)) ys" by(coinduction arbitrary: xs ys)(auto simp add: enat_min_eq_0_iff ltl_ltake epred_llength) lemma lzip_conv_lzip_ltake_min_llength: "lzip xs ys = lzip (ltake (min (llength xs) (llength ys)) xs) (ltake (min (llength xs) (llength ys)) ys)" by(coinduction arbitrary: xs ys)(auto simp add: enat_min_eq_0_iff ltl_ltake epred_llength) subsection \Taking and dropping from a lazy list: @{term "ltakeWhile"} and @{term "ldropWhile"}\ lemma ltakeWhile_simps [simp, code, nitpick_simp]: shows ltakeWhile_LNil: "ltakeWhile P LNil = LNil" and ltakeWhile_LCons: "ltakeWhile P (LCons x xs) = (if P x then LCons x (ltakeWhile P xs) else LNil)" by(auto simp add: ltakeWhile_def intro: llist.expand) lemma ldropWhile_simps [simp, code]: shows ldropWhile_LNil: "ldropWhile P LNil = LNil" and ldropWhile_LCons: "ldropWhile P (LCons x xs) = (if P x then ldropWhile P xs else LCons x xs)" by(simp_all add: ldropWhile.simps) lemma fixes f F P defines "F \ \ltakeWhile xs. case xs of LNil \ LNil | LCons x xs \ if P x then LCons x (ltakeWhile xs) else LNil" shows ltakeWhile_conv_fixp: "ltakeWhile P \ ccpo.fixp (fun_lub lSup) (fun_ord lprefix) F" (is "?lhs \ ?rhs") and ltakeWhile_mono: "\xs. mono_llist (\ltakeWhile. F ltakeWhile xs)" (is "PROP ?mono") proof(intro eq_reflection ext) show mono: "PROP ?mono" unfolding F_def by(tactic \Partial_Function.mono_tac @{context} 1\) fix xs show "?lhs xs = ?rhs xs" proof(coinduction arbitrary: xs) case Eq_llist show ?case by(subst (1 3 4) llist.mono_body_fixp[OF mono])(auto simp add: F_def split: llist.split prod.split co.enat.split) qed qed lemma mono2mono_ltakeWhile[THEN llist.mono2mono, cont_intro, simp]: shows monotone_ltakeWhile: "monotone lprefix lprefix (ltakeWhile P)" by(rule llist.fixp_preserves_mono1[OF ltakeWhile_mono ltakeWhile_conv_fixp]) simp lemma mcont2mcont_ltakeWhile [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_ltakeWhile: "mcont lSup lprefix lSup lprefix (ltakeWhile P)" by(rule llist.fixp_preserves_mcont1[OF ltakeWhile_mono ltakeWhile_conv_fixp]) simp lemma mono_llist_ltakeWhile [partial_function_mono]: "mono_llist F \ mono_llist (\f. ltakeWhile P (F f))" by(rule mono2mono_ltakeWhile) lemma mono2mono_ldropWhile [THEN llist.mono2mono, cont_intro, simp]: shows monotone_ldropWhile: "monotone (\) (\) (ldropWhile P)" by(rule llist.fixp_preserves_mono1[OF ldropWhile.mono ldropWhile_def]) simp lemma mcont2mcont_ldropWhile [THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_ldropWhile: "mcont lSup (\) lSup (\) (ldropWhile P)" by(rule llist.fixp_preserves_mcont1[OF ldropWhile.mono ldropWhile_def]) simp lemma lnull_ltakeWhile [simp]: "lnull (ltakeWhile P xs) \ (\ lnull xs \ \ P (lhd xs))" by(cases xs) simp_all lemma ltakeWhile_eq_LNil_iff: "ltakeWhile P xs = LNil \ (xs \ LNil \ \ P (lhd xs))" using lnull_ltakeWhile unfolding lnull_def . lemmas lhd_ltakeWhile = ltakeWhile.sel(1) lemma ltl_ltakeWhile: "ltl (ltakeWhile P xs) = (if P (lhd xs) then ltakeWhile P (ltl xs) else LNil)" by(cases xs) simp_all lemma lprefix_ltakeWhile: "ltakeWhile P xs \ xs" by(coinduction arbitrary: xs)(auto simp add: ltl_ltakeWhile) lemma llength_ltakeWhile_le: "llength (ltakeWhile P xs) \ llength xs" by(rule lprefix_llength_le)(rule lprefix_ltakeWhile) lemma ltakeWhile_nth: "enat i < llength (ltakeWhile P xs) \ lnth (ltakeWhile P xs) i = lnth xs i" by(rule lprefix_lnthD[OF lprefix_ltakeWhile]) lemma ltakeWhile_all: "\x\lset xs. P x \ ltakeWhile P xs = xs" by(coinduction arbitrary: xs)(auto 4 3 simp add: ltl_ltakeWhile simp del: ltakeWhile.disc_iff dest: in_lset_ltlD) lemma lset_ltakeWhileD: assumes "x \ lset (ltakeWhile P xs)" shows "x \ lset xs \ P x" using assms by(induct ys\"ltakeWhile P xs" arbitrary: xs rule: llist_set_induct)(auto simp add: ltl_ltakeWhile dest: in_lset_ltlD) lemma lset_ltakeWhile_subset: "lset (ltakeWhile P xs) \ lset xs \ {x. P x}" by(auto dest: lset_ltakeWhileD) lemma ltakeWhile_all_conv: "ltakeWhile P xs = xs \ lset xs \ {x. P x}" by (metis Int_Collect Int_absorb2 le_infE lset_ltakeWhile_subset ltakeWhile_all) lemma llength_ltakeWhile_all: "llength (ltakeWhile P xs) = llength xs \ ltakeWhile P xs = xs" by(auto intro: lprefix_llength_eq_imp_eq lprefix_ltakeWhile) lemma ldropWhile_eq_LNil_iff: "ldropWhile P xs = LNil \ (\x \ lset xs. P x)" by(induct xs) simp_all lemma lnull_ldropWhile [simp]: "lnull (ldropWhile P xs) \ (\x \ lset xs. P x)" (is "?lhs \ ?rhs") unfolding lnull_def by(simp add: ldropWhile_eq_LNil_iff) lemma lset_ldropWhile_subset: "lset (ldropWhile P xs) \ lset xs" by(induct xs) auto lemma in_lset_ldropWhileD: "x \ lset (ldropWhile P xs) \ x \ lset xs" using lset_ldropWhile_subset[of P xs] by auto lemma ltakeWhile_lmap: "ltakeWhile P (lmap f xs) = lmap f (ltakeWhile (P \ f) xs)" by(coinduction arbitrary: xs)(auto simp add: ltl_ltakeWhile) lemma ldropWhile_lmap: "ldropWhile P (lmap f xs) = lmap f (ldropWhile (P \ f) xs)" by(induct xs) simp_all lemma llength_ltakeWhile_lt_iff: "llength (ltakeWhile P xs) < llength xs \ (\x\lset xs. \ P x)" (is "?lhs \ ?rhs") proof assume ?lhs hence "ltakeWhile P xs \ xs" by auto thus ?rhs by(auto simp add: ltakeWhile_all_conv) next assume ?rhs hence "ltakeWhile P xs \ xs" by(auto simp add: ltakeWhile_all_conv) thus ?lhs unfolding llength_ltakeWhile_all[symmetric] using llength_ltakeWhile_le[of P xs] by(auto) qed lemma ltakeWhile_K_False [simp]: "ltakeWhile (\_. False) xs = LNil" by(simp add: ltakeWhile_def) lemma ltakeWhile_K_True [simp]: "ltakeWhile (\_. True) xs = xs" by(coinduction arbitrary: xs) simp lemma ldropWhile_K_False [simp]: "ldropWhile (\_. False) = id" proof fix xs show "ldropWhile (\_. False) xs = id xs" by(induct xs) simp_all qed lemma ldropWhile_K_True [simp]: "ldropWhile (\_. True) xs = LNil" by(induct xs)(simp_all) lemma lappend_ltakeWhile_ldropWhile [simp]: "lappend (ltakeWhile P xs) (ldropWhile P xs) = xs" by(coinduction arbitrary: xs rule: llist.coinduct_strong)(auto 4 4 simp add: not_lnull_conv lset_lnull intro: ccontr) lemma ltakeWhile_lappend: "ltakeWhile P (lappend xs ys) = (if \x\lset xs. \ P x then ltakeWhile P xs else lappend xs (ltakeWhile P ys))" proof(coinduction arbitrary: xs rule: llist.coinduct_strong) case (Eq_llist xs) have ?lnull by(auto simp add: lset_lnull) moreover have ?LCons by(clarsimp split: if_split_asm split del: if_split simp add: ltl_ltakeWhile)(auto 4 3 simp add: not_lnull_conv) ultimately show ?case .. qed lemma ldropWhile_lappend: "ldropWhile P (lappend xs ys) = (if \x\lset xs. \ P x then lappend (ldropWhile P xs) ys else if lfinite xs then ldropWhile P ys else LNil)" proof(cases "\x\lset xs. \ P x") case True then obtain x where "x \ lset xs" "\ P x" by blast thus ?thesis by induct simp_all next case False note xs = this show ?thesis proof(cases "lfinite xs") case False thus ?thesis using xs by(simp add: lappend_inf) next case True thus ?thesis using xs by induct simp_all qed qed lemma lfinite_ltakeWhile: "lfinite (ltakeWhile P xs) \ lfinite xs \ (\x \ lset xs. \ P x)" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by(auto simp add: ltakeWhile_all) next assume "?rhs" thus ?lhs proof assume "lfinite xs" with lprefix_ltakeWhile show ?thesis by(rule lprefix_lfiniteD) next assume "\x\lset xs. \ P x" then obtain x where "x \ lset xs" "\ P x" by blast thus ?thesis by(induct rule: lset_induct) simp_all qed qed lemma llength_ltakeWhile_eq_infinity: "llength (ltakeWhile P xs) = \ \ \ lfinite xs \ ltakeWhile P xs = xs" unfolding llength_ltakeWhile_all[symmetric] llength_eq_infty_conv_lfinite[symmetric] by(auto)(auto simp add: llength_eq_infty_conv_lfinite lfinite_ltakeWhile intro: sym) lemma llength_ltakeWhile_eq_infinity': "llength (ltakeWhile P xs) = \ \ \ lfinite xs \ (\x\lset xs. P x)" by (metis lfinite_ltakeWhile llength_eq_infty_conv_lfinite) lemma lzip_ltakeWhile_fst: "lzip (ltakeWhile P xs) ys = ltakeWhile (P \ fst) (lzip xs ys)" by(coinduction arbitrary: xs ys)(auto simp add: ltl_ltakeWhile simp del: simp del: ltakeWhile.disc_iff) lemma lzip_ltakeWhile_snd: "lzip xs (ltakeWhile P ys) = ltakeWhile (P \ snd) (lzip xs ys)" by(coinduction arbitrary: xs ys)(auto simp add: ltl_ltakeWhile) lemma ltakeWhile_lappend1: "\ x \ lset xs; \ P x \ \ ltakeWhile P (lappend xs ys) = ltakeWhile P xs" by(induct rule: lset_induct) simp_all lemma ltakeWhile_lappend2: "lset xs \ {x. P x} \ ltakeWhile P (lappend xs ys) = lappend xs (ltakeWhile P ys)" by(coinduction arbitrary: xs ys rule: llist.coinduct_strong)(auto 4 4 simp add: not_lnull_conv lappend_lnull1) lemma ltakeWhile_cong [cong, fundef_cong]: assumes xs: "xs = ys" and PQ: "\x. x \ lset ys \ P x = Q x" shows "ltakeWhile P xs = ltakeWhile Q ys" using PQ unfolding xs by(coinduction arbitrary: ys)(auto simp add: ltl_ltakeWhile not_lnull_conv) lemma lnth_llength_ltakeWhile: assumes len: "llength (ltakeWhile P xs) < llength xs" shows "\ P (lnth xs (the_enat (llength (ltakeWhile P xs))))" proof assume P: "P (lnth xs (the_enat (llength (ltakeWhile P xs))))" from len obtain len where "llength (ltakeWhile P xs) = enat len" by(cases "llength (ltakeWhile P xs)") auto with P len show False apply simp proof(induct len arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric] lnth_0_conv_lhd) next case (Suc len) thus ?case by(cases xs)(auto split: if_split_asm simp add: eSuc_enat[symmetric]) qed qed lemma assumes "\x\lset xs. \ P x" shows lhd_ldropWhile: "\ P (lhd (ldropWhile P xs))" (is ?thesis1) and lhd_ldropWhile_in_lset: "lhd (ldropWhile P xs) \ lset xs" (is ?thesis2) proof - from assms obtain x where "x \ lset xs" "\ P x" by blast thus ?thesis1 ?thesis2 by induct simp_all qed lemma ldropWhile_eq_ldrop: "ldropWhile P xs = ldrop (llength (ltakeWhile P xs)) xs" (is "?lhs = ?rhs") proof(rule lprefix_antisym) show "?lhs \ ?rhs" by(induct arbitrary: xs rule: ldropWhile.fixp_induct)(auto split: llist.split) show "?rhs \ ?lhs" proof(induct arbitrary: xs rule: ldrop.fixp_induct) case (3 ldrop xs) thus ?case by(cases xs) auto qed simp_all qed lemma ldropWhile_cong [cong]: "\ xs = ys; \x. x \ lset ys \ P x = Q x \ \ ldropWhile P xs = ldropWhile Q ys" by(simp add: ldropWhile_eq_ldrop) lemma ltakeWhile_repeat: "ltakeWhile P (repeat x) = (if P x then repeat x else LNil)" by(coinduction arbitrary: x)(auto simp add: ltl_ltakeWhile) lemma ldropWhile_repeat: "ldropWhile P (repeat x) = (if P x then LNil else repeat x)" by(simp add: ldropWhile_eq_ldrop ltakeWhile_repeat) lemma lfinite_ldropWhile: "lfinite (ldropWhile P xs) \ (\x \ lset xs. \ P x) \ lfinite xs" by(auto simp add: ldropWhile_eq_ldrop llength_eq_infty_conv_lfinite lfinite_ltakeWhile) lemma llength_ldropWhile: "llength (ldropWhile P xs) = (if \x\lset xs. \ P x then llength xs - llength (ltakeWhile P xs) else 0)" by(auto simp add: ldropWhile_eq_ldrop llength_ldrop llength_ltakeWhile_all ltakeWhile_all_conv llength_ltakeWhile_eq_infinity zero_enat_def dest!: lfinite_llength_enat) lemma lhd_ldropWhile_conv_lnth: "\x\lset xs. \ P x \ lhd (ldropWhile P xs) = lnth xs (the_enat (llength (ltakeWhile P xs)))" by(simp add: ldropWhile_eq_ldrop lhd_ldrop llength_ltakeWhile_lt_iff) subsection \@{term "llist_all2"}\ lemmas llist_all2_LNil_LNil = llist.rel_inject(1) lemmas llist_all2_LNil_LCons = llist.rel_distinct(1) lemmas llist_all2_LCons_LNil = llist.rel_distinct(2) lemmas llist_all2_LCons_LCons = llist.rel_inject(2) lemma llist_all2_LNil1 [simp]: "llist_all2 P LNil xs \ xs = LNil" by(cases xs) simp_all lemma llist_all2_LNil2 [simp]: "llist_all2 P xs LNil \ xs = LNil" by(cases xs) simp_all lemma llist_all2_LCons1: "llist_all2 P (LCons x xs) ys \ (\y ys'. ys = LCons y ys' \ P x y \ llist_all2 P xs ys')" by(cases ys) simp_all lemma llist_all2_LCons2: "llist_all2 P xs (LCons y ys) \ (\x xs'. xs = LCons x xs' \ P x y \ llist_all2 P xs' ys)" by(cases xs) simp_all lemma llist_all2_llist_of [simp]: "llist_all2 P (llist_of xs) (llist_of ys) = list_all2 P xs ys" by(induct xs ys rule: list_induct2')(simp_all) lemma llist_all2_conv_lzip: "llist_all2 P xs ys \ llength xs = llength ys \ (\(x, y) \ lset (lzip xs ys). P x y)" by(auto 4 4 elim!: GrpE simp add: llist_all2_def lmap_fst_lzip_conv_ltake lmap_snd_lzip_conv_ltake ltake_all intro!: GrpI relcomppI[of _ xs _ _ ys]) lemma llist_all2_llengthD: "llist_all2 P xs ys \ llength xs = llength ys" by(simp add: llist_all2_conv_lzip) lemma llist_all2_lnullD: "llist_all2 P xs ys \ lnull xs \ lnull ys" unfolding lnull_def by auto lemma llist_all2_all_lnthI: "\ llength xs = llength ys; \n. enat n < llength xs \ P (lnth xs n) (lnth ys n) \ \ llist_all2 P xs ys" by(auto simp add: llist_all2_conv_lzip lset_lzip) lemma llist_all2_lnthD: "\ llist_all2 P xs ys; enat n < llength xs \ \ P (lnth xs n) (lnth ys n)" by(fastforce simp add: llist_all2_conv_lzip lset_lzip) lemma llist_all2_lnthD2: "\ llist_all2 P xs ys; enat n < llength ys \ \ P (lnth xs n) (lnth ys n)" by(fastforce simp add: llist_all2_conv_lzip lset_lzip) lemma llist_all2_conv_all_lnth: "llist_all2 P xs ys \ llength xs = llength ys \ (\n. enat n < llength ys \ P (lnth xs n) (lnth ys n))" by(auto dest: llist_all2_llengthD llist_all2_lnthD2 intro: llist_all2_all_lnthI) lemma llist_all2_True [simp]: "llist_all2 (\_ _. True) xs ys \ llength xs = llength ys" by(simp add: llist_all2_conv_all_lnth) lemma llist_all2_reflI: "(\x. x \ lset xs \ P x x) \ llist_all2 P xs xs" by(auto simp add: llist_all2_conv_all_lnth lset_conv_lnth) lemma llist_all2_lmap1: "llist_all2 P (lmap f xs) ys \ llist_all2 (\x. P (f x)) xs ys" by(auto simp add: llist_all2_conv_all_lnth) lemma llist_all2_lmap2: "llist_all2 P xs (lmap g ys) \ llist_all2 (\x y. P x (g y)) xs ys" by(auto simp add: llist_all2_conv_all_lnth) lemma llist_all2_lfiniteD: "llist_all2 P xs ys \ lfinite xs \ lfinite ys" by(drule llist_all2_llengthD)(simp add: lfinite_conv_llength_enat) lemma llist_all2_coinduct[consumes 1, case_names LNil LCons, case_conclusion LCons lhd ltl, coinduct pred]: assumes major: "X xs ys" and step: "\xs ys. X xs ys \ lnull xs \ lnull ys" "\xs ys. \ X xs ys; \ lnull xs; \ lnull ys \ \ P (lhd xs) (lhd ys) \ (X (ltl xs) (ltl ys) \ llist_all2 P (ltl xs) (ltl ys))" shows "llist_all2 P xs ys" proof(rule llist_all2_all_lnthI) from major show "llength xs = llength ys" by(coinduction arbitrary: xs ys rule: enat_coinduct)(auto 4 3 dest: step llist_all2_llengthD simp add: epred_llength) fix n assume "enat n < llength xs" thus "P (lnth xs n) (lnth ys n)" using major \llength xs = llength ys\ proof(induct n arbitrary: xs ys) case 0 thus ?case by(cases "lnull xs")(auto dest: step simp add: zero_enat_def[symmetric] lnth_0_conv_lhd) next case (Suc n) from step[OF \X xs ys\] \enat (Suc n) < llength xs\ Suc show ?case by(auto 4 3 simp add: not_lnull_conv Suc_ile_eq intro: Suc.hyps llist_all2_lnthD dest: llist_all2_llengthD) qed qed lemma llist_all2_cases[consumes 1, case_names LNil LCons, cases pred]: assumes "llist_all2 P xs ys" obtains (LNil) "xs = LNil" "ys = LNil" | (LCons) x xs' y ys' where "xs = LCons x xs'" and "ys = LCons y ys'" and "P x y" and "llist_all2 P xs' ys'" using assms by(cases xs)(auto simp add: llist_all2_LCons1) lemma llist_all2_mono: "\ llist_all2 P xs ys; \x y. P x y \ P' x y \ \ llist_all2 P' xs ys" by(auto simp add: llist_all2_conv_all_lnth) lemma llist_all2_left: "llist_all2 (\x _. P x) xs ys \ llength xs = llength ys \ (\x\lset xs. P x)" by(fastforce simp add: llist_all2_conv_all_lnth lset_conv_lnth) lemma llist_all2_right: "llist_all2 (\_. P) xs ys \ llength xs = llength ys \ (\x\lset ys. P x)" by(fastforce simp add: llist_all2_conv_all_lnth lset_conv_lnth) lemma llist_all2_lsetD1: "\ llist_all2 P xs ys; x \ lset xs \ \ \y\lset ys. P x y" by(auto 4 4 simp add: llist_all2_conv_lzip lset_lzip lset_conv_lnth split_beta lnth_lzip simp del: split_paired_All) lemma llist_all2_lsetD2: "\ llist_all2 P xs ys; y \ lset ys \ \ \x\lset xs. P x y" by(auto 4 4 simp add: llist_all2_conv_lzip lset_lzip lset_conv_lnth split_beta lnth_lzip simp del: split_paired_All) lemma llist_all2_conj: "llist_all2 (\x y. P x y \ Q x y) xs ys \ llist_all2 P xs ys \ llist_all2 Q xs ys" by(auto simp add: llist_all2_conv_all_lnth) lemma llist_all2_lhdD: "\ llist_all2 P xs ys; \ lnull xs \ \ P (lhd xs) (lhd ys)" by(auto simp add: not_lnull_conv llist_all2_LCons1) lemma llist_all2_lhdD2: "\ llist_all2 P xs ys; \ lnull ys \ \ P (lhd xs) (lhd ys)" by(auto simp add: not_lnull_conv llist_all2_LCons2) lemma llist_all2_ltlI: "llist_all2 P xs ys \ llist_all2 P (ltl xs) (ltl ys)" by(cases xs)(auto simp add: llist_all2_LCons1) lemma llist_all2_lappendI: assumes 1: "llist_all2 P xs ys" and 2: "\ lfinite xs; lfinite ys \ \ llist_all2 P xs' ys'" shows "llist_all2 P (lappend xs xs') (lappend ys ys')" proof(cases "lfinite xs") case True with 1 have "lfinite ys" by(auto dest: llist_all2_lfiniteD) from 1 2[OF True this] show ?thesis by(coinduction arbitrary: xs ys)(auto dest: llist_all2_lnullD llist_all2_lhdD intro: llist_all2_ltlI simp add: lappend_eq_LNil_iff) next case False with 1 have "\ lfinite ys" by(auto dest: llist_all2_lfiniteD) with False 1 show ?thesis by(simp add: lappend_inf) qed lemma llist_all2_lappend1D: assumes "llist_all2 P (lappend xs xs') ys" shows "llist_all2 P xs (ltake (llength xs) ys)" and "lfinite xs \ llist_all2 P xs' (ldrop (llength xs) ys)" proof - from assms have len: "llength xs + llength xs' = llength ys" by(auto dest: llist_all2_llengthD) hence len_xs: "llength xs \ llength ys" and len_xs': "llength xs' \ llength ys" by (metis enat_le_plus_same llength_lappend)+ show "llist_all2 P xs (ltake (llength xs) ys)" proof(rule llist_all2_all_lnthI) show "llength xs = llength (ltake (llength xs) ys)" using len_xs by(simp add: min_def) next fix n assume n: "enat n < llength xs" also have "\ \ llength (lappend xs xs')" by(simp add: enat_le_plus_same) finally have "P (lnth (lappend xs xs') n) (lnth ys n)" using assms by -(rule llist_all2_lnthD) also from n have "lnth ys n = lnth (ltake (llength xs) ys) n" by(rule lnth_ltake[symmetric]) also from n have "lnth (lappend xs xs') n = lnth xs n" by(simp add: lnth_lappend1) finally show "P (lnth xs n) (lnth (ltake (llength xs) ys) n)" . qed assume "lfinite xs" thus "llist_all2 P xs' (ldrop (llength xs) ys)" using assms by(induct arbitrary: ys)(auto simp add: llist_all2_LCons1) qed lemma lmap_eq_lmap_conv_llist_all2: "lmap f xs = lmap g ys \ llist_all2 (\x y. f x = g y) xs ys" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by(coinduction arbitrary: xs ys)(auto simp add: neq_LNil_conv lnull_def LNil_eq_lmap lmap_eq_LNil) next assume ?rhs thus ?lhs by(coinduction arbitrary: xs ys)(auto dest: llist_all2_lnullD llist_all2_lhdD llist_all2_ltlI) qed lemma llist_all2_expand: "\ lnull xs \ lnull ys; \ \ lnull xs; \ lnull ys \ \ P (lhd xs) (lhd ys) \ llist_all2 P (ltl xs) (ltl ys) \ \ llist_all2 P xs ys" by(cases xs)(auto simp add: not_lnull_conv) lemma llist_all2_llength_ltakeWhileD: assumes major: "llist_all2 P xs ys" and Q: "\x y. P x y \ Q1 x \ Q2 y" shows "llength (ltakeWhile Q1 xs) = llength (ltakeWhile Q2 ys)" using major by(coinduction arbitrary: xs ys rule: enat_coinduct)(auto 4 3 simp add: not_lnull_conv llist_all2_LCons1 llist_all2_LCons2 dest!: Q) lemma llist_all2_lzipI: "\ llist_all2 P xs ys; llist_all2 P' xs' ys' \ \ llist_all2 (rel_prod P P') (lzip xs xs') (lzip ys ys')" by(coinduction arbitrary: xs xs' ys ys')(auto 6 6 dest: llist_all2_lhdD llist_all2_lnullD intro: llist_all2_ltlI) lemma llist_all2_ltakeI: "llist_all2 P xs ys \ llist_all2 P (ltake n xs) (ltake n ys)" by(auto simp add: llist_all2_conv_all_lnth lnth_ltake) lemma llist_all2_ldropnI: "llist_all2 P xs ys \ llist_all2 P (ldropn n xs) (ldropn n ys)" by(cases "llength ys")(auto simp add: llist_all2_conv_all_lnth) lemma llist_all2_ldropI: "llist_all2 P xs ys \ llist_all2 P (ldrop n xs) (ldrop n ys)" by(cases "llength ys")(auto simp add: llist_all2_conv_all_lnth llength_ldrop) lemma llist_all2_lSupI: assumes "Complete_Partial_Order.chain (rel_prod (\) (\)) Y" "\(xs, ys)\Y. llist_all2 P xs ys" shows "llist_all2 P (lSup (fst ` Y)) (lSup (snd ` Y))" using assms proof(coinduction arbitrary: Y) case LNil thus ?case by(auto dest: llist_all2_lnullD simp add: split_beta) next case (LCons Y) note chain = \Complete_Partial_Order.chain _ Y\ from LCons have Y: "\xs ys. (xs, ys) \ Y \ llist_all2 P xs ys" by blast from LCons obtain xs ys where xsysY: "(xs, ys) \ Y" and [simp]: "\ lnull xs" "\ lnull ys" by(auto 4 3 dest: llist_all2_lnullD simp add: split_beta) from xsysY have "lhd xs \ lhd ` (fst ` Y \ {xs. \ lnull xs})" by(auto intro: rev_image_eqI) hence "(THE x. x \ lhd ` (fst ` Y \ {xs. \ lnull xs})) = lhd xs" by(rule the_equality)(auto dest!: lprefix_lhdD chainD[OF chain xsysY]) moreover from xsysY have "lhd ys \ lhd ` (snd ` Y \ {xs. \ lnull xs})" by(auto intro: rev_image_eqI) hence "(THE x. x \ lhd ` (snd ` Y \ {xs. \ lnull xs})) = lhd ys" by(rule the_equality)(auto dest!: lprefix_lhdD chainD[OF chain xsysY]) moreover from xsysY have "llist_all2 P xs ys" by(rule Y) hence "P (lhd xs) (lhd ys)" by(rule llist_all2_lhdD) simp ultimately have ?lhd using LCons by simp moreover { let ?Y = "map_prod ltl ltl ` (Y \ {(xs, ys). \ lnull xs \ \ lnull ys})" have "Complete_Partial_Order.chain (rel_prod (\) (\)) ?Y" by(rule chainI)(auto 4 3 dest: Y chainD[OF chain] intro: lprefix_ltlI) moreover have "ltl ` (fst ` Y \ {xs. \ lnull xs}) = fst ` ?Y" and "ltl ` (snd ` Y \ {xs. \ lnull xs}) = snd ` ?Y" by(fastforce simp add: image_image dest: Y llist_all2_lnullD intro: rev_image_eqI)+ ultimately have ?ltl by(auto 4 3 intro: llist_all2_ltlI dest: Y) } ultimately show ?case .. qed lemma admissible_llist_all2 [cont_intro, simp]: assumes f: "mcont lub ord lSup (\) (\x. f x)" and g: "mcont lub ord lSup (\) (\x. g x)" shows "ccpo.admissible lub ord (\x. llist_all2 P (f x) (g x))" proof(rule ccpo.admissibleI) fix Y assume chain: "Complete_Partial_Order.chain ord Y" and Y: "\x\Y. llist_all2 P (f x) (g x)" and "Y \ {}" from chain have "Complete_Partial_Order.chain (rel_prod (\) (\)) ((\x. (f x, g x)) ` Y)" by(rule chain_imageI)(auto intro: mcont_monoD[OF f] mcont_monoD[OF g]) from llist_all2_lSupI[OF this, of P] chain Y show "llist_all2 P (f (lub Y)) (g (lub Y))" using \Y \ {}\ by(simp add: mcont_contD[OF f chain] mcont_contD[OF g chain] image_image) qed lemmas [cont_intro] = ccpo.mcont2mcont[OF llist_ccpo _ mcont_fst] ccpo.mcont2mcont[OF llist_ccpo _ mcont_snd] lemmas ldropWhile_fixp_parallel_induct = parallel_fixp_induct_1_1[OF llist_partial_function_definitions llist_partial_function_definitions ldropWhile.mono ldropWhile.mono ldropWhile_def ldropWhile_def, case_names adm LNil step] lemma llist_all2_ldropWhileI: assumes *: "llist_all2 P xs ys" and Q: "\x y. P x y \ Q1 x \ Q2 y" shows "llist_all2 P (ldropWhile Q1 xs) (ldropWhile Q2 ys)" \ \cannot prove this with parallel induction over @{term xs} and @{term ys} because @{term "\x. \ llist_all2 P (f x) (g x)"} is not admissible.\ using * by(induction arbitrary: xs ys rule: ldropWhile_fixp_parallel_induct)(auto split: llist.split dest: Q) lemma llist_all2_same [simp]: "llist_all2 P xs xs \ (\x\lset xs. P x x)" by(auto simp add: llist_all2_conv_all_lnth in_lset_conv_lnth Ball_def) lemma llist_all2_trans: "\ llist_all2 P xs ys; llist_all2 P ys zs; transp P \ \ llist_all2 P xs zs" apply(rule llist_all2_all_lnthI) apply(simp add: llist_all2_llengthD) apply(frule llist_all2_llengthD) apply(drule (1) llist_all2_lnthD) apply(drule llist_all2_lnthD) apply simp apply(erule (2) transpD) done subsection \The last element @{term "llast"}\ lemma llast_LNil: "llast LNil = undefined" by(simp add: llast_def zero_enat_def) lemma llast_LCons: "llast (LCons x xs) = (if lnull xs then x else llast xs)" by(cases "llength xs")(auto simp add: llast_def eSuc_def zero_enat_def not_lnull_conv split: enat.splits) lemma llast_linfinite: "\ lfinite xs \ llast xs = undefined" by(simp add: llast_def lfinite_conv_llength_enat) lemma [simp, code]: shows llast_singleton: "llast (LCons x LNil) = x" and llast_LCons2: "llast (LCons x (LCons y xs)) = llast (LCons y xs)" by(simp_all add: llast_LCons) lemma llast_lappend: "llast (lappend xs ys) = (if lnull ys then llast xs else if lfinite xs then llast ys else undefined)" proof(cases "lfinite xs") case True hence "\ lnull ys \ llast (lappend xs ys) = llast ys" by(induct rule: lfinite.induct)(simp_all add: llast_LCons) with True show ?thesis by(simp add: lappend_lnull2) next case False thus ?thesis by(simp add: llast_linfinite) qed lemma llast_lappend_LCons [simp]: "lfinite xs \ llast (lappend xs (LCons y ys)) = llast (LCons y ys)" by(simp add: llast_lappend) lemma llast_ldropn: "enat n < llength xs \ llast (ldropn n xs) = llast xs" proof(induct n arbitrary: xs) case 0 thus ?case by simp next case (Suc n) thus ?case by(cases xs)(auto simp add: Suc_ile_eq llast_LCons) qed lemma llast_ldrop: assumes "n < llength xs" shows "llast (ldrop n xs) = llast xs" proof - from assms obtain n' where n: "n = enat n'" by(cases n) auto show ?thesis using assms unfolding n proof(induct n' arbitrary: xs) case 0 thus ?case by(simp add: zero_enat_def[symmetric]) next case Suc thus ?case by(cases xs)(auto simp add: eSuc_enat[symmetric] llast_LCons) qed qed lemma llast_llist_of [simp]: "llast (llist_of xs) = last xs" by(induct xs)(auto simp add: last_def zero_enat_def llast_LCons llast_LNil) lemma llast_conv_lnth: "llength xs = eSuc (enat n) \ llast xs = lnth xs n" by(clarsimp simp add: llast_def zero_enat_def[symmetric] eSuc_enat split: nat.split) lemma llast_lmap: assumes "lfinite xs" "\ lnull xs" shows "llast (lmap f xs) = f (llast xs)" using assms proof induct case (lfinite_LConsI xs) thus ?case by(cases xs) simp_all qed simp subsection \Distinct lazy lists @{term "ldistinct"}\ inductive_simps ldistinct_LCons [code, simp]: "ldistinct (LCons x xs)" lemma ldistinct_LNil_code [code]: "ldistinct LNil = True" by simp lemma ldistinct_llist_of [simp]: "ldistinct (llist_of xs) \ distinct xs" by(induct xs) auto lemma ldistinct_coinduct [consumes 1, case_names ldistinct, case_conclusion ldistinct lhd ltl, coinduct pred: ldistinct]: assumes "X xs" and step: "\xs. \ X xs; \ lnull xs \ \ lhd xs \ lset (ltl xs) \ (X (ltl xs) \ ldistinct (ltl xs))" shows "ldistinct xs" using \X xs\ proof(coinduct) case (ldistinct xs) thus ?case by(cases xs)(auto dest: step) qed lemma ldistinct_lhdD: "\ ldistinct xs; \ lnull xs \ \ lhd xs \ lset (ltl xs)" by(clarsimp simp add: not_lnull_conv) lemma ldistinct_ltlI: "ldistinct xs \ ldistinct (ltl xs)" by(cases xs) simp_all lemma ldistinct_lSup: "\ Complete_Partial_Order.chain (\) Y; \xs\Y. ldistinct xs \ \ ldistinct (lSup Y)" proof(coinduction arbitrary: Y) case (ldistinct Y) hence chain: "Complete_Partial_Order.chain (\) Y" and distinct: "\xs. xs \ Y \ ldistinct xs" by blast+ have ?lhd using chain by(auto 4 4 simp add: lset_lSup chain_lprefix_ltl dest: distinct lhd_lSup_eq ldistinct_lhdD) moreover have ?ltl by(auto 4 3 simp add: chain_lprefix_ltl chain intro: ldistinct_ltlI distinct) ultimately show ?case .. qed lemma admissible_ldistinct [cont_intro, simp]: assumes mcont: "mcont lub ord lSup (\) (\x. f x)" shows "ccpo.admissible lub ord (\x. ldistinct (f x))" proof(rule ccpo.admissibleI) fix Y assume chain: "Complete_Partial_Order.chain ord Y" and distinct: "\x\Y. ldistinct (f x)" and "Y \ {}" thus "ldistinct (f (lub Y))" by(simp add: mcont_contD[OF mcont] ldistinct_lSup chain_imageI mcont_monoD[OF mcont]) qed lemma ldistinct_lappend: "ldistinct (lappend xs ys) \ ldistinct xs \ (lfinite xs \ ldistinct ys \ lset xs \ lset ys = {})" (is "?lhs = ?rhs") proof(intro iffI conjI strip) assume "?lhs" thus "ldistinct xs" by(coinduct)(auto simp add: not_lnull_conv in_lset_lappend_iff) assume "lfinite xs" thus "ldistinct ys" "lset xs \ lset ys = {}" using \?lhs\ by induct simp_all next assume "?rhs" thus ?lhs by(coinduction arbitrary: xs)(auto simp add: not_lnull_conv in_lset_lappend_iff) qed lemma ldistinct_lprefix: "\ ldistinct xs; ys \ xs \ \ ldistinct ys" by(clarsimp simp add: lprefix_conv_lappend ldistinct_lappend) lemma admissible_not_ldistinct[THEN admissible_subst, cont_intro, simp]: "ccpo.admissible lSup (\) (\x. \ ldistinct x)" by(rule ccpo.admissibleI)(auto dest: ldistinct_lprefix intro: chain_lprefix_lSup) lemma ldistinct_ltake: "ldistinct xs \ ldistinct (ltake n xs)" by (metis ldistinct_lprefix ltake_is_lprefix) lemma ldistinct_ldropn: "ldistinct xs \ ldistinct (ldropn n xs)" by(induct n arbitrary: xs)(simp, case_tac xs, simp_all) lemma ldistinct_ldrop: "ldistinct xs \ ldistinct (ldrop n xs)" proof(induct xs arbitrary: n) case (LCons x xs) thus ?case by(cases n rule: co.enat.exhaust) simp_all qed simp_all lemma ldistinct_conv_lnth: "ldistinct xs \ (\i j. enat i < llength xs \ enat j < llength xs \ i \ j \ lnth xs i \ lnth xs j)" (is "?lhs \ ?rhs") proof(intro iffI strip) assume "?rhs" thus "?lhs" proof(coinduct xs) case (ldistinct xs) from \\ lnull xs\ obtain x xs' where LCons: "xs = LCons x xs'" by(auto simp add: not_lnull_conv) have "x \ lset xs'" proof assume "x \ lset xs'" then obtain j where "enat j < llength xs'" "lnth xs' j = x" unfolding lset_conv_lnth by auto hence "enat 0 < llength xs" "enat (Suc j) < llength xs" "lnth xs (Suc j) = x" "lnth xs 0 = x" by(simp_all add: LCons Suc_ile_eq zero_enat_def[symmetric]) thus False by(auto dest: ldistinct(1)[rule_format]) qed moreover { fix i j assume "enat i < llength xs'" "enat j < llength xs'" "i \ j" hence "enat (Suc i) < llength xs" "enat (Suc j) < llength xs" by(simp_all add: LCons Suc_ile_eq) with \i \ j\ have "lnth xs (Suc i) \ lnth xs (Suc j)" by(auto dest: ldistinct(1)[rule_format]) hence "lnth xs' i \ lnth xs' j" unfolding LCons by simp } ultimately show ?case using LCons by simp qed next assume "?lhs" fix i j assume "enat i < llength xs" "enat j < llength xs" "i \ j" thus "lnth xs i \ lnth xs j" proof(induct i j rule: wlog_linorder_le) case symmetry thus ?case by simp next case (le i j) from \?lhs\ have "ldistinct (ldropn i xs)" by(rule ldistinct_ldropn) also note ldropn_Suc_conv_ldropn[symmetric] also from le have "i < j" by simp hence "lnth xs j \ lset (ldropn (Suc i) xs)" using le unfolding in_lset_conv_lnth by(cases "llength xs")(auto intro!: exI[where x="j - Suc i"]) ultimately show ?case using \enat i < llength xs\ by auto qed qed lemma ldistinct_lmap [simp]: "ldistinct (lmap f xs) \ ldistinct xs \ inj_on f (lset xs)" (is "?lhs \ ?rhs") proof(intro iffI conjI) assume dist: ?lhs thus "ldistinct xs" by(coinduct)(auto simp add: not_lnull_conv) show "inj_on f (lset xs)" proof(rule inj_onI) fix x y assume "x \ lset xs" and "y \ lset xs" and "f x = f y" then obtain i j where "enat i < llength xs" "x = lnth xs i" "enat j < llength xs" "y = lnth xs j" unfolding lset_conv_lnth by blast with dist \f x = f y\ show "x = y" unfolding ldistinct_conv_lnth by auto qed next assume ?rhs thus ?lhs by(coinduction arbitrary: xs)(auto simp add: not_lnull_conv) qed lemma ldistinct_lzipI1: "ldistinct xs \ ldistinct (lzip xs ys)" by(coinduction arbitrary: xs ys)(auto simp add: not_lnull_conv dest: lset_lzipD1) lemma ldistinct_lzipI2: "ldistinct ys \ ldistinct (lzip xs ys)" by(coinduction arbitrary: xs ys)(auto 4 3 simp add: not_lnull_conv dest: lset_lzipD2) subsection \Sortedness @{term lsorted}\ context ord begin coinductive lsorted :: "'a llist \ bool" where LNil [simp]: "lsorted LNil" | Singleton [simp]: "lsorted (LCons x LNil)" | LCons_LCons: "\ x \ y; lsorted (LCons y xs) \ \ lsorted (LCons x (LCons y xs))" inductive_simps lsorted_LCons_LCons [simp]: "lsorted (LCons x (LCons y xs))" inductive_simps lsorted_code [code]: "lsorted LNil" "lsorted (LCons x LNil)" "lsorted (LCons x (LCons y xs))" lemma lsorted_coinduct' [consumes 1, case_names lsorted, case_conclusion lsorted lhd ltl, coinduct pred: lsorted]: assumes major: "X xs" and step: "\xs. \ X xs; \ lnull xs; \ lnull (ltl xs) \ \ lhd xs \ lhd (ltl xs) \ (X (ltl xs) \ lsorted (ltl xs))" shows "lsorted xs" using major by coinduct(subst disj_commute, auto 4 4 simp add: neq_LNil_conv dest: step) lemma lsorted_ltlI: "lsorted xs \ lsorted (ltl xs)" by(erule lsorted.cases) simp_all lemma lsorted_lhdD: "\ lsorted xs; \ lnull xs; \ lnull (ltl xs) \ \ lhd xs \ lhd (ltl xs)" by(auto elim: lsorted.cases) lemma lsorted_LCons': "lsorted (LCons x xs) \ (\ lnull xs \ x \ lhd xs \ lsorted xs)" by(cases xs) auto lemma lsorted_lSup: "\ Complete_Partial_Order.chain (\) Y; \xs \ Y. lsorted xs \ \ lsorted (lSup Y)" proof(coinduction arbitrary: Y) case (lsorted Y) hence sorted: "\xs. xs \ Y \ lsorted xs" by blast note chain = \Complete_Partial_Order.chain (\) Y\ from \\ lnull (lSup Y)\ \\ lnull (ltl (lSup Y))\ obtain xs where "xs \ Y" "\ lnull xs" "\ lnull (ltl xs)" by auto hence "lhd (lSup Y) = lhd xs" "lhd (ltl (lSup Y)) = lhd (ltl xs)" "lhd xs \ lhd (ltl xs)" using chain sorted by(auto intro: lhd_lSup_eq chain_lprefix_ltl lsorted_lhdD) hence ?lhd by simp moreover have ?ltl using chain sorted by(auto intro: chain_lprefix_ltl lsorted_ltlI) ultimately show ?case .. qed lemma lsorted_lprefixD: "\ xs \ ys; lsorted ys \ \ lsorted xs" proof(coinduction arbitrary: xs ys) case (lsorted xs ys) hence "lhd xs = lhd ys" "lhd (ltl xs) = lhd (ltl ys)" by(auto dest: lprefix_lhdD lprefix_ltlI) moreover have "lhd ys \ lhd (ltl ys)" using lsorted by(auto intro: lsorted_lhdD dest: lprefix_lnullD lprefix_ltlI) ultimately have ?lhd by simp moreover have ?ltl using lsorted by(blast intro: lsorted_ltlI lprefix_ltlI) ultimately show ?case .. qed lemma admissible_lsorted [cont_intro, simp]: assumes mcont: "mcont lub ord lSup (\) (\x. f x)" and ccpo: "class.ccpo lub ord (mk_less ord)" shows "ccpo.admissible lub ord (\x. lsorted (f x))" proof(rule ccpo.admissibleI) fix Y assume chain: "Complete_Partial_Order.chain ord Y" and sorted: "\x\Y. lsorted (f x)" and "Y \ {}" thus "lsorted (f (lub Y))" by(simp add: mcont_contD[OF mcont] lsorted_lSup chain_imageI mcont_monoD[OF mcont]) qed lemma admissible_not_lsorted[THEN admissible_subst, cont_intro, simp]: "ccpo.admissible lSup (\) (\xs. \ lsorted xs)" by(rule ccpo.admissibleI)(auto dest: lsorted_lprefixD[rotated] intro: chain_lprefix_lSup) lemma lsorted_ltake [simp]: "lsorted xs \ lsorted (ltake n xs)" by(rule lsorted_lprefixD)(rule ltake_is_lprefix) lemma lsorted_ldropn [simp]: "lsorted xs \ lsorted (ldropn n xs)" by(induct n arbitrary: xs)(fastforce simp add: ldropn_Suc lsorted_LCons' ldropn_lnull split: llist.split)+ lemma lsorted_ldrop [simp]: "lsorted xs \ lsorted (ldrop n xs)" by(induct xs arbitrary: n)(auto simp add: ldrop_LCons lsorted_LCons' ldrop_lnull split: co.enat.split) end declare ord.lsorted_code [code] ord.admissible_lsorted [cont_intro, simp] ord.admissible_not_lsorted [THEN admissible_subst, cont_intro, simp] context preorder begin lemma lsorted_LCons: "lsorted (LCons x xs) \ lsorted xs \ (\y\lset xs. x \ y)" (is "?lhs \ ?rhs") proof assume ?lhs { fix y assume "y \ lset xs" hence "x \ y" using \?lhs\ by(induct arbitrary: x)(auto intro: order_trans) } with \?lhs\ show ?rhs by cases auto next assume ?rhs thus ?lhs by(cases xs) simp_all qed lemma lsorted_coinduct [consumes 1, case_names lsorted, case_conclusion lsorted lhd ltl, coinduct pred: lsorted]: assumes major: "X xs" and step: "\xs. \ X xs; \ lnull xs \ \ (\x \ lset (ltl xs). lhd xs \ x) \ (X (ltl xs) \ lsorted (ltl xs))" shows "lsorted xs" using major by(coinduct rule: lsorted_coinduct')(auto dest: step) lemma lsortedD: "\ lsorted xs; \ lnull xs; y \ lset (ltl xs) \ \ lhd xs \ y" by(clarsimp simp add: not_lnull_conv lsorted_LCons) end lemma lsorted_lmap': assumes "ord.lsorted orda xs" "monotone orda ordb f" shows "ord.lsorted ordb (lmap f xs)" using \ord.lsorted orda xs\ by(coinduction arbitrary: xs rule: ord.lsorted_coinduct')(auto intro: monotoneD[OF \monotone orda ordb f\] ord.lsorted_lhdD ord.lsorted_ltlI) lemma lsorted_lmap: assumes "lsorted xs" "monotone (\) (\) f" shows "lsorted (lmap f xs)" using \lsorted xs\ by(coinduction arbitrary: xs rule: lsorted_coinduct')(auto intro: monotoneD[OF \monotone (\) (\) f\] lsorted_lhdD lsorted_ltlI) context linorder begin lemma lsorted_ldistinct_lset_unique: "\ lsorted xs; ldistinct xs; lsorted ys; ldistinct ys; lset xs = lset ys \ \ xs = ys" proof(coinduction arbitrary: xs ys) case (Eq_llist xs ys) hence ?lnull by(cases ys)(auto simp add: lset_lnull) moreover from Eq_llist have ?LCons by(auto 4 3 intro: lsorted_ltlI ldistinct_ltlI simp add: not_lnull_conv insert_eq_iff lsorted_LCons split: if_split_asm) ultimately show ?case .. qed end lemma lsorted_llist_of[simp]: "lsorted (llist_of xs) \ sorted xs" by(induct xs)(auto simp: lsorted_LCons) subsection \Lexicographic order on lazy lists: @{term "llexord"}\ lemma llexord_coinduct [consumes 1, case_names llexord, coinduct pred: llexord]: assumes X: "X xs ys" and step: "\xs ys. \ X xs ys; \ lnull xs \ \ \ lnull ys \ (\ lnull ys \ r (lhd xs) (lhd ys) \ lhd xs = lhd ys \ (X (ltl xs) (ltl ys) \ llexord r (ltl xs) (ltl ys)))" shows "llexord r xs ys" using X proof(coinduct) case (llexord xs ys) thus ?case by(cases xs ys rule: llist.exhaust[case_product llist.exhaust])(auto dest: step) qed lemma llexord_refl [simp, intro!]: "llexord r xs xs" proof - define ys where "ys = xs" hence "xs = ys" by simp thus "llexord r xs ys" by(coinduct xs ys) auto qed lemma llexord_LCons_LCons [simp]: "llexord r (LCons x xs) (LCons y ys) \ (x = y \ llexord r xs ys \ r x y)" by(auto intro: llexord.intros(1,2) elim: llexord.cases) lemma lnull_llexord [simp]: "lnull xs \ llexord r xs ys" unfolding lnull_def by simp lemma llexord_LNil_right [simp]: "lnull ys \ llexord r xs ys \ lnull xs" by(auto elim: llexord.cases) lemma llexord_LCons_left: "llexord r (LCons x xs) ys \ (\y ys'. ys = LCons y ys' \ (x = y \ llexord r xs ys' \ r x y))" by(cases ys)(auto elim: llexord.cases) lemma lprefix_imp_llexord: assumes "xs \ ys" shows "llexord r xs ys" using assms by(coinduct)(auto simp add: not_lnull_conv LCons_lprefix_conv) lemma llexord_empty: "llexord (\x y. False) xs ys = xs \ ys" proof assume "llexord (\x y. False) xs ys" thus "xs \ ys" by(coinduct)(auto elim: llexord.cases) qed(rule lprefix_imp_llexord) lemma llexord_append_right: "llexord r xs (lappend xs ys)" by(rule lprefix_imp_llexord)(auto simp add: lprefix_conv_lappend) lemma llexord_lappend_leftI: assumes "llexord r ys zs" shows "llexord r (lappend xs ys) (lappend xs zs)" proof(cases "lfinite xs") case True thus ?thesis by induct (simp_all add: assms) next case False thus ?thesis by(simp add: lappend_inf) qed lemma llexord_lappend_leftD: assumes lex: "llexord r (lappend xs ys) (lappend xs zs)" and fin: "lfinite xs" and irrefl: "!!x. x \ lset xs \ \ r x x" shows "llexord r ys zs" using fin lex irrefl by(induct) simp_all lemma llexord_lappend_left: "\ lfinite xs; !!x. x \ lset xs \ \ r x x \ \ llexord r (lappend xs ys) (lappend xs zs) \ llexord r ys zs" by(blast intro: llexord_lappend_leftI llexord_lappend_leftD) lemma antisym_llexord: assumes r: "antisymp r" and irrefl: "\x. \ r x x" shows "antisymp (llexord r)" proof(rule antisympI) fix xs ys assume "llexord r xs ys" and "llexord r ys xs" hence "llexord r xs ys \ llexord r ys xs" by auto thus "xs = ys" by (coinduct rule: llist.coinduct) (auto 4 3 simp add: not_lnull_conv irrefl dest: antisympD[OF r, simplified]) qed lemma llexord_antisym: "\ llexord r xs ys; llexord r ys xs; !!a b. \ r a b; r b a \ \ False \ \ xs = ys" using antisympD[OF antisym_llexord, of r xs ys] by(auto intro: antisympI) lemma llexord_trans: assumes 1: "llexord r xs ys" and 2: "llexord r ys zs" and trans: "!!a b c. \ r a b; r b c \ \ r a c" shows "llexord r xs zs" proof - from 1 2 have "\ys. llexord r xs ys \ llexord r ys zs" by blast thus ?thesis by(coinduct)(auto 4 3 simp add: not_lnull_conv llexord_LCons_left dest: trans) qed lemma trans_llexord: "transp r \ transp (llexord r)" by(auto intro!: transpI elim: llexord_trans dest: transpD) lemma llexord_linear: assumes linear: "!!x y. r x y \ x = y \ r y x" shows "llexord r xs ys \ llexord r ys xs" proof(rule disjCI) assume "\ llexord r ys xs" thus "llexord r xs ys" proof(coinduct rule: llexord_coinduct) case (llexord xs ys) show ?case proof(cases xs) case LNil thus ?thesis using llexord by simp next case (LCons x xs') with \\ llexord r ys xs\ obtain y ys' where ys: "ys = LCons y ys'" "\ r y x" "y \ x \ \ llexord r ys' xs'" by(cases ys) auto with \\ r y x\ linear[of x y] ys LCons show ?thesis by auto qed qed qed lemma llexord_code [code]: "llexord r LNil ys = True" "llexord r (LCons x xs) LNil = False" "llexord r (LCons x xs) (LCons y ys) = (r x y \ x = y \ llexord r xs ys)" by auto lemma llexord_conv: "llexord r xs ys \ xs = ys \ (\zs xs' y ys'. lfinite zs \ xs = lappend zs xs' \ ys = lappend zs (LCons y ys') \ (xs' = LNil \ r (lhd xs') y))" (is "?lhs \ ?rhs") proof assume ?lhs show ?rhs (is "_ \ ?prefix") proof(rule disjCI) assume "\ ?prefix" with \?lhs\ show "xs = ys" proof(coinduction arbitrary: xs ys) case (Eq_llist xs ys) hence "llexord r xs ys" and prefix: "\zs xs' y ys'. \ lfinite zs; xs = lappend zs xs'; ys = lappend zs (LCons y ys') \ \ xs' \ LNil \ \ r (lhd xs') y" by auto from \llexord r xs ys\ show ?case proof(cases) case (llexord_LCons_eq xs' ys' x) { fix zs xs'' y ys'' assume "lfinite zs" "xs' = lappend zs xs''" and "ys' = lappend zs (LCons y ys'')" hence "lfinite (LCons x zs)" "xs = lappend (LCons x zs) xs''" and "ys = lappend (LCons x zs) (LCons y ys'')" using llexord_LCons_eq by simp_all hence "xs'' \ LNil \ \ r (lhd xs'') y" by(rule prefix) } with llexord_LCons_eq show ?thesis by auto next case (llexord_LCons_less x y xs' ys') with prefix[of LNil xs y ys'] show ?thesis by simp next case llexord_LNil thus ?thesis using prefix[of LNil xs "lhd ys" "ltl ys"] by(cases ys) simp_all qed qed qed next assume ?rhs thus ?lhs proof(coinduct xs ys) case (llexord xs ys) show ?case proof(cases xs) case LNil thus ?thesis using llexord by simp next case (LCons x xs') with llexord obtain y ys' where "ys = LCons y ys'" by(cases ys)(auto dest: sym simp add: LNil_eq_lappend_iff) show ?thesis proof(cases "x = y") case True from llexord(1) show ?thesis proof assume "xs = ys" with True LCons \ys = LCons y ys'\ show ?thesis by simp next assume "\zs xs' y ys'. lfinite zs \ xs = lappend zs xs' \ ys = lappend zs (LCons y ys') \ (xs' = LNil \ r (lhd xs') y)" then obtain zs xs' y' ys'' where "lfinite zs" "xs = lappend zs xs'" and "ys = lappend zs (LCons y' ys'')" and "xs' = LNil \ r (lhd xs') y'" by blast with True LCons \ys = LCons y ys'\ show ?thesis by(cases zs) auto qed next case False with LCons llexord \ys = LCons y ys'\ have "r x y" by(fastforce elim: lfinite.cases) with LCons \ys = LCons y ys'\ show ?thesis by auto qed qed qed qed lemma llexord_conv_ltake_index: "llexord r xs ys \ (llength xs \ llength ys \ ltake (llength xs) ys = xs) \ (\n. enat n < min (llength xs) (llength ys) \ ltake (enat n) xs = ltake (enat n) ys \ r (lnth xs n) (lnth ys n))" (is "?lhs \ ?rhs") proof(rule iffI) assume ?lhs thus ?rhs (is "?A \ ?B") unfolding llexord_conv proof assume "xs = ys" thus ?thesis by(simp add: ltake_all) next assume "\zs xs' y ys'. lfinite zs \ xs = lappend zs xs' \ ys = lappend zs (LCons y ys') \ (xs' = LNil \ r (lhd xs') y)" then obtain zs xs' y ys' where "lfinite zs" "xs' = LNil \ r (lhd xs') y" and [simp]: "xs = lappend zs xs'" "ys = lappend zs (LCons y ys')" by blast show ?thesis proof(cases xs') case LNil hence ?A by(auto intro: enat_le_plus_same simp add: ltake_lappend1 ltake_all) thus ?thesis .. next case LCons with \xs' = LNil \ r (lhd xs') y\ have "r (lhd xs') y" by simp from \lfinite zs\ obtain zs' where [simp]: "zs = llist_of zs'" unfolding lfinite_eq_range_llist_of by blast with LCons have "enat (length zs') < min (llength xs) (llength ys)" by(auto simp add: less_enat_def eSuc_def split: enat.split) moreover have "ltake (enat (length zs')) xs = ltake (enat (length zs')) ys" by(simp add: ltake_lappend1) moreover have "r (lnth xs (length zs')) (lnth ys (length zs'))" using LCons \r (lhd xs') y\ by(simp add: lappend_llist_of_LCons lnth_lappend1) ultimately have ?B by blast thus ?thesis .. qed qed next assume ?rhs (is "?A \ ?B") thus ?lhs proof assume ?A thus ?thesis proof(coinduct) case (llexord xs ys) thus ?case by(cases xs, simp)(cases ys, auto) qed next assume "?B" then obtain n where len: "enat n < min (llength xs) (llength ys)" and takexs: "ltake (enat n) xs = ltake (enat n) ys" and r: "r (lnth xs n) (lnth ys n)" by blast have "xs = lappend (ltake (enat n) xs) (ldrop (enat n) xs)" by(simp only: lappend_ltake_ldrop) moreover from takexs len have "ys = lappend (ltake (enat n) xs) (LCons (lnth ys n) (ldrop (enat (Suc n)) ys))" by(simp add: ldropn_Suc_conv_ldropn ldrop_enat) moreover from r len have "r (lhd (ldrop (enat n) xs)) (lnth ys n)" by(simp add: lhd_ldrop) moreover have "lfinite (ltake (enat n) xs)" by simp ultimately show ?thesis unfolding llexord_conv by blast qed qed lemma llexord_llist_of: - assumes "irreflp r" - shows "llexord r (llist_of xs) (llist_of ys) \ xs = ys \ (xs, ys) \ lexord {(x, y). r x y}" (is "?lhs \ ?rhs") proof assume ?lhs { fix zs xs' y ys' assume "lfinite zs" "llist_of xs = lappend zs xs'" and "llist_of ys = lappend zs (LCons y ys')" and "xs' = LNil \ r (lhd xs') y" from \lfinite zs\ obtain zs' where [simp]: "zs = llist_of zs'" unfolding lfinite_eq_range_llist_of by blast have "lfinite (llist_of xs)" by simp hence "lfinite xs'" unfolding \llist_of xs = lappend zs xs'\ by simp then obtain XS' where [simp]: "xs' = llist_of XS'" unfolding lfinite_eq_range_llist_of by blast from \llist_of xs = lappend zs xs'\ have [simp]: "xs = zs' @ XS'" by(simp add: lappend_llist_of_llist_of) have "lfinite (llist_of ys)" by simp hence "lfinite ys'" unfolding \llist_of ys = lappend zs (LCons y ys')\ by simp then obtain YS' where [simp]: "ys' = llist_of YS'" unfolding lfinite_eq_range_llist_of by blast from \llist_of ys = lappend zs (LCons y ys')\ have [simp]: "ys = zs' @ y # YS'" by(auto simp add: llist_of.simps(2)[symmetric] lappend_llist_of_llist_of simp del: llist_of.simps(2)) have "(\a ys'. ys = xs @ a # ys') \ (\zs a b. r a b \ (\xs'. xs = zs @ a # xs') \ (\ys'. ys = zs @ b # ys'))" (is "?A \ ?B") proof(cases xs') case LNil thus ?thesis by(auto simp add: llist_of_eq_LNil_conv) next case (LCons x xs'') with \xs' = LNil \ r (lhd xs') y\ have "r (lhd xs') y" by(auto simp add: llist_of_eq_LCons_conv) with LCons have ?B by(auto simp add: llist_of_eq_LCons_conv) fastforce thus ?thesis .. qed hence "(xs, ys) \ {(x, y). \a v. y = x @ a # v \ (\u a b v w. (a, b) \ {(x, y). r x y} \ x = u @ a # v \ y = u @ b # w)}" by auto } - with \?lhs\ assms show ?rhs - apply (simp add: lexord_def irreflp_def llexord_conv) - by (metis ) + with \?lhs\ show ?rhs + unfolding lexord_def llexord_conv llist_of_inject by blast next assume "?rhs" thus ?lhs proof assume "xs = ys" thus ?thesis by simp next assume "(xs, ys) \ lexord {(x, y). r x y}" thus ?thesis - by(coinduction arbitrary: xs ys)(auto, auto simp add: neq_Nil_conv split: if_split_asm) + by(coinduction arbitrary: xs ys)(auto, auto simp add: neq_Nil_conv) qed qed subsection \The filter functional on lazy lists: @{term "lfilter"}\ lemma lfilter_code [simp, code]: shows lfilter_LNil: "lfilter P LNil = LNil" and lfilter_LCons: "lfilter P (LCons x xs) = (if P x then LCons x (lfilter P xs) else lfilter P xs)" by(simp_all add: lfilter.simps) declare lfilter.mono[cont_intro] lemma mono2mono_lfilter[THEN llist.mono2mono, simp, cont_intro]: shows monotone_lfilter: "monotone (\) (\) (lfilter P)" by(rule llist.fixp_preserves_mono1[OF lfilter.mono lfilter_def]) simp lemma mcont2mcont_lfilter[THEN llist.mcont2mcont, simp, cont_intro]: shows mcont_lfilter: "mcont lSup (\) lSup (\) (lfilter P)" by(rule llist.fixp_preserves_mcont1[OF lfilter.mono lfilter_def]) simp lemma lfilter_mono [partial_function_mono]: "mono_llist A \ mono_llist (\f. lfilter P (A f))" by(rule mono2mono_lfilter) lemma lfilter_LCons_seek: "~ (p x) ==> lfilter p (LCons x l) = lfilter p l" by simp lemma lfilter_LCons_found: "P x \ lfilter P (LCons x xs) = LCons x (lfilter P xs)" by simp lemma lfilter_eq_LNil: "lfilter P xs = LNil \ (\x \ lset xs. \ P x)" by(induction xs) simp_all notepad begin fix P xs have "(lfilter P xs = LNil) \ (\x \ lset xs. \ P x)" proof(intro iffI strip) assume "\x \ lset xs. \ P x" hence "lfilter P xs \ LNil" by(induction arbitrary: xs rule: lfilter.fixp_induct)(simp_all split: llist.split del: lprefix_LNil) thus "lfilter P xs = LNil" by simp next fix x assume "x \ lset xs" "lfilter P xs = LNil" thus "\ P x" by induction(simp_all split: if_split_asm) qed end lemma diverge_lfilter_LNil [simp]: "\x\lset xs. \ P x \ lfilter P xs = LNil" by(simp add: lfilter_eq_LNil) lemmas lfilter_False = diverge_lfilter_LNil lemma lnull_lfilter [simp]: "lnull (lfilter P xs) \ (\x\lset xs. \ P x)" unfolding lnull_def by(simp add: lfilter_eq_LNil) lemmas lfilter_empty_conv = lfilter_eq_LNil lemma lhd_lfilter [simp]: "lhd (lfilter P xs) = lhd (ldropWhile (Not \ P) xs)" proof(cases "\x\lset xs. P x") case True then obtain x where "x \ lset xs" and "P x" by blast from \x \ lset xs\ show ?thesis by induct(simp_all add: \P x\) qed(simp add: o_def) lemma ltl_lfilter: "ltl (lfilter P xs) = lfilter P (ltl (ldropWhile (Not \ P) xs))" by(induct xs) simp_all lemma lfilter_eq_LCons: "lfilter P xs = LCons x xs' \ \xs''. xs' = lfilter P xs'' \ ldropWhile (Not \ P) xs = LCons x xs''" by(drule eq_LConsD)(auto intro!: exI simp add: ltl_lfilter o_def ldropWhile_eq_LNil_iff intro: llist.expand) lemma lfilter_K_True [simp]: "lfilter (%_. True) xs = xs" by(induct xs) simp_all lemma lfitler_K_False [simp]: "lfilter (\_. False) xs = LNil" by simp lemma lfilter_lappend_lfinite [simp]: "lfinite xs \ lfilter P (lappend xs ys) = lappend (lfilter P xs) (lfilter P ys)" by(induct rule: lfinite.induct) auto lemma lfinite_lfilterI [simp]: "lfinite xs \ lfinite (lfilter P xs)" by(induct rule: lfinite.induct) simp_all lemma lset_lfilter [simp]: "lset (lfilter P xs) = {x \ lset xs. P x}" by induct(auto simp add: Collect_conj_eq) notepad begin \ \show @{thm [source] lset_lfilter} by fixpoint induction\ note [simp del] = lset_lfilter fix P xs have "lset (lfilter P xs) = lset xs \ {x. P x}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by(induct arbitrary: xs rule: lfilter.fixp_induct)(auto split: llist.split) show "?rhs \ ?lhs" proof fix x assume "x \ ?rhs" hence "x \ lset xs" "P x" by simp_all thus "x \ ?lhs" by induction simp_all qed qed end lemma lfilter_lfilter: "lfilter P (lfilter Q xs) = lfilter (\x. P x \ Q x) xs" by(induction xs) simp_all notepad begin \ \show @{thm [source] lfilter_lfilter} by fixpoint induction\ fix P Q xs have "\xs. lfilter P (lfilter Q xs) \ lfilter (\x. P x \ Q x) xs" by(rule lfilter.fixp_induct)(auto split: llist.split) moreover have "\xs. lfilter (\x. P x \ Q x) xs \ lfilter P (lfilter Q xs)" by(rule lfilter.fixp_induct)(auto split: llist.split) ultimately have "lfilter P (lfilter Q xs) = lfilter (\x. P x \ Q x) xs" by(blast intro: lprefix_antisym) end lemma lfilter_idem [simp]: "lfilter P (lfilter P xs) = lfilter P xs" by(simp add: lfilter_lfilter) lemma lfilter_lmap: "lfilter P (lmap f xs) = lmap f (lfilter (P o f) xs)" by(induct xs) simp_all lemma lfilter_llist_of [simp]: "lfilter P (llist_of xs) = llist_of (filter P xs)" by(induct xs) simp_all lemma lfilter_cong [cong]: assumes xsys: "xs = ys" and set: "\x. x \ lset ys \ P x = Q x" shows "lfilter P xs = lfilter Q ys" using set unfolding xsys by(induction ys)(simp_all add: Bex_def[symmetric]) lemma llength_lfilter_ile: "llength (lfilter P xs) \ llength xs" by(induct xs)(auto intro: order_trans) lemma lfinite_lfilter: "lfinite (lfilter P xs) \ lfinite xs \ finite {n. enat n < llength xs \ P (lnth xs n)}" proof assume "lfinite (lfilter P xs)" { assume "\ lfinite xs" with \lfinite (lfilter P xs)\ have "finite {n. enat n < llength xs \ P (lnth xs n)}" proof(induct ys\"lfilter P xs" arbitrary: xs rule: lfinite_induct) case LNil hence "\x\lset xs. \ P x" by(auto) hence eq: "{n. enat n < llength xs \ P (lnth xs n)} = {}" by(auto simp add: lset_conv_lnth) show ?case unfolding eq .. next case (LCons xs) from \\ lnull (lfilter P xs)\ have exP: "\x\lset xs. P x" by simp let ?xs = "ltl (ldropWhile (\x. \ P x) xs)" let ?xs' = "ltakeWhile (\x. \ P x) xs" from exP obtain n where n: "llength ?xs' = enat n" using lfinite_conv_llength_enat[of ?xs'] by(auto simp add: lfinite_ltakeWhile) have "finite ({n. enat n < llength ?xs} \ {n. P (lnth ?xs n)})" (is "finite ?A") using LCons [[simproc add: finite_Collect]] by(auto simp add: ltl_lfilter lfinite_ldropWhile) hence "finite ((\m. n + 1 + m) ` ({n. enat n < llength ?xs} \ {n. P (lnth ?xs n)}))" (is "finite (?f ` _)") by(rule finite_imageI) moreover have xs: "xs = lappend ?xs' (LCons (lhd (ldropWhile (\x. \ P x) xs)) ?xs)" using exP by simp { fix m assume "m < n" hence "lnth ?xs' m \ lset ?xs'" using n unfolding in_lset_conv_lnth by auto hence "\ P (lnth ?xs' m)" by(auto dest: lset_ltakeWhileD) } hence "{n. enat n < llength xs \ P (lnth xs n)} \ insert (the_enat (llength ?xs')) (?f ` ?A)" using n by(subst (1 2) xs)(cases "llength ?xs", auto simp add: lnth_lappend not_less not_le lnth_LCons' eSuc_enat split: if_split_asm intro!: rev_image_eqI) ultimately show ?case by(auto intro: finite_subset) qed } thus "lfinite xs \ finite {n. enat n < llength xs \ P (lnth xs n)}" by blast next assume "lfinite xs \ finite {n. enat n < llength xs \ P (lnth xs n)}" moreover { assume "lfinite xs" with llength_lfilter_ile[of P xs] have "lfinite (lfilter P xs)" by(auto simp add: lfinite_eq_range_llist_of) } moreover { assume nfin: "\ lfinite xs" hence len: "llength xs = \" by(rule not_lfinite_llength) assume fin: "finite {n. enat n < llength xs \ P (lnth xs n)}" then obtain m where "{n. enat n < llength xs \ P (lnth xs n)} \ {..n. \ m \ n; enat n < llength xs \ \ \ P (lnth xs n)" by auto hence "lfinite (lfilter P xs)" proof(induct m arbitrary: xs) case 0 hence "lnull (lfilter P xs)" by(auto simp add: in_lset_conv_lnth) thus ?case by simp next case (Suc m) { fix n assume "m \ n" "enat n < llength (ltl xs)" hence "Suc m \ Suc n" "enat (Suc n) < llength xs" by(case_tac [!] xs)(auto simp add: Suc_ile_eq) hence "\ P (lnth xs (Suc n))" by(rule Suc) hence "\ P (lnth (ltl xs) n)" using \enat n < llength (ltl xs)\ by(cases xs) (simp_all) } hence "lfinite (lfilter P (ltl xs))" by(rule Suc) thus ?case by(cases xs) simp_all qed } ultimately show "lfinite (lfilter P xs)" by blast qed lemma lfilter_eq_LConsD: assumes "lfilter P ys = LCons x xs" shows "\us vs. ys = lappend us (LCons x vs) \ lfinite us \ (\u\lset us. \ P u) \ P x \ xs = lfilter P vs" proof - let ?us = "ltakeWhile (Not \ P) ys" and ?vs = "ltl (ldropWhile (Not \ P) ys)" from assms have "\ lnull (lfilter P ys)" by auto hence exP: "\x\lset ys. P x" by simp from eq_LConsD[OF assms] have x: "x = lhd (ldropWhile (Not \ P) ys)" and xs: "xs = ltl (lfilter P ys)" by auto from exP have "ys = lappend ?us (LCons x ?vs)" unfolding x by simp moreover have "lfinite ?us" using exP by(simp add: lfinite_ltakeWhile) moreover have "\u\lset ?us. \ P u" by(auto dest: lset_ltakeWhileD) moreover have "P x" unfolding x o_def using exP by(rule lhd_ldropWhile[where P="Not \ P", simplified]) moreover have "xs = lfilter P ?vs" unfolding xs by(simp add: ltl_lfilter) ultimately show ?thesis by blast qed lemma lfilter_eq_lappend_lfiniteD: assumes "lfilter P xs = lappend ys zs" and "lfinite ys" shows "\us vs. xs = lappend us vs \ lfinite us \ ys = lfilter P us \ zs = lfilter P vs" using \lfinite ys\ \lfilter P xs = lappend ys zs\ proof(induct arbitrary: xs zs) case lfinite_LNil hence "xs = lappend LNil xs" "LNil = lfilter P LNil" "zs = lfilter P xs" by simp_all thus ?case by blast next case (lfinite_LConsI ys y) note IH = \\xs zs. lfilter P xs = lappend ys zs \ \us vs. xs = lappend us vs \ lfinite us \ ys = lfilter P us \ zs = lfilter P vs\ from \lfilter P xs = lappend (LCons y ys) zs\ have "lfilter P xs = LCons y (lappend ys zs)" by simp from lfilter_eq_LConsD[OF this] obtain us vs where xs: "xs = lappend us (LCons y vs)" "lfinite us" "P y" "\u\lset us. \ P u" and vs: "lfilter P vs = lappend ys zs" by auto from IH[OF vs] obtain us' vs' where "vs = lappend us' vs'" "lfinite us'" and "ys = lfilter P us'" "zs = lfilter P vs'" by blast with xs show ?case by(fastforce simp add: lappend_snocL1_conv_LCons2[symmetric, where ys="lappend us' vs'"] lappend_assoc[symmetric]) qed lemma ldistinct_lfilterI: "ldistinct xs \ ldistinct (lfilter P xs)" by(induction xs) simp_all notepad begin fix P xs assume *: "ldistinct xs" from * have "ldistinct (lfilter P xs) \ lset (lfilter P xs) \ lset xs" by(induction arbitrary: xs rule: lfilter.fixp_induct)(simp_all split: llist.split, fastforce) from * have "ldistinct (lfilter P xs)" \ \only works because we use strong fixpoint induction\ by(induction arbitrary: xs rule: lfilter.fixp_induct)(simp_all split: llist.split, auto 4 4 dest: monotone_lset[THEN monotoneD] simp add: fun_ord_def) end lemma ldistinct_lfilterD: "\ ldistinct (lfilter P xs); enat n < llength xs; enat m < llength xs; P a; lnth xs n = a; lnth xs m = a \ \ m = n" proof(induct n m rule: wlog_linorder_le) case symmetry thus ?case by simp next case (le n m) thus ?case proof(induct n arbitrary: xs m) case 0 thus ?case proof(cases m) case 0 thus ?thesis . next case (Suc m') with 0 show ?thesis by(cases xs)(simp_all add: Suc_ile_eq, auto simp add: lset_conv_lnth) qed next case (Suc n) from \Suc n \ m\ obtain m' where m [simp]: "m = Suc m'" by(cases m) simp with \Suc n \ m\ have "n \ m'" by simp moreover from \enat (Suc n) < llength xs\ obtain x xs' where xs [simp]: "xs = LCons x xs'" by(cases xs) simp from \ldistinct (lfilter P xs)\ have "ldistinct (lfilter P xs')" by(simp split: if_split_asm) moreover from \enat (Suc n) < llength xs\ \enat m < llength xs\ have "enat n < llength xs'" "enat m' < llength xs'" by(simp_all add: Suc_ile_eq) moreover note \P a\ moreover have "lnth xs' n = a" "lnth xs' m' = a" using \lnth xs (Suc n) = a\ \lnth xs m = a\ by simp_all ultimately have "m' = n" by(rule Suc) thus ?case by simp qed qed lemmas lfilter_fixp_parallel_induct = parallel_fixp_induct_1_1[OF llist_partial_function_definitions llist_partial_function_definitions lfilter.mono lfilter.mono lfilter_def lfilter_def, case_names adm LNil step] lemma llist_all2_lfilterI: assumes *: "llist_all2 P xs ys" and Q: "\x y. P x y \ Q1 x \ Q2 y" shows "llist_all2 P (lfilter Q1 xs) (lfilter Q2 ys)" using * by(induction arbitrary: xs ys rule: lfilter_fixp_parallel_induct)(auto split: llist.split dest: Q) lemma distinct_filterD: "\ distinct (filter P xs); n < length xs; m < length xs; P x; xs ! n = x; xs ! m = x \ \ m = n" using ldistinct_lfilterD[of P "llist_of xs" n m x] by simp lemma lprefix_lfilterI: "xs \ ys \ lfilter P xs \ lfilter P ys" by(rule monotoneD[OF monotone_lfilter]) context preorder begin lemma lsorted_lfilterI: "lsorted xs \ lsorted (lfilter P xs)" by(induct xs)(simp_all add: lsorted_LCons) lemma lsorted_lfilter_same: "lsorted (lfilter (\x. x = c) xs)" by(induct xs)(auto simp add: lsorted_LCons) end lemma lfilter_id_conv: "lfilter P xs = xs \ (\x\lset xs. P x)" (is "?lhs \ ?rhs") proof assume ?rhs thus ?lhs by(induct xs) auto next assume ?lhs with lset_lfilter[of P xs] show ?rhs by auto qed lemma lfilter_repeat [simp]: "lfilter P (repeat x) = (if P x then repeat x else LNil)" by(simp add: lfilter_id_conv) subsection \Concatenating all lazy lists in a lazy list: @{term "lconcat"}\ lemma lconcat_simps [simp, code]: shows lconcat_LNil: "lconcat LNil = LNil" and lconcat_LCons: "lconcat (LCons xs xss) = lappend xs (lconcat xss)" by(simp_all add: lconcat.simps) declare lconcat.mono[cont_intro] lemma mono2mono_lconcat[THEN llist.mono2mono, cont_intro, simp]: shows monotone_lconcat: "monotone (\) (\) lconcat" by(rule llist.fixp_preserves_mono1[OF lconcat.mono lconcat_def]) simp lemma mcont2mcont_lconcat[THEN llist.mcont2mcont, cont_intro, simp]: shows mcont_lconcat: "mcont lSup (\) lSup (\) lconcat" by(rule llist.fixp_preserves_mcont1[OF lconcat.mono lconcat_def]) simp lemma lconcat_eq_LNil: "lconcat xss = LNil \ lset xss \ {LNil}" (is "?lhs \ ?rhs") by(induction xss)(auto simp add: lappend_eq_LNil_iff) lemma lnull_lconcat [simp]: "lnull (lconcat xss) \ lset xss \ {xs. lnull xs}" unfolding lnull_def by(simp add: lconcat_eq_LNil) lemma lconcat_llist_of: "lconcat (llist_of (map llist_of xs)) = llist_of (concat xs)" by(induct xs)(simp_all add: lappend_llist_of_llist_of) lemma lhd_lconcat [simp]: "\ \ lnull xss; \ lnull (lhd xss) \ \ lhd (lconcat xss) = lhd (lhd xss)" by(clarsimp simp add: not_lnull_conv) lemma ltl_lconcat [simp]: "\ \ lnull xss; \ lnull (lhd xss) \ \ ltl (lconcat xss) = lappend (ltl (lhd xss)) (lconcat (ltl xss))" by(clarsimp simp add: not_lnull_conv) lemma lmap_lconcat: "lmap f (lconcat xss) = lconcat (lmap (lmap f) xss)" by(induct xss)(simp_all add: lmap_lappend_distrib) lemma lconcat_lappend [simp]: assumes "lfinite xss" shows "lconcat (lappend xss yss) = lappend (lconcat xss) (lconcat yss)" using assms by induct (simp_all add: lappend_assoc) lemma lconcat_eq_LCons_conv: "lconcat xss = LCons x xs \ (\xs' xss' xss''. xss = lappend (llist_of xss') (LCons (LCons x xs') xss'') \ xs = lappend xs' (lconcat xss'') \ set xss' \ {xs. lnull xs})" (is "?lhs \ ?rhs") proof assume "?rhs" then obtain xs' xss' xss'' where "xss = lappend (llist_of xss') (LCons (LCons x xs') xss'')" and "xs = lappend xs' (lconcat xss'')" and "set xss' \ {xs. lnull xs}" by blast moreover from \set xss' \ {xs. lnull xs}\ have "lnull (lconcat (llist_of xss'))" by simp ultimately show ?lhs by(simp add: lappend_lnull1) next assume "?lhs" hence "\ lnull (lconcat xss)" by simp hence "\ lset xss \ {xs. lnull xs}" by simp hence "\ lnull (lfilter (\xs. \ lnull xs) xss)" by(auto) then obtain y ys where yys: "lfilter (\xs. \ lnull xs) xss = LCons y ys" unfolding not_lnull_conv by auto from lfilter_eq_LConsD[OF this] obtain us vs where xss: "xss = lappend us (LCons y vs)" and "lfinite us" and "lset us \ {xs. lnull xs}" "\ lnull y" and ys: "ys = lfilter (\xs. \ lnull xs) vs" by blast from \lfinite us\ obtain us' where [simp]: "us = llist_of us'" unfolding lfinite_eq_range_llist_of by blast from \lset us \ {xs. lnull xs}\ have us: "lnull (lconcat us)" unfolding lnull_lconcat . from \\ lnull y\ obtain y' ys' where y: "y = LCons y' ys'" unfolding not_lnull_conv by blast from \?lhs\ us have [simp]: "y' = x" "xs = lappend ys' (lconcat vs)" unfolding xss y by(simp_all add: lappend_lnull1) from \lset us \ {xs. lnull xs}\ ys show ?rhs unfolding xss y by simp blast qed lemma llength_lconcat_lfinite_conv_sum: assumes "lfinite xss" shows "llength (lconcat xss) = (\i | enat i < llength xss. llength (lnth xss i))" using assms proof(induct) case lfinite_LNil thus ?case by simp next case (lfinite_LConsI xss xs) have "{i. enat i \ llength xss} = insert 0 {Suc i|i. enat i < llength xss}" by(auto simp add: zero_enat_def[symmetric] Suc_ile_eq gr0_conv_Suc) also have "\ = insert 0 (Suc ` {i. enat i < llength xss})" by auto also have "0 \ Suc ` {i. enat i < llength xss}" by auto moreover from \lfinite xss\ have "finite {i. enat i < llength xss}" by(rule lfinite_finite_index) ultimately show ?case using lfinite_LConsI by(simp add: sum.reindex) qed lemma lconcat_lfilter_neq_LNil: "lconcat (lfilter (\xs. \ lnull xs) xss) = lconcat xss" by(induct xss)(simp_all add: lappend_lnull1) lemmas lconcat_fixp_parallel_induct = parallel_fixp_induct_1_1[OF llist_partial_function_definitions llist_partial_function_definitions lconcat.mono lconcat.mono lconcat_def lconcat_def, case_names adm LNil step] lemma llist_all2_lconcatI: "llist_all2 (llist_all2 A) xss yss \ llist_all2 A (lconcat xss) (lconcat yss)" by(induct arbitrary: xss yss rule: lconcat_fixp_parallel_induct)(auto split: llist.split intro: llist_all2_lappendI) lemma llength_lconcat_eqI: fixes xss :: "'a llist llist" and yss :: "'b llist llist" assumes "llist_all2 (\xs ys. llength xs = llength ys) xss yss" shows "llength (lconcat xss) = llength (lconcat yss)" apply(rule llist_all2_llengthD[where P="\_ _. True"]) apply(rule llist_all2_lconcatI) apply(simp add: llist_all2_True[abs_def] assms) done lemma lset_lconcat_lfinite: "\xs \ lset xss. lfinite xs \ lset (lconcat xss) = (\xs\lset xss. lset xs)" by(induction xss) auto lemma lconcat_ltake: "lconcat (ltake (enat n) xss) = ltake (\ii = ltake (llength xs + (\iii=1.. = (\im n'. lnth (lconcat xss) n = lnth (lnth xss m) n' \ enat n' < llength (lnth xss m) \ enat m < llength xss \ enat n = (\i lnull (lconcat xss)" by auto then obtain x xs where concat_xss: "lconcat xss = LCons x xs" unfolding not_lnull_conv by blast then obtain xs' xss' xss'' where xss: "xss = lappend (llist_of xss') (LCons (LCons x xs') xss'')" and xs: "xs = lappend xs' (lconcat xss'')" and LNil: "set xss' \ {xs. lnull xs}" unfolding lconcat_eq_LCons_conv by blast from LNil have "lnull (lconcat (llist_of xss'))" by simp moreover have [simp]: "lnth xss (length xss') = LCons x xs'" unfolding xss by(simp add: lnth_lappend2) ultimately have "lnth (lconcat xss) 0 = lnth (lnth xss (length xss')) 0" using concat_xss xss by(simp) moreover have "enat 0 < llength (lnth xss (length xss'))" by(simp add: zero_enat_def[symmetric]) moreover have "enat (length xss') < llength xss" unfolding xss by simp moreover have "(\i < length xss'. llength (lnth xss i)) = (\i < length xss'. 0)" proof(rule sum.cong) show "{..< length xss'} = {..< length xss'}" by simp next fix i assume "i \ {..< length xss'}" hence "xss' ! i \ set xss'" unfolding in_set_conv_nth by blast with LNil have "xss' ! i = LNil" by auto moreover from \i \ {..< length xss'}\ have "lnth xss i = xss' ! i" unfolding xss by(simp add: lnth_lappend1) ultimately show "llength (lnth xss i) = 0" by simp qed hence "enat 0 = (\ienat (Suc n) < llength (lconcat xss)\ have "\ lnull (lconcat xss)" by auto then obtain x xs where concat_xss: "lconcat xss = LCons x xs" unfolding not_lnull_conv by blast then obtain xs' xss' xss'' where xss: "xss = lappend (llist_of xss') (LCons (LCons x xs') xss'')" and xs: "xs = lappend xs' (lconcat xss'')" and LNil: "set xss' \ {xs. lnull xs}" unfolding lconcat_eq_LCons_conv by blast from LNil have concat_xss': "lnull (lconcat (llist_of xss'))" by simp from xs have "xs = lconcat (LCons xs' xss'')" by simp with concat_xss \enat (Suc n) < llength (lconcat xss)\ have "enat n < llength (lconcat (LCons xs' xss''))" by(simp add: Suc_ile_eq) from Suc.hyps[OF this] obtain m n' where nth_n: "lnth (lconcat (LCons xs' xss'')) n = lnth (lnth (LCons xs' xss'') m) n'" and n': "enat n' < llength (lnth (LCons xs' xss'') m)" and m': "enat m < llength (LCons xs' xss'')" and n_eq: "enat n = (\i < m. llength (lnth (LCons xs' xss'') i)) + enat n'" by blast from n_eq obtain N where N: "(\i < m. llength (lnth (LCons xs' xss'') i)) = enat N" and n: "n = N + n'" by(cases "\i < m. llength (lnth (LCons xs' xss'') i)") simp_all { fix i assume i: "i < length xss'" hence "xss' ! i = LNil" using LNil unfolding set_conv_nth by auto hence "lnth xss i = LNil" using i unfolding xss by(simp add: lnth_lappend1) } note lnth_prefix = this show ?case proof(cases "m > 0") case True then obtain m' where [simp]: "m = Suc m'" by(cases m) auto have "lnth (lconcat xss) (Suc n) = lnth (lnth xss (m + length xss')) n'" using concat_xss' nth_n unfolding xss by(simp add: lnth_lappend2 lappend_lnull1) moreover have "enat n' < llength (lnth xss (m + length xss'))" using concat_xss' n' unfolding xss by(simp add: lnth_lappend2) moreover have "enat (m + length xss') < llength xss" using concat_xss' m' unfolding xss apply (simp add: Suc_ile_eq) apply (simp add: eSuc_enat[symmetric] eSuc_plus_1 plus_enat_simps(1)[symmetric] del: plus_enat_simps(1)) done moreover have "enat (m + length xss') < llength xss" using m' unfolding xss apply(simp add: Suc_ile_eq) apply (simp add: eSuc_enat[symmetric] eSuc_plus_1 plus_enat_simps(1)[symmetric] del: plus_enat_simps(1)) done moreover { have "(\i < m + length xss'. llength (lnth xss i)) = (\i < length xss'. llength (lnth xss i)) + (\i = length xss'..i < length xss'. llength (lnth xss i)) = 0" by simp also have "{length xss'..i = 0 + length xss'..i = 0.. = (\i = 0.. = eSuc (llength xs') + (\i = Suc 0.. {Suc 0..i = Suc 0..i = Suc 0.. = 1 + (llength (lnth (LCons xs' xss'') 0) + \)" by(simp add: eSuc_plus_1 ac_simps) also note sum.atLeast_Suc_lessThan[symmetric, OF \0 < m\] finally have "enat (Suc n) = (\iii yss \ lconcat xss \ lconcat yss" by(rule monotoneD[OF monotone_lconcat]) lemma lnth_lconcat_ltake: assumes "enat w < llength (lconcat (ltake (enat n) xss))" shows "lnth (lconcat (ltake (enat n) xss)) w = lnth (lconcat xss) w" using assms by(auto intro: lprefix_lnthD lprefix_lconcatI) lemma lfinite_lconcat [simp]: "lfinite (lconcat xss) \ lfinite (lfilter (\xs. \ lnull xs) xss) \ (\xs \ lset xss. lfinite xs)" (is "?lhs \ ?rhs") proof assume "?lhs" thus "?rhs" (is "?concl xss") proof(induct "lconcat xss" arbitrary: xss) case [symmetric]: lfinite_LNil moreover hence "lnull (lfilter (\xs. \ lnull xs) xss)" by(auto simp add: lconcat_eq_LNil) ultimately show ?case by(auto) next case (lfinite_LConsI xs x) from \LCons x xs = lconcat xss\[symmetric] obtain xs' xss' xss'' where xss [simp]: "xss = lappend (llist_of xss') (LCons (LCons x xs') xss'')" and xs [simp]: "xs = lappend xs' (lconcat xss'')" and xss': "set xss' \ {xs. lnull xs}" unfolding lconcat_eq_LCons_conv by blast have "xs = lconcat (LCons xs' xss'')" by simp hence "?concl (LCons xs' xss'')" by(rule lfinite_LConsI) thus ?case using \lfinite xs\ xss' by(auto split: if_split_asm) qed next assume "?rhs" then obtain "lfinite (lfilter (\xs. \ lnull xs) xss)" and "\xs\lset xss. lfinite xs" .. thus ?lhs proof(induct "lfilter (\xs. \ lnull xs) xss" arbitrary: xss) case lfinite_LNil from \LNil = lfilter (\xs. \ lnull xs) xss\[symmetric] have "lset xss \ {xs. lnull xs}" unfolding lfilter_empty_conv by blast hence "lnull (lconcat xss)" by(simp) thus ?case by(simp) next case (lfinite_LConsI xs x) from lfilter_eq_LConsD[OF \LCons x xs = lfilter (\xs. \ lnull xs) xss\[symmetric]] obtain xss' xss'' where xss [simp]: "xss = lappend xss' (LCons x xss'')" and xss': "lfinite xss'" "lset xss' \ {xs. lnull xs}" and x: "\ lnull x" and xs [simp]: "xs = lfilter (\xs. \ lnull xs) xss''" by blast moreover from \\xs\lset xss. lfinite xs\ xss' have "\xs\lset xss''. lfinite xs" by auto with xs have "lfinite (lconcat xss'')" by(rule lfinite_LConsI) ultimately show ?case using \\xs\lset xss. lfinite xs\ by(simp) qed qed lemma list_of_lconcat: assumes "lfinite xss" and "\xs \ lset xss. lfinite xs" shows "list_of (lconcat xss) = concat (list_of (lmap list_of xss))" using assms by induct(simp_all add: list_of_lappend) lemma lfilter_lconcat_lfinite: "\xs\lset xss. lfinite xs \ lfilter P (lconcat xss) = lconcat (lmap (lfilter P) xss)" by(induct xss) simp_all lemma lconcat_repeat_LNil [simp]: "lconcat (repeat LNil) = LNil" by(simp add: lconcat_eq_LNil) lemma lconcat_lmap_singleton [simp]: "lconcat (lmap (\x. LCons (f x) LNil) xs) = lmap f xs" by(induct xs) simp_all lemma lset_lconcat_subset: "lset (lconcat xss) \ (\xs\lset xss. lset xs)" by(induct xss)(auto dest: subsetD[OF lset_lappend]) lemma ldistinct_lconcat: "\ ldistinct xss; \ys. ys \ lset xss \ ldistinct ys; \ys zs. \ ys \ lset xss; zs \ lset xss; ys \ zs \ \ lset ys \ lset zs = {} \ \ ldistinct (lconcat xss)" apply(induction arbitrary: xss rule: lconcat.fixp_induct) apply(auto simp add: ldistinct_lappend fun_ord_def split: llist.split) apply(blast dest!: subsetD[OF lprefix_lsetD] subsetD[OF lset_lconcat_subset]) done subsection \Sublist view of a lazy list: @{term "lnths"}\ lemma lnths_empty [simp]: "lnths xs {} = LNil" by(auto simp add: lnths_def split_def lfilter_empty_conv) lemma lnths_LNil [simp]: "lnths LNil A = LNil" by(simp add: lnths_def) lemma lnths_LCons: "lnths (LCons x xs) A = (if 0 \ A then LCons x (lnths xs {n. Suc n \ A}) else lnths xs {n. Suc n \ A})" proof - let ?it = "iterates Suc" let ?f = "\(x, y). (x, Suc y)" { fix n have "lzip xs (?it (Suc n)) = lmap ?f (lzip xs (?it n))" by(coinduction arbitrary: xs n)(auto) hence "lmap fst (lfilter (\(x, y). y \ A) (lzip xs (?it (Suc n)))) = lmap fst (lfilter (\(x, y). Suc y \ A) (lzip xs (?it n)))" by(simp add: lfilter_lmap o_def split_def llist.map_comp) } thus ?thesis by(auto simp add: lnths_def)(subst iterates, simp)+ qed lemma lset_lnths: "lset (lnths xs I) = {lnth xs i|i. enat i i \ I}" apply(auto simp add: lnths_def lset_lzip) apply(rule_tac x="(lnth xs i, i)" in image_eqI) apply auto done lemma lset_lnths_subset: "lset (lnths xs I) \ lset xs" by(auto simp add: lset_lnths in_lset_conv_lnth) lemma lnths_singleton [simp]: "lnths (LCons x LNil) A = (if 0 : A then LCons x LNil else LNil)" by (simp add: lnths_LCons) lemma lnths_upt_eq_ltake [simp]: "lnths xs {.. llength xs" proof - have "llength (lfilter (\(x, y). y \ A) (lzip xs (iterates Suc 0))) \ llength (lzip xs (iterates Suc 0))" by(rule llength_lfilter_ile) thus ?thesis by(simp add: lnths_def) qed lemma lnths_lmap [simp]: "lnths (lmap f xs) A = lmap f (lnths xs A)" by(simp add: lnths_def lzip_lmap1 llist.map_comp lfilter_lmap o_def split_def) lemma lfilter_conv_lnths: "lfilter P xs = lnths xs {n. enat n < llength xs \ P (lnth xs n)}" proof - have "lnths xs {n. enat n < llength xs \ P (lnth xs n)} = lmap fst (lfilter (\(x, y). enat y < llength xs \ P (lnth xs y)) (lzip xs (iterates Suc 0)))" by(simp add: lnths_def) also have "\(x, y)\lset (lzip xs (iterates Suc 0)). enat y < llength xs \ x = lnth xs y" by(auto simp add: lset_lzip) hence "lfilter (\(x, y). enat y < llength xs \ P (lnth xs y)) (lzip xs (iterates Suc 0)) = lfilter (P \ fst) (lzip xs (iterates Suc 0))" by -(rule lfilter_cong[OF refl], auto) also have "lmap fst (lfilter (P \ fst) (lzip xs (iterates Suc 0))) = lfilter P (lmap fst (lzip xs (iterates Suc 0)))" unfolding lfilter_lmap .. also have "lmap fst (lzip xs (iterates Suc 0)) = xs" by(simp add: lmap_fst_lzip_conv_ltake ltake_all) finally show ?thesis .. qed lemma ltake_iterates_Suc: "ltake (enat n) (iterates Suc m) = llist_of [m.. A})" proof - let ?it = "iterates Suc" from assms have fin: "lfinite xs" by(rule llength_eq_enat_lfiniteD) have "lnths (lappend xs ys) A = lmap fst (lfilter (\(x, y). y \ A) (lzip (lappend xs ys) (?it 0)))" by(simp add: lnths_def) also have "?it 0 = lappend (ltake (enat k) (?it 0)) (ldrop (enat k) (?it 0))" by(simp only: lappend_ltake_ldrop) also note lzip_lappend also note lfilter_lappend_lfinite also note lmap_lappend_distrib also have "lzip xs (ltake (enat k) (?it 0)) = lzip xs (?it 0)" using len by(subst (1 2) lzip_conv_lzip_ltake_min_llength) simp also note lnths_def[symmetric] also have "ldrop (enat k) (?it 0) = ?it k" by(simp add: ldrop_iterates) also { fix n m have "?it (n + m) = lmap (\n. n + m) (?it n)" by(coinduction arbitrary: n)(force)+ } from this[of 0 k] have "?it k = lmap (\n. n + k) (?it 0)" by simp also note lzip_lmap2 also note lfilter_lmap also note llist.map_comp also have "fst \ (\(x, y). (x, y + k)) = fst" by(simp add: o_def split_def) also have "(\(x, y). y \ A) \ (\(x, y). (x, y + k)) = (\(x, y). y \ {n. n + k \ A})" by(simp add: fun_eq_iff) also note lnths_def[symmetric] finally show ?thesis using len by simp qed lemma lnths_split: "lnths xs A = lappend (lnths (ltake (enat n) xs) A) (lnths (ldropn n xs) {m. n + m \ A})" proof(cases "enat n \ llength xs") case False thus ?thesis by(auto simp add: ltake_all ldropn_all) next case True have "xs = lappend (ltake (enat n) xs) (ldrop (enat n) xs)" by(simp only: lappend_ltake_ldrop) hence "xs = lappend (ltake (enat n) xs) (ldropn n xs)" by simp hence "lnths xs A = lnths (lappend (ltake (enat n) xs) (ldropn n xs)) A" by(simp) also note lnths_lappend_lfinite[where k=n] finally show ?thesis using True by(simp add: min_def ac_simps) qed lemma lnths_cong: assumes "xs = ys" and A: "\n. enat n < llength ys \ n \ A \ n \ B" shows "lnths xs A = lnths ys B" proof - have "lfilter (\(x, y). y \ A) (lzip ys (iterates Suc 0)) = lfilter (\(x, y). y \ B) (lzip ys (iterates Suc 0))" by(rule lfilter_cong[OF refl])(auto simp add: lset_lzip A) thus ?thesis unfolding \xs = ys\ lnths_def by simp qed lemma lnths_insert: assumes n: "enat n < llength xs" shows "lnths xs (insert n A) = lappend (lnths (ltake (enat n) xs) A) (LCons (lnth xs n) (lnths (ldropn (Suc n) xs) {m. Suc (n + m) \ A}))" proof - have "lnths xs (insert n A) = lappend (lnths (ltake (enat n) xs) (insert n A)) (lnths (ldropn n xs) {m. n + m \ (insert n A)})" by(rule lnths_split) also have "lnths (ltake (enat n) xs) (insert n A) = lnths (ltake (enat n) xs) A" by(rule lnths_cong[OF refl]) simp also { from n obtain X XS where "ldropn n xs = LCons X XS" by(cases "ldropn n xs")(auto simp add: ldropn_eq_LNil) moreover have "lnth (ldropn n xs) 0 = lnth xs n" using n by(simp) moreover have "ltl (ldropn n xs) = ldropn (Suc n) xs" by(cases xs)(simp_all add: ltl_ldropn) ultimately have "ldropn n xs = LCons (lnth xs n) (ldropn (Suc n) xs)" by simp hence "lnths (ldropn n xs) {m. n + m \ insert n A} = LCons (lnth xs n) (lnths (ldropn (Suc n) xs) {m. Suc (n + m) \ A})" by(simp add: lnths_LCons) } finally show ?thesis . qed lemma lfinite_lnths [simp]: "lfinite (lnths xs A) \ lfinite xs \ finite A" proof assume "lfinite (lnths xs A)" hence "lfinite xs \ finite {n. enat n < llength xs \ (\(x, y). y \ A) (lnth (lzip xs (iterates Suc 0)) n)}" by(simp add: lnths_def lfinite_lfilter) also have "{n. enat n < llength xs \ (\(x, y). y \ A) (lnth (lzip xs (iterates Suc 0)) n)} = {n. enat n < llength xs \ n \ A}" by(auto simp add: lnth_lzip) finally show "lfinite xs \ finite A" by(auto simp add: not_lfinite_llength elim: contrapos_np) next assume "lfinite xs \ finite A" moreover have "{n. enat n < llength xs \ (\(x, y). y \ A) (lnth (lzip xs (iterates Suc 0)) n)} = {n. enat n < llength xs \ n \ A}" by(auto simp add: lnth_lzip) ultimately show "lfinite (lnths xs A)" by(auto simp add: lnths_def lfinite_lfilter) qed subsection \@{const "lsum_list"}\ context monoid_add begin lemma lsum_list_0 [simp]: "lsum_list (lmap (\_. 0) xs) = 0" by(simp add: lsum_list_def) lemma lsum_list_llist_of [simp]: "lsum_list (llist_of xs) = sum_list xs" by(simp add: lsum_list_def) lemma lsum_list_lappend: "\ lfinite xs; lfinite ys \ \ lsum_list (lappend xs ys) = lsum_list xs + lsum_list ys" by(simp add: lsum_list_def list_of_lappend) lemma lsum_list_LNil [simp]: "lsum_list LNil = 0" by(simp add: lsum_list_def) lemma lsum_list_LCons [simp]: "lfinite xs \ lsum_list (LCons x xs) = x + lsum_list xs" by(simp add: lsum_list_def) lemma lsum_list_inf [simp]: "\ lfinite xs \ lsum_list xs = 0" by(simp add: lsum_list_def) end lemma lsum_list_mono: fixes f :: "'a \ 'b :: {monoid_add, ordered_ab_semigroup_add}" assumes "\x. x \ lset xs \ f x \ g x" shows "lsum_list (lmap f xs) \ lsum_list (lmap g xs)" using assms by(auto simp add: lsum_list_def intro: sum_list_mono) subsection \ Alternative view on @{typ "'a llist"} as datatype with constructors @{term "llist_of"} and @{term "inf_llist"} \ lemma lnull_inf_llist [simp]: "\ lnull (inf_llist f)" by(simp add: inf_llist_def) lemma inf_llist_neq_LNil: "inf_llist f \ LNil" using lnull_inf_llist unfolding lnull_def . lemmas LNil_neq_inf_llist = inf_llist_neq_LNil[symmetric] lemma lhd_inf_llist [simp]: "lhd (inf_llist f) = f 0" by(simp add: inf_llist_def) lemma ltl_inf_llist [simp]: "ltl (inf_llist f) = inf_llist (\n. f (Suc n))" by(simp add: inf_llist_def lmap_iterates[symmetric] llist.map_comp) lemma inf_llist_rec [code, nitpick_simp]: "inf_llist f = LCons (f 0) (inf_llist (\n. f (Suc n)))" by(rule llist.expand) simp_all lemma lfinite_inf_llist [iff]: "\ lfinite (inf_llist f)" by(simp add: inf_llist_def) lemma iterates_conv_inf_llist: "iterates f a = inf_llist (\n. (f ^^ n) a)" by(coinduction arbitrary: a)(auto simp add: funpow_swap1) lemma inf_llist_neq_llist_of [simp]: "llist_of xs \ inf_llist f" "inf_llist f \ llist_of xs" using lfinite_llist_of[of xs] lfinite_inf_llist[of f] by fastforce+ lemma lnth_inf_llist [simp]: "lnth (inf_llist f) n = f n" proof(induct n arbitrary: f) case 0 thus ?case by(subst inf_llist_rec) simp next case (Suc n) from Suc[of "\n. f (Suc n)"] show ?case by(subst inf_llist_rec) simp qed lemma inf_llist_lprefix [simp]: "inf_llist f \ xs \ xs = inf_llist f" by(auto simp add: not_lfinite_lprefix_conv_eq) lemma llength_inf_llist [simp]: "llength (inf_llist f) = \" by(simp add: inf_llist_def) lemma lset_inf_llist [simp]: "lset (inf_llist f) = range f" by(auto simp add: lset_conv_lnth) lemma inf_llist_inj [simp]: "inf_llist f = inf_llist g \ f = g" unfolding inf_llist_def lmap_eq_lmap_conv_llist_all2 by(simp add: iterates_conv_inf_llist fun_eq_iff) lemma inf_llist_lnth [simp]: "\ lfinite xs \ inf_llist (lnth xs) = xs" by(coinduction arbitrary: xs)(auto simp add: lnth_0_conv_lhd fun_eq_iff lnth_ltl) lemma llist_exhaust: obtains (llist_of) ys where "xs = llist_of ys" | (inf_llist) f where "xs = inf_llist f" proof(cases "lfinite xs") case True then obtain ys where "xs = llist_of ys" unfolding lfinite_eq_range_llist_of by auto thus ?thesis by(rule that) next case False hence "xs = inf_llist (lnth xs)" by simp thus thesis by(rule that) qed lemma lappend_inf_llist [simp]: "lappend (inf_llist f) xs = inf_llist f" by(simp add: lappend_inf) lemma lmap_inf_llist [simp]: "lmap f (inf_llist g) = inf_llist (f o g)" by(simp add: inf_llist_def llist.map_comp) lemma ltake_enat_inf_llist [simp]: "ltake (enat n) (inf_llist f) = llist_of (map f [0..m. f (m + n))" unfolding inf_llist_def ldropn_lmap ldropn_iterates by(simp add: iterates_conv_inf_llist o_def) lemma ldrop_enat_inf_llist: "ldrop (enat n) (inf_llist f) = inf_llist (\m. f (m + n))" proof(induct n arbitrary: f) case Suc thus ?case by(subst inf_llist_rec)(simp add: eSuc_enat[symmetric]) qed(simp add: zero_enat_def[symmetric]) lemma lzip_inf_llist_inf_llist [simp]: "lzip (inf_llist f) (inf_llist g) = inf_llist (\n. (f n, g n))" by(coinduction arbitrary: f g) auto lemma lzip_llist_of_inf_llist [simp]: "lzip (llist_of xs) (inf_llist f) = llist_of (zip xs (map f [0..n. f (Suc n)) [0..n. f (Suc n)"] show ?case by(subst inf_llist_rec)(simp) qed lemma lzip_inf_llist_llist_of [simp]: "lzip (inf_llist f) (llist_of xs) = llist_of (zip (map f [0..n. f (Suc n)) [0..n. f (Suc n)"] show ?case by(subst inf_llist_rec)(simp) qed lemma llist_all2_inf_llist [simp]: "llist_all2 P (inf_llist f) (inf_llist g) \ (\n. P (f n) (g n))" by(simp add: llist_all2_conv_lzip) lemma llist_all2_llist_of_inf_llist [simp]: "\ llist_all2 P (llist_of xs) (inf_llist f)" by(simp add: llist_all2_conv_lzip) lemma llist_all2_inf_llist_llist_of [simp]: "\ llist_all2 P (inf_llist f) (llist_of xs)" by(simp add: llist_all2_conv_lzip) lemma (in monoid_add) lsum_list_infllist: "lsum_list (inf_llist f) = 0" by simp subsection \Setup for lifting and transfer\ subsubsection \Relator and predicator properties\ abbreviation "llist_all == pred_llist" subsubsection \Transfer rules for the Transfer package\ context includes lifting_syntax begin lemma set1_pre_llist_transfer [transfer_rule]: "(rel_pre_llist A B ===> rel_set A) set1_pre_llist set1_pre_llist" by(auto simp add: rel_pre_llist_def vimage2p_def rel_fun_def set1_pre_llist_def rel_set_def collect_def sum_set_defs prod_set_defs elim: rel_sum.cases split: sum.split_asm) lemma set2_pre_llist_transfer [transfer_rule]: "(rel_pre_llist A B ===> rel_set B) set2_pre_llist set2_pre_llist" by(auto simp add: rel_pre_llist_def vimage2p_def rel_fun_def set2_pre_llist_def rel_set_def collect_def sum_set_defs prod_set_defs elim: rel_sum.cases split: sum.split_asm) lemma LNil_transfer [transfer_rule]: "llist_all2 P LNil LNil" by simp lemma LCons_transfer [transfer_rule]: "(A ===> llist_all2 A ===> llist_all2 A) LCons LCons" unfolding rel_fun_def by simp lemma case_llist_transfer [transfer_rule]: "(B ===> (A ===> llist_all2 A ===> B) ===> llist_all2 A ===> B) case_llist case_llist" unfolding rel_fun_def by (simp split: llist.split) lemma unfold_llist_transfer [transfer_rule]: "((A ===> (=)) ===> (A ===> B) ===> (A ===> A) ===> A ===> llist_all2 B) unfold_llist unfold_llist" proof(rule rel_funI)+ fix IS_LNIL1 :: "'a \ bool" and IS_LNIL2 LHD1 LHD2 LTL1 LTL2 x y assume rel: "(A ===> (=)) IS_LNIL1 IS_LNIL2" "(A ===> B) LHD1 LHD2" "(A ===> A) LTL1 LTL2" and "A x y" from \A x y\ show "llist_all2 B (unfold_llist IS_LNIL1 LHD1 LTL1 x) (unfold_llist IS_LNIL2 LHD2 LTL2 y)" apply(coinduction arbitrary: x y) using rel by(auto 4 4 elim: rel_funE) qed lemma llist_corec_transfer [transfer_rule]: "((A ===> (=)) ===> (A ===> B) ===> (A ===> (=)) ===> (A ===> llist_all2 B) ===> (A ===> A) ===> A ===> llist_all2 B) corec_llist corec_llist" proof(rule rel_funI)+ fix IS_LNIL1 :: "'a \ bool" and IS_LNIL2 LHD1 LHD2 and STOP1 :: "'a \ bool" and STOP2 MORE1 MORE2 LTL1 LTL2 x y assume [transfer_rule]: "(A ===> (=)) IS_LNIL1 IS_LNIL2 " "(A ===> B) LHD1 LHD2" "(A ===> (=)) STOP1 STOP2" "(A ===> llist_all2 B) MORE1 MORE2" "(A ===> A) LTL1 LTL2" assume "A x y" thus "llist_all2 B (corec_llist IS_LNIL1 LHD1 STOP1 MORE1 LTL1 x) (corec_llist IS_LNIL2 LHD2 STOP2 MORE2 LTL2 y)" proof(coinduction arbitrary: x y) case [transfer_rule]: (LNil x y) show ?case by simp transfer_prover next case (LCons x y) note [transfer_rule] = \A x y\ have ?lhd by(simp add: LCons[simplified]) transfer_prover moreover have "STOP1 x = STOP2 y" "llist_all2 B (MORE1 x) (MORE2 y)" "A (LTL1 x) (LTL2 y)" by transfer_prover+ hence ?ltl by(auto simp add: LCons[simplified]) ultimately show ?case .. qed qed lemma ltl_transfer [transfer_rule]: "(llist_all2 A ===> llist_all2 A) ltl ltl" unfolding ltl_def[abs_def] by transfer_prover lemma lset_transfer [transfer_rule]: "(llist_all2 A ===> rel_set A) lset lset" by (intro rel_funI rel_setI) (auto simp only: in_lset_conv_lnth llist_all2_conv_all_lnth Bex_def) lemma lmap_transfer [transfer_rule]: "((A ===> B) ===> llist_all2 A ===> llist_all2 B) lmap lmap" by(auto simp add: rel_fun_def llist_all2_lmap1 llist_all2_lmap2 elim: llist_all2_mono) lemma lappend_transfer [transfer_rule]: "(llist_all2 A ===> llist_all2 A ===> llist_all2 A) lappend lappend" by(auto simp add: rel_fun_def intro: llist_all2_lappendI) lemma iterates_transfer [transfer_rule]: "((A ===> A) ===> A ===> llist_all2 A) iterates iterates" unfolding iterates_def split_def corec_llist_never_stop by transfer_prover lemma lfinite_transfer [transfer_rule]: "(llist_all2 A ===> (=)) lfinite lfinite" by(auto dest: llist_all2_lfiniteD) lemma llist_of_transfer [transfer_rule]: "(list_all2 A ===> llist_all2 A) llist_of llist_of" unfolding llist_of_def by transfer_prover lemma llength_transfer [transfer_rule]: "(llist_all2 A ===> (=)) llength llength" by(auto dest: llist_all2_llengthD) lemma ltake_transfer [transfer_rule]: "((=) ===> llist_all2 A ===> llist_all2 A) ltake ltake" by(auto intro: llist_all2_ltakeI) lemma ldropn_transfer [transfer_rule]: "((=) ===> llist_all2 A ===> llist_all2 A) ldropn ldropn" unfolding ldropn_def[abs_def] by transfer_prover lemma ldrop_transfer [transfer_rule]: "((=) ===> llist_all2 A ===> llist_all2 A) ldrop ldrop" by(auto intro: llist_all2_ldropI) lemma ltakeWhile_transfer [transfer_rule]: "((A ===> (=)) ===> llist_all2 A ===> llist_all2 A) ltakeWhile ltakeWhile" proof(rule rel_funI)+ fix P :: "'a \ bool" and Q :: "'b \ bool" and xs :: "'a llist" and ys :: "'b llist" assume PQ: "(A ===> (=)) P Q" assume "llist_all2 A xs ys" thus "llist_all2 A (ltakeWhile P xs) (ltakeWhile Q ys)" apply(coinduction arbitrary: xs ys) using PQ by(auto 4 4 elim: rel_funE simp add: not_lnull_conv llist_all2_LCons2 llist_all2_LCons1) qed lemma ldropWhile_transfer [transfer_rule]: "((A ===> (=)) ===> llist_all2 A ===> llist_all2 A) ldropWhile ldropWhile" unfolding ldropWhile_eq_ldrop[abs_def] by transfer_prover lemma lzip_ltransfer [transfer_rule]: "(llist_all2 A ===> llist_all2 B ===> llist_all2 (rel_prod A B)) lzip lzip" by(auto intro: llist_all2_lzipI) lemma inf_llist_transfer [transfer_rule]: "(((=) ===> A) ===> llist_all2 A) inf_llist inf_llist" unfolding inf_llist_def[abs_def] by transfer_prover lemma lfilter_transfer [transfer_rule]: "((A ===> (=)) ===> llist_all2 A ===> llist_all2 A) lfilter lfilter" by(auto simp add: rel_fun_def intro: llist_all2_lfilterI) lemma lconcat_transfer [transfer_rule]: "(llist_all2 (llist_all2 A) ===> llist_all2 A) lconcat lconcat" by(auto intro: llist_all2_lconcatI) lemma lnths_transfer [transfer_rule]: "(llist_all2 A ===> (=) ===> llist_all2 A) lnths lnths" unfolding lnths_def[abs_def] by transfer_prover lemma llist_all_transfer [transfer_rule]: "((A ===> (=)) ===> llist_all2 A ===> (=)) llist_all llist_all" unfolding pred_llist_def[abs_def] by transfer_prover lemma llist_all2_rsp: assumes r: "\x y. R x y \ (\a b. R a b \ S x a = T y b)" and l1: "llist_all2 R x y" and l2: "llist_all2 R a b" shows "llist_all2 S x a = llist_all2 T y b" proof(cases "llength x = llength a") case True thus ?thesis using l1 l2 by(simp add: llist_all2_conv_all_lnth)(blast dest: r[rule_format]) next case False with llist_all2_llengthD[OF l1] llist_all2_llengthD[OF l2] show ?thesis by(simp add: llist_all2_conv_all_lnth) qed lemma llist_all2_transfer [transfer_rule]: "((R ===> R ===> (=)) ===> llist_all2 R ===> llist_all2 R ===> (=)) llist_all2 llist_all2" by (simp add: llist_all2_rsp rel_fun_def) end no_notation lprefix (infix "\" 65) end diff --git a/thys/Compiling-Exceptions-Correctly/Exceptions.thy b/thys/Compiling-Exceptions-Correctly/Exceptions.thy --- a/thys/Compiling-Exceptions-Correctly/Exceptions.thy +++ b/thys/Compiling-Exceptions-Correctly/Exceptions.thy @@ -1,191 +1,186 @@ (* Author: Tobias Nipkow Copyright 2004 TU Muenchen *) section \Compiling exception handling\ theory Exceptions imports Main begin subsection\The source language\ datatype expr = Val int | Add expr expr | Throw | Catch expr expr primrec eval :: "expr \ int option" where "eval (Val i) = Some i" | "eval (Add x y) = (case eval x of None \ None | Some i \ (case eval y of None \ None | Some j \ Some(i+j)))" | "eval Throw = None" | "eval (Catch x h) = (case eval x of None \ eval h | Some i \ Some i)" subsection\The target language\ datatype instr = Push int | ADD | THROW | Mark nat | Unmark | Label nat | Jump nat datatype item = VAL int | HAN nat type_synonym code = "instr list" type_synonym stack = "item list" fun jump where "jump l [] = []" | "jump l (Label l' # cs) = (if l = l' then cs else jump l cs)" | "jump l (c # cs) = jump l cs" lemma size_jump1: "size (jump l cs) < Suc (size cs)" apply(induct cs) apply simp apply(case_tac a) apply auto done lemma size_jump2: "size (jump l cs) < size cs \ jump l cs = cs" apply(induct cs) apply simp apply(case_tac a) apply auto - done - -lemma jump_neq [simp]: "jump l cs \ Jump l # cs" - using size_jump1 [of l cs] by auto +done function (sequential) exec2 :: "bool \ code \ stack \ stack" where "exec2 True [] s = s" | "exec2 True (Push i#cs) s = exec2 True cs (VAL i # s)" | "exec2 True (ADD#cs) (VAL j # VAL i # s) = exec2 True cs (VAL(i+j) # s)" | "exec2 True (THROW#cs) s = exec2 False cs s" | "exec2 True (Mark l#cs) s = exec2 True cs (HAN l # s)" | "exec2 True (Unmark#cs) (v # HAN l # s) = exec2 True cs (v # s)" | "exec2 True (Label l#cs) s = exec2 True cs s" | "exec2 True (Jump l#cs) s = exec2 True (jump l cs) s" | "exec2 False cs [] = []" | "exec2 False cs (VAL i # s) = exec2 False cs s" | "exec2 False cs (HAN l # s) = exec2 True (jump l cs) s" by pat_completeness auto -termination - apply (relation +termination by (relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(b,cs,s). (cs,s))") - using size_jump2 apply(force simp add: size_jump1)+ - done + (auto simp add: size_jump1 size_jump2) abbreviation "exec \ exec2 True" abbreviation "unwind \ exec2 False" subsection\The compiler\ primrec compile :: "nat \ expr \ code * nat" where "compile l (Val i) = ([Push i], l)" | "compile l (Add x y) = (let (xs,m) = compile l x; (ys,n) = compile m y in (xs @ ys @ [ADD], n))" | "compile l Throw = ([THROW],l)" | "compile l (Catch x h) = (let (xs,m) = compile (l+2) x; (hs,n) = compile m h in (Mark l # xs @ [Unmark, Jump (l+1), Label l] @ hs @ [Label(l+1)], n))" abbreviation cmp :: "nat \ expr \ code" where "cmp l e == fst(compile l e)" primrec isFresh :: "nat \ stack \ bool" where "isFresh l [] = True" | "isFresh l (it#s) = (case it of VAL i \ isFresh l s | HAN l' \ l' < l \ isFresh l s)" definition conv :: "code \ stack \ int option \ stack" where "conv cs s io = (case io of None \ unwind cs s | Some i \ exec cs (VAL i # s))" subsection\The proofs\ text\Lemma numbers are the same as in the paper.\ declare conv_def[simp] option.splits[split] Let_def[simp] lemma 3: "(\l. c = Label l \ isFresh l s) \ unwind (c#cs) s = unwind cs s" apply(induct s) apply simp apply(auto) apply(case_tac a) apply auto apply(case_tac c) apply auto done corollary [simp]: "(!!l. c \ Label l) \ unwind (c#cs) s = unwind cs s" by(blast intro: 3) corollary [simp]: "isFresh l s \ unwind (Label l#cs) s = unwind cs s" by(blast intro: 3) lemma 5: "\ isFresh l s; l \ m \ \ isFresh m s" apply(induct s) apply simp apply(auto split:item.split) done corollary [simp]: "isFresh l s \ isFresh (Suc l) s" by(auto intro:5) lemma 6: "\l. l \ snd(compile l e)" proof(induct e) case Val thus ?case by simp next case (Add x y) from \l \ snd (compile l x)\ and \snd (compile l x) \ snd (compile (snd (compile l x)) y)\ show ?case by(simp_all add:split_def) next case Throw thus ?case by simp next case (Catch x h) from \l+2 \ snd (compile (l+2) x)\ and \snd (compile (l+2) x) \ snd (compile (snd (compile (l+2) x)) h)\ show ?case by(simp_all add:split_def) qed corollary [simp]: "l < m \ l < snd(compile m e)" using 6[where l = m and e = e] by auto corollary [simp]: "isFresh l s \ isFresh (snd(compile l e)) s" using 5 6 by blast text\Contrary to what the paper says, the proof of lemma 4 does not just need lemma 3 but also the above corollary of 5 and 6. Hence the strange order of the lemmas in our proof.\ lemma 4 [simp]: "\l cs. isFresh l s \ unwind (cmp l e @ cs) s = unwind cs s" by (induct e) (auto simp add:split_def) lemma 7 [simp]: "l < m \ jump l (cmp m e @ cs) = jump l cs" by (induct e arbitrary: m cs) (simp_all add:split_def) text\The compiler correctness theorem:\ theorem comp_corr: "\l s cs. isFresh l s \ exec (cmp l e @ cs) s = conv cs s (eval e)" by(induct e)(auto simp add:split_def) text\The specialized and more readable version (omitted in the paper):\ corollary "exec (cmp l e) [] = (case eval e of None \ [] | Some n \ [VAL n])" by (simp add: comp_corr[where cs = "[]", simplified]) end diff --git a/thys/Containers/List_Proper_Interval.thy b/thys/Containers/List_Proper_Interval.thy --- a/thys/Containers/List_Proper_Interval.thy +++ b/thys/Containers/List_Proper_Interval.thy @@ -1,71 +1,66 @@ (* Title: Containers/List_Proper_Interval.thy Author: Andreas Lochbihler, ETH Zurich *) theory List_Proper_Interval imports "HOL-Library.List_Lexorder" Collection_Order begin section \Instantiate @{class proper_interval} of for @{typ "'a list"}\ lemma Nil_less_conv_neq_Nil: "[] < xs \ xs \ []" by(cases xs) simp_all lemma less_append_same_iff: fixes xs :: "'a :: preorder list" shows "xs < xs @ ys \ [] < ys" by(induct xs) simp_all lemma less_append_same2_iff: fixes xs :: "'a :: preorder list" shows "xs @ ys < xs @ zs \ ys < zs" by(induct xs) simp_all lemma Cons_less_iff: fixes x :: "'a :: preorder" shows "x # xs < ys \ (\y ys'. ys = y # ys' \ (x < y \ x = y \ xs < ys'))" by(cases ys) auto instantiation list :: ("{proper_interval, order}") proper_interval begin definition proper_interval_list_aux :: "'a list \ 'a list \ bool" where proper_interval_list_aux_correct: "proper_interval_list_aux xs ys \ (\zs. xs < zs \ zs < ys)" -lemma proper_interval_list_aux_simp_1: +lemma proper_interval_list_aux_simps [code]: "proper_interval_list_aux xs [] \ False" - by (simp add: proper_interval_list_aux_correct) - -lemma proper_interval_list_aux_simp_2: "proper_interval_list_aux [] (y # ys) \ ys \ [] \ proper_interval None (Some y)" - apply(auto simp add: proper_interval_list_aux_correct proper_interval_simps Nil_less_conv_neq_Nil) - apply (metis Cons_less_Cons neq_Nil_conv not_less_Nil) - by (metis Cons_less_Cons Nil_less_Cons neq_Nil_conv) - -lemma proper_interval_list_aux_simp_3: - "proper_interval_list_aux (x # xs) (y # ys) \ (if x = y then proper_interval_list_aux xs ys else x < y)" - apply(auto simp add: proper_interval_list_aux_correct proper_interval_simps Nil_less_conv_neq_Nil Cons_less_iff) - apply fastforce - by (metis Cons_less_Cons Nil_less_Cons less_append_same_iff) - -lemmas proper_interval_list_aux_simps [code] = - proper_interval_list_aux_simp_1 proper_interval_list_aux_simp_2 proper_interval_list_aux_simp_3 + "proper_interval_list_aux (x # xs) (y # ys) \ x < y \ x = y \ proper_interval_list_aux xs ys" +apply(simp_all add: proper_interval_list_aux_correct proper_interval_simps Nil_less_conv_neq_Nil) + apply(fastforce simp add: neq_Nil_conv) +apply(rule iffI) + apply(fastforce simp add: Cons_less_iff intro: less_trans) +apply(erule disjE) + apply(rule exI[where x="x # xs @ [undefined]"]) + apply(simp add: less_append_same_iff) +apply(auto 4 3 simp add: Cons_less_iff) +done fun proper_interval_list :: "'a list option \ 'a list option \ bool" where "proper_interval_list None None \ True" | "proper_interval_list None (Some xs) \ (xs \ [])" | "proper_interval_list (Some xs) None \ True" | "proper_interval_list (Some xs) (Some ys) \ proper_interval_list_aux xs ys" instance proof(intro_classes) fix xs ys :: "'a list" show "proper_interval None (Some ys) \ (\zs. zs < ys)" by(auto simp add: Nil_less_conv_neq_Nil intro: exI[where x="[]"]) show "proper_interval (Some xs) None \ (\zs. xs < zs)" by(simp add: exI[where x="xs @ [undefined]"] less_append_same_iff) show "proper_interval (Some xs) (Some ys) \ (\zs. xs < zs \ zs < ys)" by(simp add: proper_interval_list_aux_correct) qed simp end 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 simp: antisym_def) + apply (auto intro: transI) 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 simp: antisym_def) + apply (auto intro: transI) 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 simp: antisym_def) + apply (auto intro: transI) 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,1873 +1,1872 @@ 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 apply (clarsimp simp: finite_psupset_def sub_insert_same_iff) - by (metis insertI1 set_of_mem) + subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) 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/Jinja/DFA/Kildall.thy b/thys/Jinja/DFA/Kildall.thy --- a/thys/Jinja/DFA/Kildall.thy +++ b/thys/Jinja/DFA/Kildall.thy @@ -1,517 +1,513 @@ (* Title: HOL/MicroJava/BV/Kildall.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM Kildall's algorithm. *) section \Kildall's Algorithm \label{sec:Kildall}\ theory Kildall imports SemilatAlg begin primrec propa :: "'s binop \ (nat \ 's) list \ 's list \ nat set \ 's list * nat set" where "propa f [] \s w = (\s,w)" | "propa f (q'#qs) \s w = (let (q,\) = q'; u = \ \\<^bsub>f\<^esub> \s!q; w' = (if u = \s!q then w else insert q w) in propa f qs (\s[q := u]) w')" definition iter :: "'s binop \ 's step_type \ 's list \ nat set \ 's list \ nat set" where "iter f step \s w = while (\(\s,w). w \ {}) (\(\s,w). let p = SOME p. p \ w in propa f (step p (\s!p)) \s (w-{p})) (\s,w)" definition unstables :: "'s ord \ 's step_type \ 's list \ nat set" where "unstables r step \s = {p. p < size \s \ \stable r step \s p}" definition kildall :: "'s ord \ 's binop \ 's step_type \ 's list \ 's list" where "kildall r f step \s = fst(iter f step \s (unstables r step \s))" primrec merges :: "'s binop \ (nat \ 's) list \ 's list \ 's list" where "merges f [] \s = \s" | "merges f (p'#ps) \s = (let (p,\) = p' in merges f ps (\s[p := \ \\<^bsub>f\<^esub> \s!p]))" lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric] lemma (in Semilat) nth_merges: "\ss. \p < length ss; ss \ list n A; \(p,t)\set ps. p t\A \ \ (merges f ps ss)!p = map snd [(p',t') \ ps. p'=p] \\<^bsub>f\<^esub> ss!p" (is "\ss. \_; _; ?steptype ps\ \ ?P ss ps") (*<*) proof (induct ps) show "\ss. ?P ss []" by simp fix ss p' ps' assume ss: "ss \ list n A" assume l: "p < length ss" assume "?steptype (p'#ps')" then obtain a b where p': "p'=(a,b)" and ab: "aA" and ps': "?steptype ps'" by (cases p') auto assume "\ss. p< length ss \ ss \ list n A \ ?steptype ps' \ ?P ss ps'" hence IH: "\ss. ss \ list n A \ p < length ss \ ?P ss ps'" using ps' by iprover from ss ab have "ss[a := b \\<^bsub>f\<^esub> ss!a] \ list n A" by (simp add: closedD) moreover with l have "p < length (ss[a := b \\<^bsub>f\<^esub> ss!a])" by simp ultimately have "?P (ss[a := b \\<^bsub>f\<^esub> ss!a]) ps'" by (rule IH) with p' l show "?P ss (p'#ps')" by simp qed (*>*) (** merges **) lemma length_merges [simp]: "\ss. size(merges f ps ss) = size ss" (*<*) by (induct ps, auto) (*>*) lemma (in Semilat) merges_preserves_type_lemma: shows "\xs. xs \ list n A \ (\(p,x) \ set ps. p x\A) \ merges f ps xs \ list n A" (*<*) apply (insert closedI) apply (unfold closed_def) apply (induct ps) apply simp apply clarsimp done (*>*) lemma (in Semilat) merges_preserves_type [simp]: "\ xs \ list n A; \(p,x) \ set ps. p x\A \ \ merges f ps xs \ list n A" by (simp add: merges_preserves_type_lemma) lemma (in Semilat) merges_incr_lemma: "\xs. xs \ list n A \ (\(p,x)\set ps. p x \ A) \ xs [\\<^bsub>r\<^esub>] merges f ps xs" (*<*) apply (induct ps) apply simp apply simp apply clarify apply (rule order_trans) apply simp apply (erule list_update_incr) apply simp apply simp apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) done (*>*) lemma (in Semilat) merges_incr: "\ xs \ list n A; \(p,x)\set ps. p x \ A \ \ xs [\\<^bsub>r\<^esub>] merges f ps xs" by (simp add: merges_incr_lemma) lemma (in Semilat) merges_same_conv [rule_format]: "(\xs. xs \ list n A \ (\(p,x)\set ps. p x\A) \ (merges f ps xs = xs) = (\(p,x)\set ps. x \\<^bsub>r\<^esub> xs!p))" (*<*) apply (induct_tac ps) apply simp apply clarsimp apply (rename_tac p x ps xs) apply (rule iffI) apply (rule context_conjI) apply (subgoal_tac "xs[p := x \\<^bsub>f\<^esub> xs!p] [\\<^bsub>r\<^esub>] xs") apply (force dest!: le_listD simp add: nth_list_update) apply (erule subst, rule merges_incr) apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) apply clarify apply (rule conjI) apply simp apply (blast dest: boundedD) apply blast apply clarify apply (erule allE) apply (erule impE) apply assumption apply (drule bspec) apply assumption apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) apply blast apply clarify apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) done (*>*) lemma (in Semilat) list_update_le_listI [rule_format]: "set xs \ A \ set ys \ A \ xs [\\<^bsub>r\<^esub>] ys \ p < size xs \ x \\<^bsub>r\<^esub> ys!p \ x\A \ xs[p := x \\<^bsub>f\<^esub> xs!p] [\\<^bsub>r\<^esub>] ys" (*<*) apply (insert semilat) apply (simp only: Listn.le_def lesub_def semilat_def) apply (simp add: list_all2_conv_all_nth nth_list_update) done (*>*) lemma (in Semilat) merges_pres_le_ub: assumes "set ts \ A" "set ss \ A" "\(p,t)\set ps. t \\<^bsub>r\<^esub> ts!p \ t \ A \ p < size ts" "ss [\\<^bsub>r\<^esub>] ts" shows "merges f ps ss [\\<^bsub>r\<^esub>] ts" (*<*) proof - { fix t ts ps have "\qs. \set ts \ A; \(p,t)\set ps. t \\<^bsub>r\<^esub> ts!p \ t \ A \ p< size ts \ \ set qs \ set ps \ (\ss. set ss \ A \ ss [\\<^bsub>r\<^esub>] ts \ merges f qs ss [\\<^bsub>r\<^esub>] ts)" apply (induct_tac qs) apply simp apply (simp (no_asm_simp)) apply clarify apply simp apply (erule allE, erule impE, erule_tac [2] mp) apply (drule bspec, assumption) apply (simp add: closedD) apply (drule bspec, assumption) apply (simp add: list_update_le_listI) done } note this [dest] from assms show ?thesis by blast qed (*>*) (** propa **) lemma decomp_propa: "\ss w. (\(q,t)\set qs. q < size ss) \ propa f qs ss w = (merges f qs ss, {q. \t.(q,t)\set qs \ t \\<^bsub>f\<^esub> ss!q \ ss!q} \ w)" (*<*) apply (induct qs) apply simp apply (simp (no_asm)) apply clarify apply simp apply (rule conjI) apply blast apply (simp add: nth_list_update) apply blast done (*>*) (** iter **) lemma (in Semilat) stable_pres_lemma: shows "\pres_type step n A; bounded step n; ss \ list n A; p \ w; \q\w. q < n; \q. q < n \ q \ w \ stable r step ss q; q < n; \s'. (q,s') \ set (step p (ss!p)) \ s' \\<^bsub>f\<^esub> ss!q = ss!q; q \ w \ q = p \ \ stable r step (merges f (step p (ss!p)) ss) q" (*<*) apply (unfold stable_def) apply (subgoal_tac "\s'. (q,s') \ set (step p (ss!p)) \ s' : A") prefer 2 apply clarify apply (erule pres_typeD) prefer 3 apply assumption apply (rule listE_nth_in) apply assumption apply simp apply simp apply simp apply clarify apply (subst nth_merges) apply simp apply (blast dest: boundedD) apply assumption apply clarify apply (rule conjI) apply (blast dest: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply simp apply simp apply(subgoal_tac "q < length ss") prefer 2 apply simp apply (frule nth_merges [of q _ _ "step p (ss!p)"]) (* fixme: why does method subst not work?? *) apply assumption apply clarify apply (rule conjI) apply (blast dest: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply simp apply simp apply (drule_tac P = "\x. (a, b) \ set (step q x)" in subst) apply assumption apply (simp add: plusplus_empty) apply (cases "q \ w") apply simp apply (rule ub1') apply (rule Semilat.intro) apply (rule semilat) apply clarify apply (rule pres_typeD) apply assumption prefer 3 apply assumption apply (blast intro: listE_nth_in dest: boundedD) apply (blast intro: pres_typeD dest: boundedD) apply (blast intro: listE_nth_in dest: boundedD) apply assumption apply simp apply (erule allE, erule impE, assumption, erule impE, assumption) apply (rule order_trans) apply simp defer apply (rule pp_ub2)(* apply assumption*) apply simp apply clarify apply simp apply (rule pres_typeD) apply assumption prefer 3 apply assumption apply (blast intro: listE_nth_in dest: boundedD) apply (blast intro: pres_typeD dest: boundedD) apply (blast intro: listE_nth_in dest: boundedD) apply blast done (*>*) lemma (in Semilat) merges_bounded_lemma: "\ mono r step n A; bounded step n; \(p',s') \ set (step p (ss!p)). s' \ A; ss \ list n A; ts \ list n A; p < n; ss [\\<^sub>r] ts; \p. p < n \ stable r step ts p \ \ merges f (step p (ss!p)) ss [\\<^sub>r] ts" (*<*) apply (unfold stable_def) apply (rule merges_pres_le_ub) apply simp apply simp prefer 2 apply assumption apply clarsimp apply (drule boundedD, assumption+) apply (erule allE, erule impE, assumption) apply (drule bspec, assumption) apply simp apply (drule monoD [of _ _ _ _ p "ss!p" "ts!p"]) apply assumption apply simp apply (simp add: le_listD) apply (drule lesub_step_typeD, assumption) apply clarify apply (drule bspec, assumption) apply simp apply (blast intro: order_trans) done (*>*) lemma termination_lemma: assumes "Semilat A r f" shows "\ ss \ list n A; \(q,t)\set qs. q t\A; p\w \ \ ss [\\<^sub>r] merges f qs ss \ merges f qs ss = ss \ {q. \t. (q,t)\set qs \ t \\<^bsub>f\<^esub> ss!q \ ss!q} \ (w-{p}) \ w" (*<*) (is "PROP ?P") proof - interpret Semilat A r f by fact show "PROP ?P" apply(insert semilat) apply (unfold lesssub_def) apply (simp (no_asm_simp) add: merges_incr) apply (rule impI) apply (rule merges_same_conv [THEN iffD1, elim_format]) apply assumption+ defer apply (rule sym, assumption) defer apply simp apply (subgoal_tac "\q t. \((q, t) \ set qs \ t \\<^bsub>f\<^esub> ss ! q \ ss ! q)") apply (blast intro!: psubsetI elim: equalityE) apply clarsimp apply (drule bspec, assumption) apply (drule bspec, assumption) apply clarsimp done qed (*>*) -(*for lex*) -lemma ne_lesssub_iff [simp]: "y\x \ x [\\<^sub>r] y \ x [\\<^sub>r] y" - by (meson lesssub_def) - lemma iter_properties[rule_format]: assumes "Semilat A r f" shows "\ acc r; pres_type step n A; mono r step n A; bounded step n; \p\w0. p < n; ss0 \ list n A; \p w0 \ stable r step ss0 p \ \ iter f step ss0 w0 = (ss',w') \ ss' \ list n A \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" (*<*) (is "PROP ?P") proof - interpret Semilat A r f by fact show "PROP ?P" apply(insert semilat) apply (unfold iter_def stables_def) apply (rule_tac P = "\(ss,w). ss \ list n A \ (\p w \ stable r step ss p) \ ss0 [\\<^sub>r] ss \ (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts) \ (\p\w. p < n)" and r = "{(ss',ss) . ss [\\<^sub>r] ss'} <*lex*> finite_psubset" in while_rule) \ \Invariant holds initially:\ apply (simp add:stables_def) \ \Invariant is preserved:\ apply(simp add: stables_def split_paired_all) apply(rename_tac ss w) apply(subgoal_tac "(SOME p. p \ w) \ w") prefer 2 apply (fast intro: someI) apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") prefer 2 apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply (erule listE_nth_in) apply blast apply blast apply (subst decomp_propa) apply blast apply simp apply (rule conjI) apply (rule merges_preserves_type) apply blast apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply (erule listE_nth_in) apply blast apply blast apply (rule conjI) apply clarify apply (blast intro!: stable_pres_lemma) apply (rule conjI) apply (blast intro!: merges_incr intro: le_list_trans) apply (rule conjI) apply clarsimp apply (blast intro!: merges_bounded_lemma) apply (blast dest!: boundedD) \ \Postcondition holds upon termination:\ apply(clarsimp simp add: stables_def split_paired_all) \ \Well-foundedness of the termination relation:\ apply (rule wf_lex_prod) apply (insert orderI [THEN acc_le_listI]) apply (simp only: acc_def lesssub_def) apply (rule wf_finite_psubset) \ \Loop decreases along termination relation:\ apply(simp add: stables_def split_paired_all) apply(rename_tac ss w) apply(subgoal_tac "(SOME p. p \ w) \ w") prefer 2 apply (fast intro: someI) apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") prefer 2 apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply (erule listE_nth_in) apply blast apply blast apply (subst decomp_propa) apply blast apply clarify apply (simp del: listE_length add: lex_prod_def finite_psubset_def bounded_nat_set_is_finite) apply (rule termination_lemma) apply (rule assms) apply assumption+ defer apply assumption apply clarsimp done qed (*>*) lemma kildall_properties: assumes "Semilat A r f" shows "\ acc r; pres_type step n A; mono r step n A; bounded step n; ss0 \ list n A \ \ kildall r f step ss0 \ list n A \ stables r step (kildall r f step ss0) \ ss0 [\\<^sub>r] kildall r f step ss0 \ (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ kildall r f step ss0 [\\<^sub>r] ts)" (*<*) (is "PROP ?P") proof - interpret Semilat A r f by fact show "PROP ?P" apply (unfold kildall_def) apply(case_tac "iter f step ss0 (unstables r step ss0)") apply(simp) apply (rule iter_properties) apply (simp_all add: unstables_def stable_def) apply (rule Semilat.intro) apply (rule semilat) done qed lemma is_bcv_kildall: assumes "Semilat A r f" shows "\ acc r; top r T; pres_type step n A; bounded step n; mono r step n A \ \ is_bcv r T step n A (kildall r f step)" (is "PROP ?P") proof - interpret Semilat A r f by fact show "PROP ?P" apply(unfold is_bcv_def wt_step_def) apply(insert \Semilat A r f\ semilat kildall_properties[of A]) apply(simp add:stables_def) apply clarify apply(subgoal_tac "kildall r f step \s\<^sub>0 \ list n A") prefer 2 apply (simp(no_asm_simp)) apply (rule iffI) apply (rule_tac x = "kildall r f step \s\<^sub>0" in bexI) apply (rule conjI) apply (blast) apply (simp (no_asm_simp)) apply(assumption) apply clarify apply(subgoal_tac "kildall r f step \s\<^sub>0!p <=_r \s!p") apply simp apply (blast intro!: le_listD less_lengthI) done qed (*>*) end diff --git a/thys/LTL_to_GBA/LTL_to_GBA.thy b/thys/LTL_to_GBA/LTL_to_GBA.thy --- a/thys/LTL_to_GBA/LTL_to_GBA.thy +++ b/thys/LTL_to_GBA/LTL_to_GBA.thy @@ -1,3233 +1,3224 @@ section \LTL to GBA translation\ (* Author: Alexander Schimpf Modified by Peter Lammich **) theory LTL_to_GBA imports CAVA_Base.CAVA_Base LTL.LTL CAVA_Automata.Automata begin subsection \Statistics\ code_printing code_module Gerth_Statistics \ (SML) \ structure Gerth_Statistics = struct val active = Unsynchronized.ref false val data = Unsynchronized.ref (0,0,0) fun is_active () = !active fun set_data num_states num_init num_acc = ( active := true; data := (num_states, num_init, num_acc) ) fun to_string () = let val (num_states, num_init, num_acc) = !data in "Num states: " ^ IntInf.toString (num_states) ^ "\n" ^ " Num initial: " ^ IntInf.toString num_init ^ "\n" ^ " Num acc-classes: " ^ IntInf.toString num_acc ^ "\n" end val _ = Statistics.register_stat ("Gerth LTL_to_GBA",is_active,to_string) end \ code_reserved SML Gerth_Statistics consts stat_set_data_int :: "integer \ integer \ integer \ unit" code_printing constant stat_set_data_int \ (SML) "Gerth'_Statistics.set'_data" definition "stat_set_data ns ni na \ stat_set_data_int (integer_of_nat ns) (integer_of_nat ni) (integer_of_nat na)" lemma [autoref_rules]: "(stat_set_data,stat_set_data) \ nat_rel \ nat_rel \ nat_rel \ unit_rel" by auto abbreviation "stat_set_data_nres ns ni na \ RETURN (stat_set_data ns ni na)" lemma discard_stat_refine[refine]: "m1\m2 \ stat_set_data_nres ns ni na \ m1 \ m2" by simp_all subsection \Preliminaries\ text \Some very special lemmas for reasoning about the nres-monad\ lemma SPEC_rule_nested2: "\m \ SPEC P; \r1 r2. P (r1, r2) \ g (r1, r2) \ SPEC P\ \ m \ SPEC (\r'. g r' \ SPEC P)" by (simp add: pw_le_iff) blast lemma SPEC_rule_param2: assumes "f x \ SPEC (P x)" and "\r1 r2. (P x) (r1, r2) \ (P x') (r1, r2)" shows "f x \ SPEC (P x')" using assms by (simp add: pw_le_iff) lemma SPEC_rule_weak: assumes "f x \ SPEC (Q x)" and "f x \ SPEC (P x)" and "\r1 r2. \(Q x) (r1, r2); (P x) (r1, r2)\ \ (P x') (r1, r2)" shows "f x \ SPEC (P x')" using assms by (simp add: pw_le_iff) lemma SPEC_rule_weak_nested2: "\f \ SPEC Q; f \ SPEC P; \r1 r2. \Q (r1, r2); P (r1, r2)\ \ g (r1, r2) \ SPEC P\ \ f \ SPEC (\r'. g r' \ SPEC P)" by (simp add: pw_le_iff) blast subsection \Creation of States\ text \ In this section, the first part of the algorithm, which creates the states of the automaton, is formalized. \ (* FIXME: Abstraktion über node_name *) type_synonym node_name = nat type_synonym 'a frml = "'a ltlr" type_synonym 'a interprt = "'a set word" record 'a node = name :: node_name incoming :: "node_name set" new :: "'a frml set" old :: "'a frml set" "next" :: "'a frml set" context begin fun new1 where "new1 (\ and\<^sub>r \) = {\,\}" | "new1 (\ U\<^sub>r \) = {\}" | "new1 (\ R\<^sub>r \) = {\}" | "new1 (\ or\<^sub>r \) = {\}" | "new1 _ = {}" fun next1 where "next1 (X\<^sub>r \) = {\}" | "next1 (\ U\<^sub>r \) = {\ U\<^sub>r \}" | "next1 (\ R\<^sub>r \) = {\ R\<^sub>r \}" | "next1 _ = {}" fun new2 where "new2 (\ U\<^sub>r \) = {\}" | "new2 (\ R\<^sub>r \) = {\, \}" | "new2 (\ or\<^sub>r \) = {\}" | "new2 _ = {}" (* FIXME: Abstraction over concrete definition *) definition "expand_init \ 0" definition "expand_new_name \ Suc" lemma expand_new_name_expand_init: "expand_init < expand_new_name nm" by (auto simp add:expand_init_def expand_new_name_def) lemma expand_new_name_step[intro]: fixes n :: "'a node" shows "name n < expand_new_name (name n)" by (auto simp add:expand_new_name_def) lemma expand_new_name__less_zero[intro]: "0 < expand_new_name nm" proof - from expand_new_name_expand_init have "expand_init < expand_new_name nm" by auto then show ?thesis by (metis gr0I less_zeroE) qed abbreviation "upd_incoming_f n \ (\n'. if (old n' = old n \ next n' = next n) then n'\ incoming := incoming n \ incoming n' \ else n' )" definition "upd_incoming n ns \ ((upd_incoming_f n) ` ns)" lemma upd_incoming__elem: assumes "nd\upd_incoming n ns" shows "nd\ns \ (\nd'\ns. nd = nd'\ incoming := incoming n \ incoming nd' \ \ old nd' = old n \ next nd' = next n)" proof - obtain nd' where "nd'\ns" and nd_eq: "nd = upd_incoming_f n nd'" using assms unfolding upd_incoming_def by blast then show ?thesis by auto qed lemma upd_incoming__ident_node: assumes "nd\upd_incoming n ns" and "nd\ns" shows "incoming n \ incoming nd \ \ (old nd = old n \ next nd = next n)" proof (rule ccontr) assume "\ ?thesis" then have nsubset: "\ (incoming n \ incoming nd)" and old_next_eq: "old nd = old n \ next nd = next n" by auto moreover obtain nd' where "nd'\ns" and nd_eq: "nd = upd_incoming_f n nd'" using assms unfolding upd_incoming_def by blast { assume "old nd' = old n \ next nd' = next n" with nd_eq have "nd = nd'\incoming := incoming n \ incoming nd'\" by auto with nsubset have "False" by auto } moreover { assume "\ (old nd' = old n \ next nd' = next n)" with nd_eq old_next_eq have "False" by auto } ultimately show "False" by fast qed lemma upd_incoming__ident: assumes "\n\ns. P n" and "\n. \n\ns; P n\ \ (\v. P (n\ incoming := v \))" shows "\n\upd_incoming n ns. P n" proof fix nd v let ?f = "\n'. if (old n' = old n \ next n' = next n) then n'\ incoming := incoming n \ incoming n' \ else n'" assume "nd\upd_incoming n ns" then obtain nd' where "nd'\ns" and nd_eq: "nd = ?f nd'" unfolding upd_incoming_def by blast { assume "old nd' = old n \ next nd' = next n" then obtain v where "nd = nd'\ incoming := v \" using nd_eq by auto with assms \nd'\ns\ have "P nd" by auto } then show "P nd" using nd_eq \nd'\ns\ assms by auto qed lemma name_upd_incoming_f[simp]: "name (upd_incoming_f n x) = name x" by auto lemma name_upd_incoming[simp]: "name ` (upd_incoming n ns) = name ` ns" (is "?lhs = ?rhs") unfolding upd_incoming_def proof safe fix x assume "x\ns" then have "upd_incoming_f n x \ (\n'. local.upd_incoming_f n n') ` ns" by auto then have "name (upd_incoming_f n x) \ name ` (\n'. local.upd_incoming_f n n') ` ns" by blast then show "name x \ name ` (\n'. local.upd_incoming_f n n') ` ns" by simp qed auto abbreviation expand_body where "expand_body \ (\expand (n,ns). if new n = {} then ( if (\n'\ns. old n' = old n \ next n' = next n) then RETURN (name n, upd_incoming n ns) else expand ( \ name=expand_new_name (name n), incoming={name n}, new=next n, old={}, next={} \, {n} \ ns) ) else do { \ \ SPEC (\x. x\(new n)); let n = n\ new := new n - {\} \; if (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q)) then (if (not\<^sub>r \) \ old n then RETURN (name n, ns) else expand (n\ old := {\} \ old n \, ns)) else if \ = true\<^sub>r then expand (n\ old := {\} \ old n \, ns) else if \ = false\<^sub>r then RETURN (name n, ns) else if (\\ \. (\ = \ and\<^sub>r \) \ (\ = X\<^sub>r \)) then expand ( n\ new := new1 \ \ new n, old := {\} \ old n, next := next1 \ \ next n \, ns) else do { (nm, nds) \ expand ( n\ new := new1 \ \ new n, old := {\} \ old n, next := next1 \ \ next n \, ns); expand (n\ name := nm, new := new2 \ \ new n, old := {\} \ old n \, nds) } } )" lemma expand_body_mono: "trimono expand_body" by refine_mono definition expand :: "('a node \ ('a node set)) \ (node_name \ 'a node set) nres" where "expand \ REC expand_body" lemma REC_rule_old: (* TODO: Adapt proofs below, have fun with subgoal 27 ...*) fixes x::"'x" assumes M: "trimono body" assumes I0: "\ x" assumes IS: "\f x. \ \x. \ x \ f x \ M x; \ x; f \ REC body \ \ body f x \ M x" shows "REC body x \ M x" by (rule REC_rule_arb[where pre="\_. \" and M="\_. M", OF assms]) lemma expand_rec_rule: assumes I0: "\ x" assumes IS: "\f x. \ \x. f x \ expand x; \x. \ x \ f x \ M x; \ x \ \ expand_body f x \ M x" shows "expand x \ M x" unfolding expand_def apply (rule REC_rule_old[where \="\"]) apply (rule expand_body_mono) apply (rule I0) apply (rule IS[unfolded expand_def]) apply (blast dest: le_funD) apply blast apply blast done abbreviation "expand_assm_incoming n_ns \ (\nm\incoming (fst n_ns). name (fst n_ns) > nm) \ 0 < name (fst n_ns) \ (\q\snd n_ns. name (fst n_ns) > name q \ (\nm\incoming q. name (fst n_ns) > nm))" abbreviation "expand_rslt_incoming nm_nds \ (\q\snd nm_nds. (fst nm_nds > name q \ (\nm'\incoming q. fst nm_nds > nm')))" abbreviation "expand_rslt_name n_ns nm_nds \ (name (fst n_ns) \ fst nm_nds \ name ` (snd n_ns) \ name ` (snd nm_nds)) \ name ` (snd nm_nds) = name ` (snd n_ns) \ name ` {nd\snd nm_nds. name nd \ name (fst n_ns)}" abbreviation "expand_name_ident nds \ (\q\nds. \!q'\nds. name q = name q')" abbreviation "expand_assm_exist \ n_ns \ {\. \\. \ U\<^sub>r \ \ old (fst n_ns) \ \ \\<^sub>r \} \ new (fst n_ns) \ old (fst n_ns) \ (\\\new (fst n_ns). \ \\<^sub>r \) \ (\\\old (fst n_ns). \ \\<^sub>r \) \ (\\\next (fst n_ns). \ \\<^sub>r X\<^sub>r \)" abbreviation "expand_rslt_exist__node_prop \ n nd \ incoming n \ incoming nd \ (\\\old nd. \ \\<^sub>r \) \ (\\\next nd. \ \\<^sub>r X\<^sub>r \) \ {\. \\. \ U\<^sub>r \ \ old nd \ \ \\<^sub>r \} \ old nd" abbreviation "expand_rslt_exist \ n_ns nm_nds \ (\nd\snd nm_nds. expand_rslt_exist__node_prop \ (fst n_ns) nd)" abbreviation "expand_rslt_exist_eq__node n nd \ name n = name nd \ old n = old nd \ next n = next nd \ incoming n \ incoming nd" abbreviation "expand_rslt_exist_eq n_ns nm_nds \ (\n\snd n_ns. \nd\snd nm_nds. expand_rslt_exist_eq__node n nd)" lemma expand_name_propag: assumes "expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. expand_rslt_incoming r \ expand_rslt_name n_ns r \ expand_name_ident (snd r))" (is "expand _ \ SPEC (?P n_ns)") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case prems: (1 _ _ n ns) then have Q: "?Q (n, ns)" by fast let ?nds = "upd_incoming n ns" from prems have "\q\?nds. name q < name n" by (rule_tac upd_incoming__ident) auto moreover have "\q\?nds. \nm'\incoming q. nm' < name n" (is "\q\_. ?sg q") proof fix q assume q_in:"q\?nds" show "?sg q" proof (cases "q\ns") case True with prems show ?thesis by auto next case False with upd_incoming__elem[OF q_in] obtain nd' where nd'_def:"nd'\ns \ q = nd'\incoming := incoming n \ incoming nd'\" by blast { fix nm' assume "nm'\incoming q" moreover { assume "nm'\incoming n" with Q have "nm' < name n" by auto } moreover { assume "nm'\incoming nd'" with Q nd'_def have "nm' < name n" by auto } ultimately have "nm' < name n" using nd'_def by auto } then show ?thesis by fast qed qed moreover have "expand_name_ident ?nds" proof (rule upd_incoming__ident, goal_cases) case 1 show ?case proof fix q assume "q\ns" with Q have "\!q'\ns. name q = name q'" by auto then obtain q' where "q'\ns" and "name q = name q'" and q'_all: "\q''\ns. name q' = name q'' \ q' = q''" by auto let ?q' = "upd_incoming_f n q'" have P_a: "?q'\?nds \ name q = name ?q'" using \q'\ns\ \name q = name q'\ q'_all unfolding upd_incoming_def by auto have P_all: "\q''\?nds. name ?q' = name q'' \ ?q' = q''" proof(clarify) fix q'' assume "q''\?nds" and q''_name_eq: "name ?q' = name q''" { assume "q''\ns" with upd_incoming__elem[OF \q''\?nds\] obtain nd'' where "nd''\ns" and q''_is: "q'' = nd''\incoming := incoming n \ incoming nd''\ \ old nd'' = old n \ next nd'' = next n" by auto then have "name nd'' = name q'" using q''_name_eq by (cases "old q' = old n \ next q' = next n") auto with \nd''\ns\ q'_all have "nd'' = q'" by auto then have "?q' = q''" using q''_is by simp } moreover { assume "q''\ns" moreover have "name q' = name q''" using q''_name_eq by (cases "old q' = old n \ next q' = next n") auto moreover then have "incoming n \ incoming q'' \ incoming q'' = incoming n \ incoming q''" by auto ultimately have "?q' = q''" using upd_incoming__ident_node[OF \q''\?nds\] q'_all by auto } ultimately show "?q' = q''" by fast qed show "\!q'\upd_incoming n ns. name q = name q'" proof(rule ex1I[of _ ?q'], goal_cases) case 1 then show ?case using P_a by simp next case 2 then show ?case using P_all unfolding P_a[THEN conjunct2, THEN sym] by blast qed qed qed simp ultimately show ?case using prems by auto next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems have Q: "?Q (n, ns)" by auto show ?case unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2[where P = "?P"], rule_tac step, goal_cases) case 1 with expand_new_name_step[of n] show ?case using Q by (auto elim: order_less_trans[rotated]) next case prems': (2 _ nds) then have "name ` ns \ name ` nds" by auto with expand_new_name_step[of n] show ?case using prems' by auto qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (2 nm nds) then have P_x: "(name n \ nm \ name ` ns \ name ` nds) \ name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" (is "_ \ ?nodes_eq nds ns (name n)") by auto show ?case proof (rule_tac SPEC_rule_param2[where P = "?P"], goal_cases) case 1 with prems show ?case by (rule_tac step) auto next case prems': (2 nm' nds') then have "\q\nds'. name q < nm' \ (\nm''\incoming q. nm'' < nm')" by auto moreover have "?nodes_eq nds ns (name n)" and "?nodes_eq nds' nds nm" and "name n \ nm" using prems' P_x by auto then have "?nodes_eq nds' ns (name n)" by auto then have "expand_rslt_name (n, ns) (nm', nds')" using prems' P_x subset_trans[of "name ` ns" "name ` nds"] by auto ultimately show ?case using prems' by auto qed qed qed lemmas expand_name_propag__incoming = SPEC_rule_conjunct1[OF expand_name_propag] lemmas expand_name_propag__name = SPEC_rule_conjunct1[OF SPEC_rule_conjunct2[OF expand_name_propag]] lemmas expand_name_propag__name_ident = SPEC_rule_conjunct2[OF SPEC_rule_conjunct2[OF expand_name_propag]] lemma expand_rslt_exist_eq: shows "expand n_ns \ SPEC (expand_rslt_exist_eq n_ns)" (is "_ \ SPEC (?P n_ns)") proof (rule_tac expand_rec_rule[where \="\_. True"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) let ?r = "(name n, upd_incoming n ns)" have "expand_rslt_exist_eq (n, ns) ?r" unfolding snd_conv proof fix n' assume "n'\ns" { assume "old n' = old n \ next n' = next n" with \n'\ns\ have "n'\ incoming := incoming n \ incoming n' \ \ upd_incoming n ns" unfolding upd_incoming_def by auto } moreover { assume "\ (old n' = old n \ next n' = next n)" with \n'\ns\ have "n' \ upd_incoming n ns" unfolding upd_incoming_def by auto } ultimately show "\nd\upd_incoming n ns. expand_rslt_exist_eq__node n' nd" by force qed with prems show ?case by auto next case prems: (2 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns) then have step: "\x. f x \ SPEC (?P x)" by simp show ?case proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case (2 nm nds) with prems have P_x: "?P (n, ns) (nm, nds)" by fast show ?case unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_param2[where P = "?P"], goal_cases) case 1 then show ?case by (rule_tac step) next case prems': (2 nm' nds') { fix n' assume "n'\ns" with P_x obtain nd where "nd\nds" and n'_split: "expand_rslt_exist_eq__node n' nd" by auto with prems' obtain nd' where "nd'\nds'" and "expand_rslt_exist_eq__node nd nd'" by auto then have "\nd'\nds'. expand_rslt_exist_eq__node n' nd'" using n'_split subset_trans[of "incoming n'"] by auto } then have "expand_rslt_exist_eq (n, ns) (nm', nds')" by auto with prems show ?case by auto qed qed qed lemma expand_prop_exist: "expand n_ns \ SPEC (\r. expand_assm_exist \ n_ns \ expand_rslt_exist \ n_ns r)" (is "_ \ SPEC (?P n_ns)") proof (rule_tac expand_rec_rule[where \="\_. True"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) let ?nds = "upd_incoming n ns" let ?r = "(name n, ?nds)" { assume Q: "expand_assm_exist \ (n, ns)" note \\n'\ns. old n' = old n \ next n' = next n\ then obtain n' where "n'\ns" and assm_eq: "old n' = old n \ next n' = next n" by auto let ?nd = "n'\ incoming := incoming n \ incoming n'\" have "?nd \ ?nds" using \n'\ns\ assm_eq unfolding upd_incoming_def by auto moreover have "incoming n \ incoming ?nd" by auto moreover have "expand_rslt_exist__node_prop \ n ?nd" using Q assm_eq \new n = {}\ by simp ultimately have "expand_rslt_exist \ (n, ns) ?r" unfolding fst_conv snd_conv by blast } with prems show ?case by auto next case prems: (2 f x n ns) then have step: "\x. f x \ SPEC (?P x)" and f_sup: "\x. f x \ expand x" by auto show ?case unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "expand_rslt_exist_eq"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq) next case 2 then show ?case by (rule_tac step) next case prems': (3 nm nds) then have "name ` ns \ name ` nds" by auto moreover { assume assm_ex: "expand_assm_exist \ (n, ns)" with prems' obtain nd where "nd\nds" and "expand_rslt_exist_eq__node n nd" by force+ then have "expand_rslt_exist__node_prop \ n nd" using assm_ex \new n = {}\ by auto then have "expand_rslt_exist \ (n, ns) (nm, nds)" using \nd\nds\ by auto } ultimately show ?case using expand_new_name_step[of n] prems' by auto qed next case prems: (3 f x n ns \) { assume "expand_assm_exist \ (n, ns)" with prems have "\ \\<^sub>r \" and "\ \\<^sub>r not\<^sub>r \" by (metis (no_types, lifting) fstI node.select_convs(4) node.surjective node.update_convs(3))+ then have False by simp } with prems show ?case by auto next case prems: (4 f x n ns \) then have goal_assms: "\ \ new n \ (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q))" and step: "\x. f x \ SPEC (?P x)" by simp_all show ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems' have "expand_rslt_exist \ (n, ns) (nm, nds)" by auto } then show ?case by auto qed next case prems: (5 f x n ns \) then have goal_assms: "\ \ new n \ \ = true\<^sub>r" and step: "\x. f x \ SPEC (?P x)" by simp_all show ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems' have "expand_rslt_exist \ (n, ns) (nm, nds)" by auto } then show ?case by auto qed next case prems: (6 f x n ns \) { assume "expand_assm_exist \ (n, ns)" with prems have "\ \\<^sub>r false\<^sub>r" by auto } with prems show ?case by auto next case prems: (7 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \)" and step: "\x. f x \ SPEC (?P x)" by simp_all show ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems' have "expand_rslt_exist \ (n, ns) (nm, nds)" by auto } then show ?case by auto qed next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have step: "\x. f x \ SPEC (?P x)" and f_sup: "\x. f x \ expand x" by auto let ?x1 = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old n, next := next1 \ \ next n\, ns)" let ?new1_assm_sel = "\\. (case \ of \ U\<^sub>r \ => \ | \ R\<^sub>r \ \ \ | \ or\<^sub>r \ \ \)" { assume new1_assm: "\ (\ \\<^sub>r (?new1_assm_sel \))" then have ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case prems': 1 then show ?case proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems'': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems'' have "expand_assm_exist \ ?x1" unfolding fst_conv proof (cases \, goal_cases) case \: (8 \ \) then have "\ \\<^sub>r \ U\<^sub>r \" by fast then have "\ \\<^sub>r \" and "\ \\<^sub>r X\<^sub>r (\ U\<^sub>r \)" using \ ltlr_expand_Until[of \ \ \] by auto with \ show ?case by auto next case \: (9 \ \) then have *: "\ \\<^sub>r \ R\<^sub>r \" by fast with \ have "\ \\<^sub>r \" and "\ \\<^sub>r X\<^sub>r (\ R\<^sub>r \)" using ltlr_expand_Release[of \ \ \] by auto with \ * show ?case by auto qed auto with prems'' have "expand_rslt_exist \ (n, ns) (nm, nds)" by force } with prems'' show ?case by auto qed next case prems': (2 nm nds) then have P_x: "?P (n, ns) (nm, nds)" by fast show ?case unfolding case_prod_unfold proof (rule_tac SPEC_rule_weak[where P = "?P" and Q = "expand_rslt_exist_eq"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq) next case 2 then show ?case by (rule_tac step) next case prems'': (3 nm' nds') { assume "expand_assm_exist \ (n, ns)" with P_x have "expand_rslt_exist \ (n, ns) (nm, nds)" by simp then obtain nd where nd: "nd\nds" "expand_rslt_exist__node_prop \ n nd" using goal_assms by auto with prems'' obtain nd' where "nd'\nds'" and "expand_rslt_exist_eq__node nd nd'" by force with nd have "expand_rslt_exist__node_prop \ n nd'" using subset_trans[of "incoming n" "incoming nd"] by auto then have "expand_rslt_exist \ (n,ns) (nm', nds')" using \nd'\nds'\ goal_assms by auto } then show ?case by fast qed qed } moreover { assume new1_assm: "\ \\<^sub>r (?new1_assm_sel \)" let ?x2f = "\(nm::node_name, nds::'a node set). ( n\new := new n - {\}, name := nm, new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old n\, nds)" have P_x: "f ?x1 \ SPEC (?P ?x1)" by (rule step) moreover { fix r :: "node_name \ 'a node set" let ?x2 = "?x2f r" assume assm: "(?P ?x1) r" have "f ?x2 \ SPEC (?P (n, ns))" unfolding case_prod_unfold fst_conv proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm' nds') { assume "expand_assm_exist \ (n, ns)" with new1_assm goal_assms have "expand_assm_exist \ ?x2" proof (cases r, cases \, goal_cases) case prems'': (9 _ _ \ \) then have *: "\ \\<^sub>r \ R\<^sub>r \" unfolding fst_conv by fast with ltlr_expand_Release[of \ \ \] have "\ \\<^sub>r \" by auto with prems'' * show ?case by auto qed auto with prems' have "expand_rslt_exist \ ?x2 (nm', nds')" unfolding case_prod_unfold fst_conv snd_conv by fast then have "expand_rslt_exist \ (n, ns) (nm', nds')" by (cases r, auto) } then show ?case by simp qed } then have "SPEC (?P ?x1) \ SPEC (\r. (case r of (nm, nds) => f (?x2f (nm, nds))) \ SPEC (?P (n, ns)))" using goal_assms by (rule_tac SPEC_rule) force finally have ?case unfolding case_prod_unfold \x = (n, ns)\ by simp } ultimately show ?case by fast qed text \Termination proof\ definition expand\<^sub>T :: "('a node \ ('a node set)) \ (node_name \ 'a node set) nres" where "expand\<^sub>T n_ns \ REC\<^sub>T expand_body n_ns" abbreviation "old_next_pair n \ (old n, next n)" abbreviation "old_next_limit \ \ Pow (subfrmlsr \) \ Pow (subfrmlsr \)" lemma old_next_limit_finite: "finite (old_next_limit \)" using subfrmlsr_finite by auto definition "expand_ord \ \ inv_image (finite_psupset (old_next_limit \) <*lex*> less_than) (\(n, ns). (old_next_pair ` ns, size_set (new n)))" lemma expand_ord_wf[simp]: "wf (expand_ord \)" using finite_psupset_wf[OF old_next_limit_finite] unfolding expand_ord_def by auto abbreviation "expand_inv_node \ n \ finite (new n) \ finite (old n) \ finite (next n) \ (new n) \ (old n) \ (next n) \ subfrmlsr \" abbreviation "expand_inv_result \ ns \ finite ns \ (\n'\ns. (new n') \ (old n') \ (next n') \ subfrmlsr \)" definition "expand_inv \ n_ns \ (case n_ns of (n, ns) \ expand_inv_node \ n \ expand_inv_result \ ns)" lemma new1_less_sum: "size_set (new1 \) < size_set {\}" proof (cases \) case (And_ltlr \ \) thus ?thesis by (cases "\ = \"; simp) qed (simp_all) lemma new2_less_sum: "size_set (new2 \) < size_set {\}" proof (cases \) case (Release_ltlr \ \) thus ?thesis by (cases "\ = \"; simp) qed (simp_all) lemma new1_finite[intro]: "finite (new1 \)" by (cases \) auto lemma new1_subset_frmls: "\ \ new1 \ \ \ \ subfrmlsr \" by (cases \) auto lemma new2_finite[intro]: "finite (new2 \)" by (cases \) auto lemma new2_subset_frmls: "\ \ new2 \ \ \ \ subfrmlsr \" by (cases \) auto lemma next1_finite[intro]: "finite (next1 \)" by (cases \) auto lemma next1_subset_frmls: "\ \ next1 \ \ \ \ subfrmlsr \" by (cases \) auto -text \Expanding the new-style lexorder\ -lemma expand_ord_iff: "((a,b),(c,d)) \ expand_ord \ - \ ((old_next_pair ` b, old_next_pair ` d) \ finite_psupset (old_next_limit \) - \ old_next_pair ` b = old_next_pair ` d \ (\x\new a. Suc (2 * size x)) < (\x\new c. Suc (2 * size x)))" -apply (simp add: expand_ord_def) - apply (auto simp: finite_psupset_def) - apply (metis (mono_tags, lifting) image_iff)+ - done - lemma expand_inv_impl[intro!]: assumes "expand_inv \ (n, ns)" and newn: "\ \ new n" and "old_next_pair ` ns \ old_next_pair ` ns'" and "expand_inv_result \ ns'" and "(n' = n\ new := new n - {\}, old := {\} \ old n \) \ (n' = n\ new := new1 \ \ (new n - {\}), old := {\} \ old n, next := next1 \ \ next n \) \ (n' = n\ name := nm, new := new2 \ \ (new n - {\}), old := {\} \ old n \)" (is "?case1 \ ?case2 \ ?case3") shows "((n', ns'), (n, ns))\expand_ord \ \ expand_inv \ (n', ns')" (is "?concl1 \ ?concl2") proof from assms consider ?case1 | ?case2 | ?case3 by blast then show ?concl1 proof cases case n'def: 1 with assms show ?thesis - unfolding expand_ord_iff expand_inv_def finite_psupset_def + unfolding expand_ord_def expand_inv_def finite_psupset_def apply (cases "old_next_pair ` ns \ old_next_pair ` ns'") - apply (simp_all add: psubset_eq) - apply blast - apply (metis (no_types, lifting) add_Suc diff_Suc_less sum.remove sum_diff1_nat zero_less_Suc) + apply simp_all + apply auto [1] + apply (metis (no_types, lifting) add_Suc diff_Suc_less psubsetI sum.remove sum_diff1_nat zero_less_Suc) done next case n'def: 2 have \innew: "\ \ new n" and fin_new: "finite (new n)" using assms unfolding expand_inv_def by auto then have "size_set (new n - {\}) = size_set (new n) - size_set {\}" using size_set_diff by fastforce moreover from fin_new sum_Un_nat[OF new1_finite _, of "new n - {\}" size \] have "size_set (new n') \ size_set (new1 \) + size_set (new n - {\})" unfolding n'def by (simp add: new1_finite sum_Un_nat) moreover have sum_leq: "size_set (new n) \ size_set {\}" using \innew sum_mono2[OF fin_new, of "{\}"] by blast ultimately have "size_set (new n') < size_set (new n)" using new1_less_sum[of \] by auto with assms show ?thesis - unfolding expand_ord_iff finite_psupset_def by auto + unfolding expand_ord_def finite_psupset_def by auto next case n'def: 3 have \innew: "\ \ new n" and fin_new: "finite (new n)" using assms unfolding expand_inv_def by auto from \innew sum_diff1_nat[of size "new n" \] have "size_set (new n - {\}) = size_set (new n) - size_set {\}" using size_set_diff[of "new n" "{\}"] by fastforce moreover from fin_new sum_Un_nat[OF new2_finite _, of "new n - {\}" size \] have "size_set (new n') \ size_set (new2 \) + size_set (new n - {\})" unfolding n'def by (simp add: new2_finite sum_Un_nat) moreover have sum_leq:"size_set (new n) \ size_set {\}" using \innew sum_mono2[OF fin_new, of "{\}"] by blast ultimately have "size_set (new n') < size_set (new n)" using new2_less_sum[of \] sum_leq by auto with assms show ?thesis - unfolding expand_ord_iff finite_psupset_def by auto + unfolding expand_ord_def finite_psupset_def by auto qed next have "new1 \ \ subfrmlsr \" and "new2 \ \ subfrmlsr \" and "next1 \ \ subfrmlsr \" using assms subfrmlsr_subset[OF new1_subset_frmls[of _ \]] subfrmlsr_subset[of \ \, OF rev_subsetD[of _ "new n"]] subfrmlsr_subset[OF new2_subset_frmls[of _ \]] subfrmlsr_subset[OF next1_subset_frmls[of _ \]] unfolding expand_inv_def (* This proof is merely a speed optimization. A single force+ does the job, but takes very long *) apply - apply (clarsimp split: prod.splits) apply (metis in_mono new1_subset_frmls) apply (clarsimp split: prod.splits) apply (metis new2_subset_frmls rev_subsetD) apply (clarsimp split: prod.splits) apply (metis in_mono next1_subset_frmls) done with assms show ?concl2 unfolding expand_inv_def by auto qed lemma expand_term_prop_help: assumes "((n', ns'), (n, ns))\expand_ord \ \ expand_inv \ (n', ns')" and assm_rule: "\expand_inv \ (n', ns'); ((n', ns'), n, ns) \ expand_ord \\ \ f (n', ns') \ SPEC P" shows "f (n', ns') \ SPEC P" using assms by (rule_tac assm_rule) auto lemma expand_inv_upd_incoming: assumes "expand_inv \ (n, ns)" shows "expand_inv_result \ (upd_incoming n ns)" using assms unfolding expand_inv_def upd_incoming_def by force lemma upd_incoming_eq_old_next_pair: "old_next_pair ` ns = old_next_pair ` (upd_incoming n ns)" (is "?A = ?B") proof show "?A \ ?B" proof fix x let ?f = "\n'. n'\incoming := incoming n \ incoming n'\" assume "x \ ?A" then obtain n' where "n' \ ns" and xeq: "x = (old n', next n')" by auto have "x \ (old_next_pair ` (\n'. n'\incoming := incoming n \ incoming n'\) ` (ns \ {n' \ ns. old n' = old n \ next n' = next n})) \ (old_next_pair ` (ns \ {n' \ ns. old n' \ old n \ next n' \ next n}))" proof (cases "old n' = old n \ next n' = next n") case True with \n' \ ns\ have "?f n' \ ?f ` (ns \ {n'\ns. old n' = old n \ next n' = next n})" (is "_ \ ?C") by auto then have "old_next_pair (?f n') \ old_next_pair ` ?C" by (rule_tac imageI) auto with xeq have "x\old_next_pair ` ?C" by auto then show ?thesis by blast next case False with \n' \ ns\ xeq have "x \ old_next_pair ` (ns \ {n'\ns. old n' \ old n \ next n' \ next n})" by auto then show ?thesis by blast qed then show "x \ ?B" using \x \ ?A\ unfolding upd_incoming_def by auto qed show "?B \ ?A" unfolding upd_incoming_def by (force intro:imageI) qed lemma expand_term_prop: "expand_inv \ n_ns \ expand\<^sub>T n_ns \ SPEC (\(_, nds). old_next_pair ` snd n_ns\old_next_pair `nds \ expand_inv_result \ nds)" (is "_ \ _ \ SPEC (?P n_ns)") unfolding expand\<^sub>T_def apply (rule_tac RECT_rule[where pre="\(n, ns). expand_inv \ (n, ns)" and V="expand_ord \"]) apply (refine_mono) apply simp apply simp proof (intro refine_vcg, goal_cases) case prems: (1 _ _ n ns) have "old_next_pair ` ns \ old_next_pair ` (upd_incoming n ns)" by (rule equalityD1[OF upd_incoming_eq_old_next_pair]) with prems show ?case using expand_inv_upd_incoming[of \ n ns] by auto next case prems: (2 expand x n ns) let ?n' = "\ name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\" let ?ns' = "{n} \ ns" from prems have SPEC_sub:"SPEC (?P (?n', ?ns')) \ SPEC (?P x)" by (rule_tac SPEC_rule) auto from prems have "old_next_pair n \ old_next_pair ` ns" by auto then have "old_next_pair ` ns \ old_next_pair ` (insert n ns)" by auto moreover from prems have "expand_inv \ (n, ns)" by simp ultimately have "((?n', ?ns'), (n, ns)) \ expand_ord \" - by (auto simp add: expand_ord_iff finite_psupset_def expand_inv_def) + by (auto simp add: expand_ord_def finite_psupset_def expand_inv_def) moreover from prems have "expand_inv \ (?n', ?ns')" unfolding expand_inv_def by auto ultimately have "expand (?n', ?ns') \ SPEC (?P (?n', ?ns'))" using prems by fast with SPEC_sub show ?case by (rule_tac order_trans) fast+ next case 3 then show ?case by (auto simp add:expand_inv_def) next case 4 then show ?case apply (rule_tac expand_term_prop_help[OF expand_inv_impl]) apply (simp add: expand_inv_def)+ apply force done next case 5 then show ?case apply (rule_tac expand_term_prop_help[OF expand_inv_impl]) apply (simp add: expand_inv_def)+ apply force done next case 6 then show ?case by (simp add: expand_inv_def) next case 7 then show ?case apply (rule_tac expand_term_prop_help[OF expand_inv_impl]) apply (simp add: expand_inv_def)+ apply force done next case prems: (8 f x a b xa) let ?n' = "a\ new := new1 xa \ (new a - {xa}), old := insert xa (old a), next := next1 xa \ next a\" let ?n'' = "\nm. a\ name := nm, new := new2 xa \ (new a - {xa}), old := insert xa (old a)\" have step:"((?n', b), (a, b)) \ expand_ord \ \ expand_inv \ (?n', b)" using prems by (rule_tac expand_inv_impl) (auto simp add: expand_inv_def) with prems have assm1: "f (?n', b) \ SPEC (?P (a, b))" by auto moreover { fix nm::node_name and nds::"'a node set" assume assm1: "old_next_pair ` b \ old_next_pair ` nds" and assm2: "expand_inv_result \ nds" with prems step have "((?n'' nm, nds), (a, b)) \ expand_ord \ \ expand_inv \ (?n'' nm, nds)" by (rule_tac expand_inv_impl) auto with prems have "f (?n'' nm, nds) \ SPEC (?P (?n'' nm, nds))" by auto moreover have "SPEC (?P (?n'' nm, nds)) \ SPEC (?P (a, b))" using assm2 subset_trans[OF assm1] by auto ultimately have "f (?n'' nm, nds) \ SPEC (?P (a, b))" by (rule order_trans) } then have assm2: "SPEC (?P (a, b)) \ SPEC (\r. (case r of (nm, nds) \ f (?n'' nm, nds)) \ SPEC (?P (a, b)))" by (rule_tac SPEC_rule) auto from prems order_trans[OF assm1 assm2] show ?case by auto qed lemma expand_eq_expand\<^sub>T: assumes inv: "expand_inv \ n_ns" shows "expand\<^sub>T n_ns = expand n_ns" unfolding expand\<^sub>T_def expand_def apply (rule RECT_eq_REC) unfolding expand\<^sub>T_def[symmetric] using expand_term_prop[OF inv] apply auto done lemma expand_nofail: assumes inv: "expand_inv \ n_ns" shows "nofail (expand\<^sub>T n_ns)" using expand_term_prop[OF inv] by (simp add: pw_le_iff) lemma [intro!]: "expand_inv \ ( \ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \, {})" by (auto simp add: expand_inv_def) definition create_graph :: "'a frml \ 'a node set nres" where "create_graph \ \ do { (_, nds) \ expand ( \ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \::'a node, {}::'a node set); RETURN nds }" (* "expand_eq_expand\<^sub>T" *) definition create_graph\<^sub>T :: "'a frml \ 'a node set nres" where "create_graph\<^sub>T \ \ do { (_, nds) \ expand\<^sub>T ( \ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \::'a node, {}::'a node set); RETURN nds }" lemma create_graph_eq_create_graph\<^sub>T: "create_graph \ = create_graph\<^sub>T \" unfolding create_graph\<^sub>T_def create_graph_def unfolding eq_iff proof (standard, goal_cases) case 1 then show ?case by refine_mono (unfold expand_def expand\<^sub>T_def, rule REC_le_RECT) next case 2 then show ?case by (refine_mono, rule expand_eq_expand\<^sub>T[unfolded eq_iff, THEN conjunct1]) auto qed lemma create_graph_finite: "create_graph \ \ SPEC finite" unfolding create_graph_def expand_def apply (intro refine_vcg) apply (rule_tac order_trans) apply (rule_tac REC_le_RECT) apply (fold expand\<^sub>T_def) apply (rule_tac order_trans[OF expand_term_prop]) apply auto[1] apply (rule_tac SPEC_rule) apply auto done lemma create_graph_nofail: "nofail (create_graph \)" by (rule_tac pwD1[OF create_graph_finite]) auto abbreviation "create_graph_rslt_exist \ nds \ \nd\nds. expand_init\incoming nd \ (\\\old nd. \ \\<^sub>r \) \ (\\\next nd. \ \\<^sub>r X\<^sub>r \) \ {\. \\. \ U\<^sub>r \ \ old nd \ \ \\<^sub>r \} \ old nd" lemma L4_7: assumes "\ \\<^sub>r \" shows "create_graph \ \ SPEC (create_graph_rslt_exist \)" using assms unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_prop_exist) ( auto simp add:expand_new_name_expand_init) lemma expand_incoming_name_exist: assumes "name (fst n_ns) > expand_init \ (\nm\incoming (fst n_ns). nm \ expand_init \ nm\name ` (snd n_ns)) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") and "\nd\snd n_ns. name nd > expand_init \ (\nm\incoming nd. nm \ expand_init \ nm\name ` (snd n_ns))" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\nm_nds. ?P (snd nm_nds))" using assms apply (rule_tac expand_rec_rule[where \="\n_ns. ?Q n_ns \ ?P (snd n_ns)"]) apply simp apply (intro refine_vcg) proof goal_cases case (1 f x n ns) then show ?case proof (simp, clarify, goal_cases) case prems: (1 _ _ nd) { assume "nd\ns" with prems have ?case by auto } moreover { assume "nd\ns" with upd_incoming__elem[OF \nd\upd_incoming n ns\] obtain nd' where "nd'\ns" and "nd = nd'\incoming := incoming n \ incoming nd'\ \ old nd' = old n \ next nd' = next n" by auto with prems have ?case by auto } ultimately show ?case by fast qed next case (2 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x. ?P (snd x))" and QP: "?Q (n, ns) \ ?P ns" and f_sup: "\x. f x \ expand x" by auto show ?case unfolding \x = (n, ns)\ using QP expand_new_name_expand_init proof (rule_tac step, goal_cases) case prems: 1 then have name_less: "name n < expand_new_name (name n)" by auto moreover from prems name_less have "\nm\incoming n. nm < expand_new_name (name n)" by auto moreover from prems name_less have **: "\q\ns. name q < expand_new_name (name n) \ (\nm\incoming q. nm < expand_new_name (name n))" by force moreover from QP have "\!q'. (q' = n \ q' \ ns) \ name n = name q'" by force moreover have "\q\ns. \!q'. (q' = n \ q' \ ns) \ name q = name q' " using QP by auto ultimately show ?case using prems by simp qed next case 3 then show ?case by simp next case 4 then show ?case by simp next case 5 then show ?case by simp next case 6 then show ?case by simp next case 7 then show ?case by simp next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto with prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\)\, ns)" let ?props = "\x r. expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\x r. ?P (snd r)"], rule_tac step) auto next case (3 nm nds) then show ?case proof (rule_tac SPEC_rule_weak[where P = "\x r. ?P (snd r)" and Q = "\x r. expand_rslt_exist_eq x r \ ?props x r"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) force next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\x r. ?P (snd r)"], rule_tac step) force next case (3 nm' nds') then show ?case by simp qed qed qed lemma create_graph__incoming_name_exist: "create_graph \ \ SPEC (\nds. \nd\nds. expand_init < name nd \ (\nm\incoming nd. nm \ expand_init \ nm\name ` nds))" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_incoming_name_exist) ( auto simp add:expand_new_name_expand_init) abbreviation "expand_rslt_all__ex_equiv \ nd nds \ (\nd'\nds. name nd \ incoming nd' \ (\\\old nd'. suffix 1 \ \\<^sub>r \) \ (\\\next nd'. suffix 1 \ \\<^sub>r X\<^sub>r \) \ {\. \\. \ U\<^sub>r \ \ old nd' \ suffix 1 \ \\<^sub>r \} \ old nd')" abbreviation "expand_rslt_all \ n_ns nm_nds \ (\nd\snd nm_nds. name nd \ name ` (snd n_ns) \ (\\\old nd. \ \\<^sub>r \) \ (\\\next nd. \ \\<^sub>r X\<^sub>r \) \ expand_rslt_all__ex_equiv \ nd (snd nm_nds))" lemma expand_prop_all: assumes "expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (expand_rslt_all \ n_ns)" (is "_ \ SPEC (?P n_ns)") using assms apply (rule_tac expand_rec_rule[where \="?Q"]) apply simp apply (intro refine_vcg) proof goal_cases case 1 then show ?case by (simp, rule_tac upd_incoming__ident) simp_all next case (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" and Q: "?Q (n, ns)" and f_sup: "\x. f x \ expand x" by auto let ?x = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" from Q have name_le: "name n < expand_new_name (name n)" by auto show ?case unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\p r. (expand_assm_exist (suffix 1 \) ?x \ expand_rslt_exist (suffix 1 \) ?x r) \ expand_rslt_exist_eq p r \ (expand_name_ident (snd r))"], goal_cases) case 1 then show ?case proof (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_prop_exist, rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident, goal_cases) case 1 then show ?case using Q name_le by force qed next case 2 then show ?case using Q name_le by (rule_tac step) force next case prems: (3 nm nds) then obtain n' where "n'\nds" and eq_node: "expand_rslt_exist_eq__node n n'" by auto with prems have ex1_name: "\!q\nds. name n = name q" by auto then have nds_eq: "nds = {n'} \ {x \ nds. name n \ name x}" using eq_node \n'\nds\ by blast have name_notin: "name n \ name ` ns" using Q by auto from prems have P_x: "expand_rslt_all \ ?x (nm, nds)" by fast show ?case unfolding snd_conv proof clarify fix nd assume "nd \ nds" and name_img: "name nd \ name ` ns" and nd_old_equiv: "\\\old nd. \ \\<^sub>r \" and nd_next_equiv: "\\\next nd. \ \\<^sub>r X\<^sub>r \" show "expand_rslt_all__ex_equiv \ nd nds" proof (cases "name nd = name n") case True with nds_eq eq_node \nd\nds\ have "nd = n'" by auto with prems(1)[THEN conjunct1, simplified] nd_old_equiv nd_next_equiv eq_node show ?thesis by simp next case False with name_img \nd \ nds\ nd_old_equiv nd_next_equiv P_x show ?thesis by simp qed qed qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC (?P x)" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "?P"], rule_tac step) auto next case prems: (3 nm nds) then show ?case proof (rule_tac SPEC_rule_weak[where P = "?P" and Q = "\x r. expand_rslt_exist_eq x r \ ?props x r"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) auto next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "?P"], rule_tac step) auto next case prems': (3 nm' nds') then have P_x: "?P (n, ns) (nm, nds)" and P_x': "?P (n, nds) (nm', nds')" by simp_all show ?case unfolding snd_conv proof clarify fix nd assume "nd\nds'" and name_nd_notin: "name nd \ name ` ns" and old_equiv: "\\\old nd. \ \\<^sub>r \" and next_equiv: "\\\next nd. \ \\<^sub>r X\<^sub>r \" show "expand_rslt_all__ex_equiv \ nd nds'" proof (cases "name nd \ name ` nds") case True then obtain n' where "n' \ nds" and "name nd = name n'" by auto with prems' obtain nd' where "nd'\nds'" and nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto moreover from prems' have "\q\nds'. \!q'\nds'. name q = name q'" by simp ultimately have "nd' = nd" using \name nd = name n'\ \nd \ nds'\ by auto with nd'_eq have n'_eq: "expand_rslt_exist_eq__node n' nd" by simp then have "name n'\name ` ns" and "\\\old n'. \ \\<^sub>r \" and "\\\next n'. \ \\<^sub>r X\<^sub>r \" using name_nd_notin old_equiv next_equiv \n' \ nds\ by auto then have "expand_rslt_all__ex_equiv \ n' nds" (is "\nd'\nds. ?sthm n' nd'") using P_x \n' \ nds\ unfolding snd_conv by blast then obtain sucnd where sucnd: "sucnd\nds" and sthm: "?sthm n' sucnd" by blast moreover from prems' sucnd sthm obtain sucnd' where "sucnd'\nds'" and sucnd'_eq: "expand_rslt_exist_eq__node sucnd sucnd'" by auto ultimately have "?sthm n' sucnd'" by auto then show ?thesis using \sucnd' \ nds'\ unfolding \name nd = name n'\ by blast next case False with \nd \ nds'\ P_x' old_equiv next_equiv show ?thesis unfolding snd_conv by blast qed qed qed qed qed abbreviation "create_graph_rslt_all \ nds \ \q\nds. (\\\old q. \ \\<^sub>r \) \ (\\\next q. \ \\<^sub>r X\<^sub>r \) \ (\q'\nds. name q \ incoming q' \ (\\\old q'. suffix 1 \ \\<^sub>r \) \ (\\\next q'. suffix 1 \ \\<^sub>r X\<^sub>r \) \ {\. \\. \ U\<^sub>r \ \ old q' \ suffix 1 \ \\<^sub>r \} \ old q')" lemma L4_5: "create_graph \ \ SPEC (create_graph_rslt_all \)" unfolding create_graph_def apply (refine_vcg expand_prop_all) apply (auto simp add:expand_new_name_expand_init) done subsection \Creation of GBA\ text \This section formalizes the second part of the algorithm, that creates the actual generalized B\"uchi automata from the set of nodes.\ definition create_gba_from_nodes :: "'a frml \ 'a node set \ ('a node, 'a set) gba_rec" where "create_gba_from_nodes \ qs \ \ g_V = qs, g_E = {(q, q'). q\qs \ q'\qs \ name q\incoming q'}, g_V0 = {q\qs. expand_init\incoming q}, gbg_F = {{q\qs. \ U\<^sub>r \\old q \ \\old q}|\ \. \ U\<^sub>r \ \ subfrmlsr \}, gba_L = \q l. q\qs \ {p. prop\<^sub>r(p)\old q}\l \ {p. nprop\<^sub>r(p)\old q} \ l = {} \" end locale create_gba_from_nodes_precond = fixes \ :: "'a ltlr" fixes qs :: "'a node set" assumes res: "inres (create_graph \) qs" begin lemma finite_qs[simp, intro!]: "finite qs" using res create_graph_finite by (auto simp add: pw_le_iff) lemma create_gba_from_nodes__invar: "gba (create_gba_from_nodes \ qs)" using [[simproc finite_Collect]] apply unfold_locales apply (auto intro!: finite_vimageI subfrmlsr_finite injI simp: create_gba_from_nodes_def) done sublocale gba: gba "create_gba_from_nodes \ qs" by (rule create_gba_from_nodes__invar) lemma create_gba_from_nodes__fin: "finite (g_V (create_gba_from_nodes \ qs))" unfolding create_gba_from_nodes_def by auto lemma create_gba_from_nodes__ipath: "ipath gba.E r \ (\i. r i \qs \ name (r i)\incoming (r (Suc i)))" unfolding create_gba_from_nodes_def ipath_def by auto lemma create_gba_from_nodes__is_run: "gba.is_run r \ expand_init \ incoming (r 0) \ (\i. r i\qs \ name (r i)\incoming (r (Suc i)))" unfolding gba.is_run_def apply (simp add: create_gba_from_nodes__ipath) apply (auto simp: create_gba_from_nodes_def) done context begin abbreviation "auto_run_j j \ q \ (\\\old q. suffix j \ \\<^sub>r \) \ (\\\next q. suffix j \ \\<^sub>r X\<^sub>r \) \ {\. \\. \ U\<^sub>r \ \ old q \ suffix j \ \\<^sub>r \} \ old q" fun auto_run :: "['a interprt, 'a node set] \ 'a node word" where "auto_run \ nds 0 = (SOME q. q\nds \ expand_init \ incoming q \ auto_run_j 0 \ q)" | "auto_run \ nds (Suc k) = (SOME q'. q'\nds \ name (auto_run \ nds k) \ incoming q' \ auto_run_j (Suc k) \ q')" (* TODO: Remove. Only special instance of more generic principle lemma create_graph_comb: assumes "\x. inres (create_graph \) x \ P x" shows "create_graph \\SPEC P" using create_graph_nofail assms by (auto simp add: pw_le_iff refine_pw_simps) *) lemma run_propag_on_create_graph: assumes "ipath gba.E \" shows "\ k\qs \ name (\ k)\incoming (\ (k+1))" using assms by (auto simp: create_gba_from_nodes__ipath) lemma expand_false_propag: assumes "false\<^sub>r \ old (fst n_ns) \ (\nd\snd n_ns. false\<^sub>r \ old nd)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\nm_nds. \nd\snd nm_nds. false\<^sub>r \ old nd)" using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case 1 then show ?case by (simp, rule_tac upd_incoming__ident) auto next case 8 then show ?case by (rule_tac SPEC_rule_nested2) auto qed auto lemma false_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \nd\nds. false\<^sub>r \ old nd)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_false_propag) auto lemma expand_and_propag: assumes "\ and\<^sub>r \ \ old (fst n_ns) \ {\, \} \ old (fst n_ns) \ new (fst n_ns)" (is "?Q n_ns") and "\nd\snd n_ns. \ and\<^sub>r \ \ old nd \ {\, \} \ old nd" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\nm_nds. ?P (snd nm_nds))" using assms proof (rule_tac expand_rec_rule[where \="\x. ?Q x \ ?P (snd x)"], simp, intro refine_vcg, goal_cases) case 1 then show ?case by (simp, rule_tac upd_incoming__ident) auto next case prems: (4 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (5 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case (6 f x n ns) then show ?case by auto next case prems: (7 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q)) \ \ \ true\<^sub>r \ \ \ false\<^sub>r \ \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \) \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp_all show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) auto next case 2 then show ?case by (rule_tac step) auto qed qed auto lemma and_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \nd\nds. \ and\<^sub>r \ \ old nd \ {\, \} \ old nd)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_and_propag) auto lemma expand_or_propag: assumes "\ or\<^sub>r \ \ old (fst n_ns) \ {\, \} \ (old (fst n_ns) \ new (fst n_ns)) \ {}" (is "?Q n_ns") and "\nd\snd n_ns. \ or\<^sub>r \ \ old nd \ {\, \} \ old nd \ {}" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\nm_nds. ?P (snd nm_nds))" using assms proof (rule_tac expand_rec_rule[where \="\x. ?Q x \ ?P (snd x)"], simp, intro refine_vcg, goal_cases) case 1 then show ?case by (simp, rule_tac upd_incoming__ident) auto next case prems: (4 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (5 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case (6 f x n ns) then show ?case by auto next case prems: (7 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q)) \ \ \ true\<^sub>r \ \ \ false\<^sub>r \ \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \) \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp_all show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) auto next case 2 then show ?case by (rule_tac step) auto qed qed auto lemma or_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \nd\nds. \ or\<^sub>r \ \ old nd \ {\, \} \ old nd \ {})" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_or_propag) auto abbreviation "next_propag__assm \ n_ns \ (X\<^sub>r \ \ old (fst n_ns) \ \ \ next (fst n_ns)) \ (\nd\snd n_ns. X\<^sub>r \ \ old nd \ name nd \ incoming (fst n_ns) \ \ \ old (fst n_ns) \ new (fst n_ns))" abbreviation "next_propag__rslt \ ns \ \nd\ns. \nd'\ns. X\<^sub>r \ \ old nd \ name nd \ incoming nd' \ \ \ old nd'" lemma expand_next_propag: fixes n_ns :: "_ \ 'a node set" assumes "next_propag__assm \ n_ns \ next_propag__rslt \ (snd n_ns) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. next_propag__rslt \ (snd r))" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case (1 f x n ns) then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case prems: 1 { fix nd :: "'a node" and nd' :: "'a node" assume "nd\ns" and nd'_elem: "nd'\upd_incoming n ns" have "\ \ old nd'" if *: "X\<^sub>r \ \ old nd" and **: "name nd \ incoming nd'" proof (cases "nd'\ns") case True with prems * ** show ?thesis using \nd\ns\ by auto next case False with upd_incoming__elem[of nd' n ns] nd'_elem * ** obtain nd'' where "nd''\ns" and nd'_eq: "nd' = nd''\incoming := incoming n \ incoming nd''\" and old_eq: "old nd'' = old n" by auto have "\ \ old nd'" proof (cases "name nd \ incoming n") case True with prems * \nd\ns\ have "\ \ old n" by auto then show ?thesis using nd'_eq old_eq by simp next case False then have "name nd \ incoming nd''" using \name nd \incoming nd'\ nd'_eq by simp then show ?thesis unfolding nd'_eq using \nd\ns\ \nd''\ns\ * prems by auto qed then show ?thesis by auto qed } then show ?case by auto next case 2 then show ?case by simp qed next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto from prems have Q: "?Q (n, ns)" by auto from Q have name_le: "name n < expand_new_name (name n)" by auto let ?x' = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" have Q'1: "expand_assm_incoming ?x'\ expand_name_ident (snd ?x')" using \new n = {}\ Q[THEN conjunct2, THEN conjunct2] name_le by force have Q'2: "next_propag__assm \ ?x' \ next_propag__rslt \ (snd ?x')" using Q \new n = {}\ by auto show ?case using \new n = {}\ unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\_ r. expand_name_ident (snd r)" and P = "\_. ?P"], goal_cases) case 1 then show ?case using Q'1 by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident) auto next case 2 then show ?case using Q'1 Q'2 by (rule_tac step) simp next case (3 nm nds) then show ?case by simp qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_exist_eq x r \ expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp+ next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\_. ?P"], rule_tac step) auto next case prems': (3 nm nds) let ?x' = "(n\new := new n - {\}, name := fst (nm, nds), new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\)\, nds)" from prems' show ?case proof (rule_tac step, goal_cases) case prems'': 1 then have "expand_assm_incoming ?x'" by auto moreover from prems'' have nds_ident: "expand_name_ident (snd ?x')" by simp moreover have "X\<^sub>r \ \ old (fst ?x') \ \\next (fst ?x')" using Q[THEN conjunct1] goal_assms by auto moreover from prems'' have "next_propag__rslt \ (snd ?x')" by simp moreover from prems'' have name_nds_eq: "name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" by auto have "\nd\nds. (X\<^sub>r \ \ old nd \ name nd \ incoming (fst ?x')) \ \\old (fst ?x')\new (fst ?x')" (is "\nd\nds. ?assm (fst ?x') nd \ ?concl (fst ?x') nd") proof fix nd assume "nd\nds" { assume loc_assm: "name nd\name ` ns" then obtain n' where n': "n'\ns" "name n' = name nd" by auto moreover from prems'' n' obtain nd' where "nd'\nds" and n'_nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto ultimately have "nd = nd'" using nds_ident \nd\nds\ loc_assm by auto moreover from prems'' have "?assm n n' \ ?concl n n'" using \n'\ns\ by auto ultimately have "?assm (fst ?x') nd \ ?concl (fst ?x') nd" using n'_nd'_eq by auto } moreover { assume "name nd\name ` ns" with name_nds_eq \nd\nds\ have "name nd \ name n" by auto with prems'' have "\ (?assm (fst ?x') nd)" by auto } ultimately show "?assm (fst ?x') nd \ ?concl (fst ?x') nd" by auto qed ultimately show ?case by simp qed qed qed lemma next_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. \n'\nds. X\<^sub>r \\old n \ name n\incoming n' \ \\old n')" unfolding create_graph_def apply (refine_vcg expand_next_propag) apply (auto simp add:expand_new_name_expand_init) done abbreviation "release_propag__assm \ \ n_ns \ (\ R\<^sub>r \ \ old (fst n_ns) \ {\, \}\old (fst n_ns)\new (fst n_ns) \ (\\old (fst n_ns)\new (fst n_ns)) \ \ R\<^sub>r \ \ next (fst n_ns)) \ (\nd\snd n_ns. \ R\<^sub>r \ \ old nd \ name nd \ incoming (fst n_ns) \ {\, \}\old nd \ (\\old nd \ \ R\<^sub>r \ \ old (fst n_ns)\new (fst n_ns)))" abbreviation "release_propag__rslt \ \ ns \ \nd\ns. \nd'\ns. \ R\<^sub>r \ \ old nd \ name nd \ incoming nd' \ {\, \}\old nd \ (\\old nd \ \ R\<^sub>r \ \ old nd')" lemma expand_release_propag: fixes n_ns :: "_ \ 'a node set" assumes "release_propag__assm \ \ n_ns \ release_propag__rslt \ \ (snd n_ns) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. release_propag__rslt \ \ (snd r))" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case (1 f x n ns) then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case prems: 1 { fix nd :: "'a node" and nd' :: "'a node" let ?V_prop = "\ R\<^sub>r \ \ old nd \ name nd \ incoming nd' \ {\, \} \ old nd \ \ \ old nd \ \ R\<^sub>r \ \ old nd'" assume "nd\ns" and nd'_elem: "nd'\upd_incoming n ns" { assume "nd'\ns" with prems have ?V_prop using \nd\ns\ by auto } moreover { assume "nd'\ns" and V_in_nd: "\ R\<^sub>r \ \ old nd" and "name nd \incoming nd'" with upd_incoming__elem[of nd' n ns] nd'_elem obtain nd'' where "nd''\ns" and nd'_eq: "nd' = nd''\incoming := incoming n \ incoming nd''\" and old_eq: "old nd'' = old n" by auto { assume "name nd \ incoming n" with prems V_in_nd \nd\ns\ have "{\, \} \ old nd \ \ \ old nd \ \ R\<^sub>r \ \ old n" by auto then have "{\, \} \ old nd \ \ \ old nd \ \ R\<^sub>r \ \ old nd'" using nd'_eq old_eq by simp } moreover { assume "name nd \ incoming n" then have "name nd \ incoming nd''" using \name nd \incoming nd'\ nd'_eq by simp then have "{\, \} \ old nd \ \ \ old nd \ \ R\<^sub>r \ \ old nd'" unfolding nd'_eq using prems \nd\ns\ \nd''\ns\ V_in_nd by auto } ultimately have "{\, \} \ old nd \ \ \ old nd \ \ R\<^sub>r \ \ old nd'" by fast } ultimately have ?V_prop by auto } then show ?case by auto next case 2 then show ?case by simp qed next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto from prems have Q: "?Q (n, ns)" by auto from Q have name_le: "name n < expand_new_name (name n)" by auto let ?x' = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" have Q'1: "expand_assm_incoming ?x'\ expand_name_ident (snd ?x')" using \new n = {}\ Q[THEN conjunct2, THEN conjunct2] name_le by force have Q'2: "release_propag__assm \ \ ?x' \ release_propag__rslt \ \ (snd ?x')" using Q \new n = {}\ by auto show ?case using \new n = {}\ unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\_ r. expand_name_ident (snd r)" and P = "\_. ?P"], goal_cases) case 1 then show ?case using Q'1 by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident) auto next case 2 then show ?case using Q'1 Q'2 by (rule_tac step) simp next case (3 nm nds) then show ?case by simp qed next case 3 then show ?case by auto next case prems: (4 f x n ns \) then have goal_assms: "\ \ new n \ (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q))" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f x n ns \) then have goal_assms: "\ \ new n \ \ = true\<^sub>r" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \)" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_exist_eq x r \ expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp+ next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\_. ?P"], rule_tac step) auto next case prems': (3 nm nds) let ?x' = "(n\new := new n - {\}, name := fst (nm, nds), new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\)\, nds)" from prems' show ?case proof (rule_tac step, goal_cases) case prems'': 1 then have "expand_assm_incoming ?x'" by auto moreover from prems'' have nds_ident: "expand_name_ident (snd ?x')" by simp moreover have "(\ R\<^sub>r \ \ old (fst ?x') \ ({\, \}\old (fst ?x') \ new (fst ?x') \ (\\old (fst ?x')\new (fst ?x') \ \ R\<^sub>r \ \ next (fst ?x'))))" using Q[THEN conjunct1] goal_assms by auto moreover from prems'' have "release_propag__rslt \ \ (snd ?x')" by simp moreover from prems'' have name_nds_eq: "name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" by auto have "\nd\nds. (\ R\<^sub>r \ \ old nd \ name nd \ incoming (fst ?x')) \ ({\, \}\old nd \ (\\old nd \ \ R\<^sub>r \ \old (fst ?x')\new (fst ?x')))" (is "\nd\nds. ?assm (fst ?x') nd \ ?concl (fst ?x') nd") proof fix nd assume "nd\nds" { assume loc_assm: "name nd\name ` ns" then obtain n' where n': "n'\ns" "name n' = name nd" by auto with prems'' obtain nd' where "nd'\nds" and n'_nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto with n' have "nd = nd'" using nds_ident \nd\nds\ loc_assm by auto moreover from prems'' have "?assm n n' \ ?concl n n'" using \n'\ns\ by auto ultimately have "?assm (fst ?x') nd \ ?concl (fst ?x') nd" using n'_nd'_eq by auto } moreover { assume "name nd\name ` ns" with name_nds_eq \nd\nds\ have "name nd \ name n" by auto with prems'' have "\ (?assm (fst ?x') nd)" by auto } ultimately show "?assm (fst ?x') nd \ ?concl (fst ?x') nd" by auto qed ultimately show ?case by simp qed qed qed lemma release_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. \n'\nds. \ R\<^sub>r \\old n \ name n\incoming n' \ ({\, \}\old n \ \\old n \ \ R\<^sub>r \\old n'))" unfolding create_graph_def apply (refine_vcg expand_release_propag) by (auto simp add:expand_new_name_expand_init) abbreviation "until_propag__assm f g n_ns \ (f U\<^sub>r g \ old (fst n_ns) \ (g\old (fst n_ns)\new (fst n_ns) \ (f\old (fst n_ns)\new (fst n_ns) \ f U\<^sub>r g \ next (fst n_ns)))) \ (\nd\snd n_ns. f U\<^sub>r g \ old nd \ name nd \ incoming (fst n_ns) \ (g\old nd \ (f\old nd \ f U\<^sub>r g \old (fst n_ns)\new (fst n_ns))))" abbreviation "until_propag__rslt f g ns \ \n\ns. \nd\ns. f U\<^sub>r g \ old n \ name n \ incoming nd \ (g \ old n \ (f\old n \ f U\<^sub>r g \ old nd))" lemma expand_until_propag: fixes n_ns :: "_ \ 'a node set" assumes "until_propag__assm \ \ n_ns \ until_propag__rslt \ \ (snd n_ns) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. until_propag__rslt \ \ (snd r))" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case prems': 1 { fix nd :: "'a node" and nd' :: "'a node" let ?U_prop = "\ U\<^sub>r \ \ old nd \ name nd \ incoming nd' \ \ \ old nd \ \ \ old nd \ \ U\<^sub>r \ \ old nd'" assume "nd\ns" and nd'_elem: "nd'\upd_incoming n ns" { assume "nd'\ns" with prems' have ?U_prop using \nd\ns\ by auto } moreover { assume "nd'\ns" and U_in_nd: "\ U\<^sub>r \ \ old nd" and "name nd \incoming nd'" with upd_incoming__elem[of nd' n ns] nd'_elem obtain nd'' where "nd''\ns" and nd'_eq: "nd' = nd''\incoming := incoming n \ incoming nd''\" and old_eq: "old nd'' = old n" by auto { assume "name nd \ incoming n" with prems' U_in_nd \nd\ns\ have "\ \ old nd \ \ \ old nd \ \ U\<^sub>r \ \ old n" by auto then have "\ \ old nd \ \ \ old nd \ \ U\<^sub>r \ \ old nd'" using nd'_eq old_eq by simp } moreover { assume "name nd \ incoming n" then have "name nd \ incoming nd''" using \name nd \incoming nd'\ nd'_eq by simp then have "\ \ old nd \ \ \ old nd \ \ U\<^sub>r \ \ old nd'" unfolding nd'_eq using prems' \nd\ns\ \nd''\ns\ U_in_nd by auto } ultimately have "\ \ old nd \ \ \ old nd \ \ U\<^sub>r \ \ old nd'" by fast } ultimately have ?U_prop by auto } then show ?case by auto next case 2 then show ?case by simp qed next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto from prems have Q: "?Q (n, ns)" by auto from Q have name_le: "name n < expand_new_name (name n)" by auto let ?x' = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" have Q'1: "expand_assm_incoming ?x'\ expand_name_ident (snd ?x')" using \new n = {}\ Q[THEN conjunct2, THEN conjunct2] name_le by force have Q'2: "until_propag__assm \ \ ?x' \ until_propag__rslt \ \ (snd ?x')" using Q \new n = {}\ by auto show ?case using \new n = {}\ unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\_ r. expand_name_ident (snd r)" and P = "\_. ?P"], goal_cases) case 1 then show ?case using Q'1 by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident) auto next case 2 then show ?case using Q'1 Q'2 by (rule_tac step) simp next case (3 nm nds) then show ?case by simp qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_exist_eq x r \ expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp+ next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\_. ?P"], rule_tac step) auto next case prems: (3 nm nds) let ?x' = "(n\new := new n - {\}, name := fst (nm, nds), new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\)\, nds)" from prems show ?case proof (rule_tac step, goal_cases) case prems': 1 then have "expand_assm_incoming ?x'" by auto moreover from prems' have nds_ident: "expand_name_ident (snd ?x')" by simp moreover have "(\ U\<^sub>r \ \ old (fst ?x') \ (\\old (fst ?x')\new (fst ?x') \ (\\old (fst ?x')\new (fst ?x') \ \ U\<^sub>r \ \ next (fst ?x'))))" using Q[THEN conjunct1] goal_assms by auto moreover from prems' have "until_propag__rslt \ \ (snd ?x')" by simp moreover from prems' have name_nds_eq: "name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" by auto have "\nd\nds. (\ U\<^sub>r \ \ old nd \ name nd \ incoming (fst ?x')) \ (\\old nd \ (\\old nd \ \ U\<^sub>r \ \old (fst ?x')\new (fst ?x')))" (is "\nd\nds. ?assm (fst ?x') nd \ ?concl (fst ?x') nd") proof fix nd assume "nd\nds" { assume loc_assm: "name nd\name ` ns" then obtain n' where n': "n'\ns" "name n' = name nd" by auto moreover from prems' n' obtain nd' where "nd'\nds" and n'_nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto ultimately have "nd = nd'" using nds_ident \nd\nds\ loc_assm by auto moreover from prems' have "?assm n n' \ ?concl n n'" using \n'\ns\ by auto ultimately have "?assm (fst ?x') nd \ ?concl (fst ?x') nd" using n'_nd'_eq by auto } moreover { assume "name nd\name ` ns" with name_nds_eq \nd\nds\ have "name nd \ name n" by auto with prems' have "\ (?assm (fst ?x') nd)" by auto } ultimately show "?assm (fst ?x') nd \ ?concl (fst ?x') nd" by auto qed ultimately show ?case by simp qed qed qed lemma until_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. \n'\nds. \ U\<^sub>r \\old n \ name n\incoming n' \ (\\old n \ \\old n \ \ U\<^sub>r \\old n'))" unfolding create_graph_def apply (refine_vcg expand_until_propag) by (auto simp add:expand_new_name_expand_init) definition all_subfrmls :: "'a node \ 'a frml set" where "all_subfrmls n \ \(subfrmlsr ` (new n \ old n \ next n))" lemma all_subfrmls__UnionD: assumes "(\x\A. subfrmlsr x) \ B" and "x\A" and "y\subfrmlsr x" shows "y\B" proof - note subfrmlsr_id[of x] also have "subfrmlsr x \ (\x\A. subfrmlsr x)" using assms by auto finally show ?thesis using assms by auto qed lemma expand_all_subfrmls_propag: assumes "all_subfrmls (fst n_ns) \ B \ (\nd\snd n_ns. all_subfrmls nd \ B)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. \nd\snd r. all_subfrmls nd \ B)" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case 1 then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case 1 then show ?case by auto next case 2 then show ?case by (simp add: all_subfrmls_def) qed next case 2 then show ?case by (auto simp add: all_subfrmls_def) next case 3 then show ?case by (auto simp add: all_subfrmls_def) next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac step) (auto simp add: all_subfrmls_def) next case prems: (5 f _ n ns \) then have goal_assms: "\ \ new n \ \ = true\<^sub>r" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac step) (auto dest: all_subfrmls__UnionD simp add: all_subfrmls_def) next case 6 then show ?case by auto next case prems: (7 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \)" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac step) (auto dest: all_subfrmls__UnionD simp add: all_subfrmls_def) next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q)) \ \ \ true\<^sub>r \ \ \ false\<^sub>r \ \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \) \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) (auto dest: all_subfrmls__UnionD simp: all_subfrmls_def) next case 2 then show ?case by (rule_tac step) (auto dest: all_subfrmls__UnionD simp: all_subfrmls_def) qed qed lemma old_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. old n \ subfrmlsr \)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_all_subfrmls_propag[where B = "subfrmlsr \"]) (force simp add:all_subfrmls_def expand_new_name_expand_init)+ lemma L4_2__aux: assumes run: "ipath gba.E \" and "\ U\<^sub>r \ \ old (\ 0)" and "\j. (\i, \ U\<^sub>r \} \ old (\ i)) \ \ \ old (\ j)" shows "\i. {\, \ U\<^sub>r \} \ old (\ i) \ \ \ old (\ i)" proof - have "\i, \ U\<^sub>r \} \ old (\ i)" (is "?sbthm j") for j proof (induct j) show "?sbthm 0" by auto next fix k assume step: "?sbthm k" then have \_k_prop: "\ \ old (\ k) \ \ k\qs \ \ (Suc k)\qs \ name (\ k) \ incoming (\ (Suc k))" using assms run_propag_on_create_graph[OF run] by auto with inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] have "{\, \ U\<^sub>r \} \ old (\ k)" (is "?subsetthm") proof (cases k) assume "k = 0" with assms \_k_prop inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] show ?subsetthm by auto next fix l assume "k = Suc l" then have "{\, \ U\<^sub>r \}\old (\ l) \ \\old (\ l) \ \ l\qs \ \ k\qs \ name (\ l)\incoming (\ k)" using step assms run_propag_on_create_graph[OF run] by auto with inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] have "\ U\<^sub>r \\old (\ k)" by auto with \_k_prop inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] show ?subsetthm by auto qed with step show "?sbthm (Suc k)" by (metis less_SucE) qed with assms show ?thesis by auto qed lemma L4_2a: assumes "ipath gba.E \" and "\ U\<^sub>r \ \ old (\ 0)" shows "(\i. {\, \ U\<^sub>r \} \ old (\ i) \ \ \ old (\ i)) \ (\j. (\i, \ U\<^sub>r \} \ old (\ i)) \ \ \ old (\ j))" (is "?A \ ?B") proof (rule disjCI) assume "\ ?B" then show ?A using assms by (rule_tac L4_2__aux) blast+ qed lemma L4_2b: assumes run: "ipath gba.E \" and "\ U\<^sub>r \ \ old (\ 0)" and ACC: "gba.is_acc \" shows "\j. (\i, \ U\<^sub>r \} \ old (\ i)) \ \ \ old (\ j)" proof (rule ccontr) assume "\ ?thesis" then have contr: "\i. {\, \ U\<^sub>r \}\old(\ i) \ \\old(\ i)" using assms L4_2a[of \ \ \] by blast define S where "S = {q \ qs. \ U\<^sub>r \ \ old q \ \ \ old q}" from assms inres_SPEC[OF res old_propag_on_create_graph] create_gba_from_nodes__ipath have "\ U\<^sub>r \ \ subfrmlsr \" by (metis assms(2) subsetD) then have "S\gbg_F(create_gba_from_nodes \ qs)" unfolding S_def create_gba_from_nodes_def by auto with ACC have 1: "\\<^sub>\i. \ i \ S" unfolding gba.is_acc_def by blast from INFM_EX[OF 1] obtain k where "\ k \ qs" and "\ U\<^sub>r \ \ old (\ k) \ \ \ old (\ k)" unfolding S_def by auto moreover have "{\, \ U\<^sub>r \}\old(\ k) \ \\old(\ k)" using contr by auto ultimately show False by auto qed lemma L4_2c: assumes run: "ipath gba.E \" and "\ R\<^sub>r \ \ old (\ 0)" shows "\i. \ \ old (\ i) \ (\j \ old (\ j))" proof - have "{\, \ R\<^sub>r \} \ old (\ i) \ (\j \ old (\ j))" (is "?thm i") for i proof (induct i) case 0 have "\ 0\qs \ \ 1\qs \ name (\ 0) \ incoming (\ 1)" using create_gba_from_nodes__ipath assms by auto then show ?case using assms inres_SPEC[OF res release_propag_on_create_graph, of \ \] by auto next case (Suc k) note \?thm k\ moreover { assume "{\, \ R\<^sub>r \} \ old (\ k)" moreover have "\ k\qs \ \ (Suc k)\qs \ name (\ k) \ incoming (\ (Suc k))" using create_gba_from_nodes__ipath assms by auto ultimately have "\ \ old (\ k) \ \ R\<^sub>r \ \ old (\ (Suc k))" using assms inres_SPEC[OF res release_propag_on_create_graph, of \ \] by auto moreover { assume "\ \ old (\ k)" then have ?case by blast } moreover { assume "\ R\<^sub>r \ \ old (\ (Suc k))" moreover have "\ (Suc k)\qs \ \ (Suc (Suc k))\qs \ name (\ (Suc k)) \ incoming (\ (Suc (Suc k)))" using create_gba_from_nodes__ipath assms by auto ultimately have ?case using assms inres_SPEC[OF res release_propag_on_create_graph, of \ \] by auto } ultimately have ?case by fast } moreover { assume "\j \ old (\ j)" then have ?case by auto } ultimately show ?case by auto qed then show ?thesis by auto qed lemma L4_8': assumes "ipath gba.E \" (is "?inf_run \") and "gba.is_acc \" (is "?gbarel_accept \") and "\i. gba.L (\ i) (\ i)" (is "?lgbarel_accept \ \") and "\ \ old (\ 0)" shows "\ \\<^sub>r \" using assms proof (induct \ arbitrary: \ \) case True_ltlr show ?case by auto next case False_ltlr then show ?case using inres_SPEC[OF res false_propag_on_create_graph] create_gba_from_nodes__ipath by (metis) next case (Prop_ltlr p) then show ?case unfolding create_gba_from_nodes_def by auto next case (Nprop_ltlr p) then show ?case unfolding create_gba_from_nodes_def by auto next case (And_ltlr \ \) then show ?case using inres_SPEC[OF res and_propag_on_create_graph, of \ \] create_gba_from_nodes__ipath by (metis insert_subset semantics_ltlr.simps(5)) next case (Or_ltlr \ \) then have "\ \ old (\ 0) \ \ \ old (\ 0)" using inres_SPEC[OF res or_propag_on_create_graph, of \ \] create_gba_from_nodes__ipath by (metis (full_types) Int_empty_left Int_insert_left_if0) moreover have "\ \\<^sub>r \" if "\ \ old (\ 0)" using Or_ltlr that by auto moreover have "\ \\<^sub>r \" if "\ \ old (\ 0)" using Or_ltlr that by auto ultimately show ?case by auto next case (Next_ltlr \) with create_gba_from_nodes__ipath[of \] have "\ 0 \ qs \ \ 1 \ qs \ name (\ 0) \ incoming (\ 1)" by auto with inres_SPEC[OF res next_propag_on_create_graph, of \] have "\\old (suffix 1 \ 0)" using Next_ltlr by auto moreover have "?inf_run (suffix 1 \)" and "?gbarel_accept (suffix 1 \)" and "?lgbarel_accept (suffix 1 \) (suffix 1 \ )" using Next_ltlr create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done ultimately show ?case using Next_ltlr by simp next case (Until_ltlr \ \) then have "\j. (\i, \ U\<^sub>r \} \ old (\ i)) \ \ \ old (\ j)" using L4_2b by auto then obtain j where \_pre: "\i, \ U\<^sub>r \} \ old (\ i)" and "\ \ old (suffix j \ 0)" by auto moreover have "?inf_run (suffix j \)" and "?gbarel_accept (suffix j \)" and "?lgbarel_accept (suffix j \) (suffix j \)" unfolding limit_suffix using Until_ltlr create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done ultimately have "suffix j \ \\<^sub>r \" using Until_ltlr by simp moreover { fix i assume "i < j" have "?inf_run (suffix i \)" and "?gbarel_accept (suffix i \)" and "?lgbarel_accept (suffix i \) (suffix i \ )" unfolding limit_suffix using Until_ltlr create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done moreover have "\\old (suffix i \ 0)" using \_pre \i by auto ultimately have "suffix i \ \\<^sub>r \" using Until_ltlr by simp } ultimately show ?case by auto next case (Release_ltlr \ \) { fix i assume "\ \ old (\ i) \ (\j \ old (\ j))" moreover { assume *: "\ \ old (\ i)" have "?inf_run (suffix i \)" and "?gbarel_accept (suffix i \)" and "?lgbarel_accept (suffix i \) (suffix i \ )" unfolding limit_suffix using Release_ltlr create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done with * have "suffix i \ \\<^sub>r \" using Release_ltlr by auto } moreover { assume "\j \ old (\ j)" then obtain j where "j \ old (\ j)" by auto moreover have "?inf_run (suffix j \)" and "?gbarel_accept (suffix j \)" and "?lgbarel_accept (suffix j \) (suffix j \ )" unfolding limit_suffix using Release_ltlr create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done ultimately have "suffix j \ \\<^sub>r \" using Release_ltlr by auto then have "\j \\<^sub>r \" using \j by auto } ultimately have "suffix i \ \\<^sub>r \ \ (\j \\<^sub>r \)" by auto } then show ?case using Release_ltlr L4_2c by auto qed lemma L4_8: assumes "gba.is_acc_run \" and "\i. gba.L (\ i) (\ i)" and "\ \ old (\ 0)" shows "\ \\<^sub>r \" using assms unfolding gba.is_acc_run_def gba.is_run_def using L4_8' by blast lemma expand_expand_init_propag: assumes "(\nm\incoming n'. nm < name n') \ name n' \ name (fst n_ns) \ (incoming n' \ incoming (fst n_ns) \ {} \ new n' \ old (fst n_ns) \ new (fst n_ns)) " (is "?Q n_ns") and "\nd\snd n_ns. \nm\incoming n'. nm\incoming nd \ new n' \ old nd" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\r. name n'\fst r \ ?P (snd r))" using assms proof (rule_tac expand_rec_rule[where \="\x. ?Q x \ ?P (snd x)"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) then have goal_assms: "new n = {} \ ?Q (n, ns) \ ?P ns" by simp { fix nd nm assume "nd\upd_incoming n ns" and "nm\incoming n'" and "nm\incoming nd" { assume "nd\ns" with goal_assms \nm\incoming n'\ \nm\incoming nd\ have "new n' \ old nd" by auto } moreover { assume "nd\ns" with upd_incoming__elem[OF \nd\upd_incoming n ns\] obtain nd' where "nd'\ns" and nd_eq: "nd = nd'\incoming := incoming n \ incoming nd'\" and old_next_eq: "old nd' = old n \ next nd' = next n" by auto { assume "nm\incoming nd'" with goal_assms \nm\incoming n'\ \nd'\ns\ have "new n' \ old nd" unfolding nd_eq by auto } moreover { assume "nm\incoming n" with nd_eq old_next_eq goal_assms \nm\incoming n'\ have "new n' \ old nd" by auto } ultimately have "new n' \ old nd" using \nm\incoming nd\ nd_eq by auto } ultimately have "new n' \ old nd" by fast } with prems show ?case by auto next case prems: (2 f x n ns) then have goal_assms: "new n = {} \ ?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp_all then show ?case proof (rule_tac step, goal_cases) case prems': 1 have expand_new_name_less: "name n < expand_new_name (name n)" by auto moreover have "name n \ incoming n'" using goal_assms by auto ultimately show ?case using prems' by auto qed next case 3 then show ?case by auto next case prems: (4 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp from prems show ?case by (rule_tac step) auto next case prems: (5 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp from prems show ?case by (rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp from prems show ?case by (rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>r(q) \ \ = nprop\<^sub>r(q)) \ \ \ true\<^sub>r \ \ \ false\<^sub>r \ \ (\\ \. \ = \ and\<^sub>r \ \ \ = X\<^sub>r \) \ (\\ \. \ = \ or\<^sub>r \ \ \ = \ U\<^sub>r \ \ \ = \ R\<^sub>r \)" by (cases \) auto from prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp_all show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) auto next case 2 then show ?case by (rule_tac step) auto qed qed lemma expand_init_propag_on_create_graph: "create_graph \ \ SPEC(\nds. \nd\nds. expand_init\incoming nd \ \ \ old nd)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_expand_init_propag[where n' = "\ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \"]) (auto simp add:expand_new_name_expand_init) lemma L4_6: assumes "q\gba.V0" shows "\\old q" using assms inres_SPEC[OF res expand_init_propag_on_create_graph] unfolding create_gba_from_nodes_def by auto lemma L4_9: assumes acc: "gba.accept \" shows "\ \\<^sub>r \" proof - from acc obtain \ where accept: "gba.is_acc_run \ \ (\i. gba.L (\ i) (\ i))" unfolding gba.accept_def by auto then have "\ 0\gba.V0" unfolding gba.is_acc_run_def gba.is_run_def by simp with L4_6 have "\\old (\ 0)" by auto with L4_8 accept show ?thesis by auto qed lemma L4_10: assumes "\ \\<^sub>r \" shows "gba.accept \" proof - define \ where "\ = auto_run \ qs" let ?G = "create_graph \" have \_prop_0: "(\ 0)\qs \ expand_init\incoming(\ 0) \ auto_run_j 0 \ (\ 0)" (is "?sbthm") using inres_SPEC[OF res L4_7[OF \\ \\<^sub>r \\]] unfolding \_def auto_run.simps by (rule_tac someI_ex, simp) blast have \_valid: "\j. \ j \ qs \ auto_run_j j \ (\ j)" (is "\j. ?\_valid j") proof fix j show "?\_valid j" proof(induct j) from \_prop_0 show "?\_valid 0" by fast next fix k assume goal_assms: "\ k \ qs \ auto_run_j k \ (\ k)" with inres_SPEC[OF res L4_5, of "suffix k \"] have sbthm: "\q'. q'\qs \ name (\ k)\incoming q' \ auto_run_j (Suc k) \ q'" (is "\q'. ?sbthm q'") by auto have "?sbthm (\ (Suc k))" using someI_ex[OF sbthm] unfolding \_def auto_run.simps by blast then show "?\_valid (Suc k)" by fast qed qed have \_prop_Suc: "\k. \ k\ qs \ \ (Suc k)\qs \ name (\ k)\incoming (\ (Suc k)) \ auto_run_j (Suc k) \ (\ (Suc k))" proof - fix k from \_valid have "\ k \ qs" and "auto_run_j k \ (\ k)" by blast+ with inres_SPEC[OF res L4_5, of "suffix k \"] have sbthm: "\q'. q'\qs \ name (\ k)\incoming q' \ auto_run_j (Suc k) \ q'" (is "\q'. ?sbthm q'") by auto show "\ k\ qs \ ?sbthm (\ (Suc k))" using \\ k\qs\ someI_ex[OF sbthm] unfolding \_def auto_run.simps by blast qed have \_vnaccpt: "\k \ \. \ U\<^sub>r \ \ old (\ k) \ \ (\i. {\, \ U\<^sub>r \} \ old (\ (k+i)) \ \ \ old (\ (k+i)))" proof clarify fix k \ \ assume U_in: "\ U\<^sub>r \ \ old (\ k)" and cntr_prm: "\i. {\, \ U\<^sub>r \} \ old (\ (k+i)) \ \ \ old (\ (k+i))" have "suffix k \ \\<^sub>r \ U\<^sub>r \" using U_in \_valid by force then obtain i where "suffix (k+i) \ \\<^sub>r \" and "\j \\<^sub>r \" by auto moreover have "\ U\<^sub>r \ \ old (\ (k+i)) \ \ \ old (\ (k+i))" using cntr_prm by auto ultimately show False using \_valid by force qed have \_exec: "gba.is_run \" using \_prop_0 \_prop_Suc \_valid unfolding gba.is_run_def ipath_def unfolding create_gba_from_nodes_def by auto moreover have \_vaccpt: "\k \ \. \ U\<^sub>r \ \ old (\ k) \ (\j. (\i, \ U\<^sub>r \} \ old (\ (k+i))) \ \ \ old (\ (k+j)))" proof(clarify) fix k \ \ assume U_in: "\ U\<^sub>r \ \ old (\ k)" then have "\ (\i. {\, \ U\<^sub>r \} \ old (suffix k \ i) \ \ \ old (suffix k \ i))" using \_vnaccpt[THEN allE, of k] by auto moreover have "suffix k \ 0 \ qs" using \_valid by auto ultimately show "\j. (\i, \ U\<^sub>r \} \ old (\ (k+i))) \ \ \ old (\ (k+j))" apply - apply (rule make_pos_rule'[OF L4_2a]) apply (fold suffix_def) apply (rule ipath_suffix) using \_exec[unfolded gba.is_run_def] apply simp using U_in apply simp apply simp done qed have "gba.is_acc \" unfolding gba.is_acc_def proof fix S assume "S\gba.F" then obtain \ \ where S_eq: "S = {q \ qs. \ U\<^sub>r \ \ old q \ \ \ old q}" and "\ U\<^sub>r \ \ subfrmlsr \" by (auto simp add: create_gba_from_nodes_def) have range_subset: "range \ \ qs" proof fix q assume "q\range \" with full_SetCompr_eq[of \] obtain k where "q = \ k" by auto then show "q \ qs" using \_valid by auto qed with limit_nonempty[of \] limit_in_range[of \] finite_subset[OF range_subset] inres_SPEC[OF res create_graph_finite] obtain q where q_in_limit: "q \ limit \" and q_is_node: "q\qs" by auto show "\\<^sub>\i. \ i \ S" proof (cases "\ U\<^sub>r \ \ old q") case False with S_eq q_in_limit q_is_node show ?thesis by (auto simp: limit_iff_frequent intro: INFM_mono) next case True obtain k where q_eq: "q = \ k" using q_in_limit unfolding limit_iff_frequent by (metis (lifting, full_types) INFM_nat_le) have "\\<^sub>\ k. \ \ old (\ k)" unfolding INFM_nat proof (rule ccontr) assume "\ (\m. \n>m. \ \ old (\ n))" then obtain m where "\n>m. \ \ old (\ n)" by auto moreover from q_eq q_in_limit limit_iff_frequent[of q \] INFM_nat[of "\n. \ n = q"] obtain n where "mn_eq: "\ n = \ k" by auto moreover obtain j where "\ \ old (\ (n+j))" using \_vaccpt \\ U\<^sub>r \ \ old q\ unfolding q_eq by (fold \n_eq) force ultimately show False by auto qed then have "\\<^sub>\ k. \ k \ qs \ \ \ old (\ k)" using \_valid by (auto intro: INF_mono) then show "\\<^sub>\ k. \ k \ S" unfolding S_eq by (rule INFM_mono) simp qed qed moreover have "gba.L (\ i) (\ i)" for i proof - from \_valid have [simp]: "\ i \ qs" by auto have "\\\old (\ i). suffix i \ \\<^sub>r \" using \_valid by auto then show ?thesis unfolding create_gba_from_nodes_def by auto qed ultimately show ?thesis unfolding gba.accept_def gba.is_acc_run_def by blast qed end end lemma create_graph__name_ident: "create_graph \ \ SPEC (\nds. \q\nds. \!q'\nds. name q = name q')" unfolding create_graph_def apply (refine_vcg expand_name_propag__name_ident) by (auto simp add:expand_new_name_expand_init) corollary create_graph__name_inj: "create_graph \ \ SPEC (\nds. inj_on name nds)" by (rule order_trans[OF create_graph__name_ident]) (auto intro: inj_onI) definition "create_gba \ \ do { nds \ create_graph\<^sub>T \; RETURN (create_gba_from_nodes \ nds) }" lemma create_graph_precond: "create_graph \ \ SPEC (create_gba_from_nodes_precond \)" apply (clarsimp simp: pw_le_iff create_graph_nofail) apply unfold_locales apply simp done lemma create_gba__invar: "create_gba \ \ SPEC gba" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_precond]) by (rule create_gba_from_nodes_precond.create_gba_from_nodes__invar) lemma create_gba_acc: shows "create_gba \ \ SPEC(\\. \\. gba.accept \ \ \ \ \\<^sub>r \)" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_precond]) using create_gba_from_nodes_precond.L4_9 using create_gba_from_nodes_precond.L4_10 by blast lemma create_gba__name_inj: shows "create_gba \ \ SPEC(\\. (inj_on name (g_V \)))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph__name_inj]) apply (auto simp: create_gba_from_nodes_def) done lemma create_gba__fin: shows "create_gba \ \ SPEC(\\. (finite (g_V \)))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_finite]) apply (auto simp: create_gba_from_nodes_def) done lemma create_graph_old_finite: "create_graph \ \ SPEC (\nds. \nd\nds. finite (old nd))" proof - show ?thesis unfolding create_graph_def create_graph_eq_create_graph\<^sub>T[symmetric] unfolding expand_def apply (intro refine_vcg) apply (rule_tac order_trans) apply (rule_tac REC_le_RECT) apply (fold expand\<^sub>T_def) apply (rule_tac order_trans[OF expand_term_prop]) apply auto[1] apply (rule_tac SPEC_rule) apply auto by (metis infinite_super subfrmlsr_finite) qed lemma create_gba__old_fin: shows "create_gba \ \ SPEC(\\. \nd\g_V \. finite (old nd))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_old_finite]) apply (simp add: create_gba_from_nodes_def) done lemma create_gba__incoming_exists: shows "create_gba \ \ SPEC(\\. \nd\g_V \. incoming nd \ insert expand_init (name ` (g_V \)))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph__incoming_name_exist]) apply (auto simp add: create_gba_from_nodes_def) done lemma create_gba__no_init: shows "create_gba \ \ SPEC(\\. expand_init \ name ` (g_V \))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph__incoming_name_exist]) apply (auto simp add: create_gba_from_nodes_def) done definition "nds_invars nds \ inj_on name nds \ finite nds \ expand_init \ name`nds \ (\nd\nds. finite (old nd) \ incoming nd \ insert expand_init (name ` nds))" lemma create_gba_nds_invars: "create_gba \ \ SPEC (\\. nds_invars (g_V \))" using create_gba__name_inj[of \] create_gba__fin[of \] create_gba__old_fin[of \] create_gba__incoming_exists[of \] create_gba__no_init[of \] unfolding nds_invars_def by (simp add: pw_le_iff) theorem T4_1: "create_gba \ \ SPEC( \\. gba \ \ finite (g_V \) \ (\\. gba.accept \ \ \ \ \\<^sub>r \) \ (nds_invars (g_V \)))" using create_gba__invar create_gba__fin create_gba_acc create_gba_nds_invars apply (simp add: pw_le_iff) apply blast done definition "create_name_gba \ \ do { G \ create_gba \; ASSERT (nds_invars (g_V G)); RETURN (gba_rename name G) }" theorem create_name_gba_correct: "create_name_gba \ \ SPEC(\\. gba \ \ finite (g_V \) \ (\\. gba.accept \ \ \ \ \\<^sub>r \))" unfolding create_name_gba_def apply (refine_rcg refine_vcg order_trans[OF T4_1]) apply (simp_all add: nds_invars_def gba_rename_correct) done definition create_name_igba :: "'a::linorder ltlr \ _" where "create_name_igba \ \ do { A \ create_name_gba \; A' \ gba_to_idx A; stat_set_data_nres (card (g_V A)) (card (g_V0 A')) (igbg_num_acc A'); RETURN A' }" lemma create_name_igba_correct: "create_name_igba \ \ SPEC (\G. igba G \ finite (g_V G) \ (\\. igba.accept G \ \ \ \\<^sub>r \))" unfolding create_name_igba_def apply (refine_rcg order_trans[OF create_name_gba_correct] order_trans[OF gba.gba_to_idx_ext_correct] refine_vcg) apply clarsimp_all proof - fix G :: "(nat, 'a set) gba_rec" fix A :: "nat set" assume 1: "gba G" assume 2: "finite (g_V G)" "A \ gbg_F G" interpret gba G using 1 . show "finite A" using finite_V_Fe 2 . qed (* For presentation in paper*) context notes [refine_vcg] = order_trans[OF create_name_gba_correct] begin lemma "create_name_igba \ \ SPEC (\G. igba G \ (\\. igba.accept G \ \ \ \\<^sub>r \))" unfolding create_name_igba_def proof (refine_rcg refine_vcg, clarsimp_all) fix G :: "(nat, 'a set) gba_rec" assume "gba G" then interpret gba G . note [refine_vcg] = order_trans[OF gba_to_idx_ext_correct] assume "\\. gba.accept G \ = \ \\<^sub>r \" "finite (g_V G)" then show "gba_to_idx G \ SPEC (\G'. igba G' \ (\\. igba.accept G' \ = \ \\<^sub>r \))" by (refine_rcg refine_vcg) (auto intro: finite_V_Fe) qed end end diff --git a/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy b/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy --- a/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy +++ b/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy @@ -1,411 +1,408 @@ (* Title: Nested Multisets Author: Dmitriy Traytel , 2016 Author: Jasmin Blanchette , 2016 Maintainer: Dmitriy Traytel *) section \Nested Multisets\ theory Nested_Multiset imports "HOL-Library.Multiset_Order" begin declare multiset.map_comp [simp] declare multiset.map_cong [cong] subsection \Type Definition\ datatype 'a nmultiset = Elem 'a | MSet "'a nmultiset multiset" inductive no_elem :: "'a nmultiset \ bool" where "(\X. X \# M \ no_elem X) \ no_elem (MSet M)" inductive_set sub_nmset :: "('a nmultiset \ 'a nmultiset) set" where "X \# M \ (X, MSet M) \ sub_nmset" lemma wf_sub_nmset[simp]: "wf sub_nmset" proof (rule wfUNIVI) fix P :: "'a nmultiset \ bool" and M :: "'a nmultiset" assume IH: "\M. (\N. (N, M) \ sub_nmset \ P N) \ P M" show "P M" by (induct M; rule IH[rule_format]) (auto simp: sub_nmset.simps) qed primrec depth_nmset :: "'a nmultiset \ nat" ("|_|") where "|Elem a| = 0" | "|MSet M| = (let X = set_mset (image_mset depth_nmset M) in if X = {} then 0 else Suc (Max X))" lemma depth_nmset_MSet: "x \# M \ |x| < |MSet M|" by (auto simp: less_Suc_eq_le) declare depth_nmset.simps(2)[simp del] subsection \Dershowitz and Manna's Nested Multiset Order\ text \The Dershowitz--Manna extension:\ definition less_multiset_ext\<^sub>D\<^sub>M :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "less_multiset_ext\<^sub>D\<^sub>M R M N \ (\X Y. X \ {#} \ X \# N \ M = (N - X) + Y \ (\k. k \# Y \ (\a. a \# X \ R k a)))" lemma less_multiset_ext\<^sub>D\<^sub>M_imp_mult: assumes N_A: "set_mset N \ A" and M_A: "set_mset M \ A" and less: "less_multiset_ext\<^sub>D\<^sub>M R M N" shows "(M, N) \ mult {(x, y). x \ A \ y \ A \ R x y}" proof - from less obtain X Y where "X \ {#}" and "X \# N" and "M = N - X + Y" and "\k. k \# Y \ (\a. a \# X \ R k a)" unfolding less_multiset_ext\<^sub>D\<^sub>M_def by blast with N_A M_A have "(N - X + Y, N - X + X) \ mult {(x, y). x \ A \ y \ A \ R x y}" by (intro one_step_implies_mult, blast, metis (mono_tags, lifting) case_prodI mem_Collect_eq mset_subset_eqD mset_subset_eq_add_right subsetCE) with \M = N - X + Y\ \X \# N\ show "(M, N) \ mult {(x, y). x \ A \ y \ A \ R x y}" by (simp add: subset_mset.diff_add) qed lemma mult_imp_less_multiset_ext\<^sub>D\<^sub>M: assumes N_A: "set_mset N \ A" and M_A: "set_mset M \ A" and trans: "\x \ A. \y \ A. \z \ A. R x y \ R y z \ R x z" and in_mult: "(M, N) \ mult {(x, y). x \ A \ y \ A \ R x y}" shows "less_multiset_ext\<^sub>D\<^sub>M R M N" using in_mult N_A M_A unfolding mult_def less_multiset_ext\<^sub>D\<^sub>M_def proof induct case (base N) then obtain y M0 X where "N = add_mset y M0" and "M = M0 + X" and "\a. a \# X \ R a y" unfolding mult1_def by auto thus ?case by (auto intro: exI[of _ "{#y#}"]) next case (step N N') note N_N'_in_mult1 = this(2) and ih = this(3) and N'_A = this(4) and M_A = this(5) have N_A: "set_mset N \ A" using N_N'_in_mult1 N'_A unfolding mult1_def by auto obtain Y X where y_nemp: "Y \ {#}" and y_sub_N: "Y \# N" and M_eq: "M = N - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ R x y)" using ih[OF N_A M_A] by blast obtain z M0 Ya where N'_eq: "N' = M0 + {#z#}" and N_eq: "N = M0 + Ya" and z_gt: "\y. y \# Ya \ y \ A \ z \ A \ R y z" using N_N'_in_mult1[unfolded mult1_def] by auto let ?Za = "Y - Ya + {#z#}" let ?Xa = "X + Ya + (Y - Ya) - Y" have xa_sub_x_ya: "set_mset ?Xa \ set_mset (X + Ya)" by (metis diff_subset_eq_self in_diffD subsetI subset_mset.diff_diff_right) have x_A: "set_mset X \ A" using M_A M_eq by auto have ya_A: "set_mset Ya \ A" by (simp add: subsetI z_gt) have ex_y': "\y. y \# Y - Ya + {#z#} \ R x y" if x_in: "x \# X + Ya" for x proof (cases "x \# X") case True then obtain y where y_in: "y \# Y" and y_gt_x: "R x y" using ex_y by blast show ?thesis proof (cases "y \# Ya") case False hence "y \# Y - Ya + {#z#}" using y_in count_greater_zero_iff in_diff_count by fastforce thus ?thesis using y_gt_x by blast next case True hence "y \ A" and "z \ A" and "R y z" using z_gt by blast+ hence "R x z" using trans y_gt_x x_A ya_A x_in by (meson subsetCE union_iff) thus ?thesis by auto qed next case False hence "x \# Ya" using x_in by auto hence "x \ A" and "z \ A" and "R x z" using z_gt by blast+ thus ?thesis by auto qed show ?case proof (rule exI[of _ ?Za], rule exI[of _ ?Xa], intro conjI) show "Y - Ya + {#z#} \# N'" using mset_subset_eq_mono_add subset_eq_diff_conv y_sub_N N_eq N'_eq by (simp add: subset_eq_diff_conv) next show "M = N' - (Y - Ya + {#z#}) + (X + Ya + (Y - Ya) - Y)" unfolding M_eq N_eq N'_eq by (auto simp: multiset_eq_iff) next show "\x. x \# X + Ya + (Y - Ya) - Y \ (\y. y \# Y - Ya + {#z#} \ R x y)" using ex_y' xa_sub_x_ya by blast qed auto qed lemma less_multiset_ext\<^sub>D\<^sub>M_iff_mult: assumes N_A: "set_mset N \ A" and M_A: "set_mset M \ A" and trans: "\x \ A. \y \ A. \z \ A. R x y \ R y z \ R x z" shows "less_multiset_ext\<^sub>D\<^sub>M R M N \ (M, N) \ mult {(x, y). x \ A \ y \ A \ R x y}" using mult_imp_less_multiset_ext\<^sub>D\<^sub>M[OF assms] less_multiset_ext\<^sub>D\<^sub>M_imp_mult[OF N_A M_A] by blast -lemma sub_nmset_lex: "((a,b),(c,d)) \ sub_nmset <*lex*> sub_nmset \ (a,c) \ sub_nmset \ a=c \ (b,d) \ sub_nmset" - by auto - instantiation nmultiset :: (preorder) preorder begin lemma less_multiset_ext\<^sub>D\<^sub>M_cong[fundef_cong]: "(\X Y k a. X \ {#} \ X \# N \ M = (N - X) + Y \ k \# Y \ R k a = S k a) \ less_multiset_ext\<^sub>D\<^sub>M R M N = less_multiset_ext\<^sub>D\<^sub>M S M N" unfolding less_multiset_ext\<^sub>D\<^sub>M_def by metis function less_nmultiset :: "'a nmultiset \ 'a nmultiset \ bool" where "less_nmultiset (Elem a) (Elem b) \ a < b" | "less_nmultiset (Elem a) (MSet M) \ True" | "less_nmultiset (MSet M) (Elem a) \ False" | "less_nmultiset (MSet M) (MSet N) \ less_multiset_ext\<^sub>D\<^sub>M less_nmultiset M N" by pat_completeness auto termination - apply(relation "sub_nmset <*lex*> sub_nmset", force) - by (meson sub_nmset_lex sub_nmset.intros union_iff) + by (relation "sub_nmset <*lex*> sub_nmset", fastforce, + metis sub_nmset.simps in_lex_prod mset_subset_eqD mset_subset_eq_add_right) lemmas less_nmultiset_induct = less_nmultiset.induct[case_names Elem_Elem Elem_MSet MSet_Elem MSet_MSet] lemmas less_nmultiset_cases = less_nmultiset.cases[case_names Elem_Elem Elem_MSet MSet_Elem MSet_MSet] lemma trans_less_nmultiset: "X < Y \ Y < Z \ X < Z" for X Y Z :: "'a nmultiset" proof (induct "Max {|X|, |Y|, |Z|}" arbitrary: X Y Z rule: less_induct) case less from less(2,3) show ?case proof (cases X; cases Y; cases Z) fix M N N' :: "'a nmultiset multiset" define A where "A = set_mset M \ set_mset N \ set_mset N'" assume XYZ: "X = MSet M" "Y = MSet N" "Z = MSet N'" then have trans: "\x \ A. \y \ A. \z \ A. x < y \ y < z \ x < z" by (auto elim!: less(1)[rotated -1] dest!: depth_nmset_MSet simp add: A_def) have "set_mset M \ A" "set_mset N \ A" "set_mset N' \ A" unfolding A_def by auto with less(2,3) XYZ show "X < Z" by (auto simp: less_multiset_ext\<^sub>D\<^sub>M_iff_mult[OF _ _ trans] mult_def) qed (auto elim: less_trans) qed lemma irrefl_less_nmultiset: fixes X :: "'a nmultiset" shows "X < X \ False" proof (induct X) case (MSet M) from MSet(2) show ?case unfolding less_nmultiset.simps less_multiset_ext\<^sub>D\<^sub>M_def proof safe fix X Y :: "'a nmultiset multiset" define XY where "XY = {(x, y). x \# X \ y \# Y \ y < x}" then have fin: "finite XY" and trans: "trans XY" by (auto simp: trans_def intro: trans_less_nmultiset finite_subset[OF _ finite_cartesian_product]) assume "X \ {#}" "X \# M" "M = M - X + Y" then have "X = Y" by (auto simp: mset_subset_eq_exists_conv) with MSet(1) \X \# M\ have "irrefl XY" unfolding XY_def by (force dest: mset_subset_eqD simp: irrefl_def) with trans have "acyclic XY" by (simp add: acyclic_irrefl) moreover assume "\k. k \# Y \ (\a. a \# X \ k < a)" with \X = Y\ \X \ {#}\ have "\ acyclic XY" by (intro notI, elim finite_acyclic_wf[OF fin, elim_format]) (auto dest!: spec[of _ "set_mset Y"] simp: wf_eq_minimal XY_def) ultimately show False by blast qed qed simp lemma antisym_less_nmultiset: fixes X Y :: "'a nmultiset" shows "X < Y \ Y < X \ False" using trans_less_nmultiset irrefl_less_nmultiset by blast definition less_eq_nmultiset :: "'a nmultiset \ 'a nmultiset \ bool" where "less_eq_nmultiset X Y = (X < Y \ X = Y)" instance proof (intro_classes, goal_cases less_def refl trans) case (less_def x y) then show ?case unfolding less_eq_nmultiset_def by (metis irrefl_less_nmultiset antisym_less_nmultiset) next case (refl x) then show ?case unfolding less_eq_nmultiset_def by blast next case (trans x y z) then show ?case unfolding less_eq_nmultiset_def by (metis trans_less_nmultiset) qed lemma less_multiset_ext\<^sub>D\<^sub>M_less: "less_multiset_ext\<^sub>D\<^sub>M (<) = (<)" unfolding fun_eq_iff less_multiset_ext\<^sub>D\<^sub>M_def less_multiset\<^sub>D\<^sub>M by blast end instantiation nmultiset :: (order) order begin instance proof (intro_classes, goal_cases antisym) case (antisym x y) then show ?case unfolding less_eq_nmultiset_def by (metis trans_less_nmultiset irrefl_less_nmultiset) qed end instantiation nmultiset :: (linorder) linorder begin lemma total_less_nmultiset: fixes X Y :: "'a nmultiset" shows "\ X < Y \ Y \ X \ Y < X" proof (induct X Y rule: less_nmultiset_induct) case (MSet_MSet M N) then show ?case unfolding nmultiset.inject less_nmultiset.simps less_multiset_ext\<^sub>D\<^sub>M_less less_multiset\<^sub>H\<^sub>O by (metis add_diff_cancel_left' count_inI diff_add_zero in_diff_count less_imp_not_less mset_subset_eq_multiset_union_diff_commute subset_mset.order.refl) qed auto instance proof (intro_classes, goal_cases total) case (total x y) then show ?case unfolding less_eq_nmultiset_def by (metis total_less_nmultiset) qed end lemma less_depth_nmset_imp_less_nmultiset: "|X| < |Y| \ X < Y" proof (induct X Y rule: less_nmultiset_induct) case (MSet_MSet M N) then show ?case proof (cases "M = {#}") case False with MSet_MSet show ?thesis by (auto 0 4 simp: depth_nmset.simps(2) less_multiset_ext\<^sub>D\<^sub>M_def not_le Max_gr_iff intro: exI[of _ N] split: if_splits) qed (auto simp: depth_nmset.simps(2) less_multiset_ext\<^sub>D\<^sub>M_less split: if_splits) qed simp_all lemma less_nmultiset_imp_le_depth_nmset: "X < Y \ |X| \ |Y|" proof (induct X Y rule: less_nmultiset_induct) case (MSet_MSet M N) then have "M < N" by (simp add: less_multiset_ext\<^sub>D\<^sub>M_less) then show ?case proof (cases "M = {#}" "N = {#}" rule: bool.exhaust[case_product bool.exhaust]) case [simp]: False_False show ?thesis unfolding depth_nmset.simps(2) Let_def False_False Suc_le_mono set_image_mset image_is_empty set_mset_eq_empty_iff if_False proof (intro iffD2[OF Max_le_iff] ballI iffD2[OF Max_ge_iff]; (elim imageE)?; simp) fix X assume [simp]: "X \# M" with MSet_MSet(1)[of N M X, simplified] \M < N\ show "\Y\#N. |X| \ |Y|" by (meson ex_gt_imp_less_multiset less_asym' less_depth_nmset_imp_less_nmultiset not_le_imp_less) qed qed (auto simp: depth_nmset.simps(2)) qed simp_all lemma eq_mlex_I: fixes f :: "'a \ nat" and R :: "'a \ 'a \ bool" assumes "\X Y. f X < f Y \ R X Y" and "antisymp R" shows "{(X, Y). R X Y} = f <*mlex*> {(X, Y). f X = f Y \ R X Y}" proof safe fix X Y assume "R X Y" show "(X, Y) \ f <*mlex*> {(X, Y). f X = f Y \ R X Y}" proof (cases "f X" "f Y" rule: linorder_cases) case less with \R X Y\ show ?thesis by (elim mlex_less) next case equal with \R X Y\ show ?thesis by (intro mlex_leq) auto next case greater from \R X Y\ assms(1)[OF greater] \antisymp R\ greater show ?thesis unfolding antisymp_def by auto qed next fix X Y assume "(X, Y) \ f <*mlex*> {(X, Y). f X = f Y \ R X Y}" then show "R X Y" unfolding mlex_prod_def by (auto simp: assms(1)) qed instantiation nmultiset :: (wellorder) wellorder begin lemma depth_nmset_eq_0[simp]: "|X| = 0 \ (X = MSet {#} \ (\x. X = Elem x))" by (cases X; simp add: depth_nmset.simps(2)) lemma depth_nmset_eq_Suc[simp]: "|X| = Suc n \ (\N. X = MSet N \ (\Y \# N. |Y| = n) \ (\Y \# N. |Y| \ n))" by (cases X; auto simp add: depth_nmset.simps(2) intro!: Max_eqI) (metis (no_types, lifting) Max_in finite_imageI finite_set_mset imageE image_is_empty set_mset_eq_empty_iff) lemma wf_less_nmultiset_depth: "wf {(X :: 'a nmultiset, Y). |X| = i \ |Y| = i \ X < Y}" proof (induct i rule: less_induct) case (less i) define A :: "'a nmultiset set" where "A = {X. |X| < i}" from less have "wf ((depth_nmset :: 'a nmultiset \ nat) <*mlex*> (\j < i. {(X, Y). |X| = j \ |Y| = j \ X < Y}))" by (intro wf_UN wf_mlex) auto then have *: "wf (mult {(X :: 'a nmultiset, Y). X \ A \ Y \ A \ X < Y})" by (intro wf_mult, elim wf_subset) (force simp: A_def mlex_prod_def not_less_iff_gr_or_eq dest!: less_depth_nmset_imp_less_nmultiset) show ?case proof (cases i) case 0 then show ?thesis by (auto simp: inj_on_def intro!: wf_subset[OF wf_Un[OF wf_map_prod_image[OF wf, of Elem] wf_UN[of "Elem ` UNIV" "\x. {(x, MSet {#})}"]]]) next case (Suc n) then show ?thesis by (intro wf_subset[OF wf_map_prod_image[OF *, of MSet]]) (auto 0 4 simp: map_prod_def image_iff inj_on_def A_def dest!: less_multiset_ext\<^sub>D\<^sub>M_imp_mult[of _ A, rotated -1] split: prod.splits) qed qed lemma wf_less_nmultiset: "wf {(X :: 'a nmultiset, Y :: 'a nmultiset). X < Y}" (is "wf ?R") proof - have "?R = depth_nmset <*mlex*> {(X, Y). |X| = |Y| \ X < Y}" by (rule eq_mlex_I) (auto simp: antisymp_def less_depth_nmset_imp_less_nmultiset) also have "{(X, Y). |X| = |Y| \ X < Y} = (\i. {(X, Y). |X| = i \ |Y| = i \ X < Y})" by auto finally show ?thesis by (fastforce intro: wf_mlex wf_Union wf_less_nmultiset_depth) qed instance using wf_less_nmultiset unfolding wf_def mem_Collect_eq prod.case by intro_classes metis end end diff --git a/thys/Simpl/Hoare.thy b/thys/Simpl/Hoare.thy --- a/thys/Simpl/Hoare.thy +++ b/thys/Simpl/Hoare.thy @@ -1,423 +1,424 @@ (* 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 \ (a,x) \ r \ (a=x \ (b,y)\s))" + "((a,b),(x,y)) \ r <*lex*> s = ((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: order_le_less) + by (auto simp add: in_measure_iff 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 + z: "z \ range g" and + min_z: "\y. f y < f z \ y \ range g" + by (auto simp add: in_measure_iff) 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: order_less_le) + by (auto simp add: in_measure_iff 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/UTP/toolkit/List_Extra.thy b/thys/UTP/toolkit/List_Extra.thy --- a/thys/UTP/toolkit/List_Extra.thy +++ b/thys/UTP/toolkit/List_Extra.thy @@ -1,869 +1,864 @@ (******************************************************************************) (* Project: Isabelle/UTP Toolkit *) (* File: List_Extra.thy *) (* Authors: Simon Foster and Frank Zeyda *) (* Emails: simon.foster@york.ac.uk and frank.zeyda@york.ac.uk *) (******************************************************************************) section \ Lists: extra functions and properties \ theory List_Extra imports "HOL-Library.Sublist" "HOL-Library.Monad_Syntax" "HOL-Library.Prefix_Order" "Optics.Lens_Instances" begin subsection \ Useful Abbreviations \ abbreviation "list_sum xs \ foldr (+) xs 0" subsection \ Folds \ context abel_semigroup begin lemma foldr_snoc: "foldr (\<^bold>*) (xs @ [x]) k = (foldr (\<^bold>*) xs k) \<^bold>* x" by (induct xs, simp_all add: commute left_commute) end subsection \ List Lookup \ text \ The following variant of the standard \nth\ function returns \\\ if the index is out of range. \ primrec nth_el :: "'a list \ nat \ 'a option" ("_\_\\<^sub>l" [90, 0] 91) where "[]\i\\<^sub>l = None" | "(x # xs)\i\\<^sub>l = (case i of 0 \ Some x | Suc j \ xs \j\\<^sub>l)" lemma nth_el_appendl[simp]: "i < length xs \ (xs @ ys)\i\\<^sub>l = xs\i\\<^sub>l" apply (induct xs arbitrary: i) apply simp apply (case_tac i) apply simp_all done lemma nth_el_appendr[simp]: "length xs \ i \ (xs @ ys)\i\\<^sub>l = ys\i - length xs\\<^sub>l" apply (induct xs arbitrary: i) apply simp apply (case_tac i) apply simp_all done subsection \ Extra List Theorems \ subsubsection \ Map \ lemma map_nth_Cons_atLeastLessThan: "map (nth (x # xs)) [Suc m.. Suc" by auto hence "map (nth xs) [m.. Suc) [m.. Sorted Lists \ lemma sorted_last [simp]: "\ x \ set xs; sorted xs \ \ x \ last xs" by (induct xs, auto) lemma sorted_prefix: assumes "xs \ ys" "sorted ys" shows "sorted xs" proof - obtain zs where "ys = xs @ zs" using Prefix_Order.prefixE assms(1) by auto thus ?thesis using assms(2) sorted_append by blast qed lemma sorted_map: "\ sorted xs; mono f \ \ sorted (map f xs)" by (simp add: monoD sorted_iff_nth_mono) lemma sorted_distinct [intro]: "\ sorted (xs); distinct(xs) \ \ (\ i Is the given list a permutation of the given set? \ definition is_sorted_list_of_set :: "('a::ord) set \ 'a list \ bool" where "is_sorted_list_of_set A xs = ((\ i set(xs) = A)" lemma sorted_is_sorted_list_of_set: assumes "is_sorted_list_of_set A xs" shows "sorted(xs)" and "distinct(xs)" using assms proof (induct xs arbitrary: A) show "sorted []" by auto next show "distinct []" by auto next fix A :: "'a set" case (Cons x xs) note hyps = this assume isl: "is_sorted_list_of_set A (x # xs)" hence srt: "(\in. \ n < length (x # xs) - 1 \ (x # xs) ! n < (x # xs) ! (n + 1)) \ set (x # xs) = A" by (meson \is_sorted_list_of_set A (x # xs)\ is_sorted_list_of_set_def) then show ?thesis by (metis \distinct xs\ add.commute add_diff_cancel_left' distinct.simps(2) leD length_Cons length_greater_0_conv length_pos_if_in_set less_le nth_Cons_0 nth_Cons_Suc plus_1_eq_Suc set_ConsD sorted.elims(2) srtd) qed qed lemma is_sorted_list_of_set_alt_def: "is_sorted_list_of_set A xs \ sorted (xs) \ distinct (xs) \ set(xs) = A" apply (auto intro: sorted_is_sorted_list_of_set) apply (auto simp add: is_sorted_list_of_set_def) apply (metis Nat.add_0_right One_nat_def add_Suc_right sorted_distinct) done definition sorted_list_of_set_alt :: "('a::ord) set \ 'a list" where "sorted_list_of_set_alt A = (if (A = {}) then [] else (THE xs. is_sorted_list_of_set A xs))" lemma is_sorted_list_of_set: "finite A \ is_sorted_list_of_set A (sorted_list_of_set A)" by (simp add: is_sorted_list_of_set_alt_def) lemma sorted_list_of_set_other_def: "finite A \ sorted_list_of_set(A) = (THE xs. sorted(xs) \ distinct(xs) \ set xs = A)" apply (rule sym) apply (rule the_equality) apply (auto) apply (simp add: sorted_distinct_set_unique) done lemma sorted_list_of_set_alt [simp]: "finite A \ sorted_list_of_set_alt(A) = sorted_list_of_set(A)" apply (rule sym) apply (auto simp add: sorted_list_of_set_alt_def is_sorted_list_of_set_alt_def sorted_list_of_set_other_def) done text \ Sorting lists according to a relation \ definition is_sorted_list_of_set_by :: "'a rel \ 'a set \ 'a list \ bool" where "is_sorted_list_of_set_by R A xs = ((\ i R) \ set(xs) = A)" definition sorted_list_of_set_by :: "'a rel \ 'a set \ 'a list" where "sorted_list_of_set_by R A = (THE xs. is_sorted_list_of_set_by R A xs)" definition fin_set_lexord :: "'a rel \ 'a set rel" where "fin_set_lexord R = {(A, B). finite A \ finite B \ (\ xs ys. is_sorted_list_of_set_by R A xs \ is_sorted_list_of_set_by R B ys \ (xs, ys) \ lexord R)}" lemma is_sorted_list_of_set_by_mono: "\ R \ S; is_sorted_list_of_set_by R A xs \ \ is_sorted_list_of_set_by S A xs" by (auto simp add: is_sorted_list_of_set_by_def) lemma lexord_mono': "\ (\ x y. f x y \ g x y); (xs, ys) \ lexord {(x, y). f x y} \ \ (xs, ys) \ lexord {(x, y). g x y}" by (metis case_prodD case_prodI lexord_take_index_conv mem_Collect_eq) lemma fin_set_lexord_mono [mono]: "(\ x y. f x y \ g x y) \ (xs, ys) \ fin_set_lexord {(x, y). f x y} \ (xs, ys) \ fin_set_lexord {(x, y). g x y}" proof assume fin: "(xs, ys) \ fin_set_lexord {(x, y). f x y}" and hyp: "(\ x y. f x y \ g x y)" from fin have "finite xs" "finite ys" using fin_set_lexord_def by fastforce+ with fin hyp show "(xs, ys) \ fin_set_lexord {(x, y). g x y}" apply (auto simp add: fin_set_lexord_def) apply (rename_tac xs' ys') apply (rule_tac x="xs'" in exI) apply (auto) apply (metis case_prodD case_prodI is_sorted_list_of_set_by_def mem_Collect_eq) apply (metis case_prodD case_prodI is_sorted_list_of_set_by_def lexord_mono' mem_Collect_eq) done qed definition distincts :: "'a set \ 'a list set" where "distincts A = {xs \ lists A. distinct(xs)}" lemma tl_element: "\ x \ set xs; x \ hd(xs) \ \ x \ set(tl(xs))" by (metis in_set_insert insert_Nil list.collapse list.distinct(2) set_ConsD) subsubsection \ List Update \ lemma listsum_update: fixes xs :: "'a::ring list" assumes "i < length xs" shows "list_sum (xs[i := v]) = list_sum xs - xs ! i + v" using assms proof (induct xs arbitrary: i) case Nil then show ?case by (simp) next case (Cons a xs) then show ?case proof (cases i) case 0 thus ?thesis by (simp add: add.commute) next case (Suc i') with Cons show ?thesis by (auto) qed qed subsubsection \ Drop While and Take While \ lemma dropWhile_sorted_le_above: "\ sorted xs; x \ set (dropWhile (\ x. x \ n) xs) \ \ x > n" apply (induct xs) apply (auto) apply (rename_tac a xs) apply (case_tac "a \ n") apply (auto) done lemma set_dropWhile_le: "sorted xs \ set (dropWhile (\ x. x \ n) xs) = {x\set xs. x > n}" apply (induct xs) apply (simp) apply (rename_tac x xs) apply (subgoal_tac "sorted xs") apply (simp) apply (safe) apply (auto) done lemma set_takeWhile_less_sorted: "\ sorted I; x \ set I; x < n \ \ x \ set (takeWhile (\x. x < n) I)" proof (induct I arbitrary: x) case Nil thus ?case by (simp) next case (Cons a I) thus ?case by auto qed lemma nth_le_takeWhile_ord: "\ sorted xs; i \ length (takeWhile (\ x. x \ n) xs); i < length xs \ \ n \ xs ! i" apply (induct xs arbitrary: i, auto) apply (rename_tac x xs i) apply (case_tac "x \ n") apply (auto) apply (metis One_nat_def Suc_eq_plus1 le_less_linear le_less_trans less_imp_le list.size(4) nth_mem set_ConsD) done lemma length_takeWhile_less: "\ a \ set xs; \ P a \ \ length (takeWhile P xs) < length xs" by (metis in_set_conv_nth length_takeWhile_le nat_neq_iff not_less set_takeWhileD takeWhile_nth) lemma nth_length_takeWhile_less: "\ sorted xs; distinct xs; (\ a \ set xs. a \ n) \ \ xs ! length (takeWhile (\x. x < n) xs) \ n" by (induct xs, auto) subsubsection \ Last and But Last \ lemma length_gt_zero_butlast_concat: assumes "length ys > 0" shows "butlast (xs @ ys) = xs @ (butlast ys)" using assms by (metis butlast_append length_greater_0_conv) lemma length_eq_zero_butlast_concat: assumes "length ys = 0" shows "butlast (xs @ ys) = butlast xs" using assms by (metis append_Nil2 length_0_conv) lemma butlast_single_element: shows "butlast [e] = []" by (metis butlast.simps(2)) lemma last_single_element: shows "last [e] = e" by (metis last.simps) lemma length_zero_last_concat: assumes "length t = 0" shows "last (s @ t) = last s" by (metis append_Nil2 assms length_0_conv) lemma length_gt_zero_last_concat: assumes "length t > 0" shows "last (s @ t) = last t" by (metis assms last_append length_greater_0_conv) subsubsection \ Prefixes and Strict Prefixes \ lemma prefix_length_eq: "\ length xs = length ys; prefix xs ys \ \ xs = ys" by (metis not_equal_is_parallel parallel_def) lemma prefix_Cons_elim [elim]: assumes "prefix (x # xs) ys" obtains ys' where "ys = x # ys'" "prefix xs ys'" using assms by (metis append_Cons prefix_def) lemma prefix_map_inj: "\ inj_on f (set xs \ set ys); prefix (map f xs) (map f ys) \ \ prefix xs ys" apply (induct xs arbitrary:ys) apply (simp_all) apply (erule prefix_Cons_elim) apply (auto) apply (metis image_insert insertI1 insert_Diff_if singletonE) done lemma prefix_map_inj_eq [simp]: "inj_on f (set xs \ set ys) \ prefix (map f xs) (map f ys) \ prefix xs ys" using map_mono_prefix prefix_map_inj by blast lemma strict_prefix_Cons_elim [elim]: assumes "strict_prefix (x # xs) ys" obtains ys' where "ys = x # ys'" "strict_prefix xs ys'" using assms by (metis Sublist.strict_prefixE' Sublist.strict_prefixI' append_Cons) lemma strict_prefix_map_inj: "\ inj_on f (set xs \ set ys); strict_prefix (map f xs) (map f ys) \ \ strict_prefix xs ys" apply (induct xs arbitrary:ys) apply (auto) using prefix_bot.bot.not_eq_extremum apply fastforce apply (erule strict_prefix_Cons_elim) apply (auto) apply (metis (hide_lams, full_types) image_insert insertI1 insert_Diff_if singletonE) done lemma strict_prefix_map_inj_eq [simp]: "inj_on f (set xs \ set ys) \ strict_prefix (map f xs) (map f ys) \ strict_prefix xs ys" by (simp add: inj_on_map_eq_map strict_prefix_def) lemma prefix_drop: "\ drop (length xs) ys = zs; prefix xs ys \ \ ys = xs @ zs" by (metis append_eq_conv_conj prefix_def) lemma list_append_prefixD [dest]: "x @ y \ z \ x \ z" using append_prefixD less_eq_list_def by blast lemma prefix_not_empty: assumes "strict_prefix xs ys" and "xs \ []" shows "ys \ []" using Sublist.strict_prefix_simps(1) assms(1) by blast lemma prefix_not_empty_length_gt_zero: assumes "strict_prefix xs ys" and "xs \ []" shows "length ys > 0" using assms prefix_not_empty by auto lemma butlast_prefix_suffix_not_empty: assumes "strict_prefix (butlast xs) ys" shows "ys \ []" using assms prefix_not_empty_length_gt_zero by fastforce lemma prefix_and_concat_prefix_is_concat_prefix: assumes "prefix s t" "prefix (e @ t) u" shows "prefix (e @ s) u" using Sublist.same_prefix_prefix assms(1) assms(2) prefix_order.dual_order.trans by blast lemma prefix_eq_exists: "prefix s t \ (\xs . s @ xs = t)" using prefix_def by auto lemma strict_prefix_eq_exists: "strict_prefix s t \ (\xs . s @ xs = t \ (length xs) > 0)" using prefix_def strict_prefix_def by auto lemma butlast_strict_prefix_eq_butlast: assumes "length s = length t" and "strict_prefix (butlast s) t" shows "strict_prefix (butlast s) t \ (butlast s) = (butlast t)" by (metis append_butlast_last_id append_eq_append_conv assms(1) assms(2) length_0_conv length_butlast strict_prefix_eq_exists) lemma butlast_eq_if_eq_length_and_prefix: assumes "length s > 0" "length z > 0" "length s = length z" "strict_prefix (butlast s) t" "strict_prefix (butlast z) t" shows "(butlast s) = (butlast z)" using assms by (auto simp add:strict_prefix_eq_exists) lemma butlast_prefix_imp_length_not_gt: assumes "length s > 0" "strict_prefix (butlast s) t" shows "\ (length t < length s)" using assms prefix_length_less by fastforce lemma length_not_gt_iff_eq_length: assumes "length s > 0" and "strict_prefix (butlast s) t" shows "(\ (length s < length t)) = (length s = length t)" proof - have "(\ (length s < length t)) = ((length t < length s) \ (length s = length t))" by (metis not_less_iff_gr_or_eq) also have "... = (length s = length t)" using assms by (simp add:butlast_prefix_imp_length_not_gt) finally show ?thesis . qed text \ Greatest common prefix \ fun gcp :: "'a list \ 'a list \ 'a list" where "gcp [] ys = []" | "gcp (x # xs) (y # ys) = (if (x = y) then x # gcp xs ys else [])" | "gcp _ _ = []" lemma gcp_right [simp]: "gcp xs [] = []" by (induct xs, auto) lemma gcp_append [simp]: "gcp (xs @ ys) (xs @ zs) = xs @ gcp ys zs" by (induct xs, auto) lemma gcp_lb1: "prefix (gcp xs ys) xs" apply (induct xs arbitrary: ys, auto) apply (case_tac ys, auto) done lemma gcp_lb2: "prefix (gcp xs ys) ys" apply (induct ys arbitrary: xs, auto) apply (case_tac xs, auto) done interpretation prefix_semilattice: semilattice_inf gcp prefix strict_prefix proof fix xs ys :: "'a list" show "prefix (gcp xs ys) xs" by (induct xs arbitrary: ys, auto, case_tac ys, auto) show "prefix (gcp xs ys) ys" by (induct ys arbitrary: xs, auto, case_tac xs, auto) next fix xs ys zs :: "'a list" assume "prefix xs ys" "prefix xs zs" thus "prefix xs (gcp ys zs)" by (simp add: prefix_def, auto) qed subsubsection \ Lexicographic Order \ lemma lexord_append: assumes "(xs\<^sub>1 @ ys\<^sub>1, xs\<^sub>2 @ ys\<^sub>2) \ lexord R" "length(xs\<^sub>1) = length(xs\<^sub>2)" shows "(xs\<^sub>1, xs\<^sub>2) \ lexord R \ (xs\<^sub>1 = xs\<^sub>2 \ (ys\<^sub>1, ys\<^sub>2) \ lexord R)" using assms proof (induct xs\<^sub>2 arbitrary: xs\<^sub>1) case (Cons x\<^sub>2 xs\<^sub>2') note hyps = this from hyps(3) obtain x\<^sub>1 xs\<^sub>1' where xs\<^sub>1: "xs\<^sub>1 = x\<^sub>1 # xs\<^sub>1'" "length(xs\<^sub>1') = length(xs\<^sub>2')" by (auto, metis Suc_length_conv) with hyps(2) have xcases: "(x\<^sub>1, x\<^sub>2) \ R \ (xs\<^sub>1' @ ys\<^sub>1, xs\<^sub>2' @ ys\<^sub>2) \ lexord R" - by (auto split: if_split_asm) + by (auto) show ?case proof (cases "(x\<^sub>1, x\<^sub>2) \ R") case True with xs\<^sub>1 show ?thesis - using Cons.hyps hyps(2) by auto + by (auto) next case False with xcases have "(xs\<^sub>1' @ ys\<^sub>1, xs\<^sub>2' @ ys\<^sub>2) \ lexord R" by (auto) with hyps(1) xs\<^sub>1 have dichot: "(xs\<^sub>1', xs\<^sub>2') \ lexord R \ (xs\<^sub>1' = xs\<^sub>2' \ (ys\<^sub>1, ys\<^sub>2) \ lexord R)" by (auto) have "x\<^sub>1 = x\<^sub>2" - using False hyps(2) xs\<^sub>1(1) by (auto split: if_split_asm) + using False hyps(2) xs\<^sub>1(1) by auto with dichot xs\<^sub>1 show ?thesis by (simp) qed next case Nil thus ?case by auto qed -lemma prefix_lexord_left: - assumes "(xs, ys) \ lexord R" "prefix xs' xs" +lemma strict_prefix_lexord_rel: + "strict_prefix xs ys \ (xs, ys) \ lexord R" + by (metis Sublist.strict_prefixE' lexord_append_rightI) + +lemma strict_prefix_lexord_left: + assumes "trans R" "(xs, ys) \ lexord R" "strict_prefix xs' xs" shows "(xs', ys) \ lexord R" - using assms - apply (auto simp: lexord_def prefix_def) - apply (metis Nil_is_append_conv list.exhaust) - apply (simp add: append_eq_append_conv2) - apply (drule_tac x="u" in spec) - apply (drule_tac x="a" in spec) - apply (drule_tac x="b" in spec) - apply (auto simp: ) - apply (metis hd_Cons_tl hd_append list.sel(1) self_append_conv2) - apply (metis Nil_is_append_conv hd_Cons_tl) - done + by (metis assms lexord_trans strict_prefix_lexord_rel) lemma prefix_lexord_right: - assumes "(xs, ys) \ lexord R" "prefix ys ys'" + assumes "trans R" "(xs, ys) \ lexord R" "strict_prefix ys ys'" shows "(xs, ys') \ lexord R" - using assms by (fastforce simp: lexord_def prefix_def) - + by (metis assms lexord_trans strict_prefix_lexord_rel) lemma lexord_eq_length: assumes "(xs, ys) \ lexord R" "length xs = length ys" - shows "\i. (xs!i, ys!i) \ R \ xs!i\ys!i \ i < length xs \ (\ j i. (xs!i, ys!i) \ R \ i < length xs \ (\ j R \ x\y") - case True with ys Cons show ?thesis + proof (cases "(x, y) \ R") + case True with ys show ?thesis by (rule_tac x="0" in exI, simp) next case False with ys hyps(2) have xy: "x = y" "(xs, ys') \ lexord R" - by (auto split: if_split_asm) - with hyps(1,3) ys obtain i where "(xs!i, ys'!i) \ R" "xs!i\ys'!i" "i < length xs" "(\ j R" "i < length xs" "(\ j i" "length ys > i" "(xs!i, ys!i) \ R" "xs!i \ ys!i" "\j i" "length ys > i" "(xs!i, ys!i) \ R" "\ j lexord R" using assms proof (induct i arbitrary: xs ys) case 0 thus ?case by (auto, metis lexord_cons_cons list.exhaust nth_Cons_0) next case (Suc i) note hyps = this then obtain x' y' xs' ys' where "xs = x' # xs'" "ys = y' # ys'" by (metis Suc_length_conv Suc_lessE) - moreover with hyps(5,6) have "\jj Distributed Concatenation \ definition uncurry :: "('a \ 'b \ 'c) \ ('a \ 'b \ 'c)" where [simp]: "uncurry f = (\(x, y). f x y)" definition dist_concat :: "'a list set \ 'a list set \ 'a list set" (infixr "\<^sup>\" 100) where "dist_concat ls1 ls2 = (uncurry (@) ` (ls1 \ ls2))" lemma dist_concat_left_empty [simp]: "{} \<^sup>\ ys = {}" by (simp add: dist_concat_def) lemma dist_concat_right_empty [simp]: "xs \<^sup>\ {} = {}" by (simp add: dist_concat_def) lemma dist_concat_insert [simp]: "insert l ls1 \<^sup>\ ls2 = ((@) l ` ( ls2)) \ (ls1 \<^sup>\ ls2)" by (auto simp add: dist_concat_def) subsection \ List Domain and Range \ abbreviation seq_dom :: "'a list \ nat set" ("dom\<^sub>l") where "seq_dom xs \ {0.. 'a set" ("ran\<^sub>l") where "seq_ran xs \ set xs" subsection \ Extracting List Elements \ definition seq_extract :: "nat set \ 'a list \ 'a list" (infix "\\<^sub>l" 80) where "seq_extract A xs = nths xs A" lemma seq_extract_Nil [simp]: "A \\<^sub>l [] = []" by (simp add: seq_extract_def) lemma seq_extract_Cons: "A \\<^sub>l (x # xs) = (if 0 \ A then [x] else []) @ {j. Suc j \ A} \\<^sub>l xs" by (simp add: seq_extract_def nths_Cons) lemma seq_extract_empty [simp]: "{} \\<^sub>l xs = []" by (simp add: seq_extract_def) lemma seq_extract_ident [simp]: "{0..\<^sub>l xs = xs" unfolding list_eq_iff_nth_eq by (auto simp add: seq_extract_def length_nths atLeast0LessThan) lemma seq_extract_split: assumes "i \ length xs" shows "{0..\<^sub>l xs @ {i..\<^sub>l xs = xs" using assms proof (induct xs arbitrary: i) case Nil thus ?case by (simp add: seq_extract_def) next case (Cons x xs) note hyp = this have "{j. Suc j < i} = {0.. Suc j \ j < length xs} = {i - 1..\<^sub>l (xs @ ys) = (A \\<^sub>l xs) @ ({j. j + length xs \ A} \\<^sub>l ys)" by (simp add: seq_extract_def nths_append) lemma seq_extract_range: "A \\<^sub>l xs = (A \ dom\<^sub>l(xs)) \\<^sub>l xs" apply (auto simp add: seq_extract_def nths_def) apply (metis (no_types, lifting) atLeastLessThan_iff filter_cong in_set_zip nth_mem set_upt) done lemma seq_extract_out_of_range: "A \ dom\<^sub>l(xs) = {} \ A \\<^sub>l xs = []" by (metis seq_extract_def seq_extract_range nths_empty) lemma seq_extract_length [simp]: "length (A \\<^sub>l xs) = card (A \ dom\<^sub>l(xs))" proof - have "{i. i < length(xs) \ i \ A} = (A \ {0..\<^sub>l (x # xs) = (if (m = 0) then x # ({0..\<^sub>l xs) else {m-1..\<^sub>l xs)" proof - have "{j. Suc j < n} = {0.. Suc j \ Suc j < n} = {m - Suc 0..\<^sub>l xs = [xs ! i]" using assms apply (induct xs arbitrary: i) apply (auto simp add: seq_extract_Cons) apply (rename_tac xs i) apply (subgoal_tac "{j. Suc j = i} = {i - 1}") apply (auto) done lemma seq_extract_as_map: assumes "m < n" "n \ length xs" shows "{m..\<^sub>l xs = map (nth xs) [m.. (\ i\length(xs). ys = {0..\<^sub>l xs \ zs = {i..\<^sub>l xs)" proof assume xs: "xs = ys @ zs" moreover have "ys = {0..\<^sub>l (ys @ zs)" by (simp add: seq_extract_append) moreover have "zs = {length ys..\<^sub>l (ys @ zs)" proof - have "{length ys.. {0.. i\length(xs). ys = {0..\<^sub>l xs \ zs = {i..\<^sub>l xs)" by (rule_tac x="length ys" in exI, auto) next assume "\i\length xs. ys = {0..\<^sub>l xs \ zs = {i..\<^sub>l xs" thus "xs = ys @ zs" by (auto simp add: seq_extract_split) qed subsection \ Filtering a list according to a set \ definition seq_filter :: "'a list \ 'a set \ 'a list" (infix "\\<^sub>l" 80) where "seq_filter xs A = filter (\ x. x \ A) xs" lemma seq_filter_Cons_in [simp]: "x \ cs \ (x # xs) \\<^sub>l cs = x # (xs \\<^sub>l cs)" by (simp add: seq_filter_def) lemma seq_filter_Cons_out [simp]: "x \ cs \ (x # xs) \\<^sub>l cs = (xs \\<^sub>l cs)" by (simp add: seq_filter_def) lemma seq_filter_Nil [simp]: "[] \\<^sub>l A = []" by (simp add: seq_filter_def) lemma seq_filter_empty [simp]: "xs \\<^sub>l {} = []" by (simp add: seq_filter_def) lemma seq_filter_append: "(xs @ ys) \\<^sub>l A = (xs \\<^sub>l A) @ (ys \\<^sub>l A)" by (simp add: seq_filter_def) lemma seq_filter_UNIV [simp]: "xs \\<^sub>l UNIV = xs" by (simp add: seq_filter_def) lemma seq_filter_twice [simp]: "(xs \\<^sub>l A) \\<^sub>l B = xs \\<^sub>l (A \ B)" by (simp add: seq_filter_def) subsection \ Minus on lists \ instantiation list :: (type) minus begin text \ We define list minus so that if the second list is not a prefix of the first, then an arbitrary list longer than the combined length is produced. Thus we can always determined from the output whether the minus is defined or not. \ definition "xs - ys = (if (prefix ys xs) then drop (length ys) xs else [])" instance .. end lemma minus_cancel [simp]: "xs - xs = []" by (simp add: minus_list_def) lemma append_minus [simp]: "(xs @ ys) - xs = ys" by (simp add: minus_list_def) lemma minus_right_nil [simp]: "xs - [] = xs" by (simp add: minus_list_def) lemma list_concat_minus_list_concat: "(s @ t) - (s @ z) = t - z" by (simp add: minus_list_def) lemma length_minus_list: "y \ x \ length(x - y) = length(x) - length(y)" by (simp add: less_eq_list_def minus_list_def) lemma map_list_minus: "xs \ ys \ map f (ys - xs) = map f ys - map f xs" by (simp add: drop_map less_eq_list_def map_mono_prefix minus_list_def) lemma list_minus_first_tl [simp]: "[x] \ xs \ (xs - [x]) = tl xs" by (metis Prefix_Order.prefixE append.left_neutral append_minus list.sel(3) not_Cons_self2 tl_append2) text \ Extra lemmas about @{term prefix} and @{term strict_prefix} \ lemma prefix_concat_minus: assumes "prefix xs ys" shows "xs @ (ys - xs) = ys" using assms by (metis minus_list_def prefix_drop) lemma prefix_minus_concat: assumes "prefix s t" shows "(t - s) @ z = (t @ z) - s" using assms by (simp add: Sublist.prefix_length_le minus_list_def) lemma strict_prefix_minus_not_empty: assumes "strict_prefix xs ys" shows "ys - xs \ []" using assms by (metis append_Nil2 prefix_concat_minus strict_prefix_def) lemma strict_prefix_diff_minus: assumes "prefix xs ys" and "xs \ ys" shows "(ys - xs) \ []" using assms by (simp add: strict_prefix_minus_not_empty) lemma length_tl_list_minus_butlast_gt_zero: assumes "length s < length t" and "strict_prefix (butlast s) t" and "length s > 0" shows "length (tl (t - (butlast s))) > 0" using assms by (metis Nitpick.size_list_simp(2) butlast_snoc hd_Cons_tl length_butlast length_greater_0_conv length_tl less_trans nat_neq_iff strict_prefix_minus_not_empty prefix_order.dual_order.strict_implies_order prefix_concat_minus) lemma list_minus_butlast_eq_butlast_list: assumes "length t = length s" and "strict_prefix (butlast s) t" shows "t - (butlast s) = [last t]" using assms by (metis append_butlast_last_id append_eq_append_conv butlast.simps(1) length_butlast less_numeral_extra(3) list.size(3) prefix_order.dual_order.strict_implies_order prefix_concat_minus prefix_length_less) lemma butlast_strict_prefix_length_lt_imp_last_tl_minus_butlast_eq_last: assumes "length s > 0" "strict_prefix (butlast s) t" "length s < length t" shows "last (tl (t - (butlast s))) = (last t)" using assms by (metis last_append last_tl length_tl_list_minus_butlast_gt_zero less_numeral_extra(3) list.size(3) append_minus strict_prefix_eq_exists) lemma tl_list_minus_butlast_not_empty: assumes "strict_prefix (butlast s) t" and "length s > 0" and "length t > length s" shows "tl (t - (butlast s)) \ []" using assms length_tl_list_minus_butlast_gt_zero by fastforce lemma tl_list_minus_butlast_empty: assumes "strict_prefix (butlast s) t" and "length s > 0" and "length t = length s" shows "tl (t - (butlast s)) = []" using assms by (simp add: list_minus_butlast_eq_butlast_list) lemma tl_list_minus_butlast_eq_empty: assumes "strict_prefix (butlast s) t" and "length s = length t" shows "tl (t - (butlast s)) = []" using assms by (metis list.sel(3) list_minus_butlast_eq_butlast_list) lemma length_list_minus: assumes "strict_prefix s t" shows "length(t - s) = length(t) - length(s)" using assms by (simp add: minus_list_def prefix_order.dual_order.strict_implies_order) subsection \ Laws on @{term take}, @{term drop}, and @{term nths} \ lemma take_prefix: "m \ n \ take m xs \ take n xs" by (metis Prefix_Order.prefixI append_take_drop_id min_absorb2 take_append take_take) lemma nths_atLeastAtMost_0_take: "nths xs {0..m} = take (Suc m) xs" by (metis atLeast0AtMost lessThan_Suc_atMost nths_upt_eq_take) lemma nths_atLeastLessThan_0_take: "nths xs {0.. n \ nths xs {0..m} \ nths xs {0..n}" by (simp add: nths_atLeastAtMost_0_take take_prefix) lemma sorted_nths_atLeastAtMost_0: "\ m \ n; sorted (nths xs {0..n}) \ \ sorted (nths xs {0..m})" using nths_atLeastAtMost_prefix sorted_prefix by blast lemma sorted_nths_atLeastLessThan_0: "\ m \ n; sorted (nths xs {0.. \ sorted (nths xs {0.. list_augment xs k x = list_update xs k x" by (metis list_augment_def list_augment_idem list_update_overwrite) lemma nths_list_update_out: "k \ A \ nths (list_update xs k x) A = nths xs A" apply (induct xs arbitrary: k x A) apply (auto) apply (rename_tac a xs k x A) apply (case_tac k) apply (auto simp add: nths_Cons) done lemma nths_list_augment_out: "\ k < length xs; k \ A \ \ nths (list_augment xs k x) A = nths xs A" by (simp add: list_augment_as_update nths_list_update_out) end \ No newline at end of file diff --git a/thys/UTP/toolkit/Sequence.thy b/thys/UTP/toolkit/Sequence.thy --- a/thys/UTP/toolkit/Sequence.thy +++ b/thys/UTP/toolkit/Sequence.thy @@ -1,279 +1,287 @@ (******************************************************************************) (* Project: Isabelle/UTP Toolkit *) (* File: Sequence.thy *) (* Authors: Simon Foster and Frank Zeyda *) (* Emails: simon.foster@york.ac.uk and frank.zeyda@york.ac.uk *) (******************************************************************************) section \ Infinite Sequences \ theory Sequence imports HOL.Real List_Extra "HOL-Library.Sublist" "HOL-Library.Nat_Bijection" begin typedef 'a seq = "UNIV :: (nat \ 'a) set" by (auto) setup_lifting type_definition_seq definition ssubstr :: "nat \ nat \ 'a seq \ 'a list" where "ssubstr i j xs = map (Rep_seq xs) [i ..< j]" lift_definition nth_seq :: "'a seq \ nat \ 'a" (infixl "!\<^sub>s" 100) is "\ f i. f i" . abbreviation sinit :: "nat \ 'a seq \ 'a list" where "sinit i xs \ ssubstr 0 i xs" lemma sinit_len [simp]: "length (sinit i xs) = i" by (simp add: ssubstr_def) lemma sinit_0 [simp]: "sinit 0 xs = []" by (simp add: ssubstr_def) lemma prefix_upt_0 [intro]: "i \ j \ prefix [0.. j \ prefix (sinit i xs) (sinit j xs)" by (simp add: map_mono_prefix prefix_upt_0 ssubstr_def) lemma sinit_strict_prefix: "i < j \ strict_prefix (sinit i xs) (sinit j xs)" by (metis sinit_len sinit_prefix le_less nat_neq_iff prefix_order.dual_order.strict_iff_order) lemma nth_sinit: "i < n \ sinit n xs ! i = xs !\<^sub>s i" apply (auto simp add: ssubstr_def) apply (transfer, auto) done lemma sinit_append_split: assumes "i < j" shows "sinit j xs = sinit i xs @ ssubstr i j xs" proof - have "[0.. lexord R" "(sinit j ys, sinit j xs) \ lexord R" shows False proof - have sinit_xs: "sinit j xs = sinit i xs @ ssubstr i j xs" by (metis assms(2) sinit_append_split) have sinit_ys: "sinit j ys = sinit i ys @ ssubstr i j ys" by (metis assms(2) sinit_append_split) from sinit_xs sinit_ys assms(4) have "(sinit i ys, sinit i xs) \ lexord R \ (sinit i ys = sinit i xs \ (ssubstr i j ys, ssubstr i j xs) \ lexord R)" by (auto dest: lexord_append) with assms lexord_asymmetric show False by (force) qed lemma sinit_linear_asym_lemma2: assumes "asym R" "(sinit i xs, sinit i ys) \ lexord R" "(sinit j ys, sinit j xs) \ lexord R" shows False proof (cases i j rule: linorder_cases) case less with assms show ?thesis by (auto dest: sinit_linear_asym_lemma1) next case equal with assms show ?thesis by (simp add: lexord_asymmetric) next case greater with assms show ?thesis by (auto dest: sinit_linear_asym_lemma1) qed lemma range_ext: assumes "\i :: nat. \x\{0.. x" by (metis lessI) with assms show "f(x) = g(x)" by (auto) qed lemma sinit_ext: "(\ i. sinit i xs = sinit i ys) \ xs = ys" by (simp add: ssubstr_def, transfer, auto intro: range_ext) definition seq_lexord :: "'a rel \ ('a seq) rel" where "seq_lexord R = {(xs, ys). (\ i. (sinit i xs, sinit i ys) \ lexord R)}" lemma seq_lexord_irreflexive: "\x. (x, x) \ R \ (xs, xs) \ seq_lexord R" by (auto dest: lexord_irreflexive simp add: irrefl_def seq_lexord_def) lemma seq_lexord_irrefl: "irrefl R \ irrefl (seq_lexord R)" by (simp add: irrefl_def seq_lexord_irreflexive) lemma seq_lexord_transitive: - assumes "trans R" "antisym R" + assumes "trans R" shows "trans (seq_lexord R)" - unfolding seq_lexord_def +unfolding seq_lexord_def proof (rule transI, clarify) fix xs ys zs :: "'a seq" and m n :: nat assume las: "(sinit m xs, sinit m ys) \ lexord R" "(sinit n ys, sinit n zs) \ lexord R" hence inz: "m > 0" using gr0I by force - from las(1) obtain i where sinitm: "(sinit m xs!i, sinit m ys!i) \ R" "sinit m xs!i \ sinit m ys!i" - "i < m" "\ j R" "i < m" "\ j R" "sinit n ys!j \ sinit n zs!j" - "j < n" "\ k R" "j < n" "\ ki. (sinit i xs, sinit i zs) \ lexord R" proof (cases "i \ j") case True note lt = this with sinitm sinitn have "(sinit n xs!i, sinit n zs!i) \ R" by (metis assms le_eq_less_or_eq le_less_trans nth_sinit transD) - moreover have "sinit n xs!i \ sinit n zs!i" - by (metis antisymD assms(2) dual_order.strict_trans2 le_eq_less_or_eq lt nth_sinit sinitm(1) sinitm(2) sinitm(3) sinitn(1) sinitn(3) sinitn(4)) moreover from lt sinitm sinitn have "\ j lexord R" using sinitm(2) sinitn(2) lt apply (rule_tac lexord_intro_elems) - apply (auto) - by (metis less_trans nth_sinit sinitm(3)) + apply (auto) + apply (metis less_le_trans less_trans nth_sinit) + done thus ?thesis by auto next case False then have ge: "i > j" by auto with assms sinitm sinitn have "(sinit n xs!j, sinit n zs!j) \ R" by (metis less_trans nth_sinit) - moreover have "sinit n xs!j \ sinit n zs!j" - by (metis ge less_trans nth_sinit sinitm(3) sinitm(4) sinitn(2) sinitn(3)) moreover from ge sinitm sinitn have "\ k lexord R" using sinitm(2) sinitn(2) ge apply (rule_tac lexord_intro_elems) - apply (auto) - by (metis less_trans nth_sinit sinitm(3)) + apply (auto) + apply (metis less_trans nth_sinit) + done thus ?thesis by auto qed qed lemma seq_lexord_trans: - "\ (xs, ys) \ seq_lexord R; (ys, zs) \ seq_lexord R; trans R; antisym R \ \ (xs, zs) \ seq_lexord R" + "\ (xs, ys) \ seq_lexord R; (ys, zs) \ seq_lexord R; trans R \ \ (xs, zs) \ seq_lexord R" by (meson seq_lexord_transitive transE) lemma seq_lexord_antisym: "\ asym R; (a, b) \ seq_lexord R \ \ (b, a) \ seq_lexord R" by (auto dest: sinit_linear_asym_lemma2 simp add: seq_lexord_def) lemma seq_lexord_asym: assumes "asym R" shows "asym (seq_lexord R)" by (meson assms asym.simps seq_lexord_antisym seq_lexord_irrefl) lemma seq_lexord_total: assumes "total R" shows "total (seq_lexord R)" using assms by (auto simp add: total_on_def seq_lexord_def, meson lexord_linear sinit_ext) lemma seq_lexord_linear: assumes "(\ a b. (a,b)\ R \ a = b \ (b,a) \ R)" shows "(x,y) \ seq_lexord R \ x = y \ (y,x) \ seq_lexord R" proof - have "total R" using assms total_on_def by blast hence "total (seq_lexord R)" using seq_lexord_total by blast thus ?thesis by (auto simp add: total_on_def) qed instantiation seq :: (ord) ord begin definition less_seq :: "'a seq \ 'a seq \ bool" where "less_seq xs ys \ (xs, ys) \ seq_lexord {(xs, ys). xs < ys}" definition less_eq_seq :: "'a seq \ 'a seq \ bool" where "less_eq_seq xs ys = (xs = ys \ xs < ys)" instance .. end instance seq :: (order) order proof - have \
: "antisym {(xs, ys::'a). xs < ys}" "trans {(xs, ys::'a). xs < ys}" - using antisym_def trans_def by fastforce+ - show "xs \ xs" for xs :: "'a seq" - by (simp add: less_eq_seq_def) - show "xs \ zs" if "xs \ ys" and "ys \ zs" for xs ys zs :: "'a seq" - using that - by (auto intro: \
seq_lexord_trans simp: less_eq_seq_def less_seq_def) - show "xs = ys" if "xs \ ys" and "ys \ xs" for xs ys :: "'a seq" - using that + fix xs :: "'a seq" + show "xs \ xs" by (simp add: less_eq_seq_def) +next + fix xs ys zs :: "'a seq" + assume "xs \ ys" and "ys \ zs" + then show "xs \ zs" + by (force dest: seq_lexord_trans simp add: less_eq_seq_def less_seq_def trans_def) +next + fix xs ys :: "'a seq" + assume "xs \ ys" and "ys \ xs" + then show "xs = ys" apply (auto simp add: less_eq_seq_def less_seq_def) - by (metis \
case_prodD less_irrefl mem_Collect_eq seq_lexord_irreflexive seq_lexord_trans) - show "xs < ys \ xs \ ys \ \ ys \ xs" for xs ys :: "'a seq" - apply (auto simp add: less_seq_def less_eq_seq_def seq_lexord_irreflexive) - by (metis \
case_prodD less_irrefl mem_Collect_eq seq_lexord_irreflexive seq_lexord_trans) + apply (rule seq_lexord_irreflexive [THEN notE]) + defer + apply (rule seq_lexord_trans) + apply (auto intro: transI) + done +next + fix xs ys :: "'a seq" + show "xs < ys \ xs \ ys \ \ ys \ xs" + apply (auto simp add: less_seq_def less_eq_seq_def) + defer + apply (rule seq_lexord_irreflexive [THEN notE]) + apply auto + apply (rule seq_lexord_irreflexive [THEN notE]) + defer + apply (rule seq_lexord_trans) + apply (auto intro: transI) + apply (simp add: seq_lexord_irreflexive) + done qed instance seq :: (linorder) linorder proof fix xs ys :: "'a seq" have "(xs, ys) \ seq_lexord {(u, v). u < v} \ xs = ys \ (ys, xs) \ seq_lexord {(u, v). u < v}" by (rule seq_lexord_linear) auto then show "xs \ ys \ ys \ xs" by (auto simp add: less_eq_seq_def less_seq_def) qed lemma seq_lexord_mono [mono]: "(\ x y. f x y \ g x y) \ (xs, ys) \ seq_lexord {(x, y). f x y} \ (xs, ys) \ seq_lexord {(x, y). g x y}" apply (auto simp add: seq_lexord_def) apply (metis case_prodD case_prodI lexord_take_index_conv mem_Collect_eq) done fun insort_rel :: "'a rel \ 'a \ 'a list \ 'a list" where "insort_rel R x [] = [x]" | "insort_rel R x (y # ys) = (if (x = y \ (x,y) \ R) then x # y # ys else y # insort_rel R x ys)" inductive sorted_rel :: "'a rel \ 'a list \ bool" where Nil_rel [iff]: "sorted_rel R []" | Cons_rel: "\ y \ set xs. (x = y \ (x, y) \ R) \ sorted_rel R xs \ sorted_rel R (x # xs)" definition list_of_set :: "'a rel \ 'a set \ 'a list" where "list_of_set R = folding.F (insort_rel R) []" lift_definition seq_inj :: "'a seq seq \ 'a seq" is "\ f i. f (fst (prod_decode i)) (snd (prod_decode i))" . lift_definition seq_proj :: "'a seq \ 'a seq seq" is "\ f i j. f (prod_encode (i, j))" . lemma seq_inj_inverse: "seq_proj (seq_inj x) = x" by (transfer, simp) lemma seq_proj_inverse: "seq_inj (seq_proj x) = x" by (transfer, simp) lemma seq_inj: "inj seq_inj" by (metis injI seq_inj_inverse) lemma seq_inj_surj: "bij seq_inj" apply (rule bijI) apply (auto simp add: seq_inj) apply (metis rangeI seq_proj_inverse) done end \ No newline at end of file diff --git a/thys/ZFC_in_HOL/Kirby.thy b/thys/ZFC_in_HOL/Kirby.thy --- a/thys/ZFC_in_HOL/Kirby.thy +++ b/thys/ZFC_in_HOL/Kirby.thy @@ -1,1663 +1,1659 @@ section \Addition and Multiplication of Sets\ theory Kirby imports ZFC_Cardinals begin subsection \Generalised Addition\ text \Source: Laurence Kirby, Addition and multiplication of sets Math. Log. Quart. 53, No. 1, 52-65 (2007) / DOI 10.1002/malq.200610026 @{url "http://faculty.baruch.cuny.edu/lkirby/mlqarticlejan2007.pdf"}\ subsubsection \Addition is a monoid\ instantiation V :: plus begin text\This definition is credited to Tarski\ definition plus_V :: "V \ V \ V" where "plus_V x \ transrec (\f z. x \ set (f ` elts z))" instance .. end definition lift :: "V \ V \ V" where "lift x y \ set (plus x ` elts y)" lemma plus: "x + y = x \ set ((+)x ` elts y)" unfolding plus_V_def by (subst transrec) auto lemma plus_eq_lift: "x + y = x \ lift x y" unfolding lift_def using plus by blast text\Lemma 3.2\ lemma lift_sup_distrib: "lift x (a \ b) = lift x a \ lift x b" by (simp add: image_Un lift_def sup_V_def) lemma lift_Sup_distrib: "small Y \ lift x (\ Y) = \ (lift x ` Y)" by (auto simp: lift_def Sup_V_def image_Union) lemma add_Sup_distrib: fixes x::V shows "y \ 0 \ x + (SUP z\elts y. f z) = (SUP z\elts y. x + f z)" by (auto simp: plus_eq_lift SUP_sup_distrib lift_Sup_distrib image_image) lemma Limit_add_Sup_distrib: fixes x::V shows "Limit \ \ x + (SUP z\elts \. f z) = (SUP z\elts \. x + f z)" using add_Sup_distrib by force text\Proposition 3.3(ii)\ instantiation V :: monoid_add begin instance proof show "a + b + c = a + (b + c)" for a b c :: V proof (induction c rule: eps_induct) case (step c) have "(a+b) + c = a + b \ set ((+) (a + b) ` elts c)" by (metis plus) also have "\ = a \ lift a b \ set ((\u. a + (b+u)) ` elts c)" using plus_eq_lift step.IH by auto also have "\ = a \ lift a (b + c)" proof - have "lift a b \ set ((\u. a + (b + u)) ` elts c) = lift a (b + c)" unfolding lift_def by (metis elts_of_set image_image lift_def lift_sup_distrib plus_eq_lift replacement small_elts) then show ?thesis by (simp add: sup_assoc) qed also have "\ = a + (b + c)" using plus_eq_lift by auto finally show ?case . qed show "0 + x = x" for x :: V proof (induction rule: eps_induct) case (step x) then show ?case by (subst plus) auto qed show "x + 0 = x" for x :: V by (subst plus) auto qed end lemma lift_0 [simp]: "lift 0 x = x" by (simp add: lift_def) lemma lift_by0 [simp]: "lift x 0 = 0" by (simp add: lift_def) lemma lift_by1 [simp]: "lift x 1 = set{x}" by (simp add: lift_def) lemma add_eq_0_iff [simp]: fixes x y::V shows "x+y = 0 \ x=0 \ y=0" proof safe show "x = 0" if "x + y = 0" by (metis that le_imp_less_or_eq not_less_0 plus sup_ge1) then show "y = 0" if "x + y = 0" using that by auto qed auto lemma plus_vinsert: "x + vinsert z y = vinsert (x+z) (x + y)" proof - have f1: "elts (x + y) = elts x \ (+) x ` elts y" by (metis elts_of_set lift_def plus_eq_lift replacement small_Un small_elts sup_V_def) moreover have "lift x (vinsert z y) = set ((+) x ` elts (set (insert z (elts y))))" using vinsert_def lift_def by presburger ultimately show ?thesis by (simp add: vinsert_def plus_eq_lift sup_V_def) qed lemma plus_V_succ_right: "x + succ y = succ (x + y)" by (metis plus_vinsert succ_def) lemma succ_eq_add1: "succ x = x + 1" by (simp add: plus_V_succ_right one_V_def) lemma ord_of_nat_add: "ord_of_nat (m+n) = ord_of_nat m + ord_of_nat n" by (induction n) (auto simp: plus_V_succ_right) lemma succ_0_plus_eq [simp]: assumes "\ \ elts \" shows "succ 0 + \ = succ \" proof - obtain n where "\ = ord_of_nat n" using assms elts_\ by blast then show ?thesis by (metis One_nat_def ord_of_nat.simps ord_of_nat_add plus_1_eq_Suc) qed lemma omega_closed_add [intro]: assumes "\ \ elts \" "\ \ elts \" shows "\+\ \ elts \" proof - obtain m n where "\ = ord_of_nat m" "\ = ord_of_nat n" using assms elts_\ by auto then have "\+\ = ord_of_nat (m+n)" using ord_of_nat_add by auto then show ?thesis by (simp add: \_def) qed lemma mem_plus_V_E: assumes l: "l \ elts (x + y)" obtains "l \ elts x" | z where "z \ elts y" "l = x + z" using l by (auto simp: plus [of x y] split: if_split_asm) lemma not_add_less_right: assumes "Ord y" shows "\ (x + y < x)" using assms proof (induction rule: Ord_induct) case (step i) then show ?case by (metis less_le_not_le plus sup_ge1) qed lemma not_add_mem_right: "\ (x + y \ elts x)" by (metis sup_ge1 mem_not_refl plus vsubsetD) text\Proposition 3.3(iii)\ lemma add_not_less_TC_self: "\ x + y \ x" proof (induction y arbitrary: x rule: eps_induct) case (step y) then show ?case using less_TC_imp_not_le plus_eq_lift by fastforce qed lemma TC_sup_lift: "TC x \ lift x y = 0" proof - have "elts (TC x) \ elts (set ((+) x ` elts y)) = {}" using add_not_less_TC_self by (auto simp: less_TC_def) then have "TC x \ set ((+) x ` elts y) = set {}" by (metis inf_V_def) then show ?thesis using lift_def by auto qed lemma lift_lift: "lift x (lift y z) = lift (x+y) z" using add.assoc by (auto simp: lift_def) lemma lift_self_disjoint: "x \ lift x u = 0" by (metis TC_sup_lift arg_subset_TC inf.absorb_iff2 inf_assoc inf_sup_aci(3) lift_0) lemma sup_lift_eq_lift: assumes "x \ lift x u = x \ lift x v" shows "lift x u = lift x v" by (metis (no_types) assms inf_sup_absorb inf_sup_distrib2 lift_self_disjoint sup_commute sup_inf_absorb) subsubsection \Deeper properties of addition\ text\Proposition 3.4(i)\ proposition lift_eq_lift: "lift x y = lift x z \ y = z" proof (induction y arbitrary: z rule: eps_induct) case (step y) show ?case proof (intro vsubsetI order_antisym) show "u \ elts z" if "u \ elts y" for u proof - have "x+u \ elts (lift x z)" using lift_def step.prems that by fastforce then obtain v where "v \ elts z" "x+u = x+v" using lift_def by auto then have "lift x u = lift x v" using sup_lift_eq_lift by (simp add: plus_eq_lift) then have "u=v" using step.IH that by blast then show ?thesis using \v \ elts z\ by blast qed show "u \ elts y" if "u \ elts z" for u proof - have "x+u \ elts (lift x y)" using lift_def step.prems that by fastforce then obtain v where "v \ elts y" "x+u = x+v" using lift_def by auto then have "lift x u = lift x v" using sup_lift_eq_lift by (simp add: plus_eq_lift) then have "u=v" using step.IH by (metis \v \ elts y\) then show ?thesis using \v \ elts y\ by auto qed qed qed corollary inj_lift: "inj_on (lift x) A" by (auto simp: inj_on_def dest: lift_eq_lift) corollary add_right_cancel [iff]: fixes x y z::V shows "x+y = x+z \ y=z" by (metis lift_eq_lift plus_eq_lift sup_lift_eq_lift) corollary add_mem_right_cancel [iff]: fixes x y z::V shows "x+y \ elts (x+z) \ y \ elts z" apply safe apply (metis mem_plus_V_E not_add_mem_right add_right_cancel) by (metis ZFC_in_HOL.ext dual_order.antisym elts_vinsert insert_subset order_refl plus_vinsert) corollary add_le_cancel_left [iff]: fixes x y z::V shows "x+y \ x+z \ y\z" by auto (metis add_mem_right_cancel mem_plus_V_E plus sup_ge1 vsubsetD) corollary add_less_cancel_left [iff]: fixes x y z::V shows "x+y < x+z \ y x \ y = 0" by (auto simp: inf.absorb_iff2 lift_eq_lift lift_self_disjoint) lemma succ_less_\_imp: "succ x < \ \ x < \" by (metis add_le_cancel_left add.right_neutral le_0 le_less_trans succ_eq_add1) text\Proposition 3.5\ lemma card_lift: "vcard (lift x y) = vcard y" proof (rule cardinal_cong) have "bij_betw ((+)x) (elts y) (elts (lift x y))" unfolding bij_betw_def by (simp add: inj_on_def lift_def) then show "elts (lift x y) \ elts y" using eqpoll_def eqpoll_sym by blast qed lemma eqpoll_lift: "elts (lift x y) \ elts y" by (metis card_lift cardinal_eqpoll eqpoll_sym eqpoll_trans) lemma vcard_add: "vcard (x + y) = vcard x \ vcard y" using card_lift [of x y] lift_self_disjoint [of x] by (simp add: plus_eq_lift vcard_disjoint_sup) lemma countable_add: assumes "countable (elts A)" "countable (elts B)" shows "countable (elts (A+B))" proof - have "vcard A \ \0" "vcard B \ \0" using assms countable_iff_le_Aleph0 by blast+ then have "vcard (A+B) \ \0" unfolding vcard_add by (metis Aleph_0 Card_\ InfCard_cdouble_eq InfCard_def cadd_le_mono order_refl) then show ?thesis by (simp add: countable_iff_le_Aleph0) qed text\Proposition 3.6\ proposition TC_add: "TC (x + y) = TC x \ lift x (TC y)" proof (induction y rule: eps_induct) case (step y) have *: "\ (TC ` (+) x ` elts y) = TC x \ (SUP u\elts y. TC (set ((+) x ` elts u)))" if "elts y \ {}" proof - obtain w where "w \ elts y" using \elts y \ {}\ by blast then have "TC x \ TC (x + w)" by (simp add: step.IH) then have \: "TC x \ (SUP w\elts y. TC (x + w))" using \w \ elts y\ by blast show ?thesis using that apply (intro conjI ballI impI order_antisym; clarsimp simp add: image_comp \) apply(metis TC_sup_distrib Un_iff elts_sup_iff plus) by (metis TC_least Transset_TC arg_subset_TC le_sup_iff plus vsubsetD) qed have "TC (x + y) = (x + y) \ \ (TC ` elts (x + y))" using TC by blast also have "\ = x \ lift x y \ \ (TC ` elts x) \ \ ((\u. TC (x+u)) ` elts y)" apply (simp add: plus_eq_lift image_Un Sup_Un_distrib sup.left_commute sup_assoc TC_sup_distrib SUP_sup_distrib) apply (simp add: lift_def sup.commute sup_aci *) done also have "\ = x \ \ (TC ` elts x) \ lift x y \ \ ((\u. TC x \ lift x (TC u)) ` elts y)" by (simp add: sup_aci step.IH) also have "\ = TC x \ lift x y \ \ ((\u. lift x (TC u)) ` elts y)" by (simp add: sup_aci SUP_sup_distrib flip: TC [of x]) also have "\ = TC x \ lift x (y \ \ (TC ` elts y))" by (metis (no_types) elts_of_set lift_Sup_distrib image_image lift_sup_distrib replacement small_elts sup_assoc) also have "\ = TC x \ lift x (TC y)" by (simp add: TC [of y]) finally show ?case . qed corollary TC_add': "z \ x + y \ z \ x \ (\v. v \ y \ z = x + v)" using TC_add by (force simp: less_TC_def lift_def) text\Corollary 3.7\ corollary vcard_TC_add: "vcard (TC (x+y)) = vcard (TC x) \ vcard (TC y)" by (simp add: TC_add TC_sup_lift card_lift vcard_disjoint_sup) text\Corollary 3.8\ corollary TC_lift: assumes "y \ 0" shows "TC (lift x y) = TC x \ lift x (TC y)" proof - have "TC (lift x y) = lift x y \ \ ((\u. TC(x+u)) ` elts y)" unfolding TC [of "lift x y"] by (simp add: lift_def image_image) also have "\ = lift x y \ (SUP u\elts y. TC x \ lift x (TC u))" by (simp add: TC_add) also have "\ = lift x y \ TC x \ (SUP u\elts y. lift x (TC u))" using assms by (auto simp: SUP_sup_distrib) also have "\ = TC x \ lift x (TC y)" by (simp add: TC [of y] sup_aci image_image lift_sup_distrib lift_Sup_distrib) finally show ?thesis . qed proposition rank_add_distrib: "rank (x+y) = rank x + rank y" proof (induction y rule: eps_induct) case (step y) show ?case proof (cases "y=0") case False then obtain e where e: "e \ elts y" by fastforce have "rank (x+y) = (SUP u\elts (x \ ZFC_in_HOL.set ((+) x ` elts y)). succ (rank u))" by (metis plus rank_Sup) also have "\ = (SUP x\elts x. succ (rank x)) \ (SUP z\elts y. succ (rank x + rank z))" apply (simp add: Sup_Un_distrib image_Un image_image) apply (simp add: step cong: SUP_cong_simp) done also have "\ = (SUP z \ elts y. rank x + succ (rank z))" proof - have "rank x \ (SUP z\elts y. ZFC_in_HOL.succ (rank x + rank z))" using \y \ 0\ by (auto simp: plus_eq_lift intro: order_trans [OF _ cSUP_upper [OF e]]) then show ?thesis by (force simp: plus_V_succ_right simp flip: rank_Sup [of x] intro!: order_antisym) qed also have "\ = rank x + (SUP z \ elts y. succ (rank z))" by (simp add: add_Sup_distrib False) also have "\ = rank x + rank y" by (simp add: rank_Sup [of y]) finally show ?thesis . qed auto qed lemma Ord_add [simp]: "\Ord x; Ord y\ \ Ord (x+y)" proof (induction y rule: eps_induct) case (step y) then show ?case by (metis Ord_rank rank_add_distrib rank_of_Ord) qed lemma add_Sup_distrib_id: "A \ 0 \ x + \(elts A) = (SUP z\elts A. x + z)" by (metis add_Sup_distrib image_ident image_image) lemma add_Limit: "Limit \ \ x + \ = (SUP z\elts \. x + z)" by (metis Limit_add_Sup_distrib Limit_eq_Sup_self image_ident image_image) lemma add_le_left: assumes "Ord \" "Ord \" shows "\ \ \+\" using \Ord \\ proof (induction rule: Ord_induct3) case 0 then show ?case by auto next case (succ \) then show ?case by (auto simp: plus_V_succ_right Ord_mem_iff_lt assms(1)) next case (Limit \) then have k: "\ = (SUP \ \ elts \. \)" by (simp add: Limit_eq_Sup_self) also have "\ \ (SUP \ \ elts \. \ + \)" using Limit.IH by auto also have "\ = \ + (SUP \ \ elts \. \)" using Limit.hyps Limit_add_Sup_distrib by presburger finally show ?case using k by simp qed lemma plus_\_equals_\: assumes "\ \ elts \" shows "\ + \ = \" proof (rule antisym) show "\ + \ \ \" using Ord_trans assms by (auto simp: elim!: mem_plus_V_E) show "\ \ \ + \" by (simp add: add_le_left assms) qed lemma one_plus_\_equals_\ [simp]: "1 + \ = \" by (simp add: one_V_def plus_\_equals_\) subsubsection \Cancellation / set subtraction\ definition vle :: "V \ V \ bool" (infix "\" 50) where "x \ y \ \z::V. x+z = y" lemma vle_refl [iff]: "x \ x" by (metis (no_types) add.right_neutral vle_def) lemma vle_antisym: "\x \ y; y \ x\ \ x = y" by (metis V_equalityI plus_eq_lift sup_ge1 vle_def vsubsetD) lemma vle_trans [trans]: "\x \ y; y \ z\ \ x \ z" by (metis add.assoc vle_def) definition vle_comparable :: "V \ V \ bool" where "vle_comparable x y \ x \ y \ y \ x" text\Lemma 3.13\ lemma comparable: assumes "a+b = c+d" shows "vle_comparable a c" unfolding vle_comparable_def proof (rule ccontr) assume non: "\ (a \ c \ c \ a)" let ?\ = "\x. \z. a+x \ c+z" have "?\ x" for x proof (induction x rule: eps_induct) case (step x) show ?case proof (cases "x=0") case True with non nonzero_less_TC show ?thesis using vle_def by auto next case False then obtain v where "v \ elts x" using trad_foundation by blast show ?thesis proof clarsimp fix z assume eq: "a + x = c + z" then have "z \ 0" using vle_def non by auto have av: "a+v \ elts (a+x)" by (simp add: \v \ elts x\) moreover have "a+x = c \ lift c z" using eq plus_eq_lift by fastforce ultimately have "a+v \ elts (c \ lift c z)" by simp moreover define u where "u \ set (elts x - {v})" have u: "v \ elts u" and xeq: "x = vinsert v u" using \v \ elts x\ by (auto simp: u_def intro: order_antisym) have case1: "a+v \ elts c" proof assume avc: "a + v \ elts c" then have "a \ c" by clarify (metis Un_iff elts_sup_iff eq mem_not_sym mem_plus_V_E plus_eq_lift) moreover have "a \ lift a x = c \ lift c z" using eq by (simp add: plus_eq_lift) ultimately have "lift c z \ lift a x" by (metis inf.absorb_iff2 inf_commute inf_sup_absorb inf_sup_distrib2 lift_self_disjoint sup.commute) also have "\ = vinsert (a+v) (lift a u)" by (simp add: lift_def vinsert_def xeq) finally have *: "lift c z \ vinsert (a + v) (lift a u)" . have "lift c z \ lift a u" proof - have "a + v \ elts (lift c z)" using lift_self_disjoint [of c z] avc V_disjoint_iff by auto then show ?thesis using * less_eq_V_def by auto qed { fix e assume "e \ elts z" then have "c+e \ elts (lift c z)" by (simp add: lift_def) then have "c+e \ elts (lift a u)" using \lift c z \ lift a u\ by blast then obtain y where "y \ elts u" "c+e = a+y" using lift_def by auto then have False by (metis elts_vinsert insert_iff step.IH xeq) } then show False using \z \ 0\ by fastforce qed ultimately show False by (metis (no_types) \v \ elts x\ av case1 eq mem_plus_V_E step.IH) qed qed qed then show False using assms by blast qed lemma vle1: "x \ y \ x \ y" using vle_def plus_eq_lift by auto lemma vle2: "x \ y \ x \ y" by (metis (full_types) TC_add' add.right_neutral le_TC_def vle_def nonzero_less_TC) lemma vle_iff_le_Ord: assumes "Ord \" "Ord \" shows "\ \ \ \ \ \ \" proof show "\ \ \" if "\ \ \" using that by (simp add: vle1) show "\ \ \" if "\ \ \" using \Ord \\ \Ord \\ that proof (induction \ arbitrary: \ rule: Ord_induct) case (step \) then show ?case unfolding vle_def by (metis Ord_add Ord_linear add_le_left mem_not_refl mem_plus_V_E vsubsetD) qed qed lemma add_le_cancel_left0 [iff]: fixes x::V shows "x \ x+z" by (simp add: vle1 vle_def) lemma add_less_cancel_left0 [iff]: fixes x::V shows "x < x+z \ 0 \ \" "Ord \" "Ord \" obtains \ where "\+\ = \" "\ \ \" "Ord \" proof - obtain \ where \: "\+\ = \" "\ \ \" by (metis add_le_cancel_left add_le_left assms vle_def vle_iff_le_Ord) then have "Ord \" using Ord_def Transset_def \Ord \\ by force with \ that show thesis by blast qed lemma plus_Ord_le: assumes "\ \ elts \" "Ord \" shows "\+\ \ \+\" proof (cases "\ \ elts \") case True with assms have "\+\ = \+\" by (auto simp: elts_\ add.commute ord_of_nat_add [symmetric]) then show ?thesis by simp next case False then have "\ \ \" using Ord_linear2 Ord_mem_iff_lt \Ord \\ by auto then obtain \ where "\+\ = \" "\ \ \" "Ord \" using \Ord \\ le_Ord_diff by auto then have "\+\ = \" by (metis add.assoc assms(1) plus_\_equals_\) then show ?thesis by simp qed lemma add_right_mono: "\\ \ \; Ord \; Ord \; Ord \\ \ \+\ \ \+\" by (metis add_le_cancel_left add.assoc add_le_left le_Ord_diff) lemma add_strict_mono: "\\ < \; \ < \; Ord \; Ord \; Ord \; Ord \\ \ \+\ < \+\" by (metis order.strict_implies_order add_less_cancel_left add_right_mono le_less_trans) lemma add_right_strict_mono: "\\ \ \; \ < \; Ord \; Ord \; Ord \; Ord \\ \ \+\ < \+\" using add_strict_mono le_imp_less_or_eq by blast lemma Limit_add_Limit [simp]: assumes "Limit \" "Ord \" shows "Limit (\ + \)" unfolding Limit_def proof (intro conjI allI impI) show "Ord (\ + \)" using Limit_def assms by auto show "0 \ elts (\ + \)" using Limit_def add_le_left assms by auto next fix \ assume "\ \ elts (\ + \)" then consider "\ \ elts \" | \ where "\ \ elts \" "\ = \ + \" using mem_plus_V_E by blast then show "succ \ \ elts (\ + \)" proof cases case 1 then show ?thesis by (metis Kirby.add_strict_mono Limit_def Ord_add Ord_in_Ord Ord_mem_iff_lt assms one_V_def succ_eq_add1) next case 2 then show ?thesis by (metis Limit_def add_mem_right_cancel assms(1) plus_V_succ_right) qed qed subsection \Generalised Difference\ definition odiff where "odiff y x \ THE z::V. (x+z = y) \ (z=0 \ \ x \ y)" lemma vle_imp_odiff_eq: "x \ y \ x + (odiff y x) = y" by (auto simp: vle_def odiff_def) lemma not_vle_imp_odiff_0: "\ x \ y \ (odiff y x) = 0" by (auto simp: vle_def odiff_def) lemma Ord_odiff_eq: assumes "\ \ \" "Ord \" "Ord \" shows "\ + odiff \ \ = \" by (simp add: assms vle_iff_le_Ord vle_imp_odiff_eq) lemma Ord_odiff: assumes "Ord \" "Ord \" shows "Ord (odiff \ \)" proof (cases "\ \ \") case True then show ?thesis by (metis add_right_cancel assms le_Ord_diff vle1 vle_imp_odiff_eq) next case False then show ?thesis by (simp add: odiff_def vle_def) qed lemma Ord_odiff_le: assumes "Ord \" "Ord \" shows "odiff \ \ \ \" proof (cases "\ \ \") case True then show ?thesis by (metis add_right_cancel assms le_Ord_diff vle1 vle_imp_odiff_eq) next case False then show ?thesis by (simp add: odiff_def vle_def) qed lemma odiff_0_right [simp]: "odiff x 0 = x" by (metis add.left_neutral vle_def vle_imp_odiff_eq) lemma odiff_succ: "y \ x \ odiff (succ x) y = succ (odiff x y)" unfolding odiff_def by (metis add_right_cancel odiff_def plus_V_succ_right vle_def vle_imp_odiff_eq) lemma odiff_eq_iff: "z \ x \ odiff x z = y \ x = z + y" by (auto simp: odiff_def vle_def) lemma odiff_le_iff: "z \ x \ odiff x z \ y \ x \ z + y" by (auto simp: odiff_def vle_def) lemma odiff_less_iff: "z \ x \ odiff x z < y \ x < z + y" by (auto simp: odiff_def vle_def) lemma odiff_ge_iff: "z \ x \ odiff x z \ y \ x \ z + y" by (auto simp: odiff_def vle_def) lemma Ord_odiff_le_iff: "\\ \ x; Ord x; Ord \\ \ odiff x \ \ y \ x \ \ + y" by (simp add: odiff_le_iff vle_iff_le_Ord) lemma odiff_le_odiff: assumes "x \ y" shows "odiff x z \ odiff y z" proof (cases "z \ x") case True then show ?thesis using assms odiff_le_iff vle1 vle_imp_odiff_eq vle_trans by presburger next case False then show ?thesis by (simp add: not_vle_imp_odiff_0) qed lemma Ord_odiff_le_odiff: "\x \ y; Ord x; Ord y\ \ odiff x \ \ odiff y \" by (simp add: odiff_le_odiff vle_iff_le_Ord) lemma Ord_odiff_less_odiff: "\\ \ x; x < y; Ord x; Ord y; Ord \\ \ odiff x \ < odiff y \" by (metis Ord_odiff_eq Ord_odiff_le_odiff dual_order.strict_trans less_V_def) lemma Ord_odiff_less_imp_less: "\odiff x \ < odiff y \; Ord x; Ord y\ \ x < y" by (meson Ord_linear2 leD odiff_le_odiff vle_iff_le_Ord) lemma odiff_add_cancel [simp]: "odiff (x + y) x = y" by (simp add: odiff_eq_iff vle_def) lemma odiff_add_cancel_0 [simp]: "odiff x x = 0" by (simp add: odiff_eq_iff) lemma odiff_add_cancel_both [simp]: "odiff (x + y) (x + z) = odiff y z" by (simp add: add.assoc odiff_def vle_def) subsection \Generalised Multiplication\ text \Credited to Dana Scott\ instantiation V :: times begin text\This definition is credited to Tarski\ definition times_V :: "V \ V \ V" where "times_V x \ transrec (\f y. \ ((\u. lift (f u) x) ` elts y))" instance .. end lemma mult: "x * y = (SUP u\elts y. lift (x * u) x)" unfolding times_V_def by (subst transrec) (force simp:) lemma elts_multE: assumes "z \ elts (x * y)" obtains u v where "u \ elts x" "v \ elts y" "z = x*v + u" using mult [of x y] lift_def assms by auto text \Lemma 4.2\ lemma mult_zero_right [simp]: fixes x::V shows "x * 0 = 0" by (metis ZFC_in_HOL.Sup_empty elts_0 image_empty mult) lemma mult_insert: "x * (vinsert y z) = x*z \ lift (x*y) x" by (metis (no_types, lifting) elts_vinsert image_insert replacement small_elts sup_commute mult Sup_V_insert) lemma mult_succ: "x * succ y = x*y + x" by (simp add: mult_insert plus_eq_lift succ_def) lemma ord_of_nat_mult: "ord_of_nat (m*n) = ord_of_nat m * ord_of_nat n" proof (induction n) case (Suc n) then show ?case by (simp add: add.commute [of m]) (simp add: ord_of_nat_add mult_succ) qed auto lemma omega_closed_mult [intro]: assumes "\ \ elts \" "\ \ elts \" shows "\*\ \ elts \" proof - obtain m n where "\ = ord_of_nat m" "\ = ord_of_nat n" using assms elts_\ by auto then have "\*\ = ord_of_nat (m*n)" by (simp add: ord_of_nat_mult) then show ?thesis by (simp add: \_def) qed lemma zero_imp_le_mult: "0 \ elts y \ x \ x*y" by (auto simp: mult [of x y]) subsubsection\Proposition 4.3\ lemma mult_zero_left [simp]: fixes x::V shows "0 * x = 0" proof (induction x rule: eps_induct) case (step x) then show ?case by (subst mult) auto qed lemma mult_sup_distrib: fixes x::V shows "x * (y \ z) = x*y \ x*z" unfolding mult [of x "y \ z"] mult [of x y] mult [of x z] by (simp add: Sup_Un_distrib image_Un) lemma mult_Sup_distrib: "small Y \ x * (\Y) = \ ((*) x ` Y)" for Y:: "V set" unfolding mult [of x "\Y"] by (simp add: cSUP_UNION) (metis mult) lemma mult_lift_imp_distrib: "x * (lift y z) = lift (x*y) (x*z) \ x * (y+z) = x*y + x*z" by (simp add: mult_sup_distrib plus_eq_lift) lemma mult_lift: "x * (lift y z) = lift (x*y) (x*z)" proof (induction z rule: eps_induct) case (step z) have "x * lift y z = (SUP u\elts (lift y z). lift (x * u) x)" using mult by blast also have "\ = (SUP v\elts z. lift (x * (y + v)) x)" using lift_def by auto also have "\ = (SUP v\elts z. lift (x * y + x * v) x)" using mult_lift_imp_distrib step.IH by auto also have "\ = (SUP v\elts z. lift (x * y) (lift (x * v) x))" by (simp add: lift_lift) also have "\ = lift (x * y) (SUP v\elts z. lift (x * v) x)" by (simp add: image_image lift_Sup_distrib) also have "\ = lift (x*y) (x*z)" by (metis mult) finally show ?case . qed lemma mult_Limit: "Limit \ \ x * \ = \ ((*) x ` elts \)" by (metis Limit_eq_Sup_self mult_Sup_distrib small_elts) lemma add_mult_distrib: "x * (y+z) = x*y + x*z" for x::V by (simp add: mult_lift mult_lift_imp_distrib) instantiation V :: monoid_mult begin instance proof show "1 * x = x" for x :: V proof (induction x rule: eps_induct) case (step x) then show ?case by (subst mult) auto qed show "x * 1 = x" for x :: V by (subst mult) auto show "(x * y) * z = x * (y * z)" for x y z::V proof (induction z rule: eps_induct) case (step z) have "(x * y) * z = (SUP u\elts z. lift (x * y * u) (x * y))" using mult by blast also have "\ = (SUP u\elts z. lift (x * (y * u)) (x * y))" using step.IH by auto also have "\ = (SUP u\elts z. x * lift (y * u) y)" using mult_lift by auto also have "\ = x * (SUP u\elts z. lift (y * u) y)" by (simp add: image_image mult_Sup_distrib) also have "\ = x * (y * z)" by (metis mult) finally show ?case . qed qed end lemma le_mult: assumes "Ord \" "\ \ 0" shows "\ \ \ * \" using assms proof (induction rule: Ord_induct3) case (succ \) then show ?case using mult_insert succ_def by fastforce next case (Limit \) have "\ \ (*) \ ` elts \" using Limit.hyps Limit_def one_V_def by (metis imageI mult.right_neutral) then have "\ \ \ ((*) \ ` elts \)" by auto then show ?case by (simp add: Limit.hyps mult_Limit) qed auto lemma mult_sing_1 [simp]: fixes x::V shows "x * set{1} = lift x x" by (subst mult) auto lemma mult_2_right [simp]: fixes x::V shows "x * set{0,1} = x+x" by (subst mult) (auto simp: Sup_V_insert plus_eq_lift) lemma Ord_mult [simp]: "\Ord y; Ord x\ \ Ord (x*y)" proof (induction y rule: Ord_induct3) case 0 then show ?case by auto next case (succ k) then show ?case by (simp add: mult_succ) next case (Limit k) then have "Ord (x * \ (elts k))" by (metis Ord_Sup imageE mult_Sup_distrib small_elts) then show ?case using Limit.hyps Limit_eq_Sup_self by auto qed subsubsection \Proposition 4.4-5\ proposition rank_mult_distrib: "rank (x*y) = rank x * rank y" proof (induction y rule: eps_induct) case (step y) have "rank (x*y) = (SUP y\elts (SUP u\elts y. lift (x * u) x). succ (rank y))" by (metis rank_Sup mult) also have "\ = (SUP u\elts y. SUP r\elts x. succ (rank (x * u + r)))" apply (simp add: lift_def image_image image_UN) apply (simp add: Sup_V_def) done also have "\ = (SUP u\elts y. SUP r\elts x. succ (rank (x * u) + rank r))" using rank_add_distrib by auto also have "\ = (SUP u\elts y. SUP r\elts x. succ (rank x * rank u + rank r))" using step arg_cong [where f = Sup] by auto also have "\ = (SUP u\elts y. rank x * rank u + rank x)" proof (rule SUP_cong) show "(SUP r\elts x. succ (rank x * rank u + rank r)) = rank x * rank u + rank x" if "u \ elts y" for u proof (cases "x=0") case False have "(SUP r\elts x. succ (rank x * rank u + rank r)) = rank x * rank u + (SUP y\elts x. succ (rank y))" proof (rule order_antisym) show "(SUP r\elts x. succ (rank x * rank u + rank r)) \ rank x * rank u + (SUP y\elts x. succ (rank y))" by (auto simp: Sup_le_iff simp flip: plus_V_succ_right) have "rank x * rank u + (SUP y\elts x. succ (rank y)) = (SUP y\elts x. rank x * rank u + succ (rank y))" by (simp add: add_Sup_distrib False) also have "\ \ (SUP r\elts x. succ (rank x * rank u + rank r))" using plus_V_succ_right by auto finally show "rank x * rank u + (SUP y\elts x. succ (rank y)) \ (SUP r\elts x. succ (rank x * rank u + rank r))" . qed also have "\ = rank x * rank u + rank x" by (metis rank_Sup) finally show ?thesis . qed auto qed auto also have "\ = rank x * rank y" by (simp add: rank_Sup [of y] mult_Sup_distrib mult_succ image_image) finally show ?case . qed lemma mult_le1: fixes y::V assumes "y \ 0" shows "x \ x * y" proof (cases "x = 0") case False then obtain r where r: "r \ elts x" by fastforce from \y \ 0\ show ?thesis proof (induction y rule: eps_induct) case (step y) show ?case proof (cases "y = 1") case False with \y \ 0\ obtain p where p: "p \ elts y" "p \ 0" by (metis V_equalityI elts_1 insertI1 singletonD trad_foundation) then have "x*p + r \ elts (lift (x*p) x)" by (simp add: lift_def r) moreover have "lift (x*p) x \ x*y" by (metis bdd_above_iff_small cSUP_upper2 order_refl \p \ elts y\ replacement small_elts mult) ultimately have "x*p + r \ elts (x*y)" by blast moreover have "x*p \ x*p + r" by (metis TC_add' V_equalityI add.right_neutral eps_induct le_TC_refl less_TC_iff less_imp_le_TC) ultimately show ?thesis using step.IH [OF p] le_TC_trans less_TC_iff by blast qed auto qed qed auto lemma mult_eq_0_iff [simp]: fixes y::V shows "x * y = 0 \ x=0 \ y=0" proof show "x = 0 \ y = 0" if "x * y = 0" by (metis le_0 le_TC_def less_TC_imp_not_le mult_le1 that) qed auto lemma lift_lemma: assumes "x \ 0" "y \ 0" shows "\ lift (x * y) x \ x" using assms mult_le1 [of concl: x y] by (auto simp: le_TC_def TC_lift less_TC_def less_TC_imp_not_le) lemma mult_le2: fixes y::V assumes "x \ 0" "y \ 0" "y \ 1" shows "x \ x * y" proof - obtain v where v: "v \ elts y" "v \ 0" using assms by fastforce have "x \ x * y" using lift_lemma [of x v] by (metis \x \ 0\ bdd_above_iff_small cSUP_upper2 order_refl replacement small_elts mult v) then show ?thesis using assms mult_le1 [of y x] by (auto simp: le_TC_def) qed lemma elts_mult_\E: assumes "x \ elts (y * \)" obtains n where "n \ 0" "x \ elts (y * ord_of_nat n)" "\m. m < n \ x \ elts (y * ord_of_nat m)" proof - obtain k where k: "k \ 0 \ x \ elts (y * ord_of_nat k)" using assms apply (simp add: mult_Limit elts_\) by (metis mult_eq_0_iff elts_0 ex_in_conv ord_of_eq_0_iff that) define n where "n \ (LEAST k. k \ 0 \ x \ elts (y * ord_of_nat k))" show thesis proof show "n \ 0" "x \ elts (y * ord_of_nat n)" unfolding n_def by (metis (mono_tags, lifting) LeastI_ex k)+ show "\m. m < n \ x \ elts (y * ord_of_nat m)" by (metis (mono_tags, lifting) mult_eq_0_iff elts_0 empty_iff n_def not_less_Least ord_of_eq_0_iff) qed qed subsubsection\Theorem 4.6\ theorem mult_eq_imp_0: assumes "a*x = a*y + b" "b \ a" shows "b=0" proof (cases "a=0 \ x=0") case True with assms show ?thesis by (metis add_le_cancel_left mult_eq_0_iff eq_iff le_0) next case False then have "a\0" "x\0" by auto then show ?thesis proof (cases "y=0") case True then show ?thesis using assms less_asym_TC mult_le2 by force next case False have "b=0" if "Ord \" "x \ elts (Vset \)" "y \ elts (Vset \)" for \ using that assms proof (induction \ arbitrary: x y b rule: Ord_induct3) case 0 then show ?case by auto next case (succ k) define \ where "\ \ \x y. \r. 0 \ r \ r \ a \ a*x = a*y + r" show ?case proof (rule ccontr) assume "b \ 0" then have "0 \ b" by (metis nonzero_less_TC) then have "\ x y" unfolding \_def using succ.prems by blast then obtain x' where "\ x' y" "x' \ x" and min: "\x''. x'' \ x' \ \ \ x'' y" using less_TC_minimal [of "\x. \ x y" x] by blast then obtain b' where "0 \ b'" "b' \ a" and eq: "a*x' = a*y + b'" using \_def by blast have "a*y \ a*x'" using TC_add' \0 \ b'\ eq by auto then obtain p where "p \ elts (a * x')" "a * y \ p" using less_TC_iff by blast then have "p \ elts (a * y)" using less_TC_iff less_irrefl_TC by blast then have "p \ \ (elts ` (\v. lift (a * v) a) ` elts x')" by (metis \p \ elts (a * x')\ elts_Sup replacement small_elts mult) then obtain u c where "u \ elts x'" "c \ elts a" "p = a*u + c" using lift_def by auto then have "p \ elts (lift (a*y) b')" using \p \ elts (a * x')\ \p \ elts (a * y)\ eq plus_eq_lift by auto then obtain d where d: "d \ elts b'" "p = a*y + d" "p = a*u + c" by (metis \p = a * u + c\ \p \ elts (a * x')\ \p \ elts (a * y)\ eq mem_plus_V_E) have noteq: "a*y \ a*u" proof assume "a*y = a*u" then have "lift (a*y) a = lift (a*u) a" by metis also have "\ \ a*x'" unfolding mult [of _ x'] using \u \ elts x'\ by (auto intro: cSUP_upper) also have "\ = a*y \ lift (a*y) b'" by (simp add: eq plus_eq_lift) finally have "lift (a*y) a \ a*y \ lift (a*y) b'" . then have "lift (a*y) a \ lift (a*y) b'" using add_le_cancel_left less_TC_imp_not_le plus_eq_lift \b' \ a\ by auto then have "a \ b'" by (simp add: le_iff_sup lift_eq_lift lift_sup_distrib) then show False using \b' \ a\ less_TC_imp_not_le by auto qed consider "a*y \ a*u" | "a*u \ a*y" using d comparable vle_comparable_def by auto then show False proof cases case 1 then obtain e where e: "a*u = a*y + e" "e \ 0" by (metis add.right_neutral noteq vle_def) moreover have "e + c = d" by (metis e add_right_cancel \p = a * u + c\ \p = a * y + d\ add.assoc) with \d \ elts b'\ \b' \ a\ have "e \ a" by (meson less_TC_iff less_TC_trans vle2 vle_def) ultimately show False \\contradicts minimality of @{term x'}\ using min unfolding \_def by (meson \u \ elts x'\ le_TC_def less_TC_iff nonzero_less_TC) next case 2 then obtain e where e: "a*y = a*u + e" "e \ 0" by (metis add.right_neutral noteq vle_def) moreover have "e + d = c" by (metis e add_right_cancel \p = a * u + c\ \p = a * y + d\ add.assoc) with \d \ elts b'\ \b' \ a\ have "e \ a" by (metis \c \ elts a\ less_TC_iff vle2 vle_def) ultimately have "\ y u" unfolding \_def using nonzero_less_TC by blast then obtain y' where "\ y' u" "y' \ y" and min: "\x''. x'' \ y' \ \ \ x'' u" using less_TC_minimal [of "\x. \ x u" y] by blast then obtain b' where "0 \ b'" "b' \ a" and eq: "a*y' = a*u + b'" using \_def by blast have u_k: "u \ elts (Vset k)" using \u \ elts x'\ \x' \ x\ succ Vset_succ_TC less_TC_iff less_le_TC_trans by blast have "a*u \ a*y'" using TC_add' \0 \ b'\ eq by auto then obtain p where "p \ elts (a * y')" "a * u \ p" using less_TC_iff by blast then have "p \ elts (a * u)" using less_TC_iff less_irrefl_TC by blast then have "p \ \ (elts ` (\v. lift (a * v) a) ` elts y')" by (metis \p \ elts (a * y')\ elts_Sup replacement small_elts mult) then obtain v c where "v \ elts y'" "c \ elts a" "p = a*v + c" using lift_def by auto then have "p \ elts (lift (a*u) b')" using \p \ elts (a * y')\ \p \ elts (a * u)\ eq plus_eq_lift by auto then obtain d where d: "d \ elts b'" "p = a*u + d" "p = a*v + c" by (metis \p = a * v + c\ \p \ elts (a * y')\ \p \ elts (a * u)\ eq mem_plus_V_E) have v_k: "v \ elts (Vset k)" using Vset_succ_TC \v \ elts y'\ \y' \ y\ less_TC_iff less_le_TC_trans succ.hyps succ.prems(2) by blast have noteq: "a*u \ a*v" proof assume "a*u = a*v" then have "lift (a*v) a \ a*y'" unfolding mult [of _ y'] using \v \ elts y'\ by (auto intro: cSUP_upper) also have "\ = a*u \ lift (a*u) b'" by (simp add: eq plus_eq_lift) finally have "lift (a*v) a \ a*u \ lift (a*u) b'" . then have "lift (a*u) a \ lift (a*u) b'" by (metis \a * u = a * v\ le_iff_sup lift_sup_distrib sup_left_commute sup_lift_eq_lift) then have "a \ b'" by (simp add: le_iff_sup lift_eq_lift lift_sup_distrib) then show False using \b' \ a\ less_TC_imp_not_le by auto qed consider "a*u \ a*v" | "a*v \ a*u" using d comparable vle_comparable_def by auto then show False proof cases case 1 then obtain e where e: "a*v = a*u + e" "e \ 0" by (metis add.right_neutral noteq vle_def) moreover have "e + c = d" by (metis add_right_cancel \p = a * u + d\ \p = a * v + c\ add.assoc e) with \d \ elts b'\ \b' \ a\ have "e \ a" by (meson less_TC_iff less_TC_trans vle2 vle_def) ultimately show False using succ.IH u_k v_k by blast next case 2 then obtain e where e: "a*u = a*v + e" "e \ 0" by (metis add.right_neutral noteq vle_def) moreover have "e + d = c" by (metis add_right_cancel add.assoc d e) with \d \ elts b'\ \b' \ a\ have "e \ a" by (metis \c \ elts a\ less_TC_iff vle2 vle_def) ultimately show False using succ.IH u_k v_k by blast qed qed qed next case (Limit k) obtain i j where k: "i \ elts k" "j \ elts k" and x: "x \ elts (Vset i)" and y: "y \ elts (Vset j)" using that Limit by (auto simp: Limit_Vfrom_eq) show ?case proof (rule Limit.IH [of "i \ j"]) show "i \ j \ elts k" by (meson k x y Limit.hyps Limit_def Ord_in_Ord Ord_mem_iff_lt Ord_sup union_less_iff) show "x \ elts (Vset (i \ j))" "y \ elts (Vset (i \ j))" using x y by (auto simp: Vfrom_sup) qed (use Limit.prems in auto) qed then show ?thesis by (metis two_in_Vset Ord_rank Ord_VsetI rank_lt) qed qed subsubsection\Theorem 4.7\ lemma mult_cancellation_half: assumes "a*x + r \ a*y + s" "r \ a" "s \ a" shows "x \ y" proof - have "x \ y" if "Ord \" "x \ elts (Vset \)" "y \ elts (Vset \)" for \ using that assms proof (induction \ arbitrary: x y r s rule: Ord_induct3) case 0 then show ?case by auto next case (succ k) show ?case proof fix u assume u: "u \ elts x" have u_k: "u \ elts (Vset k)" using Vset_succ succ.hyps succ.prems(1) u by auto obtain r' where "r' \ elts a" "r \ r'" using less_TC_iff succ.prems(4) by blast have "a*u + r' \ elts (lift (a*u) a)" by (simp add: \r' \ elts a\ lift_def) also have "\ \ elts (a*x)" using u by (force simp: mult [of _ x]) also have "\ \ elts (a*y + s)" using plus_eq_lift succ.prems(3) by auto also have "\ = elts (a*y) \ elts (lift (a*y) s)" by (simp add: plus_eq_lift) finally have "a * u + r' \ elts (a * y) \ elts (lift (a * y) s)" . then show "u \ elts y" proof assume *: "a * u + r' \ elts (a * y)" show "u \ elts y" proof - obtain v e where v: "v \ elts y" "e \ elts a" "a * u + r' = a * v + e" using * by (auto simp: mult [of _ y] lift_def) then have v_k: "v \ elts (Vset k)" using Vset_succ_TC less_TC_iff succ.prems(2) by blast then show ?thesis by (metis \r' \ elts a\ antisym le_TC_refl less_TC_iff order_refl succ.IH u_k v) qed next assume "a * u + r' \ elts (lift (a * y) s)" then obtain t where "t \ elts s" and t: "a * u + r' = a * y + t" using lift_def by auto have noteq: "a*y \ a*u" proof assume "a*y = a*u" then have "lift (a*y) a = lift (a*u) a" by metis also have "\ \ a*x" unfolding mult [of _ x] using \u \ elts x\ by (auto intro: cSUP_upper) also have "\ \ a*y \ lift (a*y) s" using \elts (a * x) \ elts (a * y + s)\ plus_eq_lift by auto finally have "lift (a*y) a \ a*y \ lift (a*y) s" . then have "lift (a*y) a \ lift (a*y) s" using add_le_cancel_left less_TC_imp_not_le plus_eq_lift \s \ a\ by auto then have "a \ s" by (simp add: le_iff_sup lift_eq_lift lift_sup_distrib) then show False using \s \ a\ less_TC_imp_not_le by auto qed consider "a * u \ a * y" | "a * y \ a * u" using t comparable vle_comparable_def by blast then have "False" proof cases case 1 then obtain c where "a*y = a*u + c" by (metis vle_def) then have "c+t = r'" by (metis add_right_cancel add.assoc t) then have "c \ a" using \r' \ elts a\ less_TC_iff vle2 vle_def by force moreover have "c \ 0" using \a * y = a * u + c\ noteq by auto ultimately show ?thesis using \a * y = a * u + c\ mult_eq_imp_0 by blast next case 2 then obtain c where "a*u = a*y + c" by (metis vle_def) then have "c+r' = t" by (metis add_right_cancel add.assoc t) then have "c \ a" by (metis \t \ elts s\ less_TC_iff less_TC_trans \s \ a\ vle2 vle_def) moreover have "c \ 0" using \a * u = a * y + c\ noteq by auto ultimately show ?thesis using \a * u = a * y + c\ mult_eq_imp_0 by blast qed then show "u \ elts y" .. qed qed next case (Limit k) obtain i j where k: "i \ elts k" "j \ elts k" and x: "x \ elts (Vset i)" and y: "y \ elts (Vset j)" using that Limit by (auto simp: Limit_Vfrom_eq) show ?case proof (rule Limit.IH [of "i \ j"]) show "i \ j \ elts k" by (meson k x y Limit.hyps Limit_def Ord_in_Ord Ord_mem_iff_lt Ord_sup union_less_iff) show "x \ elts (Vset (i \ j))" "y \ elts (Vset (i \ j))" using x y by (auto simp: Vfrom_sup) show "a * x + r \ a * y + s" by (simp add: Limit.prems) qed (auto simp: Limit.prems) qed then show ?thesis by (metis two_in_Vset Ord_rank Ord_VsetI rank_lt) qed theorem mult_cancellation_lemma: assumes "a*x + r = a*y + s" "r \ a" "s \ a" shows "x=y \ r=s" by (metis assms leD less_V_def mult_cancellation_half odiff_add_cancel order_refl) corollary mult_cancellation [simp]: fixes a::V assumes "a \ 0" shows "a*x = a*y \ x=y" by (metis assms nonzero_less_TC mult_cancellation_lemma) corollary mult_cancellation_less: assumes lt: "a*x + r < a*y + s" and "r \ a" "s \ a" obtains "x < y" | "x = y" "r < s" proof - have "x \ y" by (meson assms dual_order.strict_implies_order mult_cancellation_half) then consider "x < y" | "x = y" using less_V_def by blast with lt that show ?thesis by blast qed corollary lift_mult_TC_disjoint: fixes x::V assumes "x \ y" shows "lift (a*x) (TC a) \ lift (a*y) (TC a) = 0" apply (rule V_equalityI) using assms by (auto simp: less_TC_def inf_V_def lift_def image_iff dest: mult_cancellation_lemma) corollary lift_mult_disjoint: fixes x::V assumes "x \ y" shows "lift (a*x) a \ lift (a*y) a = 0" proof - have "lift (a*x) a \ lift (a*y) a \ lift (a*x) (TC a) \ lift (a*y) (TC a)" by (metis TC' inf_mono lift_sup_distrib sup_ge1) then show ?thesis using assms lift_mult_TC_disjoint by auto qed lemma mult_add_mem: assumes "a*x + r \ elts (a*y)" "r \ a" shows "x \ elts y" "r \ elts a" proof - obtain v s where v: "a * x + r = a * v + s" "v \ elts y" "s \ elts a" using assms unfolding mult [of a y] lift_def by auto then show "x \ elts y" by (metis arg_subset_TC assms(2) less_TC_def mult_cancellation_lemma vsubsetD) show "r \ elts a" by (metis arg_subset_TC assms(2) less_TC_def mult_cancellation_lemma v(1) v(3) vsubsetD) qed lemma mult_add_mem_0 [simp]: "a*x \ elts (a*y) \ x \ elts y \ 0 \ elts a" proof - have "x \ elts y" if "a * x \ elts (a * y) \ 0 \ elts a" using that using mult_add_mem [of a x 0] using nonzero_less_TC by force moreover have "a * x \ elts (a * y)" if "x \ elts y" "0 \ elts a" using that by (force simp: image_iff mult [of a y] lift_def) ultimately show ?thesis by (metis mult_eq_0_iff add.right_neutral mult_add_mem(2) nonzero_less_TC) qed lemma zero_mem_mult_iff: "0 \ elts (x*y) \ 0 \ elts x \ 0 \ elts y" by (metis Kirby.mult_zero_right mult_add_mem_0) lemma zero_less_mult_iff [simp]: "0 < x*y \ 0 < x \ 0 < y" if "Ord x" using Kirby.mult_eq_0_iff ZFC_in_HOL.neq0_conv by blast lemma mult_cancel_less_iff [simp]: "\Ord \; Ord \; Ord \\ \ \*\ < \*\ \ \ < \ \ 0 < \" using mult_add_mem_0 [of \ \ \] by (meson Ord_0 Ord_mem_iff_lt Ord_mult) lemma mult_cancel_le_iff [simp]: "\Ord \; Ord \; Ord \\ \ \*\ \ \*\ \ \ \ \ \ \=0" by (metis Ord_linear2 Ord_mult eq_iff leD mult_cancel_less_iff mult_cancellation) lemma mult_Suc_add_less: "\\ < \; \ < \; Ord \; Ord \; Ord \\ \ \ * ord_of_nat m + \ < \ * ord_of_nat (Suc m) + \" apply (simp add: mult_succ add.assoc) by (meson Ord_add Ord_linear2 le_less_trans not_add_less_right) lemma mult_nat_less_add_less: assumes "m < n" "\ < \" "\ < \" and ord: "Ord \" "Ord \" "Ord \" shows "\ * ord_of_nat m + \ < \ * ord_of_nat n + \" proof - have "Suc m \ n" using \m < n\ by auto have "\ * ord_of_nat m + \ < \ * ord_of_nat (Suc m) + \" using assms mult_Suc_add_less by blast also have "\ \ \ * ord_of_nat n + \" using Ord_mult Ord_ord_of_nat add_right_mono \Suc m \ n\ ord mult_cancel_le_iff ord_of_nat_mono_iff by presburger finally show ?thesis . qed lemma add_mult_less_add_mult: assumes "x < y" "x \ elts \" "y \ elts \" "\ \ elts \" "\ \ elts \" "Ord \" "Ord \" shows "\*x + \ < \*y + \" proof - obtain "Ord x" "Ord y" using Ord_in_Ord assms by blast then obtain \ where "0 \ elts \" "y = x + \" by (metis add.right_neutral \x < y\ le_Ord_diff less_V_def mem_0_Ord) then show ?thesis apply (simp add: add_mult_distrib add.assoc) by (meson OrdmemD add_le_cancel_left0 \\ \ elts \\ \Ord \\ less_le_trans zero_imp_le_mult) qed lemma add_mult_less: assumes "\ \ elts \" "\ \ elts \" "Ord \" "Ord \" shows "\ * \ + \ \ elts (\ * \)" proof - have "Ord \" using Ord_in_Ord assms by blast with assms show ?thesis by (metis Ord_mem_iff_lt Ord_succ add_mem_right_cancel mult_cancel_le_iff mult_succ succ_le_iff vsubsetD) qed lemma Ord_add_mult_iff: assumes "\ \ elts \" "\' \ elts \" "Ord \" "Ord \'" "Ord \" shows "\ * \ + \ \ elts (\ * \' + \') \ \ \ elts \' \ \ = \' \ \ \ elts \'" (is "?lhs \ ?rhs") proof assume L: ?lhs show ?rhs proof (cases "\ \ elts \'") case False with assms have "\ = \'" by (meson L Ord_linear Ord_mult Ord_trans add_mult_less not_add_mem_right) then show ?thesis using L less_V_def by auto qed auto next assume R: ?rhs then show ?lhs proof assume "\ \ elts \'" then obtain \ where "\' = \+\" by (metis OrdmemD assms(3) assms(4) le_Ord_diff less_V_def) show ?lhs using assms by (meson \\ \ elts \'\ add_le_cancel_left0 add_mult_less vsubsetD) next assume "\ = \' \ \ \ elts \'" then show ?lhs using less_V_def by auto qed qed lemma vcard_mult: "vcard (x * y) = vcard x \ vcard y" proof - have 1: "elts (lift (x * u) x) \ elts x" if "u \ elts y" for u by (metis cardinal_eqpoll eqpoll_sym eqpoll_trans card_lift) have 2: "pairwise (\u u'. disjnt (elts (lift (x * u) x)) (elts (lift (x * u') x))) (elts y)" by (simp add: pairwise_def disjnt_def) (metis V_disjoint_iff lift_mult_disjoint) have "x * y = (SUP u\elts y. lift (x * u) x)" using mult by blast then have "elts (x * y) \ (\u\elts y. elts (lift (x * u) x))" by simp also have "\ \ elts y \ elts x" using Union_eqpoll_Times [OF 1 2] . also have "\ \ elts x \ elts y" by (simp add: times_commute_eqpoll) also have "\ \ elts (vcard x) \ elts (vcard y)" using cardinal_eqpoll eqpoll_sym times_eqpoll_cong by blast also have "\ \ elts (vcard x \ vcard y)" by (simp add: cmult_def elts_vcard_VSigma_eqpoll eqpoll_sym) finally have "elts (x * y) \ elts (vcard x \ vcard y)" . then show ?thesis by (metis cadd_cmult_distrib cadd_def cardinal_cong cardinal_idem vsum_0_eqpoll) qed proposition TC_mult: "TC(x * y) = (SUP r \ elts (TC x). SUP u \ elts (TC y). set{x * u + r})" proof (cases "x = 0") case False have *: "TC(x * y) = (SUP u \ elts (TC y). lift (x * u) (TC x))" for y proof (induction y rule: eps_induct) case (step y) have "TC(x * y) = (SUP u \ elts y. TC (lift (x * u) x))" by (simp add: mult [of x y] TC_Sup_distrib image_image) also have "\ = (SUP u \ elts y. TC(x * u) \ lift (x * u) (TC x))" by (simp add: TC_lift False) also have "\ = (SUP u \ elts y. (SUP z \ elts (TC u). lift (x * z) (TC x)) \ lift (x * u) (TC x))" by (simp add: step) also have "\ = (SUP u \ elts (TC y). lift (x * u) (TC x))" by (auto simp: TC' [of y] image_Un Sup_Un_distrib TC_Sup_distrib cSUP_UNION SUP_sup_distrib) finally show ?case . qed show ?thesis by (force simp: * lift_def) qed auto corollary vcard_TC_mult: "vcard (TC(x * y)) = vcard (TC x) \ vcard (TC y)" proof - have "(\u\elts (TC x). \v\elts (TC y). {x * v + u}) = (\u\elts (TC x). (\v. x * v + u) ` elts (TC y))" by (simp add: UNION_singleton_eq_range) also have "\ \ (\x\elts (TC x). elts (lift (TC y * x) (TC y)))" proof (rule UN_eqpoll_UN) show "(\v. x * v + u) ` elts (TC y) \ elts (lift (TC y * u) (TC y))" if "u \ elts (TC x)" for u proof - have "inj_on (\v. x * v + u) (elts (TC y))" by (meson inj_onI less_TC_def mult_cancellation_lemma that) then have "(\v. x * v + u) ` elts (TC y) \ elts (TC y)" by (rule inj_on_image_eqpoll_self) also have "\ \ elts (lift (TC y * u) (TC y))" by (simp add: eqpoll_lift eqpoll_sym) finally show ?thesis . qed show "pairwise (\u ya. disjnt ((\v. x * v + u) ` elts (TC y)) ((\v. x * v + ya) ` elts (TC y))) (elts (TC x))" apply (auto simp: pairwise_def disjnt_def) using less_TC_def mult_cancellation_lemma by blast show "pairwise (\u ya. disjnt (elts (lift (TC y * u) (TC y))) (elts (lift (TC y * ya) (TC y)))) (elts (TC x))" apply (auto simp: pairwise_def disjnt_def) by (metis Int_iff V_disjoint_iff empty_iff lift_mult_disjoint) qed also have "\ = elts (TC y * TC x)" by (metis elts_Sup image_image mult replacement small_elts) finally have "(\u\elts (TC x). \v\elts (TC y). {x * v + u}) \ elts (TC y * TC x)" . then show ?thesis apply (subst cmult_commute) by (simp add: TC_mult cardinal_cong flip: vcard_mult) qed lemma countable_mult: assumes "countable (elts A)" "countable (elts B)" shows "countable (elts (A*B))" proof - have "vcard A \ \0" "vcard B \ \0" using assms countable_iff_le_Aleph0 by blast+ then have "vcard (A*B) \ \0" unfolding vcard_mult by (metis InfCard_csquare_eq cmult_le_mono Aleph_0 Card_\ InfCard_def order_refl) then show ?thesis by (simp add: countable_iff_le_Aleph0) qed subsection \Ordertype properties\ lemma ordertype_image_plus: assumes "Ord \" shows "ordertype ((+) u ` elts \) VWF = \" proof (subst ordertype_VWF_eq_iff) have 1: "(u + x, u + y) \ VWF" if "x \ elts \" "y \ elts \" "x < y" for x y using that by (meson Ord_in_Ord Ord_mem_iff_lt add_mem_right_cancel assms mem_imp_VWF) then have 2: "x < y" if "x \ elts \" "y \ elts \" "(u + x, u + y) \ VWF" for x y using that by (metis Ord_in_Ord Ord_linear_lt VWF_asym assms) show "\f. bij_betw f ((+) u ` elts \) (elts \) \ (\x\(+) u ` elts \. \y\(+) u ` elts \. (f x < f y) = ((x, y) \ VWF))" using 1 2 unfolding bij_betw_def inj_on_def by (rule_tac x="\x. odiff x u" in exI) (auto simp: image_iff) qed (use assms in auto) lemma ordertype_diff: assumes "\ + \ = \" and \: "\ \ elts \" "Ord \" shows "ordertype (elts \ - elts \) VWF = \" proof - have *: "elts \ - elts \ = ((+)\) ` elts \" proof show "elts \ - elts \ \ (+) \ ` elts \" by clarsimp (metis assms(1) image_iff mem_plus_V_E) show "(+) \ ` elts \ \ elts \ - elts \" using assms(1) not_add_mem_right by force qed have "ordertype ((+) \ ` elts \) VWF = \" proof (subst ordertype_VWF_inc_eq) show "elts \ \ ON" "ordertype (elts \) VWF = \" using \ elts_subset_ON ordertype_eq_Ord by blast+ qed (use "*" assms elts_subset_ON in auto) then show ?thesis by (simp add: *) qed lemma ordertype_interval_eq: assumes \: "Ord \" and \: "Ord \" shows "ordertype ({\ ..< \+\} \ ON) VWF = \" proof - have ON: "(+) \ ` elts \ \ ON" using assms Ord_add Ord_in_Ord by blast have "({\ ..< \+\} \ ON) = (+) \ ` elts \" using assms apply (simp add: image_def set_eq_iff) by (metis add_less_cancel_left Ord_add Ord_in_Ord Ord_linear2 Ord_mem_iff_lt le_Ord_diff not_add_less_right) moreover have "ordertype (elts \) VWF = ordertype ((+) \ ` elts \) VWF" using ON \ elts_subset_ON ordertype_VWF_inc_eq by auto ultimately show ?thesis using \ by auto qed lemma ordertype_Times: assumes "small A" "small B" and r: "wf r" "trans r" "total_on A r" and s: "wf s" "trans s" "total_on B s" shows "ordertype (A\B) (r <*lex*> s) = ordertype B s * ordertype A r" (is "_ = ?\ * ?\") proof (subst ordertype_eq_iff) show "Ord (?\ * ?\)" by (intro wf_Ord_ordertype Ord_mult r s; simp) define f where "f \ \(x,y). ?\ * ordermap A r x + (ordermap B s y)" show "\f. bij_betw f (A \ B) (elts (?\ * ?\)) \ (\x\A \ B. \y\A \ B. (f x < f y) = ((x, y) \ (r <*lex*> s)))" unfolding bij_betw_def proof (intro exI conjI strip) show "inj_on f (A \ B)" proof (clarsimp simp: f_def inj_on_def) fix x y x' y' assume "x \ A" "y \ B" "x' \ A" "y' \ B" and eq: "?\ * ordermap A r x + ordermap B s y = ?\ * ordermap A r x' + ordermap B s y'" have "ordermap A r x = ordermap A r x' \ ordermap B s y = ordermap B s y'" proof (rule mult_cancellation_lemma [OF eq]) show "ordermap B s y \ ?\" using ordermap_in_ordertype [OF \y \ B\, of s] less_TC_iff \small B\ by blast show "ordermap B s y' \ ?\" using ordermap_in_ordertype [OF \y' \ B\, of s] less_TC_iff \small B\ by blast qed then show "x = x' \ y = y'" using \x \ A\ \x' \ A\ \y \ B\ \y' \ B\ r s \small A\ \small B\ by auto qed show "f ` (A \ B) = elts (?\ * ?\)" (is "?lhs = ?rhs") proof show "f ` (A \ B) \ elts (?\ * ?\)" apply (auto simp: f_def add_mult_less ordermap_in_ordertype wf_Ord_ordertype r s) by (simp add: add_mult_less assms ordermap_in_ordertype wf_Ord_ordertype) show "elts (?\ * ?\) \ f ` (A \ B)" proof (clarsimp simp: f_def image_iff elim !: elts_multE split: prod.split) fix u v assume u: "u \ elts (?\)" and v: "v \ elts ?\" have "inv_into B (ordermap B s) u \ B" by (simp add: inv_into_ordermap u) moreover have "inv_into A (ordermap A r) v \ A" by (simp add: inv_into_ordermap v) ultimately show "\x\A. \y\B. ?\ * v + u = ?\ * ordermap A r x + ordermap B s y" by (metis \small A\ \small B\ bij_betw_inv_into_right ordermap_bij r(1) r(3) s(1) s(3) u v) qed qed next fix p q assume "p \ A \ B" and "q \ A \ B" then obtain u v x y where \
: "p = (u,v)" "u \ A" "v \ B" "q = (x,y)" "x \ A" "y \ B" by blast show "((f p) < f q) = ((p, q) \ (r <*lex*> s))" proof assume "f p < f q" with \
assms have "(u, x) \ r \ u=x \ (v, y) \ s" apply (simp add: f_def) by (metis Ord_add Ord_add_mult_iff Ord_mem_iff_lt Ord_mult Ord_ordermap converse_ordermap_mono ordermap_eq_iff ordermap_in_ordertype wf_Ord_ordertype) then show "(p,q) \ (r <*lex*> s)" - using "\
" r(1) by auto + by (simp add: \
) next assume "(p,q) \ (r <*lex*> s)" then have "(u, x) \ r \ u = x \ (v, y) \ s" - using \
by auto + by (simp add: \
) then show "f p < f q" proof assume ux: "(u, x) \ r" have oo: "\x. Ord (ordermap A r x)" "\y. Ord (ordermap B s y)" by (simp_all add: r s) show "f p < f q" proof (clarsimp simp: f_def split: prod.split) fix a b a' b' assume "p = (a, b)" and "q = (a', b')" then have "?\ * ordermap A r a + ordermap B s b < ?\ * ordermap A r a'" using ux assms \
by (metis Ord_mult Ord_ordermap OrdmemD Pair_inject add_mult_less ordermap_in_ordertype ordermap_mono wf_Ord_ordertype) also have "\ \ ?\ * ordermap A r a' + ordermap B s b'" by simp finally show "?\ * ordermap A r a + ordermap B s b < ?\ * ordermap A r a' + ordermap B s b'" . qed next assume "u = x \ (v, y) \ s" then show "f p < f q" using \
assms by (fastforce simp: f_def split: prod.split intro: ordermap_mono_less) qed qed qed - have "antisym r" - by (meson antisymI r(1) wf_not_sym) - then show "trans (r <*lex*> s)" - by (simp add: r(2) s(2)) qed (use assms small_Times in auto) end