diff --git a/thys/AVL-Trees/AVL.thy b/thys/AVL-Trees/AVL.thy --- a/thys/AVL-Trees/AVL.thy +++ b/thys/AVL-Trees/AVL.thy @@ -1,652 +1,652 @@ (* Title: AVL Trees Author: Tobias Nipkow and Cornelia Pusch, converted to Isar by Gerwin Klein contributions by Achim Brucker, Burkhart Wolff and Jan Smaus delete formalization and a transformation to Isar by Ondrej Kuncar Maintainer: Gerwin Klein see the file Changelog for a list of changes *) section "AVL Trees" theory AVL imports Main begin text \ This is a monolithic formalization of AVL trees. \ subsection \AVL tree type definition\ datatype (set_of: 'a) tree = ET | MKT 'a "'a tree" "'a tree" nat subsection \Invariants and auxiliary functions\ primrec height :: "'a tree \ nat" where "height ET = 0" | "height (MKT x l r h) = max (height l) (height r) + 1" primrec avl :: "'a tree \ bool" where "avl ET = True" | "avl (MKT x l r h) = ((height l = height r \ height l = height r + 1 \ height r = height l + 1) \ h = max (height l) (height r) + 1 \ avl l \ avl r)" primrec is_ord :: "('a::order) tree \ bool" where "is_ord ET = True" | "is_ord (MKT n l r h) = ((\n' \ set_of l. n' < n) \ (\n' \ set_of r. n < n') \ is_ord l \ is_ord r)" subsection \AVL interface and implementation\ primrec is_in :: "('a::order) \ 'a tree \ bool" where "is_in k ET = False" | "is_in k (MKT n l r h) = (if k = n then True else if k < n then (is_in k l) else (is_in k r))" primrec ht :: "'a tree \ nat" where "ht ET = 0" | "ht (MKT x l r h) = h" definition mkt :: "'a \ 'a tree \ 'a tree \ 'a tree" where "mkt x l r = MKT x l r (max (ht l) (ht r) + 1)" fun mkt_bal_l where "mkt_bal_l n l r = ( if ht l = ht r + 2 then (case l of MKT ln ll lr _ \ (if ht ll < ht lr then case lr of MKT lrn lrl lrr _ \ mkt lrn (mkt ln ll lrl) (mkt n lrr r) else mkt ln ll (mkt n lr r))) else mkt n l r )" fun mkt_bal_r where "mkt_bal_r n l r = ( if ht r = ht l + 2 then (case r of MKT rn rl rr _ \ (if ht rl > ht rr then case rl of MKT rln rll rlr _ \ mkt rln (mkt n l rll) (mkt rn rlr rr) else mkt rn (mkt n l rl) rr)) else mkt n l r )" primrec insert :: "'a::order \ 'a tree \ 'a tree" where "insert x ET = MKT x ET ET 1" | "insert x (MKT n l r h) = (if x=n then MKT n l r h else if x 'a tree \ 'a tree" where "delete _ ET = ET" | "delete x (MKT n l r h) = ( if x = n then delete_root (MKT n l r h) else if x < n then let l' = delete x l in mkt_bal_r n l' r else let r' = delete x r in mkt_bal_l n l r' )" subsection \Correctness proof\ subsubsection \Insertion maintains AVL balance\ declare Let_def [simp] lemma [simp]: "avl t \ ht t = height t" by (induct t) simp_all lemma height_mkt_bal_l: "\ height l = height r + 2; avl l; avl r \ \ height (mkt_bal_l n l r) = height r + 2 \ height (mkt_bal_l n l r) = height r + 3" by (cases l) (auto simp:mkt_def split:tree.split) lemma height_mkt_bal_r: "\ height r = height l + 2; avl l; avl r \ \ height (mkt_bal_r n l r) = height l + 2 \ height (mkt_bal_r n l r) = height l + 3" by (cases r) (auto simp add:mkt_def split:tree.split) lemma [simp]: "height(mkt x l r) = max (height l) (height r) + 1" by (simp add: mkt_def) lemma avl_mkt: "\ avl l; avl r; height l = height r \ height l = height r + 1 \ height r = height l + 1 \ \ avl(mkt x l r)" by (auto simp add:max_def mkt_def) lemma height_mkt_bal_l2: "\ avl l; avl r; height l \ height r + 2 \ \ height (mkt_bal_l n l r) = (1 + max (height l) (height r))" by (cases l, cases r) simp_all lemma height_mkt_bal_r2: "\ avl l; avl r; height r \ height l + 2 \ \ height (mkt_bal_r n l r) = (1 + max (height l) (height r))" by (cases l, cases r) simp_all lemma avl_mkt_bal_l: assumes "avl l" "avl r" and "height l = height r \ height l = height r + 1 \ height r = height l + 1 \ height l = height r + 2" shows "avl(mkt_bal_l n l r)" proof(cases l) case ET with assms show ?thesis by (simp add: mkt_def) next case (MKT ln ll lr lh) with assms show ?thesis proof(cases "height l = height r + 2") case True - from True MKT assms show ?thesis by (auto intro!: avl_mkt split: tree.split) arith+ + from True MKT assms show ?thesis by (auto intro!: avl_mkt split: tree.split) next case False with assms show ?thesis by (simp add: avl_mkt) qed qed lemma avl_mkt_bal_r: assumes "avl l" and "avl r" and "height l = height r \ height l = height r + 1 \ height r = height l + 1 \ height r = height l + 2" shows "avl(mkt_bal_r n l r)" proof(cases r) case ET with assms show ?thesis by (simp add: mkt_def) next case (MKT rn rl rr rh) with assms show ?thesis proof(cases "height r = height l + 2") case True - from True MKT assms show ?thesis by (auto intro!: avl_mkt split: tree.split) arith+ + from True MKT assms show ?thesis by (auto intro!: avl_mkt split: tree.split) next case False with assms show ?thesis by (simp add: avl_mkt) qed qed (* It apppears that these two properties need to be proved simultaneously: *) text\Insertion maintains the AVL property:\ theorem avl_insert_aux: assumes "avl t" shows "avl(insert x t)" "(height (insert x t) = height t \ height (insert x t) = height t + 1)" using assms proof (induction t) case (MKT n l r h) case 1 with MKT show ?case proof(cases "x = n") case True with MKT 1 show ?thesis by simp next case False with MKT 1 show ?thesis proof(cases "xx\n\ show ?thesis by (auto simp add:avl_mkt_bal_r simp del:mkt_bal_r.simps) qed qed case 2 from 2 MKT show ?case proof(cases "x = n") case True with MKT 1 show ?thesis by simp next case False with MKT 1 show ?thesis proof(cases "xx < n\ show ?thesis by (auto simp del: mkt_bal_l.simps simp: height_mkt_bal_l2) next case True then consider (a) "height (mkt_bal_l n (AVL.insert x l) r) = height r + 2" | (b) "height (mkt_bal_l n (AVL.insert x l) r) = height r + 3" using MKT 2 by (atomize_elim, intro height_mkt_bal_l) simp_all then show ?thesis proof cases case a with 2 \x < n\ show ?thesis by (auto simp del: mkt_bal_l.simps) next case b with True 1 MKT(2) \x < n\ show ?thesis by (simp del: mkt_bal_l.simps) arith qed qed next case False with MKT 2 show ?thesis proof(cases "height (AVL.insert x r) = height l + 2") case False with MKT 2 \\x < n\ show ?thesis by (auto simp del: mkt_bal_r.simps simp: height_mkt_bal_r2) next case True then consider (a) "height (mkt_bal_r n l (AVL.insert x r)) = height l + 2" | (b) "height (mkt_bal_r n l (AVL.insert x r)) = height l + 3" using MKT 2 by (atomize_elim, intro height_mkt_bal_r) simp_all then show ?thesis proof cases case a with 2 \\x < n\ show ?thesis by (auto simp del: mkt_bal_r.simps) next case b with True 1 MKT(4) \\x < n\ show ?thesis by (simp del: mkt_bal_r.simps) arith qed qed qed qed qed simp_all lemmas avl_insert = avl_insert_aux(1) subsubsection \Deletion maintains AVL balance\ lemma avl_delete_max: assumes "avl x" and "x \ ET" shows "avl (snd (delete_max x))" "height x = height(snd (delete_max x)) \ height x = height(snd (delete_max x)) + 1" using assms proof (induct x rule: delete_max_induct) case (MKT n l rn rl rr rh h) case 1 with MKT have "avl l" "avl (snd (delete_max (MKT rn rl rr rh)))" by auto with 1 MKT have "avl (mkt_bal_l n l (snd (delete_max (MKT rn rl rr rh))))" by (intro avl_mkt_bal_l) fastforce+ then show ?case by (auto simp: height_mkt_bal_l height_mkt_bal_l2 linorder_class.max.absorb1 linorder_class.max.absorb2 split:prod.split simp del:mkt_bal_l.simps) next case (MKT n l rn rl rr rh h) case 2 let ?r = "MKT rn rl rr rh" let ?r' = "snd (delete_max ?r)" from \avl x\ MKT 2 have "avl l" and "avl ?r" by simp_all then show ?case using MKT 2 height_mkt_bal_l[of l ?r' n] height_mkt_bal_l2[of l ?r' n] apply (auto split:prod.splits simp del:avl.simps mkt_bal_l.simps) by arith+ qed auto lemma avl_delete_root: assumes "avl t" and "t \ ET" shows "avl(delete_root t)" using assms proof (cases t rule:delete_root_cases) case (MKT_MKT n ln ll lr lh rn rl rr rh h) let ?l = "MKT ln ll lr lh" let ?r = "MKT rn rl rr rh" let ?l' = "snd (delete_max ?l)" from \avl t\ and MKT_MKT have "avl ?r" by simp from \avl t\ and MKT_MKT have "avl ?l" by simp then have "avl(?l')" "height ?l = height(?l') \ height ?l = height(?l') + 1" by (rule avl_delete_max,simp)+ with \avl t\ MKT_MKT have "height ?l' = height ?r \ height ?l' = height ?r + 1 \ height ?r = height ?l' + 1 \ height ?r = height ?l' + 2" by fastforce with \avl ?l'\ \avl ?r\ have "avl(mkt_bal_r (fst(delete_max ?l)) ?l' ?r)" by (rule avl_mkt_bal_r) with MKT_MKT show ?thesis by (auto split:prod.splits simp del:mkt_bal_r.simps) qed simp_all lemma height_delete_root: assumes "avl t" and "t \ ET" shows "height t = height(delete_root t) \ height t = height(delete_root t) + 1" using assms proof (cases t rule: delete_root_cases) case (MKT_MKT n ln ll lr lh rn rl rr rh h) let ?l = "MKT ln ll lr lh" let ?r = "MKT rn rl rr rh" let ?l' = "snd (delete_max ?l)" let ?t' = "mkt_bal_r (fst(delete_max ?l)) ?l' ?r" from \avl t\ and MKT_MKT have "avl ?r" by simp from \avl t\ and MKT_MKT have "avl ?l" by simp then have "avl(?l')" by (rule avl_delete_max,simp) have l'_height: "height ?l = height ?l' \ height ?l = height ?l' + 1" using \avl ?l\ by (intro avl_delete_max) auto have t_height: "height t = 1 + max (height ?l) (height ?r)" using \avl t\ MKT_MKT by simp have "height t = height ?t' \ height t = height ?t' + 1" using \avl t\ MKT_MKT proof(cases "height ?r = height ?l' + 2") case False show ?thesis using l'_height t_height False by (subst height_mkt_bal_r2[OF \avl ?l'\ \avl ?r\ False])+ arith next case True show ?thesis proof(cases rule: disjE[OF height_mkt_bal_r[OF True \avl ?l'\ \avl ?r\, of "fst (delete_max ?l)"]]) case 1 then show ?thesis using l'_height t_height True by arith next case 2 then show ?thesis using l'_height t_height True by arith qed qed thus ?thesis using MKT_MKT by (auto split:prod.splits simp del:mkt_bal_r.simps) qed simp_all text\Deletion maintains the AVL property:\ theorem avl_delete_aux: assumes "avl t" shows "avl(delete x t)" and "height t = (height (delete x t)) \ height t = height (delete x t) + 1" using assms proof (induct t) case (MKT n l r h) case 1 with MKT show ?case proof(cases "x = n") case True with MKT 1 show ?thesis by (auto simp:avl_delete_root) next case False with MKT 1 show ?thesis proof(cases "xx\n\ show ?thesis by (auto simp add:avl_mkt_bal_l simp del:mkt_bal_l.simps) qed qed case 2 with MKT show ?case proof(cases "x = n") case True with 1 have "height (MKT n l r h) = height(delete_root (MKT n l r h)) \ height (MKT n l r h) = height(delete_root (MKT n l r h)) + 1" by (subst height_delete_root,simp_all) with True show ?thesis by simp next case False with MKT 1 show ?thesis proof(cases "xx < n\ show ?thesis by auto next case True then consider (a) "height (mkt_bal_r n (delete x l) r) = height (delete x l) + 2" | (b) "height (mkt_bal_r n (delete x l) r) = height (delete x l) + 3" using MKT 2 by (atomize_elim, intro height_mkt_bal_r) auto then show ?thesis proof cases case a with \x < n\ MKT 2 show ?thesis by auto next case b with \x < n\ MKT 2 show ?thesis by auto qed qed next case False show ?thesis proof(cases "height l = height (delete x r) + 2") case False with MKT 1 \\x < n\ \x \ n\ show ?thesis by auto next case True then consider (a) "height (mkt_bal_l n l (delete x r)) = height (delete x r) + 2" | (b) "height (mkt_bal_l n l (delete x r)) = height (delete x r) + 3" using MKT 2 by (atomize_elim, intro height_mkt_bal_l) auto then show ?thesis proof cases case a with \\x < n\ \x \ n\ MKT 2 show ?thesis by auto next case b with \\x < n\ \x \ n\ MKT 2 show ?thesis by auto qed qed qed qed qed simp_all lemmas avl_delete = avl_delete_aux(1) subsubsection \Correctness of insertion\ lemma set_of_mkt_bal_l: "\ avl l; avl r \ \ set_of (mkt_bal_l n l r) = Set.insert n (set_of l \ set_of r)" by (auto simp: mkt_def split:tree.splits) lemma set_of_mkt_bal_r: "\ avl l; avl r \ \ set_of (mkt_bal_r n l r) = Set.insert n (set_of l \ set_of r)" by (auto simp: mkt_def split:tree.splits) text\Correctness of @{const insert}:\ theorem set_of_insert: "avl t \ set_of(insert x t) = Set.insert x (set_of t)" by (induct t) (auto simp: avl_insert set_of_mkt_bal_l set_of_mkt_bal_r simp del:mkt_bal_l.simps mkt_bal_r.simps) subsubsection \Correctness of deletion\ fun rightmost_item :: "'a tree \ 'a" where "rightmost_item (MKT n l ET h) = n" | "rightmost_item (MKT n l r h) = rightmost_item r" lemma avl_dist: "\ avl(MKT n l r h); is_ord(MKT n l r h); x \ set_of l \ \ x \ set_of r" by fastforce lemma avl_dist2: "\ avl(MKT n l r h); is_ord(MKT n l r h); x \ set_of l \ x \ set_of r \ \ x \ n" by auto lemma ritem_in_rset: "r \ ET \ rightmost_item r \ set_of r" by(induct r rule:rightmost_item.induct) auto lemma ritem_greatest_in_rset: "\ r \ ET; is_ord r \ \ \x. x \ set_of r \ x \ rightmost_item r \ x < rightmost_item r" proof(induct r rule:rightmost_item.induct) case (2 n l rn rl rr rh h) show ?case (is "\x. ?P x") proof fix x from 2 have "is_ord (MKT rn rl rr rh)" by auto moreover from 2 have "n < rightmost_item (MKT rn rl rr rh)" by (metis is_ord.simps(2) ritem_in_rset tree.simps(2)) moreover from 2 have "x \ set_of l \ x < rightmost_item (MKT rn rl rr rh)" by (metis calculation(2) is_ord.simps(2) xt1(10)) ultimately show "?P x" using 2 by simp qed qed auto lemma ritem_not_in_ltree: "\ avl(MKT n l r h); is_ord(MKT n l r h); r \ ET \ \ rightmost_item r \ set_of l" by (metis avl_dist ritem_in_rset) lemma set_of_delete_max: "\ avl t; is_ord t; t\ET \ \ set_of (snd(delete_max t)) = (set_of t) - {rightmost_item t}" proof (induct t rule: delete_max_induct) case (MKT n l rn rl rr rh h) let ?r = "MKT rn rl rr rh" from MKT have "avl l" and "avl ?r" by simp_all let ?t' = "mkt_bal_l n l (snd (delete_max ?r))" from MKT have "avl (snd(delete_max ?r))" by (auto simp add: avl_delete_max) with MKT ritem_not_in_ltree[of n l ?r h] have "set_of ?t' = (set_of l) \ (set_of ?r) - {rightmost_item ?r} \ {n}" by (auto simp add:set_of_mkt_bal_l simp del: mkt_bal_l.simps) moreover have "n \ {rightmost_item ?r}" by (metis MKT(2) MKT(3) avl_dist2 ritem_in_rset singletonE tree.simps(3)) ultimately show ?case by (auto simp add:insert_Diff_if split:prod.splits simp del: mkt_bal_l.simps) qed auto lemma fst_delete_max_eq_ritem: "t\ET \ fst(delete_max t) = rightmost_item t" by (induct t rule:rightmost_item.induct) (auto split:prod.splits) lemma set_of_delete_root: assumes "t = MKT n l r h" and "avl t" and "is_ord t" shows "set_of (delete_root t) = (set_of t) - {n}" using assms proof(cases t rule:delete_root_cases) case(MKT_MKT n ln ll lr lh rn rl rr rh h) let ?t' = "mkt_bal_r (fst (delete_max l)) (snd (delete_max l)) r" from assms MKT_MKT have "avl l" and "avl r" and "is_ord l" and "l\ET" by auto moreover from MKT_MKT assms have "avl (snd(delete_max l))" by (auto simp add: avl_delete_max) ultimately have "set_of ?t' = (set_of l) \ (set_of r)" by (fastforce simp add: Set.insert_Diff ritem_in_rset fst_delete_max_eq_ritem set_of_delete_max set_of_mkt_bal_r simp del: mkt_bal_r.simps) moreover from MKT_MKT assms(1) have "set_of (delete_root t) = set_of ?t'" by (simp split:prod.split del:mkt_bal_r.simps) moreover from MKT_MKT assms have "(set_of t) - {n} = set_of l \ set_of r" by (metis Diff_insert_absorb UnE avl_dist2 tree.set(2) tree.inject) ultimately show ?thesis using MKT_MKT assms(1) by (simp del: delete_root.simps) qed auto text\Correctness of @{const delete}:\ theorem set_of_delete: "\ avl t; is_ord t \ \ set_of (delete x t) = (set_of t) - {x}" proof (induct t) case (MKT n l r h) then show ?case proof(cases "x = n") case True with MKT set_of_delete_root[of "MKT n l r h"] show ?thesis by simp next case False with MKT show ?thesis proof(cases "xx\n\ show ?thesis by (force simp: avl_delete set_of_mkt_bal_l[of l "(delete x r)" n] simp del:mkt_bal_l.simps) qed qed qed simp subsubsection \Correctness of lookup\ theorem is_in_correct: "is_ord t \ is_in k t = (k : set_of t)" by (induct t) auto subsubsection \Insertion maintains order\ lemma is_ord_mkt_bal_l: "is_ord(MKT n l r h) \ is_ord (mkt_bal_l n l r)" by (cases l) (auto simp: mkt_def split:tree.splits intro: order_less_trans) lemma is_ord_mkt_bal_r: "is_ord(MKT n l r h) \ is_ord (mkt_bal_r n l r)" by (cases r) (auto simp: mkt_def split:tree.splits intro: order_less_trans) text\If the order is linear, @{const insert} maintains the order:\ theorem is_ord_insert: "\ avl t; is_ord t \ \ is_ord(insert (x::'a::linorder) t)" by (induct t) (simp_all add:is_ord_mkt_bal_l is_ord_mkt_bal_r avl_insert set_of_insert linorder_not_less order_neq_le_trans del:mkt_bal_l.simps mkt_bal_r.simps) subsubsection \Deletion maintains order\ lemma is_ord_delete_max: "\ avl t; is_ord t; t\ET \ \ is_ord(snd(delete_max t))" proof(induct t rule:delete_max_induct) case(MKT n l rn rl rr rh h) let ?r = "MKT rn rl rr rh" let ?r' = "snd(delete_max ?r)" from MKT have "\h. is_ord(MKT n l ?r' h)" by (auto simp: set_of_delete_max) moreover from MKT have "avl(?r')" by (auto simp: avl_delete_max) moreover note MKT is_ord_mkt_bal_l[of n l ?r'] ultimately show ?case by (auto split:prod.splits simp del:is_ord.simps mkt_bal_l.simps) qed auto lemma is_ord_delete_root: assumes "avl t" and "is_ord t" and "t \ ET" shows "is_ord (delete_root t)" using assms proof(cases t rule:delete_root_cases) case(MKT_MKT n ln ll lr lh rn rl rr rh h) let ?l = "MKT ln ll lr lh" let ?r = "MKT rn rl rr rh" let ?l' = "snd (delete_max ?l)" let ?n' = "fst (delete_max ?l)" from assms MKT_MKT have "\h. is_ord(MKT ?n' ?l' ?r h)" proof - from assms MKT_MKT have "is_ord ?l'" by (auto simp add: is_ord_delete_max) moreover from assms MKT_MKT have "is_ord ?r" by auto moreover from assms MKT_MKT have "\x. x \ set_of ?r \ ?n' < x" by (metis fst_delete_max_eq_ritem is_ord.simps(2) order_less_trans ritem_in_rset tree.simps(3)) moreover from assms MKT_MKT ritem_greatest_in_rset have "\x. x \ set_of ?l' \ x < ?n'" by (metis Diff_iff avl.simps(2) fst_delete_max_eq_ritem is_ord.simps(2) set_of_delete_max singleton_iff tree.simps(3)) ultimately show ?thesis by auto qed moreover from assms MKT_MKT have "avl ?r" by simp moreover from assms MKT_MKT have "avl ?l'" by (simp add: avl_delete_max) moreover note MKT_MKT is_ord_mkt_bal_r[of ?n' ?l' ?r] ultimately show ?thesis by (auto simp del:mkt_bal_r.simps is_ord.simps split:prod.splits) qed simp_all text\If the order is linear, @{const delete} maintains the order:\ theorem is_ord_delete: "\ avl t; is_ord t \ \ is_ord (delete x t)" proof (induct t) case (MKT n l r h) then show ?case proof(cases "x = n") case True with MKT is_ord_delete_root[of "MKT n l r h"] show ?thesis by simp next case False with MKT show ?thesis proof(cases "xh. is_ord (MKT n (delete x l) r h)" by (auto simp:set_of_delete) with True MKT is_ord_mkt_bal_r[of n "(delete x l)" r] show ?thesis by (auto simp add: avl_delete) next case False with False MKT have "\h. is_ord (MKT n l (delete x r) h)" by (auto simp:set_of_delete) with False MKT is_ord_mkt_bal_l[of n l "(delete x r)"] \x\n\ show ?thesis by (simp add: avl_delete) qed qed qed simp end diff --git a/thys/Akra_Bazzi/Akra_Bazzi_Asymptotics.thy b/thys/Akra_Bazzi/Akra_Bazzi_Asymptotics.thy --- a/thys/Akra_Bazzi/Akra_Bazzi_Asymptotics.thy +++ b/thys/Akra_Bazzi/Akra_Bazzi_Asymptotics.thy @@ -1,375 +1,376 @@ (* File: Akra_Bazzi_Asymptotics.thy Author: Manuel Eberl Proofs for the four(ish) asymptotic inequalities required for proving the Akra Bazzi theorem with variation functions in the recursive calls. *) section \Asymptotic bounds\ theory Akra_Bazzi_Asymptotics imports Complex_Main Akra_Bazzi_Library "HOL-Library.Landau_Symbols" begin locale akra_bazzi_asymptotics_bep = fixes b e p hb :: real assumes bep: "b > 0" "b < 1" "e > 0" "hb > 0" begin context begin text \ Functions that are negligible w.r.t. @{term "ln (b*x) powr (e/2 + 1)"}. \ private abbreviation (input) negl :: "(real \ real) \ bool" where "negl f \ f \ o(\x. ln (b*x) powr (-(e/2 + 1)))" private lemma neglD: "negl f \ c > 0 \ eventually (\x. \f x\ \ c / ln (b*x) powr (e/2+1)) at_top" by (drule (1) landau_o.smallD, subst (asm) powr_minus) (simp add: field_simps) private lemma negl_mult: "negl f \ negl g \ negl (\x. f x * g x)" by (erule landau_o.small_1_mult, rule landau_o.small_imp_big, erule landau_o.small_trans) (insert bep, simp) private lemma ev4: assumes g: "negl g" shows "eventually (\x. ln (b*x) powr (-e/2) - ln x powr (-e/2) \ g x) at_top" proof (rule smallo_imp_le_real) define h1 where [abs_def]: "h1 x = (1 + ln b/ln x) powr (-e/2) - 1 + e/2 * (ln b/ln x)" for x define h2 where [abs_def]: "h2 x = ln x powr (- e / 2) * ((1 + ln b / ln x) powr (- e / 2) - 1)" for x from bep have "((\x. ln b / ln x) \ 0) at_top" by (simp add: tendsto_0_smallo_1) note one_plus_x_powr_Taylor2_bigo[OF this, of "-e/2"] also have "(\x. (1 + ln b / ln x) powr (- e / 2) - 1 - - e / 2 * (ln b / ln x)) = h1" by (simp add: h1_def) finally have "h1 \ o(\x. 1 / ln x)" by (rule landau_o.big_small_trans) (insert bep, simp add: power2_eq_square) with bep have "(\x. h1 x - e/2 * (ln b / ln x)) \ \(\x. 1 / ln x)" by simp also have "(\x. h1 x - e/2 * (ln b/ln x)) = (\x. (1 + ln b/ ln x) powr (-e/2) - 1)" by (rule ext) (simp add: h1_def) finally have "h2 \ \(\x. ln x powr (-e/2) * (1 / ln x))" unfolding h2_def by (intro landau_theta.mult) simp_all also have "(\x. ln x powr (-e/2) * (1 / ln x)) \ \(\x. ln x powr (-(e/2+1)))" by simp also from g bep have "(\x. ln x powr (-(e/2+1))) \ \(g)" by (simp add: smallomega_iff_smallo) finally have "g \ o(h2)" by (simp add: smallomega_iff_smallo) also have "eventually (\x. h2 x = ln (b*x) powr (-e/2) - ln x powr (-e/2)) at_top" using eventually_gt_at_top[of "1::real"] eventually_gt_at_top[of "1/b"] by eventually_elim (insert bep, simp add: field_simps powr_diff [symmetric] h2_def ln_mult [symmetric] powr_divide del: ln_mult) hence "h2 \ \(\x. ln (b*x) powr (-e/2) - ln x powr (-e/2))" by (rule bigthetaI_cong) finally show "g \ o(\x. ln (b * x) powr (- e / 2) - ln x powr (- e / 2))" . next show "eventually (\x. ln (b*x) powr (-e/2) - ln x powr (-e/2) \ 0) at_top" using eventually_gt_at_top[of "1/b"] eventually_gt_at_top[of "1::real"] by eventually_elim (insert bep, auto intro!: powr_mono2' simp: field_simps simp del: ln_mult) qed private lemma ev1: "negl (\x. (1 + c * inverse b * ln x powr (-(1+e))) powr p - 1)" proof- from bep have "((\x. c * inverse b * ln x powr (-(1+e))) \ 0) at_top" by (simp add: tendsto_0_smallo_1) have "(\x. (1 + c * inverse b * ln x powr (-(1+e))) powr p - 1) \ O(\x. c * inverse b * ln x powr - (1 + e))" using bep by (intro one_plus_x_powr_Taylor1_bigo) (simp add: tendsto_0_smallo_1) also from bep have "negl (\x. c * inverse b * ln x powr - (1 + e))" by simp finally show ?thesis . qed private lemma ev2_aux: defines "f \ \x. (1 + 1/ln (b*x) * ln (1 + hb / b * ln x powr (-1-e))) powr (-e/2)" obtains h where "eventually (\x. f x \ 1 + h x) at_top" "h \ o(\x. 1 / ln x)" proof (rule that[of "\x. f x - 1"]) define g where [abs_def]: "g x = 1/ln (b*x) * ln (1 + hb / b * ln x powr (-1-e))" for x have lim: "((\x. ln (1 + hb / b * ln x powr (- 1 - e))) \ 0) at_top" by (rule tendsto_eq_rhs[OF tendsto_ln[OF tendsto_add[OF tendsto_const, of _ 0]]]) (insert bep, simp_all add: tendsto_0_smallo_1) hence lim': "(g \ 0) at_top" unfolding g_def by (intro tendsto_mult_zero) (insert bep, simp add: tendsto_0_smallo_1) from one_plus_x_powr_Taylor2_bigo[OF this, of "-e/2"] have "(\x. (1 + g x) powr (-e/2) - 1 - - e/2 * g x) \ O(\x. (g x)\<^sup>2)" . also from lim' have "(\x. g x ^ 2) \ o(\x. g x * 1)" unfolding power2_eq_square by (intro landau_o.big_small_mult smalloI_tendsto) simp_all also have "o(\x. g x * 1) = o(g)" by simp also have "(\x. (1 + g x) powr (-e/2) - 1 - - e/2 * g x) = (\x. f x - 1 + e/2 * g x)" by (simp add: f_def g_def) finally have A: "(\x. f x - 1 + e / 2 * g x) \ O(g)" by (rule landau_o.small_imp_big) hence "(\x. f x - 1 + e/2 * g x - e/2 * g x) \ O(g)" by (rule sum_in_bigo) (insert bep, simp) also have "(\x. f x - 1 + e/2 * g x - e/2 * g x) = (\x. f x - 1)" by simp finally have "(\x. f x - 1) \ O(g)" . also from bep lim have "g \ o(\x. 1 / ln x)" unfolding g_def by (auto intro!: smallo_1_tendsto_0) finally show "(\x. f x - 1) \ o(\x. 1 / ln x)" . qed simp_all private lemma ev2: defines "f \ \x. ln (b * x + hb * x / ln x powr (1 + e)) powr (-e/2)" obtains h where "negl h" "eventually (\x. f x \ ln (b * x) powr (-e/2) + h x) at_top" "eventually (\x. \ln (b * x) powr (-e/2) + h x\ < 1) at_top" proof - define f' where "f' x = (1 + 1 / ln (b*x) * ln (1 + hb / b * ln x powr (-1-e))) powr (-e/2)" for x from ev2_aux obtain g where g: "eventually (\x. 1 + g x \ f' x) at_top" "g \ o(\x. 1 / ln x)" unfolding f'_def . define h where [abs_def]: "h x = ln (b*x) powr (-e/2) * g x" for x show ?thesis proof (rule that[of h]) from bep g show "negl h" unfolding h_def by (auto simp: powr_diff elim: landau_o.small_big_trans) next from g(2) have "g \ o(\x. 1)" by (rule landau_o.small_big_trans) simp with bep have "eventually (\x. \ln (b*x) powr (-e/2) * (1 + g x)\ < 1) at_top" by (intro smallo_imp_abs_less_real) simp_all thus "eventually (\x. \ln (b*x) powr (-e/2) + h x\ < 1) at_top" by (simp add: algebra_simps h_def) next from eventually_gt_at_top[of "1/b"] and g(1) show "eventually (\x. f x \ ln (b*x) powr (-e/2) + h x) at_top" proof eventually_elim case (elim x) from bep have "b * x + hb * x / ln x powr (1 + e) = b*x * (1 + hb / b * ln x powr (-1 - e))" by (simp add: field_simps powr_diff powr_add powr_minus) also from elim(1) bep have "ln \ = ln (b*x) * (1 + 1/ln (b*x) * ln (1 + hb / b * ln x powr (-1-e)))" by (subst ln_mult) (simp_all add: add_pos_nonneg field_simps) also from elim(1) bep have "\ powr (-e/2) = ln (b*x) powr (-e/2) * f' x" by (subst powr_mult) (simp_all add: field_simps f'_def) also from elim have "\ \ ln (b*x) powr (-e/2) * (1 + g x)" by (intro mult_left_mono) simp_all finally show "f x \ ln (b*x) powr (-e/2) + h x" by (simp add: f_def h_def algebra_simps) qed qed qed private lemma ev21: obtains g where "negl g" "eventually (\x. 1 + ln (b * x + hb * x / ln x powr (1 + e)) powr (-e/2) \ 1 + ln (b * x) powr (-e/2) + g x) at_top" "eventually (\x. 1 + ln (b * x) powr (-e/2) + g x > 0) at_top" proof- from ev2 guess g . note g = this from g(3) have "eventually (\x. 1 + ln (b * x) powr (-e/2) + g x > 0) at_top" by eventually_elim simp with g(1,2) show ?thesis by (intro that[of g]) simp_all qed private lemma ev22: obtains g where "negl g" "eventually (\x. 1 - ln (b * x + hb * x / ln x powr (1 + e)) powr (-e/2) \ 1 - ln (b * x) powr (-e/2) - g x) at_top" "eventually (\x. 1 - ln (b * x) powr (-e/2) - g x > 0) at_top" proof- from ev2 guess g . note g = this from g(2) have "eventually (\x. 1 - ln (b * x + hb * x / ln x powr (1 + e)) powr (-e/2) \ 1 - ln (b * x) powr (-e/2) - g x) at_top" by eventually_elim simp moreover from g(3) have "eventually (\x. 1 - ln (b * x) powr (-e/2) - g x > 0) at_top" by eventually_elim simp ultimately show ?thesis using g(1) by (intro that[of g]) simp_all qed lemma asymptotics1: shows "eventually (\x. (1 + c * inverse b * ln x powr -(1+e)) powr p * (1 + ln (b * x + hb * x / ln x powr (1 + e)) powr (- e / 2)) \ 1 + (ln x powr (-e/2))) at_top" proof- let ?f = "\x. (1 + c * inverse b * ln x powr -(1+e)) powr p" let ?g = "\x. 1 + ln (b * x + hb * x / ln x powr (1 + e)) powr (- e / 2)" define f where [abs_def]: "f x = 1 - ?f x" for x from ev1[of c] have "negl f" unfolding f_def by (subst landau_o.small.uminus_in_iff [symmetric]) simp from landau_o.smallD[OF this zero_less_one] have f: "eventually (\x. f x \ ln (b*x) powr -(e/2+1)) at_top" by eventually_elim (simp add: f_def) from ev21 guess g . note g = this define h where [abs_def]: "h x = -g x + f x + f x * ln (b*x) powr (-e/2) + f x * g x" for x have A: "eventually (\x. ?f x * ?g x \ 1 + ln (b*x) powr (-e/2) - h x) at_top" using g(2,3) f proof eventually_elim case (elim x) let ?t = "ln (b*x) powr (-e/2)" have "1 + ?t - h x = (1 - f x) * (1 + ln (b*x) powr (-e/2) + g x)" by (simp add: algebra_simps h_def) also from elim have "?f x * ?g x \ (1 - f x) * (1 + ln (b*x) powr (-e/2) + g x)" by (intro mult_mono[OF _ elim(1)]) (simp_all add: algebra_simps f_def) finally show "?f x * ?g x \ 1 + ln (b*x) powr (-e/2) - h x" . qed from bep \negl f\ g(1) have "negl h" unfolding h_def by (fastforce intro!: sum_in_smallo landau_o.small.mult simp: powr_diff intro: landau_o.small_trans)+ from ev4[OF this] A show ?thesis by eventually_elim simp qed lemma asymptotics2: shows "eventually (\x. (1 + c * inverse b * ln x powr -(1+e)) powr p * (1 - ln (b * x + hb * x / ln x powr (1 + e)) powr (- e / 2)) \ 1 - (ln x powr (-e/2))) at_top" proof- let ?f = "\x. (1 + c * inverse b * ln x powr -(1+e)) powr p" let ?g = "\x. 1 - ln (b * x + hb * x / ln x powr (1 + e)) powr (- e / 2)" define f where [abs_def]: "f x = 1 - ?f x" for x from ev1[of c] have "negl f" unfolding f_def by (subst landau_o.small.uminus_in_iff [symmetric]) simp from landau_o.smallD[OF this zero_less_one] have f: "eventually (\x. f x \ ln (b*x) powr -(e/2+1)) at_top" by eventually_elim (simp add: f_def) from ev22 guess g . note g = this define h where [abs_def]: "h x = -g x - f x + f x * ln (b*x) powr (-e/2) + f x * g x" for x have "((\x. ln (b * x + hb * x / ln x powr (1 + e)) powr - (e / 2)) \ 0) at_top" apply (insert bep, intro tendsto_neg_powr, simp) apply (rule filterlim_compose[OF ln_at_top]) apply (rule filterlim_at_top_smallomega_1, simp) using eventually_gt_at_top[of "max 1 (1/b)"] apply (auto elim!: eventually_mono intro!: add_pos_nonneg simp: field_simps) + apply (smt (z3) divide_nonneg_nonneg mult_neg_pos mult_nonneg_nonneg powr_non_neg) done hence ev_g: "eventually (\x. \1 - ?g x\ < 1) at_top" by (intro smallo_imp_abs_less_real smalloI_tendsto) simp_all have A: "eventually (\x. ?f x * ?g x \ 1 - ln (b*x) powr (-e/2) + h x) at_top" using g(2,3) ev_g f proof eventually_elim case (elim x) let ?t = "ln (b*x) powr (-e/2)" from elim have "?f x * ?g x \ (1 - f x) * (1 - ln (b*x) powr (-e/2) - g x)" by (intro mult_mono) (simp_all add: f_def) also have "... = 1 - ?t + h x" by (simp add: algebra_simps h_def) finally show "?f x * ?g x \ 1 - ln (b*x) powr (-e/2) + h x" . qed from bep \negl f\ g(1) have "negl h" unfolding h_def by (fastforce intro!: sum_in_smallo landau_o.small.mult simp: powr_diff intro: landau_o.small_trans)+ from ev4[OF this] A show ?thesis by eventually_elim simp qed lemma asymptotics3: "eventually (\x. (1 + (ln x powr (-e/2))) / 2 \ 1) at_top" (is "eventually (\x. ?f x \ 1) _") proof (rule eventually_mp[OF always_eventually], clarify) from bep have "(?f \ 1/2) at_top" by (force intro: tendsto_eq_intros tendsto_neg_powr ln_at_top) hence "\e. e>0 \ eventually (\x. \?f x - 0.5\ < e) at_top" by (subst (asm) tendsto_iff) (simp add: dist_real_def) from this[of "0.5"] show "eventually (\x. \?f x - 0.5\ < 0.5) at_top" by simp fix x assume "\?f x - 0.5\ < 0.5" thus "?f x \ 1" by simp qed lemma asymptotics4: "eventually (\x. (1 - (ln x powr (-e/2))) * 2 \ 1) at_top" (is "eventually (\x. ?f x \ 1) _") proof (rule eventually_mp[OF always_eventually], clarify) from bep have "(?f \ 2) at_top" by (force intro: tendsto_eq_intros tendsto_neg_powr ln_at_top) hence "\e. e>0 \ eventually (\x. \?f x - 2\ < e) at_top" by (subst (asm) tendsto_iff) (simp add: dist_real_def) from this[of 1] show "eventually (\x. \?f x - 2\ < 1) at_top" by simp fix x assume "\?f x - 2\ < 1" thus "?f x \ 1" by simp qed lemma asymptotics5: "eventually (\x. ln (b*x - hb*x*ln x powr -(1+e)) powr (-e/2) < 1) at_top" proof- from bep have "((\x. b - hb * ln x powr -(1+e)) \ b - 0) at_top" by (intro tendsto_intros tendsto_mult_right_zero tendsto_neg_powr ln_at_top) simp_all hence "LIM x at_top. (b - hb * ln x powr -(1+e)) * x :> at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_ident], insert bep) simp_all also have "(\x. (b - hb * ln x powr -(1+e)) * x) = (\x. b*x - hb*x*ln x powr -(1+e))" by (intro ext) (simp add: algebra_simps) finally have "filterlim ... at_top at_top" . with bep have "((\x. ln (b*x - hb*x*ln x powr -(1+e)) powr -(e/2)) \ 0) at_top" by (intro tendsto_neg_powr filterlim_compose[OF ln_at_top]) simp_all hence "eventually (\x. \ln (b*x - hb*x*ln x powr -(1+e)) powr (-e/2)\ < 1) at_top" by (subst (asm) tendsto_iff) (simp add: dist_real_def) thus ?thesis by simp qed lemma asymptotics6: "eventually (\x. hb / ln x powr (1 + e) < b/2) at_top" and asymptotics7: "eventually (\x. hb / ln x powr (1 + e) < (1 - b) / 2) at_top" and asymptotics8: "eventually (\x. x*(1 - b - hb / ln x powr (1 + e)) > 1) at_top" proof- from bep have A: "(\x. hb / ln x powr (1 + e)) \ o(\_. 1)" by simp from bep have B: "b/3 > 0" and C: "(1 - b)/3 > 0" by simp_all from landau_o.smallD[OF A B] show "eventually (\x. hb / ln x powr (1+e) < b/2) at_top" by eventually_elim (insert bep, simp) from landau_o.smallD[OF A C] show "eventually (\x. hb / ln x powr (1 + e) < (1 - b)/2) at_top" by eventually_elim (insert bep, simp) from bep have "(\x. hb / ln x powr (1 + e)) \ o(\_. 1)" "(1 - b) / 2 > 0" by simp_all from landau_o.smallD[OF this] eventually_gt_at_top[of "1::real"] have A: "eventually (\x. 1 - b - hb / ln x powr (1 + e) > 0) at_top" by eventually_elim (insert bep, simp add: field_simps) from bep have "(\x. x * (1 - b - hb / ln x powr (1+e))) \ \(\_. 1)" "(0::real) < 2" by simp_all from landau_omega.smallD[OF this] A eventually_gt_at_top[of "0::real"] show "eventually (\x. x*(1 - b - hb / ln x powr (1 + e)) > 1) at_top" by eventually_elim (simp_all add: abs_mult) qed end end definition "akra_bazzi_asymptotic1 b hb e p x \ (1 - hb * inverse b * ln x powr -(1+e)) powr p * (1 + ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2)) \ 1 + (ln x powr (-e/2) :: real)" definition "akra_bazzi_asymptotic1' b hb e p x \ (1 + hb * inverse b * ln x powr -(1+e)) powr p * (1 + ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2)) \ 1 + (ln x powr (-e/2) :: real)" definition "akra_bazzi_asymptotic2 b hb e p x \ (1 + hb * inverse b * ln x powr -(1+e)) powr p * (1 - ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2)) \ 1 - ln x powr (-e/2 :: real)" definition "akra_bazzi_asymptotic2' b hb e p x \ (1 - hb * inverse b * ln x powr -(1+e)) powr p * (1 - ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2)) \ 1 - ln x powr (-e/2 :: real)" definition "akra_bazzi_asymptotic3 e x \ (1 + (ln x powr (-e/2))) / 2 \ (1::real)" definition "akra_bazzi_asymptotic4 e x \ (1 - (ln x powr (-e/2))) * 2 \ (1::real)" definition "akra_bazzi_asymptotic5 b hb e x \ ln (b*x - hb*x*ln x powr -(1+e)) powr (-e/2::real) < 1" definition "akra_bazzi_asymptotic6 b hb e x \ hb / ln x powr (1 + e :: real) < b/2" definition "akra_bazzi_asymptotic7 b hb e x \ hb / ln x powr (1 + e :: real) < (1 - b) / 2" definition "akra_bazzi_asymptotic8 b hb e x \ x*(1 - b - hb / ln x powr (1 + e :: real)) > 1" definition "akra_bazzi_asymptotics b hb e p x \ akra_bazzi_asymptotic1 b hb e p x \ akra_bazzi_asymptotic1' b hb e p x \ akra_bazzi_asymptotic2 b hb e p x \ akra_bazzi_asymptotic2' b hb e p x \ akra_bazzi_asymptotic3 e x \ akra_bazzi_asymptotic4 e x \ akra_bazzi_asymptotic5 b hb e x \ akra_bazzi_asymptotic6 b hb e x \ akra_bazzi_asymptotic7 b hb e x \ akra_bazzi_asymptotic8 b hb e x" lemmas akra_bazzi_asymptotic_defs = akra_bazzi_asymptotic1_def akra_bazzi_asymptotic1'_def akra_bazzi_asymptotic2_def akra_bazzi_asymptotic2'_def akra_bazzi_asymptotic3_def akra_bazzi_asymptotic4_def akra_bazzi_asymptotic5_def akra_bazzi_asymptotic6_def akra_bazzi_asymptotic7_def akra_bazzi_asymptotic8_def akra_bazzi_asymptotics_def lemma akra_bazzi_asymptotics: assumes "\b. b \ set bs \ b \ {0<..<1}" assumes "hb > 0" "e > 0" shows "eventually (\x. \b\set bs. akra_bazzi_asymptotics b hb e p x) at_top" proof (intro eventually_ball_finite ballI) fix b assume "b \ set bs" with assms interpret akra_bazzi_asymptotics_bep b e p hb by unfold_locales auto show "eventually (\x. akra_bazzi_asymptotics b hb e p x) at_top" unfolding akra_bazzi_asymptotic_defs using asymptotics1[of "-c" for c] asymptotics2[of "-c" for c] by (intro eventually_conj asymptotics1 asymptotics2 asymptotics3 asymptotics4 asymptotics5 asymptotics6 asymptotics7 asymptotics8) simp_all qed simp end diff --git a/thys/BTree/BTree_ImpSet.thy b/thys/BTree/BTree_ImpSet.thy --- a/thys/BTree/BTree_ImpSet.thy +++ b/thys/BTree/BTree_ImpSet.thy @@ -1,1697 +1,1697 @@ theory BTree_ImpSet imports BTree_Imp BTree_Set begin section "Imperative Set operations" subsection "Auxiliary operations" definition "split_relation xs \ \(as,bs) i. i \ length xs \ as = take i xs \ bs = drop i xs" lemma split_relation_alt: "split_relation as (ls,rs) i = (as = ls@rs \ i = length ls)" by (auto simp add: split_relation_def) lemma split_relation_length: "split_relation xs (ls,rs) (length xs) = (ls = xs \ rs = [])" by (simp add: split_relation_def) (* auxiliary lemmas on assns *) (* simp? not sure if it always makes things more easy *) lemma list_assn_prod_map: "list_assn (A \\<^sub>a B) xs ys = list_assn B (map snd xs) (map snd ys) * list_assn A (map fst xs) (map fst ys)" apply(induct "(A \\<^sub>a B)" xs ys rule: list_assn.induct) apply(auto simp add: ab_semigroup_mult_class.mult.left_commute ent_star_mono star_aci(2) star_assoc) done (* concrete *) lemma id_assn_list: "h \ list_assn id_assn (xs::'a list) ys \ xs = ys" apply(induction "id_assn::('a \ 'a \ assn)" xs ys rule: list_assn.induct) apply(auto simp add: less_Suc_eq_0_disj pure_def) done lemma snd_map_help: "x \ length tsi \ (\j snd (tsi!x) = ((map snd tsi)!x)" by auto lemma split_ismeq: "((a::nat) \ b \ X) = ((a < b \ X) \ (a = b \ X))" by auto lemma split_relation_map: "split_relation as (ls,rs) i \ split_relation (map f as) (map f ls, map f rs) i" apply(induction as arbitrary: ls rs i) apply(auto simp add: split_relation_def take_map drop_Cons') apply(metis list.simps(9) take_map) done lemma split_relation_access: "\split_relation as (ls,rs) i; rs = r#rrs\ \ as!i = r" by (simp add: split_relation_alt) lemma index_to_elem_all: "(\jx \ set xs. P x)" by (simp add: all_set_conv_nth) lemma index_to_elem: "n < length xs \ (\jx \ set (take n xs). P x)" by (simp add: all_set_conv_nth) (* ----------------- *) definition split_half :: "('a::heap \ 'b::{heap}) pfarray \ nat Heap" where "split_half a \ do { l \ pfa_length a; return (l div 2) }" lemma split_half_rule[sep_heap_rules]: "< is_pfa c tsi a * list_assn R ts tsi> split_half a <\i. is_pfa c tsi a * list_assn R ts tsi * \(i = length ts div 2 \ split_relation ts (BTree_Set.split_half ts) i)>" unfolding split_half_def split_relation_def apply(rule hoare_triple_preI) apply(sep_auto dest!: list_assn_len mod_starD) done subsection "The imperative split locale" text "This locale extends the abstract split locale, assuming that we are provided with an imperative program that refines the abstract split function." locale imp_split = abs_split: BTree_Set.split split for split:: "('a btree \ 'a::{heap,default,linorder}) list \ 'a \ ('a btree \ 'a) list \ ('a btree \ 'a) list" + fixes imp_split:: "('a btnode ref option \ 'a::{heap,default,linorder}) pfarray \ 'a \ nat Heap" assumes imp_split_rule [sep_heap_rules]:"sorted_less (separators ts) \ imp_split (a,n) p <\i. is_pfa c tsi (a,n) * blist_assn k ts tsi * \(split_relation ts (split ts p) i)>\<^sub>t" begin subsection "Membership" partial_function (heap) isin :: "'a btnode ref option \ 'a \ bool Heap" where "isin p x = (case p of None \ return False | (Some a) \ do { node \ !a; i \ imp_split (kvs node) x; tsl \ pfa_length (kvs node); if i < tsl then do { s \ pfa_get (kvs node) i; let (sub,sep) = s in if x = sep then return True else isin sub x } else isin (last node) x } )" subsection "Insertion" datatype 'b btupi = T\<^sub>i "'b btnode ref option" | Up\<^sub>i "'b btnode ref option" "'b" "'b btnode ref option" fun btupi_assn where "btupi_assn k (abs_split.T\<^sub>i l) (T\<^sub>i li) = btree_assn k l li" | "btupi_assn k (abs_split.Up\<^sub>i l a r) (Up\<^sub>i li ai ri) = btree_assn k l li * id_assn a ai * btree_assn k r ri" | "btupi_assn _ _ _ = false" definition node\<^sub>i :: "nat \ ('a btnode ref option \ 'a) pfarray \ 'a btnode ref option \ 'a btupi Heap" where "node\<^sub>i k a ti \ do { n \ pfa_length a; if n \ 2*k then do { a' \ pfa_shrink_cap (2*k) a; l \ ref (Btnode a' ti); return (T\<^sub>i (Some l)) } else do { b \ (pfa_empty (2*k) :: ('a btnode ref option \ 'a) pfarray Heap); i \ split_half a; m \ pfa_get a i; b' \ pfa_drop a (i+1) b; a' \ pfa_shrink i a; a'' \ pfa_shrink_cap (2*k) a'; let (sub,sep) = m in do { l \ ref (Btnode a'' sub); r \ ref (Btnode b' ti); return (Up\<^sub>i (Some l) sep (Some r)) } } }" partial_function (heap) ins :: "nat \ 'a \ 'a btnode ref option \ 'a btupi Heap" where "ins k x apo = (case apo of None \ return (Up\<^sub>i None x None) | (Some ap) \ do { a \ !ap; i \ imp_split (kvs a) x; tsl \ pfa_length (kvs a); if i < tsl then do { s \ pfa_get (kvs a) i; let (sub,sep) = s in if sep = x then return (T\<^sub>i apo) else do { r \ ins k x sub; case r of (T\<^sub>i lp) \ do { pfa_set (kvs a) i (lp,sep); return (T\<^sub>i apo) } | (Up\<^sub>i lp x' rp) \ do { pfa_set (kvs a) i (rp,sep); if tsl < 2*k then do { kvs' \ pfa_insert (kvs a) i (lp,x'); ap := (Btnode kvs' (last a)); return (T\<^sub>i apo) } else do { kvs' \ pfa_insert_grow (kvs a) i (lp,x'); node\<^sub>i k kvs' (last a) } } } } else do { r \ ins k x (last a); case r of (T\<^sub>i lp) \ do { ap := (Btnode (kvs a) lp); return (T\<^sub>i apo) } | (Up\<^sub>i lp x' rp) \ if tsl < 2*k then do { kvs' \ pfa_append (kvs a) (lp,x'); ap := (Btnode kvs' rp); return (T\<^sub>i apo) } else do { kvs' \ pfa_append_grow' (kvs a) (lp,x'); node\<^sub>i k kvs' rp } } } )" (*fun tree\<^sub>i::"'a up\<^sub>i \ 'a btree" where "tree\<^sub>i (T\<^sub>i sub) = sub" | "tree\<^sub>i (Up\<^sub>i l a r) = (Node [(l,a)] r)" fun insert::"nat \ 'a \ 'a btree \ 'a btree" where "insert k x t = tree\<^sub>i (ins k x t)" *) definition insert :: "nat \ ('a::{heap,default,linorder}) \ 'a btnode ref option \ 'a btnode ref option Heap" where "insert \ \k x ti. do { ti' \ ins k x ti; case ti' of T\<^sub>i sub \ return sub | Up\<^sub>i l a r \ do { kvs \ pfa_init (2*k) (l,a) 1; t' \ ref (Btnode kvs r); return (Some t') } }" subsection "Deletion" (* rebalance middle tree gets a list of trees, an index pointing to the position of sub/sep and a last tree *) definition rebalance_middle_tree:: "nat \ (('a::{default,heap,linorder}) btnode ref option \ 'a) pfarray \ nat \ 'a btnode ref option \ 'a btnode Heap" where "rebalance_middle_tree \ \ k tsi i r_ti. ( case r_ti of None \ do { (r_sub,sep) \ pfa_get tsi i; case r_sub of None \ return (Btnode tsi r_ti) } | Some p_t \ do { (r_sub,sep) \ pfa_get tsi i; case r_sub of (Some p_sub) \ do { ti \ !p_t; sub \ !p_sub; l_sub \ pfa_length (kvs sub); l_tts \ pfa_length (kvs ti); if l_sub \ k \ l_tts \ k then do { return (Btnode tsi r_ti) } else do { l_tsi \ pfa_length tsi; if i+1 = l_tsi then do { mts' \ pfa_append_extend_grow (kvs sub) (last sub,sep) (kvs ti); res_node\<^sub>i \ node\<^sub>i k mts' (last ti); case res_node\<^sub>i of T\<^sub>i u \ do { tsi' \ pfa_shrink i tsi; return (Btnode tsi' u) } | Up\<^sub>i l a r \ do { tsi' \ pfa_set tsi i (l,a); return (Btnode tsi' r) } } else do { (r_rsub,rsep) \ pfa_get tsi (i+1); case r_rsub of Some p_rsub \ do { rsub \ !p_rsub; mts' \ pfa_append_extend_grow (kvs sub) (last sub,sep) (kvs rsub); res_node\<^sub>i \ node\<^sub>i k mts' (last rsub); case res_node\<^sub>i of T\<^sub>i u \ do { tsi' \ pfa_set tsi i (u,rsep); tsi'' \ pfa_delete tsi' (i+1); return (Btnode tsi'' r_ti) } | Up\<^sub>i l a r \ do { tsi' \ pfa_set tsi i (l,a); tsi'' \ pfa_set tsi' (i+1) (r,rsep); return (Btnode tsi'' r_ti) } } } } } }) " definition rebalance_last_tree:: "nat \ (('a::{default,heap,linorder}) btnode ref option \ 'a) pfarray \ 'a btnode ref option \ 'a btnode Heap" where "rebalance_last_tree \ \k tsi ti. do { l_tsi \ pfa_length tsi; rebalance_middle_tree k tsi (l_tsi-1) ti }" subsection "Refinement of the abstract B-tree operations" definition empty ::"('a::{default,heap,linorder}) btnode ref option Heap" where "empty = return None" lemma P_imp_Q_implies_P: "P \ (Q \ P)" by simp lemma "sorted_less (inorder t) \ isin ti x <\r. btree_assn k t ti * \(abs_split.isin t x = r)>\<^sub>t" proof(induction t x arbitrary: ti rule: abs_split.isin.induct) case (1 x) then show ?case apply(subst isin.simps) apply (cases ti) apply (auto simp add: return_cons_rule) done next case (2 ts t x) then obtain ls rs where list_split[simp]: "split ts x = (ls,rs)" by (cases "split ts x") then show ?case proof (cases rs) (* NOTE: induction condition trivial here *) case [simp]: Nil show ?thesis apply(subst isin.simps) apply(sep_auto) using "2.prems" sorted_inorder_separators apply blast apply(auto simp add: split_relation_def dest!: sym[of "[]"] mod_starD list_assn_len)[] apply(rule hoare_triple_preI) apply(auto simp add: split_relation_def dest!: sym[of "[]"] mod_starD list_assn_len)[] using 2(3) apply(sep_auto heap: "2.IH"(1)[of ls "[]"] simp add: sorted_wrt_append) done next case [simp]: (Cons h rrs) obtain sub sep where h_split[simp]: "h = (sub,sep)" by (cases h) show ?thesis proof (cases "sep = x") (* NOTE: no induction required here, only vacuous counter cases generated *) case [simp]: True then show ?thesis apply(simp split: list.splits prod.splits) apply(subst isin.simps) using "2.prems" sorted_inorder_separators apply(sep_auto) apply(rule hoare_triple_preI) apply(auto simp add: split_relation_alt list_assn_append_Cons_left dest!: mod_starD list_assn_len)[] apply(rule hoare_triple_preI) apply(auto simp add: split_relation_def dest!: sym[of "[]"] mod_starD list_assn_len)[] done next case [simp]: False show ?thesis apply(simp split: list.splits prod.splits) apply safe using False apply simp apply(subst isin.simps) using "2.prems" sorted_inorder_separators apply(sep_auto) (*eliminate vacuous case*) apply(auto simp add: split_relation_alt list_assn_append_Cons_left dest!: mod_starD list_assn_len)[] (* simplify towards induction step *) apply(auto simp add: split_relation_alt list_assn_append_Cons_left dest!: mod_starD list_assn_len)[] (* NOTE show that z = (suba, sepa) *) apply(rule norm_pre_ex_rule)+ apply(rule hoare_triple_preI) subgoal for p tsi n ti xsi suba sepa zs1 z zs2 _ apply(subgoal_tac "z = (suba, sepa)", simp) using 2(3) apply(sep_auto heap:"2.IH"(2)[of ls rs h rrs sub sep] simp add: sorted_wrt_append) using list_split Cons h_split apply simp_all (* proof that previous assumptions hold later *) apply(rule P_imp_Q_implies_P) apply(rule ent_ex_postI[where x="(tsi,n)"]) apply(rule ent_ex_postI[where x="ti"]) apply(rule ent_ex_postI[where x="(zs1 @ (suba, sepa) # zs2)"]) apply(rule ent_ex_postI[where x="zs1"]) apply(rule ent_ex_postI[where x="z"]) apply(rule ent_ex_postI[where x="zs2"]) apply sep_auto (* prove subgoal_tac assumption *) apply (metis (no_types, lifting) list_assn_aux_ineq_len list_assn_len nth_append_length star_false_left star_false_right) done (* eliminate last vacuous case *) apply(rule hoare_triple_preI) apply(auto simp add: split_relation_def dest!: mod_starD list_assn_len)[] done qed qed qed +declare abs_split.node\<^sub>i.simps [simp add] - -declare abs_split.node\<^sub>i.simps [simp add] lemma node\<^sub>i_rule: assumes c_cap: "2*k \ c" "c \ 4*k+1" shows "\<^sub>a id_assn) ts tsi * btree_assn k t ti> node\<^sub>i k (a,n) ti <\r. btupi_assn k (abs_split.node\<^sub>i k ts t) r >\<^sub>t" -proof (cases "length ts \ 2*k") +proof (cases "length ts \ 2 * k") case [simp]: True then show ?thesis apply(subst node\<^sub>i_def) apply(rule hoare_triple_preI) apply(sep_auto dest!: mod_starD list_assn_len) apply(sep_auto simp add: is_pfa_def)[] using c_cap apply(sep_auto simp add: is_pfa_def)[] apply(sep_auto dest!: mod_starD list_assn_len)[] using True apply(sep_auto dest!: mod_starD list_assn_len) done next + note max.absorb1 [simp del] max.absorb2 [simp del] max.absorb3 [simp del] max.absorb4 [simp del] + note min.absorb1 [simp del] min.absorb2 [simp del] min.absorb3 [simp del] min.absorb4 [simp del] case [simp]: False then obtain ls sub sep rs where split_half_eq: "BTree_Set.split_half ts = (ls,(sub,sep)#rs)" using abs_split.node\<^sub>i_cases by blast then show ?thesis apply(subst node\<^sub>i_def) apply(rule hoare_triple_preI) apply(sep_auto dest!: mod_starD list_assn_len) apply(sep_auto simp add: split_relation_alt split_relation_length is_pfa_def dest!: mod_starD list_assn_len) - using False apply(sep_auto simp add: split_relation_alt ) using False apply(sep_auto simp add: is_pfa_def)[] apply(sep_auto)[] apply(sep_auto simp add: is_pfa_def split_relation_alt)[] using c_cap apply(sep_auto simp add: is_pfa_def)[] - apply(sep_auto)[] + apply(sep_auto)[] using c_cap apply(sep_auto simp add: is_pfa_def)[] using c_cap apply(simp) apply(vcg) apply(simp) apply(rule impI) subgoal for _ _ _ _ rsa subi ba rn lsi al ar _ thm ent_ex_postI thm ent_ex_postI[where x="take (length tsi div 2) tsi"] (* instantiate right hand side *) apply(rule ent_ex_postI[where x="(rsa,rn)"]) apply(rule ent_ex_postI[where x="ti"]) apply(rule ent_ex_postI[where x="(drop (Suc (length tsi div 2)) tsi)"]) apply(rule ent_ex_postI[where x="lsi"]) apply(rule ent_ex_postI[where x="subi"]) apply(rule ent_ex_postI[where x="take (length tsi div 2) tsi"]) (* introduce equality between equality of split tsi/ts and original lists *) apply(simp add: split_relation_alt) apply(subgoal_tac "tsi = take (length tsi div 2) tsi @ (subi, ba) # drop (Suc (length tsi div 2)) tsi") apply(rule back_subst[where a="blist_assn k ts (take (length tsi div 2) tsi @ (subi, ba) # (drop (Suc (length tsi div 2)) tsi))" and b="blist_assn k ts tsi"]) apply(rule back_subst[where a="blist_assn k (take (length tsi div 2) ts @ (sub, sep) # rs)" and b="blist_assn k ts"]) apply(subst list_assn_aux_append_Cons) apply sep_auto apply sep_auto apply simp apply simp apply(rule back_subst[where a="tsi ! (length tsi div 2)" and b="(subi, ba)"]) apply(rule id_take_nth_drop) apply simp apply simp done done qed declare abs_split.node\<^sub>i.simps [simp del] lemma node\<^sub>i_no_split: "length ts \ 2*k \ abs_split.node\<^sub>i k ts t = abs_split.T\<^sub>i (Node ts t)" by (simp add: abs_split.node\<^sub>i.simps) lemma node\<^sub>i_rule_app: "\2*k \ c; c \ 4*k+1\ \ node\<^sub>i k (aa, al) ri i k (ls @ [(l, a)]) r)>\<^sub>t" proof - note node\<^sub>i_rule[of k c "(tsi' @ [(li, ai)])" aa al "(ls @ [(l, a)])" r ri] moreover assume "2*k \ c" "c \ 4*k+1" ultimately show ?thesis by (simp add: mult.left_assoc) qed lemma node\<^sub>i_rule_ins2: "\2*k \ c; c \ 4*k+1; length ls = length lsi\ \ node\<^sub>i k (aa, al) ti i k (ls @ (l, a) # (r,a') # rs) t)>\<^sub>t" proof - assume [simp]: "2*k \ c" "c \ 4*k+1" "length ls = length lsi" moreover note node\<^sub>i_rule[of k c "(lsi @ (li, ai) # (ri,a'i) # rsi)" aa al "(ls @ (l, a) # (r,a') # rs)" t ti] ultimately show ?thesis by (simp add: mult.left_assoc list_assn_aux_append_Cons) qed lemma ins_rule: "sorted_less (inorder t) \ ins k x ti <\r. btupi_assn k (abs_split.ins k x t) r>\<^sub>t" proof (induction k x t arbitrary: ti rule: abs_split.ins.induct) case (1 k x) then show ?case apply(subst ins.simps) apply (sep_auto simp add: pure_app_eq) done next case (2 k x ts t) obtain ls rrs where list_split: "split ts x = (ls,rrs)" by (cases "split ts x") have [simp]: "sorted_less (separators ts)" using "2.prems" sorted_inorder_separators by simp have [simp]: "sorted_less (inorder t)" using "2.prems" sorted_inorder_induct_last by simp show ?case proof (cases rrs) case Nil then show ?thesis proof (cases "abs_split.ins k x t") case (T\<^sub>i a) then show ?thesis apply(subst ins.simps) apply(sep_auto) subgoal for p tsil tsin tti using Nil list_split by (simp add: list_assn_aux_ineq_len split_relation_alt) subgoal for p tsil tsin tti tsi' i tsin' _ sub sep apply(rule hoare_triple_preI) using Nil list_split by (simp add: list_assn_aux_ineq_len split_relation_alt) subgoal for p tsil tsin tti tsi' thm "2.IH"(1)[of ls rrs tti] using Nil list_split T\<^sub>i apply(sep_auto split!: list.splits simp add: split_relation_alt heap add: "2.IH"(1)[of ls rrs tti]) subgoal for ai apply(cases ai) apply sep_auto apply sep_auto done done done next case (Up\<^sub>i l a r) then show ?thesis apply(subst ins.simps) apply(sep_auto) subgoal for p tsil tsin tti using Nil list_split by (simp add: list_assn_aux_ineq_len split_relation_alt) subgoal for p tsil tsin tti tsi' i tsin' _ sub sep using Nil list_split by (simp add: list_assn_aux_ineq_len split_relation_alt) subgoal for p tsil tsin tti tsi' i tsin' thm "2.IH"(1)[of ls rrs tti] using Nil list_split Up\<^sub>i apply(sep_auto split!: list.splits simp add: split_relation_alt heap add: "2.IH"(1)[of ls rrs tti]) subgoal for ai apply(cases ai) apply sep_auto apply(rule hoare_triple_preI) apply(sep_auto) apply(auto dest!: mod_starD simp add: is_pfa_def)[] apply (sep_auto) subgoal for li ai ri (* no split case *) apply(subgoal_tac "length (ls @ [(l,a)]) \ 2*k") apply(simp add: node\<^sub>i_no_split) apply(rule ent_ex_postI[where x="(tsil,Suc tsin)"]) apply(rule ent_ex_postI[where x="ri"]) apply(rule ent_ex_postI[where x="tsi' @ [(li, ai)]"]) apply(sep_auto) apply (sep_auto dest!: mod_starD list_assn_len simp add: is_pfa_def) done (* split case*) apply(sep_auto heap add: node\<^sub>i_rule_app) done done done qed next case (Cons a rs) obtain sub sep where a_split: "a = (sub,sep)" by (cases a) then have [simp]: "sorted_less (inorder sub)" using "2.prems" abs_split.split_axioms list_split Cons sorted_inorder_induct_subtree split_def by fastforce then show ?thesis proof(cases "x = sep") case True show ?thesis apply(subst ins.simps) apply(sep_auto) subgoal for p tsil tsin tti tsi j subi using Cons list_split a_split True by sep_auto subgoal for p tsil tsin tti tsi j _ _ subi sepi apply(rule hoare_triple_preI) using Cons list_split a_split True apply(subgoal_tac "sepi = sep") apply (sep_auto simp add: split_relation_alt) apply(sep_auto simp add: list_assn_prod_map dest!: mod_starD id_assn_list) by (metis length_map snd_conv snd_map_help(2) split_relation_access) subgoal for p tsil tsin tti tsi j apply(rule hoare_triple_preI) using Cons list_split by (sep_auto simp add: split_relation_alt dest!: mod_starD list_assn_len) done next case False then show ?thesis proof (cases "abs_split.ins k x sub") case (T\<^sub>i a') then show ?thesis apply(auto simp add: Cons list_split a_split False) using False apply simp apply(subst ins.simps) apply vcg apply auto apply(rule norm_pre_ex_rule)+ (* at this point, we want to introduce the split, and after that tease the hoare triple assumptions out of the bracket, s.t. we don't split twice *) apply vcg apply sep_auto using list_split Cons apply(simp add: split_relation_alt list_assn_append_Cons_left) apply (rule impI) apply(rule norm_pre_ex_rule)+ apply(rule hoare_triple_preI) apply sep_auto (* discard wrong branch *) subgoal for p tsil tsin ti zs1 subi sepi zs2 _ _ suba apply(subgoal_tac "sepi = x") using list_split Cons a_split apply(auto dest!: mod_starD )[] apply(auto dest!: mod_starD list_assn_len)[] done (* actual induction branch *) subgoal for p tsil tsin ti zs1 subi sepi zs2 _ _ n z suba sepa apply (cases a, simp) apply(subgoal_tac "subi = suba", simp) using list_split a_split T\<^sub>i False apply (vcg heap: 2) apply(auto split!: btupi.splits) (* careful progression for manual value insertion *) apply vcg apply simp apply vcg apply simp subgoal for a'i q r apply(rule impI) apply(simp add: list_assn_append_Cons_left) apply(rule ent_ex_postI[where x="(tsil,tsin)"]) apply(rule ent_ex_postI[where x="ti"]) apply(rule ent_ex_postI[where x="(zs1 @ (a'i, sepi) # zs2)"]) apply(rule ent_ex_postI[where x="zs1"]) apply(rule ent_ex_postI[where x="(a'i,sep)"]) apply(rule ent_ex_postI[where x="zs2"]) apply sep_auto apply (simp add: pure_app_eq) apply(sep_auto dest!: mod_starD list_assn_len)[] done apply (metis list_assn_aux_ineq_len Pair_inject list_assn_len nth_append_length star_false_left star_false_right) done subgoal for p tsil tsin ti zs1 subi sepi zs2 _ _ suba apply(auto dest!: mod_starD list_assn_len)[] done done next case (Up\<^sub>i l w r) then show ?thesis apply(auto simp add: Cons list_split a_split False) using False apply simp apply(subst ins.simps) apply vcg apply auto apply(rule norm_pre_ex_rule)+ (* at this point, we want to introduce the split, and after that tease the hoare triple assumptions out of the bracket, s.t. we don't split twice *) apply vcg apply sep_auto using list_split Cons apply(simp add: split_relation_alt list_assn_append_Cons_left) apply (rule impI) apply(rule norm_pre_ex_rule)+ apply(rule hoare_triple_preI) apply sep_auto (* discard wrong branch *) subgoal for p tsil tsin ti zs1 subi sepi zs2 _ _ suba apply(subgoal_tac "sepi = x") using list_split Cons a_split apply(auto dest!: mod_starD )[] apply(auto dest!: mod_starD list_assn_len)[] done (* actual induction branch *) subgoal for p tsil tsin ti zs1 subi sepi zs2 _ _ n z suba sepa apply(subgoal_tac "subi = suba", simp) thm 2(2)[of ls rrs a rs sub sep] using list_split a_split Cons Up\<^sub>i False apply (sep_auto heap: 2(2)) apply(auto split!: btupi.splits) (* careful progression for manual value insertion *) apply vcg apply simp subgoal for li wi ri u (* no split case *) apply (cases u,simp) apply (sep_auto dest!: mod_starD list_assn_len heap: pfa_insert_grow_rule) apply (simp add: is_pfa_def)[] apply (metis le_less_linear length_append length_take less_not_refl min.absorb2 trans_less_add1) apply(simp add: is_pfa_def) apply (metis add_Suc_right length_Cons length_append length_take min.absorb2) apply(sep_auto split: prod.splits dest!: mod_starD list_assn_len)[] (* no split case *) apply(subgoal_tac "length (ls @ [(l,w)]) \ 2*k") apply(simp add: node\<^sub>i_no_split) apply(rule ent_ex_postI[where x="(tsil,Suc tsin)"]) apply(rule ent_ex_postI[where x="ti"]) apply(rule ent_ex_postI[where x="(zs1 @ (li, wi) # (ri, sep) # zs2)"]) apply(sep_auto dest!: mod_starD list_assn_len) apply (sep_auto dest!: mod_starD list_assn_len simp add: is_pfa_def) done apply vcg apply simp subgoal for x21 x22 x23 u (* split case *) apply (cases u,simp) thm pfa_insert_grow_rule[where ?l="((zs1 @ (suba, sepi) # zs2)[length ls := (x23, sepa)])"] apply (sep_auto dest!: mod_starD list_assn_len heap: pfa_insert_grow_rule) apply (simp add: is_pfa_def)[] apply (metis le_less_linear length_append length_take less_not_refl min.absorb2 trans_less_add1) apply(auto split: prod.splits dest!: mod_starD list_assn_len)[] apply (vcg heap: node\<^sub>i_rule_ins2) apply simp apply simp apply simp apply sep_auto done apply(auto dest!: mod_starD list_assn_len)[] done subgoal for p tsil tsin ti zs1 subi sepi zs2 _ _ suba apply(auto dest!: mod_starD list_assn_len)[] done done qed qed qed qed text "The imperative insert refines the abstract insert." lemma insert_rule: assumes "k > 0" "sorted_less (inorder t)" shows " insert k x ti <\r. btree_assn k (abs_split.insert k x t) r>\<^sub>t" unfolding insert_def apply(cases "abs_split.ins k x t") apply(sep_auto split!: btupi.splits heap: ins_rule[OF assms(2)]) using assms apply(vcg heap: ins_rule[OF assms(2)]) apply(simp split!: btupi.splits) apply(vcg) apply auto[] apply vcg apply auto[] subgoal for l a r li ai ri tsa tsn ti apply(rule ent_ex_postI[where x="(tsa,tsn)"]) apply(rule ent_ex_postI[where x="ri"]) apply(rule ent_ex_postI[where x="[(li, ai)]"]) apply sep_auto done done text "The \"pure\" resulting rule follows automatically." lemma insert_rule': shows "(abs_split.invar_inorder (Suc k) t \ sorted_less (inorder t))> insert (Suc k) x ti <\ri.\\<^sub>Ar. btree_assn (Suc k) r ri * \(abs_split.invar_inorder (Suc k) r \ sorted_less (inorder r) \ inorder r = (ins_list x (inorder t)))>\<^sub>t" using abs_split.insert_bal abs_split.insert_order abs_split.insert_inorder by (sep_auto heap: insert_rule simp add: sorted_ins_list) lemma list_update_length2 [simp]: "(xs @ x # y # ys)[Suc (length xs) := z] = (xs @ x # z # ys)" by (induct xs, auto) lemma node\<^sub>i_rule_ins: "\2*k \ c; c \ 4*k+1; length ls = length lsi\ \ node\<^sub>i k (aa, al) ti i k (ls @ (l, a) # rs) t)>\<^sub>t" proof - assume [simp]: "2*k \ c" "c \ 4*k+1" "length ls = length lsi" moreover note node\<^sub>i_rule[of k c "(lsi @ (li, ai) # rsi)" aa al "(ls @ (l, a) # rs)" t ti] ultimately show ?thesis by (simp add: mult.left_assoc list_assn_aux_append_Cons) qed lemma btupi_assn_T: "h \ btupi_assn k (abs_split.node\<^sub>i k ts t) (T\<^sub>i x) \ abs_split.node\<^sub>i k ts t = abs_split.T\<^sub>i (Node ts t)" apply(auto simp add: abs_split.node\<^sub>i.simps dest!: mod_starD split!: list.splits) done lemma btupi_assn_Up: "h \ btupi_assn k (abs_split.node\<^sub>i k ts t) (Up\<^sub>i l a r) \ abs_split.node\<^sub>i k ts t = ( case BTree_Set.split_half ts of (ls, (sub,sep)#rs) \ abs_split.Up\<^sub>i (Node ls sub) sep (Node rs t))" apply(auto simp add: abs_split.node\<^sub>i.simps dest!: mod_starD split!: list.splits) done lemma second_last_access:"(xs@a#b#ys) ! Suc(length xs) = b" by (simp add: nth_via_drop) lemma pfa_assn_len:"h \ is_pfa k ls (a,n) \ n \ k \ length ls = n" by (auto simp add: is_pfa_def) (*declare "impCE"[rule del]*) lemma rebalance_middle_tree_rule: assumes "height t = height sub" and "case rs of (rsub,rsep) # list \ height rsub = height t | [] \ True" and "i = length ls" shows " rebalance_middle_tree k (a,n) i ti <\r. btnode_assn k (abs_split.rebalance_middle_tree k ls sub sep rs t) r >\<^sub>t" apply(simp add: list_assn_append_Cons_left) apply(rule norm_pre_ex_rule)+ proof(goal_cases) case (1 lsi z rsi) then show ?case proof(cases z) case z_split: (Pair subi sepi) then show ?thesis proof(cases sub) case sub_leaf[simp]: Leaf then have t_leaf[simp]: "t = Leaf" using assms by (cases t) auto show ?thesis apply (subst rebalance_middle_tree_def) apply (rule hoare_triple_preI) apply (vcg) using assms apply (auto dest!: mod_starD list_assn_len split!: option.splits) apply (vcg) apply (auto dest!: mod_starD list_assn_len split!: option.splits) apply (rule ent_ex_postI[where x=tsi]) apply sep_auto done next case sub_node[simp]: (Node mts mt) then obtain tts tt where t_node[simp]: "t = Node tts tt" using assms by (cases t) auto then show ?thesis proof(cases "length mts \ k \ length tts \ k") case True then show ?thesis apply(subst rebalance_middle_tree_def) apply(rule hoare_triple_preI) apply(sep_auto dest!: mod_starD) using assms apply (auto dest!: list_assn_len)[] using assms apply(sep_auto split!: prod.splits) using assms apply (auto simp del: height_btree.simps dest!: mod_starD list_assn_len)[] using z_split apply(auto)[] subgoal for _ _ _ _ _ _ _ _ tp tsia' tsin' _ _ _ _ _ _ _ _ _ _ tsia tsin tti ttsi sepi subp apply(auto dest!: mod_starD list_assn_len simp: prod_assn_def)[] apply(vcg) apply(auto)[] apply(rule ent_ex_postI[where x="lsi@(Some subp, sepi)#rsi"]) apply(rule ent_ex_postI[where x="(tsia, tsin)"]) apply(rule ent_ex_postI[where x="tti"]) apply(rule ent_ex_postI[where x=ttsi]) apply(sep_auto)[] apply(rule hoare_triple_preI) using True apply(auto dest!: mod_starD list_assn_len) done done next case False then show ?thesis proof(cases rs) case Nil then show ?thesis apply(subst rebalance_middle_tree_def) apply(rule hoare_triple_preI) apply(sep_auto dest!: mod_starD) using assms apply (auto dest!: list_assn_len)[] apply(sep_auto split!: prod.splits) using assms apply (auto simp del: height_btree.simps dest!: mod_starD list_assn_len)[] using z_split apply(auto)[] subgoal for _ _ _ _ _ _ _ _ tp tsia' tsin' _ _ _ _ _ _ _ _ _ _ tsia tsin tti ttsi apply(auto dest!: mod_starD list_assn_len simp: prod_assn_def)[] apply(vcg) using False apply(auto dest!: mod_starD list_assn_len) done apply(sep_auto dest!: mod_starD) using assms apply (auto dest!: list_assn_len)[] using assms apply (auto dest!: list_assn_len)[] apply(sep_auto) using assms apply (auto dest!: list_assn_len mod_starD)[] using assms apply (auto dest!: list_assn_len mod_starD)[] (* Issue: we do not know yet what 'subp is pointing at *) subgoal for _ _ _ _ _ _ tp tsia tsin tti ttsi _ _ _ _ _ _ _ _ tsia' tsin' tti' tsi' subi sepi subp apply(subgoal_tac "z = (subi, sepi)") prefer 2 apply (metis assms(3) list_assn_len nth_append_length) apply simp apply(vcg) subgoal (* still the "IF" branch *) apply(rule entailsI) (* solves impossible case*) using False apply (auto dest!: list_assn_len mod_starD)[] done apply (auto del: impCE) apply(thin_tac "_ \ _")+ apply(rule hoare_triple_preI) (* for each possible combination of \ and \\, a subgoal is created *) apply(sep_auto heap add: node\<^sub>i_rule_ins dest!: mod_starD del: impCE) apply (auto dest!: pfa_assn_len)[] apply (auto dest!: pfa_assn_len list_assn_len)[] subgoal apply(thin_tac "_ \ _")+ apply(rule hoare_triple_preI) apply(sep_auto split!: btupi.splits del: impCE) apply(auto dest!: btupi_assn_T mod_starD del: impCE)[] apply(rule ent_ex_postI[where x="lsi"]) apply sep_auto apply (sep_auto del: impCE) apply(auto dest!: btupi_assn_Up mod_starD split!: list.splits del: impCE)[] subgoal for li ai ri apply(rule ent_ex_postI[where x="lsi @ [(li, ai)]"]) apply sep_auto done done apply (sep_auto del: impCE) using assms apply(auto dest!: pfa_assn_len list_assn_len mod_starD)[] using assms apply(auto dest!: pfa_assn_len list_assn_len mod_starD)[] done done next case (Cons rss rrs) then show ?thesis apply(subst rebalance_middle_tree_def) apply(rule hoare_triple_preI) apply(sep_auto dest!: mod_starD) using assms apply (auto dest!: list_assn_len)[] apply(sep_auto split!: prod.splits) using assms apply (auto simp del: height_btree.simps dest!: mod_starD list_assn_len)[] apply(auto)[] subgoal for _ _ _ _ _ _ _ _ tp tsia' tsin' _ _ _ _ _ _ _ _ _ _ tsia tsin tti ttsi apply(auto dest!: mod_starD list_assn_len simp: prod_assn_def)[] apply(vcg) using False apply(auto dest!: mod_starD list_assn_len) done apply(sep_auto dest!: mod_starD del: impCE) using assms apply (auto dest!: list_assn_len)[] apply(sep_auto del: impCE) using assms apply (auto dest!: list_assn_len mod_starD)[] (* Issue: we do not know yet what 'xa' is pointing at *) subgoal for list_heap1 list_heap2 _ _ _ _ _ _ tp ttsia' ttsin' tti' ttsi' _ _ _ _ _ _ _ _ ttsia ttsin tti ttsi subi sepi subp apply(subgoal_tac "z = (subi, sepi)") prefer 2 apply (metis assms(3) list_assn_len nth_append_length) apply simp apply(vcg) subgoal (* still the "IF" branch *) apply(rule entailsI) (* solves impossible case*) using False apply (auto dest!: list_assn_len mod_starD)[] done apply simp subgoal for subtsi subti subts ti subi subtsl ttsl (* TODO different nodei rule here *) supply R = node\<^sub>i_rule_ins[where k=k and c="(max (2 * k) (Suc (_ + ttsin)))" and lsi=subts] thm R apply(cases subtsi) apply(sep_auto heap add: R pfa_append_extend_grow_rule dest!: mod_starD del: impCE) (* all of these cases are vacuous *) using assms apply (auto dest!: list_assn_len pfa_assn_len)[] using assms apply (auto dest!: list_assn_len pfa_assn_len)[] using assms apply (auto dest!: list_assn_len pfa_assn_len)[] apply(sep_auto split!: btupi.splits del: impCE) using assms apply (auto dest!: list_assn_len pfa_assn_len)[] apply(thin_tac "_ \ _")+ apply(rule hoare_triple_preI) apply (cases rsi) apply(auto dest!: list_assn_len mod_starD)[] (* TODO avoid creating subgoals here but still split the heap? do we need to do that anyways *) subgoal for subtsa subtsn mtsa mtsn mtt mtsi _ _ _ _ _ _ _ _ rsubsep _ rrsi rssi (* ensuring that the tree to the right is not none *) apply (cases rsubsep) apply(subgoal_tac "rsubsep = rrsi") prefer 2 using assms apply(auto dest!: list_assn_len mod_starD del: impCE simp add: second_last_access)[] apply (simp add: prod_assn_def) apply(cases rss) apply simp subgoal for rsubi rsepi rsub rsep apply(subgoal_tac "height rsub \ 0") prefer 2 using assms apply(auto)[] apply(cases rsubi; cases rsub) apply simp+ (* now we may proceed *) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) subgoal for rsubi rsubts rsubt rsubtsi' rsubti rsubtsi subnode apply(cases "kvs subnode") apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) subgoal for _ rsubtsn subtsmergedi apply (cases subtsmergedi) apply simp apply (vcg (ss)) subgoal for subtsmergeda _ supply R = node\<^sub>i_rule_ins[where k=k and c="max (2*k) (Suc (subtsn + rsubtsn))" and ls="mts" and al="Suc (subtsn+rsubtsn)" and aa=subtsmergeda and ti=rsubti and rsi=rsubtsi and li=subti and a=sep and ai=sep ] thm R apply(rule P_imp_Q_implies_P) apply(auto del: impCE dest!: mod_starD list_assn_len)[] apply(rule hoare_triple_preI) apply(subgoal_tac "subtsn \ 2*k \ rsubtsn \ 2*k") prefer 2 apply (auto simp add: is_pfa_def)[] apply (sep_auto heap add: R del: impCE) apply(sep_auto split!: btupi.splits del: impCE) using assms apply(auto dest!: mod_starD list_assn_len)[] apply(sep_auto del: impCE) using assms apply(auto dest!: mod_starD list_assn_len pfa_assn_len del: impCE)[] apply(thin_tac "_ \ _")+ apply(rule hoare_triple_preI) apply (drule btupi_assn_T mod_starD | erule conjE exE)+ apply vcg apply simp subgoal for rsubtsi ai tsian apply(cases tsian) apply simp apply(rule P_imp_Q_implies_P) apply(rule ent_ex_postI[where x="lsi @ (ai, rsep) # rssi"]) apply(rule ent_ex_postI[where x="(ttsia, ttsin)"]) apply(rule ent_ex_postI[where x="tti"]) apply(rule ent_ex_postI[where x="ttsi"]) using assms apply (sep_auto dest!: list_assn_len) done subgoal for _ _ rsubp rsubtsa _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ li ai ri apply(sep_auto del: impCE) using assms apply(auto dest!: list_assn_len)[] apply(sep_auto del: impCE) using assms apply(auto dest!: list_assn_len)[] apply(thin_tac "_ \ _")+ apply(rule hoare_triple_preI) apply (drule btupi_assn_Up mod_starD | erule conjE exE)+ apply vcg (* generates two identical subgoals ? *) apply(simp split!: list.split) apply(rule ent_ex_postI[where x="(lsi @ (li, ai) # (ri, rsepi) # rssi)"]) apply(rule ent_ex_postI[where x="(ttsia, ttsin)"]) apply(rule ent_ex_postI[where x="tti"]) apply(rule ent_ex_postI[where x="ttsi"]) using assms apply (sep_auto dest!: list_assn_len) apply(rule ent_ex_postI[where x="(lsi @ (li, ai) # (ri, rsepi) # rssi)"]) apply(rule ent_ex_postI[where x="(ttsia, ttsin)"]) apply(rule ent_ex_postI[where x="tti"]) apply(rule ent_ex_postI[where x="ttsi"]) using assms apply (sep_auto dest!: list_assn_len) done done done done done done done done done qed qed qed qed qed lemma rebalance_last_tree_rule: assumes "height t = height sub" and "ts = list@[(sub,sep)]" shows " rebalance_last_tree k tsia ti <\r. btnode_assn k (abs_split.rebalance_last_tree k ts t) r >\<^sub>t" apply(subst rebalance_last_tree_def) apply(rule hoare_triple_preI) using assms apply(auto dest!: mod_starD) apply(subgoal_tac "length tsi - Suc 0 = length list") prefer 2 apply(auto dest!: list_assn_len)[] using assms apply(sep_auto) supply R = rebalance_middle_tree_rule[where ls="list" and rs="[]" and i="length tsi - 1", simplified] apply(cases tsia) using R by blast partial_function (heap) split_max ::"nat \ ('a::{default,heap,linorder}) btnode ref option \ ('a btnode ref option \ 'a) Heap" where "split_max k r_t = (case r_t of Some p_t \ do { t \ !p_t; (case (last t) of None \ do { (sub,sep) \ pfa_last (kvs t); tsi' \ pfa_butlast (kvs t); p_t := Btnode tsi' sub; return (Some p_t, sep) } | Some x \ do { (sub,sep) \ split_max k (Some x); p_t' \ rebalance_last_tree k (kvs t) sub; p_t := p_t'; return (Some p_t, sep) }) }) " declare abs_split.split_max.simps [simp del] abs_split.rebalance_last_tree.simps [simp del] height_btree.simps [simp del] lemma split_max_rule: assumes "abs_split.nonempty_lasttreebal t" and "t \ Leaf" shows " split_max k ti <((btree_assn k) \\<^sub>a id_assn) (abs_split.split_max k t)>\<^sub>t" using assms proof(induction k t arbitrary: ti rule: abs_split.split_max.induct) case (2 Leaf) then show ?case by auto next case (1 k ts tt) then show ?case proof(cases tt) case Leaf then show ?thesis apply(subst split_max.simps) apply (vcg) using assms apply auto[] apply (vcg (ss)) apply simp apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply(rule hoare_triple_preI) apply (vcg (ss)) using 1 apply(auto dest!: mod_starD)[] apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) subgoal for tp tsi tti tsi' tnode subsep sub sep apply(cases tsi) apply(rule hoare_triple_preI) apply (vcg) apply(auto simp add: prod_assn_def abs_split.split_max.simps split!: prod.splits) subgoal for tsia tsin _ _ tsin' lastsep lastsub apply(rule ent_ex_postI[where x="(tsia, tsin')"]) apply(rule ent_ex_postI[where x="sub"]) apply(rule ent_ex_postI[where x="(butlast tsi')"]) using 1 apply (auto dest!: mod_starD simp add: list_assn_append_Cons_left) apply sep_auto done done apply(sep_auto) done next case (Node tts ttt) have IH_help: "abs_split.nonempty_lasttreebal tt \ tt \ Leaf \ split_max k (Some ttp) <(btree_assn k \\<^sub>a id_assn) (abs_split.split_max k tt)>\<^sub>t" for ttp using "1.IH" Node by blast obtain butlasttts l_sub l_sep where ts_split:"tts = butlasttts@[(l_sub, l_sep)]" using 1 Node by auto from Node show ?thesis apply(subst split_max.simps) apply (vcg) using 1 apply auto[] apply (vcg (ss)) apply simp apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) apply (vcg (ss)) using 1 apply(auto dest!: mod_starD)[] apply (vcg (ss)) subgoal for tp tsi tti tsi' tnode ttp using "1.prems" apply (vcg heap add: IH_help) apply simp apply simp apply(subst prod_assn_def) apply(cases "abs_split.split_max k tt") apply (auto simp del: abs_split.split_max.simps abs_split.rebalance_last_tree.simps height_btree.simps)[] subgoal for ttsubi ttmaxi ttsub ttmax butlasttsi' lasttssubi butlastts lasttssub lasttssepi lasttssep apply(rule hoare_triple_preI) supply R = rebalance_last_tree_rule[where k=k and tsia=tsi and ti=ttsubi and t=ttsub and tsi=tsi' and ts=" (butlasttsi' @ [(lasttssubi, lasttssepi)])" and list=butlasttsi' and sub=lasttssubi and sep=lasttssepi] thm R using ts_split (*TODO weird post conditions... *) apply (sep_auto heap add: R simp del: abs_split.split_max.simps abs_split.rebalance_last_tree.simps height_btree.simps dest!: mod_starD) apply (metis abs_split.nonempty_lasttreebal.simps(2) abs_split.split_max_height btree.distinct(1)) apply simp apply(rule hoare_triple_preI) apply (simp add: prod_assn_def) apply vcg apply(subst abs_split.split_max.simps) using "1.prems" apply(auto dest!: mod_starD split!: prod.splits btree.splits) subgoal for _ _ _ _ _ _ _ _ _ _ tp' apply(cases "abs_split.rebalance_last_tree k (butlasttsi' @ [(lasttssubi, lasttssepi)]) ttsub"; cases tp') apply auto apply(rule ent_ex_preI) subgoal for _ _ tsia' tsin' tt' _ tsi' apply(rule ent_ex_postI[where x="(tsia', tsin')"]) apply(rule ent_ex_postI[where x="tt'"]) apply(rule ent_ex_postI[where x="tsi'"]) apply sep_auto done done done done done qed qed partial_function (heap) del ::"nat \ 'a \ ('a::{default,heap,linorder}) btnode ref option \ 'a btnode ref option Heap" where "del k x ti = (case ti of None \ return None | Some p \ do { node \ !p; i \ imp_split (kvs node) x; tsl \ pfa_length (kvs node); if i < tsl then do { (sub,sep) \ pfa_get (kvs node) i; if sep \ x then do { sub' \ del k x sub; kvs' \ pfa_set (kvs node) i (sub',sep); node' \ rebalance_middle_tree k kvs' i (last node); p := node'; return (Some p) } else if sub = None then do{ kvs' \ pfa_delete (kvs node) i; p := (Btnode kvs' (last node)); return (Some p) } else do { sm \ split_max k sub; kvs' \ pfa_set (kvs node) i sm; node' \ rebalance_middle_tree k kvs' i (last node); p := node'; return (Some p) } } else do { t' \ del k x (last node); node' \ rebalance_last_tree k (kvs node) t'; p := node'; return (Some p) } })" lemma rebalance_middle_tree_update_rule: assumes "height tt = height sub" and "case rs of (rsub,rsep) # list \ height rsub = height tt | [] \ True" and "i = length ls" shows " rebalance_middle_tree k a i ti \<^sub>t" proof (cases a) case [simp]: (Pair a n) note R=rebalance_middle_tree_rule[of tt sub rs i ls k "zs1@(x', sep)#zs2" a n sep ti] show ?thesis apply(rule hoare_triple_preI) using R assms apply (sep_auto dest!: mod_starD list_assn_len simp add: prod_assn_def) using assn_times_assoc star_aci(3) by auto qed lemma del_rule: assumes "bal t" and "sorted (inorder t)" and "root_order k t" and "k > 0" shows " del k x ti \<^sub>t" using assms proof (induction k x t arbitrary: ti rule: abs_split.del.induct) case (1 k x) then show ?case apply(subst del.simps) apply sep_auto done next case (2 k x ts tt ti) obtain ls rs where split_ts[simp]: "split ts x = (ls, rs)" by (cases "split ts x") obtain tss lastts_sub lastts_sep where last_ts: "ts = tss@[(lastts_sub, lastts_sep)]" using "2.prems" apply auto by (metis abs_split.isin.cases neq_Nil_rev_conv) show ?case proof(cases "rs") case Nil then show ?thesis apply(subst del.simps) apply sep_auto using "2.prems"(2) sorted_inorder_separators apply blast apply(rule hoare_triple_preI) apply (sep_auto) using Nil apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] using Nil apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] using Nil apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] apply (sep_auto heap add: "2.IH"(1)) using "2.prems" apply (auto dest!: mod_starD)[] using "2.prems" apply (auto dest!: mod_starD simp add: sorted_wrt_append)[] using "2.prems" order_impl_root_order apply (auto dest!: mod_starD)[] using "2.prems" apply (auto)[] subgoal for tp tsia tsin tti tsi i _ _ tti' apply(rule hoare_triple_preI) supply R = rebalance_last_tree_rule[where t="(abs_split.del k x tt)" and ti=tti' and ts=ts and sub=lastts_sub and list=tss and sep=lastts_sep] thm R using last_ts apply(sep_auto heap add: R) using "2.prems" abs_split.del_height[of k tt x] order_impl_root_order[of k tt] apply (auto dest!: mod_starD)[] apply simp apply(rule hoare_triple_preI) apply (sep_auto) apply(cases "abs_split.rebalance_last_tree k ts (abs_split.del k x tt)") apply(auto simp add: split_relation_alt dest!: mod_starD list_assn_len) subgoal for tnode apply (cases tnode; sep_auto) done done done next case [simp]: (Cons rrs rss) then obtain sub sep where [simp]: "rrs = (sub, sep)" by (cases rrs) consider (sep_n_x) "sep \ x" | (sep_x_Leaf) "sep = x \ sub = Leaf" | (sep_x_Node) "sep = x \ (\ts t. sub = Node ts t)" using btree.exhaust by blast then show ?thesis proof(cases) case sep_n_x then show ?thesis apply(subst del.simps) apply sep_auto using "2.prems"(2) sorted_inorder_separators apply blast apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply simp apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) subgoal for tp tsi ti' tsi' tnode i tsi'l subsep subi sepi (* TODO this causes 4 subgoals *) apply(auto simp add: split_relation_alt list_assn_append_Cons_left; rule norm_pre_ex_rule; rule norm_pre_ex_rule; rule norm_pre_ex_rule; rule hoare_triple_preI; auto dest!: mod_starD)[] apply (auto simp add: split_relation_alt dest!: list_assn_len)[] subgoal for lsi subi rsi apply(subgoal_tac "subi = None") prefer 2 apply(auto dest!: list_assn_len)[] supply R = "2.IH"(2)[of ls rs rrs rss sub sep] thm R using split_ts apply(sep_auto heap add: R) using "2.prems" apply auto[] apply (metis "2.prems"(2) sorted_inorder_induct_subtree) using "2.prems" apply auto[] apply (meson "2.prems"(4) order_impl_root_order) using "2.prems"(4) apply fastforce apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply (auto simp add: split_relation_alt dest!: list_assn_len)[] apply(vcg (ss)) apply(vcg (ss); simp) apply(cases tsi; simp) subgoal for subi' _ tsia' tsin' supply R = rebalance_middle_tree_update_rule thm R (* TODO create a new heap rule, in the node_i style *) apply(auto dest!: list_assn_len)[] apply(rule hoare_triple_preI) apply (sep_auto heap add: R dest!: mod_starD) using "2.prems" abs_split.del_height[of k sub x] order_impl_root_order[of k sub] apply (auto)[] using "2.prems" apply (auto split!: list.splits)[] apply auto[] apply sep_auto subgoal for _ _ _ _ _ _ _ _ _ _ _ _ _ _ tnode'' apply (cases "(abs_split.rebalance_middle_tree k ls (abs_split.del k x sub) sepi rss tt)"; cases tnode'') apply sep_auto apply sep_auto done done done apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] (* copy pasta of "none" branch *) subgoal for subnode lsi subi rsi apply(subgoal_tac "subi = Some subnode") prefer 2 apply(auto dest!: list_assn_len)[] supply R = "2.IH"(2)[of ls rs rrs rss sub sep] thm R using split_ts apply(sep_auto heap add: R) using "2.prems" apply auto[] apply (metis "2.prems"(2) sorted_inorder_induct_subtree) using "2.prems" apply auto[] apply (meson "2.prems"(4) order_impl_root_order) using "2.prems"(4) apply fastforce apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply (auto simp add: split_relation_alt dest!: list_assn_len)[] apply(vcg (ss)) apply(vcg (ss); simp) apply(cases tsi; simp) subgoal for x' xab a n supply R = rebalance_middle_tree_update_rule thm R (* TODO create a new heap rule, in the node_i style *) apply(auto dest!: list_assn_len)[] apply(rule hoare_triple_preI) apply (sep_auto heap add: R dest!: mod_starD) using "2.prems" abs_split.del_height[of k sub x] order_impl_root_order[of k sub] apply (auto)[] using "2.prems" apply (auto split!: list.splits)[] apply auto[] apply sep_auto subgoal for _ _ _ _ _ _ _ _ _ _ _ _ _ _ tnode' apply (cases "(abs_split.rebalance_middle_tree k ls (abs_split.del k x sub) sepi rss tt)"; cases tnode') apply sep_auto apply sep_auto done done done done apply(rule hoare_triple_preI) using Cons apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] done next case sep_x_Leaf then show ?thesis apply(subst del.simps) apply sep_auto using "2.prems"(2) sorted_inorder_separators apply blast apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply simp apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) subgoal for tp tsi ti' tsi' tnode i tsi'l subsep subi sepi (* TODO this causes 4 subgoals *) apply(auto simp add: split_relation_alt list_assn_append_Cons_left; rule norm_pre_ex_rule; rule norm_pre_ex_rule; rule norm_pre_ex_rule; rule hoare_triple_preI; auto dest!: mod_starD)[] (* the correct subbranch *) subgoal for lsi subi rsi apply(cases tsi) apply (sep_auto) apply(auto simp add: is_pfa_def dest!: list_assn_len)[] apply (metis add_Suc_right le_imp_less_Suc length_append length_take less_add_Suc1 less_trans_Suc list.size(4) min.cobounded2 not_less_eq) apply vcg apply auto subgoal for tsin tsia apply(rule ent_ex_postI[where x="(tsia, tsin-1)"]) apply(rule ent_ex_postI[where x="ti'"]) apply(rule ent_ex_postI[where x="lsi@rsi"]) apply (sep_auto dest!: list_assn_len) done done apply (auto simp add: split_relation_alt dest!: list_assn_len)[] apply (auto simp add: split_relation_alt dest!: list_assn_len)[] apply (auto simp add: split_relation_alt dest!: list_assn_len)[] done apply(rule hoare_triple_preI) using Cons apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] done next case sep_x_Node then show ?thesis apply(subst del.simps) apply sep_auto using "2.prems"(2) sorted_inorder_separators apply blast apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) apply simp apply(vcg (ss)) apply(vcg (ss)) apply(vcg (ss)) subgoal for subts subt tp tsi ti tsi' tnode i tsi'l subsep subi sepi (* TODO this causes 4 subgoals *) apply(auto simp add: split_relation_alt list_assn_append_Cons_left; rule norm_pre_ex_rule; rule norm_pre_ex_rule; rule norm_pre_ex_rule; rule hoare_triple_preI; auto dest!: mod_starD)[] apply (auto simp add: split_relation_alt dest!: list_assn_len)[] apply (auto simp add: split_relation_alt dest!: list_assn_len)[] (* the correct sub branch *) subgoal for subnode lsi subi rsi apply(subgoal_tac "subi = Some subnode") apply (simp del: btree_assn.simps) supply R = split_max_rule[of "Node subts subt" k "Some subnode"] thm R apply(sep_auto heap add: R simp del: btree_assn.simps) using "2.prems" apply(auto dest!: list_assn_len mod_starD simp del: bal.simps order.simps)[] subgoal proof(goal_cases) case 1 then have "order k (Node subts subt)" by blast moreover have "k > 0" by (simp add: "2.prems"(4)) ultimately obtain sub_ls lsub lsep where sub_ts_split: "subts = sub_ls@[(lsub,lsep)]" by (metis abs_split.isin.cases le_0_eq list.size(3) order.simps(2) rev_exhaust zero_less_iff_neq_zero) from 1 have "bal (Node subts subt)" by auto then have "height lsub = height subt" by (simp add: sub_ts_split) then show ?thesis using sub_ts_split by blast qed using "2.prems" abs_split.order_bal_nonempty_lasttreebal[of k subt] order_impl_root_order[of k subt] apply(auto)[] apply (auto simp add: split_relation_alt dest!: list_assn_len)[] apply vcg apply auto[] apply(cases "abs_split.split_max k (Node subts subt)"; simp) subgoal for split_res _ split_sub split_sep apply(cases split_res; simp) subgoal for split_subi split_sepi supply R = rebalance_middle_tree_update_rule[ of tt split_sub rss "length lsi" ls k lsi split_subi split_sep rsi tsi ti ] thm R (* id_assn split_sepi doesnt match yet... *) apply(auto simp add: prod_assn_def dest!: list_assn_len) apply (sep_auto) apply(rule hoare_triple_preI) apply(auto dest!: mod_starD)[] apply (sep_auto heap add: R) using "2.prems" abs_split.split_max_height[of k sub] order_impl_root_order[of k sub] abs_split.order_bal_nonempty_lasttreebal[of k sub] apply (auto)[] using "2.prems" abs_split.split_max_bal[of sub k] order_impl_root_order[of k sub] apply (auto split!: list.splits)[] apply auto[] apply(rule hoare_triple_preI) apply(auto dest!: mod_starD)[] subgoal for subtsi''a subtsi''n ti subtsi'' tnode' apply(cases "(abs_split.rebalance_middle_tree k ls split_sub split_sep rss tt)"; cases "tnode'") apply auto apply sep_auto done done done apply (auto simp add: split_relation_alt dest!: list_assn_len)[] done apply (auto simp add: split_relation_alt dest!: list_assn_len)[] done apply(rule hoare_triple_preI) using Cons apply (auto simp add: split_relation_alt dest!: mod_starD list_assn_len)[] done qed qed qed definition reduce_root ::"('a::{default,heap,linorder}) btnode ref option \ 'a btnode ref option Heap" where "reduce_root ti = (case ti of None \ return None | Some p_t \ do { node \ !p_t; tsl \ pfa_length (kvs node); case tsl of 0 \ return (last node) | _ \ return ti })" lemma reduce_root_rule: " reduce_root ti \<^sub>t" apply(subst reduce_root_def) apply(cases t; cases ti) apply (sep_auto split!: nat.splits list.splits)+ done definition delete ::"nat \ 'a \ ('a::{default,heap,linorder}) btnode ref option \ 'a btnode ref option Heap" where "delete k x ti = do { ti' \ del k x ti; reduce_root ti' }" lemma delete_rule: assumes "bal t" and "root_order k t" and "k > 0" and "sorted (inorder t)" shows " delete k x ti \<^sub>t" apply(subst delete_def) using assms apply (sep_auto heap add: del_rule reduce_root_rule) done lemma empty_rule: shows " empty <\r. btree_assn k (abs_split.empty_btree) r>" apply(subst empty_def) apply(sep_auto simp add: abs_split.empty_btree_def) done end end diff --git a/thys/BenOr_Kozen_Reif/BKR_Proofs.thy b/thys/BenOr_Kozen_Reif/BKR_Proofs.thy --- a/thys/BenOr_Kozen_Reif/BKR_Proofs.thy +++ b/thys/BenOr_Kozen_Reif/BKR_Proofs.thy @@ -1,2223 +1,2224 @@ theory BKR_Proofs imports "Matrix_Equation_Construction" begin definition well_def_signs:: "nat => rat list list \ bool" where "well_def_signs num_polys sign_conds \ \ signs \ set(sign_conds). (length signs = num_polys)" definition satisfies_properties:: "real poly \ real poly list \ nat list list \ rat list list \ rat mat \ bool" where "satisfies_properties p qs subsets signs matrix = ( all_list_constr subsets (length qs) \ well_def_signs (length qs) signs \ distinct signs \ satisfy_equation p qs subsets signs \ invertible_mat matrix \ matrix = matrix_A signs subsets \ set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs) )" section "Base Case" lemma mat_base_case: shows "matrix_A [[1],[-1]] [[],[0]] = (mat_of_rows_list 2 [[1, 1], [1, -1]])" unfolding matrix_A_def mtx_row_def z_def apply (auto) by (simp add: numeral_2_eq_2) lemma base_case_sgas: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "set (characterize_consistent_signs_at_roots_copr p [q]) \ {[1], [- 1]}" unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def by auto lemma base_case_sgas_alt: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" assumes rel_prime: "\q. (List.member qs q) \ coprime p q" shows "set (characterize_consistent_signs_at_roots_copr p qs) \ {[1], [- 1]}" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_sgas nonzero rel_prime apply (auto) using characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def by auto qed lemma base_case_satisfy_equation: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "satisfy_equation p [q] [[],[0]] [[1],[-1]]" proof - have h1: "set (characterize_consistent_signs_at_roots_copr p [q]) \ {[1], [- 1]}" using base_case_sgas assms by auto have h2: " \qa. List.member [q] qa \ coprime p qa" using rel_prime by (simp add: member_rec(1) member_rec(2)) have h3: "all_list_constr [[], [0]] (Suc 0)" unfolding all_list_constr_def by (simp add: list_constr_def member_def) then show ?thesis using matrix_equation[where p = "p", where qs = "[q]", where signs = "[[1],[-1]]", where subsets = "[[],[0]]"] nonzero h1 h2 h3 by auto qed lemma base_case_satisfy_equation_alt: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" assumes rel_prime: "\q. (List.member qs q) \ coprime p q" shows "satisfy_equation p qs [[],[0]] [[1],[-1]]" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_satisfy_equation nonzero rel_prime apply (auto) by (simp add: member_rec(1)) qed lemma base_case_matrix_eq: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "(mult_mat_vec (mat_of_rows_list 2 [[1, 1], [1, -1]]) (construct_lhs_vector p [q] [[1],[-1]]) = (construct_rhs_vector p [q] [[],[0]]))" using mat_base_case base_case_satisfy_equation unfolding satisfy_equation_def by (simp add: nonzero rel_prime) lemma less_two: shows "j < Suc (Suc 0) \ j = 0 \ j = 1" by auto lemma inverse_mat_base_case: shows "inverts_mat (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -(1/2)]]::rat mat) (mat_of_rows_list 2 [[1, 1], [1, -1]]::rat mat)" unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_two by (auto simp add: scalar_prod_def) lemma inverse_mat_base_case_2: shows "inverts_mat (mat_of_rows_list 2 [[1, 1], [1, -1]]::rat mat) (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -(1/2)]]::rat mat) " unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_two by (auto simp add: scalar_prod_def) lemma base_case_invertible_mat: shows "invertible_mat (matrix_A [[1], [- 1]] [[], [0]])" unfolding invertible_mat_def using inverse_mat_base_case inverse_mat_base_case_2 apply (auto) apply (simp add: mat_base_case mat_of_rows_list_def) using mat_base_case by auto section "Inductive Step" subsection "Lemmas on smashing subsets and signs" lemma signs_smash_property: fixes signs1 signs2 :: "'a list list" fixes a b:: "nat" shows "\ (elem :: 'a list). (elem \ (set (signs_smash signs1 signs2)) \ (\ n m :: nat. elem = ((nth signs1 n)@(nth signs2 m))))" proof clarsimp fix elem assume assum: "elem \ set (signs_smash signs1 signs2)" show "\n m. elem = signs1 ! n @ signs2 ! m" using assum unfolding signs_smash_def apply (auto) by (metis in_set_conv_nth) qed lemma signs_smash_property_set: fixes signs1 signs2 :: "'a list list" fixes a b:: "nat" shows "\ (elem :: 'a list). (elem \ (set (signs_smash signs1 signs2)) \ (\ (elem1::'a list). \ (elem2::'a list). (elem1 \ set(signs1) \ elem2 \ set(signs2) \ elem = (elem1@elem2))))" proof clarsimp fix elem assume assum: "elem \ set (signs_smash signs1 signs2)" show "\elem1. elem1 \ set signs1 \ (\elem2. elem2 \ set signs2 \ elem = elem1 @ elem2)" using assum unfolding signs_smash_def by auto qed lemma subsets_smash_property: fixes subsets1 subsets2 :: "nat list list" fixes n:: "nat" shows "\ (elem :: nat list). (List.member (subsets_smash n subsets1 subsets2) elem) \ (\ (elem1::nat list). \ (elem2::nat list). (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = (elem1 @ ((map ((+) n) elem2)))))" proof - let ?new_subsets = "(map (map ((+) n)) subsets2)" (* a slightly unorthodox use of signs_smash, but it closes the proof *) have "subsets_smash n subsets1 subsets2 = signs_smash subsets1 ?new_subsets" unfolding subsets_smash_def signs_smash_def apply (auto) by (simp add: comp_def) then show ?thesis by (smt imageE in_set_member set_map signs_smash_property_set) qed (* If subsets for smaller systems are well-defined, then subsets for combined systems are well-defined *) subsection "Well-defined subsets preserved when smashing" lemma list_constr_append: "list_constr a n \ list_constr b n \ list_constr (a@b) n" apply (auto) unfolding list_constr_def using list_all_append by blast lemma well_def_step: fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" assumes well_def_subsets1: "all_list_constr (subsets1) (length qs1)" assumes well_def_subsets2: "all_list_constr (subsets2) (length qs2)" shows "all_list_constr ((subsets_smash (length qs1) subsets1 subsets2)) (length (qs1 @ qs2))" proof - have h1: "\elem. List.member (subsets_smash (length qs1) subsets1 subsets2) elem \ (\elem1 elem2. elem1 \ set subsets1 \ elem2 \ set subsets2 \ elem = elem1 @ map ((+) (length qs1)) elem2)" using subsets_smash_property by blast have h2: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr (elem1 @ map ((+) (length qs1)) elem2) (length (qs1 @ qs2))" proof clarsimp fix elem1 fix elem2 assume e1: "elem1 \ set subsets1" assume e2: "elem2 \ set subsets2" have h1: "list_constr elem1 (length qs1 + length qs2) " proof - have h1: "list_constr elem1 (length qs1)" using e1 well_def_subsets1 unfolding all_list_constr_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list.pred_mono_strong) qed have h2: "list_constr (map ((+) (length qs1)) elem2) (length qs1 + length qs2)" proof - have h1: "list_constr elem2 (length qs2)" using e2 well_def_subsets2 unfolding all_list_constr_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list_all_length) qed show "list_constr (elem1 @ map ((+) (length qs1)) elem2) (length qs1 + length qs2)" using h1 h2 list_constr_append by blast qed then show ?thesis using h1 unfolding all_list_constr_def by auto qed subsection "Well def signs preserved when smashing" lemma well_def_signs_step: fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes "length qs1 > 0" assumes "length qs2 > 0" assumes well_def1: "well_def_signs (length qs1) signs1" assumes well_def2: "well_def_signs (length qs2) signs2" shows "well_def_signs (length (qs1@qs2)) (signs_smash signs1 signs2)" using well_def1 well_def2 unfolding well_def_signs_def using signs_smash_property_set[of signs1 signs2] length_append by auto subsection "Distinct signs preserved when smashing" lemma distinct_map_append: assumes "distinct ls" shows "distinct (map ((@) xs) ls)" unfolding distinct_map inj_on_def using assms by auto lemma length_eq_append: assumes "length y = length b" shows "(x @ y = a @ b) \ x=a \ y = b" by (simp add: assms) lemma distinct_step: fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes well_def1: "well_def_signs n1 signs1" assumes well_def2: "well_def_signs n2 signs2" assumes distinct1: "distinct signs1" assumes distinct2: "distinct signs2" shows "distinct (signs_smash signs1 signs2)" unfolding signs_smash_def using distinct1 proof(induction signs1) case Nil then show ?case by auto next case (Cons a signs1) then show ?case apply (auto simp add: distinct2 distinct_map_append) using assms unfolding well_def_signs_def by (simp add: distinct1 distinct2 length_eq_append) qed (* In this section we will show that if signs1 contains all consistent sign assignments and signs2 contains all consistent sign assignments, then their smash contains all consistent sign assignments. Intuitively, this makes sense because signs1 and signs2 are carrying information about different sets of polynomials, and their smash contains all possible combinations of that information. *) subsection "Consistent sign assignments preserved when smashing" lemma subset_smash_signs: fixes a1 b1 a2 b2:: "rat list list" assumes sub1: "set a1 \ set a2" assumes sub2: "set b1 \ set b2" shows "set (signs_smash a1 b1) \ set (signs_smash a2 b2)" proof - have h1: "\x\set (signs_smash a1 b1). x\set (signs_smash a2 b2)" proof clarsimp fix x assume x_in: "x \ set (signs_smash a1 b1)" have h1: "\ n m :: nat. x = (nth a1 n)@(nth b1 m)" using signs_smash_property[of a1 b1] x_in by blast then have h2: "\ p q::nat. x = (nth a2 p)@(nth b2 q)" using sub1 sub2 apply (auto) by (metis nth_find_first signs_smash_property_set subset_code(1) x_in) then show "x \ set (signs_smash a2 b2)" unfolding signs_smash_def apply (auto) using signs_smash_property_set sub1 sub2 x_in by fastforce qed then show ?thesis by blast qed lemma subset_helper: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" shows "\ x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots_copr p qs1). \ x2 \ set (characterize_consistent_signs_at_roots_copr p qs2). x = x1@x2" proof clarsimp fix x assume x_in: "x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2))" have x_in_csv: "x \ set(map (consistent_sign_vec_copr (qs1 @ qs2)) (characterize_root_list_p p))" using x_in unfolding characterize_consistent_signs_at_roots_copr_def by simp have csv_hyp: "\r. consistent_sign_vec_copr (qs1 @ qs2) r = (consistent_sign_vec_copr qs1 r)@(consistent_sign_vec_copr qs2 r)" unfolding consistent_sign_vec_copr_def by auto have exr_iff: "(\r \ set (characterize_root_list_p p). x = consistent_sign_vec_copr (qs1 @ qs2) r) \ (\r \ set (characterize_root_list_p p). x = (consistent_sign_vec_copr qs1 r)@(consistent_sign_vec_copr qs2 r))" using csv_hyp by auto have exr: "x \ set(map (consistent_sign_vec_copr (qs1 @ qs2)) (characterize_root_list_p p)) \ (\r \ set (characterize_root_list_p p). x = consistent_sign_vec_copr (qs1 @ qs2) r)" by auto have mem_list1: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec_copr qs1 r) \ set(map (consistent_sign_vec_copr qs1) (characterize_root_list_p p))" by simp have mem_list2: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec_copr qs2 r) \ set(map (consistent_sign_vec_copr qs2) (characterize_root_list_p p))" by simp then show "\x1\set (characterize_consistent_signs_at_roots_copr p qs1). \x2\set (characterize_consistent_signs_at_roots_copr p qs2). x = x1 @ x2" using x_in_csv exr mem_list1 mem_list2 characterize_consistent_signs_at_roots_copr_def exr_iff by auto qed lemma subset_step: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes csa1: "set (characterize_consistent_signs_at_roots_copr p qs1) \ set (signs1)" assumes csa2: "set (characterize_consistent_signs_at_roots_copr p qs2) \ set (signs2)" shows "set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)) \ set (signs_smash signs1 signs2)" proof - have h0: "\ x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots_copr p qs1). \ x2 \ set (characterize_consistent_signs_at_roots_copr p qs2). (x = x1@x2)" using subset_helper[of p qs1 qs2] by blast then have "\x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). x \ set (signs_smash (characterize_consistent_signs_at_roots_copr p qs1) (characterize_consistent_signs_at_roots_copr p qs2))" unfolding signs_smash_def apply (auto) by fastforce then have "\x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). x \ set (signs_smash signs1 signs2)" using csa1 csa2 subset_smash_signs[of _ signs1 _ signs2] apply (auto) by blast thus ?thesis by blast qed subsection "Main Results" lemma dim_row_mat_of_rows_list[simp]: shows "dim_row (mat_of_rows_list nr ls) = length ls" unfolding mat_of_rows_list_def by auto lemma dim_col_mat_of_rows_list[simp]: shows "dim_col (mat_of_rows_list nr ls) = nr" unfolding mat_of_rows_list_def by auto lemma dim_row_matrix_A[simp]: shows "dim_row (matrix_A signs subsets) = length subsets" unfolding matrix_A_def by auto lemma dim_col_matrix_A[simp]: shows "dim_col (matrix_A signs subsets) = length signs" unfolding matrix_A_def by auto lemma length_signs_smash: shows "length (signs_smash signs1 signs2) = length signs1 * length signs2" unfolding signs_smash_def length_concat by (auto simp add: o_def sum_list_triv) lemma length_subsets_smash: shows "length (subsets_smash n subs1 subs2) = length subs1 * length subs2" unfolding subsets_smash_def length_concat by (auto simp add: o_def sum_list_triv) lemma length_eq_concat: assumes "\x. x \ set ls \ length x = n" assumes "i < n * length ls" shows "concat ls ! i = ls ! (i div n) ! (i mod n)" using assms proof (induct ls arbitrary: i) case Nil then show ?case by simp next case (Cons a ls) then show ?case apply (auto simp add: nth_append) using div_if mod_geq by auto qed lemma z_append: assumes "\i. i \ set xs \ i < length as" shows "z (xs @ (map ((+) (length as)) ys)) (as @ bs) = z xs as * z ys bs" proof - have 1: "map ((!) (as @ bs)) xs = map ((!) as) xs" unfolding list_eq_iff_nth_eq using assms by (auto simp add:nth_append) have 2: "map ((!) (as @ bs) \ (+) (length as)) ys = map ((!) bs) ys" unfolding list_eq_iff_nth_eq using assms by auto show ?thesis unfolding z_def apply auto unfolding 1 2 by auto qed (* Shows that the matrix of a smashed system is the Kronecker product of the matrices of the two systems being combined *) lemma matrix_construction_is_kronecker_product: fixes qs1 :: "real poly list" fixes subs1 subs2 :: "nat list list" fixes signs1 signs2 :: "rat list list" (* n1 is the number of polynomials in the "1" sets *) assumes "\l i. l \ set subs1 \ i \ set l \ i < n1" assumes "\j. j \ set signs1 \ length j = n1" shows " (matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2)) = kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2)" unfolding mat_eq_iff dim_row_matrix_A dim_col_matrix_A length_subsets_smash length_signs_smash dim_kronecker proof safe fix i j assume i: "i < length subs1 * length subs2" and j: "j < length signs1 * length signs2" have ld: "i div length subs2 < length subs1" "j div length signs2 < length signs1" using i j less_mult_imp_div_less by auto have lm: "i mod length subs2 < length subs2" "j mod length signs2 < length signs2" using i j less_mult_imp_mod_less by auto have n1: "n1 = length (signs1 ! (j div length signs2))" using assms(2) ld(2) nth_mem by blast have 1: "matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2) $$ (i, j) = z (subsets_smash n1 subs1 subs2 ! i) (signs_smash signs1 signs2 ! j)" unfolding mat_of_rows_list_def matrix_A_def mtx_row_def using i j by (simp add: length_signs_smash length_subsets_smash) have 2: " ... = z (subs1 ! (i div length subs2) @ map ((+) n1) (subs2 ! (i mod length subs2))) (signs1 ! (j div length signs2) @ signs2 ! (j mod length signs2))" unfolding signs_smash_def subsets_smash_def apply (subst length_eq_concat) using i apply (auto simp add: mult.commute) apply (subst length_eq_concat) using j apply (auto simp add: mult.commute) using ld lm by auto have 3: "... = z (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding n1 apply (subst z_append) apply (auto simp add: n1[symmetric]) using assms(1) ld(1) nth_mem by blast have 4: "kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2) $$ (i,j) = z (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding kronecker_product_def matrix_A_def mat_of_rows_list_def mtx_row_def using i j apply (auto simp add: Let_def) apply (subst index_mat(1)[OF ld]) apply (subst index_mat(1)[OF lm]) using ld lm by auto show "matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2) $$ (i, j) = kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2) $$ (i, j)" using 1 2 3 4 by auto qed (* Given that two smaller systems satisfy some mild constraints, show that their combined system (after smashing the signs and subsets) satisfies the matrix equation, and that the matrix of the combined system is invertible. *) lemma inductive_step: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" fixes signs1 signs2 :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes welldefined_signs1: "well_def_signs (length qs1) signs1" assumes welldefined_signs2: "well_def_signs (length qs2) signs2" assumes distinct_signs1: "distinct signs1" assumes distinct_signs2: "distinct signs2" assumes all_info1: "set (characterize_consistent_signs_at_roots_copr p qs1) \ set(signs1)" assumes all_info2: "set (characterize_consistent_signs_at_roots_copr p qs2) \ set(signs2)" assumes welldefined_subsets1: "all_list_constr (subsets1) (length qs1)" assumes welldefined_subsets2: "all_list_constr (subsets2) (length qs2)" assumes invertibleMtx1: "invertible_mat (matrix_A signs1 subsets1)" assumes invertibleMtx2: "invertible_mat (matrix_A signs2 subsets2)" shows "satisfy_equation p (qs1@qs2) (subsets_smash (length qs1) subsets1 subsets2) (signs_smash signs1 signs2) \ invertible_mat (matrix_A (signs_smash signs1 signs2) (subsets_smash (length qs1) subsets1 subsets2))" proof - have h1: "invertible_mat (matrix_A (signs_smash signs1 signs2) (subsets_smash (length qs1) subsets1 subsets2))" using matrix_construction_is_kronecker_product kronecker_invertible invertibleMtx1 invertibleMtx2 welldefined_subsets1 welldefined_subsets2 unfolding all_list_constr_def list_constr_def by (smt in_set_conv_nth in_set_member list_all_length well_def_signs_def welldefined_signs1) have h2a: "distinct (signs_smash signs1 signs2)" using distinct_signs1 distinct_signs2 welldefined_signs1 welldefined_signs2 nontriv1 nontriv2 distinct_step by auto have h2ba: "\ q. List.member (qs1 @ qs2) q \ (List.member qs1 q \ List.member qs2 q)" unfolding member_def by auto have h2b: "\q. ((List.member (qs1@qs2) q) \ (coprime p q))" using h2ba pairwise_rel_prime1 pairwise_rel_prime2 by auto have h2c: "all_list_constr ((subsets_smash (length qs1) subsets1 subsets2)) (length (qs1@qs2))" using well_def_step welldefined_subsets1 welldefined_subsets2 by blast have h2d: "set (characterize_consistent_signs_at_roots_copr p (qs1@qs2)) \ set(signs_smash signs1 signs2)" using subset_step all_info1 all_info2 by simp have h2: "satisfy_equation p (qs1@qs2) (subsets_smash (length qs1) subsets1 subsets2) (signs_smash signs1 signs2)" using matrix_equation[where p="p", where qs="qs1@qs2", where subsets = "subsets_smash (length qs1) subsets1 subsets2", where signs = "signs_smash signs1 signs2"] h2a h2b h2c h2d apply (auto) using nonzero by blast show ?thesis using h1 h2 by blast qed section "Reduction Step Proofs" (* Some definitions *) definition get_matrix:: "(rat mat \ (nat list list \ rat list list)) \ rat mat" where "get_matrix data = fst(data)" definition get_subsets:: "(rat mat \ (nat list list \ rat list list)) \ nat list list" where "get_subsets data = fst(snd(data))" definition get_signs:: "(rat mat \ (nat list list \ rat list list)) \ rat list list" where "get_signs data = snd(snd(data))" definition reduction_signs:: "real poly \ real poly list \ rat list list \ nat list list \ rat mat \ rat list list" where "reduction_signs p qs signs subsets matr = (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets matr)))" definition reduction_subsets:: "real poly \ real poly list \ rat list list \ nat list list \ rat mat \ nat list list" where "reduction_subsets p qs signs subsets matr = (take_indices subsets (rows_to_keep (reduce_mat_cols matr (solve_for_lhs p qs subsets matr))))" (* Some basic lemmas *) lemma reduction_signs_is_get_signs: "reduction_signs p qs signs subsets m = get_signs (reduce_system p (qs, (m, (subsets, signs))))" unfolding reduction_signs_def get_signs_def by (metis reduce_system.simps reduction_step.elims snd_conv) lemma reduction_subsets_is_get_subsets: "reduction_subsets p qs signs subsets m = get_subsets (reduce_system p (qs, (m, (subsets, signs))))" unfolding reduction_subsets_def get_subsets_def by (metis fst_conv reduce_system.simps reduction_step.elims snd_conv) lemma find_zeros_from_vec_prop: fixes input_vec :: "rat vec" shows "\n < (dim_vec input_vec). ((input_vec $ n \ 0) \ List.member (find_nonzeros_from_input_vec input_vec) n)" proof - have h1: "\n < (dim_vec input_vec). ((input_vec $ n \ 0) \ List.member (find_nonzeros_from_input_vec input_vec) n)" unfolding find_nonzeros_from_input_vec_def List.member_def by auto have h2: "\n < (dim_vec input_vec). (List.member (find_nonzeros_from_input_vec input_vec) n \ (input_vec $ n \ 0) )" unfolding find_nonzeros_from_input_vec_def List.member_def by auto show ?thesis using h1 h2 by auto qed subsection "Showing sign conditions preserved when reducing" lemma take_indices_lem: fixes p:: "real poly" fixes qs :: "real poly list" fixes arb_list :: "'a list list" fixes index_list :: "nat list" fixes n:: "nat" assumes lest: "n < length (take_indices arb_list index_list)" assumes well_def: "\q.(List.member index_list q \ q < length arb_list)" shows "\km (dim_row A)" proof (cases "mat_inverse(A)") obtain n where n: "A \ carrier_mat n n" using assms invertible_mat_def square_mat.simps by blast case None then have "A \ Units (ring_mat TYPE('a) n n)" by (simp add: mat_inverse(1) n) thus ?thesis by (meson assms det_non_zero_imp_unit invertible_Units n unit_imp_det_non_zero) next case (Some a) then show ?thesis by (metis assms carrier_matI invertible_mat_def mat_inverse(2) matr_option.simps(2) square_mat.elims(2)) qed lemma dim_invertible: fixes A:: "'a::field mat" fixes n:: "nat" assumes assm: "invertible_mat A" assumes dim: "A \ carrier_mat n n" shows "matr_option (dim_row A) (mat_inverse (A)) \ carrier_mat n n" proof (cases "mat_inverse(A)") obtain n where n: "A \ carrier_mat n n" using assms invertible_mat_def square_mat.simps by blast case None then have "A \ Units (ring_mat TYPE('a) n n)" by (simp add: mat_inverse(1) n) thus ?thesis by (meson assms det_non_zero_imp_unit invertible_Units n unit_imp_det_non_zero) next case (Some a) then show ?thesis using dim mat_inverse(2) by auto qed lemma mult_assoc: fixes A B:: "rat mat" fixes v:: "rat vec" fixes n:: "nat" assumes "A \ carrier_mat n n" assumes "B \ carrier_mat n n" assumes "dim_vec v = n" shows "A *\<^sub>v (mult_mat_vec B v) = (A*B)*\<^sub>v v" using assms(1) assms(2) assms(3) by auto lemma size_of_mat: fixes subsets :: "nat list list" fixes signs :: "rat list list" shows "(matrix_A signs subsets) \ carrier_mat (length subsets) (length signs)" unfolding matrix_A_def carrier_mat_def by auto lemma size_of_lhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes signs :: "rat list list" shows "dim_vec (construct_lhs_vector p qs signs) = length signs" unfolding construct_lhs_vector_def by simp lemma size_of_rhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" shows "dim_vec (construct_rhs_vector p qs subsets) = length subsets" unfolding construct_rhs_vector_def by simp lemma same_size: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "length subsets = length signs" using invertible_mat unfolding invertible_mat_def using size_of_mat[of signs subsets] size_of_lhs[of p qs signs] size_of_rhs[of p qs subsets] by simp lemma mat_equal_list_lem: fixes A:: "'a::field mat" fixes B:: "'a::field mat" shows "(dim_row A = dim_row B \ dim_col A = dim_col B \ mat_to_list A = mat_to_list B) \ A = B" proof - assume hyp: "dim_row A = dim_row B \ dim_col A = dim_col B \ mat_to_list A = mat_to_list B" then have "\i j. i < dim_row A \ j < dim_col A \ B $$ (i, j) = A $$ (i, j)" unfolding mat_to_list_def by auto then show ?thesis using hyp unfolding mat_to_list_def using eq_matI[of A B] by metis qed lemma mat_inverse_same: "mat_inverse_var A = mat_inverse A" unfolding mat_inverse_var_def mat_inverse_def mat_equal_def using mat_equal_list_lem apply (simp) by (smt case_prod_beta index_one_mat(2) index_one_mat(3) mat_equal_list_lem) lemma construct_lhs_matches_solve_for_lhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" proof - have matrix_equation_hyp: "(mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs) = (construct_rhs_vector p qs subsets))" using match unfolding satisfy_equation_def by blast then have eqn_hyp: " ((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs)) = mult_mat_vec (matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) (construct_rhs_vector p qs subsets))" using invertible_mat by (simp add: matrix_equation_hyp) have match_hyp: "length subsets = length signs" using same_size invertible_mat by auto then have dim_hyp1: "matrix_A signs subsets \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible by blast have mult_assoc_hyp: "((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs))) = (((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs))" using mult_assoc dim_hyp1 dim_hyp2 size_of_lhs by auto have cancel_helper: "(((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs)) = (1\<^sub>m (length signs)) *\<^sub>v (construct_lhs_vector p qs signs)" using invertible_means_mult_id[where A= "matrix_A signs subsets"] dim_hyp1 by (simp add: invertible_mat match_hyp) then have cancel_hyp: "(((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs)) = (construct_lhs_vector p qs signs)" apply (auto) by (metis carrier_vec_dim_vec one_mult_mat_vec size_of_lhs) then have mult_hyp: "((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs))) = (construct_lhs_vector p qs signs)" using mult_assoc_hyp cancel_hyp by simp then have "(construct_lhs_vector p qs signs) = (mult_mat_vec (matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) (construct_rhs_vector p qs subsets)) " using eqn_hyp by simp then show ?thesis unfolding solve_for_lhs_def by (simp add: mat_inverse_same) qed (* If set(A) is a subset of B then for all n, nth c n = 0 means nth a n not in B *) lemma reduction_signs_set_helper_lemma: fixes A:: "rat list set" fixes C:: "rat vec" fixes B:: "rat list list" assumes "dim_vec C = length B" assumes sub: "A \ set(B)" assumes not_in_hyp: "\ n < dim_vec C. C $ n = 0 \ (nth B n) \ A" shows "A \ set (take_indices B (find_nonzeros_from_input_vec C))" proof - have unfold: "\ x. (x \ A \ x \ set (take_indices B (find_nonzeros_from_input_vec C)))" proof - fix x assume in_a: "x \ A" have "x \ set (B)" using sub in_a by blast then have "\ n < length B. nth B n = x" by (simp add: in_set_conv_nth) then have "\ n < length B. (nth B n = x \ (List.member (find_nonzeros_from_input_vec C) n) = True)" using not_in_hyp find_zeros_from_vec_prop[of C] using assms(1) in_a by auto thus "x \ set (take_indices B (find_nonzeros_from_input_vec C))" unfolding take_indices_def using member_def by fastforce qed then show ?thesis by blast qed lemma reduction_signs_set_helper_lemma2: fixes A:: "rat list set" fixes C:: "rat vec" fixes B:: "rat list list" assumes dist: "distinct B" assumes eq_len: "dim_vec C = length B" assumes sub: "A \ set(B)" assumes not_in_hyp: "\ n < dim_vec C. C $ n \ 0 \ (nth B n) \ A" shows "set (take_indices B (find_nonzeros_from_input_vec C)) \ A" proof - have unfold: "\ x. (x \ (set (take_indices B (find_nonzeros_from_input_vec C))) \ x \ A)" proof - fix x assume in_set: "x \ set (take_indices B (find_nonzeros_from_input_vec C))" have h: "\n< dim_vec C. (C $ n \ 0 \ (nth B n) = x)" proof - have h1: "\n < length B.(nth B n) = x" using in_set unfolding take_indices_def find_nonzeros_from_input_vec_def eq_len by auto have h2: "\n< length B. (nth B n = x \ List.member (find_nonzeros_from_input_vec C) n)" proof clarsimp fix n assume leq_hyp: "n < length B" assume x_eq: "x = B ! n" have h: "(B !n) \ set (map ((!) B) (find_nonzeros_from_input_vec C))" using x_eq in_set by (simp add: take_indices_def) then have h2: "List.member (map ((!) B) (find_nonzeros_from_input_vec C)) (B ! n)" using in_set by (meson in_set_member) then have h3: "\k< length B. (List.member (find_nonzeros_from_input_vec C) k \ ((B ! k) = (B ! n)))" by (smt atLeastLessThan_iff eq_len find_nonzeros_from_input_vec_def imageE in_set_member mem_Collect_eq set_filter set_map set_upt) have h4: "\v< length B. ((B ! v) = (B ! n) \ v = n)" using dist apply (auto) using leq_hyp nth_eq_iff_index_eq by blast then show "List.member (find_nonzeros_from_input_vec C) n" using h2 h3 h4 by auto qed then have "\n C $ n \ 0)" using find_zeros_from_vec_prop [of C] by (simp add: eq_len) then show ?thesis using h1 eq_len by auto qed thus "x \ A" using not_in_hyp by blast qed then show ?thesis by blast qed (* Show that dropping columns doesn't affect the consistent sign assignments *) lemma reduction_doesnt_break_things_signs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size by fastforce have "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by auto then have "(solve_for_lhs p qs subsets (matrix_A signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))" using construct_lhs_vector_cleaner assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots_copr p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)\ construct_lhs_vector_clean nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ (\ x. poly p x = 0 \ consistent_sign_vec_copr qs x = w))" proof clarsimp fix x assume card_asm: "card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x} = 0" assume zero_asm: "poly p x = 0" have card_hyp: "{xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x} = {}" using card_eq_0_iff using card_asm nonzero poly_roots_finite by fastforce have "x \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x}" using zero_asm by auto thus "False" using card_hyp by blast qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ (\List.member (characterize_consistent_signs_at_roots_copr p qs) w))" by (smt (verit, best) characterize_consistent_signs_at_roots_copr_def characterize_root_list_p_def h1 imageE in_set_member mem_Collect_eq nonzero poly_roots_finite set_map set_remdups sorted_list_of_set(1)) then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ w \ set (characterize_consistent_signs_at_roots_copr p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by blast qed then have "set (characterize_consistent_signs_at_roots_copr p qs) \ set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" using all_info reduction_signs_set_helper_lemma[where A = "set (characterize_consistent_signs_at_roots_copr p qs)", where B = "signs", where C = "(solve_for_lhs p qs subsets (matrix_A signs subsets))"] using dim_hyp2 solve_for_lhs_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_def by auto qed lemma reduction_deletes_bad_sign_conds: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (characterize_consistent_signs_at_roots_copr p qs) = set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size by fastforce have supset: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by auto then have "(solve_for_lhs p qs subsets (matrix_A signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))" using construct_lhs_vector_cleaner assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n \ 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots_copr p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)\ construct_lhs_vector_clean nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ (\ x. poly p x = 0 \ consistent_sign_vec_copr qs x = w))" proof clarsimp fix w assume card_asm: "0 < card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}" show "\x. poly p x = 0 \ consistent_sign_vec_copr qs x = w" by (metis (mono_tags, lifting) Collect_empty_eq card_asm card_eq_0_iff gr_implies_not0) qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ (List.member (characterize_consistent_signs_at_roots_copr p qs) w))" proof clarsimp fix w assume card_asm: " 0 < card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}" have h0: "\x. poly p x = 0 \ consistent_sign_vec_copr qs x = w" using card_asm by (simp add: h1) then show "List.member (characterize_consistent_signs_at_roots_copr p qs) w" unfolding characterize_consistent_signs_at_roots_copr_def using in_set_member nonzero poly_roots_finite characterize_root_list_p_def by fastforce qed then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ w \ set (characterize_consistent_signs_at_roots_copr p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by (metis (no_types, lifting) \solve_for_lhs p qs subsets (matrix_A signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))\ dim_vec_of_list length_map nth_map vec_of_list_index) qed then have "set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (characterize_consistent_signs_at_roots_copr p qs)" using all_info reduction_signs_set_helper_lemma2[where A = "set (characterize_consistent_signs_at_roots_copr p qs)", where B = "signs", where C = "(solve_for_lhs p qs subsets (matrix_A signs subsets))"] using distinct_signs dim_hyp2 solve_for_lhs_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_def by auto qed have subset: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" using reduction_doesnt_break_things_signs[of p qs signs subsets] assms by blast then show ?thesis using supset subset by auto qed theorem reduce_system_sign_conditions: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding get_signs_def using reduction_deletes_bad_sign_conds[of p qs signs subsets] apply (auto) apply (simp add: all_info distinct_signs match nonzero reduction_signs_def welldefined_signs1) using nonzero invertible_mat apply (metis snd_conv) by (metis all_info distinct_signs invertible_mat match nonzero reduction_signs_def snd_conv welldefined_signs1) subsection "Showing matrix equation preserved when reducing" lemma rows_to_keep_lem: fixes A:: "('a::field) mat" shows "\ell. ell \ set (rows_to_keep A) \ ell < dim_row A" unfolding rows_to_keep_def apply auto using rref_pivot_positions by (metis carrier_mat_triv gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(3)) lemma reduce_system_matrix_equation_preserved: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs: "well_def_signs (length qs) signs" assumes welldefined_subsets: "all_list_constr (subsets) (length qs)" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "satisfy_equation p qs (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - have poly_type_hyp: "p \ 0" using nonzero by auto have distinct_signs_hyp: "distinct (snd (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))))" proof - let ?sym = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" have h1: "\ i < length (take_indices signs ?sym). \j < length(take_indices signs ?sym). i \ j \ nth (take_indices signs ?sym) i \ nth (take_indices signs ?sym) j" using distinct_signs unfolding take_indices_def proof clarsimp fix i fix j assume "distinct signs" assume "i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume "j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume neq_hyp: "i \ j" assume "signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j)" have h1: "find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i \ find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j" unfolding find_nonzeros_from_input_vec_def using neq_hyp by (metis \i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ \j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ distinct_conv_nth distinct_filter distinct_upt find_nonzeros_from_input_vec_def) then show "False" using distinct_signs proof - have f1: "\p ns n. ((n::nat) \ {n \ set ns. p n}) = (n \ set ns \ n \ Collect p)" by simp then have f2: "filter (\n. solve_for_lhs p qs subsets (matrix_A signs subsets) $ n \ 0) [0.. set [0..i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ construct_lhs_matches_solve_for_lhs find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs) have "filter (\n. solve_for_lhs p qs subsets (matrix_A signs subsets) $ n \ 0) [0.. set [0..j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ construct_lhs_matches_solve_for_lhs find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs) then show ?thesis using f2 by (metis \signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j)\ atLeastLessThan_iff distinct_conv_nth distinct_signs find_nonzeros_from_input_vec_def h1 set_upt) qed qed then have "distinct (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" using distinct_conv_nth by blast then show ?thesis using get_signs_def reduction_signs_def reduction_signs_is_get_signs by auto qed have all_info_hyp: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(snd (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))))" using reduce_system_sign_conditions assms unfolding get_signs_def by auto have pairwise_rel_prime_hyp: "\q. ((List.member qs q) \ (coprime p q))" using pairwise_rel_prime by auto have welldefined_hyp: "all_list_constr (fst (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))) (length qs)" using welldefined_subsets rows_to_keep_lem unfolding all_list_constr_def List.member_def list_constr_def list_all_def apply (auto simp add: Let_def take_indices_def take_cols_from_matrix_def) using nth_mem by fastforce then show ?thesis using poly_type_hyp distinct_signs_hyp all_info_hyp pairwise_rel_prime_hyp welldefined_hyp using matrix_equation unfolding get_subsets_def get_signs_def by blast qed (* Show that we are tracking the correct matrix in the algorithm *) subsection "Showing matrix preserved" lemma reduce_system_matrix_signs_helper_aux: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "alt_matrix_A (take_indices signs S) subsets = take_cols_from_matrix (alt_matrix_A signs subsets) S" proof - have h0a: "dim_col (take_cols_from_matrix (alt_matrix_A signs subsets) S) = length (take_indices signs S)" unfolding take_cols_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices signs S). (col (alt_matrix_A (take_indices signs S) subsets ) i = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices signs S)" have i_lt: "i < length (map ((!) (cols (alt_matrix_A signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: " vec (length subsets) (\j. z (subsets ! j) (map ((!) signs) S ! i)) = vec (length subsets) (\j. z (subsets ! j) (signs ! (S ! i)))" using nth_map by (metis \i < length (take_indices signs S)\ length_map take_indices_def) have dim: "(map ((!) (cols (alt_matrix_A signs subsets))) S) ! i \ carrier_vec (dim_row (alt_matrix_A signs subsets))" proof - have "dim_col (alt_matrix_A signs subsets) = length (signs)" by (simp add: alt_matrix_A_def) have well_d: "S ! i < length (signs)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (cols (alt_matrix_A signs subsets))) S) ! i = nth (cols (alt_matrix_A signs subsets)) (S ! i)" using i_lt by auto have "nth (cols (alt_matrix_A signs subsets)) (S ! i) \ carrier_vec (dim_row (alt_matrix_A signs subsets))" using col_dim unfolding cols_def using nth_map well_d by (simp add: \dim_col (alt_matrix_A signs subsets) = length signs\) then show ?thesis using map_eq unfolding alt_matrix_A_def by auto qed have h1: "col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i = col (mat_of_cols (dim_row (alt_matrix_A signs subsets)) (map ((!) (cols (alt_matrix_A signs subsets))) S)) i " by (simp add: take_cols_from_matrix_def take_indices_def) have h2: "col (mat_of_cols (dim_row (alt_matrix_A signs subsets)) (map ((!) (cols (alt_matrix_A signs subsets))) S)) i = nth (map ((!) (cols (alt_matrix_A signs subsets))) S) i " using dim i_lt asm col_mat_of_cols[where j = "i", where n = "(dim_row (alt_matrix_A signs subsets))", where vs = "(map ((!) (cols (alt_matrix_A signs subsets))) S)"] by blast have h3: "col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i = (col (alt_matrix_A signs subsets) (S !i))" using h1 h2 apply (auto) by (metis alt_matrix_char asm cols_nth dim_col_mat(1) in_set_member length_map mat_of_rows_list_def matrix_A_def nth_map nth_mem take_indices_def well_def_h) have "vec (length subsets) (\j. z (subsets ! j) (signs ! (S ! i))) = (col (alt_matrix_A signs subsets) (S !i))" by (metis asm in_set_member length_map nth_mem signs_are_cols take_indices_def well_def_h) then have "vec (length subsets) (\j. z (subsets ! j) (take_indices signs S ! i)) = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i " using h0 h3 by (simp add: take_indices_def) then show "col (alt_matrix_A (take_indices signs S) subsets) i = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i " using asm signs_are_cols[where signs = "(take_indices signs S)", where subsets = "subsets"] by auto qed then show ?thesis unfolding alt_matrix_A_def take_cols_from_matrix_def apply (auto) using h0a mat_col_eqI by (metis alt_matrix_A_def dim_col_mat(1) dim_row_mat(1) h0 mat_of_cols_def take_cols_from_matrix_def) qed lemma reduce_system_matrix_signs_helper: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "matrix_A (take_indices signs S) subsets = take_cols_from_matrix (matrix_A signs subsets) S" using reduce_system_matrix_signs_helper_aux alt_matrix_char assms by auto lemma reduce_system_matrix_subsets_helper_aux: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" assumes nonzero: "p \ 0" shows "alt_matrix_A signs (take_indices subsets S) = take_rows_from_matrix (alt_matrix_A signs subsets) S" proof - have h0a: "dim_row (take_rows_from_matrix (alt_matrix_A signs subsets) S) = length (take_indices subsets S)" unfolding take_rows_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices subsets S). (row (alt_matrix_A signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices subsets S)" have i_lt: "i < length (map ((!) (rows (alt_matrix_A signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: "row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i = row (mat_of_rows (dim_col (alt_matrix_A signs subsets)) (map ((!) (rows (alt_matrix_A signs subsets))) S)) i " unfolding take_rows_from_matrix_def take_indices_def by auto have dim: "(map ((!) (rows (alt_matrix_A signs subsets))) S) ! i \ carrier_vec (dim_col (alt_matrix_A signs subsets))" proof - have "dim_col (alt_matrix_A signs subsets) = length (signs)" by (simp add: alt_matrix_A_def) then have lenh: "dim_col (alt_matrix_A signs subsets) \ length (subsets)" using assms by auto have well_d: "S ! i < length (subsets)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (rows (alt_matrix_A signs subsets))) S) ! i = nth (rows (alt_matrix_A signs subsets)) (S ! i)" using i_lt by auto have "nth (rows (alt_matrix_A signs subsets)) (S ! i) \ carrier_vec (dim_col (alt_matrix_A signs subsets))" using col_dim unfolding rows_def using nth_map well_d using lenh by (simp add: alt_matrix_A_def) then show ?thesis using map_eq unfolding alt_matrix_A_def by auto qed have h1: " row (mat_of_rows (dim_col (alt_matrix_A signs subsets)) (map ((!) (rows (alt_matrix_A signs subsets))) S)) i = (row (alt_matrix_A signs subsets) (S ! i)) " using dim i_lt mat_of_rows_row[where i = "i", where n = "(dim_col (alt_matrix_A signs subsets))", where vs = "(map ((!) (rows (alt_matrix_A signs subsets))) S)"] using alt_matrix_char in_set_member nth_mem well_def_h by fastforce have "row (alt_matrix_A signs (take_indices subsets S) ) i = (row (alt_matrix_A signs subsets) (S ! i))" using subsets_are_rows proof - have f1: "i < length S" by (metis (no_types) asm length_map take_indices_def) then have "List.member S (S ! i)" by (meson in_set_member nth_mem) then show ?thesis using f1 by (simp add: \\subsets signs. \ij. z (subsets ! i) (signs ! j))\ take_indices_def well_def_h) qed then show "(row (alt_matrix_A signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i)" using h0 h1 unfolding take_indices_def by auto qed then show ?thesis unfolding alt_matrix_A_def take_rows_from_matrix_def apply (auto) using eq_rowI by (metis alt_matrix_A_def dim_col_mat(1) dim_row_mat(1) h0 length_map mat_of_rows_def take_indices_def take_rows_from_matrix_def) qed lemma reduce_system_matrix_subsets_helper: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes nonzero: "p \ 0" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" shows "matrix_A signs (take_indices subsets S) = take_rows_from_matrix (matrix_A signs subsets) S" using assms reduce_system_matrix_subsets_helper_aux alt_matrix_char by auto lemma reduce_system_matrix_match: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes inv: "invertible_mat (matrix_A signs subsets)" shows "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - let ?A = "(matrix_A signs subsets)" let ?lhs_vec = "(solve_for_lhs p qs subsets (matrix_A signs subsets))" let ?red_mtx = "(take_rows_from_matrix (reduce_mat_cols (matrix_A signs subsets) ?lhs_vec) (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) ?lhs_vec)))" have h1: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = matrix_A (reduction_signs p qs signs subsets (matrix_A signs subsets)) (reduction_subsets p qs signs subsets (matrix_A signs subsets))" using reduction_signs_is_get_signs reduction_subsets_is_get_subsets by auto have h1_var: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = matrix_A (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)))" using h1 reduction_signs_def reduction_subsets_def by auto have h2: "?red_mtx = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" by simp have h30: "(construct_lhs_vector p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs unfolding find_nonzeros_from_input_vec_def apply (auto) by (metis atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq) have h3b_a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (subsets)" using assms h30 size_of_lhs same_size unfolding find_nonzeros_from_input_vec_def apply (auto) by (simp add: find_nonzeros_from_input_vec_def h3a) have h3ba: "length (filter (\i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0) [0.. length subsets" using length_filter_le[where P = "(\i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0)", where xs = "[0..i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0) [0.. length subsets" using h3ba by auto then have h3b: "length subsets \ length (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec))" unfolding take_indices_def find_nonzeros_from_input_vec_def by auto have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) < length signs \ matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using h3a nonzero reduce_system_matrix_signs_helper by auto then have "matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))" using f4 by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def by (metis (no_types) dim_row_matrix_A rows_to_keep_def rows_to_keep_lem) qed have h3: "matrix_A (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec))) = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" using inv h3a h3b h3c reduce_system_matrix_subsets_helper reduce_system_matrix_signs_helper assms by auto show ?thesis using h1 h2 h3 by (metis fst_conv get_matrix_def h1_var reduce_system.simps reduction_step.simps) qed (* gauss_jordan_single_rank is crucial in this section *) subsection "Showing invertibility preserved when reducing" (* Overall: Start with a matrix equation. Input a matrix, subsets, and signs. Drop columns of the matrix based on the 0's on the LHS---so extract a list of 0's. Reduce signs accordingly. Then find a list of rows to delete based on using rank (use the transpose result, pivot positions!), and delete those rows. Reduce subsets accordingly. End with a reduced system! *) lemma well_def_find_zeros_from_lhs_vec: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" proof - fix i fix j assume j_in: " j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) " let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) = (length signs)" using size_of_lhs construct_lhs_matches_solve_for_lhs assms by (metis (full_types)) then have h: "j < (length signs)" using j_in unfolding find_nonzeros_from_input_vec_def by simp then show "j < length (cols (matrix_A signs subsets))" using mat_size by auto qed lemma take_cols_subsets_og_cols: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (cols (matrix_A signs subsets))" proof - have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" using assms well_def_find_zeros_from_lhs_vec by auto have "\x. x \ set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ x \ set (cols (matrix_A signs subsets))" proof clarsimp fix x let ?og_list = "(cols (matrix_A signs subsets))" let ?ind_list = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume x_in: "x \ set (take_indices ?og_list ?ind_list)" show "x \ set (cols (matrix_A signs subsets))" using x_in unfolding take_indices_def using in_set_member apply (auto) using in_set_conv_nth well_def by fastforce qed then show ?thesis by blast qed lemma reduction_doesnt_break_things_invertibility_step1: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "vec_space.rank (length signs) (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets))) = (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" proof - let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have mat_size_alt: "?og_mat \ carrier_mat (length subsets) (length subsets)" using size_of_mat same_size assms by auto have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A signs subsets"] mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff mat_size by auto then have dist_cols: "distinct (cols ?og_mat)" using mat_size vec_space.non_distinct_low_rank[where A = ?og_mat, where n = "length signs"] by auto have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" using assms well_def_find_zeros_from_lhs_vec by auto have dist1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" unfolding find_nonzeros_from_input_vec_def by auto have clear: "set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (cols (matrix_A signs subsets))" using assms take_cols_subsets_og_cols by auto then have "distinct (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" unfolding take_indices_def using dist1 dist_cols well_def conjugatable_vec_space.distinct_map_nth[where ls = "cols (matrix_A signs subsets)", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))"] by auto then have unfold_thesis: "vec_space.rank (length signs) (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) (find_nonzeros_from_input_vec ?lhs))) = (length (find_nonzeros_from_input_vec ?lhs))" using clear inv conjugatable_vec_space.rank_invertible_subset_cols[where A= "matrix_A signs subsets", where B = "(take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))"] by (simp add: len_eq mat_size take_indices_def) then show ?thesis apply (simp) unfolding take_cols_from_matrix_def by auto qed lemma rechar_take_cols: "take_cols_var A B = take_cols_from_matrix A B" unfolding take_cols_var_def take_cols_from_matrix_def take_indices_def by auto lemma rows_and_cols_transpose: "rows M = cols M\<^sup>T" using row_transpose by simp lemma take_rows_and_take_cols: "(take_rows_from_matrix M r) = (take_cols_from_matrix M\<^sup>T r)\<^sup>T" unfolding take_rows_from_matrix_def take_cols_from_matrix_def using transpose_carrier_mat rows_and_cols_transpose apply (auto) by (simp add: transpose_mat_of_cols) lemma reduction_doesnt_break_things_invertibility: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "invertible_mat (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?step1_mat = "(reduce_mat_cols ?og_mat ?lhs)" let ?new_mat = "take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)" let ?ind_list = "(find_nonzeros_from_input_vec ?lhs)" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have og_mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_col (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list)) = (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_indices_def) then have "mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list) \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: len_eq mat_of_cols_def) then have step1_mat_size: "?step1_mat \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_cols_from_matrix_def) then have "dim_row ?step1_mat \ dim_col ?step1_mat" by (metis carrier_matD(1) carrier_matD(2) construct_lhs_matches_solve_for_lhs diff_zero find_nonzeros_from_input_vec_def inv length_filter_le length_upt match size_of_lhs) then have gt_eq_assm: "dim_col ?step1_mat\<^sup>T \ dim_row ?step1_mat\<^sup>T" by simp have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A signs subsets"] og_mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff og_mat_size by auto have rank_drop_cols: "vec_space.rank (length signs) ?step1_mat = (dim_col ?step1_mat)" using assms reduction_doesnt_break_things_invertibility_step1 step1_mat_size by auto let ?step1_T = "?step1_mat\<^sup>T" have rank_transpose: "vec_space.rank (length signs) ?step1_mat = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using transpose_rank[of ?step1_mat] using step1_mat_size by auto have obv: "?step1_T \ carrier_mat (dim_row ?step1_T) (dim_col ?step1_T)" by auto have should_have_this:"vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using obv gt_eq_assm conjugatable_vec_space.gauss_jordan_single_rank[where A = "?step1_T", where n = "dim_row ?step1_T", where nc = "dim_col ?step1_T"] by (simp add: take_cols_from_matrix_def take_indices_def) then have "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_drop_cols rank_transpose by auto then have with_take_cols_from_matrix: "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols_from_matrix ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_transpose rechar_take_cols conjugatable_vec_space.gjs_and_take_cols_var apply (auto) by (smt conjugatable_vec_space.gjs_and_take_cols_var gt_eq_assm obv rechar_take_cols reduce_mat_cols.simps) have "(take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)) = (take_cols_from_matrix ?step1_T (rows_to_keep ?step1_mat))\<^sup>T" using take_rows_and_take_cols by blast then have rank_new_mat: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_col ?step1_mat" using with_take_cols_from_matrix transpose_rank apply (auto) proof - assume a1: "vec_space.rank (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))) = dim_col (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" have f2: "\m. vec_space.rank (dim_row (m::rat mat)) m = vec_space.rank (dim_row m\<^sup>T) m\<^sup>T" by (simp add: transpose_rank) have f3: "dim_row (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using \dim_col (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ by force have "vec_space.rank (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T))))) = dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T" using a1 by (simp add: take_cols_from_matrix_def) then have "vec_space.rank (dim_row (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T)))))\<^sup>T) (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T)))))\<^sup>T = dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T" using f3 f2 by presburger then show "vec_space.rank (dim_col (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))))) (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))))\<^sup>T = dim_col (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" by (simp add: rows_to_keep_def take_cols_from_matrix_def) qed have "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ (length (find_nonzeros_from_input_vec ?lhs))" using obv length_pivot_positions_dim_row[where A = "(gauss_jordan_single (?step1_mat\<^sup>T))"] by (metis carrier_matD(1) carrier_matD(2) gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(2) step1_mat_size) then have len_lt_eq: "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ dim_col ?step1_mat" using step1_mat_size by simp have len_gt_false: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat) \ False" proof - assume length_lt: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat)" have h: "dim_row ?new_mat < (dim_col ?step1_mat)" by (metis Matrix.transpose_transpose index_transpose_mat(3) length_lt length_map mat_of_cols_carrier(3) take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) then show "False" using rank_new_mat by (metis Matrix.transpose_transpose carrier_matI index_transpose_mat(2) nat_less_le not_less_iff_gr_or_eq transpose_rank vec_space.rank_le_nc) qed then have len_gt_eq: "length (rows_to_keep ?step1_mat) \ (dim_col ?step1_mat)" using not_less by blast have len_match: "length (rows_to_keep ?step1_mat) = (dim_col ?step1_mat)" using len_lt_eq len_gt_eq by (simp add: rows_to_keep_def) have mat_match: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" using reduce_system_matrix_match[of p qs signs subsets] assms by blast have f2: "length (get_subsets (take_rows_from_matrix (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) (rows_to_keep (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))), map ((!) subsets) (rows_to_keep (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))), map ((!) signs) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" by (metis (no_types) \dim_col (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ \length (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) = dim_col (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ length_map reduce_mat_cols.simps reduce_system.simps reduction_step.simps reduction_subsets_def reduction_subsets_is_get_subsets take_cols_from_matrix_def take_indices_def) have f3: "\p ps rss nss m. map ((!) rss) (find_nonzeros_from_input_vec (solve_for_lhs p ps nss m)) = get_signs (reduce_system p (ps, m, nss, rss))" by (metis (no_types) reduction_signs_def reduction_signs_is_get_signs take_indices_def) have square_final_mat: "square_mat (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" using mat_match assms size_of_mat same_size apply (auto) using f2 f3 by (metis (no_types, lifting) Matrix.transpose_transpose fst_conv get_matrix_def index_transpose_mat(2) len_match length_map mat_of_cols_carrier(2) mat_of_cols_carrier(3) reduce_mat_cols.simps take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have rank_match: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_row ?new_mat" using len_match rank_new_mat by (simp add: take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have "invertible_mat ?new_mat" using invertible_det og_mat_size using vec_space.det_rank_iff square_final_mat by (metis dim_col_matrix_A dim_row_matrix_A fst_conv get_matrix_def mat_match rank_match reduce_system.simps reduction_step.simps size_of_mat square_mat.elims(2)) then show ?thesis apply (simp) by (metis fst_conv get_matrix_def) qed subsection "Well def signs preserved when reducing" lemma reduction_doesnt_break_length_signs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes length_init : "\ x\ set(signs). length x = length qs" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" shows "\x \ set(reduction_signs p qs signs subsets (matrix_A signs subsets)). length x = length qs" proof clarsimp fix x assume x_in_set: "x \ set (reduction_signs p qs signs subsets (matrix_A signs subsets))" have "List.member (reduction_signs p qs signs subsets (matrix_A signs subsets)) x" using x_in_set by (simp add: in_set_member) then have take_ind: "List.member (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) x" unfolding reduction_signs_def by auto have find_nz_len: "\y. List.member (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) y \ y < length signs" using size_of_lhs sat_eq inv_mat construct_lhs_matches_solve_for_lhs[of p qs subsets signs] unfolding find_nonzeros_from_input_vec_def by (metis atLeastLessThan_iff filter_is_subset in_set_member set_upt subset_code(1)) then have "\ n < length signs. x = signs ! n" using take_ind by (metis in_set_conv_nth reduction_signs_def take_indices_lem x_in_set) then show "length x = length qs" using length_init take_indices_lem using nth_mem by blast qed subsection "Distinct signs preserved when reducing" lemma reduction_signs_are_distinct: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" assumes distinct_init: "distinct signs" shows "distinct (reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have solve_construct: "construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by simp have h1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" unfolding find_nonzeros_from_input_vec_def using distinct_filter using distinct_upt by blast have h2: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length signs)" proof - fix j assume "j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" show "j < length signs" using solve_construct size_of_lhs by (metis \j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ atLeastLessThan_iff filter_is_subset find_nonzeros_from_input_vec_def set_upt subset_iff) qed then show ?thesis unfolding reduction_signs_def unfolding take_indices_def using distinct_init h1 h2 conjugatable_vec_space.distinct_map_nth[where ls = "signs", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))"] by blast qed subsection "Well def subsets preserved when reducing" lemma reduction_doesnt_break_subsets: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes length_init : "all_list_constr subsets (length qs)" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" shows "all_list_constr (reduction_subsets p qs signs subsets (matrix_A signs subsets)) (length qs)" unfolding all_list_constr_def proof clarsimp fix x assume in_red_subsets: "List.member (reduction_subsets p qs signs subsets (matrix_A signs subsets)) x " have solve_construct: "construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by simp have rows_to_keep_hyp: "\y. y \ set (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ y < length subsets" proof clarsimp fix y assume in_set: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))" have in_set_2: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (construct_lhs_vector p qs signs))))" using in_set solve_construct by simp let ?lhs_vec = "(solve_for_lhs p qs subsets (matrix_A signs subsets))" have h30: "(construct_lhs_vector p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs unfolding find_nonzeros_from_input_vec_def apply (auto) by (metis atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq) have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) < length signs \ matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using h3a nonzero reduce_system_matrix_signs_helper by auto then have "matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))" by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def using pivot_positions using dim_row_matrix_A h3a in_set_member nonzero reduce_system_matrix_signs_helper rows_to_keep_def rows_to_keep_lem apply (auto) by (smt (z3) \M_mat (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets)))) subsets = take_cols_from_matrix (M_mat signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (M_mat signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets))))\<^sup>T)))\ dim_row_matrix_A rows_to_keep_def rows_to_keep_lem) qed then show "y < length subsets" using in_set_member apply (auto) using in_set_2 solve_construct by fastforce qed show "list_constr x (length qs)" using in_red_subsets unfolding reduction_subsets_def using take_indices_lem[of _ subsets] rows_to_keep_hyp by (metis all_list_constr_def in_set_conv_nth in_set_member length_init) qed section "Overall Lemmas" lemma combining_to_smash: "combine_systems p (qs1, m1, (sub1, sgn1)) (qs2, m2, (sub2, sgn2)) = smash_systems p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2" by simp lemma getter_functions: "calculate_data p qs = (get_matrix (calculate_data p qs), (get_subsets (calculate_data p qs), get_signs (calculate_data p qs))) " unfolding get_matrix_def get_subsets_def get_signs_def by auto subsection "Key properties preserved" subsubsection "Properties preserved when combining and reducing systems" lemma combining_sys_satisfies_properties_helper: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" fixes signs1 signs2 :: "rat list list" fixes matrix1 matrix2:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes satisfies_properties_sys1: "satisfies_properties p qs1 subsets1 signs1 matrix1" assumes satisfies_properties_sys2: "satisfies_properties p qs2 subsets2 signs2 matrix2" shows "satisfies_properties p (qs1@qs2) (get_subsets (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_signs (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_matrix (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2)))))))" proof - let ?subsets = "(get_subsets (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?signs = "(get_signs (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?matrix = "(get_matrix (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" have h1: "all_list_constr ?subsets (length (qs1 @ qs2))" using well_def_step[of subsets1 qs1 subsets2 qs2] assms by (simp add: nontriv2 get_subsets_def satisfies_properties_def smash_systems_def) have h2: "well_def_signs (length (qs1 @ qs2)) ?signs" using well_def_signs_step[of qs1 qs2 signs1 signs2] using get_signs_def nontriv1 nontriv2 satisfies_properties_def satisfies_properties_sys1 satisfies_properties_sys2 smash_systems_def by auto have h3: "distinct ?signs" using distinct_step[of _ signs1 _ signs2] assms using combine_systems.simps get_signs_def satisfies_properties_def smash_systems_def snd_conv by auto have h4: "satisfy_equation p (qs1 @ qs2) ?subsets ?signs" using assms inductive_step[of p qs1 qs2 signs1 signs2 subsets1 subsets2] using get_signs_def get_subsets_def satisfies_properties_def smash_systems_def by auto have h5: " invertible_mat ?matrix" using assms inductive_step[of p qs1 qs2 signs1 signs2 subsets1 subsets2] by (metis combining_to_smash fst_conv get_matrix_def kronecker_invertible satisfies_properties_def smash_systems_def snd_conv) have h6: "?matrix = matrix_A ?signs ?subsets" unfolding get_matrix_def combine_systems.simps smash_systems_def get_signs_def get_subsets_def apply simp apply (subst matrix_construction_is_kronecker_product[of subsets1 _ signs1 signs2 subsets2]) apply (metis Ball_set all_list_constr_def in_set_member list_constr_def satisfies_properties_def satisfies_properties_sys1) using satisfies_properties_def satisfies_properties_sys1 well_def_signs_def apply blast using satisfies_properties_def satisfies_properties_sys1 satisfies_properties_sys2 by auto have h7: "set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)) \ set (?signs)" using subset_step[of p qs1 signs1 qs2 signs2] assms by (simp add: nonzero get_signs_def satisfies_properties_def smash_systems_def) then show ?thesis unfolding satisfies_properties_def using h1 h2 h3 h4 h5 h6 h7 by blast qed lemma combining_sys_satisfies_properties: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes satisfies_properties_sys1: "satisfies_properties p qs1 (get_subsets (calculate_data p qs1)) (get_signs (calculate_data p qs1)) (get_matrix (calculate_data p qs1))" assumes satisfies_properties_sys2: "satisfies_properties p qs2 (get_subsets (calculate_data p qs2)) (get_signs (calculate_data p qs2)) (get_matrix (calculate_data p qs2))" shows "satisfies_properties p (qs1@qs2) (get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) (get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) (get_matrix (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" using combining_sys_satisfies_properties_helper by (metis getter_functions nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero satisfies_properties_sys1 satisfies_properties_sys2) lemma reducing_sys_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes matrix:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv: "length qs > 0" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" assumes satisfies_properties_sys: "satisfies_properties p qs subsets signs matrix" shows "satisfies_properties p qs (get_subsets (reduce_system p (qs,matrix,subsets,signs))) (get_signs (reduce_system p (qs,matrix,subsets,signs))) (get_matrix (reduce_system p (qs,matrix,subsets,signs)))" proof - have h1: " all_list_constr (get_subsets (reduce_system p (qs, matrix, subsets, signs))) (length qs)" using reduction_doesnt_break_subsets assms reduction_subsets_is_get_subsets satisfies_properties_def satisfies_properties_sys by auto have h2: "well_def_signs (length qs) (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_length_signs[of signs qs p subsets] assms reduction_signs_is_get_signs satisfies_properties_def well_def_signs_def by auto have h3: "distinct (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_signs_are_distinct[of p qs subsets signs] assms reduction_signs_is_get_signs satisfies_properties_def by auto have h4: "satisfy_equation p qs (get_subsets (reduce_system p (qs, matrix, subsets, signs))) (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduce_system_matrix_equation_preserved[of p qs signs subsets] assms satisfies_properties_def by auto have h5: "invertible_mat (get_matrix (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_invertibility assms same_size satisfies_properties_def by auto have h6: "get_matrix (reduce_system p (qs, matrix, subsets, signs)) = matrix_A (get_signs (reduce_system p (qs, matrix, subsets, signs))) (get_subsets (reduce_system p (qs, matrix, subsets, signs)))" using reduce_system_matrix_match[of p qs signs subsets] assms satisfies_properties_def by auto have h7: "set (characterize_consistent_signs_at_roots_copr p qs) \ set (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_signs[of p qs signs subsets] assms reduction_signs_is_get_signs satisfies_properties_def by auto then show ?thesis unfolding satisfies_properties_def using h1 h2 h3 h4 h5 h6 h7 by blast qed subsubsection "For length 1 qs" lemma length_1_calculate_data_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" proof - have h1: "all_list_constr [[],[0]] (length qs)" using len1 unfolding all_list_constr_def list_constr_def apply (auto) by (metis (full_types) length_Cons less_Suc0 list.size(3) list_all_length list_all_simps(2) member_rec(1) member_rec(2) nth_Cons_0) have h2: "well_def_signs (length qs) [[1],[-1]]" unfolding well_def_signs_def using len1 in_set_member by auto have h3: "distinct ([[1],[-1]]::rat list list)" unfolding distinct_def using in_set_member by auto have h4: "satisfy_equation p qs [[],[0]] [[1],[-1]]" using assms base_case_satisfy_equation_alt[of qs p] by auto have h6: "(mat_of_rows_list 2 [[1,1], [1,-1]]::rat mat) = (matrix_A ([[1],[-1]]) ([[],[0]]) :: rat mat)" using mat_base_case by auto then have h5: "invertible_mat (mat_of_rows_list 2 [[1,1], [1,-1]]:: rat mat)" using base_case_invertible_mat by simp have h7: "set (characterize_consistent_signs_at_roots_copr p qs) \ set ([[1],[-1]])" using assms base_case_sgas_alt[of qs p] by simp have base_case_hyp: "satisfies_properties p qs [[],[0]] [[1],[-1]] (mat_of_rows_list 2 [[1,1], [1,-1]])" using h1 h2 h3 h4 h5 h6 h7 using satisfies_properties_def by blast then have key_hyp: "satisfies_properties p qs (get_subsets (reduce_system p (qs,base_case_info))) (get_signs (reduce_system p (qs,base_case_info))) (get_matrix (reduce_system p (qs,base_case_info)))" using reducing_sys_satisfies_properties by (metis base_case_info_def len1 nonzero pairwise_rel_prime nonzero zero_less_one_class.zero_less_one) show ?thesis by (simp add: key_hyp len1) qed subsubsection "For arbitrary qs" lemma append_not_distinct_helper: "(List.member l1 m \ List.member l2 m) \ (distinct (l1@l2) = False)" proof - have h1: "List.member l1 m \ (\ n. n < length l1 \ (nth l1 n) = m)" using member_def nth_find_first by (simp add: member_def in_set_conv_nth) have h2: "\n. n < length l1 \ (nth l1 n) = m \ (nth (l1@l2) n) = m" proof clarsimp fix n assume lt: "n < length l1" assume nth_l1: "m = l1 ! n" show "(l1 @ l2) ! n = l1 ! n" proof (induct l2) case Nil then show ?case by simp next case (Cons a l2) then show ?case by (simp add: lt nth_append) qed qed have h3: "List.member l1 m \ (\n. n < length l1 \ (nth (l1@l2) n) = m)" using h1 h2 by auto have h4: "List.member l2 m \ (\ n. (nth l2 n) = m)" by (meson member_def nth_find_first) have h5: "\n. (nth l2 n) = m \ (nth (l1@l2) (n + length l1)) = m" proof clarsimp fix n assume nth_l2: "m = l2 ! n" show "(l1 @ l2) ! (n + length l1) = l2 ! n" proof (induct l2) case Nil then show ?case by (metis add.commute nth_append_length_plus) next case (Cons a l2) then show ?case by (simp add: nth_append) qed qed have h6: "List.member l2 m \ (\n. (nth (l1@l2) (n + length l1)) = m)" using h4 h5 by blast show ?thesis using h3 h6 by (metis distinct_append equalityI insert_disjoint(1) insert_subset member_def order_refl) qed lemma calculate_data_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" shows "(p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" proof (induction "length qs" arbitrary: qs rule: less_induct) case less have len1_h: "length qs = 1 \ ( p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" using length_1_calculate_data_satisfies_properties by blast let ?len = "length qs" let ?q1 = "take (?len div 2) qs" let ?left = "calculate_data p ?q1" let ?q2 = "drop (?len div 2) qs" let ?right = "calculate_data p ?q2" let ?comb = "combine_systems p (?q1,?left) (?q2,?right)" let ?red = "reduce_system p ?comb" have h_q1_len: "length qs > 1 \ (length ?q1 > 0)" by auto have h_q2_len: "length qs > 1 \ (length ?q2 > 0)" by auto have h_q1_copr: "(\q. ((List.member qs q) \ (coprime p q))) \ (\q. ((List.member ?q1 q) \ (coprime p q)))" proof - have mem_hyp: "(\q. ((List.member ?q1 q) \ (List.member qs q)))" by (meson in_set_member in_set_takeD) then show ?thesis by blast qed have h_q2_copr: "(\q. ((List.member qs q) \ (coprime p q))) \ (\q. ((List.member ?q2 q) \ (coprime p q)))" proof - have mem_hyp: "(\q. ((List.member ?q2 q) \ (List.member qs q)))" by (meson in_set_dropD in_set_member) then show ?thesis by blast qed have q1_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p ?q1 (get_subsets (calculate_data p ?q1)) (get_signs (calculate_data p ?q1)) (get_matrix (calculate_data p ?q1))" - using less.hyps[of ?q1] h_q1_len h_q1_copr by auto + using less.hyps[of ?q1] h_q1_len h_q1_copr + by (auto simp add: Let_def) have q2_help: "length qs > 1 \ length (drop (length qs div 2) qs) < length qs" using h_q1_len by auto then have q2_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p ?q2 (get_subsets (calculate_data p ?q2)) (get_signs (calculate_data p ?q2)) (get_matrix (calculate_data p ?q2))" using less.hyps[of ?q2] h_q2_len h_q2_copr by blast have put_tog: "?q1@?q2 = qs" using append_take_drop_id by blast then have comb_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p (qs) (get_subsets (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))) (get_signs (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))) (get_matrix (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))))" using q1_sat_props q2_sat_props combining_sys_satisfies_properties using h_q1_copr h_q1_len h_q2_copr h_q2_len put_tog by metis then have comb_sat: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p (qs) (get_subsets (snd ?comb)) (get_signs (snd ?comb)) (get_matrix (snd ?comb)))" by blast have red_char: "?red = (reduce_system p (qs,(get_matrix (snd ?comb)),(get_subsets (snd ?comb)),(get_signs (snd ?comb))))" using getter_functions by (smt Pair_inject combining_to_smash get_matrix_def get_signs_def get_subsets_def prod.collapse put_tog smash_systems_def) then have "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p qs (get_subsets ?red) (get_signs ?red) (get_matrix ?red))" using reducing_sys_satisfies_properties comb_sat by presburger have len_gt1: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" by (smt \1 < length qs \ p \ 0 \ 0 < length qs \ (\q. List.member qs q \ coprime p q) \ satisfies_properties p qs (get_subsets (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs))))) (get_signs (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs))))) (get_matrix (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs)))))\ calculate_data.simps neq0_conv not_less) then show ?case using len1_h len_gt1 by (metis One_nat_def Suc_lessI) qed subsection "Some key results on consistent sign assignments" lemma find_consistent_signs_at_roots_len1: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "set (find_consistent_signs_at_roots p qs) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - let ?signs = "[[1],[-1]]::rat list list" let ?subsets = "[[],[0]]::nat list list" have mat_help: "matrix_A [[1],[-1]] [[],[0]] = (mat_of_rows_list 2 [[1,1], [1,-1]])" using mat_base_case by auto have well_def_signs: "well_def_signs (length qs) ?signs" unfolding well_def_signs_def using len1 by auto have distinct_signs: "distinct ?signs" unfolding distinct_def by auto have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then have all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(?signs)" using assms base_case_sgas apply (auto) by (meson in_mono insertE insert_absorb insert_not_empty member_rec(1)) have match: "satisfy_equation p qs ?subsets ?signs" using ex_q base_case_satisfy_equation nonzero pairwise_rel_prime apply (auto) by (simp add: member_rec(1)) have invertible_mat: "invertible_mat (matrix_A ?signs ?subsets)" using inverse_mat_base_case inverse_mat_base_case_2 unfolding invertible_mat_def using mat_base_case by auto have h: "set (get_signs (reduce_system p (qs, ((matrix_A ?signs ?subsets), (?subsets, ?signs))))) = set (characterize_consistent_signs_at_roots_copr p qs)" using nonzero nonzero well_def_signs distinct_signs all_info match invertible_mat reduce_system_sign_conditions[where p = "p", where qs = "qs", where signs = "[[1],[-1]]", where subsets = "[[],[0]]"] by blast then have "set (snd (snd (reduce_system p (qs, ((mat_of_rows_list 2 [[1,1], [1,-1]]), ([[],[0]], [[1],[-1]])))))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding get_signs_def using mat_help by auto then have "set (snd (snd (reduce_system p (qs, base_case_info)))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding base_case_info_def by auto then show ?thesis using len1 by (simp add: find_consistent_signs_at_roots_thm) qed lemma smaller_sys_are_good: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes "set(find_consistent_signs_at_roots p qs1) = set(characterize_consistent_signs_at_roots_copr p qs1)" assumes "set(find_consistent_signs_at_roots p qs2) = set(characterize_consistent_signs_at_roots_copr p qs2)" shows "set(snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) = set(characterize_consistent_signs_at_roots_copr p (qs1@qs2))" proof - let ?signs = "(get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) " let ?subsets = "(get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) " have h0: "satisfies_properties p (qs1@qs2) ?subsets ?signs (get_matrix (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" using calculate_data_satisfies_properties combining_sys_satisfies_properties using nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero by simp then have h1: "set(characterize_consistent_signs_at_roots_copr p (qs1@qs2)) \ set ?signs" unfolding satisfies_properties_def by linarith have h2: "well_def_signs (length (qs1@qs2)) ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h3: "distinct ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h4: "satisfy_equation p (qs1@qs2) ?subsets ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h5: "invertible_mat (matrix_A ?signs ?subsets) " using calculate_data_satisfies_properties using h0 satisfies_properties_def by auto have h: "set (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets)))) = set(characterize_consistent_signs_at_roots_copr p (qs1@qs2))" using h1 h2 h3 h4 h5 reduction_deletes_bad_sign_conds using nonzero nonzero reduction_signs_def by auto then have h: "set (characterize_consistent_signs_at_roots_copr p (qs1@qs2)) = set (reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets ))" unfolding reduction_signs_def get_signs_def by blast have help_h: "reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets) = (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets))))" unfolding reduction_signs_def by auto have clear_signs: "(signs_smash (get_signs (calculate_data p qs1)) (get_signs (calculate_data p qs2))) = (get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" by (smt combining_to_smash get_signs_def getter_functions smash_systems_def snd_conv) have clear_subsets: "(subsets_smash (length qs1) (get_subsets(calculate_data p qs1)) (get_subsets (calculate_data p qs2))) = (get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" by (smt Pair_inject combining_to_smash get_subsets_def prod.collapse smash_systems_def) have "well_def_signs (length qs1) (get_signs (calculate_data p qs1))" using calculate_data_satisfies_properties using nontriv1 nonzero pairwise_rel_prime1 nonzero satisfies_properties_def by auto then have well_def_signs1: "(\j. j \ set (get_signs (calculate_data p qs1)) \ length j = (length qs1))" using well_def_signs_def by blast have "all_list_constr (get_subsets(calculate_data p qs1)) (length qs1) " using calculate_data_satisfies_properties using nontriv1 nonzero pairwise_rel_prime1 nonzero satisfies_properties_def by auto then have well_def_subsets1: "(\l i. l \ set (get_subsets(calculate_data p qs1)) \ i \ set l \ i < (length qs1))" unfolding all_list_constr_def list_constr_def using in_set_member by (metis in_set_conv_nth list_all_length) have extra_matrix_same: "matrix_A (signs_smash (get_signs (calculate_data p qs1)) (get_signs (calculate_data p qs2))) (subsets_smash (length qs1) (get_subsets(calculate_data p qs1)) (get_subsets (calculate_data p qs2))) = kronecker_product (get_matrix (calculate_data p qs1)) (get_matrix (calculate_data p qs2))" using well_def_signs1 well_def_subsets1 using matrix_construction_is_kronecker_product using calculate_data_satisfies_properties nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero satisfies_properties_def by auto then have matrix_same: "matrix_A ?signs ?subsets = kronecker_product (get_matrix (calculate_data p qs1)) (get_matrix (calculate_data p qs2))" using clear_signs clear_subsets by simp have comb_sys_h: "snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))) = snd(snd(reduce_system p (qs1@qs2, (matrix_A ?signs ?subsets, (?subsets, ?signs)))))" unfolding get_signs_def get_subsets_def using matrix_same by (smt combining_to_smash get_signs_def get_subsets_def getter_functions prod.collapse prod.inject smash_systems_def) then have extra_h: " snd(snd(reduce_system p (qs1@qs2, (matrix_A ?signs ?subsets, (?subsets, ?signs))))) = snd(snd(reduction_step (matrix_A ?signs ?subsets) ?signs ?subsets (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets)))) " by simp then have same_h: "set(snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) = set (reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets ))" using comb_sys_h unfolding reduction_signs_def by (metis get_signs_def help_h reduction_signs_is_get_signs) then show ?thesis using h by blast qed lemma find_consistent_signs_at_roots_1: fixes p:: "real poly" fixes qs :: "real poly list" shows "(p \ 0 \ length qs > 0 \ (\q. ((List.member qs q) \ (coprime p q)))) \ set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots_copr p qs)" proof (induction "length qs" arbitrary: qs rule: less_induct) case less then show ?case proof clarsimp assume ind_hyp: "(\qsa. length qsa < length qs \ qsa \ [] \ (\q. List.member qsa q \ coprime p q) \ set (find_consistent_signs_at_roots p qsa) = set (characterize_consistent_signs_at_roots_copr p qsa))" assume nonzero: "p \ 0" assume nontriv: "qs \ []" assume copr:" \q. List.member qs q \ coprime p q" let ?len = "length qs" let ?q1 = "take ((?len) div 2) qs" let ?left = "calculate_data p ?q1" let ?q2 = "drop ((?len) div 2) qs" let ?right = "calculate_data p ?q2" have nontriv_q1: "length qs>1 \ length ?q1 > 0" by auto have qs_more_q1: "length qs>1 \ length qs > length ?q1" by auto have pairwise_rel_prime_q1: "\q. ((List.member ?q1 q) \ (coprime p q))" proof clarsimp fix q assume mem: "List.member (take (length qs div 2) qs) q" have "List.member qs q" using mem set_take_subset[where n = "length qs div 2"] proof - show ?thesis by (meson \List.member (take (length qs div 2) qs) q\ in_set_member in_set_takeD) qed then show "coprime p q" using copr by blast qed have nontriv_q2: "length qs>1 \length ?q2 > 0" by auto have qs_more_q2: "length qs>1 \ length qs > length ?q2" by auto have pairwise_rel_prime_q2: "\q. ((List.member ?q2 q) \ (coprime p q))" proof clarsimp fix q assume mem: "List.member (drop (length qs div 2) qs) q" have "List.member qs q" using mem set_take_subset[where n = "length qs div 2"] proof - show ?thesis by (meson \List.member (drop (length qs div 2) qs) q\ in_set_dropD in_set_member) qed then show "coprime p q" using copr by blast qed have key_h: "set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h_len1 : "?len = 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" using find_consistent_signs_at_roots_len1[of p qs] copr nonzero nontriv by (simp add: find_consistent_signs_at_roots_thm) have h_len_gt1 : "?len > 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h_imp_a: "?len > 1 \ set (snd (snd (reduce_system p (combine_systems p (?q1, ?left) (?q2, ?right))))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h1: "?len > 1 \ set(snd(snd(?left))) = set (characterize_consistent_signs_at_roots_copr p ?q1)" using nontriv_q1 pairwise_rel_prime_q1 ind_hyp[of ?q1] qs_more_q1 by (metis find_consistent_signs_at_roots_thm less_numeral_extra(3) list.size(3)) have h2: "?len > 1 \ set(snd(snd(?right))) = set (characterize_consistent_signs_at_roots_copr p ?q2)" using nontriv_q2 pairwise_rel_prime_q2 ind_hyp[of ?q2] qs_more_q2 by (metis (full_types) find_consistent_signs_at_roots_thm list.size(3) not_less_zero) show ?thesis using nonzero nontriv_q1 pairwise_rel_prime_q1 nontriv_q2 pairwise_rel_prime_q2 h1 h2 smaller_sys_are_good by (metis append_take_drop_id find_consistent_signs_at_roots_thm) qed then have h_imp: "?len > 1 \ set (snd (snd (Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" by auto then show ?thesis by auto qed show ?thesis using h_len1 h_len_gt1 by (meson \qs \ []\ length_0_conv less_one nat_neq_iff) qed then show "set (find_consistent_signs_at_roots p qs) = set (characterize_consistent_signs_at_roots_copr p qs)" by (smt One_nat_def calculate_data.simps find_consistent_signs_at_roots_thm length_0_conv nontriv) qed qed lemma find_consistent_signs_at_roots_0: fixes p:: "real poly" assumes "p \ 0" shows "set(find_consistent_signs_at_roots p []) = set(characterize_consistent_signs_at_roots_copr p [])" proof - obtain a b c where abc: "reduce_system p ([1], base_case_info) = (a,b,c)" using prod_cases3 by blast have "find_consistent_signs_at_roots p [1] = c" using abc by (simp add: find_consistent_signs_at_roots_thm) have *: "set (find_consistent_signs_at_roots p [1]) = set (characterize_consistent_signs_at_roots_copr p [1])" apply (subst find_consistent_signs_at_roots_1) using assms apply auto by (simp add: member_rec(1) member_rec(2)) have "set(characterize_consistent_signs_at_roots_copr p []) = drop 1 ` set(characterize_consistent_signs_at_roots_copr p [1])" unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def apply simp by (smt drop0 drop_Suc_Cons image_cong image_image) thus ?thesis using abc * apply (auto) apply (simp add: find_consistent_signs_at_roots_thm) by (simp add: find_consistent_signs_at_roots_thm) qed lemma find_consistent_signs_at_roots_copr: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots_copr p qs)" by (metis assms find_consistent_signs_at_roots_0 find_consistent_signs_at_roots_1 in_set_member length_greater_0_conv) lemma find_consistent_signs_at_roots: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots p qs)" using assms find_consistent_signs_at_roots_copr csa_list_copr_rel by (simp add: in_set_member) end \ No newline at end of file diff --git a/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy b/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy --- a/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy +++ b/thys/Berlekamp_Zassenhaus/Mahler_Measure.thy @@ -1,868 +1,869 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Mahler Measure\ text \This part contains a definition of the Mahler measure, it contains Landau's inequality and the Graeffe-transformation. We also assemble a heuristic to approximate the Mahler's measure.\ theory Mahler_Measure imports Sqrt_Babylonian.Sqrt_Babylonian Poly_Mod_Finite_Field_Record_Based (* stuff about polynomials *) Polynomial_Factorization.Fundamental_Theorem_Algebra_Factorized Polynomial_Factorization.Missing_Multiset begin context comm_monoid_list begin lemma induct_gen_abs: assumes "\ a r. a\set lst \ P (f (h a) r) (f (g a) r)" "\ x y z. P x y \ P y z \ P x z" "P (F (map g lst)) (F (map g lst))" shows "P (F (map h lst)) (F (map g lst)) " using assms proof(induct lst arbitrary:P) case (Cons a as P) have inl:"a\set (a#as)" by auto let ?uf = "\ v w. P (f (g a) v) (f (g a) w)" have p_suc:"?uf (F (map g as)) (F (map g as))" using Cons.prems(3) by auto { fix r aa assume "aa \ set as" hence ins:"aa \ set (a#as)" by auto have "P (f (g a) (f (h aa) r)) (f (g a) (f (g aa) r))" using Cons.prems(1)[of aa "f r (g a)",OF ins] by (auto simp: assoc commute left_commute) } note h = this from Cons.hyps(1)[of ?uf, OF h Cons.prems(2)[simplified] p_suc] have e1:"P (f (g a) (F (map h as))) (f (g a) (F (map g as)))" by simp have e2:"P (f (h a) (F (map h as))) (f (g a) (F (map h as)))" using Cons.prems(1)[OF inl] by blast from Cons(3)[OF e2 e1] show ?case by auto next qed auto end lemma prod_induct_gen: assumes "\ a r. f (h a * r :: 'a :: {comm_monoid_mult}) = f (g a * r)" shows "f (\v\lst. h v) = f (\v\lst. g v)" proof - let "?P x y" = "f x = f y" show ?thesis using comm_monoid_mult_class.prod_list.induct_gen_abs[of _ ?P,OF assms] by auto qed abbreviation complex_of_int::"int \ complex" where "complex_of_int \ of_int" definition l2norm_list :: "int list \ int" where "l2norm_list lst = \sqrt (sum_list (map (\ a. a * a) lst))\" abbreviation l2norm :: "int poly \ int" where "l2norm p \ l2norm_list (coeffs p)" abbreviation "norm2 p \ \a\coeffs p. (cmod a)\<^sup>2" (* the square of the Euclidean/l2-norm *) abbreviation l2norm_complex where "l2norm_complex p \ sqrt (norm2 p)" abbreviation height :: "int poly \ int" where "height p \ max_list (map (nat \ abs) (coeffs p))" definition complex_roots_complex where "complex_roots_complex (p::complex poly) = (SOME as. smult (coeff p (degree p)) (\a\as. [:- a, 1:]) = p \ length as = degree p)" lemma complex_roots: "smult (lead_coeff p) (\a\complex_roots_complex p. [:- a, 1:]) = p" "length (complex_roots_complex p) = degree p" using someI_ex[OF fundamental_theorem_algebra_factorized] unfolding complex_roots_complex_def by simp_all lemma complex_roots_c [simp]: "complex_roots_complex [:c:] = []" using complex_roots(2) [of "[:c:]"] by simp declare complex_roots(2)[simp] lemma complex_roots_1 [simp]: "complex_roots_complex 1 = []" using complex_roots_c [of 1] by (simp add: pCons_one) lemma linear_term_irreducible\<^sub>d[simp]: "irreducible\<^sub>d [: a, 1:]" by (rule linear_irreducible\<^sub>d, simp) definition complex_roots_int where "complex_roots_int (p::int poly) = complex_roots_complex (map_poly of_int p)" lemma complex_roots_int: "smult (lead_coeff p) (\a\complex_roots_int p. [:- a, 1:]) = map_poly of_int p" "length (complex_roots_int p) = degree p" proof - show "smult (lead_coeff p) (\a\complex_roots_int p. [:- a, 1:]) = map_poly of_int p" "length (complex_roots_int p) = degree p" using complex_roots[of "map_poly of_int p"] unfolding complex_roots_int_def by auto qed text \The measure for polynomials, after K. Mahler\ definition mahler_measure_poly where "mahler_measure_poly p = cmod (lead_coeff p) * (\a\complex_roots_complex p. (max 1 (cmod a)))" definition mahler_measure where "mahler_measure p = mahler_measure_poly (map_poly complex_of_int p)" definition mahler_measure_monic where "mahler_measure_monic p = (\a\complex_roots_complex p. (max 1 (cmod a)))" lemma mahler_measure_poly_via_monic : "mahler_measure_poly p = cmod (lead_coeff p) * mahler_measure_monic p" unfolding mahler_measure_poly_def mahler_measure_monic_def by simp lemma smult_inj[simp]: assumes "(a::'a::idom) \ 0" shows "inj (smult a)" proof- interpret map_poly_inj_zero_hom "(*) a" using assms by (unfold_locales, auto) show ?thesis unfolding smult_as_map_poly by (rule inj_f) qed definition reconstruct_poly::"'a::idom \ 'a list \ 'a poly" where "reconstruct_poly c roots = smult c (\a\roots. [:- a, 1:])" lemma reconstruct_is_original_poly: "reconstruct_poly (lead_coeff p) (complex_roots_complex p) = p" using complex_roots(1) by (simp add: reconstruct_poly_def) lemma reconstruct_with_type_conversion: "smult (lead_coeff (map_poly of_int f)) (prod_list (map (\ a. [:- a, 1:]) (complex_roots_int f))) = map_poly of_int f" unfolding complex_roots_int_def complex_roots(1) by simp lemma reconstruct_prod: shows "reconstruct_poly (a::complex) as * reconstruct_poly b bs = reconstruct_poly (a * b) (as @ bs)" unfolding reconstruct_poly_def by auto lemma linear_term_inj[simplified,simp]: "inj (\ a. [:- a, 1::'a::idom:])" unfolding inj_on_def by simp lemma reconstruct_poly_monic_defines_mset: assumes "(\a\as. [:- a, 1:]) = (\a\bs. [:- a, 1::'a::field:])" shows "mset as = mset bs" proof - let ?as = "mset (map (\ a. [:- a, 1:]) as)" let ?bs = "mset (map (\ a. [:- a, 1:]) bs)" have eq_smult:"prod_mset ?as = prod_mset ?bs" using assms by (metis prod_mset_prod_list) have irr:"\ as::'a list. set_mset (mset (map (\ a. [:- a, 1:]) as)) \ {q. irreducible q \ monic q}" by (auto intro!: linear_term_irreducible\<^sub>d[of "-_::'a", simplified]) from monic_factorization_unique_mset[OF eq_smult irr irr] show ?thesis apply (subst inj_eq[OF multiset.inj_map,symmetric]) by auto qed lemma reconstruct_poly_defines_mset_of_argument: assumes "(a::'a::field) \ 0" "reconstruct_poly a as = reconstruct_poly a bs" shows "mset as = mset bs" proof - have eq_smult:"smult a (\a\as. [:- a, 1:]) = smult a (\a\bs. [:- a, 1:])" using assms(2) by (auto simp:reconstruct_poly_def) from reconstruct_poly_monic_defines_mset[OF Fun.injD[OF smult_inj[OF assms(1)] eq_smult]] show ?thesis by simp qed lemma complex_roots_complex_prod [simp]: assumes "f \ 0" "g \ 0" shows "mset (complex_roots_complex (f * g)) = mset (complex_roots_complex f) + mset (complex_roots_complex g)" proof - let ?p = "f * g" let "?lc v" = "(lead_coeff (v:: complex poly))" have nonzero_prod:"?lc ?p \ 0" using assms by auto from reconstruct_prod[of "?lc f" "complex_roots_complex f" "?lc g" "complex_roots_complex g"] have "reconstruct_poly (?lc ?p) (complex_roots_complex ?p) = reconstruct_poly (?lc ?p) (complex_roots_complex f @ complex_roots_complex g)" unfolding lead_coeff_mult[symmetric] reconstruct_is_original_poly by auto from reconstruct_poly_defines_mset_of_argument[OF nonzero_prod this] show ?thesis by simp qed lemma mset_mult_add: assumes "mset (a::'a::field list) = mset b + mset c" shows "prod_list a = prod_list b * prod_list c" unfolding prod_mset_prod_list[symmetric] using prod_mset_Un[of "mset b" "mset c",unfolded assms[symmetric]]. lemma mset_mult_add_2: assumes "mset a = mset b + mset c" shows "prod_list (map i a::'b::field list) = prod_list (map i b) * prod_list (map i c)" proof - have r:"mset (map i a) = mset (map i b) + mset (map i c) " using assms by (metis map_append mset_append mset_map) show ?thesis using mset_mult_add[OF r] by auto qed lemma measure_mono_eq_prod: assumes "f \ 0" "g \ 0" shows "mahler_measure_monic (f * g) = mahler_measure_monic f * mahler_measure_monic g" unfolding mahler_measure_monic_def using mset_mult_add_2[OF complex_roots_complex_prod[OF assms],of "\ a. max 1 (cmod a)"] by simp lemma mahler_measure_poly_0[simp]: "mahler_measure_poly 0 = 0" unfolding mahler_measure_poly_via_monic by auto lemma measure_eq_prod: (* Remark 10.2 *) "mahler_measure_poly (f * g) = mahler_measure_poly f * mahler_measure_poly g" proof - consider "f = 0" | "g = 0" | (both) "f \ 0" "g \ 0" by auto thus ?thesis proof(cases) case both show ?thesis unfolding mahler_measure_poly_via_monic norm_mult lead_coeff_mult by (auto simp: measure_mono_eq_prod[OF both]) qed (simp_all) qed lemma prod_cmod[simp]: "cmod (\a\lst. f a) = (\a\lst. cmod (f a))" by(induct lst,auto simp:real_normed_div_algebra_class.norm_mult) lemma lead_coeff_of_prod[simp]: "lead_coeff (\a\lst. f a::'a::idom poly) = (\a\lst. lead_coeff (f a))" by(induct lst,auto simp:lead_coeff_mult) lemma ineq_about_squares:assumes "x \ (y::real)" shows "x \ c^2 + y" using assms by (simp add: add.commute add_increasing2) lemma first_coeff_le_tail:"(cmod (lead_coeff g))^2 \ (\a\coeffs g. (cmod a)^2)" proof(induct g) case (pCons a p) thus ?case proof(cases "p = 0") case False show ?thesis using pCons unfolding lead_coeff_pCons(1)[OF False] by(cases "a = 0",simp_all add:ineq_about_squares) qed simp qed simp lemma square_prod_cmod[simp]: "(cmod (a * b))^2 = cmod a ^ 2 * cmod b ^ 2" by (simp add: norm_mult power_mult_distrib) lemma sum_coeffs_smult_cmod: "(\a\coeffs (smult v p). (cmod a)^2) = (cmod v)^2 * (\a\coeffs p. (cmod a)^2)" (is "?l = ?r") proof - have "?l = (\a\coeffs p. (cmod v)^2 * (cmod a)^2)" by(cases "v=0";induct p,auto) thus ?thesis by (auto simp:sum_list_const_mult) qed abbreviation "linH a \ if (cmod a > 1) then [:- 1,cnj a:] else [:- a,1:]" lemma coeffs_cong_1[simp]: "cCons a v = cCons b v \ a = b" unfolding cCons_def by auto lemma strip_while_singleton[simp]: "strip_while ((=) 0) [v * a] = cCons (v * a) []" unfolding cCons_def strip_while_def by auto lemma coeffs_times_linterm: shows "coeffs (pCons 0 (smult a p) + smult b p) = strip_while (HOL.eq (0::'a::{comm_ring_1})) (map (\(c,d).b*d+c*a) (zip (0 # coeffs p) (coeffs p @ [0])))" proof - {fix v have "coeffs (smult b p + pCons (a* v) (smult a p)) = strip_while (HOL.eq 0) (map (\(c,d).b*d+c*a) (zip ([v] @ coeffs p) (coeffs p @ [0])))" proof(induct p arbitrary:v) case (pCons pa ps) thus ?case by auto qed auto (* just putting ;auto does not work *) } from this[of 0] show ?thesis by (simp add: add.commute) qed lemma filter_distr_rev[simp]: shows "filter f (rev lst) = rev (filter f lst)" by(induct lst;auto) lemma strip_while_filter: shows "filter ((\) 0) (strip_while ((=) 0) (lst::'a::zero list)) = filter ((\) 0) lst" proof - {fix lst::"'a list" have "filter ((\) 0) (dropWhile ((=) 0) lst) = filter ((\) 0) lst" by (induct lst;auto) hence "(filter ((\) 0) (strip_while ((=) 0) (rev lst))) = filter ((\) 0) (rev lst)" unfolding strip_while_def by(simp)} from this[of "rev lst"] show ?thesis by simp qed lemma sum_stripwhile[simp]: assumes "f 0 = 0" shows "(\a\strip_while ((=) 0) lst. f a) = (\a\lst. f a)" proof - {fix lst have "(\a\filter ((\) 0) lst. f a) = (\a\lst. f a)" by(induct lst,auto simp:assms)} note f=this have "sum_list (map f (filter ((\) 0) (strip_while ((=) 0) lst))) = sum_list (map f (filter ((\) 0) lst))" using strip_while_filter[of lst] by(simp) thus ?thesis unfolding f. qed lemma complex_split : "Complex a b = c \ (a = Re c \ b = Im c)" using complex_surj by auto lemma norm_times_const:"(\y\lst. (cmod (a * y))\<^sup>2) = (cmod a)\<^sup>2 * (\y\lst. (cmod y)\<^sup>2)" by(induct lst,auto simp:ring_distribs) fun bisumTail where (* Used for Landau's lemma *) "bisumTail f (Cons a (Cons b bs)) = f a b + bisumTail f (Cons b bs)" | "bisumTail f (Cons a Nil) = f a 0" | "bisumTail f Nil = f 1 0" (* never called, not used in proofs *) fun bisum where "bisum f (Cons a as) = f 0 a + bisumTail f (Cons a as)" | "bisum f Nil = f 0 0" lemma bisumTail_is_map_zip: "(\x\zip (v # l1) (l1 @ [0]). f x) = bisumTail (\x y .f (x,y)) (v#l1)" by(induct l1 arbitrary:v,auto) (* converting to and from bisum *) lemma bisum_is_map_zip: "(\x\zip (0 # l1) (l1 @ [0]). f x) = bisum (\x y. f (x,y)) l1" using bisumTail_is_map_zip[of f "hd l1" "tl l1"] by(cases l1,auto) lemma map_zip_is_bisum: "bisum f l1 = (\(x,y)\zip (0 # l1) (l1 @ [0]). f x y)" using bisum_is_map_zip[of "\(x,y). f x y"] by auto lemma bisum_outside : "(bisum (\ x y. f1 x - f2 x y + f3 y) lst :: 'a :: field) = sum_list (map f1 lst) + f1 0 - bisum f2 lst + sum_list (map f3 lst) + f3 0" proof(cases lst) case (Cons a lst) show ?thesis unfolding map_zip_is_bisum Cons by(induct lst arbitrary:a,auto) qed auto lemma Landau_lemma: "(\a\coeffs (\a\lst. [:- a, 1:]). (cmod a)\<^sup>2) = (\a\coeffs (\a\lst. linH a). (cmod a)\<^sup>2)" (is "norm2 ?l = norm2 ?r") proof - have a:"\ a. (cmod a)\<^sup>2 = Re (a * cnj a) " using complex_norm_square unfolding complex_split complex_of_real_def by simp have b:"\ x a y. (cmod (x - a * y))^2 = (cmod x)\<^sup>2 - Re (a * y * cnj x + x * cnj (a * y)) + (cmod (a * y))^2" unfolding left_diff_distrib right_diff_distrib a complex_cnj_diff by simp have c:"\ y a x. (cmod (cnj a * x - y))\<^sup>2 = (cmod (a * x))\<^sup>2 - Re (a * y * cnj x + x * cnj (a * y)) + (cmod y)^2" unfolding left_diff_distrib right_diff_distrib a complex_cnj_diff by (simp add: mult.assoc mult.left_commute) { fix f1 a have "norm2 ([:- a, 1 :] * f1) = bisum (\x y. cmod (x - a * y)^2) (coeffs f1)" by(simp add: bisum_is_map_zip[of _ "coeffs f1"] coeffs_times_linterm[of 1 _ "-a",simplified]) also have "\ = norm2 f1 + cmod a^2*norm2 f1 - bisum (\x y. Re (a * y * cnj x + x * cnj (a * y))) (coeffs f1)" unfolding b bisum_outside norm_times_const by simp also have "\ = bisum (\x y. cmod (cnj a * x - y)^2) (coeffs f1)" unfolding c bisum_outside norm_times_const by auto also have "\ = norm2 ([:- 1, cnj a :] * f1)" using coeffs_times_linterm[of "cnj a" _ "-1"] by(simp add: bisum_is_map_zip[of _ "coeffs f1"] mult.commute) finally have "norm2 ([:- a, 1 :] * f1) = \".} hence h:"\ a f1. norm2 ([:- a, 1 :] * f1) = norm2 (linH a * f1)" by auto show ?thesis by(rule prod_induct_gen[OF h]) qed lemma Landau_inequality: "mahler_measure_poly f \ l2norm_complex f" proof - let ?f = "reconstruct_poly (lead_coeff f) (complex_roots_complex f)" let ?roots = "(complex_roots_complex f)" let ?g = "\a\?roots. linH a" (* g is chosen such that lead_coeff_g holds, and its l2 norm is equal to f's l2 norm *) - have max:"\a. cmod (if 1 < cmod a then cnj a else 1) = max 1 (cmod a)" by(simp add:if_split,auto) + have max:"\a. cmod (if 1 < cmod a then cnj a else 1) = max 1 (cmod a)" + by simp have "\a. 1 < cmod a \ a \ 0" by auto hence "\a. lead_coeff (linH a) = (if (cmod a > 1) then cnj a else 1)" by(auto simp:if_split) hence lead_coeff_g:"cmod (lead_coeff ?g) = (\a\?roots. max 1 (cmod a))" by(auto simp:max) have "norm2 f = (\a\coeffs ?f. (cmod a)^2)" unfolding reconstruct_is_original_poly.. also have "\ = cmod (lead_coeff f)^2 * (\a\coeffs (\a\?roots. [:- a, 1:]). (cmod a)\<^sup>2)" unfolding reconstruct_poly_def using sum_coeffs_smult_cmod. finally have fg_norm:"norm2 f = cmod (lead_coeff f)^2 * (\a\coeffs ?g. (cmod a)^2)" unfolding Landau_lemma by auto have "(cmod (lead_coeff ?g))^2 \ (\a\coeffs ?g. (cmod a)^2)" using first_coeff_le_tail by blast from ordered_comm_semiring_class.comm_mult_left_mono[OF this] have "(cmod (lead_coeff f) * cmod (lead_coeff ?g))^2 \ (\a\coeffs f. (cmod a)^2)" unfolding fg_norm by (simp add:power_mult_distrib) hence "cmod (lead_coeff f) * (\a\?roots. max 1 (cmod a)) \ sqrt (norm2 f)" using NthRoot.real_le_rsqrt lead_coeff_g by auto thus "mahler_measure_poly f \ sqrt (norm2 f)" using reconstruct_with_type_conversion[unfolded complex_roots_int_def] by (simp add: mahler_measure_poly_via_monic mahler_measure_monic_def complex_roots_int_def) qed lemma prod_list_ge1: assumes "Ball (set x) (\ (a::real). a \ 1)" shows "prod_list x \ 1" using assms proof(induct x) case (Cons a as) have "\a\set as. 1 \ a" "1 \ a" using Cons(2) by auto thus ?case using Cons.hyps mult_mono' by fastforce qed auto lemma mahler_measure_monic_ge_1: "mahler_measure_monic p \ 1" unfolding mahler_measure_monic_def by(rule prod_list_ge1,simp) lemma mahler_measure_monic_ge_0: "mahler_measure_monic p \ 0" using mahler_measure_monic_ge_1 le_numeral_extra(1) order_trans by blast lemma mahler_measure_ge_0: "0 \ mahler_measure h" unfolding mahler_measure_def mahler_measure_poly_via_monic by (simp add: mahler_measure_monic_ge_0) lemma mahler_measure_constant[simp]: "mahler_measure_poly [:c:] = cmod c" proof - have main: "complex_roots_complex [:c:] = []" unfolding complex_roots_complex_def by (rule some_equality, auto) show ?thesis unfolding mahler_measure_poly_def main by auto qed lemma mahler_measure_factor[simplified,simp]: "mahler_measure_poly [:- a, 1:] = max 1 (cmod a)" proof - have main: "complex_roots_complex [:- a, 1:] = [a]" unfolding complex_roots_complex_def proof (rule some_equality, auto, goal_cases) case (1 as) thus ?case by (cases as, auto) qed show ?thesis unfolding mahler_measure_poly_def main by auto qed lemma mahler_measure_poly_explicit: "mahler_measure_poly (smult c (\a\as. [:- a, 1:])) = cmod c * (\a\as. (max 1 (cmod a)))" proof (cases "c = 0") case True thus ?thesis by auto next case False note c = this show ?thesis proof (induct as) case (Cons a as) have "mahler_measure_poly (smult c (\a\a # as. [:- a, 1:])) = mahler_measure_poly (smult c (\a\as. [:- a, 1:]) * [: -a, 1 :])" by (rule arg_cong[of _ _ mahler_measure_poly], unfold list.simps prod_list.Cons mult_smult_left, simp) also have "\ = mahler_measure_poly (smult c (\a\as. [:- a, 1:])) * mahler_measure_poly ([:- a, 1:])" (is "_ = ?l * ?r") by (rule measure_eq_prod) also have "?l = cmod c * (\a\as. max 1 (cmod a))" unfolding Cons by simp also have "?r = max 1 (cmod a)" by simp finally show ?case by simp next case Nil show ?case by simp qed qed lemma mahler_measure_poly_ge_1: assumes "h \ 0" shows "(1::real) \ mahler_measure h" proof - have rc: "\real_of_int i\ = of_int \i\" for i by simp from assms have "cmod (lead_coeff (map_poly complex_of_int h)) > 0" by simp hence "cmod (lead_coeff (map_poly complex_of_int h)) \ 1" by(cases "lead_coeff h = 0", auto simp del: leading_coeff_0_iff) from mult_mono[OF this mahler_measure_monic_ge_1 norm_ge_zero] show ?thesis unfolding mahler_measure_def mahler_measure_poly_via_monic by auto qed lemma mahler_measure_dvd: assumes "f \ 0" and "h dvd f" shows "mahler_measure h \ mahler_measure f" proof - from assms obtain g where f: "f = g * h" unfolding dvd_def by auto from f assms have g0: "g \ 0" by auto hence mg: "mahler_measure g \ 1" by (rule mahler_measure_poly_ge_1) have "1 * mahler_measure h \ mahler_measure f" unfolding mahler_measure_def f measure_eq_prod of_int_poly_hom.hom_mult unfolding mahler_measure_def[symmetric] by (rule mult_right_mono[OF mg mahler_measure_ge_0]) thus ?thesis by simp qed definition graeffe_poly :: "'a \ 'a :: comm_ring_1 list \ nat \ 'a poly" where "graeffe_poly c as m = smult (c ^ (2^m)) (\a\as. [:- (a ^ (2^m)), 1:])" context fixes f :: "complex poly" and c as assumes f: "f = smult c (\a\as. [:- a, 1:])" begin lemma mahler_graeffe: "mahler_measure_poly (graeffe_poly c as m) = (mahler_measure_poly f)^(2^m)" proof - have graeffe: "graeffe_poly c as m = smult (c ^ 2 ^ m) (\a\(map (\ a. a ^ 2 ^ m) as). [:- a, 1:])" unfolding graeffe_poly_def by (rule arg_cong[of _ _ "smult (c ^ 2 ^ m)"], induct as, auto) { fix n :: nat assume n: "n > 0" have id: "max 1 (cmod a ^ n) = max 1 (cmod a) ^ n" for a proof (cases "cmod a \ 1") case True hence "cmod a ^ n \ 1" by (simp add: power_le_one) with True show ?thesis by (simp add: max_def) qed (auto simp: max_def) have "(\x\as. max 1 (cmod x ^ n)) = (\a\as. max 1 (cmod a)) ^ n" by (induct as, auto simp: field_simps n id) } thus ?thesis unfolding f mahler_measure_poly_explicit graeffe by (auto simp: o_def field_simps norm_power) qed end fun drop_half :: "'a list \ 'a list" where "drop_half (x # y # ys) = x # drop_half ys" | "drop_half xs = xs" fun alternate :: "'a list \ 'a list \ 'a list" where "alternate (x # y # ys) = (case alternate ys of (evn, od) \ (x # evn, y # od))" | "alternate xs = (xs,[])" definition poly_square_subst :: "'a :: comm_ring_1 poly \ 'a poly" where "poly_square_subst f = poly_of_list (drop_half (coeffs f))" definition poly_even_odd :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "poly_even_odd f = (case alternate (coeffs f) of (evn,od) \ (poly_of_list evn, poly_of_list od))" lemma poly_square_subst_coeff: "coeff (poly_square_subst f) i = coeff f (2 * i)" proof - have id: "coeff f (2 * i) = coeff (Poly (coeffs f)) (2 * i)" by simp obtain xs where xs: "coeffs f = xs" by auto show ?thesis unfolding poly_square_subst_def poly_of_list_def coeff_Poly_eq id xs proof (induct xs arbitrary: i rule: drop_half.induct) case (1 x y ys i) thus ?case by (cases i, auto) next case ("2_2" x i) thus ?case by (cases i, auto) qed auto qed lemma poly_even_odd_coeff: assumes "poly_even_odd f = (ev,od)" shows "coeff ev i = coeff f (2 * i)" "coeff od i = coeff f (2 * i + 1)" proof - have id: "\ i. coeff f i = coeff (Poly (coeffs f)) i" by simp obtain xs where xs: "coeffs f = xs" by auto from assms[unfolded poly_even_odd_def] have ev_od: "ev = Poly (fst (alternate xs))" "od = Poly (snd (alternate xs))" by (auto simp: xs split: prod.splits) have "coeff ev i = coeff f (2 * i) \ coeff od i = coeff f (2 * i + 1)" unfolding poly_of_list_def coeff_Poly_eq id xs ev_od proof (induct xs arbitrary: i rule: alternate.induct) case (1 x y ys i) thus ?case by (cases "alternate ys"; cases i, auto) next case ("2_2" x i) thus ?case by (cases i, auto) qed auto thus "coeff ev i = coeff f (2 * i)" "coeff od i = coeff f (2 * i + 1)" by auto qed lemma poly_square_subst: "poly_square_subst (f \\<^sub>p (monom 1 2)) = f" by (rule poly_eqI, unfold poly_square_subst_coeff, subst coeff_pcompose_x_pow_n, auto) lemma poly_even_odd: assumes "poly_even_odd f = (g,h)" shows "f = g \\<^sub>p monom 1 2 + monom 1 1 * (h \\<^sub>p monom 1 2)" proof - note id = poly_even_odd_coeff[OF assms] show ?thesis proof (rule poly_eqI, unfold coeff_add coeff_monom_mult) fix n :: nat obtain m i where mi: "m = n div 2" "i = n mod 2" by auto have nmi: "n = 2 * m + i" "i < 2" "0 < (2 :: nat)" "1 < (2 :: nat)" unfolding mi by auto have "(2 :: nat) \ 0" by auto show "coeff f n = coeff (g \\<^sub>p monom 1 2) n + (if 1 \ n then 1 * coeff (h \\<^sub>p monom 1 2) (n - 1) else 0)" proof (cases "i = 1") case True hence id1: "2 * m + i - 1 = 2 * m + 0" by auto show ?thesis unfolding nmi id id1 coeff_pcompose_monom[OF nmi(2)] coeff_pcompose_monom[OF nmi(3)] unfolding True by auto next case False with nmi have i0: "i = 0" by auto show ?thesis proof (cases m) case (Suc k) hence id1: "2 * m + i - 1 = 2 * k + 1" using i0 by auto show ?thesis unfolding nmi id coeff_pcompose_monom[OF nmi(2)] coeff_pcompose_monom[OF nmi(4)] id1 unfolding Suc i0 by auto next case 0 show ?thesis unfolding nmi id coeff_pcompose_monom[OF nmi(2)] unfolding i0 0 by auto qed qed qed qed context fixes f :: "'a :: idom poly" begin lemma graeffe_0: "f = smult c (\a\as. [:- a, 1:]) \ graeffe_poly c as 0 = f" unfolding graeffe_poly_def by auto lemma graeffe_recursion: assumes "graeffe_poly c as m = f" shows "graeffe_poly c as (Suc m) = smult ((-1)^(degree f)) (poly_square_subst (f * f \\<^sub>p [:0,-1:]))" proof - let ?g = "graeffe_poly c as m" have "f * f \\<^sub>p [:0,-1:] = ?g * ?g \\<^sub>p [:0,-1:]" unfolding assms by simp also have "?g \\<^sub>p [:0,-1:] = smult ((- 1) ^ length as) (smult (c ^ 2 ^ m) (\a\as. [:a ^ 2 ^ m, 1:]))" unfolding graeffe_poly_def proof (induct as) case (Cons a as) have "?case = ((smult (c ^ 2 ^ m) ([:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:] * (\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) = smult (-1 * (- 1) ^ length as) (smult (c ^ 2 ^ m) ([: a ^ 2 ^ m, 1:] * (\a\as. [:a ^ 2 ^ m, 1:])))))" unfolding list.simps prod_list.Cons pcompose_smult pcompose_mult by simp also have "smult (c ^ 2 ^ m) ([:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:] * (\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) = smult (c ^ 2 ^ m) ((\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) * [:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:]" unfolding mult_smult_left by simp also have "smult (c ^ 2 ^ m) ((\a\as. [:- (a ^ 2 ^ m), 1:]) \\<^sub>p [:0, - 1:]) = smult ((- 1) ^ length as) (smult (c ^ 2 ^ m) (\a\as. [:a ^ 2 ^ m, 1:]))" unfolding pcompose_smult[symmetric] Cons .. also have "[:- (a ^ 2 ^ m), 1:] \\<^sub>p [:0, - 1:] = smult (-1) [: a^2^m, 1:]" by simp finally have id: "?case = (smult ((- 1) ^ length as) (smult (c ^ 2 ^ m) (\a\as. [:a ^ 2 ^ m, 1:])) * smult (- 1) [:a ^ 2 ^ m, 1:] = smult (- 1 * (- 1) ^ length as) (smult (c ^ 2 ^ m) ([:a ^ 2 ^ m, 1:] * (\a\as. [:a ^ 2 ^ m, 1:]))))" by simp obtain c d where id': "(\a\as. [:a ^ 2 ^ m, 1:]) = c" "[:a ^ 2 ^ m, 1:] = d" by auto show ?case unfolding id unfolding id' by (simp add: ac_simps) qed simp finally have "f * f \\<^sub>p [:0, - 1:] = smult ((- 1) ^ length as * (c ^ 2 ^ m * c ^ 2 ^ m)) ((\a\as. [:- (a ^ 2 ^ m), 1:]) * (\a\as. [:a ^ 2 ^ m, 1:]))" unfolding graeffe_poly_def by (simp add: ac_simps) also have "c ^ 2 ^ m * c ^ 2 ^ m = c ^ 2 ^ (Suc m)" by (simp add: semiring_normalization_rules(36)) also have "(\a\as. [:- (a ^ 2 ^ m), 1:]) * (\a\as. [:a ^ 2 ^ m, 1:]) = (\a\as. [:- (a ^ 2 ^ (Suc m)), 1:]) \\<^sub>p monom 1 2" proof (induct as) case (Cons a as) have id: "(monom 1 2 :: 'a poly) = [:0,0,1:]" by (metis monom_altdef pCons_0_as_mult power2_eq_square smult_1_left) have "(\a\a # as. [:- (a ^ 2 ^ m), 1:]) * (\a\a # as. [:a ^ 2 ^ m, 1:]) = ([:- (a ^ 2 ^ m), 1:] * [: a ^ 2 ^ m, 1:]) * ((\a\ as. [:- (a ^ 2 ^ m), 1:]) * (\a\ as. [:a ^ 2 ^ m, 1:]))" (is "_ = ?a * ?b") unfolding list.simps prod_list.Cons by (simp only: ac_simps) also have "?b = (\a\as. [:- (a ^ 2 ^ Suc m), 1:]) \\<^sub>p monom 1 2" unfolding Cons by simp also have "?a = [: - (a ^ 2 ^ (Suc m)), 0 , 1:]" by (simp add: semiring_normalization_rules(36)) also have "\ = [: - (a ^ 2 ^ (Suc m)), 1:] \\<^sub>p monom 1 2" by (simp add: id) also have "[: - (a ^ 2 ^ (Suc m)), 1:] \\<^sub>p monom 1 2 * (\a\as. [:- (a ^ 2 ^ Suc m), 1:]) \\<^sub>p monom 1 2 = (\a\a # as. [:- (a ^ 2 ^ Suc m), 1:]) \\<^sub>p monom 1 2" unfolding pcompose_mult[symmetric] by simp finally show ?case . qed simp finally have "f * f \\<^sub>p [:0, - 1:] = (smult ((- 1) ^ length as) (graeffe_poly c as (Suc m)) \\<^sub>p monom 1 2)" unfolding graeffe_poly_def pcompose_smult by simp from arg_cong[OF this, of "\ f. smult ((- 1) ^ length as) (poly_square_subst f)", unfolded poly_square_subst] have "graeffe_poly c as (Suc m) = smult ((- 1) ^ length as) (poly_square_subst (f * f \\<^sub>p [:0, - 1:]))" by simp also have "\ = smult ((- 1) ^ degree f) (poly_square_subst (f * f \\<^sub>p [:0, - 1:]))" proof (cases "f = 0") case True thus ?thesis by (auto simp: poly_square_subst_def) next case False with assms have c0: "c \ 0" unfolding graeffe_poly_def by auto from arg_cong[OF assms, of degree] have "degree f = degree (smult (c ^ 2 ^ m) (\a\as. [:- (a ^ 2 ^ m), 1:]))" unfolding graeffe_poly_def by auto also have "\ = degree (\a\as. [:- (a ^ 2 ^ m), 1:])" unfolding degree_smult_eq using c0 by auto also have "\ = length as" unfolding degree_linear_factors by simp finally show ?thesis by simp qed finally show ?thesis . qed end definition graeffe_one_step :: "'a \ 'a :: idom poly \ 'a poly" where "graeffe_one_step c f = smult c (poly_square_subst (f * f \\<^sub>p [:0,-1:]))" lemma graeffe_one_step_code[code]: fixes c :: "'a :: idom" shows "graeffe_one_step c f = (case poly_even_odd f of (g,h) \ smult c (g * g - monom 1 1 * h * h))" proof - obtain g h where eo: "poly_even_odd f = (g,h)" by force from poly_even_odd[OF eo] have fgh: "f = g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2 " by auto have m2: "monom (1 :: 'a) 2 = [:0,0,1:]" "monom (1 :: 'a) 1 = [:0,1:]" unfolding coeffs_eq_iff coeffs_monom by (auto simp add: numeral_2_eq_2) show ?thesis unfolding eo split graeffe_one_step_def proof (rule arg_cong[of _ _ "smult c"]) let ?g = "g \\<^sub>p monom 1 2" let ?h = "h \\<^sub>p monom 1 2" let ?x = "monom (1 :: 'a) 1" have 2: "2 = Suc (Suc 0)" by simp have "f * f \\<^sub>p [:0, - 1:] = (g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2) * (g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2) \\<^sub>p [:0, - 1:]" unfolding fgh by simp also have "(g \\<^sub>p monom 1 2 + monom 1 1 * h \\<^sub>p monom 1 2) \\<^sub>p [:0, - 1:] = g \\<^sub>p (monom 1 2 \\<^sub>p [:0, - 1:]) + monom 1 1 \\<^sub>p [:0, - 1:] * h \\<^sub>p (monom 1 2 \\<^sub>p [:0, - 1:])" unfolding pcompose_add pcompose_mult pcompose_assoc by simp also have "monom (1 :: 'a) 2 \\<^sub>p [:0, - 1:] = monom 1 2" unfolding m2 by auto also have "?x \\<^sub>p [:0, - 1:] = [:0, -1:]" unfolding m2 by auto also have "[:0, - 1:] * h \\<^sub>p monom 1 2 = (-?x) * ?h" unfolding m2 by simp also have "(?g + ?x * ?h) * (?g + (- ?x) * ?h) = (?g * ?g - (?x * ?x) * ?h * ?h)" by (auto simp: field_simps) also have "?x * ?x = ?x \\<^sub>p monom 1 2" unfolding mult_monom by (insert m2, simp add: 2) also have "(?g * ?g - \ * ?h * ?h) = (g * g - ?x * h * h) \\<^sub>p monom 1 2" unfolding pcompose_diff pcompose_mult by auto finally have "poly_square_subst (f * f \\<^sub>p [:0, - 1:]) = poly_square_subst ((g * g - ?x * h * h) \\<^sub>p monom 1 2)" by simp also have "\ = g * g - ?x * h * h" unfolding poly_square_subst by simp finally show "poly_square_subst (f * f \\<^sub>p [:0, - 1:]) = g * g - ?x * h * h" . qed qed fun graeffe_poly_impl_main :: "'a \ 'a :: idom poly \ nat \ 'a poly" where "graeffe_poly_impl_main c f 0 = f" | "graeffe_poly_impl_main c f (Suc m) = graeffe_one_step c (graeffe_poly_impl_main c f m)" lemma graeffe_poly_impl_main: assumes "f = smult c (\a\as. [:- a, 1:])" shows "graeffe_poly_impl_main ((-1)^degree f) f m = graeffe_poly c as m" proof (induct m) case 0 show ?case using graeffe_0[OF assms] by simp next case (Suc m) have [simp]: "degree (graeffe_poly c as m) = degree f" unfolding graeffe_poly_def degree_smult_eq assms degree_linear_factors by auto from arg_cong[OF Suc, of degree] show ?case unfolding graeffe_recursion[OF Suc[symmetric]] by (simp add: graeffe_one_step_def) qed definition graeffe_poly_impl :: "'a :: idom poly \ nat \ 'a poly" where "graeffe_poly_impl f = graeffe_poly_impl_main ((-1)^(degree f)) f" lemma graeffe_poly_impl: assumes "f = smult c (\a\as. [:- a, 1:])" shows "graeffe_poly_impl f m = graeffe_poly c as m" using graeffe_poly_impl_main[OF assms] unfolding graeffe_poly_impl_def . lemma drop_half_map: "drop_half (map f xs) = map f (drop_half xs)" by (induct xs rule: drop_half.induct, auto) lemma (in inj_comm_ring_hom) map_poly_poly_square_subst: "map_poly hom (poly_square_subst f) = poly_square_subst (map_poly hom f)" unfolding poly_square_subst_def coeffs_map_poly_hom drop_half_map poly_of_list_def by (rule poly_eqI, auto simp: nth_default_map_eq) context inj_idom_hom begin lemma graeffe_poly_impl_hom: "map_poly hom (graeffe_poly_impl f m) = graeffe_poly_impl (map_poly hom f) m" proof - interpret mh: map_poly_inj_idom_hom.. obtain c where c: "(((- 1) ^ degree f) :: 'a) = c" by auto have c': "(((- 1) ^ degree f) :: 'b) = hom c" unfolding c[symmetric] by (simp add:hom_distribs) show ?thesis unfolding graeffe_poly_impl_def degree_map_poly_hom c c' apply (induct m arbitrary: f; simp) by (unfold graeffe_one_step_def hom_distribs map_poly_poly_square_subst map_poly_pcompose,simp) qed end lemma graeffe_poly_impl_mahler: "mahler_measure (graeffe_poly_impl f m) = mahler_measure f ^ 2 ^ m" proof - let ?c = "complex_of_int" let ?cc = "map_poly ?c" let ?f = "?cc f" note eq = complex_roots(1)[of ?f] interpret inj_idom_hom complex_of_int by (standard, auto) show ?thesis unfolding mahler_measure_def mahler_graeffe[OF eq[symmetric], symmetric] graeffe_poly_impl[OF eq[symmetric], symmetric] by (simp add: of_int_hom.graeffe_poly_impl_hom) qed definition mahler_landau_graeffe_approximation :: "nat \ nat \ int poly \ int" where "mahler_landau_graeffe_approximation kk dd f = (let no = sum_list (map (\ a. a * a) (coeffs f)) in root_int_floor kk (dd * no))" lemma mahler_landau_graeffe_approximation_core: assumes g: "g = graeffe_poly_impl f k" shows "mahler_measure f \ root (2 ^ Suc k) (real_of_int (\a\coeffs g. a * a))" proof - have "mahler_measure f = root (2^k) (mahler_measure f ^ (2^k))" by (simp add: real_root_power_cancel mahler_measure_ge_0) also have "\ = root (2^k) (mahler_measure g)" unfolding graeffe_poly_impl_mahler g by simp also have "\ = root (2^k) (root 2 (((mahler_measure g)^2)))" by (simp add: real_root_power_cancel mahler_measure_ge_0) also have "\ = root (2^Suc k) (((mahler_measure g)^2))" by (metis power_Suc2 real_root_mult_exp) also have "\ \ root (2 ^ Suc k) (real_of_int (\a\coeffs g. a * a))" proof (rule real_root_le_mono, force) have square_mono: "0 \ (x :: real) \ x \ y \ x * x \ y * y" for x y by (simp add: mult_mono') obtain gs where gs: "coeffs g = gs" by auto have "(mahler_measure g)\<^sup>2 \ real_of_int \\a\coeffs g. a * a\" using square_mono[OF mahler_measure_ge_0 Landau_inequality[of "of_int_poly g", folded mahler_measure_def]] by (auto simp: power2_eq_square coeffs_map_poly o_def of_int_hom.hom_sum_list) also have "\\a\coeffs g. a * a\ = (\a\coeffs g. a * a)" unfolding gs by (induct gs, auto) finally show "(mahler_measure g)\<^sup>2 \ real_of_int (\a\coeffs g. a * a)" . qed finally show "mahler_measure f \ root (2 ^ Suc k) (real_of_int (\a\coeffs g. a * a))" . qed lemma Landau_inequality_mahler_measure: "mahler_measure f \ sqrt (real_of_int (\a\coeffs f. a * a))" by (rule order.trans[OF mahler_landau_graeffe_approximation_core[OF refl, of _ 0]], auto simp: graeffe_poly_impl_def sqrt_def) lemma mahler_landau_graeffe_approximation: assumes g: "g = graeffe_poly_impl f k" "dd = d^(2^(Suc k))" "kk = 2^(Suc k)" shows "\real d * mahler_measure f\ \ mahler_landau_graeffe_approximation kk dd g" proof - have id1: "real_of_int (int (d ^ 2 ^ Suc k)) = (real d) ^ 2 ^ Suc k" by simp have id2: "root (2 ^ Suc k) (real d ^ 2 ^ Suc k) = real d" by (simp add: real_root_power_cancel) show ?thesis unfolding mahler_landau_graeffe_approximation_def Let_def root_int_floor of_int_mult g(2-3) by (rule floor_mono, unfold real_root_mult id1 id2, rule mult_left_mono, rule mahler_landau_graeffe_approximation_core[OF g(1)], auto) qed context fixes bnd :: nat begin (* "dd = d^(2^(Suc k))" "kk = 2^(Suc k)" *) function mahler_approximation_main :: "nat \ int \ int poly \ int \ nat \ nat \ int" where "mahler_approximation_main dd c g mm k kk = (let mmm = mahler_landau_graeffe_approximation kk dd g; new_mm = (if k = 0 then mmm else min mm mmm) in (if k \ bnd then new_mm else \ \abort after \bnd\ iterations of Graeffe transformation\ mahler_approximation_main (dd * dd) c (graeffe_one_step c g) new_mm (Suc k) (2 * kk)))" by pat_completeness auto termination by (relation "measure (\ (dd,c,f,mm,k,kk). Suc bnd - k)", auto) declare mahler_approximation_main.simps[simp del] lemma mahler_approximation_main: assumes "k \ 0 \ \real d * mahler_measure f\ \ mm" and "c = (-1)^(degree f)" and "g = graeffe_poly_impl_main c f k" "dd = d^(2^(Suc k))" "kk = 2^(Suc k)" shows "\real d * mahler_measure f\ \ mahler_approximation_main dd c g mm k kk" using assms proof (induct c g mm k kk rule: mahler_approximation_main.induct) case (1 dd c g mm k kk) let ?df = "\real d * mahler_measure f\" note dd = 1(5) note kk = 1(6) note g = 1(4) note c = 1(3) note mm = 1(2) note IH = 1(1) note mahl = mahler_approximation_main.simps[of dd c g mm k kk] define mmm where "mmm = mahler_landau_graeffe_approximation kk dd g" define new_mm where "new_mm = (if k = 0 then mmm else min mm mmm)" let ?cond = "bnd \ k" have id: "mahler_approximation_main dd c g mm k kk = (if ?cond then new_mm else mahler_approximation_main (dd * dd) c (graeffe_one_step c g) new_mm (Suc k) (2 * kk))" unfolding mahl mmm_def[symmetric] Let_def new_mm_def[symmetric] by simp have gg: "g = (graeffe_poly_impl f k)" unfolding g graeffe_poly_impl_def c .. from mahler_landau_graeffe_approximation[OF gg dd kk, folded mmm_def] have mmm: "?df \ mmm" . with mm have new_mm: "?df \ new_mm" unfolding new_mm_def by auto show ?case proof (cases ?cond) case True show ?thesis unfolding id using True new_mm by auto next case False hence id: "mahler_approximation_main dd c g mm k kk = mahler_approximation_main (dd * dd) c (graeffe_one_step c g) new_mm (Suc k) (2 * kk)" unfolding id by auto have id': "graeffe_one_step c g = graeffe_poly_impl_main c f (Suc k)" unfolding g by simp have "dd * dd = d ^ 2 ^ Suc (Suc k)" "2 * kk = 2 ^ Suc (Suc k)" unfolding dd kk semiring_normalization_rules(26) by auto from IH[OF mmm_def new_mm_def False new_mm c id' this] show ?thesis unfolding id . qed qed definition mahler_approximation :: "nat \ int poly \ int" where "mahler_approximation d f = mahler_approximation_main (d * d) ((-1)^(degree f)) f (-1) 0 2" lemma mahler_approximation: "\real d * mahler_measure f\ \ mahler_approximation d f" unfolding mahler_approximation_def by (rule mahler_approximation_main, auto simp: semiring_normalization_rules(29)) end end diff --git a/thys/Bertrands_Postulate/Bertrand.thy b/thys/Bertrands_Postulate/Bertrand.thy --- a/thys/Bertrands_Postulate/Bertrand.thy +++ b/thys/Bertrands_Postulate/Bertrand.thy @@ -1,1853 +1,1853 @@ (* File: Bertrand.thy Authors: Julian Biendarra, Manuel Eberl , Larry Paulson A proof of Bertrand's postulate (based on John Harrison's HOL Light proof). Uses reflection and the approximation tactic. *) theory Bertrand imports Complex_Main "HOL-Number_Theory.Number_Theory" "HOL-Library.Discrete" "HOL-Decision_Procs.Approximation_Bounds" "HOL-Library.Code_Target_Numeral" Pratt_Certificate.Pratt_Certificate begin subsection \Auxiliary facts\ lemma ln_2_le: "ln 2 \ 355 / (512 :: real)" proof - have "ln 2 \ real_of_float (ub_ln2 12)" by (rule ub_ln2) also have "ub_ln2 12 = Float 5680 (- 13)" by code_simp finally show ?thesis by simp qed lemma ln_2_ge: "ln 2 \ (5677 / 8192 :: real)" proof - have "ln 2 \ real_of_float (lb_ln2 12)" by (rule lb_ln2) also have "lb_ln2 12 = Float 5677 (-13)" by code_simp finally show ?thesis by simp qed lemma ln_2_ge': "ln (2 :: real) \ 2/3" and ln_2_le': "ln (2 :: real) \ 16/23" using ln_2_le ln_2_ge by simp_all lemma of_nat_ge_1_iff: "(of_nat x :: 'a :: linordered_semidom) \ 1 \ x \ 1" using of_nat_le_iff[of 1 x] by (subst (asm) of_nat_1) lemma floor_conv_div_nat: "of_int (floor (real m / real n)) = real (m div n)" by (subst floor_divide_of_nat_eq) simp lemma frac_conv_mod_nat: "frac (real m / real n) = real (m mod n) / real n" by (cases "n = 0") (simp_all add: frac_def floor_conv_div_nat field_simps of_nat_mult [symmetric] of_nat_add [symmetric] del: of_nat_mult of_nat_add) lemma of_nat_prod_mset: "prod_mset (image_mset of_nat A) = of_nat (prod_mset A)" by (induction A) simp_all lemma prod_mset_pos: "(\x :: 'a :: linordered_semidom. x \# A \ x > 0) \ prod_mset A > 0" by (induction A) simp_all lemma ln_msetprod: assumes "\x. x \#I \ x > 0" shows "(\p::nat\#I. ln p) = ln (\p\#I. p)" using assms by (induction I) (simp_all add: of_nat_prod_mset ln_mult prod_mset_pos) lemma ln_fact: "ln (fact n) = (\d=1..n. ln d)" by (induction n) (simp_all add: ln_mult) lemma overpower_lemma: fixes f g :: "real \ real" assumes "f a \ g a" assumes "\x. a \ x \ ((\x. g x - f x) has_real_derivative (d x)) (at x)" assumes "\x. a \ x \ d x \ 0" assumes "a \ x" shows "f x \ g x" proof (cases "a < x") case True with assms have "\z. z > a \ z < x \ g x - f x - (g a - f a) = (x - a) * d z" by (intro MVT2) auto then obtain z where z: "z > a" "z < x" "g x - f x - (g a - f a) = (x - a) * d z" by blast hence "f x = g x + (f a - g a) + (a - x) * d z" by (simp add: algebra_simps) also from assms have "f a - g a \ 0" by (simp add: algebra_simps) also from assms z have "(a - x) * d z \ 0 * d z" by (intro mult_right_mono) simp_all finally show ?thesis by simp qed (insert assms, auto) subsection \Preliminary definitions\ definition primepow_even :: "nat \ bool" where "primepow_even q \ (\ p k. 1 \ k \ prime p \ q = p^(2*k))" definition primepow_odd :: "nat \ bool" where "primepow_odd q \ (\ p k. 1 \ k \ prime p \ q = p^(2*k+1))" abbreviation (input) isprimedivisor :: "nat \ nat \ bool" where "isprimedivisor q p \ prime p \ p dvd q" definition pre_mangoldt :: "nat \ nat" where "pre_mangoldt d = (if primepow d then aprimedivisor d else 1)" definition mangoldt_even :: "nat \ real" where "mangoldt_even d = (if primepow_even d then ln (real (aprimedivisor d)) else 0)" definition mangoldt_odd :: "nat \ real" where "mangoldt_odd d = (if primepow_odd d then ln (real (aprimedivisor d)) else 0)" definition mangoldt_1 :: "nat \ real" where "mangoldt_1 d = (if prime d then ln d else 0)" definition psi :: "nat \ real" where "psi n = (\d=1..n. mangoldt d)" definition psi_even :: "nat \ real" where "psi_even n = (\d=1..n. mangoldt_even d)" definition psi_odd :: "nat \ real" where "psi_odd n = (\d=1..n. mangoldt_odd d)" abbreviation (input) psi_even_2 :: "nat \ real" where "psi_even_2 n \ (\d=2..n. mangoldt_even d)" abbreviation (input) psi_odd_2 :: "nat \ real" where "psi_odd_2 n \ (\d=2..n. mangoldt_odd d)" definition theta :: "nat \ real" where "theta n = (\p=1..n. if prime p then ln (real p) else 0)" subsection \Properties of prime powers\ lemma primepow_even_imp_primepow: assumes "primepow_even n" shows "primepow n" proof - from assms obtain p k where "1 \ k" "prime p" "n = p ^ (2 * k)" unfolding primepow_even_def by blast moreover from \1 \ k\ have "2 * k > 0" by simp ultimately show ?thesis unfolding primepow_def by blast qed lemma primepow_odd_imp_primepow: assumes "primepow_odd n" shows "primepow n" proof - from assms obtain p k where "1 \ k" "prime p" "n = p ^ (2 * k + 1)" unfolding primepow_odd_def by blast moreover from \1 \ k\ have "Suc (2 * k) > 0" by simp ultimately show ?thesis unfolding primepow_def by (auto simp del: power_Suc) qed lemma primepow_odd_altdef: "primepow_odd n \ primepow n \ odd (multiplicity (aprimedivisor n) n) \ multiplicity (aprimedivisor n) n > 1" proof (intro iffI conjI; (elim conjE)?) assume "primepow_odd n" then obtain p k where n: "k \ 1" "prime p" "n = p ^ (2 * k + 1)" by (auto simp: primepow_odd_def) thus "odd (multiplicity (aprimedivisor n) n)" "multiplicity (aprimedivisor n) n > 1" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) next assume A: "primepow n" and B: "odd (multiplicity (aprimedivisor n) n)" and C: "multiplicity (aprimedivisor n) n > 1" from A obtain p k where n: "k \ 1" "prime p" "n = p ^ k" by (auto simp: primepow_def Suc_le_eq) with B C have "odd k" "k > 1" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) then obtain j where j: "k = 2 * j + 1" "j > 0" by (auto elim!: oddE) with n show "primepow_odd n" by (auto simp: primepow_odd_def intro!: exI[of _ p, OF exI[of _ j]]) qed (auto dest: primepow_odd_imp_primepow) lemma primepow_even_altdef: "primepow_even n \ primepow n \ even (multiplicity (aprimedivisor n) n)" proof (intro iffI conjI; (elim conjE)?) assume "primepow_even n" then obtain p k where n: "k \ 1" "prime p" "n = p ^ (2 * k)" by (auto simp: primepow_even_def) thus "even (multiplicity (aprimedivisor n) n)" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) next assume A: "primepow n" and B: "even (multiplicity (aprimedivisor n) n)" from A obtain p k where n: "k \ 1" "prime p" "n = p ^ k" by (auto simp: primepow_def Suc_le_eq) with B have "even k" by (simp_all add: aprimedivisor_primepow prime_elem_multiplicity_mult_distrib) then obtain j where j: "k = 2 * j" by (auto elim!: evenE) from j n have "j \ 0" by (intro notI) simp_all with j n show "primepow_even n" by (auto simp: primepow_even_def intro!: exI[of _ p, OF exI[of _ j]]) qed (auto dest: primepow_even_imp_primepow) lemma primepow_odd_mult: assumes "d > Suc 0" shows "primepow_odd (aprimedivisor d * d) \ primepow_even d" using assms by (auto simp: primepow_odd_altdef primepow_even_altdef primepow_mult_aprimedivisorI aprimedivisor_primepow prime_aprimedivisor' aprimedivisor_dvd' prime_elem_multiplicity_mult_distrib prime_elem_aprimedivisor_nat dest!: primepow_multD) lemma pre_mangoldt_primepow: assumes "primepow n" "aprimedivisor n = p" shows "pre_mangoldt n = p" using assms by (simp add: pre_mangoldt_def) lemma pre_mangoldt_notprimepow: assumes "\primepow n" shows "pre_mangoldt n = 1" using assms by (simp add: pre_mangoldt_def) lemma primepow_cases: "primepow d \ ( primepow_even d \ \ primepow_odd d \ \ prime d) \ (\ primepow_even d \ primepow_odd d \ \ prime d) \ (\ primepow_even d \ \ primepow_odd d \ prime d)" by (auto simp: primepow_even_altdef primepow_odd_altdef multiplicity_aprimedivisor_Suc_0_iff elim!: oddE intro!: Nat.gr0I) subsection \Deriving a recurrence for the psi function\ lemma ln_fact_bounds: assumes "n > 0" shows "abs(ln (fact n) - n * ln n + n) \ 1 + ln n" proof - have "\n\{0<..}. \z>real n. z < real (n + 1) \ real (n + 1) * ln (real (n + 1)) - real n * ln (real n) = (real (n + 1) - real n) * (ln z + 1)" by (intro ballI MVT2) (auto intro!: derivative_eq_intros) hence "\n\{0<..}. \z>real n. z < real (n + 1) \ real (n + 1) * ln (real (n + 1)) - real n * ln (real n) = (ln z + 1)" by (simp add: algebra_simps) from bchoice[OF this] obtain k :: "nat \ real" where lb: "real n < k n" and ub: "k n < real (n + 1)" and mvt: "real (n+1) * ln (real (n+1)) - real n * ln (real n) = ln (k n) + 1" if "n > 0" for n::nat by blast have *: "(n + 1) * ln (n + 1) = (\i=1..n. ln(k i) + 1)" for n::nat proof (induction n) case (Suc n) have "(\i = 1..n+1. ln (k i) + 1) = (\i = 1..n. ln (k i) + 1) + ln (k (n+1)) + 1" by simp also from Suc.IH have "(\i = 1..n. ln (k i) + 1) = real (n+1) * ln (real (n+1))" .. also from mvt[of "n+1"] have "\ = real (n+2) * ln (real (n+2)) - ln (k (n+1)) - 1" by simp finally show ?case by simp qed simp have **: "abs((\i=1..n+1. ln i) - ((n+1) * ln (n+1) - (n+1))) \ 1 + ln(n+1)" for n::nat proof - have "(\i=1..n+1. ln i) \ (\i=1..n. ln i) + ln (n+1)" by simp also have "(\i=1..n. ln i) \ (\i=1..n. ln (k i))" by (intro sum_mono, subst ln_le_cancel_iff) (auto simp: Suc_le_eq dest: lb ub) also have "\ = (\i=1..n. ln (k i) + 1) - n" by (simp add: sum.distrib) also from * have "\ = (n+1) * ln (n+1) - n" by simp finally have a_minus_b: "(\i=1..n+1. ln i) - ((n+1) * ln (n+1) - (n+1)) \ 1 + ln (n+1)" by simp from * have "(n+1) * ln (n+1) - n = (\i=1..n. ln (k i) + 1) - n" by simp also have "\ = (\i=1..n. ln (k i))" by (simp add: sum.distrib) also have "\ \ (\i=1..n. ln (i+1))" by (intro sum_mono, subst ln_le_cancel_iff) (auto simp: Suc_le_eq dest: lb ub) also from sum.shift_bounds_cl_nat_ivl[of "ln" 1 1 n] have "\ = (\i=1+1..n+1. ln i)" .. also have "\ = (\i=1..n+1. ln i)" by (rule sum.mono_neutral_left) auto finally have b_minus_a: "((n+1) * ln (n+1) - (n+1)) - (\i=1..n+1. ln i) \ 1" by simp have "0 \ ln (n+1)" by simp with b_minus_a have "((n+1) * ln (n+1) - (n+1)) - (\i=1..n+1. ln i) \ 1 + ln (n+1)" by linarith with a_minus_b show ?thesis by linarith qed from \n > 0\ have "n \ 1" by simp thus ?thesis proof (induction n rule: dec_induct) case base then show ?case by simp next case (step n) from ln_fact[of "n+1"] **[of n] show ?case by simp qed qed lemma ln_fact_diff_bounds: "abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2) \ 4 * ln (if n = 0 then 1 else n) + 3" proof (cases "n div 2 = 0") case True hence "n \ 1" by simp with ln_le_minus_one[of "2::real"] show ?thesis by (cases n) simp_all next case False then have "n > 1" by simp let ?a = "real n * ln 2" let ?b = "4 * ln (real n) + 3" let ?l1 = "ln (fact (n div 2))" let ?a1 = "real (n div 2) * ln (real (n div 2)) - real (n div 2)" let ?b1 = "1 + ln (real (n div 2))" let ?l2 = "ln (fact n)" let ?a2 = "real n * ln (real n) - real n" let ?b2 = "1 + ln (real n)" have abs_a: "abs(?a - (?a2 - 2 * ?a1)) \ ?b - 2 * ?b1 - ?b2" proof (cases "even n") case True then have "real (2 * (n div 2)) = real n" by simp then have n_div_2: "real (n div 2) = real n / 2" by simp from \n > 1\ have *: "abs(?a - (?a2 - 2 * ?a1)) = 0" by (simp add: n_div_2 ln_div algebra_simps) from \even n\ and \n > 1\ have "0 \ ln (real n) - ln (real (n div 2))" by (auto elim: evenE) also have "2 * \ \ 3 * ln (real n) - 2 * ln (real (n div 2))" using \n > 1\ by (auto intro!: ln_ge_zero) also have "\ = ?b - 2 * ?b1 - ?b2" by simp finally show ?thesis using * by simp next case False then have "real (2 * (n div 2)) = real (n - 1)" by simp with \n > 1\ have n_div_2: "real (n div 2) = (real n - 1) / 2" by simp from \odd n\ \n div 2 \ 0\ have "n \ 3" by presburger have "?a - (?a2 - 2 * ?a1) = real n * ln 2 - real n * ln (real n) + real n + 2 * real (n div 2) * ln (real (n div 2)) - 2* real (n div 2)" by (simp add: algebra_simps) also from n_div_2 have "2 * real (n div 2) = real n - 1" by simp also have "real n * ln 2 - real n * ln (real n) + real n + (real n - 1) * ln (real (n div 2)) - (real n - 1) = real n * (ln (real n - 1) - ln (real n)) - ln (real (n div 2)) + 1" using \n > 1\ by (simp add: algebra_simps n_div_2 ln_div) finally have lhs: "abs(?a - (?a2 - 2 * ?a1)) = abs(real n * (ln (real n - 1) - ln (real n)) - ln (real (n div 2)) + 1)" by simp from \n > 1\ have "real n * (ln (real n - 1) - ln (real n)) \ 0" by (simp add: algebra_simps mult_left_mono) moreover from \n > 1\ have "ln (real (n div 2)) \ ln (real n)" by simp moreover { have "exp 1 \ (3::real)" by (rule exp_le) also from \n \ 3\ have "\ \ exp (ln (real n))" by simp finally have "ln (real n) \ 1" by simp } ultimately have ub: "real n * (ln (real n - 1) - ln (real n)) - ln(real (n div 2)) + 1 \ 3 * ln (real n) - 2 * ln(real (n div 2))" by simp have mon: "real n' * (ln (real n') - ln (real n' - 1)) \ real n * (ln (real n) - ln (real n - 1))" if "n \ 3" "n' \ n" for n n'::nat proof (rule DERIV_nonpos_imp_nonincreasing[where f = "\x. x * (ln x - ln (x - 1))"]) fix t assume t: "real n \ t" "t \ real n'" with that have "1 / (t - 1) \ ln (1 + 1/(t - 1))" by (intro ln_add_one_self_le_self) simp_all also from t that have "ln (1 + 1/(t - 1)) = ln t- ln (t - 1)" by (simp add: ln_div [symmetric] field_simps) finally have "ln t - ln (t - 1) \ 1 / (t - 1)" . with that t show "\y. ((\x. x * (ln x - ln (x - 1))) has_field_derivative y) (at t) \ y \ 0" by (intro exI[of _ "1 / (1 - t) + ln t - ln (t - 1)"]) (force intro!: derivative_eq_intros simp: field_simps)+ qed (use that in simp_all) from \n > 1\ have "ln 2 = ln (real n) - ln (real n / 2)" by (simp add: ln_div) also from \n > 1\ have "\ \ ln (real n) - ln (real (n div 2))" by simp finally have *: "3*ln 2 + ln(real (n div 2)) \ 3* ln(real n) - 2* ln(real (n div 2))" by simp have "- real n * (ln (real n - 1) - ln (real n)) + ln(real (n div 2)) - 1 = real n * (ln (real n) - ln (real n - 1)) - 1 + ln(real (n div 2))" by (simp add: algebra_simps) also have "real n * (ln (real n) - ln (real n - 1)) \ 3 * (ln 3 - ln (3 - 1))" using mon[OF _ \n \ 3\] by simp also { have "Some (Float 3 (-1)) = ub_ln 1 3" by code_simp from ub_ln(1)[OF this] have "ln 3 \ (1.6 :: real)" by simp also have "1.6 - 1 / 3 \ 2 * (2/3 :: real)" by simp also have "2/3 \ ln (2 :: real)" by (rule ln_2_ge') finally have "ln 3 - 1 / 3 \ 2 * ln (2 :: real)" by simp } hence "3 * (ln 3 - ln (3 - 1)) - 1 \ 3 * ln (2 :: real)" by simp also note * finally have "- real n * (ln (real n - 1) - ln (real n)) + ln(real (n div 2)) - 1 \ 3 * ln (real n) - 2 * ln (real (n div 2))" by simp hence lhs': "abs(real n * (ln (real n - 1) - ln (real n)) - ln(real (n div 2)) + 1) \ 3 * ln (real n) - 2 * ln (real (n div 2))" using ub by simp have rhs: "?b - 2 * ?b1 - ?b2 = 3* ln (real n) - 2 * ln (real (n div 2))" by simp from \n > 1\ have "ln (real (n div 2)) \ 3* ln (real n) - 2* ln (real (n div 2))" by simp with rhs lhs lhs' show ?thesis by simp qed then have minus_a: "-?a \ ?b - 2 * ?b1 - ?b2 - (?a2 - 2 * ?a1)" by simp from abs_a have a: "?a \ ?b - 2 * ?b1 - ?b2 + ?a2 - 2 * ?a1" by (simp) from ln_fact_bounds[of "n div 2"] False have abs_l1: "abs(?l1 - ?a1) \ ?b1" by (simp add: algebra_simps) then have minus_l1: "?a1 - ?l1 \ ?b1" by linarith from abs_l1 have l1: "?l1 - ?a1 \ ?b1" by linarith from ln_fact_bounds[of n] False have abs_l2: "abs(?l2 - ?a2) \ ?b2" by (simp add: algebra_simps) then have l2: "?l2 - ?a2 \ ?b2" by simp from abs_l2 have minus_l2: "?a2 - ?l2 \ ?b2" by simp from minus_a minus_l1 l2 have "?l2 - 2 * ?l1 - ?a \ ?b" by simp moreover from a l1 minus_l2 have "- ?l2 + 2 * ?l1 + ?a \ ?b" by simp ultimately have "abs((?l2 - 2*?l1) - ?a) \ ?b" by simp then show ?thesis by simp qed lemma ln_primefact: assumes "n \ (0::nat)" shows "ln n = (\d=1..n. if primepow d \ d dvd n then ln (aprimedivisor d) else 0)" (is "?lhs = ?rhs") proof - have "?rhs = (\d\{x \ {1..n}. primepow x \ x dvd n}. ln (real (aprimedivisor d)))" unfolding primepow_factors_def by (subst sum.inter_filter [symmetric]) simp_all also have "{x \ {1..n}. primepow x \ x dvd n} = primepow_factors n" using assms by (auto simp: primepow_factors_def dest: dvd_imp_le primepow_gt_Suc_0) finally have *: "(\d\primepow_factors n. ln (real (aprimedivisor d))) = ?rhs" .. from in_prime_factors_imp_prime prime_gt_0_nat have pf_pos: "\p. p\#prime_factorization n \ p > 0" by blast from ln_msetprod[of "prime_factorization n", OF pf_pos] assms have "ln n = (\p\#prime_factorization n. ln p)" by (simp add: of_nat_prod_mset) also from * sum_prime_factorization_conv_sum_primepow_factors[of n ln, OF assms(1)] have "\ = ?rhs" by simp finally show ?thesis . qed context begin private lemma divisors: fixes x d::nat assumes "x \ {1..n}" assumes "d dvd x" shows "\k\{1..n div d}. x = d * k" proof - from assms have "x \ n" by simp then have ub: "x div d \ n div d" by (simp add: div_le_mono \x \ n\) from assms have "1 \ x div d" by (auto elim!: dvdE) with ub have "x div d \ {1..n div d}" by simp with \d dvd x\ show ?thesis by (auto intro!: bexI[of _ "x div d"]) qed lemma ln_fact_conv_mangoldt: "ln (fact n) = (\d=1..n. mangoldt d * floor (n / d))" proof - have *: "(\da=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = (\(da::nat)=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" if d: "d \ {1..n}" for d by (rule sum.mono_neutral_right, insert d) (auto dest: dvd_imp_le) have "(\d=1..n. \da=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = (\d=1..n. \da=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" by (rule sum.cong) (insert *, simp_all) also have "\ = (\da=1..n. \d=1..n. if primepow da \ da dvd d then ln (aprimedivisor da) else 0)" by (rule sum.swap) also have "\ = sum (\d. mangoldt d * floor (n/d)) {1..n}" proof (rule sum.cong) fix d assume d: "d \ {1..n}" have "(\da = 1..n. if primepow d \ d dvd da then ln (real (aprimedivisor d)) else 0) = (\da = 1..n. if d dvd da then mangoldt d else 0)" by (intro sum.cong) (simp_all add: mangoldt_def) also have "\ = mangoldt d * real (card {x. x \ {1..n} \ d dvd x})" by (subst sum.inter_filter [symmetric]) (simp_all add: algebra_simps) also { have "{x. x \ {1..n} \ d dvd x} = {x. \k \{1..n div d}. x=k*d}" proof safe fix x assume "x \ {1..n}" "d dvd x" thus "\k\{1..n div d}. x = k * d" using divisors[of x n d] by auto next fix x k assume k: "k \ {1..n div d}" from k have "k * d \ n div d * d" by (intro mult_right_mono) simp_all also have "n div d * d \ n div d * d + n mod d" by (rule le_add1) also have "\ = n" by simp finally have "k * d \ n" . thus "k * d \ {1..n}" using d k by auto qed auto also have "\ = (\k. k*d) ` {1..n div d}" by fast also have "card \ = card {1..n div d}" by (rule card_image) (simp add: inj_on_def) also have "\ = n div d" by simp also have "... = \n / d\" by (simp add: floor_divide_of_nat_eq) finally have "real (card {x. x \ {1..n} \ d dvd x}) = real_of_int \n / d\" by force } finally show "(\da = 1..n. if primepow d \ d dvd da then ln (real (aprimedivisor d)) else 0) = mangoldt d * real_of_int \real n / real d\" . qed simp_all finally have "(\d=1..n. \da=1..d. if primepow da \ da dvd d then ln (aprimedivisor da) else 0) = sum (\d. mangoldt d * floor (n/d)) {1..n}" . with ln_primefact have "(\d=1..n. ln d) = (\d=1..n. mangoldt d * floor (n/d))" by simp with ln_fact show ?thesis by simp qed end context begin private lemma div_2_mult_2_bds: fixes n d :: nat assumes "d > 0" shows "0 \ \n / d\ - 2 * \(n div 2) / d\" "\n / d\ - 2 * \(n div 2) / d\ \ 1" proof - have "\2::real\ * \(n div 2) / d\ \ \2 * ((n div 2) / d)\" by (rule le_mult_floor) simp_all also from assms have "\ \ \n / d\" by (intro floor_mono) (simp_all add: field_simps) finally show "0 \ \n / d\ - 2 * \(n div 2) / d\" by (simp add: algebra_simps) next have "real (n div d) \ real (2 * ((n div 2) div d) + 1)" by (subst div_mult2_eq [symmetric], simp only: mult.commute, subst div_mult2_eq) simp thus "\n / d\ - 2 * \(n div 2) / d\ \ 1" unfolding of_nat_add of_nat_mult floor_conv_div_nat [symmetric] by simp_all qed private lemma n_div_d_eq_1: "d \ {n div 2 + 1..n} \ \real n / real d\ = 1" by (cases "n = d") (auto simp: field_simps intro: floor_eq) lemma psi_bounds_ln_fact: shows "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n" "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" proof - fix n::nat let ?k = "n div 2" and ?d = "n mod 2" have *: "\?k / d\ = 0" if "d > ?k" for d proof - from that div_less have "0 = ?k div d" by simp also have "\ = \?k / d\" by (rule floor_divide_of_nat_eq [symmetric]) finally show "\?k / d\ = 0" by simp qed have sum_eq: "(\d=1..2*?k+?d. mangoldt d * \?k / d\) = (\d=1..?k. mangoldt d * \?k / d\)" by (intro sum.mono_neutral_right) (auto simp: *) from ln_fact_conv_mangoldt have "ln (fact n) = (\d=1..n. mangoldt d * \n / d\)" . also have "\ = (\d=1..n. mangoldt d * \(2 * (n div 2) + n mod 2) / d\)" by simp also have "\ \ (\d=1..n. mangoldt d * (2 * \?k / d\ + 1))" using div_2_mult_2_bds(2)[of _ n] by (intro sum_mono mult_left_mono, subst of_int_le_iff) (auto simp: algebra_simps mangoldt_nonneg) also have "\ = 2 * (\d=1..n. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by (simp add: algebra_simps sum.distrib sum_distrib_left) also have "\ = 2 * (\d=1..2*?k+?d. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by presburger also from sum_eq have "\ = 2 * (\d=1..?k. mangoldt d * \(n div 2) / d\) + (\d=1..n. mangoldt d)" by presburger also from ln_fact_conv_mangoldt psi_def have "\ = 2 * ln (fact ?k) + psi n" by presburger finally show "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n" by simp next fix n::nat let ?k = "n div 2" and ?d = "n mod 2" from psi_def have "psi n - psi ?k = (\d=1..2*?k+?d. mangoldt d) - (\d=1..?k. mangoldt d)" by presburger also have "\ = sum mangoldt ({1..2 * (n div 2) + n mod 2} - {1..n div 2})" by (subst sum_diff) simp_all also have "\ = (\d\({1..2 * (n div 2) + n mod 2} - {1..n div 2}). (if d \ ?k then 0 else mangoldt d))" by (intro sum.cong) simp_all also have "\ = (\d=1..2*?k+?d. (if d \ ?k then 0 else mangoldt d))" by (intro sum.mono_neutral_left) auto also have "\ = (\d=1..n. (if d \ ?k then 0 else mangoldt d))" by presburger also have "\ = (\d=1..n. (if d \ ?k then mangoldt d * 0 else mangoldt d))" by (intro sum.cong) simp_all also from div_2_mult_2_bds(1) have "\ \ (\d=1..n. (if d \ ?k then mangoldt d * (\n/d\ - 2 * \?k/d\) else mangoldt d))" by (intro sum_mono) (auto simp: algebra_simps mangoldt_nonneg intro!: mult_left_mono simp del: of_int_mult) also from n_div_d_eq_1 have "\ = (\d=1..n. (if d \ ?k then mangoldt d * (\n/d\ - 2 * \?k/d\) else mangoldt d * \n/d\))" by (intro sum.cong refl) auto also have "\ = (\d=1..n. mangoldt d * real_of_int (\real n / real d\) - (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (intro sum.cong refl) (auto simp: algebra_simps) also have "\ = (\d=1..n. mangoldt d * real_of_int (\real n / real d\)) - (\d=1..n. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (rule sum_subtractf) also have "(\d=1..n. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0)) = (\d=1..?k. (if d \ ?k then 2 * mangoldt d * real_of_int \real ?k / real d\ else 0))" by (intro sum.mono_neutral_right) auto also have "\ = (\d=1..?k. 2 * mangoldt d * real_of_int \real ?k / real d\)" by (intro sum.cong) simp_all also have "\ = 2 * (\d=1..?k. mangoldt d * real_of_int \real ?k / real d\)" by (simp add: sum_distrib_left mult_ac) also have "(\d = 1..n. mangoldt d * real_of_int \real n / real d\) - \ = ln (fact n) - 2 * ln (fact (n div 2))" by (simp add: ln_fact_conv_mangoldt) finally show "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" . qed end lemma psi_bounds_induct: "real n * ln 2 - (4 * ln (real (if n = 0 then 1 else n)) + 3) \ psi n" "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" proof - from le_imp_neg_le[OF ln_fact_diff_bounds] have "n * ln 2 - (4 * ln (if n = 0 then 1 else n) + 3) \ n * ln 2 - abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2)" by simp also have "\ \ ln (fact n) - 2 * ln (fact (n div 2))" by simp also from psi_bounds_ln_fact (1) have "\ \ psi n" by simp finally show "real n * ln 2 - (4 * ln (real (if n = 0 then 1 else n)) + 3) \ psi n" . next from psi_bounds_ln_fact (2) have "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" . also have "\ \ n * ln 2 + abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2)" by simp also from ln_fact_diff_bounds [of n] have "abs(ln (fact n) - 2 * ln (fact (n div 2)) - n * ln 2) \ (4 * ln (real (if n = 0 then 1 else n)) + 3)" by simp finally show "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" by simp qed subsection \Bounding the psi function\ text \ In this section, we will first prove the relatively tight estimate @{prop "psi n \ 3 / 2 + ln 2 * n"} for @{term "n \ 128"} and then use the recurrence we have just derived to extend it to @{prop "psi n \ 551 / 256"} for @{term "n \ 1024"}, at which point applying the recurrence can be used to prove the same bound for arbitrarily big numbers. First of all, we will prove the bound for @{term "n <= 128"} using reflection and approximation. \ context begin private lemma Ball_insertD: assumes "\x\insert y A. P x" shows "P y" "\x\A. P x" using assms by auto private lemma meta_eq_TrueE: "PROP A \ Trueprop True \ PROP A" by simp private lemma pre_mangoldt_pos: "pre_mangoldt n > 0" unfolding pre_mangoldt_def by (auto simp: primepow_gt_Suc_0) private lemma psi_conv_pre_mangoldt: "psi n = ln (real (prod pre_mangoldt {1..n}))" by (auto simp: psi_def mangoldt_def pre_mangoldt_def ln_prod primepow_gt_Suc_0 intro!: sum.cong) private lemma eval_psi_aux1: "psi 0 = ln (real (numeral Num.One))" by (simp add: psi_def) private lemma eval_psi_aux2: assumes "psi m = ln (real (numeral x))" "pre_mangoldt n = y" "m + 1 = n" "numeral x * y = z" shows "psi n = ln (real z)" proof - from assms(2) [symmetric] have [simp]: "y > 0" by (simp add: pre_mangoldt_pos) have "psi n = psi (Suc m)" by (simp add: assms(3) [symmetric]) also have "\ = ln (real y * (\x = Suc 0..m. real (pre_mangoldt x)))" using assms(2,3) [symmetric] by (simp add: psi_conv_pre_mangoldt prod.nat_ivl_Suc' mult_ac) also have "\ = ln (real y) + psi m" by (subst ln_mult) (simp_all add: pre_mangoldt_pos prod_pos psi_conv_pre_mangoldt) also have "psi m = ln (real (numeral x))" by fact also have "ln (real y) + \ = ln (real (numeral x * y))" by (simp add: ln_mult) finally show ?thesis by (simp add: assms(4) [symmetric]) qed private lemma Ball_atLeast0AtMost_doubleton: assumes "psi 0 \ 3 / 2 * ln 2 * real 0" assumes "psi 1 \ 3 / 2 * ln 2 * real 1" shows "(\x\{0..1}. psi x \ 3 / 2 * ln 2 * real x)" using assms unfolding One_nat_def atLeast0_atMost_Suc ball_simps by auto private lemma Ball_atLeast0AtMost_insert: assumes "(\x\{0..m}. psi x \ 3 / 2 * ln 2 * real x)" assumes "psi (numeral n) \ 3 / 2 * ln 2 * real (numeral n)" "m = pred_numeral n" shows "(\x\{0..numeral n}. psi x \ 3 / 2 * ln 2 * real x)" using assms by (subst numeral_eq_Suc[of n], subst atLeast0_atMost_Suc, subst ball_simps, simp only: numeral_eq_Suc [symmetric]) private lemma eval_psi_ineq_aux: assumes "psi n = x" "x \ 3 / 2 * ln 2 * n" shows "psi n \ 3 / 2 * ln 2 * n" using assms by simp_all private lemma eval_psi_ineq_aux2: assumes "numeral m ^ 2 \ (2::nat) ^ (3 * n)" shows "ln (real (numeral m)) \ 3 / 2 * ln 2 * real n" proof - have "ln (real (numeral m)) \ 3 / 2 * ln 2 * real n \ 2 * log 2 (real (numeral m)) \ 3 * real n" by (simp add: field_simps log_def) also have "2 * log 2 (real (numeral m)) = log 2 (real (numeral m ^ 2))" by (subst of_nat_power, subst log_nat_power) simp_all also have "\ \ 3 * real n \ real ((numeral m) ^ 2) \ 2 powr real (3 * n)" by (subst Transcendental.log_le_iff) simp_all also have "2 powr (3 * n) = real (2 ^ (3 * n))" by (simp add: powr_realpow [symmetric]) also have "real ((numeral m) ^ 2) \ \ \ numeral m ^ 2 \ (2::nat) ^ (3 * n)" by (rule of_nat_le_iff) finally show ?thesis using assms by blast qed private lemma eval_psi_ineq_aux_mono: assumes "psi n = x" "psi m = x" "psi n \ 3 / 2 * ln 2 * n" "n \ m" shows "psi m \ 3 / 2 * ln 2 * m" proof - from assms have "psi m = psi n" by simp also have "\ \ 3 / 2 * ln 2 * n" by fact also from \n \ m\ have "\ \ 3 / 2 * ln 2 * m" by simp finally show ?thesis . qed lemma not_primepow_1_nat: "\primepow (1 :: nat)" by auto ML_file \bertrand.ML\ (* This should not take more than 1 minute *) local_setup \fn ctxt => let fun tac {context = ctxt, ...} = let val psi_cache = Bertrand.prove_psi ctxt 129 fun prove_psi_ineqs ctxt = let fun tac {context = ctxt, ...} = HEADGOAL (resolve_tac ctxt @{thms eval_psi_ineq_aux2} THEN' Simplifier.simp_tac ctxt) fun prove_by_approx n thm = let val thm = thm RS @{thm eval_psi_ineq_aux} val [prem] = Thm.prems_of thm val prem = Goal.prove ctxt [] [] prem tac in prem RS thm end fun prove_by_mono last_thm last_thm' thm = let val thm = @{thm eval_psi_ineq_aux_mono} OF [last_thm, thm, last_thm'] val [prem] = Thm.prems_of thm val prem = Goal.prove ctxt [] [] prem (K (HEADGOAL (Simplifier.simp_tac ctxt))) in prem RS thm end fun go _ acc [] = acc | go last acc ((n, x, thm) :: xs) = let val thm' = case last of NONE => prove_by_approx n thm | SOME (last_x, last_thm, last_thm') => if last_x = x then prove_by_mono last_thm last_thm' thm else prove_by_approx n thm in go (SOME (x, thm, thm')) (thm' :: acc) xs end in rev o go NONE [] end val psi_ineqs = prove_psi_ineqs ctxt psi_cache fun prove_ball ctxt (thm1 :: thm2 :: thms) = let val thm = @{thm Ball_atLeast0AtMost_doubleton} OF [thm1, thm2] fun solve_prem thm = let fun tac {context = ctxt, ...} = HEADGOAL (Simplifier.simp_tac ctxt) val thm' = Goal.prove ctxt [] [] (Thm.cprem_of thm 1 |> Thm.term_of) tac in thm' RS thm end fun go thm thm' = (@{thm Ball_atLeast0AtMost_insert} OF [thm', thm]) |> solve_prem in fold go thms thm end | prove_ball _ _ = raise Match in HEADGOAL (resolve_tac ctxt [prove_ball ctxt psi_ineqs]) end val thm = Goal.prove @{context} [] [] @{prop "\n\{0..128}. psi n \ 3 / 2 * ln 2 * n"} tac in Local_Theory.note ((@{binding psi_ubound_log_128}, []), [thm]) ctxt |> snd end \ end context begin private lemma psi_ubound_aux: defines "f \ \x::real. (4 * ln x + 3) / (ln 2 * x)" assumes "x \ 2" "x \ y" shows "f x \ f y" using assms(3) proof (rule DERIV_nonpos_imp_nonincreasing, goal_cases) case (1 t) define f' where "f' = (\x. (1 - 4 * ln x) / x^2 / ln 2 :: real)" from 1 assms(2) have "(f has_real_derivative f' t) (at t)" unfolding f_def f'_def by (auto intro!: derivative_eq_intros simp: field_simps power2_eq_square) moreover { from ln_2_ge have "1/4 \ ln (2::real)" by simp also from assms(2) 1 have "\ \ ln t" by simp finally have "ln t \ 1/4" . } with 1 assms(2) have "f' t \ 0" by (simp add: f'_def field_simps) ultimately show ?case by (intro exI[of _ "f' t"]) simp_all qed text \ These next rules are used in combination with @{thm psi_bounds_induct} and @{thm psi_ubound_log_128} to extend the upper bound for @{term "psi"} from values no greater than 128 to values no greater than 1024. The constant factor of the upper bound changes every time, but once we have reached 1024, the recurrence is self-sustaining in the sense that we do not have to adjust the constant factor anymore in order to double the range. \ lemma psi_ubound_log_double_cases': assumes "\n. n \ m \ psi n \ c * ln 2 * real n" "n \ m'" "m' = 2*m" "c \ c'" "c \ 0" "m \ 1" "c' \ 1 + c/2 + (4 * ln (m+1) + 3) / (ln 2 * (m+1))" shows "psi n \ c' * ln 2 * real n" proof (cases "n > m") case False hence "psi n \ c * ln 2 * real n" by (intro assms) simp_all also have "c \ c'" by fact finally show ?thesis by - (simp_all add: mult_right_mono) next case True hence n: "n \ m+1" by simp from psi_bounds_induct(2)[of n] True have "psi n \ real n * ln 2 + 4 * ln (real n) + 3 + psi (n div 2)" by simp also from assms have "psi (n div 2) \ c * ln 2 * real (n div 2)" by (intro assms) simp_all also have "real (n div 2) \ real n / 2" by simp also have "c * ln 2 * \ = c / 2 * ln 2 * real n" by simp also have "real n * ln 2 + 4 * ln (real n) + 3 + \ = (1 + c/2) * ln 2 * real n + (4 * ln (real n) + 3)" by (simp add: field_simps) also { have "(4 * ln (real n) + 3) / (ln 2 * (real n)) \ (4 * ln (m+1) + 3) / (ln 2 * (m+1))" using n assms by (intro psi_ubound_aux) simp_all also from assms have "(4 * ln (m+1) + 3) / (ln 2 * (m+1)) \ c' - 1 - c/2" by (simp add: algebra_simps) finally have "4 * ln (real n) + 3 \ (c' - 1 - c/2) * ln 2 * real n" using n by (simp add: field_simps) } also have "(1 + c / 2) * ln 2 * real n + (c' - 1 - c / 2) * ln 2 * real n = c' * ln 2 * real n" by (simp add: field_simps) finally show ?thesis using \c \ 0\ by (simp_all add: mult_left_mono) qed end lemma psi_ubound_log_double_cases: assumes "\n\m. psi n \ c * ln 2 * real n" "c' \ 1 + c/2 + (4 * ln (m+1) + 3) / (ln 2 * (m+1))" "m' = 2*m" "c \ c'" "c \ 0" "m \ 1" shows "\n\m'. psi n \ c' * ln 2 * real n" using assms(1) by (intro allI impI assms psi_ubound_log_double_cases'[of m c _ m' c']) auto lemma psi_ubound_log_1024: "\n\1024. psi n \ 551 / 256 * ln 2 * real n" proof - from psi_ubound_log_128 have "\n\128. psi n \ 3 / 2 * ln 2 * real n" by simp hence "\n\256. psi n \ 1025 / 512 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 624 (- 7)) = ub_ln 9 129" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all hence "\n\512. psi n \ 549 / 256 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 180 (- 5)) = ub_ln 7 257" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all thus "\n\1024. psi n \ 551 / 256 * ln 2 * real n" proof (rule psi_ubound_log_double_cases, goal_cases) case 1 have "Some (Float 203 (- 5)) = ub_ln 7 513" by code_simp from ub_ln(1)[OF this] and ln_2_ge show ?case by (simp add: field_simps) qed simp_all qed lemma psi_bounds_sustained_induct: assumes "4 * ln (1 + 2 ^ j) + 3 \ d * ln 2 * (1 + 2^j)" assumes "4 / (1 + 2^j) \ d * ln 2" assumes "0 \ c" assumes "c / 2 + d + 1 \ c" assumes "j \ k" assumes "\n. n \ 2^k \ psi n \ c * ln 2 * n" assumes "n \ 2^(Suc k)" shows "psi n \ c * ln 2 * n" proof (cases "n \ 2^k") case True with assms(6) show ?thesis . next case False from psi_bounds_induct(2) have "psi n - psi (n div 2) \ real n * ln 2 + (4 * ln (real (if n = 0 then 1 else n)) + 3)" . also from False have "(if n = 0 then 1 else n) = n" by simp finally have "psi n \ real n * ln 2 + (4 * ln (real n) + 3) + psi (n div 2)" by simp also from assms(6,7) have "psi (n div 2) \ c * ln 2 * (n div 2)" by simp also have "real (n div 2) \ real n / 2" by simp also have "real n * ln 2 + (4 * ln (real n) + 3) + c * ln 2 * (n / 2) \ c * ln 2 * real n" proof (rule overpower_lemma[of "\x. x * ln 2 + (4 * ln x + 3) + c * ln 2 * (x / 2)" "1+2^j" "\x. c * ln 2 * x" "\x. c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2" "real n"]) from assms(1) have "4 * ln (1 + 2^j) + 3 \ d * ln 2 * (1 + 2^j)" . also from assms(4) have "d \ c - c/2 - 1" by simp also have "(\) * ln 2 * (1 + 2 ^ j) = c * ln 2 * (1 + 2 ^ j) - c / 2 * ln 2 * (1 + 2 ^ j) - (1 + 2 ^ j) * ln 2" by (simp add: left_diff_distrib) finally have "4 * ln (1 + 2^j) + 3 \ c * ln 2 * (1 + 2 ^ j) - c / 2 * ln 2 * (1 + 2 ^ j) - (1 + 2 ^ j) * ln 2" by (simp add: add_pos_pos) then show "(1 + 2 ^ j) * ln 2 + (4 * ln (1 + 2 ^ j) + 3) + c * ln 2 * ((1 + 2 ^ j) / 2) \ c * ln 2 * (1 + 2 ^ j)" by simp next fix x::real assume x: "1 + 2^j \ x" moreover have "1 + 2 ^ j > (0::real)" by (simp add: add_pos_pos) ultimately have x_pos: "x > 0" by linarith show "((\x. c * ln 2 * x - (x * ln 2 + (4 * ln x + 3) + c * ln 2 * (x / 2))) has_real_derivative c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2) (at x)" by (rule derivative_eq_intros refl | simp add: \0 < x\)+ from \0 < x\ \0 < 1 + 2^j\ have "0 < x * (1 + 2^j)" by (rule mult_pos_pos) have "4 / x \ 4 / (1 + 2^j)" by (intro divide_left_mono mult_pos_pos add_pos_pos x x_pos) simp_all also from assms(2) have "4 / (1 + 2^j) \ d * ln 2" . also from assms(4) have "d \ c - c/2 - 1" by simp also have "\ * ln 2 = c * ln 2 - c/2 * ln 2 - ln 2" by (simp add: algebra_simps) finally show "0 \ c * ln 2 - ln 2 - 4 / x - c / 2 * ln 2" by simp next have "1 + 2^j = real (1 + 2^j)" by simp also from assms(5) have "\ \ real (1 + 2^k)" by simp also from False have "2^k \ n - 1" by simp finally show "1 + 2^j \ real n" using False by simp qed finally show ?thesis using assms by - (simp_all add: mult_left_mono) qed lemma psi_bounds_sustained: assumes "\n. n \ 2^k \ psi n \ c * ln 2 * n" assumes "4 * ln (1 + 2^k) + 3 \ (c/2 - 1) * ln 2 * (1 + 2^k)" assumes "4 / (1 + 2^k) \ (c/2 - 1) * ln 2" assumes "c \ 0" shows "psi n \ c * ln 2 * n" proof - - have *: "psi n \ c * ln 2 * n" if "n \ 2^j" for j n + have "psi n \ c * ln 2 * n" if "n \ 2^j" for j n using that proof (induction j arbitrary: n) case 0 with assms(4) 0 show ?case unfolding psi_def mangoldt_def by (cases n) auto next case (Suc j) show ?case proof (cases "k \ j") case True from assms(4) have c_div_2: "c/2 + (c/2 - 1) + 1 \ c" by simp from psi_bounds_sustained_induct[of k "c/2 -1" c j, OF assms(2) assms(3) assms(4) c_div_2 True Suc.IH Suc.prems] show ?thesis by simp next case False then have j_lt_k: "Suc j \ k" by simp from Suc.prems have "n \ 2 ^ Suc j" . also have "(2::nat) ^ Suc j \ 2 ^ k" using power_increasing[of "Suc j" k "2::nat", OF j_lt_k] by simp finally show ?thesis using assms(1) by simp qed qed - have "n < 2 ^ n" by (induction n) simp_all - with *[of n n] show ?thesis by simp + from less_exp this [of n n] show ?thesis by simp qed lemma psi_ubound_log: "psi n \ 551 / 256 * ln 2 * n" proof (rule psi_bounds_sustained) show "0 \ 551 / (256 :: real)" by simp next fix n :: nat assume "n \ 2 ^ 10" with psi_ubound_log_1024 show "psi n \ 551 / 256 * ln 2 * real n" by auto next have "4 / (1 + 2 ^ 10) \ (551 / 256 / 2 - 1) * (2/3 :: real)" by simp also have "\ \ (551 / 256 / 2 - 1) * ln 2" by (intro mult_left_mono ln_2_ge') simp_all finally show "4 / (1 + 2 ^ 10) \ (551 / 256 / 2 - 1) * ln (2 :: real)" . next have "Some (Float 16 (-1)) = ub_ln 3 1025" by code_simp from ub_ln(1)[OF this] and ln_2_ge have "2048 * ln 1025 + 1536 \ 39975 * (ln 2::real)" by simp thus "4 * ln (1 + 2 ^ 10) + 3 \ (551 / 256 / 2 - 1) * ln 2 * (1 + 2 ^ 10 :: real)" by simp qed lemma psi_ubound_3_2: "psi n \ 3/2 * n" proof - have "(551 / 256) * ln 2 \ (551 / 256) * (16/23 :: real)" by (intro mult_left_mono ln_2_le') auto also have "\ \ 3 / 2" by simp finally have "551 / 256 * ln 2 \ 3/(2::real)" . with of_nat_0_le_iff mult_right_mono have "551 / 256 * ln 2 * n \ 3/2 * n" by blast with psi_ubound_log[of "n"] show ?thesis by linarith qed subsection \Doubling psi and theta\ lemma psi_residues_compare_2: "psi_odd_2 n \ psi_even_2 n" proof - have "psi_odd_2 n = (\d\{d. d \ {2..n} \ primepow_odd d}. mangoldt_odd d)" unfolding mangoldt_odd_def by (rule sum.mono_neutral_right) auto also have "\ = (\d\{d. d \ {2..n} \ primepow_odd d}. ln (real (aprimedivisor d)))" by (intro sum.cong refl) (simp add: mangoldt_odd_def) also have "\ \ (\d\{d. d \ {2..n} \ primepow_even d}. ln (real (aprimedivisor d)))" proof (rule sum_le_included [where i = "\y. y * aprimedivisor y"]; clarify?) fix d :: nat assume "d \ {2..n}" "primepow_odd d" note d = this then obtain p k where d': "k \ 1" "prime p" "d = p ^ (2*k+1)" by (auto simp: primepow_odd_def) from d' have "p ^ (2 * k) \ p ^ (2 * k + 1)" by (subst power_increasing_iff) (auto simp: prime_gt_Suc_0_nat) also from d d' have "\ \ n" by simp finally have "p ^ (2 * k) \ n" . moreover from d' have "p ^ (2 * k) > 1" by (intro one_less_power) (simp_all add: prime_gt_Suc_0_nat) ultimately have "p ^ (2 * k) \ {2..n}" by simp moreover from d' have "primepow_even (p ^ (2 * k))" by (auto simp: primepow_even_def) ultimately show "\y\{d \ {2..n}. primepow_even d}. y * aprimedivisor y = d \ ln (real (aprimedivisor d)) \ ln (real (aprimedivisor y))" using d' by (intro bexI[of _ "p ^ (2 * k)"]) (auto simp: aprimedivisor_prime_power aprimedivisor_primepow) qed (simp_all add: of_nat_ge_1_iff Suc_le_eq) also have "\ = (\d\{d. d \ {2..n} \ primepow_even d}. mangoldt_even d)" by (intro sum.cong refl) (simp add: mangoldt_even_def) also have "\ = psi_even_2 n" unfolding mangoldt_even_def by (rule sum.mono_neutral_left) auto finally show ?thesis . qed lemma psi_residues_compare: "psi_odd n \ psi_even n" proof - have "\ primepow_odd 1" by (simp add: primepow_odd_def) hence *: "mangoldt_odd 1 = 0" by (simp add: mangoldt_odd_def) have "\ primepow_even 1" using primepow_gt_Suc_0[OF primepow_even_imp_primepow, of 1] by auto with mangoldt_even_def have **: "mangoldt_even 1 = 0" by simp from psi_odd_def have "psi_odd n = (\d=1..n. mangoldt_odd d)" by simp also from * have "\ = psi_odd_2 n" by (cases "n \ 1") (simp_all add: eval_nat_numeral sum.atLeast_Suc_atMost) also from psi_residues_compare_2 have "\ \ psi_even_2 n" . also from ** have "\ = psi_even n" by (cases "n \ 1") (simp_all add: eval_nat_numeral sum.atLeast_Suc_atMost psi_even_def) finally show ?thesis . qed lemma primepow_iff_even_sqr: "primepow n \ primepow_even (n^2)" by (cases "n = 0") (auto simp: primepow_even_altdef aprimedivisor_primepow_power primepow_power_iff_nat prime_elem_multiplicity_power_distrib prime_aprimedivisor' prime_imp_prime_elem unit_factor_nat_def primepow_gt_0_nat dest: primepow_gt_Suc_0) lemma psi_sqrt: "psi (Discrete.sqrt n) = psi_even n" proof (induction n) case 0 with psi_def psi_even_def show ?case by simp next case (Suc n) then show ?case proof cases assume asm: "\ m. Suc n = m^2" with sqrt_Suc have sqrt_seq: "Discrete.sqrt(Suc n) = Suc(Discrete.sqrt n)" by simp from asm obtain "m" where " Suc n = m^2" by blast with sqrt_seq have "Suc(Discrete.sqrt n) = m" by simp with \Suc n = m^2\ have suc_sqrt_n_sqrt: "(Suc(Discrete.sqrt n))^2 = Suc n" by simp from sqrt_seq have "psi (Discrete.sqrt (Suc n)) = psi (Suc (Discrete.sqrt n))" by simp also from psi_def have "\ = psi (Discrete.sqrt n) + mangoldt (Suc (Discrete.sqrt n))" by simp also from Suc.IH have "psi (Discrete.sqrt n) = psi_even n" . also have "mangoldt (Suc (Discrete.sqrt n)) = mangoldt_even (Suc n)" proof (cases "primepow (Suc(Discrete.sqrt n))") case True with primepow_iff_even_sqr have True2: "primepow_even ((Suc(Discrete.sqrt n))^2)" by simp from suc_sqrt_n_sqrt have "mangoldt_even (Suc n) = mangoldt_even ((Suc(Discrete.sqrt n))^2)" by simp also from mangoldt_even_def True2 have "\ = ln (aprimedivisor ((Suc (Discrete.sqrt n))^2))" by simp also from True have "aprimedivisor ((Suc (Discrete.sqrt n))^2) = aprimedivisor (Suc (Discrete.sqrt n))" by (simp add: aprimedivisor_primepow_power) also from True have "ln (\) = mangoldt (Suc (Discrete.sqrt n))" by (simp add: mangoldt_def) finally show ?thesis .. next case False with primepow_iff_even_sqr have False2: "\ primepow_even ((Suc(Discrete.sqrt n))^2)" by simp from suc_sqrt_n_sqrt have "mangoldt_even (Suc n) = mangoldt_even ((Suc(Discrete.sqrt n))^2)" by simp also from mangoldt_even_def False2 have "\ = 0" by simp also from False have "\ = mangoldt (Suc (Discrete.sqrt n))" by (simp add: mangoldt_def) finally show ?thesis .. qed also from psi_even_def have "psi_even n + mangoldt_even (Suc n) = psi_even (Suc n)" by simp finally show ?case . next assume asm: "\(\m. Suc n = m^2)" with sqrt_Suc have sqrt_eq: "Discrete.sqrt (Suc n) = Discrete.sqrt n" by simp then have lhs: "psi (Discrete.sqrt (Suc n)) = psi (Discrete.sqrt n)" by simp have "\ primepow_even (Suc n)" proof assume "primepow_even (Suc n)" with primepow_even_def obtain "p" "k" where "1 \ k \ prime p \ Suc n = p ^ (2 * k)" by blast with power_even_eq have "Suc n = (p ^ k)^2" by simp with asm show False by blast qed with psi_even_def mangoldt_even_def have rhs: "psi_even (Suc n) = psi_even n" by simp from Suc.IH lhs rhs show ?case by simp qed qed lemma mangoldt_split: "mangoldt d = mangoldt_1 d + mangoldt_even d + mangoldt_odd d" proof (cases "primepow d") case False thus ?thesis by (auto simp: mangoldt_def mangoldt_1_def mangoldt_even_def mangoldt_odd_def dest: primepow_even_imp_primepow primepow_odd_imp_primepow) next case True thus ?thesis by (auto simp: mangoldt_def mangoldt_1_def mangoldt_even_def mangoldt_odd_def primepow_cases) qed lemma psi_split: "psi n = theta n + psi_even n + psi_odd n" by (induction n) (simp_all add: psi_def theta_def psi_even_def psi_odd_def mangoldt_1_def mangoldt_split) lemma psi_mono: "m \ n \ psi m \ psi n" unfolding psi_def by (intro sum_mono2 mangoldt_nonneg) auto lemma psi_pos: "0 \ psi n" by (auto simp: psi_def intro!: sum_nonneg mangoldt_nonneg) lemma mangoldt_odd_pos: "0 \ mangoldt_odd d" using aprimedivisor_gt_Suc_0[of d] by (auto simp: mangoldt_odd_def of_nat_le_iff[of 1, unfolded of_nat_1] Suc_le_eq intro!: ln_ge_zero dest!: primepow_odd_imp_primepow primepow_gt_Suc_0) lemma psi_odd_mono: "m \ n \ psi_odd m \ psi_odd n" using mangoldt_odd_pos sum_mono2[of "{1..n}" "{1..m}" "mangoldt_odd"] by (simp add: psi_odd_def) lemma psi_odd_pos: "0 \ psi_odd n" by (auto simp: psi_odd_def intro!: sum_nonneg mangoldt_odd_pos) lemma psi_theta: "theta n + psi (Discrete.sqrt n) \ psi n" "psi n \ theta n + 2 * psi (Discrete.sqrt n)" using psi_odd_pos[of n] psi_residues_compare[of n] psi_sqrt[of n] psi_split[of n] by simp_all context begin private lemma sum_minus_one: "(\x \ {1..y}. (- 1 :: real) ^ (x + 1)) = (if odd y then 1 else 0)" by (induction y) simp_all private lemma div_invert: fixes x y n :: nat assumes "x > 0" "y > 0" "y \ n div x" shows "x \ n div y" proof - from assms(1,3) have "y * x \ (n div x) * x" by simp also have "\ \ n" by (simp add: minus_mod_eq_div_mult[symmetric]) finally have "y * x \ n" . with assms(2) show ?thesis using div_le_mono[of "y*x" n y] by simp qed lemma sum_expand_lemma: "(\d=1..n. (-1) ^ (d + 1) * psi (n div d)) = (\d = 1..n. (if odd (n div d) then 1 else 0) * mangoldt d)" proof - have **: "x \ n" if "x \ n div y" for x y using div_le_dividend order_trans that by blast have "(\d=1..n. (-1)^(d+1) * psi (n div d)) = (\d=1..n. (-1)^(d+1) * (\e=1..n div d. mangoldt e))" by (simp add: psi_def) also have "\ = (\d = 1..n. \e = 1..n div d. (-1)^(d+1) * mangoldt e)" by (simp add: sum_distrib_left) also from ** have "\ = (\d = 1..n. \e\{y\{1..n}. y \ n div d}. (-1)^(d+1) * mangoldt e)" by (intro sum.cong) auto also have "\ = (\y = 1..n. \x | x \ {1..n} \ y \ n div x. (- 1) ^ (x + 1) * mangoldt y)" by (rule sum.swap_restrict) simp_all also have "\ = (\y = 1..n. \x | x \ {1..n} \ x \ n div y. (- 1) ^ (x + 1) * mangoldt y)" by (intro sum.cong) (auto intro: div_invert) also from ** have "\ = (\y = 1..n. \x \ {1..n div y}. (- 1) ^ (x + 1) * mangoldt y)" by (intro sum.cong) auto also have "\ = (\y = 1..n. (\x \ {1..n div y}. (- 1) ^ (x + 1)) * mangoldt y)" by (intro sum.cong) (simp_all add: sum_distrib_right) also have "\ = (\y = 1..n. (if odd (n div y) then 1 else 0) * mangoldt y)" by (intro sum.cong refl) (simp_all only: sum_minus_one) finally show ?thesis . qed private lemma floor_half_interval: fixes n d :: nat assumes "d \ 0" shows "real (n div d) - real (2 * ((n div 2) div d)) = (if odd (n div d) then 1 else 0)" proof - have "((n div 2) div d) = (n div (2 * d))" by (rule div_mult2_eq[symmetric]) also have "\ = ((n div d) div 2)" by (simp add: mult_ac div_mult2_eq) also have "real (n div d) - real (2 * \) = (if odd (n div d) then 1 else 0)" by (cases "odd (n div d)", cases "n div d = 0 ", simp_all) finally show ?thesis by simp qed lemma fact_expand_psi: "ln (fact n) - 2 * ln (fact (n div 2)) = (\d=1..n. (-1)^(d+1) * psi (n div d))" proof - have "ln (fact n) - 2 * ln (fact (n div 2)) = (\d=1..n. mangoldt d * \n / d\) - 2 * (\d=1..n div 2. mangoldt d * \(n div 2) / d\)" by (simp add: ln_fact_conv_mangoldt) also have "(\d=1..n div 2. mangoldt d * \real (n div 2) / d\) = (\d=1..n. mangoldt d * \real (n div 2) / d\)" by (rule sum.mono_neutral_left) (auto simp: floor_unique[of 0]) also have "2 * \ = (\d=1..n. mangoldt d * 2 * \real (n div 2) / d\)" by (simp add: sum_distrib_left mult_ac) also have "(\d=1..n. mangoldt d * \n / d\) - \ = (\d=1..n. (mangoldt d * \n / d\ - mangoldt d * 2 * \real (n div 2) / d\))" by (simp add: sum_subtractf) also have "\ = (\d=1..n. mangoldt d * (\n / d\ - 2 * \real (n div 2) / d\))" by (simp add: algebra_simps) also have "\ = (\d=1..n. mangoldt d * (if odd(n div d) then 1 else 0))" by (intro sum.cong refl) (simp_all add: floor_conv_div_nat [symmetric] floor_half_interval [symmetric]) also have "\ = (\d=1..n. (if odd(n div d) then 1 else 0) * mangoldt d)" by (simp add: mult_ac) also from sum_expand_lemma[symmetric] have "\ = (\d=1..n. (-1)^(d+1) * psi (n div d))" . finally show ?thesis . qed end lemma psi_expansion_cutoff: assumes "m \ p" shows "(\d=1..2*m. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*p. (-1)^(d+1) * psi (n div d))" "(\d=1..2*p+1. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*m+1. (-1)^(d+1) * psi (n div d))" using assms proof (induction m rule: inc_induct) case (step k) have "(\d = 1..2 * k. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * Suc k. (-1)^(d + 1) * psi (n div d))" by (simp add: psi_mono div_le_mono2) with step.IH(1) show "(\d = 1..2 * k. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * p. (-1)^(d + 1) * psi (n div d))" by simp from step.IH(2) have "(\d = 1..2 * p + 1. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * Suc k + 1. (-1)^(d + 1) * psi (n div d))" . also have "\ \ (\d = 1..2 * k + 1. (-1)^(d + 1) * psi (n div d))" by (simp add: psi_mono div_le_mono2) finally show "(\d = 1..2 * p + 1. (-1)^(d + 1) * psi (n div d)) \ (\d = 1..2 * k + 1. (-1)^(d + 1) * psi (n div d))" . qed simp_all lemma fact_psi_bound_even: assumes "even k" shows "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ ln (fact n) - 2 * ln (fact (n div 2))" proof - have "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" proof (cases "k \ n") case True with psi_expansion_cutoff(1)[of "k div 2" "n div 2" n] have "(\d=1..2*(k div 2). (-1)^(d+1) * psi (n div d)) \ (\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d))" by simp also from assms have "2*(k div 2) = k" by simp also have "(\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d)) \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" proof (cases "even n") case True then show ?thesis by simp next case False from psi_pos have "(\d = 1..2*(n div 2). (- 1) ^ (d + 1) * psi (n div d)) \ (\d = 1..2*(n div 2) + 1. (- 1) ^ (d + 1) * psi (n div d))" by simp with False show ?thesis by simp qed finally show ?thesis . next case False hence *: "n div 2 \ (k-1) div 2" by simp have "(\d=1..k. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*((k-1) div 2) + 1. (-1)^(d+1) * psi (n div d))" proof (cases "k = 0") case True with psi_pos show ?thesis by simp next case False with sum.cl_ivl_Suc[of "\d. (-1)^(d+1) * psi (n div d)" 1 "k-1"] have "(\d=1..k. (-1)^(d+1) * psi (n div d)) = (\d=1..k-1. (-1)^(d+1) * psi (n div d)) + (-1)^(k+1) * psi (n div k)" by simp also from assms psi_pos have "(-1)^(k+1) * psi (n div k) \ 0" by simp also from assms False have "k-1 = 2*((k-1) div 2) + 1" by presburger finally show ?thesis by simp qed also from * psi_expansion_cutoff(2)[of "n div 2" "(k-1) div 2" n] have "\ \ (\d=1..2*(n div 2) + 1. (-1)^(d+1) * psi (n div d))" by blast also have "\ \ (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" by (cases "even n") (simp_all add: psi_def) finally show ?thesis . qed also from fact_expand_psi have "\ = ln (fact n) - 2 * ln (fact (n div 2))" .. finally show ?thesis . qed lemma fact_psi_bound_odd: assumes "odd k" shows "ln (fact n) - 2 * ln (fact (n div 2)) \ (\d=1..k. (-1)^(d+1) * psi (n div d))" proof - from fact_expand_psi have "ln (fact n) - 2 * ln (fact (n div 2)) = (\d = 1..n. (- 1) ^ (d + 1) * psi (n div d))" . also have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" proof (cases "k \ n") case True have "(\d=1..n. (-1)^(d+1) * psi (n div d)) \ ( \d=1..2*(n div 2)+1. (-1)^(d+1) * psi (n div d))" by (cases "even n") (simp_all add: psi_pos) also from True assms psi_expansion_cutoff(2)[of "k div 2" "n div 2" n] have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" by simp finally show ?thesis . next case False have "(\d=1..n. (-1)^(d+1) * psi (n div d)) \ (\d=1..2*((n+1) div 2). (-1)^(d+1) * psi (n div d))" by (cases "even n") (simp_all add: psi_def) also from False assms psi_expansion_cutoff(1)[of "(n+1) div 2" "k div 2" n] have "(\d=1..2*((n+1) div 2). (-1)^(d+1) * psi (n div d)) \ (\d=1..2*(k div 2). (-1)^(d+1) * psi (n div d))" by simp also from assms have "\ \ (\d=1..k. (-1)^(d+1) * psi (n div d))" by (auto elim: oddE simp: psi_pos) finally show ?thesis . qed finally show ?thesis . qed lemma fact_psi_bound_2_3: "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n - psi (n div 2) + psi (n div 3)" proof - show "psi n - psi (n div 2) \ ln (fact n) - 2 * ln (fact (n div 2))" by (rule psi_bounds_ln_fact (2)) next from fact_psi_bound_odd[of 3 n] have "ln (fact n) - 2 * ln (fact (n div 2)) \ (\d = 1..3. (- 1) ^ (d + 1) * psi (n div d))" by simp also have "\ = psi n - psi (n div 2) + psi (n div 3)" by (simp add: sum.atLeast_Suc_atMost numeral_2_eq_2) finally show "ln (fact n) - 2 * ln (fact (n div 2)) \ psi n - psi (n div 2) + psi (n div 3)" . qed lemma ub_ln_1200: "ln 1200 \ 57 / (8 :: real)" proof - have "Some (Float 57 (-3)) = ub_ln 8 1200" by code_simp from ub_ln(1)[OF this] show ?thesis by simp qed lemma psi_double_lemma: assumes "n \ 1200" shows "real n / 6 \ psi n - psi (n div 2)" proof - from ln_fact_diff_bounds have "\ln (fact n) - 2 * ln (fact (n div 2)) - real n * ln 2\ \ 4 * ln (real (if n = 0 then 1 else n)) + 3" . with assms have "ln (fact n) - 2 * ln (fact (n div 2)) \ real n * ln 2 - 4 * ln (real n) - 3" by simp moreover have "real n * ln 2 - 4 * ln (real n) - 3 \ 2 / 3 * n" proof (rule overpower_lemma[of "\n. 2/3 * n" 1200]) show "2 / 3 * 1200 \ 1200 * ln 2 - 4 * ln 1200 - (3::real)" using ub_ln_1200 ln_2_ge by linarith next fix x::real assume "1200 \ x" then have "0 < x" by simp show "((\x. x * ln 2 - 4 * ln x - 3 - 2 / 3 * x) has_real_derivative ln 2 - 4 / x - 2 / 3) (at x)" by (rule derivative_eq_intros refl | simp add: \0 < x\)+ next fix x::real assume "1200 \ x" then have "12 / x \ 12 / 1200" by simp then have "0 \ 0.67 - 4 / x - 2 / 3" by simp also have "0.67 \ ln (2::real)" using ln_2_ge by simp finally show "0 \ ln 2 - 4 / x - 2 / 3" by simp next from assms show "1200 \ real n" by simp qed ultimately have "2 / 3 * real n \ ln (fact n) - 2 * ln (fact (n div 2))" by simp with psi_ubound_3_2[of "n div 3"] have "n/6 + psi (n div 3) \ ln (fact n) - 2 * ln (fact (n div 2))" by simp with fact_psi_bound_2_3[of "n"] show ?thesis by simp qed lemma theta_double_lemma: assumes "n \ 1200" shows "theta (n div 2) < theta n" proof - from psi_theta[of "n div 2"] psi_pos[of "Discrete.sqrt (n div 2)"] have theta_le_psi_n_2: "theta (n div 2) \ psi (n div 2)" by simp have "(Discrete.sqrt n * 18)^2 \ 324 * n" by simp from mult_less_cancel2[of "324" "n" "n"] assms have "324 * n < n^2" by (simp add: power2_eq_square) with \(Discrete.sqrt n * 18)^2 \ 324 * n\ have "(Discrete.sqrt n*18)^2 < n^2" by presburger with power2_less_imp_less assms have "Discrete.sqrt n * 18 < n" by blast with psi_ubound_3_2[of "Discrete.sqrt n"] have "2 * psi (Discrete.sqrt n) < n / 6" by simp with psi_theta[of "n"] have psi_lt_theta_n: "psi n - n / 6 < theta n" by simp from psi_double_lemma[OF assms(1)] have "psi (n div 2) \ psi n - n / 6" by simp with theta_le_psi_n_2 psi_lt_theta_n show ?thesis by simp qed subsection \Proof of the main result\ lemma theta_mono: "mono theta" by (auto simp: theta_def [abs_def] intro!: monoI sum_mono2) lemma theta_lessE: assumes "theta m < theta n" "m \ 1" obtains p where "p \ {m<..n}" "prime p" proof - from mono_invE[OF theta_mono assms(1)] have "m \ n" by blast hence "theta n = theta m + (\p\{m<..n}. if prime p then ln (real p) else 0)" unfolding theta_def using assms(2) by (subst sum.union_disjoint [symmetric]) (auto simp: ivl_disj_un) also note assms(1) finally have "(\p\{m<..n}. if prime p then ln (real p) else 0) \ 0" by simp from sum.not_neutral_contains_not_neutral [OF this] guess p . thus ?thesis using that[of p] by (auto intro!: exI[of _ p] split: if_splits) qed theorem bertrand: fixes n :: nat assumes "n > 1" shows "\p\{n<..<2*n}. prime p" proof cases assume n_less: "n < 600" define prime_constants where "prime_constants = {2, 3, 5, 7, 13, 23, 43, 83, 163, 317, 631::nat}" from \n > 1\ n_less have "\p \ prime_constants. n < p \ p < 2 * n" unfolding bex_simps greaterThanLessThan_iff prime_constants_def by presburger moreover have "\p\prime_constants. prime p" - unfolding prime_constants_def ball_simps HOL.simp_thms by (intro conjI; pratt (silent)) + unfolding prime_constants_def ball_simps HOL.simp_thms + by (intro conjI; pratt (silent)) ultimately show ?thesis unfolding greaterThanLessThan_def greaterThan_def lessThan_def by blast next assume n: "\(n < 600)" from n have "theta n < theta (2 * n)" using theta_double_lemma[of "2 * n"] by simp with assms obtain p where "p \ {n<..2*n}" "prime p" by (auto elim!: theta_lessE) moreover from assms have "\prime (2*n)" by (auto dest!: prime_product) with \prime p\ have "p \ 2 * n" by auto ultimately show ?thesis by auto qed subsection \Proof of Mertens' first theorem\ text \ The following proof of Mertens' first theorem was ported from John Harrison's HOL Light proof by Larry Paulson: \ lemma sum_integral_ubound_decreasing': fixes f :: "real \ real" assumes "m \ n" and der: "\x. x \ {of_nat m - 1..of_nat n} \ (g has_field_derivative f x) (at x)" and le: "\x y. \real m - 1 \ x; x \ y; y \ real n\ \ f y \ f x" shows "(\k = m..n. f (of_nat k)) \ g (of_nat n) - g (of_nat m - 1)" proof - have "(\k = m..n. f (of_nat k)) \ (\k = m..n. g (of_nat(Suc k) - 1) - g (of_nat k - 1))" proof (rule sum_mono, clarsimp) fix r assume r: "m \ r" "r \ n" hence "\z>real r - 1. z < real r \ g (real r) - g (real r - 1) = (real r - (real r - 1)) * f z" using assms by (intro MVT2) auto hence "\z\{of_nat r - 1..of_nat r}. g (real r) - g (real r - 1) = f z" by auto then obtain u::real where u: "u \ {of_nat r - 1..of_nat r}" and eq: "g r - g (of_nat r - 1) = f u" by blast have "real m \ u + 1" using r u by auto then have "f (of_nat r) \ f u" using r(2) and u by (intro le) auto then show "f (of_nat r) \ g r - g (of_nat r - 1)" by (simp add: eq) qed also have "\ \ g (of_nat n) - g (of_nat m - 1)" using \m \ n\ by (subst sum_Suc_diff) auto finally show ?thesis . qed lemma Mertens_lemma: assumes "n \ 0" shows "\(\d = 1..n. mangoldt d / real d) - ln n\ \ 4" proof - have *: "\abs(s' - nl + n) \ a; abs(s' - s) \ (k - 1) * n - a\ \ abs(s - nl) \ n * k" for s' s k nl a::real by (auto simp: algebra_simps abs_if split: if_split_asm) have le: "\(\d=1..n. mangoldt d * floor (n / d)) - n * ln n + n\ \ 1 + ln n" using ln_fact_bounds ln_fact_conv_mangoldt assms by simp have "\real n * ((\d = 1..n. mangoldt d / real d) - ln n)\ = \((\d = 1..n. real n * mangoldt d / real d) - n * ln n)\" by (simp add: algebra_simps sum_distrib_left) also have "\ \ real n * 4" proof (rule * [OF le]) have "\(\d = 1..n. mangoldt d * \n / d\) - (\d = 1..n. n * mangoldt d / d)\ = \\d = 1..n. mangoldt d * (\n / d\ - n / d)\" by (simp add: sum_subtractf algebra_simps) also have "\ \ psi n" (is "\?sm\ \ ?rhs") proof - have "-?sm = (\d = 1..n. mangoldt d * (n/d - \n/d\))" by (simp add: sum_subtractf algebra_simps) also have "\ \ (\d = 1..n. mangoldt d * 1)" by (intro sum_mono mult_left_mono mangoldt_nonneg) linarith+ finally have "-?sm \ ?rhs" by (simp add: psi_def) moreover have "?sm \ 0" using mangoldt_nonneg by (simp add: mult_le_0_iff sum_nonpos) ultimately show ?thesis by (simp add: abs_if) qed also have "\ \ 3/2 * real n" by (rule psi_ubound_3_2) also have "\\ (4 - 1) * real n - (1 + ln n)" using ln_le_minus_one [of n] assms by (simp add: divide_simps) finally show "\(\d = 1..n. mangoldt d * real_of_int \real n / real d\) - (\d = 1..n. real n * mangoldt d / real d)\ \ (4 - 1) * real n - (1 + ln n)" . qed finally have "\real n * ((\d = 1..n. mangoldt d / real d) - ln n)\ \ real n * 4" . then show ?thesis using assms mult_le_cancel_left_pos by (simp add: abs_mult) qed lemma Mertens_mangoldt_versus_ln: assumes "I \ {1..n}" shows "\(\i\I. mangoldt i / i) - (\p | prime p \ p \ I. ln p / p)\ \ 3" (is "\?lhs\ \ 3") proof (cases "n = 0") case True with assms show ?thesis by simp next case False have "finite I" using assms finite_subset by blast have "0 \ (\i\I. mangoldt i / i - (if prime i then ln i / i else 0))" using mangoldt_nonneg by (intro sum_nonneg) simp_all moreover have "\ \ (\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0))" using assms by (intro sum_mono2) (auto simp: mangoldt_nonneg) ultimately have *: "\\i\I. mangoldt i / i - (if prime i then ln i / i else 0)\ \ \\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0)\" by linarith moreover have "?lhs = (\i\I. mangoldt i / i - (if prime i then ln i / i else 0))" "(\i = 1..n. mangoldt i / i - (if prime i then ln i / i else 0)) = (\d = 1..n. mangoldt d / d) - (\p | prime p \ p \ {1..n}. ln p / p)" using sum.inter_restrict [of _ "\i. ln (real i) / i" "Collect prime", symmetric] by (force simp: sum_subtractf \finite I\ intro: sum.cong)+ ultimately have "\?lhs\ \ \(\d = 1..n. mangoldt d / d) - (\p | prime p \ p \ {1..n}. ln p / p)\" by linarith also have "\ \ 3" proof - have eq_sm: "(\i = 1..n. mangoldt i / i) = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 1}. mangoldt i / i)" proof (intro sum.mono_neutral_right ballI, goal_cases) case (3 i) hence "\primepow i" by (auto simp: primepow_def Suc_le_eq) thus ?case by (simp add: mangoldt_def) qed (auto simp: Suc_le_eq prime_gt_0_nat) have "(\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p) = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i)" proof - have eq: "{p ^ k |p k. prime p \ p ^ k \ n \ 1 \ k} = {p ^ k |p k. prime p \ p ^ k \ n \ 2 \ k} \ {p. prime p \ p \ {1..n}}" (is "?A = ?B \ ?C") proof (intro equalityI subsetI; (elim UnE)?) fix x assume "x \ ?A" then obtain p k where "x = p ^ k" "prime p" "p ^ k \ n" "k \ 1" by auto thus "x \ ?B \ ?C" by (cases "k \ 2") (auto simp: prime_power_iff Suc_le_eq) next fix x assume "x \ ?B" then obtain p k where "x = p ^ k" "prime p" "p ^ k \ n" "k \ 1" by auto thus "x \ ?A" by (auto simp: prime_power_iff Suc_le_eq) next fix x assume "x \ ?C" then obtain p where "x = p ^ 1" "1 \ (1::nat)" "prime p" "p ^ 1 \ n" by auto thus "x \ ?A" by blast qed have eqln: "(\p | prime p \ p \ {1..n}. ln p / p) = (\p | prime p \ p \ {1..n}. mangoldt p / p)" by (rule sum.cong) auto have "(\i \ {p^k |p k. prime p \ p^k \ n \ k \ 1}. mangoldt i / i) = (\i \ {p ^ k |p k. prime p \ p ^ k \ n \ 2 \ k} \ {p. prime p \ p \ {1..n}}. mangoldt i / i)" by (subst eq) simp_all also have "\ = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i) + (\p | prime p \ p \ {1..n}. mangoldt p / p)" by (intro sum.union_disjoint) (auto simp: prime_power_iff finite_nat_set_iff_bounded_le) also have "\ = (\i \ {p^k |p k. prime p \ p^k \ n \ k \ 2}. mangoldt i / i) + (\p | prime p \ p \ {1..n}. ln p / p)" by (simp only: eqln) finally show ?thesis using eq_sm by auto qed have "(\p | prime p \ p \ {1..n}. ln p / p) \ (\p | prime p \ p \ {1..n}. mangoldt p / p)" using mangoldt_nonneg by (auto intro: sum_mono) also have "\ \ (\i = Suc 0..n. mangoldt i / i)" by (intro sum_mono2) (auto simp: mangoldt_nonneg) finally have "0 \ (\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p)" by simp moreover have "(\i = 1..n. mangoldt i / i) - (\p | prime p \ p \ {1..n}. ln p / p) \ 3" (is "?M - ?L \ 3") proof - have *: "\q. \j\{1..n}. prime q \ 1 \ q \ q \ n \ (q ^ j = p ^ k \ mangoldt (p ^ k) / real p ^ k \ ln (real q) / real q ^ j)" if "prime p" "p ^ k \ n" "1 \ k" for p k proof - have "mangoldt (p ^ k) / real p ^ k \ ln p / p ^ k" using that by (simp add: divide_simps) moreover have "p \ n" using that self_le_power[of p k] by (simp add: prime_ge_Suc_0_nat) moreover have "k \ n" proof - have "k < 2^k" using of_nat_less_two_power of_nat_less_numeral_power_cancel_iff by blast also have "\ \ p^k" by (simp add: power_mono prime_ge_2_nat that) also have "\ \ n" by (simp add: that) finally show ?thesis by (simp add: that) qed ultimately show ?thesis using prime_ge_1_nat that by auto (use atLeastAtMost_iff in blast) qed have finite: "finite {p ^ k |p k. prime p \ p ^ k \ n \ 1 \ k}" by (rule finite_subset[of _ "{..n}"]) auto have "?M \ (\(x, k)\{p. prime p \ p \ {1..n}} \ {1..n}. ln (real x) / real x ^ k)" by (subst eq_sm, intro sum_le_included [where i = "\(p,k). p^k"]) (insert * finite, auto) also have "\ = (\p | prime p \ p \ {1..n}. (\k = 1..n. ln p / p^k))" by (subst sum.Sigma) auto also have "\ = ?L + (\p | prime p \ p \ {1..n}. (\k = 2..n. ln p / p^k))" by (simp add: comm_monoid_add_class.sum.distrib sum.atLeast_Suc_atMost numeral_2_eq_2) finally have "?M - ?L \ (\p | prime p \ p \ {1..n}. (\k = 2..n. ln p / p^k))" by (simp add: algebra_simps) also have "\ = (\p | prime p \ p \ {1..n}. ln p * (\k = 2..n. inverse p ^ k))" by (simp add: field_simps sum_distrib_left) also have "\ = (\p | prime p \ p \ {1..n}. ln p * (((inverse p)\<^sup>2 - inverse p ^ Suc n) / (1 - inverse p)))" by (intro sum.cong refl) (simp add: sum_gp) also have "\ \ (\p | prime p \ p \ {1..n}. ln p * inverse (real (p * (p - 1))))" by (intro sum_mono mult_left_mono) (auto simp: divide_simps power2_eq_square of_nat_diff mult_less_0_iff) also have "\ \ (\p = 2..n. ln p * inverse (real (p * (p - 1))))" by (rule sum_mono2) (use prime_ge_2_nat in auto) also have "\ \ (\i = 2..n. ln i / (i - 1)\<^sup>2)" unfolding divide_inverse power2_eq_square mult.assoc by (auto intro: sum_mono mult_left_mono mult_right_mono) also have "\ \ 3" proof (cases "n \ 3") case False then show ?thesis proof (cases "n \ 2") case False then show ?thesis by simp next case True then have "n = 2" using False by linarith with ln_le_minus_one [of 2] show ?thesis by simp qed next case True have "(\i = 3..n. ln (real i) / (real (i - Suc 0))\<^sup>2) \ (ln (of_nat n - 1)) - (ln (of_nat n)) - (ln (of_nat n) / (of_nat n - 1)) + 2 * ln 2" proof - have 1: "((\z. ln (z - 1) - ln z - ln z / (z - 1)) has_field_derivative ln x / (x - 1)\<^sup>2) (at x)" if x: "x \ {2..real n}" for x by (rule derivative_eq_intros | rule refl | (use x in \force simp: power2_eq_square divide_simps\))+ have 2: "ln y / (y - 1)\<^sup>2 \ ln x / (x - 1)\<^sup>2" if xy: "2 \ x" "x \ y" "y \ real n" for x y proof (cases "x = y") case False define f' :: "real \ real" where "f' = (\u. ((u - 1)\<^sup>2 / u - ln u * (2 * u - 2)) / (u - 1) ^ 4)" have f'_altdef: "f' u = inverse u * inverse ((u - 1)\<^sup>2) - 2 * ln u / (u - 1) ^ 3" if u: "u \ {x..y}" for u::real unfolding f'_def using u (* TODO ugly *) by (simp add: eval_nat_numeral divide_simps) (simp add: algebra_simps)? have deriv: "((\z. ln z / (z - 1)\<^sup>2) has_field_derivative f' u) (at u)" if u: "u \ {x..y}" for u::real unfolding f'_def by (rule derivative_eq_intros refl | (use u xy in \force simp: divide_simps\))+ hence "\z>x. z < y \ ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2 = (y - x) * f' z" using xy and \x \ y\ by (intro MVT2) auto then obtain \::real where "x < \" "\ < y" and \: "ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2 = (y - x) * f' \" by blast have "f' \ \ 0" proof - have "2/3 \ ln (2::real)" by (fact ln_2_ge') also have "\ \ ln \" using \x < \\ xy by auto finally have "1 \ 2 * ln \" by simp then have *: "\ \ \ * (2 * ln \)" using \x < \\ xy by auto hence "\ - 1 \ ln \ * 2 * \" by (simp add: algebra_simps) hence "1 / (\ * (\ - 1)\<^sup>2) \ ln \ * 2 / (\ - 1) ^ 3" using xy \x < \\ by (simp add: divide_simps power_eq_if) thus ?thesis using xy \x < \\ \\ < y\ by (subst f'_altdef) (auto simp: divide_simps) qed then have "(ln y / (y - 1)\<^sup>2 - ln x / (x - 1)\<^sup>2) \ 0" using \x \ y\ by (simp add: mult_le_0_iff \) then show ?thesis by simp qed simp_all show ?thesis using sum_integral_ubound_decreasing' [OF \3 \ n\, of "\z. ln(z-1) - ln z - ln z / (z - 1)" "\z. ln z / (z-1)\<^sup>2"] 1 2 \3 \ n\ by (auto simp: in_Reals_norm of_nat_diff) qed also have "\ \ 2" proof - have "ln (real n - 1) - ln n \ 0" "0 \ ln n / (real n - 1)" using \3 \ n\ by auto then have "ln (real n - 1) - ln n - ln n / (real n - 1) \ 0" by linarith with ln_2_less_1 show ?thesis by linarith qed also have "\ \ 3 - ln 2" using ln_2_less_1 by (simp add: algebra_simps) finally show ?thesis using True by (simp add: algebra_simps sum.atLeast_Suc_atMost [of 2 n]) qed finally show ?thesis . qed ultimately show ?thesis by linarith qed finally show ?thesis . qed proposition Mertens: assumes "n \ 0" shows "\(\p | prime p \ p \ n. ln p / of_nat p) - ln n\ \ 7" proof - have "\(\d = 1..n. mangoldt d / real d) - (\p | prime p \ p \ {1..n}. ln (real p) / real p)\ \ 7 - 4" using Mertens_mangoldt_versus_ln [of "{1..n}" n] by simp_all also have "{p. prime p \ p \ {1..n}} = {p. prime p \ p \ n}" using atLeastAtMost_iff prime_ge_1_nat by blast finally have "\(\d = 1..n. mangoldt d / real d) - (\p\\. ln (real p) / real p)\ \ 7 - 4" . moreover from assms have "\(\d = 1..n. mangoldt d / real d) - ln n\ \ 4" by (rule Mertens_lemma) ultimately show ?thesis by linarith qed end diff --git a/thys/Chandy_Lamport/Co_Snapshot.thy b/thys/Chandy_Lamport/Co_Snapshot.thy --- a/thys/Chandy_Lamport/Co_Snapshot.thy +++ b/thys/Chandy_Lamport/Co_Snapshot.thy @@ -1,334 +1,332 @@ theory Co_Snapshot imports Snapshot Ordered_Resolution_Prover.Lazy_List_Chain begin section \Extension to infinite traces\ text \The computation locale assumes that there already exists a known final configuration $c'$ to the given initial $c$ and trace $t$. However, we can show that the snapshot algorithm must terminate correctly even if the underlying computation itself does not terminate. We relax the trace relation to allow for a potentially infinite number of ``intermediate'' events, and show that the algorithm's correctness still holds when imposing the same constraints as in the computation locale. We use a preexisting theory of lazy list chains by Schlichtkrull, Blanchette, Traytel and Waldmann~\cite{Ordered_Resolution_Prover-AFP} to construct infinite traces.\ primrec ltake where "ltake 0 t = []" | "ltake (Suc i) t = (case t of LNil \ [] | LCons x t' \ x # ltake i t')" primrec ldrop where "ldrop 0 t = t" | "ldrop (Suc i) t = (case t of LNil \ LNil | LCons x t' \ ldrop i t')" lemma ltake_LNil[simp]: "ltake i LNil = []" by (induct i) auto lemma ltake_LCons: "0 < i \ ltake i (LCons x t) = x # ltake (i - 1) t" by (induct i) auto lemma take_ltake: "i \ j \ take i (ltake j xs) = ltake i xs" by (induct j arbitrary: i xs) (auto simp: le_Suc_eq take_Cons' ltake_LCons split: llist.splits if_splits) lemma nth_ltake [simp]: "i < min n (llength xs) \ (ltake n xs) ! i = lnth xs i" by (induct n arbitrary: i xs) (auto simp: nth_Cons' gr0_conv_Suc eSuc_enat[symmetric] split: llist.splits) lemma length_ltake[simp]: "length (ltake i xs) = (case llength xs of \ \ i | enat m \ min i m)" by (induct i arbitrary: xs) (auto simp: zero_enat_def[symmetric] eSuc_enat split: llist.splits enat.splits) lemma ltake_prepend: "ltake i (prepend xs t) = (if i \ length xs then take i xs else xs @ ltake (i - length xs) t)" proof (induct i arbitrary: xs t) case 0 then show ?case by (cases xs) auto next case (Suc i) then show ?case by (cases xs) auto qed lemma prepend_ltake_ldrop_id: "prepend (ltake i t) (ldrop i t) = t" by (induct i arbitrary: t) (auto split: llist.splits) context distributed_system begin coinductive cotrace where cotr_init: "cotrace c LNil" | cotr_step: "\ c \ ev \ c'; cotrace c' t \ \ cotrace c (LCons ev t)" lemma cotrace_trace: "cotrace c t \ \!c'. trace c (ltake i t) c'" proof (induct i arbitrary: c t) case (Suc i) then show ?case proof (cases t) case (LCons ev t') with Suc(2) obtain c' where "c \ ev \ c'" "cotrace c' t'" by (auto elim: cotrace.cases) with Suc(1)[OF \cotrace c' t'\] show ?thesis by (auto simp: LCons elim: trace.intros(2) elim: trace.cases trace_and_start_determines_end) qed (auto intro: trace.intros elim: trace.cases) qed (auto simp: zero_enat_def[symmetric] intro: trace.intros elim: trace.cases) lemma cotrace_trace': "cotrace c t \ \c'. trace c (ltake i t) c'" by (metis cotrace_trace) definition cos where "cos c t i = s c (ltake i t) i" lemma cotrace_trace_cos: "cotrace c t \ trace c (ltake i t) (cos c t i)" unfolding cos_def s_def by (subst take_ltake, auto dest!: cotrace_trace[of _ _ i] elim!: theI') lemma s_0[simp]: "s c t 0 = c" unfolding s_def by (auto intro!: the_equality[where P = "trace c []"] trace.intros elim: trace.cases) lemma s_chop: "i \ length t \ s c t i = s c (take i t) i" unfolding s_def by auto lemma cotrace_prepend: "trace c t c' \ cotrace c' u \ cotrace c (prepend t u)" by (induct c t c' rule: trace.induct) (auto intro: cotrace.intros) lemma s_Cons: "\c''. trace c' xs c'' \ c \ ev \ c' \ s c (ev # xs) (Suc i) = s c' xs i" by (smt exists_trace_for_any_i take_Suc_Cons tr_step trace_and_start_determines_end) lemma cotrace_ldrop: "cotrace c t \ i \ llength t \ cotrace (cos c t i) (ldrop i t)" proof (induct i arbitrary: c t) case (Suc i) then show ?case proof (cases t) case (LCons ev t') with Suc(2) obtain c' where "c \ ev \ c'" "cotrace c' t'" by (auto elim: cotrace.cases) with Suc(1)[OF \cotrace c' t'\] Suc(3) show ?thesis by (auto simp: LCons cos_def eSuc_enat[symmetric] s_chop[symmetric] s_Cons[OF cotrace_trace']) qed (auto intro: cotrace.intros) qed (auto simp: zero_enat_def[symmetric] cos_def intro: cotrace.intros) end locale cocomputation = distributed_system + fixes init :: "('a, 'b, 'c) configuration" assumes finite_channels: "finite {i. \p q. channel i = Some (p, q)}" and strongly_connected_raw: "\p q. (p \ q) \ (tranclp (\p q. (\i. channel i = Some (p, q)))) p q" and at_least_two_processes: "card (UNIV :: 'a set) > 1" and finite_processes: "finite (UNIV :: 'a set)" and no_initial_Marker: "\i. (\p q. channel i = Some (p, q)) \ Marker \ set (msgs init i)" and no_msgs_if_no_channel: "\i. channel i = None \ msgs init i = []" and no_initial_process_snapshot: "\p. \ has_snapshotted init p" and no_initial_channel_snapshot: "\i. channel_snapshot init i = ([], NotStarted)" and valid: "\t. cotrace init t" and l1: "\t i cid. cotrace init t \ Marker \ set (msgs (cos init t i) cid) \ (\j \ llength t. j \ i \ Marker \ set (msgs (cos init t j) cid))" and l2: "\t p. cotrace init t \ (\i \ llength t. has_snapshotted (cos init t i) p)" begin abbreviation coS where "coS \ cos init" definition "some_snapshot t p = (SOME i. has_snapshotted (coS t i) p \ i \ llength t)" lemma has_snapshotted: "cotrace init t \ has_snapshotted (coS t (some_snapshot t p)) p \ some_snapshot t p \ llength t" unfolding some_snapshot_def by (rule someI_ex) (auto dest!: l2[rule_format]) lemma cotrace_cos: "cotrace init t \ j < llength t \ (coS t j) \ lnth t j \ (coS t (Suc j))" apply (drule cotrace_trace_cos[of _ _ "Suc j"]) apply (drule step_Suc[rotated, of _ _ _ "j"]) apply (auto split: enat.splits llist.splits) [] apply (auto simp: s_chop[of j "_ # ltake j _"] cos_def nth_Cons' ltake_LCons lnth_LCons' take_Cons' take_ltake split: llist.splits enat.splits if_splits elim: order.strict_trans2[rotated]) - apply (subst (asm) nth_ltake) - apply (auto elim!: order.strict_trans2[rotated]) [] apply (subst (asm) s_chop[of j "_ # ltake j _"]) apply (auto simp: take_Cons' take_ltake split: enat.splits) done lemma snapshot_stable: "cotrace init t \ i \ j \ has_snapshotted (coS t i) p \ has_snapshotted (coS t j) p" apply (drule cotrace_trace_cos[of _ _ j]) unfolding cos_def by (metis exists_trace_for_any_i_j order_refl s_def snapshot_stable take_ltake) lemma no_markers_if_all_snapshotted: "cotrace init t \ i \ j \ \p. has_snapshotted (coS t i) p \ Marker \ set (msgs (coS t i) c) \ Marker \ set (msgs (coS t j) c)" apply (drule cotrace_trace_cos[of _ _ j]) unfolding cos_def by (metis exists_trace_for_any_i_j no_markers_if_all_snapshotted order_refl s_def take_ltake) lemma cotrace_all_have_snapshotted: assumes "cotrace init t" shows "\i \ llength t. \p. has_snapshotted (coS t i) p" proof - let ?i = "Max (range (some_snapshot t))" show ?thesis using has_snapshotted[OF assms] snapshot_stable[OF assms, of "some_snapshot t _" ?i _] apply (intro exI[of _ ?i]) apply (auto simp: finite_processes) apply (cases "llength t"; auto simp: ) apply (subst Max_le_iff) apply (auto simp: finite_processes) apply blast done qed lemma no_messages_if_no_channel: assumes "cotrace init t" shows "channel cid = None \ msgs (coS t i) cid = []" using no_messages_introduced_if_no_channel[OF assms[THEN cotrace_trace_cos, of i] no_msgs_if_no_channel, of cid i] by (auto simp: cos_def) lemma cotrace_all_have_snapshotted_and_no_markers: assumes "cotrace init t" shows "\i \ llength t. (\p. has_snapshotted (coS t i) p) \ (\c. Marker \ set (msgs (coS t i) c))" proof - from cotrace_all_have_snapshotted[OF assms] obtain j :: nat where j: "j \ llength t" "\p. has_snapshotted (coS t j) p" by blast from j(2) have *: "has_snapshotted (coS t k) p" if "k \ j" for k p using snapshot_stable[OF assms, of j k p] that by auto define C where "C = {c. Marker \ set (msgs (coS t j) c)}" have "finite C" using no_messages_if_no_channel[OF assms, of _ j] unfolding C_def by (intro finite_subset[OF _ finite_channels]) fastforce define pick where "pick = (\c. SOME k. k \ llength t \ k \ j \ Marker \ set (msgs (coS t k) c))" { fix c assume "c \ C" then have "\k \ llength t. k \ j \ Marker \ set (msgs (coS t k) c)" using l1[rule_format, of t j c] assms unfolding C_def by blast then have "pick c \ llength t \ pick c \ j \ Marker \ set (msgs (coS t (pick c)) c)" unfolding pick_def by (rule someI_ex) } note pick = conjunct1[OF this] conjunct1[OF conjunct2[OF this]] conjunct2[OF conjunct2[OF this]] show ?thesis proof (cases "C = {}") case True with j show ?thesis by (auto intro!: exI[of _ j] simp: C_def) next define m where "m = Max (pick ` C)" case False with \finite C\ have m: "m \ pick ` C" "\x \ pick ` C. m \ x" unfolding m_def by auto then have "j \ m" using pick(2) by auto from *[OF \j \ m\] have "Marker \ set (msgs (coS t m) c)" for c proof (cases "c \ C") case True then show ?thesis using no_markers_if_all_snapshotted[OF assms, of "pick c" m c] pick[of c] m * by auto next case False then show ?thesis using no_markers_if_all_snapshotted[OF assms \j \ m\ j(2), of c] by (auto simp: C_def) qed with *[OF \j \ m\] m pick show ?thesis by auto qed qed context fixes t assumes cotrace: "cotrace init t" begin definition "final_i \ (SOME i. i \ llength t \ (\p. has_snapshotted (coS t i) p) \ (\c. Marker \ set (msgs (coS t i) c)))" definition final where "final = coS t final_i" lemma final_i: "final_i \ llength t" "(\p. has_snapshotted (coS t final_i) p)" "(\c. Marker \ set (msgs (coS t final_i) c))" unfolding final_i_def by (rule someI2_ex[OF cotrace_all_have_snapshotted_and_no_markers[OF cotrace]]; auto intro: cotrace_trace_cos[OF cotrace])+ lemma final: "\t. trace init t final" "(\p. has_snapshotted final p)" "(\c. Marker \ set (msgs final c))" unfolding final_def by (rule cotrace_trace_cos[OF cotrace] final_i exI)+ interpretation computation channel trans send recv init final apply standard apply (rule finite_channels) apply (rule strongly_connected_raw) apply (rule at_least_two_processes) apply (rule finite_processes) apply (rule no_initial_Marker) apply (rule no_msgs_if_no_channel) apply (rule no_initial_process_snapshot) apply (rule no_initial_channel_snapshot) apply (rule final(1)) apply (intro allI impI) subgoal for t i cid apply (rule exI[of _ "length t"]) apply (metis exists_trace_for_any_i final(3) le_cases take_all trace_and_start_determines_end) done apply (intro allI impI) subgoal for t p apply (rule exI[of _ "length t"]) apply (metis exists_trace_for_any_i final(2) order_refl take_all trace_and_start_determines_end) done done definition coperm where "coperm l r = (\xs ys z. mset xs = mset ys \ l = prepend xs z \ r = prepend ys z)" lemma copermIL: "mset ys = mset xs \ t = prepend xs z \ coperm (prepend ys z) t" unfolding coperm_def by auto lemma snapshot_algorithm_is_cocorrect: "\t' i. cotrace init t' \ coperm t' t \ state_equal_to_snapshot (coS t' i) final \ i \ final_i" proof - define prefix where "prefix = ltake final_i t" define suffix where "suffix = ldrop final_i t" have [simp]: "prepend prefix suffix = t" unfolding prefix_def suffix_def prepend_ltake_ldrop_id .. have [simp]: "cotrace final suffix" unfolding suffix_def final_def by (auto simp: cotrace final_i(1) intro!: cotrace_ldrop) from cotrace_trace_cos[OF cotrace] have "trace init prefix final" unfolding final_def prefix_def by blast with snapshot_algorithm_is_correct obtain prefix' i where "trace init prefix' final" "mset prefix' = mset prefix" "state_equal_to_snapshot (S prefix' i) final" "i \ length prefix'" by blast moreover from \mset prefix' = mset prefix\ \i \ length prefix'\ have "i \ final_i" by (auto dest!: mset_eq_length simp: prefix_def split: enat.splits) ultimately show ?thesis by (intro exI[of _ "prepend prefix' suffix"] exI[of _ i]) (auto simp: cos_def ltake_prepend s_chop[symmetric] intro!: cotrace_prepend elim!: copermIL) qed end print_statement snapshot_algorithm_is_cocorrect end end diff --git a/thys/Closest_Pair_Points/Common.thy b/thys/Closest_Pair_Points/Common.thy --- a/thys/Closest_Pair_Points/Common.thy +++ b/thys/Closest_Pair_Points/Common.thy @@ -1,1215 +1,1216 @@ section "Common" theory Common imports "HOL-Library.Going_To_Filter" "Akra_Bazzi.Akra_Bazzi_Method" "Akra_Bazzi.Akra_Bazzi_Approximation" "HOL-Library.Code_Target_Numeral" "Root_Balanced_Tree.Time_Monad" begin type_synonym point = "int * int" subsection "Auxiliary Functions and Lemmas" subsubsection "Time Monad" lemma time_distrib_bind: "time (bind_tm tm f) = time tm + time (f (val tm))" unfolding bind_tm_def by (simp split: tm.split) lemmas time_simps = time_distrib_bind tick_def lemma bind_tm_cong[fundef_cong]: assumes "\v. v = val n \ f v = g v" "m = n" shows "bind_tm m f = bind_tm n g" using assms unfolding bind_tm_def by (auto split: tm.split) subsubsection "Landau Auxiliary" text \ The following lemma expresses a procedure for deriving complexity properties of the form @{prop"t \ O[m going_to at_top within A](f o m)"} where \<^item> \t\ is a (timing) function on same data domain (e.g. lists), \<^item> \m\ is a measure function on that data domain (e.g. length), \<^item> \t'\ is a function on @{typ nat}, \<^item> \A\ is the set of valid inputs for the data domain. One needs to show that \<^item> \t\ is bounded by @{term "t' o m"} for valid inputs \<^item> @{prop"t' \ O(f)"} to conclude the overall property @{prop"t \ O[m going_to at_top within A](f o m)"}. \ lemma bigo_measure_trans: fixes t :: "'a \ real" and t' :: "nat \ real" and m :: "'a \ nat" and f ::"nat \ real" assumes "\x. x \ A \ t x \ (t' o m) x" and "t' \ O(f)" and "\x. x \ A \ 0 \ t x" shows "t \ O[m going_to at_top within A](f o m)" proof - have 0: "\x. x \ A \ 0 \ (t' o m) x" by (meson assms(1,3) order_trans) have 1: "t \ O[m going_to at_top within A](t' o m)" apply(rule bigoI[where c=1]) using assms 0 by (simp add: eventually_inf_principal going_to_within_def) have 2: "t' o m \ O[m going_to at_top](f o m)" unfolding o_def going_to_def by(rule landau_o.big.filtercomap[OF assms(2)]) have 3: "t' o m \ O[m going_to at_top within A](f o m)" using landau_o.big.filter_mono[OF _2] going_to_mono[OF _subset_UNIV] by blast show ?thesis by(rule landau_o.big_trans[OF 1 3]) qed lemma const_1_bigo_n_ln_n: "(\(n::nat). 1) \ O(\n. n * ln n)" proof - have "\N. \(n::nat) \ N. (\x. 1 \ x * ln x) n" proof - have "\(n::nat) \ 3. (\x. 1 \ x * ln x) n" proof standard fix n show "3 \ n \ 1 \ real n * ln (real n)" proof standard assume "3 \ n" hence "1 \ real n" by simp moreover have "1 \ ln (real n)" using ln_ln_nonneg' \3 \ n\ by simp ultimately show "1 \ real n * ln (real n)" by (auto simp: order_trans) qed qed thus ?thesis by blast qed thus ?thesis by auto qed subsubsection "Miscellaneous Lemmas" lemma set_take_drop_i_le_j: "i \ j \ set xs = set (take j xs) \ set (drop i xs)" proof (induction xs arbitrary: i j) case (Cons x xs) show ?case proof (cases "i = 0") case True thus ?thesis using set_take_subset by force next case False hence "set xs = set (take (j - 1) xs) \ set (drop (i - 1) xs)" by (simp add: Cons diff_le_mono) moreover have "set (take j (x # xs)) = insert x (set (take (j - 1) xs))" using False Cons.prems by (auto simp: take_Cons') moreover have "set (drop i (x # xs)) = set (drop (i - 1) xs)" using False Cons.prems by (auto simp: drop_Cons') ultimately show ?thesis by auto qed qed simp lemma set_take_drop: "set xs = set (take n xs) \ set (drop n xs)" using set_take_drop_i_le_j by fast lemma sorted_wrt_take_drop: "sorted_wrt f xs \ \x \ set (take n xs). \y \ set (drop n xs). f x y" using sorted_wrt_append[of f "take n xs" "drop n xs"] by simp lemma sorted_wrt_hd_less: assumes "sorted_wrt f xs" "\x. f x x" shows "\x \ set xs. f (hd xs) x" using assms by (cases xs) auto lemma sorted_wrt_hd_less_take: assumes "sorted_wrt f (x # xs)" "\x. f x x" shows "\y \ set (take n (x # xs)). f x y" - using assms sorted_wrt_hd_less in_set_takeD by fastforce + using assms sorted_wrt_hd_less [of f \x # xs\] in_set_takeD [of _ n \x # xs\] + by auto lemma sorted_wrt_take_less_hd_drop: assumes "sorted_wrt f xs" "n < length xs" shows "\x \ set (take n xs). f x (hd (drop n xs))" using sorted_wrt_take_drop assms by fastforce lemma sorted_wrt_hd_drop_less_drop: assumes "sorted_wrt f xs" "\x. f x x" shows "\x \ set (drop n xs). f (hd (drop n xs)) x" using assms sorted_wrt_drop sorted_wrt_hd_less by blast lemma length_filter_P_impl_Q: "(\x. P x \ Q x) \ length (filter P xs) \ length (filter Q xs)" by (induction xs) auto lemma filter_Un: "set xs = A \ B \ set (filter P xs) = { x \ A. P x } \ { x \ B. P x }" by (induction xs) (auto, metis UnI1 insert_iff, metis UnI2 insert_iff) subsubsection \@{const length}\ fun length_tm :: "'a list \ nat tm" where "length_tm [] =1 return 0" | "length_tm (x # xs) =1 do { l <- length_tm xs; return (1 + l) }" lemma length_eq_val_length_tm: "val (length_tm xs) = length xs" by (induction xs) auto lemma time_length_tm: "time (length_tm xs) = length xs + 1" by (induction xs) (auto simp: time_simps) fun length_it' :: "nat \ 'a list \ nat" where "length_it' acc [] = acc" | "length_it' acc (x#xs) = length_it' (acc+1) xs" definition length_it :: "'a list \ nat" where "length_it xs = length_it' 0 xs" lemma length_conv_length_it': "length xs + acc = length_it' acc xs" by (induction acc xs rule: length_it'.induct) auto lemma length_conv_length_it[code_unfold]: "length xs = length_it xs" unfolding length_it_def using length_conv_length_it' add_0_right by metis subsubsection \@{const rev}\ fun rev_it' :: "'a list \ 'a list \ 'a list" where "rev_it' acc [] = acc" | "rev_it' acc (x#xs) = rev_it' (x#acc) xs" definition rev_it :: "'a list \ 'a list" where "rev_it xs = rev_it' [] xs" lemma rev_conv_rev_it': "rev xs @ acc = rev_it' acc xs" by (induction acc xs rule: rev_it'.induct) auto lemma rev_conv_rev_it[code_unfold]: "rev xs = rev_it xs" unfolding rev_it_def using rev_conv_rev_it' append_Nil2 by metis subsubsection \@{const take}\ fun take_tm :: "nat \ 'a list \ 'a list tm" where "take_tm n [] =1 return []" | "take_tm n (x # xs) =1 (case n of 0 \ return [] | Suc m \ do { ys <- take_tm m xs; return (x # ys) } )" lemma take_eq_val_take_tm: "val (take_tm n xs) = take n xs" by (induction xs arbitrary: n) (auto split: nat.split) lemma time_take_tm: "time (take_tm n xs) = min n (length xs) + 1" by (induction xs arbitrary: n) (auto simp: time_simps split: nat.split) subsubsection \@{const filter}\ fun filter_tm :: "('a \ bool) \ 'a list \ 'a list tm" where "filter_tm P [] =1 return []" | "filter_tm P (x # xs) =1 (if P x then do { ys <- filter_tm P xs; return (x # ys) } else filter_tm P xs )" lemma filter_eq_val_filter_tm: "val (filter_tm P xs) = filter P xs" by (induction xs) auto lemma time_filter_tm: "time (filter_tm P xs) = length xs + 1" by (induction xs) (auto simp: time_simps) fun filter_it' :: "'a list \ ('a \ bool) \ 'a list \ 'a list" where "filter_it' acc P [] = rev acc" | "filter_it' acc P (x#xs) = ( if P x then filter_it' (x#acc) P xs else filter_it' acc P xs )" definition filter_it :: "('a \ bool) \ 'a list \ 'a list" where "filter_it P xs = filter_it' [] P xs" lemma filter_conv_filter_it': "rev acc @ filter P xs = filter_it' acc P xs" by (induction acc P xs rule: filter_it'.induct) auto lemma filter_conv_filter_it[code_unfold]: "filter P xs = filter_it P xs" unfolding filter_it_def using filter_conv_filter_it' append_Nil rev.simps(1) by metis subsubsection \\split_at\\ fun split_at_tm :: "nat \ 'a list \ ('a list \ 'a list) tm" where "split_at_tm n [] =1 return ([], [])" | "split_at_tm n (x # xs) =1 ( case n of 0 \ return ([], x # xs) | Suc m \ do { (xs', ys') <- split_at_tm m xs; return (x # xs', ys') } )" fun split_at :: "nat \ 'a list \ 'a list \ 'a list" where "split_at n [] = ([], [])" | "split_at n (x # xs) = ( case n of 0 \ ([], x # xs) | Suc m \ let (xs', ys') = split_at m xs in (x # xs', ys') )" lemma split_at_eq_val_split_at_tm: "val (split_at_tm n xs) = split_at n xs" by (induction xs arbitrary: n) (auto split: nat.split prod.split) lemma split_at_take_drop_conv: "split_at n xs = (take n xs, drop n xs)" by (induction xs arbitrary: n) (auto simp: split: nat.split) lemma time_split_at_tm: "time (split_at_tm n xs) = min n (length xs) + 1" by (induction xs arbitrary: n) (auto simp: time_simps split: nat.split prod.split) fun split_at_it' :: "'a list \ nat \ 'a list \ ('a list * 'a list)" where "split_at_it' acc n [] = (rev acc, [])" | "split_at_it' acc n (x#xs) = ( case n of 0 \ (rev acc, x#xs) | Suc m \ split_at_it' (x#acc) m xs )" definition split_at_it :: "nat \ 'a list \ ('a list * 'a list)" where "split_at_it n xs = split_at_it' [] n xs" lemma split_at_conv_split_at_it': assumes "(ts, ds) = split_at n xs" "(ts', ds') = split_at_it' acc n xs" shows "rev acc @ ts = ts'" and "ds = ds'" using assms by (induction acc n xs arbitrary: ts rule: split_at_it'.induct) (auto simp: split: prod.splits nat.splits) lemma split_at_conv_split_at_it_prod: assumes "(ts, ds) = split_at n xs" "(ts', ds') = split_at_it n xs" shows "(ts, ds) = (ts', ds')" using assms unfolding split_at_it_def using split_at_conv_split_at_it' rev.simps(1) append_Nil by fast+ lemma split_at_conv_split_at_it[code_unfold]: "split_at n xs = split_at_it n xs" using split_at_conv_split_at_it_prod surj_pair by metis declare split_at_tm.simps [simp del] declare split_at.simps [simp del] subsection "Mergesort" subsubsection "Functional Correctness Proof" definition sorted_fst :: "point list \ bool" where "sorted_fst ps = sorted_wrt (\p\<^sub>0 p\<^sub>1. fst p\<^sub>0 \ fst p\<^sub>1) ps" definition sorted_snd :: "point list \ bool" where "sorted_snd ps = sorted_wrt (\p\<^sub>0 p\<^sub>1. snd p\<^sub>0 \ snd p\<^sub>1) ps" fun merge_tm :: "('b \ 'a::linorder) \ 'b list \ 'b list \ 'b list tm" where "merge_tm f (x # xs) (y # ys) =1 ( if f x \ f y then do { tl <- merge_tm f xs (y # ys); return (x # tl) } else do { tl <- merge_tm f (x # xs) ys; return (y # tl) } )" | "merge_tm f [] ys =1 return ys" | "merge_tm f xs [] =1 return xs" fun merge :: "('b \ 'a::linorder) \ 'b list \ 'b list \ 'b list" where "merge f (x # xs) (y # ys) = ( if f x \ f y then x # merge f xs (y # ys) else y # merge f (x # xs) ys )" | "merge f [] ys = ys" | "merge f xs [] = xs" lemma merge_eq_val_merge_tm: "val (merge_tm f xs ys) = merge f xs ys" by (induction f xs ys rule: merge.induct) auto lemma length_merge: "length (merge f xs ys) = length xs + length ys" by (induction f xs ys rule: merge.induct) auto lemma set_merge: "set (merge f xs ys) = set xs \ set ys" by (induction f xs ys rule: merge.induct) auto lemma distinct_merge: assumes "set xs \ set ys = {}" "distinct xs" "distinct ys" shows "distinct (merge f xs ys)" using assms by (induction f xs ys rule: merge.induct) (auto simp: set_merge) lemma sorted_merge: assumes "P = (\x y. f x \ f y)" shows "sorted_wrt P (merge f xs ys) \ sorted_wrt P xs \ sorted_wrt P ys" using assms by (induction f xs ys rule: merge.induct) (auto simp: set_merge) declare split_at_take_drop_conv [simp] function (sequential) mergesort_tm :: "('b \ 'a::linorder) \ 'b list \ 'b list tm" where "mergesort_tm f [] =1 return []" | "mergesort_tm f [x] =1 return [x]" | "mergesort_tm f xs =1 ( do { n <- length_tm xs; (xs\<^sub>l, xs\<^sub>r) <- split_at_tm (n div 2) xs; l <- mergesort_tm f xs\<^sub>l; r <- mergesort_tm f xs\<^sub>r; merge_tm f l r } )" by pat_completeness auto termination mergesort_tm by (relation "Wellfounded.measure (\(_, xs). length xs)") (auto simp add: length_eq_val_length_tm split_at_eq_val_split_at_tm) fun mergesort :: "('b \ 'a::linorder) \ 'b list \ 'b list" where "mergesort f [] = []" | "mergesort f [x] = [x]" | "mergesort f xs = ( let n = length xs div 2 in let (l, r) = split_at n xs in merge f (mergesort f l) (mergesort f r) )" declare split_at_take_drop_conv [simp del] lemma mergesort_eq_val_mergesort_tm: "val (mergesort_tm f xs) = mergesort f xs" by (induction f xs rule: mergesort.induct) (auto simp add: length_eq_val_length_tm split_at_eq_val_split_at_tm merge_eq_val_merge_tm split: prod.split) lemma sorted_wrt_mergesort: "sorted_wrt (\x y. f x \ f y) (mergesort f xs)" by (induction f xs rule: mergesort.induct) (auto simp: split_at_take_drop_conv sorted_merge) lemma set_mergesort: "set (mergesort f xs) = set xs" by (induction f xs rule: mergesort.induct) (simp_all add: set_merge split_at_take_drop_conv, metis list.simps(15) set_take_drop) lemma length_mergesort: "length (mergesort f xs) = length xs" by (induction f xs rule: mergesort.induct) (auto simp: length_merge split_at_take_drop_conv) lemma distinct_mergesort: "distinct xs \ distinct (mergesort f xs)" proof (induction f xs rule: mergesort.induct) case (3 f x y xs) let ?xs' = "x # y # xs" obtain l r where lr_def: "(l, r) = split_at (length ?xs' div 2) ?xs'" by (metis surj_pair) have "distinct l" "distinct r" using "3.prems" split_at_take_drop_conv distinct_take distinct_drop lr_def by (metis prod.sel)+ hence "distinct (mergesort f l)" "distinct (mergesort f r)" using "3.IH" lr_def by auto moreover have "set l \ set r = {}" using "3.prems" split_at_take_drop_conv lr_def by (metis append_take_drop_id distinct_append prod.sel) ultimately show ?case using lr_def by (auto simp: distinct_merge set_mergesort split: prod.splits) qed auto lemmas mergesort = sorted_wrt_mergesort set_mergesort length_mergesort distinct_mergesort lemma sorted_fst_take_less_hd_drop: assumes "sorted_fst ps" "n < length ps" shows "\p \ set (take n ps). fst p \ fst (hd (drop n ps))" using assms sorted_wrt_take_less_hd_drop[of "\p\<^sub>0 p\<^sub>1. fst p\<^sub>0 \ fst p\<^sub>1"] sorted_fst_def by fastforce lemma sorted_fst_hd_drop_less_drop: assumes "sorted_fst ps" shows "\p \ set (drop n ps). fst (hd (drop n ps)) \ fst p" using assms sorted_wrt_hd_drop_less_drop[of "\p\<^sub>0 p\<^sub>1. fst p\<^sub>0 \ fst p\<^sub>1"] sorted_fst_def by fastforce subsubsection "Time Complexity Proof" lemma time_merge_tm: "time (merge_tm f xs ys) \ length xs + length ys + 1" by (induction f xs ys rule: merge_tm.induct) (auto simp: time_simps) function mergesort_recurrence :: "nat \ real" where "mergesort_recurrence 0 = 1" | "mergesort_recurrence 1 = 1" | "2 \ n \ mergesort_recurrence n = 4 + 3 * n + mergesort_recurrence (nat \real n / 2\) + mergesort_recurrence (nat \real n / 2\)" by force simp_all termination by akra_bazzi_termination simp_all lemma mergesort_recurrence_nonneg[simp]: "0 \ mergesort_recurrence n" by (induction n rule: mergesort_recurrence.induct) (auto simp del: One_nat_def) lemma time_mergesort_conv_mergesort_recurrence: "time (mergesort_tm f xs) \ mergesort_recurrence (length xs)" proof (induction f xs rule: mergesort_tm.induct) case (1 f) thus ?case by (auto simp: time_simps) next case (2 f x) thus ?case using mergesort_recurrence.simps(2) by (auto simp: time_simps) next case (3 f x y xs') define xs where "xs = x # y # xs'" define n where "n = length xs" obtain l r where lr_def: "(l, r) = split_at (n div 2) xs" using prod.collapse by blast define l' where "l' = mergesort f l" define r' where "r' = mergesort f r" note defs = xs_def n_def lr_def l'_def r'_def have IHL: "time (mergesort_tm f l) \ mergesort_recurrence (length l)" using defs "3.IH"(1) by (auto simp: length_eq_val_length_tm split_at_eq_val_split_at_tm) have IHR: "time (mergesort_tm f r) \ mergesort_recurrence (length r)" using defs "3.IH"(2) by (auto simp: length_eq_val_length_tm split_at_eq_val_split_at_tm) have *: "length l = n div 2" "length r = n - n div 2" using defs by (auto simp: split_at_take_drop_conv) hence "(nat \real n / 2\) = length l" "(nat \real n / 2\) = length r" by linarith+ hence IH: "time (mergesort_tm f l) \ mergesort_recurrence (nat \real n / 2\)" "time (mergesort_tm f r) \ mergesort_recurrence (nat \real n / 2\)" using IHL IHR by simp_all have "n = length l + length r" using * by linarith hence "time (merge_tm f l' r') \ n + 1" using time_merge_tm defs by (metis length_mergesort) have "time (mergesort_tm f xs) = 1 + time (length_tm xs) + time (split_at_tm (n div 2) xs) + time (mergesort_tm f l) + time (mergesort_tm f r) + time (merge_tm f l' r')" using defs by (auto simp add: time_simps length_eq_val_length_tm mergesort_eq_val_mergesort_tm split_at_eq_val_split_at_tm split: prod.split) also have "... \ 4 + 3 * n + time (mergesort_tm f l) + time (mergesort_tm f r)" using time_length_tm[of xs] time_split_at_tm[of "n div 2" xs] n_def \time (merge_tm f l' r') \ n + 1\ by simp also have "... \ 4 + 3 * n + mergesort_recurrence (nat \real n / 2\) + mergesort_recurrence (nat \real n / 2\)" using IH by simp also have "... = mergesort_recurrence n" using defs by simp finally show ?case using defs by simp qed theorem mergesort_recurrence: "mergesort_recurrence \ \(\n. n * ln n)" by (master_theorem) auto theorem time_mergesort_tm_bigo: "(\xs. time (mergesort_tm f xs)) \ O[length going_to at_top]((\n. n * ln n) o length)" proof - have 0: "\xs. time (mergesort_tm f xs) \ (mergesort_recurrence o length) xs" unfolding comp_def using time_mergesort_conv_mergesort_recurrence by blast show ?thesis using bigo_measure_trans[OF 0] by (simp add: bigthetaD1 mergesort_recurrence) qed subsubsection "Code Export" lemma merge_xs_Nil[simp]: "merge f xs [] = xs" by (cases xs) auto fun merge_it' :: "('b \ 'a::linorder) \ 'b list \ 'b list \ 'b list \ 'b list" where "merge_it' f acc [] [] = rev acc" | "merge_it' f acc (x#xs) [] = merge_it' f (x#acc) xs []" | "merge_it' f acc [] (y#ys) = merge_it' f (y#acc) ys []" | "merge_it' f acc (x#xs) (y#ys) = ( if f x \ f y then merge_it' f (x#acc) xs (y#ys) else merge_it' f (y#acc) (x#xs) ys )" definition merge_it :: "('b \ 'a::linorder) \ 'b list \ 'b list \ 'b list" where "merge_it f xs ys = merge_it' f [] xs ys" lemma merge_conv_merge_it': "rev acc @ merge f xs ys = merge_it' f acc xs ys" by (induction f acc xs ys rule: merge_it'.induct) auto lemma merge_conv_merge_it[code_unfold]: "merge f xs ys = merge_it f xs ys" unfolding merge_it_def using merge_conv_merge_it' rev.simps(1) append_Nil by metis subsection "Minimal Distance" definition sparse :: "real \ point set \ bool" where "sparse \ ps \ (\p\<^sub>0 \ ps. \p\<^sub>1 \ ps. p\<^sub>0 \ p\<^sub>1 \ \ \ dist p\<^sub>0 p\<^sub>1)" lemma sparse_identity: assumes "sparse \ (set ps)" "\p \ set ps. \ \ dist p\<^sub>0 p" shows "sparse \ (set (p\<^sub>0 # ps))" using assms by (simp add: dist_commute sparse_def) lemma sparse_update: assumes "sparse \ (set ps)" assumes "dist p\<^sub>0 p\<^sub>1 \ \" "\p \ set ps. dist p\<^sub>0 p\<^sub>1 \ dist p\<^sub>0 p" shows "sparse (dist p\<^sub>0 p\<^sub>1) (set (p\<^sub>0 # ps))" using assms by (auto simp: dist_commute sparse_def, force+) lemma sparse_mono: "sparse \ P \ \ \ \ \ sparse \ P" unfolding sparse_def by fastforce subsection "Distance" lemma dist_transform: fixes p :: point and \ :: real and l :: int shows "dist p (l, snd p) < \ \ l - \ < fst p \ fst p < l + \" proof - have "dist p (l, snd p) = sqrt ((real_of_int (fst p) - l)\<^sup>2)" by (auto simp add: dist_prod_def dist_real_def prod.case_eq_if) thus ?thesis by auto qed fun dist_code :: "point \ point \ int" where "dist_code p\<^sub>0 p\<^sub>1 = (fst p\<^sub>0 - fst p\<^sub>1)\<^sup>2 + (snd p\<^sub>0 - snd p\<^sub>1)\<^sup>2" lemma dist_eq_sqrt_dist_code: fixes p\<^sub>0 :: point shows "dist p\<^sub>0 p\<^sub>1 = sqrt (dist_code p\<^sub>0 p\<^sub>1)" by (auto simp: dist_prod_def dist_real_def split: prod.splits) lemma dist_eq_dist_code_lt: fixes p\<^sub>0 :: point shows "dist p\<^sub>0 p\<^sub>1 < dist p\<^sub>2 p\<^sub>3 \ dist_code p\<^sub>0 p\<^sub>1 < dist_code p\<^sub>2 p\<^sub>3" using dist_eq_sqrt_dist_code real_sqrt_less_iff by presburger lemma dist_eq_dist_code_le: fixes p\<^sub>0 :: point shows "dist p\<^sub>0 p\<^sub>1 \ dist p\<^sub>2 p\<^sub>3 \ dist_code p\<^sub>0 p\<^sub>1 \ dist_code p\<^sub>2 p\<^sub>3" using dist_eq_sqrt_dist_code real_sqrt_le_iff by presburger lemma dist_eq_dist_code_abs_lt: fixes p\<^sub>0 :: point shows "\c\ < dist p\<^sub>0 p\<^sub>1 \ c\<^sup>2 < dist_code p\<^sub>0 p\<^sub>1" using dist_eq_sqrt_dist_code by (metis of_int_less_of_int_power_cancel_iff real_sqrt_abs real_sqrt_less_iff) lemma dist_eq_dist_code_abs_le: fixes p\<^sub>0 :: point shows "dist p\<^sub>0 p\<^sub>1 \ \c\ \ dist_code p\<^sub>0 p\<^sub>1 \ c\<^sup>2" using dist_eq_sqrt_dist_code by (metis of_int_power_le_of_int_cancel_iff real_sqrt_abs real_sqrt_le_iff) lemma dist_fst_abs: fixes p :: point and l :: int shows "dist p (l, snd p) = \fst p - l\" proof - have "dist p (l, snd p) = sqrt ((real_of_int (fst p) - l)\<^sup>2)" by (simp add: dist_prod_def dist_real_def prod.case_eq_if) thus ?thesis by simp qed declare dist_code.simps [simp del] subsection "Brute Force Closest Pair Algorithm" subsubsection "Functional Correctness Proof" fun find_closest_bf_tm :: "point \ point list \ point tm" where "find_closest_bf_tm _ [] =1 return undefined" | "find_closest_bf_tm _ [p] =1 return p" | "find_closest_bf_tm p (p\<^sub>0 # ps) =1 ( do { p\<^sub>1 <- find_closest_bf_tm p ps; if dist p p\<^sub>0 < dist p p\<^sub>1 then return p\<^sub>0 else return p\<^sub>1 } )" fun find_closest_bf :: "point \ point list \ point" where "find_closest_bf _ [] = undefined" | "find_closest_bf _ [p] = p" | "find_closest_bf p (p\<^sub>0 # ps) = ( let p\<^sub>1 = find_closest_bf p ps in if dist p p\<^sub>0 < dist p p\<^sub>1 then p\<^sub>0 else p\<^sub>1 )" lemma find_closest_bf_eq_val_find_closest_bf_tm: "val (find_closest_bf_tm p ps) = find_closest_bf p ps" by (induction p ps rule: find_closest_bf.induct) (auto simp: Let_def) lemma find_closest_bf_set: "0 < length ps \ find_closest_bf p ps \ set ps" by (induction p ps rule: find_closest_bf.induct) (auto simp: Let_def split: prod.splits if_splits) lemma find_closest_bf_dist: "\q \ set ps. dist p (find_closest_bf p ps) \ dist p q" by (induction p ps rule: find_closest_bf.induct) (auto split: prod.splits) fun closest_pair_bf_tm :: "point list \ (point \ point) tm" where "closest_pair_bf_tm [] =1 return undefined" | "closest_pair_bf_tm [_] =1 return undefined" | "closest_pair_bf_tm [p\<^sub>0, p\<^sub>1] =1 return (p\<^sub>0, p\<^sub>1)" | "closest_pair_bf_tm (p\<^sub>0 # ps) =1 ( do { (c\<^sub>0::point, c\<^sub>1::point) <- closest_pair_bf_tm ps; p\<^sub>1 <- find_closest_bf_tm p\<^sub>0 ps; if dist c\<^sub>0 c\<^sub>1 \ dist p\<^sub>0 p\<^sub>1 then return (c\<^sub>0, c\<^sub>1) else return (p\<^sub>0, p\<^sub>1) } )" fun closest_pair_bf :: "point list \ (point * point)" where "closest_pair_bf [] = undefined" | "closest_pair_bf [_] = undefined" | "closest_pair_bf [p\<^sub>0, p\<^sub>1] = (p\<^sub>0, p\<^sub>1)" | "closest_pair_bf (p\<^sub>0 # ps) = ( let (c\<^sub>0, c\<^sub>1) = closest_pair_bf ps in let p\<^sub>1 = find_closest_bf p\<^sub>0 ps in if dist c\<^sub>0 c\<^sub>1 \ dist p\<^sub>0 p\<^sub>1 then (c\<^sub>0, c\<^sub>1) else (p\<^sub>0, p\<^sub>1) )" lemma closest_pair_bf_eq_val_closest_pair_bf_tm: "val (closest_pair_bf_tm ps) = closest_pair_bf ps" by (induction ps rule: closest_pair_bf.induct) (auto simp: Let_def find_closest_bf_eq_val_find_closest_bf_tm split: prod.split) lemma closest_pair_bf_c0: "1 < length ps \ (c\<^sub>0, c\<^sub>1) = closest_pair_bf ps \ c\<^sub>0 \ set ps" by (induction ps arbitrary: c\<^sub>0 c\<^sub>1 rule: closest_pair_bf.induct) (auto simp: Let_def find_closest_bf_set split: if_splits prod.splits) lemma closest_pair_bf_c1: "1 < length ps \ (c\<^sub>0, c\<^sub>1) = closest_pair_bf ps \ c\<^sub>1 \ set ps" proof (induction ps arbitrary: c\<^sub>0 c\<^sub>1 rule: closest_pair_bf.induct) case (4 p\<^sub>0 p\<^sub>2 p\<^sub>3 ps) let ?ps = "p\<^sub>2 # p\<^sub>3 # ps" obtain c\<^sub>0 c\<^sub>1 where c\<^sub>0_def: "(c\<^sub>0, c\<^sub>1) = closest_pair_bf ?ps" using prod.collapse by blast define p\<^sub>1 where p\<^sub>1_def: "p\<^sub>1 = find_closest_bf p\<^sub>0 ?ps" note defs = c\<^sub>0_def p\<^sub>1_def have "c\<^sub>1 \ set ?ps" using "4.IH" defs by simp moreover have "p\<^sub>1 \ set ?ps" using find_closest_bf_set defs by blast ultimately show ?case using "4.prems"(2) defs by (auto simp: Let_def split: prod.splits if_splits) qed auto lemma closest_pair_bf_c0_ne_c1: "1 < length ps \ distinct ps \ (c\<^sub>0, c\<^sub>1) = closest_pair_bf ps \ c\<^sub>0 \ c\<^sub>1" proof (induction ps arbitrary: c\<^sub>0 c\<^sub>1 rule: closest_pair_bf.induct) case (4 p\<^sub>0 p\<^sub>2 p\<^sub>3 ps) let ?ps = "p\<^sub>2 # p\<^sub>3 # ps" obtain c\<^sub>0 c\<^sub>1 where c\<^sub>0_def: "(c\<^sub>0, c\<^sub>1) = closest_pair_bf ?ps" using prod.collapse by blast define p\<^sub>1 where p\<^sub>1_def: "p\<^sub>1 = find_closest_bf p\<^sub>0 ?ps" note defs = c\<^sub>0_def p\<^sub>1_def have "c\<^sub>0 \ c\<^sub>1" using "4.IH" "4.prems"(2) defs by simp moreover have "p\<^sub>0 \ p\<^sub>1" using find_closest_bf_set "4.prems"(2) defs by (metis distinct.simps(2) length_pos_if_in_set list.set_intros(1)) ultimately show ?case using "4.prems"(3) defs by (auto simp: Let_def split: prod.splits if_splits) qed auto lemmas closest_pair_bf_c0_c1 = closest_pair_bf_c0 closest_pair_bf_c1 closest_pair_bf_c0_ne_c1 lemma closest_pair_bf_dist: assumes "1 < length ps" "(c\<^sub>0, c\<^sub>1) = closest_pair_bf ps" shows "sparse (dist c\<^sub>0 c\<^sub>1) (set ps)" using assms proof (induction ps arbitrary: c\<^sub>0 c\<^sub>1 rule: closest_pair_bf.induct) case (4 p\<^sub>0 p\<^sub>2 p\<^sub>3 ps) let ?ps = "p\<^sub>2 # p\<^sub>3 # ps" obtain c\<^sub>0 c\<^sub>1 where c\<^sub>0_def: "(c\<^sub>0, c\<^sub>1) = closest_pair_bf ?ps" using prod.collapse by blast define p\<^sub>1 where p\<^sub>1_def: "p\<^sub>1 = find_closest_bf p\<^sub>0 ?ps" note defs = c\<^sub>0_def p\<^sub>1_def hence IH: "sparse (dist c\<^sub>0 c\<^sub>1) (set ?ps)" using 4 c\<^sub>0_def by simp have *: "\p \ set ?ps. (dist p\<^sub>0 p\<^sub>1) \ dist p\<^sub>0 p" using find_closest_bf_dist defs by blast show ?case proof (cases "dist c\<^sub>0 c\<^sub>1 \ dist p\<^sub>0 p\<^sub>1") case True hence "\p \ set ?ps. dist c\<^sub>0 c\<^sub>1 \ dist p\<^sub>0 p" using * by auto hence "sparse (dist c\<^sub>0 c\<^sub>1) (set (p\<^sub>0 # ?ps))" using sparse_identity IH by blast thus ?thesis using True "4.prems" defs by (auto split: prod.splits) next case False hence "sparse (dist p\<^sub>0 p\<^sub>1) (set (p\<^sub>0 # ?ps))" using sparse_update[of "dist c\<^sub>0 c\<^sub>1" ?ps p\<^sub>0 p\<^sub>1] IH * defs by argo thus ?thesis using False "4.prems" defs by (auto split: prod.splits) qed qed (auto simp: dist_commute sparse_def) subsubsection "Time Complexity Proof" lemma time_find_closest_bf_tm: "time (find_closest_bf_tm p ps) \ length ps + 1" by (induction p ps rule: find_closest_bf_tm.induct) (auto simp: time_simps) lemma time_closest_pair_bf_tm: "time (closest_pair_bf_tm ps) \ length ps * length ps + 1" proof (induction ps rule: closest_pair_bf_tm.induct) case (4 p\<^sub>0 p\<^sub>2 p\<^sub>3 ps) let ?ps = "p\<^sub>2 # p\<^sub>3 # ps" have "time (closest_pair_bf_tm (p\<^sub>0 # ?ps)) = 1 + time (find_closest_bf_tm p\<^sub>0 ?ps) + time (closest_pair_bf_tm ?ps)" by (auto simp: time_simps split: prod.split) also have "... \ 2 + length ?ps + time (closest_pair_bf_tm ?ps)" using time_find_closest_bf_tm[of p\<^sub>0 ?ps] by simp also have "... \ 2 + length ?ps + length ?ps * length ?ps + 1" using "4.IH" by simp also have "... \ length (p\<^sub>0 # ?ps) * length (p\<^sub>0 # ?ps) + 1" by auto finally show ?case by blast qed (auto simp: time_simps) subsubsection "Code Export" fun find_closest_bf_code :: "point \ point list \ (int * point)" where "find_closest_bf_code p [] = undefined" | "find_closest_bf_code p [p\<^sub>0] = (dist_code p p\<^sub>0, p\<^sub>0)" | "find_closest_bf_code p (p\<^sub>0 # ps) = ( let (\\<^sub>1, p\<^sub>1) = find_closest_bf_code p ps in let \\<^sub>0 = dist_code p p\<^sub>0 in if \\<^sub>0 < \\<^sub>1 then (\\<^sub>0, p\<^sub>0) else (\\<^sub>1, p\<^sub>1) )" lemma find_closest_bf_code_dist_eq: "0 < length ps \ (\, c) = find_closest_bf_code p ps \ \ = dist_code p c" by (induction p ps rule: find_closest_bf_code.induct) (auto simp: Let_def split: prod.splits if_splits) lemma find_closest_bf_code_eq: "0 < length ps \ c = find_closest_bf p ps \ (\', c') = find_closest_bf_code p ps \ c = c'" proof (induction p ps arbitrary: c \' c' rule: find_closest_bf.induct) case (3 p p\<^sub>0 p\<^sub>2 ps) define \\<^sub>0 \\<^sub>0' where \\<^sub>0_def: "\\<^sub>0 = dist p p\<^sub>0" "\\<^sub>0' = dist_code p p\<^sub>0" obtain \\<^sub>1 p\<^sub>1 \\<^sub>1' p\<^sub>1' where \\<^sub>1_def: "\\<^sub>1 = dist p p\<^sub>1" "p\<^sub>1 = find_closest_bf p (p\<^sub>2 # ps)" "(\\<^sub>1', p\<^sub>1') = find_closest_bf_code p (p\<^sub>2 # ps)" using prod.collapse by blast+ note defs = \\<^sub>0_def \\<^sub>1_def have *: "p\<^sub>1 = p\<^sub>1'" using "3.IH" defs by simp hence "\\<^sub>0 < \\<^sub>1 \ \\<^sub>0' < \\<^sub>1'" using find_closest_bf_code_dist_eq[of "p\<^sub>2 # ps" \\<^sub>1' p\<^sub>1' p] dist_eq_dist_code_lt defs by simp thus ?case using "3.prems"(2,3) * defs by (auto split: prod.splits) qed auto declare find_closest_bf_code.simps [simp del] fun closest_pair_bf_code :: "point list \ (int * point * point)" where "closest_pair_bf_code [] = undefined" | "closest_pair_bf_code [p\<^sub>0] = undefined" | "closest_pair_bf_code [p\<^sub>0, p\<^sub>1] = (dist_code p\<^sub>0 p\<^sub>1, p\<^sub>0, p\<^sub>1)" | "closest_pair_bf_code (p\<^sub>0 # ps) = ( let (\\<^sub>c, c\<^sub>0, c\<^sub>1) = closest_pair_bf_code ps in let (\\<^sub>p, p\<^sub>1) = find_closest_bf_code p\<^sub>0 ps in if \\<^sub>c \ \\<^sub>p then (\\<^sub>c, c\<^sub>0, c\<^sub>1) else (\\<^sub>p, p\<^sub>0, p\<^sub>1) )" lemma closest_pair_bf_code_dist_eq: "1 < length ps \ (\, c\<^sub>0, c\<^sub>1) = closest_pair_bf_code ps \ \ = dist_code c\<^sub>0 c\<^sub>1" proof (induction ps arbitrary: \ c\<^sub>0 c\<^sub>1 rule: closest_pair_bf_code.induct) case (4 p\<^sub>0 p\<^sub>2 p\<^sub>3 ps) let ?ps = "p\<^sub>2 # p\<^sub>3 # ps" obtain \\<^sub>c c\<^sub>0 c\<^sub>1 where \\<^sub>c_def: "(\\<^sub>c, c\<^sub>0, c\<^sub>1) = closest_pair_bf_code ?ps" by (metis prod_cases3) obtain \\<^sub>p p\<^sub>1 where \\<^sub>p_def: "(\\<^sub>p, p\<^sub>1) = find_closest_bf_code p\<^sub>0 ?ps" using prod.collapse by blast note defs = \\<^sub>c_def \\<^sub>p_def have "\\<^sub>c = dist_code c\<^sub>0 c\<^sub>1" using "4.IH" defs by simp moreover have "\\<^sub>p = dist_code p\<^sub>0 p\<^sub>1" using find_closest_bf_code_dist_eq defs by blast ultimately show ?case using "4.prems"(2) defs by (auto split: prod.splits if_splits) qed auto lemma closest_pair_bf_code_eq: assumes "1 < length ps" assumes "(c\<^sub>0, c\<^sub>1) = closest_pair_bf ps" "(\', c\<^sub>0', c\<^sub>1') = closest_pair_bf_code ps" shows "c\<^sub>0 = c\<^sub>0' \ c\<^sub>1 = c\<^sub>1'" using assms proof (induction ps arbitrary: c\<^sub>0 c\<^sub>1 \' c\<^sub>0' c\<^sub>1' rule: closest_pair_bf_code.induct) case (4 p\<^sub>0 p\<^sub>2 p\<^sub>3 ps) let ?ps = "p\<^sub>2 # p\<^sub>3 # ps" obtain c\<^sub>0 c\<^sub>1 \\<^sub>c' c\<^sub>0' c\<^sub>1' where \\<^sub>c_def: "(c\<^sub>0, c\<^sub>1) = closest_pair_bf ?ps" "(\\<^sub>c', c\<^sub>0', c\<^sub>1') = closest_pair_bf_code ?ps" by (metis prod_cases3) obtain p\<^sub>1 \\<^sub>p' p\<^sub>1' where \\<^sub>p_def: "p\<^sub>1 = find_closest_bf p\<^sub>0 ?ps" "(\\<^sub>p', p\<^sub>1') = find_closest_bf_code p\<^sub>0 ?ps" using prod.collapse by blast note defs = \\<^sub>c_def \\<^sub>p_def have A: "c\<^sub>0 = c\<^sub>0' \ c\<^sub>1 = c\<^sub>1'" using "4.IH" defs by simp moreover have B: "p\<^sub>1 = p\<^sub>1'" using find_closest_bf_code_eq defs by blast moreover have "\\<^sub>c' = dist_code c\<^sub>0' c\<^sub>1'" using defs closest_pair_bf_code_dist_eq[of ?ps] by simp moreover have "\\<^sub>p' = dist_code p\<^sub>0 p\<^sub>1'" using defs find_closest_bf_code_dist_eq by blast ultimately have "dist c\<^sub>0 c\<^sub>1 \ dist p\<^sub>0 p\<^sub>1 \ \\<^sub>c' \ \\<^sub>p'" by (simp add: dist_eq_dist_code_le) thus ?case using "4.prems"(2,3) defs A B by (auto simp: Let_def split: prod.splits) qed auto subsection "Geometry" subsubsection "Band Filter" lemma set_band_filter_aux: fixes \ :: real and ps :: "point list" assumes "p\<^sub>0 \ ps\<^sub>L" "p\<^sub>1 \ ps\<^sub>R" "p\<^sub>0 \ p\<^sub>1" "dist p\<^sub>0 p\<^sub>1 < \" "set ps = ps\<^sub>L \ ps\<^sub>R" assumes "\p \ ps\<^sub>L. fst p \ l" "\p \ ps\<^sub>R. l \ fst p" assumes "ps' = filter (\p. l - \ < fst p \ fst p < l + \) ps" shows "p\<^sub>0 \ set ps' \ p\<^sub>1 \ set ps'" proof (rule ccontr) assume "\ (p\<^sub>0 \ set ps' \ p\<^sub>1 \ set ps')" then consider (A) "p\<^sub>0 \ set ps' \ p\<^sub>1 \ set ps'" | (B) "p\<^sub>0 \ set ps' \ p\<^sub>1 \ set ps'" | (C) "p\<^sub>0 \ set ps' \ p\<^sub>1 \ set ps'" by blast thus False proof cases case A hence "fst p\<^sub>0 \ l - \ \ l + \ \ fst p\<^sub>0" "fst p\<^sub>1 \ l - \ \ l + \ \ fst p\<^sub>1" using assms(1,2,5,8) by auto hence "fst p\<^sub>0 \ l - \" "l + \ \ fst p\<^sub>1" using assms(1,2,6,7) by force+ hence "\ \ dist (fst p\<^sub>0) (fst p\<^sub>1)" using dist_real_def by simp hence "\ \ dist p\<^sub>0 p\<^sub>1" using dist_fst_le[of p\<^sub>0 p\<^sub>1] by (auto split: prod.splits) then show ?thesis using assms(4) by fastforce next case B hence "fst p\<^sub>1 \ l - \ \ l + \ \ fst p\<^sub>1" using assms(2,5,8) by auto hence "l + \ \ fst p\<^sub>1" using assms(2,7) by auto moreover have "fst p\<^sub>0 \ l" using assms(1,6) by simp ultimately have "\ \ dist (fst p\<^sub>0) (fst p\<^sub>1)" using dist_real_def by simp hence "\ \ dist p\<^sub>0 p\<^sub>1" using dist_fst_le[of p\<^sub>0 p\<^sub>1] less_le_trans by (auto split: prod.splits) thus ?thesis using assms(4) by simp next case C hence "fst p\<^sub>0 \ l - \ \ l + \ \ fst p\<^sub>0" using assms(1,2,5,8) by auto hence "fst p\<^sub>0 \ l - \" using assms(1,6) by auto moreover have "l \ fst p\<^sub>1" using assms(2,7) by simp ultimately have "\ \ dist (fst p\<^sub>0) (fst p\<^sub>1)" using dist_real_def by simp hence "\ \ dist p\<^sub>0 p\<^sub>1" using dist_fst_le[of p\<^sub>0 p\<^sub>1] less_le_trans by (auto split: prod.splits) thus ?thesis using assms(4) by simp qed qed lemma set_band_filter: fixes \ :: real and ps :: "point list" assumes "p\<^sub>0 \ set ps" "p\<^sub>1 \ set ps" "p\<^sub>0 \ p\<^sub>1" "dist p\<^sub>0 p\<^sub>1 < \" "set ps = ps\<^sub>L \ ps\<^sub>R" assumes "sparse \ ps\<^sub>L" "sparse \ ps\<^sub>R" assumes "\p \ ps\<^sub>L. fst p \ l" "\p \ ps\<^sub>R. l \ fst p" assumes "ps' = filter (\p. l - \ < fst p \ fst p < l + \) ps" shows "p\<^sub>0 \ set ps' \ p\<^sub>1 \ set ps'" proof - have "p\<^sub>0 \ ps\<^sub>L \ p\<^sub>1 \ ps\<^sub>L" "p\<^sub>0 \ ps\<^sub>R \ p\<^sub>1 \ ps\<^sub>R" using assms(3,4,6,7) sparse_def by force+ then consider (A) "p\<^sub>0 \ ps\<^sub>L \ p\<^sub>1 \ ps\<^sub>R" | (B) "p\<^sub>0 \ ps\<^sub>R \ p\<^sub>1 \ ps\<^sub>L" using assms(1,2,5) by auto thus ?thesis proof cases case A thus ?thesis using set_band_filter_aux assms(3,4,5,8,9,10) by (auto split: prod.splits) next case B moreover have "dist p\<^sub>1 p\<^sub>0 < \" using assms(4) dist_commute by metis ultimately show ?thesis using set_band_filter_aux assms(3)[symmetric] assms(5,8,9,10) by (auto split: prod.splits) qed qed subsubsection "2D-Boxes and Points" lemma cbox_2D: fixes x\<^sub>0 :: real and y\<^sub>0 :: real shows "cbox (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>1) = { (x, y). x\<^sub>0 \ x \ x \ x\<^sub>1 \ y\<^sub>0 \ y \ y \ y\<^sub>1 }" by fastforce lemma mem_cbox_2D: fixes x :: real and y :: real shows "x\<^sub>0 \ x \ x \ x\<^sub>1 \ y\<^sub>0 \ y \ y \ y\<^sub>1 \ (x, y) \ cbox (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>1)" by fastforce lemma cbox_top_un: fixes x\<^sub>0 :: real and y\<^sub>0 :: real assumes "y\<^sub>0 \ y\<^sub>1" "y\<^sub>1 \ y\<^sub>2" shows "cbox (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>1) \ cbox (x\<^sub>0, y\<^sub>1) (x\<^sub>1, y\<^sub>2) = cbox (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>2)" using assms by auto lemma cbox_right_un: fixes x\<^sub>0 :: real and y\<^sub>0 :: real assumes "x\<^sub>0 \ x\<^sub>1" "x\<^sub>1 \ x\<^sub>2" shows "cbox (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>1) \ cbox (x\<^sub>1, y\<^sub>0) (x\<^sub>2, y\<^sub>1) = cbox (x\<^sub>0, y\<^sub>0) (x\<^sub>2, y\<^sub>1)" using assms by auto lemma cbox_max_dist: assumes "p\<^sub>0 = (x, y)" "p\<^sub>1 = (x + \, y + \)" assumes "(x\<^sub>0, y\<^sub>0) \ cbox p\<^sub>0 p\<^sub>1" "(x\<^sub>1, y\<^sub>1) \ cbox p\<^sub>0 p\<^sub>1" "0 \ \" shows "dist (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>1) \ sqrt 2 * \" proof - have X: "dist x\<^sub>0 x\<^sub>1 \ \" using assms dist_real_def by auto have Y: "dist y\<^sub>0 y\<^sub>1 \ \" using assms dist_real_def by auto have "dist (x\<^sub>0, y\<^sub>0) (x\<^sub>1, y\<^sub>1) = sqrt ((dist x\<^sub>0 x\<^sub>1)\<^sup>2 + (dist y\<^sub>0 y\<^sub>1)\<^sup>2)" using dist_Pair_Pair by auto also have "... \ sqrt (\\<^sup>2 + (dist y\<^sub>0 y\<^sub>1)\<^sup>2)" using X power_mono by fastforce also have "... \ sqrt (\\<^sup>2 + \\<^sup>2)" using Y power_mono by fastforce also have "... = sqrt 2 * sqrt (\\<^sup>2)" using real_sqrt_mult by simp also have "... = sqrt 2 * \" using assms(5) by simp finally show ?thesis . qed subsubsection "Pigeonhole Argument" lemma card_le_1_if_pairwise_eq: assumes "\x \ S. \y \ S. x = y" shows "card S \ 1" proof (rule ccontr) assume "\ card S \ 1" hence "2 \ card S" by simp then obtain T where *: "T \ S \ card T = 2" using ex_card by metis then obtain x y where "x \ T \ y \ T \ x \ y" by (meson card_2_iff') then show False using * assms by blast qed lemma card_Int_if_either_in: assumes "\x \ S. \y \ S. x = y \ x \ T \ y \ T" shows "card (S \ T) \ 1" proof (rule ccontr) assume "\ (card (S \ T) \ 1)" then obtain x y where *: "x \ S \ T \ y \ S \ T \ x \ y" by (meson card_le_1_if_pairwise_eq) hence "x \ T" "y \ T" by simp_all moreover have "x \ T \ y \ T" using assms * by auto ultimately show False by blast qed lemma card_Int_Un_le_Sum_card_Int: assumes "finite S" shows "card (A \ \S) \ (\B \ S. card (A \ B))" using assms proof (induction "card S" arbitrary: S) case (Suc n) then obtain B T where *: "S = { B } \ T" "card T = n" "B \ T" by (metis card_Suc_eq Suc_eq_plus1 insert_is_Un) hence "card (A \ \S) = card (A \ \({ B } \ T))" by blast also have "... \ card (A \ B) + card (A \ \T)" by (simp add: card_Un_le inf_sup_distrib1) also have "... \ card (A \ B) + (\B \ T. card (A \ B))" using Suc * by simp also have "... \ (\B \ S. card (A \ B))" using Suc.prems * by simp finally show ?case . qed simp lemma pigeonhole: assumes "finite T" "S \ \T" "card T < card S" shows "\x \ S. \y \ S. \X \ T. x \ y \ x \ X \ y \ X" proof (rule ccontr) assume "\ (\x \ S. \y \ S. \X \ T. x \ y \ x \ X \ y \ X)" hence *: "\X \ T. card (S \ X) \ 1" using card_Int_if_either_in by metis have "card T < card (S \ \T)" using Int_absorb2 assms by fastforce also have "... \ (\X \ T. card (S \ X))" using assms(1) card_Int_Un_le_Sum_card_Int by blast also have "... \ card T" using * sum_mono by fastforce finally show False by simp qed subsubsection "Delta Sparse Points within a Square" lemma max_points_square: assumes "\p \ ps. p \ cbox (x, y) (x + \, y + \)" "sparse \ ps" "0 \ \" shows "card ps \ 4" proof (cases "\ = 0") case True hence "{ (x, y) } = cbox (x, y) (x + \, y + \)" using cbox_def by simp hence "\p \ ps. p = (x, y)" using assms(1) by blast hence "\p \ ps. \q \ ps. p = q" apply (auto split: prod.splits) by (metis of_int_eq_iff)+ thus ?thesis using card_le_1_if_pairwise_eq by force next case False hence \: "0 < \" using assms(3) by simp show ?thesis proof (rule ccontr) assume A: "\ (card ps \ 4)" define PS where PS_def: "PS = (\(x, y). (real_of_int x, real_of_int y)) ` ps" have "inj_on (\(x, y). (real_of_int x, real_of_int y)) ps" using inj_on_def by fastforce hence *: "\ (card PS \ 4)" using A PS_def by (simp add: card_image) let ?x' = "x + \ / 2" let ?y' = "y + \ / 2" let ?ll = "cbox ( x , y ) (?x' , ?y' )" let ?lu = "cbox ( x , ?y') (?x' , y + \)" let ?rl = "cbox (?x', y ) ( x + \, ?y' )" let ?ru = "cbox (?x', ?y') ( x + \, y + \)" let ?sq = "{ ?ll, ?lu, ?rl, ?ru }" have card_le_4: "card ?sq \ 4" by (simp add: card_insert_le_m1) have "cbox (x, y) (?x', y + \) = ?ll \ ?lu" using cbox_top_un assms(3) by auto moreover have "cbox (?x', y) (x + \, y + \) = ?rl \ ?ru" using cbox_top_un assms(3) by auto moreover have "cbox (x, y) (?x', y + \) \ cbox (?x', y) (x + \, y + \) = cbox (x, y) (x + \, y + \)" using cbox_right_un assms(3) by simp ultimately have "?ll \ ?lu \ ?rl \ ?ru = cbox (x, y) (x + \, y + \)" by blast hence "PS \ \(?sq)" using assms(1) PS_def by (auto split: prod.splits) moreover have "card ?sq < card PS" using * card_insert_le_m1 card_le_4 by linarith moreover have "finite ?sq" by simp ultimately have "\p\<^sub>0 \ PS. \p\<^sub>1 \ PS. \S \ ?sq. (p\<^sub>0 \ p\<^sub>1 \ p\<^sub>0 \ S \ p\<^sub>1 \ S)" using pigeonhole[of ?sq PS] by blast then obtain S p\<^sub>0 p\<^sub>1 where #: "p\<^sub>0 \ PS" "p\<^sub>1 \ PS" "S \ ?sq" "p\<^sub>0 \ p\<^sub>1" "p\<^sub>0 \ S" "p\<^sub>1 \ S" by blast have D: "0 \ \ / 2" using assms(3) by simp have LL: "\p\<^sub>0 \ ?ll. \p\<^sub>1 \ ?ll. dist p\<^sub>0 p\<^sub>1 \ sqrt 2 * (\ / 2)" using cbox_max_dist[of "(x, y)" x y "(?x', ?y')" "\ / 2"] D by auto have LU: "\p\<^sub>0 \ ?lu. \p\<^sub>1 \ ?lu. dist p\<^sub>0 p\<^sub>1 \ sqrt 2 * (\ / 2)" using cbox_max_dist[of "(x, ?y')" x ?y' "(?x', y + \)" "\ / 2"] D by auto have RL: "\p\<^sub>0 \ ?rl. \p\<^sub>1 \ ?rl. dist p\<^sub>0 p\<^sub>1 \ sqrt 2 * (\ / 2)" using cbox_max_dist[of "(?x', y)" ?x' y "(x + \, ?y')" "\ / 2"] D by auto have RU: "\p\<^sub>0 \ ?ru. \p\<^sub>1 \ ?ru. dist p\<^sub>0 p\<^sub>1 \ sqrt 2 * (\ / 2)" using cbox_max_dist[of "(?x', ?y')" ?x' ?y' "(x + \, y + \)" "\ / 2"] D by auto have "\p\<^sub>0 \ S. \p\<^sub>1 \ S. dist p\<^sub>0 p\<^sub>1 \ sqrt 2 * (\ / 2)" using # LL LU RL RU by blast hence "dist p\<^sub>0 p\<^sub>1 \ (sqrt 2 / 2) * \" using # by simp moreover have "(sqrt 2 / 2) * \ < \" using sqrt2_less_2 \ by simp ultimately have "dist p\<^sub>0 p\<^sub>1 < \" by simp moreover have "\ \ dist p\<^sub>0 p\<^sub>1" using assms(2) # sparse_def PS_def by auto ultimately show False by simp qed qed end \ No newline at end of file diff --git a/thys/Coinductive/Coinductive_Nat.thy b/thys/Coinductive/Coinductive_Nat.thy --- a/thys/Coinductive/Coinductive_Nat.thy +++ b/thys/Coinductive/Coinductive_Nat.thy @@ -1,747 +1,747 @@ (* Title: Coinductive natural numbers Author: Andreas Lochbihler Maintainer: Andreas Lochbihler *) section \Extended natural numbers as a codatatype\ theory Coinductive_Nat imports "HOL-Library.Extended_Nat" "HOL-Library.Complete_Partial_Order2" begin lemma inj_enat [simp]: "inj_on enat A" by(simp add: inj_on_def) lemma Sup_range_enat [simp]: "Sup (range enat) = \" by(auto dest: finite_imageD simp add: Sup_enat_def) lemmas eSuc_plus = iadd_Suc lemmas plus_enat_eq_0_conv = iadd_is_0 lemma enat_add_sub_same: fixes a b :: enat shows "a \ \ \ a + b - a = b" by(cases b) auto lemma enat_the_enat: "n \ \ \ enat (the_enat n) = n" by auto lemma enat_min_eq_0_iff: fixes a b :: enat shows "min a b = 0 \ a = 0 \ b = 0" by(auto simp add: min_def) lemma enat_le_plus_same: "x \ (x :: enat) + y" "x \ y + x" by(auto simp add: less_eq_enat_def plus_enat_def split: enat.split) lemma the_enat_0 [simp]: "the_enat 0 = 0" by(simp add: zero_enat_def) lemma the_enat_eSuc: "n \ \ \ the_enat (eSuc n) = Suc (the_enat n)" by(cases n)(simp_all add: eSuc_enat) coinductive_set enat_set :: "enat set" where "0 \ enat_set" | "n \ enat_set \ (eSuc n) \ enat_set" lemma enat_set_eq_UNIV [simp]: "enat_set = UNIV" proof show "enat_set \ UNIV" by blast show "UNIV \ enat_set" proof fix x :: enat assume "x \ UNIV" thus "x \ enat_set" proof coinduct case (enat_set x) show ?case proof(cases "x = 0") case True thus ?thesis by simp next case False then obtain n where "x = eSuc n" by(cases x)(fastforce simp add: eSuc_def zero_enat_def gr0_conv_Suc split: enat.splits)+ thus ?thesis by auto qed qed qed qed subsection \Case operator\ lemma enat_coexhaust: obtains (0) "n = 0" | (eSuc) n' where "n = eSuc n'" proof - have "n \ enat_set" by auto thus thesis by cases (erule that)+ qed locale co begin free_constructors (plugins del: code) case_enat for "0::enat" | eSuc epred where "epred 0 = 0" apply (erule enat_coexhaust, assumption) apply (rule eSuc_inject) by (rule zero_ne_eSuc) end lemma enat_cocase_0 [simp]: "co.case_enat z s 0 = z" by (rule co.enat.case(1)) lemma enat_cocase_eSuc [simp]: "co.case_enat z s (eSuc n) = s n" by (rule co.enat.case(2)) lemma neq_zero_conv_eSuc: "n \ 0 \ (\n'. n = eSuc n')" by(cases n rule: enat_coexhaust) simp_all lemma enat_cocase_cert: assumes "CASE \ co.case_enat c d" shows "(CASE 0 \ c) &&& (CASE (eSuc n) \ d n)" using assms by simp_all lemma enat_cosplit_asm: "P (co.case_enat c d n) = (\ (n = 0 \ \ P c \ (\m. n = eSuc m \ \ P (d m))))" by (rule co.enat.split_asm) lemma enat_cosplit: "P (co.case_enat c d n) = ((n = 0 \ P c) \ (\m. n = eSuc m \ P (d m)))" by (rule co.enat.split) abbreviation epred :: "enat => enat" where "epred \ co.epred" lemma epred_0 [simp]: "epred 0 = 0" by(rule co.enat.sel(1)) lemma epred_eSuc [simp]: "epred (eSuc n) = n" by(rule co.enat.sel(2)) declare co.enat.collapse[simp] lemma epred_conv_minus: "epred n = n - 1" by(cases n rule: co.enat.exhaust)(simp_all) subsection \Corecursion for @{typ enat}\ lemma case_enat_numeral [simp]: "case_enat f i (numeral v) = (let n = numeral v in f n)" by(simp add: numeral_eq_enat) lemma case_enat_0 [simp]: "case_enat f i 0 = f 0" by(simp add: zero_enat_def) lemma [simp]: shows max_eSuc_eSuc: "max (eSuc n) (eSuc m) = eSuc (max n m)" and min_eSuc_eSuc: "min (eSuc n) (eSuc m) = eSuc (min n m)" by(simp_all add: eSuc_def split: enat.split) definition epred_numeral :: "num \ enat" where [code del]: "epred_numeral = enat \ pred_numeral" lemma numeral_eq_eSuc: "numeral k = eSuc (epred_numeral k)" by(simp add: numeral_eq_Suc eSuc_def epred_numeral_def numeral_eq_enat) lemma epred_numeral_simps [simp]: "epred_numeral num.One = 0" "epred_numeral (num.Bit0 k) = numeral (Num.BitM k)" "epred_numeral (num.Bit1 k) = numeral (num.Bit0 k)" by(simp_all add: epred_numeral_def enat_numeral zero_enat_def) lemma [simp]: shows eq_numeral_eSuc: "numeral k = eSuc n \ epred_numeral k = n" and Suc_eq_numeral: "eSuc n = numeral k \ n = epred_numeral k" and less_numeral_Suc: "numeral k < eSuc n \ epred_numeral k < n" and less_eSuc_numeral: "eSuc n < numeral k \ n < epred_numeral k" and le_numeral_eSuc: "numeral k \ eSuc n \ epred_numeral k \ n" and le_eSuc_numeral: "eSuc n \ numeral k \ n \ epred_numeral k" and diff_eSuc_numeral: "eSuc n - numeral k = n - epred_numeral k" and diff_numeral_eSuc: "numeral k - eSuc n = epred_numeral k - n" and max_eSuc_numeral: "max (eSuc n) (numeral k) = eSuc (max n (epred_numeral k))" and max_numeral_eSuc: "max (numeral k) (eSuc n) = eSuc (max (epred_numeral k) n)" and min_eSuc_numeral: "min (eSuc n) (numeral k) = eSuc (min n (epred_numeral k))" and min_numeral_eSuc: "min (numeral k) (eSuc n) = eSuc (min (epred_numeral k) n)" by(simp_all add: numeral_eq_eSuc) lemma enat_cocase_numeral [simp]: "co.case_enat a f (numeral v) = (let pv = epred_numeral v in f pv)" by(simp add: numeral_eq_eSuc) lemma enat_cocase_add_eq_if [simp]: "co.case_enat a f ((numeral v) + n) = (let pv = epred_numeral v in f (pv + n))" by(simp add: numeral_eq_eSuc iadd_Suc) lemma [simp]: shows epred_1: "epred 1 = 0" and epred_numeral: "epred (numeral i) = epred_numeral i" and epred_Infty: "epred \ = \" and epred_enat: "epred (enat m) = enat (m - 1)" by(simp_all add: epred_conv_minus one_enat_def zero_enat_def eSuc_def epred_numeral_def numeral_eq_enat split: enat.split) lemmas epred_simps = epred_0 epred_1 epred_numeral epred_eSuc epred_Infty epred_enat lemma epred_iadd1: "a \ 0 \ epred (a + b) = epred a + b" apply(cases a b rule: enat.exhaust[case_product enat.exhaust]) apply(simp_all add: epred_conv_minus eSuc_def one_enat_def zero_enat_def split: enat.splits) done lemma epred_min [simp]: "epred (min a b) = min (epred a) (epred b)" by(cases a b rule: enat_coexhaust[case_product enat_coexhaust]) simp_all lemma epred_le_epredI: "n \ m \ epred n \ epred m" by(cases m n rule: enat_coexhaust[case_product enat_coexhaust]) simp_all lemma epred_minus_epred [simp]: "m \ 0 \ epred n - epred m = n - m" by(cases n m rule: enat_coexhaust[case_product enat_coexhaust])(simp_all add: epred_conv_minus) lemma eSuc_epred: "n \ 0 \ eSuc (epred n) = n" by(cases n rule: enat_coexhaust) simp_all lemma epred_inject: "\ x \ 0; y \ 0 \ \ epred x = epred y \ x = y" by(cases x y rule: enat.exhaust[case_product enat.exhaust])(auto simp add: zero_enat_def) lemma monotone_fun_eSuc[partial_function_mono]: "monotone (fun_ord (\y x. x \ y)) (\y x. x \ y) (\f. eSuc (f x))" by (auto simp: monotone_def fun_ord_def) partial_function (gfp) enat_unfold :: "('a \ bool) \ ('a \ 'a) \ 'a \ enat" where enat_unfold [code, nitpick_simp]: "enat_unfold stop next a = (if stop a then 0 else eSuc (enat_unfold stop next (next a)))" lemma enat_unfold_stop [simp]: "stop a \ enat_unfold stop next a = 0" by(simp add: enat_unfold) lemma enat_unfold_next: "\ stop a \ enat_unfold stop next a = eSuc (enat_unfold stop next (next a))" by(simp add: enat_unfold) lemma enat_unfold_eq_0 [simp]: "enat_unfold stop next a = 0 \ stop a" by(simp add: enat_unfold) lemma epred_enat_unfold [simp]: "epred (enat_unfold stop next a) = (if stop a then 0 else enat_unfold stop next (next a))" by(simp add: enat_unfold_next) lemma epred_max: "epred (max x y) = max (epred x) (epred y)" by(cases x y rule: enat.exhaust[case_product enat.exhaust]) simp_all lemma epred_Max: assumes "finite A" "A \ {}" shows "epred (Max A) = Max (epred ` A)" using assms proof induction case (insert x A) thus ?case by(cases "A = {}")(simp_all add: epred_max) qed simp lemma finite_imageD2: "\ finite (f ` A); inj_on f (A - B); finite B \ \ finite A" by (metis Diff_subset finite_Diff2 image_mono inj_on_finite) lemma epred_Sup: "epred (Sup A) = Sup (epred ` A)" by(auto 4 4 simp add: bot_enat_def Sup_enat_def epred_Max inj_on_def neq_zero_conv_eSuc dest: finite_imageD2[where B="{0}"]) subsection \Less as greatest fixpoint\ coinductive_set Le_enat :: "(enat \ enat) set" where Le_enat_zero: "(0, n) \ Le_enat" | Le_enat_add: "\ (m, n) \ Le_enat; k \ 0 \ \ (eSuc m, n + k) \ Le_enat" lemma ile_into_Le_enat: "m \ n \ (m, n) \ Le_enat" proof - assume "m \ n" hence "(m, n) \ {(m, n)|m n. m \ n}" by simp thus ?thesis proof coinduct case (Le_enat m n) hence "m \ n" by simp show ?case proof(cases "m = 0") case True hence ?Le_enat_zero by simp thus ?thesis .. next case False with \m \ n\ obtain m' n' where "m = eSuc m'" "n = n' + 1" "m' \ n'" by(cases m rule: enat_coexhaust, simp) (cases n rule: enat_coexhaust, auto simp add: eSuc_plus_1[symmetric]) hence ?Le_enat_add by fastforce thus ?thesis .. qed qed qed lemma Le_enat_imp_ile_enat_k: "(m, n) \ Le_enat \ n < enat l \ m < enat l" proof(induct l arbitrary: m n) case 0 thus ?case by(simp add: zero_enat_def[symmetric]) next case (Suc l) from \(m, n) \ Le_enat\ show ?case proof cases case Le_enat_zero with \n < enat (Suc l)\ show ?thesis by auto next case (Le_enat_add M N K) from \n = N + K\ \n < enat (Suc l)\ \K \ 0\ have "N < enat l" by(cases N)(cases K, auto simp add: zero_enat_def) with \(M, N) \ Le_enat\ have "M < enat l" by(rule Suc) with \m = eSuc M\ show ?thesis by(simp add: eSuc_enat[symmetric]) qed qed lemma enat_less_imp_le: assumes k: "!!k. n < enat k \ m < enat k" shows "m \ n" proof(cases n) case (enat n') with k[of "Suc n'"] show ?thesis by(cases m) simp_all qed simp lemma Le_enat_imp_ile: "(m, n) \ Le_enat \ m \ n" by(rule enat_less_imp_le)(erule Le_enat_imp_ile_enat_k) lemma Le_enat_eq_ile: "(m, n) \ Le_enat \ m \ n" by(blast intro: Le_enat_imp_ile ile_into_Le_enat) lemma enat_leI [consumes 1, case_names Leenat, case_conclusion Leenat zero eSuc]: assumes major: "(m, n) \ X" and step: "\m n. (m, n) \ X \ m = 0 \ (\m' n' k. m = eSuc m' \ n = n' + enat k \ k \ 0 \ ((m', n') \ X \ m' \ n'))" shows "m \ n" apply(rule Le_enat.coinduct[unfolded Le_enat_eq_ile, where X="\x y. (x, y) \ X"]) apply(fastforce simp add: zero_enat_def dest: step intro: major)+ done lemma enat_le_coinduct[consumes 1, case_names le, case_conclusion le 0 eSuc]: assumes P: "P m n" and step: "\m n. P m n \ (n = 0 \ m = 0) \ (m \ 0 \ n \ 0 \ (\k n'. P (epred m) n' \ epred n = n' + k) \ epred m \ epred n)" shows "m \ n" proof - from P have "(m, n) \ {(m, n). P m n}" by simp thus ?thesis proof(coinduct rule: enat_leI) case (Leenat m n) hence "P m n" by simp show ?case proof(cases m rule: enat_coexhaust) case 0 thus ?thesis by simp next case (eSuc m') with step[OF \P m n\] have "n \ 0" and disj: "(\k n'. P m' n' \ epred n = n' + k) \ m' \ epred n" by auto from \n \ 0\ obtain n' where n': "n = eSuc n'" by(cases n rule: enat_coexhaust) auto from disj show ?thesis proof assume "m' \ epred n" with eSuc n' show ?thesis by(auto 4 3 intro: exI[where x="Suc 0"] simp add: eSuc_enat[symmetric] iadd_Suc_right zero_enat_def[symmetric]) next assume "\k n'. P m' n' \ epred n = n' + k" then obtain k n'' where n'': "epred n = n'' + k" and k: "P m' n''" by blast show ?thesis using eSuc k n' n'' by(cases k)(auto 4 3 simp add: iadd_Suc_right[symmetric] eSuc_enat intro: exI[where x=\]) qed qed qed qed subsection \Equality as greatest fixpoint\ lemma enat_equalityI [consumes 1, case_names Eq_enat, case_conclusion Eq_enat zero eSuc]: assumes major: "(m, n) \ X" and step: "\m n. (m, n) \ X \ m = 0 \ n = 0 \ (\m' n'. m = eSuc m' \ n = eSuc n' \ ((m', n') \ X \ m' = n'))" shows "m = n" proof(rule antisym) from major show "m \ n" by(coinduct rule: enat_leI) (drule step, auto simp add: eSuc_plus_1 enat_1[symmetric]) from major have "(n, m) \ {(n, m). (m, n) \ X}" by simp thus "n \ m" proof(coinduct rule: enat_leI) case (Leenat n m) hence "(m, n) \ X" by simp from step[OF this] show ?case by(auto simp add: eSuc_plus_1 enat_1[symmetric]) qed qed lemma enat_coinduct [consumes 1, case_names Eq_enat, case_conclusion Eq_enat zero eSuc]: assumes major: "P m n" and step: "\m n. P m n \ (m = 0 \ n = 0) \ (m \ 0 \ n \ 0 \ P (epred m) (epred n) \ epred m = epred n)" shows "m = n" proof - from major have "(m, n) \ {(m, n). P m n}" by simp thus ?thesis proof(coinduct rule: enat_equalityI) case (Eq_enat m n) hence "P m n" by simp from step[OF this] show ?case by(cases m n rule: enat_coexhaust[case_product enat_coexhaust]) auto qed qed lemma enat_coinduct2 [consumes 1, case_names zero eSuc]: "\ P m n; \m n. P m n \ m = 0 \ n = 0; \m n. \ P m n; m \ 0; n \ 0 \ \ P (epred m) (epred n) \ epred m = epred n \ \ m = n" by(erule enat_coinduct) blast subsection \Uniqueness of corecursion\ lemma enat_unfold_unique: assumes h: "!!x. h x = (if stop x then 0 else eSuc (h (next x)))" shows "h x = enat_unfold stop next x" by(coinduction arbitrary: x rule: enat_coinduct)(subst (1 3) h, auto) subsection \Setup for partial\_function\ lemma enat_diff_cancel_left: "\ m \ x; m \ y \ \ x - m = y - m \ x = (y :: enat)" by(cases x y m rule: enat.exhaust[case_product enat.exhaust[case_product enat.exhaust]])(simp_all, arith) lemma finite_lessThan_enatI: assumes "m \ \" shows "finite {.. enat ` {.. finite {..<\ :: enat}" proof have "range enat \ {..<\}" by auto moreover assume "finite {..<\ :: enat}" ultimately have "finite (range enat)" by(rule finite_subset) hence "finite (UNIV :: nat set)" by(rule finite_imageD)(simp add: inj_on_def) thus False by simp qed lemma finite_lessThan_enat_iff: "finite {.. m \ \" by(cases m)(auto intro: finite_lessThan_enatI simp add: infinite_lessThan_infty) lemma enat_minus_mono1: "x \ y \ x - m \ y - (m :: enat)" by(cases m x y rule: enat.exhaust[case_product enat.exhaust[case_product enat.exhaust]]) simp_all lemma max_enat_minus1: "max n m - k = max (n - k) (m - k :: enat)" by(cases n m k rule: enat.exhaust[case_product enat.exhaust[case_product enat.exhaust]]) simp_all lemma Max_enat_minus1: assumes "finite A" "A \ {}" shows "Max A - m = Max ((\n :: enat. n - m) ` A)" using assms proof induct case (insert x A) thus ?case by(cases "A = {}")(simp_all add: max_enat_minus1) qed simp lemma Sup_enat_minus1: assumes "m \ \" shows "\A - m = \((\n :: enat. n - m) ` A)" proof - from assms obtain m' where "m = enat m'" by auto thus ?thesis by(auto simp add: Sup_enat_def Max_enat_minus1 finite_lessThan_enat_iff enat_diff_cancel_left inj_on_def dest!: finite_imageD2[where B="{.. {}" shows "Sup ((\y :: enat. y+x) ` Y) = Sup Y + x" proof(cases "finite Y") case True thus ?thesis by(simp add: Sup_enat_def Max_add_commute assms) next case False thus ?thesis proof(cases x) case (enat x') hence "\ finite ((\y. y+x) ` Y)" using False by(auto dest!: finite_imageD intro: inj_onI) with False show ?thesis by(simp add: Sup_enat_def assms) next case infinity hence "(+) x ` Y = {\}" using assms by auto thus ?thesis using infinity by(simp add: image_constant_conv assms) qed qed lemma Sup_image_eadd2: "Y \ {} \ Sup ((\y :: enat. x + y) ` Y) = x + Sup Y" by(simp add: Sup_image_eadd1 add.commute) lemma mono2mono_eSuc [THEN lfp.mono2mono, cont_intro, simp]: shows monotone_eSuc: "monotone (\) (\) eSuc" by(rule monotoneI) simp lemma mcont2mcont_eSuc [THEN lfp.mcont2mcont, cont_intro, simp]: shows mcont_eSuc: "mcont Sup (\) Sup (\) eSuc" by(intro mcontI contI)(simp_all add: monotone_eSuc eSuc_Sup) lemma mono2mono_epred [THEN lfp.mono2mono, cont_intro, simp]: shows monotone_epred: "monotone (\) (\) epred" by(rule monotoneI)(simp add: epred_le_epredI) lemma mcont2mcont_epred [THEN lfp.mcont2mcont, cont_intro, simp]: shows mcont_epred: "mcont Sup (\) Sup (\) epred" by(simp add: mcont_def monotone_epred cont_def epred_Sup) lemma enat_cocase_mono [partial_function_mono, cont_intro]: "\ monotone orda ordb zero; \n. monotone orda ordb (\f. esuc f n) \ \ monotone orda ordb (\f. co.case_enat (zero f) (esuc f) x)" by(cases x rule: co.enat.exhaust) simp_all lemma enat_cocase_mcont [cont_intro, simp]: "\ mcont luba orda lubb ordb zero; \n. mcont luba orda lubb ordb (\f. esuc f n) \ \ mcont luba orda lubb ordb (\f. co.case_enat (zero f) (esuc f) x)" by(cases x rule: co.enat.exhaust) simp_all lemma eSuc_mono [partial_function_mono]: "monotone (fun_ord (\)) (\) f \ monotone (fun_ord (\)) (\) (\x. eSuc (f x))" by(rule mono2mono_eSuc) lemma mono2mono_enat_minus1 [THEN lfp.mono2mono, cont_intro, simp]: shows monotone_enat_minus1: "monotone (\) (\) (\n. n - m :: enat)" by(rule monotoneI)(rule enat_minus_mono1) lemma mcont2mcont_enat_minus [THEN lfp.mcont2mcont, cont_intro, simp]: shows mcont_enat_minus: "m \ \ \ mcont Sup (\) Sup (\) (\n. n - m :: enat)" by(rule mcontI)(simp_all add: monotone_enat_minus1 contI Sup_enat_minus1) lemma monotone_eadd1: "monotone (\) (\) (\x. x + y :: enat)" by(auto intro!: monotoneI) lemma monotone_eadd2: "monotone (\) (\) (\y. x + y :: enat)" by(auto intro!: monotoneI) lemma mono2mono_eadd[THEN lfp.mono2mono2, cont_intro, simp]: shows monotone_eadd: "monotone (rel_prod (\) (\)) (\) (\(x, y). x + y :: enat)" by(simp add: monotone_eadd1 monotone_eadd2) lemma mcont_eadd2: "mcont Sup (\) Sup (\) (\y. x + y :: enat)" by(auto intro: mcontI monotone_eadd2 contI Sup_image_eadd2[symmetric]) lemma mcont_eadd1: "mcont Sup (\) Sup (\) (\x. x + y :: enat)" by(auto intro: mcontI monotone_eadd1 contI Sup_image_eadd1[symmetric]) lemma mcont2mcont_eadd [cont_intro, simp]: "\ mcont lub ord Sup (\) (\x. f x); mcont lub ord Sup (\) (\x. g x) \ \ mcont lub ord Sup (\) (\x. f x + g x :: enat)" by(best intro: ccpo.mcont2mcont'[OF complete_lattice_ccpo] mcont_eadd1 mcont_eadd2 ccpo.mcont_const[OF complete_lattice_ccpo]) lemma eadd_partial_function_mono [partial_function_mono]: "\ monotone (fun_ord (\)) (\) f; monotone (fun_ord (\)) (\) g \ \ monotone (fun_ord (\)) (\) (\x. f x + g x :: enat)" by(rule mono2mono_eadd) lemma monotone_max_enat1: "monotone (\) (\) (\x. max x y :: enat)" by(auto intro!: monotoneI simp add: max_def) lemma monotone_max_enat2: "monotone (\) (\) (\y. max x y :: enat)" by(auto intro!: monotoneI simp add: max_def) lemma mono2mono_max_enat[THEN lfp.mono2mono2, cont_intro, simp]: shows monotone_max_enat: "monotone (rel_prod (\) (\)) (\) (\(x, y). max x y :: enat)" by(simp add: monotone_max_enat1 monotone_max_enat2) lemma max_Sup_enat2: assumes "Y \ {}" shows "max x (Sup Y) = Sup ((\y :: enat. max x y) ` Y)" proof(cases "finite Y") case True hence "max x (Max Y) = Max (max x ` Y)" using assms proof(induct) case (insert y Y) thus ?case by(cases "Y = {}")(simp_all, metis max.assoc max.left_commute max.left_idem) qed simp thus ?thesis using True by(simp add: Sup_enat_def assms) next case False show ?thesis proof(cases x) case infinity hence "max x ` Y = {\}" using assms by auto thus ?thesis using False by(simp add: Sup_enat_def assms) next case (enat x') { assume "finite (max x ` Y)" hence "finite (max x ` {y \ Y. y > x})" - by(rule finite_subset[rotated]) auto + by (rule rev_finite_subset) (auto simp add: image_iff dest: max.absorb4 sym) hence "finite {y \ Y. y > x}" by(rule finite_imageD)(auto intro!: inj_onI simp add: max_def split: if_split_asm) moreover have "finite {y \ Y. y \ x}" by(rule finite_enat_bounded)(auto simp add: enat) ultimately have "finite ({y \ Y. y > x} \ {y \ Y. y \ x})" by simp also have "{y \ Y. y > x} \ {y \ Y. y \ x} = Y" by auto finally have "finite Y" . } thus ?thesis using False by(auto simp add: Sup_enat_def assms) qed qed lemma max_Sup_enat1: "Y \ {} \ max (Sup Y) x = Sup ((\y :: enat. max y x) ` Y)" by(subst (1 2) max.commute)(rule max_Sup_enat2) lemma mcont_max_enat1: "mcont Sup (\) Sup (\) (\x. max x y :: enat)" by(auto intro!: mcontI contI max_Sup_enat1 simp add: monotone_max_enat1) lemma mcont_max_enat2: "mcont Sup (\) Sup (\) (\y. max x y :: enat)" by(auto intro!: mcontI contI max_Sup_enat2 simp add: monotone_max_enat2) lemma mcont2mcont_max_enat [cont_intro, simp]: "\ mcont lub ord Sup (\) (\x. f x); mcont lub ord Sup (\) (\x. g x) \ \ mcont lub ord Sup (\) (\x. max (f x) (g x) :: enat)" by(best intro: ccpo.mcont2mcont'[OF complete_lattice_ccpo] mcont_max_enat1 mcont_max_enat2 ccpo.mcont_const[OF complete_lattice_ccpo]) lemma max_enat_partial_function_mono [partial_function_mono]: "\ monotone (fun_ord (\)) (\) f; monotone (fun_ord (\)) (\) g \ \ monotone (fun_ord (\)) (\) (\x. max (f x) (g x) :: enat)" by(rule mono2mono_max_enat) lemma chain_epredI: "Complete_Partial_Order.chain (\) Y \ Complete_Partial_Order.chain (\) (epred ` (Y \ {x. x \ 0}))" by(auto intro: chainI dest: chainD) lemma monotone_enat_le_case: fixes bot assumes mono: "monotone (\) ord (\x. f x (eSuc x))" and ord: "\x. ord bot (f x (eSuc x))" and bot: "ord bot bot" shows "monotone (\) ord (\x. case x of 0 \ bot | eSuc x' \ f x' x)" proof - have "monotone (\) ord (\x. if x \ 0 then bot else f (epred x) x)" proof(rule monotone_if_bot) fix x y :: enat assume "x \ y" "\ x \ 0" thus "ord (f (epred x) x) (f (epred y) y)" by(cases x y rule: co.enat.exhaust[case_product co.enat.exhaust])(auto intro: monotoneD[OF mono]) next fix x :: enat assume "\ x \ 0" thus "ord bot (f (epred x) x)" by(cases x rule: co.enat.exhaust)(auto intro: ord) qed(rule bot) also have "(\x. if x \ 0 then bot else f (epred x) x) = (\x. case x of 0 \ bot | eSuc x' \ f x' x)" by(auto simp add: fun_eq_iff split: co.enat.split) finally show ?thesis . qed lemma mcont_enat_le_case: fixes bot assumes ccpo: "class.ccpo lub ord (mk_less ord)" and mcont: "mcont Sup (\) lub ord (\x. f x (eSuc x))" and ord: "\x. ord bot (f x (eSuc x))" shows "mcont Sup (\) lub ord (\x. case x of 0 \ bot | eSuc x' \ f x' x)" proof - from ccpo have "mcont Sup (\) lub ord (\x. if x \ 0 then bot else f (epred x) x)" proof(rule mcont_if_bot) fix x y :: enat assume "x \ y" "\ x \ 0" thus "ord (f (epred x) x) (f (epred y) y)" by(cases x y rule: co.enat.exhaust[case_product co.enat.exhaust])(auto intro: mcont_monoD[OF mcont]) next fix Y :: "enat set" assume chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" "\x. x \ Y \ \ x \ 0" from Y have Y': "Y \ {x. x \ 0} \ {}" by auto from Y(2) have eq: "Y = eSuc ` (epred ` (Y \ {x. x \ 0}))" by(fastforce intro: rev_image_eqI) let ?Y = "epred ` (Y \ {x. x \ 0})" from chain_epredI [OF chain] Y' have "f (\?Y) (eSuc (\?Y)) = lub ((\x. f x (eSuc x)) ` ?Y)" using mcont [THEN mcont_contD] by blast moreover from chain_epredI [OF chain] Y' have "\(eSuc ` ?Y) = eSuc (\?Y)" using mcont_eSuc [THEN mcont_contD, symmetric] by blast ultimately show "f (epred (Sup Y)) (Sup Y) = lub ((\x. f (epred x) x) ` Y)" by (subst (1 2 3) eq) (simp add: image_image) next fix x :: enat assume "\ x \ 0" thus "ord bot (f (epred x) x)" by(cases x rule: co.enat.exhaust)(auto intro: ord) qed also have "(\x. if x \ 0 then bot else f (epred x) x) = (\x. case x of 0 \ bot | eSuc x' \ f x' x)" by(auto simp add: fun_eq_iff split: co.enat.split) finally show ?thesis . qed subsection \Misc.\ lemma enat_add_mono [simp]: "enat x + y < enat x + z \ y < z" by(cases y)(case_tac [!] z, simp_all) lemma enat_add1_eq [simp]: "enat x + y = enat x + z \ y = z" by (metis enat_add_mono add.commute neq_iff) lemma enat_add2_eq [simp]: "y + enat x = z + enat x \ y = z" by (metis enat_add1_eq add.commute) lemma enat_less_enat_plusI: "x < y \ enat x < enat y + z" by(cases z) simp_all lemma enat_less_enat_plusI2: "enat y < z \ enat (x + y) < enat x + z" by (metis enat_add_mono plus_enat_simps(1)) lemma min_enat1_conv_enat: "\a b. min (enat a) b = enat (case b of enat b' \ min a b' | \ \ a)" and min_enat2_conv_enat: "\a b. min a (enat b) = enat (case a of enat a' \ min a' b | \ \ b)" by(simp_all split: enat.split) lemma eSuc_le_iff: "eSuc x \ y \ (\y'. y = eSuc y' \ x \ y')" by(cases y rule: co.enat.exhaust) simp_all lemma eSuc_eq_infinity_iff: "eSuc n = \ \ n = \" by(cases n)(simp_all add: zero_enat_def eSuc_enat) lemma infinity_eq_eSuc_iff: "\ = eSuc n \ n = \" by(cases n)(simp_all add: zero_enat_def eSuc_enat) lemma enat_cocase_inf: "(case \ of 0 \ a | eSuc b \ f b) = f \" by(auto split: co.enat.split simp add: infinity_eq_eSuc_iff) lemma eSuc_Inf: "eSuc (Inf A) = Inf (eSuc ` A)" proof - { assume "A \ {}" then obtain a where "a \ A" by blast then have "eSuc (LEAST a. a \ A) = (LEAST a. a \ eSuc ` A)" proof (rule LeastI2_wellorder) fix a assume "a \ A" and b: "\b. b \ A \ a \ b" then have a: "eSuc a \ eSuc ` A" by auto then show "eSuc a = (LEAST a. a \ eSuc ` A)" by (rule LeastI2_wellorder) (metis (full_types) b a antisym eSuc_le_iff imageE) qed } then show ?thesis by (simp add: Inf_enat_def) qed end diff --git a/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy b/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy --- a/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy +++ b/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy @@ -1,3088 +1,3090 @@ (* File: Dirichlet_Series_Analysis.thy Author: Manuel Eberl, TU München *) section \Analytic properties of Dirichlet series\ theory Dirichlet_Series_Analysis imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Library.Going_To_Filter" "HOL-Real_Asymp.Real_Asymp" Dirichlet_Series Moebius_Mu Partial_Summation Euler_Products begin (* TODO Move *) lemma frequently_eventually_frequently: "frequently P F \ eventually Q F \ frequently (\x. P x \ Q x) F" by (erule frequently_rev_mp, erule eventually_mono) auto text \ The following illustrates a concept we will need later on: A property holds for \f\ going to \F\ if we can find e.\,g.\ a sequence that tends to \F\ and whose elements eventually satisfy \P\. \ lemma frequently_going_toI: assumes "filterlim (\n. f (g n)) F G" assumes "eventually (\n. P (g n)) G" assumes "eventually (\n. g n \ A) G" assumes "G \ bot" shows "frequently P (f going_to F within A)" unfolding frequently_def proof assume "eventually (\x. \P x) (f going_to F within A)" hence "eventually (\x. \P x) (inf (filtercomap f F) (principal A))" by (simp add: going_to_within_def) moreover have "filterlim (\n. g n) (inf (filtercomap f F) (principal A)) G" using assms unfolding filterlim_inf filterlim_principal by (auto simp add: filterlim_iff_le_filtercomap filtercomap_filtercomap) ultimately have "eventually (\n. \P (g n)) G" by (rule eventually_compose_filterlim) with assms(2) have "eventually (\_. False) G" by eventually_elim auto with assms(4) show False by simp qed lemma frequently_filtercomapI: assumes "filterlim (\n. f (g n)) F G" assumes "eventually (\n. P (g n)) G" assumes "G \ bot" shows "frequently P (filtercomap f F)" using frequently_going_toI[of f g F G P UNIV] assms by (simp add: going_to_def) lemma frequently_going_to_at_topE: fixes f :: "'a \ real" assumes "frequently P (f going_to at_top)" obtains g where "\n. P (g n)" and "filterlim (\n. f (g n)) at_top sequentially" proof - from assms have "\k. \x. f x \ real k \ P x" by (auto simp: frequently_def eventually_going_to_at_top_linorder) hence "\g. \k. f (g k) \ real k \ P (g k)" by metis then obtain g where g: "\k. f (g k) \ real k" "\k. P (g k)" by blast have "filterlim (\n. f (g n)) at_top sequentially" by (rule filterlim_at_top_mono[OF filterlim_real_sequentially]) (use g in auto) from g(2) and this show ?thesis using that[of g] by blast qed text \ Apostol often uses statements like `$P(s_k)$ for all $k$ in an infinite sequence $s_k$ such that $\mathfrak{R}(s_k)\longrightarrow\infty$ as $k\to\infty$'. Instead, we write @{prop "frequently P (Re going_to at_top)"}. This lemma shows that our statement is equivalent to his. \ lemma frequently_going_to_at_top_iff: "frequently P (f going_to (at_top :: real filter)) \ (\g. \n. P (g n) \ filterlim (\n. f (g n)) at_top sequentially)" by (auto intro: frequently_going_toI elim!: frequently_going_to_at_topE) (* END TODO *) lemma surj_bullet_1: "surj (\s::'a::{real_normed_algebra_1, real_inner}. s \ 1)" proof (rule surjI) fix x :: real show "(x *\<^sub>R 1) \ (1 :: 'a) = x" by (simp add: dot_square_norm) qed lemma bullet_1_going_to_at_top_neq_bot [simp]: "((\s::'a::{real_normed_algebra_1, real_inner}. s \ 1) going_to at_top) \ bot" unfolding going_to_def by (rule filtercomap_neq_bot_surj[OF _ surj_bullet_1]) auto lemma fds_abs_converges_altdef: "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on {1..}" by (auto simp add: fds_abs_converges_def abs_summable_on_nat_iff intro!: summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) lemma fds_abs_converges_altdef': "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on UNIV" by (subst fds_abs_converges_altdef, rule abs_summable_on_cong_neutral) (auto simp: Suc_le_eq) lemma eval_fds_altdef: assumes "fds_abs_converges f s" shows "eval_fds f s = (\\<^sub>an. fds_nth f n / nat_power n s)" proof - have "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on UNIV" unfolding fds_abs_converges_altdef by (intro abs_summable_on_cong_neutral) (auto simp: Suc_le_eq) with assms show ?thesis unfolding eval_fds_def fds_abs_converges_altdef by (intro infsetsum_nat' [symmetric]) simp_all qed lemma multiplicative_function_divide_nat_power: fixes f :: "nat \ 'a :: {nat_power, field}" assumes "multiplicative_function f" shows "multiplicative_function (\n. f n / nat_power n s)" proof interpret f: multiplicative_function f by fact show "f 0 / nat_power 0 s = 0" "f 1 / nat_power 1 s = 1" by simp_all fix a b :: nat assume "a > 1" "b > 1" "coprime a b" thus "f (a * b) / nat_power (a * b) s = f a / nat_power a s * (f b / nat_power b s)" by (simp_all add: f.mult_coprime nat_power_mult_distrib) qed lemma completely_multiplicative_function_divide_nat_power: fixes f :: "nat \ 'a :: {nat_power, field}" assumes "completely_multiplicative_function f" shows "completely_multiplicative_function (\n. f n / nat_power n s)" proof interpret f: completely_multiplicative_function f by fact show "f 0 / nat_power 0 s = 0" "f (Suc 0) / nat_power (Suc 0) s = 1" by simp_all fix a b :: nat assume "a > 1" "b > 1" thus "f (a * b) / nat_power (a * b) s = f a / nat_power a s * (f b / nat_power b s)" by (simp_all add: f.mult nat_power_mult_distrib) qed subsection \Convergence and absolute convergence\ class nat_power_normed_field = nat_power_field + real_normed_field + real_inner + real_algebra_1 + fixes real_power :: "real \ 'a \ 'a" assumes real_power_nat_power: "n > 0 \ real_power (real n) c = nat_power n c" assumes real_power_1_right_aux: "d > 0 \ real_power d 1 = d *\<^sub>R 1" assumes real_power_add: "d > 0 \ real_power d (a + b) = real_power d a * real_power d b" assumes real_power_nonzero [simp]: "d > 0 \ real_power d a \ 0" assumes norm_real_power: "x > 0 \ norm (real_power x c) = x powr (c \ 1)" assumes nat_power_of_real_aux: "nat_power n (x *\<^sub>R 1) = ((real n powr x) *\<^sub>R 1)" assumes has_field_derivative_nat_power_aux: "\x::'a. n > 0 \ LIM y inf_class.inf (Inf (principal ` {S. open S \ x \ S})) (principal (UNIV - {x})). (nat_power n y - nat_power n x - ln (real n) *\<^sub>R nat_power n x * (y - x)) /\<^sub>R norm (y - x) :> Inf (principal ` {S. open S \ 0 \ S})" assumes has_vector_derivative_real_power_aux: "x > 0 \ filterlim (\y. (real_power y c - real_power x (c :: 'a) - (y - x) *\<^sub>R (c * real_power x (c - 1))) /\<^sub>R norm (y - x)) (INF S\{S. open S \ 0 \ S}. principal S) (at x)" assumes norm_nat_power: "n > 0 \ norm (nat_power n y) = real n powr (y \ 1)" begin lemma real_power_diff: "d > 0 \ real_power d (a - b) = real_power d a / real_power d b" using real_power_add[of d b "a - b"] by (simp add: field_simps) end lemma real_power_1_right [simp]: "d > 0 \ real_power d 1 = of_real d" using real_power_1_right_aux[of d] by (simp add: scaleR_conv_of_real) lemma has_vector_derivative_real_power [derivative_intros]: "x > 0 \ ((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x within A)" by (rule has_vector_derivative_at_within) (insert has_vector_derivative_real_power_aux[of x c], simp add: has_vector_derivative_def has_derivative_def nhds_def bounded_linear_scaleR_left) lemma has_field_derivative_nat_power [derivative_intros]: "n > 0 \ ((\y. nat_power n y) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at (x :: 'a :: nat_power_normed_field) within A)" by (rule has_field_derivative_at_within) (insert has_field_derivative_nat_power_aux[of n x], simp only: has_field_derivative_def has_derivative_def netlimit_at, simp add: nhds_def at_within_def bounded_linear_mult_right) lemma continuous_on_real_power [continuous_intros]: "A \ {0<..} \ continuous_on A (\x. real_power x s)" by (rule continuous_on_vector_derivative has_vector_derivative_real_power)+ auto instantiation real :: nat_power_normed_field begin definition real_power_real :: "real \ real \ real" where [simp]: "real_power_real = (powr)" instance proof (standard, goal_cases) case (7 n x) hence "((\x. nat_power n x) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at x)" by (auto intro!: derivative_eq_intros simp: powr_def) thus ?case unfolding has_field_derivative_def netlimit_at has_derivative_def by (simp add: nhds_def at_within_def) next case (8 x c) hence "((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x)" by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) thus ?case by (simp add: has_vector_derivative_def has_derivative_def nhds_def) qed (simp_all add: powr_add) end instantiation complex :: nat_power_normed_field begin definition nat_power_complex :: "nat \ complex \ complex" where [simp]: "nat_power_complex n z = of_nat n powr z" definition real_power_complex :: "real \ complex \ complex" where [simp]: "real_power_complex = (\x y. of_real x powr y)" instance proof fix m n :: nat and z :: complex assume "m > 0" "n > 0" thus "nat_power (m * n) z = nat_power m z * nat_power n z" unfolding nat_power_complex_def of_nat_mult by (subst powr_times_real) simp_all next fix n :: nat and z :: complex assume "n > 0" show "norm (nat_power n z) = real n powr (z \ 1)" unfolding nat_power_complex_def using norm_powr_real_powr[of "of_nat n" z] by simp next fix n :: nat and x :: complex assume n: "n > 0" hence "((\x. nat_power n x) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at x)" by (auto intro!: derivative_eq_intros simp: powr_def scaleR_conv_of_real mult_ac) thus "LIM y inf_class.inf (Inf (principal ` {S. open S \ x \ S})) (principal (UNIV - {x})). (nat_power n y - nat_power n x - ln (real n) *\<^sub>R nat_power n x * (y - x)) /\<^sub>R cmod (y - x) :> (Inf (principal ` {S. open S \ 0 \ S}))" unfolding has_field_derivative_def netlimit_at has_derivative_def by (simp add: nhds_def at_within_def) next fix x :: real and c :: complex assume "x > 0" hence "((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x)" by (auto intro!: derivative_eq_intros has_vector_derivative_real_field) thus "LIM y at x. (real_power y c - real_power x c - (y - x) *\<^sub>R (c * real_power x (c - 1))) /\<^sub>R norm (y - x) :> INF S\{S. open S \ 0 \ S}. principal S" by (simp add: has_vector_derivative_def has_derivative_def nhds_def) next fix n :: nat and x :: real show "nat_power n (x *\<^sub>R 1 :: complex) = (real n powr x) *\<^sub>R 1" by (simp add: powr_Reals_eq scaleR_conv_of_real) qed (auto simp: powr_def exp_add exp_of_nat_mult [symmetric] algebra_simps scaleR_conv_of_real simp del: Ln_of_nat) end lemma nat_power_of_real [simp]: "nat_power n (of_real x :: 'a :: nat_power_normed_field) = of_real (real n powr x)" using nat_power_of_real_aux[of n x] by (simp add: scaleR_conv_of_real) lemma fds_abs_converges_of_real [simp]: "fds_abs_converges (fds_of_real f) (of_real s :: 'a :: {nat_power_normed_field,banach}) \ fds_abs_converges f s" unfolding fds_abs_converges_def by (subst (1 2) summable_Suc_iff [symmetric]) (simp add: norm_divide norm_nat_power) lemma eval_fds_of_real [simp]: assumes "fds_converges f s" shows "eval_fds (fds_of_real f) (of_real s :: 'a :: {nat_power_normed_field,banach}) = of_real (eval_fds f s)" using assms unfolding eval_fds_def by (auto simp: fds_converges_def suminf_of_real) lemma fds_abs_summable_zeta_iff [simp]: fixes s :: "'a :: {banach, nat_power_normed_field}" shows "fds_abs_converges fds_zeta s \ s \ 1 > (1 :: real)" proof - have "fds_abs_converges fds_zeta s \ summable (\n. real n powr -(s \ 1))" unfolding fds_abs_converges_def by (intro summable_cong always_eventually) (auto simp: norm_divide fds_nth_zeta powr_minus norm_nat_power divide_simps) also have "\ \ s \ 1 > 1" by (simp add: summable_real_powr_iff) finally show ?thesis . qed lemma fds_abs_summable_zeta: "(s :: 'a :: {banach, nat_power_normed_field}) \ 1 > 1 \ fds_abs_converges fds_zeta s" by simp lemma fds_abs_converges_moebius_mu: fixes s :: "'a :: {banach,nat_power_normed_field}" assumes "s \ 1 > 1" shows "fds_abs_converges (fds moebius_mu) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (norm (fds_nth (fds moebius_mu) n / nat_power n s)) \ real n powr (-s \ 1)" by (auto simp: powr_minus divide_simps abs_moebius_mu_le norm_nat_power norm_divide moebius_mu_def norm_power) next from assms show "summable (\n. real n powr (-s \ 1))" by (simp add: summable_real_powr_iff) qed definition conv_abscissa :: "'a :: {nat_power,banach,real_normed_field, real_inner} fds \ ereal" where "conv_abscissa f = (INF s\{s. fds_converges f s}. ereal (s \ 1))" definition abs_conv_abscissa :: "'a :: {nat_power,banach,real_normed_field, real_inner} fds \ ereal" where "abs_conv_abscissa f = (INF s\{s. fds_abs_converges f s}. ereal (s \ 1))" lemma conv_abscissa_mono: assumes "\s. fds_converges g s \ fds_converges f s" shows "conv_abscissa f \ conv_abscissa g" unfolding conv_abscissa_def by (rule INF_mono) (use assms in auto) lemma abs_conv_abscissa_mono: assumes "\s. fds_abs_converges g s \ fds_abs_converges f s" shows "abs_conv_abscissa f \ abs_conv_abscissa g" unfolding abs_conv_abscissa_def by (rule INF_mono) (use assms in auto) class dirichlet_series = euclidean_space + real_normed_field + nat_power_normed_field + assumes one_in_Basis: "1 \ Basis" instance real :: dirichlet_series by standard simp_all instance complex :: dirichlet_series by standard (simp_all add: Basis_complex_def) context assumes "SORT_CONSTRAINT('a :: dirichlet_series)" begin lemma fds_abs_converges_Re_le: fixes f :: "'a fds" assumes "fds_abs_converges f z" "z \ 1 \ z' \ 1" shows "fds_abs_converges f z'" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat assume n: "n \ 1" thus "norm (norm (fds_nth f n / nat_power n z')) \ norm (fds_nth f n / nat_power n z)" using assms(2) by (simp add: norm_divide norm_nat_power divide_simps powr_mono mult_left_mono) qed (insert assms(1), simp add: fds_abs_converges_def) lemma fds_abs_converges: assumes "s \ 1 > abs_conv_abscissa (f :: 'a fds)" shows "fds_abs_converges f s" proof - from assms obtain s0 where "fds_abs_converges f s0" "s0 \ 1 < s \ 1" by (auto simp: INF_less_iff abs_conv_abscissa_def) with fds_abs_converges_Re_le[OF this(1), of s] this(2) show ?thesis by simp qed lemma fds_abs_diverges: assumes "s \ 1 < abs_conv_abscissa (f :: 'a fds)" shows "\fds_abs_converges f s" proof assume "fds_abs_converges f s" hence "abs_conv_abscissa f \ s \ 1" unfolding abs_conv_abscissa_def by (intro INF_lower) auto with assms show False by simp qed lemma uniformly_Cauchy_eval_fds_aux: fixes s0 :: "'a :: dirichlet_series" assumes bounded: "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_Cauchy_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (cases "B = {}") case False show ?thesis proof (rule uniformly_Cauchy_onI', goal_cases) case (1 \) define \ where "\ = Inf ((\s. s \ 1) ` B)" have \_le: "s \ 1 \ \" if "s \ B" for s unfolding \_def using that by (intro cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded B) auto have "\ \ ((\s. s \ 1) ` B)" unfolding \_def using B \B \ {}\ by (intro closed_contains_Inf bounded_inner_imp_bdd_below compact_imp_bounded B compact_imp_closed compact_continuous_image continuous_intros) auto with B(2) have \_gt: "\ > s0 \ 1" by auto define \ where "\ = \ - s0 \ 1" have "bounded B" by (rule compact_imp_bounded) fact then obtain norm_B_aux where norm_B_aux: "\s. s \ B \ norm s \ norm_B_aux" by (auto simp: bounded_iff) define norm_B where "norm_B = norm_B_aux + norm s0" from norm_B_aux have norm_B: "norm (s - s0) \ norm_B" if "s \ B" for s using norm_triangle_ineq4[of s s0] norm_B_aux[OF that] by (simp add: norm_B_def) then have "0 \ norm_B" by (meson \\ \ (\s. s \ 1) ` B\ imageE norm_ge_zero order.trans) define A where "A = sum_upto (\k. fds_nth f k / nat_power k s0)" from bounded obtain C_aux where C_aux: "\n. norm (\k\n. fds_nth f k / nat_power k s0) \ C_aux" by (auto simp: Bseq_def) define C where "C = max C_aux 1" have C_pos: "C > 0" by (simp add: C_def) have C: "norm (A x) \ C" for x proof - have "A x = (\k\nat \x\. fds_nth f k / nat_power k s0)" unfolding A_def sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C_aux" by (rule C_aux) also have "\ \ C" by (simp add: C_def) finally show ?thesis . qed have "(\m. 2 * C * (1 + norm_B / \) * real m powr (-\)) \ 0" unfolding \_def using \_gt by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially) simp_all from order_tendstoD(2)[OF this \\ > 0\] obtain M where M: "\m. m \ M \ 2 * C * (1 + norm_B / \) * real m powr - \ < \" by (auto simp: eventually_at_top_linorder) show ?case proof (intro exI[of _ "max M 1"] ballI allI impI, goal_cases) case (1 s m n) from 1 have s: "s \ 1 > s0 \ 1" using B(2)[of s] by simp have mn: "m \ M" "m < n" "m > 0" "n > 0" using 1 by (simp_all add: ) have "dist (\n\m. fds_nth f n / nat_power n s) (\n\n. fds_nth f n / nat_power n s) = dist (\n\n. fds_nth f n / nat_power n s) (\n\m. fds_nth f n / nat_power n s)" by (simp add: dist_commute) also from 1 have "\ = norm (\k\{..n}-{..m}. fds_nth f k / nat_power k s)" by (subst Groups_Big.sum_diff) (simp_all add: dist_norm) also from 1 have "{..n} - {..m} = real -` {real m<..real n}" by auto also have "(\k\\. fds_nth f k / nat_power k s) = (\k\\. fds_nth f k / nat_power k s0 * real_power (real k) (s0 - s))" (is "_ = ?S") by (intro sum.cong refl) (simp_all add: nat_power_diff real_power_nat_power) also have *: "((\t. A t * ((s0 - s) * real_power t (s0 - s - 1))) has_integral (A (real n) * real_power n (s0 - s) - A (real m) * real_power m (s0 - s) - ?S)) {real m..real n}" (is "(?h has_integral _) _") unfolding A_def using mn by (intro partial_summation_strong[of "{}"]) (auto intro!: derivative_eq_intros continuous_intros) hence "?S = A (real n) * nat_power n (s0 - s) - A (real m) * nat_power m (s0 - s) - integral {real m..real n} ?h" using mn by (simp add: has_integral_iff real_power_nat_power) also have "norm \ \ norm (A (real n) * nat_power n (s0 - s)) + norm (A (real m) * nat_power m (s0 - s)) + norm (integral {real m..real n} ?h)" by (intro order.trans[OF norm_triangle_ineq4] add_right_mono order.refl) also have "norm (A (real n) * nat_power n (s0 - s)) \ C * nat_power m ((s0 - s) \ 1)" using mn \s \ B\ C_pos s by (auto simp: norm_mult norm_nat_power algebra_simps intro!: mult_mono C powr_mono2') also have "norm (A (real m) * nat_power m (s0 - s)) \ C * nat_power m ((s0 - s) \ 1)" using mn by (auto simp: norm_mult norm_nat_power intro!: mult_mono C) also have "norm (integral {real m..real n} ?h) \ integral {real m..real n} (\t. C * (norm (s0 - s) * t powr ((s0 - s) \ 1 - 1)))" proof (intro integral_norm_bound_integral ballI, goal_cases) case 1 with * show ?case by (simp add: has_integral_iff) next case 2 from mn show ?case by (auto intro!: integrable_continuous_real continuous_intros) next case (3 t) thus ?case unfolding norm_mult using C_pos mn by (intro mult_mono C) (auto simp: norm_real_power dot_square_norm algebra_simps) qed also have "\ = C * norm (s0 - s) * integral {real m..real n} (\t. t powr ((s0 - s) \ 1 - 1))" by (simp add: algebra_simps dot_square_norm) also { have "((\t. t powr ((s0 - s) \ 1 - 1)) has_integral (real n powr ((s0 - s) \ 1) / ((s0 - s) \ 1) - real m powr ((s0 - s) \ 1) / ((s0 - s) \ 1))) {m..n}" (is "(?l has_integral ?I) _") using mn s by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] inner_diff_left) hence "integral {real m..real n} ?l = ?I" by (simp add: has_integral_iff) also have "\ \ -(real m powr ((s0 - s) \ 1) / ((s0 - s) \ 1))" using s mn by (simp add: divide_simps inner_diff_left) also have "\ = 1 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1))" using s by (simp add: field_simps inner_diff_left) also have "\ \ 2 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1))" using mn s by (intro mult_right_mono divide_nonneg_pos) (simp_all add: inner_diff_left) finally have "integral {m..n} ?l \ \" . } hence "C * norm (s0 - s) * integral {real m..real n} (\t. t powr ((s0 - s) \ 1 - 1)) \ C * norm (s0 - s) * (2 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1)))" using C_pos mn by (intro mult_mono mult_nonneg_nonneg integral_nonneg integrable_continuous_real continuous_intros) auto also have "C * nat_power m ((s0 - s) \ 1) + C * nat_power m ((s0 - s) \ 1) + \ = 2 * C * nat_power m ((s0 - s) \ 1) * (1 + norm (s - s0) / ((s - s0) \ 1))" by (simp add: algebra_simps norm_minus_commute) also have "\ \ 2 * C * nat_power m (-\) * (1 + norm_B / \)" using C_pos s mn \_le[of s] \s \ B\ \_gt \0 \ norm_B\ unfolding nat_power_real_def \_def by (intro mult_mono powr_mono frac_le add_mono norm_B; simp add: inner_diff_left) also have "\ = 2 * C * (1 + norm_B / \) * real m powr (-\)" by simp also from \m \ M\ have "\ < \" by (rule M) finally show ?case by - simp_all qed qed qed (auto simp: uniformly_Cauchy_on_def) lemma uniformly_convergent_eval_fds_aux: assumes "Bseq (\n. \k\n. fds_nth f k / nat_power k (s0 :: 'a))" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" by (rule Cauchy_uniformly_convergent uniformly_Cauchy_eval_fds_aux assms)+ lemma uniformly_convergent_eval_fds_aux': assumes conv: "fds_converges f (s0 :: 'a)" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (rule uniformly_convergent_eval_fds_aux) from conv have "convergent (\n. \k\n. fds_nth f k / nat_power k s0)" by (simp add: fds_converges_def summable_iff_convergent') thus "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" by (rule convergent_imp_Bseq) qed (insert assms, auto) lemma bounded_partial_sums_imp_fps_converges: fixes s0 :: "'a :: dirichlet_series" assumes "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" and "s \ 1 > s0 \ 1" shows "fds_converges f s" proof - have "uniformly_convergent_on {s} (\N z. \n\N. fds_nth f n / nat_power n z)" using assms(2) by (intro uniformly_convergent_eval_fds_aux[OF assms(1)]) auto thus ?thesis by (auto simp: fds_converges_def summable_iff_convergent' dest: uniformly_convergent_imp_convergent) qed theorem fds_converges_Re_le: assumes "fds_converges f (s0 :: 'a)" "s \ 1 > s0 \ 1" shows "fds_converges f s" proof - have "uniformly_convergent_on {s} (\N z. \n\N. fds_nth f n / nat_power n z)" by (rule uniformly_convergent_eval_fds_aux' assms)+ (insert assms(2), auto) then obtain l where "uniform_limit {s} (\N z. \n\N. fds_nth f n / nat_power n z) l at_top" by (auto simp: uniformly_convergent_on_def) from tendsto_uniform_limitI[OF this, of s] have "(\n. fds_nth f n / nat_power n s) sums l s" unfolding sums_def' by (simp add: atLeast0AtMost) thus ?thesis by (simp add: fds_converges_def sums_iff) qed lemma fds_converges: assumes "s \ 1 > conv_abscissa (f :: 'a fds)" shows "fds_converges f s" proof - from assms obtain s0 where "fds_converges f s0" "s0 \ 1 < s \ 1" by (auto simp: INF_less_iff conv_abscissa_def) with fds_converges_Re_le[OF this(1), of s] this(2) show ?thesis by simp qed lemma fds_diverges: assumes "s \ 1 < conv_abscissa (f :: 'a fds)" shows "\fds_converges f s" proof assume "fds_converges f s" hence "conv_abscissa f \ s \ 1" unfolding conv_abscissa_def by (intro INF_lower) auto with assms show False by simp qed theorem fds_converges_imp_abs_converges: assumes "fds_converges (f :: 'a fds) s" "s' \ 1 > s \ 1 + 1" shows "fds_abs_converges f s'" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from assms(2) show "summable (\n. real n powr ((s - s') \ 1))" by (subst summable_real_powr_iff) (simp_all add: inner_diff_left) next from assms(1) have "(\n. fds_nth f n / nat_power n s) \ 0" unfolding fds_converges_def by (rule summable_LIMSEQ_zero) from tendsto_norm[OF this] have "(\n. norm (fds_nth f n / nat_power n s)) \ 0" by simp hence "eventually (\n. norm (fds_nth f n / nat_power n s) < 1) at_top" by (rule order_tendstoD) simp_all thus "eventually (\n. norm (norm (fds_nth f n / nat_power n s')) \ real n powr ((s - s') \ 1)) at_top" proof eventually_elim case (elim n) thus ?case proof (cases "n = 0") case False have "norm (fds_nth f n / nat_power n s') = norm (fds_nth f n) / real n powr (s' \ 1)" using False by (simp add: norm_divide norm_nat_power) also have "\ = norm (fds_nth f n / nat_power n s) / real n powr ((s' - s) \ 1)" using False by (simp add: norm_divide norm_nat_power inner_diff_left powr_diff) also have "\ \ 1 / real n powr ((s' - s) \ 1)" using elim by (intro divide_right_mono elim) simp_all also have "\ = real n powr ((s - s') \ 1)" using False by (simp add: field_simps inner_diff_left powr_diff) finally show ?thesis by simp qed simp_all qed qed lemma conv_le_abs_conv_abscissa: "conv_abscissa f \ abs_conv_abscissa f" unfolding conv_abscissa_def abs_conv_abscissa_def by (intro INF_superset_mono) auto lemma conv_abscissa_PInf_iff: "conv_abscissa f = \ \ (\s. \fds_converges f s)" unfolding conv_abscissa_def by (subst Inf_eq_PInfty) auto lemma conv_abscissa_PInfI [intro]: "(\s. \fds_converges f s) \ conv_abscissa f = \" by (subst conv_abscissa_PInf_iff) auto lemma conv_abscissa_MInf_iff: "conv_abscissa (f :: 'a fds) = -\ \ (\s. fds_converges f s)" proof safe assume *: "\s. fds_converges f s" have "conv_abscissa f \ B" for B :: real using spec[OF *, of "of_real B"] fds_diverges[of "of_real B" f] by (cases "conv_abscissa f \ B") simp_all thus "conv_abscissa f = -\" by (rule ereal_bot) qed (auto intro: fds_converges) lemma conv_abscissa_MInfI [intro]: "(\s. fds_converges (f::'a fds) s) \ conv_abscissa f = -\" by (subst conv_abscissa_MInf_iff) auto lemma abs_conv_abscissa_PInf_iff: "abs_conv_abscissa f = \ \ (\s. \fds_abs_converges f s)" unfolding abs_conv_abscissa_def by (subst Inf_eq_PInfty) auto lemma abs_conv_abscissa_PInfI [intro]: "(\s. \fds_converges f s) \ abs_conv_abscissa f = \" by (subst abs_conv_abscissa_PInf_iff) auto lemma abs_conv_abscissa_MInf_iff: "abs_conv_abscissa (f :: 'a fds) = -\ \ (\s. fds_abs_converges f s)" proof safe assume *: "\s. fds_abs_converges f s" have "abs_conv_abscissa f \ B" for B :: real using spec[OF *, of "of_real B"] fds_abs_diverges[of "of_real B" f] by (cases "abs_conv_abscissa f \ B") simp_all thus "abs_conv_abscissa f = -\" by (rule ereal_bot) qed (auto intro: fds_abs_converges) lemma abs_conv_abscissa_MInfI [intro]: "(\s. fds_abs_converges (f::'a fds) s) \ abs_conv_abscissa f = -\" by (subst abs_conv_abscissa_MInf_iff) auto lemma conv_abscissa_geI: assumes "\c'. ereal c' < c \ \s. s \ 1 = c' \ \fds_converges f s" shows "conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\conv_abscissa f \ c" hence "c > conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c > ereal c'" "c' > conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "\fds_converges f s" by blast ultimately show False using fds_converges[of f s] by auto qed lemma conv_abscissa_leI: assumes "\c'. ereal c' > c \ \s. s \ 1 = c' \ fds_converges f s" shows "conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\conv_abscissa f \ c" hence "c < conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c < ereal c'" "c' < conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "fds_converges f s" by blast ultimately show False using fds_diverges[of s f] by auto qed lemma abs_conv_abscissa_geI: assumes "\c'. ereal c' < c \ \s. s \ 1 = c' \ \fds_abs_converges f s" shows "abs_conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\abs_conv_abscissa f \ c" hence "c > abs_conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c > ereal c'" "c' > abs_conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "\fds_abs_converges f s" by blast ultimately show False using fds_abs_converges[of f s] by auto qed lemma abs_conv_abscissa_leI: assumes "\c'. ereal c' > c \ \s. s \ 1 = c' \ fds_abs_converges f s" shows "abs_conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\abs_conv_abscissa f \ c" hence "c < abs_conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c < ereal c'" "c' < abs_conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "fds_abs_converges f s" by blast ultimately show False using fds_abs_diverges[of s f] by auto qed lemma conv_abscissa_leI_weak: assumes "\x. ereal x > d \ fds_converges f (of_real x)" shows "conv_abscissa (f :: 'a fds) \ d" proof (rule conv_abscissa_leI) fix x assume "d < ereal x" from assms[OF this] show "\s. s \ 1 = x \ fds_converges f s" by (intro exI[of _ "of_real x"]) auto qed lemma abs_conv_abscissa_leI_weak: assumes "\x. ereal x > d \ fds_abs_converges f (of_real x)" shows "abs_conv_abscissa (f :: 'a fds) \ d" proof (rule abs_conv_abscissa_leI) fix x assume "d < ereal x" from assms[OF this] show "\s. s \ 1 = x \ fds_abs_converges f s" by (intro exI[of _ "of_real x"]) auto qed lemma conv_abscissa_truncate [simp]: "conv_abscissa (fds_truncate m (f :: 'a fds)) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_truncate [simp]: "abs_conv_abscissa (fds_truncate m (f :: 'a fds)) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) theorem abs_conv_le_conv_abscissa_plus_1: "abs_conv_abscissa (f :: 'a fds) \ conv_abscissa f + 1" proof (rule abs_conv_abscissa_leI) fix c assume less: "conv_abscissa f + 1 < ereal c" define c' where "c' = (if conv_abscissa f = -\ then c - 2 else (c - 1 + real_of_ereal (conv_abscissa f)) / 2)" from less have c': "conv_abscissa f < ereal c' \ c' < c - 1" by (cases "conv_abscissa f") (simp_all add: c'_def field_simps) from c' have "fds_converges f (of_real c')" by (intro fds_converges) (simp_all add: inner_diff_left dot_square_norm) hence "fds_abs_converges f (of_real c)" by (rule fds_converges_imp_abs_converges) (insert c', simp_all) thus "\s. s \ 1 = c \ fds_abs_converges f s" by (intro exI[of _ "of_real c"]) auto qed lemma uniformly_convergent_eval_fds: assumes B: "compact B" "\z. z \ B \ z \ 1 > conv_abscissa (f :: 'a fds)" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (cases "B = {}") case False define \ where "\ = Inf ((\s. s \ 1) ` B)" have \_le: "s \ 1 \ \" if "s \ B" for s unfolding \_def using that by (intro cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded B) auto have "\ \ ((\s. s \ 1) ` B)" unfolding \_def using B \B \ {}\ by (intro closed_contains_Inf bounded_inner_imp_bdd_below compact_imp_bounded B compact_imp_closed compact_continuous_image continuous_intros) auto with B(2) have \_gt: "\ > conv_abscissa f" by auto define s where "s = (if conv_abscissa f = -\ then \ - 1 else (\ + real_of_ereal (conv_abscissa f)) / 2)" from \_gt have s: "conv_abscissa f < s \ s < \" by (cases "conv_abscissa f") (auto simp: s_def) show ?thesis using s \compact B\ by (intro uniformly_convergent_eval_fds_aux'[of f "of_real s"] fds_converges) (auto dest: \_le) qed auto corollary uniformly_convergent_eval_fds': assumes B: "compact B" "\z. z \ B \ z \ 1 > conv_abscissa (f :: 'a fds)" shows "uniformly_convergent_on B (\N z. \nN z. \n\N. fds_nth f n / nat_power n z) l at_top" by (auto simp: uniformly_convergent_on_def) also have "(\N z. \n\N. fds_nth f n / nat_power n z) = (\N z. \nN z. \nDerivative of a Dirichlet series\ lemma fds_converges_deriv_aux: assumes conv: "fds_converges f (s0 :: 'a)" and gt: "s \ 1 > s0 \ 1" shows "fds_converges (fds_deriv f) s" proof - have "Cauchy (\n. \k\n. (-ln (real k) *\<^sub>R fds_nth f k) / nat_power k s)" proof (rule CauchyI', goal_cases) case (1 \) define \ where "\ = s \ 1 - s0 \ 1" define \' where "\' = \ / 2" from gt have \_pos: "\ > 0" by (simp add: \_def) define A where "A = sum_upto (\k. fds_nth f k / nat_power k s0)" from conv have "convergent (\n. \k\n. fds_nth f k / nat_power k s0)" by (simp add: fds_converges_def summable_iff_convergent') hence "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" by (rule convergent_imp_Bseq) then obtain C_aux where C_aux: "\n. norm (\k\n. fds_nth f k / nat_power k s0) \ C_aux" by (auto simp: Bseq_def) define C where "C = max C_aux 1" have C_pos: "C > 0" by (simp add: C_def) have C: "norm (A x) \ C" for x proof - have "A x = (\k\nat \x\. fds_nth f k / nat_power k s0)" unfolding A_def sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C_aux" by (rule C_aux) also have "\ \ C" by (simp add: C_def) finally show ?thesis . qed define C' where "C' = 2 * C + C * (norm (s0 - s) * (1 + 1 / \) + 1) / \" have "(\m. C' * real m powr (-\')) \ 0" unfolding \'_def using gt \_pos by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially) simp_all from order_tendstoD(2)[OF this \\ > 0\] obtain M1 where M1: "\m. m \ M1 \ C' * real m powr - \' < \" by (auto simp: eventually_at_top_linorder) have "((\x. ln (real x) / real x powr \') \ 0) at_top" using \_pos by (intro lim_ln_over_power) (simp_all add: \'_def) from order_tendstoD(2)[OF this zero_less_one] eventually_gt_at_top[of "1::nat"] have "eventually (\n. ln (real n) \ n powr \') at_top" by eventually_elim simp_all then obtain M2 where M2: "\n. n \ M2 \ ln (real n) \ n powr \'" by (auto simp: eventually_at_top_linorder) let ?f' = "\k. -ln (real k) *\<^sub>R fds_nth f k" show ?case proof (intro exI[of _ "max (max M1 M2) 1"] allI impI, goal_cases) case (1 m n) hence mn: "m \ M1" "m \ M2" "m > 0" "m < n" by simp_all define g :: "real \ 'a" where "g = (\t. real_power t (s0 - s) * of_real (ln t))" define g' :: "real \ 'a" where "g' = (\t. real_power t (s0 - s - 1) * ((s0 - s) * of_real (ln t) + 1))" define norm_g' :: "real \ real" where "norm_g' = (\t. t powr (-\ - 1) * (norm (s0 - s) * ln t + 1))" define norm_g :: "real \ real" where "norm_g = (\t. -(t powr -\) * (norm (s0 - s) * (\ * ln t + 1) + \) / \^2)" have g_g': "(g has_vector_derivative g' t) (at t)" if "t \ {real m..real n}" for t using mn that by (auto simp: g_def g'_def real_power_diff field_simps real_power_add intro!: derivative_eq_intros) have [continuous_intros]: "continuous_on {real m..real n} g" using mn by (auto simp: g_def intro!: continuous_intros) let ?S = "\k\real -` {real m<..real n}. fds_nth f k / nat_power k s0 * g k" have "dist (\k\m. ?f' k / nat_power k s) (\k\n. ?f' k / nat_power k s) = norm (\k\{..n} - {..m}. fds_nth f k / nat_power k s * of_real (ln (real k)))" using mn by (subst sum_diff) (simp_all add: dist_norm norm_minus_commute sum_negf scaleR_conv_of_real mult_ac) also have "{..n} - {..m} = real -` {real m<..real n}" by auto also have "(\k\\. fds_nth f k / nat_power k s * of_real (ln (real k))) = (\k\\. fds_nth f k / nat_power k s0 * g k)" using mn unfolding g_def by (intro sum.cong refl) (auto simp: real_power_nat_power field_simps nat_power_diff) also have *: "((\t. A t * g' t) has_integral (A (real n) * g n - A (real m) * g m - ?S)) {real m..real n}" (is "(?h has_integral _) _") unfolding A_def using mn by (intro partial_summation_strong[of "{}"]) (auto intro!: g_g' simp: field_simps continuous_intros) hence "?S = A (real n) * g n - A (real m) * g m - integral {real m..real n} ?h" using mn by (simp add: has_integral_iff field_simps) also have "norm \ \ norm (A (real n) * g n) + norm (A (real m) * g m) + norm (integral {real m..real n} ?h)" by (intro order.trans[OF norm_triangle_ineq4] add_right_mono order.refl) also have "norm (A (real n) * g n) \ C * norm (g n)" unfolding norm_mult using mn C_pos by (intro mult_mono C) auto also have "norm (g n) \ n powr -\ * n powr \'" using mn M2[of n] by (simp add: g_def norm_real_power norm_mult \_def inner_diff_left) also have "\ = n powr -\'" using mn by (simp add: \'_def powr_minus field_simps powr_add [symmetric]) also have "norm (A (real m) * g m) \ C * norm (g m)" unfolding norm_mult using mn C_pos by (intro mult_mono C) auto also have "norm (g m) \ m powr -\ * m powr \'" using mn M2[of m] by (simp add: g_def norm_real_power norm_mult \_def inner_diff_left) also have "\ = m powr -\'" using mn by (simp add: \'_def powr_minus field_simps powr_add [symmetric]) also have "C * real n powr - \' \ C * real m powr - \'" using \_pos mn C_pos by (intro mult_left_mono powr_mono2') (simp_all add: \'_def) also have "\ + \ = 2 * \" by simp also have "norm (integral {m..n} ?h) \ integral {m..n} (\t. C * norm_g' t)" proof (intro integral_norm_bound_integral ballI, goal_cases) case 1 with * show ?case by (simp add: has_integral_iff) next case 2 from mn show ?case by (auto intro!: integrable_continuous_real continuous_intros simp: norm_g'_def) next case (3 t) have "norm (g' t) \ norm_g' t" unfolding g'_def norm_g'_def using 3 mn unfolding norm_mult by (intro mult_mono order.trans[OF norm_triangle_ineq]) (auto simp: norm_real_power inner_diff_left dot_square_norm norm_mult \_def intro!: mult_left_mono) thus ?case unfolding norm_mult using C_pos mn by (intro mult_mono C) simp_all qed also have "\ = C * integral {m..n} norm_g'" unfolding norm_g'_def by (simp add: norm_g'_def \_def inner_diff_left) also { have "(norm_g' has_integral (norm_g n - norm_g m)) {m..n}" unfolding norm_g'_def norm_g_def power2_eq_square using mn \_pos by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps powr_diff intro!: derivative_eq_intros) hence "integral {m..n} norm_g' = norm_g n - norm_g m" by (simp add: has_integral_iff) also have "norm_g n \ 0" unfolding norm_g_def using \_pos mn by (intro divide_nonpos_pos mult_nonpos_nonneg add_nonneg_nonneg mult_nonneg_nonneg) simp_all hence "norm_g n - norm_g m \ -norm_g m" by simp also have "\ = real m powr -\ * ln (real m) * (norm (s0 - s)) / \ + real m powr -\ * ((norm (s0 - s) / \ + 1) / \)" using \_pos by (simp add: field_simps norm_g_def power2_eq_square) also { have "ln (real m) \ real m powr \'" using M2[of m] mn by simp also have "real m powr -\ * \ = real m powr -\'" by (simp add: powr_add [symmetric] \'_def) finally have "real m powr -\ * ln (real m) * (norm (s0 - s)) / \ \ \ * (norm (s0 - s)) / \" using \_pos by (intro divide_right_mono mult_right_mono) (simp_all add: mult_left_mono) } also have "real m powr -\ * ((norm (s0 - s) / \ + 1) / \) \ real m powr -\' * ((norm (s0 - s) / \ + 1) / \)" using mn \_pos by (intro mult_right_mono powr_mono) (simp_all add: \'_def) also have "real m powr - \' * norm (s0 - s) / \ + \ = real m powr -\' * (norm (s0 - s) * (1 + 1 / \) + 1) / \" using \_pos by (simp add: field_simps power2_eq_square) finally have "integral {real m..real n} norm_g' \ real m powr - \' * (norm (s0 - s) * (1 + 1 / \) + 1) / \" by - simp_all } also have "2 * (C * m powr - \') + C * (m powr - \' * (norm (s0 - s) * (1 + 1 / \) + 1) / \) = C' * m powr -\'" by (simp add: algebra_simps C'_def) also have "\ < \" using M1[of m] mn by simp finally show ?case using C_pos by - simp_all qed qed from Cauchy_convergent[OF this] show ?thesis by (simp add: summable_iff_convergent' fds_converges_def fds_nth_deriv) qed theorem assumes "s \ 1 > conv_abscissa (f :: 'a fds)" shows fds_converges_deriv: "fds_converges (fds_deriv f) s" and has_field_derivative_eval_fds [derivative_intros]: "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s within A)" proof - define s1 :: real where "s1 = (if conv_abscissa f = -\ then s \ 1 - 2 else (s \ 1 * 1 / 3 + real_of_ereal (conv_abscissa f) * 2 / 3))" define s2 :: real where "s2 = (if conv_abscissa f = -\ then s \ 1 - 1 else (s \ 1 * 2 / 3 + real_of_ereal (conv_abscissa f) * 1 / 3))" from assms have s: "conv_abscissa f < s1 \ s1 < s2 \ s2 < s \ 1" by (cases "conv_abscissa f") (auto simp: s1_def s2_def field_simps) from s have *: "fds_converges f (of_real s1)" by (intro fds_converges) simp_all thus conv': "fds_converges (fds_deriv f) s" by (rule fds_converges_deriv_aux) (insert s, simp_all) from * have conv: "fds_converges (fds_deriv f) (of_real s2)" by (rule fds_converges_deriv_aux) (insert s, simp_all) define \ :: real where "\ = (s \ 1 - s2) / 2" from s have \_pos: "\ > 0" by (simp add: \_def) have "uniformly_convergent_on (cball s \) (\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s)" proof (intro uniformly_convergent_eval_fds_aux'[OF conv]) fix s'' :: 'a assume s'': "s'' \ cball s \" have "dist (s \ 1) (s'' \ 1) \ dist s s''" by (intro Euclidean_dist_upper) (simp_all add: one_in_Basis) also from s'' have "\ \ \" by simp finally show "s'' \ 1 > (of_real s2 :: 'a) \ 1" using s by (auto simp: \_def dist_real_def abs_if split: if_splits) qed (insert \_pos, auto) then obtain l where "uniform_limit (cball s \) (\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s) l at_top" by (auto simp: uniformly_convergent_on_def) also have "(\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s) = (\n s. \k) (\n s. \k) (\n s. \k cball s \" "s \ interior (cball s \)" using s by (simp_all add: \_def) show "summable (\n. fds_nth f n / nat_power n s)" using assms fds_converges[of f s] by (simp add: fds_converges_def) next fix s' :: 'a and n :: nat show "((\s. fds_nth f n / nat_power n s) has_field_derivative fds_nth (fds_deriv f) n / nat_power n s') (at s' within cball s \)" by (cases "n = 0") (simp, auto intro!: derivative_eq_intros simp: fds_nth_deriv field_simps) qed (auto simp: fds_nth_deriv intro!: derivative_eq_intros) thus "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s within A)" by (rule has_field_derivative_at_within) qed lemmas has_field_derivative_eval_fds' [derivative_intros] = DERIV_chain2[OF has_field_derivative_eval_fds] lemma continuous_eval_fds [continuous_intros]: assumes "s \ 1 > conv_abscissa f" shows "continuous (at s within A) (eval_fds (f :: 'a :: dirichlet_series fds))" proof - have "isCont (eval_fds f) s" by (rule has_field_derivative_eval_fds DERIV_isCont assms)+ thus ?thesis by (rule continuous_within_subset) auto qed lemma continuous_eval_fds' [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "continuous (at s within A) g" "g s \ 1 > conv_abscissa f" shows "continuous (at s within A) (\x. eval_fds f (g x))" by (rule continuous_within_compose3[OF _ assms(1)] continuous_intros assms)+ lemma continuous_on_eval_fds [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "A \ {s. s \ 1 > conv_abscissa f}" shows "continuous_on A (eval_fds f)" by (rule DERIV_continuous_on derivative_intros)+ (insert assms, auto) lemma continuous_on_eval_fds' [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "continuous_on A g" "g ` A \ {s. s \ 1 > conv_abscissa f}" shows "continuous_on A (\x. eval_fds f (g x))" by (rule continuous_on_compose2[OF continuous_on_eval_fds assms(1)]) (insert assms, auto simp: image_iff) lemma conv_abscissa_deriv_le: fixes f :: "'a fds" shows "conv_abscissa (fds_deriv f) \ conv_abscissa f" proof (rule conv_abscissa_leI) fix c' :: real assume "ereal c' > conv_abscissa f" thus "\s. s \ 1 = c' \ fds_converges (fds_deriv f) s" by (intro exI[of _ "of_real c'"]) (auto simp: fds_converges_deriv) qed lemma abs_conv_abscissa_integral: fixes f :: "'a fds" shows "abs_conv_abscissa (fds_integral a f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa (fds_integral a f) \ abs_conv_abscissa f" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) have "fds_abs_converges (fds_integral a f) (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from 1 have "fds_abs_converges f (of_real c)" by (intro fds_abs_converges) auto thus "summable (\n. norm (fds_nth f n / nat_power n (of_real c)))" by (simp add: fds_abs_converges_def) next show "\\<^sub>F n in sequentially. norm (norm (fds_nth (fds_integral a f) n / nat_power n (of_real c))) \ norm (fds_nth f n / nat_power n (of_real c))" using eventually_gt_at_top[of 3] proof eventually_elim case (elim n) from elim and exp_le have "ln (exp 1) \ ln (real n)" by (subst ln_le_cancel_iff) auto hence "1 * norm (fds_nth f n) \ ln (real n) * norm (fds_nth f n)" by (intro mult_right_mono) auto with elim show ?case by (simp add: norm_divide norm_nat_power fds_integral_def field_simps) qed qed thus ?case by (intro exI[of _ "of_real c"]) auto qed next show "abs_conv_abscissa f \ abs_conv_abscissa (fds_integral a f)" (is "_ \ ?s0") proof (cases "abs_conv_abscissa (fds_integral a f) = \") case False show ?thesis proof (rule abs_conv_abscissa_leI) fix c :: real define \ where "\ = (if ?s0 = -\ then 1 else (c - real_of_ereal ?s0) / 2)" assume "ereal c > ?s0" with False have \: "\ > 0" "c - \ > ?s0" by (cases ?s0; force simp: \_def field_simps)+ have "fds_abs_converges f (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from \ have "fds_abs_converges (fds_integral a f) (of_real (c - \))" by (intro fds_abs_converges) (auto simp: algebra_simps) thus "summable (\n. norm (fds_nth (fds_integral a f) n / nat_power n (of_real (c - \))))" by (simp add: fds_abs_converges_def) next have "\\<^sub>F n in at_top. ln (real n) / real n powr \ < 1" by (rule order_tendstoD lim_ln_over_power \\ > 0\ zero_less_one)+ thus "\\<^sub>F n in sequentially. norm (norm (fds_nth f n / nat_power n (of_real c))) \ norm (fds_nth (fds_integral a f) n / nat_power n (of_real (c - \)))" using eventually_gt_at_top[of 1] proof eventually_elim case (elim n) hence "ln (real n) * norm (fds_nth f n) \ real n powr \ * norm (fds_nth f n)" by (intro mult_right_mono) auto with elim show ?case by (simp add: norm_divide norm_nat_power field_simps powr_diff inner_diff_left fds_integral_def) qed qed thus "\s. s \ 1 = c \ fds_abs_converges f s" by (intro exI[of _ "of_real c"]) auto qed qed auto qed lemma abs_conv_abscissa_ln: "abs_conv_abscissa (fds_ln l (f :: 'a :: dirichlet_series fds)) = abs_conv_abscissa (fds_deriv f / f)" by (simp add: fds_ln_def abs_conv_abscissa_integral) lemma abs_conv_abscissa_deriv: fixes f :: "'a fds" shows "abs_conv_abscissa (fds_deriv f) = abs_conv_abscissa f" proof - have "abs_conv_abscissa (fds_deriv f) = abs_conv_abscissa (fds_integral (fds_nth f 1) (fds_deriv f))" by (rule abs_conv_abscissa_integral [symmetric]) also have "fds_integral (fds_nth f 1) (fds_deriv f) = f" by (rule fds_integral_fds_deriv) finally show ?thesis . qed lemma abs_conv_abscissa_higher_deriv: "abs_conv_abscissa ((fds_deriv ^^ n) f) = abs_conv_abscissa (f :: 'a :: dirichlet_series fds)" by (induction n) (simp_all add: abs_conv_abscissa_deriv) lemma conv_abscissa_higher_deriv_le: "conv_abscissa ((fds_deriv ^^ n) f) \ conv_abscissa (f :: 'a :: dirichlet_series fds)" by (induction n) (auto intro: order.trans[OF conv_abscissa_deriv_le]) lemma abs_conv_abscissa_restrict: "abs_conv_abscissa (fds_subseries P f) \ abs_conv_abscissa f" by (rule abs_conv_abscissa_mono) auto lemma eval_fds_deriv: fixes f :: "'a fds" assumes "s \ 1 > conv_abscissa f" shows "eval_fds (fds_deriv f) s = deriv (eval_fds f) s" by (intro DERIV_imp_deriv [symmetric] derivative_intros assms) lemma eval_fds_higher_deriv: assumes "(s :: 'a :: dirichlet_series) \ 1 > conv_abscissa f" shows "eval_fds ((fds_deriv ^^ n) f) s = (deriv ^^ n) (eval_fds f) s" using assms proof (induction n arbitrary: f s) case (Suc n f s) have ev: "eventually (\s. s \ {s. s \ 1 > conv_abscissa f}) (nhds s)" using Suc.prems open_halfspace_gt[of _ "1::'a"] by (intro eventually_nhds_in_open, cases "conv_abscissa f") (auto simp: open_halfspace_gt inner_commute) have "eval_fds ((fds_deriv ^^ Suc n) f) s = eval_fds ((fds_deriv ^^ n) (fds_deriv f)) s" by (subst funpow_Suc_right) simp also have "\ = (deriv ^^ n) (eval_fds (fds_deriv f)) s" by (intro Suc.IH le_less_trans[OF conv_abscissa_deriv_le] Suc.prems) also have "\ = (deriv ^^ n) (deriv (eval_fds f)) s" by (intro higher_deriv_cong_ev refl eventually_mono[OF ev] eval_fds_deriv) auto also have "\ = (deriv ^^ Suc n) (eval_fds f) s" by (subst funpow_Suc_right) simp finally show ?case . qed auto end subsection \Multiplication of two series\ lemma fixes f g :: "nat \ 'a :: {banach, real_normed_field, second_countable_topology, nat_power}" fixes s :: 'a assumes [simp]: "f 0 = 0" "g 0 = 0" assumes summable: "summable (\n. norm (f n / nat_power n s))" "summable (\n. norm (g n / nat_power n s))" shows summable_dirichlet_prod: "summable (\n. norm (dirichlet_prod f g n / nat_power n s))" and suminf_dirichlet_prod: "(\n. dirichlet_prod f g n / nat_power n s) = (\n. f n / nat_power n s) * (\n. g n / nat_power n s)" proof - have summable': "(\n. f n / nat_power n s) abs_summable_on A" "(\n. g n / nat_power n s) abs_summable_on A" for A by ((rule abs_summable_on_subset[OF _ subset_UNIV], insert summable, simp add: abs_summable_on_nat_iff'); fail)+ have f_g: "f a / nat_power a s * (g b / nat_power b s) = f a * g b / nat_power (a * b) s" for a b by (cases "a * b = 0") (auto simp: nat_power_mult_distrib) have eq: "(\\<^sub>a(m, n)\{(m, n). m * n = x}. f m * g n / nat_power x s) = dirichlet_prod f g x / nat_power x s" for x :: nat proof (cases "x > 0") case False hence "(\\<^sub>a(m,n) | m * n = x. f m * g n / nat_power x s) = (\\<^sub>a(m,n) | m * n = x. 0)" by (intro infsetsum_cong) auto with False show ?thesis by simp next case True from finite_divisors_nat'[OF this] show ?thesis by (simp add: dirichlet_prod_altdef2 case_prod_unfold sum_divide_distrib) qed have "(\(m,n). (f m / nat_power m s) * (g n / nat_power n s)) abs_summable_on UNIV \ UNIV" using summable' by (intro abs_summable_on_product) auto also have "?this \ (\(m,n). f m * g n / nat_power (m*n) s) abs_summable_on UNIV" using f_g by (intro abs_summable_on_cong) auto also have "\ \ (\(x,(m,n)). f m * g n / nat_power (m*n) s) abs_summable_on (SIGMA x:UNIV. {(m,n). m * n = x})" unfolding case_prod_unfold by (rule abs_summable_on_reindex_bij_betw [symmetric]) (auto simp: bij_betw_def inj_on_def image_iff) also have "\ \ (\(x,(m,n)). f m * g n / nat_power x s) abs_summable_on (SIGMA x:UNIV. {(m,n). m * n = x})" by (intro abs_summable_on_cong) auto finally have summable'': \ . from abs_summable_on_Sigma_project1'[OF this] show summable''': "summable (\n. norm (dirichlet_prod f g n / nat_power n s))" by (simp add: eq abs_summable_on_nat_iff') have "(\n. f n / nat_power n s) * (\n. g n / nat_power n s) = (\\<^sub>an. f n / nat_power n s) * (\\<^sub>an. g n / nat_power n s)" using summable' by (simp add: infsetsum_nat') also have "\ = (\\<^sub>a(m,n). (f m / nat_power m s) * (g n / nat_power n s))" using summable' by (subst infsetsum_product [symmetric]) simp_all also have "\ = (\\<^sub>a(m,n). f m * g n / nat_power (m * n) s)" using f_g by (intro infsetsum_cong refl) auto also have "\ = (\\<^sub>a(x,(m,n))\(SIGMA x:UNIV. {(m,n). m * n = x}). f m * g n / nat_power (m * n) s)" unfolding case_prod_unfold by (rule infsetsum_reindex_bij_betw [symmetric]) (auto simp: bij_betw_def inj_on_def image_iff) also have "\ = (\\<^sub>a(x,(m,n))\(SIGMA x:UNIV. {(m,n). m * n = x}). f m * g n / nat_power x s)" by (intro infsetsum_cong refl) (auto simp: case_prod_unfold) also have "\ = (\\<^sub>ax. dirichlet_prod f g x / nat_power x s)" (is "_ = infsetsum ?T _") using summable'' by (subst infsetsum_Sigma) (auto simp: eq) also have "\ = (\x. dirichlet_prod f g x / nat_power x s)" using summable''' by (intro infsetsum_nat') (simp_all add: abs_summable_on_nat_iff') finally show "\ = (\n. f n / nat_power n s) * (\n. g n / nat_power n s)" .. qed lemma fixes f g :: "nat \ real" fixes s :: real assumes "f 0 = 0" "g 0 = 0" assumes summable: "summable (\n. norm (f n / real n powr s))" "summable (\n. norm (g n / real n powr s))" shows summable_dirichlet_prod_real: "summable (\n. norm (dirichlet_prod f g n / real n powr s))" and suminf_dirichlet_prod_real: "(\n. dirichlet_prod f g n / real n powr s) = (\n. f n / nat_power n s) * (\n. g n / real n powr s)" using summable_dirichlet_prod[of f g s] suminf_dirichlet_prod[of f g s] assms by simp_all lemma fds_abs_converges_mult: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges g s" shows "fds_abs_converges (f * g) s" using summable_dirichlet_prod[OF _ _ assms[unfolded fds_abs_converges_def]] by (simp add: fds_abs_converges_def fds_nth_mult) lemma fds_abs_converges_power: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" shows "fds_abs_converges f s \ fds_abs_converges (f ^ n) s" by (induction n) (auto intro!: fds_abs_converges_mult) lemma fds_abs_converges_prod: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" shows "(\x. x \ A \ fds_abs_converges (f x) s) \ fds_abs_converges (prod f A) s" by (induction A rule: infinite_finite_induct) (auto intro!: fds_abs_converges_mult) lemma abs_conv_abscissa_mult_le: "abs_conv_abscissa (f * g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c') thus ?case by (auto intro!: exI[of _ "of_real c'"] fds_abs_converges_mult intro: fds_abs_converges) qed lemma abs_conv_abscissa_mult_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f * g) \ d" using abs_conv_abscissa_mult_le[of f g] by (auto simp add: le_max_iff_disj) lemma abs_conv_abscissa_shift [simp]: "abs_conv_abscissa (fds_shift c f) = abs_conv_abscissa (f :: 'a :: dirichlet_series fds) + c \ 1" proof - have "abs_conv_abscissa (fds_shift c f) \ abs_conv_abscissa f + c \ 1" for c :: 'a and f proof (rule abs_conv_abscissa_leI) fix d assume "abs_conv_abscissa f + c \ 1 < ereal d" hence "abs_conv_abscissa f < ereal (d - c \ 1)" by (cases "abs_conv_abscissa f") auto hence "fds_abs_converges (fds_shift c f) (of_real d)" by (auto intro!: fds_abs_converges_shift fds_abs_converges simp: algebra_simps) thus "\s. s \ 1 = d \ fds_abs_converges (fds_shift c f) s" by (auto intro!: exI[of _ "of_real d"]) qed note * = this[of c f] this[of "-c" "fds_shift c f"] show ?thesis by (cases "abs_conv_abscissa (fds_shift c f)"; cases "abs_conv_abscissa f") (insert *, auto intro!: antisym) qed lemma eval_fds_mult: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges g s" shows "eval_fds (f * g) s = eval_fds f s * eval_fds g s" using suminf_dirichlet_prod[OF _ _ assms[unfolded fds_abs_converges_def]] by (simp_all add: eval_fds_def fds_nth_mult) lemma eval_fds_power: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" shows "eval_fds (f ^ n) s = eval_fds f s ^ n" using assms by (induction n) (simp_all add: eval_fds_mult fds_abs_converges_power) lemma eval_fds_prod: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "(\x. x \ A \ fds_abs_converges (f x) s)" shows "eval_fds (prod f A) s = (\x\A. eval_fds (f x) s)" using assms by (induction A rule: infinite_finite_induct) (auto simp: eval_fds_mult fds_abs_converges_prod) lemma eval_fds_inverse: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges (inverse f) s" "fds_nth f 1 \ 0" shows "eval_fds (inverse f) s = inverse (eval_fds f s)" proof - have "eval_fds (inverse f * f) s = eval_fds (inverse f) s * eval_fds f s" by (intro eval_fds_mult assms) also have "inverse f * f = 1" by (intro fds_left_inverse assms) also have "eval_fds 1 s = 1" by simp finally show ?thesis by (auto simp: divide_simps) qed lemma eval_fds_integral_has_field_derivative: fixes s :: "'a :: dirichlet_series" assumes "ereal (s \ 1) > abs_conv_abscissa f" assumes "fds_nth f 1 = 0" shows "(eval_fds (fds_integral c f) has_field_derivative eval_fds f s) (at s)" proof - have "conv_abscissa (fds_integral c f) \ abs_conv_abscissa (fds_integral c f)" by (rule conv_le_abs_conv_abscissa) also from assms have "\ < ereal (s \ 1)" by (simp add: abs_conv_abscissa_integral) finally have "(eval_fds (fds_integral c f) has_field_derivative eval_fds (fds_deriv (fds_integral c f)) s) (at s)" by (intro derivative_eq_intros) auto also from assms have "fds_deriv (fds_integral c f) = f" by simp finally show ?thesis . qed lemma holomorphic_fds_eval [holomorphic_intros]: "A \ {z. Re z > conv_abscissa f} \ eval_fds f holomorphic_on A" unfolding holomorphic_on_def field_differentiable_def by (rule ballI exI derivative_intros)+ auto lemma analytic_fds_eval [holomorphic_intros]: assumes "A \ {z. Re z > conv_abscissa f}" shows "eval_fds f analytic_on A" proof - have "eval_fds f analytic_on {z. Re z > conv_abscissa f}" proof (subst analytic_on_open) show "open {z. Re z > conv_abscissa f}" by (cases "conv_abscissa f") (simp_all add: open_halfspace_Re_gt) qed (intro holomorphic_intros, simp_all) from analytic_on_subset[OF this assms] show ?thesis . qed lemma conv_abscissa_0 [simp]: "conv_abscissa (0 :: 'a :: dirichlet_series fds) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_0 [simp]: "abs_conv_abscissa (0 :: 'a :: dirichlet_series fds) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_1 [simp]: "conv_abscissa (1 :: 'a :: dirichlet_series fds) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_1 [simp]: "abs_conv_abscissa (1 :: 'a :: dirichlet_series fds) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_const [simp]: "conv_abscissa (fds_const (c :: 'a :: dirichlet_series)) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_const [simp]: "abs_conv_abscissa (fds_const (c :: 'a :: dirichlet_series)) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_numeral [simp]: "conv_abscissa (numeral n :: 'a :: dirichlet_series fds) = -\" by (auto simp: numeral_fds) lemma abs_conv_abscissa_numeral [simp]: "abs_conv_abscissa (numeral n :: 'a :: dirichlet_series fds) = -\" by (auto simp: numeral_fds) lemma abs_conv_abscissa_power_le: "abs_conv_abscissa (f ^ n :: 'a :: dirichlet_series fds) \ abs_conv_abscissa f" by (induction n) (auto intro!: order.trans[OF abs_conv_abscissa_mult_le]) lemma abs_conv_abscissa_power_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa (f ^ n) \ d" by (rule order.trans[OF abs_conv_abscissa_power_le]) lemma abs_conv_abscissa_prod_le: assumes "\x. x \ A \ abs_conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "abs_conv_abscissa (prod f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: abs_conv_abscissa_mult_leI) lemma conv_abscissa_add_le: "conv_abscissa (f + g :: 'a :: dirichlet_series fds) \ max (conv_abscissa f) (conv_abscissa g)" by (rule conv_abscissa_leI_weak) (auto intro!: fds_converges_add intro: fds_converges) lemma conv_abscissa_add_leI: "conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ conv_abscissa g \ d \ conv_abscissa (f + g) \ d" using conv_abscissa_add_le[of f g] by (auto simp: le_max_iff_disj) lemma conv_abscissa_sum_leI: assumes "\x. x \ A \ conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "conv_abscissa (sum f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: conv_abscissa_add_leI) lemma abs_conv_abscissa_add_le: "abs_conv_abscissa (f + g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" by (rule abs_conv_abscissa_leI_weak) (auto intro!: fds_abs_converges_add intro: fds_abs_converges) lemma abs_conv_abscissa_add_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f + g) \ d" using abs_conv_abscissa_add_le[of f g] by (auto simp: le_max_iff_disj) lemma abs_conv_abscissa_sum_leI: assumes "\x. x \ A \ abs_conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "abs_conv_abscissa (sum f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: abs_conv_abscissa_add_leI) lemma fds_converges_cmult_left [intro]: assumes "fds_converges f s" shows "fds_converges (fds_const c * f) s" proof - from assms have "summable (\n. c * (fds_nth f n / nat_power n s))" by (intro summable_mult) (auto simp: fds_converges_def) thus ?thesis by (simp add: fds_converges_def mult_ac) qed lemma fds_converges_cmult_right [intro]: assumes "fds_converges f s" shows "fds_converges (f * fds_const c) s" using fds_converges_cmult_left[OF assms] by (simp add: mult_ac) lemma conv_abscissa_cmult_left [simp]: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "conv_abscissa (fds_const c * f) = conv_abscissa f" proof - have "fds_converges (fds_const c * f) s \ fds_converges f s" for s proof assume "fds_converges (fds_const c * f) s" hence "fds_converges (fds_const (inverse c) * (fds_const c * f)) s" by (rule fds_converges_cmult_left) also have "fds_const (inverse c) * (fds_const c * f) = fds_const (inverse c * c) * f" by simp also have "inverse c * c = 1" using assms by simp finally show "fds_converges f s" by simp qed auto thus ?thesis by (simp add: conv_abscissa_def) qed lemma conv_abscissa_cmult_right [simp]: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "conv_abscissa (f * fds_const c) = conv_abscissa f" using assms by (subst mult.commute) auto lemma abs_conv_abscissa_cmult: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "abs_conv_abscissa (fds_const c * f) = abs_conv_abscissa f" proof (intro antisym) have "abs_conv_abscissa (fds_const (inverse c) * (fds_const c * f)) \ abs_conv_abscissa (fds_const c * f)" using abs_conv_abscissa_mult_le[of "fds_const (inverse c)" "fds_const c * f"] by (auto simp: max_def) also have "fds_const (inverse c) * (fds_const c * f) = fds_const (inverse c * c) * f" by (simp add: mult_ac) also have "inverse c * c = 1" using assms by simp finally show "abs_conv_abscissa f \ abs_conv_abscissa (fds_const c * f)" by simp qed (insert abs_conv_abscissa_mult_le[of "fds_const c" f], auto simp: max_def) lemma conv_abscissa_minus [simp]: fixes f :: "'a :: dirichlet_series fds" shows "conv_abscissa (-f) = conv_abscissa f" using conv_abscissa_cmult_left[of "-1" f] by simp lemma abs_conv_abscissa_minus [simp]: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (-f) = abs_conv_abscissa f" using abs_conv_abscissa_cmult[of "-1" f] by simp lemma conv_abscissa_diff_le: "conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (conv_abscissa f) (conv_abscissa g)" using conv_abscissa_add_le[of f "-g"] by simp lemma conv_abscissa_diff_leI: "conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ conv_abscissa g \ d \ conv_abscissa (f - g) \ d" using conv_abscissa_add_le[of f "-g"] by (auto simp: le_max_iff_disj) lemma abs_conv_abscissa_diff_le: "abs_conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" using abs_conv_abscissa_add_le[of f "-g"] by simp lemma abs_conv_abscissa_diff_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f - g) \ d" using abs_conv_abscissa_add_le[of f "-g"] by (auto simp: le_max_iff_disj) lemmas eval_fds_integral_has_field_derivative' [derivative_intros] = DERIV_chain'[OF _ eval_fds_integral_has_field_derivative] lemma abs_conv_abscissa_completely_multiplicative_log_deriv: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "abs_conv_abscissa (fds_deriv f / f) \ abs_conv_abscissa f" proof - have "fds_deriv f = - fds (\n. fds_nth f n * mangoldt n) * f" using assms by (subst completely_multiplicative_fds_deriv') simp_all also have "\ / f = - fds (\n. fds_nth f n * mangoldt n) * (f * inverse f)" by (simp add: divide_fds_def) also have "f * inverse f = 1" using assms by (intro fds_right_inverse) finally have "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n)" by simp also have "abs_conv_abscissa \ = abs_conv_abscissa (fds (\n. fds_nth f n * mangoldt n))" (is "_ = abs_conv_abscissa ?f") by (rule abs_conv_abscissa_minus) also have "\ \ abs_conv_abscissa f" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) have "fds_abs_converges ?f (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from 1 have "fds_abs_converges (fds_deriv f) (of_real c)" by (intro fds_abs_converges) (auto simp: abs_conv_abscissa_deriv) thus "summable (\n. \ln (real n)\ * norm (fds_nth f n) / norm (nat_power n (of_real c :: 'a)))" by (simp add: fds_abs_converges_def fds_deriv_def fds_nth_fds' scaleR_conv_of_real powr_minus norm_mult norm_divide norm_nat_power) next show "\\<^sub>F n in sequentially. norm (norm (fds_nth (fds (\n. fds_nth f n * mangoldt n)) n / nat_power n (of_real c))) \ \ln (real n)\ * norm (fds_nth f n) / norm (nat_power n (of_real c) :: 'a)" using eventually_gt_at_top[of 0] proof eventually_elim case (elim n) have "norm (norm (fds_nth (fds (\n. fds_nth f n * mangoldt n)) n / nat_power n (of_real c))) = norm (fds_nth f n) * mangoldt n / real n powr c" using elim by (simp add: fds_nth_fds' norm_mult norm_divide norm_nat_power abs_mult mangoldt_nonneg) also have "\ \ norm (fds_nth f n) * ln n / real n powr c" using elim by (intro mult_left_mono divide_right_mono mangoldt_le) (simp_all add: mangoldt_def) finally show ?case using elim by (simp add: norm_nat_power algebra_simps) qed qed thus ?case by (intro exI[of _ "of_real c"]) auto qed finally show ?thesis . qed subsection \Uniqueness\ context assumes "SORT_CONSTRAINT ('a :: dirichlet_series)" begin lemma norm_dirichlet_series_cutoff_le: assumes "fds_abs_converges f (s0 :: 'a)" "N > 0" "s \ 1 \ c" "c \ s0 \ 1" shows "summable (\n. fds_nth f (n + N) / nat_power (n + N) s)" "summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" and "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c) / nat_power N (s \ 1 - c)" proof - from assms have "fds_abs_converges f (of_real c)" using fds_abs_converges_Re_le[of f s0 "of_real c"] by auto hence "summable (\n. norm (fds_nth f (n + N) / nat_power (n + N) (of_real c)))" unfolding fds_abs_converges_def by (rule summable_ignore_initial_segment) also have "?this \ summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_divide norm_nat_power) finally show *: "summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" . from assms have "fds_abs_converges f s" using fds_abs_converges_Re_le[of f s0 s] by auto hence **: "summable (\n. norm (fds_nth f (n + N) / nat_power (n + N) s))" unfolding fds_abs_converges_def by (rule summable_ignore_initial_segment) thus "summable (\n. fds_nth f (n + N) / nat_power (n + N) s)" by (rule summable_norm_cancel) have "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ (\n. norm (fds_nth f (n + N) / nat_power (n + N) s))" by (intro summable_norm **) also have "\ \ (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c / nat_power N (s \ 1 - c))" proof (intro suminf_le * ** summable_divide allI) fix n :: nat have "real N powr (s \ 1 - c) \ real (n + N) powr (s \ 1 - c)" using assms by (intro powr_mono2) simp_all also have "real (n + N) powr c * \ = real (n + N) powr (s \ 1)" by (simp add: powr_diff) finally have "norm (fds_nth f (n + N)) / real (n + N) powr (s \ 1) \ norm (fds_nth f (n + N)) / (real (n + N) powr c * real N powr (s \ 1 - c))" using \N > 0\ by (intro divide_left_mono) (simp_all add: mult_left_mono) thus "norm (fds_nth f (n + N) / nat_power (n + N) s) \ norm (fds_nth f (n + N)) / nat_power (n + N) c / nat_power N (s \ 1 - c)" using \N > 0\ by (simp add: norm_divide norm_nat_power ) qed also have "\ = (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c) / nat_power N (s \ 1 - c)" using * by (rule suminf_divide) finally show "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ \" . qed lemma eval_fds_zeroD_aux: fixes h :: "'a fds" assumes conv: "fds_abs_converges h (s0 :: 'a)" assumes freq: "frequently (\s. eval_fds h s = 0) ((\s. s \ 1) going_to at_top)" shows "h = 0" proof (rule ccontr) assume "h \ 0" hence ex: "\n>0. fds_nth h n \ 0" by (auto simp: fds_eq_iff) define N :: nat where "N = (LEAST n. n > 0 \ fds_nth h n \ 0)" have N: "N > 0" "fds_nth h N \ 0" using LeastI_ex[OF ex, folded N_def] by auto have less_N: "fds_nth h n = 0" if "n < N" for n using Least_le[of "\n. n > 0 \ fds_nth h n \ 0" n, folded N_def] that by (cases "n = 0") (auto simp: not_less) define c where "c = s0 \ 1" define remainder where "remainder = (\s. (\n. fds_nth h (n + Suc N) / nat_power (n + Suc N) s))" define A where "A = (\n. norm (fds_nth h (n + Suc N)) / nat_power (n + Suc N) c) * nat_power (Suc N) c" have eq: "fds_nth h N = nat_power N s * eval_fds h s - nat_power N s * remainder s" if "s \ 1 \ c" for s :: 'a proof - from conv and that have conv': "fds_abs_converges h s" unfolding c_def by (rule fds_abs_converges_Re_le) hence conv'': "fds_converges h s" by blast from conv'' have "(\n. fds_nth h n / nat_power n s) sums eval_fds h s" by (simp add: fds_converges_iff) hence "(\n. fds_nth h (n + Suc N) / nat_power (n + Suc N) s) sums (eval_fds h s - (\nnn = fds_nth h N / nat_power N s" by (subst sum.delta) auto finally show ?thesis unfolding remainder_def using \N > 0\ by (auto simp: sums_iff field_simps) qed have remainder_bound: "norm (remainder s) \ A / real (Suc N) powr (s \ 1)" if "s \ 1 \ c" for s :: 'a proof - note * = norm_dirichlet_series_cutoff_le[of h s0 "Suc N" c s, folded remainder_def] have "norm (remainder s) \ (\n. norm (fds_nth h (n + Suc N)) / nat_power (n + Suc N) c) / nat_power (Suc N) (s \ 1 - c)" using that assms unfolding remainder_def by (intro *) (simp_all add: c_def) also have "\ = A / real (Suc N) powr (s \ 1)" by (simp add: A_def powr_diff) finally show ?thesis . qed from freq have "\c. \s. s \ 1 \ c \ eval_fds h s = 0" unfolding frequently_def by (auto simp: eventually_going_to_at_top_linorder) hence "\k. \s. s \ 1 \ real k \ eval_fds h s = 0" by blast then obtain S where S: "\k. S k \ 1 \ real k \ eval_fds h (S k) = 0" by metis have S_limit: "filterlim (\k. S k \ 1) at_top sequentially" by (rule filterlim_at_top_mono[OF filterlim_real_sequentially]) (use S in auto) have "eventually (\k. real k \ c) sequentially" by real_asymp hence "eventually (\k. norm (fds_nth h N) \ (real N / real (Suc N)) powr (S k \ 1) * A) sequentially" proof eventually_elim case (elim k) hence "norm (fds_nth h N) = real N powr (S k \ 1) * norm (remainder (S k))" (is "_ = _ * ?X") using \N > 0\ S[of k] eq[of "S k"] by (auto simp: norm_mult norm_nat_power c_def) also have "norm (remainder (S k)) \ A / real (Suc N) powr (S k \ 1)" using elim S[of k] by (intro remainder_bound) (simp_all add: c_def) finally show ?case using N by (simp add: mult_left_mono powr_divide field_simps del: of_nat_Suc) qed moreover have "((\k. (real N / real (Suc N)) powr (S k \ 1) * A) \ 0) sequentially" by (rule filterlim_compose[OF _ S_limit]) (use \N > 0\ in real_asymp) ultimately have "((\_. fds_nth h N) \ 0) sequentially" by (rule Lim_null_comparison) hence "fds_nth h N = 0" by (simp add: tendsto_const_iff) with \fds_nth h N \ 0\ show False by contradiction qed lemma eval_fds_zeroD: fixes h :: "'a fds" assumes conv: "conv_abscissa h < \" assumes freq: "frequently (\s. eval_fds h s = 0) ((\s. s \ 1) going_to at_top)" shows "h = 0" proof - have [simp]: "2 \ (1 :: 'a) = 2" using of_real_inner_1[of 2] unfolding of_real_numeral by simp from conv obtain s where "fds_converges h s" by auto hence "fds_abs_converges h (s + 2)" by (rule fds_converges_imp_abs_converges) (auto simp: algebra_simps) from this assms(2-) show ?thesis by (rule eval_fds_zeroD_aux) qed lemma eval_fds_eqD: fixes f g :: "'a fds" assumes conv: "conv_abscissa f < \" "conv_abscissa g < \" assumes eq: "frequently (\s. eval_fds f s = eval_fds g s) ((\s. s \ 1) going_to at_top)" shows "f = g" proof - have conv': "conv_abscissa (f - g) < \" using assms by (intro le_less_trans[OF conv_abscissa_diff_le]) (auto simp: max_def) have "max (conv_abscissa f) (conv_abscissa g) < \" using conv by (auto simp: max_def) from ereal_dense2[OF this] obtain c where c: "max (conv_abscissa f) (conv_abscissa g) < ereal c" by auto (* TODO: something like "frequently_elim" would be great here *) have "frequently (\s. eval_fds f s = eval_fds g s \ s \ 1 \ c) ((\s. s \ 1) going_to at_top)" using eq by (rule frequently_eventually_frequently) auto hence *: "frequently (\s. eval_fds (f - g) s = 0) ((\s. s \ 1) going_to at_top)" proof (rule frequently_mono [rotated], safe, goal_cases) case (1 s) thus ?case using c by (subst eval_fds_diff) (auto intro!: fds_converges intro: less_le_trans) qed have "f - g = 0" by (rule eval_fds_zeroD fds_abs_converges_diff assms * conv')+ thus ?thesis by simp qed end subsection \Limit at infinity\ lemma eval_fds_at_top_tail_bound: fixes f :: "'a :: dirichlet_series fds" assumes c: "ereal c > abs_conv_abscissa f" defines "B \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" assumes s: "s \ 1 \ c" shows "norm (eval_fds f s - fds_nth f 1) \ B / 2 powr (s \ 1)" proof - from c have "fds_abs_converges f (of_real c)" by (intro fds_abs_converges) simp_all also have "?this \ summable (\n. norm (fds_nth f n) / real n powr c)" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_divide norm_nat_power norm_powr_real_powr) finally have summable_c: \ . note c also from s have "ereal c \ ereal (s \ 1)" by simp finally have "fds_abs_converges f s" by (intro fds_abs_converges) auto hence summable: "summable (\n. norm (fds_nth f n / nat_power n s))" by (simp add: fds_abs_converges_def) from summable_norm_cancel[OF this] have "(\n. fds_nth f n / nat_power n s) sums eval_fds f s" by (simp add: eval_fds_def sums_iff) from sums_split_initial_segment[OF this, of "Suc (Suc 0)"] have "norm (eval_fds f s - fds_nth f 1) = norm (\n. fds_nth f (n+2) / nat_power (n+2) s)" by (auto simp: sums_iff) also have "\ \ (\n. norm (fds_nth f (n+2) / nat_power (n+2) s))" by (intro summable_norm summable_ignore_initial_segment summable) also have "\ \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c / 2 powr (s \ 1 - c))" proof (intro suminf_le allI) fix n :: nat have "norm (fds_nth f (n + 2) / nat_power (n + 2) s) = norm (fds_nth f (n + 2)) / real (n+2) powr c / real (n+2) powr (s \ 1 - c)" by (simp add: field_simps powr_diff norm_divide norm_nat_power) also have "\ \ norm (fds_nth f (n + 2)) / real (n+2) powr c / 2 powr (s \ 1 - c)" using s by (intro divide_left_mono divide_nonneg_pos powr_mono2 mult_pos_pos) simp_all finally show "norm (fds_nth f (n + 2) / nat_power (n + 2) s) \ \" . qed (intro summable_ignore_initial_segment summable summable_divide summable_c)+ also have "\ = (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) / 2 powr (s \ 1 - c)" by (intro suminf_divide summable_ignore_initial_segment summable_c) also have "\ = B / 2 powr (s \ 1)" by (simp add: B_def powr_diff) finally show ?thesis . qed lemma tendsto_eval_fds_Re_at_top: assumes "conv_abscissa (f :: 'a :: dirichlet_series fds) \ \" assumes lim: "filterlim (\x. S x \ 1) at_top F" shows "((\x. eval_fds f (S x)) \ fds_nth f 1) F" proof - from assms(1) have "abs_conv_abscissa f < \" using abs_conv_le_conv_abscissa_plus_1[of f] by auto from ereal_dense2[OF this] obtain c where c: "abs_conv_abscissa f < ereal c" by auto define B where "B = (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" have *: "norm (eval_fds f s - fds_nth f 1) \ B / 2 powr (s \ 1)" if s: "s \ 1 \ c" for s using eval_fds_at_top_tail_bound[of f c s] that c by (simp add: B_def) moreover from lim have "eventually (\x. S x \ 1 \ c) F" by (auto simp: filterlim_at_top) ultimately have "eventually (\x. norm (eval_fds f (S x) - fds_nth f 1) \ B / 2 powr (S x \ 1)) F" by (auto elim!: eventually_mono) moreover have "((\x. B / 2 powr (S x \ 1)) \ 0) F" using filterlim_tendsto_pos_mult_at_top[OF tendsto_const[of "ln 2"] _ lim] by (intro real_tendsto_divide_at_top[OF tendsto_const]) (auto simp: powr_def mult_ac intro!: filterlim_compose[OF exp_at_top]) ultimately have "((\x. eval_fds f (S x) - fds_nth f 1) \ 0) F" by (rule Lim_null_comparison) thus ?thesis by (subst (asm) Lim_null [symmetric]) qed lemma tendsto_eval_fds_Re_at_top': assumes "conv_abscissa (f :: complex fds) \ \" shows "uniform_limit UNIV (\\ t. eval_fds f (of_real \ + of_real t * \) ) (\_ .fds_nth f 1) at_top" proof - from assms(1) have "abs_conv_abscissa f < \" using abs_conv_le_conv_abscissa_plus_1[of f] by auto from ereal_dense2[OF this] obtain c where c: "abs_conv_abscissa f < ereal c" by auto define B where "B \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" show ?thesis unfolding uniform_limit_iff proof safe fix \ :: real assume "\ > 0" hence "eventually (\\. B / 2 powr \ < \) at_top" by real_asymp thus "eventually (\\. \t\UNIV. dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) < \) at_top" using eventually_ge_at_top[of c] proof eventually_elim case (elim \) show ?case proof fix t :: real have "dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) \ B / 2 powr \" using eval_fds_at_top_tail_bound[of f c "of_real \ + of_real t * \"] elim c by (simp add: dist_norm B_def) also have "\ < \" by fact finally show "dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) < \" . qed qed qed qed lemma tendsto_eval_fds_Re_going_to_at_top: assumes "conv_abscissa (f :: 'a :: dirichlet_series fds) \ \" shows "((\s. eval_fds f s) \ fds_nth f 1) ((\s. s \ 1) going_to at_top)" using assms by (rule tendsto_eval_fds_Re_at_top) auto lemma tendsto_eval_fds_Re_going_to_at_top': assumes "conv_abscissa (f :: complex fds) \ \" shows "((\s. eval_fds f s) \ fds_nth f 1) (Re going_to at_top)" using assms by (rule tendsto_eval_fds_Re_at_top) auto text \ Any Dirichlet series that is not identically zero and does not diverge everywhere has a half-plane in which it converges and is non-zero. \ theorem fds_nonzero_halfplane_exists: fixes f :: "'a :: dirichlet_series fds" assumes "conv_abscissa f < \" "f \ 0" shows "eventually (\s. fds_converges f s \ eval_fds f s \ 0) ((\s. s \ 1) going_to at_top)" proof - from ereal_dense2[OF assms(1)] obtain c where c: "conv_abscissa f < ereal c" by auto have "eventually (\s::'a. s \ 1 > c) ((\s. s \ 1) going_to at_top)" using eventually_gt_at_top[of c] by auto hence "eventually (\s. fds_converges f s) ((\s. s \ 1) going_to at_top)" by eventually_elim (use c in \auto intro!: fds_converges simp: less_le_trans\) moreover have "eventually (\s. eval_fds f s \ 0) ((\s. s \ 1) going_to at_top)" using eval_fds_zeroD[OF assms(1)] assms(2) by (auto simp: frequently_def) ultimately show ?thesis by (rule eventually_conj) qed subsection \Normed series\ lemma fds_converges_norm_iff [simp]: fixes s :: "'a :: {nat_power_normed_field,banach}" shows "fds_converges (fds_norm f) (s \ 1) \ fds_abs_converges f s" unfolding fds_converges_def fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_abs_converges_norm_iff [simp]: fixes s :: "'a :: {nat_power_normed_field,banach}" shows "fds_abs_converges (fds_norm f) (s \ 1) \ fds_abs_converges f s" unfolding fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_converges_norm_iff': fixes f :: "'a :: {nat_power_normed_field,banach} fds" shows "fds_converges (fds_norm f) s \ fds_abs_converges f (of_real s)" unfolding fds_converges_def fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_abs_converges_norm_iff': fixes f :: "'a :: {nat_power_normed_field,banach} fds" shows "fds_abs_converges (fds_norm f) s \ fds_abs_converges f (of_real s)" unfolding fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma abs_conv_abscissa_norm [simp]: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (fds_norm f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa f \ abs_conv_abscissa (fds_norm f)" proof (rule abs_conv_abscissa_leI_weak) fix x assume "abs_conv_abscissa (fds_norm f) < ereal x" hence "fds_abs_converges (fds_norm f) (of_real x)" by (intro fds_abs_converges) auto thus "fds_abs_converges f (of_real x)" by (simp add: fds_abs_converges_norm_iff') qed qed (auto intro!: abs_conv_abscissa_leI_weak simp: fds_abs_converges_norm_iff' fds_abs_converges) lemma conv_abscissa_norm [simp]: fixes f :: "'a :: dirichlet_series fds" shows "conv_abscissa (fds_norm f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa f \ conv_abscissa (fds_norm f)" proof (rule abs_conv_abscissa_leI_weak) fix x assume "conv_abscissa (fds_norm f) < ereal x" hence "fds_converges (fds_norm f) (of_real x)" by (intro fds_converges) auto thus "fds_abs_converges f (of_real x)" by (simp add: fds_converges_norm_iff') qed qed (auto intro!: conv_abscissa_leI_weak simp: fds_abs_converges) lemma fixes f g :: "'a :: dirichlet_series fds" assumes "fds_abs_converges (fds_norm f) s" "fds_abs_converges (fds_norm g) s" shows fds_abs_converges_norm_mult: "fds_abs_converges (fds_norm (f * g)) s" and eval_fds_norm_mult_le: "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" proof - show conv: "fds_abs_converges (fds_norm (f * g)) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) have "fds_abs_converges (fds_norm f * fds_norm g) s" by (rule fds_abs_converges_mult assms)+ thus "summable (\n. norm (fds_nth (fds_norm f * fds_norm g) n) / nat_power n s)" by (simp add: fds_abs_converges_def) qed (auto intro!: always_eventually divide_right_mono order.trans[OF fds_nth_norm_mult_le] simp: norm_divide) have conv': "fds_abs_converges (fds_norm f * fds_norm g) s" by (intro fds_abs_converges_mult assms) hence "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f * fds_norm g) s" using conv unfolding eval_fds_def fds_abs_converges_def norm_divide by (intro suminf_le allI divide_right_mono) (simp_all add: norm_mult fds_nth_norm_mult_le) also have "\ = eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" by (intro eval_fds_mult assms) finally show "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" . qed lemma eval_fds_norm_nonneg: assumes "fds_abs_converges (fds_norm f) s" shows "eval_fds (fds_norm f) s \ 0" using assms unfolding eval_fds_def fds_abs_converges_def by (intro suminf_nonneg) auto lemma fixes f :: "'a :: dirichlet_series fds" assumes "fds_abs_converges (fds_norm f) s" shows fds_abs_converges_norm_power: "fds_abs_converges (fds_norm (f ^ n)) s" and eval_fds_norm_power_le: "eval_fds (fds_norm (f ^ n)) s \ eval_fds (fds_norm f) s ^ n" proof - show *: "fds_abs_converges (fds_norm (f ^ n)) s" for n by (induction n) (auto intro!: fds_abs_converges_norm_mult assms) show "eval_fds (fds_norm (f ^ n)) s \ eval_fds (fds_norm f) s ^ n" by (induction n) (auto intro!: order.trans[OF eval_fds_norm_mult_le] assms * mult_left_mono eval_fds_norm_nonneg) qed subsection \Logarithms of Dirichlet series\ (* TODO: Move ? *) lemma eventually_gt_ereal_at_top: "c \ \ \ eventually (\x. ereal x > c) at_top" by (cases c) auto lemma eval_fds_log_deriv: fixes s :: "'a :: dirichlet_series" assumes "fds_nth f 1 \ 0" "s \ 1 > abs_conv_abscissa f" "s \ 1 > abs_conv_abscissa (fds_deriv f / f)" assumes "eval_fds f s \ 0" shows "eval_fds (fds_deriv f / f) s = eval_fds (fds_deriv f) s / eval_fds f s" proof - have "eval_fds (fds_deriv f / f * f) s = eval_fds (fds_deriv f / f) s * eval_fds f s" using assms by (intro eval_fds_mult fds_abs_converges) auto also have "fds_deriv f / f * f = fds_deriv f * (f * inverse f)" by (simp add: divide_fds_def algebra_simps) also have "f * inverse f = 1" using assms by (intro fds_right_inverse) finally show ?thesis using assms by simp qed text \ Given a sufficiently nice absolutely convergent Dirichlet series that converges to some function $f(s)$ and a holomorphic branch of $\ln f(s)$, we can construct a Dirichlet series that absolutely converges to that logarithm. \ lemma eval_fds_ln: fixes s0 :: ereal assumes nz: "\s. Re s > s0 \ eval_fds f s \ 0" "fds_nth f 1 \ 0" assumes l: "exp l = fds_nth f 1" "((g \ of_real) \ l) at_top" assumes g: "\s. Re s > s0 \ exp (g s) = eval_fds f s" assumes holo_g: "g holomorphic_on {s. Re s > s0}" assumes "ereal (Re s) > s0" assumes "s0 \ abs_conv_abscissa f" and "s0 \ abs_conv_abscissa (fds_deriv f / f)" shows "eval_fds (fds_ln l f) s = g s" proof - let ?s0 = "abs_conv_abscissa f" and ?s1 = "abs_conv_abscissa (inverse f)" let ?h = "\s. eval_fds (fds_ln l f) s - g s" let ?A = "{s. Re s > s0}" have open_A: "open ?A" by (cases s0) (auto simp: open_halfspace_Re_gt) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) moreover from assms have "\ \ \" by auto ultimately have "conv_abscissa f \ \" by auto have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) finally have "conv_abscissa (fds_ln l f) \ \" using assms by (auto simp: max_def abs_conv_abscissa_deriv split: if_splits) have deriv_g [derivative_intros]: "(g has_field_derivative eval_fds (fds_deriv f) s / eval_fds f s) (at s within B)" if s: "Re s > s0" for s B proof - have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ s0" using assms by simp also have "\ < Re s" by fact finally have s': "Re s > conv_abscissa f" . have deriv_g: "(g has_field_derivative deriv g s) (at s)" using holomorphic_derivI[OF holo_g open_A, of s] s by (auto simp: at_within_open[OF _ open_A]) have "((\s. exp (g s)) has_field_derivative eval_fds f s * deriv g s) (at s)" (is ?P) by (rule derivative_eq_intros deriv_g s)+ (insert s, simp_all add: g) also from s have ev: "eventually (\t. t \ ?A) (nhds s)" by (intro eventually_nhds_in_open open_A) auto have "?P \ (eval_fds f has_field_derivative eval_fds f s * deriv g s) (at s)" by (intro DERIV_cong_ev refl eventually_mono[OF ev]) (auto simp: g) finally have "(eval_fds f has_field_derivative eval_fds f s * deriv g s) (at s)" . moreover have "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s)" using s' assms by (intro derivative_intros) auto ultimately have "eval_fds f s * deriv g s = eval_fds (fds_deriv f) s" by (rule DERIV_unique) hence "deriv g s = eval_fds (fds_deriv f) s / eval_fds f s" using s nz by (simp add: field_simps) with deriv_g show ?thesis by (auto intro: has_field_derivative_at_within) qed have "\c. \z\{z. Re z > s0}. ?h z = c" proof (rule has_field_derivative_zero_constant, goal_cases) case 1 show ?case using convex_halfspace_gt[of _ "1::complex"] by (cases s0) auto next case (2 z) have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" by (simp add: abs_conv_abscissa_ln) also have "\ < Re z" using 2 assms by (auto simp: abs_conv_abscissa_deriv) finally have s1: "conv_abscissa (fds_ln l f) < ereal (Re z)" . have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ < Re z" using 2 assms by auto finally have s2: "conv_abscissa f < ereal (Re z)" . from l have "fds_nth f 1 \ 0" by auto with 2 assms have *: "eval_fds (fds_deriv f / f) z = eval_fds (fds_deriv f) z / (eval_fds f z)" by (auto simp: eval_fds_log_deriv) have "eval_fds f z \ 0" using 2 assms by auto show ?case using s1 s2 2 nz by (auto intro!: derivative_eq_intros simp: * field_simps) qed then obtain c where c: "\z. Re z > s0 \ ?h z = c" by blast have "(at_top :: real filter) \ bot" by simp moreover from assms have "s0 \ \" by auto have "eventually (\x. c = (?h \ of_real) x) at_top" using eventually_gt_ereal_at_top[OF \s0 \ \\] by eventually_elim (simp add: c) hence "((?h \ of_real) \ c) at_top" by (force intro: Lim_transform_eventually) moreover have "((?h \ of_real) \ fds_nth (fds_ln l f) 1 - l) at_top" using \conv_abscissa (fds_ln l f) \ \\ and l unfolding o_def by (intro tendsto_intros tendsto_eval_fds_Re_at_top) (auto simp: filterlim_ident) ultimately have "c = fds_nth (fds_ln l f) 1 - l" by (rule tendsto_unique) with c[OF \Re s > s0\] and l and nz show ?thesis by (simp add: exp_minus field_simps) qed text \ Less explicitly: For a sufficiently nice absolutely convergent Dirichlet series converging to a function $f(s)$, the formal logarithm absolutely converges to some logarithm of $f(s)$. \ lemma eval_fds_ln': fixes s0 :: ereal assumes "ereal (Re s) > s0" assumes "s0 \ abs_conv_abscissa f" and "s0 \ abs_conv_abscissa (fds_deriv f / f)" and nz: "\s. Re s > s0 \ eval_fds f s \ 0" "fds_nth f 1 \ 0" assumes l: "exp l = fds_nth f 1" shows "exp (eval_fds (fds_ln l f) s) = eval_fds f s" proof - let ?s0 = "abs_conv_abscissa f" and ?s1 = "abs_conv_abscissa (inverse f)" let ?h = "\s. eval_fds f s * exp (-eval_fds (fds_ln l f) s)" have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) moreover from assms have "\ \ \" by auto ultimately have "conv_abscissa f \ \" by auto have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) finally have "conv_abscissa (fds_ln l f) \ \" using assms by (auto simp: max_def abs_conv_abscissa_deriv split: if_splits) have "\c. \z\{z. Re z > s0}. ?h z = c" proof (rule has_field_derivative_zero_constant, goal_cases) case 1 show ?case using convex_halfspace_gt[of _ "1::complex"] by (cases s0) auto next case (2 z) have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) also have "\ < Re z" using 2 assms by (auto simp: abs_conv_abscissa_deriv) finally have s1: "conv_abscissa (fds_ln l f) < ereal (Re z)" . have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ < Re z" using 2 assms by auto finally have s2: "conv_abscissa f < ereal (Re z)" . from l have "fds_nth f 1 \ 0" by auto with 2 assms have *: "eval_fds (fds_deriv f / f) z = eval_fds (fds_deriv f) z / (eval_fds f z)" by (subst eval_fds_log_deriv) auto have "eval_fds f z \ 0" using 2 assms by auto thus ?case using s1 s2 by (auto intro!: derivative_eq_intros simp: *) qed then obtain c where c: "\z. Re z > s0 \ ?h z = c" by blast have "(at_top :: real filter) \ bot" by simp moreover from assms have "s0 \ \" by auto have "eventually (\x. c = (?h \ of_real) x) at_top" using eventually_gt_ereal_at_top[OF \s0 \ \\] by eventually_elim (simp add: c) hence "((?h \ of_real) \ c) at_top" by (force intro: Lim_transform_eventually) moreover have "((?h \ of_real) \ fds_nth f 1 * exp (-fds_nth (fds_ln l f) 1)) at_top" unfolding o_def using \conv_abscissa (fds_ln l f) \ \\ and \conv_abscissa f \ \\ by (intro tendsto_intros tendsto_eval_fds_Re_at_top) (auto simp: filterlim_ident) ultimately have "c = fds_nth f 1 * exp (-fds_nth (fds_ln l f) 1)" by (rule tendsto_unique) with c[OF \Re s > s0\] and l and nz show ?thesis by (simp add: exp_minus field_simps) qed lemma fds_ln_completely_multiplicative: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" assumes "fds_nth f 1 \ 0" shows "fds_ln l f = fds (\n. if n = 1 then l else fds_nth f n * mangoldt n /\<^sub>R ln n)" proof - have "fds_ln l f = fds_integral l (fds_deriv f / f)" by (simp add: fds_ln_def) also have "fds_deriv f = -fds (\n. fds_nth f n * mangoldt n) * f" by (intro completely_multiplicative_fds_deriv' assms) also have "\ / f = -fds (\n. fds_nth f n * mangoldt n) * (f * inverse f)" by (simp add: divide_fds_def) also from assms have "f * inverse f = 1" by (simp add: fds_right_inverse) also have "fds_integral l (- fds (\n. fds_nth f n * mangoldt n) * 1) = fds (\n. if n = 1 then l else fds_nth f n * mangoldt n /\<^sub>R ln n)" by (simp add: fds_integral_def cong: if_cong) finally show ?thesis . qed lemma eval_fds_ln_completely_multiplicative_strong: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" and g :: "nat \ 'a" defines "h \ fds (\n. fds_nth (fds_ln l f) n * g n)" assumes "fds_abs_converges h s" assumes "completely_multiplicative_function (fds_nth f)" and "fds_nth f 1 \ 0" shows "(\(p,k). (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k)) abs_summable_on ({p. prime p} \ UNIV)" (is ?th1) and "eval_fds h s = l * g 1 + (\\<^sub>a(p, k)\{p. prime p}\UNIV. (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k))" (is ?th2) proof - let ?P = "{p::nat. prime p}" interpret f: completely_multiplicative_function "fds_nth f" by fact from assms have *: "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV" by (auto simp: abs_summable_on_nat_iff' fds_abs_converges_def) have eq: "h = fds (\n. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n))" using fds_ln_completely_multiplicative [OF assms(3), of l] by (simp add: h_def fds_eq_iff) note * also have "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV \ (\x. if x = Suc 0 then l * g 1 else fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on {1} \ Collect primepow" using eq by (intro abs_summable_on_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) finally have sum1: "(\x. if x = Suc 0 then l * g 1 else fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on Collect primepow" by (rule abs_summable_on_subset) auto also have "?this \ (\x. fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on Collect primepow" by (intro abs_summable_on_cong) (insert primepow_gt_Suc_0, auto) also have "\ \ (\(p,k). fds_nth f (p ^ Suc k) * g (p ^ Suc k) * mangoldt (p ^ Suc k) /\<^sub>R ln (real (p ^ Suc k)) / nat_power (p ^ Suc k) s) abs_summable_on (?P \ UNIV)" using bij_betw_primepows unfolding case_prod_unfold by (intro abs_summable_on_reindex_bij_betw [symmetric]) also have "\ \ ?th1" by (intro abs_summable_on_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps scaleR_conv_of_real simp del: power_Suc) finally show ?th1 . have "eval_fds h s = (\\<^sub>an. fds_nth h n / nat_power n s)" using * unfolding eval_fds_def by (subst infsetsum_nat') auto also have "\ = (\\<^sub>an \ {1} \ {n. primepow n}. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" by (intro infsetsum_cong_neutral) (auto simp: eq fds_nth_fds mangoldt_def) also have "\ = l * g 1 + (\\<^sub>an | primepow n. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" (is "_ = _ + ?x") using sum1 primepow_gt_Suc_0 by (subst infsetsum_Un_disjoint) auto also have "?x = (\\<^sub>an\Collect primepow. fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" (is "_ = infsetsum ?f _") by (intro infsetsum_cong refl) (insert primepow_gt_Suc_0, auto) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). fds_nth f (p ^ Suc k) * g (p ^ Suc k) * mangoldt (p ^ Suc k) /\<^sub>R ln (p ^ Suc k) / nat_power (p ^ Suc k) s)" using bij_betw_primepows unfolding case_prod_unfold by (intro infsetsum_reindex_bij_betw [symmetric]) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k))" by (intro infsetsum_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps scaleR_conv_of_real simp del: power_Suc) finally show ?th2 . qed lemma eval_fds_ln_completely_multiplicative: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_nth f 1 \ 0" assumes "s \ 1 > abs_conv_abscissa (fds_deriv f / f)" shows "(\(p,k). (fds_nth f p / nat_power p s) ^ Suc k / of_nat (Suc k)) abs_summable_on ({p. prime p} \ UNIV)" (is ?th1) and "eval_fds (fds_ln l f) s = l + (\\<^sub>a(p, k)\{p. prime p}\UNIV. (fds_nth f p / nat_power p s) ^ Suc k / of_nat (Suc k))" (is ?th2) proof - from assms have "fds_abs_converges (fds_ln l f) s" by (intro fds_abs_converges_ln) (auto intro!: fds_abs_converges_mult intro: fds_abs_converges) hence "fds_abs_converges (fds (\n. fds_nth (fds_ln l f) n * 1)) s" by simp from eval_fds_ln_completely_multiplicative_strong [OF this assms(1,2)] show ?th1 ?th2 by simp_all qed subsection \Exponential and logarithm\ lemma summable_fds_exp_aux: assumes "fds_nth f' 1 = (0 :: 'a :: real_normed_algebra_1)" shows "summable (\k. fds_nth (f' ^ k) n /\<^sub>R fact k)" proof (rule summable_finite) fix k assume "k \ {..n}" hence "n < k" by simp - also have "\ < 2 ^ k" by (induction k) auto + also have "\ < 2 ^ k" + by (rule less_exp) finally have "fds_nth (f' ^ k) n = 0" using assms by (intro fds_nth_power_eq_0) auto thus "fds_nth (f' ^ k) n /\<^sub>R fact k = 0" by simp qed auto lemma fixes f :: "'a :: dirichlet_series fds" assumes "fds_abs_converges f s" shows fds_abs_converges_exp: "fds_abs_converges (fds_exp f) s" and eval_fds_exp: "eval_fds (fds_exp f) s = exp (eval_fds f s)" proof - have conv: "fds_abs_converges (fds_exp f) s" and ev: "eval_fds (fds_exp f) s = exp (eval_fds f s)" if "fds_abs_converges f s" and [simp]: "fds_nth f (Suc 0) = 0" for f proof - have [simp]: "fds (\n. if n = Suc 0 then 0 else fds_nth f n) = f" by (intro fds_eqI) simp_all have "(\(k,n). fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on (UNIV \ {1..})" proof (subst abs_summable_on_Sigma_iff, safe, goal_cases) case (3 k) from that have "fds_abs_converges (f ^ k) s" by (intro fds_abs_converges_power) hence "(\n. fds_nth (f ^ k) n / nat_power n s * inverse (fact k)) abs_summable_on {1..}" unfolding fds_abs_converges_altdef by (intro abs_summable_on_cmult_left) thus ?case by (simp add: field_simps) next case 4 show ?case unfolding abs_summable_on_nat_iff' proof (rule summable_comparison_test_ev[OF always_eventually[OF allI]]) fix k :: nat from that have *: "fds_abs_converges (fds_norm (f ^ k)) (s \ 1)" by (auto simp: fds_abs_converges_power) have "(\\<^sub>an\{1..}. norm (fds_nth (f ^ k) n / fact k / nat_power n s)) = (\\<^sub>an\{1..}. fds_nth (fds_norm (f ^ k)) n / nat_power n (s \ 1) / fact k)" (is "?S = _") by (intro infsetsum_cong) (simp_all add: norm_divide norm_mult norm_nat_power) also have "\ = (\\<^sub>an\{1..}. fds_nth (fds_norm (f ^ k)) n / nat_power n (s \ 1)) /\<^sub>R fact k" (is "_ = ?S' /\<^sub>R _") using * unfolding fds_abs_converges_altdef by (subst infsetsum_cdiv) (auto simp: abs_summable_on_nat_iff scaleR_conv_of_real divide_simps) also have "?S' = eval_fds (fds_norm (f ^ k)) (s \ 1)" using * unfolding fds_abs_converges_altdef eval_fds_def by (subst infsetsum_nat) (auto intro!: suminf_cong) finally have eq: "?S = \ /\<^sub>R fact k" . note eq also have "?S \ 0" by (intro infsetsum_nonneg) auto hence "?S = norm (norm ?S)" by simp also have "eval_fds (fds_norm (f ^ k)) (s \ 1) \ eval_fds (fds_norm f) (s \ 1) ^ k" using that by (intro eval_fds_norm_power_le) auto finally show "norm (norm (\\<^sub>an\{1..}. norm (fds_nth (f ^ k) n / fact k / nat_power n s))) \ eval_fds (fds_norm f) (s \ 1) ^ k /\<^sub>R fact k" by (simp add: divide_right_mono) next from exp_converges[of "eval_fds (fds_norm f) (s \ 1)"] show "summable (\x. eval_fds (fds_norm f) (s \ 1) ^ x /\<^sub>R fact x)" by (simp add: sums_iff) qed qed auto hence summable: "(\(n,k). fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on {1..} \ UNIV" by (subst abs_summable_on_Times_swap) (simp add: case_prod_unfold) have summable': "(\k. fds_nth (f ^ k) n / fact k) abs_summable_on UNIV" for n using abs_summable_on_cmult_left[of "nat_power n s", OF abs_summable_on_Sigma_project2 [OF summable, of n]] by (cases "n = 0") simp_all have "(\n. \\<^sub>ak. fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on {1..}" using summable by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\n. (\k. fds_nth (f ^ k) n / fact k) * inverse (nat_power n s)) abs_summable_on {1..}" proof (intro abs_summable_on_cong refl, goal_cases) case (1 n) hence "(\\<^sub>ak. fds_nth (f ^ k) n / fact k / nat_power n s) = (\\<^sub>ak. fds_nth (f ^ k) n / fact k) * inverse (nat_power n s)" using summable'[of n] by (subst infsetsum_cmult_left [symmetric]) (auto simp: field_simps) also have "(\\<^sub>ak. fds_nth (f ^ k) n / fact k) = (\k. fds_nth (f ^ k) n / fact k)" using summable'[of n] 1 by (intro abs_summable_on_cong refl infsetsum_nat') auto finally show ?case . qed finally show "fds_abs_converges (fds_exp f) s" by (simp add: fds_exp_def fds_nth_fds' abs_summable_on_Sigma_iff scaleR_conv_of_real fds_abs_converges_altdef field_simps) have "eval_fds (fds_exp f) s = (\n. (\k. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" by (simp add: fds_exp_def eval_fds_def fds_nth_fds') also have "\ = (\n. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" proof (intro suminf_cong, goal_cases) case (1 n) show ?case proof (cases "n = 0") case False have "(\k. fds_nth (f ^ k) n /\<^sub>R fact k) = (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k)" using summable'[of n] False by (intro infsetsum_nat' [symmetric]) (auto simp: scaleR_conv_of_real field_simps) thus ?thesis by simp qed simp_all qed also have "\ = (\\<^sub>an. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" proof (intro infsetsum_nat' [symmetric], goal_cases) case 1 have *: "UNIV - {Suc 0..} = {0}" by auto have "(\x. \\<^sub>ay. fds_nth (f ^ y) x / fact y / nat_power x s) abs_summable_on {1..}" by (intro abs_summable_on_Sigma_project1'[OF summable]) auto also have "?this \ (\x. (\\<^sub>ay. fds_nth (f ^ y) x / fact y) * inverse (nat_power x s)) abs_summable_on {1..}" using summable' by (intro abs_summable_on_cong refl, subst infsetsum_cmult_left [symmetric]) (auto simp: field_simps) also have "\ \ (\x. (\\<^sub>ay. fds_nth (f ^ y) x /\<^sub>R fact y) / (nat_power x s)) abs_summable_on {1..}" by (simp add: field_simps scaleR_conv_of_real) finally show ?case by (rule abs_summable_on_finite_diff) (use * in auto) qed also have "\ = (\\<^sub>an. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k * inverse (nat_power n s)))" using summable' by (subst infsetsum_cmult_left) (auto simp: field_simps scaleR_conv_of_real) also have "\ = (\\<^sub>an\{1..}. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k * inverse (nat_power n s)))" by (intro infsetsum_cong_neutral) (auto simp: Suc_le_eq) also have "\ = (\\<^sub>ak. \\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s /\<^sub>R fact k)" using summable by (subst infsetsum_swap) (auto simp: field_simps scaleR_conv_of_real case_prod_unfold) also have "\ = (\\<^sub>ak. (\\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s) /\<^sub>R fact k)" by (subst infsetsum_scaleR_right) simp also have "\ = (\\<^sub>ak. eval_fds f s ^ k /\<^sub>R fact k)" proof (intro infsetsum_cong refl, goal_cases) case (1 k) have *: "fds_abs_converges (f ^ k) s" by (intro fds_abs_converges_power that) have "(\\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s) = (\\<^sub>an. fds_nth (f ^ k) n / nat_power n s)" by (intro infsetsum_cong_neutral) (auto simp: Suc_le_eq) also have "\ = eval_fds (f ^ k) s" using * unfolding eval_fds_def by (intro infsetsum_nat') (auto simp: fds_abs_converges_def abs_summable_on_nat_iff') also from that have "\ = eval_fds f s ^ k" by (simp add: eval_fds_power) finally show ?case by simp qed also have "\ = (\k. eval_fds f s ^ k /\<^sub>R fact k)" using exp_converges[of "norm (eval_fds f s)"] by (intro infsetsum_nat') (auto simp: abs_summable_on_nat_iff' sums_iff field_simps norm_power) also have "\ = exp (eval_fds f s)" by (simp add: exp_def) finally show "eval_fds (fds_exp f) s = exp (eval_fds f s)" . qed define f' where "f' = f - fds_const (fds_nth f 1)" have *: "fds_abs_converges (fds_exp f') s" by (auto simp: f'_def intro!: fds_abs_converges_diff conv assms) have "fds_abs_converges (fds_const (exp (fds_nth f 1)) * fds_exp f') s" unfolding f'_def by (intro fds_abs_converges_mult conv fds_abs_converges_diff assms) auto thus "fds_abs_converges (fds_exp f) s" unfolding f'_def by (simp add: fds_exp_times_fds_nth_0) have "eval_fds (fds_exp f) s = eval_fds (fds_const (exp (fds_nth f 1)) * fds_exp f') s" by (simp add: f'_def fds_exp_times_fds_nth_0) also have "\ = exp (fds_nth f (Suc 0)) * eval_fds (fds_exp f') s" using * using assms by (subst eval_fds_mult) (simp_all) also have "\ = exp (eval_fds f s)" using ev[of f'] assms unfolding f'_def by (auto simp: fds_abs_converges_diff eval_fds_diff fds_abs_converges_imp_converges exp_diff) finally show "eval_fds (fds_exp f) s = exp (eval_fds f s)" . qed lemma fds_exp_add: fixes f :: "'a :: dirichlet_series fds" shows "fds_exp (f + g) = fds_exp f * fds_exp g" proof (rule fds_eqI_truncate) fix m :: nat assume m: "m > 0" let ?T = "fds_truncate m" have "?T (fds_exp (f + g)) = ?T (fds_exp (?T f + ?T g))" by (simp add: fds_truncate_exp fds_truncate_add_strong [symmetric]) also have "fds_exp (?T f + ?T g) = fds_exp (?T f) * fds_exp (?T g)" proof (rule eval_fds_eqD) have "fds_abs_converges (fds_exp (?T f + ?T g)) 0" by (intro fds_abs_converges_exp fds_abs_converges_add) auto thus "conv_abscissa (fds_exp (?T f + ?T g)) < \" using conv_abscissa_PInf_iff by blast hence "fds_abs_converges (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) 0" by (intro fds_abs_converges_mult fds_abs_converges_exp) auto thus "conv_abscissa (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) < \" using conv_abscissa_PInf_iff by blast show "frequently (\s. eval_fds (fds_exp (fds_truncate m f + fds_truncate m g)) s = eval_fds (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) s) ((\s. s \ 1) going_to at_top)" by (auto simp: eval_fds_add eval_fds_mult eval_fds_exp fds_abs_converges_add fds_abs_converges_exp exp_add) qed also have "?T \ = ?T (fds_exp f * fds_exp g)" by (subst fds_truncate_mult [symmetric], subst (1 2) fds_truncate_exp) (simp add: fds_truncate_mult) finally show "?T (fds_exp (f + g)) = \" . qed lemma fds_exp_minus: fixes f :: "'a :: dirichlet_series fds" shows "fds_exp (-f) = inverse (fds_exp f)" proof (rule fds_right_inverse_unique) have "fds_exp f * fds_exp (- f) = fds_exp (f + (-f))" by (subst fds_exp_add) simp_all also have "f + (-f) = 0" by simp also have "fds_exp \ = 1" by simp finally show "fds_exp f * fds_exp (-f) = 1" . qed lemma abs_conv_abscissa_exp: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (fds_exp f) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono fds_abs_converges_exp) lemma fds_deriv_exp [simp]: fixes f :: "'a :: dirichlet_series fds" shows "fds_deriv (fds_exp f) = fds_exp f * fds_deriv f" proof (rule fds_eqI_truncate) fix m :: nat assume m: "m > 0" let ?T = "fds_truncate m" have "abs_conv_abscissa (fds_deriv (?T f)) = -\" by (simp add: abs_conv_abscissa_deriv) have "?T (fds_deriv (fds_exp f)) = ?T (fds_deriv (fds_exp (?T f)))" by (simp add: fds_truncate_deriv fds_truncate_exp) also have "fds_deriv (fds_exp (?T f)) = fds_exp (?T f) * fds_deriv (?T f)" proof (rule eval_fds_eqD) note abscissa = conv_le_abs_conv_abscissa abs_conv_abscissa_exp note abscissa' = abscissa[THEN le_less_trans] have "fds_abs_converges (fds_deriv (fds_exp (fds_truncate m f))) 0" by (intro fds_abs_converges ) (auto simp: abs_conv_abscissa_deriv intro: le_less_trans[OF abs_conv_abscissa_exp]) thus "conv_abscissa (fds_deriv (fds_exp (fds_truncate m f))) < \" using conv_abscissa_PInf_iff by blast have "fds_abs_converges (fds_exp (fds_truncate m f) * fds_deriv (fds_truncate m f)) 0" by (intro fds_abs_converges_mult fds_abs_converges_exp) (auto intro: fds_abs_converges simp add: fds_truncate_deriv [symmetric]) thus "conv_abscissa (fds_exp (fds_truncate m f) * fds_deriv (fds_truncate m f)) < \" using conv_abscissa_PInf_iff by blast show "\\<^sub>F s in (\s. s \ 1) going_to at_top. eval_fds (fds_deriv (fds_exp (?T f))) s = eval_fds (fds_exp (?T f) * fds_deriv (?T f)) s" proof (intro always_eventually eventually_frequently allI, goal_cases) case (2 s) have "eval_fds (fds_deriv (fds_exp (?T f))) s = deriv (eval_fds (fds_exp (?T f))) s" by (auto simp: eval_fds_exp eval_fds_mult fds_abs_converges_mult fds_abs_converges_exp fds_abs_converges eval_fds_deriv abscissa') also have "eval_fds (fds_exp (?T f)) = (\s. exp (eval_fds (?T f) s))" by (intro ext eval_fds_exp) auto also have "deriv \ = (\s. exp (eval_fds (?T f) s) * deriv (eval_fds (?T f)) s)" by (auto intro!: DERIV_imp_deriv derivative_eq_intros simp: eval_fds_deriv) also have "\ = eval_fds (fds_exp (?T f) * fds_deriv (?T f))" by (auto simp: eval_fds_exp eval_fds_mult fds_abs_converges_mult fds_abs_converges_exp fds_abs_converges eval_fds_deriv abs_conv_abscissa_deriv) finally show ?case . qed auto qed also have "?T \ = ?T (fds_exp f * fds_deriv f)" by (subst fds_truncate_mult [symmetric]) (simp add: fds_truncate_exp fds_truncate_deriv [symmetric], simp add: fds_truncate_mult) finally show "?T (fds_deriv (fds_exp f)) = \" . qed lemma fds_exp_ln_strong: fixes f :: "'a :: dirichlet_series fds" assumes "fds_nth f (Suc 0) \ 0" shows "fds_exp (fds_ln l f) = fds_const (exp l / fds_nth f (Suc 0)) * f" proof - let ?c = "exp l / fds_nth f (Suc 0)" have "f * fds_const ?c = f * (fds_exp (-fds_ln l f) * fds_exp (fds_ln l f)) * fds_const ?c" (is "_ = _ * (?g * ?h) * _") by (subst fds_exp_add [symmetric]) simp also have "\ = fds_const ?c * (f * ?g) * ?h" by (simp add: mult_ac) also have "f * ?g = fds_const (inverse ?c)" proof (rule fds_deriv_eq_imp_eq) have "fds_deriv (f * fds_exp (-fds_ln l f)) = fds_exp (- fds_ln l f) * fds_deriv f * (1 - f / f)" by (simp add: divide_fds_def algebra_simps) also from assms have "f / f = 1" by (simp add: divide_fds_def fds_right_inverse) finally show "fds_deriv (f * fds_exp (-fds_ln l f)) = fds_deriv (fds_const (inverse ?c))" by simp qed (insert assms, auto simp: exp_minus field_simps) also have "fds_const ?c * fds_const (inverse ?c) = 1" using assms by (subst fds_const_mult [symmetric]) (simp add: divide_simps) finally show ?thesis by (simp add: mult_ac) qed lemma fds_exp_ln [simp]: fixes f :: "'a :: dirichlet_series fds" assumes "exp l = fds_nth f (Suc 0)" shows "fds_exp (fds_ln l f) = f" using assms by (subst fds_exp_ln_strong) auto lemma fds_ln_exp [simp]: fixes f :: "'a :: dirichlet_series fds" assumes "l = fds_nth f (Suc 0)" shows "fds_ln l (fds_exp f) = f" proof (rule fds_deriv_eq_imp_eq) have "fds_deriv (fds_ln l (fds_exp f)) = fds_deriv f * (fds_exp f / fds_exp f)" by (simp add: algebra_simps divide_fds_def) also have "fds_exp f / fds_exp f = 1" by (simp add: divide_fds_def fds_right_inverse) finally show "fds_deriv (fds_ln l (fds_exp f)) = fds_deriv f" by simp qed (insert assms, auto simp: field_simps) subsection \Euler products\ lemma fds_euler_product_LIMSEQ: fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "(\n. \p\n. if prime p then \i. fds_nth f (p ^ i) / nat_power (p ^ i) s else 1) \ eval_fds f s" unfolding eval_fds_def proof (rule euler_product_LIMSEQ) show "multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_euler_product_LIMSEQ': fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "(\n. \p\n. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1) \ eval_fds f s" unfolding eval_fds_def proof (rule euler_product_LIMSEQ') show "completely_multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule completely_multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_euler_product: fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "abs_convergent_prod (\p. if prime p then \i. fds_nth f (p ^ i) / nat_power (p ^ i) s else 1)" unfolding eval_fds_def proof (rule abs_convergent_euler_product) show "multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_euler_product': fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "abs_convergent_prod (\p. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1)" unfolding eval_fds_def proof (rule abs_convergent_euler_product') show "completely_multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule completely_multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_zero_iff: fixes f :: "'a :: {nat_power_field, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" assumes "fds_abs_converges f s" shows "eval_fds f s = 0 \ (\p. prime p \ fds_nth f p = nat_power p s)" proof - let ?g = "\p. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1" have lim: "(\n. \p\n. ?g p) \ eval_fds f s" by (intro fds_euler_product_LIMSEQ' assms) have conv: "convergent_prod ?g" by (intro abs_convergent_prod_imp_convergent_prod fds_abs_convergent_euler_product' assms) { assume "eval_fds f s = 0" from convergent_prod_to_zero_iff[OF conv] and this and lim have "\p. prime p \ fds_nth f p = nat_power p s" by (auto split: if_splits) } moreover { assume "\p. prime p \ fds_nth f p = nat_power p s" then obtain p where "prime p" "fds_nth f p = nat_power p s" by blast moreover from this have "nat_power p s \ 0" by (intro nat_power_nonzero) (auto simp: prime_gt_0_nat) ultimately have "(\n. \p\n. ?g p) \ 0" using convergent_prod_to_zero_iff[OF conv] by (auto intro!: exI[of _ p] split: if_splits) from tendsto_unique[OF _ lim this] have "eval_fds f s = 0" by simp } ultimately show ?thesis by blast qed lemma fixes s :: "'a :: {nat_power_normed_field,banach,euclidean_space}" assumes "s \ 1 > 1" shows euler_product_fds_zeta: "(\n. \p\n. if prime p then inverse (1 - 1 / nat_power p s) else 1) \ eval_fds fds_zeta s" (is ?th1) and eval_fds_zeta_nonzero: "eval_fds fds_zeta s \ 0" proof - have *: "completely_multiplicative_function (fds_nth fds_zeta)" by standard auto have lim: "(\n. \p\n. if prime p then inverse (1 - fds_nth fds_zeta p / nat_power p s) else 1) \ eval_fds fds_zeta s" (is "filterlim ?g _ _") using assms by (intro fds_euler_product_LIMSEQ' * fds_abs_summable_zeta) also have "?g = (\n. \p\n. if prime p then inverse (1 - 1 / nat_power p s) else 1)" by (intro ext prod.cong refl) (auto simp: fds_zeta_def fds_nth_fds) finally show ?th1 . { fix p :: nat assume "prime p" from this have "p > 1" by (simp add: prime_gt_Suc_0_nat) hence "norm (nat_power p s) = real p powr (s \ 1)" by (simp add: norm_nat_power) also have "\ > real p powr 0" using assms and \p > 1\ by (intro powr_less_mono) auto finally have "nat_power p s \ 1" using \p > 1\ by auto } hence **: "\p. prime p \ fds_nth fds_zeta p = nat_power p s" by (auto simp: fds_zeta_def fds_nth_fds) show "eval_fds fds_zeta s \ 0" using assms * ** by (subst fds_abs_convergent_zero_iff) simp_all qed lemma fds_primepow_subseries_euler_product_cm: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" "prime p" assumes "s \ 1 > abs_conv_abscissa f" shows "eval_fds (fds_primepow_subseries p f) s = 1 / (1 - fds_nth f p / nat_power p s)" proof - let ?f = "(\n. \pa\n. if prime pa then inverse (1 - fds_nth (fds_primepow_subseries p f) pa / nat_power pa s) else 1)" have "sequentially \ bot" by simp moreover have "?f \ eval_fds (fds_primepow_subseries p f) s" by (intro fds_euler_product_LIMSEQ' completely_multiplicative_function_only_pows assms fds_abs_converges_subseries) (insert assms, auto intro!: fds_abs_converges) moreover have "eventually (\n. ?f n = 1 / (1 - fds_nth f p / nat_power p s)) at_top" using eventually_ge_at_top[of p] proof eventually_elim case (elim n) have "(\pa\n. if prime pa then inverse (1 - fds_nth (fds_primepow_subseries p f) pa / nat_power pa s) else 1) = (\q\n. if q = p then inverse (1 - fds_nth f p / nat_power p s) else 1)" using \prime p\ by (intro prod.cong) (auto simp: fds_nth_subseries prime_prime_factors) also have "\ = 1 / (1 - fds_nth f p / nat_power p s)" using elim by (subst prod.delta) (auto simp: divide_simps) finally show ?case . qed hence "?f \ 1 / (1 - fds_nth f p / nat_power p s)" by (rule tendsto_eventually) ultimately show ?thesis by (rule tendsto_unique) qed subsection \Non-negative Dirichlet series\ lemma nonneg_Reals_sum: "(\x. x \ A \ f x \ \\<^sub>\\<^sub>0) \ sum f A \ \\<^sub>\\<^sub>0" by (induction A rule: infinite_finite_induct) auto locale nonneg_dirichlet_series = fixes f :: "'a :: dirichlet_series fds" assumes nonneg_coeffs_aux: "n > 0 \ fds_nth f n \ \\<^sub>\\<^sub>0" begin lemma nonneg_coeffs: "fds_nth f n \ \\<^sub>\\<^sub>0" using nonneg_coeffs_aux[of n] by (cases "n = 0") auto end lemma nonneg_dirichlet_series_0 [simp,intro]: "nonneg_dirichlet_series 0" by standard (auto simp: zero_fds_def) lemma nonneg_dirichlet_series_1 [simp,intro]: "nonneg_dirichlet_series 1" by standard (auto simp: one_fds_def) lemma nonneg_dirichlet_series_const [simp,intro]: "c \ \\<^sub>\\<^sub>0 \ nonneg_dirichlet_series (fds_const c)" by standard (auto simp: fds_const_def) lemma nonneg_dirichlet_series_add [intro]: assumes "nonneg_dirichlet_series f" "nonneg_dirichlet_series g" shows "nonneg_dirichlet_series (f + g)" proof - interpret f: nonneg_dirichlet_series f by fact interpret g: nonneg_dirichlet_series g by fact show ?thesis by standard (auto intro!: nonneg_Reals_add_I f.nonneg_coeffs g.nonneg_coeffs) qed lemma nonneg_dirichlet_series_mult [intro]: assumes "nonneg_dirichlet_series f" "nonneg_dirichlet_series g" shows "nonneg_dirichlet_series (f * g)" proof - interpret f: nonneg_dirichlet_series f by fact interpret g: nonneg_dirichlet_series g by fact show ?thesis by standard (auto intro!: nonneg_Reals_sum nonneg_Reals_mult_I f.nonneg_coeffs g.nonneg_coeffs simp: fds_nth_mult dirichlet_prod_def) qed lemma nonneg_dirichlet_series_power [intro]: assumes "nonneg_dirichlet_series f" shows "nonneg_dirichlet_series (f ^ n)" using assms by (induction n) auto context nonneg_dirichlet_series begin lemma nonneg_exp [intro]: "nonneg_dirichlet_series (fds_exp f)" proof fix n :: nat assume "n > 0" define c where "c = exp (fds_nth f (Suc 0))" define f' where "f' = fds (\n. if n = Suc 0 then 0 else fds_nth f n)" from nonneg_coeffs[of 1] obtain c' where "fds_nth f (Suc 0) = of_real c'" by (auto elim!: nonneg_Reals_cases) hence "c = of_real (exp c')" by (simp add: c_def exp_of_real) hence c: "c \ \\<^sub>\\<^sub>0" by simp have less: "n < 2 ^ k" if "n < k" for k proof - have "n < k" by fact - also have "\ < 2 ^ k" by (induction k) auto + also have "\ < 2 ^ k" + by (rule less_exp) finally show ?thesis . qed have nonneg_power: "fds_nth (f' ^ k) n \ \\<^sub>\\<^sub>0" for k proof - have "nonneg_dirichlet_series f'" by standard (insert nonneg_coeffs, auto simp: f'_def) interpret nonneg_dirichlet_series "f' ^ k" by (intro nonneg_dirichlet_series_power) fact+ from nonneg_coeffs[of n] show ?thesis . qed hence "fds_nth (fds_exp f) n = c * (\k. fds_nth (f' ^ k) n /\<^sub>R fact k)" by (simp add: fds_exp_def fds_nth_fds' f'_def c_def) also have "(\k. fds_nth (f' ^ k) n /\<^sub>R fact k) = (\k\n. fds_nth (f' ^ k) n /\<^sub>R fact k)" by (intro suminf_finite) (auto intro!: fds_nth_power_eq_0 less simp: f'_def not_le) also have "c * \ \ \\<^sub>\\<^sub>0" unfolding scaleR_conv_of_real by (intro nonneg_Reals_mult_I nonneg_Reals_sum nonneg_power, unfold nonneg_Reals_of_real_iff ) (auto simp: c) finally show "fds_nth (fds_exp f) n \ \\<^sub>\\<^sub>0" . qed end lemma nonneg_dirichlet_series_lnD: assumes "nonneg_dirichlet_series (fds_ln l f)" "exp l = fds_nth f (Suc 0)" shows "nonneg_dirichlet_series f" proof - from assms have "nonneg_dirichlet_series (fds_exp (fds_ln l f))" by (intro nonneg_dirichlet_series.nonneg_exp) thus ?thesis using assms by simp qed context nonneg_dirichlet_series begin lemma fds_of_real_norm: "fds_of_real (fds_norm f) = f" proof (rule fds_eqI) fix n :: nat assume n: "n > 0" show "fds_nth (fds_of_real (fds_norm f)) n = fds_nth f n" using nonneg_coeffs[of n] by (auto elim!: nonneg_Reals_cases) qed end (* TODO: Simplify proof *) lemma pringsheim_landau_aux: fixes c :: real and f :: "complex fds" assumes "nonneg_dirichlet_series f" assumes abscissa: "c \ abs_conv_abscissa f" assumes g: "\s. s \ A \ Re s > c \ g s = eval_fds f s" assumes "g holomorphic_on A" "open A" "c \ A" shows "\x. x < c \ fds_abs_converges f (of_real x)" proof - interpret nonneg_dirichlet_series f by fact define a where "a = 1 + c" define g' where "g' = (\s. if s \ {s. Re s > c} then eval_fds f s else g s)" \ \We can find some $\varepsilon > 0$ such that the Dirichlet series can be continued analytically in a ball of radius $1 + \varepsilon$ around $a$.\ from \open A\ \c \ A\ obtain \ where \: "\ > 0" "ball c \ \ A" by (auto simp: open_contains_ball) define \ where "\ = sqrt (1 + \^2) - 1" from \ have \: "\ > 0" by (simp add: \_def) have ball_a_subset: "ball a (1 + \) \ {s. Re s > c} \ A" proof (intro subsetI) fix s :: complex assume s: "s \ ball a (1 + \)" define x y where "x = Re s" and "y = Im s" have [simp]: "s = x + \ * y" by (simp add: complex_eq_iff x_def y_def) show "s \ {s. Re s > c} \ A" proof (cases "Re s \ c") case True hence "(c - x)\<^sup>2 + y\<^sup>2 \ (1 + c - x)\<^sup>2 + y\<^sup>2 - 1" by (simp add: power2_eq_square algebra_simps) also from s have "(1 + c - x)\<^sup>2 + y\<^sup>2 - 1 < \\<^sup>2" by (auto simp: dist_norm cmod_def a_def \_def) finally have "sqrt ((c - x)\<^sup>2 + y\<^sup>2) < \" using \ by (intro real_less_lsqrt) auto hence "s \ ball c \" by (auto simp: dist_norm cmod_def) also have "\ \ A" by fact finally show ?thesis .. qed auto qed have holo: "g' holomorphic_on ball a (1 + \)" unfolding g'_def proof (intro holomorphic_on_subset[OF _ ball_a_subset] holomorphic_on_If_Un) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ ereal c" by fact finally have*: "conv_abscissa f \ ereal c" . show "eval_fds f holomorphic_on {s. c < Re s}" by (intro holomorphic_intros) (auto intro: le_less_trans[OF *]) qed (insert assms, auto intro!: holomorphic_intros open_halfspace_Re_gt) define f' where "f' = fds_norm f" have f_f': "f = fds_of_real f'" by (simp add: f'_def fds_of_real_norm) have f'_nonneg: "fds_nth f' n \ 0" for n using nonneg_coeffs[of n] by (auto elim!: nonneg_Reals_cases simp: f'_def) have deriv: "(\n. (deriv ^^ n) g' a) = (\n. eval_fds ((fds_deriv ^^ n) f) a)" proof fix n :: nat have ev: "eventually (\s. s \ {s. Re s > c}) (nhds (complex_of_real a))" by (intro eventually_nhds_in_open open_halfspace_Re_gt) (auto simp: a_def) have "(deriv ^^ n) g' a = (deriv ^^ n) (eval_fds f) a" by (intro higher_deriv_cong_ev refl eventually_mono[OF ev]) (auto simp: g'_def) also have "\ = eval_fds ((fds_deriv ^^ n) f) a" proof (intro eval_fds_higher_deriv [symmetric]) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ ereal c" by (rule assms) also have "\ < a" by (simp add: a_def) finally show "conv_abscissa f < ereal (complex_of_real a \ 1)" by simp qed finally show "(deriv ^^ n) g' a = eval_fds ((fds_deriv ^^ n) f) a" . qed have nth_deriv_conv: "fds_abs_converges ((fds_deriv ^^ n) f) (of_real a)" for n by (intro fds_abs_converges) (auto simp: abs_conv_abscissa_higher_deriv a_def intro!: le_less_trans[OF abscissa]) have nth_deriv_eq: "(fds_deriv ^^ n) f = fds (\k. (-1) ^ n * fds_nth f k * ln (real k) ^ n)" for n proof - have "fds_nth ((fds_deriv ^^ n) f) k = (-1) ^ n * fds_nth f k * ln (real k) ^ n" for k by (induction n) (simp_all add: fds_deriv_def fds_eq_iff fds_nth_fds' scaleR_conv_of_real) thus ?thesis by (intro fds_eqI) simp_all qed have deriv': "(\n. eval_fds ((fds_deriv ^^ n) f) (complex_of_real a)) = (\n. (- 1) ^ n * complex_of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / real k powr a))" proof fix n have "eval_fds ((fds_deriv ^^ n) f) (of_real a) = (\\<^sub>ak. fds_nth ((fds_deriv ^^ n) f) k / of_nat k powr complex_of_real a)" using nth_deriv_conv by (subst eval_fds_altdef) auto hence "eval_fds ((fds_deriv ^^ n) f) (of_real a) = (\\<^sub>ak. (- 1) ^ n *\<^sub>R (fds_nth f k * ln (real k) ^ n / k powr a))" by (simp add: nth_deriv_eq fds_nth_fds' powr_Reals_eq scaleR_conv_of_real algebra_simps) also have "\ = (- 1) ^ n * (\\<^sub>ak. of_real (fds_nth f' k * ln (real k) ^ n / k powr a))" by (subst infsetsum_scaleR_right) (simp_all add: scaleR_conv_of_real f_f') also have "\ = (- 1) ^ n * of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / k powr a)" by (subst infsetsum_of_real) (rule refl) finally show "eval_fds ((fds_deriv ^^ n) f) (complex_of_real a) = (- 1) ^ n * complex_of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / real k powr a)" . qed define s :: complex where "s = c - \ / 2" have s: "Re s < c" using assms \ by (simp_all add: s_def \_def field_simps) have "s \ ball a (1 + \)" using s by (simp add: a_def dist_norm cmod_def s_def) from holomorphic_power_series[OF holo this] have sums: "(\n. (deriv ^^ n) g' a / fact n * (s - a) ^ n) sums g' s" by simp also note deriv also have "s - a = -of_real (1 + \ / 2)" by (simp add: s_def a_def) also have "(\n. \ ^ n) = (\n. of_real ((-1) ^ n * (1 + \ / 2) ^ n))" by (intro ext) (subst power_minus, auto) also have "(\n. eval_fds ((fds_deriv ^^ n) f) a / fact n * \ n) = (\n. of_real ((-1) ^ n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2) ^ n))" using nth_deriv_conv by (simp add: f_f' fds_abs_converges_imp_converges mult_ac) finally have "summable \" by (simp add: sums_iff) hence summable: "summable (\n. (-1)^n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2)^n)" by (subst (asm) summable_of_real_iff) have "(\(n,k). (-1)^n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s-a) ^ n / fact n)) abs_summable_on (UNIV \ UNIV)" proof (subst abs_summable_on_Sigma_iff, safe, goal_cases) case (3 n) from nth_deriv_conv[of n] show ?case unfolding fds_abs_converges_altdef' by (intro abs_summable_on_cmult_left) (simp add: nth_deriv_eq fds_nth_fds' powr_Reals_eq) next case 4 have nth_deriv_f_f': "(fds_deriv ^^ n) f = fds_of_real ((fds_deriv ^^ n) f')" for n by (induction n) (auto simp: f'_def fds_of_real_norm) have norm_nth_deriv_f: "norm (fds_nth ((fds_deriv ^^ n) f) k) = (-1) ^ n * of_real (fds_nth ((fds_deriv ^^ n) f') k)" for n k proof (induction n) case (Suc n) thus ?case by (cases k) (auto simp: f_f' fds_nth_deriv scaleR_conv_of_real norm_mult) qed (auto simp: f'_nonneg f_f') note summable also have "(\n. (-1)^n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2)^n) = (\n. \\<^sub>ak. norm ((- 1) ^ n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s - a) ^ n / fact n)))" (is "_ = ?h") proof (rule ext, goal_cases) case (1 n) have "(\\<^sub>ak. norm ((- 1) ^ n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s - a) ^ n / fact n))) = (norm ((s - a) ^ n / fact n) * (-1) ^ n) *\<^sub>R (\\<^sub>ak. (-1) ^ n * norm (fds_nth ((fds_deriv ^^ n) f) k / real k powr a))" (is "_ = _ *\<^sub>R ?S") by (subst infsetsum_scaleR_right [symmetric]) (auto simp: norm_mult norm_divide norm_power mult_ac nth_deriv_eq fds_nth_fds') also have "?S = (\\<^sub>ak. fds_nth ((fds_deriv ^^ n) f') k / real k powr a)" by (intro infsetsum_cong) (auto simp: norm_mult norm_divide norm_power norm_nth_deriv_f) also have "\ = eval_fds ((fds_deriv ^^ n) f') a" using nth_deriv_conv[of n] by (subst eval_fds_altdef) (auto simp: f'_def nth_deriv_f_f') also have "(norm ((s - a) ^ n / fact n) * (- 1) ^ n) *\<^sub>R eval_fds ((fds_deriv ^^ n) f') a = (-1) ^ n * eval_fds ((fds_deriv ^^ n) f') a / fact n * norm (s - a) ^ n" by (simp add: norm_divide norm_power) also have s_a: "s - a = -of_real (1 + \ / 2)" by (simp add: s_def a_def) have "norm (s - a) = 1 + \ / 2" unfolding s_a norm_minus_cancel norm_of_real using \ by simp finally show ?case .. qed also have "?h n \ 0" for n by (intro infsetsum_nonneg) auto hence "?h = (\n. norm (?h n))" by simp finally show ?case unfolding abs_summable_on_nat_iff' . qed auto hence "(\(k,n). (-1)^n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s-a) ^ n / fact n)) abs_summable_on (UNIV \ UNIV)" by (subst (asm) abs_summable_on_Times_swap) (simp add: case_prod_unfold) hence "(\k. \\<^sub>an. (- 1) ^ n * fds_nth f k * ln (real k) ^ n / (k powr a) * ((s - a) ^ n / fact n)) abs_summable_on UNIV" (is "?h abs_summable_on _") by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\k. fds_nth f k / nat_power k s) abs_summable_on UNIV" proof (intro abs_summable_on_cong refl, goal_cases) case (1 k) have "?h k = (fds_nth f' k / k powr a) *\<^sub>R (\\<^sub>an. (-ln (real k) * (s - a)) ^ n / fact n)" by (subst infsetsum_scaleR_right [symmetric], rule infsetsum_cong) (simp_all add: scaleR_conv_of_real f_f' power_minus' power_mult_distrib divide_simps) also have "(\\<^sub>an. (-ln (real k) * (s - a)) ^ n / fact n) = exp (-ln (real k) * (s - a))" using exp_converges[of "-ln k * (s - a)"] exp_converges[of "norm (-ln k * (s - a))"] by (subst infsetsum_nat') (auto simp: abs_summable_on_nat_iff' sums_iff scaleR_conv_of_real divide_simps norm_divide norm_mult norm_power) also have "(fds_nth f' k / k powr a) *\<^sub>R \ = fds_nth f k / nat_power k s" by (auto simp: scaleR_conv_of_real f_f' powr_def exp_minus field_simps exp_of_real [symmetric] exp_diff) finally show ?case . qed finally have "fds_abs_converges f s" by (simp add: fds_abs_converges_def abs_summable_on_nat_iff') thus ?thesis by (intro exI[of _ "(c - \ / 2)"]) (auto simp: s_def a_def \) qed theorem pringsheim_landau: fixes c :: real and f :: "complex fds" assumes "nonneg_dirichlet_series f" assumes abscissa: "abs_conv_abscissa f = c" assumes g: "\s. s \ A \ Re s > c \ g s = eval_fds f s" assumes "g holomorphic_on A" "open A" "c \ A" shows False proof - have "\x complex_of_real x \ 1" unfolding abs_conv_abscissa_def by (intro Inf_lower) (auto simp: image_iff intro!: exI[of _ "of_real x"]) also have "\ < abs_conv_abscissa f" using assms x by simp finally show False by simp qed corollary entire_continuation_imp_abs_conv_abscissa_MInfty: assumes "nonneg_dirichlet_series f" assumes c: "c \ abs_conv_abscissa f" assumes g: "\s. Re s > c \ g s = eval_fds f s" assumes holo: "g holomorphic_on UNIV" shows "abs_conv_abscissa f = -\" proof (rule ccontr) assume "abs_conv_abscissa f \ -\" with c obtain a where abscissa [simp]: "abs_conv_abscissa f = ereal a" by (cases "abs_conv_abscissa f") auto show False proof (rule pringsheim_landau[OF assms(1) abscissa _ holo]) fix s assume s: "Re s > a" show "g s = eval_fds f s" proof (rule sym, rule analytic_continuation_open[of _ _ _ g]) show "g holomorphic_on {s. Re s > a}" by (rule holomorphic_on_subset[OF holo]) auto from assms show "{s. Re s > c} \ {s. Re s > a}" by auto next have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ = ereal a" by simp finally show "eval_fds f holomorphic_on {s. Re s > a}" by (intro holomorphic_intros) (auto intro: le_less_trans) qed (insert assms s, auto intro!: exI[of _ "of_real (c + 1)"] open_halfspace_Re_gt convex_connected convex_halfspace_Re_gt) qed auto qed subsection \Convergence of the $\zeta$ and M\"{o}bius $\mu$ series\ lemma fds_abs_summable_zeta_real_iff [simp]: "fds_abs_converges fds_zeta s \ s > (1 :: real)" proof - have "fds_abs_converges fds_zeta s \ summable (\n. real n powr -s)" unfolding fds_abs_converges_def by (intro summable_cong always_eventually) (auto simp: fds_nth_zeta powr_minus divide_simps) also have "\ \ s > 1" by (simp add: summable_real_powr_iff) finally show ?thesis . qed lemma fds_abs_summable_zeta_real: "s > (1 :: real) \ fds_abs_converges fds_zeta s" by simp lemma fds_abs_converges_moebius_mu_real: assumes "s > (1 :: real)" shows "fds_abs_converges (fds moebius_mu) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (norm (fds_nth (fds moebius_mu) n / nat_power n s)) \ n powr (-s)" by (simp add: powr_minus divide_simps abs_moebius_mu_le) next from assms show "summable (\n. real n powr -s)" by (simp add: summable_real_powr_iff) qed subsection \Application to the M\"{o}bius $\mu$ function\ lemma inverse_squares_sums': "(\n. 1 / real n ^ 2) sums (pi ^ 2 / 6)" using inverse_squares_sums sums_Suc_iff[of "\n. 1 / real n ^ 2" "pi^2 / 6"] by simp lemma norm_summable_moebius_over_square: "summable (\n. norm (moebius_mu n / real n ^ 2))" proof (subst summable_Suc_iff [symmetric], rule summable_comparison_test) show "summable (\n. 1 / real (Suc n) ^ 2)" using inverse_squares_sums by (simp add: sums_iff) qed (auto simp del: of_nat_Suc simp: field_simps abs_moebius_mu_le) lemma summable_moebius_over_square: "summable (\n. moebius_mu n / real n ^ 2)" using norm_summable_moebius_over_square by (rule summable_norm_cancel) lemma moebius_over_square_sums: "(\n. moebius_mu n / n\<^sup>2) sums (6 / pi\<^sup>2)" proof - have "1 = eval_fds (1 :: real fds) 2" by simp also have "(1 :: real fds) = fds_zeta * fds moebius_mu" by (rule fds_zeta_times_moebius_mu [symmetric]) also have "eval_fds \ 2 = eval_fds fds_zeta 2 * eval_fds (fds moebius_mu) 2" by (intro eval_fds_mult fds_abs_converges_moebius_mu_real) simp_all also have "\ = pi ^ 2 / 6 * (\n. moebius_mu n / (real n)\<^sup>2)" using inverse_squares_sums' by (simp add: eval_fds_at_numeral suminf_fds_zeta_aux sums_iff) finally have "(\n. moebius_mu n / (real n)\<^sup>2) = 6 / pi ^ 2" by (simp add: field_simps) with summable_moebius_over_square show ?thesis by (simp add: sums_iff) qed end diff --git a/thys/Ergodic_Theory/Kingman.thy b/thys/Ergodic_Theory/Kingman.thy --- a/thys/Ergodic_Theory/Kingman.thy +++ b/thys/Ergodic_Theory/Kingman.thy @@ -1,2106 +1,2121 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Subcocycles, subadditive ergodic theory\ theory Kingman imports Ergodicity Fekete begin text \Subadditive ergodic theory is the branch of ergodic theory devoted to the study of subadditive cocycles (named subcocycles in what follows), i.e., functions such that $u(n+m, x) \leq u(n, x) + u(m, T^n x)$ for all $x$ and $m,n$. For instance, Birkhoff sums are examples of such subadditive cocycles (in fact, they are additive), but more interesting examples are genuinely subadditive. The main result of the theory is Kingman's theorem, asserting the almost sure convergence of $u_n / n$ (this is a generalization of Birkhoff theorem). If the asymptotic average $\lim \int u_n / n$ (which exists by subadditivity and Fekete lemma) is not $-\infty$, then the convergence takes also place in $L^1$. We prove all this below. \ context mpt begin subsection \Definition and basic properties\ definition subcocycle::"(nat \ 'a \ real) \ bool" where "subcocycle u = ((\n. integrable M (u n)) \ (\n m x. u (n+m) x \ u n x + u m ((T^^n) x)))" lemma subcocycle_ineq: assumes "subcocycle u" shows "u (n+m) x \ u n x + u m ((T^^n) x)" using assms unfolding subcocycle_def by blast lemma subcocycle_0_nonneg: assumes "subcocycle u" shows "u 0 x \ 0" proof - have "u (0+0) x \ u 0 x + u 0 ((T^^0) x)" using assms unfolding subcocycle_def by blast then show ?thesis by auto qed lemma subcocycle_integrable: assumes "subcocycle u" shows "integrable M (u n)" "u n \ borel_measurable M" using assms unfolding subcocycle_def by auto lemma subcocycle_birkhoff: assumes "integrable M f" shows "subcocycle (birkhoff_sum f)" unfolding subcocycle_def by (auto simp add: assms birkhoff_sum_integral(1) birkhoff_sum_cocycle) text \The set of subcocycles is stable under addition, multiplication by positive numbers, and $\max$.\ lemma subcocycle_add: assumes "subcocycle u" "subcocycle v" shows "subcocycle (\n x. u n x + v n x)" unfolding subcocycle_def proof (auto) fix n show "integrable M (\x. u n x + v n x)" using assms unfolding subcocycle_def by simp next fix n m x have "u (n+m) x \ u n x + u m ((T ^^ n) x)" using assms(1) subcocycle_def by simp moreover have "v (n+m) x \ v n x + v m ((T ^^ n) x)" using assms(2) subcocycle_def by simp ultimately show "u (n + m) x + v (n + m) x \ u n x + v n x + (u m ((T ^^ n) x) + v m ((T ^^ n) x))" by simp qed lemma subcocycle_cmult: assumes "subcocycle u" "c \ 0" shows "subcocycle (\n x. c * u n x)" using assms unfolding subcocycle_def by (auto, metis distrib_left mult_left_mono) lemma subcocycle_max: assumes "subcocycle u" "subcocycle v" shows "subcocycle (\n x. max (u n x) (v n x))" unfolding subcocycle_def proof (auto) fix n show "integrable M (\x. max (u n x) (v n x))" using assms unfolding subcocycle_def by auto next fix n m x have "u (n+m) x \ u n x + u m ((T^^n) x)" using assms(1) unfolding subcocycle_def by auto then show "u (n + m) x \ max (u n x) (v n x) + max (u m ((T ^^ n) x)) (v m ((T ^^ n) x))" by simp next fix n m x have "v (n+m) x \ v n x + v m ((T^^n) x)" using assms(2) unfolding subcocycle_def by auto then show "v (n + m) x \ max (u n x) (v n x) + max (u m ((T ^^ n) x)) (v m ((T ^^ n) x))" by simp qed text \Applying inductively the subcocycle equation, it follows that a subcocycle is bounded by the Birkhoff sum of the subcocycle at time $1$.\ lemma subcocycle_bounded_by_birkhoff1: assumes "subcocycle u" "n > 0" shows "u n x \ birkhoff_sum (u 1) n x" using \n > 0\ proof (induction rule: ind_from_1) case 1 show ?case by auto next case (Suc p) have "u (Suc p) x \ u p x + u 1 ((T^^p)x)" using assms(1) subcocycle_def by (metis Suc_eq_plus1) then show ?case using Suc birkhoff_sum_cocycle[where ?n = p and ?m = 1] \ p>0 \ by (simp add: birkhoff_sum_def) qed text \It is often important to bound a cocycle $u_n(x)$ by the Birkhoff sums of $u_N/N$. Compared to the trivial upper bound for $u_1$, there are additional boundary errors that make the estimate more cumbersome (but these terms only come from a $N$-neighborhood of $0$ and $n$, so they are negligible if $N$ is fixed and $n$ tends to infinity.\ lemma subcocycle_bounded_by_birkhoffN: assumes "subcocycle u" "n > 2*N" "N>0" shows "u n x \ birkhoff_sum (\x. u N x / real N) (n - 2 * N) x + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. \u 1 ((T ^^ (n - (2 * N - i))) x)\)" proof - have Iar: "u (a*N+r) x \ u r x + (\i u(a*N+r) x + u N ((T^^(a*N+r))x)" using assms(1) unfolding subcocycle_def by auto also have "... \ u r x + (\ii (\i0" for a using that proof (induction rule: ind_from_1) case 1 show ?case by auto next case (Suc a) have "u ((a+1)*N) x = u((a*N) + N) x" by (simp add: semiring_normalization_rules(2) semiring_normalization_rules(23)) also have "... \ u(a*N) x + u N ((T^^(a*N))x)" using assms(1) unfolding subcocycle_def by auto also have "... \ (\iiii<2*N. abs(u 1 ((T^^(n-(2*N-i))) x)))" have "E2 \ 0" unfolding E2_def by auto obtain a0 s0 where 0: "s0 < N" "n = a0 * N + s0" using \0 < N\ mod_div_decomp mod_less_divisor by blast then have "a0 \ 1" using \n > 2 * N\ \N>0\ by (metis Nat.add_0_right add.commute add_lessD1 add_mult_distrib comm_monoid_mult_class.mult_1 eq_imp_le less_imp_add_positive less_imp_le_nat less_one linorder_neqE_nat mult.left_neutral mult_not_zero not_add_less1 one_add_one) define a s where "a = a0-1" and "s = s0+N" then have as: "n = a * N + s" unfolding a_def s_def using \a0 \ 1\ 0 by (simp add: mult_eq_if) have s: "s \ N" "s < 2*N" using 0 unfolding s_def by auto have a: "a*N > n - 2*N" "a*N \ n - N" using as s \n > 2*N\ by auto then have "(a*N - (n-2*N)) \ N" using \n > 2*N\ by auto have "a*N \ n - 2*N" using a by simp { fix r::nat assume "r < N" have "a*N+r > n - 2*N" using \n>2*N\ as s by auto define tr where "tr = n-(a*N+r)" have "tr > 0" unfolding tr_def using as s \r by auto then have *: "n = (a*N+r) + tr" unfolding tr_def by auto have "birkhoff_sum (u 1) tr ((T^^(a*N+r))x) = (\ii\{a*N+r.. (\i\{a*N+r.. (\i\{n-2*N..r tr_def by auto also have "... = E2" unfolding E2_def apply (rule sum.reindex_bij_betw[symmetric], rule bij_betw_byWitness[where ?f' = "\i. i - (n-2*N)"]) using \n > 2*N\ by auto finally have A: "birkhoff_sum (u 1) tr ((T^^(a*N+r))x) \ E2" by simp have "u n x \ u (a*N+r) x + u tr ((T^^(a*N+r))x)" using assms(1) * unfolding subcocycle_def by auto also have "... \ u (a*N+r) x + birkhoff_sum (u 1) tr ((T^^(a*N+r))x)" using subcocycle_bounded_by_birkhoff1[OF assms(1)] \tr > 0\ by auto finally have B: "u n x \ u (a*N+r) x + E2" using A by auto have "u (a*N+r) x \ (\ii0" using \a*N+r > n - 2*N\ not_less by fastforce have "u(a*N+r) x \ (\ia>0\] True by auto moreover have "0 \ (\i u r x + (\i (\i (\i (\ir by auto finally show ?thesis using I by auto qed then have "u n x \ E1 + (\i E2" if "j\{n-2*N.. (\iN>0\] unfolding birkhoff_sum_def by auto also have "... = (\i (\ik\{j.. (\i\{n-2*N..j\{n-2*N.. \a*N \ n - N\ by auto also have "... = E2" unfolding E2_def apply (rule sum.reindex_bij_betw[symmetric], rule bij_betw_byWitness[where ?f' = "\i. i - (n-2*N)"]) using \n > 2*N\ by auto finally show ?thesis by auto qed have "(\jjj\{n-2*N..a*N \ n - 2*N\, of 0 "\j. u N ((T^^j) x)", symmetric] atLeast0LessThan by simp also have "... \ (\j\{n-2*N.. N * E2" using \(a*N - (n-2*N)) \ N\ \E2 \ 0\ by (simp add: mult_right_mono) finally have J: "(\j (\jr (\rirrij N *(E1+E2) + (\jjN>0\ by (simp add: sum_distrib_left) also have "... = N*(E1 + E2 + (\j E1 + 2*E2 + birkhoff_sum (\x. u N x / N) (n-2*N) x" unfolding birkhoff_sum_def using \N>0\ by auto then show ?thesis unfolding E1_def E2_def by auto qed text \Many natural cocycles are only defined almost everywhere, and then the subadditivity property only makes sense almost everywhere. We will now show that such an a.e.-subcocycle coincides almost everywhere with a genuine subcocycle in the above sense. Then, all the results for subcocycles will apply to such a.e.-subcocycles. (Usually, in ergodic theory, subcocycles only satisfy the subadditivity property almost everywhere, but we have requested it everywhere for simplicity in the proofs.) The subcocycle will be defined in a recursive way. This means that is can not be defined in a proof (since complicated function definitions are not available inside proofs). Since it is defined in terms of $u$, then $u$ has to be available at the top level, which is most conveniently done using a context. \ context fixes u::"nat \ 'a \ real" assumes H: "\m n. AE x in M. u (n+m) x \ u n x + u m ((T^^n) x)" "\n. integrable M (u n)" begin private fun v :: "nat \ 'a \ real" where "v n x = ( if n = 0 then max (u 0 x) 0 else if n = 1 then u 1 x else min (u n x) (Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<..v 0 x = max (u 0 x) 0\ + by simp + +private lemma v1 [simp]: + \v (Suc 0) x = u 1 x\ + by simp + +private lemma v2 [simp]: + \v n x = min (u n x) (Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<.. if \n \ 2\ + using that by (subst v.simps) (simp del: v.simps) + +declare v.simps [simp del] + private lemma integrable_v: "integrable M (v n)" for n proof (induction n rule: nat_less_induct) case (1 n) consider "n = 0" | "n = 1" | "n>1" by linarith then show ?case proof (cases) assume "n = 0" have "v 0 x = max (u 0 x) 0" for x by simp then show ?thesis using integrable_max[OF H(2)[of 0]] \n = 0\ by auto next assume "n = 1" have "v 1 x = u 1 x" for x by simp then show ?thesis using H(2)[of 1] \n = 1\ by auto next assume "n > 1" hence "v n x = min (u n x) (MIN k \ {0<..x. min (u n x) (MIN k \ {0<..n >1\ apply auto[1] apply (rule Bochner_Integration.integrable_add) using "1.IH" apply auto[1] apply (rule Tn_integral_preserving(1)) using "1.IH" by (metis \1 < n\ diff_less greaterThanLessThan_iff max_0_1(2) max_less_iff_conj) ultimately show ?case by auto qed qed private lemma u_eq_v_AE: "AE x in M. v n x = u n x" for n proof (induction n rule: nat_less_induct) case (1 n) consider "n = 0" | "n = 1" | "n>1" by linarith then show ?case proof (cases) assume "n = 0" have "AE x in M. u 0 x \ u 0 x + u 0 x" using H(1)[of 0 0] by auto then have "AE x in M. u 0 x \ 0" by auto moreover have "v 0 x = max (u 0 x) 0" for x by simp ultimately show ?thesis using \n = 0\ by auto next assume "n = 1" have "v 1 x = u 1 x" for x by simp then show ?thesis using \n = 1\ by simp next assume "n > 1" { fix k assume "ks. v k ((T^^s) x) = u k ((T^^s) x)" by simp } note * = this have "AE x in M. \k \ {..s. v k ((T^^s) x) = u k ((T^^s) x)" apply (rule AE_finite_allI) using * by simp_all moreover have "AE x in M. \i j. u (i+j) x \ u i x + u j ((T^^i) x)" apply (subst AE_all_countable, intro allI)+ using H(1) by simp moreover { fix x assume "\k \ {..s. v k ((T^^s) x) = u k ((T^^s) x)" "\i j. u (i+j) x \ u i x + u j ((T^^i) x)" then have Hx: "\k s. k < n \ v k ((T^^s) x) = u k ((T^^s) x)" "\i j. u (i+j) x \ u i x + u j ((T^^i) x)" by auto { fix k assume "k \ {0<.. u k x + u (n-k) ((T^^k) x)" using Hx(2) K by (metis le_add_diff_inverse less_imp_le_nat) also have "... = v k x + v (n-k) ((T^^k)x)" using Hx(1)[OF \k , of 0] Hx(1)[OF \n-k , of k] by auto finally have "u n x \ v k x + v (n-k) ((T^^k)x)" by simp } then have *: "\z. z \ (\k. v k x + v (n-k) ((T^^k) x))`{0<.. u n x \ z" by blast have "u n x \ Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<..n>1\ * by auto moreover have "v n x = min (u n x) (Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<..1 by auto ultimately have "v n x = u n x" by auto } ultimately show ?thesis by auto qed qed private lemma subcocycle_v: "v (n+m) x \ v n x + v m ((T^^n) x)" proof - consider "n = 0" | "m = 0" | "n>0 \ m >0" by auto then show ?thesis proof (cases) case 1 then have "v n x \ 0" by simp then show ?thesis using \n = 0\ by auto next case 2 then have "v m x \ 0" by simp then show ?thesis using \m = 0\ by auto next case 3 then have "n+m > 1" by simp then have "v (n+m) x = min (u(n+m) x) (Min ((\k. v k x + v ((n+m)-k) ((T^^k) x))`{0<.. Min ((\k. v k x + v ((n+m)-k) ((T^^k) x))`{0<.. v n x + v ((n+m)-n) ((T^^n) x)" apply (rule Min_le, simp) by (metis (lifting) \0 < n \ 0 < m\ add.commute greaterThanLessThan_iff image_iff less_add_same_cancel2) finally show ?thesis by simp qed qed lemma subcocycle_AE_in_context: "\w. subcocycle w \ (AE x in M. \n. w n x = u n x)" proof - have "subcocycle v" using subcocycle_v integrable_v unfolding subcocycle_def by auto moreover have "AE x in M. \n. v n x = u n x" by (subst AE_all_countable, intro allI, rule u_eq_v_AE) ultimately show ?thesis by blast qed end lemma subcocycle_AE: fixes u::"nat \ 'a \ real" assumes "\m n. AE x in M. u (n+m) x \ u n x + u m ((T^^n) x)" "\n. integrable M (u n)" shows "\w. subcocycle w \ (AE x in M. \n. w n x = u n x)" using subcocycle_AE_in_context assms by blast subsection \The asymptotic average\ text \In this subsection, we define the asymptotic average of a subcocycle $u$, i.e., the limit of $\int u_n(x)/n$ (the convergence follows from subadditivity of $\int u_n$) and study its basic properties, especially in terms of operations on subcocycles. In general, it can be $-\infty$, so we define it in the extended reals.\ definition subcocycle_avg_ereal::"(nat \ 'a \ real) \ ereal" where "subcocycle_avg_ereal u = Inf {ereal((\x. u n x \M) / n) |n. n > 0}" lemma subcocycle_avg_finite: "subcocycle_avg_ereal u < \" unfolding subcocycle_avg_ereal_def using Inf_less_iff less_ereal.simps(4) by blast lemma subcocycle_avg_subadditive: assumes "subcocycle u" shows "subadditive (\n. (\x. u n x \M))" unfolding subadditive_def proof (intro allI) have int_u [measurable]: "\n. integrable M (u n)" using assms unfolding subcocycle_def by auto fix m n have "(\x. u (n+m) x \M) \ (\x. u n x + u m ((T^^n) x) \M)" apply (rule integral_mono) using int_u apply (auto simp add: Tn_integral_preserving(1)) using assms unfolding subcocycle_def by auto also have "... \ (\x. u n x \M) + (\x. u m ((T^^n) x) \M)" using int_u by (auto simp add: Tn_integral_preserving(1)) also have "... = (\x. u n x \M) + (\x. u m x \M)" using int_u by (auto simp add: Tn_integral_preserving(2)) finally show "(\x. u (n+m) x \M) \ (\x. u n x \M) + (\x. u m x \M)" by simp qed lemma subcocycle_int_tendsto_avg_ereal: assumes "subcocycle u" shows "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal u" unfolding subcocycle_avg_ereal_def using subadditive_converges_ereal[OF subcocycle_avg_subadditive[OF assms]] by auto text \The average behaves well under addition, scalar multiplication and max, trivially.\ lemma subcocycle_avg_ereal_add: assumes "subcocycle u" "subcocycle v" shows "subcocycle_avg_ereal (\n x. u n x + v n x) = subcocycle_avg_ereal u + subcocycle_avg_ereal v" proof - have int [simp]: "\n. integrable M (u n)" "\n. integrable M (v n)" using assms unfolding subcocycle_def by auto { fix n have "(\x. u n x / n \M) + (\x. v n x / n \M) = (\x. u n x / n + v n x / n \M)" by (rule Bochner_Integration.integral_add[symmetric], auto) also have "... = (\x. (u n x + v n x) / n \M)" by (rule Bochner_Integration.integral_cong, auto simp add: add_divide_distrib) finally have "ereal (\x. u n x / n \M) + (\x. v n x / n \M) = (\x. (u n x + v n x) / n \M)" by auto } moreover have "(\n. ereal (\x. u n x / n \M) + (\x. v n x / n \M)) \ subcocycle_avg_ereal u + subcocycle_avg_ereal v" apply (intro tendsto_intros subcocycle_int_tendsto_avg_ereal[OF assms(1)] subcocycle_int_tendsto_avg_ereal[OF assms(2)]) using subcocycle_avg_finite by auto ultimately have "(\n. (\x. (u n x + v n x) / n \M)) \ subcocycle_avg_ereal u + subcocycle_avg_ereal v" by auto moreover have "(\n. (\x. (u n x + v n x) / n \M)) \ subcocycle_avg_ereal (\n x. u n x + v n x)" by (rule subcocycle_int_tendsto_avg_ereal[OF subcocycle_add[OF assms]]) ultimately show ?thesis using LIMSEQ_unique by blast qed lemma subcocycle_avg_ereal_cmult: assumes "subcocycle u" "c \ (0::real)" shows "subcocycle_avg_ereal (\n x. c * u n x) = c * subcocycle_avg_ereal u" proof (cases "c = 0") case True have *: "ereal (\x. (c * u n x) / n \M) = 0" if "n>0" for n by (subst True, auto) have "(\n. ereal (\x. (c * u n x) / n \M)) \ 0" by (subst lim_explicit, metis * less_le_trans zero_less_one) moreover have "(\n. ereal (\x. (c * u n x) / n \M)) \ subcocycle_avg_ereal (\n x. c * u n x)" using subcocycle_int_tendsto_avg_ereal[OF subcocycle_cmult[OF assms]] by auto ultimately have "subcocycle_avg_ereal (\n x. c * u n x) = 0" using LIMSEQ_unique by blast then show ?thesis using True by auto next case False have int: "\n. integrable M (u n)" using assms unfolding subcocycle_def by auto have "ereal (\x. c * u n x / n \M) = c * ereal (\x. u n x / n \M)" for n by auto then have "(\n. c * ereal (\x. u n x / n \M)) \ subcocycle_avg_ereal (\n x. c * u n x)" using subcocycle_int_tendsto_avg_ereal[OF subcocycle_cmult[OF assms]] by auto moreover have "(\n. c * ereal (\x. u n x / n \M)) \ c * subcocycle_avg_ereal u" apply (rule tendsto_mult_ereal) using False subcocycle_int_tendsto_avg_ereal[OF assms(1)] by auto ultimately show ?thesis using LIMSEQ_unique by blast qed lemma subcocycle_avg_ereal_max: assumes "subcocycle u" "subcocycle v" shows "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) \ max (subcocycle_avg_ereal u) (subcocycle_avg_ereal v)" proof (auto) have int: "integrable M (u n)" "integrable M (v n)" for n using assms unfolding subcocycle_def by auto have int2: "integrable M (\x. max (u n x) (v n x))" for n using integrable_max int by auto have "(\x. u n x / n \M) \ (\x. max (u n x) (v n x) / n \M)" for n apply (rule integral_mono) using int int2 by (auto simp add: divide_simps) then show "subcocycle_avg_ereal u \ subcocycle_avg_ereal (\n x. max (u n x) (v n x))" using LIMSEQ_le[OF subcocycle_int_tendsto_avg_ereal[OF assms(1)] subcocycle_int_tendsto_avg_ereal[OF subcocycle_max[OF assms]]] by auto have "(\x. v n x / n \M) \ (\x. max (u n x) (v n x) / n \M)" for n apply (rule integral_mono) using int int2 by (auto simp add: divide_simps) then show "subcocycle_avg_ereal v \ subcocycle_avg_ereal (\n x. max (u n x) (v n x))" using LIMSEQ_le[OF subcocycle_int_tendsto_avg_ereal[OF assms(2)] subcocycle_int_tendsto_avg_ereal[OF subcocycle_max[OF assms]]] by auto qed text \For a Birkhoff sum, the average at each time is the same, equal to the average of the function, so the asymptotic average is also equal to this common value.\ lemma subcocycle_avg_ereal_birkhoff: assumes "integrable M u" shows "subcocycle_avg_ereal (birkhoff_sum u) = (\x. u x \M)" proof - have *: "ereal (\x. (birkhoff_sum u n x) / n \M) = (\x. u x \M)" if "n>0" for n using birkhoff_sum_integral(2)[OF assms] that by auto have "(\n. ereal (\x. (birkhoff_sum u n x) / n \M)) \ (\x. u x \M)" by (subst lim_explicit, metis * less_le_trans zero_less_one) moreover have "(\n. ereal (\x. (birkhoff_sum u n x) / n \M)) \ subcocycle_avg_ereal (birkhoff_sum u)" using subcocycle_int_tendsto_avg_ereal[OF subcocycle_birkhoff[OF assms]] by auto ultimately show ?thesis using LIMSEQ_unique by blast qed text \In nice situations, where one can avoid the use of ereal, the following definition is more convenient. The kind of statements we are after is as follows: if the ereal average is finite, then something holds, likely involving the real average. In particular, we show in this setting what we have proved above under this new assumption: convergence (in real numbers) of the average to the asymptotic average, as well as good behavior under sum, scalar multiplication by positive numbers, max, formula for Birkhoff sums.\ definition subcocycle_avg::"(nat \ 'a \ real) \ real" where "subcocycle_avg u = real_of_ereal(subcocycle_avg_ereal u)" lemma subcocycle_avg_real_ereal: assumes "subcocycle_avg_ereal u > - \" shows "subcocycle_avg_ereal u = ereal(subcocycle_avg u)" unfolding subcocycle_avg_def using assms subcocycle_avg_finite[of u] ereal_real by auto lemma subcocycle_int_tendsto_avg: assumes "subcocycle u" "subcocycle_avg_ereal u > - \" shows "(\n. (\x. u n x / n \M)) \ subcocycle_avg u" using subcocycle_avg_real_ereal[OF assms(2)] subcocycle_int_tendsto_avg_ereal[OF assms(1)] by auto lemma subcocycle_avg_add: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > - \" "subcocycle_avg_ereal v > - \" shows "subcocycle_avg_ereal (\n x. u n x + v n x) > -\" "subcocycle_avg (\n x. u n x + v n x) = subcocycle_avg u + subcocycle_avg v" using assms subcocycle_avg_finite real_of_ereal_add unfolding subcocycle_avg_def subcocycle_avg_ereal_add[OF assms(1) assms(2)] by auto lemma subcocycle_avg_cmult: assumes "subcocycle u" "c \ (0::real)" "subcocycle_avg_ereal u > - \" shows "subcocycle_avg_ereal (\n x. c * u n x) > - \" "subcocycle_avg (\n x. c * u n x) = c * subcocycle_avg u" using assms subcocycle_avg_finite unfolding subcocycle_avg_def subcocycle_avg_ereal_cmult[OF assms(1) assms(2)] by auto lemma subcocycle_avg_max: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > - \" "subcocycle_avg_ereal v > - \" shows "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > -\" "subcocycle_avg (\n x. max (u n x) (v n x)) \ max (subcocycle_avg u) (subcocycle_avg v)" proof - show *: "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > - \" using assms(3) subcocycle_avg_ereal_max[OF assms(1) assms(2)] by auto have "ereal (subcocycle_avg (\n x. max (u n x) (v n x))) \ max (ereal(subcocycle_avg u)) (ereal(subcocycle_avg v))" using subcocycle_avg_real_ereal[OF assms(3)] subcocycle_avg_real_ereal[OF assms(4)] subcocycle_avg_real_ereal[OF *] subcocycle_avg_ereal_max[OF assms(1) assms(2)] by auto then show "subcocycle_avg (\n x. max (u n x) (v n x)) \ max (subcocycle_avg u) (subcocycle_avg v)" by auto qed lemma subcocycle_avg_birkhoff: assumes "integrable M u" shows "subcocycle_avg_ereal (birkhoff_sum u) > - \" "subcocycle_avg (birkhoff_sum u) = (\x. u x \M)" unfolding subcocycle_avg_def subcocycle_avg_ereal_birkhoff[OF assms(1)] by auto end subsection \Almost sure convergence of subcocycles\ text \In this paragraph, we prove Kingman's theorem, i.e., the almost sure convergence of subcocycles. Their limit is almost surely invariant. There is no really easy proof. The one we use below is arguably the simplest known one, due to Steele (1989). The idea is to show that the limsup of the subcocycle is bounded by the liminf (which is almost surely constant along trajectories), by using subadditivity along time intervals where the liminf is almost reached, of length at most $N$. For some points, the liminf takes a large time $>N$ to be reached. We neglect those times, introducing an additional error that gets smaller with $N$, thanks to Birkhoff ergodic theorem applied to the set of bad points. The error is most easily managed if the subcocycle is assumed to be nonpositive, which one can assume in a first step. The general case is reduced to this one by replacing $u_n$ with $u_n - S_n u_1 \leq 0$, and using Birkhoff theorem to control $S_n u_1$.\ context fmpt begin text \First, as explained above, we prove the theorem for nonpositive subcocycles.\ lemma kingman_theorem_AE_aux1: assumes "subcocycle u" "\x. u 1 x \ 0" shows "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" proof - define l where "l = (\x. liminf (\n. u n x / n))" have u_meas [measurable]: "\n. u n \ borel_measurable M" using assms(1) unfolding subcocycle_def by auto have l_meas [measurable]: "l \ borel_measurable M" unfolding l_def by auto { fix x assume *: "(\n. birkhoff_sum (u 1) n x / n) \ real_cond_exp M Invariants (u 1) x" then have "(\n. birkhoff_sum (u 1) n x / n) \ ereal(real_cond_exp M Invariants (u 1) x)" by auto then have a: "liminf (\n. birkhoff_sum (u 1) n x / n) = ereal(real_cond_exp M Invariants (u 1) x)" using lim_imp_Liminf by force have "ereal(u n x / n) \ ereal(birkhoff_sum (u 1) n x / n)" if "n>0" for n using subcocycle_bounded_by_birkhoff1[OF assms(1) that, of x] that by (simp add: divide_right_mono) with eventually_mono[OF eventually_gt_at_top[of 0] this] have "eventually (\n. ereal(u n x / n) \ ereal(birkhoff_sum (u 1) n x / n)) sequentially" by auto then have "liminf (\n. u n x / n) \ liminf (\n. birkhoff_sum (u 1) n x / n)" by (simp add: Liminf_mono) then have "l x < \" unfolding l_def using a by auto } then have "AE x in M. l x < \" using birkhoff_theorem_AE_nonergodic[of "u 1"] subcocycle_def assms(1) by auto have l_dec: "l x \ l (T x)" for x proof - have "l x = liminf (\n. ereal ((u (n+1) x)/(n+1)))" unfolding l_def by (rule liminf_shift[of "\n. ereal (u n x / real n)", symmetric]) also have "... \ liminf (\n. ereal((u 1 x)/(n+1)) + ereal((u n (T x))/(n+1)))" proof (rule Liminf_mono[OF eventuallyI]) fix n have "u (1+n) x \ u 1 x + u n ((T^^1) x)" using assms(1) unfolding subcocycle_def by blast then have "u (n+1) x \ u 1 x + u n (T x)" by auto then have "(u (n+1) x)/(n+1) \ (u 1 x)/(n+1) + (u n (T x))/(n+1)" by (metis add_divide_distrib divide_right_mono of_nat_0_le_iff) then show "ereal ((u (n+1) x)/(n+1)) \ ereal((u 1 x)/(n+1)) + ereal((u n (T x))/(n+1))" by auto qed also have "... = 0 + liminf(\n. ereal((u n (T x))/(n+1)))" proof (rule ereal_liminf_lim_add[of "\n. ereal((u 1 x)/real(n+1))" 0 "\n. ereal((u n (T x))/(n+1))"]) have "(\n. ereal((u 1 x)*(1/real(n+1)))) \ ereal((u 1 x) * 0)" by (intro tendsto_intros LIMSEQ_ignore_initial_segment) then show "(\n. ereal((u 1 x)/real(n+1))) \ 0" by (simp add: zero_ereal_def) qed (simp) also have "... = 1 * liminf(\n. ereal((u n (T x))/(n+1)))" by simp also have "... = liminf(\n. (n+1)/n * ereal((u n (T x))/(n+1)))" proof (rule ereal_liminf_lim_mult[symmetric]) have "real (n+1) / real n = 1 + 1/real n" if "n>0" for n by (simp add: divide_simps mult.commute that) with eventually_mono[OF eventually_gt_at_top[of "0::nat"] this] have "eventually (\n. real (n+1) / real n = 1 + 1/real n) sequentially" by simp moreover have "(\n. 1 + 1/real n) \ 1 + 0" by (intro tendsto_intros) ultimately have "(\n. real (n+1) / real n) \ 1" using Lim_transform_eventually by (simp add: filterlim_cong) then show "(\n. ereal(real (n+1) / real n)) \ 1" by (simp add: one_ereal_def) qed (auto) also have "... = l (T x)" unfolding l_def by auto finally show "l x \ l (T x)" by simp qed have "AE x in M. l (T x) = l x" apply (rule AE_increasing_then_invariant) using l_dec by auto then obtain g0 where g0: "g0 \ borel_measurable Invariants" "AE x in M. l x = g0 x" using Invariants_quasi_Invariants_functions[OF l_meas] by auto define g where "g = (\x. if g0 x = \ then 0 else g0 x)" have g: "g \ borel_measurable Invariants" "AE x in M. g x = l x" unfolding g_def using g0(1) \AE x in M. l x = g0 x\ \AE x in M. l x < \\ by auto have [measurable]: "g \ borel_measurable M" using g(1) Invariants_measurable_func by blast have "\x. g x < \" unfolding g_def by auto define A where "A = {x \ space M. l x < \ \ (\n. l ((T^^n) x) = g ((T^^n) x))}" have A_meas [measurable]: "A \ sets M" unfolding A_def by auto have "AE x in M. x \ A" unfolding A_def using T_AE_iterates[OF g(2)] \AE x in M. l x < \\ by auto then have "space M - A \ null_sets M" by (simp add: AE_iff_null set_diff_eq) have l_inv: "l((T^^n) x) = l x" if "x \ A" for x n proof - have "l((T^^n) x) = g((T^^n) x)" using \ x \ A \ unfolding A_def by blast also have "... = g x" using g(1) A_def Invariants_func_is_invariant_n that by blast also have "... = g((T^^0) x)" by auto also have "... = l((T^^0) x)" using \ x \ A \ unfolding A_def by (metis (mono_tags, lifting) mem_Collect_eq) finally show ?thesis by auto qed define F where "F = (\ K e x. real_of_ereal(max (l x) (-ereal K)) + e)" have F_meas [measurable]: "F K e \ borel_measurable M" for K e unfolding F_def by auto define B where "B = (\N K e. {x \ A. \n\{1..N}. u n x - n * F K e x < 0})" have B_meas [measurable]: "B N K e \ sets M" for N K e unfolding B_def by (measurable) define I where "I = (\N K e x. (indicator (- B N K e) x)::real)" have I_meas [measurable]: "I N K e \ borel_measurable M" for N K e unfolding I_def by auto have I_int: "integrable M (I N K e)" for N K e unfolding I_def apply (subst integrable_cong[where ?g = "indicator (space M - B N K e)::_ \ real"], auto) by (auto split: split_indicator simp: less_top[symmetric]) have main: "AE x in M. limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I N K e) x)" if "N>(1::nat)" "K>(0::real)" "e>(0::real)" for N K e proof - let ?B = "B N K e" and ?I = "I N K e" and ?F = "F K e" define t where "t = (\x. if x \ ?B then Min {n\{1..N}. u n x - n * ?F x < 0} else 1)" have [measurable]: "t \ measurable M (count_space UNIV)" unfolding t_def by measurable have t1: "t x \ {1..N}" for x proof (cases "x \ ?B") case False then have "t x = 1" by (simp add: t_def) then show ?thesis using \N>1\by auto next case True let ?A = "{n\{1..N}. u n x - n * ?F x < 0}" have "t x = Min ?A" using True by (simp add: t_def) moreover have "Min ?A \ ?A" apply (rule Min_in, simp) using True B_def by blast ultimately show ?thesis by auto qed have bound1: "u (t x) x \ t x * ?F x + birkhoff_sum ?I (t x) x * abs(?F x)" for x proof (cases "x \ ?B") case True let ?A = "{n\{1..N}. u n x - n * F K e x < 0}" have "t x = Min ?A" using True by (simp add: t_def) moreover have "Min ?A \ ?A" apply (rule Min_in, simp) using True B_def by blast ultimately have "u (t x) x \ (t x) * ?F x" by auto moreover have "0 \ birkhoff_sum ?I (t x) x * abs(?F x)" unfolding birkhoff_sum_def I_def by (simp add: sum_nonneg) ultimately show ?thesis by auto next case False then have "0 \ ?F x + ?I x * abs(?F x)" unfolding I_def by auto then have "u 1 x \ ?F x + ?I x * abs(?F x)" using assms(2)[of x] by auto moreover have "t x = 1" unfolding t_def using False by auto ultimately show ?thesis by auto qed define TB where "TB = (\x. (T^^(t x)) x)" have [measurable]: "TB \ measurable M M" unfolding TB_def by auto define S where "S = (\n x. (\i measurable M (count_space UNIV)" for n unfolding S_def by measurable have TB_pow: "(TB^^n) x = (T^^(S n x)) x" for n x unfolding S_def TB_def by (induction n, auto, metis (mono_tags, lifting) add.commute funpow_add o_apply) have uS: "u (S n x) x \ (S n x) * ?F x + birkhoff_sum ?I (S n x) x * abs(?F x)" if "x \ A" "n>0" for x n using \n > 0\ proof (induction rule: ind_from_1) case 1 show ?case unfolding S_def using bound1 by auto next case (Suc n) have *: "?F((TB^^n) x) = ?F x" apply (subst TB_pow) unfolding F_def using l_inv[OF \x\A\] by auto have **: "S n x + t ((TB^^n) x) = S (Suc n) x" unfolding S_def by auto have "u (S (Suc n) x) x = u (S n x + t((TB^^n) x)) x" unfolding S_def by auto also have "... \ u (S n x) x + u (t((TB^^n) x)) ((T^^(S n x)) x)" using assms(1) unfolding subcocycle_def by auto also have "... \ u (S n x) x + u (t((TB^^n) x)) ((TB^^n) x)" using TB_pow by auto also have "... \ (S n x) * ?F x + birkhoff_sum ?I (S n x) x * abs(?F x) + t ((TB^^n) x) * ?F ((TB^^n) x) + birkhoff_sum ?I (t ((TB^^n) x)) ((TB^^n) x) * abs(?F ((TB^^n) x))" using Suc bound1[of "((TB^^n) x)"] by auto also have "... = (S n x) * ?F x + birkhoff_sum ?I (S n x) x * abs(?F x) + t ((TB^^n) x) * ?F x + birkhoff_sum ?I (t ((TB^^n) x)) ((T^^(S n x)) x) * abs(?F x)" using * TB_pow by auto also have "... = (real(S n x) + t ((TB^^n) x)) * ?F x + (birkhoff_sum ?I (S n x) x + birkhoff_sum ?I (t ((TB^^n) x)) ((T^^(S n x)) x)) * abs(?F x)" by (simp add: mult.commute ring_class.ring_distribs(1)) also have "... = (S n x + t ((TB^^n) x)) * ?F x + (birkhoff_sum ?I (S n x) x + birkhoff_sum ?I (t ((TB^^n) x)) ((T^^(S n x)) x)) * abs(?F x)" by simp also have "... = (S (Suc n) x) * ?F x + birkhoff_sum ?I (S (Suc n) x) x * abs(?F x)" by (subst birkhoff_sum_cocycle[symmetric], subst **, subst **, simp) finally show ?case by simp qed have un: "u n x \ n * ?F x + N * abs(?F x) + birkhoff_sum ?I n x * abs(?F x)" if "x \ A" "n>N" for x n proof - let ?A = "{i. S i x > n}" let ?iA = "Inf ?A" have "n < (\i S (n+1) x" unfolding S_def apply (rule sum_mono) using t1 by auto finally have "?A \ {}" by blast then have "?iA \ ?A" by (meson Inf_nat_def1) moreover have "0 \ ?A" unfolding S_def by auto ultimately have "?iA \ 0" by fastforce define j where "j = ?iA - 1" then have "j < ?iA" using \?iA \ 0\ by auto then have "j \ ?A" by (meson bdd_below_def cInf_lower le0 not_less) then have "S j x \ n" by auto define k where "k = n - S j x" have "n = S j x + k" unfolding k_def using \S j x \ n\ by auto have "n < S (j+1) x" unfolding j_def using \?iA \ 0\ \?iA \ ?A\ by auto also have "... = S j x + t((TB^^j) x)" unfolding S_def by auto also have "... \ S j x + N" using t1 by auto finally have "k \ N" unfolding k_def using \n > N\ by auto then have "S j x > 0" unfolding k_def using \n > N\ by auto then have "j > 0" unfolding S_def using not_gr0 by fastforce have "birkhoff_sum ?I (S j x) x \ birkhoff_sum ?I n x" unfolding birkhoff_sum_def I_def using \S j x \ n\ by (metis finite_Collect_less_nat indicator_pos_le lessThan_def lessThan_subset_iff sum_mono2) have "u n x \ u (S j x) x" proof (cases "k = 0") case True show ?thesis using True unfolding k_def using \S j x \ n\ by auto next case False then have "k > 0" by simp have "u k ((T^^(S j x)) x) \ birkhoff_sum (u 1) k ((T ^^ S j x) x)" using subcocycle_bounded_by_birkhoff1[OF assms(1) \k>0\, of "(T^^(S j x)) x"] by simp also have "... \ 0" unfolding birkhoff_sum_def using sum_mono assms(2) by (simp add: sum_nonpos) also have "u n x \ u (S j x) x + u k ((T^^(S j x)) x)" apply (subst \n = S j x + k\) using assms(1) subcocycle_def by auto ultimately show ?thesis by auto qed also have "... \ (S j x) * ?F x + birkhoff_sum ?I (S j x) x * abs(?F x)" using uS[OF \x \ A\ \j>0\] by simp also have "... \ (S j x) * ?F x + birkhoff_sum ?I n x * abs(?F x)" using \birkhoff_sum ?I (S j x) x \ birkhoff_sum ?I n x\ by (simp add: mult_right_mono) also have "... = n * ?F x - k * ?F x + birkhoff_sum ?I n x * abs(?F x)" by (metis \n = S j x + k\ add_diff_cancel_right' le_add2 left_diff_distrib' of_nat_diff) also have "... \ n * ?F x + k * abs(?F x) + birkhoff_sum ?I n x * abs(?F x)" by (auto, metis abs_ge_minus_self abs_mult abs_of_nat) also have "... \ n * ?F x + N * abs(?F x) + birkhoff_sum ?I n x * abs(?F x)" using \k \ N\ by (simp add: mult_right_mono) finally show ?thesis by simp qed have "limsup (\n. u n x / n) \ ?F x + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" if "x \ A" for x proof - have "(\n. ereal(?F x + N * abs(?F x) * (1 / n))) \ ereal(?F x + N * abs (?F x) * 0)" by (intro tendsto_intros) then have *: "limsup (\n. ?F x + N * abs(?F x)/n) = ?F x" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by force { fix n assume "n > N" have "u n x / real n \ ?F x + N * abs(?F x) / n + abs(?F x) * birkhoff_sum ?I n x / n" using un[OF \x \ A\ \n > N\] using \n>N\ by (auto simp add: divide_simps mult.commute) then have "ereal(u n x/n) \ ereal(?F x + N * abs(?F x) / n) + abs(?F x) * ereal(birkhoff_sum ?I n x / n)" by auto } then have "eventually (\n. ereal(u n x / n) \ ereal(?F x + N * abs(?F x) / n) + abs(?F x) * ereal(birkhoff_sum ?I n x / n)) sequentially" using eventually_mono[OF eventually_gt_at_top[of N]] by auto with Limsup_mono[OF this] have "limsup (\n. u n x / n) \ limsup (\n. ereal(?F x + N * abs(?F x) / n) + abs(?F x) * ereal(birkhoff_sum ?I n x / n))" by auto also have "... \ limsup (\n. ?F x + N * abs(?F x) / n) + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" by (rule ereal_limsup_add_mono) also have "... = ?F x + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" using * by auto finally show ?thesis by auto qed then have *: "AE x in M. limsup (\n. u n x / n) \ ?F x + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" using \AE x in M. x \ A\ by auto { fix x assume H: "(\n. birkhoff_sum ?I n x / n) \ real_cond_exp M Invariants ?I x" have "(\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n)) \ abs(?F x) * ereal(real_cond_exp M Invariants ?I x)" by (rule tendsto_mult_ereal, auto simp add: H) then have "limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n)) = abs(?F x) * ereal(real_cond_exp M Invariants ?I x)" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast } moreover have "AE x in M. (\n. birkhoff_sum ?I n x / n) \ real_cond_exp M Invariants ?I x" by (rule birkhoff_theorem_AE_nonergodic[OF I_int]) ultimately have "AE x in M. limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n)) = abs(?F x) * ereal(real_cond_exp M Invariants ?I x)" by auto then show ?thesis using * by auto qed have bound2: "AE x in M. limsup (\n. u n x / n) \ F K e x" if "K > 0" "e > 0" for K e proof - define C where "C = (\N. A - B N K e)" have C_meas [measurable]: "\N. C N \ sets M" unfolding C_def by auto { fix x assume "x \ A" have "F K e x > l x" using \x \ A\ \e > 0\ unfolding F_def A_def by (cases "l x", auto, metis add.commute ereal_max less_add_same_cancel2 max_less_iff_conj real_of_ereal.simps(1)) then have "\n>0. ereal(u n x / n) < F K e x" unfolding l_def using liminf_upper_bound by fastforce then obtain n where "n>0" "ereal(u n x / n) < F K e x" by auto then have "u n x - n * F K e x < 0" by (simp add: divide_less_eq mult.commute) then have "x \ C n" unfolding C_def B_def using \x \ A\ \n>0\ by auto then have "x \ (\n. C n)" by auto } then have "(\n. C n) = {}" unfolding C_def by auto then have *: "0 = measure M (\n. C n)" by auto have "(\n. measure M (C n)) \ 0" apply (subst *, rule finite_Lim_measure_decseq, auto) unfolding C_def B_def decseq_def by auto moreover have "measure M (C n) = (\x. norm(real_cond_exp M Invariants (I n K e) x) \M)" for n proof - have *: "AE x in M. 0 \ real_cond_exp M Invariants (I n K e) x" apply (rule real_cond_exp_pos, auto) unfolding I_def by auto have "measure M (C n) = (\x. indicator (C n) x \M)" by auto also have "... = (\x. I n K e x \M)" apply (rule integral_cong_AE, auto) unfolding C_def I_def indicator_def using \AE x in M. x \ A\ by auto also have "... = (\x. real_cond_exp M Invariants (I n K e) x \M)" by (rule real_cond_exp_int(2)[symmetric, OF I_int]) also have "... = (\x. norm(real_cond_exp M Invariants (I n K e) x) \M)" apply (rule integral_cong_AE, auto) using * by auto finally show ?thesis by auto qed ultimately have *: "(\n. (\x. norm(real_cond_exp M Invariants (I n K e) x) \M)) \ 0" by simp have "\r. strict_mono r \ (AE x in M. (\n. real_cond_exp M Invariants (I (r n) K e) x) \ 0)" apply (rule tendsto_L1_AE_subseq) using * real_cond_exp_int[OF I_int] by auto then obtain r where "strict_mono r" "AE x in M. (\n. real_cond_exp M Invariants (I (r n) K e) x) \ 0" by auto moreover have "AE x in M. \N \ {1<..}. limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I N K e) x)" apply (rule AE_ball_countable') using main[OF _ \K>0\ \e>0\] by auto moreover { fix x assume H: "(\n. real_cond_exp M Invariants (I (r n) K e) x) \ 0" "\N. N > 1 \ limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I N K e) x)" have 1: "eventually (\N. limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I (r N) K e) x)) sequentially" apply (rule eventually_mono[OF eventually_gt_at_top[of 1] H(2)]) using \strict_mono r\ less_le_trans seq_suble by blast have 2: "(\N. F K e x + (abs(F K e x) * ereal(real_cond_exp M Invariants (I (r N) K e) x))) \ ereal(F K e x) + (abs(F K e x) * ereal 0)" by (intro tendsto_intros) (auto simp add: H(1)) have "limsup (\n. u n x / n) \ F K e x" apply (rule LIMSEQ_le_const) using 1 2 by (auto simp add: eventually_at_top_linorder) } ultimately show "AE x in M. limsup (\n. u n x / n) \ F K e x" by auto qed have "AE x in M. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))" if "K>(0::nat)" for K apply (rule AE_upper_bound_inf_ereal) using bound2 \K>0\ unfolding F_def by auto then have "AE x in M. \K\{(0::nat)<..}. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))" by (rule AE_ball_countable', auto) moreover have "(\n. u n x / n) \ l x" if H: "\K\{(0::nat)<..}. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))" for x proof - have "limsup (\n. u n x / n) \ l x" proof (cases "l x = \") case False then have "(\K. real_of_ereal(max (l x) (-ereal K))) \ l x" using ereal_truncation_real_bottom by auto moreover have "eventually (\K. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))) sequentially" using H by (metis (no_types, lifting) eventually_at_top_linorder eventually_gt_at_top greaterThan_iff) ultimately show "limsup (\n. u n x / n) \ l x" using Lim_bounded2 eventually_sequentially by auto qed (simp) then have "limsup (\n. ereal (u n x / real n)) = l x" using Liminf_le_Limsup l_def eq_iff sequentially_bot by blast then show "(\n. u n x / n) \ l x" by (simp add: l_def tendsto_iff_Liminf_eq_Limsup) qed ultimately have "AE x in M. (\n. u n x / n) \ l x" by auto then have "AE x in M. (\n. u n x / n) \ g x" using g(2) by auto then show "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" using g(1) \\x. g x < \\ by auto qed text \We deduce it for general subcocycles, by reducing to nonpositive subcocycles by subtracting the Birkhoff sum of $u_1$ (for which the convergence follows from Birkhoff theorem).\ theorem kingman_theorem_AE_aux2: assumes "subcocycle u" shows "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" proof - define v where "v = (\n x. u n x + birkhoff_sum (\x. - u 1 x) n x)" have "subcocycle v" unfolding v_def apply (rule subcocycle_add[OF assms], rule subcocycle_birkhoff) using assms unfolding subcocycle_def by auto have "\(gv::'a\ereal). (gv\borel_measurable Invariants \ (\x. gv x < \) \ (AE x in M. (\n. v n x / n) \ gv x))" apply (rule kingman_theorem_AE_aux1[OF \subcocycle v\]) unfolding v_def by auto then obtain gv where gv: "gv \ borel_measurable Invariants" "AE x in M. (\n. v n x / n) \ (gv x::ereal)" "\x. gv x < \" by blast define g where "g = (\x. gv x + ereal(real_cond_exp M Invariants (u 1) x))" have g_meas: "g \ borel_measurable Invariants" unfolding g_def using gv(1) by auto have g_fin: "\x. g x < \" unfolding g_def using gv(3) by auto have "AE x in M. (\n. birkhoff_sum (u 1) n x / n) \ real_cond_exp M Invariants (u 1) x" apply (rule birkhoff_theorem_AE_nonergodic) using assms unfolding subcocycle_def by auto moreover { fix x assume H: "(\n. v n x / n) \ (gv x)" "(\n. birkhoff_sum (u 1) n x / n) \ real_cond_exp M Invariants (u 1) x" then have "(\n. ereal(birkhoff_sum (u 1) n x / n)) \ ereal(real_cond_exp M Invariants (u 1) x)" by auto { fix n have "u n x = v n x + birkhoff_sum (u 1) n x" unfolding v_def birkhoff_sum_def apply auto by (simp add: sum_negf) then have "u n x / n = v n x / n + birkhoff_sum (u 1) n x / n" by (simp add: add_divide_distrib) then have "ereal(u n x / n) = ereal(v n x / n) + ereal(birkhoff_sum (u 1) n x / n)" by auto } note * = this have "(\n. ereal(u n x / n)) \ g x" unfolding * g_def apply (intro tendsto_intros) using H by auto } ultimately have "AE x in M. (\n. ereal(u n x / n)) \ g x" using gv(2) by auto then show ?thesis using g_meas g_fin by blast qed text \For applications, it is convenient to have a limit which is really measurable with respect to the invariant sigma algebra and does not come from a hard to use abstract existence statement. Hence we introduce the following definition for the would-be limit -- Kingman's theorem shows that it is indeed a limit. We introduce the definition for any function, not only subcocycles, but it will only be usable for subcocycles. We introduce an if clause in the definition so that the limit is always measurable, even when $u$ is not a subcocycle and there is no convergence.\ definition subcocycle_lim_ereal::"(nat \ 'a \ real) \ ('a \ ereal)" where "subcocycle_lim_ereal u = ( if (\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))) then (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x)) else (\_. 0))" definition subcocycle_lim::"(nat \ 'a \ real) \ ('a \ real)" where "subcocycle_lim u = (\x. real_of_ereal(subcocycle_lim_ereal u x))" lemma subcocycle_lim_meas_Inv [measurable]: "subcocycle_lim_ereal u \ borel_measurable Invariants" "subcocycle_lim u \ borel_measurable Invariants" proof - show "subcocycle_lim_ereal u \ borel_measurable Invariants" proof (cases "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))") case True then have "subcocycle_lim_ereal u = (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" unfolding subcocycle_lim_ereal_def by auto then show ?thesis using someI_ex[OF True] by auto next case False then have "subcocycle_lim_ereal u = (\_. 0)" unfolding subcocycle_lim_ereal_def by auto then show ?thesis by auto qed then show "subcocycle_lim u \ borel_measurable Invariants" unfolding subcocycle_lim_def by auto qed lemma subcocycle_lim_meas [measurable]: "subcocycle_lim_ereal u \ borel_measurable M" "subcocycle_lim u \ borel_measurable M" using subcocycle_lim_meas_Inv Invariants_measurable_func apply blast using subcocycle_lim_meas_Inv Invariants_measurable_func by blast lemma subcocycle_lim_ereal_not_PInf: "subcocycle_lim_ereal u x < \" proof (cases "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))") case True then have "subcocycle_lim_ereal u = (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" unfolding subcocycle_lim_ereal_def by auto then show ?thesis using someI_ex[OF True] by auto next case False then have "subcocycle_lim_ereal u = (\_. 0)" unfolding subcocycle_lim_ereal_def by auto then show ?thesis by auto qed text \We reformulate the subadditive ergodic theorem of Kingman with this definition. From this point on, the technical definition of \verb+subcocycle_lim_ereal+ will never be used, only the following property will be relevant.\ theorem kingman_theorem_AE_nonergodic_ereal: assumes "subcocycle u" shows "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" proof - have *: "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" using kingman_theorem_AE_aux2[OF assms] by auto then have "subcocycle_lim_ereal u = (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" unfolding subcocycle_lim_ereal_def by auto then show ?thesis using someI_ex[OF *] by auto qed text \The subcocycle limit behaves well under addition, multiplication by a positive scalar, max, and is simply the conditional expectation with respect to invariants for Birkhoff sums, thanks to Birkhoff theorem.\ lemma subcocycle_lim_ereal_add: assumes "subcocycle u" "subcocycle v" shows "AE x in M. subcocycle_lim_ereal (\n x. u n x + v n x) x = subcocycle_lim_ereal u x + subcocycle_lim_ereal v x" proof - have "AE x in M. (\n. (u n x + v n x)/n) \ subcocycle_lim_ereal (\n x. u n x + v n x) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_add[OF assms]]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim_ereal v x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(2)]) moreover { fix x assume H: "(\n. (u n x + v n x)/n) \ subcocycle_lim_ereal (\n x. u n x + v n x) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" "(\n. v n x / n) \ subcocycle_lim_ereal v x" have *: "(u n x + v n x)/n = ereal (u n x / n) + (v n x / n)" for n by (simp add: add_divide_distrib) have "(\n. (u n x + v n x)/n) \ subcocycle_lim_ereal u x + subcocycle_lim_ereal v x" unfolding * apply (intro tendsto_intros H(2) H(3)) using subcocycle_lim_ereal_not_PInf by auto then have "subcocycle_lim_ereal (\n x. u n x + v n x) x = subcocycle_lim_ereal u x + subcocycle_lim_ereal v x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed lemma subcocycle_lim_ereal_cmult: assumes "subcocycle u" "c\(0::real)" shows "AE x in M. subcocycle_lim_ereal (\n x. c * u n x) x = c * subcocycle_lim_ereal u x" proof - have "AE x in M. (\n. (c * u n x)/n) \ subcocycle_lim_ereal (\n x. c * u n x) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_cmult[OF assms]]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover { fix x assume H: "(\n. (c * u n x)/n) \ subcocycle_lim_ereal (\n x. c * u n x) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" have "(\n. c * ereal (u n x / n)) \ c * subcocycle_lim_ereal u x" by (rule tendsto_cmult_ereal[OF _ H(2)], auto) then have "subcocycle_lim_ereal (\n x. c * u n x) x = c * subcocycle_lim_ereal u x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed lemma subcocycle_lim_ereal_max: assumes "subcocycle u" "subcocycle v" shows "AE x in M. subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" proof - have "AE x in M. (\n. max (u n x) (v n x) / n) \ subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_max[OF assms]]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim_ereal v x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(2)]) moreover { fix x assume H: "(\n. max (u n x) (v n x) / n) \ subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" "(\n. v n x / n) \ subcocycle_lim_ereal v x" have "(\n. max (ereal(u n x / n)) (ereal(v n x / n))) \ max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" apply (rule tendsto_max) using H by auto moreover have "max (ereal(u n x / n)) (ereal(v n x / n)) = max (u n x) (v n x) / n" for n by (simp del: ereal_max add:ereal_max[symmetric] max_divide_distrib_right) ultimately have "(\n. max (u n x) (v n x) / n) \ max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" by auto then have "subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" using H(1) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed lemma subcocycle_lim_ereal_birkhoff: assumes "integrable M u" shows "AE x in M. subcocycle_lim_ereal (birkhoff_sum u) x = ereal(real_cond_exp M Invariants u x)" proof - have "AE x in M. (\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" by (rule birkhoff_theorem_AE_nonergodic[OF assms]) moreover have "AE x in M. (\n. birkhoff_sum u n x / n) \ subcocycle_lim_ereal (birkhoff_sum u) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_birkhoff[OF assms]]) moreover { fix x assume H: "(\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" "(\n. birkhoff_sum u n x / n) \ subcocycle_lim_ereal (birkhoff_sum u) x" have "(\n. birkhoff_sum u n x / n) \ ereal(real_cond_exp M Invariants u x)" using H(1) by auto then have "subcocycle_lim_ereal (birkhoff_sum u) x = ereal(real_cond_exp M Invariants u x)" using H(2) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed subsection \$L^1$ and a.e.\ convergence of subcocycles with finite asymptotic average\ text \In this subsection, we show that the almost sure convergence in Kingman theorem also takes place in $L^1$ if the limit is integrable, i.e., if the asymptotic average of the subcocycle is $> -\infty$. To deduce it from the almost sure convergence, we only need to show that there is no loss of mass, i.e., that the integral of the limit is not strictly larger than the limit of the integrals (thanks to Scheffe criterion). This follows from comparison to Birkhoff sums, for which we know that the average of the limit is the same as the average of the function.\ text \First, we show that the subcocycle limit is bounded by the limit of the Birkhoff sums of $u_N$, i.e., its conditional expectation. This follows from the fact that $u_n$ is bounded by the Birkhoff sum of $u_N$ (up to negligible boundary terms).\ lemma subcocycle_lim_ereal_atmost_uN_invariants: assumes "subcocycle u" "N>(0::nat)" shows "AE x in M. subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" proof - have "AE x in M. (\n. u 1 ((T^^n) x) / n) \ 0" apply (rule limit_foTn_over_n') using assms(1) unfolding subcocycle_def by auto moreover have "AE x in M. (\n. birkhoff_sum (\x. u N x/N) n x / n) \ real_cond_exp M Invariants (\x. u N x / N) x" apply (rule birkhoff_theorem_AE_nonergodic) using assms(1) unfolding subcocycle_def by auto moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover { fix x assume H: "(\n. u 1 ((T^^n) x) / n) \ 0" "(\n. birkhoff_sum (\x. u N x/N) n x / n) \ real_cond_exp M Invariants (\x. u N x / N) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" let ?f = "\n. birkhoff_sum (\x. u N x / real N) (n - 2 * N) x / n + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. \u 1 ((T ^^ (n - (2 * N - i))) x)\ / n)" { fix n assume "n\2*N+1" then have "n > 2 * N" by simp have "u n x / n \ (birkhoff_sum (\x. u N x / real N) (n - 2 * N) x + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. \u 1 ((T ^^ (n - (2 * N - i))) x)\)) / n" using subcocycle_bounded_by_birkhoffN[OF assms(1) \n>2*N\ \N>0\, of x] \n>2*N\ by (simp add: divide_right_mono) also have "... = ?f n" apply (subst add_divide_distrib)+ by (auto simp add: sum_divide_distrib[symmetric]) finally have "u n x / n \ ?f n" by simp then have "u n x / n \ ereal(?f n)" by simp } have "(\n. ?f n) \ real_cond_exp M Invariants (\x. u N x / N) x + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. 0)" apply (intro tendsto_intros) using H(2) tendsto_norm[OF H(1)] by auto then have "(\n. ereal(?f n)) \ real_cond_exp M Invariants (\x. u N x / N) x" by auto with lim_mono[OF \\n. n \ 2*N+1 \ u n x / n \ ereal(?f n)\ H(3) this] have "subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" by simp } ultimately show ?thesis by auto qed text \To apply Scheffe criterion, we need to deal with nonnegative functions, or equivalently with nonpositive functions after a change of sign. Hence, as in the proof of the almost sure version of Kingman theorem above, we first give the proof assuming that the subcocycle is nonpositive, and deduce the general statement by adding a suitable Birkhoff sum.\ lemma kingman_theorem_L1_aux: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" "\x. u 1 x \ 0" shows "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "integrable M (subcocycle_lim u)" "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)) \ 0" proof - have int_u [measurable]: "\n. integrable M (u n)" using assms(1) subcocycle_def by auto then have int_F [measurable]: "\n. integrable M (\x. - u n x/ n)" by auto have F_pos: "- u n x / n \ 0" for x n proof (cases "n > 0") case True have "u n x \ (\in>0\] unfolding birkhoff_sum_def by simp also have "... \ 0" using sum_mono[OF assms(3)] by auto finally have "u n x \ 0" by simp then have "-u n x \ 0" by simp with divide_nonneg_nonneg[OF this] show "- u n x / n \ 0" using \n>0\ by auto qed (auto) { fix x assume *: "(\n. u n x / n) \ subcocycle_lim_ereal u x" have H: "(\n. - u n x / n) \ - subcocycle_lim_ereal u x" using tendsto_cmult_ereal[OF _ *, of "-1"] by auto have "liminf (\n. -u n x / n) = - subcocycle_lim_ereal u x" "(\n. - u n x / n) \ - subcocycle_lim_ereal u x" "- subcocycle_lim_ereal u x \ 0" using H apply (simp add: tendsto_iff_Liminf_eq_Limsup, simp) apply (rule LIMSEQ_le_const[OF H]) using F_pos by auto } then have AE_1: "AE x in M. liminf (\n. -u n x / n) = - subcocycle_lim_ereal u x" "AE x in M. (\n. - u n x / n) \ - subcocycle_lim_ereal u x" "AE x in M. - subcocycle_lim_ereal u x \ 0" using kingman_theorem_AE_nonergodic_ereal[OF assms(1)] by auto have "(\\<^sup>+ x. -subcocycle_lim_ereal u x \M) = (\\<^sup>+ x. liminf (\n. -u n x / n) \M)" apply (rule nn_integral_cong_AE) using AE_1(1) by auto also have "... \ liminf (\n. \\<^sup>+ x. -u n x / n \M)" apply (subst e2ennreal_Liminf) apply (simp_all add: e2ennreal_ereal) using F_pos by (intro nn_integral_liminf) (simp add: int_F) also have "... = - subcocycle_avg_ereal u" proof - have "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal u" by (rule subcocycle_int_tendsto_avg_ereal[OF assms(1)]) with tendsto_cmult_ereal[OF _ this, of "-1"] have "(\n. (\x. - u n x / n \M)) \ - subcocycle_avg_ereal u" by simp then have "- subcocycle_avg_ereal u = liminf (\n. (\x. - u n x / n \M))" by (simp add: tendsto_iff_Liminf_eq_Limsup) moreover have "(\\<^sup>+ x. ennreal (-u n x / n) \M) = ennreal(\x. - u n x / n \M)" for n apply (rule nn_integral_eq_integral[OF int_F]) using F_pos by auto ultimately show ?thesis by (auto simp: e2ennreal_Liminf e2ennreal_ereal) qed finally have "(\\<^sup>+ x. -subcocycle_lim_ereal u x \M) \ - subcocycle_avg_ereal u" by simp also have "\ < \" using assms(2) by (cases "subcocycle_avg_ereal u") (auto simp: e2ennreal_ereal e2ennreal_neg) finally have *: "(\\<^sup>+ x. -subcocycle_lim_ereal u x \M) < \" . have "AE x in M. e2ennreal (- subcocycle_lim_ereal u x) \ \" apply (rule nn_integral_PInf_AE) using * by auto then have **: "AE x in M. - subcocycle_lim_ereal u x \ \" using AE_1(3) by eventually_elim simp { fix x assume H: "- subcocycle_lim_ereal u x \ \" "(\n. u n x / n) \ subcocycle_lim_ereal u x" "- subcocycle_lim_ereal u x \ 0" then have 1: "abs(subcocycle_lim_ereal u x) \ \" by auto then have 2: "(\n. u n x / n) \ subcocycle_lim u x" using H(2) unfolding subcocycle_lim_def by auto then have 3: "(\n. - (u n x / n)) \ - subcocycle_lim u x" using tendsto_mult[OF _ 2, of "\_. -1", of "-1"] by auto have 4: "-subcocycle_lim u x \ 0" using H(3) unfolding subcocycle_lim_def by auto have "abs(subcocycle_lim_ereal u x) \ \" "(\n. u n x / n) \ subcocycle_lim u x" "(\n. - (u n x / n)) \ - subcocycle_lim u x" "-subcocycle_lim u x \ 0" using 1 2 3 4 by auto } then have AE_2: "AE x in M. abs(subcocycle_lim_ereal u x) \ \" "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "AE x in M. (\n. - (u n x / n)) \ - subcocycle_lim u x" "AE x in M. -subcocycle_lim u x \ 0" using kingman_theorem_AE_nonergodic_ereal[OF assms(1)] ** AE_1(3) by auto then show "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by simp have "(\\<^sup>+x. abs(subcocycle_lim u x) \M) = (\\<^sup>+x. -subcocycle_lim_ereal u x \M)" apply (rule nn_integral_cong_AE) using AE_2 unfolding subcocycle_lim_def abs_real_of_ereal apply eventually_elim by (auto simp: e2ennreal_ereal) then have A: "(\\<^sup>+x. abs(subcocycle_lim u x) \M) < \" using * by auto show int_Gr: "integrable M (subcocycle_lim u)" apply (rule integrableI_bounded) using A by auto have B: "(\n. (\\<^sup>+ x. norm((- u n x /n) - (-subcocycle_lim u x)) \M)) \ 0" proof (rule Scheffe_lemma1, auto simp add: int_Gr int_u AE_2(2) AE_2(3)) { fix n assume "n>(0::nat)" have *: "AE x in M. subcocycle_lim u x \ real_cond_exp M Invariants (\x. u n x / n) x" using subcocycle_lim_ereal_atmost_uN_invariants[OF assms(1) \n>0\] AE_2(1) unfolding subcocycle_lim_def by auto have "(\x. subcocycle_lim u x \M) \ (\x. real_cond_exp M Invariants (\x. u n x / n) x \M)" apply (rule integral_mono_AE[OF int_Gr _ *], rule real_cond_exp_int(1)) using int_u by auto also have "... = (\x. u n x / n \M)" apply (rule real_cond_exp_int(2)) using int_u by auto finally have A: "(\x. subcocycle_lim u x \M) \ (\x. u n x / n \M)" by auto have "(\\<^sup>+x. abs(u n x) / n \M) = (\\<^sup>+x. - u n x / n \M)" apply (rule nn_integral_cong) using F_pos abs_of_nonneg by (intro arg_cong[where f = ennreal]) fastforce also have "... = (\x. - u n x / n \M)" apply (rule nn_integral_eq_integral) using F_pos int_F by auto also have "... \ (\x. - subcocycle_lim u x \M)" using A by (auto intro!: ennreal_leI) also have "... = (\\<^sup>+x. - subcocycle_lim u x \M)" apply (rule nn_integral_eq_integral[symmetric]) using int_Gr AE_2(4) by auto also have "... = (\\<^sup>+x. abs(subcocycle_lim u x) \M)" apply (rule nn_integral_cong_AE) using AE_2(4) by auto finally have "(\\<^sup>+x. abs(u n x) / n \M) \ (\\<^sup>+x. abs(subcocycle_lim u x) \M)" by simp } with eventually_mono[OF eventually_gt_at_top[of 0] this] have "eventually (\n. (\\<^sup>+x. abs(u n x) / n \M) \ (\\<^sup>+x. abs(subcocycle_lim u x) \M)) sequentially" by fastforce then show "limsup (\n. \\<^sup>+ x. abs(u n x) / n \M) \ \\<^sup>+ x. abs(subcocycle_lim u x) \M" using Limsup_bounded by fastforce qed moreover have "norm((- u n x /n) - (-subcocycle_lim u x)) = abs(u n x / n - subcocycle_lim u x)" for n x by auto ultimately show "(\n. \\<^sup>+ x. ennreal \u n x / real n - subcocycle_lim u x\ \M) \ 0" by auto qed text \We can then remove the nonpositivity assumption, by subtracting the Birkhoff sums of $u_1$ to a general subcocycle $u$.\ theorem kingman_theorem_nonergodic: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "integrable M (subcocycle_lim u)" "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)) \ 0" proof - have [measurable]: "u n \ borel_measurable M" for n using assms(1) unfolding subcocycle_def by auto have int_u [measurable]: "integrable M (u 1)" using assms(1) subcocycle_def by auto define v where "v = (\n x. u n x + birkhoff_sum (\x. - u 1 x) n x)" have [measurable]: "v n \ borel_measurable M" for n unfolding v_def by auto define w where "w = birkhoff_sum (u 1)" have [measurable]: "w n \ borel_measurable M" for n unfolding w_def by auto have "subcocycle v" unfolding v_def apply (rule subcocycle_add[OF assms(1)], rule subcocycle_birkhoff) using assms unfolding subcocycle_def by auto have "subcocycle w" unfolding w_def by (rule subcocycle_birkhoff[OF int_u]) have uvw: "u n x = v n x + w n x" for n x unfolding v_def w_def birkhoff_sum_def by (auto simp add: sum_negf) then have "subcocycle_avg_ereal (\n x. u n x) = subcocycle_avg_ereal v + subcocycle_avg_ereal w" using subcocycle_avg_ereal_add[OF \subcocycle v\ \subcocycle w\] by auto then have "subcocycle_avg_ereal u = subcocycle_avg_ereal v + subcocycle_avg_ereal w" by auto then have "subcocycle_avg_ereal v > -\" unfolding w_def using subcocycle_avg_ereal_birkhoff[OF int_u] assms(2) by auto have "subcocycle_avg_ereal w > - \" unfolding w_def using subcocycle_avg_birkhoff[OF int_u] by auto have "\x. v 1 x \ 0" unfolding v_def by auto have v: "AE x in M. (\n. v n x / n) \ subcocycle_lim v x" "integrable M (subcocycle_lim v)" "(\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M)) \ 0" using kingman_theorem_L1_aux[OF \subcocycle v\ \subcocycle_avg_ereal v > -\\ \\x. v 1 x \ 0\] by auto have w: "AE x in M. (\n. w n x / n) \ subcocycle_lim w x" "integrable M (subcocycle_lim w)" "(\n. (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0" proof - show "AE x in M. (\n. w n x / n) \ subcocycle_lim w x" unfolding w_def subcocycle_lim_def using subcocycle_lim_ereal_birkhoff[OF int_u] birkhoff_theorem_AE_nonergodic[OF int_u] by auto show "integrable M (subcocycle_lim w)" apply (subst integrable_cong_AE[where ?g = "\x. real_cond_exp M Invariants (u 1) x"]) unfolding w_def subcocycle_lim_def using subcocycle_lim_ereal_birkhoff[OF int_u] real_cond_exp_int(1)[OF int_u] by auto have "(\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M) = (\\<^sup>+x. abs(birkhoff_sum (u 1) n x / n - real_cond_exp M Invariants (u 1) x) \M)" for n apply (rule nn_integral_cong_AE) unfolding w_def subcocycle_lim_def using subcocycle_lim_ereal_birkhoff[OF int_u] by auto then show "(\n. (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0" using birkhoff_theorem_L1_nonergodic[OF int_u] by auto qed { fix x assume H: "(\n. v n x / n) \ subcocycle_lim v x" "(\n. w n x / n) \ subcocycle_lim w x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" then have "(\n. v n x / n + w n x / n) \ subcocycle_lim v x + subcocycle_lim w x" using tendsto_add[OF H(1) H(2)] by simp then have *: "(\n. ereal(u n x / n)) \ ereal(subcocycle_lim v x + subcocycle_lim w x)" unfolding uvw by (simp add: add_divide_distrib) then have "subcocycle_lim_ereal u x = ereal(subcocycle_lim v x + subcocycle_lim w x)" using H(3) LIMSEQ_unique by blast then have **: "subcocycle_lim u x = subcocycle_lim v x + subcocycle_lim w x" using subcocycle_lim_def by auto have "u n x / n - subcocycle_lim u x = v n x / n - subcocycle_lim v x + w n x / n - subcocycle_lim w x" for n apply (subst **, subst uvw) using add_divide_distrib add.commute by auto then have "(\n. u n x / n) \ subcocycle_lim u x \ subcocycle_lim u x = subcocycle_lim v x + subcocycle_lim w x \ (\n. u n x / n - subcocycle_lim u x = v n x / n - subcocycle_lim v x + w n x / n - subcocycle_lim w x)" using * ** by auto } then have AE: "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "AE x in M. subcocycle_lim u x = subcocycle_lim v x + subcocycle_lim w x" "AE x in M. \n. u n x / n - subcocycle_lim u x = v n x / n - subcocycle_lim v x + w n x / n - subcocycle_lim w x" using v(1) w(1) kingman_theorem_AE_nonergodic_ereal[OF assms(1)] by auto then show "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by simp show "integrable M (subcocycle_lim u)" apply (subst integrable_cong_AE[where ?g = "\x. subcocycle_lim v x + subcocycle_lim w x"]) by (auto simp add: AE(2) v(2) w(2)) show "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)) \ 0" proof (rule tendsto_sandwich[where ?f = "\_. 0" and ?h = "\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)"], auto) { fix n have "(\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M) = (\\<^sup>+x. abs((v n x / n - subcocycle_lim v x) + (w n x / n - subcocycle_lim w x)) \M)" apply (rule nn_integral_cong_AE) using AE(3) by auto also have "... \ (\\<^sup>+x. ennreal(abs(v n x / n - subcocycle_lim v x)) + abs(w n x / n - subcocycle_lim w x) \M)" by (rule nn_integral_mono, auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus) also have "... = (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)" by (rule nn_integral_add, auto, measurable) finally have "(\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M) \ (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)" using tendsto_sandwich by simp } then show "eventually (\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M) \ (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) sequentially" by auto have "(\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0 + 0" by (rule tendsto_add[OF v(3) w(3)]) then show "(\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0" by simp qed qed text \From the almost sure convergence, we can prove the basic properties of the (real) subcocycle limit: relationship to the asymptotic average, behavior under sum, multiplication, max, behavior for Birkhoff sums.\ lemma subcocycle_lim_avg: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "(\x. subcocycle_lim u x \M) = subcocycle_avg u" proof - have H: "(\n. (\\<^sup>+x. norm(u n x / n - subcocycle_lim u x) \M)) \ 0" "integrable M (subcocycle_lim u)" using kingman_theorem_nonergodic[OF assms] by auto have "(\n. (\x. u n x / n \M)) \ (\x. subcocycle_lim u x \M)" apply (rule tendsto_L1_int[OF _ H(2) H(1)]) using subcocycle_integrable[OF assms(1)] by auto then have "(\n. (\x. u n x / n \M)) \ ereal (\x. subcocycle_lim u x \M)" by auto moreover have "(\n. (\x. u n x / n \M)) \ ereal (subcocycle_avg u)" using subcocycle_int_tendsto_avg[OF assms] by auto ultimately show ?thesis using LIMSEQ_unique by blast qed lemma subcocycle_lim_real_ereal: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" proof - { fix x assume H: "(\n. u n x / n) \ subcocycle_lim_ereal u x" "(\n. u n x / n) \ subcocycle_lim u x" then have "(\n. u n x / n) \ ereal(subcocycle_lim u x)" by auto then have "subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" using H(1) LIMSEQ_unique by blast } then show ?thesis using kingman_theorem_AE_nonergodic_ereal[OF assms(1)] kingman_theorem_nonergodic(1)[OF assms] by auto qed lemma subcocycle_lim_add: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > -\" "subcocycle_avg_ereal v > -\" shows "subcocycle_avg_ereal (\n x. u n x + v n x) > - \" "AE x in M. subcocycle_lim (\n x. u n x + v n x) x = subcocycle_lim u x + subcocycle_lim v x" proof - show *: "subcocycle_avg_ereal (\n x. u n x + v n x) > - \" using subcocycle_avg_add[OF assms(1) assms(2)] assms(3) assms(4) by auto have "AE x in M. (\n. (u n x + v n x)/n) \ subcocycle_lim (\n x. u n x + v n x) x" by (rule kingman_theorem_nonergodic(1)[OF subcocycle_add[OF assms(1) assms(2)] *]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic[OF assms(1) assms(3)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim v x" by (rule kingman_theorem_nonergodic[OF assms(2) assms(4)]) moreover { fix x assume H: "(\n. (u n x + v n x)/n) \ subcocycle_lim (\n x. u n x + v n x) x" "(\n. u n x / n) \ subcocycle_lim u x" "(\n. v n x / n) \ subcocycle_lim v x" have *: "(u n x + v n x)/n = (u n x / n) + (v n x / n)" for n by (simp add: add_divide_distrib) have "(\n. (u n x + v n x)/n) \ subcocycle_lim u x + subcocycle_lim v x" unfolding * by (intro tendsto_intros H) then have "subcocycle_lim (\n x. u n x + v n x) x = subcocycle_lim u x + subcocycle_lim v x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (\n x. u n x + v n x) x = subcocycle_lim u x + subcocycle_lim v x" by auto qed lemma subcocycle_lim_cmult: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" "c\(0::real)" shows "subcocycle_avg_ereal (\n x. c * u n x) > - \" "AE x in M. subcocycle_lim (\n x. c * u n x) x = c * subcocycle_lim u x" proof - show *: "subcocycle_avg_ereal (\n x. c * u n x) > - \" using subcocycle_avg_cmult[OF assms(1) assms(3)] assms(2) assms(3) by auto have "AE x in M. (\n. (c * u n x)/n) \ subcocycle_lim (\n x. c * u n x) x" by (rule kingman_theorem_nonergodic(1)[OF subcocycle_cmult[OF assms(1) assms(3)] *]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic(1)[OF assms(1) assms(2)]) moreover { fix x assume H: "(\n. (c * u n x)/n) \ subcocycle_lim (\n x. c * u n x) x" "(\n. u n x / n) \ subcocycle_lim u x" have "(\n. c * (u n x / n)) \ c * subcocycle_lim u x" by (rule tendsto_mult[OF _ H(2)], auto) then have "subcocycle_lim (\n x. c * u n x) x = c * subcocycle_lim u x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (\n x. c * u n x) x = c * subcocycle_lim u x" by auto qed lemma subcocycle_lim_max: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > -\" "subcocycle_avg_ereal v > -\" shows "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > - \" "AE x in M. subcocycle_lim (\n x. max (u n x) (v n x)) x = max (subcocycle_lim u x) (subcocycle_lim v x)" proof - show *: "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > - \" using subcocycle_avg_max(1)[OF assms(1) assms(2)] assms(3) assms(4) by auto have "AE x in M. (\n. max (u n x) (v n x) / n) \ subcocycle_lim (\n x. max (u n x) (v n x)) x" by (rule kingman_theorem_nonergodic[OF subcocycle_max[OF assms(1) assms(2)] *]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic[OF assms(1) assms(3)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim v x" by (rule kingman_theorem_nonergodic[OF assms(2) assms(4)]) moreover { fix x assume H: "(\n. max (u n x) (v n x) / n) \ subcocycle_lim (\n x. max (u n x) (v n x)) x" "(\n. u n x / n) \ subcocycle_lim u x" "(\n. v n x / n) \ subcocycle_lim v x" have "(\n. max (u n x / n) (v n x / n)) \ max (subcocycle_lim u x) (subcocycle_lim v x)" apply (rule tendsto_max) using H by auto moreover have "max (u n x / n) (v n x / n) = max (u n x) (v n x) / n" for n by (simp add: max_divide_distrib_right) ultimately have "(\n. max (u n x) (v n x) / n) \ max (subcocycle_lim u x) (subcocycle_lim v x)" by auto then have "subcocycle_lim (\n x. max (u n x) (v n x)) x = max (subcocycle_lim u x) (subcocycle_lim v x)" using H(1) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (\n x. max (u n x) (v n x)) x = max (subcocycle_lim u x) (subcocycle_lim v x)" by auto qed lemma subcocycle_lim_birkhoff: assumes "integrable M u" shows "subcocycle_avg_ereal (birkhoff_sum u) > -\" "AE x in M. subcocycle_lim (birkhoff_sum u) x = real_cond_exp M Invariants u x" proof - show *: "subcocycle_avg_ereal (birkhoff_sum u) > -\" using subcocycle_avg_birkhoff[OF assms] by auto have "AE x in M. (\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" by (rule birkhoff_theorem_AE_nonergodic[OF assms]) moreover have "AE x in M. (\n. birkhoff_sum u n x / n) \ subcocycle_lim (birkhoff_sum u) x" by (rule kingman_theorem_nonergodic(1)[OF subcocycle_birkhoff[OF assms] *]) moreover { fix x assume H: "(\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" "(\n. birkhoff_sum u n x / n) \ subcocycle_lim (birkhoff_sum u) x" then have "subcocycle_lim (birkhoff_sum u) x = real_cond_exp M Invariants u x" using H(2) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (birkhoff_sum u) x = real_cond_exp M Invariants u x" by auto qed subsection \Conditional expectations of subcocycles\ text \In this subsection, we show that the conditional expectations of a subcocycle (with respect to the invariant subalgebra) also converge, with the same limit as the cocycle. Note that the conditional expectation of a subcocycle $u$ is still a subcocycle, with the same average at each step so with the same asymptotic average. Kingman theorem can be applied to it, and what we have to show is that the limit of this subcocycle is the same as the limit of the original subcocycle. When the asymptotic average is $>-\infty$, both limits have the same integral, and moreover the domination of the subcocycle by the Birkhoff sums of $u_n$ for fixed $n$ (which converge to the conditional expectation of $u_n$) implies that one limit is smaller than the other. Hence, they coincide almost everywhere. The case when the asymptotic average is $-\infty$ is deduced from the previous one by truncation. \ text \First, we prove the result when the asymptotic average with finite.\ theorem kingman_theorem_nonergodic_invariant: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim u x" "(\n. (\\<^sup>+x. abs(real_cond_exp M Invariants (u n) x / n - subcocycle_lim u x) \M)) \ 0" proof - have int [simp]: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by auto then have int2: "integrable M (real_cond_exp M Invariants (u n))" for n using real_cond_exp_int by auto { fix n m have "u (n+m) x \ u n x + u m ((T^^n) x)" for x using subcocycle_ineq[OF assms(1)] by auto have "AE x in M. real_cond_exp M Invariants (u (n+m)) x \ real_cond_exp M Invariants (\x. u n x + u m ((T^^n) x)) x" apply (rule real_cond_exp_mono) using subcocycle_ineq[OF assms(1)] apply auto by (rule Bochner_Integration.integrable_add, auto simp add: Tn_integral_preserving) moreover have "AE x in M. real_cond_exp M Invariants (\x. u n x + u m ((T^^n) x)) x = real_cond_exp M Invariants (u n) x + real_cond_exp M Invariants (\x. u m ((T^^n) x)) x" by (rule real_cond_exp_add, auto simp add: Tn_integral_preserving) moreover have "AE x in M. real_cond_exp M Invariants (u m \ ((T^^n))) x = real_cond_exp M Invariants (u m) x" by (rule Invariants_of_foTn, simp) moreover have "AE x in M. real_cond_exp M Invariants (u m) x = real_cond_exp M Invariants (u m) ((T^^n) x)" using Invariants_func_is_invariant_n[symmetric, of "real_cond_exp M Invariants (u m)"] by auto ultimately have "AE x in M. real_cond_exp M Invariants (u (n+m)) x \ real_cond_exp M Invariants (u n) x + real_cond_exp M Invariants (u m) ((T^^n) x)" unfolding o_def by auto } with subcocycle_AE[OF this int2] obtain w where w: "subcocycle w" "AE x in M. \n. w n x = real_cond_exp M Invariants (u n) x" by blast have [measurable]: "integrable M (w n)" for n using subcocycle_integrable[OF w(1)] by simp { fix n::nat have "(\x. w n x / n \M) = (\x. real_cond_exp M Invariants (u n) x / n \M)" apply (rule integral_cong_AE) using w(2) by auto also have "... = (\x. real_cond_exp M Invariants (u n) x \M) / n" by (rule integral_divide_zero) also have "... = (\x. u n x \M) / n" by (simp add: divide_simps real_cond_exp_int(2)[OF int[of n]]) also have "... = (\x. u n x / n \M)" by (rule integral_divide_zero[symmetric]) finally have "ereal (\x. w n x / n \M) = ereal (\x. u n x / n \M)" by simp } note * = this have "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal w" apply (rule Lim_transform_eventually[OF subcocycle_int_tendsto_avg_ereal[OF w(1)]]) using * by auto then have "subcocycle_avg_ereal u = subcocycle_avg_ereal w" using subcocycle_int_tendsto_avg_ereal[OF assms(1)] LIMSEQ_unique by auto then have "subcocycle_avg_ereal w > -\" using assms(2) by simp have "subcocycle_avg u = subcocycle_avg w" using \subcocycle_avg_ereal u = subcocycle_avg_ereal w\ unfolding subcocycle_avg_def by simp have *: "AE x in M. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" for N by (cases "N = 0", auto simp add: subcocycle_lim_ereal_atmost_uN_invariants[OF assms(1)]) have "AE x in M. \N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" by (subst AE_all_countable, intro allI, simp add: *) moreover have "AE x in M. subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" by (rule subcocycle_lim_real_ereal[OF assms]) moreover have "AE x in M. (\N. u N x / N) \ subcocycle_lim u x" using kingman_theorem_nonergodic[OF assms] by simp moreover have "AE x in M. (\N. w N x / N) \ subcocycle_lim w x" using kingman_theorem_nonergodic[OF w(1) \subcocycle_avg_ereal w > -\\ ] by simp moreover have "AE x in M. \n. w n x = real_cond_exp M Invariants (u n) x" using w(2) by simp moreover have "AE x in M. \n. real_cond_exp M Invariants (u n) x / n = real_cond_exp M Invariants (\x. u n x / n) x" apply (subst AE_all_countable, intro allI) using AE_symmetric[OF real_cond_exp_cdiv[OF int]] by auto moreover { fix x assume x: "\N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" "subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" "(\N. u N x / N) \ subcocycle_lim u x" "(\N. w N x / N) \ subcocycle_lim w x" "\n. w n x = real_cond_exp M Invariants (u n) x" "\n. real_cond_exp M Invariants (u n) x / n = real_cond_exp M Invariants (\x. u n x / n) x" { fix N::nat assume "N\1" have "subcocycle_lim u x \ real_cond_exp M Invariants (\x. u N x / N) x" using x(1) x(2) \N\1\ by auto also have "... = real_cond_exp M Invariants (u N) x / N" using x(6) by simp also have "... = w N x / N" using x(5) by simp finally have "subcocycle_lim u x \ w N x / N" by simp } note * = this have "subcocycle_lim u x \ subcocycle_lim w x" apply (rule LIMSEQ_le_const[OF x(4)]) using * by auto } ultimately have *: "AE x in M. subcocycle_lim u x \ subcocycle_lim w x" by auto have **: "(\x. subcocycle_lim u x \M) = (\x. subcocycle_lim w x \M)" using subcocycle_lim_avg[OF assms] subcocycle_lim_avg[OF w(1) \subcocycle_avg_ereal w > -\\] \subcocycle_avg u = subcocycle_avg w\ by simp have AE_eq: "AE x in M. subcocycle_lim u x = subcocycle_lim w x" by (rule integral_ineq_eq_0_then_AE[OF * kingman_theorem_nonergodic(2)[OF assms] kingman_theorem_nonergodic(2)[OF w(1) \subcocycle_avg_ereal w > -\\] **]) moreover have "AE x in M. (\n. w n x / n) \ subcocycle_lim w x" by (rule kingman_theorem_nonergodic(1)[OF w(1) \subcocycle_avg_ereal w > -\\]) moreover have "AE x in M. \n. w n x = real_cond_exp M Invariants (u n) x" using w(2) by auto moreover { fix x assume H: "subcocycle_lim u x = subcocycle_lim w x" "(\n. w n x / n) \ subcocycle_lim w x" "\n. w n x = real_cond_exp M Invariants (u n) x" then have "(\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim u x" by auto } ultimately show "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim u x" by auto { fix n::nat have "AE x in M. subcocycle_lim u x = subcocycle_lim w x" using AE_eq by simp moreover have "AE x in M. w n x = real_cond_exp M Invariants (u n) x" using w(2) by auto moreover { fix x assume H: "subcocycle_lim u x = subcocycle_lim w x" "w n x = real_cond_exp M Invariants (u n) x" then have "ennreal \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ = ennreal \w n x / real n - subcocycle_lim w x\" by auto } ultimately have "AE x in M. ennreal \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ = ennreal \w n x / real n - subcocycle_lim w x\" by auto then have "(\\<^sup>+ x. ennreal \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ \M) = (\\<^sup>+ x. ennreal \w n x / real n - subcocycle_lim w x\ \M)" by (rule nn_integral_cong_AE) } moreover have "(\n. (\\<^sup>+ x. \w n x / real n - subcocycle_lim w x\ \M)) \ 0" by (rule kingman_theorem_nonergodic(3)[OF w(1) \subcocycle_avg_ereal w > -\\]) ultimately show "(\n. (\\<^sup>+ x. \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ \M)) \ 0" by auto qed text \Then, we extend it by truncation to the general case, i.e., to the asymptotic limit in extended reals.\ theorem kingman_theorem_AE_nonergodic_invariant_ereal: assumes "subcocycle u" shows "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" proof - have [simp]: "subcocycle u" using assms by simp have int [simp]: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by auto have limsup_ineq_K: "AE x in M. limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" for K::nat proof - define v where "v = (\ (n::nat) (x::'a). (-n * real K))" have [simp]: "subcocycle v" unfolding v_def subcocycle_def by (auto simp add: algebra_simps) have "ereal (\x. v n x / n \M) = ereal(- real K * measure M (space M))" if "n\1" for n unfolding v_def using that by simp then have "(\n. ereal (\x. v n x / n \M)) \ ereal(- real K * measure M (space M))" using lim_explicit by force moreover have "(\n. ereal (\x. v n x / n \M)) \ subcocycle_avg_ereal v" using subcocycle_int_tendsto_avg_ereal[OF \subcocycle v\] by auto ultimately have "subcocycle_avg_ereal v = - real K * measure M (space M)" using LIMSEQ_unique by blast then have "subcocycle_avg_ereal v > -\" by auto { fix x assume H: "(\n. v n x / n) \ subcocycle_lim_ereal v x" have "ereal(v n x / n) = -real K" if "n\1" for n unfolding v_def using that by auto then have "(\n. ereal(v n x / n)) \ - real K" using lim_explicit by force then have "subcocycle_lim_ereal v x = -real K" using H LIMSEQ_unique by blast } then have "AE x in M. subcocycle_lim_ereal v x = -real K" using kingman_theorem_AE_nonergodic_ereal[OF \subcocycle v\] by auto define w where "w = (\n x. max (u n x) (v n x))" have [simp]: "subcocycle w" unfolding w_def by (rule subcocycle_max, auto) have "subcocycle_avg_ereal w \ subcocycle_avg_ereal v" unfolding w_def using subcocycle_avg_ereal_max by auto then have "subcocycle_avg_ereal w > -\" using \subcocycle_avg_ereal v > -\\ by auto have *: "AE x in M. real_cond_exp M Invariants (u n) x \ real_cond_exp M Invariants (w n) x" for n apply (rule real_cond_exp_mono) using subcocycle_integrable[OF assms, of n] subcocycle_integrable[OF \subcocycle w\, of n] apply auto unfolding w_def by auto have "AE x in M. \n. real_cond_exp M Invariants (u n) x \ real_cond_exp M Invariants (w n) x" apply (subst AE_all_countable) using * by auto moreover have "AE x in M. (\n. real_cond_exp M Invariants (w n) x / n) \ subcocycle_lim w x" apply (rule kingman_theorem_nonergodic_invariant(1)) using \subcocycle_avg_ereal w > -\\ by auto moreover have "AE x in M. subcocycle_lim_ereal w x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" unfolding w_def using subcocycle_lim_ereal_max by auto moreover { fix x assume H: "(\n. real_cond_exp M Invariants (w n) x / n) \ subcocycle_lim w x" "subcocycle_lim_ereal w x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" "subcocycle_lim_ereal v x = - real K" "\n. real_cond_exp M Invariants (u n) x \ real_cond_exp M Invariants (w n) x" have "subcocycle_lim_ereal w x > -\" - using H(2) H(3) MInfty_neq_ereal(1) ereal_MInfty_lessI max.cobounded2 by fastforce + using H(2) H(3) + by auto (metis MInfty_neq_ereal(1) ereal_infty_less_eq2(2) max.cobounded2) then have "subcocycle_lim_ereal w x = ereal(subcocycle_lim w x)" unfolding subcocycle_lim_def using subcocycle_lim_ereal_not_PInf[of w x] ereal_real by force moreover have "(\n. real_cond_exp M Invariants (w n) x / n) \ ereal(subcocycle_lim w x)" using H(1) by auto ultimately have "(\n. real_cond_exp M Invariants (w n) x / n) \ subcocycle_lim_ereal w x" by auto then have *: "limsup (\n. real_cond_exp M Invariants (w n) x / n) = subcocycle_lim_ereal w x" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast have "ereal(real_cond_exp M Invariants (u n) x / n) \ real_cond_exp M Invariants (w n) x / n" for n using H(4) by (auto simp add: divide_simps) then have "eventually (\n. ereal(real_cond_exp M Invariants (u n) x / n) \ real_cond_exp M Invariants (w n) x / n) sequentially" by auto then have "limsup (\n. real_cond_exp M Invariants (u n) x / n) \ limsup (\n. real_cond_exp M Invariants (w n) x / n)" using Limsup_mono[of _ _ sequentially] by force then have "limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" using * H(2) H(3) by auto } ultimately show ?thesis using \AE x in M. subcocycle_lim_ereal v x = -real K\ by auto qed have "AE x in M. \K::nat. limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" apply (subst AE_all_countable) using limsup_ineq_K by auto moreover have "AE x in M. liminf (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" proof - have *: "AE x in M. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" for N by (cases "N = 0", auto simp add: subcocycle_lim_ereal_atmost_uN_invariants[OF assms(1)]) have "AE x in M. \N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" by (subst AE_all_countable, intro allI, simp add: *) moreover have "AE x in M. \n. real_cond_exp M Invariants (\x. u n x / n) x = real_cond_exp M Invariants (u n) x / n" apply (subst AE_all_countable, intro allI) using real_cond_exp_cdiv by auto moreover { fix x assume x: "\N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" "\n. real_cond_exp M Invariants (\x. u n x / n) x = real_cond_exp M Invariants (u n) x / n" then have *: "subcocycle_lim_ereal u x \ real_cond_exp M Invariants (u n) x / n" if "n \ 1" for n using that by auto have "subcocycle_lim_ereal u x \ liminf (\n. real_cond_exp M Invariants (u n) x / n)" apply (subst liminf_bounded_iff) using * less_le_trans by blast } ultimately show ?thesis by auto qed moreover { fix x assume H: "\K::nat. limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" "liminf (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" have "(\K::nat. max (subcocycle_lim_ereal u x) (-real K)) \ subcocycle_lim_ereal u x" by (rule ereal_truncation_bottom) with LIMSEQ_le_const[OF this] have *: "limsup (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" using H(1) by auto have "(\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" apply (subst tendsto_iff_Liminf_eq_Limsup[OF trivial_limit_at_top_linorder]) using H(2) * Liminf_le_Limsup[OF trivial_limit_at_top_linorder, of "(\n. real_cond_exp M Invariants (u n) x / n)"] by auto } ultimately show ?thesis by auto qed end subsection \Subcocycles in the ergodic case\ text \In this subsection, we describe how all the previous results simplify in the ergodic case. Indeed, subcocycle limits are almost surely constant, given by the asymptotic average.\ context ergodic_pmpt begin lemma subcocycle_ergodic_lim_avg: assumes "subcocycle u" shows "AE x in M. subcocycle_lim_ereal u x = subcocycle_avg_ereal u" "AE x in M. subcocycle_lim u x = subcocycle_avg u" proof - have I: "integrable M (u N)" for N using subcocycle_integrable[OF assms]by auto obtain c::ereal where c: "AE x in M. subcocycle_lim_ereal u x = c" using Invariant_func_is_AE_constant[OF subcocycle_lim_meas_Inv(1)] by blast have "c = subcocycle_avg_ereal u" proof (cases "subcocycle_avg_ereal u = - \") case True { fix N assume "N > (0::nat)" have "AE x in M. real_cond_exp M Invariants (\x. u N x / N) x = (\ x. u N x / N \M)" apply (rule Invariants_cond_exp_is_integral) using I by auto moreover have "AE x in M. subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" using subcocycle_lim_ereal_atmost_uN_invariants[OF assms \N>0\] by simp ultimately have "AE x in M. c \ (\x. u N x / N \M)" using c by force then have "c \ (\x. u N x / N \M)" by auto } then have "\N\1. c \ (\x. u N x / N \M)" by auto with Lim_bounded2[OF subcocycle_int_tendsto_avg_ereal[OF assms] this] have "c \ subcocycle_avg_ereal u" by simp then show ?thesis using True by auto next case False then have fin: "subcocycle_avg_ereal u > - \" by simp obtain cr::real where cr: "AE x in M. subcocycle_lim u x = cr" using Invariant_func_is_AE_constant[OF subcocycle_lim_meas_Inv(2)] by blast have "AE x in M. c = ereal cr" using c cr subcocycle_lim_real_ereal[OF assms fin] by force then have "c = ereal cr" by auto have "subcocycle_avg u = (\x. subcocycle_lim u x \M)" using subcocycle_lim_avg[OF assms fin] by auto also have "... = (\x. cr \M)" apply (rule integral_cong_AE) using cr by auto also have "... = cr" by (simp add: prob_space.prob_space prob_space_axioms) finally have "ereal(subcocycle_avg u) = ereal cr" by simp then show ?thesis using \ c = ereal cr \ subcocycle_avg_real_ereal[OF fin] by auto qed then show "AE x in M. subcocycle_lim_ereal u x = subcocycle_avg_ereal u" using c by auto then show "AE x in M. subcocycle_lim u x = subcocycle_avg u" unfolding subcocycle_lim_def subcocycle_avg_def by auto qed theorem kingman_theorem_AE_ereal: assumes "subcocycle u" shows "AE x in M. (\n. u n x / n) \ subcocycle_avg_ereal u" using kingman_theorem_AE_nonergodic_ereal[OF assms] subcocycle_ergodic_lim_avg(1)[OF assms] by auto theorem kingman_theorem: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. (\n. u n x / n) \ subcocycle_avg u" "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_avg u) \M)) \ 0" proof - have *: "AE x in M. subcocycle_lim u x = subcocycle_avg u" using subcocycle_ergodic_lim_avg(2)[OF assms(1)] by auto then show "AE x in M. (\n. u n x / n) \ subcocycle_avg u" using kingman_theorem_nonergodic(1)[OF assms] by auto have "(\\<^sup>+x. abs(u n x / n - subcocycle_avg u) \M) = (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)" for n apply (rule nn_integral_cong_AE) using * by auto then show "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_avg u) \M)) \ 0" using kingman_theorem_nonergodic(3)[OF assms] by auto qed end subsection \Subocycles for invertible maps\ text \If $T$ is invertible, then a subcocycle $u_n$ for $T$ gives rise to another subcocycle for $T^{-1}$. Intuitively, if $u$ is subadditive along the time interval $[0,n)$, then it should also be subadditive along the time interval $[-n,0)$. This is true, and formalized with the following statement.\ proposition (in mpt) subcocycle_u_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "mpt.subcocycle M Tinv (\n x. u n (((Tinv)^^n) x))" proof - have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by auto have int: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by simp interpret I: mpt M Tinv using Tinv_mpt[OF assms(2)] by simp show "I.subcocycle (\n x. u n (((Tinv)^^n) x))" unfolding I.subcocycle_def proof(auto) show "integrable M (\x. u n ((Tinv ^^ n) x))" for n using I.Tn_integral_preserving(1)[OF int[of n]] by simp fix n m::nat and x::'a define y where "y = (Tinv^^(m+n)) x" have "(T^^m) y = (T^^m) ((Tinv^^m) ((Tinv^^n) x))" unfolding y_def by (simp add: funpow_add) then have *: "(T^^m) y = (Tinv^^n) x" using fn_o_inv_fn_is_id[OF bij, of m] by (metis Tinv_def comp_def) have "u (n + m) ((Tinv ^^ (n + m)) x) = u (m+n) y" unfolding y_def by (simp add: add.commute[of n m]) also have "... \ u m y + u n ((T^^m) y)" using subcocycle_ineq[OF \subcocycle u\, of m n y] by simp also have "... = u m ((Tinv^^(m+n)) x) + u n ((Tinv^^n) x)" using * y_def by auto finally show "u (n + m) ((Tinv ^^ (n + m)) x) \ u n ((Tinv ^^ n) x) + u m ((Tinv ^^ m) ((Tinv ^^ n) x))" by (simp add: funpow_add) qed qed text \The subcocycle averages for $T$ and $T^{-1}$ coincide.\ proposition (in mpt) subcocycle_avg_ereal_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "mpt.subcocycle_avg_ereal M (\n x. u n (((Tinv)^^n) x)) = subcocycle_avg_ereal u" proof - have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by auto have int: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by simp interpret I: mpt M Tinv using Tinv_mpt[OF assms(2)] by simp have "(\n. (\x. u n (((Tinv)^^n) x) / n \M)) \ I.subcocycle_avg_ereal (\n x. u n (((Tinv)^^n) x))" using I.subcocycle_int_tendsto_avg_ereal[OF subcocycle_u_Tinv[OF assms]] by simp moreover have "(\x. u n x / n \M) = ereal (\x. u n (((Tinv)^^n) x) / n \M)" for n apply (simp) apply (rule disjI2) apply (rule I.Tn_integral_preserving(2)[symmetric]) apply (simp add: int) done ultimately have "(\n. (\x. u n x / n \M)) \ I.subcocycle_avg_ereal (\n x. u n (((Tinv)^^n) x))" by presburger moreover have "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal u" using subcocycle_int_tendsto_avg_ereal[OF \subcocycle u\] by simp ultimately show ?thesis using LIMSEQ_unique by simp qed text \The asymptotic limit of the subcocycle is the same for $T$ and $T^{-1}$. This is clear in the ergodic case, and follows from the ergodic decomposition in the general case (on a standard probability space). We give a direct proof below (on a general probability space) using the fact that the asymptotic limit is the same for the subcocycle conditioned by the invariant sigma-algebra, which is clearly the same for $T$ and $T^{-1}$ as it is constant along orbits.\ proposition (in fmpt) subcocycle_lim_ereal_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "AE x in M. fmpt.subcocycle_lim_ereal M Tinv (\n x. u n (((Tinv)^^n) x)) x = subcocycle_lim_ereal u x" proof - have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by auto have int: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by simp interpret I: fmpt M Tinv using Tinv_fmpt[OF assms(2)] by simp have *: "AE x in M. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x = real_cond_exp M I.Invariants (u n) x" for n using I.Invariants_of_foTn int unfolding o_def by simp have "AE x in M. \n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x = real_cond_exp M I.Invariants (u n) x" apply (subst AE_all_countable) using * by simp moreover have "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" using kingman_theorem_AE_nonergodic_invariant_ereal[OF \subcocycle u\] by simp moreover have "AE x in M. (\n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x / n) \ I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x" using I.kingman_theorem_AE_nonergodic_invariant_ereal[OF subcocycle_u_Tinv[OF assms]] by simp moreover { fix x assume H: "\n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x = real_cond_exp M I.Invariants (u n) x" "(\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" "(\n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x / n) \ I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x" have "ereal(real_cond_exp M Invariants (u n) x / n) = ereal(real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x / n)" for n using H(1) Invariants_Tinv[OF \invertible_qmpt\] by auto then have "(\n. real_cond_exp M Invariants (u n) x / n) \ I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x" using H(3) by presburger then have "I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x = subcocycle_lim_ereal u x" using H(2) LIMSEQ_unique by auto } ultimately show ?thesis by auto qed proposition (in fmpt) subcocycle_lim_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "AE x in M. fmpt.subcocycle_lim M Tinv (\n x. u n (((Tinv)^^n) x)) x = subcocycle_lim u x" proof - interpret I: fmpt M Tinv using Tinv_fmpt[OF assms(2)] by simp show ?thesis unfolding subcocycle_lim_def I.subcocycle_lim_def using subcocycle_lim_ereal_Tinv[OF assms] by auto qed end (*of Kingman.thy*) diff --git a/thys/Extended_Finite_State_Machines/GExp.thy b/thys/Extended_Finite_State_Machines/GExp.thy --- a/thys/Extended_Finite_State_Machines/GExp.thy +++ b/thys/Extended_Finite_State_Machines/GExp.thy @@ -1,837 +1,827 @@ subsection \Guards Expressions\ text\ This theory defines the guard language of EFSMs which can be translated directly to and from contexts. Boolean values true and false respectively represent the guards which are always and never satisfied. Guards may test for (in)equivalence of two arithmetic expressions or be connected using \textsc{nor} logic into compound expressions. The use of \textsc{nor} logic reduces the number of subgoals when inducting over guard expressions. We also define syntax hacks for the relations less than, less than or equal to, greater than or equal to, and not equal to as well as the expression of logical conjunction, disjunction, and negation in terms of nor logic.\ theory GExp imports AExp Trilean begin text_raw\\snip{gexptype}{1}{2}{%\ datatype 'a gexp = Bc bool | Eq "'a aexp" "'a aexp" | Gt "'a aexp" "'a aexp" | In 'a "value list" | Nor "'a gexp" "'a gexp" text_raw\}%endsnip\ fun gval :: "'a gexp \ 'a datastate \ trilean" where "gval (Bc True) _ = true" | "gval (Bc False) _ = false" | "gval (Gt a1 a2) s = value_gt (aval a1 s) (aval a2 s)" | "gval (Eq a1 a2) s = value_eq (aval a1 s) (aval a2 s)" | "gval (In v l) s = (case s v of None \ invalid | Some vv \ if vv \ set l then true else false)" | "gval (Nor a1 a2) s = \? ((gval a1 s) \? (gval a2 s))" text_raw\\snip{connectives}{1}{2}{%\ definition gNot :: "'a gexp \ 'a gexp" where "gNot g \ Nor g g" definition gOr :: "'a gexp \ 'a gexp \ 'a gexp" (*infix "\" 60*) where "gOr v va \ Nor (Nor v va) (Nor v va)" definition gAnd :: "'a gexp \ 'a gexp \ 'a gexp" (*infix "\" 60*) where "gAnd v va \ Nor (Nor v v) (Nor va va)" definition gImplies :: "'a gexp \ 'a gexp \ 'a gexp" where "gImplies p q \ gOr (gNot p) q" definition Lt :: "'a aexp \ 'a aexp \ 'a gexp" (*infix "<" 60*) where "Lt a b \ Gt b a" definition Le :: "'a aexp \ 'a aexp \ 'a gexp" (*infix "\" 60*) where "Le v va \ gNot (Gt v va)" definition Ge :: "'a aexp \ 'a aexp \ 'a gexp" (*infix "\" 60*) where "Ge v va \ gNot (Lt v va)" definition Ne :: "'a aexp \ 'a aexp \ 'a gexp" (*infix "\" 60*) where "Ne v va \ gNot (Eq v va)" text_raw\}%endsnip\ lemma gval_Lt [simp]: "gval (Lt a1 a2) s = value_gt (aval a2 s) (aval a1 s)" by (simp add: Lt_def) lemma gval_Le [simp]: "gval (Le a1 a2) s = \? (value_gt (aval a1 s) (aval a2 s))" by (simp add: Le_def value_gt_def gNot_def maybe_or_idempotent) lemma gval_Ge [simp]: "gval (Ge a1 a2) s = \? (value_gt (aval a2 s) (aval a1 s))" by (simp add: Ge_def value_gt_def gNot_def maybe_or_idempotent) lemma gval_Ne [simp]: "gval (Ne a1 a2) s = \? (value_eq (aval a1 s) (aval a2 s))" by (simp add: Ne_def value_gt_def gNot_def maybe_or_idempotent) lemmas connectives = gAnd_def gOr_def gNot_def Lt_def Le_def Ge_def Ne_def lemma gval_gOr [simp]: "gval (gOr x y) r = (gval x r) \? (gval y r)" by (simp add: maybe_double_negation maybe_or_idempotent gOr_def) lemma gval_gNot [simp]: "gval (gNot x) s = \? (gval x s)" by (simp add: maybe_or_idempotent gNot_def) lemma gval_gAnd [simp]: "gval (gAnd g1 g2) s = (gval g1 s) \? (gval g2 s)" by (simp add: de_morgans_1 maybe_double_negation maybe_or_idempotent gAnd_def) lemma gAnd_commute: "gval (gAnd a b) s = gval (gAnd b a) s" by (simp add: times_trilean_commutative) lemma gOr_commute: "gval (gOr a b) s = gval (gOr b a) s" by (simp add: plus_trilean_commutative gOr_def) lemma gval_gAnd_True: "(gval (gAnd g1 g2) s = true) = ((gval g1 s = true) \ gval g2 s = true)" by (simp add: maybe_and_true) lemma nor_equiv: "gval (gNot (gOr a b)) s = gval (Nor a b) s" by simp definition satisfiable :: "vname gexp \ bool" where "satisfiable g \ (\i r. gval g (join_ir i r) = true)" definition "satisfiable_list l = satisfiable (fold gAnd l (Bc True))" lemma unsatisfiable_false: "\ satisfiable (Bc False)" by (simp add: satisfiable_def) lemma satisfiable_true: "satisfiable (Bc True)" by (simp add: satisfiable_def) definition valid :: "vname gexp \ bool" where "valid g \ (\s. gval g s = true)" lemma valid_true: "valid (Bc True)" by (simp add: valid_def) fun gexp_constrains :: "'a gexp \ 'a aexp \ bool" where "gexp_constrains (Bc _) _ = False" | "gexp_constrains (Eq a1 a2) a = (aexp_constrains a1 a \ aexp_constrains a2 a)" | "gexp_constrains (Gt a1 a2) a = (aexp_constrains a1 a \ aexp_constrains a2 a)" | "gexp_constrains (Nor g1 g2) a = (gexp_constrains g1 a \ gexp_constrains g2 a)" | "gexp_constrains (In v l) a = aexp_constrains (V v) a" fun contains_bool :: "'a gexp \ bool" where "contains_bool (Bc _) = True" | "contains_bool (Nor g1 g2) = (contains_bool g1 \ contains_bool g2)" | "contains_bool _ = False" fun gexp_same_structure :: "'a gexp \ 'a gexp \ bool" where "gexp_same_structure (Bc b) (Bc b') = (b = b')" | "gexp_same_structure (Eq a1 a2) (Eq a1' a2') = (aexp_same_structure a1 a1' \ aexp_same_structure a2 a2')" | "gexp_same_structure (Gt a1 a2) (Gt a1' a2') = (aexp_same_structure a1 a1' \ aexp_same_structure a2 a2')" | "gexp_same_structure (Nor g1 g2) (Nor g1' g2') = (gexp_same_structure g1 g1' \ gexp_same_structure g2 g2')" | "gexp_same_structure (In v l) (In v' l') = (v = v' \ l = l')" | "gexp_same_structure _ _ = False" lemma gval_foldr_true: "(gval (foldr gAnd G (Bc True)) s = true) = (\g \ set G. gval g s = true)" proof(induct G) case (Cons a G) then show ?case apply (simp only: foldr.simps comp_def gval_gAnd maybe_and_true) by simp qed auto fun enumerate_gexp_inputs :: "vname gexp \ nat set" where "enumerate_gexp_inputs (Bc _) = {}" | "enumerate_gexp_inputs (Eq v va) = enumerate_aexp_inputs v \ enumerate_aexp_inputs va" | "enumerate_gexp_inputs (Gt v va) = enumerate_aexp_inputs v \ enumerate_aexp_inputs va" | "enumerate_gexp_inputs (In v va) = enumerate_aexp_inputs (V v)" | "enumerate_gexp_inputs (Nor v va) = enumerate_gexp_inputs v \ enumerate_gexp_inputs va" lemma enumerate_gexp_inputs_list: "\l. enumerate_gexp_inputs g = set l" proof(induct g) case (Eq x1a x2) then show ?case by (metis enumerate_aexp_inputs_list enumerate_gexp_inputs.simps(2) set_append) next case (Gt x1a x2) then show ?case by (metis enumerate_aexp_inputs_list enumerate_gexp_inputs.simps(3) set_append) next case (In x1a x2) then show ?case by (simp add: enumerate_aexp_inputs_list) next case (Nor g1 g2) then show ?case by (metis enumerate_gexp_inputs.simps(5) set_append) qed auto definition max_input :: "vname gexp \ nat option" where - "max_input g = (let inputs = (enumerate_gexp_inputs g) in if inputs = {} then None else Some (Max inputs))" + "max_input g = (let inputs = enumerate_gexp_inputs g in if inputs = {} then None else Some (Max inputs))" definition max_input_list :: "vname gexp list \ nat option" where - "max_input_list g = fold max (map (\g. max_input g) g) None" + "max_input_list g = fold max (map max_input g) None" lemma max_input_list_cons: "max_input_list (a # G) = max (max_input a) (max_input_list G)" - apply (simp add: max_input_list_def) - apply (cases "max_input a") - apply (simp add: max_def_raw) - by (metis (no_types, lifting) List.finite_set Max.insert Max.set_eq_fold fold_simps(1) list.set(2) max.assoc set_empty) + apply (cases \max_input a\) + apply (simp_all add: max_input_list_def flip: Max.set_eq_fold) + apply (metis (mono_tags, lifting) List.finite_set Max_insert Max_singleton bot_option_def finite_imageI max.assoc max_bot2) + done fun enumerate_regs :: "vname gexp \ nat set" where "enumerate_regs (Bc _) = {}" | "enumerate_regs (Eq v va) = AExp.enumerate_regs v \ AExp.enumerate_regs va" | "enumerate_regs (Gt v va) = AExp.enumerate_regs v \ AExp.enumerate_regs va" | "enumerate_regs (In v va) = AExp.enumerate_regs (V v)" | "enumerate_regs (Nor v va) = enumerate_regs v \ enumerate_regs va" lemma finite_enumerate_regs: "finite (enumerate_regs g)" using AExp.finite_enumerate_regs by (induct g, auto) definition max_reg :: "vname gexp \ nat option" where "max_reg g = (let regs = (enumerate_regs g) in if regs = {} then None else Some (Max regs))" lemma max_reg_gNot: "max_reg (gNot x) = max_reg x" by (simp add: max_reg_def gNot_def) lemma max_reg_Eq: "max_reg (Eq a b) = max (AExp.max_reg a) (AExp.max_reg b)" apply (simp add: max_reg_def AExp.max_reg_def Let_def max_absorb2) by (metis AExp.finite_enumerate_regs Max.union bot_option_def max_bot2 sup_Some sup_max) lemma max_reg_Gt: "max_reg (Gt a b) = max (AExp.max_reg a) (AExp.max_reg b)" apply (simp add: max_reg_def AExp.max_reg_def Let_def max_absorb2) by (metis AExp.finite_enumerate_regs Max.union bot_option_def max_bot2 sup_Some sup_max) lemma max_reg_Nor: "max_reg (Nor a b) = max (max_reg a) (max_reg b)" apply (simp add: max_reg_def AExp.max_reg_def Let_def max_absorb2) by (metis GExp.finite_enumerate_regs Max.union bot_option_def max_bot2 sup_Some sup_max) lemma gval_In_cons: "gval (In v (a # as)) s = (gval (Eq (V v) (L a)) s \? gval (In v as) s)" by (cases "s v", auto) lemma possible_to_be_in: "s \ [] \ satisfiable (In v s)" proof- assume "s \ []" have aux: "\v' i r. join_ir i r v = Some v' \ v' \ set s \ \i r. (case join_ir i r v of None \ false | Some v \ if v \ set s then true else false) = true" by (metis (mono_tags, lifting) option.simps(5)) show ?thesis apply (simp add: satisfiable_def gval_In_cons) apply (cases s) apply (simp add: \s \ []\) apply (cases v) apply (case_tac "\(i::value list). length i > x1 \ i ! x1 = a") apply clarsimp subgoal for _ _ i by (rule exI[of _ i], intro exI, simp) apply (metis gt_ex length_list_update length_repeat nth_list_update_eq) apply (rule_tac exI) apply (case_tac "\r. r $ x2 = Some a") apply clarsimp subgoal for _ _ _ r by (rule exI[of _ r], simp) by (metis join_ir_R join_ir_double_exists) qed definition max_reg_list :: "vname gexp list \ nat option" where - "max_reg_list g = (fold max (map (\g. max_reg g) g) None)" + "max_reg_list g = fold max (map max_reg g) None" lemma max_reg_list_cons: "max_reg_list (a # G) = max (max_reg a) (max_reg_list G)" - apply (simp add: max_reg_list_def) - by (metis (no_types, lifting) List.finite_set Max.insert Max.set_eq_fold fold.simps(1) id_apply list.simps(15) max.assoc set_empty) + apply (simp add: max_reg_list_def flip: Max.set_eq_fold) + apply (metis List.finite_set Max_insert bot_option_def empty_not_insert finite_imageI finite_insert insert_commute max.commute max_bot2) + done lemma max_reg_list_append_singleton: "max_reg_list (as@[bs]) = max (max_reg_list as) (max_reg_list [bs])" apply (simp add: max_reg_list_def) by (metis max.commute sup_None_2 sup_max) lemma max_reg_list_append: "max_reg_list (as@bs) = max (max_reg_list as) (max_reg_list bs)" proof(induct bs rule: rev_induct) case Nil then show ?case by (metis append_Nil2 fold_simps(1) list.simps(8) max_reg_list_def sup_None_2 sup_max) next case (snoc x xs) then show ?case by (metis append_assoc max.assoc max_reg_list_append_singleton) qed definition apply_guards :: "vname gexp list \ vname datastate \ bool" where "apply_guards G s = (\g \ set (map (\g. gval g s) G). g = true)" lemma apply_guards_singleton[simp]: "(apply_guards [g] s) = (gval g s = true)" by (simp add: apply_guards_def) lemma apply_guards_empty [simp]: "apply_guards [] s" by (simp add: apply_guards_def) lemma apply_guards_cons: "apply_guards (a # G) c = (gval a c = true \ apply_guards G c)" by (simp add: apply_guards_def) lemma apply_guards_double_cons: "apply_guards (y # x # G) s = (gval (gAnd y x) s = true \ apply_guards G s)" using apply_guards_cons gval_gAnd_True by blast lemma apply_guards_append: "apply_guards (a@a') s = (apply_guards a s \ apply_guards a' s)" using apply_guards_def by auto lemma apply_guards_foldr: "apply_guards G s = (gval (foldr gAnd G (Bc True)) s = true)" proof(induct G) case Nil then show ?case by (simp add: apply_guards_def) next case (Cons a G) then show ?case by (metis apply_guards_cons foldr.simps(2) gval_gAnd_True o_apply) qed lemma rev_apply_guards: "apply_guards (rev G) s = apply_guards G s" by (simp add: apply_guards_def) lemma apply_guards_fold: "apply_guards G s = (gval (fold gAnd G (Bc True)) s = true)" using rev_apply_guards[symmetric] by (simp add: foldr_conv_fold apply_guards_foldr) lemma fold_apply_guards: "(gval (fold gAnd G (Bc True)) s = true) = apply_guards G s" by (simp add: apply_guards_fold) lemma foldr_apply_guards: "(gval (foldr gAnd G (Bc True)) s = true) = apply_guards G s" by (simp add: apply_guards_foldr) lemma apply_guards_subset: "set g' \ set g \ apply_guards g c \ apply_guards g' c" proof(induct g) case (Cons a g) then show ?case using apply_guards_def by auto qed auto lemma apply_guards_subset_append: "set G \ set G' \ apply_guards (G @ G') s = apply_guards (G') s" using apply_guards_append apply_guards_subset by blast lemma apply_guards_rearrange: "x \ set G \ apply_guards G s = apply_guards (x#G) s" using apply_guards_def by auto lemma apply_guards_condense: "\g. apply_guards G s = (gval g s = true)" using apply_guards_fold by blast lemma apply_guards_false_condense: "\g. (\apply_guards G s) = (gval g s = false)" using foldr_apply_guards gval.simps(2) not_true by blast lemma max_input_Bc: "max_input (Bc x) = None" by (simp add: max_input_def) lemma max_input_Eq: "max_input (Eq a1 a2) = max (AExp.max_input a1) (AExp.max_input a2)" apply (simp add: AExp.max_input_def max_input_def Let_def max_absorb2) by (metis List.finite_set Max.union bot_option_def enumerate_aexp_inputs_not_empty max_bot2 sup_Some sup_max) lemma max_input_Gt: "max_input (Gt a1 a2) = max (AExp.max_input a1) (AExp.max_input a2)" apply (simp add: AExp.max_input_def max_input_def Let_def max_absorb2) by (metis List.finite_set Max.union bot_option_def enumerate_aexp_inputs_not_empty max_bot2 sup_Some sup_max) lemma gexp_max_input_Nor: "max_input (Nor g1 g2) = max (max_input g1) (max_input g2)" apply (simp add: AExp.max_input_def max_input_def Let_def max_absorb2) by (metis List.finite_set Max.union enumerate_gexp_inputs_list less_eq_option_Some_None max_def sup_Some sup_max) lemma gexp_max_input_In: "max_input (In v l) = AExp.max_input (V v)" by (simp add: AExp.max_input_def GExp.max_input_def) lemma gval_foldr_gOr_invalid: "(gval (fold gOr l g) s = invalid) = (\g' \ (set (g#l)). gval g' s = invalid)" proof(induct l rule: rev_induct) case (snoc x xs) then show ?case by (simp, metis gval_gOr maybe_or_invalid) qed auto lemma gval_foldr_gOr_true: "(gval (fold gOr l g) s = true) = ((\g' \ (set (g#l)). gval g' s = true) \ (\g' \ (set (g#l)). gval g' s \ invalid))" proof(induct l rule: rev_induct) case (snoc x xs) then show ?case apply (simp add: maybe_or_true) using gval_foldr_gOr_invalid by auto qed auto lemma gval_foldr_gOr_false: "(gval (fold gOr l g) s = false) = (\g' \ (set (g#l)). gval g' s = false)" proof(induct l rule: rev_induct) case (snoc x xs) then show ?case by (auto simp add: maybe_or_false) qed auto lemma gval_fold_gOr_rev: "gval (fold gOr (rev l) g) s = gval (fold gOr l g) s" apply (cases "gval (fold gOr l g) s") apply (simp, simp add: gval_foldr_gOr_true) apply (simp, simp add: gval_foldr_gOr_false) by (simp, simp add: gval_foldr_gOr_invalid) lemma gval_fold_gOr_foldr: "gval (fold gOr l g) s = gval (foldr gOr l g) s" by (simp add: foldr_conv_fold gval_fold_gOr_rev) lemma gval_fold_gOr: "gval (fold gOr (a # l) g) s = (gval a s \? gval (fold gOr l g) s)" by (simp only: gval_fold_gOr_foldr foldr.simps comp_def gval_gOr) lemma gval_In_fold: "gval (In v l) s = (if s v = None then invalid else gval (fold gOr (map (\x. Eq (V v) (L x)) l) (Bc False)) s)" proof(induct l) case Nil then show ?case apply simp apply (cases "s v") apply simp by auto next case (Cons a l) then show ?case apply (simp only: gval_In_cons) apply (cases "s v") apply simp by (simp add: gval_fold_gOr del: fold.simps) qed fun fold_In :: "'a \ value list \ 'a gexp" where "fold_In _ [] = Bc False" | "fold_In v (l#t) = gOr (Eq (V v) (L l)) (fold_In v t)" lemma gval_fold_In: "l \ [] \ gval (In v l) s = gval (fold_In v l) s" proof(induct l) next case (Cons a l) then show ?case apply (case_tac "s v") apply simp apply simp apply safe apply simp apply (metis fold_In.simps(1) gval.simps(2) plus_trilean.simps(4) plus_trilean.simps(5)) apply fastforce apply fastforce by fastforce qed auto lemma fold_maybe_or_invalid_base: "fold (\?) l invalid = invalid" proof(induct l) case (Cons a l) then show ?case by (metis fold_simps(2) maybe_or_valid) qed auto lemma fold_maybe_or_true_base_never_false: "fold (\?) l true \ false" proof(induct l) case (Cons a l) then show ?case by (metis fold_maybe_or_invalid_base fold_simps(2) maybe_not.cases maybe_or_valid plus_trilean.simps(4) plus_trilean.simps(6)) qed auto lemma fold_true_fold_false_not_invalid: "fold (\?) l true = true \ fold (\?) (rev l) false \ invalid" proof(induct l) case (Cons a l) then show ?case apply simp by (metis fold_maybe_or_invalid_base maybe_or_invalid maybe_or_true) qed auto lemma fold_true_invalid_fold_rev_false_invalid: "fold (\?) l true = invalid \ fold (\?) (rev l) false = invalid" proof(induct l) case (Cons a l) then show ?case apply simp by (metis maybe_or_true maybe_or_valid) qed auto lemma fold_maybe_or_rev: "fold (\?) l b = fold (\?) (rev l) b" proof(induct l) case (Cons a l) then show ?case proof(induction a b rule: plus_trilean.induct) case (1 uu) then show ?case by (simp add: fold_maybe_or_invalid_base) next case "2_1" then show ?case by (simp add: fold_maybe_or_invalid_base) next case "2_2" then show ?case by (simp add: fold_maybe_or_invalid_base) next case "3_1" then show ?case apply simp by (metis add.assoc fold_maybe_or_true_base_never_false maybe_not.cases maybe_or_idempotent maybe_or_true) next case "3_2" then show ?case apply simp apply (case_tac "fold (\?) l true") apply (simp add: eq_commute[of true]) apply (case_tac "fold (\?) (rev l) false") apply simp apply simp apply (simp add: fold_true_fold_false_not_invalid) apply (simp add: fold_maybe_or_true_base_never_false) by (simp add: fold_true_invalid_fold_rev_false_invalid) next case 4 then show ?case by (simp add: maybe_or_zero) next case 5 then show ?case by (simp add: maybe_or_zero) qed qed auto lemma fold_maybe_or_cons: "fold (\?) (a#l) b = a \? (fold (\?) l b)" by (metis fold_maybe_or_rev foldr.simps(2) foldr_conv_fold o_apply) lemma gval_fold_gOr_map: "gval (fold gOr l (Bc False)) s = fold (\?) (map (\g. gval g s) l) (false)" proof(induct l) case (Cons a l) then show ?case by (metis fold_maybe_or_cons gval_fold_gOr list.simps(9)) qed auto lemma gval_unfold_first: "gval (fold gOr (map (\x. Eq (V v) (L x)) ls) (Eq (V v) (L l))) s = gval (fold gOr (map (\x. Eq (V v) (L x)) (l#ls)) (Bc False)) s" proof(induct ls) case Nil then show ?case apply (cases "s v") apply simp by (simp add: gOr_def) next case (Cons a ls) then show ?case proof - have "gval (fold gOr (map (\va. Eq (V v) (L va)) ls) (gOr (Eq (V v) (L l)) (Bc False))) s = gval (fold gOr (map (\va. Eq (V v) (L va)) (l # ls)) (Bc False)) s" by simp then have "gval (fold gOr (map (\va. Eq (V v) (L va)) (a # ls)) (Eq (V v) (L l))) s = gval (fold gOr (Eq (V v) (L a) # map (\va. Eq (V v) (L va)) ls) (gOr (Eq (V v) (L l)) (Bc False))) s" by (metis (no_types) Cons.hyps gval_fold_gOr list.simps(9)) then show ?thesis by force qed qed lemma fold_Eq_true: "\v. fold (\?) (map (\x. if v = x then true else false) vs) true = true" by(induct vs, auto) lemma x_in_set_fold_eq: "x \ set ll \ fold (\?) (map (\xa. if x = xa then true else false) ll) false = true" proof(induct ll) case (Cons a ll) then show ?case apply simp apply standard apply (simp add: fold_Eq_true) by auto qed auto lemma x_not_in_set_fold_eq: "s v \ Some ` set ll \ false = fold (\?) (map (\x. if s v = Some x then true else false) ll) false" by(induct ll, auto) lemma gval_take: "max_input g < Some a \ gval g (join_ir i r) = gval g (join_ir (take a i) r)" proof(induct g) case (Bc x) then show ?case by (metis (full_types) gval.simps(1) gval.simps(2)) next case (Eq x1a x2) then show ?case by (metis aval_take gval.simps(4) max_input_Eq max_less_iff_conj) next case (Gt x1a x2) then show ?case by (metis aval_take gval.simps(3) max_input_Gt max_less_iff_conj) next case (Nor g1 g2) then show ?case by (simp add: maybe_not_eq gexp_max_input_Nor) next case (In v l) then show ?case apply (simp add: gexp_max_input_In) using aval_take by fastforce qed lemma gval_fold_gAnd_append_singleton: "gval (fold gAnd (a @ [G]) (Bc True)) s = gval (fold gAnd a (Bc True)) s \? gval G s" apply simp using times_trilean_commutative by blast lemma gval_fold_rev_true: "gval (fold gAnd (rev G) (Bc True)) s = true \ gval (fold gAnd G (Bc True)) s = true" by (metis foldr_conv_fold gval_foldr_true rev_rev_ident set_rev) lemma gval_fold_not_invalid_all_valid_contra: "\g \ set G. gval g s = invalid \ gval (fold gAnd G (Bc True)) s = invalid" proof(induct G rule: rev_induct) case (snoc a G) then show ?case apply (simp only: gval_fold_gAnd_append_singleton) apply simp using maybe_and_valid by blast qed auto lemma gval_fold_not_invalid_all_valid: "gval (fold gAnd G (Bc True)) s \ invalid \ \g \ set G. gval g s \ invalid" using gval_fold_not_invalid_all_valid_contra by blast lemma all_gval_not_false: "(\g \ set G. gval g s \ false) = (\g \ set G. gval g s = true) \ (\g \ set G. gval g s = invalid)" using trilean.exhaust by auto lemma must_have_one_false_contra: "\g \ set G. gval g s \ false \ gval (fold gAnd G (Bc True)) s \ false" using all_gval_not_false[of G s] apply simp apply (case_tac "(\g\set G. gval g s = true)") apply (metis (full_types) foldr_conv_fold gval_fold_rev_true gval_foldr_true not_true) by (simp add: gval_fold_not_invalid_all_valid_contra) lemma must_have_one_false: "gval (fold gAnd G (Bc True)) s = false \ \g \ set G. gval g s = false" using must_have_one_false_contra by blast lemma all_valid_fold: "\g \ set G. gval g s \ invalid \ gval (fold gAnd G (Bc True)) s \ invalid" apply (induct G rule: rev_induct) apply simp by (simp add: maybe_and_invalid) lemma one_false_all_valid_false: "\g\set G. gval g s = false \ \g\set G. gval g s \ invalid \ gval (fold gAnd G (Bc True)) s = false" by (metis (full_types) all_valid_fold foldr_conv_fold gval_foldr_true not_true rev_rev_ident set_rev) lemma gval_fold_rev_false: "gval (fold gAnd (rev G) (Bc True)) s = false \ gval (fold gAnd G (Bc True)) s = false" using must_have_one_false[of "rev G" s] gval_fold_not_invalid_all_valid[of "rev G" s] by (simp add: one_false_all_valid_false) lemma fold_invalid_means_one_invalid: "gval (fold gAnd G (Bc True)) s = invalid \ \g \ set G. gval g s = invalid" using all_valid_fold by blast lemma gval_fold_rev_invalid: "gval (fold gAnd (rev G) (Bc True)) s = invalid \ gval (fold gAnd G (Bc True)) s = invalid" using fold_invalid_means_one_invalid[of "rev G" s] by (simp add: gval_fold_not_invalid_all_valid_contra) lemma gval_fold_rev_equiv_fold: "gval (fold gAnd (rev G) (Bc True)) s = gval (fold gAnd G (Bc True)) s" apply (cases "gval (fold gAnd (rev G) (Bc True)) s") apply (simp add: gval_fold_rev_true) apply (simp add: gval_fold_rev_false) by (simp add: gval_fold_rev_invalid) lemma gval_fold_equiv_fold_rev: "gval (fold gAnd G (Bc True)) s = gval (fold gAnd (rev G) (Bc True)) s" by (simp add: gval_fold_rev_equiv_fold) lemma gval_fold_equiv_gval_foldr: "gval (fold gAnd G (Bc True)) s = gval (foldr gAnd G (Bc True)) s" proof - have "gval (fold gAnd G (Bc True)) s = gval (fold gAnd (rev G) (Bc True)) s" using gval_fold_equiv_fold_rev by force then show ?thesis by (simp add: foldr_conv_fold) qed lemma gval_foldr_equiv_gval_fold: "gval (foldr gAnd G (Bc True)) s = gval (fold gAnd G (Bc True)) s" by (simp add: gval_fold_equiv_gval_foldr) lemma gval_fold_cons: "gval (fold gAnd (g # gs) (Bc True)) s = gval g s \? gval (fold gAnd gs (Bc True)) s" apply (simp only: apply_guards_fold gval_fold_equiv_gval_foldr) by (simp only: foldr.simps comp_def gval_gAnd) lemma gval_fold_take: "max_input_list G < Some a \ a \ length i \ max_input_list G \ Some (length i) \ gval (fold gAnd G (Bc True)) (join_ir i r) = gval (fold gAnd G (Bc True)) (join_ir (take a i) r)" proof(induct G) case (Cons g gs) then show ?case apply (simp only: gval_fold_cons) apply (simp add: max_input_list_cons) using gval_take[of g a i r] by simp qed auto primrec padding :: "nat \ 'a list" where "padding 0 = []" | "padding (Suc m) = (Eps (\x. True))#(padding m)" definition take_or_pad :: "'a list \ nat \ 'a list" where "take_or_pad a n = (if length a \ n then take n a else a@(padding (n-length a)))" lemma length_padding: "length (padding n) = n" by (induct n, auto) lemma length_take_or_pad: "length (take_or_pad a n) = n" -proof(induct n) - case 0 - then show ?case - by (simp add: take_or_pad_def) -next - case (Suc n) - then show ?case - apply (simp add: take_or_pad_def) - apply standard - apply auto[1] - by (simp add: length_padding) -qed + by (induction n) (simp_all add: take_or_pad_def length_padding) fun enumerate_gexp_strings :: "'a gexp \ String.literal set" where "enumerate_gexp_strings (Bc _) = {}" | "enumerate_gexp_strings (Eq a1 a2) = enumerate_aexp_strings a1 \ enumerate_aexp_strings a2" | "enumerate_gexp_strings (Gt a1 a2) = enumerate_aexp_strings a1 \ enumerate_aexp_strings a2" | "enumerate_gexp_strings (In v l) = fold (\x acc. case x of Num n \ acc | Str s \ insert s acc) l {}" | "enumerate_gexp_strings (Nor g1 g2) = enumerate_gexp_strings g1 \ enumerate_gexp_strings g2" fun enumerate_gexp_ints :: "'a gexp \ int set" where "enumerate_gexp_ints (Bc _) = {}" | "enumerate_gexp_ints (Eq a1 a2) = enumerate_aexp_ints a1 \ enumerate_aexp_ints a2" | "enumerate_gexp_ints (Gt a1 a2) = enumerate_aexp_ints a1 \ enumerate_aexp_ints a2" | "enumerate_gexp_ints (In v l) = fold (\x acc. case x of Str s \ acc | Num n \ insert n acc) l {}" | "enumerate_gexp_ints (Nor g1 g2) = enumerate_gexp_ints g1 \ enumerate_gexp_ints g2" definition restricted_once :: "'a \ 'a gexp list \ bool" where "restricted_once v G = (length (filter (\g. gexp_constrains g (V v)) G) = 1)" definition not_restricted :: "'a \ 'a gexp list \ bool" where "not_restricted v G = (length (filter (\g. gexp_constrains g (V v)) G) = 0)" lemma restricted_once_cons: "restricted_once v (g#gs) = ((gexp_constrains g (V v) \ not_restricted v gs) \ ((\ gexp_constrains g (V v)) \ restricted_once v gs))" by (simp add: restricted_once_def not_restricted_def) lemma not_restricted_cons: "not_restricted v (g#gs) = ((\ gexp_constrains g (V v)) \ not_restricted v gs)" by (simp add: not_restricted_def) definition enumerate_vars :: "vname gexp \ vname list" where "enumerate_vars g = sorted_list_of_set ((image R (enumerate_regs g)) \ (image I (enumerate_gexp_inputs g)))" fun rename_regs :: "(nat \ nat) \ vname gexp \ vname gexp" where "rename_regs _ (Bc b) = Bc b" | "rename_regs f (Eq a1 a2) = Eq (AExp.rename_regs f a1) (AExp.rename_regs f a2)" | "rename_regs f (Gt a1 a2) = Gt (AExp.rename_regs f a1) (AExp.rename_regs f a2)" | "rename_regs f (In (R r) vs) = In (R (f r)) vs" | "rename_regs f (In v vs) = In v vs" | "rename_regs f (Nor g1 g2) = Nor (rename_regs f g1) (rename_regs f g2)" definition eq_upto_rename :: "vname gexp \ vname gexp \ bool" where "eq_upto_rename g1 g2 = (\f. bij f \ rename_regs f g1 = g2)" lemma gval_reg_some_superset: "\a. (r $ a \ None) \ r $ a = r' $ a \ x \ invalid \ gval a (join_ir i r) = x \ gval a (join_ir i r') = x" proof(induct a arbitrary: x) case (Bc b) then show ?case by (cases b, auto) next case (Eq x1a x2) then show ?case apply (cases x) apply simp using value_eq_true[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"] apply clarsimp apply (simp add: aval_reg_some_superset) apply simp using value_eq_false[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"] apply clarsimp apply (simp add: aval_reg_some_superset) by simp next case (Gt x1a x2) then show ?case apply (cases x) apply simp using value_gt_true_Some[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"] apply clarsimp apply (simp add: aval_reg_some_superset) apply simp using value_gt_false_Some[of "aval x1a (join_ir i r)" "aval x2 (join_ir i r)"] apply clarsimp apply (simp add: aval_reg_some_superset) by simp next case (In x1a x2) then show ?case apply simp apply (case_tac "join_ir i r x1a") apply simp apply (case_tac "join_ir i r' x1a") apply simp apply (metis aval_reg_some_superset In.prems(1) aval.simps(2) option.distinct(1)) apply simp by (metis (full_types) aval_reg_some_superset In.prems(1) aval.simps(2) option.inject) next case (Nor a1 a2) then show ?case apply simp apply (cases x) apply (simp add: maybe_negate_true maybe_or_false) apply (simp add: maybe_negate_false maybe_or_true) apply presburger by simp qed lemma apply_guards_reg_some_superset: "\a. (r $ a \ None) \ r $ a = r' $ a \ apply_guards G (join_ir i r) \ apply_guards G (join_ir i r')" apply (induct G) apply simp apply (simp add: apply_guards_cons) using gval_reg_some_superset by simp end diff --git a/thys/Gromov_Hyperbolicity/Gromov_Hyperbolicity.thy b/thys/Gromov_Hyperbolicity/Gromov_Hyperbolicity.thy --- a/thys/Gromov_Hyperbolicity/Gromov_Hyperbolicity.thy +++ b/thys/Gromov_Hyperbolicity/Gromov_Hyperbolicity.thy @@ -1,1308 +1,1308 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Gromov hyperbolic spaces\ theory Gromov_Hyperbolicity imports Isometries Metric_Completion begin subsection \Definition, basic properties\ text \Although we will mainly work with type classes later on, we introduce the definition of hyperbolicity on subsets of a metric space. A set is $\delta$-hyperbolic if it satisfies the following inequality. It is very obscure at first sight, but we will see several equivalent characterizations later on. For instance, a space is hyperbolic (maybe for a different constant $\delta$) if all geodesic triangles are thin, i.e., every side is close to the union of the two other sides. This definition captures the main features of negative curvature at a large scale, and has proved extremely fruitful and influential. Two important references on this topic are~\cite{ghys_hyperbolique} and~\cite{bridson_haefliger}. We will sometimes follow them, sometimes depart from them.\ definition Gromov_hyperbolic_subset::"real \ ('a::metric_space) set \ bool" where "Gromov_hyperbolic_subset delta A = (\x\A. \y\A. \z\A. \t\A. dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * delta)" lemma Gromov_hyperbolic_subsetI [intro]: assumes "\x y z t. x \ A \ y \ A \ z \ A \ t \ A \ dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * delta" shows "Gromov_hyperbolic_subset delta A" using assms unfolding Gromov_hyperbolic_subset_def by auto text \When the four points are not all distinct, the above inequality is always satisfied for $\delta = 0$.\ lemma Gromov_hyperbolic_ineq_not_distinct: assumes "x = y \ x = z \ x = t \ y = z \ y = t \ z = (t::'a::metric_space)" shows "dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z)" using assms by (auto simp add: dist_commute, simp add: dist_triangle add.commute, simp add: dist_triangle3) text \It readily follows from the definition that hyperbolicity passes to the closure of the set.\ lemma Gromov_hyperbolic_closure: assumes "Gromov_hyperbolic_subset delta A" shows "Gromov_hyperbolic_subset delta (closure A)" unfolding Gromov_hyperbolic_subset_def proof (auto) fix x y z t assume H: "x \ closure A" "y \ closure A" "z \ closure A" "t \ closure A" obtain X::"nat \ 'a" where X: "\n. X n \ A" "X \ x" using H closure_sequential by blast obtain Y::"nat \ 'a" where Y: "\n. Y n \ A" "Y \ y" using H closure_sequential by blast obtain Z::"nat \ 'a" where Z: "\n. Z n \ A" "Z \ z" using H closure_sequential by blast obtain T::"nat \ 'a" where T: "\n. T n \ A" "T \ t" using H closure_sequential by blast have *: "max (dist (X n) (Z n) + dist (Y n) (T n)) (dist (X n) (T n) + dist (Y n) (Z n)) + 2 * delta - dist (X n) (Y n) - dist (Z n) (T n) \ 0" for n using assms X(1)[of n] Y(1)[of n] Z(1)[of n] T(1)[of n] unfolding Gromov_hyperbolic_subset_def by (auto simp add: algebra_simps) have **: "(\n. max (dist (X n) (Z n) + dist (Y n) (T n)) (dist (X n) (T n) + dist (Y n) (Z n)) + 2 * delta - dist (X n) (Y n) - dist (Z n) (T n)) \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * delta - dist x y - dist z t" apply (auto intro!: tendsto_intros) using X Y Z T by auto have "max (dist x z + dist y t) (dist x t + dist y z) + 2 * delta - dist x y - dist z t \ 0" apply (rule LIMSEQ_le_const[OF **]) using * by auto then show "dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * delta" by auto qed text \A good formulation of hyperbolicity is in terms of Gromov products. Intuitively, the Gromov product of $x$ and $y$ based at $e$ is the distance between $e$ and the geodesic between $x$ and $y$. It is also the time after which the geodesics from $e$ to $x$ and from $e$ to $y$ stop travelling together.\ definition Gromov_product_at::"('a::metric_space) \ 'a \ 'a \ real" where "Gromov_product_at e x y = (dist e x + dist e y - dist x y) / 2" lemma Gromov_hyperbolic_subsetI2: fixes delta::real assumes "\e x y z. e \ A \ x \ A \ y \ A \ z \ A \ Gromov_product_at (e::'a::metric_space) x z \ min (Gromov_product_at e x y) (Gromov_product_at e y z) - delta" shows "Gromov_hyperbolic_subset delta A" proof (rule Gromov_hyperbolic_subsetI) fix x y z t assume H: "x \ A" "z \ A" "y \ A" "t \ A" show "dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * delta" using assms[OF H] unfolding Gromov_product_at_def min_def max_def by (auto simp add: divide_simps algebra_simps dist_commute) qed lemma Gromov_product_nonneg [simp, mono_intros]: "Gromov_product_at e x y \ 0" unfolding Gromov_product_at_def by (simp add: dist_triangle3) lemma Gromov_product_commute: "Gromov_product_at e x y = Gromov_product_at e y x" unfolding Gromov_product_at_def by (auto simp add: dist_commute) lemma Gromov_product_le_dist [simp, mono_intros]: "Gromov_product_at e x y \ dist e x" "Gromov_product_at e x y \ dist e y" unfolding Gromov_product_at_def by (auto simp add: diff_le_eq dist_triangle dist_triangle2) lemma Gromov_product_le_infdist [mono_intros]: assumes "geodesic_segment_between G x y" shows "Gromov_product_at e x y \ infdist e G" proof - have [simp]: "G \ {}" using assms by auto have "Gromov_product_at e x y \ dist e z" if "z \ G" for z proof - have "dist e x + dist e y \ (dist e z + dist z x) + (dist e z + dist z y)" by (intro add_mono dist_triangle) also have "... = 2 * dist e z + dist x y" apply (auto simp add: dist_commute) using \z \ G\ assms by (metis dist_commute geodesic_segment_dist) finally show ?thesis unfolding Gromov_product_at_def by auto qed then show ?thesis apply (subst infdist_notempty) by (auto intro: cINF_greatest) qed lemma Gromov_product_add: "Gromov_product_at e x y + Gromov_product_at x e y = dist e x" unfolding Gromov_product_at_def by (auto simp add: algebra_simps divide_simps dist_commute) lemma Gromov_product_geodesic_segment: assumes "geodesic_segment_between G x y" "t \ {0..dist x y}" shows "Gromov_product_at x y (geodesic_segment_param G x t) = t" proof - have "dist x (geodesic_segment_param G x t) = t" using assms(1) assms(2) geodesic_segment_param(6) by auto moreover have "dist y (geodesic_segment_param G x t) = dist x y - t" by (metis \dist x (geodesic_segment_param G x t) = t\ add_diff_cancel_left' assms(1) assms(2) dist_commute geodesic_segment_dist geodesic_segment_param(3)) ultimately show ?thesis unfolding Gromov_product_at_def by auto qed lemma Gromov_product_e_x_x [simp]: "Gromov_product_at e x x = dist e x" unfolding Gromov_product_at_def by auto lemma Gromov_product_at_diff: "\Gromov_product_at x y z - Gromov_product_at a b c\ \ dist x a + dist y b + dist z c" unfolding Gromov_product_at_def abs_le_iff apply (auto simp add: divide_simps) by (smt dist_commute dist_triangle4)+ lemma Gromov_product_at_diff1: "\Gromov_product_at a x y - Gromov_product_at b x y\ \ dist a b" using Gromov_product_at_diff[of a x y b x y] by auto lemma Gromov_product_at_diff2: "\Gromov_product_at e x z - Gromov_product_at e y z\ \ dist x y" using Gromov_product_at_diff[of e x z e y z] by auto lemma Gromov_product_at_diff3: "\Gromov_product_at e x y - Gromov_product_at e x z\ \ dist y z" using Gromov_product_at_diff[of e x y e x z] by auto text \The Gromov product is continuous in its three variables. We formulate it in terms of sequences, as it is the way it will be used below (and moreover continuity for functions of several variables is very poor in the library).\ lemma Gromov_product_at_continuous: assumes "(u \ x) F" "(v \ y) F" "(w \ z) F" shows "((\n. Gromov_product_at (u n) (v n) (w n)) \ Gromov_product_at x y z) F" proof - have "((\n. abs(Gromov_product_at (u n) (v n) (w n) - Gromov_product_at x y z)) \ 0 + 0 + 0) F" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. dist (u n) x + dist (v n) y + dist (w n) z", OF always_eventually always_eventually]) apply (simp, simp add: Gromov_product_at_diff, simp, intro tendsto_intros) using assms tendsto_dist_iff by auto then show ?thesis apply (subst tendsto_dist_iff) unfolding dist_real_def by auto qed subsection \Typeclass for Gromov hyperbolic spaces\ text \We could (should?) just derive \verb+Gromov_hyperbolic_space+ from \verb+metric_space+. However, in this case, properties of metric spaces are not available when working in the locale! It is more efficient to ensure that we have a metric space by putting a type class restriction in the definition. The $\delta$ in Gromov-hyperbolicity type class is called \verb+deltaG+ to avoid name clashes. \ class metric_space_with_deltaG = metric_space + fixes deltaG::"('a::metric_space) itself \ real" class Gromov_hyperbolic_space = metric_space_with_deltaG + assumes hyperb_quad_ineq0: "Gromov_hyperbolic_subset (deltaG(TYPE('a::metric_space))) (UNIV::'a set)" class Gromov_hyperbolic_space_geodesic = Gromov_hyperbolic_space + geodesic_space lemma (in Gromov_hyperbolic_space) hyperb_quad_ineq [mono_intros]: shows "dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * deltaG(TYPE('a))" using hyperb_quad_ineq0 unfolding Gromov_hyperbolic_subset_def by auto text \It readily follows from the definition that the completion of a $\delta$-hyperbolic space is still $\delta$-hyperbolic.\ instantiation metric_completion :: (Gromov_hyperbolic_space) Gromov_hyperbolic_space begin definition deltaG_metric_completion::"('a metric_completion) itself \ real" where "deltaG_metric_completion _ = deltaG(TYPE('a))" instance proof (standard, rule Gromov_hyperbolic_subsetI) have "Gromov_hyperbolic_subset (deltaG(TYPE('a))) (range (to_metric_completion::'a \ _))" unfolding Gromov_hyperbolic_subset_def apply (auto simp add: isometry_onD[OF to_metric_completion_isometry]) by (metis hyperb_quad_ineq) then have "Gromov_hyperbolic_subset (deltaG TYPE('a metric_completion)) (UNIV::'a metric_completion set)" unfolding deltaG_metric_completion_def to_metric_completion_dense'[symmetric] using Gromov_hyperbolic_closure by auto then show "dist x y + dist z t \ max (dist x z + dist y t) (dist x t + dist y z) + 2 * deltaG TYPE('a metric_completion)" for x y z t::"'a metric_completion" unfolding Gromov_hyperbolic_subset_def by auto qed end (*of instantiation metric_completion (of Gromov_hyperbolic_space) is Gromov_hyperbolic*) context Gromov_hyperbolic_space begin lemma delta_nonneg [simp, mono_intros]: "deltaG(TYPE('a)) \ 0" proof - obtain x::'a where True by auto show ?thesis using hyperb_quad_ineq[of x x x x] by auto qed lemma hyperb_ineq [mono_intros]: "Gromov_product_at (e::'a) x z \ min (Gromov_product_at e x y) (Gromov_product_at e y z) - deltaG(TYPE('a))" using hyperb_quad_ineq[of e y x z] unfolding Gromov_product_at_def min_def max_def by (auto simp add: divide_simps algebra_simps metric_space_class.dist_commute) lemma hyperb_ineq' [mono_intros]: "Gromov_product_at (e::'a) x z + deltaG(TYPE('a)) \ min (Gromov_product_at e x y) (Gromov_product_at e y z)" using hyperb_ineq[of e x y z] by auto lemma hyperb_ineq_4_points [mono_intros]: "Min {Gromov_product_at (e::'a) x y, Gromov_product_at e y z, Gromov_product_at e z t} - 2 * deltaG(TYPE('a)) \ Gromov_product_at e x t" using hyperb_ineq[of e x y z] hyperb_ineq[of e x z t] apply auto using delta_nonneg by linarith lemma hyperb_ineq_4_points' [mono_intros]: "Min {Gromov_product_at (e::'a) x y, Gromov_product_at e y z, Gromov_product_at e z t} \ Gromov_product_at e x t + 2 * deltaG(TYPE('a))" using hyperb_ineq_4_points[of e x y z t] by auto text \In Gromov-hyperbolic spaces, geodesic triangles are thin, i.e., a point on one side of a geodesic triangle is close to the union of the two other sides (where the constant in "close" is $4\delta$, independent of the size of the triangle). We prove this basic property (which, in fact, is a characterization of Gromov-hyperbolic spaces: a geodesic space in which triangles are thin is hyperbolic).\ lemma thin_triangles1: assumes "geodesic_segment_between G x y" "geodesic_segment_between H x (z::'a)" "t \ {0..Gromov_product_at x y z}" shows "dist (geodesic_segment_param G x t) (geodesic_segment_param H x t) \ 4 * deltaG(TYPE('a))" proof - have *: "Gromov_product_at x z (geodesic_segment_param H x t) = t" apply (rule Gromov_product_geodesic_segment[OF assms(2)]) using assms(3) Gromov_product_le_dist(2) by (metis atLeastatMost_subset_iff subset_iff) have "Gromov_product_at x y (geodesic_segment_param H x t) \ min (Gromov_product_at x y z) (Gromov_product_at x z (geodesic_segment_param H x t)) - deltaG(TYPE('a))" by (rule hyperb_ineq) then have I: "Gromov_product_at x y (geodesic_segment_param H x t) \ t - deltaG(TYPE('a))" using assms(3) unfolding * by auto have *: "Gromov_product_at x (geodesic_segment_param G x t) y = t" apply (subst Gromov_product_commute) apply (rule Gromov_product_geodesic_segment[OF assms(1)]) using assms(3) Gromov_product_le_dist(1) by (metis atLeastatMost_subset_iff subset_iff) have "t - 2 * deltaG(TYPE('a)) = min t (t- deltaG(TYPE('a))) - deltaG(TYPE('a))" unfolding min_def using antisym by fastforce also have "... \ min (Gromov_product_at x (geodesic_segment_param G x t) y) (Gromov_product_at x y (geodesic_segment_param H x t)) - deltaG(TYPE('a))" - using I * by auto + using I * by (simp add: algebra_simps) also have "... \ Gromov_product_at x (geodesic_segment_param G x t) (geodesic_segment_param H x t)" by (rule hyperb_ineq) finally have I: "Gromov_product_at x (geodesic_segment_param G x t) (geodesic_segment_param H x t) \ t - 2 * deltaG(TYPE('a))" by simp have A: "dist x (geodesic_segment_param G x t) = t" by (meson assms(1) assms(3) atLeastatMost_subset_iff geodesic_segment_param(6) Gromov_product_le_dist(1) subset_eq) have B: "dist x (geodesic_segment_param H x t) = t" by (meson assms(2) assms(3) atLeastatMost_subset_iff geodesic_segment_param(6) Gromov_product_le_dist(2) subset_eq) show ?thesis using I unfolding Gromov_product_at_def A B by auto qed theorem thin_triangles: assumes "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" "geodesic_segment_between Gyz y z" "(w::'a) \ Gyz" shows "infdist w (Gxy \ Gxz) \ 4 * deltaG(TYPE('a))" proof - obtain t where w: "t \ {0..dist y z}" "w = geodesic_segment_param Gyz y t" using geodesic_segment_param[OF assms(3)] assms(4) by (metis imageE) show ?thesis proof (cases "t \ Gromov_product_at y x z") case True have *: "dist w (geodesic_segment_param Gxy y t) \ 4 * deltaG(TYPE('a))" unfolding w(2) apply (rule thin_triangles1[of _ _ z _ x]) using True assms(1) assms(3) w(1) by (auto simp add: geodesic_segment_commute Gromov_product_commute) show ?thesis apply (rule infdist_le2[OF _ *]) by (metis True assms(1) box_real(2) geodesic_segment_commute geodesic_segment_param(3) Gromov_product_le_dist(1) mem_box_real(2) order_trans subset_eq sup.cobounded1 w(1)) next case False define s where "s = dist y z - t" have s: "s \ {0..Gromov_product_at z y x}" unfolding s_def using Gromov_product_add[of y z x] w(1) False by (auto simp add: Gromov_product_commute) have w2: "w = geodesic_segment_param Gyz z s" unfolding s_def w(2) apply (rule geodesic_segment_reverse_param[symmetric]) using assms(3) w(1) by auto have *: "dist w (geodesic_segment_param Gxz z s) \ 4 * deltaG(TYPE('a))" unfolding w2 apply (rule thin_triangles1[of _ _ y _ x]) using s assms by (auto simp add: geodesic_segment_commute) show ?thesis apply (rule infdist_le2[OF _ *]) by (metis Un_iff assms(2) atLeastAtMost_iff geodesic_segment_commute geodesic_segment_param(3) Gromov_product_commute Gromov_product_le_dist(1) order_trans s) qed qed text \A consequence of the thin triangles property is that, although the geodesic between two points is in general not unique in a Gromov-hyperbolic space, two such geodesics are within $O(\delta)$ of each other.\ lemma geodesics_nearby: assumes "geodesic_segment_between G x y" "geodesic_segment_between H x y" "(z::'a) \ G" shows "infdist z H \ 4 * deltaG(TYPE('a))" using thin_triangles[OF geodesic_segment_between_x_x(1) assms(2) assms(1) assms(3)] geodesic_segment_endpoints(1)[OF assms(2)] insert_absorb by fastforce text \A small variant of the property of thin triangles is that triangles are slim, i.e., there is a point which is close to the three sides of the triangle (a "center" of the triangle, but only defined up to $O(\delta)$). And one can take it on any side, and its distance to the corresponding vertices is expressed in terms of a Gromov product.\ lemma slim_triangle: assumes "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" "geodesic_segment_between Gyz y (z::'a)" shows "\w. infdist w Gxy \ 4 * deltaG(TYPE('a)) \ infdist w Gxz \ 4 * deltaG(TYPE('a)) \ infdist w Gyz \ 4 * deltaG(TYPE('a)) \ dist w x = (Gromov_product_at x y z) \ w \ Gxy" proof - define w where "w = geodesic_segment_param Gxy x (Gromov_product_at x y z)" have "w \ Gxy" unfolding w_def by (rule geodesic_segment_param(3)[OF assms(1)], auto) then have xy: "infdist w Gxy \ 4 * deltaG(TYPE('a))" by simp have *: "dist w x = (Gromov_product_at x y z)" unfolding w_def using assms(1) by (metis Gromov_product_le_dist(1) Gromov_product_nonneg atLeastAtMost_iff geodesic_segment_param(6) metric_space_class.dist_commute) define w2 where "w2 = geodesic_segment_param Gxz x (Gromov_product_at x y z)" have "w2 \ Gxz" unfolding w2_def by (rule geodesic_segment_param(3)[OF assms(2)], auto) moreover have "dist w w2 \ 4 * deltaG(TYPE('a))" unfolding w_def w2_def by (rule thin_triangles1[OF assms(1) assms(2)], auto) ultimately have xz: "infdist w Gxz \ 4 * deltaG(TYPE('a))" using infdist_le2 by blast have "w = geodesic_segment_param Gxy y (dist x y - Gromov_product_at x y z)" unfolding w_def by (rule geodesic_segment_reverse_param[OF assms(1), symmetric], auto) then have w: "w = geodesic_segment_param Gxy y (Gromov_product_at y x z)" using Gromov_product_add[of x y z] by (metis add_diff_cancel_left') define w3 where "w3 = geodesic_segment_param Gyz y (Gromov_product_at y x z)" have "w3 \ Gyz" unfolding w3_def by (rule geodesic_segment_param(3)[OF assms(3)], auto) moreover have "dist w w3 \ 4 * deltaG(TYPE('a))" unfolding w w3_def by (rule thin_triangles1[OF geodesic_segment_commute[OF assms(1)] assms(3)], auto) ultimately have yz: "infdist w Gyz \ 4 * deltaG(TYPE('a))" using infdist_le2 by blast show ?thesis using xy xz yz * \w \ Gxy\ by force qed text \The distance of a vertex of a triangle to the opposite side is essentially given by the Gromov product, up to $2\delta$.\ lemma dist_triangle_side_middle: assumes "geodesic_segment_between G x (y::'a)" shows "dist z (geodesic_segment_param G x (Gromov_product_at x z y)) \ Gromov_product_at z x y + 2 * deltaG(TYPE('a))" proof - define m where "m = geodesic_segment_param G x (Gromov_product_at x z y)" have "m \ G" unfolding m_def using assms(1) by auto have A: "dist x m = Gromov_product_at x z y" unfolding m_def by (rule geodesic_segment_param(6)[OF assms(1)], auto) have B: "dist y m = dist x y - dist x m" using geodesic_segment_dist[OF assms \m \ G\] by (auto simp add: metric_space_class.dist_commute) have *: "dist x z + dist y m = Gromov_product_at z x y + dist x y" "dist x m + dist y z = Gromov_product_at z x y + dist x y" unfolding B A Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute divide_simps) have "dist x y + dist z m \ max (dist x z + dist y m) (dist x m + dist y z) + 2 * deltaG(TYPE('a))" by (rule hyperb_quad_ineq) then have "dist z m \ Gromov_product_at z x y + 2 * deltaG(TYPE('a))" unfolding * by auto then show ?thesis unfolding m_def by auto qed lemma infdist_triangle_side [mono_intros]: assumes "geodesic_segment_between G x (y::'a)" shows "infdist z G \ Gromov_product_at z x y + 2 * deltaG(TYPE('a))" proof - have "infdist z G \ dist z (geodesic_segment_param G x (Gromov_product_at x z y))" using assms by (auto intro!: infdist_le) then show ?thesis using dist_triangle_side_middle[OF assms, of z] by auto qed text \The distance of a point on a side of triangle to the opposite vertex is controlled by the length of the opposite sides, up to $\delta$.\ lemma dist_le_max_dist_triangle: assumes "geodesic_segment_between G x y" "m \ G" shows "dist m z \ max (dist x z) (dist y z) + deltaG(TYPE('a))" proof - consider "dist m x \ deltaG(TYPE('a))" | "dist m y \ deltaG(TYPE('a))" | "dist m x \ deltaG(TYPE('a)) \ dist m y \ deltaG(TYPE('a)) \ Gromov_product_at z x m \ Gromov_product_at z m y" | "dist m x \ deltaG(TYPE('a)) \ dist m y \ deltaG(TYPE('a)) \ Gromov_product_at z m y \ Gromov_product_at z x m" by linarith then show ?thesis proof (cases) case 1 have "dist m z \ dist m x + dist x z" by (intro mono_intros) then show ?thesis using 1 by auto next case 2 have "dist m z \ dist m y + dist y z" by (intro mono_intros) then show ?thesis using 2 by auto next case 3 then have "Gromov_product_at z x m = min (Gromov_product_at z x m) (Gromov_product_at z m y)" by auto also have "... \ Gromov_product_at z x y + deltaG(TYPE('a))" by (intro mono_intros) finally have "dist z m \ dist z y + dist x m - dist x y + 2 * deltaG(TYPE('a))" unfolding Gromov_product_at_def by (auto simp add: divide_simps algebra_simps) also have "... = dist z y - dist m y + 2 * deltaG(TYPE('a))" using geodesic_segment_dist[OF assms] by auto also have "... \ dist z y + deltaG(TYPE('a))" using 3 by auto finally show ?thesis by (simp add: metric_space_class.dist_commute) next case 4 then have "Gromov_product_at z m y = min (Gromov_product_at z x m) (Gromov_product_at z m y)" by auto also have "... \ Gromov_product_at z x y + deltaG(TYPE('a))" by (intro mono_intros) finally have "dist z m \ dist z x + dist m y - dist x y + 2 * deltaG(TYPE('a))" unfolding Gromov_product_at_def by (auto simp add: divide_simps algebra_simps) also have "... = dist z x - dist x m + 2 * deltaG(TYPE('a))" using geodesic_segment_dist[OF assms] by auto also have "... \ dist z x + deltaG(TYPE('a))" using 4 by (simp add: metric_space_class.dist_commute) finally show ?thesis by (simp add: metric_space_class.dist_commute) qed qed end (* of locale Gromov_hyperbolic_space *) text \A useful variation around the previous properties is that quadrilaterals are thin, in the following sense: if one has a union of three geodesics from $x$ to $t$, then a geodesic from $x$ to $t$ remains within distance $8\delta$ of the union of these 3 geodesics. We formulate the statement in geodesic hyperbolic spaces as the proof requires the construction of an additional geodesic, but in fact the statement is true without this assumption, thanks to the Bonk-Schramm extension theorem.\ lemma (in Gromov_hyperbolic_space_geodesic) thin_quadrilaterals: assumes "geodesic_segment_between Gxy x y" "geodesic_segment_between Gyz y z" "geodesic_segment_between Gzt z t" "geodesic_segment_between Gxt x t" "(w::'a) \ Gxt" shows "infdist w (Gxy \ Gyz \ Gzt) \ 8 * deltaG(TYPE('a))" proof - have I: "infdist w ({x--z} \ Gzt) \ 4 * deltaG(TYPE('a))" apply (rule thin_triangles[OF _ assms(3) assms(4) assms(5)]) by (simp add: geodesic_segment_commute) have "\u \ {x--z} \ Gzt. infdist w ({x--z} \ Gzt) = dist w u" apply (rule infdist_proper_attained, auto intro!: proper_Un simp add: geodesic_segment_topology(7)) by (meson assms(3) geodesic_segmentI geodesic_segment_topology) then obtain u where u: "u \ {x--z} \ Gzt" "infdist w ({x--z} \ Gzt) = dist w u" by auto have "infdist u (Gxy \ Gyz \ Gzt) \ 4 * deltaG(TYPE('a))" proof (cases "u \ {x--z}") case True have "infdist u (Gxy \ Gyz \ Gzt) \ infdist u (Gxy \ Gyz)" apply (intro mono_intros) using assms(1) by auto also have "... \ 4 * deltaG(TYPE('a))" using thin_triangles[OF geodesic_segment_commute[OF assms(1)] assms(2) _ True] by auto finally show ?thesis by auto next case False then have *: "u \ Gzt" using u(1) by auto have "infdist u (Gxy \ Gyz \ Gzt) \ infdist u Gzt" apply (intro mono_intros) using assms(3) by auto also have "... = 0" using * by auto finally show ?thesis using local.delta_nonneg by linarith qed moreover have "infdist w (Gxy \ Gyz \ Gzt) \ infdist u (Gxy \ Gyz \ Gzt) + dist w u" by (intro mono_intros) ultimately show ?thesis using I u(2) by auto qed text \There are converses to the above statements: if triangles are thin, or slim, then the space is Gromov-hyperbolic, for some $\delta$. We prove these criteria here, following the proofs in Ghys (with a simplification in the case of slim triangles.\ text \The basic result we will use twice below is the following: if points on sides of triangles at the same distance of the basepoint are close to each other up to the Gromov product, then the space is hyperbolic. The proof goes as follows. One wants to show that $(x,z)_e \geq \min((x,y)_e, (y,z)_e) - \delta = t-\delta$. On $[ex]$, $[ey]$ and $[ez]$, consider points $wx$, $wy$ and $wz$ at distance $t$ of $e$. Then $wx$ and $wy$ are $\delta$-close by assumption, and so are $wy$ and $wz$. Then $wx$ and $wz$ are $2\delta$-close. One can use these two points to express $(x,z)_e$, and the result follows readily.\ lemma (in geodesic_space) controlled_thin_triangles_implies_hyperbolic: assumes "\(x::'a) y z t Gxy Gxz. geodesic_segment_between Gxy x y \ geodesic_segment_between Gxz x z \ t \ {0..Gromov_product_at x y z} \ dist (geodesic_segment_param Gxy x t) (geodesic_segment_param Gxz x t) \ delta" shows "Gromov_hyperbolic_subset delta (UNIV::'a set)" proof (rule Gromov_hyperbolic_subsetI2) fix e x y z::'a define t where "t = min (Gromov_product_at e x y) (Gromov_product_at e y z)" define wx where "wx = geodesic_segment_param {e--x} e t" define wy where "wy = geodesic_segment_param {e--y} e t" define wz where "wz = geodesic_segment_param {e--z} e t" have "dist wx wy \ delta" unfolding wx_def wy_def t_def by (rule assms[of _ _ x _ y], auto) have "dist wy wz \ delta" unfolding wy_def wz_def t_def by (rule assms[of _ _ y _ z], auto) have "t + dist wy x = dist e wx + dist wy x" unfolding wx_def apply (auto intro!: geodesic_segment_param_in_geodesic_spaces(6)[symmetric]) unfolding t_def by (auto, meson Gromov_product_le_dist(1) min.absorb_iff2 min.left_idem order.trans) also have "... \ dist e wx + (dist wy wx + dist wx x)" by (intro mono_intros) also have "... \ dist e wx + (delta + dist wx x)" using \dist wx wy \ delta\ by (auto simp add: metric_space_class.dist_commute) also have "... = delta + dist e x" apply auto apply (rule geodesic_segment_dist[of "{e--x}"]) unfolding wx_def t_def by (auto simp add: geodesic_segment_param_in_segment) finally have *: "t + dist wy x - delta \ dist e x" by simp have "t + dist wy z = dist e wz + dist wy z" unfolding wz_def apply (auto intro!: geodesic_segment_param_in_geodesic_spaces(6)[symmetric]) unfolding t_def by (auto, meson Gromov_product_le_dist(2) min.absorb_iff1 min.right_idem order.trans) also have "... \ dist e wz + (dist wy wz + dist wz z)" by (intro mono_intros) also have "... \ dist e wz + (delta + dist wz z)" using \dist wy wz \ delta\ by (auto simp add: metric_space_class.dist_commute) also have "... = delta + dist e z" apply auto apply (rule geodesic_segment_dist[of "{e--z}"]) unfolding wz_def t_def by (auto simp add: geodesic_segment_param_in_segment) finally have "t + dist wy z - delta \ dist e z" by simp then have "(t + dist wy x - delta) + (t + dist wy z - delta) \ dist e x + dist e z" using * by simp also have "... = dist x z + 2 * Gromov_product_at e x z" unfolding Gromov_product_at_def by (auto simp add: algebra_simps divide_simps) also have "... \ dist wy x + dist wy z + 2 * Gromov_product_at e x z" using metric_space_class.dist_triangle[of x z wy] by (auto simp add: metric_space_class.dist_commute) finally have "2 * t - 2 * delta \ 2 * Gromov_product_at e x z" by auto then show "min (Gromov_product_at e x y) (Gromov_product_at e y z) - delta \ Gromov_product_at e x z" unfolding t_def by auto qed text \We prove that if triangles are thin, i.e., they satisfy the Rips condition, i.e., every side of a triangle is included in the $\delta$-neighborhood of the union of the other triangles, then the space is hyperbolic. If a point $w$ on $[xy]$ satisfies $d(x,w) < (y,z)_x - \delta$, then its friend on $[xz] \cup [yz]$ has to be on $[xz]$, and roughly at the same distance of the origin. Then it follows that the point on $[xz]$ with $d(x,w') = d(x,w)$ is close to $w$, as desired. If $d(x,w) \in [(y,z)_x - \delta, (y,z)_x)$, we argue in the same way but for the point which is closer to $x$ by an amount $\delta$. Finally, the last case $d(x,w) = (y,z)_x$ follows by continuity.\ proposition (in geodesic_space) thin_triangles_implies_hyperbolic: assumes "\(x::'a) y z w Gxy Gyz Gxz. geodesic_segment_between Gxy x y \ geodesic_segment_between Gxz x z \ geodesic_segment_between Gyz y z \ w \ Gxy \ infdist w (Gxz \ Gyz) \ delta" shows "Gromov_hyperbolic_subset (4 * delta) (UNIV::'a set)" proof - obtain x0::'a where True by auto have "infdist x0 ({x0} \ {x0}) \ delta" by (rule assms[of "{x0}" x0 x0 "{x0}" x0 "{x0}" x0], auto) then have [simp]: "delta \ 0" using infdist_nonneg by auto have "dist (geodesic_segment_param Gxy x t) (geodesic_segment_param Gxz x t) \ 4 * delta" if H: "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" "t \ {0..Gromov_product_at x y z}" for x y z t Gxy Gxz proof - have Main: "dist (geodesic_segment_param Gxy x u) (geodesic_segment_param Gxz x u) \ 4 * delta" if "u \ {delta..delta \ 0\ by linarith+ then have I1: "dist wy (geodesic_segment_param Gxy x u) = delta" by auto have "infdist wy (Gxz \ {y--z}) \ delta" unfolding wy_def apply (rule assms[of Gxy x y _ z]) using H by (auto simp add: geodesic_segment_param_in_segment) moreover have "\wz \ Gxz \ {y--z}. infdist wy (Gxz \ {y--z}) = dist wy wz" apply (rule infdist_proper_attained, intro proper_Un) using H(2) by (auto simp add: geodesic_segment_topology) ultimately obtain wz where wz: "wz \ Gxz \ {y--z}" "dist wy wz \ delta" by force have "dist wz x \ dist wz wy + dist wy x" by (rule metric_space_class.dist_triangle) also have "... \ delta + (u-delta)" apply (intro add_mono) using wz(2) unfolding wy_def apply (auto simp add: metric_space_class.dist_commute) apply (intro eq_refl geodesic_segment_param(6)[OF H(1)]) using that apply auto by (metis diff_0_right diff_mono dual_order.trans Gromov_product_le_dist(1) less_eq_real_def metric_space_class.dist_commute metric_space_class.zero_le_dist wy_def) finally have "dist wz x \ u" by auto also have "... < Gromov_product_at x y z" using that by auto also have "... \ infdist x {y--z}" by (rule Gromov_product_le_infdist, auto) finally have "dist x wz < infdist x {y--z}" by (simp add: metric_space_class.dist_commute) then have "wz \ {y--z}" by (metis add.left_neutral infdist_triangle infdist_zero leD) then have "wz \ Gxz" using wz by auto have "u - delta = dist x wy" unfolding wy_def apply (rule geodesic_segment_param(6)[symmetric, OF H(1)]) using that apply auto using Gromov_product_le_dist(1)[of x y z] \delta \ 0\ by linarith also have "... \ dist x wz + dist wz wy" by (rule metric_space_class.dist_triangle) also have "... \ dist x wz + delta" using wz(2) by (simp add: metric_space_class.dist_commute) finally have "dist x wz \ u - 2 * delta" by auto define dz where "dz = dist x wz" have *: "wz = geodesic_segment_param Gxz x dz" unfolding dz_def using \wz \ Gxz\ H(2) by auto have "dist wz (geodesic_segment_param Gxz x u) = abs(dz - u)" unfolding * apply (rule geodesic_segment_param(7)[OF H(2)]) unfolding dz_def using \dist wz x \ u\ that apply (auto simp add: metric_space_class.dist_commute) using Gromov_product_le_dist(2)[of x y z] \delta \ 0\ by linarith+ also have "... \ 2 * delta" unfolding dz_def using \dist wz x \ u\ \dist x wz \ u - 2 * delta\ by (auto simp add: metric_space_class.dist_commute) finally have I3: "dist wz (geodesic_segment_param Gxz x u) \ 2 * delta" by simp have "dist (geodesic_segment_param Gxy x u) (geodesic_segment_param Gxz x u) \ dist (geodesic_segment_param Gxy x u) wy + dist wy wz + dist wz (geodesic_segment_param Gxz x u)" by (rule dist_triangle4) also have "... \ delta + delta + (2 * delta)" using I1 wz(2) I3 by (auto simp add: metric_space_class.dist_commute) finally show ?thesis by simp qed have "t \ {0..dist x y}" "t \ {0..dist x z}" "t \ 0" using \t \ {0..Gromov_product_at x y z}\ apply auto using Gromov_product_le_dist[of x y z] by linarith+ consider "t \ delta" | "t \ {delta.. t > delta" using \t \ {0..Gromov_product_at x y z}\ by (auto, linarith) then show ?thesis proof (cases) case 1 have "dist (geodesic_segment_param Gxy x t) (geodesic_segment_param Gxz x t) \ dist x (geodesic_segment_param Gxy x t) + dist x (geodesic_segment_param Gxz x t)" by (rule metric_space_class.dist_triangle3) also have "... = t + t" using geodesic_segment_param(6)[OF H(1) \t \ {0..dist x y}\] geodesic_segment_param(6)[OF H(2) \t \ {0..dist x z}\] by auto also have "... \ 4 * delta" using 1 \delta \ 0\ by linarith finally show ?thesis by simp next case 2 show ?thesis using Main[OF 2] by simp next case 3 text \In this case, we argue by approximating $t$ by a slightly smaller parameter, for which the result has already been proved above. We need to argue that all functions are continuous on the sets we are considering, which is straightforward but tedious.\ define u::"nat \ real" where "u = (\n. t-1/n)" have "u \ t - 0" unfolding u_def by (intro tendsto_intros) then have "u \ t" by simp then have *: "eventually (\n. u n > delta) sequentially" using 3 by (auto simp add: order_tendsto_iff) have **: "eventually (\n. u n \ 0) sequentially" apply (rule eventually_elim2[OF *, of "(\n. delta \ 0)"]) apply auto using \delta \ 0\ by linarith have ***: "u n \ t" for n unfolding u_def by auto have A: "eventually (\n. u n \ {delta..n. dist (geodesic_segment_param Gxy x (u n)) (geodesic_segment_param Gxz x (u n)) \ 4 * delta) sequentially" by (rule eventually_mono[OF A Main], simp) have C: "(\n. dist (geodesic_segment_param Gxy x (u n)) (geodesic_segment_param Gxz x (u n))) \ dist (geodesic_segment_param Gxy x t) (geodesic_segment_param Gxz x t)" apply (intro tendsto_intros) apply (rule continuous_on_tendsto_compose[OF _ \u \ t\ \t \ {0..dist x y}\]) apply (simp add: isometry_on_continuous H(1)) using ** *** \t \ {0..dist x y}\ apply (simp, intro eventually_conj, simp, meson dual_order.trans eventually_mono) apply (rule continuous_on_tendsto_compose[OF _ \u \ t\ \t \ {0..dist x z}\]) apply (simp add: isometry_on_continuous H(2)) using ** *** \t \ {0..dist x z}\ apply (simp, intro eventually_conj, simp, meson dual_order.trans eventually_mono) done show ?thesis using B unfolding eventually_sequentially using LIMSEQ_le_const2[OF C] by simp qed qed with controlled_thin_triangles_implies_hyperbolic[OF this] show ?thesis by auto qed text \Then, we prove that if triangles are slim (i.e., there is a point that is $\delta$-close to all sides), then the space is hyperbolic. Using the previous statement, we should show that points on $[xy]$ and $[xz]$ at the same distance $t$ of the origin are close, if $t \leq (y,z)_x$. There are two steps: - for $t = (y,z)_x$, then the two points are in fact close to the middle of the triangle (as this point satisfies $d(x,y) = d(x,w) + d(w,y) + O(\delta)$, and similarly for the other sides, one gets readily $d(x,w) = (y,z)_w + O(\delta)$ by expanding the formula for the Gromov product). Hence, they are close together. - For $t < (y,z)_x$, we argue that there are points $y' \in [xy]$ and $z' \in [xz]$ for which $t = (y',z')_x$, by a continuity argument and the intermediate value theorem. Then the result follows from the first step in the triangle $xy'z'$. The proof we give is simpler than the one in~\cite{ghys_hyperbolique}, and gives better constants.\ proposition (in geodesic_space) slim_triangles_implies_hyperbolic: assumes "\(x::'a) y z Gxy Gyz Gxz. geodesic_segment_between Gxy x y \ geodesic_segment_between Gxz x z \ geodesic_segment_between Gyz y z \ \w. infdist w Gxy \ delta \ infdist w Gxz \ delta \ infdist w Gyz \ delta" shows "Gromov_hyperbolic_subset (6 * delta) (UNIV::'a set)" proof - text \First step: the result is true for $t = (y,z)_x$.\ have Main: "dist (geodesic_segment_param Gxy x (Gromov_product_at x y z)) (geodesic_segment_param Gxz x (Gromov_product_at x y z)) \ 6 * delta" if H: "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" for x y z Gxy Gxz proof - obtain w where w: "infdist w Gxy \ delta" "infdist w Gxz \ delta" "infdist w {y--z} \ delta" using assms[OF H, of "{y--z}"] by auto have "\wxy \ Gxy. infdist w Gxy = dist w wxy" apply (rule infdist_proper_attained) using H(1) by (auto simp add: geodesic_segment_topology) then obtain wxy where wxy: "wxy \ Gxy" "dist w wxy \ delta" using w by auto have "\wxz \ Gxz. infdist w Gxz = dist w wxz" apply (rule infdist_proper_attained) using H(2) by (auto simp add: geodesic_segment_topology) then obtain wxz where wxz: "wxz \ Gxz" "dist w wxz \ delta" using w by auto have "\wyz \ {y--z}. infdist w {y--z} = dist w wyz" apply (rule infdist_proper_attained) by (auto simp add: geodesic_segment_topology) then obtain wyz where wyz: "wyz \ {y--z}" "dist w wyz \ delta" using w by auto have I: "dist wxy wxz \ 2 * delta" "dist wxy wyz \ 2 * delta" "dist wxz wyz \ 2 * delta" using metric_space_class.dist_triangle[of wxy wxz w] metric_space_class.dist_triangle[of wxy wyz w] metric_space_class.dist_triangle[of wxz wyz w] wxy(2) wyz(2) wxz(2) by (auto simp add: metric_space_class.dist_commute) text \We show that $d(x, wxy)$ is close to the Gromov product of $y$ and $z$ seen from $x$. This follows from the fact that $w$ is essentially on all geodesics, so that everything simplifies when one writes down the Gromov products, leaving only $d(x, w)$ up to $O(\delta)$. To get the right $O(\delta)$, one has to be a little bit careful, using the triangular inequality when possible. This means that the computations for the upper and lower bounds are different, making them a little bit tedious, although straightforward.\ have "dist y wxy -4 * delta + dist wxy z \ dist y wxy - dist wxy wyz + dist wxy z - dist wxy wyz" using I by simp also have "... \ dist wyz y + dist wyz z" using metric_space_class.dist_triangle[of y wxy wyz] metric_space_class.dist_triangle[of wxy z wyz] by (auto simp add: metric_space_class.dist_commute) also have "... = dist y z" using wyz(1) by (metis geodesic_segment_dist local.some_geodesic_is_geodesic_segment(1) metric_space_class.dist_commute) finally have *: "dist y wxy + dist wxy z - 4 * delta \ dist y z" by simp have "2 * Gromov_product_at x y z = dist x y + dist x z - dist y z" unfolding Gromov_product_at_def by simp also have "... \ dist x wxy + dist wxy y + dist x wxy + dist wxy z - (dist y wxy + dist wxy z - 4 * delta)" using metric_space_class.dist_triangle[of x y wxy] metric_space_class.dist_triangle[of x z wxy] * by (auto simp add: metric_space_class.dist_commute) also have "... = 2 * dist x wxy + 4 * delta" by (auto simp add: metric_space_class.dist_commute) finally have A: "Gromov_product_at x y z \ dist x wxy + 2 * delta" by simp have "dist x wxy -4 * delta + dist wxy z \ dist x wxy - dist wxy wxz + dist wxy z - dist wxy wxz" using I by simp also have "... \ dist wxz x + dist wxz z" using metric_space_class.dist_triangle[of x wxy wxz] metric_space_class.dist_triangle[of wxy z wxz] by (auto simp add: metric_space_class.dist_commute) also have "... = dist x z" using wxz(1) H(2) by (metis geodesic_segment_dist metric_space_class.dist_commute) finally have *: "dist x wxy + dist wxy z - 4 * delta \ dist x z" by simp have "2 * dist x wxy - 4 * delta = (dist x wxy + dist wxy y) + (dist x wxy + dist wxy z - 4 * delta) - (dist y wxy + dist wxy z)" by (auto simp add: metric_space_class.dist_commute) also have "... \ dist x y + dist x z - dist y z" using * metric_space_class.dist_triangle[of y z wxy] geodesic_segment_dist[OF H(1) wxy(1)] by auto also have "... = 2 * Gromov_product_at x y z" unfolding Gromov_product_at_def by simp finally have B: "Gromov_product_at x y z \ dist x wxy - 2 * delta" by simp define dy where "dy = dist x wxy" have *: "wxy = geodesic_segment_param Gxy x dy" unfolding dy_def using \wxy \ Gxy\ H(1) by auto have "dist wxy (geodesic_segment_param Gxy x (Gromov_product_at x y z)) = abs(dy - Gromov_product_at x y z)" unfolding * apply (rule geodesic_segment_param(7)[OF H(1)]) unfolding dy_def using that geodesic_segment_dist_le[OF H(1) wxy(1), of x] by (auto simp add: metric_space_class.dist_commute) also have "... \ 2 * delta" using A B unfolding dy_def by auto finally have Iy: "dist wxy (geodesic_segment_param Gxy x (Gromov_product_at x y z)) \ 2 * delta" by simp text \We need the same estimate for $wxz$. The proof is exactly the same, copied and pasted. It would be better to have a separate statement, but since its assumptions would be rather cumbersome I decided to keep the two proofs.\ have "dist z wxz -4 * delta + dist wxz y \ dist z wxz - dist wxz wyz + dist wxz y - dist wxz wyz" using I by simp also have "... \ dist wyz z + dist wyz y" using metric_space_class.dist_triangle[of z wxz wyz] metric_space_class.dist_triangle[of wxz y wyz] by (auto simp add: metric_space_class.dist_commute) also have "... = dist z y" using \dist wyz y + dist wyz z = dist y z\ by (auto simp add: metric_space_class.dist_commute) finally have *: "dist z wxz + dist wxz y - 4 * delta \ dist z y" by simp have "2 * Gromov_product_at x y z = dist x z + dist x y - dist z y" unfolding Gromov_product_at_def by (simp add: metric_space_class.dist_commute) also have "... \ dist x wxz + dist wxz z + dist x wxz + dist wxz y - (dist z wxz + dist wxz y - 4 * delta)" using metric_space_class.dist_triangle[of x z wxz] metric_space_class.dist_triangle[of x y wxz] * by (auto simp add: metric_space_class.dist_commute) also have "... = 2 * dist x wxz + 4 * delta" by (auto simp add: metric_space_class.dist_commute) finally have A: "Gromov_product_at x y z \ dist x wxz + 2 * delta" by simp have "dist x wxz -4 * delta + dist wxz y \ dist x wxz - dist wxz wxy + dist wxz y - dist wxz wxy" using I by (simp add: metric_space_class.dist_commute) also have "... \ dist wxy x + dist wxy y" using metric_space_class.dist_triangle[of x wxz wxy] metric_space_class.dist_triangle[of wxz y wxy] by (auto simp add: metric_space_class.dist_commute) also have "... = dist x y" using wxy(1) H(1) by (metis geodesic_segment_dist metric_space_class.dist_commute) finally have *: "dist x wxz + dist wxz y - 4 * delta \ dist x y" by simp have "2 * dist x wxz - 4 * delta = (dist x wxz + dist wxz z) + (dist x wxz + dist wxz y - 4 * delta) - (dist z wxz + dist wxz y)" by (auto simp add: metric_space_class.dist_commute) also have "... \ dist x z + dist x y - dist z y" using * metric_space_class.dist_triangle[of z y wxz] geodesic_segment_dist[OF H(2) wxz(1)] by auto also have "... = 2 * Gromov_product_at x y z" unfolding Gromov_product_at_def by (simp add: metric_space_class.dist_commute) finally have B: "Gromov_product_at x y z \ dist x wxz - 2 * delta" by simp define dz where "dz = dist x wxz" have *: "wxz = geodesic_segment_param Gxz x dz" unfolding dz_def using \wxz \ Gxz\ H(2) by auto have "dist wxz (geodesic_segment_param Gxz x (Gromov_product_at x y z)) = abs(dz - Gromov_product_at x y z)" unfolding * apply (rule geodesic_segment_param(7)[OF H(2)]) unfolding dz_def using that geodesic_segment_dist_le[OF H(2) wxz(1), of x] by (auto simp add: metric_space_class.dist_commute) also have "... \ 2 * delta" using A B unfolding dz_def by auto finally have Iz: "dist wxz (geodesic_segment_param Gxz x (Gromov_product_at x y z)) \ 2 * delta" by simp have "dist (geodesic_segment_param Gxy x (Gromov_product_at x y z)) (geodesic_segment_param Gxz x (Gromov_product_at x y z)) \ dist (geodesic_segment_param Gxy x (Gromov_product_at x y z)) wxy + dist wxy wxz + dist wxz (geodesic_segment_param Gxz x (Gromov_product_at x y z))" by (rule dist_triangle4) also have "... \ 2 * delta + 2 * delta + 2 * delta" using Iy Iz I by (auto simp add: metric_space_class.dist_commute) finally show ?thesis by simp qed text \Second step: the result is true for $t \leq (y,z)_x$, by a continuity argument and a reduction to the first step.\ have "dist (geodesic_segment_param Gxy x t) (geodesic_segment_param Gxz x t) \ 6 * delta" if H: "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" "t \ {0..Gromov_product_at x y z}" for x y z t Gxy Gxz proof - define ys where "ys = (\s. geodesic_segment_param Gxy x (s * dist x y))" define zs where "zs = (\s. geodesic_segment_param Gxz x (s * dist x z))" define F where "F = (\s. Gromov_product_at x (ys s) (zs s))" have "\s. 0 \ s \ s \ 1 \ F s = t" proof (rule IVT') show "F 0 \ t" "t \ F 1" unfolding F_def using that unfolding ys_def zs_def by (auto simp add: Gromov_product_e_x_x) show "continuous_on {0..1} F" unfolding F_def Gromov_product_at_def ys_def zs_def apply (intro continuous_intros continuous_on_compose2[of "{0..dist x y}" _ _ "\t. t * dist x y"] continuous_on_compose2[of "{0..dist x z}" _ _ "\t. t * dist x z"]) apply (auto intro!: isometry_on_continuous geodesic_segment_param(4) that) using metric_space_class.zero_le_dist mult_left_le_one_le by blast+ qed (simp) then obtain s where s: "s \ {0..1}" "t = Gromov_product_at x (ys s) (zs s)" unfolding F_def by auto have a: "x = geodesic_segment_param Gxy x 0" using H(1) by auto have b: "x = geodesic_segment_param Gxz x 0" using H(2) by auto have dy: "dist x (ys s) = s * dist x y" unfolding ys_def apply (rule geodesic_segment_param[OF H(1)]) using s(1) by (auto simp add: mult_left_le_one_le) have dz: "dist x (zs s) = s * dist x z" unfolding zs_def apply (rule geodesic_segment_param[OF H(2)]) using s(1) by (auto simp add: mult_left_le_one_le) define Gxy2 where "Gxy2 = geodesic_subsegment Gxy x 0 (s * dist x y)" define Gxz2 where "Gxz2 = geodesic_subsegment Gxz x 0 (s * dist x z)" have "dist (geodesic_segment_param Gxy2 x t) (geodesic_segment_param Gxz2 x t) \ 6 * delta" unfolding s(2) proof (rule Main) show "geodesic_segment_between Gxy2 x (ys s)" apply (subst a) unfolding Gxy2_def ys_def apply (rule geodesic_subsegment[OF H(1)]) using s(1) by (auto simp add: mult_left_le_one_le) show "geodesic_segment_between Gxz2 x (zs s)" apply (subst b) unfolding Gxz2_def zs_def apply (rule geodesic_subsegment[OF H(2)]) using s(1) by (auto simp add: mult_left_le_one_le) qed moreover have "geodesic_segment_param Gxy2 x (t-0) = geodesic_segment_param Gxy x t" apply (subst a) unfolding Gxy2_def apply (rule geodesic_subsegment(3)[OF H(1)]) using s(1) H(3) unfolding s(2) apply (auto simp add: mult_left_le_one_le) unfolding dy[symmetric] by (rule Gromov_product_le_dist) moreover have "geodesic_segment_param Gxz2 x (t-0) = geodesic_segment_param Gxz x t" apply (subst b) unfolding Gxz2_def apply (rule geodesic_subsegment(3)[OF H(2)]) using s(1) H(3) unfolding s(2) apply (auto simp add: mult_left_le_one_le) unfolding dz[symmetric] by (rule Gromov_product_le_dist) ultimately show ?thesis by simp qed with controlled_thin_triangles_implies_hyperbolic[OF this] show ?thesis by auto qed section \Metric trees\ text \Metric trees have several equivalent definitions. The simplest one is probably that it is a geodesic space in which the union of two geodesic segments intersecting only at one endpoint is still a geodesic segment. Metric trees are Gromov hyperbolic, with $\delta = 0$.\ class metric_tree = geodesic_space + assumes geod_union: "geodesic_segment_between G x y \ geodesic_segment_between H y z \ G \ H = {y} \ geodesic_segment_between (G \ H) x z" text \We will now show that the real line is a metric tree, by identifying its geodesic segments, i.e., the compact intervals.\ lemma geodesic_segment_between_real: assumes "x \ (y::real)" shows "geodesic_segment_between (G::real set) x y = (G = {x..y})" proof assume H: "geodesic_segment_between G x y" then have "connected G" "x \ G" "y \ G" using geodesic_segment_topology(2) geodesic_segmentI geodesic_segment_endpoints by auto then have *: "{x..y} \ G" by (simp add: connected_contains_Icc) moreover have "G \ {x..y}" proof fix s assume "s \ G" have "abs(s-x) + abs(s-y) = abs(x-y)" using geodesic_segment_dist[OF H \s \ G\] unfolding dist_real_def by auto then show "s \ {x..y}" using \x \ y\ by auto qed ultimately show "G = {x..y}" by auto next assume H: "G = {x..y}" define g where "g = (\t. t + x)" have "g 0 = x \ g (dist x y) = y \ isometry_on {0..dist x y} g \ G = g ` {0..dist x y}" unfolding g_def isometry_on_def H using \x \ y\ by (auto simp add: dist_real_def) then have "\g. g 0 = x \ g (dist x y) = y \ isometry_on {0..dist x y} g \ G = g ` {0..dist x y}" by auto then show "geodesic_segment_between G x y" unfolding geodesic_segment_between_def by auto qed lemma geodesic_segment_between_real': "{x--y} = {min x y..max x (y::real)}" by (metis geodesic_segment_between_real geodesic_segment_commute some_geodesic_is_geodesic_segment(1) max_def min.cobounded1 min_def) lemma geodesic_segment_real: "geodesic_segment (G::real set) = (\x y. x \ y \ G = {x..y})" proof assume "geodesic_segment G" then obtain x y where *: "geodesic_segment_between G x y" unfolding geodesic_segment_def by auto have "(x \ y \ G = {x..y}) \ (y \ x \ G = {y..x})" apply (rule le_cases[of x y]) using geodesic_segment_between_real * geodesic_segment_commute apply simp using geodesic_segment_between_real * geodesic_segment_commute by metis then show "\x y. x \ y \ G = {x..y}" by auto next assume "\x y. x \ y \ G = {x..y}" then show "geodesic_segment G" unfolding geodesic_segment_def using geodesic_segment_between_real by metis qed instance real::metric_tree proof fix G H::"real set" and x y z::real assume GH: "geodesic_segment_between G x y" "geodesic_segment_between H y z" "G \ H = {y}" have G: "G = {min x y..max x y}" using GH by (metis geodesic_segment_between_real geodesic_segment_commute inf_real_def inf_sup_ord(2) max.coboundedI2 max_def min_def) have H: "H = {min y z..max y z}" using GH by (metis geodesic_segment_between_real geodesic_segment_commute inf_real_def inf_sup_ord(2) max.coboundedI2 max_def min_def) have *: "(x \ y \ y \ z) \ (z \ y \ y \ x)" using G H \G \ H = {y}\ unfolding min_def max_def apply auto apply (metis (mono_tags, hide_lams) min_le_iff_disj order_refl) by (metis (full_types) less_eq_real_def max_def) show "geodesic_segment_between (G \ H) x z" using * apply rule using \G \ H = {y}\ unfolding G H apply (metis G GH(1) GH(2) H geodesic_segment_between_real ivl_disj_un_two_touch(4) order_trans) using \G \ H = {y}\ unfolding G H by (metis (full_types) Un_commute geodesic_segment_between_real geodesic_segment_commute ivl_disj_un_two_touch(4) le_max_iff_disj max.absorb_iff2 max.commute min_absorb2) qed context metric_tree begin text \We show that a metric tree is uniquely geodesic.\ subclass uniquely_geodesic_space proof fix x y G H assume H: "geodesic_segment_between G x y" "geodesic_segment_between H x (y::'a)" show "G = H" proof (rule uniquely_geodesic_spaceI[OF _ H]) fix G H x y assume "geodesic_segment_between G x y" "geodesic_segment_between H x y" "G \ H = {x, (y::'a)}" show "x = y" proof (rule ccontr) assume "x \ y" then have "dist x y > 0" by auto obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "G = g`{0..dist x y}" by (meson \geodesic_segment_between G x y\ geodesic_segment_between_def) define G2 where "G2 = g`{0..dist x y/2}" have "G2 \ G" unfolding G2_def g(4) by auto define z where "z = g(dist x y/2)" have "dist x z = dist x y/2" using isometry_onD[OF g(3), of 0 "dist x y/2"] g(1) z_def unfolding dist_real_def by auto have "dist y z = dist x y/2" using isometry_onD[OF g(3), of "dist x y" "dist x y/2"] g(2) z_def unfolding dist_real_def by auto have G2: "geodesic_segment_between G2 x z" unfolding \g 0 = x\[symmetric] z_def G2_def apply (rule geodesic_segmentI2) by (rule isometry_on_subset[OF g(3)], auto simp add: \g 0 = x\) have [simp]: "x \ G2" "z \ G2" using geodesic_segment_endpoints G2 by auto have "dist x a \ dist x z" if "a \ G2" for a apply (rule geodesic_segment_dist_le) using G2 that by auto also have "... < dist x y" unfolding \dist x z = dist x y/2\ using \dist x y > 0\ by auto finally have "y \ G2" by auto then have "G2 \ H = {x}" using \G2 \ G\ \x \ G2\ \G \ H = {x, y}\ by auto have *: "geodesic_segment_between (G2 \ H) z y" apply (rule geod_union[of _ _ x]) using \G2 \ H = {x}\ \geodesic_segment_between H x y\ G2 by (auto simp add: geodesic_segment_commute) have "dist x y \ dist z x + dist x y" by auto also have "... = dist z y" apply (rule geodesic_segment_dist[OF *]) using \G \ H = {x, y}\ by auto also have "... = dist x y / 2" by (simp add: \dist y z = dist x y / 2\ metric_space_class.dist_commute) finally show False using \dist x y > 0\ by auto qed qed qed text \An important property of metric trees is that any geodesic triangle is degenerate, i.e., the three sides intersect at a unique point, the center of the triangle, that we introduce now.\ definition center::"'a \ 'a \ 'a \ 'a" where "center x y z = (SOME t. t \ {x--y} \ {x--z} \ {y--z})" lemma center_as_intersection: "{x--y} \ {x--z} \ {y--z} = {center x y z}" proof - obtain g where g: "g 0 = x" "g (dist x y) = y" "isometry_on {0..dist x y} g" "{x--y} = g`{0..dist x y}" by (meson geodesic_segment_between_def some_geodesic_is_geodesic_segment(1)) obtain h where h: "h 0 = x" "h (dist x z) = z" "isometry_on {0..dist x z} h" "{x--z} = h`{0..dist x z}" by (meson geodesic_segment_between_def some_geodesic_is_geodesic_segment(1)) define Z where "Z = {t \ {0..min (dist x y) (dist x z)}. g t = h t}" have "0 \ Z" unfolding Z_def using g(1) h(1) by auto have [simp]: "closed Z" proof - have *: "Z = (\s. dist (g s) (h s))-`{0} \ {0..min (dist x y) (dist x z)}" unfolding Z_def by auto show ?thesis unfolding * apply (rule closed_vimage_Int) using continuous_on_subset[OF isometry_on_continuous[OF g(3)], of "{0..min (dist x y) (dist x z)}"] continuous_on_subset[OF isometry_on_continuous[OF h(3)], of "{0..min (dist x y) (dist x z)}"] continuous_on_dist by auto qed define a where "a = Sup Z" have "a \ Z" unfolding a_def apply (rule closed_contains_Sup, auto) using \0 \ Z\ Z_def by auto define c where "c = h a" then have a: "g a = c" "h a = c" "a \ 0" "a \ dist x y" "a \ dist x z" using \a \ Z\ unfolding Z_def c_def by auto define G2 where "G2 = g`{a..dist x y}" have G2: "geodesic_segment_between G2 (g a) (g (dist x y))" unfolding G2_def apply (rule geodesic_segmentI2) using isometry_on_subset[OF g(3)] \a \ Z\ unfolding Z_def by auto define H2 where "H2 = h`{a..dist x z}" have H2: "geodesic_segment_between H2 (h a) (h (dist x z))" unfolding H2_def apply (rule geodesic_segmentI2) using isometry_on_subset[OF h(3)] \a \ Z\ unfolding Z_def by auto have "G2 \ H2 \ {c}" proof fix w assume w: "w \ G2 \ H2" obtain sg where sg: "w = g sg" "sg \ {a..dist x y}" using w unfolding G2_def by auto obtain sh where sh: "w = h sh" "sh \ {a..dist x z}" using w unfolding H2_def by auto have "dist w x = sg" unfolding g(1)[symmetric] sg(1) using isometry_onD[OF g(3), of 0 sg] sg(2) unfolding dist_real_def using a by (auto simp add: metric_space_class.dist_commute) moreover have "dist w x = sh" unfolding h(1)[symmetric] sh(1) using isometry_onD[OF h(3), of 0 sh] sh(2) unfolding dist_real_def using a by (auto simp add: metric_space_class.dist_commute) ultimately have "sg = sh" by simp have "sh \ Z" unfolding Z_def using sg sh \a \ 0\ unfolding \sg = sh\ by auto then have "sh \ a" unfolding a_def apply (rule cSup_upper) unfolding Z_def by auto then have "sh = a" using sh(2) by auto then show "w \ {c}" unfolding sh(1) using a(2) by auto qed then have *: "G2 \ H2 = {c}" unfolding G2_def H2_def using a by (auto simp add: image_iff, force) have "geodesic_segment_between (G2 \ H2) y z" apply (subst g(2)[symmetric], subst h(2)[symmetric]) apply(rule geod_union[of _ _ "h a"]) using geodesic_segment_commute G2 H2 a * by force+ then have "G2 \ H2 = {y--z}" using geodesic_segment_unique by auto then have "c \ {y--z}" using * by auto then have *: "c \ {x--y} \ {x--z} \ {y--z}" using g(4) h(4) c_def a by force have center: "center x y z \ {x--y} \ {x--z} \ {y--z}" unfolding center_def using someI[of "\p. p \ {x--y} \ {x--z} \ {y--z}", OF *] by blast have *: "dist x d = Gromov_product_at x y z" if "d \ {x--y} \ {x--z} \ {y--z}" for d proof - have "dist x y = dist x d + dist d y" "dist x z = dist x d + dist d z" "dist y z = dist y d + dist d z" using that by (auto simp add: geodesic_segment_dist geodesic_segment_unique) then show ?thesis unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute) qed have "d = center x y z" if "d \ {x--y} \ {x--z} \ {y--z}" for d apply (rule geodesic_segment_dist_unique[of "{x--y}" x y]) using *[OF that] *[OF center] that center by auto then show "{x--y} \ {x--z} \ {y--z} = {center x y z}" using center by blast qed lemma center_on_geodesic [simp]: "center x y z \ {x--y}" "center x y z \ {x--z}" "center x y z \ {y--z}" "center x y z \ {y--x}" "center x y z \ {z--x}" "center x y z \ {z--y}" using center_as_intersection by (auto simp add: some_geodesic_commute) lemma center_commute: "center x y z = center x z y" "center x y z = center y x z" "center x y z = center y z x" "center x y z = center z x y" "center x y z = center z y x" using center_as_intersection some_geodesic_commute by blast+ lemma center_dist: "dist x (center x y z) = Gromov_product_at x y z" proof - have "dist x y = dist x (center x y z) + dist (center x y z) y" "dist x z = dist x (center x y z) + dist (center x y z) z" "dist y z = dist y (center x y z) + dist (center x y z) z" by (auto simp add: geodesic_segment_dist geodesic_segment_unique) then show ?thesis unfolding Gromov_product_at_def by (auto simp add: metric_space_class.dist_commute) qed lemma geodesic_intersection: "{x--y} \ {x--z} = {x--center x y z}" proof - have "{x--y} = {x--center x y z} \ {center x y z--y}" using center_as_intersection geodesic_segment_split by blast moreover have "{x--z} = {x--center x y z} \ {center x y z--z}" using center_as_intersection geodesic_segment_split by blast ultimately have "{x--y} \ {x--z} = {x--center x y z} \ ({center x y z--y} \ {x--center x y z}) \ ({center x y z--y} \ {x--center x y z}) \ ({center x y z--y} \ {center x y z--z})" by auto moreover have "{center x y z--y} \ {x--center x y z} = {center x y z}" using geodesic_segment_split(2) center_as_intersection[of x y z] by auto moreover have "{center x y z--y} \ {x--center x y z} = {center x y z}" using geodesic_segment_split(2) center_as_intersection[of x y z] by auto moreover have "{center x y z--y} \ {center x y z--z} = {center x y z}" using geodesic_segment_split(2)[of "center x y z" y z] center_as_intersection[of x y z] by (auto simp add: some_geodesic_commute) ultimately show "{x--y} \ {x--z} = {x--center x y z}" by auto qed end (*of context metric_tree*) text \We can now prove that a metric tree is Gromov hyperbolic, for $\delta = 0$. The simplest proof goes through the slim triangles property: it suffices to show that, given a geodesic triangle, there is a point at distance at most $0$ of each of its sides. This is the center we have constructed above.\ class metric_tree_with_delta = metric_tree + metric_space_with_deltaG + assumes delta0: "deltaG(TYPE('a::metric_space)) = 0" class Gromov_hyperbolic_space_0 = Gromov_hyperbolic_space + assumes delta0 [simp]: "deltaG(TYPE('a::metric_space)) = 0" class Gromov_hyperbolic_space_0_geodesic = Gromov_hyperbolic_space_0 + geodesic_space text \Isabelle does not accept cycles in the class graph. So, we will show that \verb+metric_tree_with_delta+ is a subclass of \verb+Gromov_hyperbolic_space_0_geodesic+, and conversely that \verb+Gromov_hyperbolic_space_0_geodesic+ is a subclass of \verb+metric_tree+. In a tree, we have already proved that triangles are $0$-slim (the center is common to all sides of the triangle). The $0$-hyperbolicity follows from one of the equivalent characterizations of hyperbolicity (the other characterizations could be used as well, but the proofs would be less immediate.)\ subclass (in metric_tree_with_delta) Gromov_hyperbolic_space_0 proof (standard) show "deltaG TYPE('a) = 0" unfolding delta0 by auto have "Gromov_hyperbolic_subset (6 * 0) (UNIV::'a set)" proof (rule slim_triangles_implies_hyperbolic) fix x::'a and y z Gxy Gyz Gxz define w where "w = center x y z" assume "geodesic_segment_between Gxy x y" "geodesic_segment_between Gxz x z" "geodesic_segment_between Gyz y z" then have "Gxy = {x--y}" "Gyz = {y--z}" "Gxz = {x--z}" by (auto simp add: local.geodesic_segment_unique) then have "w \ Gxy" "w \ Gyz" "w \ Gxz" unfolding w_def by auto then have "infdist w Gxy \ 0 \ infdist w Gxz \ 0 \ infdist w Gyz \ 0" by auto then show "\w. infdist w Gxy \ 0 \ infdist w Gxz \ 0 \ infdist w Gyz \ 0" by blast qed then show "Gromov_hyperbolic_subset (deltaG TYPE('a)) (UNIV::'a set)" unfolding delta0 by auto qed text \To use the fact that reals are Gromov hyperbolic, given that they are a metric tree, we need to instantiate them as \verb+metric_tree_with_delta+.\ instantiation real::metric_tree_with_delta begin definition deltaG_real::"real itself \ real" where "deltaG_real _ = 0" instance apply standard unfolding deltaG_real_def by auto end text \Let us now prove the converse: a geodesic space which is $\delta$-hyperbolic for $\delta = 0$ is a metric tree. For the proof, we consider two geodesic segments $G = [x,y]$ and $H = [y,z]$ with a common endpoint, and we have to show that their union is still a geodesic segment from $x$ to $z$. For this, introduce a geodesic segment $L = [x,z]$. By the property of thin triangles, $G$ is included in $H \cup L$. In particular, a point $Y$ close to $y$ but different from $y$ on $G$ is on $L$, and therefore realizes the equality $d(x,z) = d(x, Y) + d(Y, z)$. Passing to the limit, $y$ also satisfies this equality. The conclusion readily follows thanks to Lemma \verb+geodesic_segment_union+. \ subclass (in Gromov_hyperbolic_space_0_geodesic) metric_tree proof fix G H x y z assume A: "geodesic_segment_between G x y" "geodesic_segment_between H y z" "G \ H = {y::'a}" show "geodesic_segment_between (G \ H) x z" proof (cases "x = y") case True then show ?thesis by (metis A Un_commute geodesic_segment_between_x_x(3) inf.commute sup_inf_absorb) next case False define D::"nat \ real" where "D = (\n. dist x y - (dist x y) * (1/(real(n+1))))" have D: "D n \ {0..< dist x y}" "D n \ {0..dist x y}" for n unfolding D_def by (auto simp add: False divide_simps algebra_simps) have Dlim: "D \ dist x y - dist x y * 0" unfolding D_def by (intro tendsto_intros LIMSEQ_ignore_initial_segment[OF lim_1_over_n, of 1]) define Y::"nat \ 'a" where "Y = (\n. geodesic_segment_param G x (D n))" have *: "Y \ y" unfolding Y_def apply (subst geodesic_segment_param(2)[OF A(1), symmetric]) using isometry_on_continuous[OF geodesic_segment_param(4)[OF A(1)]] unfolding continuous_on_sequentially comp_def using D(2) Dlim by auto have "dist x z = dist x (Y n) + dist (Y n) z" for n proof - obtain L where L: "geodesic_segment_between L x z" using geodesic_subsetD[OF geodesic] by blast have "Y n \ G" unfolding Y_def apply (rule geodesic_segment_param(3)[OF A(1)]) using D[of n] by auto have "dist x (Y n) = D n" unfolding Y_def apply (rule geodesic_segment_param[OF A(1)]) using D[of n] by auto then have "Y n \ y" using D[of n] by auto then have "Y n \ H" using A(3) \Y n \ G\ by auto have "infdist (Y n) (H \ L) \ 4 * deltaG(TYPE('a))" apply (rule thin_triangles[OF geodesic_segment_commute[OF A(2)] geodesic_segment_commute[OF L] geodesic_segment_commute[OF A(1)]]) using \Y n \ G\ by simp then have "infdist (Y n) (H \ L) = 0" using infdist_nonneg[of "Y n" "H \ L"] unfolding delta0 by auto have "Y n \ H \ L" proof (subst in_closed_iff_infdist_zero) have "closed H" using A(2) geodesic_segment_topology geodesic_segment_def by fastforce moreover have "closed L" using L geodesic_segment_topology geodesic_segment_def by fastforce ultimately show "closed (H \ L)" by auto show "H \ L \ {}" using A(2) geodesic_segment_endpoints(1) by auto qed (fact) then have "Y n \ L" using \Y n \ H\ by simp show ?thesis using geodesic_segment_dist[OF L \Y n \ L\] by simp qed moreover have "(\n. dist x (Y n) + dist (Y n) z) \ dist x y + dist y z" by (intro tendsto_intros *) ultimately have "(\n. dist x z) \ dist x y + dist y z" using filterlim_cong eventually_sequentially by auto then have *: "dist x z = dist x y + dist y z" using LIMSEQ_unique by auto show "geodesic_segment_between (G \ H) x z" by (rule geodesic_segment_union[OF * A(1) A(2)]) qed qed end (*of theory Gromov_Hyperbolic*) diff --git a/thys/Group-Ring-Module/Algebra5.thy b/thys/Group-Ring-Module/Algebra5.thy --- a/thys/Group-Ring-Module/Algebra5.thy +++ b/thys/Group-Ring-Module/Algebra5.thy @@ -1,5061 +1,5064 @@ (** Algebra5 author Hidetsune Kobayashi Group You Santo Department of Mathematics Nihon University h_coba@math.cst.nihon-u.ac.jp May 3, 2004. April 6, 2007 (revised). chapter 4. Ring theory (continued) section 6. operation of ideals section 7. direct product1, general case section 8. Chinese remainder theorem section 9. addition of finite elements of a ring and ideal_multiplication section 10. extension and contraction section 11. complete system of representatives section 12. Polynomial ring section 13. addition and multiplication of polyn_exprs section 14. the degree of a polynomial **) theory Algebra5 imports Algebra4 begin section "Operation of ideals" lemma (in Ring) ideal_sumTr1:"\ideal R A; ideal R B\ \ A \ B = \ {J. ideal R J \ (A \ B) \ J}" apply (frule sum_ideals[of "A" "B"], assumption, frule sum_ideals_la1[of "A" "B"], assumption, frule sum_ideals_la2[of "A" "B"], assumption) apply (rule equalityI) (* A \\<^sub>R B \ \{J. ideal R J \ A \ B \ J} *) apply (rule_tac A = "{J. ideal R J \ (A \ B) \ J}" and C = "A \ B" in Inter_greatest) apply (simp, (erule conjE)+) apply (rule_tac A = A and B = B and I = X in sum_ideals_cont, simp add:ideal_subset1, simp add:ideal_subset1, assumption+) apply (rule_tac B = "A \ B" and A = "{J. ideal R J \ (A \ B) \ J}" in Inter_lower) apply simp done lemma (in Ring) sum_ideals_commute:"\ideal R A; ideal R B\ \ A \ B = B \ A" apply (frule ideal_sumTr1 [of "B"], assumption+, frule ideal_sumTr1 [of "A" "B"], assumption+) apply (simp add:Un_commute[of "B" "A"]) done lemma (in Ring) ideal_prod_mono1:"\ideal R A; ideal R B; ideal R C; A \ B \ \ A \\<^sub>r C \ B \\<^sub>r C" apply (frule ideal_prod_ideal[of "B" "C"], assumption+) apply (rule ideal_prod_subTr[of "A" "C" "B \\<^sub>r C"], assumption+) apply (rule ballI)+ apply (frule_tac c = i in subsetD[of "A" "B"], assumption+) apply (rule_tac i = i and j = j in prod_mem_prod_ideals[of "B" "C"], assumption+) done lemma (in Ring) ideal_prod_mono2:"\ideal R A; ideal R B; ideal R C; A \ B \ \ C \\<^sub>r A \ C \\<^sub>r B" apply (frule ideal_prod_mono1[of "A" "B" "C"], assumption+) apply (simp add:ideal_prod_commute) done lemma (in Ring) cont_ideal_prod:"\ideal R A; ideal R B; ideal R C; A \ C; B \ C \ \ A \\<^sub>r B \ C" apply (simp add:ideal_prod_def) apply (rule subsetI, simp) apply (frule ideal_prod_ideal[of "A" "B"], assumption, frule_tac a = "A \\<^sub>r B" in forall_spec, thin_tac "\xa. ideal R xa \ {x. \i\A. \j\B. x = i \\<^sub>r j} \ xa \ x \ xa", simp) apply (rule subsetI, simp, (erule bexE)+, simp add:prod_mem_prod_ideals) apply (frule ideal_prod_la1[of "A" "B"], assumption, frule_tac c = x in subsetD[of "A \\<^sub>r B" "A"], assumption+, simp add:subsetD) done lemma (in Ring) ideal_distrib:"\ideal R A; ideal R B; ideal R C\ \ A \\<^sub>r (B \ C) = A \\<^sub>r B \ A \\<^sub>r C" apply (frule sum_ideals[of "B" "C"], assumption, frule ideal_prod_ideal[of "A" "B \ C"], assumption+, frule ideal_prod_ideal[of "A" "B"], assumption+, frule ideal_prod_ideal[of "A" "C"], assumption+, frule sum_ideals[of "A \\<^sub>r B" "A \\<^sub>r C"], assumption) apply (rule equalityI) apply (rule ideal_prod_subTr[of "A" "B \ C" "A \\<^sub>r B \ A \\<^sub>r C"], assumption+) apply (rule ballI)+ apply (frule_tac x = j in ideals_set_sum[of B C], assumption+, (erule bexE)+, simp) apply ( thin_tac "j = h \ k", frule_tac h = i in ideal_subset[of "A"], assumption+, frule_tac h = h in ideal_subset[of "B"], assumption+, frule_tac h = k in ideal_subset[of "C"], assumption+) apply (simp add:ring_distrib1) apply (frule_tac i = i and j = h in prod_mem_prod_ideals[of "A" "B"], assumption+, frule_tac i = i and j = k in prod_mem_prod_ideals[of "A" "C"], assumption+) apply (frule sum_ideals_la1[of "A \\<^sub>r B" "A \\<^sub>r C"], assumption+, frule sum_ideals_la2[of "A \\<^sub>r B" "A \\<^sub>r C"], assumption+) apply (frule_tac c = "i \\<^sub>r h" in subsetD[of "A \\<^sub>r B" "A \\<^sub>r B \ A \\<^sub>r C"], assumption+, frule_tac c = "i \\<^sub>r k" in subsetD[of "A \\<^sub>r C" "A \\<^sub>r B \ A \\<^sub>r C"], assumption+) apply (simp add:ideal_pOp_closed) apply (rule sum_ideals_cont[of "A \\<^sub>r (B \ C)" "A \\<^sub>r B" "A \\<^sub>r C" ], assumption+) apply (rule ideal_prod_subTr[of "A" "B" "A \\<^sub>r (B \ C)"], assumption+) apply (rule ballI)+ apply (frule sum_ideals_la1[of "B" "C"], assumption+, frule_tac c = j in subsetD[of "B" "B \ C"], assumption+, rule_tac i = i and j = j in prod_mem_prod_ideals[of "A" "B \ C"], assumption+) apply (rule ideal_prod_subTr[of "A" "C" "A \\<^sub>r (B \ C)"], assumption+) apply (rule ballI)+ apply (frule sum_ideals_la2[of "B" "C"], assumption+, frule_tac c = j in subsetD[of "C" "B \ C"], assumption+, rule_tac i = i and j = j in prod_mem_prod_ideals[of "A" "B \ C"], assumption+) done definition coprime_ideals::"[_, 'a set, 'a set] \ bool" where "coprime_ideals R A B \ A \\<^bsub>R\<^esub> B = carrier R" lemma (in Ring) coprimeTr:"\ideal R A; ideal R B\ \ coprime_ideals R A B = (\a \ A. \b \ B. a \ b = 1\<^sub>r)" apply (rule iffI) apply (simp add:coprime_ideals_def) apply (cut_tac ring_one, frule sym, thin_tac "A \ B = carrier R", simp, thin_tac "carrier R = A \ B", frule ideals_set_sum[of A B], assumption+, (erule bexE)+, frule sym, blast) apply (erule bexE)+ apply (frule ideal_subset1[of A], frule ideal_subset1[of B]) apply (frule_tac a = a and b = b in set_sum_mem[of _ A _ B], assumption+) apply (simp add:coprime_ideals_def) apply (frule sum_ideals[of "A" "B"], assumption+, frule ideal_inc_one[of "A \ B"], assumption) apply (simp add:coprime_ideals_def) done lemma (in Ring) coprime_int_prod:"\ideal R A; ideal R B; coprime_ideals R A B\ \ A \ B = A \\<^sub>r B" apply (frule ideal_prod_la1[of "A" "B"], assumption+, frule ideal_prod_la2[of "A" "B"], assumption+) apply (rule equalityI) defer (** A \\<^bsub>R\<^esub> B \ A \ B **) apply simp apply (simp add:coprime_ideals_def) apply (frule int_ideal[of "A" "B"], assumption) apply (frule idealprod_whole_r[of "A \ B"]) apply (frule sym, thin_tac "A \ B = carrier R", simp) apply (simp add:ideal_distrib) apply (simp add:ideal_prod_commute[of "A \ B" "A"]) apply (cut_tac Int_lower1[of "A" "B"], cut_tac Int_lower2[of "A" "B"]) apply (frule ideal_prod_mono2[of "A \ B" "B" "A"], assumption+, frule ideal_prod_mono1[of "A \ B" "A" "B"], assumption+) apply (frule ideal_prod_ideal[of "A \ B" "B"], assumption+, frule ideal_prod_ideal[of "A" "A \ B"], assumption+, frule ideal_subset1[of "A \\<^sub>r (A \ B)"], frule ideal_subset1[of "(A \ B) \\<^sub>r B"]) apply (frule ideal_prod_ideal[of A B], assumption+, frule sum_ideals_cont[of "A \\<^sub>r B" "A \\<^sub>r (A \ B)" "(A \ B) \\<^sub>r B"], assumption+) apply simp done lemma (in Ring) coprime_elems:"\ideal R A; ideal R B; coprime_ideals R A B\ \ \a\A. \b\B. a \ b = 1\<^sub>r" by (simp add:coprimeTr) lemma (in Ring) coprime_elemsTr:"\ideal R A; ideal R B; a\A; b\B; a \ b = 1\<^sub>r\ \ pj R A b = 1\<^sub>r\<^bsub>(qring R A)\<^esub> \ pj R B a = 1\<^sub>r\<^bsub>(qring R B)\<^esub>" apply (frule pj_Hom [OF Ring, of "A"], frule pj_Hom [OF Ring, of "B"]) apply (frule qring_ring[of "A"], frule qring_ring[of "B"]) apply (cut_tac ring_is_ag, frule Ring.ring_is_ag[of "R /\<^sub>r A"], frule Ring.ring_is_ag[of "R /\<^sub>r B"]) apply (frule_tac a = a and b = b in aHom_add[of "R" "R /\<^sub>r A" "pj R A"], assumption+, simp add:rHom_def, simp add:ideal_subset, simp add:ideal_subset, simp) apply (frule ideal_subset[of "A" "a"], assumption, frule ideal_subset[of "B" "b"], assumption+) apply (simp only:pj_zero[OF Ring, THEN sym, of "A" "a"], frule rHom_mem[of "pj R A" "R" "qring R A" "b"], assumption+, simp add:aGroup.l_zero[of "R /\<^sub>r A"]) apply (simp add:rHom_one[OF Ring, of "qring R A"]) apply (frule_tac a = a and b = b in aHom_add[of "R" "R /\<^sub>r B" "pj R B"], assumption+, simp add:rHom_def, simp add:ideal_subset, simp add:ideal_subset, simp) apply (simp only:pj_zero[OF Ring, THEN sym, of "B" "b"], frule rHom_mem[of "pj R B" "R" "qring R B" "a"], assumption+, simp add:aGroup.ag_r_zero[of "R /\<^sub>r B"]) apply (simp add:rHom_one[OF Ring, of "qring R B"]) done lemma (in Ring) partition_of_unity:"\ideal R A; a \ A; b \ carrier R; a \ b = 1\<^sub>r; u \ carrier R; v \ carrier R\ \ pj R A (a \\<^sub>r v \ b \\<^sub>r u ) = pj R A u" apply (frule pj_Hom [OF Ring, of "A"], frule ideal_subset [of "A" "a"], assumption+, frule ring_tOp_closed[of "a" "v"], assumption+, frule ring_tOp_closed[of "b" "u"], assumption+, frule qring_ring[of "A"]) apply (simp add:ringhom1[OF Ring, of "qring R A" "a \\<^sub>r v" "b \\<^sub>r u" "pj R A"]) apply (frule ideal_ring_multiple1[of "A" "a" "v"], assumption+, simp add:pj_zero[OF Ring, THEN sym, of "A" "a \\<^sub>r v"], frule rHom_mem[of "pj R A" "R" "qring R A" "b \\<^sub>r u"], assumption+, simp add:Ring.l_zero, simp add:rHom_tOp[OF Ring]) apply (frule ringhom1[OF Ring, of "qring R A" "a" "b" "pj R A"], assumption+, simp, simp add:pj_zero[OF Ring, THEN sym, of "A" "a"], frule rHom_mem[of "pj R A" "R" "qring R A" "b"], assumption+, simp add:Ring.l_zero) apply (simp add:rHom_one[OF Ring, of "qring R A" "pj R A"], rotate_tac -2, frule sym, thin_tac "1\<^sub>r\<^bsub>R /\<^sub>r A\<^esub> = pj R A b", simp, rule Ring.ring_l_one[of "qring R A" "pj R A u"], assumption) apply (simp add:rHom_mem) done lemma (in Ring) coprimes_commute:"\ideal R A; ideal R B; coprime_ideals R A B \ \ coprime_ideals R B A" apply (simp add:coprime_ideals_def) apply (simp add:sum_ideals_commute) done lemma (in Ring) coprime_surjTr:"\ideal R A; ideal R B; coprime_ideals R A B; X \ carrier (qring R A); Y \ carrier (qring R B) \ \ \r\carrier R. pj R A r = X \ pj R B r = Y" apply (frule qring_ring [of "A"], frule qring_ring [of "B"], frule coprime_elems [of "A" "B"], assumption+, frule pj_Hom [OF Ring, of "A"], frule pj_Hom [OF Ring, of "B"]) apply (erule bexE)+ apply (simp add:qring_carrier[of "A"], simp add:qring_carrier[of "B"], (erule bexE)+, rename_tac a b u v) apply (rotate_tac -1, frule sym, thin_tac "v \\<^bsub>R\<^esub> B = Y", rotate_tac -3, frule sym, thin_tac "u \\<^bsub>R\<^esub> A = X", simp) apply (frule_tac h = a in ideal_subset[of "A"], assumption+, frule_tac h = b in ideal_subset[of "B"], assumption+, frule_tac x = a and y = v in ring_tOp_closed, assumption+, frule_tac x = b and y = u in ring_tOp_closed, assumption+, cut_tac ring_is_ag, frule_tac x = "a \\<^sub>r v" and y = "b \\<^sub>r u" in aGroup.ag_pOp_closed[of "R"], assumption+) apply (frule_tac a = a and b = b and u = u and v = v in partition_of_unity[of "A"], assumption+, simp add:pj_mem[OF Ring, of "A"]) apply (frule_tac a = b and b = a and u = v and v = u in partition_of_unity[of "B"], assumption+, subst aGroup.ag_pOp_commute[of "R"], assumption+, simp add:pj_mem[OF Ring, of "B"]) apply (frule_tac x = "b \\<^sub>r u" and y = "a \\<^sub>r v" in aGroup.ag_pOp_commute[of "R"], assumption+, simp) apply (simp add:pj_mem[OF Ring], blast) done lemma (in Ring) coprime_n_idealsTr0:"\ideal R A; ideal R B; ideal R C; coprime_ideals R A C; coprime_ideals R B C \ \ coprime_ideals R (A \\<^sub>r B) C" apply (simp add:coprimeTr[of "A" "C"], simp add:coprimeTr[of "B" "C"], (erule bexE)+) apply (cut_tac ring_is_ag, frule_tac h = a in ideal_subset[of "A"], assumption+, frule_tac h = aa in ideal_subset[of "B"], assumption+, frule_tac h = b in ideal_subset[of "C"], assumption+, frule_tac h = ba in ideal_subset[of "C"], assumption+, frule_tac x = a and y = b in aGroup.ag_pOp_closed, assumption+) apply (frule_tac x = "a \ b" and y = aa and z = ba in ring_distrib1, assumption+) apply ( rotate_tac 6, frule sym, thin_tac "a \ b = 1\<^sub>r", frule sym, thin_tac "aa \ ba = 1\<^sub>r") apply (simp only:ring_distrib2) apply (rotate_tac -1, frule sym, thin_tac "1\<^sub>r = a \ b", simp, thin_tac "aa \ ba = 1\<^sub>r") apply (simp add:ring_r_one, thin_tac "a \\<^sub>r aa \ b \\<^sub>r aa \ (a \\<^sub>r ba \ b \\<^sub>r ba) \ carrier R", thin_tac "a \ b = a \\<^sub>r aa \ b \\<^sub>r aa \ (a \\<^sub>r ba \ b \\<^sub>r ba)") apply (frule_tac x = a and y = aa in ring_tOp_closed, assumption+, frule_tac x = b and y = aa in ring_tOp_closed, assumption+, frule_tac x = a and y = ba in ring_tOp_closed, assumption+, frule_tac x = b and y = ba in ring_tOp_closed, assumption+, frule_tac x = "a \\<^sub>r ba" and y = "b \\<^sub>r ba" in aGroup.ag_pOp_closed, assumption+) apply (simp add:aGroup.ag_pOp_assoc, frule sym, thin_tac "1\<^sub>r = a \\<^sub>r aa \ (b \\<^sub>r aa \ (a \\<^sub>r ba \ b \\<^sub>r ba))") apply (frule_tac i = a and j = aa in prod_mem_prod_ideals[of "A" "B"], assumption+) apply (frule_tac x = ba and r = a in ideal_ring_multiple[of "C"], assumption+, frule_tac x = ba and r = b in ideal_ring_multiple[of "C"], assumption+, frule_tac x = b and r = aa in ideal_ring_multiple1[of "C"], assumption+) apply (frule_tac I = C and x = "a \\<^sub>r ba" and y = "b \\<^sub>r ba" in ideal_pOp_closed, assumption+, frule_tac I = C and x = "b \\<^sub>r aa" and y = "a \\<^sub>r ba \ b \\<^sub>r ba" in ideal_pOp_closed, assumption+) apply (frule ideal_prod_ideal[of "A" "B"], assumption) apply (subst coprimeTr, assumption+, blast) done lemma (in Ring) coprime_n_idealsTr1:"ideal R C \ (\k \ n. ideal R (J k)) \ (\i \ n. coprime_ideals R (J i) C) \ coprime_ideals R (i\\<^bsub>R,n\<^esub> J) C" apply (induct_tac n) apply (rule impI) apply (erule conjE) apply simp apply (rule impI) apply (erule conjE) apply (cut_tac n = "Suc n" in n_in_Nsetn) apply (cut_tac n = n in Nset_Suc) apply simp apply (cut_tac n = n and J = J in n_prod_ideal, rule allI, simp) apply (rule_tac A = "i\\<^bsub>R,n\<^esub> J" and B = "J (Suc n)" in coprime_n_idealsTr0[of _ _ "C"], simp+) done lemma (in Ring) coprime_n_idealsTr2:"\ideal R C; (\k \ n. ideal R (J k)); (\i \ n. coprime_ideals R (J i) C) \ \ coprime_ideals R (i\\<^bsub>R,n\<^esub> J) C" by (simp add:coprime_n_idealsTr1) lemma (in Ring) coprime_n_idealsTr3:"(\k \ (Suc n). ideal R (J k)) \ (\i \ (Suc n). \j \ (Suc n). i \ j \ coprime_ideals R (J i) (J j)) \ coprime_ideals R (i\\<^bsub>R,n\<^esub> J) (J (Suc n))" apply (rule impI, erule conjE) apply (case_tac "n = 0", simp) apply (simp add:Nset_1) apply (cut_tac nat_eq_le[of "Suc n" "Suc n"], frule_tac a = "Suc n" in forall_spec, assumption) apply (rotate_tac 1, frule_tac a = "Suc n" in forall_spec, assumption, thin_tac "\j\Suc n. Suc n \ j \ coprime_ideals R (J (Suc n)) (J j)") apply (rule_tac C = "J (Suc n)" and n = n and J = J in coprime_n_idealsTr2, assumption, rule allI, simp) apply (rule allI, rule impI) apply (frule_tac a = i in forall_spec, simp, thin_tac "\i\Suc n. \j\Suc n. i \ j \ coprime_ideals R (J i) (J j)", rotate_tac -1, frule_tac a = "Suc n" in forall_spec, assumption+) apply simp+ done lemma (in Ring) coprime_n_idealsTr4:"\(\k \ (Suc n). ideal R (J k)) \ (\i \ (Suc n). \j \ (Suc n). i \ j \ coprime_ideals R (J i) (J j))\ \ coprime_ideals R (i\\<^bsub>R,n\<^esub> J) (J (Suc n))" apply (simp add:coprime_n_idealsTr3) done section "Direct product1, general case" definition prod_tOp :: "['i set, 'i \ ('a, 'm) Ring_scheme] \ ('i \ 'a) \ ('i \ 'a) \ ('i \ 'a)" where "prod_tOp I A = (\f\carr_prodag I A. \g\carr_prodag I A. \x\I. (f x) \\<^sub>r\<^bsub>(A x)\<^esub> (g x))" (** Let x \ I, then A x is a ring, {A x | x \ I} is a set of rings. **) definition prod_one::"['i set, 'i \ ('a, 'm) Ring_scheme] \ ('i \ 'a)" where "prod_one I A == \x\I. 1\<^sub>r\<^bsub>(A x)\<^esub>" (** 1\<^sub>(A x) is the unit of a ring A x. **) definition prodrg :: "['i set, 'i \ ('a, 'more) Ring_scheme] \ ('i \ 'a) Ring" where "prodrg I A = \carrier = carr_prodag I A, pop = prod_pOp I A, mop = prod_mOp I A, zero = prod_zero I A, tp = prod_tOp I A, un = prod_one I A \" (** I is the index set **) abbreviation PRODRING ("(r\\<^bsub>_\<^esub>/ _)" [72,73]72) where "r\\<^bsub>I\<^esub> A == prodrg I A" definition augm_func :: "[nat, nat \ 'a,'a set, nat, nat \ 'a, 'a set] \ nat \ 'a" where "augm_func n f A m g B = (\i\{j. j \ (n + m)}. if i \ n then f i else if (Suc n) \ i \ i \ n + m then g ((sliden (Suc n)) i) else undefined)" (* Remark. g is a function of Nset (m - 1) \ B *) definition ag_setfunc :: "[nat, nat \ ('a, 'more) Ring_scheme, nat, nat \ ('a, 'more) Ring_scheme] \ (nat \ 'a) set \ (nat \ 'a) set \ (nat \ 'a) set" where "ag_setfunc n B1 m B2 X Y = {f. \g. \h. (g\X) \(h\Y) \(f = (augm_func n g (Un_carrier {j. j \ n} B1) m h (Un_carrier {j. j \ (m - 1)} B2)))}" (* Later we consider X = ac_Prod_Rg (Nset n) B1 and Y = ac_Prod_Rg (Nset (m - 1)) B2 *) primrec ac_fProd_Rg :: "[nat, nat \ ('a, 'more) Ring_scheme] \ (nat \ 'a) set" where fprod_0: "ac_fProd_Rg 0 B = carr_prodag {0::nat} B" | frpod_n: "ac_fProd_Rg (Suc n) B = ag_setfunc n B (Suc 0) (compose {0::nat} B (slide (Suc n))) (carr_prodag {j. j \ n} B) (carr_prodag {0} (compose {0} B (slide (Suc n))))" definition prodB1 :: "[('a, 'm) Ring_scheme, ('a, 'm) Ring_scheme] \ (nat \ ('a, 'm) Ring_scheme)" where "prodB1 R S = (\k. if k=0 then R else if k=Suc 0 then S else undefined)" definition Prod2Rg :: "[('a, 'm) Ring_scheme, ('a, 'm) Ring_scheme] \ (nat \ 'a) Ring" (infixl "\\<^sub>r" 80) where "A1 \\<^sub>r A2 = prodrg {0, Suc 0} (prodB1 A1 A2)" text \Don't try \(Prod_ring (Nset n) B) \\<^sub>r (B (Suc n))\\ lemma carr_prodrg_mem_eq:"\f \ carrier (r\\<^bsub>I\<^esub> A); g \ carrier (r\\<^bsub>I\<^esub> A); \i\I. f i = g i \ \ f = g" by (simp add:prodrg_def carr_prodag_def, (erule conjE)+, rule funcset_eq[of _ I], assumption+) lemma prod_tOp_mem:"\\k\I. Ring (A k); X \ carr_prodag I A; Y \ carr_prodag I A\ \ prod_tOp I A X Y \ carr_prodag I A" apply (subst carr_prodag_def, simp) apply (rule conjI) apply (simp add:prod_tOp_def restrict_def extensional_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:Un_carrier_def prod_tOp_def) apply (simp add:carr_prodag_def, (erule conjE)+) apply (blast dest: Ring.ring_tOp_closed del:PiE) apply (rule ballI) apply (simp add:prod_tOp_def) apply (simp add:carr_prodag_def, (erule conjE)+) apply (simp add:Ring.ring_tOp_closed) done lemma prod_tOp_func:"\k\I. Ring (A k) \ prod_tOp I A \ carr_prodag I A \ carr_prodag I A \ carr_prodag I A" by (simp add:prod_tOp_mem) lemma prod_one_func:"\k\I. Ring (A k) \ prod_one I A \ carr_prodag I A" apply (simp add:prod_one_def carr_prodag_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:Un_carrier_def) apply (blast dest: Ring.ring_one) apply (simp add: Ring.ring_one) done lemma prodrg_carrier:"\k\I. Ring (A k) \ carrier (prodrg I A) = carrier (prodag I A)" by (simp add:prodrg_def prodag_def) lemma prodrg_ring:"\k\I. Ring (A k) \ Ring (prodrg I A)" apply (rule Ring.intro) apply (simp add:prodrg_def) apply (rule prod_pOp_func, rule ballI, simp add:Ring.ring_is_ag) apply (simp add:prodrg_def, rule prod_pOp_assoc, simp add:Ring.ring_is_ag, assumption+) apply (simp add:prodrg_def, rule prod_pOp_commute, simp add:Ring.ring_is_ag, assumption+) apply (simp add:prodrg_def, rule prod_mOp_func, simp add:Ring.ring_is_ag) apply (simp add:prodrg_def) apply (cut_tac X = a in prod_mOp_mem[of "I" "A"]) apply (simp add:Ring.ring_is_ag, assumption) apply (cut_tac X = "prod_mOp I A a" and Y = a in prod_pOp_mem[of "I" "A"], simp add:Ring.ring_is_ag, assumption+, cut_tac prod_zero_func[of "I" "A"]) apply (rule carr_prodag_mem_eq[of "I" "A"], simp add:Ring.ring_is_ag, assumption+, rule ballI, simp add:prod_pOp_def, subst prod_mOp_mem_i, simp add:Ring.ring_is_ag, assumption+, subst prod_zero_i, simp add:Ring.ring_is_ag, assumption+, rule aGroup.l_m, simp add:Ring.ring_is_ag, simp add:prodag_comp_i, simp add:Ring.ring_is_ag) apply (simp add:prodrg_def, rule prod_zero_func, simp add:Ring.ring_is_ag) apply (simp add:prodrg_def, cut_tac prod_zero_func[of "I" "A"], cut_tac X = "prod_zero I A" and Y = a in prod_pOp_mem[of "I" "A"], simp add:Ring.ring_is_ag, assumption+, rule carr_prodag_mem_eq[of "I" "A"], simp add:Ring.ring_is_ag, assumption+) apply (rule ballI) apply (simp add:prod_pOp_def prod_zero_def) apply (rule aGroup.ag_l_zero, simp add:Ring.ring_is_ag) apply (simp add:prodag_comp_i, simp add:Ring.ring_is_ag) apply (simp add:prodrg_def, rule prod_tOp_func, simp add:Ring.ring_is_ag) apply (simp add:prodrg_def) apply (frule_tac X = a and Y = b in prod_tOp_mem[of "I" "A"], assumption+, frule_tac X = "prod_tOp I A a b" and Y = c in prod_tOp_mem[of "I" "A"], assumption+, frule_tac X = b and Y = c in prod_tOp_mem[of "I" "A"], assumption+, frule_tac X = a and Y = "prod_tOp I A b c" in prod_tOp_mem[of "I" "A"], assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], simp add:Ring.ring_is_ag, assumption+, rule ballI) apply (simp add:prod_tOp_def) apply (rule Ring.ring_tOp_assoc, simp, (simp add:prodag_comp_i)+) apply (simp add:prodrg_def) apply (frule_tac X = a and Y = b in prod_tOp_mem[of "I" "A"], assumption+, frule_tac X = b and Y = a in prod_tOp_mem[of "I" "A"], assumption+, rule carr_prodag_mem_eq[of "I" "A"], simp add:Ring.ring_is_ag, assumption+) apply (rule ballI, simp add:prod_tOp_def) apply (rule Ring.ring_tOp_commute, (simp add:prodag_comp_i)+) apply (simp add:prodrg_def, rule prod_one_func, assumption) apply (simp add:prodrg_def) apply (cut_tac X = b and Y = c in prod_pOp_mem[of "I" "A"], simp add:Ring.ring_is_ag, assumption+, cut_tac X = a and Y = b in prod_tOp_mem[of "I" "A"], assumption+, cut_tac X = a and Y = c in prod_tOp_mem[of "I" "A"], assumption+, cut_tac X = "prod_tOp I A a b" and Y = "prod_tOp I A a c" in prod_pOp_mem[of "I" "A"], simp add:Ring.ring_is_ag, assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], simp add:Ring.ring_is_ag, rule prod_tOp_mem[of "I" "A"], assumption+) apply (rule ballI, simp add:prod_tOp_def prod_pOp_def) apply (rule Ring.ring_distrib1, (simp add:prodag_comp_i)+) apply (simp add:prodrg_def, cut_tac prod_one_func[of "I" "A"], cut_tac X = "prod_one I A" and Y = a in prod_tOp_mem[of "I" "A"], assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], simp add:Ring.ring_is_ag, assumption+, rule ballI, simp add:prod_tOp_def prod_one_def, rule Ring.ring_l_one, simp, simp add:prodag_comp_i) apply simp done lemma prodrg_elem_extensional:"\\k\I. Ring (A k); f \ carrier (prodrg I A)\ \ f \ extensional I" apply (simp add:prodrg_carrier) apply (simp add:prodag_def carr_prodag_def) done lemma prodrg_pOp:"\k\I. Ring (A k) \ pop (prodrg I A) = prod_pOp I A" apply (simp add:prodrg_def) done lemma prodrg_mOp:"\k\I. Ring (A k) \ mop (prodrg I A) = prod_mOp I A" apply (simp add:prodrg_def) done lemma prodrg_zero:"\k\I. Ring (A k) \ zero (prodrg I A) = prod_zero I A" apply (simp add:prodrg_def) done lemma prodrg_tOp:"\k\I. Ring (A k) \ tp (prodrg I A) = prod_tOp I A" apply (simp add:prodrg_def) done lemma prodrg_one:"\k\I. Ring (A k) \ un (prodrg I A) = prod_one I A" apply (simp add:prodrg_def) done lemma prodrg_sameTr5:"\\k\I. Ring (A k); \k\I. A k = B k\ \ prod_tOp I A = prod_tOp I B" apply (frule prod_tOp_func) apply (subgoal_tac "carr_prodag I A = carr_prodag I B", simp) apply (frule prod_tOp_func [of "I" "B"]) apply (rule funcset_eq [of _ "carr_prodag I B" _]) apply (simp add:prod_tOp_def extensional_def) apply (simp add:prod_tOp_def extensional_def) apply (rule ballI) apply (frule_tac x = x in funcset_mem [of "prod_tOp I A" "carr_prodag I B" "carr_prodag I B \ carr_prodag I B"], assumption+) apply (frule_tac x = x in funcset_mem [of "prod_tOp I B" "carr_prodag I B" "carr_prodag I B \ carr_prodag I B"], assumption+) apply (thin_tac " prod_tOp I A \ carr_prodag I B \ carr_prodag I B \ carr_prodag I B") apply (thin_tac "prod_tOp I B \ carr_prodag I B \ carr_prodag I B \ carr_prodag I B") apply (rule funcset_eq [of _ "carr_prodag I B"]) apply (simp add:prod_tOp_def extensional_def) apply (simp add:prod_tOp_def extensional_def) apply (rule ballI) apply (frule_tac f = "prod_tOp I A x" and A = "carr_prodag I B" and x = xa in funcset_mem, assumption+) apply (frule_tac f = "prod_tOp I B x" and A = "carr_prodag I B" and x = xa in funcset_mem, assumption+) apply (subgoal_tac "\k\I. aGroup (B k)") apply (rule carr_prodag_mem_eq, assumption+) apply (rule ballI, simp add:prod_tOp_def) apply (rule ballI, rule Ring.ring_is_ag, simp) apply (subgoal_tac "\k\I. aGroup (A k)") apply (simp add:prodag_sameTr1) apply (rule ballI, rule Ring.ring_is_ag, simp) done lemma prodrg_sameTr6:"\\k\I. Ring (A k); \k\I. A k = B k\ \ prod_one I A = prod_one I B" apply (frule prod_one_func [of "I" "A"]) apply (cut_tac prodag_sameTr1[of "I" "A" "B"]) apply (rule carr_prodag_mem_eq[of I A "prod_one I A" "prod_one I B"]) apply (simp add:Ring.ring_is_ag, assumption) apply (cut_tac prod_one_func [of "I" "B"], simp) apply simp apply (rule ballI, simp add:prod_one_def) apply (simp add:Ring.ring_is_ag, simp) done lemma prodrg_same:"\\k\I. Ring (A k); \k\I. A k = B k\ \ prodrg I A = prodrg I B" apply (subgoal_tac "\k\I. aGroup (A k)") apply (frule prodag_sameTr1, assumption+) apply (frule prodag_sameTr2, assumption+) apply (frule prodag_sameTr3, assumption+) apply (frule prodag_sameTr4, assumption+) apply (frule prodrg_sameTr5, assumption+) apply (frule prodrg_sameTr6, assumption+) apply (simp add:prodrg_def) apply (rule ballI, rule Ring.ring_is_ag, simp) done lemma prodrg_component:"\f \ carrier (prodrg I A); i \ I\ \ f i \ carrier (A i)" by (simp add:prodrg_def carr_prodag_def) lemma project_rhom:"\\k\I. Ring (A k); j \ I\ \ PRoject I A j \ rHom ( prodrg I A) (A j)" apply (simp add:rHom_def) apply (rule conjI) apply (simp add:aHom_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:prodrg_carrier) apply (cut_tac prodag_carrier[of I A], simp) apply (simp add:PRoject_def) apply (cut_tac prodag_carrier[of I A], simp, thin_tac "carrier (a\\<^bsub>I\<^esub> A) = carr_prodag I A") apply (simp add:prodag_comp_i) apply (simp add:Ring.ring_is_ag) apply (simp add:Ring.ring_is_ag) apply (subgoal_tac "\k\I. aGroup (A k)") apply (frule project_aHom [of "I" "A" "j"], assumption+) apply (rule conjI) apply (simp add:prodrg_carrier aHom_def) apply (simp add:prodrg_carrier) apply (simp add:prodrg_pOp) apply (simp add:prodag_pOp[THEN sym]) apply (simp add:aHom_def) apply (rule ballI, simp add:Ring.ring_is_ag) apply (rule conjI) apply (rule ballI)+ apply (simp add:prodrg_carrier) apply (cut_tac prodag_carrier[of I A], simp) apply (frule_tac X = x and Y = y in prod_tOp_mem[of I A], assumption+) apply (simp add:prodrg_tOp) apply (simp add:PRoject_def) apply (simp add:prod_tOp_def) apply (rule ballI) apply (simp add:Ring.ring_is_ag) apply (frule prodrg_ring [of "I" "A"]) apply (frule Ring.ring_one[of "r\\<^bsub>I\<^esub> A"]) apply (simp add:prodrg_carrier) apply (cut_tac prodag_carrier[of I A], simp) apply (simp add:PRoject_def) apply (simp add:prodrg_def) apply (fold prodrg_def) apply (simp add:prod_one_def) apply (rule ballI) apply (simp add:Ring.ring_is_ag) done lemma augm_funcTr:"\\k \(Suc n). Ring (B k); f \ carr_prodag {i. i \ (Suc n)} B\ \ f = augm_func n (restrict f {i. i \ n}) (Un_carrier {i. i \ n} B) (Suc 0) (\x\{0::nat}. f (x + Suc n)) (Un_carrier {0} (compose {0} B (slide (Suc n))))" apply (rule carr_prodag_mem_eq [of "{i. i \ (Suc n)}" "B" "f" "augm_func n (restrict f {i. i \ n}) (Un_carrier {i. i \ n} B) (Suc 0) (\x\{0}. f (x + Suc n)) (Un_carrier {0} (compose {0} B (slide (Suc n))))"]) apply (rule ballI, simp add:Ring.ring_is_ag) apply (simp add:augm_func_def) apply (simp add:carr_prodag_def) apply (rule conjI) apply (simp add:augm_func_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:augm_func_def sliden_def) apply (erule conjE)+ apply (frule_tac x = x in funcset_mem[of f "{i. i \ Suc n}" "Un_carrier {i. i \ Suc n} B"]) apply simp apply simp apply (rule impI) apply (rule_tac x = "Suc n" in funcset_mem[of f "{i. i \ Suc n}" "Un_carrier {i. i \ Suc n} B"], assumption) apply simp apply (rule allI, (erule conjE)+) apply (simp add:augm_func_def) apply (case_tac "i \ n", simp add:sliden_def) apply (simp add:sliden_def, rule impI) apply (simp add:nat_not_le_less, frule_tac m = n and n = i in Suc_leI, frule_tac m = i and n = "Suc n" in Nat.le_antisym, assumption+, simp) apply (rule ballI, simp) apply (simp add:augm_func_def sliden_def) apply (rule impI, simp add:nat_not_le_less) apply (frule_tac n = l in Suc_leI[of n _]) apply (frule_tac m = l and n = "Suc n" in Nat.le_antisym, assumption+, simp) done lemma A_to_prodag_mem:"\Ring A; \k\I. Ring (B k); \k\I. (S k) \ rHom A (B k); x \ carrier A \ \ A_to_prodag A I S B x \ carr_prodag I B" apply (simp add:carr_prodag_def) apply (rule conjI) apply (simp add:A_to_prodag_def extensional_def) apply (rule conjI) apply (rule Pi_I) apply (simp add: A_to_prodag_def) apply (subgoal_tac "(S xa) \ rHom A (B xa)") prefer 2 apply simp apply (thin_tac "\k\I. S k \ rHom A (B k)") apply (frule_tac f = "S xa" and A = A and R = "B xa" in rHom_mem, assumption+) apply (simp add:Un_carrier_def) apply blast apply (rule ballI) apply (simp add:A_to_prodag_def) apply (rule_tac f = "S i" and A = A and R = "B i" and a = x in rHom_mem) apply simp apply assumption done lemma A_to_prodag_rHom:"\Ring A; \k\I. Ring (B k); \k\I. (S k) \ rHom A (B k) \ \ A_to_prodag A I S B \ rHom A (r\\<^bsub>I\<^esub> B)" apply (simp add:rHom_def [of "A" "r\\<^bsub>I\<^esub> B"]) apply (rule conjI) apply (cut_tac A_to_prodag_aHom[of A I B S]) apply (subst aHom_def, simp) apply (simp add:prodrg_carrier) apply (simp add:aHom_def) apply (simp add:prodrg_def) apply (cut_tac prodag_pOp[of I B], simp) apply (rule ballI, simp add:Ring.ring_is_ag, simp add:Ring.ring_is_ag, rule ballI, simp add:Ring.ring_is_ag) apply (rule ballI) apply (simp add:rHom_def) apply (rule conjI) apply (rule ballI)+ apply (frule_tac x = x and y = y in Ring.ring_tOp_closed[of A], assumption+) apply (frule_tac x = "x \\<^sub>r\<^bsub>A\<^esub> y" in A_to_prodag_mem[of A I B S], assumption+, frule_tac x = x in A_to_prodag_mem[of A I B S], assumption+, frule_tac x = y in A_to_prodag_mem[of A I B S], assumption+) apply (simp add:prodrg_tOp[of I B]) apply (frule_tac X = "A_to_prodag A I S B x " and Y = "A_to_prodag A I S B y" in prod_tOp_mem[of I B], assumption+) apply (cut_tac X = "A_to_prodag A I S B (x \\<^sub>r\<^bsub>A\<^esub> y)" and Y = "prod_tOp I B (A_to_prodag A I S B x) (A_to_prodag A I S B y)" in carr_prodag_mem_eq[of I B]) apply (rule ballI, simp add:Ring.ring_is_ag) apply assumption+ apply (rule ballI, simp add:prod_tOp_def A_to_prodag_def) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. Ring (B k)", frule_tac x = l in bspec, assumption, thin_tac "\k\I. S k \ rHom A (B k)") apply (simp add:rHom_tOp) apply simp apply (simp add:prodrg_one[of I B]) apply (frule prod_one_func[of I B]) apply (frule Ring.ring_one[of A], frule_tac x = "1\<^sub>r\<^bsub>A\<^esub>" in A_to_prodag_mem[of A I B S], assumption+) apply (cut_tac X = "A_to_prodag A I S B 1\<^sub>r\<^bsub>A\<^esub>" and Y = "prod_one I B" in carr_prodag_mem_eq[of I B]) apply (rule ballI, simp add:Ring.ring_is_ag) apply assumption+ apply (rule ballI) apply (subst A_to_prodag_def, simp add:prod_one_def) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. Ring (B k)", frule_tac x = l in bspec, assumption, thin_tac "\k\I. S k \ rHom A (B k)") apply (simp add:rHom_one) apply assumption done lemma ac_fProd_ProdTr1:"\k \ (Suc n). Ring (B k) \ ag_setfunc n B (Suc 0) (compose {0::nat} B (slide (Suc n))) (carr_prodag {i. i \ n} B) (carr_prodag {0} (compose {0} B (slide (Suc n)))) \ carr_prodag {i. i \ (Suc n)} B" supply [[simproc del: defined_all]] apply (rule subsetI) apply (simp add:ag_setfunc_def) apply (erule exE, erule conjE, erule exE, erule conjE) apply (simp, thin_tac "x = augm_func n g (Un_carrier {j. j \ n} B) (Suc 0) h (Un_carrier {0} (compose {0} B (slide (Suc n))))") apply (simp add:carr_prodag_def [of "{j. j \ (Suc n)}" "B"]) apply (rule conjI) apply (simp add:augm_func_def) apply (rule conjI) apply (simp add:Pi_def) apply (rule allI) apply (rule impI) apply (simp add:augm_func_def) apply (case_tac "x \ n") apply simp apply (simp add:carr_prodag_def) apply (erule conjE)+ apply (frule_tac x = x in mem_of_Nset [of _ "n"]) apply (frule_tac f = g and x = x in funcset_mem[of _ "{j. j \ n}" "Un_carrier {j. j \ n} B"], assumption+) apply (simp add:Un_carrier_def, erule exE, erule conjE, erule exE, simp, erule conjE, frule_tac x = i and y = n and z = "Suc n" in le_less_trans, simp, frule_tac x = i and y = "Suc n" in less_imp_le, blast) apply (simp add:sliden_def) apply (simp add:carr_prodag_def Un_carrier_def, (erule conjE)+) apply (simp add:compose_def slide_def) apply (cut_tac n_in_Nsetn[of "Suc n"], blast) apply (rule allI, rule impI) apply (simp add:augm_func_def) apply (case_tac "i \ n", simp) apply (simp add:carr_prodag_def [of "{i. i \ n}" _]) apply simp apply (thin_tac "g \ carr_prodag {i. i \ n} B") apply (simp add: not_less [symmetric, of _ n], frule_tac n = i in Suc_leI[of n], frule_tac m = i and n = "Suc n" in le_antisym, assumption+, simp) apply (simp add:carr_prodag_def compose_def slide_def sliden_def) done lemma ac_fProd_Prod:"\k \ n. Ring (B k) \ ac_fProd_Rg n B = carr_prodag {j. j \ n} B" apply (case_tac "n = 0") apply simp apply (subgoal_tac "\m. n = Suc m") apply (subgoal_tac "\m. n = Suc m \ ac_fProd_Rg n B = carr_prodag {j. j \ n} B") apply blast apply (thin_tac "\m. n = Suc m") apply (rule allI, rule impI) apply (simp, thin_tac "n = Suc m") apply (rule equalityI) apply (simp add:ac_fProd_ProdTr1) apply (rule subsetI) apply (rename_tac m f) apply (frule augm_funcTr, assumption+) apply (simp add:ag_setfunc_def) apply (subgoal_tac "(restrict f {j. j \ m}) \ carr_prodag {j. j \ m} B") apply (subgoal_tac "(\x\{0::nat}. f (Suc (x + m))) \ carr_prodag {0} (compose {0} B (slide (Suc m)))") apply blast apply (thin_tac "f = augm_func m (restrict f {i. i \ m}) (Un_carrier {i. i \ m} B) (Suc 0) (\x\{0}. f (Suc (x + m))) (Un_carrier {0} (compose {0} B (slide (Suc m))))") apply (simp add:carr_prodag_def) apply (rule conjI) apply (simp add:Pi_def restrict_def) apply (simp add:Un_carrier_def compose_def slide_def) apply (simp add:compose_def slide_def) apply (thin_tac "f = augm_func m (restrict f {i. i \ m}) (Un_carrier {i. i \ m} B) (Suc 0) (\x\{0}. f (Suc (x + m))) (Un_carrier {0} (compose {0} B (slide (Suc m))))") apply (simp add:carr_prodag_def) apply (simp add:Un_carrier_def) apply (simp add:Pi_def) apply (rule allI) apply (rule impI) apply (erule conjE)+ apply (rotate_tac 1) apply (frule_tac a = x in forall_spec, simp) apply (erule exE, thin_tac "\x\Suc m. \xa. (\i\Suc m. xa = carrier (B i)) \ f x \ xa") apply (frule_tac a = x in forall_spec, simp) apply blast apply (cut_tac t = n in Suc_pred[THEN sym], simp) apply blast done text\A direct product of a finite number of rings defined with \ac_fProd_Rg\ is equal to that defined by using \carr_prodag\.\ definition fprodrg :: "[nat, nat \ ('a, 'more) Ring_scheme] \ \carrier:: (nat \ 'a) set, pop::[(nat \ 'a), (nat \ 'a)] \ (nat \ 'a), mop:: (nat \ 'a) \ (nat \ 'a), zero::(nat \ 'a), tp :: [(nat \ 'a), (nat \ 'a)] \ (nat \ 'a), un :: (nat \ 'a) \" where "fprodrg n B = \ carrier = ac_fProd_Rg n B, pop = \f. \g. prod_pOp {i. i \ n} B f g, mop = \f. prod_mOp {i. i \ n} B f, zero = prod_zero {i. i \ n} B, tp = \f. \g. prod_tOp {i. i \ n} B f g, un = prod_one {i. i \ n} B \" definition fPRoject ::"[nat, nat \ ('a, 'more) Ring_scheme, nat] \ (nat \ 'a) \ 'a" where "fPRoject n B x = (\f\ac_fProd_Rg n B. f x)" lemma fprodrg_ring:"\k \ n. Ring (B k) \ Ring (fprodrg n B)" apply (simp add:fprodrg_def) apply (frule ac_fProd_Prod) apply simp apply (fold prodrg_def) apply (simp add:prodrg_ring) done section "Chinese remainder theorem" lemma Chinese_remTr1:"\Ring A; \k \ (n::nat). ideal A (J k); \k \ n. B k = qring A (J k); \k \ n. S k = pj A (J k) \ \ ker\<^bsub>A,(r\\<^bsub>{j. j \ n}\<^esub> B)\<^esub> (A_to_prodag A {j. j \ n} S B) = \ {I. \k\{j. j \ n}. I = (J k)}" apply (rule equalityI) apply (rule subsetI) apply (simp add:ker_def) apply (rule allI, rule impI) apply (rename_tac a K, erule conjE) apply (simp add:prodrg_def, simp add:A_to_prodag_def prod_one_def) apply (erule exE, erule conjE) apply (subgoal_tac "(\k\{j. j \ n}. S k a) k = (\x\{j. j \ n}. \\<^bsub>B x\<^esub>) k") apply (thin_tac "(\k\{j. j \ n}. S k a) = prod_zero {j. j \ n} B") apply simp apply (frule_tac I = "J k" in Ring.qring_zero [of "A"]) apply simp apply (frule_tac I = "J k" and x = a in pj_mem [of "A"]) apply simp apply assumption apply simp apply (frule_tac I = "J k" and a = a in Ring.a_in_ar_coset [of "A"]) apply simp apply assumption apply simp apply (simp add:prod_zero_def) apply (rule subsetI) apply (simp add:CollectI ker_def) apply (cut_tac Nset_inc_0[of n]) apply (frule_tac a = "J 0" in forall_spec, blast) apply (frule_tac x = 0 in spec, simp) apply (frule_tac h = x in Ring.ideal_subset [of "A" "J 0"], simp+) apply (thin_tac "x \ J 0") apply (simp add:A_to_prodag_def prodrg_def) apply (simp add:prod_zero_def) apply (rule funcset_eq [of _ "{j. j \ n}"]) apply (simp add:extensional_def restrict_def)+ apply (rule allI, rule impI) apply (simp add:Ring.qring_zero) apply (frule_tac a = xa in forall_spec, assumption, thin_tac "\k \ n. ideal A (J k)") apply (subst pj_mem [of "A"], assumption+) apply (frule_tac I = "J xa" and a = x in Ring.a_in_ar_coset [of "A"], assumption+) apply (rule_tac a = x and I = "J xa" in Ring.Qring_fix1 [of "A"], assumption+) apply blast done lemma (in Ring) coprime_prod_int2Tr: "((\k \ (Suc n). ideal R (J k)) \ (\i \ (Suc n). \j \ (Suc n). (i \j \ coprime_ideals R (J i) (J j)))) \ (\ {I. \k \ (Suc n). I = (J k)} = ideal_n_prod R (Suc n) J)" apply (induct_tac n) apply (rule impI) apply (erule conjE) apply (simp add:Nset_1) apply (subst coprime_int_prod [THEN sym, of "J 0" "J (Suc 0)"], simp+) apply (rule equalityI, rule subsetI) apply (simp, blast) apply (rule subsetI, simp, rule allI, rule impI, erule exE, (erule conjE)+) apply simp apply (simp add:Nset_1_1, erule disjE, (simp del:ideal_n_prodSn)+) (** n **) apply (rule impI) apply (subgoal_tac "\{I. \k \ (Suc (Suc n)). I = J k} = (\{I. \k \ (Suc n). I = J k}) \ (J (Suc (Suc n)))") apply (subgoal_tac "\{I. \k \ (Suc n). I = J k} = (i\\<^bsub>R,(Suc n)\<^esub> J)") (* apply (simp del:ideal_n_prodSn)*) apply (frule_tac n = "Suc n" and J = J in coprime_n_idealsTr4) apply (simp del:ideal_n_prodSn) apply (subst coprime_int_prod) apply (rule n_prod_ideal) apply (rule allI, simp, simp, assumption) apply simp apply (cut_tac n = "Suc n" in Nsetn_sub_mem1) apply simp apply (thin_tac "(\k\Suc n. ideal R (J k)) \ (\i\Suc n. \j\Suc n. i \ j \ coprime_ideals R (J i) (J j)) \ \{I. \k\Suc n. I = J k} = i\\<^bsub>R,Suc n\<^esub> J", thin_tac "(\k\Suc (Suc n). ideal R (J k)) \ (\i\Suc (Suc n). \j\Suc (Suc n). i \ j \ coprime_ideals R (J i) (J j))") apply (rule equalityI, rule subsetI, simp) apply (rule conjI, rule allI, rule impI, erule exE, erule conjE, simp, frule_tac a = xa in forall_spec, frule_tac x = k and y = "Suc n" and z = "Suc (Suc n)" in le_less_trans, simp, frule_tac x = k and y = "Suc (Suc n)" in less_imp_le, blast) apply simp apply (frule_tac a = "J (Suc (Suc n))" in forall_spec, cut_tac n = "Suc (Suc n)" in le_refl, blast, simp) apply (rule subsetI, simp, rule allI, rule impI) apply (erule exE, erule conjE) apply (erule conjE, case_tac "k = Suc (Suc n)", simp) apply (frule_tac m = k and n = "Suc (Suc n)" in noteq_le_less, assumption, thin_tac "k \ Suc (Suc n)") apply (frule_tac x = k and n = "Suc n" in Suc_less_le) apply (frule_tac a = xa in forall_spec, blast, thin_tac "\xa. (\k\Suc n. xa = J k) \ x \ xa", simp) done lemma (in Ring) coprime_prod_int2:"\ \k \ (Suc n). ideal R (J k); \i \ (Suc n). \j \ (Suc n). (i \j \ coprime_ideals R (J i) (J j))\ \ (\ {I. \k \ (Suc n). I = (J k)} = ideal_n_prod R (Suc n) J)" apply (simp add:coprime_prod_int2Tr) done lemma (in Ring) coprime_2_n:"\ideal R A; ideal R B\ \ (qring R A) \\<^sub>r (qring R B) = r\\<^bsub>{j. j \ (Suc 0)}\<^esub> (prodB1 (qring R A) (qring R B))" apply (simp add:Prod2Rg_def Nset_1) done text\In this and following lemmata, ideals A and B are of type \('a, 'more) RingType_scheme\. Don't try \(r\\<^sub>(Nset n) B) \\<^sub>r B (Suc n)\\ lemma (in Ring) A_to_prodag2_hom:"\ideal R A; ideal R B; S 0 = pj R A; S (Suc 0) = pj R B\ \ A_to_prodag R {j. j \ (Suc 0)} S (prodB1 (qring R A) (qring R B)) \ rHom R (qring R A \\<^sub>r qring R B)" apply (subst coprime_2_n [of "A" "B"], assumption+) apply (rule A_to_prodag_rHom, rule Ring_axioms) apply (rule ballI) apply (case_tac "k = 0") apply (simp add:prodB1_def) apply (simp add:qring_ring) apply (simp) apply (frule_tac n = k in Suc_leI[of 0], thin_tac "0 < k") apply (frule_tac m = k and n = "Suc 0" in le_antisym, assumption) apply (simp, simp add:prodB1_def, simp add:qring_ring) apply (rule ballI) apply (simp add:Nset_1) apply (erule disjE) apply (simp add:prodB1_def, rule pj_Hom, rule Ring_axioms, assumption) apply (simp, simp add:prodB1_def) apply (rule pj_Hom, rule Ring_axioms, assumption+) done lemma (in Ring) A2coprime_rsurjecTr:"\ideal R A; ideal R B; S 0 = pj R A; S (Suc 0) = pj R B\ \ (carrier (qring R A \\<^sub>r qring R B)) = carr_prodag {j. j \ (Suc 0)} (prodB1 (qring R A) (qring R B))" apply (simp add:Prod2Rg_def prodrg_def Nset_1) done lemma (in Ring) A2coprime_rsurjec:"\ideal R A; ideal R B; S 0 = pj R A; S (Suc 0) = pj R B; coprime_ideals R A B\ \ surjec\<^bsub>R,((qring R A) \\<^sub>r (qring R B))\<^esub> (A_to_prodag R {j. j\(Suc 0)} S (prodB1 (qring R A) (qring R B)))" apply (frule A_to_prodag2_hom [of "A" "B" "S"], assumption+) apply (simp add:surjec_def) apply (rule conjI, simp add:rHom_def) apply (rule surj_to_test[of "A_to_prodag R {j. j \ (Suc 0)} S (prodB1 (qring R A) (qring R B))" "carrier R" "carrier (qring R A \\<^sub>r qring R B)"]) apply (simp add:rHom_def aHom_def) apply (rule ballI) apply (frule rHom_func[of "A_to_prodag R {j. j \ (Suc 0)} S (prodB1 (R /\<^sub>r A) (R /\<^sub>r B))" R], thin_tac "A_to_prodag R {j. j \ (Suc 0)} S (prodB1 (R /\<^sub>r A) (R /\<^sub>r B)) \ rHom R (R /\<^sub>r A \\<^sub>r R /\<^sub>r B)") apply (simp add:A2coprime_rsurjecTr[of A B S]) apply (simp add:Nset_1) apply (frule_tac X = "b 0" and Y = "b (Suc 0)" in coprime_surjTr[of A B], assumption+) apply (simp add:carr_prodag_def prodB1_def, simp add:carr_prodag_def prodB1_def) apply (erule bexE) apply (frule_tac x = r in funcset_mem[of "A_to_prodag R {0, Suc 0} S (prodB1 (R /\<^sub>r A) (R /\<^sub>r B))" "carrier R" "carr_prodag {0, Suc 0} (prodB1 (R /\<^sub>r A) (R /\<^sub>r B))"], assumption+) apply (cut_tac X = "A_to_prodag R {0, Suc 0} S (prodB1 (R /\<^sub>r A) (R /\<^sub>r B)) r" and Y = b in carr_prodag_mem_eq[of "{0, Suc 0}" "prodB1 (R /\<^sub>r A) (R /\<^sub>r B)"]) apply (rule ballI) apply (simp, erule disjE) apply (simp add:prodB1_def, fold prodB1_def, simp add:qring_ring Ring.ring_is_ag) apply (simp add:prodB1_def, fold prodB1_def, simp add:qring_ring Ring.ring_is_ag) apply assumption+ apply (rule ballI, simp, erule disjE, simp) apply (subst A_to_prodag_def, simp) apply (subst A_to_prodag_def, simp) apply blast done lemma (in Ring) prod2_n_Tr1:"\\k \ (Suc 0). ideal R (J k); \k \ (Suc 0). B k = qring R (J k); \k \ (Suc 0). S k = pj R (J k) \ \ A_to_prodag R {j. j \ (Suc 0)} S (prodB1 (qring R (J 0)) (qring R (J (Suc 0)))) = A_to_prodag R {j. j \ (Suc 0)} S B" apply (subgoal_tac "\k \ (Suc 0). (prodB1 (qring R (J 0)) (qring R (J (Suc 0)))) k = B k") apply (simp add:A_to_prodag_def) apply (rule allI, rule impI) apply (case_tac "k = 0", simp add:Nset_1_1) apply (simp add:prodB1_def) apply (simp add:Nset_1_1) apply (simp add:prodB1_def) done lemma (in aGroup) restrict_prod_Suc:"\\k \ (Suc (Suc n)). ideal R (J k); \k \ (Suc (Suc n)). B k = R /\<^sub>r J k; \k \ (Suc (Suc n)). S k = pj R (J k); f \ carrier (r\\<^bsub>{j. j \ (Suc (Suc n))}\<^esub> B)\ \ restrict f {j. j \ (Suc n)} \ carrier (r\\<^bsub>{j. j \ (Suc n)}\<^esub> B)" apply (cut_tac Nsetn_sub_mem1[of "Suc n"]) apply (simp add:prodrg_def) apply (simp add:carr_prodag_def, (erule conjE)+) apply (simp add:Un_carrier_def) apply (rule Pi_I) apply simp apply (frule_tac x = x in funcset_mem[of f "{j. j \ (Suc (Suc n))}" "\{X. \i \ (Suc (Suc n)). X = carrier (R /\<^sub>r J i)}"], simp) apply simp apply (erule exE, erule conjE, erule exE, erule conjE, simp) apply (rotate_tac -5) apply (frule_tac a = x in forall_spec) apply simp apply blast done lemma (in Ring) Chinese_remTr2:"(\k \ (Suc n). ideal R (J k)) \ (\k\(Suc n). B k = qring R (J k)) \ (\k\(Suc n). S k = pj R (J k)) \ (\i\(Suc n). \j\ (Suc n). (i \j \ coprime_ideals R (J i) (J j))) \ surjec\<^bsub>R,(r\\<^bsub>{j. j\ (Suc n)}\<^esub> B)\<^esub> (A_to_prodag R {j. j\(Suc n)} S B)" apply (cut_tac Ring) apply (induct_tac n) (* case n = 0, i.e. two coprime_ideals *) (** use coprime_surjTr **) apply (rule impI) apply (erule conjE)+ apply (frule A_to_prodag_rHom [of R "{j. j \ Suc 0}" "B" "S"]) apply (rule ballI, simp add:Ring.qring_ring) apply (rule ballI, simp add:pj_Hom) apply (frule rHom_func[of "A_to_prodag R {j. j \ (Suc 0)} S B" R "r\\<^bsub>{j. j \ (Suc 0)}\<^esub> B"]) apply (simp add:surjec_def) apply (rule conjI) apply (simp add:rHom_def) apply (rule surj_to_test, assumption+) apply (rule ballI) apply (simp add:Nset_1) apply (cut_tac coprime_elems [of "J 0" "J (Suc 0)"]) apply (rename_tac f) apply (erule bexE, erule bexE) apply (simp add:prodrg_def) apply (fold prodrg_def) apply (cut_tac X = "f 0" and Y = "f (Suc 0)" in coprime_surjTr[of "J 0" "J (Suc 0)"], simp+) apply (simp add:carr_prodag_def, simp add:carr_prodag_def) apply (erule bexE, (erule conjE)+) apply (frule_tac x = r in funcset_mem[of "A_to_prodag R {0, Suc 0} S B" "carrier R" "carr_prodag {0, Suc 0} B"], assumption+) apply (cut_tac X = "A_to_prodag R {0, Suc 0} S B r" and Y = f in carr_prodag_mem_eq[of "{0, Suc 0}" B]) apply (rule ballI, simp, erule disjE, simp add:qring_ring Ring.ring_is_ag, simp add:Ring.qring_ring Ring.ring_is_ag) apply assumption+ apply (rule ballI, simp, erule disjE, simp) apply (simp add:A_to_prodag_def, simp add:A_to_prodag_def) apply blast apply simp+ apply (rule impI, (erule conjE)+) (**** n ****) apply (cut_tac n = "Suc n" in Nsetn_sub_mem1) apply (subgoal_tac "\k\{i. i \ Suc (Suc n)}. Ring (B k)") apply (frule_tac I = "{i. i \ Suc (Suc n)}" in A_to_prodag_rHom [of "R" _ "B" "S"]) apply assumption apply (rule ballI) apply (simp add:pj_Hom) apply simp apply (subst surjec_def, rule conjI, simp add:rHom_def) apply (cut_tac n = "Suc n" in coprime_n_idealsTr4[of _ J]) apply blast apply (frule_tac f = "A_to_prodag R {j. j \ (Suc (Suc n))} S B" and A = R in rHom_func) apply (rule_tac f = "A_to_prodag R {j. j \ (Suc (Suc n))} S B" and A = "carrier R" and B = "carrier (r\\<^bsub>{j. j \ (Suc (Suc n))}\<^esub> B)" in surj_to_test, assumption+) apply (rule ballI) apply (cut_tac n = "Suc n" in n_prod_ideal[of _ J]) apply (rule allI, simp) apply (frule_tac A = "i\\<^bsub>R,(Suc n)\<^esub> J" and B = "J (Suc (Suc n))" in coprime_elems, cut_tac n = "Suc (Suc n)" in n_in_Nsetn, blast, assumption) apply (erule bexE, erule bexE) apply (rename_tac n f a b) apply (thin_tac " coprime_ideals R (i\\<^bsub>R,(Suc n)\<^esub> J) (J (Suc (Suc n)))") apply (cut_tac n = "Suc n" and a = a and J = J in ele_n_prod, rule allI, simp, assumption) apply (cut_tac ring_is_ag) apply (frule_tac n = n and f = f in aGroup.restrict_prod_Suc[of R _ R J B S], assumption+) apply (frule_tac S = "r\\<^bsub>{j. j \ (Suc n)}\<^esub> B" and f = "A_to_prodag R {j. j \ (Suc n)} S B" in surjec_surj_to[of R]) apply (frule_tac f = "A_to_prodag R {j. j \ (Suc n)} S B" and A = "carrier R" and B = "carrier (r\\<^bsub>{j. j \ (Suc n)}\<^esub> B)" and b = "restrict f {j. j \ (Suc n)}" in surj_to_el2, assumption) apply (erule bexE, rename_tac n f a b u) apply (cut_tac n = "Suc (Suc n)" in n_in_Nsetn, frule_tac f = f and I = "{j. j \ (Suc (Suc n))}" and A = B and i = "Suc (Suc n)" in prodrg_component, assumption) apply simp apply (frule_tac J = "J (Suc (Suc n))" and X = "f (Suc (Suc n))" in pj_surj_to[of R], simp, assumption) apply (erule bexE, rename_tac n f a b u v) apply (frule_tac a = "Suc (Suc n)" in forall_spec, simp, frule_tac I = "J (Suc (Suc n))" and h = b in Ring.ideal_subset[of R], assumption+, cut_tac I = "i\\<^bsub>R,n\<^esub> J \\<^sub>r\<^bsub>R\<^esub> J (Suc n)" and h = a in Ring.ideal_subset[of R], assumption+) apply (frule_tac x = b and y = u in Ring.ring_tOp_closed[of R], assumption+, frule_tac x = a and y = v in Ring.ring_tOp_closed[of R], assumption+, frule Ring.ring_is_ag[of R], frule_tac x = "b \\<^sub>r\<^bsub>R\<^esub> u" and y = "a \\<^sub>r\<^bsub>R\<^esub> v" in aGroup.ag_pOp_closed[of R], assumption+) apply (frule_tac f = "A_to_prodag R {j. j \ (Suc (Suc n))} S B" and A = "carrier R" and B = "carrier (r\\<^bsub>{j. j \ (Suc (Suc n))}\<^esub> B)" and x = "b \\<^sub>r\<^bsub>R\<^esub> u \\<^bsub>R\<^esub> a \\<^sub>r\<^bsub>R\<^esub> v" in funcset_mem, assumption+) apply (frule_tac f = "A_to_prodag R {j. j \ (Suc (Suc n))} S B (b \\<^sub>r\<^bsub>R\<^esub> u \\<^bsub>R\<^esub> a \\<^sub>r\<^bsub>R\<^esub> v)" and I = "{j. j \ (Suc (Suc n))}" and g = f in carr_prodrg_mem_eq, simp) apply (rule ballI) apply (subst A_to_prodag_def, simp) apply (frule_tac I = "J i" in pj_Hom[of R], simp) apply (simp add: rHom_add) apply (frule_tac a = i in forall_spec, assumption, thin_tac "\k \ (Suc (Suc n)). ideal R (J k)", frule_tac I = "J i" in Ring.qring_ring[of R], assumption, thin_tac "\k \ (Suc (Suc n)). Ring (R /\<^sub>r J k)", frule_tac R = "R /\<^sub>r (J i)" and x = b and y = u and f = "pj R (J i)" in rHom_tOp[of R], assumption+, simp, thin_tac "pj R (J i) (b \\<^sub>r\<^bsub>R\<^esub> u) = pj R (J i) b \\<^sub>r\<^bsub>R /\<^sub>r J i\<^esub> pj R (J i) u", frule_tac R = "R /\<^sub>r (J i)" and x = a and y = v and f = "pj R (J i)" in rHom_tOp[of R], simp add:Ring.qring_ring, assumption+) apply (frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = v in rHom_mem[of _ R], assumption+, frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = u in rHom_mem[of _ R], assumption+, frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = b in rHom_mem[of _ R], assumption+, frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = a in rHom_mem[of _ R], assumption+) apply (frule_tac R = "R /\<^sub>r J i" in Ring.ring_is_ag) apply (case_tac "i \ (Suc n)") apply (frule_tac I1 = "J i" and x1 = a in pj_zero[THEN sym, of R ], assumption+, simp, thin_tac "pj R (J i) (a \\<^sub>r\<^bsub>R\<^esub> v) = \\<^bsub>R /\<^sub>r J i\<^esub> \\<^sub>r\<^bsub>R /\<^sub>r J i\<^esub> pj R (J i) v") apply (simp add:Ring.ring_times_0_x) apply (frule_tac f = "pj R (J i)" and A = R and R = "R /\<^sub>r J i" and x = a and y = b in rHom_add, assumption+, simp, thin_tac "A_to_prodag R {j. j \ Suc (Suc n)} S B (b \\<^sub>r u) \\<^bsub>r\\<^bsub>{i. i \ Suc (Suc n)}\<^esub> B\<^esub> A_to_prodag R {j. j \ Suc (Suc n)} S B (a \\<^sub>r v) \ carrier (r\\<^bsub>{j. j \ Suc (Suc n)}\<^esub> B)") apply (simp add:aGroup.ag_l_zero) apply (rotate_tac -1, frule sym, thin_tac " pj R (J i) 1\<^sub>r\<^bsub>R\<^esub> = pj R (J i) b", simp add:rHom_one) apply (simp add:Ring.ring_l_one) apply (simp add:aGroup.ag_r_zero) apply (frule_tac f = "A_to_prodag R {j. j \ (Suc n)} S B u" and g = "restrict f {j. j \ (Suc n)}" and x = i in eq_fun_eq_val, thin_tac "A_to_prodag R {j. j\(Suc n)} S B u = restrict f {j. j\(Suc n)}") apply (simp add:A_to_prodag_def) apply simp apply (frule_tac y = i and x = "Suc n" in not_le_imp_less, frule_tac m = "Suc n" and n = i in Suc_leI, frule_tac m = i and n = "Suc (Suc n)" in Nat.le_antisym, assumption+, simp) apply (frule_tac I1 = "J (Suc (Suc n))" and x1 = b in pj_zero[THEN sym, of R ], assumption+, simp add:Ring.ring_times_0_x) apply (frule_tac f = "pj R (J (Suc (Suc n)))" and A = R and R = "R /\<^sub>r J (Suc (Suc n))" and x = a and y = b in rHom_add, assumption+, simp) apply (simp add:aGroup.ag_r_zero) apply (rotate_tac -1, frule sym, thin_tac "pj R (J (Suc (Suc n))) 1\<^sub>r\<^bsub>R\<^esub> = pj R (J (Suc (Suc n))) a", simp add:rHom_one, simp add:Ring.ring_l_one) apply (simp add:aGroup.ag_l_zero) apply blast apply (rule ballI, simp add:Ring.qring_ring) done lemma (in Ring) Chinese_remTr3:"\\k \ (Suc n). ideal R (J k); \k \ (Suc n). B k = qring R (J k); \k\ (Suc n). S k = pj R (J k); \i \ (Suc n). \j \ (Suc n). (i \j \ coprime_ideals R (J i) (J j))\ \ surjec\<^bsub>R,(r\\<^bsub>{j. j \ (Suc n)}\<^esub> B)\<^esub> (A_to_prodag R {j. j \ (Suc n)} S B)" apply (simp add:Chinese_remTr2 [of "n" "J" "B" "S"]) done lemma (in Ring) imset:"\\k\ (Suc n). ideal R (J k)\ \ {I. \k\ (Suc n). I = J k} = {J k| k. k \ {j. j \ (Suc n)}}" apply blast done theorem (in Ring) Chinese_remThm:"\(\k \ (Suc n). ideal R (J k)); \k\(Suc n). B k = qring R (J k); \k \ (Suc n). S k = pj R (J k); \i \ (Suc n). \j \ (Suc n). (i \j \ coprime_ideals R (J i) (J j))\ \ bijec\<^bsub>(qring R (\ {J k | k. k\{j. j \ (Suc n)}})),(r\\<^bsub>{j. j \ (Suc n)}\<^esub> B)\<^esub> ((A_to_prodag R {j. j \ (Suc n)} S B)\\<^bsub>R,(prodrg {j. j \ (Suc n)} B)\<^esub>)" apply (frule Chinese_remTr3, assumption+) apply (cut_tac I = "{j. j \ (Suc n)}" and A = B in prodrg_ring) apply (rule ballI, simp add:qring_ring) apply (cut_tac surjec_ind_bijec [of "R" "r\\<^bsub>{j. j \ (Suc n)}\<^esub> B" "A_to_prodag R {j. j \ (Suc n)} S B"]) apply (cut_tac Ring, simp add:Chinese_remTr1 [of "R" "Suc n" "J" "B" "S"]) apply (simp add:imset, rule Ring_axioms, assumption+) apply (rule A_to_prodag_rHom, rule Ring_axioms) apply (rule ballI) apply (simp add:qring_ring) apply (rule ballI, simp, rule pj_Hom, rule Ring_axioms, simp) apply assumption done lemma (in Ring) prod_prime:"\ideal R A; \k\(Suc n). prime_ideal R (P k); \l\(Suc n). \ (A \ P l); \k\ (Suc n). \l\ (Suc n). k = l \ \ (P k) \ (P l)\ \ \i \ (Suc n). (nprod R (ppa R P A i) n \ A \ (\l \ {j. j\(Suc n)} - {i}. nprod R (ppa R P A i) n \ P l) \ (nprod R (ppa R P A i) n \ P i))" apply (rule allI, rule impI) apply (rule conjI) apply (frule_tac i = i in prod_primeTr1[of n P A], assumption+) apply (frule_tac n = n and f = "ppa R P A i" in ideal_nprod_inc[of A]) apply (rule allI, rule impI) apply (rotate_tac -2, frule_tac a = ia in forall_spec, assumption, thin_tac "\l \ n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", simp add:ideal_subset) apply (rotate_tac -1, frule_tac a = n in forall_spec, simp, thin_tac "\l\ n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", (erule conjE)+, blast, assumption) apply (frule_tac i = i in prod_primeTr1[of n P A], assumption+) apply (rule conjI) apply (rule ballI) apply (frule_tac a = l in forall_spec, simp, frule_tac I = "P l" in prime_ideal_ideal) apply (frule_tac n = n and f = "ppa R P A i" and A = "P l" in ideal_nprod_inc) apply (rule allI) apply (simp add:ppa_mem, simp) apply (case_tac "l < i", thin_tac "\l\ (Suc n). \ A \ P l", thin_tac "\k\ (Suc n). \l \ (Suc n). k = l \ \ P k \ P l") apply (erule conjE, frule_tac x = l and y = i and z = "Suc n" in less_le_trans, assumption, frule_tac x = l and n = n in Suc_less_le) apply (rotate_tac 2, frule_tac a = l in forall_spec, assumption, thin_tac "\l\n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", thin_tac "l < Suc n") apply (simp only:skip_im_Tr1_2, blast) apply (frule_tac x = l and y = i in leI, thin_tac "\ l < i", cut_tac x = l and A = "{j. j \ (Suc n)}" and a = i in in_diff1) apply simp apply (erule conjE, frule not_sym, thin_tac "l \ i", frule_tac x = i and y = l in le_imp_less_or_eq, erule disjE, thin_tac "i \ l", frule_tac x = i and n = l in less_le_diff) apply (cut_tac i = i and n = n and x = "l - Suc 0" in skip_im_Tr2_1, simp, assumption+, simp, frule_tac x = l and n = n in le_Suc_diff_le) apply (rotate_tac -7, frule_tac a = "l - Suc 0" in forall_spec, assumption, thin_tac "\l\n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", simp, (erule conjE)+) apply blast apply simp apply assumption apply (frule_tac a = i in forall_spec, assumption, thin_tac "\k\ (Suc n). prime_ideal R (P k)") apply (rule_tac P = "P i" and n = n and f = "ppa R P A i" in prime_nprod_exc, assumption+) apply (rule allI, rule impI) apply (rotate_tac -3, frule_tac a = ia in forall_spec, assumption, thin_tac "\l \ n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", simp add:ideal_subset) apply (rule allI, rule impI) apply ( rotate_tac 4, frule_tac a = l in forall_spec, assumption, thin_tac "\l\ n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", simp) done lemma skip_im1:"\i \ (Suc n); P \ {j. j \ (Suc n)} \ Collect (prime_ideal R)\ \ compose {j. j \ n} P (skip i) ` {j. j \ n} = P ` ({j. j \ (Suc n)} - {i})" apply (cut_tac skip_fun[of i n]) apply (subst setim_cmpfn[of _ _ _ _ "{X. prime_ideal R X}"], assumption+) apply simp apply (simp add:skip_fun_im) done lemma (in Ring) mutch_aux1:"\ideal R A; i \ (Suc n); P \ {j. j \ (Suc n)} \ Collect (prime_ideal R)\ \ compose {j. j \ n} P (skip i) \ {j. j \ n} \ Collect (prime_ideal R)" apply (cut_tac skip_fun[of i n]) apply (simp add:composition[of "skip i" "{j. j \ n}" "{j. j \ (Suc n)}" P "Collect (prime_ideal R)"]) done lemma (in Ring) prime_ideal_cont1Tr:"ideal R A \ \P. ((P \ {j. j \ (n::nat)} \ {X. prime_ideal R X}) \ (A \ \ (P ` {j. j \ n}))) \ (\i\ n. A \ (P i))" apply (induct_tac n) apply (rule allI, rule impI) apply (erule conjE) apply simp (** n **) apply (rule allI, rule impI) apply (erule conjE)+ apply (case_tac "\i \ (Suc n). \j\ (Suc n). (i \ j \ P i \ P j)") apply ((erule exE, erule conjE)+, erule conjE) apply (frule_tac f = P and n = n and X = "{X. prime_ideal R X}" and A = A and i = i and j = j in Un_less_Un, assumption+, simp+) apply (frule mutch_aux1, assumption+) apply (frule_tac a = "compose {j. j \ n} P (skip i)" in forall_spec, simp, erule exE) apply (cut_tac i = i and n = n and x = ia in skip_fun_im1, simp+, erule conjE, simp add:compose_def,blast) (** last_step induction assumption is of no use **) apply (thin_tac "\P. P \ {j. j \ n} \ {X. prime_ideal R X} \ A \ \(P ` {j. j \ n}) \ (\i\n. A \ P i)", rule contrapos_pp, simp+) apply (cut_tac n = n and P = P in prod_prime [of A], assumption) apply (rule allI, rule impI, frule_tac f = P and A = "{j. j \ (Suc n)}" and B = "{X. prime_ideal R X}" and x = k in funcset_mem, simp, simp, assumption+) apply (frule_tac n = "Suc n" and f = "\i\{j. j \ (Suc n)}. (nprod R (ppa R P A i) n)" in nsum_ideal_inc[of A], rule allI, rule impI, simp) apply (subgoal_tac "(nsum R (\i\{j. j \ (Suc n)}. nprod R (ppa R P A i) n) (Suc n)) \ (\x\{j. j \ (Suc n)}. P x)") apply blast apply (simp del:nsum_suc) apply (rule allI, rule impI) apply (rename_tac n P l) apply (frule_tac f = P and A = "{j. j \ (Suc n)}" and B = "{X. prime_ideal R X}" and x = l in funcset_mem, simp, simp del:nsum_suc, frule_tac I = "P l" in prime_ideal_ideal) apply (rule_tac A = "P l" and n = "Suc n" and f = "\i\{j. j \ (Suc n)}. (nprod R (ppa R P A i) n)" in nsum_ideal_exc, simp+, rule allI, simp add:ideal_subset) apply (rule contrapos_pp, simp+) apply (rotate_tac -1, frule_tac a = l in forall_spec, simp, thin_tac "\j\Suc n. (\la\{i. i \ Suc n} - {j}. e\\<^bsub>R,n\<^esub> ppa R P A la \ P l) \ e\\<^bsub>R,n\<^esub> ppa R P A j \ P l", thin_tac "\i\Suc n. \j\Suc n. i = j \ \ P i \ P j", thin_tac "\i\Suc n. \ A \ P i") apply (erule disjE, erule bexE) apply (frule_tac a = la in forall_spec, simp, thin_tac "\i\Suc n. e\\<^bsub>R,n\<^esub> ppa R P A i \ A \ (\l\{j. j \ Suc n} - {i}. e\\<^bsub>R,n\<^esub> ppa R P A i \ P l) \ e\\<^bsub>R,n\<^esub> ppa R P A i \ P i", (erule conjE)+) apply blast apply blast done lemma (in Ring) prime_ideal_cont1:"\ideal R A; \i \ (n::nat). prime_ideal R (P i); A \ \ {X. (\i \ n. X = (P i))} \ \ \i\ n. A\(P i)" apply (frule prime_ideal_cont1Tr[of A n]) apply (frule_tac a = P in forall_spec, thin_tac "\P. P \ {j. j \ n} \ {X. prime_ideal R X} \ A \ \(P ` {j. j \ n}) \ (\i\n. A \ P i)") apply (rule conjI, simp, rule subsetI, simp, frule_tac c = x in subsetD[of A "\{X. \i\n. X = P i}"], assumption+, simp, blast) apply assumption done lemma (in Ring) prod_n_ideal_contTr0:"(\l\ n. ideal R (J l)) \ i\\<^bsub>R,n\<^esub> J \ \{X. (\k\n. X = (J k))}" apply (induct_tac n) apply simp apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (cut_tac n = n in n_prod_ideal[of _ J], simp) apply (cut_tac I = "i\\<^bsub>R,n\<^esub> J" and J = "J (Suc n)" in ideal_prod_sub_Int) apply assumption apply simp apply (frule_tac A = "i\\<^bsub>R,n\<^esub> J \\<^sub>r\<^bsub>R\<^esub> J (Suc n)" and B = "i\\<^bsub>R,n\<^esub> J \ J (Suc n)" and C = "\{X. \k\ n. X = J k} \ J (Suc n)" in subset_trans) apply (rule_tac A = "i\\<^bsub>R,n\<^esub> J" and B = "\{X. \k\n. X = J k}" and C = "J (Suc n)" in inter_mono, assumption) apply (rule_tac A = "i\\<^bsub>R,n\<^esub> J \\<^sub>r J (Suc n)" and B = "\{X. \k\ n. X = J k} \ J (Suc n)" and C = "\{X. \k\ (Suc n). X = J k}" in subset_trans, assumption) apply (rule subsetI) apply simp apply (rule allI, rule impI) apply (erule exE, (erule conjE)+) apply (case_tac "k = Suc n", simp) apply (frule_tac m = k and n = "Suc n" in noteq_le_less, assumption) apply (thin_tac " k \ Suc n") apply (frule_tac x = k and n = "Suc n" in less_le_diff, thin_tac "k < Suc n", simp, thin_tac "\l\Suc n. ideal R (J l)") apply (frule_tac a = xa in forall_spec, blast, thin_tac "\xa. (\k\n. xa = J k) \ x \ xa", simp) done lemma (in Ring) prod_n_ideal_contTr:"\\l\ n. ideal R (J l)\ \ i\\<^bsub>R,n\<^esub> J \ \{X. (\k \ n. X = (J k))}" apply (simp add:prod_n_ideal_contTr0) done lemma (in Ring) prod_n_ideal_cont2:"\\l\ (n::nat). ideal R (J l); prime_ideal R P; \{X. (\k\ n. X = (J k))} \ P\ \ \l\ n. (J l) \ P" apply (frule prod_n_ideal_contTr[of n J]) apply (frule_tac A = "i\\<^bsub>R,n\<^esub> J" and B = "\{X. \k\ n. X = J k}" and C = P in subset_trans, assumption+) apply (thin_tac "\{X. \k\ n. X = J k} \ P", thin_tac "i\\<^bsub>R,n\<^esub> J \ \{X. \k\ n. X = J k}") apply (simp add:ideal_n_prod_prime) done lemma (in Ring) prod_n_ideal_cont3:"\\l\ (n::nat). ideal R (J l); prime_ideal R P; \{X. (\k\ n. X = (J k))} = P\ \ \l\ n. (J l) = P" apply (frule prod_n_ideal_cont2[of n J P], assumption+) apply simp apply (erule exE) apply (subgoal_tac "J l = P") apply blast apply (rule equalityI, simp) apply (rule subsetI) apply (rotate_tac -4, frule sym, thin_tac "\{X. \k\ n. X = J k} = P") apply simp apply blast done definition ideal_quotient :: "[_ , 'a set, 'a set] \ 'a set" where "ideal_quotient R A B = {x| x. x \ carrier R \ (\b\B. x \\<^sub>r\<^bsub>R\<^esub> b \ A)}" abbreviation IDEALQT ("(3_/ \\<^sub>_/ _)" [82,82,83]82) where "A \\<^sub>R B == ideal_quotient R A B" lemma (in Ring) ideal_quotient_is_ideal: "\ideal R A; ideal R B\ \ ideal R (ideal_quotient R A B)" apply (rule ideal_condition) apply (rule subsetI) apply (simp add:ideal_quotient_def CollectI) apply (simp add:ideal_quotient_def) apply (cut_tac ring_zero) apply (subgoal_tac "\b\B. \ \\<^sub>r b \ A") apply blast apply (rule ballI) apply (frule_tac h = b in ideal_subset[of B], assumption) apply (frule_tac x = b in ring_times_0_x ) apply (simp add:ideal_zero) apply (rule ballI)+ apply (simp add:ideal_quotient_def, (erule conjE)+, rule conjI) apply (rule ideal_pOp_closed) apply (simp add:whole_ideal, assumption+) apply (cut_tac ring_is_ag) apply (simp add:aGroup.ag_mOp_closed) apply (rule ballI) apply (subst ring_distrib2) apply (simp add:ideal_subset, assumption) apply (cut_tac ring_is_ag, simp add: aGroup.ag_mOp_closed) apply (frule_tac a1 = y and b1 = b in ring_inv1_1 [THEN sym]) apply (simp add:ideal_subset, simp) apply (rule ideal_pOp_closed, assumption+, simp) apply (rule ideal_inv1_closed, assumption+, simp) apply (rule ballI)+ apply (simp add:ideal_quotient_def) apply (rule conjI) apply (erule conjE) apply (simp add:ring_tOp_closed) apply (erule conjE) apply (rule ballI) apply (subst ring_tOp_assoc, assumption+, simp add:ideal_subset) apply (simp add:ideal_ring_multiple [of "A"]) done section \Addition of finite elements of a ring and \ideal_multiplication\\ text\We consider sum in an abelian group\ lemma (in aGroup) nsum_mem1Tr:" A +> J \ (\j \ n. f j \ J) \ nsum A f n \ J" apply (induct_tac n) apply (rule impI) apply simp apply (rule impI) apply simp apply (rule asubg_pOp_closed, assumption+) apply simp done lemma (in aGroup) fSum_mem:"\\j \ nset (Suc n) m. f j \ carrier A; n < m\ \ fSum A f (Suc n) m \ carrier A" apply (simp add:fSum_def) apply (rule nsum_mem) apply (rule allI, simp add:cmp_def slide_def) apply (rule impI) apply (frule_tac x = "Suc (n + j)" in bspec) apply (simp add:nset_def, arith) done lemma (in aGroup) nsum_mem1:"\A +> J; \j \ n. f j \ J\ \ nsum A f n \ J" apply (simp add:nsum_mem1Tr) done lemma (in aGroup) nsum_eq_i:"\\j\n. f j \ carrier A; \j\n. g j \ carrier A; i \ n; \l \ i. f l = g l\ \ nsum A f i = nsum A g i" apply (rule nsum_eq) apply (rule allI, rule impI, simp)+ done lemma (in aGroup) nsum_cmp_eq:"\f \ {j. j\(n::nat)} \ carrier A; h1 \ {j. j \ n} \ {j. j \ n}; h2 \ {j. j \ n} \ {j. j \ n}; i \ n\ \ nsum A (cmp f (cmp h2 h1)) i = nsum A (cmp (cmp f h2) h1) i" apply (rule nsum_eq_i [of n "cmp f (cmp h2 h1)" "cmp (cmp f h2) h1" i]) apply (rule allI, rule impI, simp add:cmp_def) apply ((rule funcset_mem, assumption)+, simp) apply (rule allI, rule impI, simp add:cmp_def, (rule funcset_mem, assumption)+, simp+) apply (rule allI, rule impI, simp add:cmp_def) done lemma (in aGroup) nsum_cmp_eq_transpos:"\ \j\(Suc n). f j \ carrier A; i \ n;i \ n \ \ nsum A (cmp f (cmp (transpos i n) (cmp (transpos n (Suc n)) (transpos i n)))) (Suc n) = nsum A (cmp f (transpos i (Suc n))) (Suc n)" apply (rule nsum_eq [of "Suc n" "cmp f (cmp (transpos i n) (cmp (transpos n (Suc n)) (transpos i n)))" "cmp f (transpos i (Suc n))"]) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac i = i and n = "Suc n" and j = n and l = j in transpos_mem, simp+) apply (cut_tac i = n and n = "Suc n" and j = "Suc n" and l = "transpos i n j" in transpos_mem, simp+) apply (cut_tac i = i and n = "Suc n" and j = n and l = "transpos n (Suc n) (transpos i n j)" in transpos_mem, simp+) apply (rule allI, rule impI, simp add:cmp_def) apply (cut_tac i = i and n = "Suc n" and j = "Suc n" and l = j in transpos_mem, simp+) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (thin_tac "\j\Suc n. f j \ carrier A", rule eq_elems_eq_val[of _ _ f]) apply (simp add:transpos_def) done lemma transpos_Tr_n1:"Suc (Suc 0) \ n \ transpos (n - Suc 0) n n = n - Suc 0" apply (simp add:transpos_def) done lemma transpos_Tr_n2:"Suc (Suc 0) \ n \ transpos (n - (Suc 0)) n (n - (Suc 0)) = n" apply (simp add:transpos_def) done lemma (in aGroup) additionTr0:"\0 < n; \j \ n. f j \ carrier A\ \ nsum A (cmp f (transpos (n - 1) n)) n = nsum A f n" apply (case_tac "n \ 1") apply simp apply (frule Suc_leI [of "0" "n"]) apply (frule le_antisym [of "n" "Suc 0"], assumption+, simp) apply (simp add:cmp_def) apply (subst transpos_ij_1[of 0 "Suc 0"], simp+) apply (subst transpos_ij_2[of 0 "Suc 0"], simp+) apply (rule ag_pOp_commute, simp+) apply (frule not_le_imp_less[of n "Suc 0"]) apply (frule_tac Suc_leI [of "Suc 0" "n"], thin_tac "\ n \ Suc 0") apply (cut_tac nsum_suc[of A f "n - Suc 0"], simp) apply (cut_tac nsum_suc[of A "cmp f (transpos (n - Suc 0) n)" "n - Suc 0"], simp, thin_tac "\\<^sub>e A f n = \\<^sub>e A f (n - Suc 0) \ f n", thin_tac "\\<^sub>e A (cmp f (transpos (n - Suc 0) n)) n = \\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc 0) \ (cmp f (transpos (n - Suc 0) n) n)") apply (case_tac "n = Suc (Suc 0)", simp) apply (cut_tac transpos_id_1[of "Suc 0" "Suc (Suc 0)" "Suc (Suc 0)" 0], cut_tac transpos_ij_1[of "Suc 0" "Suc (Suc 0)" "Suc (Suc 0)"], cut_tac transpos_ij_2[of "Suc 0" "Suc (Suc 0)" "Suc (Suc 0)"], simp add:cmp_def, thin_tac "n = Suc (Suc 0)", thin_tac "transpos (Suc 0) (Suc (Suc 0)) 0 = 0", thin_tac "transpos (Suc 0) (Suc (Suc 0)) (Suc 0) = Suc (Suc 0)", thin_tac "transpos (Suc 0) (Suc (Suc 0)) (Suc (Suc 0)) = Suc 0") apply (subst ag_pOp_assoc, simp+) apply (subst ag_pOp_commute[of "f (Suc (Suc 0))" "f (Suc 0)"], simp+) apply (subst ag_pOp_assoc[THEN sym], simp+) apply (frule not_sym) apply (frule noteq_le_less[of "Suc (Suc 0)" n], assumption, thin_tac "Suc (Suc 0) \ n") apply (cut_tac nsum_suc[of A f "n - Suc 0 - Suc 0"]) apply (cut_tac Suc_pred[of "n - Suc 0"], simp del:nsum_suc) apply (cut_tac nsum_suc[of A "cmp f (transpos (n - Suc 0) n)" "n - Suc (Suc 0)"], simp del:nsum_suc, thin_tac "\\<^sub>e A f (n - Suc 0) = \\<^sub>e A f (n - Suc (Suc 0)) \ f (n - Suc 0)", thin_tac "Suc (n - Suc (Suc 0)) = n - Suc 0", thin_tac "\\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc 0) = \\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc (Suc 0)) \ (cmp f (transpos (n - Suc 0) n)) (n - Suc 0)") apply (cut_tac nsum_eq_i[of n "cmp f (transpos (n - Suc 0) n)" f "n - Suc (Suc 0)"], simp, thin_tac "\\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc (Suc 0)) = \\<^sub>e A f (n - Suc (Suc 0))") apply (simp add:cmp_def) apply (cut_tac transpos_ij_1[of "n - Suc 0" n n], simp) apply (cut_tac transpos_ij_2[of "n - Suc 0" n n], simp) apply (subst ag_pOp_assoc, rule nsum_mem, rule allI, rule impI) apply (frule_tac x = j and y = "n - Suc (Suc 0)" and z = n in le_less_trans, simp, frule_tac x = j and y = n in less_imp_le) apply simp+ apply (subst ag_pOp_commute[of "f n"], simp, simp) apply (subst ag_pOp_assoc[THEN sym], rule nsum_mem, rule allI, rule impI, frule_tac x = j and y = "n - Suc (Suc 0)" and z = n in le_less_trans, simp, frule_tac x = j and y = n in less_imp_le) apply simp+ apply (rule allI, rule impI, simp add:cmp_def) apply (cut_tac i = "n - Suc 0" and n = n and j = n and l = j in transpos_mem, simp+) apply (rule allI, rule impI) apply (simp add:cmp_def) apply (cut_tac i = "n - Suc 0" and n = n and j = n and x = l in transpos_id, simp+) apply (cut_tac x = l and y = "n - Suc (Suc 0)" and z = n in le_less_trans, assumption) apply simp apply arith apply simp apply arith done lemma (in aGroup) additionTr1:"\ \f. \h. f \ {j. j\(Suc n)} \ carrier A \ h \ {j. j\(Suc n)} \ {j. j\(Suc n)} \ inj_on h {j. j\(Suc n)} \ nsum A (cmp f h) (Suc n) = nsum A f (Suc n); f \ {j. j\(Suc (Suc n))} \ carrier A; h \ {j. j\(Suc (Suc n))} \ {j. j\(Suc (Suc n))}; inj_on h {j. j\(Suc (Suc n))}; h (Suc (Suc n)) = Suc (Suc n)\ \ nsum A (cmp f h) (Suc (Suc n)) = nsum A f (Suc (Suc n))" apply (subgoal_tac "f \ {j. j\(Suc n)} \ carrier A") apply (subgoal_tac "h \ {j. j\(Suc n)} \ {j. j\(Suc n)}") apply (subgoal_tac "inj_on h {j. j\(Suc n)}") apply (subgoal_tac "nsum A (cmp f h) (Suc n) = nsum A f (Suc n)") apply (thin_tac "\f. \h. f \ {j. j\(Suc n)} \ carrier A \ h \ {j. j\(Suc n)} \ {j. j\(Suc n)} \ inj_on h {j. j\(Suc n)} \ nsum A (cmp f h) (Suc n) = nsum A f (Suc n)") prefer 2 apply simp apply simp apply (thin_tac "nsum A (cmp f h) n \ (cmp f h (Suc n)) = nsum A f n \ (f (Suc n))") apply (simp add:cmp_def) apply (thin_tac "\f h. (f \ {j. j \ Suc n} \ carrier A) \ (h \ {j. j\Suc n} \ {j. j\Suc n}) \ (inj_on h {j. j\Suc n}) \ \\<^sub>e A (cmp f h) (Suc n) = \\<^sub>e A f (Suc n)") apply (frule Nset_injTr0 [of "h" "Suc n"], assumption+, simp) apply (frule Nset_injTr0 [of "h" "Suc n"], assumption+, simp) apply (simp add:Pi_def) done lemma (in aGroup) additionTr1_1:"\\f. \h. f \ {j. j\Suc n} \ carrier A \ h \ {j. j\Suc n} \ {j. j\Suc n} \ inj_on h {j. j\Suc n} \ nsum A (cmp f h) (Suc n) = nsum A f (Suc n); f \ {j. j\Suc (Suc n)} \ carrier A; i \ n\ \ nsum A (cmp f (transpos i (Suc n))) (Suc (Suc n)) = nsum A f (Suc (Suc n))" apply (rule additionTr1 [of "n" "f" "transpos i (Suc n)"], assumption+) apply (rule transpos_hom [of "i" "Suc (Suc n)" "Suc n"]) apply simp+ apply (rule transpos_inj [of "i" "Suc (Suc n)" "Suc n"]) apply simp+ apply (subst transpos_id[of i "Suc (Suc n)" "Suc n" "Suc (Suc n)"]) apply simp+ done lemma (in aGroup) additionTr1_2:"\\f. \h. f \ {j. j\Suc n} \ carrier A \ h \ {j. j\Suc n} \ {j. j\Suc n} \ inj_on h {j. j\Suc n} \ nsum A (cmp f h) (Suc n) = nsum A f (Suc n); f \ {j. j\ Suc (Suc n)} \ carrier A; i \ (Suc n)\ \ nsum A (cmp f (transpos i (Suc (Suc n)))) (Suc (Suc n)) = nsum A f (Suc (Suc n))" apply (case_tac "i = Suc n") apply (simp del:nsum_suc) apply (cut_tac additionTr0 [of "Suc (Suc n)" "f"], simp, simp, rule allI, rule impI, rule funcset_mem[of f "{j. j \ Suc (Suc n)}" "carrier A"], (simp del:nsum_suc)+) apply (subst nsum_cmp_eq_transpos [THEN sym, of "Suc n" f i], rule allI, rule impI, rule funcset_mem[of f "{j. j \ Suc (Suc n)}" "carrier A"], assumption+, simp, assumption+) apply (subst nsum_cmp_eq [of "f" "Suc (Suc n)" "cmp (transpos (Suc n) (Suc(Suc n))) (transpos i (Suc n))" "transpos i (Suc n)" "Suc (Suc n)"], assumption+, rule Pi_I, simp add:cmp_def, rule transpos_mem, (simp del:nsum_suc)+, rule transpos_mem, (simp del:nsum_suc)+, rule Pi_I, simp, rule transpos_mem, (simp del:nsum_suc)+) apply (subst nsum_cmp_eq [of "cmp f (transpos i (Suc n))" "Suc (Suc n)" "(transpos i (Suc n))" "transpos (Suc n) (Suc (Suc n))" "Suc (Suc n)"], rule Pi_I, simp add:cmp_def, rule funcset_mem[of f "{j. j \ Suc (Suc n)}" "carrier A"], assumption, simp, rule transpos_mem, (simp del:nsum_suc)+, (rule Pi_I, simp, rule transpos_mem, (simp del:nsum_suc)+)+) apply (subst additionTr1_1 [of "n" "cmp (cmp f (transpos i (Suc n))) (transpos (Suc n) (Suc (Suc n)))" "i"], assumption+, rule cmp_fun [of _ "{j. j \ (Suc (Suc n))}" "{j. j \ (Suc (Suc n))}" _ "carrier A"], rule transpos_hom, simp, simp, simp, rule cmp_fun [of _ "{j. j \ (Suc (Suc n))}" "{j. j \ (Suc (Suc n))}" "f" "carrier A"], rule transpos_hom, simp, simp, assumption+, arith) apply (cut_tac additionTr0 [of "Suc (Suc n)" "cmp f (transpos i (Suc n))"], simp del:nsum_suc, thin_tac "nsum A (cmp (cmp f (transpos i (Suc n))) (transpos (Suc n) (Suc (Suc n)))) (Suc (Suc n)) = nsum A (cmp f (transpos i (Suc n))) (Suc (Suc n))") apply (rule additionTr1_1, assumption+, arith, simp, rule allI, rule impI, simp add:cmp_def, rule funcset_mem[of f "{j. j \ Suc (Suc n)}" "carrier A"], assumption) apply (simp add:transpos_mem) done lemma (in aGroup) additionTr2:" \f. \h. f \ {j. j \ (Suc n)} \ carrier A \ h \ {j. j \ (Suc n)} \ {j. j \ (Suc n)} \ inj_on h {j. j \ (Suc n)} \ nsum A (cmp f h) (Suc n) = nsum A f (Suc n)" apply (induct_tac n) apply (rule allI)+ apply (rule impI, (erule conjE)+) apply (simp add:cmp_def) apply (case_tac "h 0 = 0") apply (simp add:Nset_1) apply (simp add:Nset_1 ag_pOp_commute) (************* n *****************) apply (rule allI)+ apply (rule impI, (erule conjE)+) apply (case_tac "h (Suc (Suc n)) = Suc (Suc n)") apply (rule additionTr1, assumption+) apply (frule_tac f = h and n = "Suc (Suc n)" in inj_surj, assumption+) apply (frule sym, thin_tac "h ` {i. i \ Suc (Suc n)} = {i. i \ Suc (Suc n)}") apply (cut_tac n = "Suc (Suc n)" in n_in_Nsetn) apply (frule_tac a = "Suc (Suc n)" and A = "{i. i \ Suc (Suc n)}" and B = "h ` {i. i \ Suc (Suc n)}" in eq_set_inc, assumption+) apply (thin_tac "{i. i \ Suc (Suc n)} = h ` {i. i \ Suc (Suc n)}") apply (simp del:nsum_suc add:image_def) apply (erule exE, erule conjE) apply (frule sym, thin_tac "Suc (Suc n) = h x") apply (frule_tac i = x and n = "Suc (Suc n)" and j = "Suc (Suc n)" in transpos_ij_2, simp del:nsum_suc add:n_in_Nsetn) apply (rule contrapos_pp, (simp del:nsum_suc)+) apply (frule_tac x = "transpos x (Suc (Suc n)) (Suc (Suc n))" and y = x and f = h in eq_elems_eq_val, thin_tac "transpos x (Suc (Suc n)) (Suc (Suc n)) = x", simp del:nsum_suc) apply (frule_tac f = h and A = "{i. i \ Suc (Suc n)}" and x = x and y = "Suc (Suc n)" in inj_onTr2, simp, simp, frule not_sym, simp) apply (cut_tac f1 = "cmp f h" and n1 = n and i1 = x in additionTr1_2[THEN sym], assumption) apply (rule cmp_fun, simp, assumption, arith) apply (simp del:nsum_suc, thin_tac "\\<^sub>e A (cmp f h) (Suc (Suc n)) = \\<^sub>e A (cmp (cmp f h) (transpos x (Suc (Suc n)))) (Suc (Suc n))") apply (frule_tac f = f and n = "Suc n" and A = "carrier A" in func_pre) apply (cut_tac f = "cmp h (transpos x (Suc (Suc n)))" and A = "{j. j \ (Suc ( Suc n))}" and ?A1.0 = "{j. j \ (Suc n)}" in restrict_inj) apply (rule_tac f = "transpos x (Suc (Suc n))" and A = "{j. j \ Suc (Suc n)}" and B = "{j. j \ Suc (Suc n)}" and g = h and C = "{j. j \ Suc (Suc n)}" in cmp_inj, simp, rule transpos_hom, simp, simp, assumption+, rule transpos_inj, simp, simp, assumption+, rule subsetI, simp) apply (subst nsum_cmp_assoc, rule allI, rule impI, simp add:Pi_def, rule transpos_hom, assumption, simp, assumption+) apply (cut_tac f = "cmp f (cmp h (transpos x (Suc (Suc n))))" and n = "Suc n" in nsum_suc[of A ], simp del:nsum_suc, thin_tac "\\<^sub>e A (cmp f (cmp h (transpos x (Suc (Suc n))))) (Suc (Suc n)) = \\<^sub>e A (cmp f (cmp h (transpos x (Suc (Suc n))))) (Suc n) \ cmp f (cmp h (transpos x (Suc (Suc n)))) (Suc (Suc n))") apply (frule_tac x = f in spec, thin_tac "\f h. f \ {j. j \ Suc n} \ carrier A \ h \ {j. j \ Suc n} \ {j. j \ Suc n} \ inj_on h {j. j \ Suc n} \ \\<^sub>e A (cmp f h) (Suc n) = \\<^sub>e A f (Suc n)") apply (frule_tac a = "cmp h (transpos x (Suc (Suc n)))" in forall_spec, thin_tac "\h. f \ {j. j \ Suc n} \ carrier A \ h \ {j. j \ Suc n} \ {j. j \ Suc n} \ inj_on h {j. j \ Suc n} \ \\<^sub>e A (cmp f h) (Suc n) = \\<^sub>e A f (Suc n)") apply simp apply (rule Pi_I) apply (simp add:cmp_def) apply (case_tac "xa = x", simp) apply (cut_tac i = x and n = "Suc (Suc n)" and j = "Suc (Suc n)" in transpos_ij_1, simp, simp, simp, simp, frule_tac x = "Suc (Suc n)" and f = h and A = "{j. j \ Suc (Suc n)}" and B = "{j. j \ Suc (Suc n)}" in funcset_mem, simp, thin_tac "h \ {j. j \ Suc (Suc n)} \ {j. j \ Suc (Suc n)}") apply (cut_tac m = "h (Suc (Suc n))" and n = "Suc (Suc n)" in noteq_le_less, simp, simp, rule_tac x = "h (Suc (Suc n))" and n = "Suc n" in Suc_less_le, assumption) apply (subst transpos_id, simp, simp, simp, simp, frule_tac x = xa and f = h and A = "{j. j \ Suc (Suc n)}" and B = "{j. j \ Suc (Suc n)}" in funcset_mem, simp) apply (frule_tac f = h and A = "{j. j \ Suc (Suc n)}" and x = xa and y = x in injective, simp, simp, assumption) apply (cut_tac m = "h xa" and n = "Suc (Suc n)" in noteq_le_less, simp, simp) apply (rule Suc_less_le, assumption, thin_tac "\h. f \ {j. j \ Suc n} \ carrier A \ h \ {j. j \ Suc n} \ {j. j \ Suc n} \ inj_on h {j. j \ Suc n} \ \\<^sub>e A (cmp f h) (Suc n) = \\<^sub>e A f (Suc n)") apply (simp del:nsum_suc add:cmp_def) apply simp done lemma (in aGroup) addition2:"\f \ {j. j \ (Suc n)} \ carrier A; h \ {j. j \ (Suc n)} \ {j. j \ (Suc n)}; inj_on h {j. j \ (Suc n)}\ \ nsum A (cmp f h) (Suc n) = nsum A f (Suc n)" apply (simp del:nsum_suc add:additionTr2) done lemma (in aGroup) addition21:"\f \ {j. j \ n} \ carrier A; h \ {j. j \ n} \ {j. j \ n}; inj_on h {j. j \ n}\ \ nsum A (cmp f h) n = nsum A f n" apply (case_tac "n = 0") apply (simp add: cmp_def) apply (cut_tac f = f and n = "n - Suc 0" and h = h in addition2) apply simp+ done lemma (in aGroup) addition3:"\\j \ (Suc n). f j \ carrier A; j \ (Suc n); j \ Suc n\ \ nsum A f (Suc n) = nsum A (cmp f (transpos j (Suc n))) (Suc n)" apply (rule addition2 [THEN sym,of "f" "n" "transpos j (Suc n)"]) apply (simp) apply (rule transpos_hom, assumption+, simp, assumption) apply (rule transpos_inj, simp+) done lemma (in aGroup) nsum_splitTr:"(\j \ (Suc (n + m)). f j \ carrier A) \ nsum A f (Suc (n + m)) = nsum A f n \ (nsum A (cmp f (slide (Suc n))) m)" apply (induct_tac m) apply (rule impI) apply (simp add:slide_def cmp_def) apply (rule impI, simp del:nsum_suc) apply (cut_tac n = "Suc (n + na)" in nsum_suc[of A f], simp del:nsum_suc, thin_tac "\\<^sub>e A f (Suc (Suc (n + na))) = \\<^sub>e A f n \ \\<^sub>e A (cmp f (slide (Suc n))) na \ f (Suc (Suc (n + na)))") apply (cut_tac f = "cmp f (slide (Suc n))" and n = na in nsum_suc[of A], simp del:nsum_suc) apply (subst ag_pOp_assoc) apply (rule nsum_mem, rule allI, simp) apply (rule_tac n = na in nsum_mem, thin_tac "\\<^sub>e A (cmp f (slide (Suc n))) (Suc na) = \\<^sub>e A (cmp f (slide (Suc n))) na \ (cmp f (slide (Suc n)) (Suc na))") apply (rule allI, rule impI, simp add:cmp_def slide_def, simp) apply (simp add:cmp_def slide_def) done lemma (in aGroup) nsum_split:"\j \ (Suc (n + m)). f j \ carrier A \ nsum A f (Suc (n + m)) = nsum A f n \ (nsum A (cmp f (slide (Suc n))) m)" by (simp del:nsum_suc add:nsum_splitTr) lemma (in aGroup) nsum_split1:"\\j \ m. f j \ carrier A; n < m\ \ nsum A f m = nsum A f n \ (fSum A f (Suc n) m)" apply (cut_tac nsum_split[of n "m - n - Suc 0" f]) apply simp apply (simp add:fSum_def) apply simp done lemma (in aGroup) nsum_minusTr:" (\j \ n. f j \ carrier A) \ -\<^sub>a (nsum A f n) = nsum A (\x\{j. j \ n}. -\<^sub>a (f x)) n" apply (induct_tac n) apply (rule impI, simp) apply (rule impI) apply (subst nsum_suc, subst nsum_suc) apply (subst ag_p_inv) apply (rule_tac n = n in nsum_mem [of _ f], rule allI, simp, simp) apply (subgoal_tac "\j\n. f j \ carrier A", simp) apply (rule_tac a = "\\<^sub>e A (\u. if u \ n then -\<^sub>a (f u) else undefined) n" and b = "\\<^sub>e A (\x\{j. j \ (Suc n)}. -\<^sub>a (f x)) n" and c = "-\<^sub>a (f (Suc n))" in ag_pOp_add_r, rule_tac n = n in nsum_mem, rule allI, rule impI, simp, rule ag_mOp_closed, simp) apply (rule_tac n = n in nsum_mem, rule allI, rule impI, simp, rule ag_mOp_closed, simp, rule ag_mOp_closed, simp) apply (rule_tac f = "\u. if u \ n then -\<^sub>a (f u) else undefined" and n = n and g = "\x\{j. j \ (Suc n)}. -\<^sub>a (f x)" in nsum_eq, rule allI, rule impI, simp, rule ag_mOp_closed, simp, rule allI, simp, rule impI, rule ag_mOp_closed, simp) apply (rule allI, simp) apply (rule allI, simp) done lemma (in aGroup) nsum_minus:"\j \ n. f j \ carrier A \ -\<^sub>a (nsum A f n) = nsum A (\x\{j. j \ n}. -\<^sub>a (f x)) n" apply (simp add:nsum_minusTr) done lemma (in aGroup) ring_nsum_zeroTr:"(\j \ (n::nat). f j \ carrier A) \ (\j \ n. f j = \) \ nsum A f n = \" apply (induct_tac n) apply (rule impI) apply (erule conjE)+ apply simp apply (rule impI, (erule conjE)+) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (simp add:ag_inc_zero) apply (cut_tac ag_inc_zero, simp add:ag_r_zero) done lemma (in aGroup) ring_nsum_zero:"\j \ (n::nat). f j = \ \ \\<^sub>e A f n = \" apply (cut_tac ring_nsum_zeroTr[of n f]) apply (simp add:ag_inc_zero) done lemma (in aGroup) ag_nsum_1_nonzeroTr: "\f. (\j \ n. f j \ carrier A) \ (l \ n \ (\j \ {j. j \ n} - {l}. f j = \)) \ nsum A f n = f l" apply (induct_tac n) apply simp apply (rule allI, rule impI, (erule conjE)+) apply (case_tac "l = Suc n") apply simp apply (subgoal_tac "{j. j \ Suc n} - {Suc n} = {j. j \ n}", simp, frule ring_nsum_zero, simp) apply (rule ag_l_zero, simp) apply (rule equalityI, rule subsetI, simp, rule subsetI, simp) apply (frule_tac m = l and n = "Suc n" in noteq_le_less, assumption, thin_tac "l \ Suc n", frule_tac x = l and n = n in Suc_less_le) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (thin_tac "\f. (\j\n. f j \ carrier A) \ (\j\{j. j \ n} - {l}. f j = \) \ \\<^sub>e A f n = f l", frule_tac a = l in forall_spec, simp) apply (simp add:ag_r_zero) done lemma (in aGroup) ag_nsum_1_nonzero:"\\j \ n. f j \ carrier A; l \ n; \j\({j. j \ n} - {l}). f j = \ \ \ nsum A f n = f l" apply (simp add:ag_nsum_1_nonzeroTr[of n l]) done definition set_mult :: "[_ , 'a set, 'a set] \ 'a set" where "set_mult R A B = {z. \x\A. \y\B. x \\<^sub>r\<^bsub>R\<^esub> y = z}" definition sum_mult :: "[_ , 'a set, 'a set] \ 'a set" where "sum_mult R A B = {x. \n. \f \ {j. j \ (n::nat)} \ set_mult R A B. nsum R f n = x}" (* zero_fini::"[_ , 'a set, 'a set] \ (nat \ 'a)" "zero_fini R A B i == \ \\<^sub>r \" *) lemma (in Ring) set_mult_sub:"\A \ carrier R; B \ carrier R\ \ set_mult R A B \ carrier R" apply (rule subsetI, simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \\<^sub>r y = x", simp) apply (rule ring_tOp_closed, (simp add:subsetD)+) done lemma (in Ring) set_mult_mono:"\A1 \ carrier R; A2 \ carrier R; A1 \ A2; B \ carrier R\ \ set_mult R A1 B \ set_mult R A2 B" apply (rule subsetI) apply (simp add:set_mult_def, (erule bexE)+) apply (frule_tac c = xa in subsetD[of A1 A2], assumption+) apply blast done lemma (in Ring) sum_mult_Tr1:"\A \ carrier R; B \ carrier R\ \ (\j \ n. f j \ set_mult R A B) \ nsum R f n \ carrier R" apply (cut_tac ring_is_ag) apply (induct_tac n) apply (rule impI, simp) apply (frule set_mult_sub[of A B], assumption, simp add:subsetD) apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (frule set_mult_sub[of A B], assumption) apply (frule_tac a = "Suc n" in forall_spec, simp, frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "carrier R"], assumption) apply (rule aGroup.ag_pOp_closed, assumption+) done lemma (in Ring) sum_mult_mem:"\A \ carrier R; B \ carrier R; \j \ n. f j \ set_mult R A B\ \ nsum R f n \ carrier R" apply (cut_tac ring_is_ag) apply (simp add:sum_mult_Tr1) done lemma (in Ring) sum_mult_mem1:"\A \ carrier R; B \ carrier R; x \ sum_mult R A B\ \ \n. \f\{j. j \ (n::nat)} \ set_mult R A B. nsum R f n = x" by (simp add:sum_mult_def) lemma (in Ring) sum_mult_subR:"\A \ carrier R; B \ carrier R\ \ sum_mult R A B \ carrier R" apply (rule subsetI) apply (frule_tac x = x in sum_mult_mem1[of A B], assumption+) apply (erule exE, erule bexE, frule sym, thin_tac "\\<^sub>e R f n = x", simp) apply (cut_tac ring_is_ag) apply (rule aGroup.nsum_mem[of R], assumption) apply (rule allI, rule impI) apply (frule_tac f = f and A = "{j. j \ n}" and B = "set_mult R A B" and x = j in funcset_mem, simp) apply (frule set_mult_sub[of A B], assumption) apply (simp add:subsetD) done lemma (in Ring) times_mem_sum_mult:"\A \ carrier R; B \ carrier R; a \ A; b \ B \ \ a \\<^sub>r b \ sum_mult R A B" apply (simp add:sum_mult_def) apply (subgoal_tac "(\i\{j. j \ (0::nat)}. a \\<^sub>r b) \ {j. j \ 0} \ set_mult R A B") apply (subgoal_tac "nsum R (\i\{j. j \ (0::nat)}. a \\<^sub>r b) 0 = a \\<^sub>r b") apply blast apply simp apply (rule Pi_I, simp add:set_mult_def, blast) done lemma (in Ring) mem_minus_sum_multTr2:"\A \ carrier R; B \ carrier R; \j \ n. f j \ set_mult R A B; i \ n\ \ f i \ carrier R" apply (frule_tac a = i in forall_spec, simp) apply (frule set_mult_sub[of A B], assumption, simp add:subsetD) done lemma (in aGroup) nsum_jointfun:"\\j \ n. f j \ carrier A; \j \ m. g j \ carrier A\ \ \\<^sub>e A (jointfun n f m g) (Suc (n + m)) = \\<^sub>e A f n \ (\\<^sub>e A g m)" apply (subst nsum_split) apply (rule allI, rule impI) apply (frule_tac f = f and n = n and A = "carrier A" and g = g and m = m and B = "carrier A" in jointfun_mem, assumption+, simp) apply (subgoal_tac "nsum A (jointfun n f m g) n = nsum A f n") apply (subgoal_tac "nsum A (cmp (jointfun n f m g) (slide (Suc n))) m = nsum A g m") apply simp apply (thin_tac "nsum A (jointfun n f m g) n = nsum A f n") apply (rule nsum_eq) apply (rule allI, rule impI, simp add:cmp_def jointfun_def slide_def sliden_def, assumption) apply (rule allI, simp add:cmp_def jointfun_def slide_def sliden_def) apply (rule nsum_eq) apply (rule allI, rule impI, simp add:jointfun_def, assumption) apply (rule allI, rule impI) apply (simp add:jointfun_def) done lemma (in Ring) sum_mult_pOp_closed:"\A \ carrier R; B \ carrier R; a \ sum_mult R A B; b \ sum_mult R A B \ \ a \\<^bsub>R\<^esub> b \ sum_mult R A B" apply (cut_tac ring_is_ag) apply (simp add:sum_mult_def) apply ((erule exE)+, (erule bexE)+) apply (rename_tac n m f g) apply (frule sym, thin_tac "\\<^sub>e R f n = a", frule sym, thin_tac "\\<^sub>e R g m = b", simp) apply (frule set_mult_sub[of A B], assumption) apply (subst aGroup.nsum_jointfun[THEN sym, of R], assumption) apply (rule allI, rule impI, frule_tac f = f and A = "{j. j \ n}" and B = "set_mult R A B" and x = j in funcset_mem, simp, simp add:subsetD) apply (rule allI, rule impI, frule_tac f = g and A = "{j. j \ m}" and B = "set_mult R A B" and x = j in funcset_mem, simp, simp add:subsetD) apply (frule_tac f = f and n = n and A = "set_mult R A B" and g = g and m = m and B = "set_mult R A B" in jointfun_hom, assumption+) apply (simp del:nsum_suc) apply blast done lemma (in Ring) set_mult_mOp_closed:"\A \ carrier R; ideal R B; x \ set_mult R A B\ \ -\<^sub>a x \ set_mult R A B" apply (cut_tac ring_is_ag, simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \\<^sub>r y = x", simp, frule_tac c = xa in subsetD[of A "carrier R"], assumption+, frule ideal_subset1[of B], frule_tac c = y in subsetD[of B "carrier R"], assumption+, simp add:ring_inv1_2, frule_tac I = B and x = y in ideal_inv1_closed, assumption+) apply blast done lemma (in Ring) set_mult_ring_times_closed:"\A \ carrier R; ideal R B; x \ set_mult R A B; r \ carrier R\ \ r \\<^sub>r x \ set_mult R A B" apply (cut_tac ring_is_ag, simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \\<^sub>r y = x", simp, frule_tac c = xa in subsetD[of A "carrier R"], assumption+, frule ideal_subset1[of B], frule_tac c = y in subsetD[of B "carrier R"], assumption, frule_tac x = r and y = "xa \\<^sub>r y" in ring_tOp_commute, simp add:ring_tOp_closed, simp, subst ring_tOp_assoc, assumption+) apply (frule_tac x = y and y = r in ring_tOp_commute, assumption+, simp, frule_tac x = y and r = r in ideal_ring_multiple [of B], assumption+) apply blast done lemma (in Ring) set_mult_sub_sum_mult:"\A \ carrier R; ideal R B\ \ set_mult R A B \ sum_mult R A B" apply (rule subsetI) apply (simp add:sum_mult_def) apply (cut_tac f = "(\i\{j. j \ (0::nat)}. x)" in nsum_0[of R]) apply (cut_tac n_in_Nsetn[of 0], simp del:nsum_0) apply (cut_tac f = "\i\{j. j \ (0::nat)}. x" and B = "%_. set_mult R A B" in Pi_I[of "{j. j \ 0}"], simp) apply (subgoal_tac "\\<^sub>e R (\i\{j. j \ 0}. x) 0 = x") apply blast apply simp done lemma (in Ring) sum_mult_pOp_closedn:"\A \ carrier R; ideal R B\ \ (\j \ n. f j \ set_mult R A B) \ \\<^sub>e R f n \ sum_mult R A B" apply (induct_tac n) apply (rule impI, simp) apply (frule set_mult_sub_sum_mult[of A B], assumption+, simp add:subsetD) apply (rule impI) apply simp apply (frule_tac a = "Suc n" in forall_spec, simp) apply (frule set_mult_sub_sum_mult[of A B], assumption+, frule_tac c= "f (Suc n)" in subsetD[of "set_mult R A B" "sum_mult R A B"], assumption+) apply (rule sum_mult_pOp_closed, assumption, simp add:ideal_subset1, assumption+) done lemma (in Ring) mem_minus_sum_multTr4:"\A \ carrier R; ideal R B\ \ (\j \ n. f j \ set_mult R A B) \ -\<^sub>a (nsum R f n) \ sum_mult R A B" apply (cut_tac ring_is_ag) apply (induct_tac n) apply (rule impI) apply (cut_tac n_in_Nsetn[of 0]) apply (frule_tac x = "f 0" in set_mult_mOp_closed[of A B], assumption+) apply simp apply (frule set_mult_sub_sum_mult[of A B], assumption+, simp add:subsetD) apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (frule sum_mult_subR[of A B], simp add:ideal_subset1) apply (frule_tac n = n and f = f in sum_mult_pOp_closedn[of A B], assumption, cut_tac n = n in Nsetn_sub_mem1, simp) apply (frule_tac c = "\\<^sub>e R f n" in subsetD[of "sum_mult R A B" "carrier R"], assumption+, frule_tac a = "Suc n" in forall_spec, simp, thin_tac "\j\Suc n. f j \ set_mult R A B", frule set_mult_sub[of A B], simp add:ideal_subset1, frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "carrier R"], assumption+ ) apply (frule_tac x = "\\<^sub>e R f n" and y = "f (Suc n)" in aGroup.ag_p_inv[of R], assumption+, simp) apply (rule_tac a = "-\<^sub>a (\\<^sub>e R f n)" and b = "-\<^sub>a (f (Suc n))" in sum_mult_pOp_closed[of A B], assumption+, simp add:ideal_subset1, assumption) apply (frule_tac x = "f (Suc n)" in set_mult_mOp_closed[of A B], assumption+, frule set_mult_sub_sum_mult[of A B], assumption+) apply (simp add:subsetD) done lemma (in Ring) sum_mult_iOp_closed:"\A \ carrier R; ideal R B; x \ sum_mult R A B \ \ -\<^sub>a x \ sum_mult R A B" apply (frule sum_mult_mem1 [of A B x], simp add:ideal_subset1, assumption) apply (erule exE, erule bexE, frule sym, thin_tac "\\<^sub>e R f n = x") apply simp apply (frule_tac n = n and f = f in mem_minus_sum_multTr4[of A B], assumption+) apply (simp add:Pi_def) done lemma (in Ring) sum_mult_ring_multiplicationTr: "\A \ carrier R; ideal R B; r \ carrier R\ \ (\j \ n. f j \ set_mult R A B) \ r \\<^sub>r (nsum R f n) \ sum_mult R A B" apply (cut_tac ring_is_ag) apply (induct_tac n) apply (rule impI, simp) apply (simp add:set_mult_def) apply ((erule bexE)+, frule sym, thin_tac "x \\<^sub>r y = f 0", simp) apply (frule_tac c = x in subsetD[of A "carrier R"], assumption+) apply (frule ideal_subset1[of B], frule_tac c = y in subsetD[of B "carrier R"], assumption, frule_tac x = r and y = "x \\<^sub>r y" in ring_tOp_commute, simp add:ring_tOp_closed, simp, subst ring_tOp_assoc, assumption+) apply (frule_tac x = y and y = r in ring_tOp_commute, assumption+, simp, frule_tac x = y and r = r in ideal_ring_multiple [of B], assumption+) apply (rule times_mem_sum_mult, assumption+) apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (frule_tac f = f and n = n in aGroup.nsum_mem, frule set_mult_sub [of "A" "B"], simp add:ideal_subset1, rule allI, rule impI, cut_tac n = n in Nsetn_sub_mem1, simp add: subsetD, frule_tac a = "Suc n" in forall_spec, simp) apply (frule set_mult_sub[of A B], simp add:ideal_subset1, frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "carrier R"], assumption) apply (simp add: ring_distrib1) apply (rule sum_mult_pOp_closed[of A B], assumption+, simp add:ideal_subset1, assumption) apply (frule_tac x = "f (Suc n)" in set_mult_ring_times_closed [of A B _ r], assumption+, simp, assumption, frule set_mult_sub_sum_mult[of A B], assumption+, simp add:subsetD) done lemma (in Ring) sum_mult_ring_multiplication:"\A \ carrier R; ideal R B; r \ carrier R; a \ sum_mult R A B\ \ r \\<^sub>r a \ sum_mult R A B" apply (cut_tac ring_is_ag) apply (frule sum_mult_mem1[of A B a], simp add:ideal_subset1, assumption) apply (erule exE, erule bexE, frule sym, thin_tac "\\<^sub>e R f n = a", simp) apply (subgoal_tac "\j \ n. f j \ set_mult R A B") apply (simp add:sum_mult_ring_multiplicationTr) apply (simp add:Pi_def) done lemma (in Ring) ideal_sum_mult:"\A \ carrier R; A \ {}; ideal R B\ \ ideal R (sum_mult R A B)" apply (simp add:ideal_def [of _ "sum_mult R A B"]) apply (cut_tac ring_is_ag) apply (rule conjI) apply (rule aGroup.asubg_test, assumption+) apply (rule subsetI) apply (frule_tac x = x in sum_mult_mem1[of A B], simp add:ideal_subset1, assumption, erule exE, erule bexE, frule sym, thin_tac "\\<^sub>e R f n = x", simp) apply (rule_tac f = f and n = n in sum_mult_mem[of A B _ _], assumption+) apply (simp add:ideal_subset1) apply (simp add:Pi_def) apply (frule nonempty_ex[of A], erule exE) apply (frule ideal_zero[of B]) apply (frule_tac a = x and b = \ in times_mem_sum_mult[of A B], simp add:ideal_subset1, assumption+) apply blast apply (rule ballI)+ apply (rule_tac a = a and b = "-\<^sub>a b" in sum_mult_pOp_closed[of A B], assumption, simp add:ideal_subset1, assumption+, rule_tac x = b in sum_mult_iOp_closed[of A B], assumption+) apply (rule ballI)+ apply (rule sum_mult_ring_multiplication, assumption+) done lemma (in Ring) ideal_inc_set_multTr:"\A \ carrier R; ideal R B; ideal R C; set_mult R A B \ C \ \ \f \ {j. j \ (n::nat)} \ set_mult R A B. \\<^sub>e R f n \ C" apply (induct_tac n) apply (simp add:subsetD) apply (rule ballI) apply ( frule_tac f = f and A = "{j. j \ Suc n}" and x = "Suc n" and B = "set_mult R A B"in funcset_mem, simp, frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "C"], assumption+, simp) apply (rule ideal_pOp_closed[of C], assumption+, cut_tac n = n in Nsetn_sub_mem1, frule_tac x = f in bspec, simp) apply (simp add:Pi_def, assumption+) done lemma (in Ring) ideal_inc_set_mult:"\A \ carrier R; ideal R B; ideal R C; set_mult R A B \ C \ \ sum_mult R A B \ C" apply (rule subsetI) apply (frule_tac x = x in sum_mult_mem1[of A B], simp add:ideal_subset1, assumption+) apply (erule exE, erule bexE, frule sym, thin_tac "\\<^sub>e R f n = x", simp, thin_tac "x = \\<^sub>e R f n", simp add:subsetD) apply (simp add:ideal_inc_set_multTr) done lemma (in Ring) AB_inc_sum_mult:"\ideal R A; ideal R B\ \ sum_mult R A B \ A \ B" apply (frule ideal_subset1[of A], frule ideal_subset1[of B]) apply (frule ideal_inc_set_mult [of "A" "B" "A"], assumption+) apply (rule subsetI, simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \\<^sub>r y = x", simp, frule_tac c = xa in subsetD[of A "carrier R"], assumption+, frule_tac c = y in subsetD[of B "carrier R"], assumption+, subst ring_tOp_commute, assumption+, simp add:ideal_ring_multiple) apply (frule ideal_inc_set_mult [of "A" "B" "B"], assumption+) apply (rule subsetI, simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \\<^sub>r y = x", simp, frule_tac c = xa in subsetD[of A "carrier R"], assumption+, simp add:ideal_ring_multiple) apply simp done lemma (in Ring) sum_mult_is_ideal_prod:"\ideal R A; ideal R B\ \ sum_mult R A B = A \\<^sub>r B" apply (rule equalityI) apply (frule ideal_prod_ideal [of "A" "B"], assumption+) apply (rule ideal_inc_set_mult) apply (simp add:ideal_subset1)+ apply (rule subsetI) apply (simp add:set_mult_def ideal_prod_def) apply (auto del:subsetI) apply (rule subsetI) apply (simp add:ideal_prod_def) apply (frule ideal_subset1[of A], frule ideal_sum_mult[of A B], frule ideal_zero[of A], blast, assumption) apply (frule_tac a = "sum_mult R A B" in forall_spec, simp) apply (rule subsetI, simp, thin_tac "\xa. ideal R xa \ {x. \i\A. \j\B. x = i \\<^sub>r j} \ xa \ x \ xa", (erule bexE)+, simp) apply (rule times_mem_sum_mult, assumption, simp add:ideal_subset1, assumption+) done lemma (in Ring) ideal_prod_assocTr0:"\ideal R A; ideal R B; ideal R C; y \ C; z \ set_mult R A B\ \ z \\<^sub>r y \ sum_mult R A (B \\<^sub>r C)" apply (simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "x \\<^sub>r ya = z", simp) apply (frule_tac h = x in ideal_subset[of A], assumption, frule_tac h = ya in ideal_subset[of B], assumption, frule_tac h = y in ideal_subset[of C], assumption, subst ring_tOp_assoc, assumption+) apply (frule ideal_subset1[of A], frule ideal_subset1[of B], frule ideal_subset1[of C], frule ideal_prod_ideal[of B C], assumption, frule ideal_subset1[of "B \\<^sub>r C"]) apply (rule times_mem_sum_mult[of A "B \\<^sub>r C"], assumption+, subst sum_mult_is_ideal_prod[of B C, THEN sym], assumption+, rule times_mem_sum_mult[of B C], assumption+) done lemma (in Ring) ideal_prod_assocTr1:"\ideal R A; ideal R B; ideal R C; y \ C\ \ \f \ {j. j\(n::nat)} \ set_mult R A B. (\\<^sub>e R f n) \\<^sub>r y \ A \\<^sub>r (B \\<^sub>r C)" apply (cut_tac ring_is_ag) apply (frule ideal_prod_ideal[of "B" "C"], assumption+, subst sum_mult_is_ideal_prod[of A "B \\<^sub>r C", THEN sym], assumption+) apply (induct_tac n) apply simp apply (simp add:ideal_prod_assocTr0) apply (rule ballI, frule_tac x = f in bspec, thin_tac "\f\{j. j \ n} \ set_mult R A B. \\<^sub>e R f n \\<^sub>r y \ sum_mult R A (B \\<^sub>r C)", rule Pi_I, simp, frule_tac f = f and A = "{j. j \ Suc n}" and B = "set_mult R A B" and x = x in funcset_mem, simp, assumption) apply simp apply (frule ideal_subset1[of A], frule ideal_subset1[of B], frule set_mult_sub[of A B], assumption, frule_tac f = f and A = "{j. j \ (Suc n)}" in extend_fun[of _ _ "set_mult R A B" "carrier R"], assumption, subst ring_distrib2, simp add:ideal_subset) apply (rule aGroup.nsum_mem, assumption) apply (simp add:Pi_def) apply (simp add:funcset_mem del:Pi_I', frule_tac f = f and A = "{j. j \ Suc n}" and B = "set_mult R A B" and x = "Suc n" in funcset_mem, simp) apply (frule ideal_subset1[of A], frule ideal_zero[of A], frule ideal_sum_mult[of A "B \\<^sub>r C"], blast, assumption) apply (rule ideal_pOp_closed, assumption+) apply (simp add:ideal_prod_assocTr0) done lemma (in Ring) ideal_quotient_idealTr:"\ideal R A; ideal R B; ideal R C; x \ carrier R;\c\C. x \\<^sub>r c \ ideal_quotient R A B\ \ f\{j. j \ n} \ set_mult R B C \ x \\<^sub>r (nsum R f n) \ A" apply (frule ideal_subset1 [of "A"], frule ideal_subset1 [of "B"]) apply (induct_tac n) apply (rule impI) apply (cut_tac n_in_Nsetn[of 0]) apply (frule funcset_mem, assumption+) apply (thin_tac "f \ {j. j \ 0} \ set_mult R B C") apply (simp add:set_mult_def) apply (erule bexE)+ apply (frule sym, thin_tac "xa \\<^sub>r y = f 0", simp) apply (frule_tac h = xa in ideal_subset[of B], assumption, frule_tac h = y in ideal_subset[of C], assumption) apply (frule_tac x = xa and y = y in ring_tOp_commute, assumption+, simp) apply (subst ring_tOp_assoc[THEN sym], assumption+) apply (frule_tac x = y in bspec, assumption, thin_tac "\c\C. x \\<^sub>r c \ A \\<^sub>R B") apply (simp add:ideal_quotient_def) (****** n ********) apply (rule impI) apply (frule func_pre) apply simp apply (cut_tac ring_is_ag) apply (frule ideal_subset1[of B], frule ideal_subset1[of C], frule set_mult_sub[of B C], assumption+) apply (cut_tac n = n in nsum_memr [of _ "f"], rule allI, rule impI, frule_tac x = i in funcset_mem, simp, simp add:subsetD) apply (frule_tac a = n in forall_spec, simp) apply (thin_tac "\l\n. \\<^sub>e R f l \ carrier R", frule_tac f = f and A = "{j. j \ Suc n}" and B = "set_mult R B C" and x = "Suc n" in funcset_mem, simp, frule_tac c = "f (Suc n)" in subsetD[of "set_mult R B C" " carrier R"], assumption+) apply (subst ring_distrib1, assumption+) apply (rule ideal_pOp_closed[of A], assumption+) apply (simp add: set_mult_def, (erule bexE)+, fold set_mult_def, frule sym, thin_tac "xa \\<^sub>r y = f (Suc n)", simp) apply (frule_tac c = xa in subsetD[of B "carrier R"], assumption+, frule_tac c = y in subsetD[of C "carrier R"], assumption+, frule_tac x = xa and y = y in ring_tOp_commute, assumption, simp, subst ring_tOp_assoc[THEN sym], assumption+) apply (simp add:ideal_quotient_def) done lemma (in Ring) ideal_quotient_ideal:"\ideal R A; ideal R B; ideal R C\ \ A \\<^sub>R B \\<^sub>R C = A \\<^sub>R B \\<^sub>r C" apply (rule equalityI) apply (rule subsetI) apply (simp add:ideal_quotient_def [of _ _ "C"]) apply (erule conjE) apply (simp add:ideal_quotient_def [of _ _ "B \\<^sub>r C"]) apply (rule ballI) apply (simp add:sum_mult_is_ideal_prod [THEN sym]) apply (simp add:sum_mult_def) apply (erule exE, erule bexE) apply (rename_tac x c n f) apply (frule sym) apply simp apply (simp add:ideal_quotient_idealTr) apply (rule subsetI) apply (simp add:sum_mult_is_ideal_prod [THEN sym]) apply (simp add:ideal_quotient_def) apply (erule conjE) apply (rule ballI) apply (rename_tac x c) apply (frule ideal_subset [of "C"], assumption+) apply (simp add:ring_tOp_closed) apply (rule ballI) apply (rename_tac x v u) apply (frule ideal_subset [of "B"], assumption+) apply (subst ring_tOp_assoc, assumption+) apply (frule ideal_subset1[of B], frule ideal_subset1[of C], frule_tac a = u and b = v in times_mem_sum_mult[of B C], assumption+) apply (frule_tac x = u and y = v in ring_tOp_commute, assumption, simp) done lemma (in Ring) ideal_prod_assocTr:"\ideal R A; ideal R B; ideal R C\ \ \f. (f \ {j. j \ (n::nat)} \ set_mult R (A \\<^sub>r B) C \ (\\<^sub>e R f n) \ A \\<^sub>r (B \\<^sub>r C))" apply (subgoal_tac "\x\(A \\<^sub>r B). \y\C. x \\<^sub>r y \ A \\<^sub>r (B \\<^sub>r C)") apply (induct_tac n) apply (rule allI) apply (rule impI) apply (frule_tac f = f and A = "{j. j \ 0}" and B = "set_mult R (A \\<^sub>r B) C" and x = 0 in funcset_mem, simp, simp) apply (simp add:set_mult_def) apply ((erule bexE)+, frule sym, thin_tac "x \\<^sub>r y = f 0", simp) apply (rule allI, rule impI) apply (frule func_pre) apply (frule_tac a = f in forall_spec, simp, thin_tac "\f. f \ {j. j \ n} \ set_mult R (A \\<^sub>r B) C \ \\<^sub>e R f n \ A \\<^sub>r (B \\<^sub>r C)", frule ideal_prod_ideal[of "B" "C"], assumption+, frule ideal_prod_ideal[of "A" "B \\<^sub>r C"], assumption+, simp) apply (rule ideal_pOp_closed[of "A \\<^sub>r (B \\<^sub>r C)"], assumption+) apply (cut_tac n = "Suc n" in n_in_Nsetn, frule_tac f = f and A = "{j. j \ Suc n}" and B = "set_mult R (A \\<^sub>r B) C" and x = "Suc n" in funcset_mem, assumption) apply (thin_tac "f \ {j. j \ n} \ set_mult R (A \\<^sub>r B) C", thin_tac "f \ {j. j \ Suc n} \ set_mult R (A \\<^sub>r B) C") apply (simp add:set_mult_def) apply ((erule bexE)+, frule sym, thin_tac "x \\<^sub>r y = f (Suc n)", simp) apply (rule ballI)+ apply (simp add:sum_mult_is_ideal_prod[of A B, THEN sym]) apply (frule ideal_subset1[of A], frule ideal_subset1[of B], frule_tac x = x in sum_mult_mem1[of A B], assumption+) apply (erule exE, erule bexE, frule sym, thin_tac "\\<^sub>e R f n = x", simp) apply (simp add:ideal_prod_assocTr1) done lemma (in Ring) ideal_prod_assoc:"\ideal R A; ideal R B; ideal R C\ \ (A \\<^sub>r B) \\<^sub>r C = A \\<^sub>r (B \\<^sub>r C)" apply (rule equalityI) apply (rule subsetI) apply (frule ideal_prod_ideal[of "A" "B"], assumption+) apply (frule sum_mult_is_ideal_prod[of "A \\<^sub>r B" "C"], assumption+) apply (frule sym) apply (thin_tac "sum_mult R (A \\<^sub>r B) C = (A \\<^sub>r B) \\<^sub>r C") apply simp apply (thin_tac "(A \\<^sub>r B) \\<^sub>r C = sum_mult R (A \\<^sub>r B) C") apply (thin_tac "ideal R (A \\<^sub>r B)") apply (frule ideal_prod_ideal[of "B" "C"], assumption+) apply (simp add:sum_mult_def) apply (erule exE, erule bexE) apply (frule sym, thin_tac "\\<^sub>e R f n = x", simp) apply (simp add:ideal_prod_assocTr) apply (rule subsetI) apply (frule ideal_prod_ideal[of "B" "C"], assumption+) apply (simp add:ideal_prod_commute [of "A" "B \\<^sub>r C"]) apply (frule ideal_prod_ideal[of "A" "B"], assumption+) apply (simp add:ideal_prod_commute[of "A \\<^sub>r B" "C"]) apply (simp add:ideal_prod_commute[of "A" "B"]) apply (simp add:ideal_prod_commute[of "B" "C"]) apply (frule ideal_prod_ideal[of "C" "B"], assumption+) apply (frule sum_mult_is_ideal_prod[of "C \\<^sub>r B" "A"], assumption+) apply (frule sym) apply (thin_tac "sum_mult R (C \\<^sub>r B) A = (C \\<^sub>r B) \\<^sub>r A") apply simp apply (thin_tac "(C \\<^sub>r B) \\<^sub>r A = sum_mult R (C \\<^sub>r B) A") apply (thin_tac "ideal R (C \\<^sub>r B)") apply (frule ideal_prod_ideal[of "B" "A"], assumption+) apply (simp add:sum_mult_def) apply (erule exE, erule bexE) apply (frule sym, thin_tac "\\<^sub>e R f n = x", simp) apply (simp add:ideal_prod_assocTr) done lemma (in Ring) prod_principal_idealTr0:" \a \ carrier R; b \ carrier R; z \ set_mult R (R \\<^sub>p a) (R \\<^sub>p b)\ \ z \ R \\<^sub>p (a \\<^sub>r b)" apply (simp add:set_mult_def, (erule bexE)+, simp add:Rxa_def, (erule bexE)+, simp) apply (frule_tac x = r and y = a and z = "ra \\<^sub>r b" in ring_tOp_assoc, assumption+, simp add:ring_tOp_closed, simp) apply (simp add:ring_tOp_assoc[THEN sym, of a _ b]) apply (frule_tac x = a and y = ra in ring_tOp_commute, assumption+, simp) apply (simp add:ring_tOp_assoc[of _ a b], frule_tac x = a and y = b in ring_tOp_closed, assumption) apply (simp add:ring_tOp_assoc[THEN sym, of _ _ "a \\<^sub>r b"], frule sym, thin_tac "r \\<^sub>r ra \\<^sub>r (a \\<^sub>r b) = z", simp, frule_tac x = r and y = ra in ring_tOp_closed, assumption+) apply blast done lemma (in Ring) prod_principal_idealTr1:" \a \ carrier R; b \ carrier R\ \ \f \ {j. j \ (n::nat)} \ set_mult R (R \\<^sub>p a) (R \\<^sub>p b). \\<^sub>e R f n \ R \\<^sub>p (a \\<^sub>r b)" apply (induct_tac n) apply (rule ballI, frule_tac f = f in funcset_mem[of _ "{j. j \ 0}" "set_mult R (R \\<^sub>p a) (R \\<^sub>p b)"], simp) apply (simp add:prod_principal_idealTr0) apply (rule ballI, frule func_pre, frule_tac x = f in bspec, assumption, thin_tac "\f\{j. j \ n} \ set_mult R (R \\<^sub>p a) (R \\<^sub>p b). \\<^sub>e R f n \ R \\<^sub>p (a \\<^sub>r b)") apply (frule ring_tOp_closed[of a b], assumption) apply (frule principal_ideal[of "a \\<^sub>r b"], simp, rule ideal_pOp_closed, assumption+) apply (cut_tac n = "Suc n" in n_in_Nsetn, frule_tac f = f and A = "{j. j \ Suc n}" and B = "set_mult R (R \\<^sub>p a) (R \\<^sub>p b)" in funcset_mem, assumption) apply (simp add:prod_principal_idealTr0) done lemma (in Ring) prod_principal_ideal:"\a \ carrier R; b \ carrier R\ \ (Rxa R a) \\<^sub>r (Rxa R b) = Rxa R (a \\<^sub>r b)" apply (frule principal_ideal[of "a"], frule principal_ideal[of "b"]) apply (subst sum_mult_is_ideal_prod[THEN sym, of "Rxa R a" "Rxa R b"], assumption+) apply (rule equalityI) apply (rule subsetI) apply (simp add:sum_mult_def) apply (erule exE, erule bexE) apply (frule sym, thin_tac "\\<^sub>e R f n = x", simp, thin_tac "x = \\<^sub>e R f n") apply (simp add:prod_principal_idealTr1) apply (rule subsetI) apply (simp add:Rxa_def, fold Rxa_def) apply (erule bexE) apply (simp add:ring_tOp_assoc[THEN sym]) apply (frule ideal_subset1[of "R \\<^sub>p a"], frule ideal_subset1[of "R \\<^sub>p b"]) apply (rule_tac a = "r \\<^sub>r a" and b = b in times_mem_sum_mult[of "R \\<^sub>p a" "R \\<^sub>p b"], assumption+) apply (simp add:Rxa_def, blast) apply (simp add:a_in_principal) done lemma (in Ring) principal_ideal_n_pow1:"a \ carrier R \ (Rxa R a)\<^bsup>\R n\<^esup> = Rxa R (a^\<^bsup>R n \<^esup>)" apply (cut_tac ring_one) apply (induct_tac n) apply simp apply (cut_tac a_in_principal[of "1\<^sub>r"]) apply (frule principal_ideal[of "1\<^sub>r"]) apply (frule ideal_inc_one, assumption, simp) apply (simp add:ring_one) apply simp apply (frule_tac n = n in npClose[of a], subst prod_principal_ideal, assumption+) apply (simp add:ring_tOp_commute) done lemma (in Ring) principal_ideal_n_pow:"\a \ carrier R; I = Rxa R a\ \ I \<^bsup>\R n\<^esup> = Rxa R (a^\<^bsup>R n\<^esup>)" apply simp apply (rule principal_ideal_n_pow1[of "a" "n"], assumption+) done text\more about \ideal_n_prod\\ lemma (in Ring) nprod_eqTr:" f \ {j. j \ (n::nat)} \ carrier R \ g \ {j. j \ n} \ carrier R \ (\j \ n. f j = g j) \ nprod R f n = nprod R g n" apply (induct_tac n) apply simp apply (rule impI, (erule conjE)+) apply (frule func_pre[of f], frule func_pre[of g], cut_tac n = n in Nsetn_sub_mem1, simp) done lemma (in Ring) nprod_eq:"\\j \ n. f j \ carrier R; \j \ n. g j \ carrier R; (\j \ (n::nat). f j = g j)\ \ nprod R f n = nprod R g n" apply (cut_tac nprod_eqTr[of f n g]) apply simp done definition mprod_expR :: "[('b, 'm) Ring_scheme, nat \ nat, nat \ 'b, nat] \ 'b" where "mprod_expR R e f n = nprod R (\j. ((f j)^\<^bsup>R (e j)\<^esup>)) n" (** Note that e j is a natural number for all j in Nset n **) lemma (in Ring) mprodR_Suc:"\e \ {j. j \ (Suc n)} \ {j. (0::nat) \ j}; f \ {j. j \ (Suc n)} \ carrier R\ \ mprod_expR R e f (Suc n) = (mprod_expR R e f n) \\<^sub>r ((f (Suc n))^\<^bsup>R (e (Suc n))\<^esup>)" apply (simp add:mprod_expR_def) done lemma (in Ring) mprod_expR_memTr:"e \ {j. j \ n} \ {j. (0::nat) \ j} \ f \ {j. j \ n} \ carrier R \ mprod_expR R e f n \ carrier R" apply (induct_tac n) apply (rule impI, (erule conjE)+) apply (cut_tac n_in_Nsetn[of 0], simp add: mprod_expR_def) apply (rule npClose, simp add:Pi_def) apply (rule impI, (erule conjE)+) apply (frule func_pre[of "e"], frule func_pre[of "f"]) apply simp apply (simp add:mprodR_Suc) apply (rule ring_tOp_closed, assumption+) apply (rule npClose, cut_tac n = "Suc n" in n_in_Nsetn) apply (simp add:Pi_def) done lemma (in Ring) mprod_expR_mem:"\ e \ {j. j \ n} \ {j. (0::nat) \ j}; f \ {j. j \ n} \ carrier R\ \ mprod_expR R e f n \ carrier R" apply (simp add:mprod_expR_memTr) done lemma (in Ring) prod_n_principal_idealTr:"e \ {j. j\n} \ {j. (0::nat)\j} \ f \ {j. j\n} \ carrier R \ (\k \ n. J k = (Rxa R (f k))\<^bsup>\R (e k)\<^esup>) \ ideal_n_prod R n J = Rxa R (mprod_expR R e f n)" apply (induct_tac n) apply (rule impI) apply (erule conjE)+ apply (simp add:mprod_expR_def) apply (subgoal_tac "J 0 = R \\<^sub>p (f 0) \<^bsup>\R (e 0)\<^esup>") apply simp apply (rule principal_ideal_n_pow[of "f 0" "R \\<^sub>p (f 0)"]) apply (cut_tac n_in_Nsetn[of 0], simp add:Pi_def) apply simp apply (cut_tac n_in_Nsetn[of 0], simp) apply (rule impI, (erule conjE)+) apply (frule func_pre[of "e"], frule func_pre[of "f"]) apply (cut_tac n = n in Nsetn_sub_mem1, simp add:mprodR_Suc) apply (cut_tac n = "Suc n" in n_in_Nsetn, simp) apply (frule_tac A = "{j. j \ Suc n}" and x = "Suc n" in funcset_mem[of "f" _ "carrier R"], simp) apply (frule_tac a = "f (Suc n)" and I = "R \\<^sub>p (f (Suc n))" and n = "e (Suc n)" in principal_ideal_n_pow) apply simp apply (subst prod_principal_ideal[THEN sym]) apply (simp add:mprod_expR_mem) apply (rule npClose, assumption+) apply simp done (************* used in Valuation2.thy *****************) lemma (in Ring) prod_n_principal_ideal:"\e \ {j. j\n} \ {j. (0::nat)\j}; f \ {j. j\n} \ carrier R; \k\ n. J k = (Rxa R (f k))\<^bsup>\R (e k)\<^esup>\ \ ideal_n_prod R n J = Rxa R (mprod_expR R e f n)" apply (simp add:prod_n_principal_idealTr[of e n f J]) done (*******************************************************) lemma (in Idomain) a_notin_n_pow1:"\a \ carrier R; \ Unit R a; a \ \; 0 < n\ \ a \ (Rxa R a) \<^bsup>\R (Suc n)\<^esup>" apply (rule contrapos_pp) apply (simp del:ipSuc) apply (simp del:ipSuc) apply (frule principal_ideal[of "a"]) apply (frule principal_ideal_n_pow[of "a" "R \\<^sub>p a" "Suc n"]) apply simp apply (simp del:ipSuc) apply (thin_tac "R \\<^sub>p a \<^bsup>\R (Suc n)\<^esup> = R \\<^sub>p (a^\<^bsup>R n\<^esup> \\<^sub>r a)") apply (thin_tac "ideal R (R \\<^sub>p a)") apply (simp add:Rxa_def) apply (erule bexE) apply (frule npClose[of "a" "n"]) apply (simp add:ring_tOp_assoc[THEN sym]) apply (frule ring_l_one[THEN sym, of "a"]) apply (subgoal_tac "1\<^sub>r \\<^sub>r a = r \\<^sub>r a^\<^bsup>R n\<^esup> \\<^sub>r a") apply (cut_tac b = "r \\<^sub>r (a^\<^bsup>R n\<^esup>)" in idom_mult_cancel_r[of "1\<^sub>r" _ "a"]) apply (simp add:ring_one) apply (simp add:ring_tOp_closed) apply assumption+ apply (thin_tac "1\<^sub>r \\<^sub>r a = r \\<^sub>r a^\<^bsup>R n\<^esup> \\<^sub>r a", thin_tac "a = 1\<^sub>r \\<^sub>r a", thin_tac "a = r \\<^sub>r a^\<^bsup>R n\<^esup> \\<^sub>r a") apply (subgoal_tac "1\<^sub>r = r \\<^sub>r (a^\<^bsup>R (Suc (n - Suc 0))\<^esup>)") prefer 2 apply (simp del:ipSuc) apply (thin_tac "1\<^sub>r = r \\<^sub>r a^\<^bsup>R n\<^esup>") apply (simp del:Suc_pred) apply (frule npClose[of "a" "n - Suc 0"]) apply (simp add:ring_tOp_assoc[THEN sym]) apply (frule_tac x = r and y = "a^\<^bsup>R (n - Suc 0)\<^esup>" in ring_tOp_closed, assumption) apply (simp add:ring_tOp_commute[of _ a]) apply (simp add:Unit_def) apply blast apply simp done lemma (in Idomain) a_notin_n_pow2:"\a \ carrier R; \ Unit R a; a \ \; 0 < n\ \ a^\<^bsup>R n\<^esup> \ (Rxa R a) \<^bsup>\R (Suc n)\<^esup>" apply (rule contrapos_pp) apply (simp del:ipSuc, simp del:ipSuc) apply (frule principal_ideal[of "a"]) apply (frule principal_ideal_n_pow[of "a" "R \\<^sub>p a" "Suc n"]) apply (simp, simp del:ipSuc) apply (thin_tac "R \\<^sub>p a \<^bsup>\R (Suc n)\<^esup> = R \\<^sub>p (a^\<^bsup>R n\<^esup> \\<^sub>r a)") apply (thin_tac "ideal R (R \\<^sub>p a)") apply (simp add:Rxa_def) apply (erule bexE) apply (frule idom_potent_nonzero[of "a" "n"], assumption+) apply (frule npClose[of "a" "n"]) apply (frule ring_l_one[THEN sym, of "a^\<^bsup>R n\<^esup> "]) apply (subgoal_tac "1\<^sub>r \\<^sub>r (a^\<^bsup>R n\<^esup>) = r \\<^sub>r ((a^\<^bsup>R n\<^esup>) \\<^sub>r a)") prefer 2 apply simp apply (thin_tac "a^\<^bsup>R n\<^esup> = 1\<^sub>r \\<^sub>r a^\<^bsup>R n\<^esup>", thin_tac "a^\<^bsup>R n\<^esup> = r \\<^sub>r (a^\<^bsup>R n\<^esup> \\<^sub>r a)") apply (simp add:ring_tOp_commute[of "a^\<^bsup>R n\<^esup>" a]) apply (simp add:ring_tOp_assoc[THEN sym]) apply (cut_tac ring_one, frule_tac b = "r \\<^sub>r a" in idom_mult_cancel_r[of "1\<^sub>r" _ "a^\<^bsup>R n\<^esup>"], simp add:ring_tOp_closed, assumption+) apply (simp add:ring_tOp_commute[of _ a]) apply (simp add:Unit_def, blast) done lemma (in Idomain) n_pow_not_prime:"\a \ carrier R; a \ \; 0 < n\ \ \ prime_ideal R ((Rxa R a) \<^bsup>\R (Suc n)\<^esup>)" apply (case_tac "n = 0") apply simp apply (case_tac "Unit R a") apply (simp del:ipSuc add:prime_ideal_def, rule impI) apply (frule principal_ideal[of "a"]) apply (frule principal_ideal_n_pow[of "a" "R \\<^sub>p a" "Suc n"]) apply simp apply (simp del:npow_suc) apply (simp del:npow_suc add:idom_potent_unit [of "a" "Suc n"]) apply (thin_tac "R \\<^sub>p a \\<^sub>r R \\<^sub>p a \<^bsup>\R n\<^esup> = R \\<^sub>p (a^\<^bsup>R (Suc n)\<^esup>)") apply (frule npClose[of "a" "Suc n"]) apply (frule a_in_principal[of "a^\<^bsup>R (Suc n)\<^esup>"]) apply (simp add: ideal_inc_unit) apply (frule a_notin_n_pow1[of "a" "n"], assumption+) apply (frule a_notin_n_pow2[of "a" "n"], assumption+) apply (frule npClose[of "a" "n"]) apply (frule principal_ideal[of "a"]) apply (frule principal_ideal_n_pow[of "a" "R \\<^sub>p a" "Suc n"]) apply simp apply (simp del:ipSuc npow_suc) apply (thin_tac "R \\<^sub>p a \<^bsup>\R (Suc n)\<^esup> = R \\<^sub>p (a^\<^bsup>R (Suc n)\<^esup>)") apply (subst prime_ideal_def) apply (simp del:npow_suc) apply (rule impI) apply (subgoal_tac "(a^\<^bsup>R n\<^esup>) \\<^sub>r a \ R \\<^sub>p (a^\<^bsup>R (Suc n)\<^esup>)") apply blast apply (simp add:Rxa_def) apply (frule ring_tOp_closed[of "a" "a^\<^bsup>R n\<^esup>"], assumption+) apply (frule ring_l_one[THEN sym, of "a \\<^sub>r (a^\<^bsup>R n\<^esup>)"]) apply (cut_tac ring_one) apply (simp add:ring_tOp_commute[of _ a], blast) done lemma (in Idomain) principal_pow_prime_condTr: "\a \ carrier R; a \ \; prime_ideal R ((Rxa R a) \<^bsup>\R (Suc n)\<^esup>)\ \ n = 0" apply (rule contrapos_pp, (simp del:ipSuc)+) apply (frule n_pow_not_prime[of "a" "n"], assumption+) apply (simp del:ipSuc) done lemma (in Idomain) principal_pow_prime_cond: "\a \ carrier R; a \ \; prime_ideal R ((Rxa R a) \<^bsup>\R n\<^esup>)\ \ n = Suc 0" apply (case_tac "n = 0") apply simp apply (simp add:prime_ideal_def) apply (erule conjE) apply (cut_tac ring_one, simp) apply (subgoal_tac "prime_ideal R (R \\<^sub>p a \<^bsup>\R (Suc (n - Suc 0))\<^esup>)") apply (frule principal_pow_prime_condTr[of "a" "n - Suc 0"], assumption+) apply simp apply simp done section "Extension and contraction" locale TwoRings = Ring + fixes R' (structure) assumes secondR: "Ring R'" definition i_contract :: "['a \ 'b, ('a, 'm1) Ring_scheme, ('b, 'm2) Ring_scheme, 'b set] \ 'a set" where "i_contract f R R' J = invim f (carrier R) J" definition i_extension :: "['a \ 'b, ('a, 'm1) Ring_scheme, ('b, 'm2) Ring_scheme, 'a set] \ 'b set" where "i_extension f R R' I = sum_mult R' (f ` I) (carrier R')" lemma (in TwoRings) i_contract_sub:"\f \ rHom R R'; ideal R' J \ \ (i_contract f R R' J) \ carrier R" by (auto simp add:i_contract_def invim_def) lemma (in TwoRings) i_contract_ideal:"\f \ rHom R R'; ideal R' J \ \ ideal R (i_contract f R R' J)" apply (cut_tac Ring, cut_tac secondR) apply (rule ideal_condition) apply (simp add:i_contract_sub) apply (simp add:i_contract_def invim_def) apply (cut_tac ring_zero) apply (cut_tac Ring) apply (frule rHom_0_0[of R R' f], assumption+, cut_tac Ring.ideal_zero[of R' J]) apply (frule sym, thin_tac "f \ = \\<^bsub>R'\<^esub>", simp, blast, assumption+) apply (rule ballI)+ apply (simp add:i_contract_def invim_def, (erule conjE)+) apply (cut_tac ring_is_ag, frule_tac x = y in aGroup.ag_mOp_closed[of R], assumption) apply (simp add:aGroup.ag_pOp_closed) apply (simp add:rHom_add) apply (frule_tac x = y in rHom_inv_inv[of R R' _ f], assumption+, simp, thin_tac "f (-\<^sub>a y) = -\<^sub>a\<^bsub>R'\<^esub> (f y)", frule_tac x = "f y" in Ring.ideal_inv1_closed[of R' J], assumption+, rule Ring.ideal_pOp_closed[of R'], assumption+) apply ((rule ballI)+, simp add:i_contract_def invim_def, erule conjE, simp add:ring_tOp_closed, simp add:rHom_tOp) apply (frule_tac a = r in rHom_mem[of f R R'], assumption, simp add:Ring.ideal_ring_multiple[of R' J]) done lemma (in TwoRings) i_contract_mono:"\f \ rHom R R'; ideal R' J1; ideal R' J2; J1 \ J2 \ \ i_contract f R R' J1 \ i_contract f R R' J2" apply (rule subsetI) apply (simp add:i_contract_def invim_def) apply (erule conjE) apply (rule subsetD, assumption+) done lemma (in TwoRings) i_contract_prime:"\f \ rHom R R'; prime_ideal R' P\ \ prime_ideal R (i_contract f R R' P)" apply (cut_tac Ring, cut_tac secondR) apply (simp add:prime_ideal_def, (erule conjE)+) apply (simp add:i_contract_ideal) apply (rule conjI) apply (rule contrapos_pp, simp+) apply (simp add:i_contract_def invim_def, erule conjE) apply (simp add:rHom_one) apply (rule ballI)+ apply (frule_tac a = x in rHom_mem[of "f" "R" "R'"], assumption+, frule_tac a = y in rHom_mem[of "f" "R" "R'"], assumption+) apply (rule impI) apply (simp add:i_contract_def invim_def, erule conjE) apply (simp add:rHom_tOp) done lemma (in TwoRings) i_extension_ideal:"\f \ rHom R R'; ideal R I \ \ ideal R' (i_extension f R R' I)" apply (cut_tac Ring, cut_tac secondR) apply (simp add:i_extension_def) apply (rule Ring.ideal_sum_mult [of "R'" "f ` I" "carrier R'"], assumption+) apply (rule subsetI) apply (simp add:image_def) apply (erule bexE, frule_tac a = xa in rHom_mem[of f R R'], rule ideal_subset, assumption+, simp) apply (frule ideal_zero, simp, blast) apply (simp add:Ring.whole_ideal[of R']) done lemma (in TwoRings) i_extension_mono:"\f \ rHom R R'; ideal R I1; ideal R I2; I1 \ I2 \ \ (i_extension f R R' I1) \ (i_extension f R R' I2)" apply (rule subsetI) apply (simp add:i_extension_def) apply (simp add:sum_mult_def) apply (erule exE, erule bexE) apply (cut_tac Ring.set_mult_mono[of R' "f ` I1" "f ` I2" "carrier R'"]) apply (frule_tac f = fa and A = "{j. j \ n}" in extend_fun[of _ _ "set_mult R' (f ` I1) (carrier R')" "set_mult R' (f ` I2) (carrier R')"], assumption+) apply blast apply (simp add:secondR) apply (simp add:image_def, rule subsetI, simp, erule bexE, frule_tac h = xb in ideal_subset[of I1], assumption, simp add:rHom_mem) apply (simp add:image_def, rule subsetI, simp, erule bexE, frule_tac h = xb in ideal_subset[of I2], assumption, simp add:rHom_mem) apply (rule subsetI, simp add:image_def, erule bexE, frule_tac c = xb in subsetD[of I1 I2], assumption+, blast) apply simp done lemma (in TwoRings) e_c_inc_self:"\f \ rHom R R'; ideal R I\ \ I \ i_contract f R R' (i_extension f R R' I)" apply (rule subsetI) apply (simp add:i_contract_def i_extension_def invim_def) apply (simp add:ideal_subset) apply (cut_tac secondR, frule Ring.ring_one [of "R'"]) apply (frule_tac h = x in ideal_subset[of I], assumption, frule_tac f = f and A = R and R = R' and a = x in rHom_mem, assumption) apply (frule_tac t = "f x" in Ring.ring_r_one[THEN sym, of R'], assumption) apply (frule_tac a = "f x" and b = "1\<^sub>r\<^bsub>R'\<^esub>" in Ring.times_mem_sum_mult[of R' "f ` I" "carrier R'"], rule subsetI, simp add:image_def, erule bexE, frule_tac h = xb in ideal_subset[of I], assumption, simp add:rHom_mem, simp, simp add:image_def, blast, assumption+) apply simp done lemma (in TwoRings) c_e_incd_self:"\f \ rHom R R'; ideal R' J \ \ i_extension f R R' (i_contract f R R' J) \ J" apply (rule subsetI) apply (simp add:i_extension_def) apply (simp add:sum_mult_def) apply (erule exE, erule bexE) apply (cut_tac secondR, frule_tac n = n and f = fa in Ring.ideal_nsum_closed[of R' J ], assumption) apply (rule allI, rule impI) apply ( frule_tac f = fa and A = "{j. j \ n}" and B = "set_mult R' (f ` i_contract f R R' J) (carrier R')" and x = j in funcset_mem, simp) apply ( thin_tac "fa \ {j. j \ n} \ set_mult R' (f ` i_contract f R R' J) (carrier R')") apply (simp add:set_mult_def, (erule bexE)+, simp add:i_contract_def invim_def, erule conjE) apply (frule_tac x = "f xa" and r = y in Ring.ideal_ring_multiple1[of R' J], assumption+, simp) apply simp done lemma (in TwoRings) c_e_c_eq_c:"\f \ rHom R R'; ideal R' J \ \ i_contract f R R' (i_extension f R R' (i_contract f R R' J)) = i_contract f R R' J" apply (frule i_contract_ideal [of "f" "J"], assumption) apply (frule e_c_inc_self [of "f" "i_contract f R R' J"], assumption+) apply (frule c_e_incd_self [of "f" "J"], assumption+) apply (frule i_contract_mono [of "f" "i_extension f R R' (i_contract f R R' J)" "J"]) apply (rule i_extension_ideal, assumption+) apply (rule equalityI, assumption+) done lemma (in TwoRings) e_c_e_eq_e:"\f \ rHom R R'; ideal R I \ \ i_extension f R R' (i_contract f R R' (i_extension f R R' I)) = i_extension f R R' I" apply (frule i_extension_ideal [of "f" "I"], assumption+) apply (frule c_e_incd_self [of "f" "i_extension f R R' I"], assumption+) apply (rule equalityI, assumption+) apply (thin_tac "i_extension f R R' (i_contract f R R' (i_extension f R R' I)) \ i_extension f R R' I") apply (frule e_c_inc_self [of "f" "I"], assumption+) apply (rule i_extension_mono [of "f" "I" "i_contract f R R' (i_extension f R R' I)"], assumption+) apply (rule i_contract_ideal, assumption+) done section "Complete system of representatives" definition csrp_fn :: "[_, 'a set] \ 'a set \ 'a" where "csrp_fn R I = (\x\carrier (R /\<^sub>r I). (if x = I then \\<^bsub>R\<^esub> else SOME y. y \ x))" definition csrp :: "[_ , 'a set] \ 'a set" where "csrp R I == (csrp_fn R I) ` (carrier (R /\<^sub>r I))" (** complete system of representatives having 1-1 correspondence with carrier (R /\<^sub>r I) **) lemma (in Ring) csrp_mem:"\ideal R I; a \ carrier R\ \ csrp_fn R I (a \\<^bsub>R\<^esub> I) \ a \\<^bsub>R\<^esub> I" apply (simp add:csrp_fn_def qring_carrier) apply (case_tac "a \\<^bsub>R\<^esub> I = I") apply simp apply (rule conjI, rule impI) apply (simp add:ideal_zero) apply (rule impI) apply (cut_tac ring_zero) apply (frule_tac x = \ in bspec, assumption+) apply (thin_tac "\a\carrier R. a \\<^bsub>R\<^esub> I \ I") apply (frule ideal_zero[of "I"]) apply (frule ar_coset_same4[of "I" "\"], assumption+, simp) apply simp apply (rule conjI) apply (rule impI, rule someI2_ex) apply (frule a_in_ar_coset[of "I" "a"], assumption+, blast, assumption+) apply (rule impI) apply (frule_tac x = a in bspec, assumption+, thin_tac "\aa\carrier R. aa \\<^bsub>R\<^esub> I \ a \\<^bsub>R\<^esub> I", simp) done lemma (in Ring) csrp_same:"\ideal R I; a \ carrier R\ \ csrp_fn R I (a \\<^bsub>R\<^esub> I) \\<^bsub>R\<^esub> I = a \\<^bsub>R\<^esub> I" apply (frule csrp_mem[of "I" "a"], assumption+) apply (rule ar_cos_same[of "a" "I" "csrp_fn R I (a \\<^bsub>R\<^esub> I)"], assumption+) done lemma (in Ring) csrp_mem1:"\ideal R I; x \ carrier (R /\<^sub>r I)\ \ csrp_fn R I x \ x" apply (simp add:qring_carrier, erule bexE, frule sym, thin_tac "a \\<^bsub>R\<^esub> I = x", simp) apply (simp add:csrp_mem) done lemma (in Ring) csrp_fn_mem:"\ideal R I; x \ carrier (R /\<^sub>r I)\ \ (csrp_fn R I x) \ carrier R" apply (simp add:qring_carrier, erule bexE, frule sym, thin_tac "a \\<^bsub>R\<^esub> I = x", simp, frule_tac a = a in csrp_mem[of "I"], assumption+) apply (rule_tac a = a and x = "csrp_fn R I (a \\<^bsub>R\<^esub> I)" in ar_coset_subsetD[of "I"], assumption+) done lemma (in Ring) csrp_eq_coset:"\ideal R I; x \ carrier (R /\<^sub>r I)\ \ (csrp_fn R I x) \\<^bsub>R\<^esub> I = x" apply (simp add:qring_carrier, erule bexE) apply (frule sym, thin_tac "a \\<^bsub>R\<^esub> I = x", simp) apply (frule_tac a = a in csrp_mem[of "I"], assumption+) apply (rule ar_cos_same, assumption+) done lemma (in Ring) csrp_nz_nz:"\ideal R I; x \ carrier (R /\<^sub>r I); x \ \\<^bsub>(R /\<^sub>r I)\<^esub>\ \ (csrp_fn R I x) \ \" apply (rule contrapos_pp, simp+) apply (frule csrp_eq_coset[of "I" "x"], assumption+, simp) apply (simp add:qring_zero[of "I"]) apply (frule ideal_zero[of "I"]) apply ( cut_tac ring_zero) apply (simp add:Qring_fix1 [of "\" "I"]) done lemma (in Ring) csrp_diff_in_vpr:"\ideal R I; x \ carrier R\ \ x \ (-\<^sub>a (csrp_fn R I (pj R I x))) \ I" apply (frule csrp_mem[of "I" "x"], frule csrp_same[of "I" "x"], simp add:pj_mem, assumption, frule ar_coset_subsetD[of I x "csrp_fn R I (x \\<^bsub>R\<^esub> I)"], assumption+) apply (frule belong_ar_coset2[of I x "csrp_fn R I (x \\<^bsub>R\<^esub> I)"], assumption+, frule ideal_inv1_closed[of I "csrp_fn R I (x \\<^bsub>R\<^esub> I) \ -\<^sub>a x"], assumption+, cut_tac ring_is_ag, frule aGroup.ag_mOp_closed[of R x], assumption, simp add:aGroup.ag_pOp_commute[of R "csrp_fn R I (x \\<^bsub>R\<^esub> I)" "-\<^sub>a x"]) apply (simp add:aGroup.ag_p_inv[of R "-\<^sub>a x" "csrp_fn R I (x \\<^bsub>R\<^esub> I)"], simp add:aGroup.ag_inv_inv, cut_tac Ring, simp add:pj_mem[of R I x]) done lemma (in Ring) csrp_pj:"\ideal R I; x \ carrier (R /\<^sub>r I)\ \ (pj R I) (csrp_fn R I x) = x" apply(cut_tac Ring, frule csrp_fn_mem[of "I" "x"], assumption+, simp add:pj_mem[of "R" "I" "csrp_fn R I x"], simp add:csrp_eq_coset) done section "Polynomial ring" text\In this section, we treat a ring of polynomials over a ring S. Numbers are of type ant\ definition pol_coeff :: "[('a, 'more) Ring_scheme, (nat \ (nat \ 'a))] \ bool" where "pol_coeff S c \ (\j \ (fst c). (snd c) j \ carrier S)" definition c_max :: "[('a, 'more) Ring_scheme, nat \ (nat \ 'a)] \ nat" where "c_max S c = (if {j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>} = {} then 0 else n_max {j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>})" definition polyn_expr :: "[('a, 'more) Ring_scheme, 'a, nat, nat \ (nat \ 'a)] \ 'a" where "polyn_expr R X k c == nsum R (\j. ((snd c) j) \\<^sub>r\<^bsub>R\<^esub> (X^\<^bsup>R j\<^esup>)) k" definition algfree_cond :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a] \ bool" where "algfree_cond R S X \ (\c. pol_coeff S c \ (\k \ (fst c). (nsum R (\j. ((snd c) j) \\<^sub>r\<^bsub>R\<^esub> (X^\<^bsup>R j\<^esup>)) k = \\<^bsub>R\<^esub> \ (\j \ k. (snd c) j = \\<^bsub>S\<^esub>))))" locale PolynRg = Ring + fixes S (structure) fixes X (structure) assumes X_mem_R:"X \ carrier R" and not_zeroring:"\ zeroring S" and subring: "Subring R S" and algfree: "algfree_cond R S X" and S_X_generate:"x \ carrier R \ \f. pol_coeff S f \ x = polyn_expr R X (fst f) f" (** a polynomial is an element of a polynomial ring **) section \Addition and multiplication of \polyn_exprs\\ subsection \Simple properties of a \polyn_ring\\ lemma Subring_subset:"Subring R S \ carrier S \ carrier R" by (simp add:Subring_def) lemma (in Ring) subring_Ring:"Subring R S \ Ring S" by (simp add:Subring_def) lemma (in Ring) mem_subring_mem_ring:"\Subring R S; x \ carrier S\ \ x \ carrier R" by (simp add:Subring_def, (erule conjE)+, simp add: subsetD) lemma (in Ring) Subring_pOp_ring_pOp:"\Subring R S; a \ carrier S; b \ carrier S \ \ a \\<^bsub>S\<^esub> b = a \ b" apply (simp add:Subring_def, (erule conjE)+) apply (frule rHom_add[of "ridmap S" S R a b], assumption+) apply (cut_tac Ring.ring_is_ag[of S], frule aGroup.ag_pOp_closed[of S a b], assumption+, simp add:ridmap_def, assumption) done lemma (in Ring) Subring_tOp_ring_tOp:"\Subring R S; a \ carrier S; b \ carrier S \ \ a \\<^sub>r\<^bsub>S\<^esub> b = a \\<^sub>r b" apply (simp add:Subring_def, (erule conjE)+) apply (frule rHom_tOp[of "S" "R" "a" "b" "ridmap S"], rule Ring_axioms, assumption+) apply (frule Ring.ring_tOp_closed[of "S" "a" "b"], assumption+, simp add:ridmap_def) done lemma (in Ring) Subring_one_ring_one:"Subring R S \ 1\<^sub>r\<^bsub>S\<^esub> = 1\<^sub>r" apply (simp add:Subring_def, (erule conjE)+) apply (frule rHom_one[of "S" "R" "ridmap S"], rule Ring_axioms, assumption+) apply (simp add:ridmap_def, simp add:Ring.ring_one[of S]) done lemma (in Ring) Subring_zero_ring_zero:"Subring R S \ \\<^bsub>S\<^esub> = \" apply (simp add:Subring_def, (erule conjE)+, frule rHom_0_0[of "S" "R" "ridmap S"], rule Ring_axioms, assumption+, simp add:ridmap_def, simp add:Ring.ring_zero[of "S"]) done lemma (in Ring) Subring_minus_ring_minus:"\Subring R S; x \ carrier S\ \ -\<^sub>a\<^bsub>S\<^esub> x = -\<^sub>a x" apply (simp add:Subring_def, (erule conjE)+, simp add:rHom_def, (erule conjE)+) apply (cut_tac ring_is_ag, frule Ring.ring_is_ag[of "S"]) apply (frule aHom_inv_inv[of "S" "R" "ridmap S" "x"], assumption+, frule aGroup.ag_mOp_closed[of "S" "x"], assumption+) apply (simp add:ridmap_def) done lemma (in PolynRg) Subring_pow_ring_pow:"x \ carrier S \ x^\<^bsup>S n\<^esup> = x^\<^bsup>R n\<^esup>" apply (cut_tac subring, frule subring_Ring) apply (induct_tac n) apply (simp, simp add:Subring_one_ring_one) apply (frule_tac n = n in Ring.npClose[of S x], assumption+) apply (simp add:Subring_tOp_ring_tOp) done lemma (in PolynRg) is_Ring: "Ring R" .. lemma (in PolynRg) polyn_ring_nonzero:"1\<^sub>r \ \" apply (cut_tac Ring, cut_tac subring) apply (simp add:Subring_zero_ring_zero[THEN sym]) apply (simp add:Subring_one_ring_one[THEN sym]) using Ring.Zero_ring1 not_zeroring subring_Ring apply blast done lemma (in PolynRg) polyn_ring_S_nonzero:"1\<^sub>r\<^bsub>S\<^esub> \ \\<^bsub>S\<^esub>" apply (cut_tac subring) apply (simp add:Subring_zero_ring_zero) apply (simp add:Subring_one_ring_one) apply (simp add:polyn_ring_nonzero) done lemma (in PolynRg) polyn_ring_X_nonzero:"X \ \" apply (cut_tac algfree, cut_tac subring) apply (simp add:algfree_cond_def) apply (rule contrapos_pp, simp+) apply (drule_tac x = "Suc 0" in spec) apply (subgoal_tac "pol_coeff S ((Suc 0), (\j\{l. l \ (Suc 0)}. if j = 0 then \\<^bsub>S\<^esub> else 1\<^sub>r\<^bsub>S\<^esub>))") apply (drule_tac x = "\j\{l. l \ (Suc 0)}. if j = 0 then \\<^bsub>S\<^esub> else 1\<^sub>r\<^bsub>S\<^esub>" in spec) apply (erule conjE, simp) apply (simp only:Nset_1) apply (drule_tac a = "Suc 0" in forall_spec, simp) apply simp apply (cut_tac subring, simp add:Subring_zero_ring_zero, simp add:Subring_one_ring_one, cut_tac ring_zero, cut_tac ring_one, simp add:ring_r_one, simp add:ring_times_x_0, cut_tac ring_is_ag, simp add:aGroup.ag_r_zero, drule_tac a = "Suc 0" in forall_spec, simp, simp) apply (cut_tac polyn_ring_S_nonzero, simp add:Subring_zero_ring_zero) apply (thin_tac "\b. pol_coeff S (Suc 0, b) \ (\k\Suc 0. \\<^sub>e R (\j. b j \\<^sub>r \^\<^bsup>R j\<^esup>) k = \ \ (\j\k. b j = \\<^bsub>S\<^esub>))", simp add:pol_coeff_def, rule allI, simp add:Subring_def, simp add:Ring.ring_zero, (rule impI)+, simp add:Ring.ring_one) done subsection "Coefficients of a polynomial" lemma (in PolynRg) pol_coeff_split:"pol_coeff S f = pol_coeff S (fst f, snd f)" by simp lemma (in PolynRg) pol_coeff_cartesian:"pol_coeff S c \ (fst c, snd c) = c" by simp lemma (in PolynRg) split_pol_coeff:"\pol_coeff S c; k \ (fst c)\ \ pol_coeff S (k, snd c)" by (simp add:pol_coeff_def) lemma (in PolynRg) pol_coeff_pre:"pol_coeff S ((Suc n), f) \ pol_coeff S (n, f)" apply (simp add:pol_coeff_def) done lemma (in PolynRg) pol_coeff_le:"\pol_coeff S c; n \ (fst c)\ \ pol_coeff S (n, (snd c))" apply (simp add:pol_coeff_def) done lemma (in PolynRg) pol_coeff_mem:"\pol_coeff S c; j \ (fst c)\ \ ((snd c) j) \ carrier S" by (simp add:pol_coeff_def) lemma (in PolynRg) pol_coeff_mem_R:"\pol_coeff S c; j \ (fst c)\ \ ((snd c) j) \ carrier R" apply (cut_tac subring, frule subring_Ring) apply (frule pol_coeff_mem[of c "j"], assumption+, simp add:mem_subring_mem_ring) done lemma (in PolynRg) Slide_pol_coeff:"\pol_coeff S c; n < (fst c)\ \ pol_coeff S (((fst c) - Suc n), (\x. (snd c) (Suc (n + x))))" apply (simp add: pol_coeff_def) done subsection \Addition of \polyn_exprs\\ lemma (in PolynRg) monomial_mem:"pol_coeff S c \ \j \ (fst c). (snd c) j \\<^sub>r X^\<^bsup>R j\<^esup> \ carrier R" apply (rule allI, rule impI) apply (rule ring_tOp_closed) apply (simp add:pol_coeff_mem_R[of c], cut_tac X_mem_R, simp add:npClose) done lemma (in PolynRg) polyn_mem:"\pol_coeff S c; k \ (fst c)\ \ polyn_expr R X k c \ carrier R" apply (simp add:polyn_expr_def, cut_tac ring_is_ag) apply (rule aGroup.nsum_mem[of R k "\j. (snd c) j \\<^sub>r X^\<^bsup>R j\<^esup>"], assumption+) apply (simp add:monomial_mem) done lemma (in PolynRg) polyn_exprs_eq:"\pol_coeff S c; pol_coeff S d; k \ (min (fst c) (fst d)); \j \ k. (snd c) j = (snd d) j\ \ polyn_expr R X k c = polyn_expr R X k d" apply (cut_tac ring_is_ag, simp add:polyn_expr_def, cut_tac subring, cut_tac X_mem_R) apply (rule aGroup.nsum_eq[of R k "\j. (snd c) j \\<^sub>r X^\<^bsup>R j\<^esup>" "\j. (snd d) j \\<^sub>r X^\<^bsup>R j\<^esup>"], assumption) apply (simp add:monomial_mem)+ done lemma (in PolynRg) polyn_expr_restrict:"pol_coeff S (Suc n, f) \ polyn_expr R X n (Suc n, f) = polyn_expr R X n (n, f)" apply (cut_tac subring, frule subring_Ring, cut_tac pol_coeff_le[of "(Suc n, f)" n]) apply (cut_tac polyn_exprs_eq[of "(Suc n, f)" "(n, f)" n], (simp add:pol_coeff_split[THEN sym])+) done lemma (in PolynRg) polyn_expr_short:"\pol_coeff S c; k \ (fst c)\ \ polyn_expr R X k c = polyn_expr R X k (k, snd c)" apply (rule polyn_exprs_eq[of c "(k, snd c)" k], assumption+) apply (simp add:pol_coeff_def) apply (simp) apply simp done lemma (in PolynRg) polyn_expr0:"pol_coeff S c \ polyn_expr R X 0 c = (snd c) 0" apply (simp add:polyn_expr_def) apply (cut_tac subring, cut_tac subring_Ring[of S]) apply (frule pol_coeff_mem[of c 0], simp) apply (frule mem_subring_mem_ring [of S "(snd c) 0"], assumption) apply (simp add:ring_r_one, assumption) done lemma (in PolynRg) polyn_expr_split:" polyn_expr R X k f = polyn_expr R X k (fst f, snd f)" by simp lemma (in PolynRg) polyn_Suc:"Suc n \ (fst c) \ polyn_expr R X (Suc n) ((Suc n), (snd c)) = polyn_expr R X n c \ ((snd c) (Suc n)) \\<^sub>r (X^\<^bsup>R (Suc n)\<^esup>)" by (simp add:polyn_expr_def) lemma (in PolynRg) polyn_Suc_split:"pol_coeff S (Suc n, f) \ polyn_expr R X (Suc n) ((Suc n), f) = polyn_expr R X n (n, f) \ (f (Suc n)) \\<^sub>r (X^\<^bsup>R (Suc n)\<^esup>)" apply (cut_tac polyn_Suc[of n "(Suc n, f)"]) apply (simp del:npow_suc) apply (subst polyn_expr_short[of "(Suc n, f)" n], assumption+, simp) apply (simp del:npow_suc) apply simp done lemma (in PolynRg) polyn_n_m:"\pol_coeff S c; n < m; m \ (fst c)\ \ polyn_expr R X m (m, (snd c)) = polyn_expr R X n (n, (snd c)) \ (fSum R (\j. ((snd c) j) \\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m)" apply (simp add:polyn_expr_def, cut_tac ring_is_ag) apply (rule aGroup.nsum_split1[of "R" m "\j. ((snd c) j) \\<^sub>r (X^\<^bsup>R j\<^esup>)" n], assumption+) apply (rule allI, rule impI) apply (frule_tac monomial_mem[of c], frule_tac i = j and j = m and k = "(fst c)" in le_trans, assumption+, simp+) done lemma (in PolynRg) polyn_n_m1:"\pol_coeff S c; n < m; m \ (fst c)\ \ polyn_expr R X m c = polyn_expr R X n c \ (fSum R (\j. ((snd c) j) \\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m)" apply (subst polyn_expr_short[of c n], assumption) apply (frule_tac x = n and y = m and z = "fst c" in less_le_trans, assumption, simp add:less_imp_le) apply (subst polyn_expr_short[of c m], assumption+) apply (simp add:polyn_n_m) done lemma (in PolynRg) polyn_n_m_mem:"\pol_coeff S c; n < m; m \ (fst c)\ \ (fSum R (\j. ((snd c) j) \\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m) \ carrier R" apply (simp add:fSum_def) apply (cut_tac ring_is_ag, rule_tac n = "m - Suc n" in aGroup.nsum_mem, assumption+) apply (rule allI, rule impI, simp del:npow_suc add:cmp_def slide_def) apply (rule ring_tOp_closed) apply (simp add:pol_coeff_def) apply (frule_tac a = "Suc (n + j)" in forall_spec, arith) apply (cut_tac subring) apply (simp add:mem_subring_mem_ring) apply (rule npClose) apply (cut_tac X_mem_R, simp del:npow_suc add:npClose) done lemma (in PolynRg) polyn_n_ms_eq:"\pol_coeff S c; pol_coeff S d; m \ min (fst c) (fst d); n < m; \j\nset (Suc n) m. (snd c) j = (snd d) j\ \ (fSum R (\j. ((snd c) j) \\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m) = (fSum R (\j. ((snd d) j) \\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m)" apply (cut_tac ring_is_ag) apply (cut_tac aGroup.fSum_eq1[of R "Suc n" m "\j. (snd c) j \\<^sub>r X^\<^bsup>R j\<^esup>" "\j. (snd d) j \\<^sub>r X^\<^bsup>R j\<^esup>"], assumption+) apply (rule Suc_leI, assumption, simp add:nset_def, simp add:monomial_mem) apply (frule Suc_leI, rule ballI, simp add:nset_def) apply (simp add:monomial_mem) apply simp done lemma (in PolynRg) polyn_addTr: "(pol_coeff S (n, f)) \ (pol_coeff S (n, g)) \ (polyn_expr R X n (n, f)) \ (polyn_expr R X n (n, g)) = nsum R (\j. ((f j) \\<^bsub>S\<^esub> (g j)) \\<^sub>r (X^\<^bsup>R j\<^esup>)) n" apply (cut_tac subring, frule subring_Ring[of S]) apply (induct_tac n) apply (rule impI, simp, erule conjE) apply (simp add:polyn_expr0) apply (cut_tac pol_coeff_mem[of "(0, f)" 0], simp, cut_tac pol_coeff_mem[of "(0, g)" 0], simp, frule mem_subring_mem_ring[of S "f 0"], assumption+, frule mem_subring_mem_ring[of S "g 0"], assumption+, frule Ring.ring_is_ag[of S], frule aGroup.ag_pOp_closed[of S "f 0" "g 0"], assumption+, frule mem_subring_mem_ring[of S "f 0 \\<^bsub>S\<^esub> g 0"], assumption+) apply (simp add:ring_r_one) apply (simp add:Subring_pOp_ring_pOp[of S "f 0" "g 0"]) apply (simp del:npow_suc)+ apply (rule impI, erule conjE) apply (frule_tac n = n in pol_coeff_pre[of _ f], frule_tac n = n in pol_coeff_pre[of _ g], simp del:npow_suc) apply (cut_tac n = n and c = "(Suc n, f)" in polyn_Suc, simp del:npow_suc, simp del:npow_suc, thin_tac "polyn_expr R X (Suc n) (Suc n, f) = polyn_expr R X n (Suc n, f) \ f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") apply (cut_tac n = n and c = "(Suc n, g)" in polyn_Suc, simp del:npow_suc, simp del:npow_suc, thin_tac "polyn_expr R X (Suc n) (Suc n, g) = polyn_expr R X n (Suc n, g) \ g (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") apply (cut_tac c = "(Suc n, f)" and k = n in polyn_mem, assumption, simp del:npow_suc, cut_tac k = n and c = "(Suc n, g)" in polyn_mem, assumption, simp del:npow_suc) apply (frule_tac j = "Suc n" and c = "(Suc n, f)" in pol_coeff_mem_R, simp, frule_tac j = "Suc n" and c = "(Suc n, g)" in pol_coeff_mem_R, simp, cut_tac X_mem_R, frule_tac n = "Suc n" in npClose[of "X"], simp del:npow_suc) apply (frule_tac x = "f (Suc n)" and y = "X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed, assumption+, frule_tac x = "g (Suc n)" and y = "X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed, assumption+) apply (cut_tac ring_is_ag, subst aGroup.pOp_assocTr43, assumption+) apply (frule_tac x = "f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>" and y = "polyn_expr R X n (Suc n, g)" in aGroup.ag_pOp_commute[of R], assumption+, simp del:npow_suc, thin_tac "f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup> \ polyn_expr R X n (Suc n, g) = polyn_expr R X n (Suc n, g) \ f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") apply (subst aGroup.pOp_assocTr43[THEN sym], assumption+, simp del:npow_suc add:polyn_expr_restrict) apply (frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp, frule_tac c = "(Suc n, g)" and j = "Suc n" in pol_coeff_mem, simp) apply (subst ring_distrib2[THEN sym], assumption+) apply (frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp, frule_tac c = "(Suc n, g)" and j = "Suc n" in pol_coeff_mem, simp) apply (frule_tac a = "f (Suc n)" and b = "g (Suc n)" in Subring_pOp_ring_pOp[of S], simp, simp) apply simp done lemma (in PolynRg) polyn_add_n:"\pol_coeff S (n, f); pol_coeff S (n, g)\ \ (polyn_expr R X n (n, f)) \ (polyn_expr R X n (n, g)) = nsum R (\j. ((f j) \\<^bsub>S\<^esub> (g j)) \\<^sub>r (X^\<^bsup>R j\<^esup>)) n" by (simp add:polyn_addTr) definition add_cf :: "[('a, 'm) Ring_scheme, nat \ (nat \ 'a), nat \ (nat \ 'a)] \ nat \ (nat \ 'a)" where "add_cf S c d = (if (fst c) < (fst d) then ((fst d), \j. (if j \ (fst c) then (((snd c) j) \\<^bsub>S\<^esub> ((snd d) j)) else ((snd d) j))) else if (fst c) = (fst d) then ((fst c), \j. ((snd c) j \\<^bsub>S\<^esub> (snd d) j)) else ((fst c), \j. (if j \ (fst d) then ((snd c) j \\<^bsub>S\<^esub> (snd d) j) else ((snd c) j))))" lemma (in PolynRg) add_cf_pol_coeff:"\pol_coeff S c; pol_coeff S d\ \ pol_coeff S (add_cf S c d)" apply (cut_tac subring, frule subring_Ring[of S], frule Ring.ring_is_ag[of S]) apply (simp add:pol_coeff_def) apply (rule allI, rule impI) apply (case_tac "(fst c) < (fst d)", simp add:add_cf_def) apply (rule impI, rule aGroup.ag_pOp_closed, assumption+, simp+) apply (drule leI[of "fst c" "fst d"], drule le_imp_less_or_eq[of "fst d" "fst c"]) apply (erule disjE) apply (simp add:add_cf_def, rule impI) apply (frule Ring.ring_is_ag[of S], rule aGroup.ag_pOp_closed, assumption, simp+) apply (simp add:add_cf_def) apply (frule Ring.ring_is_ag[of S], rule aGroup.ag_pOp_closed, assumption, simp+) done lemma (in PolynRg) add_cf_len:"\pol_coeff S c; pol_coeff S d\ \ fst (add_cf S c d) = (max (fst c) (fst d))" by (simp add: add_cf_def max.absorb1 max.absorb2) lemma (in PolynRg) polyn_expr_restrict1:"\pol_coeff S (n, f); pol_coeff S (Suc (m + n), g)\ \ polyn_expr R X (m + n) (add_cf S (n, f) (m + n, g)) = polyn_expr R X (m + n) (m + n, snd (add_cf S (n, f) (Suc (m + n), g)))" apply (frule pol_coeff_pre[of "m+n" g]) apply (frule add_cf_pol_coeff[of "(n, f)" "(Suc (m + n), g)"], assumption+, frule add_cf_pol_coeff[of "(n, f)" "(m + n, g)"], assumption+) apply (rule polyn_exprs_eq[of "add_cf S (n, f) (m + n, g)" "(m + n, snd (add_cf S (n, f) (Suc (m + n), g)))" "m + n"], assumption+) apply (rule split_pol_coeff[of "add_cf S (n, f) (Suc (m + n), g)" "m + n"], assumption, simp add:add_cf_len) apply (simp add:add_cf_len) apply (rule allI, rule impI) apply (simp add:add_cf_def) done lemma (in PolynRg) polyn_add_n1:"\pol_coeff S (n, f); pol_coeff S (n, g)\ \ (polyn_expr R X n (n, f)) \ (polyn_expr R X n (n, g)) = polyn_expr R X n (add_cf S (n, f) (n, g))" apply (subst polyn_add_n, assumption+) apply (simp add:polyn_expr_def add_cf_def) done lemma (in PolynRg) add_cf_val_hi:"(fst c) < (fst d) \ snd (add_cf S c d) (fst d) = (snd d) (fst d)" by (simp add:add_cf_def) lemma (in PolynRg) add_cf_commute:"\pol_coeff S c; pol_coeff S d\ \ \j \ (max (fst c) (fst d)). snd (add_cf S c d) j = snd (add_cf S d c) j" apply (cut_tac subring, frule subring_Ring, frule Ring.ring_is_ag[of S]) apply (simp add: add_cf_def max.absorb1 max.absorb2) apply (case_tac "(fst c) = (fst d)", simp add: pol_coeff_def) apply (rule allI, rule impI, rule aGroup.ag_pOp_commute[of S], simp+) apply (case_tac "(fst d) < (fst c)", simp, rule allI, rule impI, rule aGroup.ag_pOp_commute, assumption+) apply (frule_tac x = j and y = "fst d" and z = "fst c" in le_less_trans, assumption+, frule_tac x = j and y = "fst c" in less_imp_le, thin_tac "j < fst c", simp add:pol_coeff_mem, simp add:pol_coeff_mem) apply simp apply (frule leI[of "fst d" "fst c"], frule noteq_le_less[of "fst c" "fst d"], assumption, rule allI, rule impI, simp) apply (rule aGroup.ag_pOp_commute, assumption+, simp add:pol_coeff_mem, frule_tac x = j and y = "fst c" and z = "fst d" in le_less_trans, assumption+, frule_tac x = j and y = "fst d" in less_imp_le, thin_tac "j < fst d", simp add:pol_coeff_mem) done lemma (in PolynRg) polyn_addTr1:"pol_coeff S (n, f) \ \g. pol_coeff S (n + m, g) \ (polyn_expr R X n (n, f) \ (polyn_expr R X (n + m) ((n + m), g)) = polyn_expr R X (n + m) (add_cf S (n, f) ((n + m), g)))" apply (cut_tac subring, frule subring_Ring) apply (induct_tac m) apply (rule allI, rule impI, simp) apply (simp add:polyn_add_n1) apply (simp add:add.commute[of n]) apply (rule allI, rule impI) apply (frule_tac n = "na + n" and f = g in pol_coeff_pre) apply (drule_tac a = g in forall_spec, assumption) apply (cut_tac n = "na + n" and c = "(Suc (na + n), g)" in polyn_Suc, simp, simp del:npow_suc, thin_tac "polyn_expr R X (Suc (na + n)) (Suc (na + n), g) = polyn_expr R X (na + n) (Suc (na + n), g) \ g (Suc (na + n)) \\<^sub>r X^\<^bsup>R (Suc (na + n))\<^esup>") apply (frule_tac c = "(n, f)" and k = n in polyn_mem, simp, frule_tac c = "(Suc (na + n), g)" and k = "na + n" in polyn_mem, simp, frule_tac c = "(Suc (na + n), g)" in monomial_mem) apply (drule_tac a = "Suc (na + n)" in forall_spec, simp del:npow_suc, cut_tac ring_is_ag, subst aGroup.ag_pOp_assoc[THEN sym], assumption+, simp del:npow_suc) apply (simp del:npow_suc add:polyn_expr_restrict) apply (frule_tac c = "(n, f)" and d = "(Suc (na + n), g)" in add_cf_pol_coeff, assumption+, frule_tac c = "(n, f)" and d = "(na + n, g)" in add_cf_pol_coeff, assumption+) apply (frule_tac c = "add_cf S (n, f) (Suc (na + n), g)" and n = "na + n" and m = "Suc (na + n)" in polyn_n_m, simp, subst add_cf_len, assumption+, simp) apply (cut_tac k = "Suc (na + n)" and f = "add_cf S (n, f) (Suc (na + n), g)" in polyn_expr_split) apply (frule_tac c = "(n, f)" and d = "(Suc (na + n), g)" in add_cf_len, assumption+, simp del: npow_suc add: max.absorb1 max.absorb2) apply (thin_tac "polyn_expr R X (Suc (na + n)) (Suc (na + n), snd (add_cf S (n, f) (Suc (na + n), g))) = polyn_expr R X (na + n) (na + n, snd (add_cf S (n, f) (Suc (na + n), g))) \ \\<^sub>f R (\j. snd (add_cf S (n, f) (Suc (na + n), g)) j \\<^sub>r X^\<^bsup>R j\<^esup>) (Suc (na + n)) (Suc (na + n))", thin_tac "polyn_expr R X (Suc (na + n)) (add_cf S (n, f) (Suc (na + n), g)) = polyn_expr R X (na + n) (na + n, snd (add_cf S (n, f) (Suc (na + n), g))) \ \\<^sub>f R (\j. snd (add_cf S (n, f) (Suc (na + n), g)) j \\<^sub>r X^\<^bsup>R j\<^esup>) (Suc (na + n)) (Suc (na + n))") apply (simp del:npow_suc add:fSum_def cmp_def slide_def) apply (cut_tac d = "(Suc (na + n), g)" in add_cf_val_hi[of "(n, f)"], simp, simp del:npow_suc, thin_tac "snd (add_cf S (n, f) (Suc (na + n), g)) (Suc (na + n)) = g (Suc (na + n))") apply (frule_tac c = "add_cf S (n, f) (Suc (na + n), g)" and k = "na + n" in polyn_mem, simp, frule_tac c = "add_cf S (n, f) (na + n, g)" and k = "na + n" in polyn_mem, simp ) apply (subst add_cf_len, assumption+, simp del:npow_suc) apply (frule_tac a = "polyn_expr R X (na + n) (add_cf S (n, f) (na + n, g))" and b = "polyn_expr R X (na + n) (add_cf S (n, f) (Suc (na + n), g))" and c = "g (Suc (na + n)) \\<^sub>r X^\<^bsup>R (Suc (na + n))\<^esup>" in aGroup.ag_pOp_add_r[of R], assumption+) apply (rule_tac c = "add_cf S (n, f) (na + n, g)" and d = "add_cf S (n, f) (Suc (na + n), g)" and k = "na + n" in polyn_exprs_eq, assumption+, simp, subst add_cf_len, assumption+) apply (simp) apply (rule allI, rule impI, (subst add_cf_def)+, simp, frule_tac m = na and g = g in polyn_expr_restrict1[of n f], assumption, simp del:npow_suc) done lemma (in PolynRg) polyn_add:"\pol_coeff S (n, f); pol_coeff S (m, g)\ \ polyn_expr R X n (n, f) \ (polyn_expr R X m (m, g)) = polyn_expr R X (max n m) (add_cf S (n, f) (m, g))" apply (cut_tac less_linear[of n m]) apply (erule disjE, frule polyn_addTr1[of n f "m - n"], drule_tac a = g in forall_spec, simp, simp add: max.absorb1 max.absorb2) apply (erule disjE, simp add:polyn_add_n1) apply (frule polyn_mem[of "(n, f)" n], simp, frule polyn_mem[of "(m, g)" m], simp) apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_commute) apply (frule polyn_addTr1[of m g "n - m"], drule_tac a = f in forall_spec, simp, simp, frule add_cf_commute[of "(m, g)" "(n, f)"], assumption+, simp add:max_def, frule add_cf_pol_coeff[of "(n, f)" "(m, g)"], assumption+, frule add_cf_pol_coeff[of "(m, g)" "(n, f)"], assumption+) apply (rule polyn_exprs_eq[of "add_cf S (m, g) (n, f)" "add_cf S (n, f) (m, g)" n], assumption+) apply (simp add:add_cf_len, simp) done lemma (in PolynRg) polyn_add1:"\pol_coeff S c; pol_coeff S d\ \ polyn_expr R X (fst c) c \ (polyn_expr R X (fst d) d) = polyn_expr R X (max (fst c) (fst d)) (add_cf S c d)" apply (cases c) apply (cases d) apply (simp add: polyn_add) done lemma (in PolynRg) polyn_minus_nsum:"\pol_coeff S c; k \ (fst c)\ \ -\<^sub>a (polyn_expr R X k c) = nsum R (\j. ((-\<^sub>a\<^bsub>S\<^esub> ((snd c) j)) \\<^sub>r (X^\<^bsup>R j\<^esup>))) k" apply (cut_tac subring, frule subring_Ring[of S], frule Ring.ring_is_ag[of S], cut_tac ring_is_ag, cut_tac X_mem_R) apply (simp add:polyn_expr_def, subst aGroup.nsum_minus[of R], assumption) apply (frule monomial_mem[of c], rule allI, rule impI, frule_tac i = j and j = k and k = "fst c" in le_trans, assumption+, simp) apply (rule aGroup.nsum_eq, assumption, rule allI, rule impI, simp, rule aGroup.ag_mOp_closed, assumption) apply (frule monomial_mem[of c], frule_tac i = j and j = k and k = "fst c" in le_trans, assumption+, simp) apply (rule allI, rule impI, rule ring_tOp_closed) apply (frule_tac j = j in pol_coeff_mem[of c]) apply (frule_tac i = j and j = k and k = "fst c" in le_trans, assumption+, simp add:Subring_minus_ring_minus, frule_tac x = "(snd c) j" in mem_subring_mem_ring[of S], assumption, simp add:aGroup.ag_mOp_closed, simp add:npClose) apply (rule allI, rule impI, simp, cut_tac j = j in pol_coeff_mem[of c], assumption, rule_tac i = j and j = k and k = "fst c" in le_trans, assumption+) apply (simp add:Subring_minus_ring_minus, frule_tac x = "(snd c) j" in mem_subring_mem_ring[of S], assumption) apply (subst ring_inv1_1, assumption+) apply (simp add:npClose, simp) done lemma (in PolynRg) minus_pol_coeff:"pol_coeff S c \ pol_coeff S ((fst c), (\j. (-\<^sub>a\<^bsub>S\<^esub> ((snd c) j))))" apply (simp add:pol_coeff_def) apply (rule allI, rule impI) apply (cut_tac subring, frule subring_Ring) apply (frule Ring.ring_is_ag[of "S"]) apply (rule aGroup.ag_mOp_closed, assumption) apply simp done lemma (in PolynRg) polyn_minus:"\pol_coeff S c; k \ (fst c)\ \ -\<^sub>a (polyn_expr R X k c) = polyn_expr R X k (fst c, (\j. (-\<^sub>a\<^bsub>S\<^esub> ((snd c) j))))" apply (cases c) apply (subst polyn_minus_nsum) apply (simp_all add: polyn_expr_def) done definition m_cf :: "[('a, 'm) Ring_scheme, nat \ (nat \ 'a)] \ nat \ (nat \ 'a)" where "m_cf S c = (fst c, (\j. (-\<^sub>a\<^bsub>S\<^esub> ((snd c) j))))" lemma (in PolynRg) m_cf_pol_coeff:"pol_coeff S c \ pol_coeff S (m_cf S c)" by (simp add:m_cf_def, simp add:minus_pol_coeff) lemma (in PolynRg) m_cf_len:"pol_coeff S c \ fst (m_cf S c) = fst c" by (simp add:m_cf_def) lemma (in PolynRg) polyn_minus_m_cf:"\pol_coeff S c; k \ (fst c)\ \ -\<^sub>a (polyn_expr R X k c) = polyn_expr R X k (m_cf S c)" by (simp add:m_cf_def polyn_minus) lemma (in PolynRg) polyn_zero_minus_zero:"\pol_coeff S c; k \ (fst c)\ \ (polyn_expr R X k c = \) = (polyn_expr R X k (m_cf S c) = \)" apply (cut_tac ring_is_ag) apply (simp add:polyn_minus_m_cf[THEN sym]) apply (rule iffI, simp) apply (simp add:aGroup.ag_inv_zero) apply (frule polyn_mem[of c k], assumption) apply (frule aGroup.ag_inv_inv[of "R" "polyn_expr R X k c"], assumption) apply (simp add:aGroup.ag_inv_zero) done lemma (in PolynRg) coeff_0_pol_0:"\pol_coeff S c; k \ fst c\ \ (\j\ k. (snd c) j = \\<^bsub>S\<^esub>) = (polyn_expr R X k c = \)" apply (rule iffI) apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring) apply (simp add:Subring_zero_ring_zero) apply (simp add:polyn_expr_def, rule aGroup.nsum_zeroA[of R], assumption) apply (rule allI, rule impI, cut_tac X_mem_R) apply (drule_tac a = j in forall_spec, simp, frule_tac n = j in npClose[of X], simp) apply (simp add:ring_times_0_x) apply (cases c) using algfree [simplified algfree_cond_def] by (auto simp add: polyn_expr_def) subsection \Multiplication of \pol_exprs\\ subsection "Multiplication" definition ext_cf :: "[('a, 'm) Ring_scheme, nat, nat \ (nat \ 'a)] \ nat \ (nat \ 'a)" where "ext_cf S n c = (n + fst c, \i. if n \ i then (snd c) (sliden n i) else \\<^bsub>S\<^esub>)" (* 0 0 g(0) g(m) 0 n m+n , where (m, g) is a pol_coeff **) definition sp_cf :: "[('a, 'm) Ring_scheme, 'a, nat \ (nat \ 'a)] \ nat \ (nat \ 'a)" where "sp_cf S a c = (fst c, \j. a \\<^sub>r\<^bsub>S\<^esub> ((snd c) j))" (* scalar times cf *) definition special_cf :: "('a, 'm) Ring_scheme \ nat \ (nat \ 'a)" ("C\<^sub>0") where "C\<^sub>0 S = (0, \j. 1\<^sub>r\<^bsub>S\<^esub>)" lemma (in PolynRg) special_cf_pol_coeff:"pol_coeff S (C\<^sub>0 S)" apply (cut_tac subring, frule subring_Ring) apply (simp add:pol_coeff_def special_cf_def) apply (simp add:Ring.ring_one) done lemma (in PolynRg) special_cf_len:"fst (C\<^sub>0 S) = 0" apply (simp add:special_cf_def) done lemma (in PolynRg) ext_cf_pol_coeff:"pol_coeff S c \ pol_coeff S (ext_cf S n c)" apply (simp add: pol_coeff_def ext_cf_def sliden_def) apply (rule impI) apply (rule Ring.ring_zero) apply (rule subring_Ring) apply (rule subring) done lemma (in PolynRg) ext_cf_len:"pol_coeff S c \ fst (ext_cf S m c) = m + fst c" by (simp add:ext_cf_def) lemma (in PolynRg) ext_special_cf_len:"fst (ext_cf S m (C\<^sub>0 S)) = m" apply (cut_tac special_cf_pol_coeff) apply (simp add:ext_cf_len special_cf_def) done lemma (in PolynRg) ext_cf_self:"pol_coeff S c \ \j \ (fst c). snd (ext_cf S 0 c) j = (snd c) j" apply (rule allI, rule impI, simp add:ext_cf_def sliden_def) done lemma (in PolynRg) ext_cf_hi:"pol_coeff S c \ (snd c) (fst c) = snd (ext_cf S n c) (n + (fst c))" apply (subst ext_cf_def) apply (simp add:sliden_def) done lemma (in PolynRg) ext_special_cf_hi:"snd (ext_cf S n (C\<^sub>0 S)) n = 1\<^sub>r\<^bsub>S\<^esub>" apply (cut_tac special_cf_pol_coeff) apply (cut_tac ext_cf_hi[of "C\<^sub>0 S" n, THEN sym]) apply (simp add:special_cf_def, assumption) done lemma (in PolynRg) ext_cf_lo_zero:"\pol_coeff S c; 0 < n; x \ (n - Suc 0)\ \ snd (ext_cf S n c) x = \\<^bsub>S\<^esub>" apply (cut_tac Suc_le_mono[THEN sym, of x "n - Suc 0"], simp, cut_tac x = x and y = "Suc x" and z = n in less_le_trans, simp, assumption, simp add:nat_not_le_less[THEN sym, of x n], thin_tac "x \ n - Suc 0") apply (simp add:ext_cf_def) done lemma (in PolynRg) ext_special_cf_lo_zero:"\0 < n; x \ (n - Suc 0)\ \ snd (ext_cf S n (C\<^sub>0 S)) x = \\<^bsub>S\<^esub>" by (cut_tac special_cf_pol_coeff, frule ext_cf_lo_zero[of "C\<^sub>0 S" n], assumption+) lemma (in PolynRg) sp_cf_pol_coeff:"\pol_coeff S c; a \ carrier S\ \ pol_coeff S (sp_cf S a c)" apply (cut_tac subring, frule subring_Ring) apply (simp add:pol_coeff_def sp_cf_def, rule allI, rule impI, rule Ring.ring_tOp_closed, assumption+) apply simp done lemma (in PolynRg) sp_cf_len:"\pol_coeff S c; a \ carrier S\ \ fst (sp_cf S a c) = fst c" by (simp add:sp_cf_def) lemma (in PolynRg) sp_cf_val:"\pol_coeff S c; j \ (fst c); a \ carrier S\ \ snd (sp_cf S a c) j = a \\<^sub>r\<^bsub>S\<^esub> ((snd c) j)" by (simp add:sp_cf_def) lemma (in PolynRg) polyn_ext_cf_lo_zero:"\pol_coeff S c; 0 < j\ \ polyn_expr R X (j - Suc 0) (ext_cf S j c) = \" apply (simp add:polyn_expr_def, cut_tac ring_is_ag, rule aGroup.nsum_zeroA, assumption) apply (rule allI, rule impI) apply (frule_tac x = ja in ext_cf_lo_zero [of c j], assumption+) apply (cut_tac X_mem_R, frule_tac n = ja in npClose[of X]) apply (cut_tac subring, simp add:Subring_zero_ring_zero, simp add:ring_times_0_x) done lemma (in PolynRg) monomial_d:"pol_coeff S c \ polyn_expr R X d (ext_cf S d c) = ((snd c) 0) \\<^sub>r X^\<^bsup>R d\<^esup>" apply (cut_tac ring_is_ag, cut_tac subring, cut_tac X_mem_R, frule subring_Ring[of S]) apply (frule pol_coeff_mem [of c 0], simp) apply (case_tac "d = 0") apply simp apply (simp add:polyn_expr_def ext_cf_def sliden_def) apply (frule ext_cf_pol_coeff[of c d]) apply (cut_tac polyn_Suc[of "d - Suc 0" "ext_cf S d c"]) apply (simp) apply (cut_tac polyn_ext_cf_lo_zero[of c d], simp, thin_tac "polyn_expr R X (d - Suc 0) (ext_cf S d c) = \") apply (frule monomial_mem[of "ext_cf S d c"], simp add:ext_cf_len, drule_tac a = d in forall_spec, simp, simp add:aGroup.ag_l_zero) apply (subst polyn_expr_short[of "ext_cf S d c" d], assumption, simp add:ext_cf_len) apply (simp, subst ext_cf_def, simp add:sliden_def , assumption+, simp add:ext_cf_len) done lemma (in PolynRg) X_to_d:" X^\<^bsup>R d\<^esup> = polyn_expr R X d (ext_cf S d (C\<^sub>0 S))" apply (cut_tac special_cf_pol_coeff) apply (subst monomial_d[of "C\<^sub>0 S" d], assumption+) apply (subst special_cf_def, simp) apply (cut_tac subring, frule subring_Ring) apply (simp add:Subring_one_ring_one) apply (cut_tac X_mem_R, frule_tac n = d in npClose[of X]) apply (simp add:ring_l_one) done lemma (in PolynRg) c_max_ext_special_cf:"c_max S (ext_cf S n (C\<^sub>0 S)) = n" apply (cut_tac polyn_ring_S_nonzero, cut_tac subring, frule subring_Ring) apply (simp add:c_max_def special_cf_def ext_cf_def) apply (cut_tac n_max[of "{j. (n \ j \ j = n) \ n \ j}" n]) apply (erule conjE)+ apply simp apply (rule subsetI, simp, erule conjE, simp) apply (cut_tac le_refl[of n], blast) done lemma (in PolynRg) scalar_times_polynTr:"a \ carrier S \ \f. pol_coeff S (n, f) \ a \\<^sub>r (polyn_expr R X n (n, f)) = polyn_expr R X n (sp_cf S a (n, f))" apply (cut_tac subring, cut_tac X_mem_R, frule_tac x = a in mem_subring_mem_ring, assumption) apply (induct_tac n, rule allI, rule impI, simp add:polyn_expr_def sp_cf_def, cut_tac n_in_Nsetn[of "0"]) apply (cut_tac subring_Ring, frule_tac c = "(0, f)" in pol_coeff_mem[of _ "0"], simp) apply (simp, frule_tac x = "f 0" in mem_subring_mem_ring, assumption) apply ( simp add:Subring_tOp_ring_tOp, frule_tac y = "f 0" in ring_tOp_closed[of a], assumption+, cut_tac ring_one, simp add:ring_tOp_assoc, assumption) apply (rule allI, rule impI, frule subring_Ring, frule_tac n = n and f = f in pol_coeff_pre, drule_tac x = f in spec, simp) apply (cut_tac n = n and c = "(Suc n, f)" in polyn_Suc, simp, simp del:npow_suc, thin_tac "polyn_expr R X (Suc n) (Suc n, f) = polyn_expr R X n (Suc n, f) \ f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") apply (cut_tac n = n and c = "sp_cf S a (Suc n, f)" in polyn_Suc, simp add:sp_cf_len) apply (frule_tac c = "(Suc n, f)" and a = a in sp_cf_len, assumption+, simp only:fst_conv) apply (cut_tac k = "Suc n" and f = "sp_cf S a (Suc n, f)" in polyn_expr_split, simp del:npow_suc, thin_tac "polyn_expr R X (Suc n) (Suc n, snd (sp_cf S a (Suc n, f))) = polyn_expr R X n (sp_cf S a (Suc n, f)) \ snd (sp_cf S a (Suc n, f)) (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>", thin_tac "polyn_expr R X (Suc n) (sp_cf S a (Suc n, f)) = polyn_expr R X n (sp_cf S a (Suc n, f)) \ snd (sp_cf S a (Suc n, f)) (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") apply (frule_tac c = "(Suc n, f)" and a = a in sp_cf_pol_coeff, assumption) apply (frule_tac c = "(Suc n, f)" and k = n in polyn_mem, simp, frule_tac c = "(Suc n, f)" in monomial_mem, drule_tac a = "Suc n" in forall_spec, simp, simp only:snd_conv) apply (subst ring_distrib1, assumption+, subst polyn_expr_restrict, assumption+, simp del:npow_suc, subst sp_cf_val, assumption, simp, assumption, simp only:snd_conv, frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp, simp only:snd_conv, simp del:npow_suc add:Subring_tOp_ring_tOp, subst ring_tOp_assoc[THEN sym, of a], assumption+, simp add:mem_subring_mem_ring, rule npClose, assumption) apply (cut_tac ring_is_ag, rule aGroup.ag_pOp_add_r, assumption+, rule polyn_mem, rule sp_cf_pol_coeff, assumption+, simp add:sp_cf_len, rule polyn_mem, assumption, simp add:sp_cf_len, frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem_R, simp, simp only:snd_conv, (rule ring_tOp_closed)+, assumption+, rule npClose, assumption) apply (rule_tac c = "sp_cf S a (n, f)" and d = "sp_cf S a (Suc n, f)" and k = n in polyn_exprs_eq, rule sp_cf_pol_coeff, assumption+, simp add:sp_cf_len) apply (rule allI, rule impI, (subst sp_cf_def)+, simp) done lemma (in PolynRg) scalar_times_pol_expr:"\a \ carrier S; pol_coeff S c; n \ fst c\ \ a \\<^sub>r (polyn_expr R X n c) = polyn_expr R X n (sp_cf S a c)" apply (cases c) apply (simp only:) apply (rename_tac m g) apply (thin_tac "c = (m, g)") apply (frule_tac c = "(m, g)" and k = n in polyn_expr_short, simp, simp) apply (frule scalar_times_polynTr[of a n], drule_tac x = g in spec) apply (frule_tac c = "(m, g)" and n = n in pol_coeff_le, simp, simp, thin_tac "polyn_expr R X n (m, g) = polyn_expr R X n (n, g)", thin_tac "a \\<^sub>r polyn_expr R X n (n, g) = polyn_expr R X n (sp_cf S a (n, g))") apply (frule_tac c = "(m, g)" and n = n in pol_coeff_le, simp, simp, frule_tac c = "(n, g)" and a = a in sp_cf_pol_coeff, assumption, frule_tac c = "(m, g)" and a = a in sp_cf_pol_coeff, assumption) apply (rule_tac c = "sp_cf S a (n, g)" and d = "sp_cf S a (m, g)" and k = n in polyn_exprs_eq, assumption+) apply (simp add:sp_cf_len) apply (rule allI, (subst sp_cf_def)+, simp) done lemma (in PolynRg) sp_coeff_nonzero:"\Idomain S; a \ carrier S; a \ \\<^bsub>S\<^esub>; pol_coeff S c; (snd c) j \ \\<^bsub>S\<^esub>; j \ (fst c)\ \ snd (sp_cf S a c) j \ \\<^bsub>S\<^esub>" apply (simp add:sp_cf_def) apply (frule_tac y = "(snd c) j" in Idomain.idom_tOp_nonzeros[of S a], assumption+, simp add:pol_coeff_def, simp add:Pi_def, assumption+) done lemma (in PolynRg) ext_cf_inductTl:"pol_coeff S (Suc n, f) \ polyn_expr R X (n + j) (ext_cf S j (Suc n, f)) = polyn_expr R X (n + j) (ext_cf S j (n, f))" apply (frule pol_coeff_pre[of n f], frule ext_cf_pol_coeff[of "(Suc n, f)" j], frule ext_cf_pol_coeff[of "(n, f)" j], rule polyn_exprs_eq[of "ext_cf S j (Suc n, f)" "ext_cf S j (n, f)" "n + j"], assumption+) apply (simp add:ext_cf_len) apply (rule allI, (subst ext_cf_def)+, simp add:sliden_def) done lemma (in PolynRg) low_deg_terms_zeroTr:" pol_coeff S (n, f) \ polyn_expr R X (n + j) (ext_cf S j (n, f)) = (X^\<^bsup>R j\<^esup>) \\<^sub>r (polyn_expr R X n (n, f))" apply (cut_tac ring_is_ag, cut_tac X_mem_R, frule npClose[of "X" "j"]) apply (induct_tac n) apply (rule impI, simp) apply (case_tac "j = 0", simp add:ext_cf_def sliden_def polyn_expr_def) apply (frule_tac c = "(0, f)" and j = 0 in pol_coeff_mem_R, simp, simp) apply (simp add:ring_r_one ring_l_one) apply (cut_tac polyn_Suc[of "j - Suc 0" "ext_cf S j (0, f)"], simp del:npow_suc) apply (frule ext_cf_len[of "(0, f)" j], cut_tac polyn_expr_split[of j "ext_cf S j (0, f)"], simp, thin_tac "polyn_expr R X j (ext_cf S j (0, f)) = polyn_expr R X (j - Suc 0) (ext_cf S j (0, f)) \ snd (ext_cf S j (0, f)) j \\<^sub>r X^\<^bsup>R j\<^esup>") apply (simp add:polyn_ext_cf_lo_zero[of "(0, f)" j], thin_tac "polyn_expr R X j (j, snd (ext_cf S j (0, f))) = \ \ snd (ext_cf S j (0, f)) j \\<^sub>r X^\<^bsup>R j\<^esup>", frule ext_cf_hi[THEN sym, of "(0, f)" j], simp add:polyn_expr_def) apply (frule_tac c = "(0, f)" and j = 0 in pol_coeff_mem_R, simp, simp) apply (subst aGroup.ag_l_zero, assumption, simp add:ring_tOp_closed, simp add:ring_r_one, subst ring_tOp_commute, assumption+, simp) apply (simp add:ext_cf_len) apply (rule impI, cut_tac subring, cut_tac subring_Ring[of S], frule_tac n = n in pol_coeff_pre[of _ "f"]) apply simp apply (subst polyn_expr_split) apply (cut_tac n = "n + j" and c = "ext_cf S j (Suc n, f)" in polyn_Suc, simp add:ext_cf_len) apply (subst ext_cf_len, assumption+, simp del:npow_suc add:add.commute[of j], thin_tac "polyn_expr R X (Suc (n + j)) (Suc (n + j), snd (ext_cf S j (Suc n, f))) = polyn_expr R X (n + j) (ext_cf S j (Suc n, f)) \ snd (ext_cf S j (Suc n, f)) (Suc (n + j)) \\<^sub>r X^\<^bsup>R (Suc (n + j))\<^esup>", subst ext_cf_inductTl, assumption+, simp del:npow_suc, thin_tac "polyn_expr R X (n + j) (ext_cf S j (n, f)) = X^\<^bsup>R j\<^esup> \\<^sub>r polyn_expr R X n (n, f)") apply (cut_tac c1 = "(Suc n, f)" and n1 = j in ext_cf_hi[THEN sym], assumption+, simp del:npow_suc add:add.commute[of j]) apply (cut_tac n = n and c = "(Suc n, f)" in polyn_Suc, simp, simp del:npow_suc) apply (frule_tac c = "(Suc n, f)" and k = n in polyn_mem, simp, frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem_R, simp, simp del:npow_suc, frule_tac x = "f (Suc n)" and y = "X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed, rule npClose, assumption, subst ring_distrib1, assumption+) apply (subst polyn_expr_restrict, assumption+) apply (rule_tac a = "f (Suc n) \\<^sub>r X^\<^bsup>R (Suc (n + j))\<^esup> " and b = "X^\<^bsup>R j\<^esup> \\<^sub>r (f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>)" and c = "X^\<^bsup>R j\<^esup> \\<^sub>r polyn_expr R X n (n, f)" in aGroup.ag_pOp_add_l, assumption+, rule ring_tOp_closed, assumption+, rule npClose, assumption, (rule ring_tOp_closed, assumption+)+, simp add:polyn_mem, frule_tac n = "Suc n" in npClose[of X], subst ring_tOp_assoc[THEN sym], assumption+, subst ring_tOp_commute[of "X^\<^bsup>R j\<^esup>"], assumption, simp add:pol_coeff_mem, subst ring_tOp_assoc, assumption+, subst npMulDistr[of X], assumption, simp add:add.commute[of j]) apply simp done lemma (in PolynRg) low_deg_terms_zero:"pol_coeff S (n, f) \ polyn_expr R X (n + j) (ext_cf S j (n, f)) = (X^\<^bsup>R j\<^esup>) \\<^sub>r (polyn_expr R X n (n, f))" by (simp add:low_deg_terms_zeroTr) lemma (in PolynRg) low_deg_terms_zero1:"pol_coeff S c \ polyn_expr R X ((fst c) + j) (ext_cf S j c) = (X^\<^bsup>R j\<^esup>) \\<^sub>r (polyn_expr R X (fst c) c)" by (cases c) (simp add: low_deg_terms_zeroTr) lemma (in PolynRg) polyn_expr_tOpTr:"pol_coeff S (n, f) \ \g. (pol_coeff S (m, g) \ (\h. pol_coeff S ((n + m), h) \ h (n + m) = (f n) \\<^sub>r\<^bsub>S\<^esub> (g m) \ (polyn_expr R X (n + m) (n + m, h) = (polyn_expr R X n (n, f)) \\<^sub>r (polyn_expr R X m (m, g)))))" apply (cut_tac subring, cut_tac X_mem_R, frule subring_Ring[of S]) apply (induct_tac m) apply (rule allI, rule impI, simp) apply (simp add:polyn_expr_def [of R X 0]) apply (frule_tac c = "(0,g)" in pol_coeff_mem[of _ 0], simp, simp, frule_tac c = "(0,g)" in pol_coeff_mem_R[of _ 0], simp, simp) apply (simp add:ring_r_one, frule_tac c = "(n, f)" and k = n in polyn_mem, simp, simp only:ring_tOp_commute[of "polyn_expr R X n (n, f)"], subst scalar_times_pol_expr, assumption+, simp) apply (cut_tac f = "sp_cf S (g 0) (n, f)" in pol_coeff_split) apply (simp add:sp_cf_len) apply (cut_tac f = "sp_cf S (g 0) (n, f)" in polyn_expr_split[of n], simp only:sp_cf_len, simp only:fst_conv, frule_tac a = "g 0" in sp_cf_pol_coeff[of "(n, f)"], assumption+, simp, subgoal_tac "snd (sp_cf S (g 0) (n, f)) n = (f n) \\<^sub>r\<^bsub>S\<^esub> (g 0)", blast) apply (thin_tac "pol_coeff S (n, snd (sp_cf S (g 0) (n, f)))", thin_tac "polyn_expr R X n (sp_cf S (g 0) (n, f)) = polyn_expr R X n (n, snd (sp_cf S (g 0) (n, f)))", thin_tac "pol_coeff S (sp_cf S (g 0) (n, f))") apply (subst sp_cf_val[of "(n, f)" n], assumption+, simp, assumption, simp, frule_tac c = "(n,f)" in pol_coeff_mem[of _ n], simp, simp, simp add:Ring.ring_tOp_commute) apply (rule allI, rule impI) apply (frule_tac n = na and f = g in pol_coeff_pre, drule_tac a = g in forall_spec, assumption+) apply (erule exE, (erule conjE)+) apply (cut_tac n = na and c = "(Suc na, g)" in polyn_Suc, (simp del:npow_suc)+, thin_tac "polyn_expr R X (Suc na) (Suc na, g) = polyn_expr R X na (Suc na, g) \ g (Suc na) \\<^sub>r X^\<^bsup>R (Suc na)\<^esup>", subst polyn_expr_restrict, assumption) apply (frule_tac c = "(n, f)" and k = n in polyn_mem,simp del:npow_suc, frule_tac c = "(na, g)" and k = na in polyn_mem, simp del:npow_suc, frule_tac c = "(Suc na, g)" in monomial_mem, simp del:npow_suc, drule_tac a = "Suc na" in forall_spec, simp del:npow_suc) apply (subst ring_distrib1, assumption+) apply (rotate_tac 8, drule sym, simp del:npow_suc) apply (thin_tac "polyn_expr R X n (n, f) \\<^sub>r polyn_expr R X na (na, g) = polyn_expr R X (n + na) (n + na, h)") apply (frule_tac c = "(Suc na, g)" and j ="Suc na" in pol_coeff_mem_R, simp, simp del:npow_suc, frule_tac c = "(Suc na, g)" and j ="Suc na" in pol_coeff_mem, simp, simp del:npow_suc) apply (subst ring_tOp_commute, assumption+, subst ring_tOp_assoc, assumption+, rule npClose, assumption+, subst low_deg_terms_zero[THEN sym], assumption+) apply (frule_tac c = "(n, f)" and n = "Suc na" in ext_cf_pol_coeff) apply (frule_tac c = "ext_cf S (Suc na) (n, f)" and a = "g (Suc na)" in sp_cf_pol_coeff, assumption) apply (subst scalar_times_pol_expr, assumption+, simp add:ext_cf_len, cut_tac k = "n + Suc na" and f = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in polyn_expr_split, simp only:sp_cf_len, thin_tac "polyn_expr R X (n + Suc na) (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))) = polyn_expr R X (n + Suc na) (fst (ext_cf S (Suc na) (n, f)), snd (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))))", simp only:ext_cf_len, simp only:fst_conv, simp add:add.commute[of _ n]) apply (subst polyn_add, assumption+, cut_tac f = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in pol_coeff_split, simp only:sp_cf_len, simp only:ext_cf_len, simp add:add.commute[of _ n], simp add: max_def, frule_tac c = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in pol_coeff_cartesian, simp only:sp_cf_len, simp only:ext_cf_len, simp add:add.commute[of _ n], thin_tac "(Suc (n + na), snd (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))) = sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))", frule_tac c = "(n + na, h)" and d = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in add_cf_pol_coeff, assumption) -apply (cut_tac k = "Suc (n + na)" and f = "add_cf S (n + na, h) - (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))" in polyn_expr_split, - simp only:mp, + apply (cut_tac k = "Suc (n + na)" and f = "add_cf S (n + na, h) + (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))" in polyn_expr_split) + apply (simp only: mp) + + apply ( thin_tac "polyn_expr R X (Suc (n + na)) (add_cf S (n + na, h) (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))) = polyn_expr R X (Suc (n + na)) (fst (add_cf S (n + na, h) (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))), snd (add_cf S (n + na, h) - (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))))", - subst add_cf_len, assumption+, - simp add:sp_cf_len, simp add:ext_cf_len max_def, - cut_tac f = "add_cf S (n + na, h) + (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))))") + apply (subst add_cf_len) + apply assumption+ + apply (simp add: sp_cf_len) + apply (simp add: ext_cf_len max_def) + apply (cut_tac f = "add_cf S (n + na, h) (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))" in - pol_coeff_split, - simp only:add_cf_len, - simp only:sp_cf_len, simp add:ext_cf_len, simp add:max_def, - thin_tac "pol_coeff S + pol_coeff_split) + apply (simp only: add_cf_len) + apply (simp only: sp_cf_len) + apply (simp add: ext_cf_len) + apply (thin_tac "pol_coeff S (add_cf S (n + na, h) (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))))", subgoal_tac "snd (add_cf S (n + na, h) (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))) (Suc (n + na)) = f n \\<^sub>r\<^bsub>S\<^esub> g (Suc na)", simp add:add.commute[of _ n], blast) apply (subst add_cf_def, simp add:sp_cf_len ext_cf_len, subst sp_cf_def, simp add:ext_cf_len, subst ext_cf_def, simp add:sliden_def, frule pol_coeff_mem[of "(n, f)" n], simp, simp add:Ring.ring_tOp_commute) done lemma (in PolynRg) polyn_expr_tOp:"\ pol_coeff S (n, f); pol_coeff S (m, g)\ \ \e. pol_coeff S ((n + m), e) \ e (n + m) = (f n) \\<^sub>r\<^bsub>S\<^esub> (g m) \ polyn_expr R X (n + m)(n + m, e) = (polyn_expr R X n (n, f)) \\<^sub>r (polyn_expr R X m (m, g))" by (simp add:polyn_expr_tOpTr) lemma (in PolynRg) polyn_expr_tOp_c:"\pol_coeff S c; pol_coeff S d\ \ \e. pol_coeff S e \ (fst e = fst c + fst d) \ (snd e) (fst e) = (snd c (fst c)) \\<^sub>r\<^bsub>S\<^esub> (snd d) (fst d) \ polyn_expr R X (fst e) e = (polyn_expr R X (fst c) c) \\<^sub>r (polyn_expr R X (fst d) d)" by (cases c, cases d) (simp add: polyn_expr_tOpTr) section "The degree of a polynomial" lemma (in PolynRg) polyn_degreeTr:"\pol_coeff S c; k \ (fst c)\ \ (polyn_expr R X k c = \ ) = ({j. j \ k \ (snd c) j \ \\<^bsub>S\<^esub>} = {})" apply (subst coeff_0_pol_0[THEN sym, of c k], assumption+) apply blast done lemma (in PolynRg) higher_part_zero:"\pol_coeff S c; k < fst c; \j\nset (Suc k) (fst c). snd c j = \\<^bsub>S\<^esub>\ \ \\<^sub>f R (\j. snd c j \\<^sub>r X^\<^bsup>R j\<^esup>) (Suc k) (fst c) = \" apply (cut_tac ring_is_ag, rule aGroup.fSum_zero1[of R k "fst c" "\j. snd c j \\<^sub>r X^\<^bsup>R j\<^esup>"], assumption+) apply (rule ballI, drule_tac x = j in bspec, assumption, simp) apply (cut_tac subring, simp add:Subring_zero_ring_zero, cut_tac X_mem_R, frule_tac n = j in npClose[of X], simp add:ring_times_0_x) done lemma (in PolynRg) coeff_nonzero_polyn_nonzero:"\pol_coeff S c; k \ (fst c)\ \ (polyn_expr R X k c \ \) = (\j\k. (snd c) j \ \\<^bsub>S\<^esub> )" by (simp add:coeff_0_pol_0[THEN sym, of c k]) lemma (in PolynRg) pol_expr_unique:"\p \ carrier R; p \ \; pol_coeff S c; p = polyn_expr R X (fst c) c; (snd c) (fst c) \ \\<^bsub>S\<^esub>; pol_coeff S d; p = polyn_expr R X (fst d) d; (snd d) (fst d) \ \\<^bsub>S\<^esub>\ \ (fst c) = (fst d) \ (\j \ (fst c). (snd c) j = (snd d) j)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule Ring.ring_is_ag[of S]) apply (frule m_cf_pol_coeff[of d]) apply (frule polyn_minus_m_cf[of d "fst d"], simp) apply (drule sym, drule sym, simp) apply (rotate_tac -2, drule sym, drule sym) apply (frule_tac x = p in aGroup.ag_r_inv1[of R], assumption, simp, thin_tac "p = polyn_expr R X (fst c) c", thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) c", thin_tac "-\<^sub>a (polyn_expr R X (fst c) c) = polyn_expr R X (fst d) (m_cf S d)") apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp add:m_cf_len, thin_tac "polyn_expr R X (fst c) c \ polyn_expr R X (fst d) (m_cf S d) = polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d))", thin_tac "polyn_expr R X (fst c) c \ polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d))") apply (frule_tac add_cf_pol_coeff[of c "m_cf S d"], simp add:m_cf_len) apply (cut_tac coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)" "max (fst c) (fst d)"], drule sym, simp, thin_tac "polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d)) = \", thin_tac "pol_coeff S (add_cf S c (m_cf S d))", thin_tac "pol_coeff S (m_cf S d)") apply (case_tac "fst c = fst d", simp) apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption) apply (simp add:add_cf_def m_cf_def m_cf_len) apply (frule_tac j = j in pol_coeff_mem[of c], simp, frule_tac j = j in pol_coeff_mem[of d], simp) apply (subst aGroup.ag_eq_diffzero[of S], assumption+) apply (simp add:add_cf_def) apply (case_tac "\ (fst c) \ (fst d)", simp) apply (simp add:m_cf_len) apply (drule_tac a = "fst c" in forall_spec, simp, simp) apply simp apply (drule_tac a = "fst d" in forall_spec, simp, simp add:m_cf_len) apply (case_tac "fst c \ fst d", frule noteq_le_less[of "fst c" "fst d"], assumption, simp) apply (simp add:m_cf_def) apply (frule pol_coeff_mem[of d "fst d"], simp) apply (frule Ring.ring_is_ag[of S], frule aGroup.ag_inv_inv[of S "snd d (fst d)"], assumption) apply (simp add:aGroup.ag_inv_zero) apply simp apply simp apply (simp add:add_cf_len m_cf_len) done lemma (in PolynRg) pol_expr_unique2:"\pol_coeff S c; pol_coeff S d; fst c = fst d\ \ (polyn_expr R X (fst c) c = polyn_expr R X (fst d) d ) = (\j \ (fst c). (snd c) j = (snd d) j)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule Ring.ring_is_ag[of S]) apply (rule iffI) apply (frule m_cf_pol_coeff[of d]) apply (frule polyn_mem[of c "fst c"], simp, frule polyn_mem[of d "fst d"], simp) apply (frule aGroup.ag_eq_diffzero[of R "polyn_expr R X (fst c) c" "polyn_expr R X (fst d) d"], assumption+, simp, simp only:polyn_minus_m_cf[of d "fst d"], drule sym, simp) apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp add:m_cf_len) apply (thin_tac "polyn_expr R X (fst c) d \ polyn_expr R X (fst c) (m_cf S d) = polyn_expr R X (fst c) (add_cf S c (m_cf S d))", thin_tac "polyn_expr R X (fst c) c = polyn_expr R X (fst c) d", thin_tac "polyn_expr R X (fst c) d \ carrier R", drule sym) apply (frule_tac add_cf_pol_coeff[of c "m_cf S d"], simp add:m_cf_len) apply (frule coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)" "fst c"], simp add:add_cf_len, simp add:m_cf_len, thin_tac "\ = polyn_expr R X (fst d) (add_cf S c (m_cf S d))", thin_tac "pol_coeff S (add_cf S c (m_cf S d))") apply (simp add:add_cf_def m_cf_def) apply (rule allI, rule impI) apply (drule_tac a = j in forall_spec, assumption) apply (frule_tac j = j in pol_coeff_mem[of c], simp, frule_tac j = j in pol_coeff_mem[of d], simp) apply (simp add:aGroup.ag_eq_diffzero[THEN sym]) apply simp apply (rule polyn_exprs_eq[of c d "fst d"], assumption+) apply (simp, assumption+) done lemma (in PolynRg) pol_expr_unique3:"\pol_coeff S c; pol_coeff S d; fst c < fst d\ \ (polyn_expr R X (fst c) c = polyn_expr R X (fst d) d ) = ((\j \ (fst c). (snd c) j = (snd d) j) \ (\j\nset (Suc (fst c)) (fst d). (snd d) j = \\<^bsub>S\<^esub>))" apply (rule iffI) apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule Ring.ring_is_ag[of S]) apply (frule m_cf_pol_coeff[of d]) apply (frule polyn_mem[of c "fst c"], simp, frule polyn_mem[of d "fst d"], simp) apply (frule aGroup.ag_eq_diffzero[of R "polyn_expr R X (fst c) c" "polyn_expr R X (fst d) d"], assumption+, simp, simp only:polyn_minus_m_cf[of d "fst d"], drule sym, simp) - apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp add:m_cf_len, - thin_tac "polyn_expr R X (fst c) c \ polyn_expr R X (fst d) - (m_cf S d) = - polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d))", - thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) c", - thin_tac "polyn_expr R X (fst c) c \ carrier R", drule sym) + apply (frule polyn_add1 [of c "m_cf S d"]) + apply assumption+ + apply (simp add: m_cf_len) + apply (thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) c") + apply (thin_tac "polyn_expr R X (fst c) c \ carrier R", drule sym) apply (frule_tac add_cf_pol_coeff[of c "m_cf S d"], simp add:m_cf_len) apply (frule coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)" - "max (fst c) (fst d)"], - simp add:add_cf_len m_cf_len, simp, - thin_tac "polyn_expr R X (max (fst c) (fst d)) - (add_cf S c (m_cf S d)) = \", - thin_tac "pol_coeff S (add_cf S c (m_cf S d))") + "max (fst c) (fst d)"]) + apply (simp add:add_cf_len m_cf_len, simp) + apply (thin_tac "pol_coeff S (add_cf S c (m_cf S d))") apply (simp add:add_cf_def m_cf_def max_def) apply (rule conjI) apply (rule allI, rule impI, frule_tac x = j and y = "fst c" and z = "fst d" in le_less_trans, assumption+, frule_tac x = j and y = "fst d" in less_imp_le) apply (drule_tac a = j in forall_spec, simp, simp) apply (frule_tac j = j in pol_coeff_mem[of c], simp, frule_tac j = j in pol_coeff_mem[of d], simp) apply (simp add:aGroup.ag_eq_diffzero[THEN sym]) apply (rule ballI, simp add:nset_def, erule conjE) apply (cut_tac x = "fst c" and y = "Suc (fst c)" and z = j in less_le_trans, simp, assumption) apply (cut_tac m1 = "fst c" and n1 = j in nat_not_le_less[THEN sym], simp) apply (drule_tac a = j in forall_spec, assumption, simp, frule_tac j = j in pol_coeff_mem[of d], simp) apply (frule_tac x = "snd d j" in aGroup.ag_inv_inv[of S], assumption, simp add:aGroup.ag_inv_inv aGroup.ag_inv_zero) apply (cut_tac polyn_n_m[of d "fst c" "fst d"]) apply (subst polyn_expr_split[of "fst d" d], simp, thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) (fst c, snd d) \ \\<^sub>f R (\j. snd d j \\<^sub>r X^\<^bsup>R j\<^esup>) (Suc (fst c)) (fst d)", erule conjE) apply (subst higher_part_zero[of d "fst c"], assumption+) apply (frule pol_coeff_le[of d "fst c"], simp add:less_imp_le, frule polyn_mem[of "(fst c, snd d)" "fst c"], simp, cut_tac ring_is_ag, simp add:aGroup.ag_r_zero, subst polyn_expr_short[THEN sym, of d "fst c"], assumption+, simp add:less_imp_le) apply (rule polyn_exprs_eq[of c d "fst c"], assumption+) apply (simp, assumption+) apply (simp add:less_imp_le) done lemma (in PolynRg) polyn_degree_unique:"\pol_coeff S c; pol_coeff S d; polyn_expr R X (fst c) c = polyn_expr R X (fst d) d\ \ c_max S c = c_max S d" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule Ring.ring_is_ag[of S]) apply (case_tac "polyn_expr R X (fst d) d = \\<^bsub>R\<^esub>") apply (cut_tac coeff_0_pol_0[THEN sym, of d "fst d"], simp, cut_tac coeff_0_pol_0[THEN sym, of c "fst c"], simp) apply (simp add:c_max_def, assumption, simp, assumption, simp) apply (frule polyn_mem[of c "fst c"], simp, frule polyn_mem[of d "fst d"], simp) apply (frule aGroup.ag_eq_diffzero[of "R" "polyn_expr R X (fst c) c" "polyn_expr R X (fst d) d"], assumption+) apply (simp only:polyn_minus_m_cf[of d "fst d"], frule m_cf_pol_coeff [of d]) apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp only:m_cf_len) apply (rotate_tac -1, drule sym, simp, thin_tac "polyn_expr R X (fst d) d \ polyn_expr R X (fst d) (m_cf S d) = \", frule add_cf_pol_coeff[of c "m_cf S d"], assumption+) apply (cut_tac coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)" "fst (add_cf S c (m_cf S d))"], simp add:add_cf_len m_cf_len, thin_tac "polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d)) = \", thin_tac "pol_coeff S (add_cf S c (m_cf S d))", thin_tac "pol_coeff S (m_cf S d)") apply (frule coeff_nonzero_polyn_nonzero[of d "fst d"], simp, simp) apply (drule sym, simp) apply (frule coeff_nonzero_polyn_nonzero[of c "fst c"], simp, simp) apply (simp add:c_max_def, rule conjI, rule impI, blast, rule conjI, rule impI, blast) apply (rule n_max_eq_sets) apply (rule equalityI) apply (rule subsetI, simp) apply (erule conjE) apply (case_tac "fst c \ fst d") - apply (frule_tac i = x in le_trans[of _ "fst c" "fst d"], assumption+, simp, - rule contrapos_pp, simp+, simp add:max_def, + apply (frule_tac i = x in le_trans[of _ "fst c" "fst d"], assumption+, simp) +apply ( + rule contrapos_pp, simp+, frule_tac i = x in le_trans[of _ "fst c" "fst d"], assumption+, drule_tac a = x in forall_spec, assumption, drule le_imp_less_or_eq[of "fst c" "fst d"], erule disjE, simp add:add_cf_def m_cf_len m_cf_def, frule_tac j = x in pol_coeff_mem[of c], assumption+, simp add:aGroup.ag_inv_zero aGroup.ag_r_zero[of S]) apply (simp add:add_cf_def m_cf_len m_cf_def, rotate_tac -1, drule sym, simp, frule_tac j = x in pol_coeff_mem[of c], simp, simp add:aGroup.ag_inv_zero aGroup.ag_r_zero[of S]) apply (simp add:nat_not_le_less) (* \ fst c \ fst d *) apply (case_tac "\ x \ (fst d)", simp, simp add:nat_not_le_less, frule_tac x = "fst d" and y = x and z = "fst c" in less_le_trans, assumption+, drule_tac x = x in spec, simp add:max_def, simp add:add_cf_def m_cf_len m_cf_def) apply (simp, drule_tac x = x in spec, simp add:max_def, rule contrapos_pp, simp+, simp add:add_cf_def m_cf_len m_cf_def, frule_tac j = x in pol_coeff_mem[of c], frule_tac x = x and y = "fst d" and z = "fst c" in le_less_trans, assumption+, simp add:less_imp_le, simp add:aGroup.ag_inv_zero aGroup.ag_r_zero[of S]) apply (rule subsetI, simp, erule conjE, case_tac "fst d \ fst c", frule_tac i = x and j = "fst d" and k = "fst c" in le_trans, assumption+, simp, drule_tac x = x in spec, simp add:max_def, rule contrapos_pp, simp+, simp add:add_cf_def m_cf_len m_cf_def) apply (case_tac "fst d = fst c", simp, rotate_tac -1, drule sym, simp, frule_tac j = x in pol_coeff_mem[of d], assumption, frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+, simp add:aGroup.ag_l_zero, frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S], assumption, simp add:aGroup.ag_inv_zero) apply (drule noteq_le_less[of "fst d" "fst c"], assumption, simp, frule_tac j = x in pol_coeff_mem[of d], assumption, frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+, simp add:aGroup.ag_l_zero, frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S], assumption, simp add:aGroup.ag_inv_zero) apply (simp add:nat_not_le_less, case_tac "\ x \ fst c", simp, simp add:nat_not_le_less, drule_tac x = x in spec, simp add:max_def, simp add:add_cf_def m_cf_len m_cf_def, frule_tac j = x in pol_coeff_mem[of d], assumption, frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+, simp add:aGroup.ag_l_zero, frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S], assumption, simp add:aGroup.ag_inv_zero) apply (simp, drule_tac x = x in spec, simp add:max_def, rule contrapos_pp, simp+, simp add:add_cf_def m_cf_len m_cf_def, frule_tac x = x and y = "fst c" and z = "fst d" in le_less_trans, assumption+, frule_tac x = x and y = "fst d" in less_imp_le, frule_tac j = x in pol_coeff_mem[of d], assumption, frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+, simp add:aGroup.ag_l_zero, frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S], assumption, simp add:aGroup.ag_inv_zero) apply (thin_tac "\j\max (fst c) (fst d). snd (add_cf S c (m_cf S d)) j = \\<^bsub>S\<^esub>") apply (rotate_tac -1, drule sym, simp) apply (simp add:coeff_0_pol_0[THEN sym, of c "fst c"]) apply blast apply simp+ done lemma (in PolynRg) ex_polyn_expr:"p \ carrier R \ \c. pol_coeff S c \ p = polyn_expr R X (fst c) c" apply (cut_tac S_X_generate[of p], blast) apply assumption done lemma (in PolynRg) c_max_eqTr0:"\pol_coeff S c; k \ (fst c); polyn_expr R X k c = polyn_expr R X (fst c) c; \j\k. (snd c) j \ \\<^bsub>S\<^esub>\ \ c_max S (k, snd c) = c_max S c" apply (simp add:polyn_expr_short[of c k], frule pol_coeff_le[of c k], assumption+, rule polyn_degree_unique[of "(k, snd c)" c], assumption+, simp) done definition cf_sol :: "[('a, 'b) Ring_scheme, ('a, 'b1) Ring_scheme, 'a, 'a, nat \ (nat \ 'a)] \ bool" where "cf_sol R S X p c \ pol_coeff S c \ (p = polyn_expr R X (fst c) c)" definition deg_n ::"[('a, 'b) Ring_scheme, ('a, 'b1) Ring_scheme, 'a, 'a] \ nat" where "deg_n R S X p = c_max S (SOME c. cf_sol R S X p c)" definition deg ::"[('a, 'b) Ring_scheme, ('a, 'b1) Ring_scheme, 'a, 'a] \ ant" where "deg R S X p = (if p = \\<^bsub>R\<^esub> then -\ else (an (deg_n R S X p)))" lemma (in PolynRg) ex_cf_sol:"p \ carrier R \ \c. cf_sol R S X p c" apply (unfold cf_sol_def) apply (frule ex_polyn_expr[of p], (erule exE)+) apply (cut_tac n = "fst c" in le_refl, blast) done lemma (in PolynRg) deg_in_aug_minf:"p \ carrier R \ deg R S X p \ Z\<^sub>-\<^sub>\" apply (simp add:aug_minf_def deg_def an_def) done lemma (in PolynRg) deg_noninf:"p \ carrier R \ deg R S X p \ \" apply (cut_tac deg_in_aug_minf[of p], simp add:deg_def, simp add:aug_minf_def) apply (case_tac "p = \\<^bsub>R\<^esub>", simp+) done lemma (in PolynRg) deg_ant_int:"\p \ carrier R; p \ \\ \ deg R S X p = ant (int (deg_n R S X p))" by (simp add:deg_def an_def) lemma (in PolynRg) deg_an:"\p \ carrier R; p \ \\ \ deg R S X p = an (deg_n R S X p)" by (simp add:deg_def) lemma (in PolynRg) pol_SOME_1:"p \ carrier R \ cf_sol R S X p (SOME f. cf_sol R S X p f)" apply (frule ex_cf_sol[of p]) apply (rule_tac P = "cf_sol R S X p" in someI_ex, assumption) done lemma (in PolynRg) pol_SOME_2:"p \ carrier R \ pol_coeff S (SOME c. cf_sol R S X p c) \ p = polyn_expr R X (fst (SOME c. cf_sol R S X p c)) (SOME c. cf_sol R S X p c)" apply (frule pol_SOME_1[of p]) apply (simp add:cf_sol_def) done lemma (in PolynRg) coeff_max_zeroTr:"pol_coeff S c \ \j. j \ (fst c) \ (c_max S c) < j \ (snd c) j = \\<^bsub>S\<^esub>" apply (case_tac "\j \ (fst c). (snd c) j = \\<^bsub>S\<^esub>", rule allI, rule impI, erule conjE, simp) apply simp apply (frule coeff_nonzero_polyn_nonzero[THEN sym, of c "fst c"], simp, simp) apply (rule allI, rule impI, erule conjE, simp add:c_max_def, simp add:polyn_degreeTr[of c "fst c"]) apply (subgoal_tac "{j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>} \ {j. j \ (fst c)}", frule n_max[of "{j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>}" "fst c"], blast) apply (case_tac "\x\fst c. snd c x = \\<^bsub>S\<^esub> ", blast, simp) apply (erule conjE) apply (rule contrapos_pp, simp+, thin_tac "\x\fst c. snd c x \ \\<^bsub>S\<^esub>", thin_tac "{j. j \ fst c \ snd c j \ \\<^bsub>S\<^esub>} \ {j. j \ fst c}", thin_tac "snd c (n_max {j. j \ fst c \ snd c j \ \\<^bsub>S\<^esub>}) \ \\<^bsub>S\<^esub>", drule_tac a = j in forall_spec, simp) apply simp apply (rule subsetI, simp) done lemma (in PolynRg) coeff_max_nonzeroTr:"\pol_coeff S c; \j \ (fst c). (snd c) j \ \\<^bsub>S\<^esub>\ \ (snd c) (c_max S c) \ \\<^bsub>S\<^esub>" apply (simp add:c_max_def) apply (subgoal_tac "{j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>} \ {j. j \ (fst c)}", frule n_max[of "{j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>}" "fst c"], blast) apply (erule conjE, simp) apply (rule subsetI, simp) done lemma (in PolynRg) coeff_max_bddTr:"pol_coeff S c \ c_max S c \ (fst c)" apply (case_tac "\j\(fst c). (snd c) j = \\<^bsub>S\<^esub>", simp add:c_max_def) apply (simp add:c_max_def, frule polyn_degreeTr[of c "fst c"], simp, simp, subgoal_tac "{j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>} \ {j. j \ (fst c)}", frule n_max[of "{j. j \ (fst c) \ (snd c) j \ \\<^bsub>S\<^esub>}" "fst c"], blast, erule conjE, simp) apply (rule subsetI, simp) done lemma (in PolynRg) pol_coeff_max:"pol_coeff S c \ pol_coeff S ((c_max S c), snd c)" apply (rule pol_coeff_le[of c "c_max S c"], assumption) apply (simp add:coeff_max_bddTr) done lemma (in PolynRg) polyn_c_max:"pol_coeff S c \ polyn_expr R X (fst c) c = polyn_expr R X (c_max S c) c" apply (case_tac "(c_max S c) = (fst c)", simp) apply (frule coeff_max_bddTr[of c], frule noteq_le_less[of "c_max S c" "fst c"], assumption) apply (subst polyn_n_m1[of c "c_max S c" "fst c"], assumption+, simp) apply (frule_tac polyn_mem[of c "c_max S c"], assumption+) apply (subst higher_part_zero[of c "c_max S c"], assumption+) apply (frule coeff_max_zeroTr[of c], rule ballI, simp add:nset_def) apply (cut_tac ring_is_ag, simp add:aGroup.ag_r_zero) done lemma (in PolynRg) pol_deg_eq_c_max:"\p \ carrier R; pol_coeff S c; p = polyn_expr R X (fst c) c\ \ deg_n R S X p = c_max S c" apply (cut_tac subring, frule subring_Ring) apply (frule polyn_c_max[of c]) apply (frule pol_SOME_2[of p]) apply (subst deg_n_def, erule conjE) apply (rule polyn_degree_unique[of "Eps (cf_sol R S X p)" "c"], simp, assumption) apply simp done lemma (in PolynRg) pol_deg_le_n:"\p \ carrier R; pol_coeff S c; p = polyn_expr R X (fst c) c\ \ deg_n R S X p \ (fst c)" apply (frule pol_deg_eq_c_max[of p c], assumption+, frule coeff_max_bddTr[of c]) apply simp done lemma (in PolynRg) pol_deg_le_n1:"\p \ carrier R; pol_coeff S c; k \ (fst c); p = polyn_expr R X k c\ \ deg_n R S X p \ k" apply (simp add:deg_n_def, drule sym, simp) apply (frule pol_SOME_2[of p], erule conjE) apply (frule pol_coeff_le[of c k], assumption) apply (simp only:polyn_expr_short[of c k]) apply (drule sym) apply (subst polyn_degree_unique[of "SOME c. cf_sol R S X p c" "(k, snd c)"], assumption+, simp) apply (frule coeff_max_bddTr[of "(k, snd c)"], simp) done lemma (in PolynRg) pol_len_gt_deg:"\p \ carrier R; pol_coeff S c; p = polyn_expr R X (fst c) c; deg R S X p < (an j); j \ (fst c)\ \ (snd c) j = \\<^bsub>S\<^esub>" apply (case_tac "p = \\<^bsub>R\<^esub>", simp, drule sym) apply (simp add:coeff_0_pol_0[THEN sym, of c "fst c"]) apply (simp add:deg_def, simp add:aless_natless) apply (drule sym, simp) apply (frule coeff_max_zeroTr[of c]) apply (simp add:pol_deg_eq_c_max) done lemma (in PolynRg) pol_diff_deg_less:"\p \ carrier R; pol_coeff S c; p = polyn_expr R X (fst c) c; pol_coeff S d; fst c = fst d; (snd c) (fst c) = (snd d) (fst d)\ \ p \ (-\<^sub>a (polyn_expr R X (fst d) d)) = \ \ deg_n R S X (p \ (-\<^sub>a (polyn_expr R X (fst d) d))) < (fst c)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring) apply (case_tac "p \\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (polyn_expr R X (fst d) d)) = \\<^bsub>R\<^esub>", simp) apply simp apply (simp add:polyn_minus_m_cf[of d "fst d"], frule m_cf_pol_coeff[of d]) apply (cut_tac polyn_add1[of c "m_cf S d"], simp add:m_cf_len, thin_tac "polyn_expr R X (fst d) c \ polyn_expr R X (fst d) (m_cf S d) = polyn_expr R X (fst d) (add_cf S c (m_cf S d))") apply (frule add_cf_pol_coeff[of c "m_cf S d"], assumption+) apply (cut_tac polyn_mem[of "add_cf S c (m_cf S d)" "fst d"], frule pol_deg_le_n[of "polyn_expr R X (fst d) (add_cf S c (m_cf S d))" "add_cf S c (m_cf S d)"], assumption+, simp add:add_cf_len m_cf_len, simp add:add_cf_len m_cf_len) apply (rule noteq_le_less[of "deg_n R S X (polyn_expr R X (fst d) (add_cf S c (m_cf S d)))" "fst d"], assumption) apply (rule contrapos_pp, simp+) apply (cut_tac pol_deg_eq_c_max[of "polyn_expr R X (fst d) (add_cf S c (m_cf S d))" "add_cf S c (m_cf S d)"], simp, thin_tac "deg_n R S X (polyn_expr R X (fst d) (add_cf S c (m_cf S d))) = fst d") apply (frule coeff_nonzero_polyn_nonzero[of "add_cf S c (m_cf S d)" "fst d"], simp add:add_cf_len m_cf_len, simp, thin_tac "polyn_expr R X (fst d) (add_cf S c (m_cf S d)) \ \", frule coeff_max_nonzeroTr[of "add_cf S c (m_cf S d)"], simp add:add_cf_len m_cf_len, thin_tac "\j\fst d. snd (add_cf S c (m_cf S d)) j \ \\<^bsub>S\<^esub>", thin_tac "pol_coeff S (m_cf S d)", thin_tac "pol_coeff S (add_cf S c (m_cf S d))", thin_tac "polyn_expr R X (fst d) (add_cf S c (m_cf S d)) \ carrier R", simp, thin_tac "c_max S (add_cf S c (m_cf S d)) = fst d") apply (simp add:add_cf_def m_cf_def, frule pol_coeff_mem[of d "fst d"], simp, frule Ring.ring_is_ag[of S], simp add:aGroup.ag_r_inv1, assumption+, simp add:add_cf_len m_cf_len, assumption, simp add:add_cf_len m_cf_len, assumption+) done lemma (in PolynRg) pol_pre_lt_deg:"\p \ carrier R; pol_coeff S c; deg_n R S X p \ (fst c); (deg_n R S X p) \ 0; p = polyn_expr R X (deg_n R S X p) c \ \ (deg_n R S X (polyn_expr R X ((deg_n R S X p) - Suc 0) c)) < (deg_n R S X p)" apply (frule polyn_expr_short[of c "deg_n R S X p"], assumption) apply (cut_tac pol_deg_le_n[of "polyn_expr R X (deg_n R S X p - Suc 0) c" "(deg_n R S X p - Suc 0, snd c)"], simp) apply (rule polyn_mem[of c "deg_n R S X p - Suc 0"], assumption+, arith, rule pol_coeff_le[of c "deg_n R S X p - Suc 0"], assumption, arith, simp) apply (subst polyn_expr_short[of c "deg_n R S X p - Suc 0"], assumption+, arith, simp) done lemma (in PolynRg) pol_deg_n:"\p \ carrier R; pol_coeff S c; n \ fst c; p = polyn_expr R X n c; (snd c) n \ \\<^bsub>S\<^esub>\ \ deg_n R S X p = n" apply (simp add:polyn_expr_short[of c n]) apply (frule pol_coeff_le[of c n], assumption+, cut_tac pol_deg_eq_c_max[of p "(n, snd c)"], drule sym, simp, simp add:c_max_def) apply (rule conjI, rule impI, cut_tac le_refl[of n], thin_tac "deg_n R S X p = (if \x\n. snd c x = \\<^bsub>S\<^esub> then 0 else n_max {j. j \ fst (n, snd c) \ snd (n, snd c) j \ \\<^bsub>S\<^esub>})", drule_tac a = n in forall_spec, assumption, simp) apply (rule impI) apply (cut_tac n_max[of "{j. j \ n \ snd c j \ \\<^bsub>S\<^esub>}" n], erule conjE, drule_tac x = n in bspec, simp, simp) apply (rule subsetI, simp, blast, drule sym, simp, assumption) apply simp done lemma (in PolynRg) pol_expr_deg:"\p \ carrier R; p \ \\ \ \c. pol_coeff S c \ deg_n R S X p \ (fst c) \ p = polyn_expr R X (deg_n R S X p) c \ (snd c) (deg_n R S X p) \ \\<^bsub>S\<^esub>" apply (cut_tac subring, frule subring_Ring) apply (frule ex_polyn_expr[of p], erule exE, erule conjE) apply (frule_tac c = c in polyn_c_max) apply (frule_tac c = c in pol_deg_le_n[of p], assumption+) apply (frule_tac c1 = c and k1 ="fst c" in coeff_0_pol_0[THEN sym], simp) apply (subgoal_tac "p = polyn_expr R X (deg_n R S X p) c \ snd c (deg_n R S X p) \ \\<^bsub>S\<^esub>", blast) apply (subst pol_deg_eq_c_max, assumption+)+ apply simp apply (cut_tac c = c in coeff_max_nonzeroTr, simp+) done lemma (in PolynRg) deg_n_pos:"p \ carrier R \ 0 \ deg_n R S X p" by simp lemma (in PolynRg) pol_expr_deg1:"\p \ carrier R; d = na (deg R S X p)\ \ \c. (pol_coeff S c \ p = polyn_expr R X d c)" apply (cut_tac subring, frule subring_Ring) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:deg_def na_minf, subgoal_tac "pol_coeff S (0, (\j. \\<^bsub>S\<^esub>))", subgoal_tac "\ = polyn_expr R X d (0, (\j. \\<^bsub>S\<^esub>))", blast, cut_tac coeff_0_pol_0[of "(d, \j. \\<^bsub>S\<^esub>)" d], simp+, simp add:pol_coeff_def, simp add:Ring.ring_zero) apply (simp add:deg_def na_an, frule pol_expr_deg[of p], assumption, erule exE, (erule conjE)+, unfold split_paired_all, simp, blast) done end diff --git a/thys/Group-Ring-Module/Algebra6.thy b/thys/Group-Ring-Module/Algebra6.thy --- a/thys/Group-Ring-Module/Algebra6.thy +++ b/thys/Group-Ring-Module/Algebra6.thy @@ -1,5224 +1,5221 @@ (** Algebra6 author Hidetsune Kobayashi Group You Santo Department of Mathematics Nihon University hikoba@math.cst.nihon-u.ac.jp May 3, 2004. April 6, 2007 (revised) chapter 4. Ring theory section 14. the degree of a polynomial(continued) section 15. homomorphism of polynomial rings section 16. relatively prime polynomials **) theory Algebra6 imports Algebra5 begin definition s_cf :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a] \ nat \ (nat \ 'a)" where "s_cf R S X p = (if p = \\<^bsub>R\<^esub> then (0, \j. \\<^bsub>S\<^esub>) else SOME c. (pol_coeff S c \ p = polyn_expr R X (fst c) c \ (snd c) (fst c) \ \\<^bsub>S\<^esub>))" (* special coefficients for p *) definition lcf :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a] \ 'a" where "lcf R S X p = (snd (s_cf R S X p)) (fst (s_cf R S X p))" lemma (in PolynRg) lcf_val_0:"lcf R S X \ = \\<^bsub>S\<^esub>" by (simp add:lcf_def s_cf_def) lemma (in PolynRg) lcf_val:"\p \ carrier R; p \ \ \ \ lcf R S X p = (snd (s_cf R S X p)) (fst (s_cf R S X p))" by (simp add:lcf_def) lemma (in PolynRg) s_cf_pol_coeff:"p \ carrier R \ pol_coeff S (s_cf R S X p)" apply (simp add:s_cf_def) apply (case_tac "p = \\<^bsub>R\<^esub>", simp) apply (cut_tac subring, frule subring_Ring, simp add:pol_coeff_def Ring.ring_zero) apply simp apply (rule someI2_ex) apply (frule ex_polyn_expr[of p], erule exE, erule conjE) apply (frule_tac c = c in coeff_max_bddTr) apply (frule_tac c = c and n = "c_max S c" in pol_coeff_le, assumption) apply (subgoal_tac "p = polyn_expr R X (fst (c_max S c, snd c)) (c_max S c, snd c) \ snd (c_max S c, snd c) (fst (c_max S c, snd c)) \ \\<^bsub>S\<^esub>", blast) apply (rule conjI, simp) apply (subst polyn_expr_short[THEN sym], assumption+) apply (simp add:polyn_c_max) apply simp apply (rule coeff_max_nonzeroTr, assumption) apply (simp add:coeff_0_pol_0[THEN sym]) apply simp done lemma (in PolynRg) lcf_mem:"p \ carrier R \ (lcf R S X p) \ carrier S" apply (cut_tac subring, frule subring_Ring) apply (simp add:lcf_def) apply (cut_tac pol_coeff_mem[of "s_cf R S X p" "fst (s_cf R S X p)"], assumption, rule s_cf_pol_coeff, assumption, simp) done lemma (in PolynRg) s_cf_expr0:"p \ carrier R \ pol_coeff S (s_cf R S X p) \ p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)" apply (cut_tac subring, frule subring_Ring) apply (simp add:s_cf_def) apply (case_tac "p = \\<^bsub>R\<^esub>", simp) apply (rule conjI, simp add:pol_coeff_def, simp add:Ring.ring_zero) apply (simp add:polyn_expr_def, simp add:Subring_zero_ring_zero, simp add:ring_r_one) apply simp apply (rule someI2_ex) apply (frule ex_polyn_expr[of p], erule exE, erule conjE) apply (frule_tac c = c in coeff_max_bddTr) apply (frule_tac c = c and n = "c_max S c" in pol_coeff_le, assumption) apply (subgoal_tac "p = polyn_expr R X (fst (c_max S c, snd c)) (c_max S c, snd c) \ snd (c_max S c, snd c) (fst (c_max S c, snd c)) \ \\<^bsub>S\<^esub>", blast) apply (rule conjI, simp) apply (subst polyn_expr_short[THEN sym], assumption+) apply (simp add:polyn_c_max) apply simp apply (rule coeff_max_nonzeroTr, assumption) apply (simp add:coeff_0_pol_0[THEN sym]) apply simp done lemma (in PolynRg) pos_deg_nonzero:"\p \ carrier R; 0 < deg_n R S X p\ \ p \ \" apply (cut_tac s_cf_expr0[of p], (erule conjE)+) apply (frule pol_deg_eq_c_max[of p "s_cf R S X p"], assumption+) apply (simp, thin_tac "deg_n R S X p = c_max S (s_cf R S X p)") apply (simp add:c_max_def) apply (case_tac "\x\fst (s_cf R S X p). snd (s_cf R S X p) x = \\<^bsub>S\<^esub> ", simp) apply (thin_tac "0 < (if \x\fst (s_cf R S X p). snd (s_cf R S X p) x = \\<^bsub>S\<^esub> then 0 else n_max {j. j \ fst (s_cf R S X p) \ snd (s_cf R S X p) j \ \\<^bsub>S\<^esub>})") apply (simp add:coeff_0_pol_0[of "s_cf R S X p" "fst (s_cf R S X p)"]) apply assumption done lemma (in PolynRg) s_cf_expr:"\p \ carrier R; p \ \\ \ pol_coeff S (s_cf R S X p) \ p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) \ (snd (s_cf R S X p)) (fst (s_cf R S X p)) \ \\<^bsub>S\<^esub>" apply (simp add:s_cf_def) apply (rule someI2_ex) apply (frule ex_polyn_expr[of p], erule exE, erule conjE) apply (frule_tac c = c in coeff_max_bddTr) apply (frule_tac c = c and n = "c_max S c" in pol_coeff_le, assumption) apply (subgoal_tac "p = polyn_expr R X (fst (c_max S c, snd c)) (c_max S c, snd c) \ snd (c_max S c, snd c) (fst (c_max S c, snd c)) \ \\<^bsub>S\<^esub>", blast) apply (rule conjI, simp) apply (subst polyn_expr_short[THEN sym], assumption+) apply (simp add:polyn_c_max) apply simp apply (rule coeff_max_nonzeroTr, assumption) apply (simp add:coeff_0_pol_0[THEN sym]) apply simp done lemma (in PolynRg) lcf_nonzero:"\p \ carrier R; p \ \ \ \ lcf R S X p \ \\<^bsub>S\<^esub>" apply (frule s_cf_expr[of p], assumption) apply (simp add:lcf_def) done lemma (in PolynRg) s_cf_deg:"\p \ carrier R; p \ \\ \ deg_n R S X p = fst (s_cf R S X p)" apply (frule s_cf_expr[of p], assumption, (erule conjE)+) apply (simp add:pol_deg_n[of p "s_cf R S X p" "fst (s_cf R S X p)"]) done lemma (in PolynRg) pol_expr_edeg:"\p \ carrier R; deg R S X p \ (an d)\ \ \f. (pol_coeff S f \ fst f = d \ p = polyn_expr R X d f)" apply (case_tac "p = \\<^bsub>R\<^esub>") apply (subgoal_tac "pol_coeff S (d, \j. \\<^bsub>S\<^esub>) \ fst (d, \j. \\<^bsub>S\<^esub>) = d \ p = polyn_expr R X d (d, \j. \\<^bsub>S\<^esub>)", blast) apply (rule conjI) apply (simp add:pol_coeff_def, cut_tac Ring.ring_zero[of S], simp, cut_tac subring, simp add:subring_Ring) apply (cut_tac coeff_0_pol_0[of "(d, \j. \\<^bsub>S\<^esub>)" d], simp) apply (simp add:pol_coeff_def, cut_tac Ring.ring_zero[of S], simp, cut_tac subring, simp add:subring_Ring) apply simp apply (frule s_cf_expr[of p], assumption+, (erule conjE)+) apply (simp add:deg_def na_an, simp add:ale_natle) apply (simp add:s_cf_deg) apply (subgoal_tac "pol_coeff S (d, \j. (if j \ (fst (s_cf R S X p)) then (snd (s_cf R S X p) j) else \\<^bsub>S\<^esub>)) \ p = polyn_expr R X d (d, \j. (if j \ (fst (s_cf R S X p)) then (snd (s_cf R S X p) j) else \\<^bsub>S\<^esub>))", blast) apply (rule conjI) apply (simp add:pol_coeff_def, rule allI, rule impI, rule impI) apply (cut_tac subring, simp add:subring_Ring, simp add:subring_Ring[of S] Ring.ring_zero) apply (case_tac "fst (s_cf R S X p) = d", simp) apply (subst polyn_exprs_eq[of "(d, \j. if j \ d then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>)" "s_cf R S X p" d]) apply (simp add:pol_coeff_def, cut_tac Ring.ring_zero[of S], simp, cut_tac subring, simp add:subring_Ring, simp) apply (rule allI, rule impI, simp, assumption+) apply (drule noteq_le_less[of "fst (s_cf R S X p)" d], assumption+) apply (cut_tac polyn_n_m1[of "(d, \j. if j \ fst (s_cf R S X p) then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>)" "fst (s_cf R S X p)" d], simp) apply (cut_tac higher_part_zero[of "(d, \j. if j \ fst (s_cf R S X p) then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>)" "fst (s_cf R S X p)"], simp, thin_tac "polyn_expr R X d (d, \j. if j \ fst (s_cf R S X p) then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>) = polyn_expr R X (fst (s_cf R S X p)) (d, \j. if j \ fst (s_cf R S X p) then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>) \ \", thin_tac "\\<^sub>f R (\j. (if j \ fst (s_cf R S X p) then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>) \\<^sub>r X^\<^bsup>R j\<^esup>) (Suc (fst (s_cf R S X p))) d = \") apply (subst polyn_exprs_eq[of "(d, \j. if j \ fst (s_cf R S X p) then snd (s_cf R S X p) j else \\<^bsub>S\<^esub>)" "s_cf R S X p" "fst (s_cf R S X p)"]) apply (simp add:pol_coeff_def, rule allI, rule impI, rule impI, cut_tac subring, simp add:subring_Ring, simp add:subring_Ring[of S] Ring.ring_zero, assumption+) apply (simp) apply (rule allI, rule impI, simp) apply (frule polyn_mem[of "s_cf R S X p" "fst (s_cf R S X p)"], simp+) apply (cut_tac ring_is_ag, simp add:aGroup.ag_r_zero) apply (simp add:pol_coeff_def, rule allI, rule impI, rule impI, cut_tac subring, simp add:subring_Ring, simp add:subring_Ring[of S] Ring.ring_zero) apply simp apply (rule ballI, simp add:nset_def) apply (simp add:pol_coeff_def, rule allI, rule impI, rule impI, cut_tac subring, simp add:subring_Ring, simp add:subring_Ring[of S] Ring.ring_zero) apply assumption apply simp done lemma (in PolynRg) cf_scf:"\pol_coeff S c; k \ fst c; polyn_expr R X k c \ \\ \ \j \ fst (s_cf R S X (polyn_expr R X k c)). snd (s_cf R S X (polyn_expr R X k c)) j = snd c j" apply (frule polyn_mem[of c k], assumption+) apply (simp add:polyn_expr_short[of c k], rule allI, rule impI) apply (cut_tac pol_deg_le_n1[of "polyn_expr R X k c" c k], frule s_cf_expr0[of "polyn_expr R X k (k, snd c)"], erule conjE) apply (rotate_tac -1, drule sym) apply (case_tac "fst (s_cf R S X (polyn_expr R X k (k, snd c))) = k", simp, cut_tac c = "s_cf R S X (polyn_expr R X k (k, snd c))" and d = "(k, snd c)" in pol_expr_unique2, simp add:s_cf_pol_coeff, simp add:split_pol_coeff, simp, simp, simp add:polyn_expr_short[THEN sym, of c k]) apply (simp add:s_cf_deg[of "polyn_expr R X k c"], drule noteq_le_less[of "fst (s_cf R S X (polyn_expr R X k c))" k], assumption) apply (frule pol_expr_unique3[of "s_cf R S X (polyn_expr R X k c)" "(k, snd c)"], simp add:split_pol_coeff, simp, simp add:polyn_expr_short[THEN sym, of c k]) apply (simp add:polyn_expr_short[THEN sym, of c k], assumption+, simp) done definition scf_cond :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a, nat, nat \ (nat \ 'a)] \ bool" where "scf_cond R S X p d c \ pol_coeff S c \ fst c = d \ p = polyn_expr R X d c" definition scf_d :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a, nat] \ nat \ (nat \ 'a)" where "scf_d R S X p d = (SOME f. scf_cond R S X p d f)" (** system of coefficients, coeff_length d **) lemma (in PolynRg) scf_d_polTr:"\p \ carrier R; deg R S X p \ an d\ \ scf_cond R S X p d (scf_d R S X p d)" apply (simp add:scf_d_def) apply (rule_tac P = "scf_cond R S X p d" in someI2_ex) apply (frule pol_expr_edeg[of "p" "d"], assumption+) apply (simp add:scf_cond_def, assumption) done lemma (in PolynRg) scf_d_pol:"\p \ carrier R; deg R S X p \ an d\ \ pol_coeff S (scf_d R S X p d) \ fst (scf_d R S X p d) = d \ p = polyn_expr R X d (scf_d R S X p d)" apply (frule scf_d_polTr[of "p" "d"], assumption+) apply (simp add:scf_cond_def) done lemma (in PolynRg) pol_expr_of_X: "X = polyn_expr R X (Suc 0) (ext_cf S (Suc 0) (C\<^sub>0 S))" apply (cut_tac X_mem_R, cut_tac subring) apply (cut_tac X_to_d[of "Suc 0"]) apply (simp add:ring_l_one) done lemma (in PolynRg) deg_n_of_X:"deg_n R S X X = Suc 0" apply (cut_tac X_mem_R, cut_tac polyn_ring_S_nonzero, cut_tac subring) apply (cut_tac pol_expr_of_X) apply (cut_tac special_cf_pol_coeff) apply (frule ext_cf_pol_coeff[of "C\<^sub>0 S" "Suc 0"]) apply (frule pol_deg_eq_c_max[of X "ext_cf S (Suc 0) (C\<^sub>0 S)"], assumption) apply (simp add:ext_cf_len special_cf_len) apply (simp add:c_max_ext_special_cf) done lemma (in PolynRg) pol_X:"cf_sol R S X X c \ snd c 0 = \\<^bsub>S\<^esub> \ snd c (Suc 0) = 1\<^sub>r\<^bsub>S\<^esub>" apply (simp add:cf_sol_def, erule conjE) apply (cut_tac pol_expr_of_X) apply (cut_tac special_cf_pol_coeff, frule ext_cf_pol_coeff[of "C\<^sub>0 S" "Suc 0"]) apply (cut_tac X_mem_R, cut_tac polyn_ring_X_nonzero, cut_tac subring) apply (frule pol_deg_le_n[of X c], assumption+, simp add:deg_n_of_X) apply (case_tac "fst c = Suc 0") apply (frule box_equation[of X "polyn_expr R X (Suc 0) (ext_cf S (Suc 0) (C\<^sub>0 S))" "polyn_expr R X (fst c) c"], assumption+, thin_tac "X = polyn_expr R X (Suc 0) (ext_cf S (Suc 0) (C\<^sub>0 S))", thin_tac "X = polyn_expr R X (fst c) c") apply (cut_tac pol_expr_unique2[of "ext_cf S (Suc 0) (C\<^sub>0 S)" c], simp, simp add:ext_cf_len special_cf_len, rule conjI, drule_tac a = 0 in forall_spec, simp, simp add:ext_special_cf_lo_zero) apply( drule_tac a = "Suc 0" in forall_spec, simp, simp add:ext_special_cf_hi, assumption+, simp add:ext_cf_len special_cf_len) apply (frule noteq_le_less[of "Suc 0" "fst c"],rule not_sym, assumption, cut_tac pol_expr_unique3[of "ext_cf S (Suc 0) (C\<^sub>0 S)" c], simp add:ext_cf_len special_cf_len, erule conjE, thin_tac "\j\nset (Suc (Suc 0)) (fst c). snd c j = \\<^bsub>S\<^esub>") apply (rule conjI, drule_tac a = 0 in forall_spec, simp, simp add:ext_special_cf_lo_zero, drule_tac a = "Suc 0" in forall_spec, simp, simp add:ext_special_cf_hi, assumption+) apply (simp add:ext_cf_len special_cf_len) done lemma (in PolynRg) pol_of_deg0:"\p \ carrier R; p \ \\ \ (deg_n R S X p = 0) = (p \ carrier S)" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag, frule Ring.ring_is_ag[of S]) apply (rule iffI) apply (frule s_cf_expr[of p], assumption) apply (simp add:s_cf_deg, (erule conjE)+, simp add:polyn_expr_def) apply (frule pol_coeff_mem[of "s_cf R S X p" 0], simp) apply (cut_tac mem_subring_mem_ring[of S "snd (s_cf R S X p) 0"], simp add:ring_r_one, assumption+) apply (frule s_cf_expr[of p], assumption+, (erule conjE)+, simp add:s_cf_deg) apply (rule contrapos_pp, simp+) apply (subgoal_tac "pol_coeff S (0, (\j. p)) \ p = polyn_expr R X 0 (0, (\j. p))", erule conjE) apply (cut_tac a = "polyn_expr R X 0 (0, \j. p)" and b = "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)" in aGroup.ag_eq_diffzero[of R], assumption+, simp, simp) apply (frule_tac c = "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)" in box_equation[of p "polyn_expr R X 0 (0, \j. p)"], assumption, thin_tac "p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)", thin_tac "p = polyn_expr R X 0 (0, \j. p)", simp) apply (simp only:polyn_minus_m_cf) apply (rotate_tac -2, drule sym, simp, thin_tac "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) = polyn_expr R X 0 (0, \j. p)") apply (frule_tac c = "s_cf R S X p" in m_cf_pol_coeff) apply (frule_tac d = "m_cf S (s_cf R S X p)" in polyn_add1[of "(0, \j. p)"], assumption, simp add:m_cf_len, thin_tac "polyn_expr R X 0 (0, \j. p) \ polyn_expr R X (fst (s_cf R S X p)) (m_cf S (s_cf R S X p)) = polyn_expr R X (fst (s_cf R S X p)) (add_cf S (0, \j. p) (m_cf S (s_cf R S X p)))") apply (rotate_tac -1, drule sym, simp) apply (frule_tac d = "m_cf S (s_cf R S X p)" in add_cf_pol_coeff[of "(0, \j. p)"], assumption) apply (frule_tac c1 = "add_cf S (0, \j. p) (m_cf S (s_cf R S X p))" and k1 = "fst (s_cf R S X p)" in coeff_0_pol_0[THEN sym], simp add:add_cf_len m_cf_len, simp, thin_tac "pol_coeff S (add_cf S (0, \j. p) (m_cf S (s_cf R S X p)))", thin_tac "polyn_expr R X (fst (s_cf R S X p)) (add_cf S (0, \j. p) (m_cf S (s_cf R S X p))) = \") apply (drule_tac a = "fst (s_cf R S X p)" in forall_spec, simp, simp add:add_cf_def m_cf_len m_cf_def) apply (frule_tac c = "s_cf R S X p" and j = "fst (s_cf R S X p)" in pol_coeff_mem, simp, frule_tac x = "snd (s_cf R S X p) (fst (s_cf R S X p))" in aGroup.ag_inv_inv[of S], assumption, simp add:aGroup.ag_inv_zero) apply (subst polyn_expr_def, simp add:ring_r_one) apply (simp add:pol_coeff_def) done lemma (in PolynRg) pols_const:"\p \ carrier R; (deg R S X p) \ 0\ \ p \ carrier S" apply (case_tac "p = \\<^bsub>R\<^esub>") apply (cut_tac subring) apply (frule Subring_zero_ring_zero[THEN sym, of S], simp, cut_tac subring, frule subring_Ring[of S], rule Ring.ring_zero[of S], assumption) apply (subst pol_of_deg0[THEN sym], assumption+, simp add:deg_def, simp only:an_0[THEN sym], simp add:an_inj) done lemma (in PolynRg) less_deg_add_nonzero:"\p \ carrier R; p \ \; q \ carrier R; q \ \; (deg_n R S X p) < (deg_n R S X q)\ \ p \ q \ \" apply (frule ex_polyn_expr[of p], erule exE, erule conjE, frule ex_polyn_expr[of q], erule exE, erule conjE, rename_tac c d) apply (frule_tac c = c in pol_deg_eq_c_max[of p], assumption+, frule_tac c = d in pol_deg_eq_c_max[of q], assumption+, frule_tac c = c in coeff_max_bddTr, frule_tac c = d in coeff_max_bddTr) apply (frule_tac c = c and n = "c_max S c" in pol_coeff_le, assumption, frule_tac c = d and n = "c_max S d" in pol_coeff_le, assumption+) apply simp apply (subst polyn_c_max, assumption, subst polyn_c_max, assumption, subst polyn_expr_short, assumption+) apply (frule_tac c = d and k = "c_max S d" in polyn_expr_short, assumption+, simp, thin_tac "polyn_expr R X (c_max S d) d = polyn_expr R X (c_max S d) (c_max S d, snd d)", thin_tac "deg_n R S X (polyn_expr R X (fst c) c) = c_max S c", thin_tac "deg_n R S X (polyn_expr R X (fst d) d) = c_max S d") apply (subst polyn_add, assumption+, simp add: max.absorb1 max.absorb2) apply (rule contrapos_pp, simp+, frule_tac c = "(c_max S c, snd c)" and d = "(c_max S d, snd d)" in add_cf_pol_coeff, assumption+, frule_tac c1 = "add_cf S (c_max S c, snd c) (c_max S d, snd d)" and k1 = "c_max S d" in coeff_0_pol_0[THEN sym], simp add:add_cf_len, simp, thin_tac "pol_coeff S (add_cf S (c_max S c, snd c) (c_max S d, snd d))", thin_tac "polyn_expr R X (c_max S d) (add_cf S (c_max S c, snd c) (c_max S d, snd d)) = \") apply (drule_tac a = "c_max S d" in forall_spec, simp, simp add:add_cf_def, frule_tac c = d and k = "fst d" in coeff_nonzero_polyn_nonzero, simp, simp, frule_tac c = d in coeff_max_nonzeroTr, assumption+, simp) done lemma (in PolynRg) polyn_deg_add1:"\p \ carrier R; p \ \; q \ carrier R; q \ \; (deg_n R S X p) < (deg_n R S X q)\ \ deg_n R S X (p \ q) = (deg_n R S X q)" apply (cut_tac subring) apply (frule less_deg_add_nonzero[of p q], assumption+) apply (frule ex_polyn_expr[of p], erule exE, erule conjE, frule ex_polyn_expr[of q], erule exE, erule conjE, rename_tac c d) apply (simp only:pol_deg_eq_c_max) apply (frule_tac c = c in coeff_max_bddTr, frule_tac c = d in coeff_max_bddTr) apply (frule_tac c = c and n = "c_max S c" in pol_coeff_le, assumption, frule_tac c = d and n = "c_max S d" in pol_coeff_le, assumption) apply (simp add:polyn_c_max) apply (frule_tac c = c and k = "c_max S c" in polyn_expr_short, simp, frule_tac c = d and k = "c_max S d" in polyn_expr_short, simp, simp) apply (frule_tac c = "(c_max S c, snd c)" and d = "(c_max S d, snd d)" in polyn_add1, assumption+, simp, thin_tac "polyn_expr R X (c_max S c) c = polyn_expr R X (c_max S c) (c_max S c, snd c)", thin_tac "polyn_expr R X (c_max S d) d = polyn_expr R X (c_max S d) (c_max S d, snd d)") - apply (simp add: max.absorb1 max.absorb2) apply (frule_tac c = "(c_max S c, snd c)" and d = "(c_max S d, snd d)" in add_cf_pol_coeff, assumption+, rule_tac p = "polyn_expr R X (c_max S d) (add_cf S (c_max S c, snd c) (c_max S d, snd d))" and c = "add_cf S (c_max S c, snd c) (c_max S d, snd d)" and n = "c_max S d" in pol_deg_n) apply (rule_tac polyn_mem, simp, simp add:add_cf_len, assumption, simp add:add_cf_len, simp, subst add_cf_def, simp) apply (frule_tac c1 = "(c_max S d, snd d)" and k1 = "c_max S d" in coeff_0_pol_0[THEN sym], simp, simp, rule coeff_max_nonzeroTr, assumption+, erule exE, erule conjE, frule_tac i = j and j = "c_max S d" and k = "fst d" in le_trans, assumption+, blast) done lemma (in PolynRg) polyn_deg_add2:"\p \ carrier R; p \ \; q \ carrier R; q \ \; p \ q \ \; (deg_n R S X p) = (deg_n R S X q)\ \ deg_n R S X (p \ q) \ (deg_n R S X q)" apply (cut_tac subring) apply (frule ex_polyn_expr[of p], erule exE, erule conjE, frule ex_polyn_expr[of q], erule exE, erule conjE, rename_tac c d) apply (simp only:pol_deg_eq_c_max) apply (frule_tac c = c in coeff_max_bddTr, frule_tac c = d in coeff_max_bddTr) apply (frule_tac c = c and n = "c_max S c" in pol_coeff_le, assumption, frule_tac c = d and n = "c_max S d" in pol_coeff_le, assumption) apply (simp add:polyn_c_max) apply (frule_tac c = c and k = "c_max S c" in polyn_expr_short, simp, frule_tac c = d and k = "c_max S d" in polyn_expr_short, simp, simp, frule_tac c = "(c_max S d, snd c)" and d = "(c_max S d, snd d)" in polyn_add1, simp) apply (thin_tac "polyn_expr R X (c_max S d) d = polyn_expr R X (c_max S d) (c_max S d, snd d)", thin_tac "polyn_expr R X (c_max S d) c = polyn_expr R X (c_max S d) (c_max S d, snd c)", simp) apply (thin_tac "polyn_expr R X (c_max S d) (c_max S d, snd c) \ polyn_expr R X (c_max S d) (c_max S d, snd d) = polyn_expr R X (c_max S d) (add_cf S (c_max S d, snd c) (c_max S d, snd d))") apply (frule_tac c = "(c_max S d, snd c)" and d = "(c_max S d, snd d)" in add_cf_pol_coeff, assumption+) apply (cut_tac p = "polyn_expr R X (c_max S d) (add_cf S (c_max S d, snd c) (c_max S d, snd d))" and c = "add_cf S (c_max S d, snd c) (c_max S d, snd d)" in pol_deg_eq_c_max) apply (rule_tac polyn_mem, simp) apply (simp add:add_cf_len, assumption, simp add:add_cf_len, simp) apply (thin_tac "deg_n R S X (polyn_expr R X (c_max S d) (add_cf S (c_max S d, snd c) (c_max S d, snd d))) = c_max S (add_cf S (c_max S d, snd c) (c_max S d, snd d))", thin_tac "polyn_expr R X (c_max S d) (add_cf S (c_max S d, snd c) (c_max S d, snd d)) \ \") apply (frule_tac c = "add_cf S (c_max S d, snd c) (c_max S d, snd d)" in coeff_max_bddTr) apply (simp add:add_cf_len) done lemma (in PolynRg) polyn_deg_add3:"\p \ carrier R; p \ \; q \ carrier R; q \ \; p \ q \ \; (deg_n R S X p) \ n; (deg_n R S X q) \ n\ \ deg_n R S X (p \ q) \ n" apply (case_tac "(deg_n R S X p) = (deg_n R S X q)", frule polyn_deg_add2[of "p" "q"], assumption+, simp) apply (cut_tac less_linear[of "deg_n R S X p" "deg_n R S X q"], simp, thin_tac "deg_n R S X p \ deg_n R S X q", erule disjE, simp add:polyn_deg_add1, cut_tac ring_is_ag, simp add:aGroup.ag_pOp_commute[of "R" "p" "q"], simp add:polyn_deg_add1) done lemma (in PolynRg) polyn_deg_add4:"\p \ carrier R; q \ carrier R; (deg R S X p) \ (an n); (deg R S X q) \ (an n)\ \ deg R S X (p \ q) \ (an n)" apply (cut_tac ring_is_ag) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:aGroup.ag_l_zero) apply (case_tac "q = \\<^bsub>R\<^esub>", simp add:aGroup.ag_r_zero) apply (case_tac "p \\<^bsub>R\<^esub> q = \\<^bsub>R\<^esub>", simp add:deg_def) apply (frule aGroup.ag_pOp_closed[of R p q], assumption+) apply (simp add:deg_an) apply (simp only:ale_natle) apply (simp add:polyn_deg_add3) done lemma (in PolynRg) polyn_deg_add5:"\p \ carrier R; q \ carrier R; (deg R S X p) \ a; (deg R S X q) \ a\ \ deg R S X (p \ q) \ a" apply (case_tac "a = \", simp) apply (cut_tac ring_is_ag, case_tac "p = \\<^bsub>R\<^esub>", simp add:aGroup.ag_l_zero[of R], case_tac "q = \\<^bsub>R\<^esub>", simp add:aGroup.ag_r_zero, simp add:deg_an[of p]) apply (cut_tac an_nat_pos[of "deg_n R S X p"], frule ale_trans[of "0" "an (deg_n R S X p)" "a"], assumption+, subgoal_tac "an (deg_n R S X p) \ an (na a)", simp only:ale_natle[of "deg_n R S X p" "na a"]) apply (simp add:deg_an[of q]) apply (cut_tac an_nat_pos[of "deg_n R S X q"], frule ale_trans[of "0" "an (deg_n R S X q)" "a"], assumption+, subgoal_tac "an (deg_n R S X q) \ an (na a)", simp only:ale_natle[of "deg_n R S X q" "na a"]) apply (frule polyn_deg_add4[of p q "na a"], assumption+, simp add:an_na, simp add:deg_an, simp add:deg_an an_na, simp add:an_na) apply (simp add:deg_an an_na, simp add:deg_an an_na) done lemma (in PolynRg) lower_deg_part:"\p \ carrier R; p \ \; 0 < deg_n R S X p\ \ deg R S X (polyn_expr R X (deg_n R S X p - Suc 0)(SOME f. cf_sol R S X p f)) < deg R S X p" apply (case_tac "polyn_expr R X (deg_n R S X p - Suc 0) (SOME f. cf_sol R S X p f) = \\<^bsub>R\<^esub>") apply (simp add:deg_def, cut_tac minf_le_any[of "an (deg_n R S X p)"]) apply (subst less_le, simp, simp add:an_def) apply (rule not_sym, rule contrapos_pp, simp+) apply (simp add:deg_def, simp add:aless_natless) apply (frule pol_SOME_2[of p], erule conjE) apply (simp add:pol_deg_eq_c_max[of p "SOME c. cf_sol R S X p c"]) apply (frule_tac c = "SOME c. cf_sol R S X p c" in coeff_max_bddTr) apply (cut_tac p = "polyn_expr R X (c_max S (SOME c. cf_sol R S X p c) - Suc 0) (Eps (cf_sol R S X p))" and c = "(c_max S (SOME c. cf_sol R S X p c) - Suc 0, snd (SOME c. cf_sol R S X p c))" in pol_deg_eq_c_max) apply (rule polyn_mem, simp, arith) apply (rule pol_coeff_le, assumption, arith) apply (subst polyn_expr_short, arith, arith, simp) apply simp apply (cut_tac c = "(c_max S (SOME c. cf_sol R S X p c) - Suc 0, snd (SOME c. cf_sol R S X p c))" in coeff_max_bddTr, rule pol_coeff_le, assumption, arith, simp) done definition ldeg_p :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, nat, 'a] \ 'a" where "ldeg_p R S X d p = polyn_expr R X d (scf_d R S X p (Suc d))" (** deg R S X p \ (Suc d) **) definition hdeg_p :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, nat, 'a] \ 'a" where "hdeg_p R S X d p = (snd (scf_d R S X p d) d) \\<^sub>r\<^bsub>R\<^esub> (X^\<^bsup>R d\<^esup>)" (** deg R S X p \ d **) lemma (in PolynRg) ldeg_p_mem:"\p \ carrier R; deg R S X p \ an (Suc d) \ \ ldeg_p R S X d p \ carrier R" apply (frule scf_d_pol[of "p" "Suc d"], assumption+, erule conjE) apply (simp add:ldeg_p_def) apply (rule polyn_mem[of "scf_d R S X p (Suc d)" d], assumption+) apply simp done lemma (in PolynRg) ldeg_p_zero:"p = \\<^bsub>R\<^esub> \ ldeg_p R S X d p = \\<^bsub>R\<^esub>" apply (subgoal_tac "deg R S X p \ an (Suc d)", subgoal_tac "p \ carrier R") apply (frule scf_d_pol[of "p" "Suc d"], assumption+, erule conjE) apply simp apply (frule coeff_0_pol_0[of "scf_d R S X \ (Suc d)" "Suc d"], simp) apply (simp add:ldeg_p_def) apply (subst coeff_0_pol_0[THEN sym, of "scf_d R S X \ (Suc d)"], assumption+, simp) apply (rule allI, rule impI, simp) apply (simp, simp add:ring_zero) apply (simp add:deg_def) done lemma (in PolynRg) hdeg_p_mem:"\p \ carrier R; deg R S X p \ an (Suc d)\ \ hdeg_p R S X (Suc d) p \ carrier R" apply (frule scf_d_pol[of p "Suc d"], assumption+, erule conjE) apply (simp only:hdeg_p_def, (erule conjE)+) apply (cut_tac Ring) apply (rule Ring.ring_tOp_closed[of "R"], assumption) apply (frule pol_coeff_mem[of "scf_d R S X p (Suc d)" "Suc d"], simp) apply (cut_tac subring) apply (simp add:Ring.mem_subring_mem_ring) apply (rule Ring.npClose[of "R"], assumption+) apply (rule X_mem_R) done (* ***************************************************************** definition ldeg_p :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a] \ 'a" where "ldeg_p R S X p == if p = \\<^bsub>R\<^esub> then \\<^bsub>R\<^esub> else if deg_n R S X p = 0 then p else polyn_expr R X (fst (s_cf R S X p) - Suc 0) (s_cf R S X p)" *) (** deg R S X p \ (Suc d), lower degree part **) (* definition hdeg_p :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a,'a] \ 'a" where "hdeg_p R S X p == if p = \\<^bsub>R\<^esub> then \\<^bsub>R\<^esub> else (if (deg_n R S X p) = 0 then \\<^bsub>R\<^esub> else (snd (s_cf R S X p) (fst (s_cf R S X p))) \\<^sub>r\<^bsub>R\<^esub> X^\<^bsup>R (fst (s_cf R S X p))\<^esup>)" *) (** deg R S X p \ d, the highest degree term **) (* lemma (in PolynRg) ldeg_p_mem:"p \ carrier R \ ldeg_p R S X p \ carrier R" apply (simp add:ldeg_p_def) apply (simp add:ring_zero) apply (rule impI, rule impI) apply (frule s_cf_pol_coeff[of p]) apply (simp add:polyn_mem) done lemma (in PolynRg) ldeg_p_zero:"ldeg_p R S X \ = \" apply (simp add:ldeg_p_def) done lemma (in PolynRg) ldeg_p_zero1:"\p \ carrier R; p \ \; deg_n R S X p = 0\ \ ldeg_p R S X p = p" by (simp add:ldeg_p_def) lemma (in PolynRg) hdeg_p_mem:"p \ carrier R \ hdeg_p R S X p \ carrier R" apply (cut_tac ring_is_ag) apply (cut_tac subring) apply (simp add:hdeg_p_def) apply (case_tac "deg_n R S X p = 0", simp add:aGroup.ag_inc_zero) apply simp apply (simp add:aGroup.ag_inc_zero) apply (rule impI) apply (frule s_cf_pol_coeff[of p]) apply (cut_tac X_mem_R, rule ring_tOp_closed) apply (simp add:pol_coeff_mem mem_subring_mem_ring) apply (rule npClose, assumption) done *) lemma (in PolynRg) hdeg_p_zero:"p = \ \ hdeg_p R S X (Suc d) p = \" apply (cut_tac X_mem_R) apply (subgoal_tac "deg R S X p \ an (Suc d)", subgoal_tac "p \ carrier R") apply (frule scf_d_pol[of p "Suc d"], assumption+, erule conjE) apply simp apply (frule coeff_0_pol_0[of "scf_d R S X \ (Suc d)" "Suc d"], (erule conjE)+, simp) apply (simp only:hdeg_p_def) apply (rotate_tac -1, drule sym, simp del:npow_suc) apply (cut_tac subring, simp del:npow_suc add:Subring_zero_ring_zero, rule ring_times_0_x, rule npClose, assumption) apply (simp add:ring_zero) apply (simp add:deg_def) done lemma (in PolynRg) decompos_p:"\p \ carrier R; deg R S X p \ an (Suc d)\ \ p = (ldeg_p R S X d p) \ (hdeg_p R S X (Suc d) p)" apply (frule scf_d_pol[of p "Suc d"], assumption+, erule conjE) apply (cut_tac subring, (erule conjE)+) apply (cut_tac polyn_Suc[of d "scf_d R S X p (Suc d)"]) apply (simp only:ldeg_p_def hdeg_p_def) apply (rotate_tac -1, drule sym, simp del:npow_suc) apply (thin_tac "polyn_expr R X d (scf_d R S X p (Suc d)) \ snd (scf_d R S X p (Suc d)) (Suc d) \\<^sub>r X^\<^bsup>R (Suc d)\<^esup> = polyn_expr R X (Suc d) (Suc d, snd (scf_d R S X p (Suc d)))") apply (simp add:polyn_expr_split[of "Suc d" "scf_d R S X p (Suc d)"], simp) done lemma (in PolynRg) deg_ldeg_p:"\p \ carrier R; deg R S X p \ an (Suc d)\ \ deg R S X (ldeg_p R S X d p) \ an d" apply (cut_tac subring, frule subring_Ring) apply (case_tac "p = \\<^bsub>R\<^esub>") apply (simp add:ldeg_p_zero, simp add:deg_def) apply (frule scf_d_pol[of p "Suc d"], assumption+, (erule conjE)+) apply (simp only:ldeg_p_def) apply (case_tac "polyn_expr R X d (scf_d R S X p (Suc d)) = \\<^bsub>R\<^esub>") apply (simp add:deg_def) apply (simp add:deg_an) apply (simp add:ale_natle) apply (cut_tac pol_deg_le_n1[of "polyn_expr R X d (scf_d R S X p (Suc d))" "scf_d R S X p (Suc d)" d], simp add:deg_def ale_natle) apply (rule polyn_mem, assumption+, simp+) done lemma (in PolynRg) deg_minus_eq:"\p \ carrier R; p \ \\ \ deg_n R S X (-\<^sub>a p) = deg_n R S X p" apply (cut_tac subring, cut_tac ring_is_ag, frule subring_Ring) apply (cut_tac ring_is_ag) apply (frule s_cf_expr[of p], assumption, (erule conjE)+, frule polyn_minus_m_cf[of "s_cf R S X p" "fst (s_cf R S X p)"], simp, drule sym, simp) apply (frule_tac x = p in aGroup.ag_mOp_closed, assumption+, frule m_cf_pol_coeff [of "s_cf R S X p"], frule pol_deg_n[of "-\<^sub>a p" "m_cf S (s_cf R S X p)" "fst (s_cf R S X p)"], assumption, simp add:m_cf_len, assumption+) apply (simp add:m_cf_def, frule pol_coeff_mem[of "s_cf R S X p" "fst (s_cf R S X p)"], simp, frule Ring.ring_is_ag[of S]) apply (rule contrapos_pp, simp+) apply (frule aGroup.ag_inv_inv[THEN sym, of S "snd (s_cf R S X p) (fst (s_cf R S X p))"], assumption, simp add:aGroup.ag_inv_zero) apply (drule sym, simp, simp add:s_cf_deg) done lemma (in PolynRg) deg_minus_eq1:"p \ carrier R \ deg R S X (-\<^sub>a p) = deg R S X p" apply (cut_tac ring_is_ag) apply (case_tac "p = \\<^bsub>R\<^esub>") apply (simp add:aGroup.ag_inv_zero) apply (frule deg_minus_eq[of p], assumption+, frule aGroup.ag_inv_inj[of "R" "p" "\"], assumption, simp add:ring_zero, assumption, simp add:aGroup.ag_inv_zero) apply (frule aGroup.ag_mOp_closed[of R p], assumption, simp add:deg_an) done lemma (in PolynRg) ldeg_p_pOp:"\p \ carrier R; q \ carrier R; deg R S X p \ an (Suc d); deg R S X q \ an (Suc d)\ \ (ldeg_p R S X d p) \\<^bsub>R\<^esub> (ldeg_p R S X d q) = ldeg_p R S X d (p \\<^bsub>R\<^esub> q)" apply (simp add:ldeg_p_def, cut_tac ring_is_ag, cut_tac subring, frule subring_Ring[of S], frule scf_d_pol[of p "Suc d"], assumption+, frule scf_d_pol[of q "Suc d"], assumption+, (erule conjE)+) apply (frule polyn_add1[of "scf_d R S X p (Suc d)" "scf_d R S X q (Suc d)"], assumption+, rotate_tac -2, drule sym, frule aGroup.ag_pOp_closed[of "R" "p" "q"], assumption+, frule polyn_deg_add4 [of p q "Suc d"], assumption+, rotate_tac -5, drule sym) apply simp apply (rotate_tac 4, drule sym, simp) apply (rotate_tac -1, drule sym, frule scf_d_pol[of "p \ q" "Suc d"], assumption+, (erule conjE)+, frule box_equation[of "p \ q" "polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d)))" "polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))"], assumption+, thin_tac "p \ q = polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))") apply (frule add_cf_pol_coeff[of "scf_d R S X p (Suc d)" "scf_d R S X q (Suc d)"], assumption+) apply (frule pol_expr_unique2[of "add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d))" "scf_d R S X (p \ q) (Suc d)"], assumption+) apply (subst add_cf_len[of "scf_d R S X p (Suc d)" "scf_d R S X q (Suc d)"], assumption+) apply (thin_tac "polyn_expr R X (Suc d) (scf_d R S X p (Suc d)) = p", thin_tac "polyn_expr R X (Suc d) (scf_d R S X q (Suc d)) = q", thin_tac "p \ q = polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d)))", thin_tac "polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d))) = polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))") apply simp apply (thin_tac "polyn_expr R X (Suc d) (scf_d R S X p (Suc d)) = p", thin_tac "polyn_expr R X (Suc d) (scf_d R S X q (Suc d)) = q", thin_tac "p \ q = polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d)))") apply (simp add:add_cf_len, thin_tac "polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d))) = polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))") apply (subst polyn_expr_short[of "scf_d R S X p (Suc d)" d], assumption, simp) apply (subst polyn_expr_short[of "scf_d R S X q (Suc d)" d], assumption, simp) apply (subst polyn_add_n[of d "snd (scf_d R S X p (Suc d))" "snd (scf_d R S X q (Suc d))"]) apply (simp add:split_pol_coeff, simp add:split_pol_coeff, subst polyn_expr_def) apply (rule aGroup.nsum_eq, assumption+, rule allI, rule impI, frule_tac j = j in pol_coeff_mem[of "scf_d R S X p (Suc d)"], simp, frule_tac j = j in pol_coeff_mem[of "scf_d R S X q (Suc d)"], simp, cut_tac Ring, rule Ring.ring_tOp_closed, assumption+, rule Ring.mem_subring_mem_ring[of R S], assumption+, frule Ring.ring_is_ag[of S], rule aGroup.ag_pOp_closed[of S], assumption+, rule Ring.npClose, assumption, simp add:X_mem_R) apply (rule allI, rule impI, frule_tac j = j in pol_coeff_mem[of "scf_d R S X (p \ q) (Suc d)"], simp, cut_tac Ring, subst Ring.ring_tOp_closed, assumption, rule Ring.mem_subring_mem_ring[of R S], assumption+, rule Ring.npClose, assumption, simp add:X_mem_R, simp) apply (rule allI, rule impI, drule_tac a = j in forall_spec, simp, thin_tac " pol_coeff S (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d)))") apply (simp add:add_cf_def) done lemma (in PolynRg) hdeg_p_pOp:"\p \ carrier R; q \ carrier R; deg R S X p \ an (Suc d); deg R S X q \ an (Suc d)\ \ (hdeg_p R S X (Suc d) p) \ (hdeg_p R S X (Suc d) q) = hdeg_p R S X (Suc d) (p \ q)" apply (cut_tac Ring, frule Ring.ring_is_ag[of "R"]) apply (cut_tac subring, frule subring_Ring[of S]) apply (frule scf_d_pol[of p "Suc d"], assumption+, frule scf_d_pol[of q "Suc d"], assumption+, (erule conjE)+) apply (frule polyn_add1[of "scf_d R S X p (Suc d)" "scf_d R S X q (Suc d)"], assumption+, rotate_tac -2, drule sym, frule aGroup.ag_pOp_closed[of "R" "p" "q"], assumption+, frule polyn_deg_add4 [of p q "Suc d"], assumption+, rotate_tac -5, drule sym) apply simp apply (rotate_tac -13, drule sym, simp) apply (rotate_tac -1, drule sym) apply (frule scf_d_pol[of "p \ q" "Suc d"], assumption+, (erule conjE)+) apply (drule box_equation[of "p \ q" "polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d)))" "polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))"], assumption+) apply ( thin_tac "p \ q = polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))") apply (frule add_cf_pol_coeff[of "scf_d R S X p (Suc d)" "scf_d R S X q (Suc d)"], assumption+) apply (cut_tac pol_expr_unique2[of "add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d))" "scf_d R S X (p \ q) (Suc d)"], simp add:add_cf_len) apply (drule_tac a = "Suc d" in forall_spec, simp) apply (simp del:npow_suc add:hdeg_p_def) apply (rotate_tac -1, drule sym, simp del:npow_suc) apply (subst add_cf_def, simp del:npow_suc) apply (thin_tac "polyn_expr R X (Suc d) (scf_d R S X p (Suc d)) = p", thin_tac "polyn_expr R X (Suc d) (scf_d R S X q (Suc d)) = q", thin_tac "polyn_expr R X (Suc d) (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d))) = polyn_expr R X (Suc d) (scf_d R S X (p \ q) (Suc d))", thin_tac "pol_coeff S (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d)))", thin_tac "snd (scf_d R S X (p \ q) (Suc d)) (Suc d) = snd (add_cf S (scf_d R S X p (Suc d)) (scf_d R S X q (Suc d))) (Suc d)") apply (frule pol_coeff_mem[of "scf_d R S X p (Suc d)" "Suc d"], simp del:npow_suc, frule pol_coeff_mem[of "scf_d R S X q (Suc d)" "Suc d"], simp del:npow_suc) apply (simp del:npow_suc add:Subring_pOp_ring_pOp) apply (frule mem_subring_mem_ring[of S "snd (scf_d R S X p (Suc d)) (Suc d)"], assumption, frule mem_subring_mem_ring[of S "snd (scf_d R S X q (Suc d)) (Suc d)"], assumption, cut_tac X_mem_R, frule Ring.npClose[of R X "Suc d"], assumption+) apply (subst Ring.ring_distrib2[THEN sym], assumption+, simp) apply (simp add:add_cf_pol_coeff, simp) apply (simp add:add_cf_len) done lemma (in PolynRg) ldeg_p_mOp:"\p \ carrier R; deg R S X p \ an (Suc d)\ \ -\<^sub>a (ldeg_p R S X d p) = ldeg_p R S X d (-\<^sub>a p)" apply (cut_tac Ring, frule Ring.ring_is_ag[of "R"], cut_tac subring, frule subring_Ring[of S], frule scf_d_pol[of p "Suc d"], assumption+, (erule conjE)+, frule aGroup.ag_mOp_closed[of R p], assumption, frule scf_d_pol[of "-\<^sub>a p" "Suc d"]) apply (simp add:deg_minus_eq1, (erule conjE)+) apply (frule polyn_minus[of "scf_d R S X p (Suc d)" "Suc d"], simp) apply (drule box_equation[of "-\<^sub>a p" "polyn_expr R X (Suc d) (scf_d R S X (-\<^sub>a p) (Suc d))" "polyn_expr R X (Suc d) (fst (scf_d R S X p (Suc d)), \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)"]) apply (rotate_tac 8, drule sym, simp, thin_tac "-\<^sub>a polyn_expr R X (Suc d) (scf_d R S X p (Suc d)) = polyn_expr R X (Suc d) (fst (scf_d R S X p (Suc d)), \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)") apply simp apply (frule pol_expr_unique2[of "scf_d R S X (-\<^sub>a p) (Suc d)" "(Suc d, \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)"]) apply (subst pol_coeff_def, rule allI, rule impI, simp) apply (frule_tac j = j in pol_coeff_mem[of "scf_d R S X p (Suc d)"], simp, frule Ring.ring_is_ag[of S], rule aGroup.ag_mOp_closed[of S], assumption+, simp) apply simp apply (simp add:ldeg_p_def) apply (subst polyn_minus[of "scf_d R S X p (Suc d)" d], assumption, simp, simp) apply (subst polyn_expr_short[of "(Suc d, \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)" d]) apply (subst pol_coeff_def, rule allI, rule impI, simp, frule_tac j = j in pol_coeff_mem[of "scf_d R S X p (Suc d)"], simp, frule Ring.ring_is_ag[of S], rule aGroup.ag_mOp_closed[of S], assumption+, simp) apply (subst polyn_expr_short[of "scf_d R S X (-\<^sub>a p) (Suc d)" d], assumption, simp) apply (cut_tac pol_expr_unique2[of "(d, snd (Suc d, \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j))" "(d, snd (scf_d R S X (-\<^sub>a p) (Suc d)))"]) apply (thin_tac "p = polyn_expr R X (Suc d) (scf_d R S X p (Suc d))", thin_tac "polyn_expr R X (Suc d) (scf_d R S X (-\<^sub>a p) (Suc d)) = polyn_expr R X (Suc d) (Suc d, \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)", simp) apply (subst pol_coeff_def, rule allI, rule impI, simp, frule Ring.ring_is_ag[of S], rule aGroup.ag_mOp_closed, assumption, simp add:pol_coeff_mem) apply (subst pol_coeff_def, rule allI, rule impI, simp, frule Ring.ring_is_ag[of S], rule aGroup.ag_mOp_closed, assumption, simp add:pol_coeff_mem) apply simp done lemma (in PolynRg) hdeg_p_mOp:"\p \ carrier R;deg R S X p \ an (Suc d)\ \ -\<^sub>a (hdeg_p R S X (Suc d) p) = hdeg_p R S X (Suc d) (-\<^sub>a p)" apply (cut_tac Ring, frule Ring.ring_is_ag[of "R"], cut_tac subring, frule subring_Ring[of S], frule scf_d_pol[of p "Suc d"], assumption+, (erule conjE)+, frule aGroup.ag_mOp_closed[of R p], assumption) apply ( frule scf_d_pol[of "-\<^sub>a p" "Suc d"]) apply (simp add:deg_minus_eq1, (erule conjE)+) apply (frule polyn_minus[of "scf_d R S X p (Suc d)" "Suc d"], simp) apply (drule box_equation[of "-\<^sub>a p" "polyn_expr R X (Suc d) (scf_d R S X (-\<^sub>a p) (Suc d))" "polyn_expr R X (Suc d) (fst (scf_d R S X p (Suc d)), \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)"]) apply (rotate_tac 8, drule sym, simp, thin_tac "-\<^sub>a polyn_expr R X (Suc d) (scf_d R S X p (Suc d)) = polyn_expr R X (Suc d) (fst (scf_d R S X p (Suc d)), \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)") apply simp apply (frule pol_expr_unique2[of "scf_d R S X (-\<^sub>a p) (Suc d)" "(Suc d, \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)"]) apply (subst pol_coeff_def, rule allI, rule impI, simp) apply (frule_tac j = j in pol_coeff_mem[of "scf_d R S X p (Suc d)"], simp, frule Ring.ring_is_ag[of S], rule aGroup.ag_mOp_closed[of S], assumption+, simp) apply simp apply (drule_tac a = "Suc d" in forall_spec, simp) apply (simp del:npow_suc add:hdeg_p_def) apply (thin_tac "p = polyn_expr R X (Suc d) (scf_d R S X p (Suc d))", thin_tac "polyn_expr R X (Suc d) (scf_d R S X (-\<^sub>a p) (Suc d)) = polyn_expr R X (Suc d) (Suc d, \j. -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) j)", thin_tac "snd (scf_d R S X (-\<^sub>a p) (Suc d)) (Suc d) = -\<^sub>a\<^bsub>S\<^esub> snd (scf_d R S X p (Suc d)) (Suc d)") apply (frule pol_coeff_mem[of "scf_d R S X p (Suc d)" "Suc d"], simp, frule mem_subring_mem_ring[of S "snd (scf_d R S X p (Suc d)) (Suc d)"], assumption+, frule Ring.npClose[of R X "Suc d"], simp add:X_mem_R) apply (subst Ring.ring_inv1_1[of "R"], assumption+) apply (simp del:npow_suc add:Subring_minus_ring_minus) done subsection "Multiplication of polynomials" lemma (in PolynRg) deg_mult_pols:"\Idomain S; p \ carrier R; p \ \; q \ carrier R; q \ \ \ \ p \\<^sub>r q \ \ \ deg_n R S X (p \\<^sub>r q) = deg_n R S X p + deg_n R S X q" apply (frule Idomain.idom_is_ring[of "S"], frule_tac x = p and y = q in ring_tOp_closed, assumption+) apply (frule s_cf_expr[of p], assumption, frule s_cf_expr[of q], assumption, (erule conjE)+) apply (frule_tac c = "s_cf R S X p" and d = "s_cf R S X q" in polyn_expr_tOp_c, assumption, erule exE, (erule conjE)+) apply (drule sym, drule sym, simp, thin_tac "polyn_expr R X (fst (s_cf R S X q)) (s_cf R S X q) = q", thin_tac "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) = p") apply (rotate_tac -1, drule sym, frule ring_tOp_closed[of p q], assumption+, frule pol_coeff_mem[of "s_cf R S X p" "fst (s_cf R S X p)"], simp, frule pol_coeff_mem[of "s_cf R S X q" "fst (s_cf R S X q)"], simp, frule_tac x = "snd (s_cf R S X p) (fst (s_cf R S X p))" and y = "snd (s_cf R S X q) (fst (s_cf R S X q))" in Idomain.idom_tOp_nonzeros[of "S"], assumption+, frule_tac c = e in coeff_nonzero_polyn_nonzero[ of _ "deg_n R S X p + deg_n R S X q"], simp) apply (simp add:s_cf_deg, simp add:s_cf_deg) apply (cut_tac n = "fst (s_cf R S X p) + fst (s_cf R S X q)" in le_refl) apply (subgoal_tac "\j\fst (s_cf R S X p) + fst (s_cf R S X q). snd e j \ \\<^bsub>S\<^esub>", simp) apply (cut_tac c = e in pol_deg_n[of "p \\<^sub>r q" _ "fst (s_cf R S X p) + fst (s_cf R S X q)"], simp+) apply (thin_tac "(polyn_expr R X (fst (s_cf R S X p) + fst (s_cf R S X q)) e \ \) = (\j\fst (s_cf R S X p) + fst (s_cf R S X q). snd e j \ \\<^bsub>S\<^esub>)", thin_tac "p \\<^sub>r q = polyn_expr R X (fst (s_cf R S X p) + fst (s_cf R S X q)) e", thin_tac "polyn_expr R X (fst (s_cf R S X p) + fst (s_cf R S X q)) e \ carrier R", thin_tac "snd (s_cf R S X p) (fst (s_cf R S X p)) \ carrier S", thin_tac "snd (s_cf R S X q) (fst (s_cf R S X q)) \ carrier S") apply (drule sym, drule sym, simp) apply (cut_tac n = "fst e" in le_refl, blast) done lemma (in PolynRg) deg_mult_pols1:"\Idomain S; p \ carrier R; q \ carrier R\ \ deg R S X (p \\<^sub>r q) = deg R S X p + deg R S X q" apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:ring_times_0_x, simp add:deg_def, rule impI) apply (simp add:an_def) apply (case_tac "q = \\<^bsub>R\<^esub>", simp add:ring_times_x_0, simp add:deg_def) apply (simp add:an_def) apply (frule deg_mult_pols[of p q], assumption+, erule conjE) apply (frule Idomain.idom_is_ring[of "S"]) apply (frule ring_tOp_closed[of p q], assumption+) apply (simp add:deg_an an_def a_zpz) done lemma (in PolynRg) const_times_polyn:"\Idomain S; c \ carrier S; c \ \\<^bsub>S\<^esub>; p \ carrier R; p \ \\ \ (c \\<^sub>r p) \ \ \ deg_n R S X (c \\<^sub>r p) = deg_n R S X p" apply (frule Idomain.idom_is_ring[of "S"], cut_tac subring, frule mem_subring_mem_ring[of S c], assumption+, simp add:Subring_zero_ring_zero) apply (frule deg_mult_pols[of c p], assumption+, erule conjE, simp, simp add:pol_of_deg0[THEN sym, of "c"]) done (*lemma (in PolynRg) deg_minus_eq:"\ring R; integral_domain S; polyn_ring R S X; p \ carrier R; p \ 0\<^sub>R\ \ deg_n R S X (-\<^sub>R p) = deg_n R S X p" *) lemma (in PolynRg) p_times_monomial_nonzero:"\p \ carrier R; p \ \\ \ (X^\<^bsup>R j\<^esup>) \\<^sub>r p \ \" apply (cut_tac subring, frule subring_Ring) apply (frule s_cf_expr[of p], assumption+, (erule conjE)+) apply (frule low_deg_terms_zero1[THEN sym, of "s_cf R S X p" j]) apply (drule sym, simp, thin_tac "X^\<^bsup>R j\<^esup> \\<^sub>r p = polyn_expr R X (fst (s_cf R S X p) + j) (ext_cf S j (s_cf R S X p))", thin_tac "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) = p") apply (frule ext_cf_pol_coeff[of "s_cf R S X p" j]) apply(frule coeff_nonzero_polyn_nonzero[of "ext_cf S j (s_cf R S X p)" "fst (ext_cf S j (s_cf R S X p))"], simp) apply (simp add:ext_cf_len add.commute[of j], thin_tac "(polyn_expr R X (fst (s_cf R S X p) + j) (ext_cf S j (s_cf R S X p)) \ \) = (\ja\fst (s_cf R S X p) + j. snd (ext_cf S j (s_cf R S X p)) ja \ \\<^bsub>S\<^esub>)") apply (cut_tac ext_cf_hi[of "s_cf R S X p" j], simp, thin_tac "snd (s_cf R S X p) (fst (s_cf R S X p)) = snd (ext_cf S j (s_cf R S X p)) (j + fst (s_cf R S X p))", simp add:add.commute[of _ j]) apply (cut_tac n = "j + fst (s_cf R S X p)" in le_refl, blast) apply assumption done lemma (in PolynRg) p_times_monomial_nonzero1:"\Idomain S; p \ carrier R; p \ \; c \ carrier S; c \ \\<^bsub>S\<^esub>\ \(c \\<^sub>r (X^\<^bsup>R j\<^esup>)) \\<^sub>r p \ \" apply (frule Idomain.idom_is_ring[of "S"], cut_tac subring, cut_tac X_mem_R , frule mem_subring_mem_ring[of S c], assumption+, frule npClose[of X j]) apply (simp add:ring_tOp_commute[of c], simp add:ring_tOp_assoc, frule const_times_polyn[of c p], assumption+, erule conjE, frule ring_tOp_closed[of c p], assumption+, simp add:p_times_monomial_nonzero[of "c \\<^sub>r p"]) done lemma (in PolynRg) polyn_ring_integral:"Idomain S = Idomain R" apply (cut_tac subring, frule subring_Ring) apply (rule iffI) apply (subst Idomain_def) apply (cut_tac Ring, simp) apply (rule Idomain_axioms.intro, rule contrapos_pp, simp+, erule conjE, frule_tac p = a and q = b in deg_mult_pols, assumption+, erule conjE, simp) apply (subst Idomain_def) apply (cut_tac Ring, simp) apply (rule Idomain_axioms.intro, rule contrapos_pp, simp+, erule conjE) apply (simp add:Subring_tOp_ring_tOp) apply (frule_tac x = a in mem_subring_mem_ring[of S], assumption, frule_tac x = b in mem_subring_mem_ring[of S], assumption) apply (frule_tac a = a and b = b in Idomain.idom[of R], assumption+) apply (simp add:Subring_zero_ring_zero) apply (erule disjE, simp add:Subring_zero_ring_zero) apply (simp add:Subring_zero_ring_zero) done lemma (in PolynRg) deg_to_X_d:"Idomain S \ deg_n R S X (X^\<^bsup>R d\<^esup>) = d" apply (cut_tac subring, frule subring_Ring, cut_tac polyn_ring_S_nonzero, cut_tac X_mem_R, cut_tac polyn_ring_X_nonzero, cut_tac polyn_ring_integral) apply (induct_tac d) apply (cut_tac ring_one, simp add:Subring_one_ring_one[THEN sym], simp add:Subring_zero_ring_zero) apply (subst pol_of_deg0[of "1\<^sub>r\<^bsub>S\<^esub>"], assumption+, simp add:Ring.ring_one[of S]) apply simp apply (subst deg_mult_pols, assumption+, simp add:npClose, rule Idomain.idom_potent_nonzero, assumption+) apply (simp add:deg_n_of_X) done subsection \Degree with value in \aug_minf\\ lemma (in PolynRg) nonzero_deg_pos:"\p \ carrier R; p \ \\ \ 0 \ deg R S X p" by (simp add:deg_def) lemma (in PolynRg) deg_minf_pol_0:"p \ carrier R \ (deg R S X p = -\) = (p = \)" apply (rule iffI) apply (rule contrapos_pp, simp+, frule nonzero_deg_pos[of p], assumption+, simp add:deg_def an_def) apply (simp add:deg_def) done lemma (in PolynRg) pol_nonzero:"p \ carrier R \ (0 \ deg R S X p) = (p \ \)" apply (rule iffI) apply (rule contrapos_pp, simp+, simp add:deg_def, cut_tac minf_le_any[of "0"], frule ale_antisym[of "0" "-\"], assumption+, simp only:an_0[THEN sym], simp only:an_def, simp del:of_nat_0) apply (simp add:deg_def) done lemma (in PolynRg) minus_deg_in_aug_minf:"\p \ carrier R; p \ \\ \ - (deg R S X p) \ Z\<^sub>-\<^sub>\" apply (frule deg_in_aug_minf[of p], frule pol_nonzero[THEN sym, of p], simp add:aug_minf_def, rule contrapos_pp, simp+, cut_tac a_minus_minus[of "deg R S X p"], simp) apply (thin_tac "- deg R S X p = \", frule sym, thin_tac "- \ = deg R S X p", frule deg_minf_pol_0[of p], simp) done lemma (in PolynRg) deg_of_X:"deg R S X X = 1" (* the degree of the polyn. X *) apply (cut_tac X_mem_R, cut_tac polyn_ring_X_nonzero, cut_tac subring) apply (simp add:deg_def, simp only:an_1[THEN sym], rule nat_eq_an_eq, simp add:deg_n_of_X) done lemma (in PolynRg) pol_deg_0:"\p \ carrier R; p \ \\ \ (deg R S X p = 0) = (p \ carrier S)" apply (simp add:deg_def, simp only:an_0[THEN sym], rule iffI, frule an_inj[of "deg_n R S X p" "0"], simp, simp add:pol_of_deg0, rule nat_eq_an_eq, simp add:pol_of_deg0[of p]) done lemma (in PolynRg) deg_of_X2n:"Idomain S \ deg R S X (X^\<^bsup>R n\<^esup>) = an n" apply (frule Idomain.idom_is_ring[of "S"]) apply (cut_tac subring, cut_tac X_mem_R, cut_tac polyn_ring_X_nonzero, cut_tac polyn_ring_integral, simp) apply (induct_tac n) apply simp apply (simp add:Subring_one_ring_one[THEN sym, of "S"]) apply (subst pol_deg_0[of "1\<^sub>r\<^bsub>S\<^esub>"]) apply (simp add:Subring_one_ring_one, simp add:ring_one) apply (simp add:Subring_one_ring_one[of S] polyn_ring_nonzero) apply (simp add:Ring.ring_one[of S]) apply (frule_tac n = n in npClose[of X]) apply (simp add:deg_def) apply (simp add:Idomain.idom_potent_nonzero, frule_tac n = "Suc n" in Idomain.idom_potent_nonzero[of R X], assumption+, simp) apply (rule nat_eq_an_eq) apply (frule_tac n = n in Idomain.idom_potent_nonzero[of R X], assumption+) apply (frule_tac n = "deg_n R S X (X^\<^bsup>R n\<^esup>)" and m = n in an_inj, thin_tac "an (deg_n R S X (X^\<^bsup>R n\<^esup>)) = an n") apply (cut_tac deg_of_X) apply (simp add:deg_def, simp only:an_1[THEN sym]) apply (frule_tac n = "deg_n R S X X" and m = 1 in an_inj) apply (simp add:deg_mult_pols) done lemma (in PolynRg) add_pols_nonzero:"\p \ carrier R; q \ carrier R; (deg R S X p) \ (deg R S X q)\ \ p \ q \ \" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:deg_minf_pol_0[THEN sym], simp add:aGroup.ag_l_zero, rule contrapos_pp, simp+, case_tac "q = \\<^bsub>R\<^esub>", simp add:aGroup.ag_r_zero) apply (simp add:deg_def, simp only:aneq_natneq[of "deg_n R S X p" "deg_n R S X q"], cut_tac less_linear[of "deg_n R S X p" "deg_n R S X q"], simp, erule disjE, rule less_deg_add_nonzero[of p q], assumption+, frule less_deg_add_nonzero[of q p], assumption+, simp add:aGroup.ag_pOp_commute) done lemma (in PolynRg) deg_pols_add1:"\p \ carrier R; q \ carrier R; (deg R S X p) < (deg R S X q)\ \ deg R S X (p \ q) = deg R S X q" apply (cut_tac ring_is_ag, case_tac "p = \\<^bsub>R\<^esub>", simp add:deg_def aGroup.ag_l_zero, case_tac "q = \\<^bsub>R\<^esub>", simp add:deg_def) apply (rule impI) apply (simp add:an_def z_neq_minf) apply (fold an_def, frule aless_imp_le[of "an (deg_n R S X p)" " - \"], cut_tac minf_le_any[of "an (deg_n R S X p)"], frule ale_antisym[of "an (deg_n R S X p)" "- \"], assumption+, simp add:an_neq_minf) apply (simp add:deg_def, simp add:aless_nat_less, frule less_deg_add_nonzero[of p q], assumption+, simp, frule polyn_deg_add1[of p q], assumption+, simp) done lemma (in PolynRg) deg_pols_add2:"\p \ carrier R; q \ carrier R; (deg R S X p) = (deg R S X q)\ \ deg R S X (p \ q) \ (deg R S X q)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:aGroup.ag_l_zero) apply (case_tac "q = \\<^bsub>R\<^esub>", simp add:aGroup.ag_r_zero) apply (simp add:deg_def, frule an_inj[of "deg_n R S X p" "deg_n R S X q"], simp, rule impI, subst ale_natle, simp add:polyn_deg_add2) done lemma (in PolynRg) deg_pols_add3:"\p \ carrier R; q \ carrier R; (deg R S X p) \ an n; (deg R S X q) \ an n\ \ deg R S X (p \ q) \ an n" apply (case_tac "(deg R S X p) = (deg R S X q)", frule deg_pols_add2[of p q], assumption+, simp) apply (cut_tac aless_linear[of "deg R S X p" "deg R S X q"], simp, thin_tac "deg R S X p \ deg R S X q", erule disjE, simp add:deg_pols_add1, cut_tac ring_is_ag, simp add:aGroup.ag_pOp_commute[of "R" "p" "q"], simp add:deg_pols_add1) done lemma (in PolynRg) const_times_polyn1:"\Idomain S; p\ carrier R; c \ carrier S; c \ \\<^bsub>S\<^esub>\ \ deg R S X (c \\<^sub>r p) = deg R S X p" apply (frule Idomain.idom_is_ring, cut_tac subring, frule mem_subring_mem_ring[of S c], assumption, simp add:Subring_zero_ring_zero) apply (subst deg_mult_pols1[of c p], assumption+, simp add: pol_deg_0[THEN sym, of "c"], simp add:aadd_0_l) done section "Homomorphism of polynomial rings" definition cf_h :: " ('a \ 'b) \ nat \ (nat \ 'a) \ nat \ (nat \ 'b)" where "cf_h f = (\c. (fst c, cmp f (snd c)))" definition polyn_Hom :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, ('b, 'n) Ring_scheme, ('b, 'n1) Ring_scheme, 'b] \ ('a \ 'b) set" ("(pHom _ _ _, _ _ _)" [67,67,67,67,67,68]67) where "pHom R S X, A B Y = {f. f \ rHom R A \ f`(carrier S) \ carrier B \ f X = Y}" (* Hom from a polynomial ring to a polynomial ring *) definition erh :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, ('b, 'n) Ring_scheme, ('b, 'n1) Ring_scheme, 'b, 'a \ 'b, nat, nat \ (nat \ 'a)] \ 'b" where "erh R S X A B Y f n c = polyn_expr A Y n (cf_h f c)" (* extension of a ring hom. *) lemma (in PolynRg) cf_h_len:"\PolynRg A B Y; f \ rHom S B; pol_coeff S c\ \ fst (cf_h f c) = fst c" by (simp add:cf_h_def) lemma (in PolynRg) cf_h_coeff:"\PolynRg A B Y; f \ rHom S B; pol_coeff S c\ \ pol_coeff B (cf_h f c)" apply (subst pol_coeff_def) apply (rule allI, rule impI, simp add:cf_h_len cf_h_def) apply (frule_tac j = j in pol_coeff_mem[of c], assumption) apply (simp add:cmp_def rHom_mem) done lemma (in PolynRg) cf_h_cmp:"\PolynRg A B Y; pol_coeff S (n, f); h \ rHom S B; j \ n\ \ (snd (cf_h h (n, f))) j = (cmp h f) j" by (simp add:cf_h_def) lemma (in PolynRg) cf_h_special_cf:"\PolynRg A B Y; h \ rHom S B\ \ polyn_expr A Y (Suc 0) (cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S))) = polyn_expr A Y (Suc 0) (ext_cf B (Suc 0) (C\<^sub>0 B))" apply (cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (cut_tac special_cf_pol_coeff, frule ext_cf_pol_coeff[of "C\<^sub>0 S" "Suc 0"], frule cf_h_coeff[of A B Y h "ext_cf S (Suc 0) (C\<^sub>0 S)"], assumption+) apply (frule PolynRg.special_cf_pol_coeff, frule PolynRg.ext_cf_pol_coeff[of A B Y "C\<^sub>0 B" "Suc 0"], assumption) apply (frule PolynRg.pol_expr_unique2[of A B Y "cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S))" "ext_cf B (Suc 0) (C\<^sub>0 B)"], assumption+, simp add:cf_h_len PolynRg.ext_cf_len, simp add:ext_cf_len special_cf_len PolynRg.special_cf_len, simp add:cf_h_len PolynRg.ext_cf_len, simp add:ext_cf_len special_cf_len PolynRg.special_cf_len, thin_tac "(polyn_expr A Y (Suc 0) (cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S))) = polyn_expr A Y (Suc 0) (ext_cf B (Suc 0) (C\<^sub>0 B))) = (\j\Suc 0. snd (cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S))) j = snd (ext_cf B (Suc 0) (C\<^sub>0 B)) j)", thin_tac "pol_coeff S (C\<^sub>0 S)", thin_tac "pol_coeff S (ext_cf S (Suc 0) (C\<^sub>0 S))", thin_tac "pol_coeff B (cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S)))") apply (rule allI, rule impI) apply (case_tac "j = 0", simp add:cf_h_def cmp_def ext_cf_def sliden_def) apply (simp add:rHom_0_0) apply (simp) apply (frule_tac n = j in Suc_leI[of 0], frule_tac m = j and n = "Suc 0" in le_antisym, assumption+, thin_tac "j \ Suc 0", thin_tac "Suc 0 \ j", simp add:cf_h_def cmp_def ext_cf_def sliden_def special_cf_def, simp add:rHom_one) done lemma (in PolynRg) polyn_Hom_coeff_to_coeff: "\PolynRg A B Y; f \ pHom R S X, A B Y; pol_coeff S c\ \ pol_coeff B (cf_h f c)" apply (subst pol_coeff_def) apply (rule allI, rule impI, simp add:cf_h_len cf_h_def) apply (frule_tac j = j in pol_coeff_mem[of c], assumption) apply (simp add:cmp_def polyn_Hom_def, (erule conjE)+) apply (simp add:image_def) apply (rule subsetD[of "{y. \x\carrier S. y = f x}" "carrier B"], assumption, simp) apply blast done (* old name is cmp_pol_coeff *) lemma (in PolynRg) cf_h_len1:"\PolynRg A B Y; h \ rHom S B; f \ pHom R S X, A B Y; \x\carrier S. f x = h x; pol_coeff S c\ \ fst (cf_h f c) = fst (cf_h h c)" by (simp add:cf_h_def) lemma (in PolynRg) cf_h_len2:"\PolynRg A B Y; f \ pHom R S X, A B Y; pol_coeff S c\ \ fst (cf_h f c) = fst c" by (simp add:cf_h_def) lemma (in PolynRg) cmp_pol_coeff:"\f \ rHom S B; pol_coeff S (n, c)\ \ pol_coeff B (n,(cmp f c))" apply (simp add:pol_coeff_def, rule allI, rule impI, simp add:cmp_def, frule_tac a = j in forall_spec, simp, thin_tac "\j\n. c j \ carrier S") apply (simp add:rHom_mem) done lemma (in PolynRg) cmp_pol_coeff_e:"\PolynRg A B Y; f \ pHom R S X, A B Y; pol_coeff S (n, c)\ \ pol_coeff B (n, (cmp f c))" apply (subst pol_coeff_def) apply (rule allI, rule impI, simp) apply (frule_tac j = j in pol_coeff_mem[of "(n, c)"], simp) apply (simp add:polyn_Hom_def cmp_def, (erule conjE)+) apply (simp add:image_def) apply (rule_tac c = "f (c j)" in subsetD[of "{y. \x\carrier S. y = f x}" "carrier B"], assumption+) apply (simp, blast) done lemma (in PolynRg) cf_h_pol_coeff:"\PolynRg A B Y; h \ rHom S B; pol_coeff S (n, f)\ \ cf_h h (n, f) = (n, cmp h f)" by (simp add:cf_h_def) lemma (in PolynRg) cf_h_polyn:"\PolynRg A B Y; h \ rHom S B; pol_coeff S (n, f)\ \ polyn_expr A Y n (cf_h h (n, f)) = polyn_expr A Y n (n, (cmp h f))" apply (frule cf_h_coeff[of A B Y h "(n, f)"], assumption+, frule cmp_pol_coeff[of h B n f], assumption+) apply (rule PolynRg.polyn_exprs_eq[of A B Y "cf_h h (n, f)" "(n, cmp h f)" n], assumption+, simp add:cf_h_len, rule allI, rule impI, simp add:cf_h_def) done lemma (in PolynRg) pHom_rHom:"\PolynRg A B Y; f \ pHom R S X, A B Y\ \ f \ rHom R A" by (simp add:polyn_Hom_def) lemma (in PolynRg) pHom_X_Y:"\PolynRg A B Y; f \ pHom R S X, A B Y\ \ f X = Y" by (simp add:polyn_Hom_def) lemma (in PolynRg) pHom_memTr:"\PolynRg A B Y; f \ pHom R S X, A B Y\ \ \c. pol_coeff S (n, c) \ f (polyn_expr R X n (n, c)) = polyn_expr A Y n (n, cmp f c)" apply (cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (induct_tac n) apply (rule allI, rule impI) apply (simp add:polyn_expr_def cmp_def) apply (frule_tac c = "(0, c)" and j = 0 in pol_coeff_mem, simp, simp, frule_tac x = "c 0" in mem_subring_mem_ring, assumption+, simp add:polyn_Hom_def, (erule conjE)+, frule rHom_func[of f R A], frule_tac x = "c 0" in funcset_mem[of f "carrier R" "carrier A"], assumption+, simp add:ring_r_one, simp add:Ring.ring_r_one) apply (rule allI, rule impI) apply (cut_tac n = n and f = c in pol_coeff_pre, assumption) apply ( drule_tac a = c in forall_spec, assumption) apply (cut_tac n = n and c = "(Suc n, c)" in polyn_Suc, simp, simp del:npow_suc, thin_tac "polyn_expr R X (Suc n) (Suc n, c) = polyn_expr R X n (Suc n, c) \ c (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") apply (frule_tac c = "(Suc n, c)" and k = n in polyn_expr_short, simp) apply (simp del:npow_suc, thin_tac "polyn_expr R X n (Suc n, c) = polyn_expr R X n (n, c)") apply (frule_tac c = "(Suc n, c)" in polyn_Hom_coeff_to_coeff[of A B Y f], assumption+, simp del:npow_suc add:cf_h_def) apply (frule_tac c = "(Suc n, cmp f c)" and n = n in PolynRg.polyn_Suc[of A B Y], simp, simp del:npow_suc, thin_tac "polyn_expr A Y (Suc n) (Suc n, cmp f c) = polyn_expr A Y n (Suc n, cmp f c) \\<^bsub>A\<^esub> cmp f c (Suc n) \\<^sub>r\<^bsub>A\<^esub> Y^\<^bsup>A (Suc n)\<^esup>") apply (frule_tac k = n and c = "(n, c)" in polyn_mem, simp) apply (frule_tac c = "(Suc n, c)" in monomial_mem, drule_tac a = "Suc n" in forall_spec, simp, simp del:npow_suc) apply (frule pHom_rHom[of A B Y f], assumption+, simp del:npow_suc add:rHom_add) apply (frule_tac c = "(Suc n, c)" and j = "Suc n" in pol_coeff_mem_R, simp, simp del:npow_suc) apply (cut_tac X_mem_R, frule_tac n = "Suc n" in npClose[of X], cut_tac Ring, subst rHom_tOp[of R A _ _ f], assumption+) apply (frule_tac c = "(Suc n, cmp f c)" and k = n in PolynRg.polyn_expr_short[of A B Y], assumption+, simp, simp del:npow_suc) apply (cut_tac c = "(Suc n, cmp f c)" and n = n in PolynRg.pol_coeff_le[of A B Y], assumption+, simp, simp del:npow_suc add:PolynRg.pol_coeff_le[of A B Y]) apply (subst rHom_npow[of R A X f], assumption+, simp del:npow_suc add:pHom_X_Y cmp_def) done lemma (in PolynRg) pHom_mem:"\PolynRg A B Y; f \ pHom R S X, A B Y; pol_coeff S (n, c)\ \ f (polyn_expr R X n (n, c)) = polyn_expr A Y n (n, cmp f c)" apply (simp add:pHom_memTr) done lemma (in PolynRg) pHom_memc:"\PolynRg A B Y; f \ pHom R S X, A B Y; pol_coeff S c\ \ f (polyn_expr R X (fst c) c) = polyn_expr A Y (fst c) (cf_h f c)" by (cases c) (simp add: cf_h_def pHom_mem) lemma (in PolynRg) pHom_mem1:"\PolynRg A B Y; f \ pHom R S X, A B Y; p \ carrier R\ \ f p \ carrier A" apply (simp add:polyn_Hom_def, (erule conjE)+) apply (simp add:rHom_mem) done lemma (in PolynRg) pHom_pol_mem:"\PolynRg A B Y; f \ pHom R S X, A B Y; p \ carrier R; p \ \\ \ f p = polyn_expr A Y (deg_n R S X p) (cf_h f (s_cf R S X p))" apply (frule s_cf_expr[of p], assumption, (erule conjE)+) apply (subst s_cf_deg[of p], assumption+) apply (subst pHom_memc[THEN sym, of A B Y f], assumption+) apply (drule sym, simp) done lemma (in PolynRg) erh_rHom_coeff:"\PolynRg A B Y; h \ rHom S B; pol_coeff S c\ \ erh R S X A B Y h 0 c = (cmp h (snd c)) 0" apply (cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (simp add:erh_def polyn_expr_def cf_h_def) apply (frule pol_coeff_mem [of c 0], simp) apply (simp add:cmp_def, frule rHom_mem[of h S B "snd c 0"], assumption) apply (frule Ring.mem_subring_mem_ring[of A B "h (snd c 0)"], assumption+, simp add:Ring.ring_r_one) done lemma (in PolynRg) erh_polyn_exprs:"\PolynRg A B Y; h \ rHom S B; pol_coeff S c; pol_coeff S d; polyn_expr R X (fst c) c = polyn_expr R X (fst d) d\ \ erh R S X A B Y h (fst c) c = erh R S X A B Y h (fst d) d" apply (cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption+) apply (simp add:erh_def) apply (cut_tac less_linear[of "fst c" "fst d"]) apply (erule disjE, frule pol_expr_unique3[of c d], assumption+, simp, thin_tac "polyn_expr R X (fst c) c = polyn_expr R X (fst d) d", frule cf_h_coeff[of A B Y h c], assumption+, frule cf_h_coeff[of A B Y h d], assumption+) apply (frule PolynRg.pol_expr_unique3[of A B Y "cf_h h c" "cf_h h d"], assumption+, simp add:cf_h_len, simp add:cf_h_len, thin_tac "(polyn_expr A Y (fst c) (cf_h h c) = polyn_expr A Y (fst d) (cf_h h d)) = ((\j\fst c. snd (cf_h h c) j = snd (cf_h h d) j) \ (\j\nset (Suc (fst c)) (fst d). snd (cf_h h d) j = \\<^bsub>B\<^esub>))", simp add:cf_h_def cmp_def, simp add:rHom_0_0) apply (erule disjE, frule pol_expr_unique2[of c d], assumption+, simp, thin_tac "polyn_expr R X (fst d) c = polyn_expr R X (fst d) d", frule cf_h_coeff[of A B Y h c], assumption+, frule cf_h_coeff[of A B Y h d], assumption+) apply (frule PolynRg.pol_expr_unique2[of A B Y "cf_h h c" "cf_h h d"], assumption+, simp add:cf_h_len, simp add:cf_h_len, thin_tac "(polyn_expr A Y (fst d) (cf_h h c) = polyn_expr A Y (fst d) (cf_h h d)) = (\j\fst d. snd (cf_h h c) j = snd (cf_h h d) j)", simp add:cf_h_def cmp_def) apply (drule sym, rule sym, frule pol_expr_unique3[of d c], assumption+, simp, thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) c", frule cf_h_coeff[of A B Y h c], assumption+, frule cf_h_coeff[of A B Y h d], assumption+) apply (frule PolynRg.pol_expr_unique3[of A B Y "cf_h h d" "cf_h h c"], assumption+, simp add:cf_h_len, simp add:cf_h_len, thin_tac "(polyn_expr A Y (fst d) (cf_h h d) = polyn_expr A Y (fst c) (cf_h h c)) = ((\j\fst d. snd (cf_h h d) j = snd (cf_h h c) j) \ (\j\nset (Suc (fst d)) (fst c). snd (cf_h h c) j = \\<^bsub>B\<^esub>))", simp add:cf_h_def cmp_def, simp add:rHom_0_0) done definition erH :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, ('b, 'n) Ring_scheme, ('b, 'n1) Ring_scheme, 'b, 'a \ 'b] \ 'a \ 'b" where "erH R S X A B Y h = (\x\carrier R. erh R S X A B Y h (fst (s_cf R S X x)) (s_cf R S X x))" (* lemma (in PolynRg) erH_phom:"\PolynRg A B y; h \ rHom S B\ \ erH R S X A B Y h \ pHom R S X, A B Y" *) lemma (in PolynRg) erH_rHom_0:"\PolynRg A B Y; h \ rHom S B\ \ (erH R S X A B Y h) \ = \\<^bsub>A\<^esub>" apply (cut_tac subring, frule subring_Ring, cut_tac PolynRg.is_Ring[of A B Y], cut_tac PolynRg.subring[of A B Y], cut_tac Ring.subring_Ring[of A B]) apply (simp add:erH_def erh_def s_cf_def polyn_expr_def) apply (cut_tac ring_zero, simp add:cf_h_def cmp_def rHom_0_0, simp add:Ring.Subring_zero_ring_zero, frule Ring.ring_zero[of A], simp add:Ring.ring_r_one, assumption+) done lemma (in PolynRg) erH_mem:"\PolynRg A B Y; h \ rHom S B; p \ carrier R\ \ erH R S X A B Y h p \ carrier A" apply (cut_tac subring, frule subring_Ring, cut_tac PolynRg.is_Ring[of A B Y], cut_tac PolynRg.subring[of A B Y], cut_tac Ring.subring_Ring[of A B]) apply (case_tac "p = \\<^bsub>R\<^esub>") apply (simp add:erH_rHom_0, simp add:Ring.ring_zero) apply (simp add:erH_def erh_def) apply (frule s_cf_expr[of p], assumption, (erule conjE)+) apply (rule PolynRg.polyn_mem[of A B Y "cf_h h (s_cf R S X p)"], assumption+) apply (simp add:cf_h_coeff) apply (simp add:cf_h_len, assumption+) done lemma (in PolynRg) erH_rHom_nonzero:"\PolynRg A B Y; f \ rHom S B; p \ carrier R; (erH R S X A B Y f) p \ \\<^bsub>A\<^esub>\ \ p \ \" apply (rule contrapos_pp, simp+) apply (simp add:erH_rHom_0) done lemma (in PolynRg) erH_rHomTr2:"\PolynRg A B Y; h \ rHom S B\ \ (erH R S X A B Y h) (1\<^sub>r) = (1\<^sub>r\<^bsub>A\<^esub>)" apply (cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption, cut_tac polyn_ring_nonzero) apply (cut_tac Subring_one_ring_one[THEN sym, of S], frule Ring.ring_one[of S], cut_tac ring_one) apply (frule s_cf_expr[of "1\<^sub>r"], assumption+, (erule conjE)+) apply (frule s_cf_deg[THEN sym, of "1\<^sub>r"], assumption, simp) apply (cut_tac pol_of_deg0[THEN sym, of "1\<^sub>r"], simp, simp add:erH_def erh_def cf_h_def polyn_expr_def, frule pol_coeff_mem[of "s_cf R S X 1\<^sub>r\<^bsub>S\<^esub>" 0], simp) apply (simp add:Subring_tOp_ring_tOp[THEN sym], simp add:Ring.ring_r_one cmp_def) apply (simp add:rHom_one, simp add:Ring.Subring_one_ring_one[of A B], frule Ring.ring_one[of A], simp add:Ring.ring_r_one) apply (simp add:ring_one) apply simp apply assumption done declare max.absorb1 [simp] max.absorb2 [simp] lemma (in PolynRg) erH_multTr:"\PolynRg A B Y; h \ rHom S B; pol_coeff S c\ \ \f g. pol_coeff S (m, f) \ pol_coeff S (((fst c) + m), g) \ (polyn_expr R X (fst c) c) \\<^sub>r (polyn_expr R X m (m, f)) = (polyn_expr R X ((fst c) + m) ((fst c) + m, g)) \ (polyn_expr A Y (fst c) (cf_h h c)) \\<^sub>r\<^bsub>A\<^esub> (polyn_expr A Y m (cf_h h (m, f))) = (polyn_expr A Y ((fst c) + m) (cf_h h ((fst c)+m, g)))" apply (cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (cases c) apply (simp only:) apply (rename_tac l u) apply (thin_tac "c = (l, u)") apply (induct_tac m) apply ((rule allI)+, rule impI, (erule conjE)+, simp) apply (simp add:cf_h_polyn[of A B Y h]) apply (simp add:polyn_expr_def[of _ _ 0]) apply (frule_tac c = "(0, f)" and j = 0 in pol_coeff_mem, simp, simp, frule_tac c = "(0, f)" and j = 0 in pol_coeff_mem_R, simp, simp, frule_tac c = "(l, u)" and k = l in polyn_mem, simp, simp add:ring_r_one, simp add:ring_tOp_commute, simp add:scalar_times_pol_expr) apply (frule_tac c = "(0, f)" in cf_h_coeff[of A B Y h], assumption+, frule_tac c = "(l, u)" in cf_h_coeff[of A B Y h], assumption+) apply (frule_tac c = "cf_h h (0, f)" in PolynRg.pol_coeff_mem[of A B Y _ 0], assumption+, simp, simp add:cf_h_cmp, frule_tac c = "cf_h h (0, f)" in PolynRg.pol_coeff_mem_R[of A B Y _ 0], assumption+, simp, simp add:cf_h_cmp, frule_tac c = "cf_h h (l, u)" and k = l in PolynRg.polyn_mem, simp, simp add:cf_h_len, simp add:cf_h_polyn, simp add:Ring.ring_r_one, simp add:Ring.ring_tOp_commute[of A], frule_tac n = l and f = u in cf_h_pol_coeff[of A B Y h], assumption+, simp) apply (simp add:PolynRg.scalar_times_pol_expr, frule_tac c = "(l, u)" and a = "f 0" in sp_cf_pol_coeff, assumption+, frule_tac c = "(l, cmp h u)" and a = "(cmp h f) 0" in PolynRg.sp_cf_pol_coeff, assumption+, frule_tac c = "(l, g)" in cf_h_coeff[of A B Y h], assumption+, simp add:cf_h_pol_coeff) apply (rule_tac c = "sp_cf B (cmp h f 0) (l, cmp h u)" and d = "(l, cmp h g)" and k = l in PolynRg.polyn_exprs_eq[of A B Y], assumption+, simp add:sp_cf_len, simp add:PolynRg.sp_cf_len) apply (rule allI, rule impI) apply (frule_tac c = "sp_cf S (f 0) (l, u)" and d = "(l, g)" in pol_expr_unique2, assumption+, simp add:sp_cf_len, simp add:sp_cf_len) apply (drule_tac a = j in forall_spec, simp) apply (simp add:sp_cf_def) apply (rotate_tac -1, drule sym, simp add:cmp_def, thin_tac "pol_coeff B (l, \x. h (g x))", thin_tac "pol_coeff B (l, \j. h (f 0) \\<^sub>r\<^bsub>B\<^esub> h (u j))", thin_tac "pol_coeff S (l, \j. f 0 \\<^sub>r\<^bsub>S\<^esub> u j)", thin_tac "polyn_expr A Y l (l, \x. h (u x)) \ carrier A", thin_tac "pol_coeff B (l, \x. h (u x))", thin_tac "polyn_expr R X l (l, \j. f 0 \\<^sub>r\<^bsub>S\<^esub> u j) = polyn_expr R X l (l, g)") apply (frule_tac c = "(l, u)" and j = j in pol_coeff_mem, simp, simp) apply (simp add:rHom_tOp) apply ((rule allI)+, (rule impI), (erule conjE)+) apply (simp add:cf_h_polyn, frule_tac n = n and f = f in pol_coeff_pre, frule_tac n = "l + n" and f = g in pol_coeff_pre, frule_tac n = l and f = u and m = n and g = f in polyn_expr_tOp, assumption+, erule exE, (erule conjE)+, rotate_tac -1, drule sym) apply (drule_tac x = f in spec, drule_tac x = e in spec, simp, frule_tac n = n and f = f in polyn_Suc_split, simp del:npow_suc, thin_tac "polyn_expr R X (Suc n) (Suc n, f) = polyn_expr R X n (n, f) \ f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>") (* got polyn_expr A Y l (l, cmp h u) \\<^sub>r\<^bsub>A\<^esub> polyn_expr A Y n (n, cmp h f) = polyn_expr A Y (l + n) (l + n, cmp h e) *) apply (frule_tac c = "(Suc n, f)" in cf_h_coeff[of A B Y h], assumption+, simp del:npow_suc add:cf_h_pol_coeff) apply (frule_tac n = n and f = "cmp h f" in PolynRg.polyn_Suc_split[of A B Y], simp add:cf_h_pol_coeff, simp del:npow_suc, thin_tac "polyn_expr A Y (Suc n) (Suc n, cmp h f) = polyn_expr A Y n (n, cmp h f) \\<^bsub>A\<^esub> cmp h f (Suc n) \\<^sub>r\<^bsub>A\<^esub> Y^\<^bsup>A (Suc n)\<^esup>") apply (frule_tac c = "(l, u)" and k = l in polyn_mem, simp, frule_tac n = n and f = f in pol_coeff_pre, frule_tac c = "(n, f)" and k = n in polyn_mem, simp, frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem_R, simp, cut_tac X_mem_R, simp del:npow_suc, cut_tac n = "Suc n" in npClose[of X], assumption, frule_tac x = "f (Suc n)" and y = " X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed, assumption+, simp del:npow_suc add:ring_distrib1) apply (frule_tac c = "(l, u)" in cf_h_coeff[of A B Y h], assumption+, frule_tac c = "(Suc n, f)" in cf_h_coeff[of A B Y h], assumption+, frule_tac n = n and f = f in pol_coeff_pre, frule_tac c = "(n, f)" in cf_h_coeff[of A B Y h], assumption+, simp del:npow_suc add:cf_h_pol_coeff) apply (frule_tac c = "(l, cmp h u)" and k = l in PolynRg.polyn_mem, simp, simp, frule_tac c = "(n, cmp h f)" and k = n in PolynRg.polyn_mem, simp, simp) apply (frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem_R, simp, cut_tac X_mem_R, simp del:npow_suc, cut_tac n = "Suc n" in npClose[of X], assumption, frule_tac x = "f (Suc n)" and y = " X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed, assumption+, simp del:npow_suc add:ring_distrib1) apply (frule_tac c = "(Suc n, cmp h f)" and j = "Suc n" in PolynRg.pol_coeff_mem_R[of A B Y], simp del:npow_suc, simp del:npow_suc, frule_tac PolynRg.X_mem_R[of A B Y], simp del:npow_suc, frule_tac n = "Suc n" in Ring.npClose[of A Y], assumption, frule_tac x = "cmp h f (Suc n)" and y = " Y^\<^bsup>A (Suc n)\<^esup>" in Ring.ring_tOp_closed[of A], assumption+, simp del:npow_suc add:Ring.ring_distrib1) apply (frule_tac x1 = "polyn_expr R X l (l, u)" and y1 = "f (Suc n)" and z1 = " X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_assoc[THEN sym], assumption+, simp del:npow_suc, thin_tac "polyn_expr R X l (l, u) \\<^sub>r polyn_expr R X n (n, f) = polyn_expr R X (l + n) (l + n, e)", thin_tac "polyn_expr R X l (l, u) \\<^sub>r (f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>) = polyn_expr R X l (l, u) \\<^sub>r f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup>", simp only:ring_tOp_commute, frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp, simp del:npow_suc, simp del:npow_suc add:scalar_times_pol_expr) apply (frule_tac c = "(l, u)" and a = "f (Suc n)" in sp_cf_pol_coeff, assumption+, frule_tac c = "sp_cf S (f (Suc n)) (l, u)" and k = l in polyn_mem, simp add:sp_cf_len, simp only:ring_tOp_commute, frule_tac c1 = "sp_cf S (f (Suc n)) (l, u)" and j1 = "Suc n" in low_deg_terms_zero1[THEN sym], simp only:sp_cf_len, simp del: npow_suc) apply (frule_tac c = "sp_cf S (f (Suc n)) (l, u)" and n = "Suc n" in ext_cf_pol_coeff, frule_tac c = "(l + n, e)" and d = "ext_cf S (Suc n) (sp_cf S ( f (Suc n)) (l, u))" in polyn_add1, assumption+, simp del:npow_suc add:ext_cf_len sp_cf_len, cut_tac a = l and b = n in add.commute, simp del:npow_suc, thin_tac "polyn_expr R X (n + l) (n + l, e) \ polyn_expr R X (Suc (n + l)) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u))) = polyn_expr R X (Suc (n + l)) (add_cf S (n + l, e) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u))))", thin_tac "polyn_expr R X l (l, u) \ carrier R", thin_tac "polyn_expr R X n (n, f) \ carrier R", thin_tac "f (Suc n) \ carrier R", thin_tac "X \ carrier R", thin_tac "X^\<^bsup>R (Suc n)\<^esup> \ carrier R", thin_tac "f (Suc n) \\<^sub>r X^\<^bsup>R (Suc n)\<^esup> \ carrier R", thin_tac "polyn_expr R X l (sp_cf S (f (Suc n)) (l, u)) \ carrier R", thin_tac "X^\<^bsup>R (Suc n)\<^esup> \\<^sub>r polyn_expr R X l (sp_cf S (f (Suc n)) (l, u)) = polyn_expr R X (Suc (n + l)) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u)))") (* got polyn_expr R X (Suc (n + l)) (Suc (n + l), g) = polyn_expr R X (Suc (n + l)) (add_cf S (n + l, e) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u)))) *) apply (subst Ring.ring_tOp_assoc[THEN sym], assumption+, subst Ring.ring_tOp_commute, assumption+) apply (frule_tac c = "(Suc n, f)" in cf_h_coeff[of A B Y h], assumption+, simp only:cf_h_pol_coeff) apply (frule_tac c = "(Suc n, cmp h f)" and j = "Suc n" in PolynRg.pol_coeff_mem[of A B Y], assumption+, simp, simp del:npow_suc) apply (subst PolynRg.scalar_times_pol_expr[of A B Y], assumption+, simp) apply (frule_tac a = "cmp h f (Suc n)" and c = "(l, cmp h u)" in PolynRg.sp_cf_pol_coeff[of A B Y], assumption+) apply (frule_tac c = "sp_cf B (cmp h f (Suc n)) (l, cmp h u)" and k = l in PolynRg.polyn_mem[of A B Y], assumption, simp) apply (simp add:PolynRg.sp_cf_len) apply (subst Ring.ring_tOp_commute[of A], assumption+) apply (frule_tac c1 = "sp_cf B (cmp h f (Suc n)) (l, cmp h u)" and j1 = "Suc n" in PolynRg.low_deg_terms_zero1[of A B Y, THEN sym], assumption+) apply (simp del:npow_suc add:sp_cf_len PolynRg.sp_cf_len, thin_tac "Y^\<^bsup>A (Suc n)\<^esup> \\<^sub>r\<^bsub>A\<^esub> polyn_expr A Y l (sp_cf B (cmp h f (Suc n)) (l, cmp h u)) = polyn_expr A Y (Suc (l + n)) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u)))", thin_tac "polyn_expr A Y l (sp_cf B (cmp h f (Suc n)) (l, cmp h u)) \ carrier A") apply (frule_tac c = "sp_cf B (cmp h f (Suc n)) (l, cmp h u)" and n = "Suc n" in PolynRg.ext_cf_pol_coeff[of A B Y], assumption+) apply (frule_tac c = "(l + n, cmp h e)" and d = "ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))" in PolynRg.polyn_add1[of A B Y]) apply (frule_tac c = "(n + l, e)" in cf_h_coeff[of A B Y h], assumption+) apply (simp add:cf_h_pol_coeff[of A B Y h] add.commute, assumption) apply (simp add:PolynRg.ext_cf_len PolynRg.sp_cf_len) apply (cut_tac a = n and b = l in add.commute, simp) (** Now we got polyn_expr A Y (max (l + n) (Suc (l + n))) (add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u)))) = polyn_expr A Y (Suc (l + n)) (Suc (l + n), cmp h g) *) apply (frule_tac c = "(l + n, e)" and d = "ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u))" in add_cf_pol_coeff, assumption+) apply (frule_tac c = "(l + n, cmp h e)" and d = "ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))" in PolynRg.add_cf_pol_coeff[of A B Y]) apply (frule_tac c = "(l + n, e)" in cf_h_coeff[of A B Y h], assumption+) apply (simp add:cf_h_pol_coeff) apply simp apply (thin_tac "polyn_expr A Y l (l, cmp h u) \\<^sub>r\<^bsub>A\<^esub> polyn_expr A Y n (n, cmp h f) = polyn_expr A Y (l + n) (l + n, cmp h e)", thin_tac "polyn_expr A Y l (l, cmp h u) \ carrier A", thin_tac "polyn_expr A Y n (n, cmp h f) \ carrier A", thin_tac "Y \ carrier A", thin_tac "Y^\<^bsup>A n\<^esup> \\<^sub>r\<^bsub>A\<^esub> Y \ carrier A", thin_tac "cmp h f (Suc n) \\<^sub>r\<^bsub>A\<^esub> (Y^\<^bsup>A n\<^esup> \\<^sub>r\<^bsub>A\<^esub> Y) \ carrier A", thin_tac "polyn_expr A Y (l + n) (l + n, cmp h e) \\<^bsub>A\<^esub> polyn_expr A Y (Suc (l + n)) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))) = polyn_expr A Y (Suc (l + n)) (add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))))") apply (frule_tac c = "(Suc (l + n), g)" in cf_h_coeff[of A B Y h], assumption+) apply (simp add:cf_h_pol_coeff) apply (frule_tac c = "(Suc (l + n), g)" and d = "add_cf S (l + n, e) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u)))" in pol_expr_unique2) apply assumption apply (simp add:add_cf_len) apply (simp add:ext_cf_len sp_cf_len) apply (simp add:add_cf_len ext_cf_len sp_cf_len) apply (cut_tac a = n and b = l in add.commute, simp, thin_tac "pol_coeff B (add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))))", thin_tac "pol_coeff B (Suc (l + n), cmp h g)", thin_tac "pol_coeff S (add_cf S (l + n, e) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u))))", thin_tac "pol_coeff B (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u)))", thin_tac "pol_coeff B (sp_cf B (cmp h f (Suc n)) (l, cmp h u))", thin_tac "polyn_expr R X (Suc (l + n)) (Suc (l + n), g) = polyn_expr R X (Suc (l + n)) (add_cf S (l + n, e) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u))))") apply (rule sym) apply (frule_tac c = "(Suc (l + n), g)" in cf_h_coeff[of A B Y h], assumption+, frule_tac c = "(l + n, e)" in cf_h_coeff[of A B Y h], assumption+) apply (frule_tac c = "(l, cmp h u)" and a = "cmp h f (Suc n)" in PolynRg.sp_cf_pol_coeff[of A B Y], assumption+) apply (cut_tac c = "(l + n, cmp h e)" and d = "ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))" in PolynRg.add_cf_pol_coeff, assumption+) apply (simp add:cf_h_pol_coeff) apply (rule PolynRg.ext_cf_pol_coeff, assumption+) apply (frule_tac c = "(Suc (l + n), cmp h g)" and d = "add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u)))" in PolynRg.pol_expr_unique2[of A B Y]) apply (simp add:cf_h_pol_coeff) apply assumption apply (simp add:add_cf_len) apply (frule_tac n = "l + n" and f = e in cf_h_pol_coeff[of A B Y h], assumption+) apply (frule_tac c = "sp_cf B (cmp h f (Suc n)) (l, cmp h u)" and n = "Suc n" in PolynRg.ext_cf_pol_coeff[of A B Y], assumption+) apply (simp add:PolynRg.add_cf_len) apply (simp add:PolynRg.ext_cf_len) apply (simp add:PolynRg.sp_cf_len) apply (simp add:PolynRg.add_cf_len) apply (frule_tac c = "sp_cf B (cmp h f (Suc n)) (l, cmp h u)" and n = "Suc n" in PolynRg.ext_cf_pol_coeff[of A B Y], assumption+) apply (frule_tac c = "(l + n, cmp h e)" and d = "ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))" in PolynRg.add_cf_pol_coeff[of A B Y]) apply (simp add:cf_h_pol_coeff, assumption) apply (simp add:cf_h_pol_coeff) apply (simp add:PolynRg.add_cf_len) apply (simp add:PolynRg.ext_cf_len) apply (simp add:PolynRg.sp_cf_len) apply (cut_tac a = n and b = l in add.commute, simp) (* we got \j\Suc (l + n). cmp h g j = snd (add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u)))) j *) apply (thin_tac "(polyn_expr A Y (Suc (l + n)) (Suc (l + n), cmp h g) = polyn_expr A Y (Suc (l + n)) (add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u))))) = (\j\Suc (l + n). cmp h g j = snd (add_cf B (l + n, cmp h e) (ext_cf B (Suc n) (sp_cf B (cmp h f (Suc n)) (l, cmp h u)))) j)") apply (rule allI, rule impI) apply (subst cmp_def)+ apply (drule_tac a = j in forall_spec, simp, simp, thin_tac "g j = snd (add_cf S (l + n, e) (ext_cf S (Suc n) (sp_cf S (f (Suc n)) (l, u)))) j") apply (case_tac "j = Suc (l+n)", simp) apply ((subst add_cf_def)+, simp add:ext_cf_len, simp add:sp_cf_len, simp add:cmp_def PolynRg.ext_cf_len, simp add:PolynRg.sp_cf_len, (subst ext_cf_def)+, simp add:sp_cf_len sliden_def, (subst sp_cf_def)+, simp, frule_tac c = "(l, u)" and j = l in pol_coeff_mem, simp, simp, simp add:rHom_tOp) apply (frule_tac m = j and n = "Suc (l + n)" in noteq_le_less, assumption, thin_tac "j \ Suc (l + n)", thin_tac "j \ Suc (l + n)", (subst add_cf_def)+, simp add:ext_cf_len sp_cf_len, simp add:cmp_def, simp add:PolynRg.ext_cf_len PolynRg.sp_cf_len, (subst ext_cf_def)+, simp add:sp_cf_len, (subst sp_cf_def)+, simp add:sliden_def, frule_tac x = j and n = "l + n" in Suc_less_le) apply (rule conjI) apply (rule impI, frule_tac x = j and y = "Suc (l + n)" in less_imp_le, frule_tac m = j and n = "Suc (l + n)" and l = "Suc n" in diff_le_mono, simp, frule_tac c = "(l, u)" and j = "j - Suc n" in pol_coeff_mem, simp, frule_tac c = "(l + n, e)" and j = j in pol_coeff_mem, simp, simp, frule_tac x = "f (Suc n)" and y = "u (j - Suc n)" in Ring.ring_tOp_closed[of S], assumption+, simp add:rHom_add rHom_tOp) apply (rule impI) apply (frule_tac c = "(l + n, e)" and j = j in pol_coeff_mem, simp, simp, frule_tac Ring.ring_zero[of S], simp add:rHom_add rHom_0_0) done lemma (in PolynRg) erH_multTr1:"\PolynRg A B Y; h \ rHom S B; pol_coeff S c; pol_coeff S d; pol_coeff S e; fst e = fst c + fst d; (polyn_expr R X (fst c) c) \\<^sub>r (polyn_expr R X (fst d) d) = polyn_expr R X ((fst c) + (fst d)) e \ \ (polyn_expr A Y (fst c) (cf_h h c)) \\<^sub>r\<^bsub>A\<^esub> (polyn_expr A Y (fst d) (cf_h h d)) = (polyn_expr A Y (fst e) (cf_h h e))" by (cases d, cases e) (simp add: erH_multTr) lemma (in PolynRg) erHomTr0:"\PolynRg A B Y; h \ rHom S B; x \ carrier R\ \ erH R S X A B Y h (-\<^sub>a x) = -\<^sub>a\<^bsub>A\<^esub> (erH R S X A B Y h x)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule PolynRg.is_Ring[of A B Y], frule Ring.ring_is_ag[of A], frule PolynRg.subring[of A B Y], frule Ring.subring_Ring[of A B], assumption+) apply (case_tac "x = \\<^bsub>R\<^esub>", simp add:aGroup.ag_inv_zero, simp add:erH_rHom_0[of A B Y h], frule Ring.ring_is_ag[of A], simp add:aGroup.ag_inv_zero) apply (simp add:erH_def erh_def) apply (simp add:aGroup.ag_mOp_closed) apply (frule_tac p = x in s_cf_expr, assumption+, (erule conjE)+) apply (frule_tac x = x in aGroup.ag_mOp_closed, assumption+, frule_tac p = "-\<^sub>a x" in s_cf_expr, thin_tac "x = polyn_expr R X (fst (s_cf R S X x)) (s_cf R S X x)", rule contrapos_pp, simp+, frule_tac x = x in aGroup.ag_inv_inv, simp, simp add:aGroup.ag_inv_zero, (erule conjE)+) apply (frule_tac c = "s_cf R S X (-\<^sub>a x)" in cf_h_coeff[of A B Y h], assumption+, frule_tac c = "s_cf R S X x" in cf_h_coeff[of A B Y h], assumption+) apply (frule polyn_minus_m_cf[of "s_cf R S X x" "fst (s_cf R S X x)"], simp) apply (cut_tac a = "-\<^sub>a x" and b = "polyn_expr R X (fst (s_cf R S X (-\<^sub>a x))) (s_cf R S X (-\<^sub>a x))" and c = "polyn_expr R X (fst (s_cf R S X x)) (m_cf S (s_cf R S X x))" in box_equation, assumption, simp, thin_tac "x = polyn_expr R X (fst (s_cf R S X x)) (s_cf R S X x)", thin_tac "-\<^sub>a x = polyn_expr R X (fst (s_cf R S X (-\<^sub>a x))) (s_cf R S X (-\<^sub>a x))", thin_tac "-\<^sub>a (polyn_expr R X (fst (s_cf R S X x)) (s_cf R S X x)) = polyn_expr R X (fst (s_cf R S X x)) (m_cf S (s_cf R S X x))", frule m_cf_pol_coeff[of "s_cf R S X x"]) apply (subst PolynRg.polyn_minus_m_cf[of A B Y], assumption+, simp add:cf_h_len) apply (frule_tac c = "cf_h h (s_cf R S X x)" in PolynRg.m_cf_pol_coeff[of A B Y], assumption, frule PolynRg.pol_expr_unique2[of A B Y "cf_h h (s_cf R S X (-\<^sub>a x))" "m_cf B (cf_h h (s_cf R S X x))"], assumption+) apply (simp add:cf_h_len) apply (simp add:PolynRg.m_cf_len cf_h_len) apply (simp add:s_cf_deg[THEN sym, of x], cut_tac ring_zero, frule aGroup.ag_inv_inj[of R x \], assumption+, simp only:aGroup.ag_inv_zero, subst s_cf_deg[THEN sym, of "-\<^sub>a x"], assumption+, simp add:deg_minus_eq) apply (simp add:cf_h_len PolynRg.m_cf_len, thin_tac "(polyn_expr A Y (fst (s_cf R S X (-\<^sub>a x))) (cf_h h (s_cf R S X (-\<^sub>a x))) = polyn_expr A Y (fst (s_cf R S X x)) (m_cf B (cf_h h (s_cf R S X x)))) = (\j\fst (s_cf R S X (-\<^sub>a x)). snd (cf_h h (s_cf R S X (-\<^sub>a x))) j = snd (m_cf B (cf_h h (s_cf R S X x))) j)") apply (rule allI, rule impI, subst m_cf_def) apply ((subst cf_h_def)+, simp add:cmp_def) apply (thin_tac "snd (s_cf R S X (-\<^sub>a x)) (fst (s_cf R S X (-\<^sub>a x))) \ \\<^bsub>S\<^esub>", thin_tac "pol_coeff B (cf_h h (s_cf R S X (-\<^sub>a x)))", thin_tac "pol_coeff B (cf_h h (s_cf R S X x))", thin_tac "pol_coeff B (m_cf B (cf_h h (s_cf R S X x)))") apply (frule m_cf_pol_coeff[of "s_cf R S X x"]) apply (frule pol_expr_unique2[of "s_cf R S X (-\<^sub>a x)" "m_cf S (s_cf R S X x)"], assumption+, simp add:m_cf_len cf_h_len) apply (simp add:s_cf_deg[THEN sym, of x], cut_tac ring_zero, frule aGroup.ag_inv_inj[of R x \], assumption+, simp only:aGroup.ag_inv_zero, subst s_cf_deg[THEN sym, of "-\<^sub>a x"], assumption+, simp add:deg_minus_eq, simp add:m_cf_len) apply (drule_tac a = j in forall_spec, assumption, thin_tac "snd (s_cf R S X (-\<^sub>a x)) j = snd (m_cf S (s_cf R S X x)) j", thin_tac "polyn_expr R X (fst (s_cf R S X (-\<^sub>a x))) (s_cf R S X (-\<^sub>a x)) = polyn_expr R X (fst (s_cf R S X x)) (m_cf S (s_cf R S X x))") apply (cut_tac ring_zero, frule aGroup.ag_inv_inj[of R x \], assumption+, simp only:aGroup.ag_inv_zero, (simp add:s_cf_deg[THEN sym, of "-\<^sub>a x"] deg_minus_eq, simp add:s_cf_deg[of x]) ) apply (frule_tac j = j in pol_coeff_mem[of "s_cf R S X x"], assumption+) apply (subst m_cf_def, simp) apply (simp add:rHom_inv_inv) done lemma (in PolynRg) erHomTr1:"\PolynRg A B Y; h \ rHom S B; a \ carrier R; b \ carrier R; a \ \; b \ \; a \ b \ \; deg_n R S X a = deg_n R S X b\ \ erH R S X A B Y h (a \ b) = erH R S X A B Y h a \\<^bsub>A\<^esub> (erH R S X A B Y h b)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule PolynRg.subring[of A B Y], frule_tac x = a and y = b in aGroup.ag_pOp_closed[of "R"], assumption+, simp add:erH_def erh_def) apply (frule s_cf_expr[of a], assumption, frule s_cf_expr[of b], assumption, frule s_cf_expr[of "a \ b"], assumption, (erule conjE)+, simp add:s_cf_deg) apply (frule polyn_add1[of "s_cf R S X a" "s_cf R S X b"], assumption+) apply (cut_tac a = "a \ b" and b = "polyn_expr R X (fst (s_cf R S X (a \ b))) (s_cf R S X (a \ b))" and c = "polyn_expr R X (max (fst (s_cf R S X a)) (fst (s_cf R S X b))) (add_cf S (s_cf R S X a) (s_cf R S X b))" in box_equation, drule sym, drule sym, drule sym, simp, drule sym, drule sym, drule sym, simp, thin_tac "polyn_expr R X (fst (s_cf R S X a)) (s_cf R S X a) \ polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X b) = polyn_expr R X (max (fst (s_cf R S X a)) (fst (s_cf R S X b))) (add_cf S (s_cf R S X a) (s_cf R S X b))", thin_tac "a = polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X a)", thin_tac "b = polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X b)", thin_tac "a \ b = polyn_expr R X (fst (s_cf R S X (a \ b))) (s_cf R S X (a \ b))") apply simp apply (frule cf_h_coeff[of A B Y h "s_cf R S X a"], assumption+, frule cf_h_coeff[of A B Y h "s_cf R S X b"], assumption+, frule cf_h_coeff[of A B Y h "s_cf R S X (a \ b)"], assumption+) apply (frule PolynRg.polyn_add1[of A B Y "cf_h h (s_cf R S X a)" "cf_h h (s_cf R S X b)"], assumption+, simp add:cf_h_len, thin_tac "polyn_expr A Y (fst (s_cf R S X b)) (cf_h h (s_cf R S X a)) \\<^bsub>A\<^esub> polyn_expr A Y (fst (s_cf R S X b)) (cf_h h (s_cf R S X b)) = polyn_expr A Y (fst (s_cf R S X b)) (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b)))", frule PolynRg.add_cf_pol_coeff[of A B Y "cf_h h (s_cf R S X a)" "cf_h h (s_cf R S X b)"], assumption+) apply (case_tac "fst (s_cf R S X (a \\<^bsub>R\<^esub> b)) = fst (s_cf R S X b)") apply (frule PolynRg.pol_expr_unique2[of A B Y "cf_h h (s_cf R S X (a \ b))" "add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))"], assumption+) apply (simp add:cf_h_len add_cf_len, simp add:PolynRg.add_cf_len cf_h_len) apply (simp add:PolynRg.add_cf_len cf_h_len, thin_tac "(polyn_expr A Y (fst (s_cf R S X b)) (cf_h h (s_cf R S X (a \ b))) = polyn_expr A Y (fst (s_cf R S X b)) (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b)))) = (\j\fst (s_cf R S X b). snd (cf_h h (s_cf R S X (a \ b))) j = snd (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))) j)") apply (frule pol_expr_unique2[of "s_cf R S X (a \ b)" "add_cf S (s_cf R S X a) (s_cf R S X b)"]) apply (simp add:add_cf_pol_coeff) apply (simp add:add_cf_len, simp add:add_cf_len, thin_tac "polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X (a \ b)) = polyn_expr R X (fst (s_cf R S X b)) (add_cf S (s_cf R S X a) (s_cf R S X b))") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, subst add_cf_def, simp add:cf_h_len, (subst cf_h_def)+, (subst cmp_def)+, simp, thin_tac "snd (s_cf R S X (a \ b)) j = snd (add_cf S (s_cf R S X a) (s_cf R S X b)) j") apply (subst add_cf_def, simp) apply (frule_tac j = j in pol_coeff_mem[of "s_cf R S X a"], simp, frule_tac j = j in pol_coeff_mem[of "s_cf R S X b"], simp, simp add:rHom_add) apply (frule s_cf_deg[of a], assumption, frule s_cf_deg[of b], assumption, frule s_cf_deg[of "a \ b"], assumption, frule deg_pols_add2[of a b], assumption+, simp add:deg_def, simp add:deg_def ale_natle, frule_tac m = "fst (s_cf R S X (a \ b))" and n = "fst (s_cf R S X b)" in noteq_le_less, assumption+) apply (frule pol_expr_unique3[of "s_cf R S X (a \ b)" "add_cf S (s_cf R S X a) (s_cf R S X b)"], simp add:add_cf_pol_coeff, simp add:add_cf_len, simp add:add_cf_len, thin_tac "polyn_expr R X (fst (s_cf R S X (a \ b))) (s_cf R S X (a \ b)) = polyn_expr R X (fst (s_cf R S X b)) (add_cf S (s_cf R S X a) (s_cf R S X b))") apply (frule PolynRg.pol_expr_unique3[of A B Y "cf_h h (s_cf R S X (a \ b))" "add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))"], assumption+, simp add:cf_h_len PolynRg.add_cf_len, simp add:PolynRg.add_cf_len cf_h_len, thin_tac "(polyn_expr A Y (fst (s_cf R S X (a \ b))) (cf_h h (s_cf R S X (a \ b))) = polyn_expr A Y (fst (s_cf R S X b)) (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b)))) = ((\j\fst (s_cf R S X (a \ b)). snd (cf_h h (s_cf R S X (a \ b))) j = snd (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))) j) \ (\j\nset (Suc (fst (s_cf R S X (a \ b)))) (fst (s_cf R S X b)). snd (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))) j = \\<^bsub>B\<^esub>))", thin_tac "pol_coeff B (cf_h h (s_cf R S X a))", thin_tac "pol_coeff B (cf_h h (s_cf R S X b))", thin_tac "pol_coeff B (cf_h h (s_cf R S X (a \ b)))", thin_tac "pol_coeff B (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b)))", thin_tac "deg_n R S X a = fst (s_cf R S X b)", thin_tac "deg_n R S X b = fst (s_cf R S X b)", thin_tac "deg_n R S X (a \ b) = fst (s_cf R S X (a \ b))") apply (rule conjI, erule conjE, thin_tac "\j\nset (Suc (fst (s_cf R S X (a \ b)))) (fst (s_cf R S X b)). snd (add_cf S (s_cf R S X a) (s_cf R S X b)) j = \\<^bsub>S\<^esub>") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, (subst cf_h_def)+, (subst cmp_def)+, simp, (subst add_cf_def)+, simp, frule_tac j = j in pol_coeff_mem[of "s_cf R S X a"], simp, frule_tac j = j in pol_coeff_mem[of "s_cf R S X b"], simp, simp add:rHom_add) apply (erule conjE, thin_tac "\j\fst (s_cf R S X (a \ b)). snd (s_cf R S X (a \ b)) j = snd (add_cf S (s_cf R S X a) (s_cf R S X b)) j") apply (rule ballI, drule_tac x = j in bspec, assumption, simp add:add_cf_def cf_h_len, simp add:cf_h_def cmp_def, frule_tac j = j in pol_coeff_mem[of "s_cf R S X a"], simp add:nset_def, frule_tac j = j in pol_coeff_mem[of "s_cf R S X b"], simp add:nset_def) apply (subst rHom_add[THEN sym, of h S B], assumption+, simp, (frule PolynRg.is_Ring[of A B Y], frule Ring.subring_Ring[of A B], assumption), simp add:rHom_0_0[of S B]) done lemma (in PolynRg) erHomTr2:"\PolynRg A B Y; h \ rHom S B; a \ carrier R; b \ carrier R; a \ \; b \ \; a \ b \ \; deg_n R S X a < deg_n R S X b\ \ erH R S X A B Y h (a \ b) = erH R S X A B Y h a \\<^bsub>A\<^esub> (erH R S X A B Y h b)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring, frule PolynRg.subring[of A B Y], frule_tac x = a and y = b in aGroup.ag_pOp_closed[of "R"], assumption+, simp add:erH_def erh_def) apply (frule s_cf_expr[of a], assumption, frule s_cf_expr[of b], assumption, frule s_cf_expr[of "a \ b"], assumption, (erule conjE)+, frule polyn_deg_add1[of a b], assumption+, simp add:s_cf_deg) apply (frule polyn_add1[of "s_cf R S X a" "s_cf R S X b"], assumption+) apply (cut_tac a = "a \ b" and b = "polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X (a \ b))" and c = "polyn_expr R X (max (fst (s_cf R S X a)) (fst (s_cf R S X b))) (add_cf S (s_cf R S X a) (s_cf R S X b))" in box_equation, drule sym, drule sym, drule sym, simp, drule sym, drule sym, drule sym, simp, thin_tac "polyn_expr R X (fst (s_cf R S X a)) (s_cf R S X a) \ polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X b) = polyn_expr R X (max (fst (s_cf R S X a)) (fst (s_cf R S X b))) (add_cf S (s_cf R S X a) (s_cf R S X b))", thin_tac "a = polyn_expr R X (fst (s_cf R S X a)) (s_cf R S X a)", thin_tac "b = polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X b)", thin_tac "a \ b = polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X (a \ b))", simp) apply (frule cf_h_coeff[of A B Y h "s_cf R S X a"], assumption+, frule cf_h_coeff[of A B Y h "s_cf R S X b"], assumption+, frule cf_h_coeff[of A B Y h "s_cf R S X (a \ b)"], assumption+) apply (frule PolynRg.polyn_add1[of A B Y "cf_h h (s_cf R S X a)" "cf_h h (s_cf R S X b)"], assumption+, simp add:cf_h_len) apply (thin_tac "polyn_expr A Y (fst (s_cf R S X a)) (cf_h h (s_cf R S X a)) \\<^bsub>A\<^esub> polyn_expr A Y (fst (s_cf R S X b)) (cf_h h (s_cf R S X b)) = polyn_expr A Y (fst (s_cf R S X b)) (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b)))", frule PolynRg.add_cf_pol_coeff[of A B Y "cf_h h (s_cf R S X a)" "cf_h h (s_cf R S X b)"], assumption+) apply (frule PolynRg.pol_expr_unique2[of A B Y "cf_h h (s_cf R S X (a \ b))" "add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))"], assumption+, simp add:cf_h_len add_cf_len, simp add:PolynRg.add_cf_len cf_h_len, simp add:PolynRg.add_cf_len cf_h_len, thin_tac "(polyn_expr A Y (fst (s_cf R S X b)) (cf_h h (s_cf R S X (a \ b))) = polyn_expr A Y (fst (s_cf R S X b)) (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b)))) = (\j\fst (s_cf R S X b). snd (cf_h h (s_cf R S X (a \ b))) j = snd (add_cf B (cf_h h (s_cf R S X a)) (cf_h h (s_cf R S X b))) j)") apply (frule pol_expr_unique2[of "s_cf R S X (a \ b)" "add_cf S (s_cf R S X a) (s_cf R S X b)"], simp add:add_cf_pol_coeff, simp add:add_cf_len, simp add:add_cf_len) apply (thin_tac "polyn_expr R X (fst (s_cf R S X b)) (s_cf R S X (a \ b)) = polyn_expr R X (fst (s_cf R S X b)) (add_cf S (s_cf R S X a) (s_cf R S X b))") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, subst add_cf_def, simp add:cf_h_len, (subst cf_h_def)+, (subst cmp_def)+, simp, subst add_cf_def, simp) apply (case_tac "j \ fst (s_cf R S X a)", simp) apply (frule_tac j = j in pol_coeff_mem[of "s_cf R S X a"], simp, frule_tac j = j in pol_coeff_mem[of "s_cf R S X b"], simp, simp add:rHom_add) apply simp apply (subst add_cf_def, simp) done lemma (in PolynRg) erH_rHom:"\Idomain S; PolynRg A B Y; h \ rHom S B\ \ erH R S X A B Y h \ pHom R S X, A B Y" apply (frule Idomain.idom_is_ring[of "S"], cut_tac subring, cut_tac polyn_ring_integral, simp, frule PolynRg.subring[of A B Y], frule PolynRg.is_Ring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (simp add:polyn_Hom_def, rule conjI) prefer 2 apply (cut_tac polyn_ring_X_nonzero, cut_tac X_mem_R, rule conjI, rule subsetI, simp add:image_def, erule bexE) apply (case_tac "xa = \\<^bsub>S\<^esub>", simp add:Subring_zero_ring_zero, simp add:erH_rHom_0, simp add:Ring.Subring_zero_ring_zero[THEN sym, of A B], simp add:Ring.ring_zero[of B]) apply (simp add:Subring_zero_ring_zero, frule_tac x = xa in mem_subring_mem_ring, assumption, frule_tac p = xa in s_cf_expr, assumption+, (erule conjE)+, frule_tac p1 = xa in s_cf_deg[THEN sym], assumption+, frule_tac p1 = xa in pol_of_deg0[THEN sym], assumption+, simp) apply (simp add:erH_def erh_def, subst polyn_expr_def, simp, frule_tac c = "s_cf R S X xa" in cf_h_coeff[of A B Y h], assumption+, frule_tac c = "cf_h h (s_cf R S X xa)" and j = 0 in PolynRg.pol_coeff_mem[of A B Y], assumption, simp, frule_tac c = "cf_h h (s_cf R S X xa)" and j = 0 in PolynRg.pol_coeff_mem_R[of A B Y], assumption, simp, simp add:Ring.ring_r_one[of A]) apply (cut_tac pol_expr_of_X, cut_tac special_cf_pol_coeff, frule ext_cf_pol_coeff[of "C\<^sub>0 S" "Suc 0"]) apply (simp add:erH_def erh_def) apply (cut_tac deg_n_of_X) apply (frule s_cf_expr[of X], assumption+, (erule conjE)+, frule_tac a = X and b = "polyn_expr R X (Suc 0) (ext_cf S (Suc 0) (C\<^sub>0 S))" and c = "polyn_expr R X (fst (s_cf R S X X)) (s_cf R S X X)" in box_equation, assumption+, thin_tac "X = polyn_expr R X (Suc 0) (ext_cf S (Suc 0) (C\<^sub>0 S))", thin_tac "X = polyn_expr R X (fst (s_cf R S X X)) (s_cf R S X X)") apply (rule sym, subst PolynRg.pol_expr_of_X[of A B Y], assumption+, frule s_cf_deg[of X], assumption+, simp) apply (frule pol_expr_unique2[of "ext_cf S (Suc 0) (C\<^sub>0 S)" "s_cf R S X X"], assumption+, simp add:ext_cf_len special_cf_len, simp add:ext_cf_len special_cf_len) apply (frule cf_h_coeff[of A B Y h "ext_cf S (Suc 0) (C\<^sub>0 S)"], assumption+, frule cf_h_coeff[of A B Y h "s_cf R S X X"], assumption+) apply (frule PolynRg.pol_expr_unique2[of A B Y "cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S))" "cf_h h (s_cf R S X X)"], assumption+, simp add:cf_h_len ext_cf_len special_cf_len, simp add:cf_h_len ext_cf_len special_cf_len) apply (simp add:cf_h_special_cf) apply (thin_tac "(polyn_expr A Y (Suc 0) (ext_cf B (Suc 0) (C\<^sub>0 B)) = polyn_expr A Y (Suc 0) (cf_h h (s_cf R S X X))) = (\j\Suc 0. snd (cf_h h (ext_cf S (Suc 0) (C\<^sub>0 S))) j = snd (cf_h h (s_cf R S X X)) j)") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, (subst cf_h_def)+, (subst cmp_def)+, simp) apply (subst rHom_def, simp, cut_tac ring_is_ag, frule Ring.ring_is_ag[of A]) apply (rule conjI) apply (subst aHom_def, simp) apply (rule conjI) apply (simp add:erH_mem) apply (rule conjI, simp add:erH_def erh_def extensional_def) apply (rule ballI)+ apply (case_tac "a = \\<^bsub>R\<^esub>", case_tac "b = \\<^bsub>R\<^esub>", simp) apply (simp add:erH_rHom_0, frule Ring.ring_is_ag[of A], frule Ring.ring_zero[of A], simp add:aGroup.ag_r_zero) apply (simp add:erH_rHom_0, simp add:erH_rHom_0) apply (frule_tac p = b in erH_mem[of A B Y h], assumption+) apply (simp add:aGroup.ag_l_zero) apply (case_tac "b = \\<^bsub>R\<^esub>", simp) apply (simp add:erH_rHom_0, frule_tac p = a in erH_mem[of A B Y h], assumption+, simp add:aGroup.ag_r_zero) apply (case_tac "a \\<^bsub>R\<^esub> b = \\<^bsub>R\<^esub>", simp add:erH_rHom_0) apply (frule_tac x = a and y = b in aGroup.ag_inv_unique[of R], assumption+, simp, thin_tac "b = -\<^sub>a a") apply (subst erHomTr0[of A B Y h], assumption+, frule_tac p = a in erH_mem[of A B Y h], assumption+, simp add:aGroup.ag_r_inv1) apply (case_tac "deg_n R S X a = deg_n R S X b", simp add:erHomTr1[of A B Y h]) apply (cut_tac y = "deg_n R S X a" and x = "deg_n R S X b" in less_linear, simp) apply (erule disjE) apply (subst aGroup.ag_pOp_commute, assumption+, frule_tac p = a in erH_mem[of A B Y h], assumption+, frule_tac p = b in erH_mem[of A B Y h], assumption+, subst aGroup.ag_pOp_commute[of A], assumption+, rule erHomTr2[of A B Y h], assumption+, simp add:aGroup.ag_pOp_commute, assumption) apply(rule erHomTr2[of A B Y h], assumption+) apply (simp add:erH_rHomTr2) apply (rule ballI)+ apply (case_tac "x = \\<^bsub>R\<^esub>", simp add:ring_times_0_x erH_rHom_0, frule_tac p = y in erH_mem[of A B Y h], assumption+, simp add:Ring.ring_times_0_x[of A]) apply (case_tac "y = \\<^bsub>R\<^esub>", simp add:ring_times_x_0 erH_rHom_0, frule_tac p = x in erH_mem[of A B Y h], assumption+, simp add:Ring.ring_times_x_0[of A]) apply (frule_tac p = x in s_cf_expr, assumption+, frule_tac p = y in s_cf_expr, assumption+, frule_tac x = x and y = y in ring_tOp_closed, assumption+, frule_tac p = "x \\<^sub>r y" in s_cf_expr, simp add:Idomain.idom_tOp_nonzeros, (erule conjE)+) apply (frule_tac c = "s_cf R S X x" and d = "s_cf R S X y" in polyn_expr_tOp_c, assumption+, erule exE, (erule conjE)+, cut_tac a = "x \\<^sub>r y" and b = "polyn_expr R X (fst (s_cf R S X (x \\<^sub>r y))) (s_cf R S X (x \\<^sub>r y))" and c = "polyn_expr R X (fst e) e" in box_equation) apply assumption apply (thin_tac "x \\<^sub>r y = polyn_expr R X (fst (s_cf R S X (x \\<^sub>r y))) (s_cf R S X (x \\<^sub>r y))") apply (drule sym, drule sym, simp) apply (thin_tac "x = polyn_expr R X (fst (s_cf R S X x)) (s_cf R S X x)", thin_tac "y = polyn_expr R X (fst (s_cf R S X y)) (s_cf R S X y)", thin_tac "x \\<^sub>r y = polyn_expr R X (fst (s_cf R S X (x \\<^sub>r y))) (s_cf R S X (x \\<^sub>r y))") apply ((subst erH_def)+, (subst erh_def)+, simp) apply (frule_tac c = "s_cf R S X x" and d = "s_cf R S X y" and e = e in erH_multTr1[of A B Y h], assumption+, simp, simp, thin_tac "polyn_expr A Y (fst (s_cf R S X x)) (cf_h h (s_cf R S X x)) \\<^sub>r\<^bsub>A\<^esub> polyn_expr A Y (fst (s_cf R S X y)) (cf_h h (s_cf R S X y)) = polyn_expr A Y (fst (s_cf R S X x) + fst (s_cf R S X y)) (cf_h h e)") apply (rotate_tac -1, drule sym, simp, thin_tac "polyn_expr R X (fst (s_cf R S X x)) (s_cf R S X x) \\<^sub>r polyn_expr R X (fst (s_cf R S X y)) (s_cf R S X y) = polyn_expr R X (fst (s_cf R S X (x \\<^sub>r y))) (s_cf R S X (x \\<^sub>r y))", rotate_tac -1, drule sym) apply (frule_tac p = x in s_cf_deg, assumption, frule_tac p = y in s_cf_deg, assumption, frule_tac x = x and y = y in Idomain.idom_tOp_nonzeros[of R], assumption+, frule_tac p = "x \\<^sub>r y" in s_cf_deg, assumption+) apply (frule_tac p = x and q = y in deg_mult_pols, assumption+, (erule conjE)+, simp, thin_tac "deg_n R S X x = fst (s_cf R S X x)", thin_tac "deg_n R S X y = fst (s_cf R S X y)", thin_tac "deg_n R S X (x \\<^sub>r y) = fst (s_cf R S X (x \\<^sub>r y))", rotate_tac -2, drule sym) apply (simp add:cf_h_len) apply (frule_tac c = "s_cf R S X (x \\<^sub>r y)" in cf_h_coeff[of A B Y h], assumption+, frule_tac c = e in cf_h_coeff[of A B Y h], assumption+) apply (frule_tac c = "cf_h h (s_cf R S X (x \\<^sub>r y))" and d = "cf_h h e" in PolynRg.pol_expr_unique2[of A B Y], assumption+, simp add:cf_h_len, simp add:cf_h_len, thin_tac "(polyn_expr A Y (fst (s_cf R S X x) + fst (s_cf R S X y)) (cf_h h (s_cf R S X (x \\<^sub>r y))) = polyn_expr A Y (fst (s_cf R S X x) + fst (s_cf R S X y)) (cf_h h e)) = (\j\fst (s_cf R S X x) + fst (s_cf R S X y). snd (cf_h h (s_cf R S X (x \\<^sub>r y))) j = snd (cf_h h e) j)") apply (frule_tac c = "s_cf R S X (x \\<^sub>r y)" and d = e in pol_expr_unique2, assumption+, simp, simp, thin_tac "polyn_expr R X (fst (s_cf R S X x) + fst (s_cf R S X y)) (s_cf R S X (x \\<^sub>r y)) = polyn_expr R X (fst (s_cf R S X x) + fst (s_cf R S X y)) e") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, subst cf_h_def, subst cmp_def, simp, subst cf_h_def, subst cmp_def, simp) done lemma (in PolynRg) erH_q_rHom:"\Idomain S; maximal_ideal S P; PolynRg R' (S /\<^sub>r P) Y\ \ erH R S X R' (S /\<^sub>r P) Y (pj S P) \ pHom R S X, R' (S /\<^sub>r P) Y" apply (frule Idomain.idom_is_ring[of "S"], frule Ring.qring_ring[of S P], simp add:Ring.maximal_ideal_ideal, rule erH_rHom[of R' "S /\<^sub>r P" Y "pj S P"], assumption+) apply (rule pj_Hom[of S P], assumption+, simp add:Ring.maximal_ideal_ideal) done lemma (in PolynRg) erH_add:"\Idomain S; PolynRg A B Y; h \ rHom S B; p \ carrier R; q \ carrier R\ \ erH R S X A B Y h (p \ q) = (erH R S X A B Y h p) \\<^bsub>A\<^esub> (erH R S X A B Y h q)" apply (frule erH_rHom[of A B Y h], assumption+) apply (simp add:polyn_Hom_def, (erule conjE)+) apply (simp add:rHom_add) done lemma (in PolynRg) erH_minus:"\Idomain S; PolynRg A B Y; h \ rHom S B; p \ carrier R\ \ erH R S X A B Y h (-\<^sub>a p) = -\<^sub>a\<^bsub>A\<^esub> (erH R S X A B Y h p)" apply (frule erH_rHom[of A B Y h], assumption+, simp add:polyn_Hom_def, (erule conjE)+) apply (frule PolynRg.is_Ring[of A B Y]) apply (rule rHom_inv_inv[of R A p "erH R S X A B Y h"]) apply (cut_tac is_Ring, assumption+) done lemma (in PolynRg) erH_mult:"\Idomain S; PolynRg A B Y; h \ rHom S B; p \ carrier R; q \ carrier R\ \ erH R S X A B Y h (p \\<^sub>r q) = (erH R S X A B Y h p) \\<^sub>r\<^bsub>A\<^esub> (erH R S X A B Y h q)" apply (frule erH_rHom[of A B Y h], assumption+, simp add:polyn_Hom_def, (erule conjE)+, frule PolynRg.is_Ring[of A B Y], cut_tac is_Ring, rule rHom_tOp[of R A p q "erH R S X A B Y h"], assumption+) done lemma (in PolynRg) erH_rHom_cf:"\Idomain S; PolynRg A B Y; h \ rHom S B; s \ carrier S\ \ erH R S X A B Y h s = h s" apply (cut_tac subring, frule subring_Ring, frule PolynRg.subring[of A B Y], frule PolynRg.is_Ring[of A B Y], frule Ring.subring_Ring[of A B], assumption, frule mem_subring_mem_ring[of S s], assumption+) apply (case_tac "s = \\<^bsub>S\<^esub>", simp add:Subring_zero_ring_zero, simp add:erH_rHom_0, simp add:Subring_zero_ring_zero[THEN sym], simp add:rHom_0_0, simp add:Ring.Subring_zero_ring_zero) apply (frule s_cf_expr[of s],simp add:Subring_zero_ring_zero, (erule conjE)+, simp add:Subring_zero_ring_zero) apply (frule s_cf_deg[of s], assumption+, frule pol_of_deg0[of s], assumption+, simp) apply (subst erH_def, simp, subst erh_rHom_coeff[of A B Y h "s_cf R S X s"], assumption+, simp add:cmp_def polyn_expr_def, frule_tac c = "s_cf R S X s" and j = 0 in pol_coeff_mem, simp, frule mem_subring_mem_ring[of S "snd (s_cf R S X s) 0"], assumption+, simp add:ring_r_one) done lemma (in PolynRg) erH_rHom_coeff:"\Idomain S; PolynRg A B Y; h \ rHom S B; pol_coeff S (n, f)\ \ pol_coeff B (n, cmp h f)" apply (simp add:pol_coeff_def) apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption) apply (simp add:cmp_def rHom_mem) done lemma (in PolynRg) erH_rHom_unique:"\Idomain S; PolynRg A B Y; h \ rHom S B\ \ \!g. g \ pHom R S X, A B Y \ (\x\carrier S. h x = g x)" apply (cut_tac subring, frule subring_Ring, cut_tac is_Ring, frule PolynRg.subring[of A B Y], frule PolynRg.is_Ring[of A B Y], frule Ring.subring_Ring[of A B], assumption, frule Idomain.idom_is_ring[of S]) apply (rule ex_ex1I) apply (frule erH_rHom[of A B Y h], assumption+) apply (subgoal_tac "\x\carrier S. h x = (erH R S X A B Y h) x", blast) apply (rule ballI, simp add:erH_rHom_cf, (erule conjE)+) apply (frule pHom_rHom[of A B Y], assumption+, frule pHom_rHom[of A B Y], assumption+, frule_tac f = g in rHom_func[of _ R A], frule_tac f = y in rHom_func[of _ R A]) apply (rule_tac f = g and g = y in funcset_eq[of _ "carrier R"], simp add:rHom_def aHom_def, simp add:rHom_def aHom_def) apply (rule ballI, thin_tac "g \ carrier R \ carrier A", thin_tac "y \ carrier R \ carrier A") apply (case_tac "x = \\<^bsub>R\<^esub>", simp, subst rHom_0_0[of R A], assumption+, rule sym, subst rHom_0_0[of R A], assumption+, simp) apply (subst pHom_pol_mem[of A B Y], assumption+) apply (frule_tac f = y and p = x in pHom_pol_mem[of A B Y], assumption+, simp, frule_tac f = g and c = "s_cf R S X x" in polyn_Hom_coeff_to_coeff, assumption+, simp add:s_cf_pol_coeff, frule_tac f = y and c = "s_cf R S X x" in polyn_Hom_coeff_to_coeff, assumption+, simp add:s_cf_pol_coeff) apply (simp add:s_cf_deg) apply (frule_tac f = g and c = "s_cf R S X x" in cf_h_len1[of A B Y h], assumption+, rule ballI, rule sym, simp, rule s_cf_pol_coeff, assumption+) apply (frule_tac f = y and c = "s_cf R S X x" in cf_h_len1[of A B Y h], assumption+, rule ballI, rule sym, simp, rule s_cf_pol_coeff, assumption+) apply (frule_tac c = "cf_h g (s_cf R S X x)" and d = "cf_h y (s_cf R S X x)" in PolynRg.pol_expr_unique2[of A B Y], assumption+, simp) apply (frule_tac p = x in s_cf_pol_coeff, simp add:cf_h_len, thin_tac "(polyn_expr A Y (fst (s_cf R S X x)) (cf_h g (s_cf R S X x)) = polyn_expr A Y (fst (s_cf R S X x)) (cf_h y (s_cf R S X x))) = (\j\fst (s_cf R S X x). snd (cf_h g (s_cf R S X x)) j = snd (cf_h y (s_cf R S X x)) j)") apply (rule allI, rule impI, (subst cf_h_def)+, (subst cmp_def)+, simp, frule_tac c = "s_cf R S X x" and j = j in pol_coeff_mem, assumption+) apply simp done lemma (in PolynRg) erH_rHom_unique1:"\Idomain S; PolynRg A B Y; h \ rHom S B; f \ pHom R S X, A B Y; \x \ carrier S. f x = h x\ \ f = (erH R S X A B Y h)" apply (frule erH_rHom_unique[of A B Y h], assumption+, erule ex1E, frule_tac x = f in spec, drule_tac x = "erH R S X A B Y h" in spec) apply (frule erH_rHom[of A B Y h], assumption+, simp add:erH_rHom_cf[THEN sym]) done lemma (in PolynRg) pHom_dec_deg:"\PolynRg A B Y; f \ pHom R S X, A B Y; p \ carrier R\ \ deg A B Y (f p) \ deg R S X p" apply (cut_tac subring, frule subring_Ring, frule PolynRg.subring[of A B Y], cut_tac is_Ring) apply (frule PolynRg.is_Ring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (case_tac "f p = \\<^bsub>A\<^esub>", case_tac "p = \\<^bsub>R\<^esub>", simp add:deg_def, simp add:deg_def an_def, simp add:deg_def, subst ale_natle) apply (case_tac "p = \\<^bsub>R\<^esub>", frule pHom_rHom[of A B Y f], assumption+, rule conjI, rule impI, frule rHom_0_0[of R A f], assumption+, simp, rule impI, simp) apply simp apply (frule pHom_pol_mem[of A B Y f p], assumption+) apply (cut_tac polyn_Hom_coeff_to_coeff[of A B Y f "s_cf R S X p"]) apply (frule PolynRg.pol_deg_le_n[of A B Y "f p" "cf_h f (s_cf R S X p)"], frule pHom_rHom[of A B Y f], assumption+, rule rHom_mem[of f R A p], assumption+, frule s_cf_pol_coeff[of p], subst cf_h_len2[of A B Y f "s_cf R S X p"], assumption+, simp add:s_cf_deg, thin_tac "f p = polyn_expr A Y (deg_n R S X p) (cf_h f (s_cf R S X p))") apply (frule s_cf_pol_coeff[of p], simp add:cf_h_len2, simp add:s_cf_deg[THEN sym], assumption+, simp add:s_cf_pol_coeff) done lemma (in PolynRg) erH_map:"\Idomain S; PolynRg A B Y; h \ rHom S B; pol_coeff S (n, c)\ \ (erH R S X A B Y h) (polyn_expr R X n (n, c)) = polyn_expr A Y n (n, (cmp h c))" apply (cut_tac subring, frule subring_Ring, frule PolynRg.subring[of A B Y], cut_tac is_Ring, frule PolynRg.is_Ring[of A B Y], frule Ring.subring_Ring[of A B], assumption) apply (case_tac "polyn_expr R X n (n, c) = \\<^bsub>R\<^esub>", simp add:erH_rHom_0) apply (frule coeff_0_pol_0[THEN sym, of "(n, c)" n], simp, simp, thin_tac "polyn_expr R X n (n, c) = \") apply (frule cf_h_coeff[of A B Y h "(n, c)"], assumption+, simp add:cf_h_pol_coeff) apply (rule sym, frule_tac PolynRg.coeff_0_pol_0[THEN sym, of A B Y "(n, cmp h c)" n], simp+) apply (rule allI, rule impI, simp add:cmp_def, simp add:rHom_0_0) apply (frule erH_rHom[of A B Y h], assumption+) apply (subst pHom_mem[of A B Y "erH R S X A B Y h" n c], assumption+) apply (frule PolynRg.pol_expr_unique2[of A B Y "(n, cmp (erH R S X A B Y h) c)" "(n, cmp h c)"], simp add:cmp_pol_coeff_e, simp add:cmp_pol_coeff) apply (simp, simp, thin_tac "(polyn_expr A Y n (n, cmp (erH R S X A B Y h) c) = polyn_expr A Y n (n, cmp h c)) = (\j\n. cmp (erH R S X A B Y h) c j = cmp h c j)") apply (rule allI, rule impI, frule_tac j = j in pol_coeff_mem[of "(n, c)"], simp, simp add:cmp_def) apply (simp add:erH_rHom_cf) done section "Relatively prime polynomials" definition rel_prime_pols :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a, 'a ] \ bool" where "rel_prime_pols R S X p q \ (1\<^sub>r\<^bsub>R\<^esub>) \ ((Rxa R p) \\<^bsub>R\<^esub> (Rxa R q))" definition div_condn :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, nat, 'a, 'a ] \ bool" where "div_condn R S X n g f \ f \ carrier R \ n = deg_n R S X f \ (\q. q \ carrier R \ ((f \\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (q \\<^sub>r\<^bsub>R\<^esub> g)) = \\<^bsub>R\<^esub>) \ (deg_n R S X (f \\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (q \\<^sub>r\<^bsub>R\<^esub> g))) < deg_n R S X g)))" lemma (in PolynRg) divisionTr0:"\Idomain S; p \ carrier R; c \ carrier S; c \ \\<^bsub>S\<^esub>\ \ lcf R S X (c \\<^sub>r X^\<^bsup>R n\<^esup> \\<^sub>r p) = c \\<^sub>r\<^bsub>S\<^esub> (lcf R S X p)" apply (cut_tac polyn_ring_integral, simp, cut_tac subring, frule subring_Ring, cut_tac polyn_ring_X_nonzero, cut_tac X_mem_R) apply (frule mem_subring_mem_ring[of S c], assumption+, frule npClose[of X n]) apply (case_tac "p = \\<^bsub>R\<^esub>", simp, frule ring_tOp_closed[of c "X^\<^bsup>R n\<^esup>"], assumption+, simp add:ring_times_x_0 lcf_val_0, simp add:Ring.ring_times_x_0[of S]) apply (frule_tac x = c and y = " X^\<^bsup>R n\<^esup>" in Idomain.idom_tOp_nonzeros[of R], assumption+, simp add:Subring_zero_ring_zero, frule Idomain.idom_potent_nonzero[of R X n], assumption+, frule_tac x = c and y = " X^\<^bsup>R n\<^esup>" in ring_tOp_closed, assumption+, frule_tac x = "c \\<^sub>r X^\<^bsup>R n\<^esup>" and y = p in Idomain.idom_tOp_nonzeros[of R], assumption+, frule_tac x = "c \\<^sub>r X^\<^bsup>R n\<^esup>" and y = p in ring_tOp_closed, assumption+) apply (simp add:lcf_val) apply (frule s_cf_expr[of p], assumption, (erule conjE)+, simp add:ring_tOp_assoc[of c _ p], frule low_deg_terms_zero1[THEN sym, of "s_cf R S X p" n]) apply (cut_tac a = "X^\<^bsup>R n\<^esup> \\<^sub>r polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)" and b = "X^\<^bsup>R n\<^esup> \\<^sub>r p" and c = "polyn_expr R X (fst (s_cf R S X p) + n) (ext_cf S n (s_cf R S X p))" in box_equation, simp, assumption, thin_tac "p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)", thin_tac "X^\<^bsup>R n\<^esup> \\<^sub>r polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) = polyn_expr R X (fst (s_cf R S X p) + n) (ext_cf S n (s_cf R S X p))") apply (frule ext_cf_pol_coeff[of "s_cf R S X p" n], frule scalar_times_pol_expr[of c "ext_cf S n (s_cf R S X p)" "fst (s_cf R S X p) + n"], assumption+, simp add:ext_cf_len) apply (frule sp_cf_pol_coeff[of "ext_cf S n (s_cf R S X p)" c], assumption+, cut_tac a = "c \\<^sub>r polyn_expr R X (fst (s_cf R S X p) + n) (ext_cf S n (s_cf R S X p))" and b = "c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p)" and c = "polyn_expr R X (fst (s_cf R S X p) + n) (sp_cf S c (ext_cf S n (s_cf R S X p)))" in box_equation, simp, simp, thin_tac "X^\<^bsup>R n\<^esup> \\<^sub>r p = polyn_expr R X (fst (s_cf R S X p) + n) (ext_cf S n (s_cf R S X p))", thin_tac "c \\<^sub>r polyn_expr R X (fst (s_cf R S X p) + n) (ext_cf S n (s_cf R S X p)) = polyn_expr R X (fst (s_cf R S X p) + n) (sp_cf S c (ext_cf S n (s_cf R S X p)))") apply (frule s_cf_expr[of "c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p)"], assumption+, (erule conjE)+, drule_tac a = "c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p)" and b = "polyn_expr R X (fst (s_cf R S X (c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p)))) (s_cf R S X (c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p)))" and c = "polyn_expr R X (fst (s_cf R S X p) + n) (sp_cf S c (ext_cf S n (s_cf R S X p)))" in box_equation, assumption, thin_tac "c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p) = polyn_expr R X (fst (s_cf R S X p) + n) (sp_cf S c (ext_cf S n (s_cf R S X p)))", frule pol_expr_unique2[of "s_cf R S X (c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p))" "sp_cf S c (ext_cf S n (s_cf R S X p))"], assumption+, subst s_cf_deg[THEN sym], assumption+, frule_tac Idomain.idom_potent_nonzero[of R X n], assumption+, frule_tac x = "X^\<^bsup>R n\<^esup>" and y = p in Idomain.idom_tOp_nonzeros[of R], assumption+) apply (subst deg_mult_pols, assumption+, simp add:Subring_zero_ring_zero, simp add:ring_tOp_closed, assumption+, subst deg_mult_pols, assumption+, simp add:deg_to_X_d, cut_tac pol_of_deg0[THEN sym, of c], simp, simp add:sp_cf_len ext_cf_len s_cf_deg, assumption+, simp add:Subring_zero_ring_zero, simp add:sp_cf_len ext_cf_len) apply (subst s_cf_deg[THEN sym], assumption+, frule_tac Idomain.idom_potent_nonzero[of R X n], assumption+, frule_tac x = "X^\<^bsup>R n\<^esup>" and y = p in Idomain.idom_tOp_nonzeros[of R], assumption+, simp add:s_cf_deg[THEN sym], frule_tac x = "X^\<^bsup>R n\<^esup>" and y = p in ring_tOp_closed, assumption+) apply (frule deg_mult_pols[of c "X^\<^bsup>R n\<^esup> \\<^sub>r p"], assumption+, simp add:Subring_zero_ring_zero, assumption+, (erule conjE)+, simp, thin_tac "deg_n R S X (c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p)) = deg_n R S X c + deg_n R S X (X^\<^bsup>R n\<^esup> \\<^sub>r p)", cut_tac pol_of_deg0[THEN sym, of c], simp, frule deg_mult_pols[of "X^\<^bsup>R n\<^esup>" p], assumption+, (erule conjE)+, simp, thin_tac "deg_n R S X (X^\<^bsup>R n\<^esup> \\<^sub>r p) = deg_n R S X (X^\<^bsup>R n\<^esup>) + deg_n R S X p", simp add:deg_to_X_d, simp add:add.commute[of n], thin_tac "polyn_expr R X (deg_n R S X p + n) (s_cf R S X (c \\<^sub>r (X^\<^bsup>R n\<^esup> \\<^sub>r p))) = polyn_expr R X (deg_n R S X p + n) (sp_cf S c (ext_cf S n (s_cf R S X p)))") apply (subst sp_cf_def, simp) apply (subst ext_cf_def, simp add:sliden_def, assumption) apply (simp add:Subring_zero_ring_zero) done lemma (in PolynRg) divisionTr1:"\Corps S; g \ carrier R; g \ \; 0 < deg_n R S X g; f \ carrier R; f \ \; deg_n R S X g \ deg_n R S X f\ \ f \ -\<^sub>a ((lcf R S X f) \\<^sub>r\<^bsub>S\<^esub> ((lcf R S X g)\<^bsup>\ S\<^esup>) \\<^sub>r (X^\<^bsup>R ((deg_n R S X f) - (deg_n R S X g))\<^esup>) \\<^sub>r g) = \ \ deg_n R S X (f \ -\<^sub>a ((lcf R S X f) \\<^sub>r\<^bsub>S\<^esub> ((lcf R S X g)\<^bsup>\ S\<^esup>) \\<^sub>r (X^\<^bsup>R ((deg_n R S X f) - (deg_n R S X g))\<^esup>) \\<^sub>r g)) < deg_n R S X f" apply (cut_tac ring_is_ag, cut_tac subring, frule Corps.field_is_idom[of "S"], frule subring_Ring, cut_tac subring, cut_tac polyn_ring_X_nonzero, cut_tac X_mem_R, cut_tac polyn_ring_integral, simp) apply (frule npClose[of X "deg_n R S X f - deg_n R S X g"], frule_tac Idomain.idom_potent_nonzero[of R X "fst (s_cf R S X f) - fst (s_cf R S X g)"], assumption+, frule s_cf_expr[of f], assumption+, (erule conjE)+, frule s_cf_expr[of g], assumption+, (erule conjE)+, simp add:s_cf_deg, simp add:lcf_val[THEN sym], frule Corps.invf_closed1[of S "lcf R S X g"], simp, simp add:lcf_mem, frule lcf_mem[of f], simp, frule subring_Ring, frule Ring.ring_is_ag[of S], (erule conjE)+, frule_tac x = "lcf R S X f" and y = "lcf R S X g\<^bsup>\ S\<^esup>" in Ring.ring_tOp_closed[of S], assumption+, frule mem_subring_mem_ring[of S " lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> "], assumption+, frule_tac x = "lcf R S X f" and y = "lcf R S X g\<^bsup>\ S\<^esup>" in Idomain.idom_tOp_nonzeros[of S], assumption+, simp add:Subring_zero_ring_zero) apply(frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>" and y = "X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup>" in Idomain.idom_tOp_nonzeros[of R], assumption+, frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>" and y = "X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup>" in ring_tOp_closed, assumption+, frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup>" and y = g in Idomain.idom_tOp_nonzeros[of R], assumption+, frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup>" and y = g in ring_tOp_closed, assumption+) apply (frule pol_diff_deg_less[of f "s_cf R S X f" "s_cf R S X (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup> \\<^sub>r g)"], assumption+, simp add:s_cf_pol_coeff) apply (simp add:s_cf_deg[THEN sym]) apply (frule deg_mult_pols[of "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> (lcf R S X g)\<^bsup>\ S\<^esup>" "(X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> )\\<^sub>r g"], assumption+, rule ring_tOp_closed, assumption+, rule Idomain.idom_tOp_nonzeros[of R], assumption+, (erule conjE)+, subst ring_tOp_assoc[of "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>" _ g], assumption+, simp, cut_tac pol_of_deg0[THEN sym, of "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>"], simp, thin_tac "deg_n R S X (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r (X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g)) = deg_n R S X (X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g)") apply (frule deg_mult_pols[of "(X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> )" g], assumption+, simp, simp add:deg_to_X_d, assumption+, fold lcf_def, subst divisionTr0[of g "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>" "fst (s_cf R S X f) - fst (s_cf R S X g)"], assumption+, simp add:Subring_zero_ring_zero) apply (subst Ring.ring_tOp_assoc, assumption+, simp add:lcf_mem, frule Corps.invf_inv[of S "lcf R S X g"], simp add:lcf_mem, simp add:Subring_zero_ring_zero, simp add:Ring.ring_r_one, frule s_cf_expr[of "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup> \\<^sub>r g"], rule Idomain.idom_tOp_nonzeros[of R], assumption+) apply ((erule conjE)+, thin_tac "snd (s_cf R S X (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup> \\<^sub>r g)) (fst (s_cf R S X (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (fst (s_cf R S X f) - fst (s_cf R S X g))\<^esup> \\<^sub>r g))) \ \\<^bsub>S\<^esub>") apply (rotate_tac -1, drule sym, simp) done lemma (in PolynRg) divisionTr2:"\Corps S; g \ carrier R; g \ \; 0 < deg_n R S X g\ \ \f. div_condn R S X n g f" apply (cut_tac ring_is_ag, frule Corps.field_is_idom[of "S"], cut_tac subring, frule subring_Ring, cut_tac polyn_ring_integral, simp, cut_tac X_mem_R) apply (rule nat_less_induct) apply (rule allI) apply (subst div_condn_def, rule impI, (erule conjE)+) apply (case_tac "f = \\<^bsub>R\<^esub>", cut_tac ring_zero, subgoal_tac " f \ -\<^sub>a (\ \\<^sub>r g) = \", blast, simp add:ring_times_0_x, simp add:aGroup.ag_inv_zero[of "R"], simp add:aGroup.ag_r_zero) apply (case_tac "n < deg_n R S X g") apply (cut_tac ring_zero, subgoal_tac "deg_n R S X (f \ -\<^sub>a (\ \\<^sub>r g)) < deg_n R S X g", blast) apply ( simp add:ring_times_0_x, simp add:aGroup.ag_inv_zero, simp add:aGroup.ag_r_zero) apply (frule_tac x = n and y = "deg_n R S X g" in leI, thin_tac "\ n < deg_n R S X g") (** deg_n R S X g \ deg_n R S X f **) apply (frule_tac f = f in divisionTr1[of g], assumption+, simp) apply (frule_tac p = f in lcf_mem, frule lcf_mem[of g], frule lcf_nonzero[of g], assumption+, frule Corps.invf_closed1[of S "lcf R S X g"], simp) apply (frule_tac x = "lcf R S X f" and y = "lcf R S X g\<^bsup>\ S\<^esup>" in Ring.ring_tOp_closed[of S], assumption+, simp) apply (frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>" in mem_subring_mem_ring, assumption) apply (frule_tac n = "deg_n R S X f - deg_n R S X g" in npClose[of X], frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup>" and y = "X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup>" in ring_tOp_closed, assumption+, frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup>" and y = g in ring_tOp_closed, assumption+) apply (erule disjE, blast) apply (drule_tac a = "deg_n R S X (f \ -\<^sub>a (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g))" in forall_spec, simp) apply (simp add:div_condn_def) apply (drule_tac x = "f \ -\<^sub>a (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g)" in spec) apply (frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g" in aGroup.ag_mOp_closed, assumption) apply (frule_tac x = f and y = "-\<^sub>a (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g)" in aGroup.ag_pOp_closed, assumption+, simp, thin_tac "deg_n R S X (f \ -\<^sub>a (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g)) < deg_n R S X f") apply (erule exE, thin_tac "f \ -\<^sub>a (lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup> \\<^sub>r g) \ carrier R") apply ((erule conjE)+, frule_tac x = q and y = g in ring_tOp_closed, assumption+, frule_tac x = "q \\<^sub>r g" in aGroup.ag_mOp_closed, assumption+, simp add:aGroup.ag_pOp_assoc, simp add:aGroup.ag_p_inv[THEN sym], simp add:ring_distrib2[THEN sym]) apply (frule_tac x = "lcf R S X f \\<^sub>r\<^bsub>S\<^esub> lcf R S X g\<^bsup>\ S\<^esup> \\<^sub>r X^\<^bsup>R (deg_n R S X f - deg_n R S X g)\<^esup>" and y = q in aGroup.ag_pOp_closed, assumption+) apply (erule disjE) apply blast apply blast done lemma (in PolynRg) divisionTr3:"\Corps S; g \ carrier R; g \ \; 0 < deg_n R S X g; f \ carrier R\ \ \q\carrier R. (f \ -\<^sub>a (q \\<^sub>r g) = \) \ ( f \ -\<^sub>a (q \\<^sub>r g) \ \ \ deg_n R S X (f \ -\<^sub>a (q \\<^sub>r g)) < (deg_n R S X g))" apply (frule divisionTr2[of g "deg_n R S X f"], assumption+) apply (drule_tac x = f in spec) apply (simp add:div_condn_def, blast) done lemma (in PolynRg) divisionTr4:"\Corps S; g \ carrier R; g \ \; 0 < deg_n R S X g; f \ carrier R\ \ \q\carrier R. (f = q \\<^sub>r g) \ (\r\carrier R. r \ \ \ (f = (q \\<^sub>r g) \ r) \ (deg_n R S X r) < (deg_n R S X g))" apply (cut_tac is_Ring, cut_tac ring_is_ag) apply (frule divisionTr3[of g f], assumption+, erule bexE, frule_tac x = q in ring_tOp_closed[of _ g], assumption+, erule disjE) apply (simp add:aGroup.ag_eq_diffzero[THEN sym, of "R" "f"], blast) apply (subgoal_tac "f = q \\<^sub>r g \ (f \ -\<^sub>a (q \\<^sub>r g))", subgoal_tac "(f \ -\<^sub>a (q \\<^sub>r g)) \ carrier R", blast, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (frule_tac x = "q \\<^sub>r g" in aGroup.ag_mOp_closed, assumption+, subst aGroup.ag_pOp_assoc[THEN sym], assumption+, subst aGroup.ag_pOp_commute[of R _ f], assumption+, subst aGroup.ag_pOp_assoc, assumption+, simp add:aGroup.ag_r_inv1, simp add:aGroup.ag_r_zero) done lemma (in PolynRg) divisionTr:"\Corps S; g \ carrier R; 0 < deg R S X g; f \ carrier R\ \ \q\carrier R. (\r\carrier R. (f = (q \\<^sub>r g) \ r) \ (deg R S X r) < (deg R S X g))" apply (subgoal_tac "g \ \", frule divisionTr4[of g f], assumption+, simp add:deg_def, simp only:an_0[THEN sym], cut_tac aless_nat_less[of "0" "deg_n R S X g"], simp, assumption) apply (erule bexE, erule disjE, cut_tac ring_is_ag, frule aGroup.ag_r_zero[of "R" "f"], simp, simp, rotate_tac -1, frule sym, cut_tac ring_zero, subgoal_tac "deg R S X \ < deg R S X g", blast, simp add:deg_def an_def) apply (erule bexE, (erule conjE)+, cut_tac n1 = "deg_n R S X r" and m1 = "deg_n R S X g" in aless_natless[THEN sym], simp add:deg_def, drule sym, simp, rotate_tac -1, drule sym, blast) apply (rule contrapos_pp, simp+, simp add:deg_def, frule aless_imp_le[of "0" "-\"], cut_tac minf_le_any[of "0"]) apply (frule ale_antisym[of "0" "-\"], assumption) apply simp done lemma (in PolynRg) rel_prime_equation:"\Corps S; f \ carrier R; g \ carrier R; 0 < deg R S X f; 0 < deg R S X g; rel_prime_pols R S X f g; h \ carrier R\ \ \u \ carrier R. \v \ carrier R. (deg R S X u \ amax ((deg R S X h) - (deg R S X f)) (deg R S X g)) \ (deg R S X v \ (deg R S X f)) \ (u \\<^sub>r f \ (v \\<^sub>r g) = h)" apply (cut_tac ring_is_ag, cut_tac ring_zero, cut_tac subring, frule subring_Ring, frule aless_imp_le [of "0" "deg R S X f"], frule pol_nonzero[of f], simp, frule aless_imp_le [of "0" "deg R S X g"], frule pol_nonzero[of g], simp, frule Corps.field_is_idom[of "S"], cut_tac polyn_ring_integral, simp, frule Idomain.idom_tOp_nonzeros[of R f g], assumption+) apply (case_tac "h = \\<^bsub>R\<^esub>") apply (cut_tac ring_is_ag, cut_tac ring_zero, subgoal_tac "deg R S X \ \ amax (deg R S X h - deg R S X f) (deg R S X g)", subgoal_tac "deg R S X \ \ deg R S X f \ \ \\<^sub>r f \ \ \\<^sub>r g = h", blast) apply (simp add:ring_times_0_x, simp add:aGroup.ag_r_zero, simp add:deg_def) apply (simp add:deg_def amax_def) apply (simp add:rel_prime_pols_def, frule principal_ideal[of f], frule principal_ideal[of g], frule ideals_set_sum[of "R \\<^sub>p f" "R \\<^sub>p g" "1\<^sub>r"], assumption+, thin_tac "1\<^sub>r \ R \\<^sub>p f \ R \\<^sub>p g", (erule bexE)+, thin_tac "ideal R (R \\<^sub>p f)", thin_tac "ideal R (R \\<^sub>p g)", simp add:Rxa_def, (erule bexE)+, simp, thin_tac "ha = r \\<^sub>r f", thin_tac "k = ra \\<^sub>r g") apply (frule_tac x = r in ring_tOp_closed[of _ f], assumption+, frule_tac x = ra in ring_tOp_closed[of _ g], assumption+, frule_tac y1 = "r \\<^sub>r f" and z1 = "ra \\<^sub>r g" in ring_distrib1[THEN sym, of "h"], assumption+, simp add:ring_r_one, drule sym, simp, thin_tac "r \\<^sub>r f \ ra \\<^sub>r g = 1\<^sub>r", simp add:ring_tOp_assoc[THEN sym], simp add:ring_r_one) apply (frule_tac f = "h \\<^sub>r r" in divisionTr[of g], assumption+, simp add:ring_tOp_closed, frule_tac f = "h \\<^sub>r ra" in divisionTr[of f], assumption+, simp add:ring_tOp_closed, (erule bexE)+, (erule conjE)+) (** final **) apply (thin_tac " r \ carrier R", thin_tac "ra \ carrier R", thin_tac "r \\<^sub>r f \ carrier R", thin_tac "ra \\<^sub>r g \ carrier R") apply (frule_tac x = q in ring_tOp_closed[of _ g], assumption+, frule_tac x = qa in ring_tOp_closed[of _ f], assumption+, frule_tac x = "q \\<^sub>r g" and y = rb in aGroup.ag_pOp_commute, assumption+, simp, thin_tac "q \\<^sub>r g \ rb = rb \ q \\<^sub>r g", thin_tac "h \\<^sub>r r = rb \ q \\<^sub>r g", thin_tac "h \\<^sub>r ra = qa \\<^sub>r f \ rc") apply (simp add:ring_distrib2[of g], frule_tac x = rb and y = "q \\<^sub>r g" in aGroup.ag_pOp_closed[of R], assumption+, frule_tac x = "rb \ q \\<^sub>r g" and y = f in ring_tOp_closed, assumption+, frule_tac x = "qa \\<^sub>r f" and y = g in ring_tOp_closed, assumption+, frule_tac x = rc and y = g in ring_tOp_closed, assumption+, simp add:aGroup.ag_pOp_assoc[THEN sym, of "R"], simp add:ring_tOp_assoc[of _ f g], simp add:ring_tOp_commute[of f g], simp add:ring_tOp_assoc[THEN sym, of _ g f], frule_tac x = qa and y = g in ring_tOp_closed, assumption+, simp add:ring_distrib2[THEN sym], simp add:aGroup.ag_pOp_assoc, simp add:ring_distrib2[THEN sym], case_tac "q \\<^bsub>R\<^esub> qa = \\<^bsub>R\<^esub>", simp add:ring_times_0_x, simp add:aGroup.ag_r_zero, subgoal_tac "deg R S X rb \ amax (deg R S X h - deg R S X f) (deg R S X g)", subgoal_tac "deg R S X rc \ deg R S X f", blast, simp add:aless_imp_le, frule_tac x = "deg R S X rb" and y = "deg R S X g" in aless_imp_le, rule_tac i = "deg R S X rb" in ale_trans[of _ "deg R S X g" "amax (deg R S X h - deg R S X f) (deg R S X g)"], assumption, simp add:amax_def, simp add:aneg_le aless_imp_le) apply (subgoal_tac "rb \ (q \ qa) \\<^sub>r g \ carrier R", subgoal_tac "deg R S X (rb \ (q \ qa) \\<^sub>r g) \ amax (deg R S X h - deg R S X f) (deg R S X g)", subgoal_tac "deg R S X rc \ deg R S X f", blast, simp add:aless_imp_le, frule_tac x = q and y = qa in aGroup.ag_pOp_closed[of "R"], assumption+, frule_tac p = rb and q = "(q \ qa) \\<^sub>r g" in deg_pols_add1, rule ring_tOp_closed, assumption+, simp add:deg_mult_pols1, frule_tac p1 = "q \ qa" in pol_nonzero[THEN sym], simp) apply (frule_tac y = "deg R S X (q \ qa)" and z = "deg R S X rb" in aadd_le_mono[of "0"], simp add:aadd_0_l) apply (frule_tac p = "q \ qa" in deg_ant_int, assumption+, frule_tac x = "deg R S X rb" and y = "deg R S X g" and z = "int (deg_n R S X ( q \ qa))" in aadd_less_mono_z, simp add:aadd_commute) apply (simp add:deg_mult_pols1, frule_tac p = "rb \ (q \ qa) \\<^sub>r g" and q = f in deg_mult_pols1, assumption+, simp, thin_tac "deg R S X (rb \ (q \ qa) \\<^sub>r g) = deg R S X (q \ qa) + deg R S X g", frule_tac x = "rb \ (q \ qa) \\<^sub>r g" and y = f in ring_tOp_closed, assumption+, simp only:aGroup.ag_pOp_commute, frule_tac p = "rc \\<^sub>r g" and q = "(rb \ (q \ qa) \\<^sub>r g) \\<^sub>r f" in deg_pols_add1, assumption+, simp, thin_tac "deg R S X ((rb \ (q \ qa) \\<^sub>r g) \\<^sub>r f) = deg R S X (q \ qa) + deg R S X g + deg R S X f") apply (simp add:deg_mult_pols1, frule_tac p1 = "q \ qa" in pol_nonzero[THEN sym], simp, simp add:deg_ant_int[of g]) apply (frule_tac x = "deg R S X rc" and y = "deg R S X f" and z = "int (deg_n R S X g)" in aadd_less_mono_z, frule_tac a = "deg R S X ( q \ qa)" in aadd_pos_le[of _ "deg R S X f + ant (int (deg_n R S X g))"], frule_tac x = "deg R S X rc + ant (int (deg_n R S X g))" and y = "deg R S X f + ant (int (deg_n R S X g))" and z = "deg R S X ( q \ qa) + (deg R S X f + ant (int (deg_n R S X g)))" in aless_le_trans, assumption+, thin_tac "deg R S X rc + ant (int (deg_n R S X g)) < deg R S X f + ant (int (deg_n R S X g))", thin_tac "deg R S X f + ant (int (deg_n R S X g)) \ deg R S X ( q \ qa) + (deg R S X f + ant (int (deg_n R S X g)))") apply (simp add:deg_ant_int[THEN sym]) apply (frule_tac p = "q \ qa" in deg_in_aug_minf, frule_tac p = "g" in deg_in_aug_minf, frule_tac p = "f" in deg_in_aug_minf, simp add:aadd_commute[of "deg R S X f" "deg R S X g"], simp only:aadd_assoc_m[THEN sym], simp) apply (frule_tac p = "g" in deg_in_aug_minf, frule_tac p = "f" in deg_in_aug_minf, frule_tac p = "q \ qa" in deg_in_aug_minf, simp add:diff_ant_def, subgoal_tac "-(deg R S X f) \ Z\<^sub>-\<^sub>\") apply (subst aadd_assoc_m[of _ "deg R S X f" "- deg R S X f"], simp add:Zminf_pOp_closed, assumption+, (simp add:aadd_minus_r, simp add:aadd_0_r), simp add:amax_ge_l, simp add:deg_ant_int, simp add:aminus, simp add:z_in_aug_minf) apply (rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, rule aGroup.ag_pOp_closed, assumption+) done subsection "Polynomial, coeff mod P" definition P_mod :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme, 'a, 'a set, 'a] \ bool" where "P_mod R S X P p \ p = \\<^bsub>R\<^esub> \ (\j \ (fst (s_cf R S X p)). (snd (s_cf R S X p) j) \ P)" lemma (in PolynRg) P_mod_whole:"p \ carrier R \ P_mod R S X (carrier S) p" apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:P_mod_def) apply (simp add:P_mod_def, rule allI, rule impI, rule pol_coeff_mem, simp add:s_cf_pol_coeff, assumption) done lemma (in PolynRg) zero_P_mod:"ideal S I \ P_mod R S X I \" by (simp add:P_mod_def) lemma (in PolynRg) P_mod_mod:"\ideal S I; p \ carrier R; pol_coeff S c; p = polyn_expr R X (fst c) c\ \ (\j \ (fst c). (snd c) j \ I) = (P_mod R S X I p)" apply (cut_tac subring, frule subring_Ring) apply (case_tac "p = \\<^bsub>R\<^esub>") apply (simp add:P_mod_def, drule sym, frule coeff_0_pol_0[THEN sym, of c "fst c"], simp, simp) apply (rule impI, simp add:Ring.ideal_zero) apply (frule s_cf_expr[of p], simp add:P_mod_def, (erule conjE)+) apply (frule polyn_c_max[of c]) apply (frule coeff_nonzero_polyn_nonzero[of c "fst c"], simp) apply (frule coeff_max_nonzeroTr[of c], simp) apply (thin_tac "(polyn_expr R X (fst c) c \ \) = (\j\fst c. snd c j \ \\<^bsub>S\<^esub>)") apply (frule coeff_max_bddTr[of c]) apply (frule polyn_expr_short[of c "c_max S c"], assumption+) apply (frule pol_expr_unique[of p "(c_max S c, snd c)" "s_cf R S X p"], assumption+, rule split_pol_coeff[of c], assumption+, simp, simp, assumption+) apply (thin_tac "p = polyn_expr R X (fst c) c", thin_tac "p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)", thin_tac "polyn_expr R X (fst c) c = polyn_expr R X (c_max S c) c", thin_tac "polyn_expr R X (c_max S c) c = polyn_expr R X (c_max S c) (c_max S c, snd c)") apply (frule coeff_max_zeroTr[of c], (erule conjE)+) apply (subst P_mod_def, simp) apply (rule iffI, rule allI, rule impI) apply (rotate_tac 10, drule_tac a = j in forall_spec, drule_tac x = j in spec, assumption, frule_tac i = j and j = "fst (s_cf R S X p)" and k = "fst c" in le_trans, assumption+, drule_tac a = j in forall_spec, assumption, simp) apply (rule allI, rule impI) apply (case_tac "fst (s_cf R S X p) < j", drule_tac a = j in forall_spec, simp, simp add:Ring.ideal_zero, frule_tac x = "fst (s_cf R S X p)" and y = j in leI, thin_tac "\j. j \ fst c \ fst (s_cf R S X p) < j \ snd c j = \\<^bsub>S\<^esub>", drule_tac a = j in forall_spec, assumption, drule_tac a = j in forall_spec, assumption, simp) done lemma (in PolynRg) monomial_P_mod_mod:"\ideal S I; c \ carrier S; p = c \\<^sub>r (X^\<^bsup>R d\<^esup>)\ \ (c \ I) = (P_mod R S X I p)" apply (cut_tac subring, frule subring_Ring) apply (cut_tac monomial_d[THEN sym, of "(0, \j. c)" "d"], simp) apply (drule sym, simp) apply (subst P_mod_mod[THEN sym, of I p "ext_cf S d (0, \j. c)"], assumption+) apply (frule mem_subring_mem_ring[of S c], assumption, cut_tac X_mem_R, frule npClose[of X d], drule sym, simp add:ring_tOp_closed) apply (simp add:pol_coeff_def, rule allI, rule impI, simp add:ext_cf_def sliden_def, rule impI, simp add:Ring.ring_zero, subst ext_cf_len, simp add:pol_coeff_def, simp) apply (subst ext_cf_len, simp add:pol_coeff_def, simp add:ext_cf_def) apply (rule iffI) apply (simp add:Ring.ideal_zero, drule_tac x = d in spec, simp, simp add:pol_coeff_def) done lemma (in PolynRg) P_mod_add:"\ideal S I; p \ carrier R; q \ carrier R; P_mod R S X I p; P_mod R S X I q\ \ P_mod R S X I (p \ q)" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:aGroup.ag_l_zero, case_tac "q = \\<^bsub>R\<^esub>", simp add:aGroup.ag_r_zero) apply (case_tac "p \\<^bsub>R\<^esub> q = \\<^bsub>R\<^esub>", simp add:P_mod_def) apply (frule s_cf_expr[of p], assumption, frule s_cf_expr[of q], assumption, (erule conjE)+) apply (frule polyn_add1[of "s_cf R S X p" "s_cf R S X q"], assumption+, drule sym, drule sym, simp, drule sym, simp, rotate_tac -1, drule sym) apply (frule P_mod_mod[THEN sym, of I p "s_cf R S X p"], assumption+, simp, thin_tac "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) = p", frule P_mod_mod[THEN sym, of I q "s_cf R S X q"], assumption+, simp, thin_tac "polyn_expr R X (fst (s_cf R S X q)) (s_cf R S X q) = q") apply (frule aGroup.ag_pOp_closed[of R p q], assumption+) apply (subst P_mod_mod[THEN sym, of I "p \ q" "add_cf S (s_cf R S X p) (s_cf R S X q)"], assumption+, simp add:add_cf_pol_coeff, simp, simp add:add_cf_len, thin_tac "p \ q = polyn_expr R X (max (fst (s_cf R S X p)) (fst (s_cf R S X q))) (add_cf S (s_cf R S X p) (s_cf R S X q))") apply simp apply (subst add_cf_len, assumption+) apply (rule allI, rule impI) apply (cut_tac x = "fst (s_cf R S X p)" and y = "fst (s_cf R S X q)" in less_linear) apply (erule disjE) apply (simp, subst add_cf_def, simp, (rule impI, drule_tac a = j in forall_spec, assumption, drule_tac x = j in spec, frule_tac x = j and y = "fst (s_cf R S X p)" and z = "fst (s_cf R S X q)" in le_less_trans, assumption+, frule_tac x = j and y = "fst (s_cf R S X q)" in less_imp_le, simp)) apply (rule Ring.ideal_pOp_closed[of S I], assumption+) apply (erule disjE) apply (simp, subst add_cf_def, simp, drule_tac a = j in forall_spec, assumption, drule_tac a = j in forall_spec, assumption) apply (rule Ring.ideal_pOp_closed[of S I], assumption+) apply (simp add: max.absorb1 max.absorb2, subst add_cf_def, simp, rule impI, drule_tac x = j in spec, drule_tac a = j in forall_spec, assumption, frule_tac x = j and y = "fst (s_cf R S X q)" and z = "fst (s_cf R S X p)" in le_less_trans, assumption+, frule_tac x = j and y = "fst (s_cf R S X p)" in less_imp_le, simp) apply (rule Ring.ideal_pOp_closed[of S I], assumption+) done lemma (in PolynRg) P_mod_minus:"\ideal S I; p \ carrier R; P_mod R S X I p\ \ P_mod R S X I (-\<^sub>a p)" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:aGroup.ag_inv_zero) apply (frule s_cf_expr[of p], assumption+, (erule conjE)+, frule polyn_minus_m_cf[of "s_cf R S X p" "fst (s_cf R S X p)"], simp, frule aGroup.ag_inv_inj[of R p \], assumption, simp add:ring_zero, assumption, simp add:aGroup.ag_inv_zero, frule m_cf_pol_coeff[of "s_cf R S X p"], drule sym, drule sym, simp) apply (subst P_mod_mod[THEN sym, of I "-\<^sub>a p" "m_cf S (s_cf R S X p)"], assumption+, rule aGroup.ag_mOp_closed[of R p], assumption+, simp add:m_cf_len, thin_tac "polyn_expr R X (fst (s_cf R S X p)) (m_cf S (s_cf R S X p)) = -\<^sub>a p") apply (frule P_mod_mod[THEN sym, of I p "s_cf R S X p"], assumption+, simp, thin_tac "polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p) = p", simp) apply (rule allI, rule impI, drule_tac a = j in forall_spec, simp add:m_cf_len, subst m_cf_def, simp, rule Ring.ideal_inv1_closed[of S I], assumption+) done lemma (in PolynRg) P_mod_pre:"\ideal S I; pol_coeff S ((Suc n), f); P_mod R S X I (polyn_expr R X (Suc n) (Suc n, f))\ \ P_mod R S X I (polyn_expr R X n (n, f))" apply (frule pol_coeff_pre[of n f], frule polyn_mem[of "(n, f)" n],simp, frule polyn_mem[of "(Suc n, f)" "Suc n"], simp) apply (case_tac "polyn_expr R X n (n, f) = \\<^bsub>R\<^esub>", simp add:P_mod_def) apply (subst P_mod_mod[THEN sym, of I "polyn_expr R X n (n, f)" "(n, f)"], assumption+, simp, frule P_mod_mod[THEN sym, of I "polyn_expr R X (Suc n) (Suc n, f)" "(Suc n, f)"], assumption+, simp, simp) done lemma (in PolynRg) P_mod_pre1:"\ideal S I; pol_coeff S ((Suc n), f); P_mod R S X I (polyn_expr R X (Suc n) (Suc n, f))\ \ P_mod R S X I (polyn_expr R X n (Suc n, f))" by (simp add:polyn_expr_restrict[of n f], simp add:P_mod_pre) lemma (in PolynRg) P_mod_coeffTr:"\ideal S I; d \ carrier S\ \ (P_mod R S X I d) = (d \ I)" apply (cut_tac subring, frule subring_Ring, subst monomial_P_mod_mod[of I d "d \\<^sub>r X^\<^bsup>R 0\<^esup>" 0], assumption+, simp, simp, frule mem_subring_mem_ring[of _ d], assumption+, simp add:ring_r_one) done lemma (in PolynRg) P_mod_mult_const:"\ideal S I; ideal S J; pol_coeff S (n, f); P_mod R S X I (polyn_expr R X n (n, f)); pol_coeff S (0, g); P_mod R S X J (polyn_expr R X 0 (0, g))\ \ P_mod R S X (I \\<^sub>r\<^bsub>S\<^esub> J) ((polyn_expr R X n (n, f)) \\<^sub>r (polyn_expr R X 0 (0, g)))" apply (cut_tac subring, frule subring_Ring) apply (frule_tac c = "(n, f)" in polyn_mem[of _ n], simp) apply (frule Ring.ideal_prod_ideal[of S I J], assumption+) apply (case_tac "polyn_expr R X n (n, f) = \\<^bsub>R\<^esub>", simp) apply (frule_tac c = "(0, g)" in polyn_mem[of _ 0], simp, simp add:ring_times_0_x, simp add:P_mod_def) apply (simp add:polyn_expr_def [of _ _ "0"]) apply (frule pol_coeff_mem[of "(0, g)" 0], simp, simp, frule mem_subring_mem_ring[of S "g 0"], assumption, simp add:ring_r_one, simp add:ring_tOp_commute[of _ "g 0"]) apply (frule sp_cf_pol_coeff[of "(n, f)" "g 0"], assumption+) apply (subst scalar_times_pol_expr[of "g 0" "(n, f)" n], assumption+, simp) apply (subst P_mod_mod[THEN sym, of "I \\<^sub>r\<^bsub>S\<^esub> J" "polyn_expr R X n (sp_cf S (g 0) (n, f))" "sp_cf S (g 0) (n, f)"], assumption+, simp add:polyn_mem, simp add:sp_cf_pol_coeff, rule polyn_mem, simp add:sp_cf_pol_coeff, simp add:sp_cf_len, simp, simp add:sp_cf_len) apply (frule P_mod_mod[THEN sym, of I "polyn_expr R X n (n, f)" "(n, f)"], assumption+, simp, simp, simp add:sp_cf_len, subst sp_cf_def, simp, simp add:P_mod_coeffTr[of J "g 0"]) apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, frule_tac h = "f j" in Ring.ideal_subset[of S I], assumption+, simp add:Ring.ring_tOp_commute[of S "g 0"]) apply (simp add:Ring.prod_mem_prod_ideals[of S I J]) done lemma (in PolynRg) P_mod_mult_const1:"\ideal S I; ideal S J; pol_coeff S (n, f); P_mod R S X I (polyn_expr R X n (n, f)); d \ J\ \ P_mod R S X (I \\<^sub>r\<^bsub>S\<^esub> J) ((polyn_expr R X n (n, f)) \\<^sub>r d)" apply (cut_tac subring, frule subring_Ring) apply (frule P_mod_coeffTr[THEN sym, of J d], simp add:Ring.ideal_subset, simp) apply (frule P_mod_mult_const[of I J n f "\j. d"], assumption+, simp add:pol_coeff_def, simp add:Ring.ideal_subset) apply (subst polyn_expr_def, simp, frule Ring.ideal_subset[of S J d], assumption+, frule mem_subring_mem_ring[of S d], assumption, simp add:ring_r_one) apply (simp add:polyn_expr_def[of _ _ 0], frule Ring.ideal_subset[of S J d], assumption+, frule mem_subring_mem_ring[of S d], assumption, simp add:ring_r_one) done lemma (in PolynRg) P_mod_mult_monomial:"\ideal S I; p \ carrier R\ \ (P_mod R S X I p ) = (P_mod R S X I (p \\<^sub>r X^\<^bsup>R m\<^esup>))" apply (cut_tac X_mem_R, cut_tac subring, frule subring_Ring) apply (frule npClose[of X m], simp add:ring_tOp_commute[of p ]) apply (case_tac "p = \\<^bsub>R\<^esub>", simp add:ring_times_x_0) apply (rule iffI) apply (frule s_cf_expr[of p], assumption+, (erule conjE)+, cut_tac low_deg_terms_zero[THEN sym, of "fst (s_cf R S X p)" "snd (s_cf R S X p)" m], simp add:polyn_expr_split[THEN sym], thin_tac "X^\<^bsup>R m\<^esup> \\<^sub>r p = polyn_expr R X (fst (s_cf R S X p) + m) (ext_cf S m (s_cf R S X p))", frule ext_cf_pol_coeff[of "s_cf R S X p" m]) apply (frule P_mod_mod[THEN sym, of I p "s_cf R S X p"], assumption+, thin_tac "p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)", simp) apply (frule P_mod_mod[THEN sym, of I "polyn_expr R X (fst (s_cf R S X p) + m) (ext_cf S m (s_cf R S X p))" "ext_cf S m (s_cf R S X p)"], rule polyn_mem, assumption, simp add:ext_cf_def, assumption, simp add:ext_cf_len add.commute, simp, thin_tac "P_mod R S X I (polyn_expr R X (fst (s_cf R S X p) + m) (ext_cf S m (s_cf R S X p))) = (\j\fst (ext_cf S m (s_cf R S X p)). snd (ext_cf S m (s_cf R S X p)) j \ I)", thin_tac "snd (s_cf R S X p) (fst (s_cf R S X p)) \ \\<^bsub>S\<^esub>") apply (rule allI, rule impI, simp add:ext_cf_len) apply ( subst ext_cf_def, simp add:sliden_def) apply (rule impI, simp add:Ring.ideal_zero[of S]) apply (simp add:pol_coeff_split[THEN sym]) apply (frule s_cf_expr[of p], assumption+, (erule conjE)+, cut_tac low_deg_terms_zero[THEN sym, of "fst (s_cf R S X p)" "snd (s_cf R S X p)" m], simp add:polyn_expr_split[THEN sym], thin_tac "X^\<^bsup>R m\<^esup> \\<^sub>r p = polyn_expr R X (fst (s_cf R S X p) + m) (ext_cf S m (s_cf R S X p))", frule ext_cf_pol_coeff[of "s_cf R S X p" m]) apply (frule P_mod_mod[THEN sym, of I "polyn_expr R X (fst (s_cf R S X p) + m) (ext_cf S m (s_cf R S X p))" "ext_cf S m (s_cf R S X p)"], rule polyn_mem, assumption, simp add:ext_cf_def, assumption, simp add:ext_cf_len add.commute, simp, thin_tac "p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)", thin_tac "P_mod R S X I (polyn_expr R X (fst (s_cf R S X p) + m) (ext_cf S m (s_cf R S X p)))", thin_tac "snd (s_cf R S X p) (fst (s_cf R S X p)) \ \\<^bsub>S\<^esub>") apply (subst P_mod_mod[THEN sym, of I p "s_cf R S X p"], assumption+, cut_tac s_cf_expr[of p], simp, assumption+, rule allI, rule impI, thin_tac "pol_coeff S (ext_cf S m (s_cf R S X p))", simp add:ext_cf_len, simp add:ext_cf_def) apply (drule_tac x = "m + j" in spec, frule_tac i = j and j = "fst (s_cf R S X p)" and k = m and l = m in add_le_mono, simp, simp only:add.commute[of _ m], thin_tac "j \ fst (s_cf R S X p)", simp, simp add:sliden_def) apply simp done lemma (in PolynRg) P_mod_multTr:"\ideal S I; ideal S J; pol_coeff S (n, f); P_mod R S X I (polyn_expr R X n (n, f))\ \ \g. ((pol_coeff S (m, g) \ (P_mod R S X J (polyn_expr R X m (m, g)))) \ P_mod R S X (I \\<^sub>r\<^bsub>S\<^esub> J) ((polyn_expr R X n (n, f)) \\<^sub>r (polyn_expr R X m (m, g))))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag, cut_tac X_mem_R) apply (frule polyn_mem[of "(n, f)" n], simp) apply (frule Ring.ideal_prod_ideal[of "S" "I" "J"], assumption+) apply (case_tac "polyn_expr R X n (n, f) = \\<^bsub>R\<^esub>", simp) apply (rule allI, rule impI, erule conjE) apply (frule_tac c = "(m, g)" in polyn_mem[of _ m], simp, simp add:ring_times_0_x, simp add:P_mod_def) apply (induct_tac m) apply (rule allI, rule impI, erule conjE, rule_tac g = g in P_mod_mult_const[of I J n f], assumption+) (* case m = 0 done *) apply (rule allI, rule impI, erule conjE) apply (frule_tac n = na and f = g in pol_coeff_pre, frule_tac n = na and f = g in P_mod_pre[of J], assumption+) apply (drule_tac a = g in forall_spec, simp) apply (frule_tac n = na and f = g in polyn_Suc_split, simp del:npow_suc) apply (thin_tac "polyn_expr R X (Suc na) (Suc na, g) = polyn_expr R X na (na, g) \ g (Suc na) \\<^sub>r X^\<^bsup>R (Suc na)\<^esup>") apply (frule_tac c = "(na, g)" and k = na in polyn_mem, simp, subgoal_tac "(g (Suc na)) \\<^sub>r (X^\<^bsup>R (Suc na)\<^esup>) \ carrier R", subst ring_distrib1, assumption+) apply (frule_tac p = "(polyn_expr R X n (n, f)) \\<^sub>r (polyn_expr R X na (na, g))" and q = "(polyn_expr R X n (n, f)) \\<^sub>r ((g (Suc na)) \\<^sub>r (X^\<^bsup>R (Suc na)\<^esup>))" in P_mod_add[of "I \\<^sub>r\<^bsub>S\<^esub> J"]) apply (simp add:ring_tOp_closed, rule ring_tOp_closed, assumption+) apply (frule_tac c = "(Suc na, g)" and j = "Suc na" in pol_coeff_mem_R, simp) apply (subst ring_tOp_assoc[THEN sym], assumption+, simp, rule npClose, assumption+) apply (subst P_mod_mult_monomial[THEN sym, of "I \\<^sub>r\<^bsub>S\<^esub> J"], assumption, rule ring_tOp_closed, assumption+, simp add:pol_coeff_mem_R) apply (rule P_mod_mult_const1, assumption+, thin_tac "P_mod R S X (I \\<^sub>r\<^bsub>S\<^esub> J) (polyn_expr R X n (n, f) \\<^sub>r polyn_expr R X na (na, g))") apply (cut_tac n1 = na and c1 = "(Suc na, g)" in polyn_Suc[THEN sym], simp, simp, frule_tac c = "(Suc na, g)" and k = na in polyn_expr_short, simp, simp, thin_tac "P_mod R S X J (polyn_expr R X na (na, g))", thin_tac "polyn_expr R X na (na, g) \ carrier R", thin_tac "polyn_expr R X na (na, g) \ g (Suc na) \\<^sub>r (X^\<^bsup>R na\<^esup> \\<^sub>r X) = polyn_expr R X (Suc na) (Suc na, g)", thin_tac "polyn_expr R X na (Suc na, g) = polyn_expr R X na (na, g)") apply (frule_tac p1 = "polyn_expr R X (Suc na) (Suc na, g)" and c1 = "(Suc na, g)" in P_mod_mod[THEN sym, of J], simp add:polyn_mem, assumption, simp, simp) apply (simp, rule ring_tOp_closed, cut_tac c = "(Suc na, g)" and j = "Suc na" in pol_coeff_mem_R, assumption, simp, simp, rule npClose, assumption+) done lemma (in PolynRg) P_mod_mult:"\ideal S I; ideal S J; pol_coeff S (n, c); pol_coeff S (m, d); P_mod R S X I (polyn_expr R X n (n, c)); P_mod R S X J (polyn_expr R X m (m, d))\ \ P_mod R S X (I \\<^sub>r\<^bsub>S\<^esub> J) ((polyn_expr R X n (n, c)) \\<^sub>r (polyn_expr R X m (m, d)))" apply (simp add:P_mod_multTr) done lemma (in PolynRg) P_mod_mult1:"\ideal S I; ideal S J; p \ carrier R; q \ carrier R; P_mod R S X I p; P_mod R S X J q\ \ P_mod R S X (I \\<^sub>r\<^bsub>S\<^esub> J) (p \\<^sub>r q)" apply (case_tac "p = \\<^bsub>R\<^esub>") apply (simp add:ring_times_0_x, simp add:P_mod_def) apply (case_tac "q = \\<^bsub>R\<^esub>") apply (simp add:ring_times_x_0, simp add:P_mod_def) apply (frule s_cf_expr[of p], assumption+, frule s_cf_expr[of q], assumption+, (erule conjE)+) apply (cut_tac P_mod_mult[of I J "fst (s_cf R S X p)" "snd (s_cf R S X p)" "fst (s_cf R S X q)" "snd (s_cf R S X q)"]) apply (simp add:polyn_expr_split[THEN sym], assumption+) apply (simp add:pol_coeff_split[THEN sym]) apply (simp add:polyn_expr_split[THEN sym])+ done lemma (in PolynRg) P_mod_mult2l:"\ideal S I; p \ carrier R; q \ carrier R; P_mod R S X I p\ \ P_mod R S X I (p \\<^sub>r q)" apply (cut_tac subring, frule subring_Ring[of S], frule Ring.whole_ideal[of S]) apply (frule P_mod_whole[of q]) apply (frule P_mod_mult1[of I "carrier S" p q], assumption+) apply (simp add:Ring.idealprod_whole_r) done lemma (in PolynRg) P_mod_mult2r:"\ideal S I; p \ carrier R; q \ carrier R; P_mod R S X I q\ \ P_mod R S X I (p \\<^sub>r q)" apply (cut_tac subring, frule subring_Ring[of S], frule Ring.whole_ideal[of S]) apply (frule P_mod_whole[of p]) apply (frule P_mod_mult1[of "carrier S" I p q], assumption+) apply (simp add:Ring.idealprod_whole_l) done lemma (in PolynRg) csrp_fn_pol_coeff:"\ideal S P; PolynRg R' (S /\<^sub>r P) Y; pol_coeff (S /\<^sub>r P) (n, c')\ \ pol_coeff S (n, (cmp (csrp_fn S P) c'))" apply (cut_tac subring, frule subring_Ring) apply (simp add:pol_coeff_def) apply (rule allI, rule impI, simp add:cmp_def) apply (rule Ring.csrp_fn_mem[of S P], assumption+) apply simp done lemma (in PolynRg) pj_csrp_mem_coeff:"\ideal S P; pol_coeff (S /\<^sub>r P) (n, c')\ \ \j \ n. (pj S P) ((csrp_fn S P) (c' j)) = c' j" apply (cut_tac subring, frule subring_Ring) apply (rule allI, rule impI, simp add:pol_coeff_def) apply (simp add:Ring.csrp_pj) done lemma (in PolynRg) pHom_pj_csrp:"\Idomain S; ideal S P; PolynRg R' (S /\<^sub>r P) Y; pol_coeff (S /\<^sub>r P) (n, c')\ \ erH R S X R' (S /\<^sub>r P) Y (pj S P) (polyn_expr R X n (n, (cmp (csrp_fn S P) c'))) = polyn_expr R' Y n (n, c')" apply (cut_tac subring, frule subring_Ring, frule Ring.qring_ring[of "S" "P"], assumption+) apply (subst pHom_mem[of R' "(S /\<^sub>r P)" Y "erH R S X R' (S /\<^sub>r P) Y (pj S P)" n "cmp (csrp_fn S P) c'"], assumption+, rule erH_rHom[of R' "S /\<^sub>r P" Y "pj S P"], assumption+, simp add:pj_Hom, simp add:csrp_fn_pol_coeff) apply (rule PolynRg.polyn_exprs_eq[of R' "S /\<^sub>r P" Y "(n, cmp (erH R S X R' (S /\<^sub>r P) Y (pj S P)) (cmp (csrp_fn S P) c'))" "(n, c')" n], assumption+) apply (frule csrp_fn_pol_coeff[of P R' Y n c'], assumption+, frule erH_rHom [of R' "S /\<^sub>r P" Y "pj S P"], assumption+, simp add:pj_Hom, rule cmp_pol_coeff_e[of R' "S /\<^sub>r P" Y "erH R S X R' (S /\<^sub>r P) Y (pj S P)" n "cmp (csrp_fn S P) c'"], assumption+, simp) apply (rule allI, rule impI, simp add:cmp_def, frule_tac c = "(n, c')" and j = j in PolynRg.pol_coeff_mem[of R' "S /\<^sub>r P" Y], assumption+, simp+, frule_tac x = "c' j" in Ring.csrp_fn_mem[of S P], assumption+, frule_tac s = "csrp_fn S P (c' j)" in erH_rHom_cf[of R' "S /\<^sub>r P" Y "pj S P"], assumption+, simp add:pj_Hom, assumption+) apply (simp add:pj_csrp_mem_coeff) done lemma (in PolynRg) ext_csrp_fn_nonzero:"\Idomain S; ideal S P; PolynRg R' (S /\<^sub>r P) Y; g' \ carrier R'; g' \ \\<^bsub>R'\<^esub> \ \ polyn_expr R X (deg_n R' (S /\<^sub>r P) Y g') ((deg_n R' (S /\<^sub>r P) Y g'), (cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))) \ \" apply (cut_tac subring, frule subring_Ring, frule Ring.qring_ring[of "S" "P"], assumption+, frule pj_Hom[of "S" "P"], assumption+, frule PolynRg.s_cf_expr[of R' "S /\<^sub>r P" Y g'], assumption+, (erule conjE)+) apply (simp add:PolynRg.s_cf_deg[THEN sym, of R' "S /\<^sub>r P" Y g'], frule csrp_fn_pol_coeff[of P R' Y "deg_n R' (S /\<^sub>r P) Y g'" "snd (s_cf R' (S /\<^sub>r P) Y g')"], assumption+, simp add:PolynRg.s_cf_deg[of R' "S /\<^sub>r P" Y g']) apply (subst coeff_nonzero_polyn_nonzero[of "(deg_n R' (S /\<^sub>r P) Y g', cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))" "deg_n R' (S /\<^sub>r P) Y g'"], assumption+, simp) apply (simp add:cmp_def, rule contrapos_pp, simp+) apply (drule_tac a = "deg_n R' (S /\<^sub>r P) Y g'" in forall_spec, simp, frule pj_csrp_mem_coeff[of P "deg_n R' (S /\<^sub>r P) Y g'" "snd (s_cf R' (S /\<^sub>r P) Y g')"], simp add:PolynRg.s_cf_deg[of R' "S /\<^sub>r P" Y g']) apply (drule_tac a = "deg_n R' (S /\<^sub>r P) Y g'" in forall_spec, simp, simp, frule pj_Hom[of S P], assumption, simp add:rHom_0_0) done lemma (in PolynRg) erH_inv:"\Idomain S; ideal S P; Ring R'; PolynRg R' (S /\<^sub>r P) Y; g' \ carrier R'\ \ \g\carrier R. deg R S X g \ (deg R' (S /\<^sub>r P) Y g') \ (erH R S X R' (S /\<^sub>r P) Y (pj S P)) g = g'" apply (cut_tac subring, frule subring_Ring, frule Ring.qring_ring[of "S" "P"], assumption+, frule pj_Hom[of "S" "P"], assumption+) apply (frule erH_rHom[of R' "S /\<^sub>r P" Y "pj S P"], assumption+) apply (case_tac "g' = \\<^bsub>R'\<^esub>", simp, frule erH_rHom_0[of R' "S /\<^sub>r P" Y "pj S P"], assumption+, cut_tac ring_zero, subgoal_tac "deg R S X (\) \ deg R' (S /\<^sub>r P) Y g'", blast, simp add:deg_def) apply (frule PolynRg.s_cf_expr [of R' "S /\<^sub>r P" Y g'], assumption+, (erule conjE)+) apply (frule pHom_pj_csrp[of P R' Y "fst (s_cf R' (S /\<^sub>r P) Y g')" "snd (s_cf R' (S /\<^sub>r P) Y g')"], assumption+, simp add:PolynRg.pol_coeff_split[THEN sym], drule sym, simp) apply (subgoal_tac "deg R S X (polyn_expr R X (fst (s_cf R' (S /\<^sub>r P) Y g')) (fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))) \ deg R' (S /\<^sub>r P) Y g'", subgoal_tac "polyn_expr R X (fst (s_cf R' (S /\<^sub>r P) Y g')) (fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g'))) \ carrier R", blast) apply(thin_tac " deg R S X (polyn_expr R X (fst (s_cf R' (S /\<^sub>r P) Y g')) (fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))) \ deg R' (S /\<^sub>r P) Y g'", thin_tac "polyn_expr R' Y (fst (s_cf R' (S /\<^sub>r P) Y g')) (s_cf R' (S /\<^sub>r P) Y g') = g'", thin_tac "erH R S X R' (S /\<^sub>r P) Y (pj S P) (polyn_expr R X (fst (s_cf R' (S /\<^sub>r P) Y g')) (fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))) = g'", thin_tac "snd (s_cf R' (S /\<^sub>r P) Y g') (fst (s_cf R' (S /\<^sub>r P) Y g')) \ \\<^bsub>S /\<^sub>r P\<^esub>") apply (rule_tac c = "(fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))" and k = "fst (s_cf R' (S /\<^sub>r P) Y g')" in polyn_mem) apply (rule csrp_fn_pol_coeff, assumption+, simp, simp, cut_tac pol_deg_le_n[of "polyn_expr R X (fst (s_cf R' (S /\<^sub>r P) Y g')) (fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))" "(fst (s_cf R' (S /\<^sub>r P) Y g'), cmp (csrp_fn S P) (snd (s_cf R' (S /\<^sub>r P) Y g')))"]) apply (simp, simp add:PolynRg.s_cf_deg[THEN sym, of R' "S /\<^sub>r P" Y g'], frule ext_csrp_fn_nonzero[of P R' Y g'], assumption+, simp add:deg_def, simp add:ale_natle, rule polyn_mem, simp add:csrp_fn_pol_coeff, simp, simp add:csrp_fn_pol_coeff, simp) done lemma (in PolynRg) P_mod_0:"\Idomain S; ideal S P; PolynRg R' (S /\<^sub>r P) Y; g \ carrier R\ \ (erH R S X R' (S /\<^sub>r P) Y (pj S P) g = \\<^bsub>R'\<^esub>) = (P_mod R S X P g)" apply (cut_tac subring, frule subring_Ring, frule Ring.qring_ring[of "S" "P"], assumption+, frule pj_Hom[of "S" "P"], assumption+) apply (case_tac "g = \\<^bsub>R\<^esub>", cut_tac ring_zero, simp add:P_mod_def, rule erH_rHom_0[of R' "S /\<^sub>r P" Y "pj S P"], assumption+) apply (frule s_cf_expr[of g], assumption+, (erule conjE)+, cut_tac polyn_expr_split[of "fst (s_cf R S X g)" "s_cf R S X g"]) apply (frule erH_map[of R' "S /\<^sub>r P" Y "pj S P" "fst (s_cf R S X g)" "snd (s_cf R S X g)"], assumption+) apply (subst pol_coeff_split[THEN sym], assumption) apply (drule sym, simp) apply (thin_tac "erH R S X R' (S /\<^sub>r P) Y (pj S P) g = polyn_expr R' Y (fst (s_cf R S X g)) (fst (s_cf R S X g), cmp (pj S P) (snd (s_cf R S X g)))") apply (rotate_tac -1, drule sym) apply (subst P_mod_mod[THEN sym, of P g "s_cf R S X g"], assumption+, thin_tac "g = polyn_expr R X (fst (s_cf R S X g)) (s_cf R S X g)", frule erH_rHom_coeff[of R' "S /\<^sub>r P" Y "pj S P" "fst (s_cf R S X g)" "snd (s_cf R S X g)"], assumption+, simp) apply (subst PolynRg.coeff_0_pol_0[THEN sym, of R' "S /\<^sub>r P" Y "(fst (s_cf R S X g), cmp (pj S P) (snd (s_cf R S X g)))" "fst (s_cf R S X g)"], assumption+, simp, thin_tac "pol_coeff (S /\<^sub>r P) (fst (s_cf R S X g), cmp (pj S P) (snd (s_cf R S X g)))") apply (simp add:cmp_def) apply (rule iffI) apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, frule_tac j = j in pol_coeff_mem[of "s_cf R S X g"], assumption+, simp add:pj_zero[of S P]) apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, frule_tac j = j in pol_coeff_mem[of "s_cf R S X g"], assumption+, simp add:pj_zero[THEN sym, of S P]) done lemma (in PolynRg) P_mod_I_J:"\p \ carrier R; ideal S I; ideal S J; I \ J; P_mod R S X I p\ \ P_mod R S X J p" apply (case_tac "p = \\<^bsub>R\<^esub>", simp) apply (simp add:P_mod_def) apply (frule s_cf_expr[of p], assumption, (erule conjE)+) apply (frule P_mod_mod[THEN sym, of I p "s_cf R S X p"], assumption+) apply (subst P_mod_mod[THEN sym, of J p "s_cf R S X p"], assumption+, thin_tac "p = polyn_expr R X (fst (s_cf R S X p)) (s_cf R S X p)", simp) apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption, simp add:subsetD) done lemma (in PolynRg) P_mod_n_1:"\Idomain S; t \ carrier S; g \ carrier R; P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc n)\<^esup>)) g\ \ P_mod R S X (S \\<^sub>p t) g" apply (cut_tac subring, frule subring_Ring, frule Ring.npClose[of S t n], assumption+, frule Ring.npClose[of S t "Suc n"], assumption+, frule Ring.principal_ideal[of S t], assumption+, frule Ring.principal_ideal[of S "t^\<^bsup>S (Suc n)\<^esup>"], assumption+) apply (case_tac "g = \\<^bsub>R\<^esub>", simp add:P_mod_def) apply (frule s_cf_expr[of g], assumption, (erule conjE)+, subst P_mod_mod[THEN sym, of "S \\<^sub>p t" "g" "s_cf R S X g"], assumption+) apply (frule_tac P_mod_mod[THEN sym, of "S \\<^sub>p (t^\<^bsup>S (Suc n)\<^esup>)" g "s_cf R S X g"], assumption+) apply (simp del:npow_suc, thin_tac "g = polyn_expr R X (fst (s_cf R S X g)) (s_cf R S X g)") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption+) apply (simp add:Rxa_def, erule bexE, simp, simp add:Ring.ring_tOp_assoc[THEN sym, of S], frule_tac x = r and y = "t^\<^bsup>S n\<^esup>" in Ring.ring_tOp_closed, assumption+, blast) done lemma (in PolynRg) P_mod_n_m:"\Idomain S; t \ carrier S; g \ carrier R; m \ n; P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc n)\<^esup>)) g\ \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) g" apply (cut_tac subring, frule subring_Ring) apply (rule P_mod_I_J[of g "S \\<^sub>p (t^\<^bsup>S (Suc n)\<^esup>)" "S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)"], assumption) apply (rule Ring.principal_ideal, assumption+, rule Ring.npClose, assumption+) apply (rule Ring.principal_ideal, assumption+, rule Ring.npClose, assumption+) apply (thin_tac "P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc n)\<^esup>)) g") apply (rule subsetI) apply (simp del:npow_suc add:Rxa_def, erule bexE, simp del:npow_suc) apply (frule Ring.npMulDistr[THEN sym, of S t "Suc n - Suc m" "Suc m"], assumption) apply (simp del:npow_suc, thin_tac "t^\<^bsup>S (Suc n)\<^esup> = t^\<^bsup>S (n - m)\<^esup> \\<^sub>r\<^bsub>S\<^esub> t^\<^bsup>S (Suc m)\<^esup>", thin_tac "x = r \\<^sub>r\<^bsub>S\<^esub> (t^\<^bsup>S (n - m)\<^esup> \\<^sub>r\<^bsub>S\<^esub> t^\<^bsup>S (Suc m)\<^esup>)") apply (subst Ring.ring_tOp_assoc[THEN sym, of S], assumption+, (rule Ring.npClose, assumption+)+) apply (frule_tac x = r and y = "t^\<^bsup>S (n - m)\<^esup>" in Ring.ring_tOp_closed, assumption+, rule Ring.npClose, assumption+, blast) apply assumption done lemma (in PolynRg) P_mod_diff:"\Idomain S; ideal S P; PolynRg R' (S /\<^sub>r P) Y; g \ carrier R; h \ carrier R\ \ (erH R S X R' (S /\<^sub>r P) Y (pj S P) g = (erH R S X R' (S /\<^sub>r P) Y (pj S P) h)) = (P_mod R S X P (g \ -\<^sub>a h))" apply (cut_tac ring_is_ag, frule PolynRg.is_Ring[of R'], cut_tac subring, frule subring_Ring, frule Ring.qring_ring[of S P], assumption+, frule pj_Hom[of "S" "P"], assumption+, frule erH_rHom[of R' "S /\<^sub>r P" Y "pj S P"], assumption+, frule Ring.ring_is_ag[of R']) apply (frule erH_mem[of R' "S /\<^sub>r P" Y "pj S P" g], assumption+, frule erH_mem[of R' "S /\<^sub>r P" Y "pj S P" h], assumption+) apply (rule iffI) apply (frule_tac a = "erH R S X R' (S /\<^sub>r P) Y (pj S P) g" and b = "erH R S X R' (S /\<^sub>r P) Y (pj S P) h" in aGroup.ag_eq_diffzero[of R'], assumption+, simp, simp add:erH_minus[THEN sym, of R' "S /\<^sub>r P" Y "pj S P" h], drule sym, simp, thin_tac "erH R S X R' (S /\<^sub>r P) Y (pj S P) h = erH R S X R' (S /\<^sub>r P) Y (pj S P) g", frule_tac x = h in aGroup.ag_mOp_closed, assumption+, simp add:erH_add[THEN sym, of R' "S /\<^sub>r P" Y "pj S P" g "-\<^sub>a h"]) apply (subst P_mod_0[THEN sym, of P R' Y "g \ -\<^sub>a h"], assumption+, rule aGroup.ag_pOp_closed, assumption+) apply (frule_tac a = "erH R S X R' (S /\<^sub>r P) Y (pj S P) g" and b = "erH R S X R' (S /\<^sub>r P) Y (pj S P) h" in aGroup.ag_eq_diffzero[of R'], assumption+, simp, simp add:erH_minus[THEN sym, of R' "S /\<^sub>r P" Y "pj S P" h]) apply (subst erH_add[THEN sym, of R' "S /\<^sub>r P" Y "pj S P" g "-\<^sub>a h"], assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (subst P_mod_0[of P R' Y "g \ -\<^sub>a h"], assumption+, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) done lemma (in PolynRg) P_mod_erH:"\Idomain S; ideal S P; PolynRg R' (S /\<^sub>r P) Y; g \ carrier R; v \ carrier R; t \ P \ \ (erH R S X R' (S /\<^sub>r P) Y (pj S P) g = (erH R S X R' (S /\<^sub>r P) Y (pj S P) (g \ (t \\<^sub>r v))))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag, frule Ring.ideal_subset[of S P t], assumption+, frule mem_subring_mem_ring[of S t], assumption+, frule ring_tOp_closed[of t v], assumption+) apply (subst P_mod_diff[of P R' Y g "g \ (t \\<^sub>r v)"], assumption+, rule aGroup.ag_pOp_closed, assumption+) apply (simp add:aGroup.ag_p_inv, frule aGroup.ag_mOp_closed[of R g], assumption+, frule aGroup.ag_mOp_closed[of R "t \\<^sub>r v"], assumption+, subst aGroup.ag_pOp_assoc[THEN sym], assumption+, simp add:aGroup.ag_r_inv1, simp add:aGroup.ag_l_zero) apply (rule P_mod_minus[of P "t \\<^sub>r v"], assumption+, frule P_mod_mult1[of P "carrier S" t v], simp add:Ring.whole_ideal, assumption+, subst P_mod_coeffTr[of P t], assumption+, rule P_mod_whole[of v], assumption+, simp add:Ring.idealprod_whole_r[of S P]) done lemma (in PolynRg) coeff_principalTr:"\t \ carrier S\ \ \f. pol_coeff S (n, f) \ (\j \ n. f j \ S \\<^sub>p t) \ (\f'. pol_coeff S (n, f') \ (\j \ n. f j = t \\<^sub>r\<^bsub>S\<^esub> (f' j)))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag) apply (induct_tac n, rule allI, rule impI, erule conjE, simp add:Rxa_def, erule bexE, simp add:Ring.ring_tOp_commute[of S _ t], subgoal_tac "pol_coeff S (0, (\j. r))", subgoal_tac "t \\<^sub>r\<^bsub>S\<^esub> r = t \\<^sub>r\<^bsub>S\<^esub> ((\j. r) 0)", blast, simp, simp add:pol_coeff_def) apply (rule allI, rule impI, erule conjE, frule_tac n = n and f = f in pol_coeff_pre, subgoal_tac "\j \ n. f j \ S \\<^sub>p t", drule_tac a = f in forall_spec, simp, erule exE, erule conjE, frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp, simp, drule_tac x = "Suc n" in spec, simp, simp add:Rxa_def, erule bexE, simp add:Ring.ring_tOp_commute[of "S" _ "t"]) apply (subgoal_tac "pol_coeff S ((Suc n), (\j. if j \ n then (f' j) else r)) \ (\j \ (Suc n). f j = t \\<^sub>r\<^bsub>S\<^esub> ((\j. if j \ n then (f' j) else r) j))", blast) apply (rule conjI, simp add:pol_coeff_def, rule allI, rule impI, case_tac "j \ n", simp) apply simp apply (drule_tac y = j and x = n in not_le_imp_less, drule_tac m = n and n = j in Suc_leI) apply (frule_tac m = j and n = "Suc n" in le_antisym, assumption, simp, thin_tac "\f. pol_coeff S (n, f) \ (\j\n. f j \ S \\<^sub>p t) \ (\f'. pol_coeff S (n, f') \ (\j\n. f j = t \\<^sub>r\<^bsub>S\<^esub> f' j))") apply (rule allI, rule impI, drule_tac a = j in forall_spec, simp+) done lemma (in PolynRg) coeff_principal:"\t \ carrier S; pol_coeff S (n, f); \j \ n. f j \ S \\<^sub>p t\ \ \f'. pol_coeff S (n, f') \ (\j \ n. f j = t \\<^sub>r\<^bsub>S\<^esub> (f' j))" apply (simp add:coeff_principalTr) done lemma (in PolynRg) Pmod_0_principal:"\Idomain S; t \ carrier S; g \ carrier R; P_mod R S X (S \\<^sub>p t) g\ \ \h\ carrier R. g = t \\<^sub>r h" apply (cut_tac subring, frule subring_Ring) apply (case_tac "g = \\<^bsub>R\<^esub>", cut_tac ring_zero, frule mem_subring_mem_ring[of S t], assumption+, frule ring_times_x_0[THEN sym, of t], blast) apply (frule s_cf_expr[of g], assumption+, (erule conjE)+, frule Ring.principal_ideal[of S t], assumption, simp add:P_mod_mod[THEN sym, of "S \\<^sub>p t" g], frule coeff_principal[of t "fst (s_cf R S X g)" "snd (s_cf R S X g)"], simp add:pol_coeff_split[THEN sym], assumption+, erule exE, erule conjE) apply (frule_tac c = "(fst (s_cf R S X g), f')" and k = "fst (s_cf R S X g)" in polyn_mem, simp, subgoal_tac "g = t \\<^sub>r (polyn_expr R X (fst (s_cf R S X g)) (fst (s_cf R S X g), f'))", blast) apply (subst scalar_times_pol_expr[of t "(fst (s_cf R S X g), f')" "fst (s_cf R S X g)"], assumption+, simp, drule sym, subgoal_tac "polyn_expr R X (fst (s_cf R S X g)) (s_cf R S X g) = polyn_expr R X (fst (s_cf R S X g)) (sp_cf S t (fst (s_cf R S X g), f'))", simp) apply (frule_tac c = "(fst (s_cf R S X g), f')" in sp_cf_pol_coeff[of _ t], assumption+, frule_tac d = "sp_cf S t (fst (s_cf R S X g), f')" in pol_expr_unique2[of "s_cf R S X g"], assumption+, simp, simp add:sp_cf_len, simp add:sp_cf_len, thin_tac "(g = polyn_expr R X (fst (s_cf R S X g)) (sp_cf S t (fst (s_cf R S X g), f'))) = (\j\fst (s_cf R S X g). t \\<^sub>r\<^bsub>S\<^esub> f' j = snd (sp_cf S t (fst (s_cf R S X g), f')) j)", thin_tac "polyn_expr R X (fst (s_cf R S X g)) (s_cf R S X g) = g", thin_tac "polyn_expr R X (fst (s_cf R S X g)) (fst (s_cf R S X g), f') \ carrier R") apply (rule allI, rule impI, drule_tac a = j in forall_spec, assumption+, simp add:sp_cf_def) done lemma (in PolynRg) Pmod0_principal_rev:"\Idomain S; t \ carrier S; g \ carrier R; \h\ carrier R. g = t \\<^sub>r h\ \ P_mod R S X (S \\<^sub>p t) g" apply (cut_tac subring, frule subring_Ring) apply (erule bexE) apply (case_tac "t = \\<^bsub>S\<^esub>", frule Subring_zero_ring_zero, simp) apply (simp add:ring_times_0_x, simp add:P_mod_def) apply (case_tac "h = \\<^bsub>R\<^esub>", simp, frule mem_subring_mem_ring[of S t], assumption+, simp add:ring_times_x_0, simp add:P_mod_def, cut_tac polyn_ring_integral, simp) apply (frule_tac p = h in s_cf_expr, assumption+, (erule conjE)+, frule_tac c = "s_cf R S X h" and n = "fst (s_cf R S X h)" in scalar_times_pol_expr[of t], assumption+, simp, thin_tac "g = t \\<^sub>r h", drule sym, simp) apply (frule Ring.principal_ideal[of S t], assumption+, frule_tac c1 = "sp_cf S t (s_cf R S X h)" and p1 = "t \\<^sub>r h" in P_mod_mod[THEN sym], frule_tac x = t in mem_subring_mem_ring, assumption, rule ring_tOp_closed, assumption+, simp add:sp_cf_pol_coeff, simp add:sp_cf_len) apply (drule sym, simp, thin_tac "P_mod R S X (S \\<^sub>p t) (t \\<^sub>r h) = (\j\fst (sp_cf S t (s_cf R S X h)). snd (sp_cf S t (s_cf R S X h)) j \ S \\<^sub>p t)", thin_tac "polyn_expr R X (fst (s_cf R S X h)) (sp_cf S t (s_cf R S X h)) = t \\<^sub>r h", thin_tac "polyn_expr R X (fst (s_cf R S X h)) (s_cf R S X h) = h") apply (rule allI, rule impI, simp add:sp_cf_len, subst sp_cf_def, simp, subst Rxa_def, simp, frule_tac c = "s_cf R S X h" and j = j in pol_coeff_mem, assumption) apply (simp add:Ring.ring_tOp_commute[of S t], blast) done (** NOTE. if t \ 0\<^sub>S then, deg g = deg h, because deg t = 0 **) lemma (in PolynRg) Pmod0_principal_rev1:"\Idomain S; t \ carrier S; h \ carrier R\ \ P_mod R S X (S \\<^sub>p t) (t \\<^sub>r h)" apply (rule Pmod0_principal_rev[of t "t \\<^sub>r h"], assumption+) apply (cut_tac subring, frule mem_subring_mem_ring[of S t], assumption+, simp add:ring_tOp_closed) apply blast done lemma (in PolynRg) Pmod0_principal_erH_vanish_t:"\Idomain S; ideal S (S \\<^sub>p t); t \ carrier S; t \ \\<^bsub>S\<^esub>; PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y \ \ erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) t = \\<^bsub>R'\<^esub>" apply (cut_tac subring, frule subring_Ring, frule mem_subring_mem_ring[of S t], assumption+) apply (subst P_mod_0[of "S \\<^sub>p t" R' Y t], assumption+) apply (rule Pmod0_principal_rev[of t t], assumption+) apply (cut_tac ring_one, frule ring_r_one[THEN sym, of t], blast) done lemma (in PolynRg) P_mod_diffxxx1:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; g \ carrier R; h \ carrier R; f \ \; g \ \; h \ \; u \ carrier R; v \ carrier R; erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \ \\<^bsub>R'\<^esub>; erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h \ \\<^bsub>R'\<^esub>; ra \ carrier R; f \ -\<^sub>a (g \\<^sub>r h) = t^\<^bsup>S m\<^esup> \\<^sub>r ra; 0 < m; (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) u) \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \\<^bsub>R'\<^esub> (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) v) \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra\ \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) (f \ -\<^sub>a ((g \ t^\<^bsup>S m\<^esup> \\<^sub>r v) \\<^sub>r (h \ t^\<^bsup>S m\<^esup> \\<^sub>r u)))" apply (cut_tac is_Ring, cut_tac subring, frule subring_Ring, cut_tac ring_is_ag, frule PolynRg.is_Ring[of R' "S /\<^sub>r (S \\<^sub>p t)" Y], frule Ring.ring_is_ag[of R'], frule Ring.maximal_ideal_ideal[of "S" "S \\<^sub>p t"], assumption+, frule Ring.qring_ring[of S "S \\<^sub>p t"], assumption+, frule erH_rHom[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)"], assumption+, frule mem_subring_mem_ring[of S t], assumption+) apply (rule pj_Hom[of S "S \\<^sub>p t"], assumption+, frule pHom_rHom[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t))"], assumption+) apply (simp del:npow_suc add:rHom_tOp[THEN sym]) apply (frule_tac ring_tOp_closed[of u g], assumption, frule_tac ring_tOp_closed[of v h], assumption) apply (simp del:npow_suc add:rHom_add[THEN sym]) apply (rotate_tac 17, drule sym) apply (frule P_mod_diff[of "S \\<^sub>p t" R' Y ra "u \\<^sub>r g \ v \\<^sub>r h"], assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, simp del:npow_suc) apply (frule Pmod_0_principal[of t "ra \ -\<^sub>a (u \\<^sub>r g \ v \\<^sub>r h)"], assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption, rule aGroup.ag_pOp_closed, assumption+, erule bexE) apply (frule Ring.npClose[of S t m], assumption+, frule mem_subring_mem_ring[of S "t^\<^bsup>S m\<^esup>"], assumption+, subst ring_distrib1, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, simp add:mem_subring_mem_ring, assumption+) apply (rule ring_tOp_closed, assumption+) apply (subst ring_distrib2, assumption+, rule ring_tOp_closed, assumption+ ) apply (frule_tac x = g and y = h in ring_tOp_closed, assumption+, frule_tac x = "t^\<^bsup>S m\<^esup>" and y = v in ring_tOp_closed, assumption+, frule_tac x = "t^\<^bsup>S m\<^esup>" and y = u in ring_tOp_closed, assumption+, frule_tac x = "t^\<^bsup>S m\<^esup> \\<^sub>r v" and y = h in ring_tOp_closed, assumption+) apply (subst ring_distrib2, assumption+, frule_tac x = "t^\<^bsup>S m\<^esup> \\<^sub>r v" and y = "t^\<^bsup>S m\<^esup> \\<^sub>r u" in ring_tOp_closed, assumption+) apply (subst aGroup.ag_p_inv, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+) apply (subst aGroup.ag_p_inv, assumption+, frule aGroup.ag_mOp_closed[of R "g \\<^sub>r h"], assumption+, frule aGroup.ag_mOp_closed[of R "t^\<^bsup>S m\<^esup> \\<^sub>r v \\<^sub>r h"], assumption+) apply (subst aGroup.ag_pOp_assoc[of R "-\<^sub>a (g \\<^sub>r h)" " -\<^sub>a (t^\<^bsup>S m\<^esup> \\<^sub>r v \\<^sub>r h)"], assumption+) apply (rule aGroup.ag_mOp_closed, assumption, rule aGroup.ag_pOp_closed, assumption, rule ring_tOp_closed, assumption+) apply (subst aGroup.ag_pOp_assoc[THEN sym], assumption+, rule aGroup.ag_pOp_closed, assumption, rule aGroup.ag_mOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+, simp del:npow_suc) apply (subst aGroup.ag_p_inv, assumption, rule ring_tOp_closed, assumption+, simp del:npow_suc add:ring_tOp_assoc[of "t^\<^bsup>S m\<^esup>" v h], simp add:del:npow_suc add:ring_tOp_commute[of g "t^\<^bsup>S m\<^esup> \\<^sub>r u"], simp del:npow_suc add:ring_tOp_assoc[of "t^\<^bsup>S m\<^esup>" u g], simp del:npow_suc add:ring_inv1_2, subst aGroup.ag_pOp_assoc[THEN sym], assumption+, rule ring_tOp_closed, assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, (rule ring_tOp_closed, assumption+)+, rule aGroup.ag_mOp_closed, assumption+, (rule ring_tOp_closed, assumption+)+, rule aGroup.ag_mOp_closed, assumption+) apply (subst ring_distrib1[THEN sym, of "t^\<^bsup>S m\<^esup>" ra "v \\<^sub>r (-\<^sub>a h)"], assumption+, rule ring_tOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (subst aGroup.ag_pOp_assoc[THEN sym], assumption+, rule ring_tOp_closed, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply ((rule ring_tOp_closed, assumption+)+, rule aGroup.ag_mOp_closed, assumption+, (rule ring_tOp_closed, assumption+)+, rule aGroup.ag_mOp_closed, assumption+) apply (subst ring_distrib1[THEN sym, of "t^\<^bsup>S m\<^esup>"], assumption+, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, rule ring_tOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (subst ring_tOp_assoc[of "t^\<^bsup>S m\<^esup>" v], assumption+, rule ring_tOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (subst ring_distrib1[THEN sym, of "t^\<^bsup>S m\<^esup>"], assumption+, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption, rule aGroup.ag_mOp_closed, assumption+, rule ring_tOp_closed, assumption, rule aGroup.ag_mOp_closed, assumption+, (rule ring_tOp_closed, assumption+)+, rule aGroup.ag_mOp_closed, assumption+) apply (frule ring_tOp_closed[of u g], assumption+, frule ring_tOp_closed[of v h], assumption+, simp del:npow_suc add:aGroup.ag_p_inv[of R "u \\<^sub>r g" "v \\<^sub>r h"], simp del:npow_suc add:add:ring_inv1_2, frule aGroup.ag_mOp_closed[of R g], assumption+, frule aGroup.ag_mOp_closed[of R h], assumption+, frule ring_tOp_closed[of u "-\<^sub>a g"], assumption+, frule ring_tOp_closed[of v "-\<^sub>a h"], assumption+, simp del:npow_suc add:aGroup.ag_pOp_commute[of R "u \\<^sub>r (-\<^sub>a g)" "v \\<^sub>r (-\<^sub>a h)"], simp del:npow_suc add:aGroup.ag_pOp_assoc[THEN sym, of R ra "v \\<^sub>r (-\<^sub>a h)" "u \\<^sub>r (-\<^sub>a g)"]) apply (subst ring_tOp_assoc[THEN sym, of v "t^\<^bsup>S m\<^esup>" "-\<^sub>a u"], assumption+, rule aGroup.ag_mOp_closed, assumption+, simp only:ring_tOp_commute[of v "t^\<^bsup>S m\<^esup>"], subgoal_tac "t^\<^bsup>S m\<^esup> \\<^sub>r v = t \\<^sub>r (t^\<^bsup>S (m - Suc 0)\<^esup> \\<^sub>r v)", simp del:npow_suc) apply (subst ring_tOp_assoc[of t], frule mem_subring_mem_ring[of S t], assumption+, rule ring_tOp_closed) apply (frule Ring.npClose[of S t "m - Suc 0"], assumption+, simp add:mem_subring_mem_ring, assumption, rule aGroup.ag_mOp_closed, assumption+, subst ring_distrib1[THEN sym, of t], simp add:mem_subring_mem_ring, assumption+) apply ((rule ring_tOp_closed)+, frule Ring.npClose[of S t "m - Suc 0"], assumption+, simp add:mem_subring_mem_ring, assumption, rule aGroup.ag_mOp_closed, assumption+) apply (subst ring_tOp_assoc[THEN sym], frule Ring.npClose[of S t "m - Suc 0"], assumption+, simp add:mem_subring_mem_ring) apply (rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, frule Ring.npClose[of S t "m - Suc 0"], assumption+, rule ring_tOp_closed, simp add:mem_subring_mem_ring, assumption, rule aGroup.ag_mOp_closed, assumption+, simp add:Subring_tOp_ring_tOp[THEN sym], simp only:npow_suc[THEN sym, of S t m]) apply (rule Pmod0_principal_rev1[of "t^\<^bsup>S (Suc m)\<^esup>"], assumption+, rule Ring.npClose, assumption+, rule aGroup.ag_pOp_closed, assumption+, (rule ring_tOp_closed)+, frule Ring.npClose[of S t "m - Suc 0"], assumption+, simp add:mem_subring_mem_ring, assumption, rule aGroup.ag_mOp_closed, assumption+) apply (frule Ring.npClose[of S t "m - Suc 0"], assumption+, frule mem_subring_mem_ring[of S t], assumption, frule mem_subring_mem_ring[of S "t^\<^bsup>S (m - Suc 0)\<^esup>"], assumption, simp add:ring_tOp_assoc[THEN sym], simp add:ring_tOp_commute[of t "t^\<^bsup>S (m - Suc 0)\<^esup>"], subgoal_tac "t^\<^bsup>S m\<^esup> = t^\<^bsup>S (Suc (m - Suc 0))\<^esup>", simp del:Suc_pred add:Subring_tOp_ring_tOp, simp only:Suc_pred) done lemma (in PolynRg) P_mod_diffxxx2:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; g \ carrier R; h \ carrier R; deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (f \ -\<^sub>a (g \\<^sub>r h)); 0 < m\ \ \g1 h1. g1 \carrier R \ h1 \ carrier R \ (deg R S X g1 \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g1)) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (g \ -\<^sub>a g1) \ (deg R S X h1 + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g1) \ deg R S X f) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (h \ -\<^sub>a h1) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) (f \ (-\<^sub>a (g1 \\<^sub>r h1)))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag, frule Ring.residue_field_cd[of S "S \\<^sub>p t"], assumption+, frule Ring.maximal_ideal_ideal[of "S" "S \\<^sub>p t"], assumption+, frule pj_Hom[of "S" "S \\<^sub>p t"], assumption+, frule mem_subring_mem_ring[of S t], assumption+, frule Ring.qring_ring[of S "S \\<^sub>p t"], assumption+, frule PolynRg.pol_nonzero[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g"], rule erH_mem, assumption+, frule erH_rHom_nonzero[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)" "g"], assumption+, simp add:aless_imp_le) apply (frule PolynRg.pol_nonzero[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h"], rule erH_mem, assumption+, frule erH_rHom_nonzero[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)" "h"], assumption+, simp add:aless_imp_le, simp del:npow_suc add:aless_imp_le) apply ( frule pol_nonzero[THEN sym, of "h"], simp del:npow_suc, frule aadd_pos_poss[of "deg R S X h" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)"], assumption+, frule aless_le_trans[of "0" "(deg R S X h) + (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g))" "deg R S X f"], assumption+, frule pol_nonzero[of f], simp del:npow_suc add:aless_imp_le) apply (thin_tac "0 < deg R S X f", thin_tac "0 < deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)") apply (frule Pmod_0_principal[of "t^\<^bsup>S m\<^esup>" "f \ -\<^sub>a (g \\<^sub>r h)"], rule Ring.npClose, assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, rule ring_tOp_closed, assumption+, erule bexE, rename_tac ra) (******* deg (t^\<^sup>S\<^sup> \<^sup>m) ra \ deg f ******) apply (frule deg_mult_pols1 [of g h], assumption+, frule aadd_le_mono[of "deg R S X g" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X h"]) apply (simp only:aadd_commute[of "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X h"]) apply (frule ale_trans[of "deg R S X g + deg R S X h" "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X f"], assumption+) apply (thin_tac "deg R S X g + deg R S X h \ deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)") apply (frule_tac ring_tOp_closed[of g h], assumption+, frule deg_minus_eq1[of "g \\<^sub>r h"], frule polyn_deg_add4[of "f" "-\<^sub>a (g \\<^sub>r h)" "deg_n R S X f"], rule aGroup.ag_mOp_closed, assumption+) apply (subst deg_an[THEN sym], assumption+, simp del:npow_suc) apply (simp add:deg_an[THEN sym], simp del:npow_suc add:deg_an[THEN sym], thin_tac "deg R S X (g \\<^sub>r h) = deg R S X g + deg R S X h", thin_tac "deg R S X g + deg R S X h \ deg R S X f", thin_tac "deg R S X (-\<^sub>a (g \\<^sub>r h)) = deg R S X g + deg R S X h") (******* deg (t^\<^sup>S\<^sup> \<^sup>m) ra \ deg f done *** next show deg ra \ deg f ***) apply (frule Ring.npClose[of S t m], assumption, frule Idomain.idom_potent_nonzero[of S t m], assumption+, frule_tac p = ra in const_times_polyn1[of _ "t^\<^bsup>S m\<^esup>"], assumption+, simp del:npow_suc) (****************** got deg ra \ deg f ***********************) (****** make g1 and h1 ******) apply (frule_tac h = "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra" in PolynRg.rel_prime_equation[of R' "(S /\<^sub>r (S \\<^sub>p t))" "Y" "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g" "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h"], assumption+, simp del:npow_suc add:erH_mem, simp del:npow_suc add:erH_mem, assumption+, simp del:npow_suc add:erH_mem) apply (erule bexE, erule bexE, (erule conjE)+, frule_tac erH_mem[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)" "g"], assumption+, frule_tac erH_mem[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)" "h"], assumption+) apply (rename_tac ra u' v') apply (frule_tac g' = v' in erH_inv[of "S \\<^sub>p t" R' Y], assumption+, simp add:PolynRg.is_Ring[of R'], assumption+) apply (frule_tac g' = u' in erH_inv[of "S \\<^sub>p t" R' Y ], assumption+, simp add:PolynRg.is_Ring[of R'], assumption+) apply ((erule bexE)+, rename_tac ra u' v' v u, (erule conjE)+) apply ( frule_tac p1 = u in erH_mult[THEN sym, of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)" _ "g"], assumption+, frule_tac p1 = v in erH_mult[THEN sym, of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)" _ "h"], assumption+, thin_tac "0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)", thin_tac "rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)") apply (subgoal_tac "g \ (t^\<^bsup>S m\<^esup>) \\<^sub>r v \ carrier R \ h \ (t^\<^bsup>S m\<^esup>) \\<^sub>r u \ carrier R \ deg R S X (g \ (t^\<^bsup>S m\<^esup>) \\<^sub>r v) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (g \ (t^\<^bsup>S m\<^esup>) \\<^sub>r v)) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (g \ -\<^sub>a (g \ (t^\<^bsup>S m\<^esup>) \\<^sub>r v)) \ deg R S X (h \ (t^\<^bsup>S m\<^esup>) \\<^sub>r u) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (g \ (t^\<^bsup>S m\<^esup>) \\<^sub>r v)) \ deg R S X f \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ( h \ -\<^sub>a (h \ (t^\<^bsup>S m\<^esup>) \\<^sub>r u)) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) ( f \ -\<^sub>a ((g \ (t^\<^bsup>S m\<^esup>) \\<^sub>r v) \\<^sub>r (h \ (t^\<^bsup>S m\<^esup>) \\<^sub>r u)))") apply (thin_tac "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f", thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y u' \ amax (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)) (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h))", thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y v' \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "u' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \\<^bsub>R'\<^esub> v' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \ carrier R'", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h \ carrier R'", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) v = v'", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) u = u'", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) u \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (u \\<^sub>r g)", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) v \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (v \\<^sub>r h)") apply blast apply (frule mem_subring_mem_ring[of "S" "t^\<^bsup>S m\<^esup>"], assumption) apply (rule conjI) apply (rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+) apply (rule conjI, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+) apply (frule Ring.a_in_principal[of "S" "t"], assumption+, frule Ring.maximal_ideal_ideal[of "S" "S \\<^sub>p t"], assumption+, frule Ring.ideal_npow_closed[of "S" "S \\<^sub>p t" "t" "m"], assumption+, frule PolynRg.is_Ring[of R' "S /\<^sub>r (S \\<^sub>p t)" Y], frule Ring.ring_is_ag[of R'], frule erH_rHom[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "pj S (S \\<^sub>p t)"], assumption+) apply (rule conjI) apply (frule pHom_dec_deg[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t))" g], assumption+, frule_tac i = "deg R S X v" and j = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y v'" and k = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" in ale_trans, assumption, thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y u' \ amax (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)) (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h))", thin_tac "u' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \\<^bsub>R'\<^esub> v' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra", thin_tac "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f") apply (subst P_mod_erH[THEN sym, of "S \\<^sub>p t" "R'" "Y" "g" _ "t^\<^bsup>S m\<^esup>"], assumption+, thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X g") apply (frule_tac p = v in const_times_polyn1[of _ "t^\<^bsup>S m\<^esup>"], assumption+, frule_tac q = "(t^\<^bsup>S m\<^esup>) \\<^sub>r v" in polyn_deg_add4[of "g" _ "deg_n R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)"]) apply (rule ring_tOp_closed, assumption+, simp del:npow_suc add:PolynRg.deg_an[THEN sym], simp add:PolynRg.deg_an[THEN sym], simp add:PolynRg.deg_an[THEN sym]) apply (rule conjI) apply (frule Ring.principal_ideal[of S "t^\<^bsup>S m\<^esup>"], assumption+, frule Ring.a_in_principal[of S "t^\<^bsup>S m\<^esup>"], assumption+) apply (frule_tac y = v in ring_tOp_closed[of "t^\<^bsup>S m\<^esup>"], assumption+, subst aGroup.ag_p_inv, assumption+, frule aGroup.ag_mOp_closed[of "R" "g"], assumption+, frule_tac x = "(t^\<^bsup>S m\<^esup>) \\<^sub>r v" in aGroup.ag_mOp_closed[of "R"], assumption+) apply (subst aGroup.ag_pOp_assoc[THEN sym], assumption+, subst aGroup.ag_r_inv1[of "R"], assumption+, subst aGroup.ag_l_zero[of "R"], assumption+, rule P_mod_minus, assumption+) apply (rule_tac g = "(t^\<^bsup>S m\<^esup>) \\<^sub>r v" in Pmod0_principal_rev[of "t^\<^bsup>S m\<^esup>"], assumption+) apply blast apply (rule conjI) apply (subst P_mod_erH[THEN sym, of "S \\<^sub>p t" R' Y g _ "t^\<^bsup>S m\<^esup>"], assumption+, thin_tac "P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (t^\<^bsup>S m\<^esup> \\<^sub>r ra)", thin_tac "f \ -\<^sub>a (g \\<^sub>r h) = t^\<^bsup>S m\<^esup> \\<^sub>r ra", thin_tac "u' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \\<^bsub>R'\<^esub> v' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra") apply (case_tac " (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)) \ (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h))") apply (simp add:amax_def) apply (frule_tac i = "deg R S X u" and j = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y u'" and k = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)" in ale_trans, assumption+, frule pHom_dec_deg[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t))" h], assumption+, frule_tac i = "deg R S X u" and j = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)" and k = "deg R S X h" in ale_trans, assumption+, frule_tac p = u and c = "t^\<^bsup>S m\<^esup>" in const_times_polyn1, assumption+) apply (frule_tac q = "(t^\<^bsup>S m\<^esup>) \\<^sub>r u" in polyn_deg_add4[of h _ "deg_n R S X h"], rule ring_tOp_closed, assumption+, subst deg_an[THEN sym], assumption+, rule ale_refl, subst deg_an[THEN sym], assumption+, simp, frule deg_an[THEN sym, of h], assumption+, simp) apply (frule_tac x = "deg R S X ( h \ (t^\<^bsup>S m\<^esup>) \\<^sub>r u)" in aadd_le_mono[of _ "deg R S X h" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)"], rule_tac i = "deg R S X ( h \ (t^\<^bsup>S m\<^esup>) \\<^sub>r u) + (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g))" in ale_trans[of _ "deg R S X h + (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g))" "deg R S X f"], assumption+) apply (simp add:amax_def) apply (thin_tac "\ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)") apply (subst aplus_le_aminus[of _ "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X f"]) apply (rule deg_in_aug_minf, rule aGroup.ag_pOp_closed, assumption+, rule ring_tOp_closed, assumption+, rule PolynRg.deg_in_aug_minf, assumption+, rule deg_in_aug_minf, assumption+) apply (subst PolynRg.deg_an, assumption+, simp add:minus_an_in_aug_minf, frule_tac y = u in ring_tOp_closed[of "t^\<^bsup>S m\<^esup>"], assumption+, frule_tac q = "(t^\<^bsup>S m\<^esup>) \\<^sub>r u" in polyn_deg_add5[of h _ "deg R S X f - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)"], assumption+, frule deg_in_aug_minf[of h], subst aplus_le_aminus[THEN sym, of "deg R S X h" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X f"], assumption+, rule PolynRg.deg_in_aug_minf, assumption+, rule deg_in_aug_minf, assumption+, subst PolynRg.deg_an, assumption+, simp add:minus_an_in_aug_minf, assumption) apply (subst const_times_polyn1, assumption+, frule_tac i = "deg R S X u" and j = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y u'" and k = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" in ale_trans, assumption+, frule_tac p = ra in pHom_dec_deg[of R' "S /\<^sub>r (S \\<^sub>p t)" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t))"], assumption+, frule_tac i = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra)" and j = "deg R S X ra" and k = "deg R S X f" in ale_trans, assumption+, frule_tac a = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra)" and a' = "deg R S X f" and b = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" in adiff_le_adiff, frule_tac i = "deg R S X u" and j = "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" and k = "deg R S X f - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" in ale_trans, assumption+) apply (rule conjI) apply (frule Ring.principal_ideal[of "S" "t^\<^bsup>S m\<^esup>"], assumption+, frule Ring.a_in_principal[of "S" "t^\<^bsup>S m\<^esup>"], assumption+, frule_tac y = u in ring_tOp_closed[of "t^\<^bsup>S m\<^esup>"], assumption+, subst aGroup.ag_p_inv, assumption+, frule aGroup.ag_mOp_closed[of "R" "g"], assumption+, frule_tac x = "(t^\<^bsup>S m\<^esup>) \\<^sub>r u" in aGroup.ag_mOp_closed[of "R"], assumption+, subst aGroup.ag_pOp_assoc[THEN sym], assumption+, rule aGroup.ag_mOp_closed, assumption+, subst aGroup.ag_r_inv1[of "R"], assumption+, subst aGroup.ag_l_zero[of "R"], assumption+, rule P_mod_minus, assumption+, rule_tac g = "(t^\<^bsup>S m\<^esup>) \\<^sub>r u" in Pmod0_principal_rev[of "t^\<^bsup>S m\<^esup>"], assumption+, thin_tac "deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f", thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y u' \ amax (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra) - deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)) (deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h))", thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y v' \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac " u' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g \\<^bsub>R'\<^esub> v' \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) ra", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) u \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (u \\<^sub>r g)", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) v \\<^sub>r\<^bsub>R'\<^esub> erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (v \\<^sub>r h)") apply blast apply (rule_tac u = u and v = v and ra = ra in P_mod_diffxxx1[of t R' Y f g h], assumption+) apply (rotate_tac -12, drule sym, drule sym, simp) done (** Hensel_next R S X t R' Y f m gh **) definition Hensel_next :: "[('a, 'b) Ring_scheme, ('a, 'c) Ring_scheme, 'a, 'a, ('a set, 'm) Ring_scheme, 'a set,'a, nat] \ ('a \ 'a) \ ('a \ 'a)" ("(9Hen\<^bsub> _ _ _ _ _ _ _\<^esub> _ _)" [67,67,67,67,67,67,67,68]67) where "Hen\<^bsub>R S X t R' Y f \<^esub> m gh = (SOME gh1. gh1 \ carrier R \ carrier R \ (deg R S X (fst gh1) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh1))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((fst gh) \\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (fst gh1)) \ (deg R S X (snd gh1) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh1)) \ deg R S X f) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((snd gh) \\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (snd gh1)) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) (f \\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> ((fst gh1) \\<^sub>r\<^bsub>R\<^esub> (snd gh1)))))" lemma cart_prod_fst:"x \ A \ B \ fst x \ A" by auto lemma cart_prod_snd:"x \ A \ B \ snd x \ B" by auto lemma cart_prod_split:"((x,y) \ A \ B) = (x \ A \ y \ B)" by auto lemma (in PolynRg) P_mod_diffxxx3:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; gh \ carrier R \ carrier R; deg R S X (fst gh) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)); deg R S X (snd gh) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd gh)); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd gh)); P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (f \ -\<^sub>a ((fst gh) \\<^sub>r (snd gh))); 0 < m\ \ \gh1. gh1 \carrier R \ carrier R \ (deg R S X (fst gh1) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh1))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((fst gh) \ -\<^sub>a (fst gh1)) \ (deg R S X (snd gh1) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh1)) \ deg R S X f) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((snd gh) \ -\<^sub>a (snd gh1)) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) (f \ (-\<^sub>a ((fst gh1) \\<^sub>r (snd gh1))))" apply (cases gh) apply (simp del: npow_suc) apply (rename_tac g h) apply (erule conjE, frule_tac g = g and h = h and f = f in P_mod_diffxxx2[of t R' Y], assumption+) apply blast done lemma (in PolynRg) P_mod_diffxxx4:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; gh \ carrier R \ carrier R; deg R S X (fst gh) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)); deg R S X (snd gh) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd gh)); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst gh)) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd gh)); P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) (f \ -\<^sub>a ((fst gh) \\<^sub>r (snd gh))); 0 < m\ \ (Hen\<^bsub>R S X t R' Y f \<^esub> m gh) \ carrier R \ carrier R \ (deg R S X (fst (Hen\<^bsub>R S X t R' Y f \<^esub> m gh)) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hen\<^bsub>R S X t R' Y f \<^esub> m gh)))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((fst gh) \ -\<^sub>a (fst (Hen\<^bsub>R S X t R' Y f \<^esub> m gh))) \ (deg R S X (snd (Hen\<^bsub>R S X t R' Y f \<^esub> m gh)) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hen\<^bsub>R S X t R' Y f \<^esub> m gh))) \ deg R S X f) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((snd gh) \ -\<^sub>a (snd (Hen\<^bsub>R S X t R' Y f \<^esub> m gh))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) (f \ (-\<^sub>a ((fst (Hen\<^bsub>R S X t R' Y f \<^esub> m gh)) \\<^sub>r (snd (Hen\<^bsub>R S X t R' Y f \<^esub> m gh)))))" apply (unfold Hensel_next_def) apply (rule someI2_ex) apply (rule P_mod_diffxxx3, assumption+) done (* Hensel_pair R S X t R' Y f g h m *) primrec Hensel_pair :: "[('a, 'b) Ring_scheme, ('a, 'c) Ring_scheme, 'a, 'a, ('a set, 'm) Ring_scheme, 'a set, 'a, 'a, 'a, nat] \ ('a \ 'a)" ("(10Hpr\<^bsub> _ _ _ _ _ _ _ _ _\<^esub> _)" [67,67,67,67,67,67,67,67,67,68]67) where Hpr_0: "Hpr\<^bsub>R S X t R' Y f g h\<^esub> 0 = (g, h)" | Hpr_Suc: "Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m) = Hen\<^bsub>R S X t R' Y f \<^esub> (Suc m) (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)" lemma (in PolynRg) fst_xxx:" \t \ carrier S; t \ \\<^bsub>S\<^esub>; ideal S (S \\<^sub>p t); \(n::nat). (F n) \ carrier R \ carrier R; \m. P_mod R S X (S \\<^sub>p t) (fst (F m) \ -\<^sub>a (fst (F (Suc m))))\ \ P_mod R S X (S \\<^sub>p t) (fst (F 0) \ -\<^sub>a (fst (F n)))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag) apply (induct_tac n) apply (drule_tac x = 0 in spec) apply (frule cart_prod_fst[of "F 0" "carrier R" "carrier R"]) apply (simp add:aGroup.ag_r_inv1) apply (simp add:P_mod_def) apply (frule_tac x = 0 in spec, frule_tac x = n in spec, drule_tac x = "Suc n" in spec) apply (frule_tac x = "F 0" in cart_prod_fst[of _ "carrier R" "carrier R"], frule_tac x = "F n" in cart_prod_fst[of _ "carrier R" "carrier R"], frule_tac x = "F (Suc n)" in cart_prod_fst[of _ "carrier R" "carrier R"]) apply (drule_tac x = n in spec) apply (frule_tac p = "fst (F 0) \ -\<^sub>a (fst (F n))" and q = "fst (F n) \ -\<^sub>a (fst (F (Suc n)))" in P_mod_add[of "S \\<^sub>p t"]) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+)+ apply (frule_tac x = "fst (F n)" in aGroup.ag_mOp_closed, assumption+, frule_tac x = "fst (F (Suc n))" in aGroup.ag_mOp_closed, assumption+) apply (simp add:aGroup.pOp_assocTr41[of "R", THEN sym], simp add:aGroup.pOp_assocTr42[of "R"], simp add:aGroup.ag_l_inv1, simp add:aGroup.ag_r_zero) done lemma (in PolynRg) snd_xxx:"\t \ carrier S; t \ \\<^bsub>S\<^esub>; ideal S (S \\<^sub>p t); \(n::nat). (F n) \ carrier R \ carrier R; \m. P_mod R S X (S \\<^sub>p t) (snd (F m) \ -\<^sub>a (snd (F (Suc m))))\ \ P_mod R S X (S \\<^sub>p t) (snd (F 0) \ -\<^sub>a (snd (F n)))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag) apply (induct_tac n) apply (drule_tac x = 0 in spec) apply (frule cart_prod_snd[of "F 0" "carrier R" "carrier R"]) apply (simp add:aGroup.ag_r_inv1) apply (simp add:P_mod_def) apply (frule_tac x = 0 in spec, frule_tac x = n in spec, drule_tac x = "Suc n" in spec) apply (frule_tac x = "F 0" in cart_prod_snd[of _ "carrier R" "carrier R"], frule_tac x = "F n" in cart_prod_snd[of _ "carrier R" "carrier R"], frule_tac x = "F (Suc n)" in cart_prod_snd[of _ "carrier R" "carrier R"]) apply (drule_tac x = n in spec) apply (frule_tac p = "snd (F 0) \ -\<^sub>a (snd (F n))" and q = "snd (F n) \ -\<^sub>a (snd (F (Suc n)))" in P_mod_add[of "S \\<^sub>p t"]) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+)+ apply (frule_tac x = "snd (F n)" in aGroup.ag_mOp_closed, assumption+, frule_tac x = "snd (F (Suc n))" in aGroup.ag_mOp_closed, assumption+) apply (simp add:aGroup.pOp_assocTr41[of "R", THEN sym], simp add:aGroup.pOp_assocTr42[of "R"], simp add:aGroup.ag_l_inv1, simp add:aGroup.ag_r_zero) done lemma (in PolynRg) P_mod_diffxxx5:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; (g, h) \ carrier R \ carrier R; deg R S X (fst (g, h)) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h))); deg R S X (snd (g, h)) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h))) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h))); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (g, h))); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h))) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (g, h))); P_mod R S X (S \\<^sub>p t) (f \ -\<^sub>a (g \\<^sub>r h))\ \ (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)) \ carrier R \ carrier R \ erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h)) \ erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (g, h)) \ (deg R S X (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) ((fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ -\<^sub>a (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))) \ (deg R S X (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))) \ deg R S X f) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) ((snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ -\<^sub>a (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc (Suc m))\<^esup>)) (f \ -\<^sub>a ((fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) \\<^sub>r (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))))" apply (cut_tac subring, frule subring_Ring, cut_tac ring_is_ag, frule mem_subring_mem_ring[of S t], assumption+, frule Ring.maximal_ideal_ideal[of S "S \\<^sub>p t"], assumption+) apply (induct_tac m) apply (simp del:Hpr_0 npow_suc) apply (simp only:Hpr_0) apply (frule P_mod_diffxxx4[of t R' Y f "(g, h)" "Suc 0"], assumption+) apply (simp add:cart_prod_split, simp+) apply (simp add:Ring.ring_l_one, simp) apply (simp add:Ring.ring_l_one, (erule conjE)+) apply (frule P_mod_diff[THEN sym, of "S \\<^sub>p t" R' Y g "fst (Hen\<^bsub> R S X t R' Y f\<^esub> (Suc 0) (g, h))"], assumption+, simp add:cart_prod_fst, rotate_tac -1, drule sym, simp) apply (frule P_mod_diff[THEN sym, of "S \\<^sub>p t" R' Y h "snd (Hen\<^bsub> R S X t R' Y f\<^esub> (Suc 0) (g, h))"], assumption+, simp add:cart_prod_snd, rotate_tac -1, drule sym, simp) apply ((erule conjE)+, rename_tac m) apply (frule_tac m = "Suc (Suc m)" and gh = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc m)" in P_mod_diffxxx4[of t R' Y f], assumption+) apply (simp, simp, simp, simp del:npow_suc, simp) apply (erule conjE)+ apply (simp del:npow_suc del:Hpr_Suc add:Hpr_Suc[THEN sym, of R S X t R' Y f _ g h]) apply (thin_tac "deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f", thin_tac "0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)", thin_tac "rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)", thin_tac "deg R S X (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc m)) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "deg R S X (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc m)) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f", thin_tac "deg R S X (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (Suc m))) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (Suc m))))", thin_tac "deg R S X (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (Suc m))) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (Suc m)))) \ deg R S X f") apply (frule_tac g = "fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc m)) \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (Suc m))))" and n = "Suc m" in P_mod_n_1 [of t], assumption+, (rule aGroup.ag_pOp_closed, assumption+, simp add:cart_prod_fst, rule aGroup.ag_mOp_closed, assumption+, simp add:cart_prod_fst, assumption), (frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc m" in cart_prod_fst[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (Suc m))" in cart_prod_fst[of _ "carrier R" "carrier R"]), (frule_tac g1 = "fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc m))" and h1 = "fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (Suc m)))" in P_mod_diff[THEN sym, of "S \\<^sub>p t" R' Y], assumption+)) apply (frule_tac g = "snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc m)) \ -\<^sub>a (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (Suc m))))" and n = "Suc m" in P_mod_n_1 [of t], assumption+, (rule aGroup.ag_pOp_closed, assumption+, simp add:cart_prod_snd, rule aGroup.ag_mOp_closed, assumption+, simp add:cart_prod_snd, assumption), (frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc m" in cart_prod_snd[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (Suc m))" in cart_prod_snd[of _ "carrier R" "carrier R"]), (frule_tac g1 = "snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc m))" and h1 = "snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (Suc m)))" in P_mod_diff[THEN sym, of "S \\<^sub>p t" R' Y], assumption+)) apply simp done (*** Hensel_pair basic ***) lemma (in PolynRg) P_mod_diffxxx5_1:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; g \ carrier R; h \ carrier R; deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); P_mod R S X (S \\<^sub>p t) (f \ -\<^sub>a (g \\<^sub>r h))\ \ (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)) \ carrier R \ carrier R \ erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h)) \ erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (g, h)) \ (deg R S X (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) ((fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ -\<^sub>a (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))) \ (deg R S X (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))) \ deg R S X f) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc m)\<^esup>)) ((snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ -\<^sub>a (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc (Suc m))\<^esup>)) (f \ -\<^sub>a ((fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m))) \\<^sub>r (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (Suc m)))))" apply (frule P_mod_diffxxx5[of t R' Y f g h m], assumption+) apply (simp add:cart_prod_split, simp, simp, simp, simp, simp, assumption+) done (*** Hpr sequence of polynomial pair ***) lemma (in PolynRg) P_mod_diffxxx5_2:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; g \ carrier R; h \ carrier R; deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); P_mod R S X (S \\<^sub>p t) (f \ -\<^sub>a (g \\<^sub>r h))\ \ (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m) \ carrier R \ carrier R" apply (case_tac "m = 0", simp, simp) apply (frule P_mod_diffxxx5_1[of t R' Y f g h "m - Suc 0"], assumption+) apply (erule conjE)+ apply simp done (*** Cauchy 1***) lemma (in PolynRg) P_mod_diffxxx5_3:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; g \ carrier R; h \ carrier R; deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); P_mod R S X (S \\<^sub>p t) (f \ -\<^sub>a (g \\<^sub>r h))\ \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ -\<^sub>a (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (m + n)))) \ P_mod R S X (S \\<^sub>p (t^\<^bsup>S m\<^esup>)) ((snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ -\<^sub>a (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> (m + n))))" apply (cut_tac ring_is_ag, cut_tac subring, frule subring_Ring) apply (induct_tac n) apply (simp del:npow_suc Hpr_Suc) apply (frule P_mod_diffxxx5_2[of t R' Y f g h m], assumption+) apply (frule cart_prod_fst[of "Hpr\<^bsub> R S X t R' Y f g h\<^esub> m" "carrier R" "carrier R"], frule cart_prod_snd[of "Hpr\<^bsub> R S X t R' Y f g h\<^esub> m" "carrier R" "carrier R"]) apply (simp add:aGroup.ag_r_inv1, simp add:P_mod_def) apply (frule_tac m = "m + n" in P_mod_diffxxx5_1[of t R' Y f g h], assumption+, (erule conjE)+) apply (frule_tac m = m in P_mod_diffxxx5_2[of t R' Y f g h], assumption+) apply (frule_tac m = "m + n" in P_mod_diffxxx5_2[of t R' Y f g h], assumption+) apply (thin_tac "deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f", thin_tac "0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)", thin_tac "rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)", thin_tac "P_mod R S X (S \\<^sub>p t) ( f \ -\<^sub>a (g \\<^sub>r h))", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n))) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (g, h))", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n))) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (g, h))", thin_tac "deg R S X (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n))) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n))))", thin_tac "deg R S X (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n))) + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n)))) \ deg R S X f", thin_tac "P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc (Suc (m + n)))\<^esup>)) (f \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n)) \\<^sub>r snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n))))") apply (simp del:npow_suc Hpr_Suc) apply (frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> m" in cart_prod_fst[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)" in cart_prod_fst[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (m + n))" in cart_prod_fst[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> m" in cart_prod_snd[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)" in cart_prod_snd[of _ "carrier R" "carrier R"], frule_tac x = "Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc (m + n))" in cart_prod_snd[of _ "carrier R" "carrier R"]) apply (case_tac "m = 0", simp del:npow_suc Hpr_Suc) apply (simp only:Ring.Rxa_one) apply (rule conjI) apply (rule_tac p = "g \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc n)))" in P_mod_whole, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (rule_tac p = "h \ -\<^sub>a (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (Suc n)))" in P_mod_whole, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (frule_tac g = "fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)) \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n)))" and n = "m + n" in P_mod_n_m[of t _ "m - Suc 0"], assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, arith, simp del:npow_suc Hpr_Suc, simp del:npow_suc Hpr_Suc) apply (frule Ring.npClose[of S t m], assumption, frule Ring.principal_ideal[of S "t^\<^bsup>S m\<^esup>"], assumption) apply (frule_tac p = "fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m) \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)))" and q = "fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)) \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n)))" in P_mod_add[of "S \\<^sub>p (t^\<^bsup>S m\<^esup>)"], (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+)+, simp del:npow_suc Hpr_Suc add:aGroup.pOp_assoc_cancel) apply (frule_tac g = "snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)) \ -\<^sub>a (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n)))" and n = "m + n" in P_mod_n_m[of t _ "m - Suc 0"], assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, arith, simp del:npow_suc Hpr_Suc, simp del:npow_suc Hpr_Suc) apply (frule_tac p = "snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m) \ -\<^sub>a (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)))" and q = "snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m + n)) \ -\<^sub>a (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m + n)))" in P_mod_add[of "S \\<^sub>p (t^\<^bsup>S m\<^esup>)"], (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+)+, simp del:npow_suc Hpr_Suc add:aGroup.pOp_assoc_cancel) done (*** Cauchy, deg bounded ****) lemma (in PolynRg) P_mod_diffxxx5_4:"\Idomain S; t \ carrier S; t \ \\<^bsub>S\<^esub>; maximal_ideal S (S \\<^sub>p t); PolynRg R' (S /\<^sub>r (S \\<^sub>p t)) Y; f \ carrier R; g \ carrier R; h \ carrier R; deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f; 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g); 0 < deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h); P_mod R S X (S \\<^sub>p t) (f \ -\<^sub>a (g \\<^sub>r h))\ \ deg R S X (fst (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ deg R S X g \ deg R S X (snd (Hpr\<^bsub>R S X t R' Y f g h\<^esub> m)) \ deg R S X f" apply (cut_tac subring, frule subring_Ring, frule Ring.maximal_ideal_ideal[of S "S \\<^sub>p t"], assumption) apply (case_tac "m = 0") apply simp apply (frule aless_imp_le[of "0" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)"]) apply (frule aadd_le_mono[of "0" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X h"]) apply (simp add:aadd_0_l, simp add:aadd_commute[of _ "deg R S X h"]) apply (frule P_mod_diffxxx5_1[of t R' Y f g h "m - Suc 0"], assumption+, (erule conjE)+) apply (thin_tac "deg R S X g \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)") apply (thin_tac "deg R S X h + deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X f") apply (thin_tac "rel_prime_pols R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h)", thin_tac "P_mod R S X (S \\<^sub>p t) ( f \ -\<^sub>a (g \\<^sub>r h))", thin_tac "P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc (m - Suc 0))\<^esup>)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m - Suc 0)) \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m - Suc 0))))", thin_tac "P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc (m - Suc 0))\<^esup>)) ( snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> (m - Suc 0)) \ -\<^sub>a (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m - Suc 0))))", thin_tac "P_mod R S X (S \\<^sub>p (t^\<^bsup>S (Suc (Suc (m - Suc 0)))\<^esup>)) (f \ -\<^sub>a (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m - Suc 0)) \\<^sub>r snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> Suc (m - Suc 0))))") apply (simp del:npow_suc Hpr_Suc) apply (thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m)) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g", thin_tac "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m)) = erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) h") apply (frule_tac p = g in pHom_dec_deg[of R' "(S /\<^sub>r (S \\<^sub>p t))" Y "erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t))"]) apply (frule Ring.qring_ring[of "S" "S \\<^sub>p t"], assumption+) apply (rule erH_rHom[of R' "(S /\<^sub>r (S \\<^sub>p t))" Y "pj S (S \\<^sub>p t)"], assumption+, simp add:pj_Hom, assumption+) apply (frule ale_trans[of "deg R S X (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m))" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X g"], assumption+) apply simp apply (thin_tac "deg R S X (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m)) \ deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)", thin_tac "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g) \ deg R S X g", thin_tac "deg R S X (fst (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m)) \ deg R S X g") apply (frule aless_imp_le[of "0" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)"]) apply (frule aadd_le_mono[of "0" "deg R' (S /\<^sub>r (S \\<^sub>p t)) Y (erH R S X R' (S /\<^sub>r (S \\<^sub>p t)) Y (pj S (S \\<^sub>p t)) g)" "deg R S X (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m))"]) apply (simp add:aadd_0_l, simp add:aadd_commute[of _ "deg R S X (snd (Hpr\<^bsub> R S X t R' Y f g h\<^esub> m))"]) done -declare max.absorb1 [simp del] max.absorb2 [simp del] - end diff --git a/thys/HereditarilyFinite/HF.thy b/thys/HereditarilyFinite/HF.thy --- a/thys/HereditarilyFinite/HF.thy +++ b/thys/HereditarilyFinite/HF.thy @@ -1,1107 +1,1107 @@ chapter \The Hereditarily Finite Sets\ theory HF imports "HOL-Library.Nat_Bijection" abbrevs "<:" = "\<^bold>\" and "~<:" = "\<^bold>\" begin text \From "Finite sets and Gödel's Incompleteness Theorems" by S. Swierczkowski. Thanks for Brian Huffman for this development, up to the cases and induct rules.\ section \Basic Definitions and Lemmas\ typedef hf = "UNIV :: nat set" .. definition hfset :: "hf \ hf set" where "hfset a = Abs_hf ` set_decode (Rep_hf a)" definition HF :: "hf set \ hf" where "HF A = Abs_hf (set_encode (Rep_hf ` A))" definition hinsert :: "hf \ hf \ hf" where "hinsert a b = HF (insert a (hfset b))" definition hmem :: "hf \ hf \ bool" (infixl "\<^bold>\" 50) where "hmem a b \ a \ hfset b" abbreviation not_hmem :: "hf \ hf \ bool" (infixl "\<^bold>\" 50) where "a \<^bold>\ b \ \ a \<^bold>\ b" notation (ASCII) hmem (infixl "<:" 50) instantiation hf :: zero begin definition Zero_hf_def: "0 = HF {}" instance .. end lemma Abs_hf_0 [simp]: "Abs_hf 0 = 0" by (simp add: HF_def Zero_hf_def) text \HF Set enumerations\ abbreviation inserthf :: "hf \ hf \ hf" (infixl "\" 60) where "y \ x \ hinsert x y" syntax (ASCII) "_HFinset" :: "args \ hf" ("{|(_)|}") syntax "_HFinset" :: "args \ hf" ("\_\") translations "\x, y\" \ "\y\ \ x" "\x\" \ "0 \ x" lemma finite_hfset [simp]: "finite (hfset a)" unfolding hfset_def by simp lemma HF_hfset [simp]: "HF (hfset a) = a" unfolding HF_def hfset_def by (simp add: image_image Abs_hf_inverse Rep_hf_inverse) lemma hfset_HF [simp]: "finite A \ hfset (HF A) = A" unfolding HF_def hfset_def by (simp add: image_image Abs_hf_inverse Rep_hf_inverse) lemma inj_on_HF: "inj_on HF (Collect finite)" by (metis hfset_HF inj_onI mem_Collect_eq) lemma hmem_hempty [simp]: "a \<^bold>\ 0" unfolding hmem_def Zero_hf_def by simp lemmas hemptyE [elim!] = hmem_hempty [THEN notE] lemma hmem_hinsert [iff]: "hmem a (c \ b) \ a = b \ a \<^bold>\ c" unfolding hmem_def hinsert_def by simp lemma hf_ext: "a = b \ (\x. x \<^bold>\ a \ x \<^bold>\ b)" unfolding hmem_def set_eq_iff [symmetric] by (metis HF_hfset) lemma finite_cases [consumes 1, case_names empty insert]: "\finite F; F = {} \ P; \A x. \F = insert x A; x \ A; finite A\ \ P\ \ P" by (induct F rule: finite_induct, simp_all) lemma hf_cases [cases type: hf, case_names 0 hinsert]: obtains "y = 0" | a b where "y = b \ a" and "a \<^bold>\ b" proof - have "finite (hfset y)" by (rule finite_hfset) thus thesis by (metis Zero_hf_def finite_cases hf_ext hfset_HF hinsert_def hmem_def that) qed lemma Rep_hf_hinsert: "a \<^bold>\ b \ Rep_hf (hinsert a b) = 2 ^ (Rep_hf a) + Rep_hf b" unfolding hinsert_def HF_def hfset_def apply (simp add: image_image Abs_hf_inverse Rep_hf_inverse) apply (subst set_encode_insert, simp) apply (clarsimp simp add: hmem_def hfset_def image_def Rep_hf_inject [symmetric] Abs_hf_inverse, simp) done lemma less_two_power: "n < 2 ^ n" - by (induct n, auto) + by (fact less_exp) section \Verifying the Axioms of HF\ text \HF1\ lemma hempty_iff: "z=0 \ (\x. x \<^bold>\ z)" by (simp add: hf_ext) text \HF2\ lemma hinsert_iff: "z = x \ y \ (\u. u \<^bold>\ z \ u \<^bold>\ x \ u = y)" by (auto simp: hf_ext) text \HF induction\ lemma hf_induct [induct type: hf, case_names 0 hinsert]: assumes [simp]: "P 0" "\x y. \P x; P y; x \<^bold>\ y\ \ P (y \ x)" shows "P z" proof (induct z rule: wf_induct [where r="measure Rep_hf", OF wf_measure]) case (1 x) show ?case proof (cases x rule: hf_cases) case 0 thus ?thesis by simp next case (hinsert a b) thus ?thesis using 1 by (simp add: Rep_hf_hinsert less_le_trans [OF less_two_power le_add1]) qed qed text \HF3\ lemma hf_induct_ax: "\P 0; \x. P x \ (\y. P y \ P (x \ y))\ \ P x" by (induct x, auto) lemma hf_equalityI [intro]: "(\x. x \<^bold>\ a \ x \<^bold>\ b) \ a = b" by (simp add: hf_ext) lemma hinsert_nonempty [simp]: "A \ a \ 0" by (auto simp: hf_ext) lemma hinsert_commute: "(z \ y) \ x = (z \ x) \ y" by (auto simp: hf_ext) lemma hmem_HF_iff [simp]: "x \<^bold>\ HF A \ x \ A \ finite A" apply (cases "finite A", auto) apply (simp add: hmem_def) apply (simp add: hmem_def) apply (metis HF_def Rep_hf_inject Abs_hf_0 finite_imageD hempty_iff inj_onI set_encode_inf) done section \Ordered Pairs, from ZF/ZF.thy\ lemma singleton_eq_iff [iff]: "\a\ = \b\ \ a=b" by (metis hmem_hempty hmem_hinsert) lemma doubleton_eq_iff: "\a,b\ = \c,d\ \ (a=c \ b=d) \ (a=d \ b=c)" by auto (metis hmem_hempty hmem_hinsert)+ definition hpair :: "hf \ hf \ hf" where "hpair a b = \\a\,\a,b\\" definition hfst :: "hf \ hf" where "hfst p \ THE x. \y. p = hpair x y" definition hsnd :: "hf \ hf" where "hsnd p \ THE y. \x. p = hpair x y" definition hsplit :: "[[hf, hf] \ 'a, hf] \ 'a::{}" \ \for pattern-matching\ where "hsplit c \ \p. c (hfst p) (hsnd p)" text \Ordered Pairs, from ZF/ZF.thy\ nonterminal hfs syntax (ASCII) "_Tuple" :: "[hf, hfs] \ hf" ("<(_,/ _)>") "_hpattern" :: "[pttrn, patterns] \ pttrn" ("<_,/ _>") syntax "" :: "hf \ hfs" ("_") "_Enum" :: "[hf, hfs] \ hfs" ("_,/ _") "_Tuple" :: "[hf, hfs] \ hf" ("\(_,/ _)\") "_hpattern" :: "[pttrn, patterns] \ pttrn" ("\_,/ _\") translations "" \ ">" "" \ "CONST hpair x y" "" \ ">" "\. b" \ "CONST hsplit(\x . b)" "\. b" \ "CONST hsplit(\x y. b)" lemma hpair_def': "hpair a b = \\a,a\,\a,b\\" by (auto simp: hf_ext hpair_def) lemma hpair_iff [simp]: "hpair a b = hpair a' b' \ a=a' \ b=b'" by (auto simp: hpair_def' doubleton_eq_iff) lemmas hpair_inject = hpair_iff [THEN iffD1, THEN conjE, elim!] lemma hfst_conv [simp]: "hfst \a,b\ = a" by (simp add: hfst_def) lemma hsnd_conv [simp]: "hsnd \a,b\ = b" by (simp add: hsnd_def) lemma hsplit [simp]: "hsplit c \a,b\ = c a b" by (simp add: hsplit_def) section \Unions, Comprehensions, Intersections\ subsection \Unions\ text \Theorem 1.5 (Existence of the union of two sets).\ lemma binary_union: "\z. \u. u \<^bold>\ z \ u \<^bold>\ x \ u \<^bold>\ y" proof (induct x rule: hf_induct) case 0 thus ?case by auto next case (hinsert a b) thus ?case by (metis hmem_hinsert) qed text \Theorem 1.6 (Existence of the union of a set of sets).\ lemma union_of_set: "\z. \u. u \<^bold>\ z \ (\y. y \<^bold>\ x \ u \<^bold>\ y)" proof (induct x rule: hf_induct) case 0 thus ?case by (metis hmem_hempty) next case (hinsert a b) then show ?case by (metis hmem_hinsert binary_union [of a]) qed subsection \Set comprehensions\ text \Theorem 1.7, comprehension scheme\ lemma comprehension: "\z. \u. u \<^bold>\ z \ u \<^bold>\ x \ P u" proof (induct x rule: hf_induct) case 0 thus ?case by (metis hmem_hempty) next case (hinsert a b) thus ?case by (metis hmem_hinsert) qed definition HCollect :: "(hf \ bool) \ hf \ hf" \ \comprehension\ where "HCollect P A = (THE z. \u. u \<^bold>\ z = (P u \ u \<^bold>\ A))" syntax (ASCII) "_HCollect" :: "idt \ hf \ bool \ hf" ("(1\_ <:/ _./ _\)") syntax "_HCollect" :: "idt \ hf \ bool \ hf" ("(1\_ \<^bold>\/ _./ _\)") translations "\x \<^bold>\ A. P\" \ "CONST HCollect (\x. P) A" lemma HCollect_iff [iff]: "hmem x (HCollect P A) \ P x \ x \<^bold>\ A" apply (insert comprehension [of A P], clarify) apply (simp add: HCollect_def) apply (rule theI2, blast) apply (auto simp: hf_ext) done lemma HCollectI: "a \<^bold>\ A \ P a \ hmem a \x \<^bold>\ A. P x\" by simp lemma HCollectE: assumes "a \<^bold>\ \x \<^bold>\ A. P x\" obtains "a \<^bold>\ A" "P a" using assms by auto lemma HCollect_hempty [simp]: "HCollect P 0 = 0" by (simp add: hf_ext) subsection \Union operators\ instantiation hf :: sup begin definition "sup a b = (THE z. \u. u \<^bold>\ z \ u \<^bold>\ a \ u \<^bold>\ b)" instance .. end abbreviation hunion :: "hf \ hf \ hf" (infixl "\" 65) where "hunion \ sup" lemma hunion_iff [iff]: "hmem x (a \ b) \ x \<^bold>\ a \ x \<^bold>\ b" apply (insert binary_union [of a b], clarify) apply (simp add: sup_hf_def) apply (rule theI2) apply (auto simp: hf_ext) done definition HUnion :: "hf \ hf" ("\_" [900] 900) where "HUnion A = (THE z. \u. u \<^bold>\ z \ (\y. y \<^bold>\ A \ u \<^bold>\ y))" lemma HUnion_iff [iff]: "hmem x (\ A) \ (\y. y \<^bold>\ A \ x \<^bold>\ y)" apply (insert union_of_set [of A], clarify) apply (simp add: HUnion_def) apply (rule theI2) apply (auto simp: hf_ext) done lemma HUnion_hempty [simp]: "\ 0 = 0" by (simp add: hf_ext) lemma HUnion_hinsert [simp]: "\(A \ a) = a \ \A" by (auto simp: hf_ext) lemma HUnion_hunion [simp]: "\(A \ B) = \A \ \B" by blast subsection \Definition 1.8, Intersections\ instantiation hf :: inf begin definition "inf a b = \x \<^bold>\ a. x \<^bold>\ b\" instance .. end abbreviation hinter :: "hf \ hf \ hf" (infixl "\" 70) where "hinter \ inf" lemma hinter_iff [iff]: "hmem u (x \ y) \ u \<^bold>\ x \ u \<^bold>\ y" by (metis HCollect_iff inf_hf_def) definition HInter :: "hf \ hf" ("\_" [900] 900) where "HInter(A) = \x \<^bold>\ HUnion(A). \y. y \<^bold>\ A \ x \<^bold>\ y\" lemma HInter_hempty [iff]: "\ 0 = 0" by (metis HCollect_hempty HUnion_hempty HInter_def) lemma HInter_iff [simp]: "A\0 \ hmem x (\ A) \ (\y. y \<^bold>\ A \ x \<^bold>\ y)" by (auto simp: HInter_def) lemma HInter_hinsert [simp]: "A\0 \ \(A \ a) = a \ \A" by (auto simp: hf_ext HInter_iff [OF hinsert_nonempty]) subsection \Set Difference\ instantiation hf :: minus begin definition "A - B = \x \<^bold>\ A. x \<^bold>\ B\" instance .. end lemma hdiff_iff [iff]: "hmem u (x - y) \ u \<^bold>\ x \ u \<^bold>\ y" by (auto simp: minus_hf_def) lemma hdiff_zero [simp]: fixes x :: hf shows "(x - 0) = x" by blast lemma zero_hdiff [simp]: fixes x :: hf shows "(0 - x) = 0" by blast lemma hdiff_insert: "A - (B \ a) = A - B - \a\" by blast lemma hinsert_hdiff_if: "(A \ x) - B = (if x \<^bold>\ B then A - B else (A - B) \ x)" by auto section \Replacement\ text \Theorem 1.9 (Replacement Scheme).\ lemma replacement: "(\u v v'. u \<^bold>\ x \ R u v \ R u v' \ v'=v) \ \z. \v. v \<^bold>\ z \ (\u. u \<^bold>\ x \ R u v)" proof (induct x rule: hf_induct) case 0 thus ?case by (metis hmem_hempty) next case (hinsert a b) thus ?case by simp (metis hmem_hinsert) qed lemma replacement_fun: "\z. \v. v \<^bold>\ z \ (\u. u \<^bold>\ x \ v = f u)" by (rule replacement [where R = "\u v. v = f u"]) auto definition PrimReplace :: "hf \ (hf \ hf \ bool) \ hf" where "PrimReplace A R = (THE z. \v. v \<^bold>\ z \ (\u. u \<^bold>\ A \ R u v))" definition Replace :: "hf \ (hf \ hf \ bool) \ hf" where "Replace A R = PrimReplace A (\x y. (\!z. R x z) \ R x y)" definition RepFun :: "hf \ (hf \ hf) \ hf" where "RepFun A f = Replace A (\x y. y = f x)" syntax (ASCII) "_HReplace" :: "[pttrn, pttrn, hf, bool] \ hf" ("(1{|_ ./ _<: _, _|})") "_HRepFun" :: "[hf, pttrn, hf] \ hf" ("(1{|_ ./ _<: _|})" [51,0,51]) "_HINTER" :: "[pttrn, hf, hf] \ hf" ("(3INT _<:_./ _)" 10) "_HUNION" :: "[pttrn, hf, hf] \ hf" ("(3UN _<:_./ _)" 10) syntax "_HReplace" :: "[pttrn, pttrn, hf, bool] \ hf" ("(1\_ ./ _ \<^bold>\ _, _\)") "_HRepFun" :: "[hf, pttrn, hf] \ hf" ("(1\_ ./ _ \<^bold>\ _\)" [51,0,51]) "_HINTER" :: "[pttrn, hf, hf] \ hf" ("(3\_\<^bold>\_./ _)" 10) "_HUNION" :: "[pttrn, hf, hf] \ hf" ("(3\_\<^bold>\_./ _)" 10) translations "\y. x\<^bold>\A, Q\" \ "CONST Replace A (\x y. Q)" "\b. x\<^bold>\A\" \ "CONST RepFun A (\x. b)" "\x\<^bold>\A. B" \ "CONST HInter(CONST RepFun A (\x. B))" "\x\<^bold>\A. B" \ "CONST HUnion(CONST RepFun A (\x. B))" lemma PrimReplace_iff: assumes sv: "\u v v'. u \<^bold>\ A \ R u v \ R u v' \ v'=v" shows "v \<^bold>\ (PrimReplace A R) \ (\u. u \<^bold>\ A \ R u v)" apply (insert replacement [OF sv], clarify) apply (simp add: PrimReplace_def) apply (rule theI2) apply (auto simp: hf_ext) done lemma Replace_iff [iff]: "v \<^bold>\ Replace A R \ (\u. u \<^bold>\ A \ R u v \ (\y. R u y \ y=v))" apply (simp add: Replace_def) apply (subst PrimReplace_iff, auto) done lemma Replace_0 [simp]: "Replace 0 R = 0" by blast lemma Replace_hunion [simp]: "Replace (A \ B) R = Replace A R \ Replace B R" by blast lemma Replace_cong [cong]: "\ A=B; \x y. x \<^bold>\ B \ P x y \ Q x y \ \ Replace A P = Replace B Q" by (simp add: hf_ext cong: conj_cong) lemma RepFun_iff [iff]: "v \<^bold>\ (RepFun A f) \ (\u. u \<^bold>\ A \ v = f u)" by (auto simp: RepFun_def) lemma RepFun_cong [cong]: "\ A=B; \x. x \<^bold>\ B \ f(x)=g(x) \ \ RepFun A f = RepFun B g" by (simp add: RepFun_def) lemma triv_RepFun [simp]: "RepFun A (\x. x) = A" by blast lemma RepFun_0 [simp]: "RepFun 0 f = 0" by blast lemma RepFun_hinsert [simp]: "RepFun (hinsert a b) f = hinsert (f a) (RepFun b f)" by blast lemma RepFun_hunion [simp]: "RepFun (A \ B) f = RepFun A f \ RepFun B f" by blast lemma HF_HUnion: "\finite A; \x. x\A \ finite (B x)\ \ HF (\x \ A. B x) = (\x\<^bold>\HF A. HF (B x))" by (rule hf_equalityI) (auto) section \Subset relation and the Lattice Properties\ text \Definition 1.10 (Subset relation).\ instantiation hf :: order begin definition less_eq_hf where "A \ B \ (\x. x \<^bold>\ A \ x \<^bold>\ B)" definition less_hf where "A < B \ A \ B \ A \ (B::hf)" instance by standard (auto simp: less_eq_hf_def less_hf_def) end subsection \Rules for subsets\ lemma hsubsetI [intro!]: "(\x. x\<^bold>\A \ x\<^bold>\B) \ A \ B" by (simp add: less_eq_hf_def) text \Classical elimination rule\ lemma hsubsetCE [elim]: "\ A \ B; c\<^bold>\A \ P; c\<^bold>\B \ P \ \ P" by (auto simp: less_eq_hf_def) text \Rule in Modus Ponens style\ lemma hsubsetD [elim]: "\ A \ B; c\<^bold>\A \ \ c\<^bold>\B" by (simp add: less_eq_hf_def) text \Sometimes useful with premises in this order\ lemma rev_hsubsetD: "\ c\<^bold>\A; A\B \ \ c\<^bold>\B" by blast lemma contra_hsubsetD: "\ A \ B; c \ B \ \ c \ A" by blast lemma rev_contra_hsubsetD: "\ c \ B; A \ B \ \ c \ A" by blast lemma hf_equalityE: fixes A :: hf shows "A = B \ (A \ B \ B \ A \ P) \ P" by (metis order_refl) subsection \Lattice properties\ instantiation hf :: distrib_lattice begin instance by standard (auto simp: less_eq_hf_def less_hf_def inf_hf_def) end instantiation hf :: bounded_lattice_bot begin definition "bot = (0::hf)" instance by standard (auto simp: less_eq_hf_def bot_hf_def) end lemma hinter_hempty_left [simp]: "0 \ A = 0" by (metis bot_hf_def inf_bot_left) lemma hinter_hempty_right [simp]: "A \ 0 = 0" by (metis bot_hf_def inf_bot_right) lemma hunion_hempty_left [simp]: "0 \ A = A" by (metis bot_hf_def sup_bot_left) lemma hunion_hempty_right [simp]: "A \ 0 = A" by (metis bot_hf_def sup_bot_right) lemma less_eq_hempty [simp]: "u \ 0 \ u = (0::hf)" by (metis hempty_iff less_eq_hf_def) lemma less_eq_insert1_iff [iff]: "(hinsert x y) \ z \ x \<^bold>\ z \ y \ z" by (auto simp: less_eq_hf_def) lemma less_eq_insert2_iff: "z \ (hinsert x y) \ z \ y \ (\u. hinsert x u = z \ x \<^bold>\ u \ u \ y)" proof (cases "x \<^bold>\ z") case True hence u: "hinsert x (z - \x\) = z" by auto show ?thesis proof assume "z \ (hinsert x y)" thus "z \ y \ (\u. hinsert x u = z \ x \<^bold>\ u \ u \ y)" by (simp add: less_eq_hf_def) (metis u hdiff_iff hmem_hinsert) next assume "z \ y \ (\u. hinsert x u = z \ x \<^bold>\ u \ u \ y)" thus "z \ (hinsert x y)" by (auto simp: less_eq_hf_def) qed next case False thus ?thesis by (metis hmem_hinsert less_eq_hf_def) qed lemma zero_le [simp]: "0 \ (x::hf)" by blast lemma hinsert_eq_sup: "b \ a = b \ \a\" by blast lemma hunion_hinsert_left: "hinsert x A \ B = hinsert x (A \ B)" by blast lemma hunion_hinsert_right: "B \ hinsert x A = hinsert x (B \ A)" by blast lemma hinter_hinsert_left: "hinsert x A \ B = (if x \<^bold>\ B then hinsert x (A \ B) else A \ B)" by auto lemma hinter_hinsert_right: "B \ hinsert x A = (if x \<^bold>\ B then hinsert x (B \ A) else B \ A)" by auto section \Foundation, Cardinality, Powersets\ subsection \Foundation\ text \Theorem 1.13: Foundation (Regularity) Property.\ lemma foundation: assumes z: "z \ 0" shows "\w. w \<^bold>\ z \ w \ z = 0" proof - { fix x assume z: "(\w. w \<^bold>\ z \ w \ z \ 0)" have "x \<^bold>\ z \ x \ z = 0" proof (induction x rule: hf_induct) case 0 thus ?case by (metis hinter_hempty_left z) next case (hinsert x y) thus ?case by (metis hinter_hinsert_left z) qed } thus ?thesis using z by (metis z hempty_iff) qed lemma hmem_not_refl: "x \<^bold>\ x" using foundation [of "\x\"] by (metis hinter_iff hmem_hempty hmem_hinsert) lemma hmem_not_sym: "\ (x \<^bold>\ y \ y \<^bold>\ x)" using foundation [of "\x,y\"] by (metis hinter_iff hmem_hempty hmem_hinsert) lemma hmem_ne: "x \<^bold>\ y \ x \ y" by (metis hmem_not_refl) lemma hmem_Sup_ne: "x \<^bold>\ y \ \x \ y" by (metis HUnion_iff hmem_not_sym) lemma hpair_neq_fst: "\a,b\ \ a" by (metis hpair_def hinsert_iff hmem_not_sym) lemma hpair_neq_snd: "\a,b\ \ b" by (metis hpair_def hinsert_iff hmem_not_sym) lemma hpair_nonzero [simp]: "\x,y\ \ 0" by (auto simp: hpair_def) lemma zero_notin_hpair: "0 \<^bold>\ \x,y\" by (auto simp: hpair_def) subsection \Cardinality\ text \First we need to hack the underlying representation\ lemma hfset_0 [simp]: "hfset 0 = {}" by (metis Zero_hf_def finite.emptyI hfset_HF) lemma hfset_hinsert: "hfset (b \ a) = insert a (hfset b)" by (metis finite_insert hinsert_def HF.finite_hfset hfset_HF) lemma hfset_hdiff: "hfset (x - y) = hfset x - hfset y" proof (induct x arbitrary: y rule: hf_induct) case 0 thus ?case by simp next case (hinsert a b) thus ?case by (simp add: hfset_hinsert Set.insert_Diff_if hinsert_hdiff_if hmem_def) qed definition hcard :: "hf \ nat" where "hcard x = card (hfset x)" lemma hcard_0 [simp]: "hcard 0 = 0" by (simp add: hcard_def) lemma hcard_hinsert_if: "hcard (hinsert x y) = (if x \<^bold>\ y then hcard y else Suc (hcard y))" by (simp add: hcard_def hfset_hinsert card_insert_if hmem_def) lemma hcard_union_inter: "hcard (x \ y) + hcard (x \ y) = hcard x + hcard y" apply (induct x arbitrary: y rule: hf_induct) apply (auto simp: hcard_hinsert_if hunion_hinsert_left hinter_hinsert_left) done lemma hcard_hdiff1_less: "x \<^bold>\ z \ hcard (z - \x\) < hcard z" by (simp add: hcard_def hfset_hdiff hfset_hinsert) (metis card_Diff1_less finite_hfset hmem_def) subsection \Powerset Operator\ text \Theorem 1.11 (Existence of the power set).\ lemma powerset: "\z. \u. u \<^bold>\ z \ u \ x" proof (induction x rule: hf_induct) case 0 thus ?case by (metis hmem_hempty hmem_hinsert less_eq_hempty) next case (hinsert a b) then obtain Pb where Pb: "\u. u \<^bold>\ Pb \ u \ b" by auto obtain RPb where RPb: "\v. v \<^bold>\ RPb \ (\u. u \<^bold>\ Pb \ v = hinsert a u)" using replacement_fun .. thus ?case using Pb binary_union [of Pb RPb] apply (simp add: less_eq_insert2_iff, clarify) apply (rule_tac x=z in exI) apply (metis hinsert.hyps less_eq_hf_def) done qed definition HPow :: "hf \ hf" where "HPow x = (THE z. \u. u \<^bold>\ z \ u \ x)" lemma HPow_iff [iff]: "u \<^bold>\ HPow x \ u \ x" apply (insert powerset [of x], clarify) apply (simp add: HPow_def) apply (rule theI2) apply (auto simp: hf_ext) done lemma HPow_mono: "x \ y \ HPow x \ HPow y" by (metis HPow_iff less_eq_hf_def order_trans) lemma HPow_mono_strict: "x < y \ HPow x < HPow y" by (metis HPow_iff HPow_mono less_le_not_le order_eq_iff) lemma HPow_mono_iff [simp]: "HPow x \ HPow y \ x \ y" by (metis HPow_iff HPow_mono hsubsetCE order_refl) lemma HPow_mono_strict_iff [simp]: "HPow x < HPow y \ x < y" by (metis HPow_mono_iff less_le_not_le) section \Bounded Quantifiers\ definition HBall :: "hf \ (hf \ bool) \ bool" where "HBall A P \ (\x. x \<^bold>\ A \ P x)" \ \bounded universal quantifiers\ definition HBex :: "hf \ (hf \ bool) \ bool" where "HBex A P \ (\x. x \<^bold>\ A \ P x)" \ \bounded existential quantifiers\ syntax (ASCII) "_HBall" :: "pttrn \ hf \ bool \ bool" ("(3ALL _<:_./ _)" [0, 0, 10] 10) "_HBex" :: "pttrn \ hf \ bool \ bool" ("(3EX _<:_./ _)" [0, 0, 10] 10) "_HBex1" :: "pttrn \ hf \ bool \ bool" ("(3EX! _<:_./ _)" [0, 0, 10] 10) syntax "_HBall" :: "pttrn \ hf \ bool \ bool" ("(3\_\<^bold>\_./ _)" [0, 0, 10] 10) "_HBex" :: "pttrn \ hf \ bool \ bool" ("(3\_\<^bold>\_./ _)" [0, 0, 10] 10) "_HBex1" :: "pttrn \ hf \ bool \ bool" ("(3\!_\<^bold>\_./ _)" [0, 0, 10] 10) translations "\x\<^bold>\A. P" \ "CONST HBall A (\x. P)" "\x\<^bold>\A. P" \ "CONST HBex A (\x. P)" "\!x\<^bold>\A. P" \ "\!x. x\A \ P" lemma hball_cong [cong]: "\ A=A'; \x. x \<^bold>\ A' \ P(x) \ P'(x) \ \ (\x\<^bold>\A. P(x)) \ (\x\<^bold>\A'. P'(x))" by (simp add: HBall_def) lemma hballI [intro!]: "(\x. x\<^bold>\A \ P x) \ \x\<^bold>\A. P x" by (simp add: HBall_def) lemma hbspec [dest?]: "\x\<^bold>\A. P x \ x\<^bold>\A \ P x" by (simp add: HBall_def) lemma hballE [elim]: "\x\<^bold>\A. P x \ (P x \ Q) \ (x \<^bold>\ A \ Q) \ Q" by (unfold HBall_def) blast lemma hbex_cong [cong]: "\ A=A'; \x. x \<^bold>\ A' \ P(x) \ P'(x) \ \ (\x\<^bold>\A. P(x)) \ (\x\<^bold>\A'. P'(x))" by (simp add: HBex_def cong: conj_cong) lemma hbexI [intro]: "P x \ x\<^bold>\A \ \x\<^bold>\A. P x" by (unfold HBex_def) blast lemma rev_hbexI [intro?]: "x\<^bold>\A \ P x \ \x\<^bold>\A. P x" by (unfold HBex_def) blast lemma bexCI: "(\x\<^bold>\A. \ P x \ P a) \ a\<^bold>\A \ \x\<^bold>\A. P x" by (unfold HBex_def) blast lemma hbexE [elim!]: "\x\<^bold>\A. P x \ (\x. x\<^bold>\A \ P x \ Q) \ Q" by (unfold HBex_def) blast lemma hball_triv [simp]: "(\x\<^bold>\A. P) = ((\x. x\<^bold>\A) \ P)" \ \trivial rewrite rule.\ by (simp add: HBall_def) lemma hbex_triv [simp]: "(\x\<^bold>\A. P) = ((\x. x\<^bold>\A) \ P)" \ \Dual form for existentials.\ by (simp add: HBex_def) lemma hbex_triv_one_point1 [simp]: "(\x\<^bold>\A. x = a) = (a\<^bold>\A)" by blast lemma hbex_triv_one_point2 [simp]: "(\x\<^bold>\A. a = x) = (a\<^bold>\A)" by blast lemma hbex_one_point1 [simp]: "(\x\<^bold>\A. x = a \ P x) = (a\<^bold>\A \ P a)" by blast lemma hbex_one_point2 [simp]: "(\x\<^bold>\A. a = x \ P x) = (a\<^bold>\A \ P a)" by blast lemma hball_one_point1 [simp]: "(\x\<^bold>\A. x = a \ P x) = (a\<^bold>\A \ P a)" by blast lemma hball_one_point2 [simp]: "(\x\<^bold>\A. a = x \ P x) = (a\<^bold>\A \ P a)" by blast lemma hball_conj_distrib: "(\x\<^bold>\A. P x \ Q x) \ ((\x\<^bold>\A. P x) \ (\x\<^bold>\A. Q x))" by blast lemma hbex_disj_distrib: "(\x\<^bold>\A. P x \ Q x) \ ((\x\<^bold>\A. P x) \ (\x\<^bold>\A. Q x))" by blast lemma hb_all_simps [simp, no_atp]: "\A P Q. (\x \<^bold>\ A. P x \ Q) \ ((\x \<^bold>\ A. P x) \ Q)" "\A P Q. (\x \<^bold>\ A. P \ Q x) \ (P \ (\x \<^bold>\ A. Q x))" "\A P Q. (\x \<^bold>\ A. P \ Q x) \ (P \ (\x \<^bold>\ A. Q x))" "\A P Q. (\x \<^bold>\ A. P x \ Q) \ ((\x \<^bold>\ A. P x) \ Q)" "\P. (\x \<^bold>\ 0. P x) \ True" "\a B P. (\x \<^bold>\ B \ a. P x) \ (P a \ (\x \<^bold>\ B. P x))" "\P Q. (\x \<^bold>\ HCollect Q A. P x) \ (\x \<^bold>\ A. Q x \ P x)" "\A P. (\ (\x \<^bold>\ A. P x)) \ (\x \<^bold>\ A. \ P x)" by auto lemma hb_ex_simps [simp, no_atp]: "\A P Q. (\x \<^bold>\ A. P x \ Q) \ ((\x \<^bold>\ A. P x) \ Q)" "\A P Q. (\x \<^bold>\ A. P \ Q x) \ (P \ (\x \<^bold>\ A. Q x))" "\P. (\x \<^bold>\ 0. P x) \ False" "\a B P. (\x \<^bold>\ B \ a. P x) \ (P a \ (\x \<^bold>\ B. P x))" "\P Q. (\x \<^bold>\ HCollect Q A. P x) \ (\x \<^bold>\ A. Q x \ P x)" "\A P. (\(\x \<^bold>\ A. P x)) \ (\x \<^bold>\ A. \ P x)" by auto lemma le_HCollect_iff: "A \ \x \<^bold>\ B. P x\ \ A \ B \ (\x \<^bold>\ A. P x)" by blast section \Relations and Functions\ definition is_hpair :: "hf \ bool" where "is_hpair z = (\x y. z = \x,y\)" definition hconverse :: "hf \ hf" where "hconverse(r) = \z. w \<^bold>\ r, \x y. w = \x,y\ \ z = \y,x\\" definition hdomain :: "hf \ hf" where "hdomain(r) = \x. w \<^bold>\ r, \y. w = \x,y\\" definition hrange :: "hf \ hf" where "hrange(r) = hdomain(hconverse(r))" definition hrelation :: "hf \ bool" where "hrelation(r) = (\z. z \<^bold>\ r \ is_hpair z)" definition hrestrict :: "hf \ hf \ hf" \ \Restrict the relation r to the domain A\ where "hrestrict r A = \z \<^bold>\ r. \x \<^bold>\ A. \y. z = \x,y\\" definition nonrestrict :: "hf \ hf \ hf" where "nonrestrict r A = \z \<^bold>\ r. \x \<^bold>\ A. \y. z \ \x,y\\" definition hfunction :: "hf \ bool" where "hfunction(r) = (\x y. \x,y\ \<^bold>\ r \ (\y'. \x,y'\ \<^bold>\ r \ y=y'))" definition app :: "hf \ hf \ hf" where "app f x = (THE y. \x, y\ \<^bold>\ f)" lemma hrestrict_iff [iff]: "z \<^bold>\ hrestrict r A \ z \<^bold>\ r \ (\ x y. z = \x, y\ \ x \<^bold>\ A)" by (auto simp: hrestrict_def) lemma hrelation_0 [simp]: "hrelation 0" by (force simp add: hrelation_def) lemma hrelation_restr [iff]: "hrelation (hrestrict r x)" by (metis hrelation_def hrestrict_iff is_hpair_def) lemma hrelation_hunion [simp]: "hrelation (f \ g) \ hrelation f \ hrelation g" by (auto simp: hrelation_def) lemma hfunction_restr: "hfunction r \ hfunction (hrestrict r x)" by (auto simp: hfunction_def hrestrict_def) lemma hdomain_restr [simp]: "hdomain (hrestrict r x) = hdomain r \ x" by (force simp add: hdomain_def hrestrict_def) lemma hdomain_0 [simp]: "hdomain 0 = 0" by (force simp add: hdomain_def) lemma hdomain_ins [simp]: "hdomain (r \ \x, y\) = hdomain r \ x" by (force simp add: hdomain_def) lemma hdomain_hunion [simp]: "hdomain (f \ g) = hdomain f \ hdomain g" by (simp add: hdomain_def) lemma hdomain_not_mem [iff]: "\hdomain r, a\ \<^bold>\ r" by (metis hdomain_ins hinter_hinsert_right hmem_hinsert hmem_not_refl hunion_hinsert_right sup_inf_absorb) lemma app_singleton [simp]: "app \\x, y\\ x = y" by (simp add: app_def) lemma app_equality: "hfunction f \ \x, y\ \<^bold>\ f \ app f x = y" by (auto simp: app_def hfunction_def intro: the1I2) lemma app_ins2: "x' \ x \ app (f \ \x, y\) x' = app f x'" by (simp add: app_def) lemma hfunction_0 [simp]: "hfunction 0" by (force simp add: hfunction_def) lemma hfunction_ins: "hfunction f \ x \<^bold>\ hdomain f \ hfunction (f\ \x, y\)" by (auto simp: hfunction_def hdomain_def) lemma hdomainI: "\x, y\ \<^bold>\ f \ x \<^bold>\ hdomain f" by (auto simp: hdomain_def) lemma hfunction_hunion: "hdomain f \ hdomain g = 0 \ hfunction (f \ g) \ hfunction f \ hfunction g" by (auto simp: hfunction_def) (metis hdomainI hinter_iff hmem_hempty)+ lemma app_hrestrict [simp]: "x \<^bold>\ A \ app (hrestrict f A) x = app f x" by (simp add: hrestrict_def app_def) section \Operations on families of sets\ definition HLambda :: "hf \ (hf \ hf) \ hf" where "HLambda A b = RepFun A (\x. \x, b x\)" definition HSigma :: "hf \ (hf \ hf) \ hf" where "HSigma A B = (\x\<^bold>\A. \y\<^bold>\B(x). \\x,y\\)" definition HPi :: "hf \ (hf \ hf) \ hf" where "HPi A B = \ f \<^bold>\ HPow(HSigma A B). A \ hdomain(f) \ hfunction(f)\" syntax (ASCII) "_PROD" :: "[pttrn, hf, hf] \ hf" ("(3PROD _<:_./ _)" 10) "_SUM" :: "[pttrn, hf, hf] \ hf" ("(3SUM _<:_./ _)" 10) "_lam" :: "[pttrn, hf, hf] \ hf" ("(3lam _<:_./ _)" 10) syntax "_PROD" :: "[pttrn, hf, hf] \ hf" ("(3\_\<^bold>\_./ _)" 10) "_SUM" :: "[pttrn, hf, hf] \ hf" ("(3\_\<^bold>\_./ _)" 10) "_lam" :: "[pttrn, hf, hf] \ hf" ("(3\_\<^bold>\_./ _)" 10) translations "\x\<^bold>\A. B" \ "CONST HPi A (\x. B)" "\x\<^bold>\A. B" \ "CONST HSigma A (\x. B)" "\x\<^bold>\A. f" \ "CONST HLambda A (\x. f)" subsection \Rules for Unions and Intersections of families\ lemma HUN_iff [simp]: "b \<^bold>\ (\x\<^bold>\A. B(x)) \ (\x\<^bold>\A. b \<^bold>\ B(x))" by auto (*The order of the premises presupposes that A is rigid; b may be flexible*) lemma HUN_I: "\ a \<^bold>\ A; b \<^bold>\ B(a) \ \ b \<^bold>\ (\x\<^bold>\A. B(x))" by auto lemma HUN_E [elim!]: assumes "b \<^bold>\ (\x\<^bold>\A. B(x))" obtains x where "x \<^bold>\ A" "b \<^bold>\ B(x)" using assms by blast lemma HINT_iff: "b \<^bold>\ (\x\<^bold>\A. B(x)) \ (\x\<^bold>\A. b \<^bold>\ B(x)) \ A\0" by (simp add: HInter_def HBall_def) (metis foundation hmem_hempty) lemma HINT_I: "\ \x. x \<^bold>\ A \ b \<^bold>\ B(x); A\0 \ \ b \<^bold>\ (\x\<^bold>\A. B(x))" by (simp add: HINT_iff) lemma HINT_E: "\ b \<^bold>\ (\x\<^bold>\A. B(x)); a \<^bold>\ A \ \ b \<^bold>\ B(a)" by (auto simp: HINT_iff) subsection \Generalized Cartesian product\ lemma HSigma_iff [simp]: "\a,b\ \<^bold>\ HSigma A B \ a \<^bold>\ A \ b \<^bold>\ B(a)" by (force simp add: HSigma_def) lemma HSigmaI [intro!]: "\ a \<^bold>\ A; b \<^bold>\ B(a) \ \ \a,b\ \<^bold>\ HSigma A B" by simp lemmas HSigmaD1 = HSigma_iff [THEN iffD1, THEN conjunct1] lemmas HSigmaD2 = HSigma_iff [THEN iffD1, THEN conjunct2] text \The general elimination rule\ lemma HSigmaE [elim!]: assumes "c \<^bold>\ HSigma A B" obtains x y where "x \<^bold>\ A" "y \<^bold>\ B(x)" "c=\x,y\" using assms by (force simp add: HSigma_def) lemma HSigmaE2 [elim!]: assumes "\a,b\ \<^bold>\ HSigma A B" obtains "a \<^bold>\ A" and "b \<^bold>\ B(a)" using assms by auto lemma HSigma_empty1 [simp]: "HSigma 0 B = 0" by blast instantiation hf :: times begin definition "A * B = HSigma A (\x. B)" instance .. end lemma times_iff [simp]: "\a,b\ \<^bold>\ A * B \ a \<^bold>\ A \ b \<^bold>\ B" by (simp add: times_hf_def) lemma timesI [intro!]: "\ a \<^bold>\ A; b \<^bold>\ B \ \ \a,b\ \<^bold>\ A * B" by simp lemmas timesD1 = times_iff [THEN iffD1, THEN conjunct1] lemmas timesD2 = times_iff [THEN iffD1, THEN conjunct2] text \The general elimination rule\ lemma timesE [elim!]: assumes c: "c \<^bold>\ A * B" obtains x y where "x \<^bold>\ A" "y \<^bold>\ B" "c=\x,y\" using c by (auto simp: times_hf_def) text \...and a specific one\ lemma timesE2 [elim!]: assumes "\a,b\ \<^bold>\ A * B" obtains "a \<^bold>\ A" and "b \<^bold>\ B" using assms by auto lemma times_empty1 [simp]: "0 * B = (0::hf)" by auto lemma times_empty2 [simp]: "A*0 = (0::hf)" by blast lemma times_empty_iff: "A*B=0 \ A=0 \ B=(0::hf)" by (auto simp: times_hf_def hf_ext) instantiation hf :: mult_zero begin instance by standard auto end section \Disjoint Sum\ instantiation hf :: zero_neq_one begin definition One_hf_def: "1 = \0\" instance by standard (auto simp: One_hf_def) end instantiation hf :: plus begin definition "A + B = (\0\ * A) \ (\1\ * B)" instance .. end definition Inl :: "hf\hf" where "Inl(a) \ \0,a\" definition Inr :: "hf\hf" where "Inr(b) \ \1,b\" lemmas sum_defs = plus_hf_def Inl_def Inr_def lemma Inl_nonzero [simp]:"Inl x \ 0" by (metis Inl_def hpair_nonzero) lemma Inr_nonzero [simp]:"Inr x \ 0" by (metis Inr_def hpair_nonzero) text\Introduction rules for the injections (as equivalences)\ lemma Inl_in_sum_iff [iff]: "Inl(a) \<^bold>\ A+B \ a \<^bold>\ A" by (auto simp: sum_defs) lemma Inr_in_sum_iff [iff]: "Inr(b) \<^bold>\ A+B \ b \<^bold>\ B" by (auto simp: sum_defs) text \Elimination rule\ lemma sumE [elim!]: assumes u: "u \<^bold>\ A+B" obtains x where "x \<^bold>\ A" "u=Inl(x)" | y where "y \<^bold>\ B" "u=Inr(y)" using u by (auto simp: sum_defs) text \Injection and freeness equivalences, for rewriting\ lemma Inl_iff [iff]: "Inl(a)=Inl(b) \ a=b" by (simp add: sum_defs) lemma Inr_iff [iff]: "Inr(a)=Inr(b) \ a=b" by (simp add: sum_defs) lemma Inl_Inr_iff [iff]: "Inl(a)=Inr(b) \ False" by (simp add: sum_defs) lemma Inr_Inl_iff [iff]: "Inr(b)=Inl(a) \ False" by (simp add: sum_defs) lemma sum_empty [simp]: "0+0 = (0::hf)" by (auto simp: sum_defs) lemma sum_iff: "u \<^bold>\ A+B \ (\x. x \<^bold>\ A \ u=Inl(x)) \ (\y. y \<^bold>\ B \ u=Inr(y))" by blast lemma sum_subset_iff: fixes A :: hf shows "A+B \ C+D \ A\C \ B\D" by blast lemma sum_equal_iff: fixes A :: hf shows "A+B = C+D \ A=C \ B=D" by (auto simp: hf_ext sum_subset_iff) definition is_hsum :: "hf \ bool" where "is_hsum z = (\x. z = Inl x \ z = Inr x)" definition sum_case :: "(hf \ 'a) \ (hf \ 'a) \ hf \ 'a" where "sum_case f g a \ THE z. (\x. a = Inl x \ z = f x) \ (\y. a = Inr y \ z = g y) \ (\ is_hsum a \ z = undefined)" lemma sum_case_Inl [simp]: "sum_case f g (Inl x) = f x" by (simp add: sum_case_def is_hsum_def) lemma sum_case_Inr [simp]: "sum_case f g (Inr y) = g y" by (simp add: sum_case_def is_hsum_def) lemma sum_case_non [simp]: "\ is_hsum a \ sum_case f g a = undefined" by (simp add: sum_case_def is_hsum_def) lemma is_hsum_cases: "(\x. z = Inl x \ z = Inr x) \ \ is_hsum z" by (auto simp: is_hsum_def) lemma sum_case_split: "P (sum_case f g a) \ (\x. a = Inl x \ P(f x)) \ (\y. a = Inr y \ P(g y)) \ (\ is_hsum a \ P undefined)" by (cases "is_hsum a") (auto simp: is_hsum_def) lemma sum_case_split_asm: "P (sum_case f g a) \ \ ((\x. a = Inl x \ \ P(f x)) \ (\y. a = Inr y \ \ P(g y)) \ (\ is_hsum a \ \ P undefined))" by (auto simp add: sum_case_split) end diff --git a/thys/IP_Addresses/IPv6.thy b/thys/IP_Addresses/IPv6.thy --- a/thys/IP_Addresses/IPv6.thy +++ b/thys/IP_Addresses/IPv6.thy @@ -1,956 +1,958 @@ (* Title: IPv6.thy Authors: Cornelius Diekmann *) theory IPv6 imports IP_Address NumberWang_IPv6 (* include "HOL-Library.Code_Target_Nat" if you need to work with actual numbers.*) begin section \IPv6 Addresses\ text\An IPv6 address is basically a 128 bit unsigned integer. RFC 4291, Section 2.\ type_synonym ipv6addr = "128 word" text\Conversion between natural numbers and IPv6 adresses\ definition nat_of_ipv6addr :: "ipv6addr \ nat" where "nat_of_ipv6addr a = unat a" definition ipv6addr_of_nat :: "nat \ ipv6addr" where "ipv6addr_of_nat n = of_nat n" lemma "ipv6addr_of_nat n = word_of_int (int n)" by(simp add: ipv6addr_of_nat_def) text\The maximum IPv6 address\ definition max_ipv6_addr :: "ipv6addr" where "max_ipv6_addr \ ipv6addr_of_nat ((2^128) - 1)" lemma max_ipv6_addr_number: "max_ipv6_addr = 0xFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFF" unfolding max_ipv6_addr_def ipv6addr_of_nat_def by(simp) lemma "max_ipv6_addr = 340282366920938463463374607431768211455" by(fact max_ipv6_addr_number) lemma max_ipv6_addr_max_word: "max_ipv6_addr = - 1" by(simp add: max_ipv6_addr_number) lemma max_ipv6_addr_max: "\a. a \ max_ipv6_addr" by(simp add: max_ipv6_addr_max_word) lemma UNIV_ipv6addrset: "UNIV = {0 .. max_ipv6_addr}" (*not in the simp set, for a reason*) by(simp add: max_ipv6_addr_max_word) fastforce text\identity functions\ lemma nat_of_ipv6addr_ipv6addr_of_nat_mod: "nat_of_ipv6addr (ipv6addr_of_nat n) = n mod 2^128" by (simp add: ipv6addr_of_nat_def nat_of_ipv6addr_def unat_of_nat take_bit_eq_mod) lemma nat_of_ipv6addr_ipv6addr_of_nat: "n \ nat_of_ipv6addr max_ipv6_addr \ nat_of_ipv6addr (ipv6addr_of_nat n) = n" by (simp add: nat_of_ipv6addr_ipv6addr_of_nat_mod max_ipv6_addr_def) lemma ipv6addr_of_nat_nat_of_ipv6addr: "ipv6addr_of_nat (nat_of_ipv6addr addr) = addr" by(simp add: ipv6addr_of_nat_def nat_of_ipv6addr_def) subsection\Syntax of IPv6 Adresses\ text\RFC 4291, Section 2.2.: Text Representation of Addresses\ text\Quoting the RFC (note: errata exists):\ text_raw\ \begin{verbatim} 1. The preferred form is x:x:x:x:x:x:x:x, where the 'x's are one to four hexadecimal digits of the eight 16-bit pieces of the address. Examples: ABCD:EF01:2345:6789:ABCD:EF01:2345:6789 2001:DB8:0:0:8:800:200C:417A \end{verbatim} \ datatype ipv6addr_syntax = IPv6AddrPreferred "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" text_raw\ \begin{verbatim} 2. [...] In order to make writing addresses containing zero bits easier, a special syntax is available to compress the zeros. The use of "::" indicates one or more groups of 16 bits of zeros. The "::" can only appear once in an address. The "::" can also be used to compress leading or trailing zeros in an address. For example, the following addresses 2001:DB8:0:0:8:800:200C:417A a unicast address FF01:0:0:0:0:0:0:101 a multicast address 0:0:0:0:0:0:0:1 the loopback address 0:0:0:0:0:0:0:0 the unspecified address may be represented as 2001:DB8::8:800:200C:417A a unicast address FF01::101 a multicast address ::1 the loopback address :: the unspecified address \end{verbatim} \ (*datatype may take some minutes to load*) datatype ipv6addr_syntax_compressed = \ \using @{typ unit} for the omission @{text "::"}. Naming convention of the datatype: The first number is the position where the omission occurs. The second number is the length of the specified address pieces. I.e. `8 minus the second number' pieces are omitted.\ IPv6AddrCompressed1_0 unit | IPv6AddrCompressed1_1 unit "16 word" | IPv6AddrCompressed1_2 unit "16 word" "16 word" | IPv6AddrCompressed1_3 unit "16 word" "16 word" "16 word" | IPv6AddrCompressed1_4 unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_5 unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_6 unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_7 unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_1 "16 word" unit | IPv6AddrCompressed2_2 "16 word" unit "16 word" | IPv6AddrCompressed2_3 "16 word" unit "16 word" "16 word" | IPv6AddrCompressed2_4 "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed2_5 "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_6 "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_7 "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed3_2 "16 word" "16 word" unit | IPv6AddrCompressed3_3 "16 word" "16 word" unit "16 word" | IPv6AddrCompressed3_4 "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed3_5 "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed3_6 "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed3_7 "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed4_3 "16 word" "16 word" "16 word" unit | IPv6AddrCompressed4_4 "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed4_5 "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed4_6 "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed4_7 "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed5_4 "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed5_5 "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed5_6 "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed5_7 "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed6_5 "16 word" "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed6_6 "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed6_7 "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed7_6 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed7_7 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed8_7 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit (*RFC 5952: """ 4. A Recommendation for IPv6 Text Representation 4.2.2. Handling One 16-Bit 0 Field The symbol "::" MUST NOT be used to shorten just one 16-bit 0 field. For example, the representation 2001:db8:0:1:1:1:1:1 is correct, but 2001:db8::1:1:1:1:1 is not correct. """ So we could remove all IPv6AddrCompressed*_7 constructors. But these are `recommendations', we might still see these non-recommended definitions. "[...] all implementations must accept and be able to handle any legitimate RFC 4291 format." *) (*More convenient parser helper function for compressed IPv6 addresses: Input list (from parser): Some 16word \ address piece None \ omission '::' Basically, the parser must only do the following (python syntax): split the string which is an ipv6 address at ':' map empty string to None map everything else to Some (string_to_16word str) sanitize empty strings at the start and the end (see toString and parser theories) Example: "1:2:3".split(":") = ['1', '2', '3'] ":2:3:4".split(":") = ['', '2', '3', '4'] ":2::3".split(":") = ['', '2', '', '3'] "1:2:3:".split(":") = ['1', '2', '3', ''] *) definition parse_ipv6_address_compressed :: "((16 word) option) list \ ipv6addr_syntax_compressed option" where "parse_ipv6_address_compressed as = (case as of [None] \ Some (IPv6AddrCompressed1_0 ()) | [None, Some a] \ Some (IPv6AddrCompressed1_1 () a) | [None, Some a, Some b] \ Some (IPv6AddrCompressed1_2 () a b) | [None, Some a, Some b, Some c] \ Some (IPv6AddrCompressed1_3 () a b c) | [None, Some a, Some b, Some c, Some d] \ Some (IPv6AddrCompressed1_4 () a b c d) | [None, Some a, Some b, Some c, Some d, Some e] \ Some (IPv6AddrCompressed1_5 () a b c d e) | [None, Some a, Some b, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed1_6 () a b c d e f) | [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed1_7 () a b c d e f g) | [Some a, None] \ Some (IPv6AddrCompressed2_1 a ()) | [Some a, None, Some b] \ Some (IPv6AddrCompressed2_2 a () b) | [Some a, None, Some b, Some c] \ Some (IPv6AddrCompressed2_3 a () b c) | [Some a, None, Some b, Some c, Some d] \ Some (IPv6AddrCompressed2_4 a () b c d) | [Some a, None, Some b, Some c, Some d, Some e] \ Some (IPv6AddrCompressed2_5 a () b c d e) | [Some a, None, Some b, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed2_6 a () b c d e f) | [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed2_7 a () b c d e f g) | [Some a, Some b, None] \ Some (IPv6AddrCompressed3_2 a b ()) | [Some a, Some b, None, Some c] \ Some (IPv6AddrCompressed3_3 a b () c) | [Some a, Some b, None, Some c, Some d] \ Some (IPv6AddrCompressed3_4 a b () c d) | [Some a, Some b, None, Some c, Some d, Some e] \ Some (IPv6AddrCompressed3_5 a b () c d e) | [Some a, Some b, None, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed3_6 a b () c d e f) | [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed3_7 a b () c d e f g) | [Some a, Some b, Some c, None] \ Some (IPv6AddrCompressed4_3 a b c ()) | [Some a, Some b, Some c, None, Some d] \ Some (IPv6AddrCompressed4_4 a b c () d) | [Some a, Some b, Some c, None, Some d, Some e] \ Some (IPv6AddrCompressed4_5 a b c () d e) | [Some a, Some b, Some c, None, Some d, Some e, Some f] \ Some (IPv6AddrCompressed4_6 a b c () d e f) | [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed4_7 a b c () d e f g) | [Some a, Some b, Some c, Some d, None] \ Some (IPv6AddrCompressed5_4 a b c d ()) | [Some a, Some b, Some c, Some d, None, Some e] \ Some (IPv6AddrCompressed5_5 a b c d () e) | [Some a, Some b, Some c, Some d, None, Some e, Some f] \ Some (IPv6AddrCompressed5_6 a b c d () e f) | [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g] \ Some (IPv6AddrCompressed5_7 a b c d () e f g) | [Some a, Some b, Some c, Some d, Some e, None] \ Some (IPv6AddrCompressed6_5 a b c d e ()) | [Some a, Some b, Some c, Some d, Some e, None, Some f] \ Some (IPv6AddrCompressed6_6 a b c d e () f) | [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g] \ Some (IPv6AddrCompressed6_7 a b c d e () f g) | [Some a, Some b, Some c, Some d, Some e, Some f, None] \ Some (IPv6AddrCompressed7_6 a b c d e f ()) | [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g] \ Some (IPv6AddrCompressed7_7 a b c d e f () g) | [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None] \ Some (IPv6AddrCompressed8_7 a b c d e f g ()) | _ \ None \ \invalid ipv6 copressed address.\ )" fun ipv6addr_syntax_compressed_to_list :: "ipv6addr_syntax_compressed \ ((16 word) option) list" where "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_0 _) = [None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_1 () a) = [None, Some a]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_2 () a b) = [None, Some a, Some b]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_3 () a b c) = [None, Some a, Some b, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_4 () a b c d) = [None, Some a, Some b, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_5 () a b c d e) = [None, Some a, Some b, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_6 () a b c d e f) = [None, Some a, Some b, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_7 () a b c d e f g) = [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_1 a ()) = [Some a, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_2 a () b) = [Some a, None, Some b]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_3 a () b c) = [Some a, None, Some b, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_4 a () b c d) = [Some a, None, Some b, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_5 a () b c d e) = [Some a, None, Some b, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_6 a () b c d e f) = [Some a, None, Some b, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_7 a () b c d e f g) = [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_2 a b ()) = [Some a, Some b, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_3 a b () c) = [Some a, Some b, None, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_4 a b () c d) = [Some a, Some b, None, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_5 a b () c d e) = [Some a, Some b, None, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_6 a b () c d e f) = [Some a, Some b, None, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_7 a b () c d e f g) = [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_3 a b c ()) = [Some a, Some b, Some c, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_4 a b c () d) = [Some a, Some b, Some c, None, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_5 a b c () d e) = [Some a, Some b, Some c, None, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_6 a b c () d e f) = [Some a, Some b, Some c, None, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_7 a b c () d e f g) = [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_4 a b c d ()) = [Some a, Some b, Some c, Some d, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_5 a b c d () e) = [Some a, Some b, Some c, Some d, None, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_6 a b c d () e f) = [Some a, Some b, Some c, Some d, None, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_7 a b c d () e f g) = [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_5 a b c d e ()) = [Some a, Some b, Some c, Some d, Some e, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_6 a b c d e () f) = [Some a, Some b, Some c, Some d, Some e, None, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_7 a b c d e () f g) = [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed7_6 a b c d e f ()) = [Some a, Some b, Some c, Some d, Some e, Some f, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed7_7 a b c d e f () g) = [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed8_7 a b c d e f g ()) = [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None]" (*for all ipv6_syntax, there is a corresponding list representation*) lemma parse_ipv6_address_compressed_exists: obtains ss where "parse_ipv6_address_compressed ss = Some ipv6_syntax" proof define ss where "ss = ipv6addr_syntax_compressed_to_list ipv6_syntax" thus "parse_ipv6_address_compressed ss = Some ipv6_syntax" by (cases ipv6_syntax; simp add: parse_ipv6_address_compressed_def) qed lemma parse_ipv6_address_compressed_identity: "parse_ipv6_address_compressed (ipv6addr_syntax_compressed_to_list (ipv6_syntax)) = Some ipv6_syntax" by(cases ipv6_syntax; simp add: parse_ipv6_address_compressed_def) lemma parse_ipv6_address_compressed_someE: assumes "parse_ipv6_address_compressed as = Some ipv6" obtains "as = [None]" "ipv6 = (IPv6AddrCompressed1_0 ())" | a where "as = [None, Some a]" "ipv6 = (IPv6AddrCompressed1_1 () a)" | a b where "as = [None, Some a, Some b]" "ipv6 = (IPv6AddrCompressed1_2 () a b)" | a b c where "as = [None, Some a, Some b, Some c]" "ipv6 = (IPv6AddrCompressed1_3 () a b c)" | a b c d where "as = [None, Some a, Some b, Some c, Some d]" "ipv6 = (IPv6AddrCompressed1_4 () a b c d)" | a b c d e where "as = [None, Some a, Some b, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed1_5 () a b c d e)" | a b c d e f where "as = [None, Some a, Some b, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed1_6 () a b c d e f)" | a b c d e f g where "as = [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed1_7 () a b c d e f g)" | a where "as = [Some a, None]" "ipv6 = (IPv6AddrCompressed2_1 a ())" | a b where "as = [Some a, None, Some b]" "ipv6 = (IPv6AddrCompressed2_2 a () b)" | a b c where "as = [Some a, None, Some b, Some c]" "ipv6 = (IPv6AddrCompressed2_3 a () b c)" | a b c d where "as = [Some a, None, Some b, Some c, Some d]" "ipv6 = (IPv6AddrCompressed2_4 a () b c d)" | a b c d e where "as = [Some a, None, Some b, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed2_5 a () b c d e)" | a b c d e f where "as = [Some a, None, Some b, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed2_6 a () b c d e f)" | a b c d e f g where "as = [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed2_7 a () b c d e f g)" | a b where "as = [Some a, Some b, None]" "ipv6 = (IPv6AddrCompressed3_2 a b ())" | a b c where "as = [Some a, Some b, None, Some c]" "ipv6 = (IPv6AddrCompressed3_3 a b () c)" | a b c d where "as = [Some a, Some b, None, Some c, Some d]" "ipv6 = (IPv6AddrCompressed3_4 a b () c d)" | a b c d e where "as = [Some a, Some b, None, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed3_5 a b () c d e)" | a b c d e f where "as = [Some a, Some b, None, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed3_6 a b () c d e f)" | a b c d e f g where "as = [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed3_7 a b () c d e f g)" | a b c where "as = [Some a, Some b, Some c, None]" "ipv6 = (IPv6AddrCompressed4_3 a b c ())" | a b c d where "as = [Some a, Some b, Some c, None, Some d]" "ipv6 = (IPv6AddrCompressed4_4 a b c () d)" | a b c d e where "as = [Some a, Some b, Some c, None, Some d, Some e]" "ipv6 = (IPv6AddrCompressed4_5 a b c () d e)" | a b c d e f where "as = [Some a, Some b, Some c, None, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed4_6 a b c () d e f)" | a b c d e f g where "as = [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed4_7 a b c () d e f g)" | a b c d where "as = [Some a, Some b, Some c, Some d, None]" "ipv6 = (IPv6AddrCompressed5_4 a b c d ())" | a b c d e where "as = [Some a, Some b, Some c, Some d, None, Some e]" "ipv6 = (IPv6AddrCompressed5_5 a b c d () e)" | a b c d e f where "as = [Some a, Some b, Some c, Some d, None, Some e, Some f]" "ipv6 = (IPv6AddrCompressed5_6 a b c d () e f)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed5_7 a b c d () e f g)" | a b c d e where "as = [Some a, Some b, Some c, Some d, Some e, None]" "ipv6 = (IPv6AddrCompressed6_5 a b c d e ())" | a b c d e f where "as = [Some a, Some b, Some c, Some d, Some e, None, Some f]" "ipv6 = (IPv6AddrCompressed6_6 a b c d e () f)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g]" "ipv6 = (IPv6AddrCompressed6_7 a b c d e () f g)" | a b c d e f where "as = [Some a, Some b, Some c, Some d, Some e, Some f, None]" "ipv6 = (IPv6AddrCompressed7_6 a b c d e f ())" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g]" "ipv6 = (IPv6AddrCompressed7_7 a b c d e f () g)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None]" "ipv6 = (IPv6AddrCompressed8_7 a b c d e f g ())" using assms unfolding parse_ipv6_address_compressed_def by (auto split: list.split_asm option.split_asm) (* takes a minute *) lemma parse_ipv6_address_compressed_identity2: "ipv6addr_syntax_compressed_to_list ipv6_syntax = ls \ (parse_ipv6_address_compressed ls) = Some ipv6_syntax" (is "?lhs = ?rhs") proof assume ?rhs thus ?lhs by (auto elim: parse_ipv6_address_compressed_someE) next assume ?lhs thus ?rhs by (cases ipv6_syntax) (auto simp: parse_ipv6_address_compressed_def) qed text\Valid IPv6 compressed notation: \<^item> at most one omission \<^item> at most 7 pieces \ lemma RFC_4291_format: "parse_ipv6_address_compressed as \ None \ length (filter (\p. p = None) as) = 1 \ length (filter (\p. p \ None) as) \ 7" (is "?lhs = ?rhs") proof assume ?lhs then obtain addr where "parse_ipv6_address_compressed as = Some addr" by blast thus ?rhs by (elim parse_ipv6_address_compressed_someE; simp) next assume ?rhs thus ?lhs unfolding parse_ipv6_address_compressed_def by (auto split: option.split list.split if_split_asm) qed text_raw\ \begin{verbatim} 3. An alternative form that is sometimes more convenient when dealing with a mixed environment of IPv4 and IPv6 nodes is x:x:x:x:x:x:d.d.d.d, where the 'x's are the hexadecimal values of the six high-order 16-bit pieces of the address, and the 'd's are the decimal values of the four low-order 8-bit pieces of the address (standard IPv4 representation). Examples: 0:0:0:0:0:0:13.1.68.3 0:0:0:0:0:FFFF:129.144.52.38 or in compressed form: ::13.1.68.3 ::FFFF:129.144.52.38 \end{verbatim} This is currently not supported by our library! \ (*TODO*) (*TODO: oh boy, they can also be compressed*) subsection\Semantics\ fun ipv6preferred_to_int :: "ipv6addr_syntax \ ipv6addr" where "ipv6preferred_to_int (IPv6AddrPreferred a b c d e f g h) = (ucast a << (16 * 7)) OR (ucast b << (16 * 6)) OR (ucast c << (16 * 5)) OR (ucast d << (16 * 4)) OR (ucast e << (16 * 3)) OR (ucast f << (16 * 2)) OR (ucast g << (16 * 1)) OR (ucast h << (16 * 0))" lemma "ipv6preferred_to_int (IPv6AddrPreferred 0x2001 0xDB8 0x0 0x0 0x8 0x800 0x200C 0x417A) = 42540766411282592856906245548098208122" by eval lemma "ipv6preferred_to_int (IPv6AddrPreferred 0xFF01 0x0 0x0 0x0 0x0 0x0 0x0 0x101) = 338958331222012082418099330867817087233" by eval declare ipv6preferred_to_int.simps[simp del] definition int_to_ipv6preferred :: "ipv6addr \ ipv6addr_syntax" where "int_to_ipv6preferred i = IPv6AddrPreferred (ucast ((i AND 0xFFFF0000000000000000000000000000) >> 16*7)) (ucast ((i AND 0xFFFF000000000000000000000000) >> 16*6)) (ucast ((i AND 0xFFFF00000000000000000000) >> 16*5)) (ucast ((i AND 0xFFFF0000000000000000) >> 16*4)) (ucast ((i AND 0xFFFF000000000000) >> 16*3)) (ucast ((i AND 0xFFFF00000000) >> 16*2)) (ucast ((i AND 0xFFFF0000) >> 16*1)) (ucast ((i AND 0xFFFF)))" lemma "int_to_ipv6preferred 42540766411282592856906245548098208122 = IPv6AddrPreferred 0x2001 0xDB8 0x0 0x0 0x8 0x800 0x200C 0x417A" by eval lemma word128_masks_ipv6pieces: "(0xFFFF0000000000000000000000000000::ipv6addr) = (mask 16) << 112" "(0xFFFF000000000000000000000000::ipv6addr) = (mask 16) << 96" "(0xFFFF00000000000000000000::ipv6addr) = (mask 16) << 80" "(0xFFFF0000000000000000::ipv6addr) = (mask 16) << 64" "(0xFFFF000000000000::ipv6addr) = (mask 16) << 48" "(0xFFFF00000000::ipv6addr) = (mask 16) << 32" "(0xFFFF0000::ipv6addr) = (mask 16) << 16" "(0xFFFF::ipv6addr) = (mask 16)" by (simp_all add: mask_eq push_bit_of_1) text\Correctness: round trip property one\ lemma ipv6preferred_to_int_int_to_ipv6preferred: "ipv6preferred_to_int (int_to_ipv6preferred ip) = ip" proof - have and_mask_shift_helper: "w AND (mask m << n) >> n << n = w AND (mask m << n)" for m n::nat and w::ipv6addr by (metis is_aligned_shift is_aligned_shiftr_shiftl shiftr_and_eq_shiftl) have ucast_ipv6_piece_rule: "length (dropWhile Not (to_bl w)) \ 16 \ (ucast::16 word \ 128 word) ((ucast::128 word \ 16 word) w) = w" for w::ipv6addr by(rule ucast_short_ucast_long_ingoreLeadingZero) (simp_all) have ucast_ipv6_piece: "16 \ 128 - n \ (ucast::16 word \ 128 word) ((ucast::128 word \ 16 word) (w AND (mask 16 << n) >> n)) << n = w AND (mask 16 << n)" for w::ipv6addr and n::nat apply(subst ucast_ipv6_piece_rule) apply(rule length_drop_mask_inner) apply(simp; fail) apply(subst and_mask_shift_helper) apply simp done have ucast16_ucast128_masks_highest_bits: "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000000000000000000000000000 >> 112)) << 112) = (ip AND 0xFFFF0000000000000000000000000000)" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF000000000000000000000000 >> 96)) << 96) = ip AND 0xFFFF000000000000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF00000000000000000000 >> 80)) << 80) = ip AND 0xFFFF00000000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000000000000000 >> 64)) << 64) = ip AND 0xFFFF0000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF000000000000 >> 48)) << 48) = ip AND 0xFFFF000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF00000000 >> 32)) << 32) = ip AND 0xFFFF00000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000 >> 16)) << 16) = ip AND 0xFFFF0000" apply (simp_all add: word128_masks_ipv6pieces ucast_ipv6_piece and_mask2 word_size flip: take_bit_eq_mask) apply (simp_all add: bit_eq_iff) apply (auto simp add: bit_simps) done have ucast16_ucast128_masks_highest_bits0: "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF))) = ip AND 0xFFFF" apply(subst word128_masks_ipv6pieces)+ apply(subst ucast_short_ucast_long_ingoreLeadingZero) apply simp_all by (simp add: length_drop_mask) have mask_len_word:"n = (LENGTH('a)) \ w AND mask n = w" for n and w::"'a::len word" by (simp add: mask_eq_iff) have ipv6addr_16word_pieces_compose_or: "ip && (mask 16 << 112) || ip && (mask 16 << 96) || ip && (mask 16 << 80) || ip && (mask 16 << 64) || ip && (mask 16 << 48) || ip && (mask 16 << 32) || ip && (mask 16 << 16) || ip && mask 16 = ip" apply(subst word_ao_dist2[symmetric])+ apply(simp add: mask_eq push_bit_of_1) apply(subst mask128) apply(rule mask_len_word) apply simp done show ?thesis apply (simp add: ipv6preferred_to_int.simps int_to_ipv6preferred_def) apply (simp add: ucast16_ucast128_masks_highest_bits ucast16_ucast128_masks_highest_bits0) apply (simp add: word128_masks_ipv6pieces ucast_ucast_mask) using ipv6addr_16word_pieces_compose_or apply (simp flip: push_bit_and add: shiftr_and_eq_shiftl) done qed text\Correctness: round trip property two\ lemma int_to_ipv6preferred_ipv6preferred_to_int: "int_to_ipv6preferred (ipv6preferred_to_int ip) = ip" proof - note ucast_shift_simps=helper_masked_ucast_generic helper_masked_ucast_reverse_generic helper_masked_ucast_generic[where n=0, simplified] helper_masked_ucast_equal_generic note ucast_simps=helper_masked_ucast_reverse_generic[where m=0, simplified] helper_masked_ucast_equal_generic[where n=0, simplified] show ?thesis apply (cases ip, rename_tac a b c d e f g h) apply (simp add: ipv6preferred_to_int.simps int_to_ipv6preferred_def) apply (simp add: word128_masks_ipv6pieces) apply (simp add: word_ao_dist ucast_shift_simps ucast_simps) apply (simp add: unsigned_or_eq) apply (simp flip: take_bit_eq_mask add: unsigned_take_bit_eq take_bit_word_eq_self) apply (simp add: shiftl_shiftr1 shiftl_shiftr2) apply (simp flip: take_bit_eq_mask push_bit_and add: word_size) apply (simp add: unsigned_take_bit_eq take_bit_word_eq_self) apply (simp flip: unsigned_take_bit_eq) apply (simp add: unsigned_ucast_eq) apply (simp add: unsigned_push_bit_eq take_bit_word_eq_self) apply (simp flip: ucast_drop_bit_eq) done qed text\compressed to preferred format\ fun ipv6addr_c2p :: "ipv6addr_syntax_compressed \ ipv6addr_syntax" where "ipv6addr_c2p (IPv6AddrCompressed1_0 ()) = IPv6AddrPreferred 0 0 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed1_1 () h) = IPv6AddrPreferred 0 0 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed1_2 () g h) = IPv6AddrPreferred 0 0 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed1_3 () f g h) = IPv6AddrPreferred 0 0 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_4 () e f g h) = IPv6AddrPreferred 0 0 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_5 () d e f g h) = IPv6AddrPreferred 0 0 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_6 () c d e f g h) = IPv6AddrPreferred 0 0 c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_7 () b c d e f g h) = IPv6AddrPreferred 0 b c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_1 a ()) = IPv6AddrPreferred a 0 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed2_2 a () h) = IPv6AddrPreferred a 0 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed2_3 a () g h) = IPv6AddrPreferred a 0 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed2_4 a () f g h) = IPv6AddrPreferred a 0 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_5 a () e f g h) = IPv6AddrPreferred a 0 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_6 a () d e f g h) = IPv6AddrPreferred a 0 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_7 a () c d e f g h) = IPv6AddrPreferred a 0 c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_2 a b ()) = IPv6AddrPreferred a b 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed3_3 a b () h) = IPv6AddrPreferred a b 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed3_4 a b () g h) = IPv6AddrPreferred a b 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed3_5 a b () f g h) = IPv6AddrPreferred a b 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_6 a b () e f g h) = IPv6AddrPreferred a b 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_7 a b () d e f g h) = IPv6AddrPreferred a b 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed4_3 a b c ()) = IPv6AddrPreferred a b c 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed4_4 a b c () h) = IPv6AddrPreferred a b c 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed4_5 a b c () g h) = IPv6AddrPreferred a b c 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed4_6 a b c () f g h) = IPv6AddrPreferred a b c 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed4_7 a b c () e f g h) = IPv6AddrPreferred a b c 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed5_4 a b c d ()) = IPv6AddrPreferred a b c d 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed5_5 a b c d () h) = IPv6AddrPreferred a b c d 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed5_6 a b c d () g h) = IPv6AddrPreferred a b c d 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed5_7 a b c d () f g h) = IPv6AddrPreferred a b c d 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed6_5 a b c d e ()) = IPv6AddrPreferred a b c d e 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed6_6 a b c d e () h) = IPv6AddrPreferred a b c d e 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed6_7 a b c d e () g h) = IPv6AddrPreferred a b c d e 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed7_6 a b c d e f ()) = IPv6AddrPreferred a b c d e f 0 0" | "ipv6addr_c2p (IPv6AddrCompressed7_7 a b c d e f () h) = IPv6AddrPreferred a b c d e f 0 h" | "ipv6addr_c2p (IPv6AddrCompressed8_7 a b c d e f g ()) = IPv6AddrPreferred a b c d e f g 0" definition ipv6_unparsed_compressed_to_preferred :: "((16 word) option) list \ ipv6addr_syntax option" where "ipv6_unparsed_compressed_to_preferred ls = ( if length (filter (\p. p = None) ls) \ 1 \ length (filter (\p. p \ None) ls) > 7 then None else let before_omission = map the (takeWhile (\x. x \ None) ls); after_omission = map the (drop 1 (dropWhile (\x. x \ None) ls)); num_omissions = 8 - (length before_omission + length after_omission); expanded = before_omission @ (replicate num_omissions 0) @ after_omission in case expanded of [a,b,c,d,e,f,g,h] \ Some (IPv6AddrPreferred a b c d e f g h) | _ \ None )" lemma "ipv6_unparsed_compressed_to_preferred [Some 0x2001, Some 0xDB8, None, Some 0x8, Some 0x800, Some 0x200C, Some 0x417A] = Some (IPv6AddrPreferred 0x2001 0xDB8 0 0 8 0x800 0x200C 0x417A)" by eval lemma "ipv6_unparsed_compressed_to_preferred [None] = Some (IPv6AddrPreferred 0 0 0 0 0 0 0 0)" by eval lemma "ipv6_unparsed_compressed_to_preferred [] = None" by eval lemma ipv6_unparsed_compressed_to_preferred_identity1: "ipv6_unparsed_compressed_to_preferred (ipv6addr_syntax_compressed_to_list ipv6compressed) = Some ipv6prferred \ ipv6addr_c2p ipv6compressed = ipv6prferred" by (cases ipv6compressed) (simp_all add: ipv6_unparsed_compressed_to_preferred_def numeral_eq_Suc) (*1s*) lemma ipv6_unparsed_compressed_to_preferred_identity2: "ipv6_unparsed_compressed_to_preferred ls = Some ipv6prferred \ (\ipv6compressed. parse_ipv6_address_compressed ls = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ipv6prferred)" apply(rule iffI) apply(subgoal_tac "parse_ipv6_address_compressed ls \ None") prefer 2 apply(subst RFC_4291_format) apply(simp add: ipv6_unparsed_compressed_to_preferred_def split: if_split_asm; fail) apply(simp) apply(erule exE, rename_tac ipv6compressed) apply(rule_tac x="ipv6compressed" in exI) apply(simp) apply(subgoal_tac "(ipv6addr_syntax_compressed_to_list ipv6compressed = ls)") prefer 2 using parse_ipv6_address_compressed_identity2 apply presburger using ipv6_unparsed_compressed_to_preferred_identity1 apply blast apply(erule exE, rename_tac ipv6compressed) apply(subgoal_tac "(ipv6addr_syntax_compressed_to_list ipv6compressed = ls)") prefer 2 using parse_ipv6_address_compressed_identity2 apply presburger using ipv6_unparsed_compressed_to_preferred_identity1 apply blast done subsection\IPv6 Pretty Printing (converting to compressed format)\ text_raw\ RFC5952: \begin{verbatim} 4. A Recommendation for IPv6 Text Representation A recommendation for a canonical text representation format of IPv6 addresses is presented in this section. The recommendation in this document is one that complies fully with [RFC4291], is implemented by various operating systems, and is human friendly. The recommendation in this section SHOULD be followed by systems when generating an address to be represented as text, but all implementations MUST accept and be able to handle any legitimate [RFC4291] format. It is advised that humans also follow these recommendations when spelling an address. 4.1. Handling Leading Zeros in a 16-Bit Field Leading zeros MUST be suppressed. For example, 2001:0db8::0001 is not acceptable and must be represented as 2001:db8::1. A single 16- bit 0000 field MUST be represented as 0. 4.2. "::" Usage 4.2.1. Shorten as Much as Possible The use of the symbol "::" MUST be used to its maximum capability. For example, 2001:db8:0:0:0:0:2:1 must be shortened to 2001:db8::2:1. Likewise, 2001:db8::0:1 is not acceptable, because the symbol "::" could have been used to produce a shorter representation 2001:db8::1. 4.2.2. Handling One 16-Bit 0 Field The symbol "::" MUST NOT be used to shorten just one 16-bit 0 field. For example, the representation 2001:db8:0:1:1:1:1:1 is correct, but 2001:db8::1:1:1:1:1 is not correct. 4.2.3. Choice in Placement of "::" When there is an alternative choice in the placement of a "::", the longest run of consecutive 16-bit 0 fields MUST be shortened (i.e., the sequence with three consecutive zero fields is shortened in 2001: 0:0:1:0:0:0:1). When the length of the consecutive 16-bit 0 fields are equal (i.e., 2001:db8:0:0:1:0:0:1), the first sequence of zero bits MUST be shortened. For example, 2001:db8::1:0:0:1 is correct representation. 4.3. Lowercase The characters "a", "b", "c", "d", "e", and "f" in an IPv6 address MUST be represented in lowercase. \end{verbatim} \ text\See @{file \IP_Address_toString.thy\} for examples and test cases.\ context begin private function goup_by_zeros :: "16 word list \ 16 word list list" where "goup_by_zeros [] = []" | "goup_by_zeros (x#xs) = ( if x = 0 then takeWhile (\x. x = 0) (x#xs) # (goup_by_zeros (dropWhile (\x. x = 0) xs)) else [x]#(goup_by_zeros xs))" by(pat_completeness, auto) termination goup_by_zeros apply(relation "measure (\xs. length xs)") apply(simp_all) by (simp add: le_imp_less_Suc length_dropWhile_le) private lemma "goup_by_zeros [0,1,2,3,0,0,0,0,3,4,0,0,0,2,0,0,2,0,3,0] = [[0], [1], [2], [3], [0, 0, 0, 0], [3], [4], [0, 0, 0], [2], [0, 0], [2], [0], [3], [0]]" by eval private lemma "concat (goup_by_zeros ls) = ls" by(induction ls rule:goup_by_zeros.induct) simp+ private lemma "[] \ set (goup_by_zeros ls)" by(induction ls rule:goup_by_zeros.induct) simp+ private primrec List_replace1 :: "'a \ 'a \ 'a list \ 'a list" where "List_replace1 _ _ [] = []" | "List_replace1 a b (x#xs) = (if a = x then b#xs else x#List_replace1 a b xs)" private lemma "List_replace1 a a ls = ls" by(induction ls) simp_all private lemma "a \ set ls \ List_replace1 a b ls = ls" by(induction ls) simp_all private lemma "a \ set ls \ b \ set (List_replace1 a b ls)" apply(induction ls) apply(simp) apply(simp) by blast private fun List_explode :: "'a list list \ ('a option) list" where "List_explode [] = []" | "List_explode ([]#xs) = None#List_explode xs" | "List_explode (xs1#xs2) = map Some xs1@List_explode xs2" private lemma "List_explode [[0::int], [2,3], [], [3,4]] = [Some 0, Some 2, Some 3, None, Some 3, Some 4]" by eval private lemma List_explode_def: "List_explode xss = concat (map (\xs. if xs = [] then [None] else map Some xs) xss)" by(induction xss rule: List_explode.induct) simp+ private lemma List_explode_no_empty: "[] \ set xss \ List_explode xss = map Some (concat xss)" by(induction xss rule: List_explode.induct) simp+ private lemma List_explode_replace1: "[] \ set xss \ foo \ set xss \ List_explode (List_replace1 foo [] xss) = map Some (concat (takeWhile (\xs. xs \ foo) xss)) @ [None] @ map Some (concat (tl (dropWhile (\xs. xs \ foo) xss)))" apply(induction xss rule: List_explode.induct) apply(simp; fail) apply(simp; fail) apply(simp) apply safe apply(simp_all add: List_explode_no_empty) done fun ipv6_preferred_to_compressed :: "ipv6addr_syntax \ ((16 word) option) list" where "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = ( let lss = goup_by_zeros [a,b,c,d,e,f,g,h]; max_zero_seq = foldr (\xs. max (length xs)) lss 0; shortened = if max_zero_seq > 1 then List_replace1 (replicate max_zero_seq 0) [] lss else lss in List_explode shortened )" declare ipv6_preferred_to_compressed.simps[simp del] private lemma foldr_max_length: "foldr (\xs. max (length xs)) lss n = fold max (map length lss) n" apply(subst List.foldr_fold) apply fastforce apply(induction lss arbitrary: n) apply(simp; fail) apply(simp) done private lemma List_explode_goup_by_zeros: "List_explode (goup_by_zeros xs) = map Some xs" apply(induction xs rule: goup_by_zeros.induct) apply(simp; fail) apply(simp) apply(safe) apply(simp) by (metis map_append takeWhile_dropWhile_id) private definition "max_zero_streak xs \ foldr (\xs. max (length xs)) (goup_by_zeros xs) 0" private lemma max_zero_streak_def2: "max_zero_streak xs = fold max (map length (goup_by_zeros xs)) 0" unfolding max_zero_streak_def by(simp add: foldr_max_length) private lemma ipv6_preferred_to_compressed_pull_out_if: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = ( if max_zero_streak [a,b,c,d,e,f,g,h] > 1 then List_explode (List_replace1 (replicate (max_zero_streak [a,b,c,d,e,f,g,h]) 0) [] (goup_by_zeros [a,b,c,d,e,f,g,h])) else map Some [a,b,c,d,e,f,g,h] )" - by(simp add: ipv6_preferred_to_compressed.simps max_zero_streak_def List_explode_goup_by_zeros) + apply (simp only: ipv6_preferred_to_compressed.simps Let_def max_zero_streak_def List_explode_goup_by_zeros) + using List_explode_goup_by_zeros by presburger + private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0 0 0 0 0 0 0 0) = [None]" by eval private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0x2001 0xDB8 0 0 8 0x800 0x200C 0x417A) = [Some 0x2001, Some 0xDB8, None, Some 8, Some 0x800, Some 0x200C, Some 0x417A]" by eval private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0x2001 0xDB8 0 3 8 0x800 0x200C 0x417A) = [Some 0x2001, Some 0xDB8, Some 0, Some 3, Some 8, Some 0x800, Some 0x200C, Some 0x417A]" by eval (*the output should even conform to RFC5952, ...*) lemma ipv6_preferred_to_compressed_RFC_4291_format: "ipv6_preferred_to_compressed ip = as \ length (filter (\p. p = None) as) = 0 \ length as = 8 \ length (filter (\p. p = None) as) = 1 \ length (filter (\p. p \ None) as) \ 7" apply(cases ip) apply(simp add: ipv6_preferred_to_compressed_pull_out_if) apply(simp only: split: if_split_asm) subgoal for a b c d e f g h apply(rule disjI2) apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(auto simp add: max_zero_streak_def) (*1min*) subgoal apply(rule disjI1) apply(simp) by force done \ \Idea for the following proof:\ private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = None#xs \ xs = map Some (dropWhile (\x. x=0) [a,b,c,d,e,f,g,h])" apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*20s*) lemma ipv6_preferred_to_compressed: assumes "ipv6_unparsed_compressed_to_preferred (ipv6_preferred_to_compressed ip) = Some ip'" shows "ip = ip'" proof - from assms have 1: "\ipv6compressed. parse_ipv6_address_compressed (ipv6_preferred_to_compressed ip) = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ip'" using ipv6_unparsed_compressed_to_preferred_identity2 by simp obtain a b c d e f g h where ip: "ip = IPv6AddrPreferred a b c d e f g h" by(cases ip) have ipv6_preferred_to_compressed_None1: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = None#xs \ (map Some (dropWhile (\x. x=0) [a,b,c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a b c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None2: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#None#xs \ (map Some (dropWhile (\x. x=0) [b,c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None3: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#None#xs \ (map Some (dropWhile (\x. x=0) [c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None4: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#None#xs \ (map Some (dropWhile (\x. x=0) [d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None5: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#None#xs \ (map Some (dropWhile (\x. x=0) [e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None6: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#None#xs \ (map Some (dropWhile (\x. x=0) [f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None7: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#(Some f')#None#xs \ (map Some (dropWhile (\x. x=0) [g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f' g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' f' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None8: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#(Some f')#(Some g')#None#xs \ (map Some (dropWhile (\x. x=0) [h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f' g' h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' f' g' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have 2: "parse_ipv6_address_compressed (ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h)) = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ip' \ IPv6AddrPreferred a b c d e f g h = ip'" for ipv6compressed apply(erule parse_ipv6_address_compressed_someE) apply(simp_all) apply(erule ipv6_preferred_to_compressed_None1, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None2, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None3, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None4, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None5, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None6, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None7, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None8, simp split: if_split_asm) done from 1 2 ip show ?thesis by(elim exE conjE, simp) qed end end diff --git a/thys/Inductive_Inference/Universal.thy b/thys/Inductive_Inference/Universal.thy --- a/thys/Inductive_Inference/Universal.thy +++ b/thys/Inductive_Inference/Universal.thy @@ -1,2538 +1,2539 @@ section \A universal partial recursive function\ theory Universal imports Partial_Recursive begin text \The main product of this section is a universal partial recursive function, which given a code $i$ of an $n$-ary partial recursive function $f$ and an encoded list @{term xs} of $n$ arguments, computes @{term "eval f xs"}. From this we can derive fixed-arity universal functions satisfying the usual results such as the $s$-$m$-$n$ theorem. To represent the code $i$, we need a way to encode @{typ recf}s as natural numbers (Section~\ref{s:recf_enc}). To construct the universal function, we devise a ternary function taking $i$, $xs$, and a step bound $t$ and simulating the execution of $f$ on input $xs$ for $t$ steps. This function is useful in its own right, enabling techniques like dovetailing or ``concurrent'' evaluation of partial recursive functions. The notion of a ``step'' is not part of the definition of (the evaluation of) partial recursive functions, but one can simulate the evaluation on an abstract machine (Section~\ref{s:step}). This machine's configurations can be encoded as natural numbers, and this leads us to a step function @{typ "nat \ nat"} on encoded configurations (Section~\ref{s:step_enc}). This function in turn can be computed by a primitive recursive function, from which we develop the aforementioned ternary function of $i$, @{term xs}, and $t$ (Section~\ref{s:step_recf}). From this we can finally derive a universal function (Section~\ref{s:the_universal}).\ subsection \A step function\label{s:step}\ text \We simulate the stepwise execution of a partial recursive function in a fairly straightforward way reminiscent of the execution of function calls in an imperative programming language. A configuration of the abstract machine is a pair consisting of: \begin{enumerate} \item A stack of frames. A frame represents the execution of a function and is a triple @{term "(f, xs, locals)"} of \begin{enumerate} \item a @{typ recf} @{term f} being executed, \item a @{typ "nat list"} of arguments of @{term f}, \item a @{typ "nat list"} of local variables, which holds intermediate values when @{term f} is of the form @{term Cn}, @{term Pr}, or @{term Mn}. \end{enumerate} \item A register of type @{typ "nat option"} representing the return value of the last function call: @{term None} signals that in the previous step the stack was not popped and hence no value was returned, whereas @{term "Some v"} means that in the previous step a function returned @{term v}. \end{enumerate} For computing @{term h} on input @{term xs}, the initial configuration is @{term "([(h, xs, [])], None)"}. When the computation for a frame ends, it is popped off the stack, and its return value is put in the register. The entire computation ends when the stack is empty. In such a final configuration the register contains the value of @{term h} at @{term xs}. If no final configuration is ever reached, @{term h} diverges at @{term xs}. The execution of one step depends on the topmost (that is, active) frame. In the step when a frame @{term "(h, xs, locals)"} is pushed onto the stack, the local variables are @{term "locals = []"}. The following happens until the frame is popped off the stack again (if it ever is): \begin{itemize} \item For the base functions @{term "h = Z"}, @{term "h = S"}, @{term[names_short] "h = Id m n"}, the frame is popped off the stack right away, and the return value is placed in the register. \item For @{term "h = Cn n f gs"}, for each function $g$ in @{term gs}: \begin{enumerate} \item A new frame of the form @{term "(g, xs, [])"} is pushed onto the stack. \item When (and if) this frame is eventually popped, the value in the register is @{term "eval g xs"}. This value is appended to the list @{term locals} of local variables. \end{enumerate} When all $g$ in $gs$ have been evaluated in this manner, $f$ is evaluated on the local variables by pushing @{term "(f, locals, [])"}. The resulting register value is kept and the active frame for $h$ is popped off the stack. \item For @{text "h = Pr n f g"}, let @{term "xs = y # ys"}. First @{term "(f, ys, [])"} is pushed and the return value stored in the @{term locals}. Then @{term "(g, x # v # ys, [])"} is pushed, where $x$ is the length of @{term locals} and $v$ the most recently appended value. The return value is appended to @{term locals}. This is repeated until the length of @{term locals} reaches @{term y}. Then the most recently appended local is placed in the register, and the stack is popped. \item For @{text "h = Mn n f"}, frames @{term "(f, x # xs, [])"} are pushed for $x = 0, 1, 2, \ldots$ until one of them returns $0$. Then this $x$ is placed in the register and the stack is popped. Until then $x$ is stored in @{term locals}. If none of these evaluations return $0$, the stack never shrinks, and thus the machine never reaches a final state. \end{itemize}\ type_synonym frame = "recf \ nat list \ nat list" type_synonym configuration = "frame list \ nat option" subsubsection \Definition of the step function\ fun step :: "configuration \ configuration" where "step ([], rv) = ([], rv)" | "step (((Z, _, _) # fs), rv) = (fs, Some 0)" | "step (((S, xs, _) # fs), rv) = (fs, Some (Suc (hd xs)))" | "step (((Id m n, xs, _) # fs), rv) = (fs, Some (xs ! n))" | "step (((Cn n f gs, xs, ls) # fs), rv) = (if length ls = length gs then if rv = None then ((f, ls, []) # (Cn n f gs, xs, ls) # fs, None) else (fs, rv) else if rv = None then if length ls < length gs then ((gs ! (length ls), xs, []) # (Cn n f gs, xs, ls) # fs, None) else (fs, rv) \\cannot occur, so don't-care term\ else ((Cn n f gs, xs, ls @ [the rv]) # fs, None))" | "step (((Pr n f g, xs, ls) # fs), rv) = (if ls = [] then if rv = None then ((f, tl xs, []) # (Pr n f g, xs, ls) # fs, None) else ((Pr n f g, xs, [the rv]) # fs, None) else if length ls = Suc (hd xs) then (fs, Some (hd ls)) else if rv = None then ((g, (length ls - 1) # hd ls # tl xs, []) # (Pr n f g, xs, ls) # fs, None) else ((Pr n f g, xs, (the rv) # ls) # fs, None))" | "step (((Mn n f, xs, ls) # fs), rv) = (if ls = [] then ((f, 0 # xs, []) # (Mn n f, xs, [0]) # fs, None) else if rv = Some 0 then (fs, Some (hd ls)) else ((f, (Suc (hd ls)) # xs, []) # (Mn n f, xs, [Suc (hd ls)]) # fs, None))" definition reachable :: "configuration \ configuration \ bool" where "reachable x y \ \t. iterate t step x = y" lemma step_reachable [intro]: assumes "step x = y" shows "reachable x y" unfolding reachable_def using assms by (metis iterate.simps(1,2) comp_id) lemma reachable_transitive [trans]: assumes "reachable x y" and "reachable y z" shows "reachable x z" using assms iterate_additive[where ?f=step] reachable_def by metis lemma reachable_refl: "reachable x x" unfolding reachable_def by (metis iterate.simps(1) eq_id_iff) text \From a final configuration, that is, when the stack is empty, only final configurations are reachable.\ lemma step_empty_stack: assumes "fst x = []" shows "fst (step x) = []" using assms by (metis prod.collapse step.simps(1)) lemma reachable_empty_stack: assumes "fst x = []" and "reachable x y" shows "fst y = []" proof - have "fst (iterate t step x) = []" for t using assms step_empty_stack by (induction t) simp_all then show ?thesis using reachable_def assms(2) by auto qed abbreviation nonterminating :: "configuration \ bool" where "nonterminating x \ \t. fst (iterate t step x) \ []" lemma reachable_nonterminating: assumes "reachable x y" and "nonterminating y" shows "nonterminating x" proof - from assms(1) obtain t\<^sub>1 where t1: "iterate t\<^sub>1 step x = y" using reachable_def by auto have "fst (iterate t step x) \ []" for t proof (cases "t \ t\<^sub>1") case True then show ?thesis using t1 assms(2) reachable_def reachable_empty_stack iterate_additive' by (metis le_Suc_ex) next case False then have "iterate t step x = iterate (t\<^sub>1 + (t - t\<^sub>1)) step x" by simp then have "iterate t step x = iterate (t - t\<^sub>1) step (iterate t\<^sub>1 step x)" by (simp add: iterate_additive') then have "iterate t step x = iterate (t - t\<^sub>1) step y" using t1 by simp then show "fst (iterate t step x) \ []" using assms(2) by simp qed then show ?thesis .. qed text \The function @{term step} is underdefined, for example, when the top frame contains a non-well-formed @{typ recf} or too few arguments. All is well, though, if every frame contains a well-formed @{typ recf} whose arity matches the number of arguments. Such stacks will be called \emph{valid}.\ definition valid :: "frame list \ bool" where "valid stack \ \s\set stack. recfn (length (fst (snd s))) (fst s)" lemma valid_frame: "valid (s # ss) \ valid ss \ recfn (length (fst (snd s))) (fst s)" using valid_def by simp lemma valid_ConsE: "valid ((f, xs, locs) # rest) \ valid rest \ recfn (length xs) f" using valid_def by simp lemma valid_ConsI: "valid rest \ recfn (length xs) f \ valid ((f, xs, locs) # rest)" using valid_def by simp text \Stacks in initial configurations are valid, and performing a step maintains the validity of the stack.\ lemma step_valid: "valid stack \ valid (fst (step (stack, rv)))" proof (cases stack) case Nil then show ?thesis using valid_def by simp next case (Cons s ss) assume valid: "valid stack" then have *: "valid ss \ recfn (length (fst (snd s))) (fst s)" using valid_frame Cons by simp show ?thesis proof (cases "fst s") case Z then show ?thesis using Cons valid * by (metis fstI prod.collapse step.simps(2)) next case S then show ?thesis using Cons valid * by (metis fst_conv prod.collapse step.simps(3)) next case Id then show ?thesis using Cons valid * by (metis fstI prod.collapse step.simps(4)) next case (Cn n f gs) then obtain xs ls where "s = (Cn n f gs, xs, ls)" using Cons by (metis prod.collapse) moreover consider "length ls = length gs \ rv \" | "length ls = length gs \ rv \" | "length ls < length gs \ rv \" | "length ls \ length gs \ rv \" | "length ls > length gs \ rv \" by linarith ultimately show ?thesis using valid Cons valid_def by (cases) auto next case (Pr n f g) then obtain xs ls where s: "s = (Pr n f g, xs, ls)" using Cons by (metis prod.collapse) consider "length ls = 0 \ rv \" | "length ls = 0 \ rv \" | "length ls \ 0 \ length ls = Suc (hd xs)" | "length ls \ 0 \ length ls \ Suc (hd xs) \ rv \" | "length ls \ 0 \ length ls \ Suc (hd xs) \ rv \" by linarith then show ?thesis using Cons * valid_def s by (cases) auto next case (Mn n f) then obtain xs ls where s: "s = (Mn n f, xs, ls)" using Cons by (metis prod.collapse) consider "length ls = 0" | "length ls \ 0 \ rv \" | "length ls \ 0 \ rv \" by linarith then show ?thesis using Cons * valid_def s by (cases) auto qed qed corollary iterate_step_valid: assumes "valid stack" shows "valid (fst (iterate t step (stack, rv)))" using assms proof (induction t) case 0 then show ?case by simp next case (Suc t) moreover have "iterate (Suc t) step (stack, rv) = step (iterate t step (stack, rv))" by simp ultimately show ?case using step_valid valid_def by (metis prod.collapse) qed subsubsection \Correctness of the step function\ text \The function @{term step} works correctly for a @{typ recf} $f$ on arguments @{term xs} in some configuration if (1) in case $f$ converges, @{term step} reaches a configuration with the topmost frame popped and @{term "eval f xs"} in the register, and (2) in case $f$ diverges, @{term step} does not reach a final configuration.\ fun correct :: "configuration \ bool" where "correct ([], r) = True" | "correct ((f, xs, ls) # rest, r) = (if eval f xs \ then reachable ((f, xs, ls) # rest, r) (rest, eval f xs) else nonterminating ((f, xs, ls) # rest, None))" lemma correct_convergI: assumes "eval f xs \" and "reachable ((f, xs, ls) # rest, None) (rest, eval f xs)" shows "correct ((f, xs, ls) # rest, None)" using assms by auto lemma correct_convergE: assumes "correct ((f, xs, ls) # rest, None)" and "eval f xs \" shows "reachable ((f, xs, ls) # rest, None) (rest, eval f xs)" using assms by simp text \The correctness proof for @{term step} is by structural induction on the @{typ recf} in the top frame. The base cases @{term Z}, @{term S}, and @{term[names_short] Id} are simple. For @{text "X = Cn, Pr, Mn"}, the lemmas named @{text reachable_X} show which configurations are reachable for @{typ recf}s of shape @{text X}. Building on those, the lemmas named @{text step_X_correct} show @{term step}'s correctness for @{text X}.\ lemma reachable_Cn: assumes "valid (((Cn n f gs), xs, []) # rest)" (is "valid ?stack") and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" and "\g xs rest. g \ set gs \ valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" and "\i" and "k \ length gs" shows "reachable (?stack, None) ((Cn n f gs, xs, take k (map (\g. the (eval g xs)) gs)) # rest, None)" using assms(4,5) proof (induction k) case 0 then show ?case using reachable_refl by simp next case (Suc k) let ?ys = "map (\g. the (eval g xs)) gs" from Suc have "k < length gs" by simp have valid: "recfn (length xs) (Cn n f gs)" "valid rest" using assms(1) valid_ConsE[of "(Cn n f gs)"] by simp_all from Suc have "reachable (?stack, None) ((Cn n f gs, xs, take k ?ys) # rest, None)" (is "_ (?stack1, None)") by simp also have "reachable ... ((gs ! k, xs, []) # ?stack1, None)" using step_reachable \k < length gs\ by (auto simp: min_absorb2) also have "reachable ... (?stack1, eval (gs ! k) xs)" (is "_ (_, ?rv)") using Suc.prems(1) \k < length gs\ assms(3) valid valid_ConsI by auto also have "reachable ... ((Cn n f gs, xs, (take (Suc k) ?ys)) # rest, None)" (is "_ (?stack2, None)") proof - have "step (?stack1, ?rv) = ((Cn n f gs, xs, (take k ?ys) @ [the ?rv]) # rest, None)" using Suc by auto also have "... = ((Cn n f gs, xs, (take (Suc k) ?ys)) # rest, None)" by (simp add: \k < length gs\ take_Suc_conv_app_nth) finally show ?thesis using step_reachable by auto qed finally show "reachable (?stack, None) (?stack2, None)" . qed lemma step_Cn_correct: assumes "valid (((Cn n f gs), xs, []) # rest)" (is "valid ?stack") and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" and "\g xs rest. g \ set gs \ valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" shows "correct (?stack, None)" proof - have valid: "recfn (length xs) (Cn n f gs)" "valid rest" using valid_ConsE[OF assms(1)] by auto let ?ys = "map (\g. the (eval g xs)) gs" consider (diverg_f) "\g\set gs. eval g xs \" and "eval f ?ys \" | (diverg_gs) "\g\set gs. eval g xs \" | (converg) "eval (Cn n f gs) xs \" using valid_ConsE[OF assms(1)] by fastforce then show ?thesis proof (cases) case diverg_f then have "\i" by simp then have "reachable (?stack, None) ((Cn n f gs, xs, ?ys) # rest, None)" (is "_ (?stack1, None)") using reachable_Cn[OF assms, where ?k="length gs"] by simp also have "reachable ... ((f, ?ys, []) # ?stack1, None)" (is "_ (?stack2, None)") by (simp add: step_reachable) finally have "reachable (?stack, None) (?stack2, None)" . moreover have "nonterminating (?stack2, None)" using diverg_f(2) assms(2)[of ?ys ?stack1] valid_ConsE[OF assms(1)] valid_ConsI by auto ultimately have "nonterminating (?stack, None)" using reachable_nonterminating by simp moreover have "eval (Cn n f gs) xs \" using diverg_f(2) assms(1) eval_Cn valid_ConsE by presburger ultimately show ?thesis by simp next case diverg_gs then have ex_i: "\i" using in_set_conv_nth[of _ gs] by auto define k where "k = (LEAST i. i < length gs \ eval (gs ! i) xs \)" (is "_ = Least ?P") then have gs_k: "eval (gs ! k) xs \" using LeastI_ex[OF ex_i] by simp have "\i" using k_def not_less_Least[of _ ?P] LeastI_ex[OF ex_i] by simp moreover from this have "k < length gs" using ex_i less_le_trans not_le by blast ultimately have "reachable (?stack, None) ((Cn n f gs, xs, take k ?ys) # rest, None)" using reachable_Cn[OF assms] by simp also have "reachable ... ((gs ! (length (take k ?ys)), xs, []) # (Cn n f gs, xs, take k ?ys) # rest, None)" (is "_ (?stack1, None)") proof - have "length (take k ?ys) < length gs" by (simp add: \k < length gs\ less_imp_le_nat min_less_iff_disj) - then show ?thesis using step_reachable by auto + then show ?thesis using step_reachable \k < length gs\ + by auto qed finally have "reachable (?stack, None) (?stack1, None)" . moreover have "nonterminating (?stack1, None)" proof - have "recfn (length xs) (gs ! k)" using \k < length gs\ valid(1) by simp then have "correct (?stack1, None)" using \k < length gs\ nth_mem valid valid_ConsI assms(3)[of "gs ! (length (take k ?ys))" xs] by auto moreover have "length (take k ?ys) = k" by (simp add: \k < length gs\ less_imp_le_nat min_absorb2) ultimately show ?thesis using gs_k by simp qed ultimately have "nonterminating (?stack, None)" using reachable_nonterminating by simp moreover have "eval (Cn n f gs) xs \" using diverg_gs valid by fastforce ultimately show ?thesis by simp next case converg then have f: "eval f ?ys \" and g: "\g. g \ set gs \ eval g xs \" using valid(1) by (metis eval_Cn)+ then have "\i" by simp then have "reachable (?stack, None) ((Cn n f gs, xs, take (length gs) ?ys) # rest, None)" using reachable_Cn assms by blast also have "reachable ... ((Cn n f gs, xs, ?ys) # rest, None)" (is "_ (?stack1, None)") by (simp add: reachable_refl) also have "reachable ... ((f, ?ys, []) # ?stack1, None)" using step_reachable by auto also have "reachable ... (?stack1, eval f ?ys)" using assms(2)[of "?ys"] correct_convergE valid f valid_ConsI by auto also have "reachable (?stack1, eval f ?ys) (rest, eval f ?ys)" using f by auto finally have "reachable (?stack, None) (rest, eval f ?ys)" . moreover have "eval (Cn n f gs) xs = eval f ?ys" using g valid(1) by auto ultimately show ?thesis using converg correct_convergI by auto qed qed text \During the execution of a frame with a partial recursive function of shape @{term "Pr n f g"} and arguments @{term "x # xs"}, the list of local variables collects all the function values up to @{term x} in reversed order. We call such a list a @{term trace} for short.\ definition trace :: "nat \ recf \ recf \ nat list \ nat \ nat list" where "trace n f g xs x \ map (\y. the (eval (Pr n f g) (y # xs))) (rev [0..xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" and "\xs rest. valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" and "y \ x" and "eval (Pr n f g) (y # xs) \" shows "reachable (?stack, None) ((Pr n f g, x # xs, trace n f g xs y) # rest, None)" using assms(4,5) proof (induction y) case 0 have valid: "recfn (length (x # xs)) (Pr n f g)" "valid rest" using valid_ConsE[OF assms(1)] by simp_all then have f: "eval f xs \" using 0 by simp let ?as = "x # xs" have "reachable (?stack, None) ((f, xs, []) # ((Pr n f g), ?as, []) # rest, None)" using step_reachable by auto also have "reachable ... (?stack, eval f xs)" using assms(2)[of xs "((Pr n f g), ?as, []) # rest"] correct_convergE[OF _ f] f valid valid_ConsI by simp also have "reachable ... ((Pr n f g, ?as, [the (eval f xs)]) # rest, None)" using step_reachable valid(1) f by auto finally have "reachable (?stack, None) ((Pr n f g, ?as, [the (eval f xs)]) # rest, None)" . then show ?case using trace_def valid(1) by simp next case (Suc y) have valid: "recfn (length (x # xs)) (Pr n f g)" "valid rest" using valid_ConsE[OF assms(1)] by simp_all let ?ls = "trace n f g xs y" have lenls: "length ?ls = Suc y" using trace_length by auto moreover have hdls: "hd ?ls = the (eval (Pr n f g) (y # xs))" using Suc trace_hd by auto ultimately have g: "eval g (y # hd ?ls # xs) \" "eval (Pr n f g) (Suc y # xs) = eval g (y # hd ?ls # xs)" using eval_Pr_Suc_converg hdls valid(1) Suc by simp_all then have "reachable (?stack, None) ((Pr n f g, x # xs, ?ls) # rest, None)" (is "_ (?stack1, None)") using Suc valid(1) by fastforce also have "reachable ... ((g, y # hd ?ls # xs, []) # (Pr n f g, x # xs, ?ls) # rest, None)" using Suc.prems lenls by fastforce also have "reachable ... (?stack1, eval g (y # hd ?ls # xs))" (is "_ (_, ?rv)") using assms(3) g(1) valid valid_ConsI by auto also have "reachable ... ((Pr n f g, x # xs, (the ?rv) # ?ls) # rest, None)" using Suc.prems(1) g(1) lenls by auto finally have "reachable (?stack, None) ((Pr n f g, x # xs, (the ?rv) # ?ls) # rest, None)" . moreover have "trace n f g xs (Suc y) = (the ?rv) # ?ls" using g(2) trace_Suc by simp ultimately show ?case by simp qed lemma step_Pr_correct: assumes "valid (((Pr n f g), xs, []) # rest)" (is "valid ?stack") and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" and "\xs rest. valid ((g, xs, []) # rest) \ correct ((g, xs, []) # rest, None)" shows "correct (?stack, None)" proof - have valid: "valid rest" "recfn (length xs) (Pr n f g)" using valid_ConsE[OF assms(1)] by simp_all then have "length xs > 0" by auto then obtain y ys where y_ys: "xs = y # ys" using list.exhaust_sel by auto let ?t = "trace n f g ys" consider (converg) "eval (Pr n f g) xs \" | (diverg_f) "eval (Pr n f g) xs \" and "eval f ys \" | (diverg) "eval (Pr n f g) xs \" and "eval f ys \" by auto then show ?thesis proof (cases) case converg then have "\z. z \ y \ reachable (?stack, None) (((Pr n f g), xs, ?t z) # rest, None)" using assms valid by (simp add: eval_Pr_converg_le reachable_Pr y_ys) then have "reachable (?stack, None) (((Pr n f g), xs, ?t y) # rest, None)" by simp moreover have "reachable (((Pr n f g), xs, ?t y) # rest, None) (rest, Some (hd (?t y)))" using trace_length step_reachable y_ys by fastforce ultimately have "reachable (?stack, None) (rest, Some (hd (?t y)))" using reachable_transitive by blast then show ?thesis using assms(1) trace_hd converg y_ys by simp next case diverg_f have *: "step (?stack, None) = ((f, ys, []) # ((Pr n f g), xs, []) # tl ?stack, None)" (is "_ = (?stack1, None)") using assms(1,2) y_ys by simp then have "reachable (?stack, None) (?stack1, None)" using step_reachable by force moreover have "nonterminating (?stack1, None)" using assms diverg_f valid valid_ConsI * by auto ultimately have "nonterminating (?stack, None)" using reachable_nonterminating by blast then show ?thesis using diverg_f(1) assms(1) by simp next case diverg let ?h = "\z. the (eval (Pr n f g) (z # ys))" let ?Q = "\z. z < y \ eval (Pr n f g) (z # ys) \" have "?Q 0" using assms diverg neq0_conv y_ys valid by fastforce define zmax where "zmax = Greatest ?Q" then have "?Q zmax" using \?Q 0\ GreatestI_nat[of ?Q 0 y] by simp have le_zmax: "\z. ?Q z \ z \ zmax" using Greatest_le_nat[of ?Q _ y] zmax_def by simp have len: "length (?t zmax) < Suc y" by (simp add: \?Q zmax\ trace_length) have "eval (Pr n f g) (y # ys) \" if "y \ zmax" for y using that zmax_def \?Q zmax\ assms eval_Pr_converg_le[of n f g ys zmax y] valid y_ys by simp then have "reachable (?stack, None) (((Pr n f g), xs, ?t y) # rest, None)" if "y \ zmax" for y using that \?Q zmax\ diverg y_ys assms reachable_Pr by simp then have "reachable (?stack, None) (((Pr n f g), xs, ?t zmax) # rest, None)" (is "reachable _ (?stack1, None)") by simp also have "reachable ... ((g, zmax # ?h zmax # tl xs, []) # (Pr n f g, xs, ?t zmax) # rest, None)" (is "_ (?stack2, None)") proof (rule step_reachable) have "length (?t zmax) \ Suc (hd xs)" using len y_ys by simp moreover have "hd (?t zmax) = ?h zmax" using trace_hd by auto moreover have "length (?t zmax) = Suc zmax" using trace_length by simp ultimately show "step (?stack1, None) = (?stack2, None)" by auto qed finally have "reachable (?stack, None) (?stack2, None)" . moreover have "nonterminating (?stack2, None)" proof - have "correct (?stack2, None)" using y_ys assms valid_ConsI valid by simp moreover have "eval g (zmax # ?h zmax # ys) \" using \?Q zmax\ diverg le_zmax len less_Suc_eq trace_length y_ys valid by fastforce ultimately show ?thesis using y_ys by simp qed ultimately have "nonterminating (?stack, None)" using reachable_nonterminating by simp then show ?thesis using diverg assms(1) by simp qed qed lemma reachable_Mn: assumes "valid ((Mn n f, xs, []) # rest)" (is "valid ?stack") and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" and "\y {None, Some 0}" shows "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" using assms(3) proof (induction z) case 0 then have "step (?stack, None) = ((f, 0 # xs, []) # (Mn n f, xs, [0]) # rest, None)" using assms by simp then show ?case using step_reachable assms(1) by force next case (Suc z) have valid: "valid rest" "recfn (length xs) (Mn n f)" using valid_ConsE[OF assms(1)] by auto have f: "eval f (z # xs) \ {None, Some 0}" using Suc by simp have "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" using Suc by simp also have "reachable ... ((Mn n f, xs, [z]) # rest, eval f (z # xs))" using f assms(2)[of "z # xs"] valid correct_convergE valid_ConsI by auto also have "reachable ... ((f, (Suc z) # xs, []) # (Mn n f, xs, [Suc z]) # rest, None)" (is "_ (?stack1, None)") using step_reachable f by force finally have "reachable (?stack, None) (?stack1, None)" . then show ?case by simp qed lemma iterate_step_empty_stack: "iterate t step ([], rv) = ([], rv)" using step_empty_stack by (induction t) simp_all lemma reachable_iterate_step_empty_stack: assumes "reachable cfg ([], rv)" shows "\t. iterate t step cfg = ([], rv) \ (\t' [])" proof - let ?P = "\t. iterate t step cfg = ([], rv)" from assms have "\t. ?P t" by (simp add: reachable_def) moreover define tmin where "tmin = Least ?P" ultimately have "?P tmin" using LeastI_ex[of ?P] by simp have "fst (iterate t' step cfg) \ []" if "t' < tmin" for t' proof assume "fst (iterate t' step cfg) = []" then obtain v where v: "iterate t' step cfg = ([], v)" by (metis prod.exhaust_sel) then have "iterate t'' step ([], v) = ([], v)" for t'' using iterate_step_empty_stack by simp then have "iterate (t' + t'') step cfg = ([], v)" for t'' using v iterate_additive by fast moreover obtain t'' where "t' + t'' = tmin" using \t' < tmin\ less_imp_add_positive by auto ultimately have "iterate tmin step cfg = ([], v)" by auto then have "v = rv" using \?P tmin\ by simp then have "iterate t' step cfg = ([], rv)" using v by simp moreover have "\t' ?P t'" unfolding tmin_def using not_less_Least[of _ ?P] by simp ultimately show False using that by simp qed then show ?thesis using \?P tmin\ by auto qed lemma step_Mn_correct: assumes "valid ((Mn n f, xs, []) # rest)" (is "valid ?stack") and "\xs rest. valid ((f, xs, []) # rest) \ correct ((f, xs, []) # rest, None)" shows "correct (?stack, None)" proof - have valid: "valid rest" "recfn (length xs) (Mn n f)" using valid_ConsE[OF assms(1)] by auto consider (diverg) "eval (Mn n f) xs \" and "\z. eval f (z # xs) \" | (diverg_f) "eval (Mn n f) xs \" and "\z. eval f (z # xs) \" | (converg) "eval (Mn n f) xs \" by fast then show ?thesis proof (cases) case diverg then have "\z. eval f (z # xs) \ Some 0" using eval_Mn_diverg[OF valid(2)] by simp then have "\y {None, Some 0}" for z using diverg by simp then have reach_z: "\z. reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" using reachable_Mn[OF assms] diverg by simp define h :: "nat \ configuration" where "h z \ ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" for z then have h_inj: "\x y. x \ y \ h x \ h y" and z_neq_Nil: "\z. fst (h z) \ []" by simp_all have z: "\z\<^sub>0. \z>z\<^sub>0. \ (\t'\t. iterate t' step (?stack, None) = h z)" for t proof (induction t) case 0 then show ?case by (metis h_inj le_zero_eq less_not_refl3) next case (Suc t) then show ?case using h_inj by (metis (no_types, hide_lams) le_Suc_eq less_not_refl3 less_trans) qed have "nonterminating (?stack, None)" proof (rule ccontr) assume "\ nonterminating (?stack, None)" then obtain t where t: "fst (iterate t step (?stack, None)) = []" by auto then obtain z\<^sub>0 where "\z>z\<^sub>0. \ (\t'\t. iterate t' step (?stack, None) = h z)" using z by auto then have not_h: "\t'\t. iterate t' step (?stack, None) \ h (Suc z\<^sub>0)" by simp have "\t'\t. fst (iterate t' step (?stack, None)) = []" using t iterate_step_empty_stack iterate_additive'[of t] by (metis le_Suc_ex prod.exhaust_sel) then have "\t'\t. iterate t' step (?stack, None) \ h (Suc z\<^sub>0)" using z_neq_Nil by auto then have "\t'. iterate t' step (?stack, None) \ h (Suc z\<^sub>0)" using not_h nat_le_linear by auto then have "\ reachable (?stack, None) (h (Suc z\<^sub>0))" using reachable_def by simp then show False using reach_z[of "Suc z\<^sub>0"] h_def by simp qed then show ?thesis using diverg by simp next case diverg_f let ?P = "\z. eval f (z # xs) \" define zmin where "zmin \ Least ?P" then have "\y {None, Some 0}" using diverg_f eval_Mn_diverg[OF valid(2)] less_trans not_less_Least[of _ ?P] by blast moreover have f_zmin: "eval f (zmin # xs) \" using diverg_f LeastI_ex[of ?P] zmin_def by simp ultimately have "reachable (?stack, None) ((f, zmin # xs, []) # (Mn n f, xs, [zmin]) # rest, None)" (is "reachable _ (?stack1, None)") using reachable_Mn[OF assms] by simp moreover have "nonterminating (?stack1, None)" using f_zmin assms valid diverg_f valid_ConsI by auto ultimately have "nonterminating (?stack, None)" using reachable_nonterminating by simp then show ?thesis using diverg_f by simp next case converg then obtain z where z: "eval (Mn n f) xs \= z" by auto have f_z: "eval f (z # xs) \= 0" and f_less_z: "\y. y < z \ eval f (y # xs) \\ 0" using eval_Mn_convergE(2,3)[OF valid(2) z] by simp_all then have "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" using reachable_Mn[OF assms] by simp also have "reachable ... ((Mn n f, xs, [z]) # rest, eval f (z # xs))" using assms(2)[of "z # xs"] valid f_z valid_ConsI correct_convergE by auto also have "reachable ... (rest, Some z)" using f_z f_less_z step_reachable by auto finally have "reachable (?stack, None) (rest, Some z)" . then show ?thesis using z by simp qed qed theorem step_correct: assumes "valid ((f, xs, []) # rest)" shows "correct ((f, xs, []) # rest, None)" using assms proof (induction f arbitrary: xs rest) case Z then show ?case using valid_ConsE[of Z] step_reachable by auto next case S then show ?case using valid_ConsE[of S] step_reachable by auto next case (Id m n) then show ?case using valid_ConsE[of "Id m n"] by auto next case Cn then show ?case using step_Cn_correct by presburger next case Pr then show ?case using step_Pr_correct by simp next case Mn then show ?case using step_Mn_correct by presburger qed subsection \Encoding partial recursive functions\label{s:recf_enc}\ text \In this section we define an injective, but not surjective, mapping from @{typ recf}s to natural numbers.\ abbreviation triple_encode :: "nat \ nat \ nat \ nat" where "triple_encode x y z \ prod_encode (x, prod_encode (y, z))" abbreviation quad_encode :: "nat \ nat \ nat \ nat \ nat" where "quad_encode w x y z \ prod_encode (w, prod_encode (x, prod_encode (y, z)))" fun encode :: "recf \ nat" where "encode Z = 0" | "encode S = 1" | "encode (Id m n) = triple_encode 2 m n" | "encode (Cn n f gs) = quad_encode 3 n (encode f) (list_encode (map encode gs))" | "encode (Pr n f g) = quad_encode 4 n (encode f) (encode g)" | "encode (Mn n f) = triple_encode 5 n (encode f)" lemma prod_encode_gr1: "a > 1 \ prod_encode (a, x) > 1" using le_prod_encode_1 less_le_trans by blast lemma encode_not_Z_or_S: "encode f = prod_encode (a, b) \ a > 1 \ f \ Z \ f \ S" by (metis encode.simps(1) encode.simps(2) less_numeral_extra(4) not_one_less_zero prod_encode_gr1) lemma encode_injective: "encode f = encode g \ f = g" proof (induction g arbitrary: f) case Z have "\a x. a > 1 \ prod_encode (a, x) > 0" using prod_encode_gr1 by (meson less_one less_trans) then have "f \ Z \ encode f > 0" by (cases f) auto then have "encode f = 0 \ f = Z" by fastforce then show ?case using Z by simp next case S have "\a x. a > 1 \ prod_encode (a, x) \ Suc 0" using prod_encode_gr1 by (metis One_nat_def less_numeral_extra(4)) then have "encode f = 1 \ f = S" by (cases f) auto then show ?case using S by simp next case Id then obtain z where *: "encode f = prod_encode (2, z)" by simp show ?case using Id by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq) next case Cn then obtain z where *: "encode f = prod_encode (3, z)" by simp show ?case proof (cases f) case Z then show ?thesis using * encode_not_Z_or_S by simp next case S then show ?thesis using * encode_not_Z_or_S by simp next case Id then show ?thesis using * by (simp add: prod_encode_eq) next case Cn then show ?thesis using * Cn.IH Cn.prems list_decode_encode by (smt encode.simps(4) fst_conv list.inj_map_strong prod_encode_eq snd_conv) next case Pr then show ?thesis using * by (simp add: prod_encode_eq) next case Mn then show ?thesis using * by (simp add: prod_encode_eq) qed next case Pr then obtain z where *: "encode f = prod_encode (4, z)" by simp show ?case using Pr by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq) next case Mn then obtain z where *: "encode f = prod_encode (5, z)" by simp show ?case using Mn by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq) qed definition encode_kind :: "nat \ nat" where "encode_kind e \ if e = 0 then 0 else if e = 1 then 1 else pdec1 e" lemma encode_kind_0: "encode_kind (encode Z) = 0" unfolding encode_kind_def by simp lemma encode_kind_1: "encode_kind (encode S) = 1" unfolding encode_kind_def by simp lemma encode_kind_2: "encode_kind (encode (Id m n)) = 2" unfolding encode_kind_def by (metis encode.simps(1-3) encode_injective fst_conv prod_encode_inverse recf.simps(16) recf.simps(8)) lemma encode_kind_3: "encode_kind (encode (Cn n f gs)) = 3" unfolding encode_kind_def by (metis encode.simps(1,2,4) encode_injective fst_conv prod_encode_inverse recf.simps(10) recf.simps(18)) lemma encode_kind_4: "encode_kind (encode (Pr n f g)) = 4" unfolding encode_kind_def by (metis encode.simps(1,2,5) encode_injective fst_conv prod_encode_inverse recf.simps(12) recf.simps(20)) lemma encode_kind_5: "encode_kind (encode (Mn n f)) = 5" unfolding encode_kind_def by (metis encode.simps(1,2,6) encode_injective fst_conv prod_encode_inverse recf.simps(14) recf.simps(22)) lemmas encode_kind_n = encode_kind_0 encode_kind_1 encode_kind_2 encode_kind_3 encode_kind_4 encode_kind_5 lemma encode_kind_Cn: assumes "encode_kind (encode f) = 3" shows "\n f' gs. f = Cn n f' gs" using assms encode_kind_n by (cases f) auto lemma encode_kind_Pr: assumes "encode_kind (encode f) = 4" shows "\n f' g. f = Pr n f' g" using assms encode_kind_n by (cases f) auto lemma encode_kind_Mn: assumes "encode_kind (encode f) = 5" shows "\n g. f = Mn n g" using assms encode_kind_n by (cases f) auto lemma pdec2_encode_Id: "pdec2 (encode (Id m n)) = prod_encode (m, n)" by simp lemma pdec2_encode_Pr: "pdec2 (encode (Pr n f g)) = triple_encode n (encode f) (encode g)" by simp subsection \The step function on encoded configurations\label{s:step_enc}\ text \In this section we construct a function @{text "estep :: nat \ nat"} that is equivalent to the function @{text "step :: configuration \ configuration"} except that it applies to encoded configurations. We start by defining an encoding for configurations.\ definition encode_frame :: "frame \ nat" where "encode_frame s \ triple_encode (encode (fst s)) (list_encode (fst (snd s))) (list_encode (snd (snd s)))" lemma encode_frame: "encode_frame (f, xs, ls) = triple_encode (encode f) (list_encode xs) (list_encode ls)" unfolding encode_frame_def by simp abbreviation encode_option :: "nat option \ nat" where "encode_option x \ if x = None then 0 else Suc (the x)" definition encode_config :: "configuration \ nat" where "encode_config cfg \ prod_encode (list_encode (map encode_frame (fst cfg)), encode_option (snd cfg))" lemma encode_config: "encode_config (ss, rv) = prod_encode (list_encode (map encode_frame ss), encode_option rv)" unfolding encode_config_def by simp text \Various projections from encoded configurations:\ definition e2stack where "e2stack e \ pdec1 e" definition e2rv where "e2rv e \ pdec2 e" definition e2tail where "e2tail e \ e_tl (e2stack e)" definition e2frame where "e2frame e \ e_hd (e2stack e)" definition e2i where "e2i e \ pdec1 (e2frame e)" definition e2xs where "e2xs e \ pdec12 (e2frame e)" definition e2ls where "e2ls e \ pdec22 (e2frame e)" definition e2lenas where "e2lenas e \ e_length (e2xs e)" definition e2lenls where "e2lenls e \ e_length (e2ls e)" lemma e2rv_rv [simp]: "e2rv (encode_config (ss, rv)) = (if rv \ then 0 else Suc (the rv))" unfolding e2rv_def using encode_config by simp lemma e2stack_stack [simp]: "e2stack (encode_config (ss, rv)) = list_encode (map encode_frame ss)" unfolding e2stack_def using encode_config by simp lemma e2tail_tail [simp]: "e2tail (encode_config (s # ss, rv)) = list_encode (map encode_frame ss)" unfolding e2tail_def using encode_config by fastforce lemma e2frame_frame [simp]: "e2frame (encode_config (s # ss, rv)) = encode_frame s" unfolding e2frame_def using encode_config by fastforce lemma e2i_f [simp]: "e2i (encode_config ((f, xs, ls) # ss, rv)) = encode f" unfolding e2i_def using encode_config e2frame_frame encode_frame by force lemma e2xs_xs [simp]: "e2xs (encode_config ((f, xs, ls) # ss, rv)) = list_encode xs" using e2xs_def e2frame_frame encode_frame by force lemma e2ls_ls [simp]: "e2ls (encode_config ((f, xs, ls) # ss, rv)) = list_encode ls" using e2ls_def e2frame_frame encode_frame by force lemma e2lenas_lenas [simp]: "e2lenas (encode_config ((f, xs, ls) # ss, rv)) = length xs" using e2lenas_def e2frame_frame encode_frame by simp lemma e2lenls_lenls [simp]: "e2lenls (encode_config ((f, xs, ls) # ss, rv)) = length ls" using e2lenls_def e2frame_frame encode_frame by simp lemma e2stack_0_iff_Nil: assumes "e = encode_config (ss, rv)" shows "e2stack e = 0 \ ss = []" using assms by (metis list_encode.simps(1) e2stack_stack list_encode_0 map_is_Nil_conv) lemma e2ls_0_iff_Nil [simp]: "list_decode (e2ls e) = [] \ e2ls e = 0" by (metis list_decode.simps(1) list_encode_decode) text \We now define @{text eterm} piecemeal by considering the more complicated cases @{text Cn}, @{text Pr}, and @{text Mn} separately.\ definition "estep_Cn e \ if e2lenls e = e_length (pdec222 (e2i e)) then if e2rv e = 0 then prod_encode (e_cons (triple_encode (pdec122 (e2i e)) (e2ls e) 0) (e2stack e), 0) else prod_encode (e2tail e, e2rv e) else if e2rv e = 0 then if e2lenls e < e_length (pdec222 (e2i e)) then prod_encode (e_cons (triple_encode (e_nth (pdec222 (e2i e)) (e2lenls e)) (e2xs e) 0) (e2stack e), 0) else prod_encode (e2tail e, e2rv e) else prod_encode (e_cons (triple_encode (e2i e) (e2xs e) (e_snoc (e2ls e) (e2rv e - 1))) (e2tail e), 0)" lemma estep_Cn: assumes "c = (((Cn n f gs, xs, ls) # fs), rv)" shows "estep_Cn (encode_config c) = encode_config (step c)" using encode_frame by (simp add: assms estep_Cn_def, simp add: encode_config assms) definition "estep_Pr e \ if e2ls e = 0 then if e2rv e = 0 then prod_encode (e_cons (triple_encode (pdec122 (e2i e)) (e_tl (e2xs e)) 0) (e2stack e), 0) else prod_encode (e_cons (triple_encode (e2i e) (e2xs e) (singleton_encode (e2rv e - 1))) (e2tail e), 0) else if e2lenls e = Suc (e_hd (e2xs e)) then prod_encode (e2tail e, Suc (e_hd (e2ls e))) else if e2rv e = 0 then prod_encode (e_cons (triple_encode (pdec222 (e2i e)) (e_cons (e2lenls e - 1) (e_cons (e_hd (e2ls e)) (e_tl (e2xs e)))) 0) (e2stack e), 0) else prod_encode (e_cons (triple_encode (e2i e) (e2xs e) (e_cons (e2rv e - 1) (e2ls e))) (e2tail e), 0)" lemma estep_Pr1: assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "ls \ []" and "length ls \ Suc (hd xs)" and "rv \ None" and "recfn (length xs) (Pr n f g)" shows "estep_Pr (encode_config c) = encode_config (step c)" proof - let ?e = "encode_config c" from assms(5) have "length xs > 0" by auto then have eq: "hd xs = e_hd (e2xs ?e)" using assms e_hd_def by auto have "step c = ((Pr n f g, xs, (the rv) # ls) # fs, None)" (is "step c = (?t # ?ss, None)") using assms by simp then have "encode_config (step c) = prod_encode (list_encode (map encode_frame (?t # ?ss)), 0)" using encode_config by simp also have "... = prod_encode (e_cons (encode_frame ?t) (list_encode (map encode_frame (?ss))), 0)" by simp also have "... = prod_encode (e_cons (encode_frame ?t) (e2tail ?e), 0)" using assms(1) by simp also have "... = prod_encode (e_cons (triple_encode (e2i ?e) (e2xs ?e) (e_cons (e2rv ?e - 1) (e2ls ?e))) (e2tail ?e), 0)" by (simp add: assms encode_frame) finally show ?thesis using assms eq estep_Pr_def by auto qed lemma estep_Pr2: assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "ls \ []" and "length ls \ Suc (hd xs)" and "rv = None" and "recfn (length xs) (Pr n f g)" shows "estep_Pr (encode_config c) = encode_config (step c)" proof - let ?e = "encode_config c" from assms(5) have "length xs > 0" by auto then have eq: "hd xs = e_hd (e2xs ?e)" using assms e_hd_def by auto have "step c = ((g, (length ls - 1) # hd ls # tl xs, []) # (Pr n f g, xs, ls) # fs, None)" (is "step c = (?t # ?ss, None)") using assms by simp then have "encode_config (step c) = prod_encode (list_encode (map encode_frame (?t # ?ss)), 0)" using encode_config by simp also have "... = prod_encode (e_cons (encode_frame ?t) (list_encode (map encode_frame (?ss))), 0)" by simp also have "... = prod_encode (e_cons (encode_frame ?t) (e2stack ?e), 0)" using assms(1) by simp also have "... = prod_encode (e_cons (triple_encode (pdec222 (e2i ?e)) (e_cons (e2lenls ?e - 1) (e_cons (e_hd (e2ls ?e)) (e_tl (e2xs ?e)))) 0) (e2stack ?e), 0)" using assms(1,2) encode_frame[of g "(length ls - 1) # hd ls # tl xs" "[]"] pdec2_encode_Pr[of n f g] e2xs_xs e2i_f e2lenls_lenls e2ls_ls e_hd by (metis list_encode.simps(1) list.collapse list_decode_encode prod_encode_inverse snd_conv) finally show ?thesis using assms eq estep_Pr_def by auto qed lemma estep_Pr3: assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "ls \ []" and "length ls = Suc (hd xs)" and "recfn (length xs) (Pr n f g)" shows "estep_Pr (encode_config c) = encode_config (step c)" proof - let ?e = "encode_config c" from assms(4) have "length xs > 0" by auto then have "hd xs = e_hd (e2xs ?e)" using assms e_hd_def by auto then have "(length ls = Suc (hd xs)) = (e2lenls ?e = Suc (e_hd (e2xs ?e)))" using assms by simp then have *: "estep_Pr ?e = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" using assms estep_Pr_def by auto have "step c = (fs, Some (hd ls))" using assms(1,2,3) by simp then have "encode_config (step c) = prod_encode (list_encode (map encode_frame fs), encode_option (Some (hd ls)))" using encode_config by simp also have "... = prod_encode (list_encode (map encode_frame fs), encode_option (Some (e_hd (e2ls ?e))))" using assms(1,2) e_hd_def by auto also have "... = prod_encode (list_encode (map encode_frame fs), Suc (e_hd (e2ls ?e)))" by simp also have "... = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" using assms(1) by simp finally have "encode_config (step c) = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" . then show ?thesis using estep_Pr_def * by presburger qed lemma estep_Pr4: assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "ls = []" shows "estep_Pr (encode_config c) = encode_config (step c)" using encode_frame by (simp add: assms estep_Pr_def, simp add: encode_config assms) lemma estep_Pr: assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "recfn (length xs) (Pr n f g)" shows "estep_Pr (encode_config c) = encode_config (step c)" using assms estep_Pr1 estep_Pr2 estep_Pr3 estep_Pr4 by auto definition "estep_Mn e \ if e2ls e = 0 then prod_encode (e_cons (triple_encode (pdec22 (e2i e)) (e_cons 0 (e2xs e)) 0) (e_cons (triple_encode (e2i e) (e2xs e) (singleton_encode 0)) (e2tail e)), 0) else if e2rv e = 1 then prod_encode (e2tail e, Suc (e_hd (e2ls e))) else prod_encode (e_cons (triple_encode (pdec22 (e2i e)) (e_cons (Suc (e_hd (e2ls e))) (e2xs e)) 0) (e_cons (triple_encode (e2i e) (e2xs e) (singleton_encode (Suc (e_hd (e2ls e))))) (e2tail e)), 0)" lemma estep_Mn: assumes "c = (((Mn n f, xs, ls) # fs), rv)" shows "estep_Mn (encode_config c) = encode_config (step c)" proof - let ?e = "encode_config c" consider "ls \ []" and "rv \ Some 0" | "ls \ []" and "rv = Some 0" | "ls = []" by auto then show ?thesis proof (cases) case 1 then have step_c: "step c = ((f, (Suc (hd ls)) # xs, []) # (Mn n f, xs, [Suc (hd ls)]) # fs, None)" (is "step c = ?cfg") using assms by simp have "estep_Mn ?e = prod_encode (e_cons (triple_encode (encode f) (e_cons (Suc (hd ls)) (list_encode xs)) 0) (e_cons (triple_encode (encode (Mn n f)) (list_encode xs) (singleton_encode (Suc (hd ls)))) (list_encode (map encode_frame fs))), 0)" using 1 assms e_hd_def estep_Mn_def by auto also have "... = encode_config ?cfg" using encode_config by (simp add: encode_frame) finally show ?thesis using step_c by simp next case 2 have "estep_Mn ?e = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" using 2 assms estep_Mn_def by auto also have "... = prod_encode (e2tail ?e, Suc (hd ls))" using 2 assms e_hd_def by auto also have "... = prod_encode (list_encode (map encode_frame fs), Suc (hd ls))" using assms by simp also have "... = encode_config (fs, Some (hd ls))" using encode_config by simp finally show ?thesis using 2 assms by simp next case 3 then show ?thesis using assms encode_frame by (simp add: estep_Mn_def, simp add: encode_config) qed qed definition "estep e \ if e2stack e = 0 then prod_encode (0, e2rv e) else if e2i e = 0 then prod_encode (e2tail e, 1) else if e2i e = 1 then prod_encode (e2tail e, Suc (Suc (e_hd (e2xs e)))) else if encode_kind (e2i e) = 2 then prod_encode (e2tail e, Suc (e_nth (e2xs e) (pdec22 (e2i e)))) else if encode_kind (e2i e) = 3 then estep_Cn e else if encode_kind (e2i e) = 4 then estep_Pr e else if encode_kind (e2i e) = 5 then estep_Mn e else 0" lemma estep_Z: assumes "c = (((Z, xs, ls) # fs), rv)" shows "estep (encode_config c) = encode_config (step c)" using encode_frame by (simp add: assms estep_def, simp add: encode_config assms) lemma estep_S: assumes "c = (((S, xs, ls) # fs), rv)" and "recfn (length xs) (fst (hd (fst c)))" shows "estep (encode_config c) = encode_config (step c)" proof - let ?e = "encode_config c" from assms have "length xs > 0" by auto then have eq: "hd xs = e_hd (e2xs ?e)" using assms(1) e_hd_def by auto then have "estep ?e = prod_encode (e2tail ?e, Suc (Suc (e_hd (e2xs ?e))))" using assms(1) estep_def by simp moreover have "step c = (fs, Some (Suc (hd xs)))" using assms(1) by simp ultimately show ?thesis using assms(1) eq estep_def encode_config[of fs "Some (Suc (hd xs))"] by simp qed lemma estep_Id: assumes "c = (((Id m n, xs, ls) # fs), rv)" and "recfn (length xs) (fst (hd (fst c)))" shows "estep (encode_config c) = encode_config (step c)" proof - let ?e = "encode_config c" from assms have "length xs = m" and "m > 0" by auto then have eq: "xs ! n = e_nth (e2xs ?e) n" using assms e_hd_def by auto moreover have "encode_kind (e2i ?e) = 2" using assms(1) encode_kind_2 by auto ultimately have "estep ?e = prod_encode (e2tail ?e, Suc (e_nth (e2xs ?e) (pdec22 (e2i ?e))))" using assms estep_def encode_kind_def by auto moreover have "step c = (fs, Some (xs ! n))" using assms(1) by simp ultimately show ?thesis using assms(1) eq encode_config[of fs "Some (xs ! n)"] by simp qed lemma estep: assumes "valid (fst c)" shows "estep (encode_config c) = encode_config (step c)" proof (cases "fst c") case Nil then show ?thesis using estep_def by (metis list_encode.simps(1) e2rv_def e2stack_stack encode_config_def map_is_Nil_conv prod.collapse prod_encode_inverse snd_conv step.simps(1)) next case (Cons s fs) then obtain f xs ls rv where c: "c = ((f, xs, ls) # fs, rv)" by (metis prod.exhaust_sel) with assms valid_def have lenas: "recfn (length xs) f" by simp show ?thesis proof (cases f) case Z then show ?thesis using estep_Z c by simp next case S then show ?thesis using estep_S c lenas by simp next case Id then show ?thesis using estep_Id c lenas by simp next case Cn then show ?thesis using estep_Cn c by (metis e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2 encode_kind_3 encode_kind_Cn estep_def list.distinct(1) recf.distinct(13) recf.distinct(19) recf.distinct(5)) next case Pr then show ?thesis using estep_Pr c lenas by (metis e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2 encode_kind_4 encode_kind_Cn encode_kind_Pr estep_def list.distinct(1) recf.distinct(15) recf.distinct(21) recf.distinct(25) recf.distinct(7)) next case Mn then show ?thesis using estep_Pr c lenas by (metis (no_types, lifting) e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2 encode_kind_5 encode_kind_Cn encode_kind_Mn encode_kind_Pr estep_Mn estep_def list.distinct(1) recf.distinct(17) recf.distinct(23) recf.distinct(27) recf.distinct(9)) qed qed subsection \The step function as a partial recursive function\label{s:step_recf}\ text \In this section we construct a primitive recursive function @{term r_step} computing @{term estep}. This will entail defining @{typ recf}s for many functions defined in the previous section.\ definition "r_e2stack \ r_pdec1" lemma r_e2stack_prim: "prim_recfn 1 r_e2stack" unfolding r_e2stack_def using r_pdec1_prim by simp lemma r_e2stack [simp]: "eval r_e2stack [e] \= e2stack e" unfolding r_e2stack_def e2stack_def using r_pdec1_prim by simp definition "r_e2rv \ r_pdec2" lemma r_e2rv_prim: "prim_recfn 1 r_e2rv" unfolding r_e2rv_def using r_pdec2_prim by simp lemma r_e2rv [simp]: "eval r_e2rv [e] \= e2rv e" unfolding r_e2rv_def e2rv_def using r_pdec2_prim by simp definition "r_e2tail \ Cn 1 r_tl [r_e2stack]" lemma r_e2tail_prim: "prim_recfn 1 r_e2tail" unfolding r_e2tail_def using r_e2stack_prim r_tl_prim by simp lemma r_e2tail [simp]: "eval r_e2tail [e] \= e2tail e" unfolding r_e2tail_def e2tail_def using r_e2stack_prim r_tl_prim by simp definition "r_e2frame \ Cn 1 r_hd [r_e2stack]" lemma r_e2frame_prim: "prim_recfn 1 r_e2frame" unfolding r_e2frame_def using r_hd_prim r_e2stack_prim by simp lemma r_e2frame [simp]: "eval r_e2frame [e] \= e2frame e" unfolding r_e2frame_def e2frame_def using r_hd_prim r_e2stack_prim by simp definition "r_e2i \ Cn 1 r_pdec1 [r_e2frame]" lemma r_e2i_prim: "prim_recfn 1 r_e2i" unfolding r_e2i_def using r_pdec12_prim r_e2frame_prim by simp lemma r_e2i [simp]: "eval r_e2i [e] \= e2i e" unfolding r_e2i_def e2i_def using r_pdec12_prim r_e2frame_prim by simp definition "r_e2xs \ Cn 1 r_pdec12 [r_e2frame]" lemma r_e2xs_prim: "prim_recfn 1 r_e2xs" unfolding r_e2xs_def using r_pdec122_prim r_e2frame_prim by simp lemma r_e2xs [simp]: "eval r_e2xs [e] \= e2xs e" unfolding r_e2xs_def e2xs_def using r_pdec122_prim r_e2frame_prim by simp definition "r_e2ls \ Cn 1 r_pdec22 [r_e2frame]" lemma r_e2ls_prim: "prim_recfn 1 r_e2ls" unfolding r_e2ls_def using r_pdec222_prim r_e2frame_prim by simp lemma r_e2ls [simp]: "eval r_e2ls [e] \= e2ls e" unfolding r_e2ls_def e2ls_def using r_pdec222_prim r_e2frame_prim by simp definition "r_e2lenls \ Cn 1 r_length [r_e2ls]" lemma r_e2lenls_prim: "prim_recfn 1 r_e2lenls" unfolding r_e2lenls_def using r_length_prim r_e2ls_prim by simp lemma r_e2lenls [simp]: "eval r_e2lenls [e] \= e2lenls e" unfolding r_e2lenls_def e2lenls_def using r_length_prim r_e2ls_prim by simp definition "r_kind \ Cn 1 r_ifz [Id 1 0, Z, Cn 1 r_ifeq [Id 1 0, r_const 1, r_const 1, r_pdec1]]" lemma r_kind_prim: "prim_recfn 1 r_kind" unfolding r_kind_def by simp lemma r_kind: "eval r_kind [e] \= encode_kind e" unfolding r_kind_def encode_kind_def by simp lemmas helpers_for_r_step_prim = r_e2i_prim r_e2lenls_prim r_e2ls_prim r_e2rv_prim r_e2xs_prim r_e2stack_prim r_e2tail_prim r_e2frame_prim text \We define primitive recursive functions @{term r_step_Id}, @{term r_step_Cn}, @{term r_step_Pr}, and @{term r_step_Mn}. The last three correspond to @{term estep_Cn}, @{term estep_Pr}, and @{term estep_Mn} from the previous section.\ definition "r_step_Id \ Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_nth [r_e2xs, Cn 1 r_pdec22 [r_e2i]]]]" lemma r_step_Id: "eval r_step_Id [e] \= prod_encode (e2tail e, Suc (e_nth (e2xs e) (pdec22 (e2i e))))" unfolding r_step_Id_def using helpers_for_r_step_prim by simp abbreviation r_triple_encode :: "recf \ recf \ recf \ recf" where "r_triple_encode x y z \ Cn 1 r_prod_encode [x, Cn 1 r_prod_encode [y, z]]" definition "r_step_Cn \ Cn 1 r_ifeq [r_e2lenls, Cn 1 r_length [Cn 1 r_pdec222 [r_e2i]], Cn 1 r_ifz [r_e2rv, Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec122 [r_e2i]) r_e2ls Z, r_e2stack], Z], Cn 1 r_prod_encode [r_e2tail, r_e2rv]], Cn 1 r_ifz [r_e2rv, Cn 1 r_ifless [r_e2lenls, Cn 1 r_length [Cn 1 r_pdec222 [r_e2i]], Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode (Cn 1 r_nth [Cn 1 r_pdec222 [r_e2i], r_e2lenls]) r_e2xs Z, r_e2stack], Z], Cn 1 r_prod_encode [r_e2tail, r_e2rv]], Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode r_e2i r_e2xs (Cn 1 r_snoc [r_e2ls, Cn 1 r_dec [r_e2rv]]), r_e2tail], Z]]]" lemma r_step_Cn_prim: "prim_recfn 1 r_step_Cn" unfolding r_step_Cn_def using helpers_for_r_step_prim by simp lemma r_step_Cn: "eval r_step_Cn [e] \= estep_Cn e" unfolding r_step_Cn_def estep_Cn_def using helpers_for_r_step_prim by simp definition "r_step_Pr \ Cn 1 r_ifz [r_e2ls, Cn 1 r_ifz [r_e2rv, Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec122 [r_e2i]) (Cn 1 r_tl [r_e2xs]) Z, r_e2stack], Z], Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Cn 1 r_dec [r_e2rv]]), r_e2tail], Z]], Cn 1 r_ifeq [r_e2lenls, Cn 1 S [Cn 1 r_hd [r_e2xs]], Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_hd [r_e2ls]]], Cn 1 r_ifz [r_e2rv, Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec222 [r_e2i]) (Cn 1 r_cons [Cn 1 r_dec [r_e2lenls], Cn 1 r_cons [Cn 1 r_hd [r_e2ls], Cn 1 r_tl [r_e2xs]]]) Z, r_e2stack], Z], Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode r_e2i r_e2xs (Cn 1 r_cons [Cn 1 r_dec [r_e2rv], r_e2ls]), r_e2tail], Z]]]]" lemma r_step_Pr_prim: "prim_recfn 1 r_step_Pr" unfolding r_step_Pr_def using helpers_for_r_step_prim by simp lemma r_step_Pr: "eval r_step_Pr [e] \= estep_Pr e" unfolding r_step_Pr_def estep_Pr_def using helpers_for_r_step_prim by simp definition "r_step_Mn \ Cn 1 r_ifz [r_e2ls, Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec22 [r_e2i]) (Cn 1 r_cons [Z, r_e2xs]) Z, Cn 1 r_cons [r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Z]), r_e2tail]], Z], Cn 1 r_ifeq [r_e2rv, r_const 1, Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_hd [r_e2ls]]], Cn 1 r_prod_encode [Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec22 [r_e2i]) (Cn 1 r_cons [Cn 1 S [Cn 1 r_hd [r_e2ls]], r_e2xs]) Z, Cn 1 r_cons [r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Cn 1 S [Cn 1 r_hd [r_e2ls]]]), r_e2tail]], Z]]]" lemma r_step_Mn_prim: "prim_recfn 1 r_step_Mn" unfolding r_step_Mn_def using helpers_for_r_step_prim by simp lemma r_step_Mn: "eval r_step_Mn [e] \= estep_Mn e" unfolding r_step_Mn_def estep_Mn_def using helpers_for_r_step_prim by simp definition "r_step \ Cn 1 r_ifz [r_e2stack, Cn 1 r_prod_encode [Z, r_e2rv], Cn 1 r_ifz [r_e2i, Cn 1 r_prod_encode [r_e2tail, r_const 1], Cn 1 r_ifeq [r_e2i, r_const 1, Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 S [Cn 1 r_hd [r_e2xs]]]], Cn 1 r_ifeq [Cn 1 r_kind [r_e2i], r_const 2, Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_nth [r_e2xs, Cn 1 r_pdec22 [r_e2i]]]], Cn 1 r_ifeq [Cn 1 r_kind [r_e2i], r_const 3, r_step_Cn, Cn 1 r_ifeq [Cn 1 r_kind [r_e2i], r_const 4, r_step_Pr, Cn 1 r_ifeq [Cn 1 r_kind [r_e2i], r_const 5, r_step_Mn, Z]]]]]]]" lemma r_step_prim: "prim_recfn 1 r_step" unfolding r_step_def using r_kind_prim r_step_Mn_prim r_step_Pr_prim r_step_Cn_prim helpers_for_r_step_prim by simp lemma r_step: "eval r_step [e] \= estep e" unfolding r_step_def estep_def using r_kind_prim r_step_Mn_prim r_step_Pr_prim r_step_Cn_prim helpers_for_r_step_prim r_kind r_step_Cn r_step_Pr r_step_Mn by simp theorem r_step_equiv_step: assumes "valid (fst c)" shows "eval r_step [encode_config c] \= encode_config (step c)" using r_step estep assms by simp subsection \The universal function\label{s:the_universal}\ text \The next function computes the configuration after arbitrarily many steps.\ definition "r_leap \ Pr 2 (Cn 2 r_prod_encode [Cn 2 r_singleton_encode [Cn 2 r_prod_encode [Id 2 0, Cn 2 r_prod_encode [Id 2 1, r_constn 1 0]]], r_constn 1 0]) (Cn 4 r_step [Id 4 1])" lemma r_leap_prim [simp]: "prim_recfn 3 r_leap" unfolding r_leap_def using r_step_prim by simp lemma r_leap_total: "eval r_leap [t, i, x] \" using prim_recfn_total[OF r_leap_prim] by simp lemma r_leap: assumes "i = encode f" and "recfn (e_length x) f" shows "eval r_leap [t, i, x] \= encode_config (iterate t step ([(f, list_decode x, [])], None))" proof (induction t) case 0 then show ?case unfolding r_leap_def using r_step_prim assms encode_config encode_frame by simp next case (Suc t) let ?c = "([(f, list_decode x, [])], None)" let ?tc = "iterate t step ?c" have "valid (fst ?c)" using valid_def assms by simp then have valid: "valid (fst ?tc)" using iterate_step_valid by simp have "eval r_leap [Suc t, i, x] = eval (Cn 4 r_step [Id 4 1]) [t, the (eval r_leap [t, i, x]), i, x]" by (smt One_nat_def Suc_eq_plus1 eq_numeral_Suc eval_Pr_converg_Suc list.size(3) list.size(4) nat_1_add_1 pred_numeral_simps(3) r_leap_def r_leap_prim r_leap_total) then have "eval r_leap [Suc t, i, x] = eval (Cn 4 r_step [Id 4 1]) [t, encode_config ?tc, i, x]" using Suc by simp then have "eval r_leap [Suc t, i, x] = eval r_step [encode_config ?tc]" using r_step_prim by simp then have "eval r_leap [Suc t, i, x] \= encode_config (step ?tc)" by (simp add: r_step_equiv_step valid) then show ?case by simp qed lemma step_leaves_empty_stack_empty: assumes "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" shows "iterate (t + t') step ([(f, list_decode x, [])], None) = ([], Some v)" using assms by (induction t') simp_all text \The next function is essentially a convenience wrapper around @{term r_leap}. It returns zero if the configuration returned by @{term r_leap} is non-final, and @{term "Suc v"} if the configuration is final with return value $v$.\ definition "r_result \ Cn 3 r_ifz [Cn 3 r_pdec1 [r_leap], Cn 3 r_pdec2 [r_leap], r_constn 2 0]" lemma r_result_prim [simp]: "prim_recfn 3 r_result" unfolding r_result_def using r_leap_prim by simp lemma r_result_total: "total r_result" using r_result_prim by blast lemma r_result_empty_stack_None: assumes "i = encode f" and "recfn (e_length x) f" and "iterate t step ([(f, list_decode x, [])], None) = ([], None)" shows "eval r_result [t, i, x] \= 0" unfolding r_result_def using assms r_leap e2stack_0_iff_Nil e2stack_def e2stack_stack r_leap_total r_leap_prim e2rv_def e2rv_rv by simp lemma r_result_empty_stack_Some: assumes "i = encode f" and "recfn (e_length x) f" and "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" shows "eval r_result [t, i, x] \= Suc v" unfolding r_result_def using assms r_leap e2stack_0_iff_Nil e2stack_def e2stack_stack r_leap_total r_leap_prim e2rv_def e2rv_rv by simp lemma r_result_empty_stack_stays: assumes "i = encode f" and "recfn (e_length x) f" and "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" shows "eval r_result [t + t', i, x] \= Suc v" using assms step_leaves_empty_stack_empty r_result_empty_stack_Some by simp lemma r_result_nonempty_stack: assumes "i = encode f" and "recfn (e_length x) f" and "fst (iterate t step ([(f, list_decode x, [])], None)) \ []" shows "eval r_result [t, i, x] \= 0" proof - obtain ss rv where "iterate t step ([(f, list_decode x, [])], None) = (ss, rv)" by fastforce moreover from this assms(3) have "ss \ []" by simp ultimately have "eval r_leap [t, i, x] \= encode_config (ss, rv)" using assms r_leap by simp then have "eval (Cn 3 r_pdec1 [r_leap]) [t, i, x] \\ 0" using \ss \ []\ r_leap_prim encode_config r_leap_total list_encode_0 by (auto, blast) then show ?thesis unfolding r_result_def using r_leap_prim by auto qed lemma r_result_Suc: assumes "i = encode f" and "recfn (e_length x) f" and "eval r_result [t, i, x] \= Suc v" shows "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)" (is "?cfg = _") proof (cases "fst ?cfg") case Nil then show ?thesis using assms r_result_empty_stack_None r_result_empty_stack_Some by (metis Zero_not_Suc nat.inject option.collapse option.inject prod.exhaust_sel) next case Cons then show ?thesis using assms r_result_nonempty_stack by simp qed lemma r_result_converg: assumes "i = encode f" and "recfn (e_length x) f" and "eval f (list_decode x) \= v" shows "\t. (\t'\t. eval r_result [t', i, x] \= Suc v) \ (\t'= 0)" proof - let ?xs = "list_decode x" let ?stack = "[(f, ?xs, [])]" have "wellf f" using assms(2) by simp moreover have "length ?xs = arity f" using assms(2) by simp ultimately have "correct (?stack, None)" using step_correct valid_def by simp with assms(3) have "reachable (?stack, None) ([], Some v)" by simp then obtain t where "iterate t step (?stack, None) = ([], Some v)" "\t' []" using reachable_iterate_step_empty_stack by blast then have t: "eval r_result [t, i, x] \= Suc v" "\t'= 0" using r_result_empty_stack_Some r_result_nonempty_stack assms(1,2) by simp_all then have "eval r_result [t + t', i, x] \= Suc v" for t' using r_result_empty_stack_stays assms r_result_Suc by simp then have "\t'\t. eval r_result [t', i, x] \= Suc v" using le_Suc_ex by blast with t(2) show ?thesis by auto qed lemma r_result_diverg: assumes "i = encode f" and "recfn (e_length x) f" and "eval f (list_decode x) \" shows "eval r_result [t, i, x] \= 0" proof - let ?xs = "list_decode x" let ?stack = "[(f, ?xs, [])]" have "recfn (length ?xs) f" using assms(2) by auto then have "correct (?stack, None)" using step_correct valid_def by simp with assms(3) have "nonterminating (?stack, None)" by simp then show ?thesis using r_result_nonempty_stack assms(1,2) by simp qed text \Now we can define the universal partial recursive function. This function executes @{term r_result} for increasing time bounds, waits for it to reach a final configuration, and then extracts its result value. If no final configuration is reached, the universal function diverges.\ definition "r_univ \ Cn 2 r_dec [Cn 2 r_result [Mn 2 (Cn 3 r_not [r_result]), Id 2 0, Id 2 1]]" lemma r_univ_recfn [simp]: "recfn 2 r_univ" unfolding r_univ_def by simp theorem r_univ: assumes "i = encode f" and "recfn (e_length x) f" shows "eval r_univ [i, x] = eval f (list_decode x)" proof - let ?cond = "Cn 3 r_not [r_result]" let ?while = "Mn 2 ?cond" let ?res = "Cn 2 r_result [?while, Id 2 0, Id 2 1]" let ?xs = "list_decode x" have *: "eval ?cond [t, i, x] \= (if eval r_result [t, i, x] \= 0 then 1 else 0)" for t proof - have "eval ?cond [t, i, x] = eval r_not [the (eval r_result [t, i, x])]" using r_result_total by simp moreover have "eval r_result [t, i, x] \" by (simp add: r_result_total) ultimately show ?thesis by auto qed show ?thesis proof (cases "eval f ?xs \") case True then show ?thesis unfolding r_univ_def using * r_result_diverg[OF assms] eval_Mn_diverg by simp next case False then obtain v where v: "eval f ?xs \= v" by auto then obtain t where t: "\t'\t. eval r_result [t', i, x] \= Suc v" "\t'= 0" using r_result_converg[OF assms] by blast then have "\t'\t. eval ?cond [t', i, x] \= 0" "\t'= 1" using * by simp_all then have "eval ?while [i, x] \= t" using eval_Mn_convergI[of 2 ?cond "[i, x]" t] by simp then have "eval ?res [i, x] = eval r_result [t, i, x]" by simp then have "eval ?res [i, x] \= Suc v" using t(1) by simp then show ?thesis unfolding r_univ_def using v by simp qed qed theorem r_univ': assumes "recfn (e_length x) f" shows "eval r_univ [encode f, x] = eval f (list_decode x)" using r_univ assms by simp text \Universal functions for every arity can be built from @{term "r_univ"}.\ definition r_universal :: "nat \ recf" where "r_universal n \ Cn (Suc n) r_univ [Id (Suc n) 0, r_shift (r_list_encode (n - 1))]" lemma r_universal_recfn [simp]: "n > 0 \ recfn (Suc n) (r_universal n)" unfolding r_universal_def by simp lemma r_universal: assumes "recfn n f" and "length xs = n" shows "eval (r_universal n) (encode f # xs) = eval f xs" unfolding r_universal_def using wellf_arity_nonzero assms r_list_encode r_univ' by fastforce text \We will mostly be concerned with computing unary functions. Hence we introduce separate functions for this case.\ definition "r_result1 \ Cn 3 r_result [Id 3 0, Id 3 1, Cn 3 r_singleton_encode [Id 3 2]]" lemma r_result1_prim [simp]: "prim_recfn 3 r_result1" unfolding r_result1_def by simp lemma r_result1_total: "total r_result1" using Mn_free_imp_total by simp lemma r_result1 [simp]: "eval r_result1 [t, i, x] = eval r_result [t, i, singleton_encode x]" unfolding r_result1_def by simp text \The following function will be our standard Gödel numbering of all unary partial recursive functions.\ definition "r_phi \ r_universal 1" lemma r_phi_recfn [simp]: "recfn 2 r_phi" unfolding r_phi_def by simp theorem r_phi: assumes "i = encode f" and "recfn 1 f" shows "eval r_phi [i, x] = eval f [x]" unfolding r_phi_def using r_universal assms by force corollary r_phi': assumes "recfn 1 f" shows "eval r_phi [encode f, x] = eval f [x]" using assms r_phi by simp lemma r_phi'': "eval r_phi [i, x] = eval r_univ [i, singleton_encode x]" unfolding r_universal_def r_phi_def using r_list_encode by simp section \Applications of the universal function\ text \In this section we shall see some ways @{term r_univ} and @{term r_result} can be used.\ subsection \Lazy conditional evaluation\ text \With the help of @{term r_univ} we can now define a \hypertarget{p:r_lifz}{lazy variant} of @{term r_ifz}, in which only one branch is evaluated.\ definition r_lazyifzero :: "nat \ nat \ nat \ recf" where "r_lazyifzero n j\<^sub>1 j\<^sub>2 \ Cn (Suc (Suc n)) r_univ [Cn (Suc (Suc n)) r_ifz [Id (Suc (Suc n)) 0, r_constn (Suc n) j\<^sub>1, r_constn (Suc n) j\<^sub>2], r_shift (r_list_encode n)]" lemma r_lazyifzero_recfn: "recfn (Suc (Suc n)) (r_lazyifzero n j\<^sub>1 j\<^sub>2)" using r_lazyifzero_def by simp lemma r_lazyifzero: assumes "length xs = Suc n" and "j\<^sub>1 = encode f\<^sub>1" and "j\<^sub>2 = encode f\<^sub>2" and "recfn (Suc n) f\<^sub>1" and "recfn (Suc n) f\<^sub>2" shows "eval (r_lazyifzero n j\<^sub>1 j\<^sub>2) (c # xs) = (if c = 0 then eval f\<^sub>1 xs else eval f\<^sub>2 xs)" proof - let ?a = "r_constn (Suc n) n" let ?b = "Cn (Suc (Suc n)) r_ifz [Id (Suc (Suc n)) 0, r_constn (Suc n) j\<^sub>1, r_constn (Suc n) j\<^sub>2]" let ?c = "r_shift (r_list_encode n)" have "eval ?a (c # xs) \= n" using assms(1) by simp moreover have "eval ?b (c # xs) \= (if c = 0 then j\<^sub>1 else j\<^sub>2)" using assms(1) by simp moreover have "eval ?c (c # xs) \= list_encode xs" using assms(1) r_list_encode r_shift by simp ultimately have "eval (r_lazyifzero n j\<^sub>1 j\<^sub>2) (c # xs) = eval r_univ [if c = 0 then j\<^sub>1 else j\<^sub>2, list_encode xs]" unfolding r_lazyifzero_def using r_lazyifzero_recfn assms(1) by simp then show ?thesis using assms r_univ by simp qed definition r_lifz :: "recf \ recf \ recf" where "r_lifz f g \ r_lazyifzero (arity f - 1) (encode f) (encode g)" lemma r_lifz_recfn [simp]: assumes "recfn n f" and "recfn n g" shows "recfn (Suc n) (r_lifz f g)" using assms r_lazyifzero_recfn r_lifz_def wellf_arity_nonzero by auto lemma r_lifz [simp]: assumes "length xs = n" and "recfn n f" and "recfn n g" shows "eval (r_lifz f g) (c # xs) = (if c = 0 then eval f xs else eval g xs)" using assms r_lazyifzero r_lifz_def wellf_arity_nonzero by (metis One_nat_def Suc_pred) subsection \Enumerating the domains of partial recursive functions\ text \In this section we define a binary function $\mathit{enumdom}$ such that for all $i$, the domain of $\varphi_i$ equals $\{\mathit{enumdom}(i, x) \mid \mathit{enumdom}(i, x)\!\downarrow\}$. In other words, the image of $\mathit{enumdom}_i$ is the domain of $\varphi_i$. First we need some more properties of @{term r_leap} and @{term r_result}.\ lemma r_leap_Suc: "eval r_leap [Suc t, i, x] = eval r_step [the (eval r_leap [t, i, x])]" proof - have "eval r_leap [Suc t, i, x] = eval (Cn 4 r_step [Id 4 1]) [t, the (eval r_leap [t, i, x]), i, x]" using r_leap_total eval_Pr_converg_Suc r_leap_def by (metis length_Cons list.size(3) numeral_2_eq_2 numeral_3_eq_3 r_leap_prim) then show ?thesis using r_step_prim by auto qed lemma r_leap_Suc_saturating: assumes "pdec1 (the (eval r_leap [t, i, x])) = 0" shows "eval r_leap [Suc t, i, x] = eval r_leap [t, i, x]" proof - let ?e = "eval r_leap [t, i, x]" have "eval r_step [the ?e] \= estep (the ?e)" using r_step by simp then have "eval r_step [the ?e] \= prod_encode (0, e2rv (the ?e))" using estep_def assms by (simp add: e2stack_def) then have "eval r_step [the ?e] \= prod_encode (pdec1 (the ?e), pdec2 (the ?e))" using assms by (simp add: e2rv_def) then have "eval r_step [the ?e] \= the ?e" by simp then show ?thesis using r_leap_total r_leap_Suc by simp qed lemma r_result_Suc_saturating: assumes "eval r_result [t, i, x] \= Suc v" shows "eval r_result [Suc t, i, x] \= Suc v" proof - let ?r = "\t. eval r_ifz [pdec1 (the (eval r_leap [t, i, x])), pdec2 (the (eval r_leap [t, i, x])), 0]" have "?r t \= Suc v" using assms unfolding r_result_def using r_leap_total r_leap_prim by simp then have "pdec1 (the (eval r_leap [t, i, x])) = 0" using option.sel by fastforce then have "eval r_leap [Suc t, i, x] = eval r_leap [t, i, x]" using r_leap_Suc_saturating by simp moreover have "eval r_result [t, i, x] = ?r t" unfolding r_result_def using r_leap_total r_leap_prim by simp moreover have "eval r_result [Suc t, i, x] = ?r (Suc t)" unfolding r_result_def using r_leap_total r_leap_prim by simp ultimately have "eval r_result [Suc t, i, x] = eval r_result [t, i, x]" by simp with assms show ?thesis by simp qed lemma r_result_saturating: assumes "eval r_result [t, i, x] \= Suc v" shows "eval r_result [t + d, i, x] \= Suc v" using r_result_Suc_saturating assms by (induction d) simp_all lemma r_result_converg': assumes "eval r_univ [i, x] \= v" shows "\t. (\t'\t. eval r_result [t', i, x] \= Suc v) \ (\t'= 0)" proof - let ?f = "Cn 3 r_not [r_result]" let ?m = "Mn 2 ?f" have "recfn 2 ?m" by simp have eval_m: "eval ?m [i, x] \" proof assume "eval ?m [i, x] \" then have "eval r_univ [i, x] \" unfolding r_univ_def by simp with assms show False by simp qed then obtain t where t: "eval ?m [i, x] \= t" by auto then have f_t: "eval ?f [t, i, x] \= 0" and f_less_t: "\y. y < t \ eval ?f [y, i, x] \\ 0" using eval_Mn_convergE[of 2 ?f "[i, x]" t] \recfn 2 ?m\ by (metis (no_types, lifting) One_nat_def Suc_1 length_Cons list.size(3))+ have eval_Cn2: "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \" proof assume "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \" then have "eval r_univ [i, x] \" unfolding r_univ_def by simp with assms show False by simp qed have "eval r_result [t, i, x] \= Suc v" proof (rule ccontr) assume neq_Suc: "\ eval r_result [t, i, x] \= Suc v" show False proof (cases "eval r_result [t, i, x] = None") case True then show ?thesis using f_t by simp next case False then obtain w where w: "eval r_result [t, i, x] \= w" "w \ Suc v" using neq_Suc by auto moreover have "eval r_result [t, i, x] \\ 0" by (rule ccontr; use f_t in auto) ultimately have "w \ 0" by simp have "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] = eval r_result [the (eval ?m [i, x]), i, x]" using eval_m by simp with w t have "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \= w" by simp moreover have "eval r_univ [i, x] = eval r_dec [the (eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x])]" unfolding r_univ_def using eval_Cn2 by simp ultimately have "eval r_univ [i, x] = eval r_dec [w]" by simp then have "eval r_univ [i, x] \= w - 1" by simp with assms \w \ 0\ w show ?thesis by simp qed qed then have "\t'\t. eval r_result [t', i, x] \= Suc v" using r_result_saturating le_Suc_ex by blast moreover have "eval r_result [y, i, x] \= 0" if "y < t" for y proof (rule ccontr) assume neq0: "eval r_result [y, i, x] \ Some 0" then show False proof (cases "eval r_result [y, i, x] = None") case True then show ?thesis using f_less_t \y < t\ by fastforce next case False then obtain v where "eval r_result [y, i, x] \= v" "v \ 0" using neq0 by auto then have "eval ?f [y, i, x] \= 0" by simp then show ?thesis using f_less_t \y < t\ by simp qed qed ultimately show ?thesis by auto qed lemma r_result_diverg': assumes "eval r_univ [i, x] \" shows "eval r_result [t, i, x] \= 0" proof (rule ccontr) let ?f = "Cn 3 r_not [r_result]" let ?m = "Mn 2 ?f" assume "eval r_result [t, i, x] \ Some 0" with r_result_total have "eval r_result [t, i, x] \\ 0" by simp then have "eval ?f [t, i, x] \= 0" by auto moreover have "eval ?f [y, i, x] \" if "y < t" for y using r_result_total by simp ultimately have "\z. eval ?f (z # [i, x]) \= 0 \ (\y)" by blast then have "eval ?m [i, x] \" by simp then have "eval r_univ [i, x] \" unfolding r_univ_def using r_result_total by simp with assms show False by simp qed lemma r_result_bivalent': assumes "eval r_univ [i, x] \= v" shows "eval r_result [t, i, x] \= Suc v \ eval r_result [t, i, x] \= 0" using r_result_converg'[OF assms] not_less by blast lemma r_result_Some': assumes "eval r_result [t, i, x] \= Suc v" shows "eval r_univ [i, x] \= v" proof (rule ccontr) assume not_v: "\ eval r_univ [i, x] \= v" show False proof (cases "eval r_univ [i, x] \") case True then show ?thesis using assms r_result_diverg' by simp next case False then obtain w where w: "eval r_univ [i, x] \= w" "w \ v" using not_v by auto then have "eval r_result [t, i, x] \= Suc w \ eval r_result [t, i, x] \= 0" using r_result_bivalent' by simp then show ?thesis using assms not_v w by simp qed qed lemma r_result1_converg': assumes "eval r_phi [i, x] \= v" shows "\t. (\t'\t. eval r_result1 [t', i, x] \= Suc v) \ (\t'= 0)" using assms r_result1 r_result_converg' r_phi'' by simp lemma r_result1_diverg': assumes "eval r_phi [i, x] \" shows "eval r_result1 [t, i, x] \= 0" using assms r_result1 r_result_diverg' r_phi'' by simp lemma r_result1_Some': assumes "eval r_result1 [t, i, x] \= Suc v" shows "eval r_phi [i, x] \= v" using assms r_result1 r_result_Some' r_phi'' by simp text \The next function performs dovetailing in order to evaluate $\varphi_i$ for every argument for arbitrarily many steps. Given $i$ and $z$, the function decodes $z$ into a pair $(x, t$) and outputs zero (meaning ``true'') iff.\ the computation of $\varphi_i$ on input $x$ halts after at most $t$ steps. Fixing $i$ and varying $z$ will eventually compute $\varphi_i$ for every argument in the domain of $\varphi_i$ sufficiently long for it to converge.\ definition "r_dovetail \ Cn 2 r_not [Cn 2 r_result1 [Cn 2 r_pdec2 [Id 2 1], Id 2 0, Cn 2 r_pdec1 [Id 2 1]]]" lemma r_dovetail_prim: "prim_recfn 2 r_dovetail" by (simp add: r_dovetail_def) lemma r_dovetail: "eval r_dovetail [i, z] \= (if the (eval r_result1 [pdec2 z, i, pdec1 z]) > 0 then 0 else 1)" unfolding r_dovetail_def using r_result_total by simp text \The function $\mathit{enumdom}$ works as follows in order to enumerate exactly the domain of $\varphi_i$. Given $i$ and $y$ it searches for the minimum $z \geq y$ for which the dovetail function returns true. This $z$ is decoded into $(x, t)$ and the $x$ is output. In this way every value output by $\mathit{enumdom}$ is in the domain of $\varphi_i$ by construction of @{term r_dovetail}. Conversely an $x$ in the domain will be output for $y = (x, t)$ where $t$ is such that $\varphi_i$ halts on $x$ within $t$ steps.\ definition "r_dovedelay \ Cn 3 r_and [Cn 3 r_dovetail [Id 3 1, Id 3 0], Cn 3 r_ifle [Id 3 2, Id 3 0, r_constn 2 0, r_constn 2 1]]" lemma r_dovedelay_prim: "prim_recfn 3 r_dovedelay" unfolding r_dovedelay_def using r_dovetail_prim by simp lemma r_dovedelay: "eval r_dovedelay [z, i, y] \= (if the (eval r_result1 [pdec2 z, i, pdec1 z]) > 0 \ y \ z then 0 else 1)" by (simp add: r_dovedelay_def r_dovetail r_dovetail_prim) definition "r_enumdom \ Cn 2 r_pdec1 [Mn 2 r_dovedelay]" lemma r_enumdom_recfn [simp]: "recfn 2 r_enumdom" by (simp add: r_enumdom_def r_dovedelay_prim) lemma r_enumdom [simp]: "eval r_enumdom [i, y] = (if \z. eval r_dovedelay [z, i, y] \= 0 then Some (pdec1 (LEAST z. eval r_dovedelay [z, i, y] \= 0)) else None)" proof - let ?h = "Mn 2 r_dovedelay" have "total r_dovedelay" using r_dovedelay_prim by blast then have "eval ?h [i, y] = (if (\z. eval r_dovedelay [z, i, y] \= 0) then Some (LEAST z. eval r_dovedelay [z, i, y] \= 0) else None)" using r_dovedelay_prim r_enumdom_recfn eval_Mn_convergI by simp then show ?thesis unfolding r_enumdom_def using r_dovedelay_prim by simp qed text \If @{term i} is the code of the empty function, @{term r_enumdom} has an empty domain, too.\ lemma r_enumdom_empty_domain: assumes "\x. eval r_phi [i, x] \" shows "\y. eval r_enumdom [i, y] \" using assms r_result1_diverg' r_dovedelay by simp text \If @{term i} is the code of a function with non-empty domain, @{term r_enumdom} enumerates its domain.\ lemma r_enumdom_nonempty_domain: assumes "eval r_phi [i, x\<^sub>0] \" shows "\y. eval r_enumdom [i, y] \" and "\x. eval r_phi [i, x] \ \ (\y. eval r_enumdom [i, y] \= x)" proof - show "eval r_enumdom [i, y] \" for y proof - obtain t where t: "\t'\t. the (eval r_result1 [t', i, x\<^sub>0]) > 0" using assms r_result1_converg' by fastforce let ?z = "prod_encode (x\<^sub>0, max t y)" have "y \ ?z" using le_prod_encode_2 max.bounded_iff by blast moreover have "pdec2 ?z \ t" by simp ultimately have "the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0" using t by simp with \y \ ?z\ r_dovedelay have "eval r_dovedelay [?z, i, y] \= 0" by presburger then show "eval r_enumdom [i, y] \" using r_enumdom by auto qed show "eval r_phi [i, x] \ = (\y. eval r_enumdom [i, y] \= x)" for x proof show "\y. eval r_enumdom [i, y] \= x" if "eval r_phi [i, x] \" for x proof - from that obtain v where "eval r_phi [i, x] \= v" by auto then obtain t where t: "the (eval r_result1 [t, i, x]) > 0" using r_result1_converg' assms by (metis Zero_not_Suc dual_order.refl option.sel zero_less_iff_neq_zero) let ?y = "prod_encode (x, t)" have "eval r_dovedelay [?y, i, ?y] \= 0" using r_dovedelay t by simp moreover from this have "(LEAST z. eval r_dovedelay [z, i, ?y] \= 0) = ?y" using gr_implies_not_zero r_dovedelay by (intro Least_equality; fastforce) ultimately have "eval r_enumdom [i, ?y] \= x" using r_enumdom by auto then show ?thesis by blast qed show "eval r_phi [i, x] \" if "\y. eval r_enumdom [i, y] \= x" for x proof - from that obtain y where y: "eval r_enumdom [i, y] \= x" by auto then have "eval r_enumdom [i, y] \" by simp then have "\z. eval r_dovedelay [z, i, y] \= 0" and *: "eval r_enumdom [i, y] \= pdec1 (LEAST z. eval r_dovedelay [z, i, y] \= 0)" (is "_ \= pdec1 ?z") using r_enumdom by metis+ then have z: "eval r_dovedelay [?z, i, y] \= 0" by (meson wellorder_Least_lemma(1)) have "the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0" proof (rule ccontr) assume "\ (the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0)" then show False using r_dovedelay z by simp qed then have "eval r_phi [i, pdec1 ?z] \" using r_result1_diverg' assms by fastforce then show ?thesis using y * by auto qed qed qed text \For every $\varphi_i$ with non-empty domain there is a total recursive function that enumerates the domain of $\varphi_i$.\ lemma nonempty_domain_enumerable: assumes "eval r_phi [i, x\<^sub>0] \" shows "\g. recfn 1 g \ total g \ (\x. eval r_phi [i, x] \ \ (\y. eval g [y] \= x))" proof - define g where "g \ Cn 1 r_enumdom [r_const i, Id 1 0]" then have "recfn 1 g" by simp moreover from this have "total g" using totalI1[of g] g_def assms r_enumdom_nonempty_domain(1) by simp moreover have "eval r_phi [i, x] \ \ (\y. eval g [y] \= x)" for x unfolding g_def using r_enumdom_nonempty_domain(2)[OF assms] by simp ultimately show ?thesis by auto qed subsection \Concurrent evaluation of functions\ text \We define a function that simulates two @{typ recf}s ``concurrently'' for the same argument and returns the result of the one converging first. If both diverge, so does the simulation function.\ definition "r_both \ Cn 4 r_ifz [Cn 4 r_result1 [Id 4 0, Id 4 1, Id 4 3], Cn 4 r_ifz [Cn 4 r_result1 [Id 4 0, Id 4 2, Id 4 3], Cn 4 r_prod_encode [r_constn 3 2, r_constn 3 0], Cn 4 r_prod_encode [r_constn 3 1, Cn 4 r_dec [Cn 4 r_result1 [Id 4 0, Id 4 2, Id 4 3]]]], Cn 4 r_prod_encode [r_constn 3 0, Cn 4 r_dec [Cn 4 r_result1 [Id 4 0, Id 4 1, Id 4 3]]]]" lemma r_both_prim [simp]: "prim_recfn 4 r_both" unfolding r_both_def by simp lemma r_both: assumes "\x. eval r_phi [i, x] = eval f [x]" and "\x. eval r_phi [j, x] = eval g [x]" shows "eval f [x] \ \ eval g [x] \ \ eval r_both [t, i, j, x] \= prod_encode (2, 0)" and "\eval r_result1 [t, i, x] \= 0; eval r_result1 [t, j, x] \= 0\ \ eval r_both [t, i, j, x] \= prod_encode (2, 0)" and "eval r_result1 [t, i, x] \= Suc v \ eval r_both [t, i, j, x] \= prod_encode (0, the (eval f [x]))" and "\eval r_result1 [t, i, x] \= 0; eval r_result1 [t, j, x] \= Suc v\ \ eval r_both [t, i, j, x] \= prod_encode (1, the (eval g [x]))" proof - have r_result_total [simp]: "eval r_result [t, k, x] \" for t k x using r_result_total by simp { assume "eval f [x] \ \ eval g [x] \" then have "eval r_result1 [t, i, x] \= 0" and "eval r_result1 [t, j, x] \= 0" using assms r_result1_diverg' by auto then show "eval r_both [t, i, j, x] \= prod_encode (2, 0)" unfolding r_both_def by simp next assume "eval r_result1 [t, i, x] \= 0" and "eval r_result1 [t, j, x] \= 0" then show "eval r_both [t, i, j, x] \= prod_encode (2, 0)" unfolding r_both_def by simp next assume "eval r_result1 [t, i, x] \= Suc v" moreover from this have "eval r_result1 [t, i, x] \= Suc (the (eval f [x]))" using assms r_result1_Some' by fastforce ultimately show "eval r_both [t, i, j, x] \= prod_encode (0, the (eval f [x]))" unfolding r_both_def by auto next assume "eval r_result1 [t, i, x] \= 0" and "eval r_result1 [t, j, x] \= Suc v" moreover from this have "eval r_result1 [t, j, x] \= Suc (the (eval g [x]))" using assms r_result1_Some' by fastforce ultimately show "eval r_both [t, i, j, x] \= prod_encode (1, the (eval g [x]))" unfolding r_both_def by auto } qed definition "r_parallel \ Cn 3 r_both [Mn 3 (Cn 4 r_le [Cn 4 r_pdec1 [r_both], r_constn 3 1]), Id 3 0, Id 3 1, Id 3 2]" lemma r_parallel_recfn [simp]: "recfn 3 r_parallel" unfolding r_parallel_def by simp lemma r_parallel: assumes "\x. eval r_phi [i, x] = eval f [x]" and "\x. eval r_phi [j, x] = eval g [x]" shows "eval f [x] \ \ eval g [x] \ \ eval r_parallel [i, j, x] \" and "eval f [x] \ \ eval g [x] \ \ eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x]))" and "eval g [x] \ \ eval f [x] \ \ eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" and "eval f [x] \ \ eval g [x] \ \ eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x])) \ eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" proof - let ?cond = "Cn 4 r_le [Cn 4 r_pdec1 [r_both], r_constn 3 1]" define m where "m = Mn 3 ?cond" then have m: "r_parallel = Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]" unfolding r_parallel_def by simp from m_def have "recfn 3 m" by simp { assume "eval f [x] \ \ eval g [x] \" then have "\t. eval r_both [t, i, j, x] \= prod_encode (2, 0)" using assms r_both by simp then have "eval ?cond [t, i, j, x] \= 1" for t by simp then have "eval m [i, j, x] \" unfolding m_def using eval_Mn_diverg by simp then have "eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x] \" using \recfn 3 m\ by simp then show "eval r_parallel [i, j, x] \" using m by simp next assume "eval f [x] \ \ eval g [x] \" then obtain vf vg where v: "eval f [x] \= vf" "eval g [x] \= vg" by auto then obtain tf where tf: "\t\tf. eval r_result1 [t, i, x] \= Suc vf" "\t= 0" using r_result1_converg' assms by metis from v obtain tg where tg: "\t\tg. eval r_result1 [t, j, x] \= Suc vg" "\t= 0" using r_result1_converg' assms by metis show "eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x])) \ eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" proof (cases "tf \ tg") case True with tg(2) have j0: "\t= 0" by simp have *: "eval r_both [tf, i, j, x] \= prod_encode (0, the (eval f [x]))" using r_both(3) assms tf(1) by simp have "eval m [i, j, x] \= tf" unfolding m_def proof (rule eval_Mn_convergI) show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp have "eval (Cn 4 r_pdec1 [r_both]) [tf, i, j, x] \= 0" using * by simp then show "eval ?cond [tf, i, j, x] \= 0" by simp have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < tf" for t using tf(2) r_both(2) assms that j0 by simp then have "eval ?cond [t, i, j, x] \= 1" if "t < tf" for t using that by simp then show "\y. y < tf \ eval ?cond [y, i, j, x] \\ 0" by simp qed moreover have "eval r_parallel [i, j, x] = eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" using m by simp ultimately have "eval r_parallel [i, j, x] = eval r_both [tf, i, j, x]" using \recfn 3 m\ by simp with * have "eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x]))" by simp then show ?thesis by simp next case False with tf(2) have i0: "\t\tg. eval r_result1 [t, i, x] \= 0" by simp then have *: "eval r_both [tg, i, j, x] \= prod_encode (1, the (eval g [x]))" using assms r_both(4) tg(1) by auto have "eval m [i, j, x] \= tg" unfolding m_def proof (rule eval_Mn_convergI) show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp have "eval (Cn 4 r_pdec1 [r_both]) [tg, i, j, x] \= 1" using * by simp then show "eval ?cond [tg, i, j, x] \= 0" by simp have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < tg" for t using tg(2) r_both(2) assms that i0 by simp then have "eval ?cond [t, i, j, x] \= 1" if "t < tg" for t using that by simp then show "\y. y < tg \ eval ?cond [y, i, j, x] \\ 0" by simp qed moreover have "eval r_parallel [i, j, x] = eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" using m by simp ultimately have "eval r_parallel [i, j, x] = eval r_both [tg, i, j, x]" using \recfn 3 m\ by simp with * have "eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" by simp then show ?thesis by simp qed next assume eval_fg: "eval g [x] \ \ eval f [x] \" then have i0: "\t. eval r_result1 [t, i, x] \= 0" using r_result1_diverg' assms by auto from eval_fg obtain v where "eval g [x] \= v" by auto then obtain t\<^sub>0 where t0: "\t\t\<^sub>0. eval r_result1 [t, j, x] \= Suc v" "\t0. eval r_result1 [t, j, x] \= 0" using r_result1_converg' assms by metis then have *: "eval r_both [t\<^sub>0, i, j, x] \= prod_encode (1, the (eval g [x]))" using r_both(4) assms i0 by simp have "eval m [i, j, x] \= t\<^sub>0" unfolding m_def proof (rule eval_Mn_convergI) show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp have "eval (Cn 4 r_pdec1 [r_both]) [t\<^sub>0, i, j, x] \= 1" using * by simp then show "eval ?cond [t\<^sub>0, i, j, x] \= 0" by simp have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < t\<^sub>0" for t using t0(2) r_both(2) assms that i0 by simp then have "eval ?cond [t, i, j, x] \= 1" if "t < t\<^sub>0" for t using that by simp then show "\y. y < t\<^sub>0 \ eval ?cond [y, i, j, x] \\ 0" by simp qed moreover have "eval r_parallel [i, j, x] = eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" using m by simp ultimately have "eval r_parallel [i, j, x] = eval r_both [t\<^sub>0, i, j, x]" using \recfn 3 m\ by simp with * show "eval r_parallel [i, j, x] \= prod_encode (1, the (eval g [x]))" by simp next assume eval_fg: "eval f [x] \ \ eval g [x] \" then have j0: "\t. eval r_result1 [t, j, x] \= 0" using r_result1_diverg' assms by auto from eval_fg obtain v where "eval f [x] \= v" by auto then obtain t\<^sub>0 where t0: "\t\t\<^sub>0. eval r_result1 [t, i, x] \= Suc v" "\t0. eval r_result1 [t, i, x] \= 0" using r_result1_converg' assms by metis then have *: "eval r_both [t\<^sub>0, i, j, x] \= prod_encode (0, the (eval f [x]))" using r_both(3) assms by blast have "eval m [i, j, x] \= t\<^sub>0" unfolding m_def proof (rule eval_Mn_convergI) show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp have "eval (Cn 4 r_pdec1 [r_both]) [t\<^sub>0, i, j, x] \= 0" using * by simp then show "eval ?cond [t\<^sub>0, i, j, x] \= 0" by simp have "eval r_both [t, i, j, x] \= prod_encode (2, 0)" if "t < t\<^sub>0" for t using t0(2) r_both(2) assms that j0 by simp then have "eval ?cond [t, i, j, x] \= 1" if "t < t\<^sub>0" for t using that by simp then show "\y. y < t\<^sub>0 \ eval ?cond [y, i, j, x] \\ 0" by simp qed moreover have "eval r_parallel [i, j, x] = eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]" using m by simp ultimately have "eval r_parallel [i, j, x] = eval r_both [t\<^sub>0, i, j, x]" using \recfn 3 m\ by simp with * show "eval r_parallel [i, j, x] \= prod_encode (0, the (eval f [x]))" by simp } qed end \ No newline at end of file diff --git a/thys/JinjaThreads/MM/SC_Completion.thy b/thys/JinjaThreads/MM/SC_Completion.thy --- a/thys/JinjaThreads/MM/SC_Completion.thy +++ b/thys/JinjaThreads/MM/SC_Completion.thy @@ -1,1284 +1,1284 @@ (* Title: JinjaThreads/MM/SC_Completion.thy Author: Andreas Lochbihler *) section \Sequentially consistent completion of executions in the JMM\ theory SC_Completion imports Non_Speculative begin subsection \Most recently written values\ fun mrw_value :: "'m prog \ (('addr \ addr_loc) \ ('addr val \ bool)) \ ('addr, 'thread_id) obs_event action \ (('addr \ addr_loc) \ ('addr val \ bool))" where "mrw_value P vs (NormalAction (WriteMem ad al v)) = vs((ad, al) \ (v, True))" | "mrw_value P vs (NormalAction (NewHeapElem ad hT)) = (\(ad', al). if ad = ad' \ al \ addr_locs P hT \ (case vs (ad, al) of None \ True | Some (v, b) \ \ b) then Some (addr_loc_default P hT al, False) else vs (ad', al))" | "mrw_value P vs _ = vs" lemma mrw_value_cases: obtains ad al v where "x = NormalAction (WriteMem ad al v)" | ad hT where "x = NormalAction (NewHeapElem ad hT)" | ad M vs v where "x = NormalAction (ExternalCall ad M vs v)" | ad al v where "x = NormalAction (ReadMem ad al v)" | t where "x = NormalAction (ThreadStart t)" | t where "x = NormalAction (ThreadJoin t)" | ad where "x = NormalAction (SyncLock ad)" | ad where "x = NormalAction (SyncUnlock ad)" | t where "x = NormalAction (ObsInterrupt t)" | t where "x = NormalAction (ObsInterrupted t)" | "x = InitialThreadAction" | "x = ThreadFinishAction" by pat_completeness abbreviation mrw_values :: "'m prog \ (('addr \ addr_loc) \ ('addr val \ bool)) \ ('addr, 'thread_id) obs_event action list \ (('addr \ addr_loc) \ ('addr val \ bool))" where "mrw_values P \ foldl (mrw_value P)" lemma mrw_values_eq_SomeD: assumes mrw: "mrw_values P vs0 obs (ad, al) = \(v, b)\" and "vs0 (ad, al) = \(v, b)\ \ \wa. wa \ set obs \ is_write_action wa \ (ad, al) \ action_loc_aux P wa \ (b \ \ is_new_action wa)" shows "\obs' wa obs''. obs = obs' @ wa # obs'' \ is_write_action wa \ (ad, al) \ action_loc_aux P wa \ value_written_aux P wa al = v \ (is_new_action wa \ \ b) \ (\ob\set obs''. is_write_action ob \ (ad, al) \ action_loc_aux P ob \ is_new_action ob \ b)" (is "?concl obs") using assms proof(induct obs rule: rev_induct) case Nil thus ?case by simp next case (snoc ob obs) note mrw = \mrw_values P vs0 (obs @ [ob]) (ad, al) = \(v, b)\\ show ?case proof(cases "is_write_action ob \ (ad, al) \ action_loc_aux P ob \ (is_new_action ob \ \ b)") case True thus ?thesis using mrw by(fastforce elim!: is_write_action.cases intro: action_loc_aux_intros split: if_split_asm) next case False with mrw have "mrw_values P vs0 obs (ad, al) = \(v, b)\" by(cases "ob" rule: mrw_value_cases)(auto split: if_split_asm simp add: addr_locs_def split: htype.split_asm) moreover { assume "vs0 (ad, al) = \(v, b)\" hence "\wa. wa \ set (obs @ [ob]) \ is_write_action wa \ (ad, al) \ action_loc_aux P wa \ (b \ \ is_new_action wa)" by(rule snoc) with False have "\wa. wa \ set obs \ is_write_action wa \ (ad, al) \ action_loc_aux P wa \ (b \ \ is_new_action wa)" by auto } ultimately have "?concl obs" by(rule snoc) thus ?thesis using False mrw by fastforce qed qed lemma mrw_values_WriteMemD: assumes "NormalAction (WriteMem ad al v') \ set obs" shows "\v. mrw_values P vs0 obs (ad, al) = Some (v, True)" using assms apply(induct obs rule: rev_induct) apply simp apply clarsimp apply(erule disjE) apply clarsimp apply clarsimp apply(case_tac x rule: mrw_value_cases) apply simp_all done lemma mrw_values_new_actionD: assumes "w \ set obs" "is_new_action w" "adal \ action_loc_aux P w" shows "\v b. mrw_values P vs0 obs adal = Some (v, b)" using assms apply(induct obs rule: rev_induct) apply simp apply clarsimp apply(erule disjE) apply(fastforce simp add: split_beta elim!: action_loc_aux_cases is_new_action.cases) apply clarsimp apply(rename_tac w' obs' v b) apply(case_tac w' rule: mrw_value_cases) apply(auto simp add: split_beta) done lemma mrw_value_dom_mono: "dom vs \ dom (mrw_value P vs ob)" by(cases ob rule: mrw_value_cases) auto lemma mrw_values_dom_mono: "dom vs \ dom (mrw_values P vs obs)" by(induct obs arbitrary: vs)(auto intro: subset_trans[OF mrw_value_dom_mono] del: subsetI) lemma mrw_values_eq_NoneD: assumes "mrw_values P vs0 obs adal = None" and "w \ set obs" and "is_write_action w" and "adal \ action_loc_aux P w" shows False using assms apply - apply(erule is_write_action.cases) apply(fastforce dest: mrw_values_WriteMemD[where ?vs0.0=vs0 and P=P] mrw_values_new_actionD[where ?vs0.0=vs0] elim: action_loc_aux_cases)+ done lemma mrw_values_mrw: assumes mrw: "mrw_values P vs0 (map snd obs) (ad, al) = \(v, b)\" and initial: "vs0 (ad, al) = \(v, b)\ \ \wa. wa \ set (map snd obs) \ is_write_action wa \ (ad, al) \ action_loc_aux P wa \ (b \ \ is_new_action wa)" shows "\i. i < length obs \ P,llist_of (obs @ [(t, NormalAction (ReadMem ad al v))]) \ length obs \mrw i \ value_written P (llist_of obs) i (ad, al) = v" proof - from mrw_values_eq_SomeD[OF mrw initial] obtain obs' wa obs'' where obs: "map snd obs = obs' @ wa # obs''" and wa: "is_write_action wa" and adal: "(ad, al) \ action_loc_aux P wa" and written: "value_written_aux P wa al = v" and new: "is_new_action wa \ \ b" and last: "\ob. \ ob \ set obs''; is_write_action ob; (ad, al) \ action_loc_aux P ob \ \ is_new_action ob \ b" by blast let ?i = "length obs'" let ?E = "llist_of (obs @ [(t, NormalAction (ReadMem ad al v))])" from obs have len: "length (map snd obs) = Suc (length obs') + length obs''" by simp hence "?i < length obs" by simp moreover hence obs_i: "action_obs ?E ?i = wa" using len obs by(auto simp add: action_obs_def map_eq_append_conv) have "P,?E \ length obs \mrw ?i" proof(rule most_recent_write_for.intros) show "length obs \ read_actions ?E" by(auto intro: read_actions.intros simp add: actions_def action_obs_def) show "(ad, al) \ action_loc P ?E (length obs)" by(simp add: action_obs_def lnth_llist_of) show "?E \ length obs' \a length obs" using len by-(rule action_orderI, auto simp add: actions_def action_obs_def nth_append) show "?i \ write_actions ?E" using len obs wa by-(rule write_actions.intros, auto simp add: actions_def action_obs_def nth_append map_eq_append_conv) show "(ad, al) \ action_loc P ?E ?i" using obs_i adal by simp fix wa' assume wa': "wa' \ write_actions ?E" and adal': "(ad, al) \ action_loc P ?E wa'" from wa' \?i \ write_actions ?E\ have "wa' \ actions ?E" "?i \ actions ?E" by simp_all hence "?E \ wa' \a ?i" proof(rule action_orderI) assume new_wa': "is_new_action (action_obs ?E wa')" and new_i: "is_new_action (action_obs ?E ?i)" from new_i obs_i new have b: "\ b" by simp show "wa' \ ?i" proof(rule ccontr) assume "\ ?thesis" hence "?i < wa'" by simp hence "snd (obs ! wa') \ set obs''" using obs wa' unfolding in_set_conv_nth by -(rule exI[where x="wa' - Suc (length obs')"], auto elim!: write_actions.cases actionsE simp add: action_obs_def lnth_llist_of actions_def nth_append map_eq_append_conv nth_Cons' split: if_split_asm) moreover from wa' have "is_write_action (snd (obs ! wa'))" by cases(auto simp add: action_obs_def nth_append actions_def split: if_split_asm) moreover from adal' wa' have "(ad, al) \ action_loc_aux P (snd (obs ! wa'))" by(auto simp add: action_obs_def nth_append nth_Cons' actions_def split: if_split_asm elim!: write_actions.cases) ultimately show False using last[of "snd (obs ! wa')"] b by simp qed next assume new_wa': "\ is_new_action (action_obs ?E wa')" with wa' adal' obtain v' where "NormalAction (WriteMem ad al v') \ set (map snd obs)" unfolding in_set_conv_nth by (fastforce elim!: write_actions.cases is_write_action.cases simp add: action_obs_def actions_def nth_append split: if_split_asm intro!: exI[where x=wa']) from mrw_values_WriteMemD[OF this, of P vs0] mrw have b by simp with new obs_i have "\ is_new_action (action_obs ?E ?i)" by simp moreover have "wa' \ ?i" proof(rule ccontr) assume "\ ?thesis" hence "?i < wa'" by simp hence "snd (obs ! wa') \ set obs''" using obs wa' unfolding in_set_conv_nth by -(rule exI[where x="wa' - Suc (length obs')"], auto elim!: write_actions.cases actionsE simp add: action_obs_def lnth_llist_of actions_def nth_append map_eq_append_conv nth_Cons' split: if_split_asm) moreover from wa' have "is_write_action (snd (obs ! wa'))" by cases(auto simp add: action_obs_def nth_append actions_def split: if_split_asm) moreover from adal' wa' have "(ad, al) \ action_loc_aux P (snd (obs ! wa'))" by(auto simp add: action_obs_def nth_append nth_Cons' actions_def split: if_split_asm elim!: write_actions.cases) ultimately have "is_new_action (snd (obs ! wa'))" using last[of "snd (obs ! wa')"] by simp moreover from new_wa' wa' have "\ is_new_action (snd (obs ! wa'))" by(auto elim!: write_actions.cases simp add: action_obs_def nth_append actions_def split: if_split_asm) ultimately show False by contradiction qed ultimately show "\ is_new_action (action_obs ?E ?i) \ wa' \ ?i" by blast qed thus "?E \ wa' \a ?i \ ?E \ length obs \a wa'" .. qed moreover from written \?i < length obs\ obs_i have "value_written P (llist_of obs) ?i (ad, al) = v" by(simp add: value_written_def action_obs_def nth_append) ultimately show ?thesis by blast qed lemma mrw_values_no_write_unchanged: assumes no_write: "\w. \ w \ set obs; is_write_action w; adal \ action_loc_aux P w \ \ case vs adal of None \ False | Some (v, b) \ b \ is_new_action w" shows "mrw_values P vs obs adal = vs adal" using assms proof(induct obs arbitrary: vs) case Nil show ?case by simp next case (Cons ob obs) from Cons.prems[of ob] have "mrw_value P vs ob adal = vs adal" apply(cases adal) apply(cases ob rule: mrw_value_cases, fastforce+) apply(auto simp add: addr_locs_def split: htype.split_asm) apply blast+ done moreover have "mrw_values P (mrw_value P vs ob) obs adal = mrw_value P vs ob adal" proof(rule Cons.hyps) fix w assume "w \ set obs" "is_write_action w" "adal \ action_loc_aux P w" with Cons.prems[of w] \mrw_value P vs ob adal = vs adal\ show "case mrw_value P vs ob adal of None \ False | \(v, b)\ \ b \ is_new_action w" by simp qed ultimately show ?case by simp qed subsection \Coinductive version of sequentially consistent prefixes\ coinductive ta_seq_consist :: "'m prog \ ('addr \ addr_loc \ 'addr val \ bool) \ ('addr, 'thread_id) obs_event action llist \ bool" for P :: "'m prog" where LNil: "ta_seq_consist P vs LNil" | LCons: "\ case ob of NormalAction (ReadMem ad al v) \ \b. vs (ad, al) = \(v, b)\ | _ \ True; ta_seq_consist P (mrw_value P vs ob) obs \ \ ta_seq_consist P vs (LCons ob obs)" inductive_simps ta_seq_consist_simps [simp]: "ta_seq_consist P vs LNil" "ta_seq_consist P vs (LCons ob obs)" lemma ta_seq_consist_lappend: assumes "lfinite obs" shows "ta_seq_consist P vs (lappend obs obs') \ ta_seq_consist P vs obs \ ta_seq_consist P (mrw_values P vs (list_of obs)) obs'" (is "?concl vs obs") using assms proof(induct arbitrary: vs) case lfinite_LNil thus ?case by simp next case (lfinite_LConsI obs ob) have "?concl (mrw_value P vs ob) obs" by fact thus ?case using \lfinite obs\ by(simp split: action.split add: list_of_LCons) qed lemma assumes "ta_seq_consist P vs obs" shows ta_seq_consist_ltake: "ta_seq_consist P vs (ltake n obs)" (is ?thesis1) and ta_seq_consist_ldrop: "ta_seq_consist P (mrw_values P vs (list_of (ltake n obs))) (ldrop n obs)" (is ?thesis2) proof - note assms also have "obs = lappend (ltake n obs) (ldrop n obs)" by(simp add: lappend_ltake_ldrop) finally have "?thesis1 \ ?thesis2" by(cases n)(simp_all add: ta_seq_consist_lappend del: lappend_ltake_enat_ldropn) thus ?thesis1 ?thesis2 by blast+ qed lemma ta_seq_consist_coinduct_append [consumes 1, case_names ta_seq_consist, case_conclusion ta_seq_consist LNil lappend]: assumes major: "X vs obs" and step: "\vs obs. X vs obs \ obs = LNil \ (\obs' obs''. obs = lappend obs' obs'' \ obs' \ LNil \ ta_seq_consist P vs obs' \ (lfinite obs' \ (X (mrw_values P vs (list_of obs')) obs'' \ ta_seq_consist P (mrw_values P vs (list_of obs')) obs'')))" (is "\vs obs. _ \ _ \ ?step vs obs") shows "ta_seq_consist P vs obs" proof - from major have "\obs' obs''. obs = lappend (llist_of obs') obs'' \ ta_seq_consist P vs (llist_of obs') \ X (mrw_values P vs obs') obs''" by(auto intro: exI[where x="[]"]) thus ?thesis proof(coinduct) case (ta_seq_consist vs obs) then obtain obs' obs'' where obs: "obs = lappend (llist_of obs') obs''" and sc_obs': "ta_seq_consist P vs (llist_of obs')" and X: "X (mrw_values P vs obs') obs''" by blast show ?case proof(cases obs') case Nil with X have "X vs obs''" by simp from step[OF this] show ?thesis proof assume "obs'' = LNil" with Nil obs show ?thesis by simp next assume "?step vs obs''" then obtain obs''' obs'''' where obs'': "obs'' = lappend obs''' obs''''" and "obs''' \ LNil" and sc_obs''': "ta_seq_consist P vs obs'''" and fin: "lfinite obs''' \ X (mrw_values P vs (list_of obs''')) obs'''' \ ta_seq_consist P (mrw_values P vs (list_of obs''')) obs''''" by blast from \obs''' \ LNil\ obtain ob obs''''' where obs''': "obs''' = LCons ob obs'''''" unfolding neq_LNil_conv by blast with Nil obs'' obs have concl1: "obs = LCons ob (lappend obs''''' obs'''')" by simp have concl2: "case ob of NormalAction (ReadMem ad al v) \ \b. vs (ad, al) = \(v, b)\ | _ \ True" using sc_obs''' obs''' by simp show ?thesis proof(cases "lfinite obs'''") case False hence "lappend obs''''' obs'''' = obs'''''" using obs''' by(simp add: lappend_inf) hence "ta_seq_consist P (mrw_value P vs ob) (lappend obs''''' obs'''')" using sc_obs''' obs''' by simp with concl1 concl2 have ?LCons by blast thus ?thesis by simp next case True with obs''' obtain obs'''''' where obs''''': "obs''''' = llist_of obs''''''" by simp(auto simp add: lfinite_eq_range_llist_of) from fin[OF True] have "?LCons" proof assume X: "X (mrw_values P vs (list_of obs''')) obs''''" hence "X (mrw_values P (mrw_value P vs ob) obs'''''') obs''''" using obs''''' obs''' by simp moreover from obs''''' have "lappend obs''''' obs'''' = lappend (llist_of obs'''''') obs''''" by simp moreover have "ta_seq_consist P (mrw_value P vs ob) (llist_of obs'''''')" using sc_obs''' obs''' obs''''' by simp ultimately show ?thesis using concl1 concl2 by blast next assume "ta_seq_consist P (mrw_values P vs (list_of obs''')) obs''''" with sc_obs''' obs''''' obs''' have "ta_seq_consist P (mrw_value P vs ob) (lappend obs''''' obs'''')" by(simp add: ta_seq_consist_lappend) with concl1 concl2 show ?thesis by blast qed thus ?thesis by simp qed qed next case (Cons ob obs''') hence "obs = LCons ob (lappend (llist_of obs''') obs'')" using obs by simp moreover from sc_obs' Cons have "case ob of NormalAction (ReadMem ad al v) \ \b. vs (ad, al) = \(v, b)\ | _ \ True" and "ta_seq_consist P (mrw_value P vs ob) (llist_of obs''')" by simp_all moreover from X Cons have "X (mrw_values P (mrw_value P vs ob) obs''') obs''" by simp ultimately show ?thesis by blast qed qed qed lemma ta_seq_consist_coinduct_append_wf [consumes 2, case_names ta_seq_consist, case_conclusion ta_seq_consist LNil lappend]: assumes major: "X vs obs a" and wf: "wf R" and step: "\vs obs a. X vs obs a \ obs = LNil \ (\obs' obs'' a'. obs = lappend obs' obs'' \ ta_seq_consist P vs obs' \ (obs' = LNil \ (a', a) \ R) \ (lfinite obs' \ X (mrw_values P vs (list_of obs')) obs'' a' \ ta_seq_consist P (mrw_values P vs (list_of obs')) obs''))" (is "\vs obs a. _ \ _ \ ?step vs obs a") shows "ta_seq_consist P vs obs" proof - { fix vs obs a assume "X vs obs a" with wf have "obs = LNil \ (\obs' obs''. obs = lappend obs' obs'' \ obs' \ LNil \ ta_seq_consist P vs obs' \ (lfinite obs' \ (\a. X (mrw_values P vs (list_of obs')) obs'' a) \ ta_seq_consist P (mrw_values P vs (list_of obs')) obs''))" (is "_ \ ?step_concl vs obs") proof(induct a arbitrary: vs obs rule: wf_induct[consumes 1, case_names wf]) case (wf a) note IH = wf.hyps[rule_format] from step[OF \X vs obs a\] show ?case proof assume "obs = LNil" thus ?thesis .. next assume "?step vs obs a" then obtain obs' obs'' a' where obs: "obs = lappend obs' obs''" and sc_obs': "ta_seq_consist P vs obs'" and decr: "obs' = LNil \ (a', a) \ R" and fin: "lfinite obs' \ X (mrw_values P vs (list_of obs')) obs'' a' \ ta_seq_consist P (mrw_values P vs (list_of obs')) obs''" by blast show ?case proof(cases "obs' = LNil") case True hence "lfinite obs'" by simp from fin[OF this] show ?thesis proof assume X: "X (mrw_values P vs (list_of obs')) obs'' a'" from True have "(a', a) \ R" by(rule decr) from IH[OF this X] show ?thesis proof assume "obs'' = LNil" with True obs have "obs = LNil" by simp thus ?thesis .. next assume "?step_concl (mrw_values P vs (list_of obs')) obs''" hence "?step_concl vs obs" using True obs by simp thus ?thesis .. qed next assume "ta_seq_consist P (mrw_values P vs (list_of obs')) obs''" thus ?thesis using obs True by cases(auto cong: action.case_cong obs_event.case_cong intro: exI[where x="LCons x LNil" for x]) qed next case False with obs sc_obs' fin show ?thesis by auto qed qed qed } note step' = this from major show ?thesis proof(coinduction arbitrary: vs obs a rule: ta_seq_consist_coinduct_append) case (ta_seq_consist vs obs a) thus ?case by simp(rule step') qed qed lemma ta_seq_consist_nthI: "(\i ad al v. \ enat i < llength obs; lnth obs i = NormalAction (ReadMem ad al v); ta_seq_consist P vs (ltake (enat i) obs) \ \ \b. mrw_values P vs (list_of (ltake (enat i) obs)) (ad, al) = \(v, b)\) \ ta_seq_consist P vs obs" proof(coinduction arbitrary: vs obs) case (ta_seq_consist vs obs) hence nth: "\i ad al v. \ enat i < llength obs; lnth obs i = NormalAction (ReadMem ad al v); ta_seq_consist P vs (ltake (enat i) obs) \ \ \b. mrw_values P vs (list_of (ltake (enat i) obs)) (ad, al) = \(v, b)\" by blast show ?case proof(cases obs) case LNil thus ?thesis by simp next case (LCons ob obs') { fix ad al v assume "ob = NormalAction (ReadMem ad al v)" with nth[of 0 ad al v] LCons have "\b. vs (ad, al) = \(v, b)\" by(simp add: zero_enat_def[symmetric]) } note base = this moreover { fix i ad al v assume "enat i < llength obs'" "lnth obs' i = NormalAction (ReadMem ad al v)" and "ta_seq_consist P (mrw_value P vs ob) (ltake (enat i) obs')" with LCons nth[of "Suc i" ad al v] base have "\b. mrw_values P (mrw_value P vs ob) (list_of (ltake (enat i) obs')) (ad, al) = \(v, b)\" by(clarsimp simp add: eSuc_enat[symmetric] split: obs_event.split action.split) } ultimately have ?LCons using LCons by(simp split: action.split obs_event.split) thus ?thesis .. qed qed lemma ta_seq_consist_into_non_speculative: "\ ta_seq_consist P vs obs; \adal. set_option (vs adal) \ vs' adal \ UNIV \ \ non_speculative P vs' obs" proof(coinduction arbitrary: vs' obs vs) case (non_speculative vs' obs vs) thus ?case apply cases apply(auto split: action.split_asm obs_event.split_asm) apply(rule exI, erule conjI, auto)+ done qed lemma llist_of_list_of_append: "lfinite xs \ llist_of (list_of xs @ ys) = lappend xs (llist_of ys)" unfolding lfinite_eq_range_llist_of by(clarsimp simp add: lappend_llist_of_llist_of) lemma ta_seq_consist_most_recent_write_for: assumes sc: "ta_seq_consist P Map.empty (lmap snd E)" and read: "r \ read_actions E" and new_actions_for_fun: "\adal a a'. \ a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" shows "\i. P,E \ r \mrw i \ i < r" proof - from read obtain t v ad al where nth_r: "lnth E r = (t, NormalAction (ReadMem ad al v))" and r: "enat r < llength E" by(cases)(cases "lnth E r", auto simp add: action_obs_def actions_def) from nth_r r have E_unfold: "E = lappend (ltake (enat r) E) (LCons (t, NormalAction (ReadMem ad al v)) (ldropn (Suc r) E))" by (metis lappend_ltake_enat_ldropn ldropn_Suc_conv_ldropn) from sc obtain b where sc': "ta_seq_consist P Map.empty (ltake (enat r) (lmap snd E))" and mrw': "mrw_values P Map.empty (map snd (list_of (ltake (enat r) E))) (ad, al) = \(v, b)\" by(subst (asm) (3) E_unfold)(auto simp add: ta_seq_consist_lappend lmap_lappend_distrib) from mrw_values_mrw[OF mrw', of t] r obtain E' w' where E': "E' = llist_of (list_of (ltake (enat r) E) @ [(t, NormalAction (ReadMem ad al v))])" and v: "v = value_written P (ltake (enat r) E) w' (ad, al)" and mrw'': "P,E' \ r \mrw w'" and w': "w' < r" by(fastforce simp add: length_list_of_conv_the_enat min_def split: if_split_asm) from E' r have sim: "ltake (enat (Suc r)) E' [\] ltake (enat (Suc r)) E" by(subst E_unfold)(simp add: ltake_lappend llist_of_list_of_append min_def, auto simp add: eSuc_enat[symmetric] zero_enat_def[symmetric] eq_into_sim_actions) from nth_r have adal_r: "(ad, al) \ action_loc P E r" by(simp add: action_obs_def) from E' r have nth_r': "lnth E' r = (t, NormalAction (ReadMem ad al v))" by(auto simp add: nth_append length_list_of_conv_the_enat min_def) with mrw'' w' r adal_r obtain "E \ w' \a r" "w' \ write_actions E" "(ad, al) \ action_loc P E w'" by cases(fastforce simp add: action_obs_def action_loc_change_prefix[OF sim[symmetric], simplified action_obs_def] intro: action_order_change_prefix[OF _ sim] write_actions_change_prefix[OF _ sim]) with read adal_r have "P,E \ r \mrw w'" proof(rule most_recent_write_for.intros) fix wa' assume write': "wa' \ write_actions E" and adal_wa': "(ad, al) \ action_loc P E wa'" show "E \ wa' \a w' \ E \ r \a wa'" proof(cases "r \ wa'") assume "r \ wa'" show ?thesis proof(cases "is_new_action (action_obs E wa')") case False with \r \ wa'\ have "E \ r \a wa'" using read write' by(auto simp add: action_order_def elim!: read_actions.cases) thus ?thesis .. next case True with write' adal_wa' have "wa' \ new_actions_for P E (ad, al)" by(simp add: new_actions_for_def) hence "w' \ new_actions_for P E (ad, al)" using r w' \r \ wa'\ by(auto dest: new_actions_for_fun) with \w' \ write_actions E\ \(ad, al) \ action_loc P E w'\ have "\ is_new_action (action_obs E w')" by(simp add: new_actions_for_def) with write' True \w' \ write_actions E\ have "E \ wa' \a w'" by(simp add: action_order_def) thus ?thesis .. qed next assume "\ r \ wa'" hence "wa' < r" by simp with write' adal_wa' have "wa' \ write_actions E'" "(ad, al) \ action_loc P E' wa'" by(auto intro: write_actions_change_prefix[OF _ sim[symmetric]] simp add: action_loc_change_prefix[OF sim]) from most_recent_write_recent[OF mrw'' _ this] nth_r' have "E' \ wa' \a w' \ E' \ r \a wa'" by(simp add: action_obs_def) thus ?thesis using \wa' < r\ w' by(auto 4 3 del: disjCI intro: disjI1 disjI2 action_order_change_prefix[OF _ sim]) qed qed with w' show ?thesis by blast qed lemma ta_seq_consist_mrw_before: assumes sc: "ta_seq_consist P Map.empty (lmap snd E)" and new_actions_for_fun: "\adal a a'. \ a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" and mrw: "P,E \ r \mrw w" shows "w < r" proof - from mrw have "r \ read_actions E" by cases with sc new_actions_for_fun obtain w' where "P,E \ r \mrw w'" "w' < r" by(auto dest: ta_seq_consist_most_recent_write_for) with mrw show ?thesis by(auto dest: most_recent_write_for_fun) qed lemma ta_seq_consist_imp_sequentially_consistent: assumes tsa_ok: "thread_start_actions_ok E" and new_actions_for_fun: "\adal a a'. \ a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" and seq: "ta_seq_consist P Map.empty (lmap snd E)" shows "\ws. sequentially_consistent P (E, ws) \ P \ (E, ws) \" proof(intro exI conjI) define ws where "ws i = (THE w. P,E \ i \mrw w)" for i from seq have ns: "non_speculative P (\_. {}) (lmap snd E)" by(rule ta_seq_consist_into_non_speculative) simp show "sequentially_consistent P (E, ws)" unfolding ws_def proof(rule sequentially_consistentI) fix r assume "r \ read_actions E" with seq new_actions_for_fun obtain w where "P,E \ r \mrw w" by(auto dest: ta_seq_consist_most_recent_write_for) thus "P,E \ r \mrw THE w. P,E \ r \mrw w" by(simp add: THE_most_recent_writeI) qed show "P \ (E, ws) \" proof(rule wf_execI) show "is_write_seen P E ws" proof(rule is_write_seenI) fix a ad al v assume a: "a \ read_actions E" and adal: "action_obs E a = NormalAction (ReadMem ad al v)" from ns have seq': "non_speculative P (\_. {}) (ltake (enat a) (lmap snd E))" by(rule non_speculative_ltake) from seq a seq new_actions_for_fun obtain w where mrw: "P,E \ a \mrw w" and "w < a" by(auto dest: ta_seq_consist_most_recent_write_for) hence w: "ws a = w" by(simp add: ws_def THE_most_recent_writeI) with mrw adal show "ws a \ write_actions E" and "(ad, al) \ action_loc P E (ws a)" and "\ P,E \ a \hb ws a" by(fastforce elim!: most_recent_write_for.cases dest: happens_before_into_action_order antisymPD[OF antisym_action_order] read_actions_not_write_actions)+ let ?between = "ltake (enat (a - Suc w)) (ldropn (Suc w) E)" let ?prefix = "ltake (enat w) E" let ?vs_prefix = "mrw_values P Map.empty (map snd (list_of ?prefix))" { fix v' assume new: "is_new_action (action_obs E w)" and vs': "?vs_prefix (ad, al) = \(v', True)\" from mrw_values_eq_SomeD[OF vs'] obtain obs' wa obs'' where split: "map snd (list_of ?prefix) = obs' @ wa # obs''" and wa: "is_write_action wa" and adal': "(ad, al) \ action_loc_aux P wa" and new_wa: "\ is_new_action wa" by blast from split have "length (map snd (list_of ?prefix)) = Suc (length obs' + length obs'')" by simp hence len_prefix: "llength ?prefix = enat \" by(simp add: length_list_of_conv_the_enat min_enat1_conv_enat) with split have "nth (map snd (list_of ?prefix)) (length obs') = wa" and "enat (length obs') < llength ?prefix" by simp_all hence "snd (lnth ?prefix (length obs')) = wa" by(simp add: list_of_lmap[symmetric] del: list_of_lmap) hence wa': "action_obs E (length obs') = wa" and "enat (length obs') < llength E" using \enat (length obs') < llength ?prefix\ by(auto simp add: action_obs_def lnth_ltake) with wa have "length obs' \ write_actions E" by(auto intro: write_actions.intros simp add: actions_def) from most_recent_write_recent[OF mrw _ this, of "(ad, al)"] adal adal' wa' have "E \ length obs' \a w \ E \ a \a length obs'" by simp hence False using new_wa new wa' adal len_prefix \w < a\ by(auto elim!: action_orderE simp add: min_enat1_conv_enat split: enat.split_asm) } hence mrw_value_w: "mrw_value P ?vs_prefix (snd (lnth E w)) (ad, al) = \(value_written P E w (ad, al), \ is_new_action (action_obs E w))\" using \ws a \ write_actions E\ \(ad, al) \ action_loc P E (ws a)\ w by(cases "snd (lnth E w)" rule: mrw_value_cases)(fastforce elim: write_actions.cases simp add: value_written_def action_obs_def)+ have "mrw_values P (mrw_value P ?vs_prefix (snd (lnth E w))) (list_of (lmap snd ?between)) (ad, al) = \(value_written P E w (ad, al), \ is_new_action (action_obs E w))\" proof(subst mrw_values_no_write_unchanged) fix wa assume "wa \ set (list_of (lmap snd ?between))" and write_wa: "is_write_action wa" and adal_wa: "(ad, al) \ action_loc_aux P wa" hence wa: "wa \ lset (lmap snd ?between)" by simp from wa obtain i_wa where "wa = lnth (lmap snd ?between) i_wa" and i_wa: "enat i_wa < llength (lmap snd ?between)" unfolding lset_conv_lnth by blast moreover hence i_wa_len: "enat (Suc (w + i_wa)) < llength E" by(cases "llength E") auto ultimately have wa': "wa = action_obs E (Suc (w + i_wa))" by(simp_all add: lnth_ltake action_obs_def ac_simps) with write_wa i_wa_len have "Suc (w + i_wa) \ write_actions E" by(auto intro: write_actions.intros simp add: actions_def) from most_recent_write_recent[OF mrw _ this, of "(ad, al)"] adal adal_wa wa' have "E \ Suc (w + i_wa) \a w \ E \ a \a Suc (w + i_wa)" by(simp) hence "is_new_action wa \ \ is_new_action (action_obs E w)" using adal i_wa wa' by(auto elim: action_orderE) thus "case (mrw_value P ?vs_prefix (snd (lnth E w)) (ad, al)) of None \ False | Some (v, b) \ b \ is_new_action wa" unfolding mrw_value_w by simp qed(simp add: mrw_value_w) moreover from a have "a \ actions E" by simp hence "enat a < llength E" by(rule actionsE) with \w < a\ have "enat (a - Suc w) < llength E - enat (Suc w)" by(cases "llength E") simp_all hence "E = lappend (lappend ?prefix (LCons (lnth E w) ?between)) (LCons (lnth (ldropn (Suc w) E) (a - Suc w)) (ldropn (Suc (a - Suc w)) (ldropn (Suc w) E)))" using \w < a\ \enat a < llength E\ unfolding lappend_assoc lappend_code apply(subst ldropn_Suc_conv_ldropn, simp) apply(subst lappend_ltake_enat_ldropn) apply(subst ldropn_Suc_conv_ldropn, simp add: less_trans[where y="enat a"]) by simp hence E': "E = lappend (lappend ?prefix (LCons (lnth E w) ?between)) (LCons (lnth E a) (ldropn (Suc a) E))" using \w < a\ \enat a < llength E\ by simp from seq have "ta_seq_consist P (mrw_values P Map.empty (list_of (lappend (lmap snd ?prefix) (LCons (snd (lnth E w)) (lmap snd ?between))))) (lmap snd (LCons (lnth E a) (ldropn (Suc a) E)))" by(subst (asm) E')(simp add: lmap_lappend_distrib ta_seq_consist_lappend) ultimately show "value_written P E (ws a) (ad, al) = v" using adal w by(clarsimp simp add: action_obs_def list_of_lappend list_of_LCons) (* assume "is_volatile P al" *) show "\ P,E \ a \so ws a" using \w < a\ w adal by(auto elim!: action_orderE sync_orderE) fix a' assume a': "a' \ write_actions E" "(ad, al) \ action_loc P E a'" { presume "E \ ws a \a a'" "E \ a' \a a" with mrw adal a' have "a' = ws a" unfolding w by cases(fastforce dest: antisymPD[OF antisym_action_order] read_actions_not_write_actions elim!: meta_allE[where x=a']) thus "a' = ws a" "a' = ws a" by - next assume "P,E \ ws a \hb a'" "P,E \ a' \hb a" thus "E \ ws a \a a'" "E \ a' \a a" using a' by(blast intro: happens_before_into_action_order)+ next assume "is_volatile P al" "P,E \ ws a \so a'" "P,E \ a' \so a" thus "E \ ws a \a a'" "E \ a' \a a" by(auto elim: sync_orderE) } qed qed(rule tsa_ok) qed subsection \Cut-and-update and sequentially consistent completion\ inductive foldl_list_all2 :: "('b \ 'c \ 'a \ 'a) \ ('b \ 'c \ 'a \ bool) \ ('b \ 'c \ 'a \ bool) \ 'b list \ 'c list \ 'a \ bool" for f and P and Q where "foldl_list_all2 f P Q [] [] s" | "\ Q x y s; P x y s \ foldl_list_all2 f P Q xs ys (f x y s) \ \ foldl_list_all2 f P Q (x # xs) (y # ys) s" inductive_simps foldl_list_all2_simps [simp]: "foldl_list_all2 f P Q [] ys s" "foldl_list_all2 f P Q xs [] s" "foldl_list_all2 f P Q (x # xs) (y # ys) s" inductive_simps foldl_list_all2_Cons1: "foldl_list_all2 f P Q (x # xs) ys s" inductive_simps foldl_list_all2_Cons2: "foldl_list_all2 f P Q xs (y # ys) s" definition eq_upto_seq_inconsist :: "'m prog \ ('addr, 'thread_id) obs_event action list \ ('addr, 'thread_id) obs_event action list \ ('addr \ addr_loc \ 'addr val \ bool) \ bool" where "eq_upto_seq_inconsist P = foldl_list_all2 (\ob ob' vs. mrw_value P vs ob) (\ob ob' vs. case ob of NormalAction (ReadMem ad al v) \ \b. vs (ad, al) = Some (v, b) | _ \ True) (\ob ob' vs. if (case ob of NormalAction (ReadMem ad al v) \ \b. vs (ad, al) = Some (v, b) | _ \ True) then ob = ob' else ob \ ob')" lemma eq_upto_seq_inconsist_simps: "eq_upto_seq_inconsist P [] obs' vs \ obs' = []" "eq_upto_seq_inconsist P obs [] vs \ obs = []" "eq_upto_seq_inconsist P (ob # obs) (ob' # obs') vs \ (case ob of NormalAction (ReadMem ad al v) \ if (\b. vs (ad, al) = \(v, b)\) then ob = ob' \ eq_upto_seq_inconsist P obs obs' (mrw_value P vs ob) else ob \ ob' | _ \ ob = ob' \ eq_upto_seq_inconsist P obs obs' (mrw_value P vs ob))" by(auto simp add: eq_upto_seq_inconsist_def split: action.split obs_event.split) lemma eq_upto_seq_inconsist_Cons1: "eq_upto_seq_inconsist P (ob # obs) obs' vs \ (\ob' obs''. obs' = ob' # obs'' \ (case ob of NormalAction (ReadMem ad al v) \ if (\b. vs (ad, al) = \(v, b)\) then ob' = ob \ eq_upto_seq_inconsist P obs obs'' (mrw_value P vs ob) else ob \ ob' | _ \ ob' = ob \ eq_upto_seq_inconsist P obs obs'' (mrw_value P vs ob)))" unfolding eq_upto_seq_inconsist_def by(auto split: obs_event.split action.split simp add: foldl_list_all2_Cons1) lemma eq_upto_seq_inconsist_appendD: assumes "eq_upto_seq_inconsist P (obs @ obs') obs'' vs" and "ta_seq_consist P vs (llist_of obs)" shows "length obs \ length obs''" (is ?thesis1) and "take (length obs) obs'' = obs" (is ?thesis2) and "eq_upto_seq_inconsist P obs' (drop (length obs) obs'') (mrw_values P vs obs)" (is ?thesis3) using assms by(induct obs arbitrary: obs'' vs)(auto split: action.split_asm obs_event.split_asm simp add: eq_upto_seq_inconsist_Cons1) lemma ta_seq_consist_imp_eq_upto_seq_inconsist_refl: "ta_seq_consist P vs (llist_of obs) \ eq_upto_seq_inconsist P obs obs vs" apply(induct obs arbitrary: vs) apply(auto simp add: eq_upto_seq_inconsist_simps split: action.split obs_event.split) done context notes split_paired_Ex [simp del] eq_upto_seq_inconsist_simps [simp] begin lemma eq_upto_seq_inconsist_appendI: "\ eq_upto_seq_inconsist P obs OBS vs; \ ta_seq_consist P vs (llist_of obs) \ \ eq_upto_seq_inconsist P obs' OBS' (mrw_values P vs OBS) \ \ eq_upto_seq_inconsist P (obs @ obs') (OBS @ OBS') vs" apply(induct obs arbitrary: vs OBS) apply simp apply(auto simp add: eq_upto_seq_inconsist_Cons1) apply(simp split: action.split obs_event.split) apply auto done lemma eq_upto_seq_inconsist_trans: "\ eq_upto_seq_inconsist P obs obs' vs; eq_upto_seq_inconsist P obs' obs'' vs \ \ eq_upto_seq_inconsist P obs obs'' vs" apply(induction obs arbitrary: obs' obs'' vs) apply(clarsimp simp add: eq_upto_seq_inconsist_Cons1)+ apply(auto split!: action.split obs_event.split if_split_asm) done lemma eq_upto_seq_inconsist_append2: "\ eq_upto_seq_inconsist P obs obs' vs; \ ta_seq_consist P vs (llist_of obs) \ \ eq_upto_seq_inconsist P obs (obs' @ obs'') vs" apply(induction obs arbitrary: obs' vs) apply(clarsimp simp add: eq_upto_seq_inconsist_Cons1)+ apply(auto split!: action.split obs_event.split if_split_asm) done end context executions_sc_hb begin lemma ta_seq_consist_mrwI: assumes E: "E \ \" and wf: "P \ (E, ws) \" and mrw: "\a. \ enat a < r; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "ta_seq_consist P Map.empty (lmap snd (ltake r E))" proof(rule ta_seq_consist_nthI) fix i ad al v assume i_len: "enat i < llength (lmap snd (ltake r E))" and E_i: "lnth (lmap snd (ltake r E)) i = NormalAction (ReadMem ad al v)" and sc: "ta_seq_consist P Map.empty (ltake (enat i) (lmap snd (ltake r E)))" from i_len have "enat i < r" by simp with sc have "ta_seq_consist P Map.empty (ltake (enat i) (lmap snd E))" by(simp add: min_def split: if_split_asm) hence ns: "non_speculative P (\_. {}) (ltake (enat i) (lmap snd E))" by(rule ta_seq_consist_into_non_speculative) simp from i_len have "i \ actions E" by(simp add: actions_def) moreover from E_i i_len have obs_i: "action_obs E i = NormalAction (ReadMem ad al v)" by(simp add: action_obs_def lnth_ltake) ultimately have read: "i \ read_actions E" .. with i_len have mrw_i: "P,E \ i \mrw ws i" by(auto intro: mrw) with E have "ws i < i" using ns by(rule mrw_before) from mrw_i obs_i obtain adal_w: "(ad, al) \ action_loc P E (ws i)" and adal_r: "(ad, al) \ action_loc P E i" and "write": "ws i \ write_actions E" by cases auto from wf have "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) from is_write_seenD[OF this read obs_i] have vw_v: "value_written P E (ws i) (ad, al) = v" by simp let ?vs = "mrw_values P Map.empty (map snd (list_of (ltake (enat (ws i)) E)))" from \ws i < i\ i_len have "enat (ws i) < llength (ltake (enat i) E)" by(simp add: less_trans[where y="enat i"]) hence "ltake (enat i) E = lappend (ltake (enat (ws i)) (ltake (enat i) E)) (LCons (lnth (ltake (enat i) E) (ws i)) (ldropn (Suc (ws i)) (ltake (enat i) E)))" by(simp only: ldropn_Suc_conv_ldropn lappend_ltake_enat_ldropn) also have "\ = lappend (ltake (enat (ws i)) E) (LCons (lnth E (ws i)) (ldropn (Suc (ws i)) (ltake (enat i) E)))" using \ws i < i\ i_len \enat (ws i) < llength (ltake (enat i) E)\ - by(simp add: lnth_ltake)(simp add: min_def) + by (simp add: lnth_ltake) finally have r_E: "ltake (enat i) E = \" . have "mrw_values P Map.empty (list_of (ltake (enat i) (lmap snd (ltake r E)))) (ad, al) = mrw_values P Map.empty (map snd (list_of (ltake (enat i) E))) (ad, al)" using \enat i < r\ by(auto simp add: min_def) also have "\ = mrw_values P (mrw_value P ?vs (snd (lnth E (ws i)))) (map snd (list_of (ldropn (Suc (ws i)) (ltake (enat i) E)))) (ad, al)" by(subst r_E)(simp add: list_of_lappend) also have "\ = mrw_value P ?vs (snd (lnth E (ws i))) (ad, al)" proof(rule mrw_values_no_write_unchanged) fix wa assume wa: "wa \ set (map snd (list_of (ldropn (Suc (ws i)) (ltake (enat i) E))))" and "is_write_action wa" "(ad, al) \ action_loc_aux P wa" from wa obtain w where "w < length (map snd (list_of (ldropn (Suc (ws i)) (ltake (enat i) E))))" and "map snd (list_of (ldropn (Suc (ws i)) (ltake (enat i) E))) ! w = wa" unfolding in_set_conv_nth by blast moreover hence "Suc (ws i + w) < i" (is "?w < _") using i_len by(cases "llength E")(simp_all add: length_list_of_conv_the_enat) ultimately have obs_w': "action_obs E ?w = wa" using i_len by(simp add: action_obs_def lnth_ltake less_trans[where y="enat i"] ac_simps) from \?w < i\ i_len have "?w \ actions E" by(simp add: actions_def less_trans[where y="enat i"]) with \is_write_action wa\ obs_w' \(ad, al) \ action_loc_aux P wa\ have write': "?w \ write_actions E" and adal': "(ad, al) \ action_loc P E ?w" by(auto intro: write_actions.intros) from \?w < i\ \i \ read_actions E\ \?w \ actions E\ have "E \ ?w \a i" by(auto simp add: action_order_def elim: read_actions.cases) from mrw_i adal_r write' adal' have "E \ ?w \a ws i \ E \ i \a ?w" by(rule most_recent_write_recent) hence "E \ ?w \a ws i" proof assume "E \ i \a ?w" with \E \ ?w \a i\ have "?w = i" by(rule antisymPD[OF antisym_action_order]) with write' read have False by(auto dest: read_actions_not_write_actions) thus ?thesis .. qed from adal_w "write" have "mrw_value P ?vs (snd (lnth E (ws i))) (ad, al) \ None" by(cases "snd (lnth E (ws i))" rule: mrw_value_cases) (auto simp add: action_obs_def split: if_split_asm elim: write_actions.cases) then obtain b v where vb: "mrw_value P ?vs (snd (lnth E (ws i))) (ad, al) = Some (v, b)" by auto moreover from \E \ ?w \a ws i\ obs_w' have "is_new_action wa" "\ is_new_action (action_obs E (ws i))" by(auto elim!: action_orderE) from \\ is_new_action (action_obs E (ws i))\ "write" adal_w obtain v' where "action_obs E (ws i) = NormalAction (WriteMem ad al v')" by(auto elim!: write_actions.cases is_write_action.cases) with vb have b by(simp add: action_obs_def) with \is_new_action wa\ vb show "case mrw_value P ?vs (snd (lnth E (ws i))) (ad, al) of None \ False | \(v, b)\ \ b \ is_new_action wa" by simp qed also { fix v assume "?vs (ad, al) = Some (v, True)" and "is_new_action (action_obs E (ws i))" from mrw_values_eq_SomeD[OF this(1)] obtain wa where "wa \ set (map snd (list_of (ltake (enat (ws i)) E)))" and "is_write_action wa" and "(ad, al) \ action_loc_aux P wa" and "\ is_new_action wa" by(fastforce simp del: set_map) moreover then obtain w where w: "w < ws i" and wa: "wa = snd (lnth E w)" unfolding in_set_conv_nth by(cases "llength E")(auto simp add: lnth_ltake length_list_of_conv_the_enat) ultimately have "w \ write_actions E" "action_obs E w = wa" "(ad, al) \ action_loc P E w" using \ws i \ write_actions E\ by(auto intro!: write_actions.intros simp add: actions_def less_trans[where y="enat (ws i)"] action_obs_def elim!: write_actions.cases) with mrw_i adal_r have "E \ w \a ws i \ E \ i \a w" by -(rule most_recent_write_recent) hence False proof assume "E \ w \a ws i" moreover from \\ is_new_action wa\ \is_new_action (action_obs E (ws i))\ "write" w wa \w \ write_actions E\ have "E \ ws i \a w" by(auto simp add: action_order_def action_obs_def) ultimately have "w = ws i" by(rule antisymPD[OF antisym_action_order]) with \w < ws i\ show False by simp next assume "E \ i \a w" moreover from \w \ write_actions E\ \w < ws i\ \ws i < i\ read have "E \ w \a i" by(auto simp add: action_order_def elim: read_actions.cases) ultimately have "i = w" by(rule antisymPD[OF antisym_action_order]) with \w < ws i\ \ws i < i\ show False by simp qed } then obtain b where "\ = Some (v, b)" using vw_v "write" adal_w apply(atomize_elim) apply(auto simp add: action_obs_def value_written_def write_actions_iff) apply(erule is_write_action.cases) apply auto done finally show "\b. mrw_values P Map.empty (list_of (ltake (enat i) (lmap snd (ltake r E)))) (ad, al) = \(v, b)\" by blast qed end context jmm_multithreaded begin definition complete_sc :: "('l,'thread_id,'x,'m,'w) state \ ('addr \ addr_loc \ 'addr val \ bool) \ ('thread_id \ ('l, 'thread_id, 'x, 'm, 'w, ('addr, 'thread_id) obs_event action) thread_action) llist" where "complete_sc s vs = unfold_llist (\(s, vs). \t ta s'. \ s -t\ta\ s') (\(s, vs). fst (SOME ((t, ta), s'). s -t\ta\ s' \ ta_seq_consist P vs (llist_of \ta\\<^bsub>o\<^esub>))) (\(s, vs). let ((t, ta), s') = SOME ((t, ta), s'). s -t\ta\ s' \ ta_seq_consist P vs (llist_of \ta\\<^bsub>o\<^esub>) in (s', mrw_values P vs \ta\\<^bsub>o\<^esub>)) (s, vs)" definition sc_completion :: "('l, 'thread_id, 'x, 'm, 'w) state \ ('addr \ addr_loc \ 'addr val \ bool) \ bool" where "sc_completion s vs \ (\ttas s' t x ta x' m'. s -\ttas\* s' \ ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) \ thr s' t = \(x, no_wait_locks)\ \ t \ (x, shr s') -ta\ (x', m') \ actions_ok s' t ta \ (\ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>)))" lemma sc_completionD: "\ sc_completion s vs; s -\ttas\* s'; ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))); thr s' t = \(x, no_wait_locks)\; t \ (x, shr s') -ta\ (x', m'); actions_ok s' t ta \ \ \ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>)" unfolding sc_completion_def by blast lemma sc_completionI: "(\ttas s' t x ta x' m'. \ s -\ttas\* s'; ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))); thr s' t = \(x, no_wait_locks)\; t \ (x, shr s') -ta\ (x', m'); actions_ok s' t ta \ \ \ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>)) \ sc_completion s vs" unfolding sc_completion_def by blast lemma sc_completion_shift: assumes sc_c: "sc_completion s vs" and \Red: "s -\ttas\* s'" and sc: "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (llist_of ttas)))" shows "sc_completion s' (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" proof(rule sc_completionI) fix ttas' s'' t x ta x' m' assume \Red': "s' -\ttas'\* s''" and sc': "ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')))" and red: "thr s'' t = \(x, no_wait_locks)\" "t \ \x, shr s''\ -ta\ \x', m'\" "actions_ok s'' t ta" from \Red \Red' have "s -\ttas @ ttas'\* s''" unfolding RedT_def by(rule rtrancl3p_trans) moreover from sc sc' have "ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (ttas @ ttas'))))" apply(simp add: lappend_llist_of_llist_of[symmetric] ta_seq_consist_lappend del: lappend_llist_of_llist_of) apply(simp add: lconcat_llist_of[symmetric] lmap_llist_of[symmetric] llist.map_comp o_def split_def del: lmap_llist_of) done ultimately show "\ta' x'' m''. t \ \x, shr s''\ -ta'\ \x'', m''\ \ actions_ok s'' t ta' \ ta_seq_consist P (mrw_values P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas'))) (llist_of \ta'\\<^bsub>o\<^esub>)" using red unfolding foldl_append[symmetric] concat_append[symmetric] map_append[symmetric] by(rule sc_completionD[OF sc_c]) qed lemma complete_sc_in_Runs: assumes cau: "sc_completion s vs" and ta_seq_consist_convert_RA: "\vs ln. ta_seq_consist P vs (llist_of (convert_RA ln))" shows "mthr.Runs s (complete_sc s vs)" proof - let ?ttas' = "\ttas' :: ('thread_id \ ('l,'thread_id,'x,'m,'w, ('addr, 'thread_id) obs_event action) thread_action) list. concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')" let "?vs ttas'" = "mrw_values P vs (?ttas' ttas')" define s' vs' and ttas :: "('thread_id \ ('l,'thread_id,'x,'m,'w, ('addr, 'thread_id) obs_event action) thread_action) list" where "s' = s" and "vs' = vs" and "ttas = []" hence "s -\ttas\* s'" "ta_seq_consist P vs (llist_of (?ttas' ttas))" by auto hence "mthr.Runs s' (complete_sc s' (?vs ttas))" proof(coinduction arbitrary: s' ttas rule: mthr.Runs.coinduct) case (Runs s' ttas') note Red = \s -\ttas'\* s'\ and sc = \ta_seq_consist P vs (llist_of (?ttas' ttas'))\ show ?case proof(cases "\t' ta' s''. s' -t'\ta'\ s''") case False hence ?Stuck by(simp add: complete_sc_def) thus ?thesis .. next case True let ?proceed = "\((t', ta'), s''). s' -t'\ta'\ s'' \ ta_seq_consist P (?vs ttas') (llist_of \ta'\\<^bsub>o\<^esub>)" from True obtain t' ta' s'' where red: "s' -t'\ta'\ s''" by(auto) then obtain ta'' s''' where "s' -t'\ta''\ s'''" and "ta_seq_consist P (?vs ttas') (llist_of \ta''\\<^bsub>o\<^esub>)" proof(cases) case (redT_normal x x' m') note red = \t' \ \x, shr s'\ -ta'\ \x', m'\\ and ts''t' = \thr s' t' = \(x, no_wait_locks)\\ and aok = \actions_ok s' t' ta'\ and s'' = \redT_upd s' t' ta' x' m' s''\ from sc_completionD[OF cau Red sc ts''t' red aok] obtain ta'' x'' m'' where red': "t' \ \x, shr s'\ -ta''\ \x'', m''\" and aok': "actions_ok s' t' ta''" and sc': "ta_seq_consist P (?vs ttas') (llist_of \ta''\\<^bsub>o\<^esub>)" by blast from redT_updWs_total obtain ws' where "redT_updWs t' (wset s') \ta''\\<^bsub>w\<^esub> ws'" .. then obtain s''' where "redT_upd s' t' ta'' x'' m'' s'''" by fastforce with red' ts''t' aok' have "s' -t'\ta''\ s'''" .. thus thesis using sc' by(rule that) next case redT_acquire thus thesis by(simp add: that[OF red] ta_seq_consist_convert_RA) qed hence "?proceed ((t', ta''), s''')" using Red by(auto) hence *: "?proceed (Eps ?proceed)" by(rule someI) moreover from Red * have "s -\ttas' @ [fst (Eps ?proceed)]\* snd (Eps ?proceed)" by(auto simp add: split_beta RedT_def intro: rtrancl3p_step) moreover from True have "complete_sc s' (?vs ttas') = LCons (fst (Eps ?proceed)) (complete_sc (snd (Eps ?proceed)) (?vs (ttas' @ [fst (Eps ?proceed)])))" unfolding complete_sc_def by(simp add: split_def) moreover from sc \?proceed (Eps ?proceed)\ have "ta_seq_consist P vs (llist_of (?ttas' (ttas' @ [fst (Eps ?proceed)])))" unfolding map_append concat_append lappend_llist_of_llist_of[symmetric] by(subst ta_seq_consist_lappend)(auto simp add: split_def) ultimately have ?Step by(fastforce intro: exI[where x="ttas' @ [fst (Eps ?proceed)]"] simp del: split_paired_Ex) thus ?thesis by simp qed qed thus ?thesis by(simp add: s'_def ttas_def) qed lemma complete_sc_ta_seq_consist: assumes cau: "sc_completion s vs" and ta_seq_consist_convert_RA: "\vs ln. ta_seq_consist P vs (llist_of (convert_RA ln))" shows "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (complete_sc s vs)))" proof - define vs' where "vs' = vs" let ?obs = "\ttas. lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) ttas)" define obs where "obs = ?obs (complete_sc s vs)" define a where "a = complete_sc s vs'" let ?ttas' = "\ttas' :: ('thread_id \ ('l,'thread_id,'x,'m,'w,('addr, 'thread_id) obs_event action) thread_action) list. concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas')" let ?vs = "\ttas'. mrw_values P vs (?ttas' ttas')" from vs'_def obs_def have "s -\[]\* s" "ta_seq_consist P vs (llist_of (?ttas' []))" "vs' = ?vs []" by(auto) hence "\s' ttas'. obs = ?obs (complete_sc s' vs') \ s -\ttas'\* s' \ ta_seq_consist P vs (llist_of (?ttas' ttas')) \ vs' = ?vs ttas' \ a = complete_sc s' vs'" unfolding obs_def vs'_def a_def by metis moreover have "wf (inv_image {(m, n). m < n} (llength \ ltakeWhile (\(t, ta). \ta\\<^bsub>o\<^esub> = [])))" (is "wf ?R") by(rule wf_inv_image)(rule wellorder_class.wf) ultimately show "ta_seq_consist P vs' obs" proof(coinduct vs' obs a rule: ta_seq_consist_coinduct_append_wf) case (ta_seq_consist vs' obs a) then obtain s' ttas' where obs_def: "obs = ?obs (complete_sc s' (?vs ttas'))" and Red: "s -\ttas'\* s'" and sc: "ta_seq_consist P vs (llist_of (?ttas' ttas'))" and vs'_def: "vs' = ?vs ttas'" and a_def: "a = complete_sc s' vs'" by blast show ?case proof(cases "\t' ta' s''. s' -t'\ta'\ s''") case False hence "obs = LNil" unfolding obs_def complete_sc_def by simp hence ?LNil unfolding obs_def by auto thus ?thesis .. next case True let ?proceed = "\((t', ta'), s''). s' -t'\ta'\ s'' \ ta_seq_consist P (?vs ttas') (llist_of \ta'\\<^bsub>o\<^esub>)" let ?tta = "fst (Eps ?proceed)" let ?s' = "snd (Eps ?proceed)" from True obtain t' ta' s'' where red: "s' -t'\ta'\ s''" by blast then obtain ta'' s''' where "s' -t'\ta''\ s'''" and "ta_seq_consist P (?vs ttas') (llist_of \ta''\\<^bsub>o\<^esub>)" proof(cases) case (redT_normal x x' m') note red = \t' \ \x, shr s'\ -ta'\ \x', m'\\ and ts''t' = \thr s' t' = \(x, no_wait_locks)\\ and aok = \actions_ok s' t' ta'\ and s''' = \redT_upd s' t' ta' x' m' s''\ from sc_completionD[OF cau Red sc ts''t' red aok] obtain ta'' x'' m'' where red': "t' \ \x, shr s'\ -ta''\ \x'', m''\" and aok': "actions_ok s' t' ta''" and sc': "ta_seq_consist P (?vs ttas') (llist_of \ta''\\<^bsub>o\<^esub>)" by blast from redT_updWs_total obtain ws' where "redT_updWs t' (wset s') \ta''\\<^bsub>w\<^esub> ws'" .. then obtain s''' where "redT_upd s' t' ta'' x'' m'' s'''" by fastforce with red' ts''t' aok' have "s' -t'\ta''\ s'''" .. thus thesis using sc' by(rule that) next case redT_acquire thus thesis by(simp add: that[OF red] ta_seq_consist_convert_RA) qed hence "?proceed ((t', ta''), s''')" by auto hence "?proceed (Eps ?proceed)" by(rule someI) show ?thesis proof(cases "obs = LNil") case True thus ?thesis .. next case False from True have csc_unfold: "complete_sc s' (?vs ttas') = LCons ?tta (complete_sc ?s' (?vs (ttas' @ [?tta])))" unfolding complete_sc_def by(simp add: split_def) hence "obs = lappend (llist_of \snd ?tta\\<^bsub>o\<^esub>) (?obs (complete_sc ?s' (?vs (ttas' @ [?tta]))))" using obs_def by(simp add: split_beta) moreover have "ta_seq_consist P vs' (llist_of \snd ?tta\\<^bsub>o\<^esub>)" using \?proceed (Eps ?proceed)\ vs'_def by(clarsimp simp add: split_beta) moreover { assume "llist_of \snd ?tta\\<^bsub>o\<^esub> = LNil" moreover from obs_def \obs \ LNil\ have "lfinite (ltakeWhile (\(t, ta). \ta\\<^bsub>o\<^esub> = []) (complete_sc s' (?vs ttas')))" unfolding lfinite_ltakeWhile by(fastforce simp add: split_def lconcat_eq_LNil) ultimately have "(complete_sc ?s' (?vs (ttas' @ [?tta])), a) \ ?R" unfolding a_def vs'_def csc_unfold by(clarsimp simp add: split_def llist_of_eq_LNil_conv)(auto simp add: lfinite_eq_range_llist_of) } moreover have "?obs (complete_sc ?s' (?vs (ttas' @ [?tta]))) = ?obs (complete_sc ?s' (mrw_values P vs' (list_of (llist_of \snd ?tta\\<^bsub>o\<^esub>))))" unfolding vs'_def by(simp add: split_def) moreover from \?proceed (Eps ?proceed)\ Red have "s -\ttas' @ [?tta]\* ?s'" by(auto simp add: RedT_def split_def intro: rtrancl3p_step) moreover from sc \?proceed (Eps ?proceed)\ have "ta_seq_consist P vs (llist_of (?ttas' (ttas' @ [?tta])))" by(clarsimp simp add: split_def ta_seq_consist_lappend lappend_llist_of_llist_of[symmetric] simp del: lappend_llist_of_llist_of) moreover have "mrw_values P vs' (list_of (llist_of \snd ?tta\\<^bsub>o\<^esub>)) = ?vs (ttas' @ [?tta])" unfolding vs'_def by(simp add: split_def) moreover have "complete_sc ?s' (?vs (ttas' @ [?tta])) = complete_sc ?s' (mrw_values P vs' (list_of (llist_of \snd ?tta\\<^bsub>o\<^esub>)))" unfolding vs'_def by(simp add: split_def) ultimately have "?lappend" by blast thus ?thesis .. qed qed qed qed lemma sequential_completion_Runs: assumes "sc_completion s vs" and "\vs ln. ta_seq_consist P vs (llist_of (convert_RA ln))" shows "\ttas. mthr.Runs s ttas \ ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) ttas))" using complete_sc_ta_seq_consist[OF assms] complete_sc_in_Runs[OF assms] by blast definition cut_and_update :: "('l, 'thread_id, 'x, 'm, 'w) state \ ('addr \ addr_loc \ 'addr val \ bool) \ bool" where "cut_and_update s vs \ (\ttas s' t x ta x' m'. s -\ttas\* s' \ ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) \ thr s' t = \(x, no_wait_locks)\ \ t \ (x, shr s') -ta\ (x', m') \ actions_ok s' t ta \ (\ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>) \ eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))))" lemma cut_and_updateI[intro?]: "(\ttas s' t x ta x' m'. \ s -\ttas\* s'; ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))); thr s' t = \(x, no_wait_locks)\; t \ (x, shr s') -ta\ (x', m'); actions_ok s' t ta \ \ \ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>) \ eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))) \ cut_and_update s vs" unfolding cut_and_update_def by blast lemma cut_and_updateD: "\ cut_and_update s vs; s -\ttas\* s'; ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))); thr s' t = \(x, no_wait_locks)\; t \ (x, shr s') -ta\ (x', m'); actions_ok s' t ta \ \ \ta' x'' m''. t \ (x, shr s') -ta'\ (x'', m'') \ actions_ok s' t ta' \ ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>) \ eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" unfolding cut_and_update_def by blast lemma cut_and_update_imp_sc_completion: "cut_and_update s vs \ sc_completion s vs" apply(rule sc_completionI) apply(drule (5) cut_and_updateD) apply blast done lemma sequential_completion: assumes cut_and_update: "cut_and_update s vs" and ta_seq_consist_convert_RA: "\vs ln. ta_seq_consist P vs (llist_of (convert_RA ln))" and Red: "s -\ttas\* s'" and sc: "ta_seq_consist P vs (llist_of (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and red: "s' -t\ta\ s''" shows "\ta' ttas'. mthr.Runs s' (LCons (t, ta') ttas') \ ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (lappend (llist_of ttas) (LCons (t, ta') ttas')))) \ eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" proof - from red obtain ta' s''' where red': "redT s' (t, ta') s'''" and sc': "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (lappend (llist_of ttas) (LCons (t, ta') LNil))))" and eq: "eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" proof cases case (redT_normal x x' m') note ts't = \thr s' t = \(x, no_wait_locks)\\ and red = \t \ \x, shr s'\ -ta\ \x', m'\\ and aok = \actions_ok s' t ta\ and s'' = \redT_upd s' t ta x' m' s''\ from cut_and_updateD[OF cut_and_update, OF Red sc ts't red aok] obtain ta' x'' m'' where red: "t \ \x, shr s'\ -ta'\ \x'', m''\" and sc': "ta_seq_consist P (mrw_values P vs (concat (map (\(t, y). \y\\<^bsub>o\<^esub>) ttas))) (llist_of \ta'\\<^bsub>o\<^esub>)" and eq: "eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta'\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" and aok: "actions_ok s' t ta'" by blast obtain ws''' where "redT_updWs t (wset s') \ta'\\<^bsub>w\<^esub> ws'''" using redT_updWs_total .. then obtain s''' where s''': "redT_upd s' t ta' x'' m'' s'''" by fastforce with red \thr s' t = \(x, no_wait_locks)\\ aok have "s' -t\ta'\ s'''" by(rule redT.redT_normal) moreover from sc sc' have "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (lappend (llist_of ttas) (LCons (t, ta') LNil))))" by(auto simp add: lmap_lappend_distrib ta_seq_consist_lappend split_def lconcat_llist_of[symmetric] o_def list_of_lconcat) ultimately show thesis using eq by(rule that) next case (redT_acquire x ln n) hence "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (lappend (llist_of ttas) (LCons (t, ta) LNil))))" and "eq_upto_seq_inconsist P \ta\\<^bsub>o\<^esub> \ta\\<^bsub>o\<^esub> (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) ttas)))" using sc by(simp_all add: lmap_lappend_distrib ta_seq_consist_lappend split_def lconcat_llist_of[symmetric] o_def list_of_lconcat ta_seq_consist_convert_RA ta_seq_consist_imp_eq_upto_seq_inconsist_refl) with red show thesis by(rule that) qed txt \Now, find a sequentially consistent completion from @{term "s'''"} onwards.\ from Red red' have Red': "s -\ttas @ [(t, ta')]\* s'''" unfolding RedT_def by(auto intro: rtrancl3p_step) from sc sc' have "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (llist_of (ttas @ [(t, ta')]))))" by(simp add: o_def split_def lappend_llist_of_llist_of[symmetric]) with cut_and_update_imp_sc_completion[OF cut_and_update] Red' have "sc_completion s''' (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (ttas @ [(t, ta')]))))" by(rule sc_completion_shift) from sequential_completion_Runs[OF this ta_seq_consist_convert_RA] obtain ttas' where \Runs: "mthr.Runs s''' ttas'" and sc'': "ta_seq_consist P (mrw_values P vs (concat (map (\(t, ta). \ta\\<^bsub>o\<^esub>) (ttas @ [(t, ta')])))) (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) ttas'))" by blast from red' \Runs have "mthr.Runs s' (LCons (t, ta') ttas')" .. moreover from sc sc' sc'' have "ta_seq_consist P vs (lconcat (lmap (\(t, ta). llist_of \ta\\<^bsub>o\<^esub>) (lappend (llist_of ttas) (LCons (t, ta') ttas'))))" unfolding lmap_lappend_distrib lconcat_lappend by(simp add: o_def ta_seq_consist_lappend split_def list_of_lconcat) ultimately show ?thesis using eq by blast qed end end diff --git a/thys/List-Infinite/ListInf/List2.thy b/thys/List-Infinite/ListInf/List2.thy --- a/thys/List-Infinite/ListInf/List2.thy +++ b/thys/List-Infinite/ListInf/List2.thy @@ -1,1194 +1,1193 @@ (* Title: List2.thy Date: Oct 2006 Author: David Trachtenherz *) section \Additional definitions and results for lists\ theory List2 imports "../CommonSet/SetIntervalCut" begin subsection \Additional definitions and results for lists\ text \ Infix syntactical abbreviations for operators @{term take} and @{term drop}. The abbreviations resemble to the operator symbols used later for take and drop operators on infinite lists in ListInf.\ abbreviation f_take' :: "'a list \ nat \ 'a list" (infixl "\" 100) where "xs \ n \ take n xs" abbreviation f_drop' :: "'a list \ nat \ 'a list" (infixl "\" 100) where "xs \ n \ drop n xs" lemma append_eq_Cons: "[x] @ xs = x # xs" by simp lemma length_Cons: "length (x # xs) = Suc (length xs)" by simp lemma length_snoc: "length (xs @ [x]) = Suc (length xs)" by simp subsubsection \Additional lemmata about list emptiness\ lemma length_greater_imp_not_empty:"n < length xs \ xs \ []" by fastforce lemma length_ge_Suc_imp_not_empty:"Suc n \ length xs \ xs \ []" by fastforce lemma length_take_le: "length (xs \ n) \ length xs" by simp lemma take_not_empty_conv:"(xs \ n \ []) = (0 < n \ xs \ [])" by simp lemma drop_not_empty_conv:"(xs \ n \ []) = (n < length xs)" by fastforce lemma zip_eq_Nil: "(zip xs ys = []) = (xs = [] \ ys = [])" by (force simp: length_0_conv[symmetric] min_def simp del: length_0_conv) lemma zip_not_empty_conv: "(zip xs ys \ []) = (xs \ [] \ ys \ [])" by (simp add: zip_eq_Nil) subsubsection \Additional lemmata about @{term take}, @{term drop}, @{term hd}, @{term last}, \nth\ and \filter\\ lemma nth_tl_eq_nth_Suc: " Suc n \ length xs \ (tl xs) ! n = xs ! Suc n" by (rule hd_Cons_tl[OF length_ge_Suc_imp_not_empty, THEN subst], simp+) corollary nth_tl_eq_nth_Suc2: " n < length xs \ (tl xs) ! n = xs ! Suc n" by (simp add: nth_tl_eq_nth_Suc) lemma hd_eq_first: "xs \ [] \ xs ! 0 = hd xs" by (induct xs, simp_all) corollary take_first:"xs \ [] \ xs \ (Suc 0) = [xs ! 0]" by (induct xs, simp_all) corollary take_hd:"xs \ [] \ xs \ (Suc 0) = [hd xs]" by (simp add: take_first hd_eq_first) theorem last_nth: "xs \ [] \ last xs = xs ! (length xs - Suc 0)" by (simp add: last_conv_nth) lemma last_take: "n < length xs \ last (xs \ Suc n) = xs ! n" by (simp add: last_nth length_greater_imp_not_empty min_eqR) corollary last_take2:" \ 0 < n; n \ length xs \ \ last (xs \ n) = xs ! (n - Suc 0)" apply (frule diff_Suc_less[THEN order_less_le_trans, of _ "length xs" 0], assumption) apply (drule last_take[of "n - Suc 0" xs]) apply simp done corollary nth_0_drop: "n \ length xs \ (xs \ n) ! 0 = xs ! n" by (cut_tac nth_drop[of n xs 0], simp+) lemma drop_eq_tl: "xs \ (Suc 0) = tl xs" by (simp add: drop_Suc) lemma drop_take_1: " n < length xs \ xs \ n \ (Suc 0) = [xs ! n]" by (simp add: take_hd hd_drop_conv_nth) lemma upt_append: "m \ n \ [0.. (xs @ ys) ! n = xs ! n" by (simp add: nth_append) lemma nth_append2: "length xs \ n \ (xs @ ys) ! n = ys ! (n - length xs)" by (simp add: nth_append) lemma list_all_conv: "list_all P xs = (\iys. (xs = ys) = (length xs = length ys \ (\i xs \ n = ys \ n; xs \ n = ys \ n \ \ xs = ys" apply (rule subst[OF append_take_drop_id[of n xs]]) apply (rule subst[OF append_take_drop_id[of n ys]]) apply simp done lemma list_take_drop_eq_conv: " (xs = ys) = (\n. (xs \ n = ys \ n \ xs \ n = ys \ n))" by (blast intro: list_take_drop_imp_eq) lemma list_take_eq_conv: "(xs = ys) = (\n. xs \ n = ys \ n)" apply (rule iffI, simp) apply (drule_tac x="max (length xs) (length ys)" in spec) apply simp done lemma list_drop_eq_conv: "(xs = ys) = (\n. xs \ n = ys \ n)" apply (rule iffI, simp) apply (drule_tac x=0 in spec) apply simp done abbreviation replicate' :: "'a \ nat \ 'a list" ("_\<^bsup>_\<^esup>" [1000,65]) where "x\<^bsup>n\<^esup> \ replicate n x" lemma replicate_snoc: "x\<^bsup>n\<^esup> @ [x] = x\<^bsup>Suc n\<^esup>" by (simp add: replicate_app_Cons_same) lemma eq_replicate_conv: "(\ilength xs\<^esup>)" apply (rule iffI) apply (simp add: expand_list_eq) apply clarsimp apply (rule ssubst[of xs "replicate (length xs) m"], assumption) apply (rule nth_replicate, simp) done lemma replicate_Cons_length: "length (x # a\<^bsup>n\<^esup>) = Suc n" by simp lemma replicate_pred_Cons_length: "0 < n \ length (x # a\<^bsup>n - Suc 0\<^esup>) = n" by simp lemma replicate_le_diff: "m \ n \ x\<^bsup>m\<^esup> @ x\<^bsup>n - m\<^esup> = x\<^bsup>n\<^esup>" by (simp add: replicate_add[symmetric]) lemma replicate_le_diff2: "\ k \ m; m \ n \ \ x\<^bsup>m - k\<^esup> @ x\<^bsup>n - m\<^esup> = x\<^bsup>n - k\<^esup>" by (subst replicate_add[symmetric], simp) lemma append_constant_length_induct_aux: "\xs. \ length xs div k = n; \ys. k = 0 \ length ys < k \ P ys; \xs ys. \ length xs = k; P ys \ \ P (xs @ ys) \ \ P xs" apply (case_tac "k = 0", blast) apply simp apply (induct n) apply (simp add: div_eq_0_conv') apply (subgoal_tac "k \ length xs") prefer 2 apply (rule div_gr_imp_gr_divisor[of 0], simp) apply (simp only: atomize_all atomize_imp, clarsimp) apply (erule_tac x="drop k xs" in allE) apply (simp add: div_diff_self2) apply (erule_tac x=undefined in allE) apply (erule_tac x="take k xs" in allE) apply (simp add: min_eqR) apply (erule_tac x="drop k xs" in allE) apply simp done lemma append_constant_length_induct: " \ \ys. k = 0 \ length ys < k \ P ys; \xs ys. \ length xs = k; P ys \ \ P (xs @ ys) \ \ P xs" by (simp add: append_constant_length_induct_aux[of _ _ "length xs div k"]) lemma zip_swap: "map (\(y,x). (x,y)) (zip ys xs) = (zip xs ys)" by (simp add: expand_list_eq) lemma zip_takeL: "(zip xs ys) \ n = zip (xs \ n) ys" by (simp add: expand_list_eq) lemma zip_takeR: "(zip xs ys) \ n = zip xs (ys \ n)" apply (subst zip_swap[of ys, symmetric]) apply (subst take_map) apply (subst zip_takeL) apply (simp add: zip_swap) done lemma zip_take: "(zip xs ys) \ n = zip (xs \ n) (ys \ n)" by (rule take_zip) lemma hd_zip: "\ xs \ []; ys \ [] \ \ hd (zip xs ys) = (hd xs, hd ys)" by (simp add: hd_conv_nth zip_not_empty_conv) lemma map_id: "map id xs = xs" by (simp add: id_def) lemma map_id_subst: "P (map id xs) \ P xs" by (subst map_id[symmetric]) lemma map_one: "map f [x] = [f x]" by simp lemma map_last: "xs \ [] \ last (map f xs) = f (last xs)" by (rule last_map) lemma filter_list_all: "list_all P xs \ filter P xs = xs" by (induct xs, simp+) lemma filter_snoc: "filter P (xs @ [x]) = (if P x then (filter P xs) @ [x] else filter P xs)" by (case_tac "P x", simp+) lemma filter_filter_eq: "list_all (\x. P x = Q x) xs \ filter P xs = filter Q xs" by (induct xs, simp+) lemma filter_nth: "\n. n < length (filter P xs) \ (filter P xs) ! n = xs ! (LEAST k. k < length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)})" apply (induct xs rule: rev_induct, simp) apply (rename_tac x xs n) apply (simp only: filter_snoc) apply (simp split del: if_split) apply (case_tac "xs = []") apply (simp split del: if_split) apply (rule_tac t = "\k i. i = 0 \ i \ k \ P ([x] ! i)" and s = "\k i. i = 0 \ P x" in subst) apply (simp add: fun_eq_iff) apply fastforce apply (fastforce simp: Least_def) apply (rule_tac t = "\k. card {i. i \ k \ i < Suc (length xs) \ P ((xs @ [x]) ! i)}" and s = "\k. (card {i. i \ k \ i < length xs \ P (xs ! i)} + (if k \ length xs \ P x then Suc 0 else 0))" in subst) apply (clarsimp simp: fun_eq_iff split del: if_split, rename_tac k) apply (simp split del: if_split add: less_Suc_eq conj_disj_distribL conj_disj_distribR Collect_disj_eq) apply (subst card_Un_disjoint) apply (rule_tac n="length xs" in bounded_nat_set_is_finite, blast) apply (rule_tac n="Suc (length xs)" in bounded_nat_set_is_finite, blast) apply blast apply (rule_tac t = "\i. i < length xs \ P ((xs @ [x]) ! i)" and s = "\i. i < length xs \ P (xs ! i)" in subst) apply (rule fun_eq_iff[THEN iffD2]) apply (fastforce simp: nth_append1) apply (rule add_left_cancel[THEN iffD2]) apply (rule_tac t = "\i. i = length xs \ i \ k \ P ((xs @ [x]) ! i)" and s = "\i. i = length xs \ i \ k \ P x" in subst) apply (rule fun_eq_iff[THEN iffD2]) apply fastforce apply (case_tac "length xs \ k") apply clarsimp apply (rule_tac t = "\i. i = length xs \ i \ k" and s = "\i. i = length xs" in subst) apply (rule fun_eq_iff[THEN iffD2]) apply fastforce apply simp apply simp apply (simp split del: if_split add: less_Suc_eq conj_disj_distribL conj_disj_distribR) apply (rule_tac t = "\k. k < length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)} + (if length xs \ k \ P x then Suc 0 else 0)" and s = "\k. k < length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)}" in subst) apply (simp add: fun_eq_iff) apply (rule_tac t = "\k. k = length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)} + (if length xs \ k \ P x then Suc 0 else 0)" and s = "\k. k = length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)} + (if P x then Suc 0 else 0)" in subst) apply (simp add: fun_eq_iff) apply (case_tac "n < length (filter P xs)") apply (rule_tac t = "(if P x then filter P xs @ [x] else filter P xs) ! n" and s = "(filter P xs) ! n" in subst) apply (simp add: nth_append1) apply (simp split del: if_split) apply (subgoal_tac "\k k \ i < length xs \ P (xs ! i)}") prefer 2 apply (rule_tac x="length xs - Suc 0" in exI) apply (simp add: length_filter_conv_card less_eq_le_pred[symmetric]) apply (subgoal_tac "\k\length xs. n < card {i. i \ k \ i < length xs \ P (xs ! i)}") prefer 2 apply (blast intro: less_imp_le) apply (subst Least_le_imp_le_disj) apply simp apply simp apply (rule sym, rule nth_append1) apply (rule LeastI2_ex, assumption) apply blast apply (simp add: linorder_not_less) apply (subgoal_tac "P x") prefer 2 apply (rule ccontr, simp) apply (simp add: length_snoc) apply (drule less_Suc_eq_le[THEN iffD1], drule_tac x=n in order_antisym, assumption) apply (simp add: nth_append2) apply (simp add: length_filter_conv_card) apply (rule_tac t = "\k. card {i. i < length xs \ P (xs ! i)} < card {i. i \ k \ i < length xs \ P (xs ! i)}" and s = "\k. False" in subst) apply (rule fun_eq_iff[THEN iffD2], rule allI, rename_tac k) apply (simp add: linorder_not_less) apply (rule card_mono) apply fastforce apply blast apply simp apply (rule_tac t = "(LEAST k. k = length xs \ card {i. i < length xs \ P (xs ! i)} < Suc (card {i. i \ k \ i < length xs \ P (xs ! i)}))" and s = "length xs" in subst) apply (rule sym, rule Least_equality) apply simp apply (rule le_imp_less_Suc) apply (rule card_mono) apply fastforce apply fastforce apply simp apply simp done subsubsection \Ordered lists\ fun list_ord :: "('a \ 'a \ bool) \ ('a::ord) list \ bool" where "list_ord ord (x1 # x2 # xs) = (ord x1 x2 \ list_ord ord (x2 # xs))" | "list_ord ord xs = True" definition list_asc :: "('a::ord) list \ bool" where "list_asc xs \ list_ord (\) xs" definition list_strict_asc :: "('a::ord) list \ bool" where "list_strict_asc xs \ list_ord (<) xs" value "list_asc [1::nat, 2, 2]" value "list_strict_asc [1::nat, 2, 2]" definition list_desc :: "('a::ord) list \ bool" where "list_desc xs \ list_ord (\) xs" definition list_strict_desc :: "('a::ord) list \ bool" where "list_strict_desc xs \ list_ord (>) xs" lemma list_ord_Nil: "list_ord ord []" by simp lemma list_ord_one: "list_ord ord [x]" by simp lemma list_ord_Cons: " list_ord ord (x # xs) = (xs = [] \ (ord x (hd xs) \ list_ord ord xs))" by (induct xs, simp+) lemma list_ord_Cons_imp: "\ list_ord ord xs; ord x (hd xs) \ \ list_ord ord (x # xs)" by (induct xs, simp+) lemma list_ord_append: "\ys. list_ord ord (xs @ ys) = (list_ord ord xs \ (ys = [] \ (list_ord ord ys \ (xs = [] \ ord (last xs) (hd ys)))))" apply (induct xs, fastforce) apply (case_tac xs, case_tac ys, fastforce+) done lemma list_ord_snoc: " list_ord ord (xs @ [x]) = (xs = [] \ (ord (last xs) x \ list_ord ord xs))" by (fastforce simp: list_ord_append) lemma list_ord_all_conv: " (list_ord ord xs) = (\n < length xs - 1. ord (xs ! n) (xs ! Suc n))" apply (rule iffI) apply (induct xs, simp) apply clarsimp apply (simp add: list_ord_Cons) apply (erule disjE, simp) apply clarsimp apply (case_tac n) apply (simp add: hd_conv_nth) apply simp apply (induct xs, simp) apply (simp add: list_ord_Cons) apply (case_tac "xs = []", simp) apply (drule meta_mp) apply (intro allI impI, rename_tac n) apply (drule_tac x="Suc n" in spec, simp) apply (drule_tac x=0 in spec) apply (simp add: hd_conv_nth) done lemma list_ord_imp: " \ \x y. ord x y \ ord' x y; list_ord ord xs \ \ list_ord ord' xs" apply (induct xs, simp) apply (simp add: list_ord_Cons) apply fastforce done corollary list_strict_asc_imp_list_asc: " list_strict_asc (xs::'a::preorder list) \ list_asc xs" by (unfold list_strict_asc_def list_asc_def, rule list_ord_imp[of "(<)"], rule order_less_imp_le) corollary list_strict_desc_imp_list_desc: " list_strict_desc (xs::'a::preorder list) \ list_desc xs" by (unfold list_strict_desc_def list_desc_def, rule list_ord_imp[of "(>)"], rule order_less_imp_le) lemma list_ord_trans_imp: "\i. \ transP ord; list_ord ord xs; j < length xs; i < j \ \ ord (xs ! i) (xs ! j)" apply (simp add: list_ord_all_conv) apply (induct j, simp) apply (case_tac "j < i", simp) apply (simp add: linorder_not_less) apply (case_tac "i = j", simp) apply (drule_tac x=i in meta_spec, simp) apply (drule_tac x=j in spec, simp add: Suc_less_pred_conv) apply (unfold trans_def) apply (drule_tac x="xs ! i" in spec, drule_tac x="xs ! j" in spec, drule_tac x="xs ! Suc j" in spec) apply simp done lemma list_ord_trans: " transP ord \ (list_ord ord xs) = (\j < length xs. \i < j. ord (xs ! i) (xs ! j))" apply (rule iffI) apply (simp add: list_ord_trans_imp) apply (simp add: list_ord_all_conv) done lemma list_ord_trans_refl_le: " \ transP ord; reflP ord \ \ (list_ord ord xs) = (\j < length xs. \i \ j. ord (xs ! i) (xs ! j))" apply (subst list_ord_trans, simp) apply (rule iffI) apply clarsimp apply (case_tac "i = j") apply (simp add: refl_on_def) apply simp+ done lemma list_ord_trans_refl_le_imp: " \ transP ord; \x y. ord x y \ ord' x y; reflP ord'; list_ord ord xs \ \ (\j < length xs. \i \ j. ord' (xs ! i) (xs ! j))" apply clarify apply (case_tac "i = j") apply (simp add: refl_on_def) apply (simp add: list_ord_trans_imp) done corollary list_asc_trans: " (list_asc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! i \ xs ! j)" and list_strict_asc_trans: " (list_strict_asc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! i < xs ! j)" and list_desc_trans: " (list_desc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! j \ xs ! i)" and list_strict_desc_trans: " (list_strict_desc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! j < xs ! i)" apply (unfold list_asc_def list_strict_asc_def list_desc_def list_strict_desc_def) apply (rule list_ord_trans, unfold trans_def, blast intro: order_trans order_less_trans)+ done corollary list_asc_trans_le: " (list_asc (xs::'a::preorder list)) = (\j < length xs. \i \ j. xs ! i \ xs ! j)" and list_desc_trans_le: " (list_desc (xs::'a::preorder list)) = (\j < length xs. \i \ j. xs ! j \ xs ! i)" apply (unfold list_asc_def list_strict_asc_def list_desc_def list_strict_desc_def) apply (rule list_ord_trans_refl_le, unfold trans_def, blast intro: order_trans, simp add: refl_on_def)+ done corollary list_strict_asc_trans_le: " (list_strict_asc (xs::'a::preorder list)) \ (\j < length xs. \i \ j. xs ! i \ xs ! j)" apply (unfold list_strict_asc_def) apply (rule list_ord_trans_refl_le_imp[where ord="(\)"]) apply (unfold trans_def, blast intro: order_trans) apply assumption apply (unfold refl_on_def, clarsimp) apply (rule list_ord_imp[where ord="(<)"], simp_all add: less_imp_le) done lemma list_ord_le_sorted_eq: "list_asc xs = sorted xs" apply (rule sym) apply (simp add: list_asc_def) apply (induct xs, simp) apply (rename_tac x xs) apply (simp add: list_ord_Cons) apply (case_tac "xs = []", simp_all) apply (case_tac "list_ord (\) xs", simp_all) apply (rule iffI) apply (drule_tac x="hd xs" in bspec, simp_all) apply clarify apply (drule in_set_conv_nth[THEN iffD1], clarsimp, rename_tac i1) apply (simp add: hd_conv_nth) apply (case_tac i1, simp) apply (rename_tac i2) apply simp apply (fold list_asc_def) apply (fastforce simp: list_asc_trans) done corollary list_asc_upto: "list_asc [m..n]" by (simp add: list_ord_le_sorted_eq) lemma list_strict_asc_upt: "list_strict_asc [m.. irrefl {(a, b). ord a b}; transP ord; list_ord ord xs; i < length xs; j < length xs; i < j \ \ xs ! i \ xs ! j" apply (subgoal_tac "\x y. ord x y \ x \ y") prefer 2 apply (rule ccontr) apply (simp add: irrefl_def) apply (simp add: list_ord_trans) done lemma list_ord_distinct: " \ irrefl {(a,b). ord a b}; transP ord; list_ord ord xs \ \ distinct xs" apply (simp add: distinct_conv_nth, intro allI impI, rename_tac i j) apply (drule neq_iff[THEN iffD1], erule disjE) apply (simp add: list_ord_distinct_aux) apply (simp add: list_ord_distinct_aux[THEN not_sym]) done lemma list_strict_asc_distinct: "list_strict_asc (xs::'a::preorder list) \ distinct xs" apply (rule_tac ord="(<)" in list_ord_distinct) apply (unfold irrefl_def list_strict_asc_def trans_def) apply (blast intro: less_trans)+ done lemma list_strict_desc_distinct: "list_strict_desc (xs::'a::preorder list) \ distinct xs" apply (rule_tac ord="(>)" in list_ord_distinct) apply (unfold irrefl_def list_strict_desc_def trans_def) apply (blast intro: less_trans)+ done subsubsection \Additional definitions and results for sublists\ primrec sublist_list :: "'a list \ nat list \ 'a list" where "sublist_list xs [] = []" | "sublist_list xs (y # ys) = (xs ! y) # (sublist_list xs ys)" value "sublist_list [0::int,10::int,20,30,40,50] [1::nat,2,3]" value "sublist_list [0::int,10::int,20,30,40,50] [1::nat,1,2,3]" value [nbe] "sublist_list [0::int,10::int,20,30,40,50] [1::nat,1,2,3,10]" lemma sublist_list_length: "length (sublist_list xs ys) = length ys" by (induct ys, simp_all) lemma sublist_list_append: " \zs. sublist_list xs (ys @ zs) = sublist_list xs ys @ sublist_list xs zs" by (induct ys, simp_all) lemma sublist_list_Nil: "sublist_list xs [] =[]" by simp lemma sublist_list_is_Nil_conv: " (sublist_list xs ys = []) = (ys = [])" apply (rule iffI) apply (rule ccontr) apply (clarsimp simp: neq_Nil_conv) apply simp done lemma sublist_list_eq_imp_length_eq: " sublist_list xs ys = sublist_list xs zs \ length ys = length zs" by (drule arg_cong[where f=length], simp add: sublist_list_length) lemma sublist_list_nth: " \n. n < length ys \ sublist_list xs ys ! n = xs ! (ys ! n)" apply (induct ys, simp) apply (case_tac n, simp_all) done lemma take_drop_eq_sublist_list: " m + n \ length xs \ xs \ m \ n = sublist_list xs [m.. nat list \ 'a list" where "sublist_list_if xs [] = []" | "sublist_list_if xs (y # ys) = (if y < length xs then (xs ! y) # (sublist_list_if xs ys) else (sublist_list_if xs ys))" value "sublist_list_if [0::int,10::int,20,30,40,50] [1::nat,2,3]" value "sublist_list_if [0::int,10::int,20,30,40,50] [1::nat,1,2,3]" value "sublist_list_if [0::int,10::int,20,30,40,50] [1::nat,1,2,3,10]" lemma sublist_list_if_sublist_list_filter_conv: "\xs. sublist_list_if xs ys = sublist_list xs (filter (\i. i < length xs) ys)" by (induct ys, simp+) corollary sublist_list_if_sublist_list_eq: "\xs. list_all (\i. i < length xs) ys \ sublist_list_if xs ys = sublist_list xs ys" by (simp add: sublist_list_if_sublist_list_filter_conv filter_list_all) corollary sublist_list_if_sublist_list_eq2: "\xs. \n sublist_list_if xs ys = sublist_list xs ys" by (rule sublist_list_if_sublist_list_eq, rule list_all_conv[THEN iffD2]) lemma sublist_list_if_Nil_left: "sublist_list_if [] ys = []" by (induct ys, simp+) lemma sublist_list_if_Nil_right: "sublist_list_if xs [] = []" by simp lemma sublist_list_if_length: " length (sublist_list_if xs ys) = length (filter (\i. i < length xs) ys)" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_length) lemma sublist_list_if_append: " sublist_list_if xs (ys @ zs) = sublist_list_if xs ys @ sublist_list_if xs zs" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_append) lemma sublist_list_if_snoc: " sublist_list_if xs (ys @ [y]) = sublist_list_if xs ys @ (if y < length xs then [xs ! y] else [])" by (simp add: sublist_list_if_append) lemma sublist_list_if_is_Nil_conv: " (sublist_list_if xs ys = []) = (list_all (\i. length xs \ i) ys)" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_is_Nil_conv filter_empty_conv list_all_iff linorder_not_less) lemma sublist_list_if_nth: " n < length ((filter (\i. i < length xs) ys)) \ sublist_list_if xs ys ! n = xs ! ((filter (\i. i < length xs) ys) ! n)" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_nth) lemma take_drop_eq_sublist_list_if: " m + n \ length xs \ xs \ m \ n = sublist_list_if xs [m..i\I. length xs \ i)" using [[hypsubst_thin = true]] by (fastforce simp: set_empty[symmetric] set_nths linorder_not_le[symmetric]) lemma nths_singleton2: "nths xs {y} = (if y < length xs then [xs ! y] else [])" apply (unfold nths_def) apply (induct xs rule: rev_induct, simp) apply (simp add: nth_append) done lemma nths_take_eq: " \ finite I; Max I < n \ \ nths (xs \ n) I = nths xs I" apply (case_tac "I = {}", simp) apply (case_tac "n < length xs") prefer 2 apply simp apply (rule_tac t = "nths xs I" and s = "nths (xs \ n @ xs \ n) I" in subst) apply simp apply (subst nths_append) apply (simp add: min_eqR) apply (rule_tac t="{j. j + n \ I}" and s="{}" in subst) apply blast apply simp done lemma nths_drop_eq: " n \ iMin I \ nths (xs \ n) {j. j + n \ I} = nths xs I" apply (case_tac "I = {}", simp) apply (case_tac "n < length xs") prefer 2 apply (simp add: nths_def filter_empty_conv linorder_not_less) apply (clarsimp, rename_tac a b) apply (drule set_zip_rightD) apply fastforce apply (rule_tac t = "nths xs I" and s = "nths (xs \ n @ xs \ n) I" in subst) apply simp apply (subst nths_append) apply (fastforce simp: nths_empty_conv min_eqR) done lemma nths_cut_less_eq: " length xs \ n \ nths xs (I \< n) = nths xs I" apply (simp add: nths_def cut_less_mem_iff) apply (rule_tac f="\xs. map fst xs" in arg_cong) apply (rule filter_filter_eq) apply (simp add: list_all_conv) done lemma nths_disjoint_Un: " \ finite A; Max A < iMin B \ \ nths xs (A \ B) = nths xs A @ nths xs B" apply (case_tac "A = {}", simp) apply (case_tac "B = {}", simp) apply (case_tac "length xs \ iMin B") apply (subst nths_cut_less_eq[of xs "iMin B", symmetric], assumption) apply (simp (no_asm_simp) add: cut_less_Un cut_less_Min_empty cut_less_Max_all) apply (simp add: nths_empty_conv iMin_ge_iff) apply (simp add: linorder_not_le) apply (rule_tac t = "nths xs (A \ B)" and s = "nths (xs \ (iMin B) @ xs \ (iMin B)) (A \ B)" in subst) apply simp apply (subst nths_append) apply (simp add: min_eqR) apply (subst nths_cut_less_eq[where xs="xs \ iMin B" and n="iMin B", symmetric], simp) apply (simp add: cut_less_Un cut_less_Min_empty cut_less_Max_all) apply (simp add: nths_take_eq) apply (rule_tac t = "\j. j + iMin B \ A \ j + iMin B \ B" and s = "\j. j + iMin B \ B" in subst) apply (force simp: fun_eq_iff) apply (simp add: nths_drop_eq) done corollary nths_disjoint_insert_left: " \ finite I; x < iMin I \ \ nths xs (insert x I) = nths xs {x} @ nths xs I" apply (rule_tac t="insert x I" and s="{x} \ I" in subst, simp) apply (subst nths_disjoint_Un) apply simp_all done corollary nths_disjoint_insert_right: " \ finite I; Max I < x \ \ nths xs (insert x I) = nths xs I @ nths xs {x}" apply (rule_tac t="insert x I" and s="I \ {x}" in subst, simp) apply (subst nths_disjoint_Un) apply simp_all done lemma nths_all: "{.. I \ nths xs I = xs" apply (case_tac "xs = []", simp) apply (rule_tac t = "I" and s = "I \< (length xs) \ I \\ (length xs)" in subst) apply (simp add: cut_less_cut_ge_ident) apply (rule_tac t = "I \< length xs" and s = "{..\ (length xs) = {}", simp) apply (subst nths_disjoint_Un[OF finite_lessThan]) apply (rule less_imp_Max_less_iMin[OF finite_lessThan]) apply blast apply blast apply (blast intro: less_le_trans) apply (fastforce simp: nths_empty_conv) done corollary nths_UNIV: "nths xs UNIV = xs" by (rule nths_all[OF subset_UNIV]) lemma sublist_list_nths_eq: "\xs. list_strict_asc ys \ sublist_list_if xs ys = nths xs (set ys)" apply (case_tac "xs = []") apply (simp add: sublist_list_if_Nil_left) apply (induct ys rule: rev_induct, simp) apply (rename_tac y ys xs) apply (case_tac "ys = []") apply (simp add: nths_singleton2) apply (unfold list_strict_asc_def) apply (simp add: sublist_list_if_snoc split del: if_split) apply (frule list_ord_append[THEN iffD1]) apply (clarsimp split del: if_split) apply (subst nths_disjoint_insert_right) apply simp apply (clarsimp simp: in_set_conv_nth, rename_tac i) apply (drule_tac i=i and j="length ys" in list_strict_asc_trans[unfolded list_strict_asc_def, THEN iffD1, rule_format]) apply (simp add: nth_append split del: if_split)+ apply (simp add: nths_singleton2) done lemma set_sublist_list_if: "\xs. set (sublist_list_if xs ys) = {xs ! i |i. i < length xs \ i \ set ys}" apply (induct ys, simp_all) apply blast done lemma set_sublist_list: " list_all (\i. i < length xs) ys \ set (sublist_list xs ys) = {xs ! i |i. i < length xs \ i \ set ys}" by (simp add: sublist_list_if_sublist_list_eq[symmetric] set_sublist_list_if) lemma set_sublist_list_if_eq_set_sublist: "set (sublist_list_if xs ys) = set (nths xs (set ys))" by (simp add: set_nths set_sublist_list_if) lemma set_sublist_list_eq_set_sublist: " list_all (\i. i < length xs) ys \ set (sublist_list xs ys) = set (nths xs (set ys))" by (simp add: sublist_list_if_sublist_list_eq[symmetric] set_sublist_list_if_eq_set_sublist) subsubsection \Natural set images with lists\ definition f_image :: "'a list \ nat set \ 'a set" (infixr "`\<^sup>f" 90) where "xs `\<^sup>f A \ {y. \n\A. n < length xs \ y = xs ! n}" abbreviation f_range :: "'a list \ 'a set" where "f_range xs \ f_image xs UNIV" lemma f_image_eqI[simp, intro]: " \ x = xs ! n; n \ A; n < length xs \ \ x \ xs `\<^sup>f A" by (unfold f_image_def, blast) lemma f_imageI: "\ n \ A; n < length xs \ \ xs ! n \ xs `\<^sup>f A" by blast lemma rev_f_imageI: "\ n \ A; n < length xs; x = xs ! n \ \ x \ xs `\<^sup>f A" by (rule f_image_eqI) lemma f_imageE[elim!]: " \ x \ xs `\<^sup>f A; \n. \ x = xs ! n; n \ A; n < length xs \ \ P \ \ P" by (unfold f_image_def, blast) lemma f_image_Un: "xs `\<^sup>f (A \ B) = xs `\<^sup>f A \ xs `\<^sup>f B" by blast lemma f_image_mono: "A \ B ==> xs `\<^sup>f A \ xs `\<^sup>f B" by blast lemma f_image_iff: "(x \ xs `\<^sup>f A) = (\n\A. n < length xs \ x = xs ! n)" by blast lemma f_image_subset_iff: " (xs `\<^sup>f A \ B) = (\n\A. n < length xs \ xs ! n \ B)" by blast lemma subset_f_image_iff: "(B \ xs `\<^sup>f A) = (\A'\A. B = xs `\<^sup>f A')" apply (rule iffI) apply (rule_tac x="{ n. n \ A \ n < length xs \ xs ! n \ B }" in exI) apply blast apply (blast intro: f_image_mono) done lemma f_image_subsetI: " \ \n. n \ A \ n < length xs \ xs ! n \ B \ \ xs `\<^sup>f A \ B" by blast lemma f_image_empty: "xs `\<^sup>f {} = {}" by blast lemma f_image_insert_if: " xs `\<^sup>f (insert n A) = ( if n < length xs then insert (xs ! n) (xs `\<^sup>f A) else (xs `\<^sup>f A))" by (split if_split, blast) lemma f_image_insert_eq1: " n < length xs \ xs `\<^sup>f (insert n A) = insert (xs ! n) (xs `\<^sup>f A)" by (simp add: f_image_insert_if) lemma f_image_insert_eq2: " length xs \ n \ xs `\<^sup>f (insert n A) = (xs `\<^sup>f A)" by (simp add: f_image_insert_if) lemma insert_f_image: " \ n \ A; n < length xs \ \ insert (xs ! n) (xs `\<^sup>f A) = (xs `\<^sup>f A)" by blast lemma f_image_is_empty: "(xs `\<^sup>f A = {}) = ({x. x \ A \ x < length xs} = {})" by blast lemma f_image_Collect: "xs `\<^sup>f {n. P n} = {xs ! n |n. P n \ n < length xs}" by blast lemma f_image_eq_set: "\n A \ xs `\<^sup>f A = set xs" by (fastforce simp: in_set_conv_nth) lemma f_range_eq_set: "f_range xs = set xs" by (simp add: f_image_eq_set) lemma f_image_eq_set_nths: "xs `\<^sup>f A = set (nths xs A)" by (unfold set_nths, blast) lemma f_image_eq_set_sublist_list_if: "xs `\<^sup>f (set ys) = set (sublist_list_if xs ys)" by (simp add: set_sublist_list_if_eq_set_sublist f_image_eq_set_nths) lemma f_image_eq_set_sublist_list: " list_all (\i. i < length xs) ys \ xs `\<^sup>f (set ys) = set (sublist_list xs ys)" by (simp add: sublist_list_if_sublist_list_eq f_image_eq_set_sublist_list_if) lemma f_range_eqI: "\ x = xs ! n; n < length xs \ \ x \ f_range xs" by blast lemma f_rangeI: "n < length xs \ xs ! n \ f_range xs" by blast lemma f_rangeE[elim?]: " \ x \ f_range xs; \n. \ n < length xs; x = xs ! n \ \ P \ \ P" by blast subsubsection \Mapping lists of functions to lists\ primrec map_list :: "('a \ 'b) list \ 'a list \ 'b list" where "map_list [] xs = []" | "map_list (f # fs) xs = f (hd xs) # map_list fs (tl xs)" lemma map_list_Nil: "map_list [] xs = []" by simp lemma map_list_Cons_Cons: " map_list (f # fs) (x # xs) = (f x) # map_list fs xs" by simp lemma map_list_length: "\xs. length (map_list fs xs) = length fs" by (induct fs, simp+) corollary map_list_empty_conv: " (map_list fs xs = []) = (fs = [])" by (simp del: length_0_conv add: length_0_conv[symmetric] map_list_length) corollary map_list_not_empty_conv: " (map_list fs xs \ []) = (fs \ [])" by (simp add: map_list_empty_conv) lemma map_list_nth: "\n xs. \ n < length fs; n < length xs \ \ (map_list fs xs ! n) = (fs ! n) (xs ! n)" apply (induct fs, simp+) apply (case_tac n) apply (simp add: hd_conv_nth) apply (simp add: nth_tl_eq_nth_Suc2) done lemma map_list_xs_take: "\n xs. length fs \ n \ map_list fs (xs \ n) = map_list fs xs" apply (induct fs, simp+) apply (rename_tac fs n xs) apply (simp add: tl_take) done lemma map_list_take: "\n xs. (map_list fs xs) \ n = (map_list (fs \ n) xs)" apply (induct fs, simp) apply (case_tac n, simp+) done lemma map_list_take_take: "\n xs. (map_list fs xs) \ n = (map_list (fs \ n) (xs \ n))" by (simp add: map_list_take map_list_xs_take) lemma map_list_drop: "\n xs. (map_list fs xs) \ n = (map_list (fs \ n) (xs \ n))" apply (induct fs, simp) apply (case_tac n) apply (simp add: drop_Suc)+ done lemma map_list_append_append: "\xs1 . length fs1 = length xs1 \ map_list (fs1 @ fs2) (xs1 @ xs2) = map_list fs1 xs1 @ map_list fs2 xs2" apply (induct fs1, simp+) apply (case_tac "xs1", simp+) done lemma map_list_snoc_snoc: " length fs = length xs \ map_list (fs @ [f]) (xs @ [x]) = map_list fs xs @ [f x]" by (simp add: map_list_append_append) lemma map_list_snoc: "\xs. length fs < length xs \ map_list (fs @ [f]) xs = map_list fs xs @ [f (xs ! (length fs))]" apply (induct fs) apply (simp add: hd_conv_nth) apply (simp add: nth_tl_eq_nth_Suc2) done lemma map_list_Cons_if: " map_list fs (x # xs) = (if (fs = []) then [] else ( ((hd fs) x) # map_list (tl fs) xs))" by (case_tac "fs", simp+) lemma map_list_Cons_not_empty: " fs \ [] \ map_list fs (x # xs) = ((hd fs) x) # map_list (tl fs) xs" by (simp add: map_list_Cons_if) lemma map_eq_map_list_take: "\xs. \ length fs \ length xs; list_all (\x. x = f) fs \ \ map_list fs xs = map f (xs \ length fs)" apply (induct fs, simp+) apply (case_tac xs, simp+) done lemma map_eq_map_list_take2: " \ length fs = length xs; list_all (\x. x = f) fs \ \ map_list fs xs = map f xs" by (simp add: map_eq_map_list_take) lemma map_eq_map_list_replicate: " map_list (f\<^bsup>length xs\<^esup>) xs = map f xs" by (induct xs, simp+) subsubsection \Mapping functions with two arguments to lists\ primrec map2 :: " \ \Function taking two parameters\ ('a \ 'b \ 'c) \ \ \Lists of parameters\ 'a list \ 'b list \ 'c list" where "map2 f [] ys = []" | "map2 f (x # xs) ys = f x (hd ys) # map2 f xs (tl ys)" lemma map2_map_list_conv: "\ys. map2 f xs ys = map_list (map f xs) ys" by (induct xs, simp+) lemma map2_Nil: "map2 f [] ys = []" by simp lemma map2_Cons_Cons: " map2 f (x # xs) (y # ys) = (f x y) # map2 f xs ys" by simp lemma map2_length: "\ys. length (map2 f xs ys) = length xs" by (induct xs, simp+) corollary map2_empty_conv: " (map2 f xs ys = []) = (xs = [])" by (simp del: length_0_conv add: length_0_conv[symmetric] map2_length) corollary map2_not_empty_conv: " (map2 f xs ys \ []) = (xs \ [])" by (simp add: map2_empty_conv) lemma map2_nth: "\n ys. \ n < length xs; n < length ys \ \ (map2 f xs ys ! n) = f (xs ! n) (ys ! n)" by (simp add: map2_map_list_conv map_list_nth) lemma map2_ys_take: "\n ys. length xs \ n \ map2 f xs (ys \ n) = map2 f xs ys" by (simp add: map2_map_list_conv map_list_xs_take) lemma map2_take: "\n ys. (map2 f xs ys) \ n = (map2 f (xs \ n) ys)" by (simp add: map2_map_list_conv take_map map_list_take) lemma map2_take_take: "\n ys. (map2 f xs ys) \ n = (map2 f (xs \ n) (ys \ n))" by (simp add: map2_take map2_ys_take) lemma map2_drop: "\n ys. (map2 f xs ys) \ n = (map2 f (xs \ n) (ys \ n))" by (simp add: map2_map_list_conv map_list_drop drop_map) lemma map2_append_append: "\ys1 . length xs1 = length ys1 \ map2 f (xs1 @ xs2) (ys1 @ ys2) = map2 f xs1 ys1 @ map2 f xs2 ys2" by (simp add: map2_map_list_conv map_list_append_append) lemma map2_snoc_snoc: " length xs = length ys \ map2 f (xs @ [x]) (ys @ [y]) = map2 f xs ys @ [f x y]" by (simp add: map2_append_append) lemma map2_snoc: "\ys. length xs < length ys \ map2 f (xs @ [x]) ys = map2 f xs ys @ [f x (ys ! (length xs))]" by (simp add: map2_map_list_conv map_list_snoc) lemma map2_Cons_if: " map2 f xs (y # ys) = (if (xs = []) then [] else ( (f (hd xs) y) # map2 f (tl xs) ys))" by (case_tac "xs", simp+) lemma map2_Cons_not_empty: " xs \ [] \ map2 f xs (y # ys) = (f (hd xs) y) # map2 f (tl xs) ys" by (simp add: map2_Cons_if) lemma map2_append1_take_drop: " length xs1 \ length ys \ map2 f (xs1 @ xs2) ys = map2 f xs1 (ys \ length xs1) @ map2 f xs2 (ys \ length xs1)" apply (rule_tac t = "map2 f (xs1 @ xs2) ys" and s = "map2 f (xs1 @ xs2) (ys \ length xs1 @ ys \ length xs1)" in subst) apply simp apply (simp add: map2_append_append del: append_take_drop_id) done lemma map2_append2_take_drop: " length ys1 \ length xs \ map2 f xs (ys1 @ ys2) = map2 f (xs \ length ys1) ys1 @ map2 f (xs \ length ys1) ys2" apply (rule_tac t = "map2 f xs (ys1 @ ys2)" and s = "map2 f (xs \ length ys1 @ xs \ length ys1) (ys1 @ ys2)" in subst) apply simp apply (simp add: map2_append_append del: append_take_drop_id) done lemma map2_cong: " \ xs1 = xs2; ys1 = ys2; length xs2 \ length ys2; \x y. \ x \ set xs2; y \ set ys2 \ \ f x y = g x y \ \ map2 f xs1 ys1 = map2 g xs2 ys2" by (simp (no_asm_simp) add: expand_list_eq map2_length map2_nth) lemma map2_eq_conv: " length xs \ length ys \ (map2 f xs ys = map2 g xs ys) = (\in\<^esup> y\<^bsup>n\<^esup> = (f x y)\<^bsup>n\<^esup>" by (induct n, simp+) lemma map2_zip_conv: "\ys. length xs \ length ys \ map2 f xs ys = map (\(x,y). f x y) (zip xs ys)" apply (induct xs, simp) apply (case_tac ys, simp+) done lemma map2_rev: "\ys. length xs = length ys \ rev (map2 f xs ys) = map2 f (rev xs) (rev ys)" apply (induct xs, simp) apply (case_tac ys, simp) apply (simp add: map2_Cons_Cons map2_snoc_snoc) done hide_const (open) map2 end diff --git a/thys/List_Update/Move_to_Front.thy b/thys/List_Update/Move_to_Front.thy --- a/thys/List_Update/Move_to_Front.thy +++ b/thys/List_Update/Move_to_Front.thy @@ -1,1018 +1,1018 @@ (* Author: Tobias Nipkow *) section "Deterministic List Update" theory Move_to_Front imports Swaps On_Off Competitive_Analysis begin declare Let_def[simp] subsection "Function \mtf\" definition mtf :: "'a \ 'a list \ 'a list" where "mtf x xs = (if x \ set xs then x # (take (index xs x) xs) @ drop (index xs x + 1) xs else xs)" lemma mtf_id[simp]: "x \ set xs \ mtf x xs = xs" by(simp add: mtf_def) lemma mtf0[simp]: "x \ set xs \ mtf x xs ! 0 = x" by(auto simp: mtf_def) lemma before_in_mtf: assumes "z \ set xs" shows "x < y in mtf z xs \ (y \ z \ (if x=z then y \ set xs else x < y in xs))" proof- have 0: "index xs z < size xs" by (metis assms index_less_size_conv) let ?xs = "take (index xs z) xs @ xs ! index xs z # drop (Suc (index xs z)) xs" have "x < y in mtf z xs = (y \ z \ (if x=z then y \ set ?xs else x < y in ?xs))" using assms - by(auto simp add: mtf_def before_in_def index_append) - (metis add_lessD1 index_less_size_conv length_take less_Suc_eq not_less_eq) + by (auto simp add: mtf_def before_in_def index_append) + (metis index_take index_take_if_set le_add1 le_trans less_imp_le_nat) with id_take_nth_drop[OF 0, symmetric] show ?thesis by(simp) qed lemma Inv_mtf: "set xs = set ys \ z : set ys \ Inv xs (mtf z ys) = Inv xs ys \ {(x,z)|x. x < z in xs \ x < z in ys} - {(z,x)|x. z < x in xs \ x < z in ys}" by(auto simp add: Inv_def before_in_mtf not_before_in dest: before_in_setD1) lemma set_mtf[simp]: "set(mtf x xs) = set xs" by(simp add: mtf_def) (metis append_take_drop_id Cons_nth_drop_Suc index_less le_refl Un_insert_right nth_index set_append set_simps(2)) lemma length_mtf[simp]: "size (mtf x xs) = size xs" by (auto simp add: mtf_def min_def) (metis index_less_size_conv leD) lemma distinct_mtf[simp]: "distinct (mtf x xs) = distinct xs" by (metis length_mtf set_mtf card_distinct distinct_card) subsection "Function \mtf2\" definition mtf2 :: "nat \ 'a \ 'a list \ 'a list" where "mtf2 n x xs = (if x : set xs then swaps [index xs x - n.. index xs x - (size xs - Suc 0) = 0" by (auto simp: less_Suc_eq_le[symmetric]) thus ?thesis by(auto simp: mtf_def mtf2_def swaps_eq_nth_take_drop) qed lemma mtf20[simp]: "mtf2 0 x xs = xs" by(auto simp add: mtf2_def) lemma length_mtf2[simp]: "length (mtf2 n x xs) = length xs" by (auto simp: mtf2_def index_less_size_conv[symmetric] simp del:index_conv_size_if_notin) lemma set_mtf2[simp]: "set(mtf2 n x xs) = set xs" by (auto simp: mtf2_def index_less_size_conv[symmetric] simp del:index_conv_size_if_notin) lemma distinct_mtf2[simp]: "distinct (mtf2 n x xs) = distinct xs" by (metis length_mtf2 set_mtf2 card_distinct distinct_card) lemma card_Inv_mtf2: "xs!j = ys!0 \ j < length xs \ dist_perm xs ys \ card (Inv (swaps [i.. j" thus ?thesis by simp next assume [arith]: "\ i > j" have 0: "Suc j < length ys" by (metis Suc.prems(2,3) distinct_card) have 1: "(ys ! 0, xs ! j) : Inv ys xs" proof (auto simp: Inv_def) show "ys ! 0 < xs ! j in ys" using Suc.prems by (metis Suc_lessD n_not_Suc_n not_before0 not_before_in nth_eq_iff_index_eq nth_mem) show "xs ! j < ys ! 0 in xs" using Suc.prems by (metis Suc_lessD before_id lessI) qed have 2: "card(Inv ys xs) \ 0" using 1 by auto have "int(card (Inv (swaps [i.. = card (Inv ys (swap j xs)) - int (j-i)" by(simp add: card_Inv_sym) also have "\ = card (Inv ys xs - {(ys ! 0, xs ! j)}) - int (j - i)" using Suc.prems 0 by(simp add: Inv_swap) also have "\ = int(card (Inv ys xs) - 1) - (j - i)" using 1 by(simp add: card_Diff_singleton) also have "\ = card (Inv ys xs) - int (Suc j - i)" using 2 by arith also have "\ = card (Inv xs ys) - int (Suc j - i)" by(simp add: card_Inv_sym) finally show ?thesis . qed qed simp subsection "Function Lxy" definition Lxy :: "'a list \ 'a set \ 'a list" where "Lxy xs S = filter (\z. z\S) xs" thm inter_set_filter lemma Lxy_length_cons: "length (Lxy xs S) \ length (Lxy (x#xs) S)" unfolding Lxy_def by(simp) lemma Lxy_empty[simp]: "Lxy [] S = []" unfolding Lxy_def by simp lemma Lxy_set_filter: "set (Lxy xs S) = S \ set xs" by (simp add: Lxy_def inter_set_filter) lemma Lxy_distinct: "distinct xs \ distinct (Lxy xs S)" by (simp add: Lxy_def) lemma Lxy_append: "Lxy (xs@ys) S = Lxy xs S @ Lxy ys S" by(simp add: Lxy_def) lemma Lxy_snoc: "Lxy (xs@[x]) S = (if x\S then Lxy xs S @ [x] else Lxy xs S)" by(simp add: Lxy_def) lemma Lxy_not: "S \ set xs = {} \ Lxy xs S = []" unfolding Lxy_def apply(induct xs) by simp_all lemma Lxy_notin: "set xs \ S = {} \ Lxy xs S = []" apply(induct xs) by(simp_all add: Lxy_def) lemma Lxy_in: "x\S \ Lxy [x] S = [x]" by(simp add: Lxy_def) lemma Lxy_project: assumes "x\y" "x \ set xs" "y\set xs" "distinct xs" and "x < y in xs" shows "Lxy xs {x,y} = [x,y]" proof - from assms have ij: "index xs x < index xs y" and xinxs: "index xs x < length xs" and yinxs: "index xs y < length xs" unfolding before_in_def by auto from xinxs obtain a as where dec1: "a @ [xs!index xs x] @ as = xs" and "a = take (index xs x) xs" and "as = drop (Suc (index xs x)) xs" and length_a: "length a = index xs x" and length_as: "length as = length xs - index xs x- 1" using id_take_nth_drop by fastforce have "index xs y\length (a @ [xs!index xs x])" using length_a ij by auto then have "((a @ [xs!index xs x]) @ as) ! index xs y = as ! (index xs y-length (a @ [xs ! index xs x]))" using nth_append[where xs="a @ [xs!index xs x]" and ys="as"] by(simp) then have xsj: "xs ! index xs y = as ! (index xs y-index xs x-1)" using dec1 length_a by auto have las: "(index xs y-index xs x-1) < length as" using length_as yinxs ij by simp obtain b c where dec2: "b @ [xs!index xs y] @ c = as" and "b = take (index xs y-index xs x-1) as" "c=drop (Suc (index xs y-index xs x-1)) as" and length_b: "length b = index xs y-index xs x-1" using id_take_nth_drop[OF las] xsj by force have xs_dec: "a @ [xs!index xs x] @ b @ [xs!index xs y] @ c = xs" using dec1 dec2 by auto from xs_dec assms(4) have "distinct ((a @ [xs!index xs x] @ b @ [xs!index xs y]) @ c)" by simp then have c_empty: "set c \ {x,y} = {}" and b_empty: "set b \ {x,y} = {}"and a_empty: "set a \ {x,y} = {}" by(auto simp add: assms(2,3)) have "Lxy (a @ [xs!index xs x] @ b @ [xs!index xs y] @ c) {x,y} = [x,y]" apply(simp only: Lxy_append) apply(simp add: assms(2,3)) using a_empty b_empty c_empty by(simp add: Lxy_notin Lxy_in) with xs_dec show ?thesis by auto qed lemma Lxy_mono: "{x,y} \ set xs \ distinct xs \ x < y in xs = x < y in Lxy xs {x,y}" apply(cases "x=y") apply(simp add: before_in_irefl) proof - assume xyset: "{x,y} \ set xs" assume dxs: "distinct xs" assume xy: "x\y" { fix x y assume 1: "{x,y} \ set xs" assume xny: "x\y" assume 3: "x < y in xs" have "Lxy xs {x,y} = [x,y]" apply(rule Lxy_project) using xny 1 3 dxs by(auto) then have "x < y in Lxy xs {x,y}" using xny by(simp add: before_in_def) } note aha=this have a: "x < y in xs \ x < y in Lxy xs {x,y}" apply(subst Lxy_project) using xy xyset dxs by(simp_all add: before_in_def) have t: "{x,y}={y,x}" by(auto) have f: "~ x < y in xs \ y < x in Lxy xs {x,y}" unfolding t apply(rule aha) using xyset apply(simp) using xy apply(simp) using xy xyset by(simp add: not_before_in) have b: "~ x < y in xs \ ~ x < y in Lxy xs {x,y}" proof - assume "~ x < y in xs" then have "y < x in Lxy xs {x,y}" using f by auto then have "~ x < y in Lxy xs {x,y}" using xy by(simp add: not_before_in) then show ?thesis . qed from a b show ?thesis by metis qed subsection "List Update as Online/Offline Algorithm" type_synonym 'a state = "'a list" type_synonym answer = "nat * nat list" definition step :: "'a state \ 'a \ answer \ 'a state" where "step s r a = (let (k,sws) = a in mtf2 k r (swaps sws s))" definition t :: "'a state \ 'a \ answer \ nat" where "t s r a = (let (mf,sws) = a in index (swaps sws s) r + 1 + size sws)" definition static where "static s rs = (set rs \ set s)" interpretation On_Off step t static . type_synonym 'a alg_off = "'a state \ 'a list \ answer list" type_synonym ('a,'is) alg_on = "('a state,'is,'a,answer) alg_on" lemma T_ge_len: "length as = length rs \ T s rs as \ length rs" by(induction arbitrary: s rule: list_induct2) (auto simp: t_def trans_le_add2) lemma T_off_neq0: "(\rs s0. size(alg s0 rs) = length rs) \ rs \ [] \ T_off alg s0 rs \ 0" apply(erule_tac x=rs in meta_allE) apply(erule_tac x=s0 in meta_allE) apply (auto simp: neq_Nil_conv length_Suc_conv t_def) done lemma length_step[simp]: "length (step s r as) = length s" by(simp add: step_def split_def) lemma step_Nil_iff[simp]: "step xs r act = [] \ xs = []" by(auto simp add: step_def mtf2_def split: prod.splits) lemma set_step2: "set(step s r (mf,sws)) = set s" by(auto simp add: step_def) lemma set_step: "set(step s r act) = set s" by(cases act)(simp add: set_step2) lemma distinct_step: "distinct(step s r as) = distinct s" by (auto simp: step_def split_def) subsection "Online Algorithm Move-to-Front is 2-Competitive" definition MTF :: "('a,unit) alg_on" where "MTF = (\_. (), \s r. ((size (fst s) - 1,[]), ()))" text\It was first proved by Sleator and Tarjan~\cite{SleatorT-CACM85} that the Move-to-Front algorithm is 2-competitive.\ (* The core idea with upper bounds: *) lemma potential: fixes t :: "nat \ 'a::linordered_ab_group_add" and p :: "nat \ 'a" assumes p0: "p 0 = 0" and ppos: "\n. p n \ 0" and ub: "\n. t n + p(n+1) - p n \ u n" shows "(\i (\iii 'a::linordered_ab_group_add" and p :: "nat \ 'a" assumes p0: "p 0 = 0" and ppos: "\n. p n \ 0" and ub: "\m. m t m + p(m+1) - p m \ u m" shows "(\i (\iii \ (\i \ (\i {y. y < x in xs}" abbreviation "after x xs \ {y. x < y in xs}" lemma finite_before[simp]: "finite (before x xs)" apply(rule finite_subset[where B = "set xs"]) apply (auto dest: before_in_setD1) done lemma finite_after[simp]: "finite (after x xs)" apply(rule finite_subset[where B = "set xs"]) apply (auto dest: before_in_setD2) done lemma before_conv_take: "x : set xs \ before x xs = set(take (index xs x) xs)" by (auto simp add: before_in_def set_take_if_index index_le_size) (metis index_take leI) lemma card_before: "distinct xs \ x : set xs \ card (before x xs) = index xs x" using index_le_size[of xs x] by(simp add: before_conv_take distinct_card[OF distinct_take] min_def) lemma before_Un: "set xs = set ys \ x : set xs \ before x ys = before x xs \ before x ys Un after x xs \ before x ys" by(auto)(metis before_in_setD1 not_before_in) lemma phi_diff_aux: "card (Inv xs ys \ {(y, x) |y. y < x in xs \ y < x in ys} - {(x, y) |y. x < y in xs \ y < x in ys}) = card(Inv xs ys) + card(before x xs \ before x ys) - int(card(after x xs \ before x ys))" (is "card(?I \ ?B - ?A) = card ?I + card ?b - int(card ?a)") proof- have 1: "?I \ ?B = {}" by(auto simp: Inv_def) (metis no_before_inI) have 2: "?A \ ?I \ ?B" by(auto simp: Inv_def) have 3: "?A \ ?I" by(auto simp: Inv_def) have "int(card(?I \ ?B - ?A)) = int(card ?I + card ?B) - int(card ?A)" using card_mono[OF _ 3] by(simp add: card_Un_disjoint[OF _ _ 1] card_Diff_subset[OF _ 2]) also have "card ?B = card (fst ` ?B)" by(auto simp: card_image inj_on_def) also have "fst ` ?B = ?b" by force also have "card ?A = card (snd ` ?A)" by(auto simp: card_image inj_on_def) also have "snd ` ?A = ?a" by force finally show ?thesis . qed lemma not_before_Cons[simp]: "\ x < y in y # xs" by (simp add: before_in_def) lemma before_Cons[simp]: "y \ set xs \ y \ x \ before y (x#xs) = insert x (before y xs)" by(auto simp: before_in_def) lemma card_before_le_index: "card (before x xs) \ index xs x" apply(cases "x \ set xs") prefer 2 apply (simp add: before_in_def) apply(induction xs) apply (simp add: before_in_def) apply (auto simp: card_insert_if) done lemma config_config_length: "length (fst (config A init qs)) = length init" apply (induct rule: config_induct) by (simp_all) lemma config_config_distinct: shows " distinct (fst (config A init qs)) = distinct init" apply (induct rule: config_induct) by (simp_all add: distinct_step) lemma config_config_set: shows "set (fst (config A init qs)) = set init" apply(induct rule: config_induct) by(simp_all add: set_step) lemma config_config: "set (fst (config A init qs)) = set init \ distinct (fst (config A init qs)) = distinct init \ length (fst (config A init qs)) = length init" using config_config_distinct config_config_set config_config_length by metis lemma config_dist_perm: "distinct init \ dist_perm (fst (config A init qs)) init" using config_config_distinct config_config_set by metis lemma config_rand_length: "\x\set_pmf (config_rand A init qs). length (fst x) = length init" apply (induct rule: config_rand_induct) by (simp_all) lemma config_rand_distinct: shows "\x \ (config_rand A init qs). distinct (fst x) = distinct init" apply (induct rule: config_rand_induct) by (simp_all add: distinct_step) lemma config_rand_set: shows " \x \ (config_rand A init qs). set (fst x) = set init" apply(induct rule: config_rand_induct) by(simp_all add: set_step) lemma config_rand: "\x \ (config_rand A init qs). set (fst x) = set init \ distinct (fst x) = distinct init \ length (fst x) = length init" using config_rand_distinct config_rand_set config_rand_length by metis lemma config_rand_dist_perm: "distinct init \ \x \ (config_rand A init qs). dist_perm (fst x) init" using config_rand_distinct config_rand_set by metis (*fixme start from Inv*) lemma amor_mtf_ub: assumes "x : set ys" "set xs = set ys" shows "int(card(before x xs Int before x ys)) - card(after x xs Int before x ys) \ 2 * int(index xs x) - card (before x ys)" (is "?m - ?n \ 2 * ?j - ?k") proof- have xxs: "x \ set xs" using assms(1,2) by simp let ?bxxs = "before x xs" let ?bxys = "before x ys" let ?axxs = "after x xs" have 0: "?bxxs \ ?axxs = {}" by (auto simp: before_in_def) hence 1: "(?bxxs \ ?bxys) \ (?axxs \ ?bxys) = {}" by blast have "(?bxxs \ ?bxys) \ (?axxs \ ?bxys) = ?bxys" using assms(2) before_Un xxs by fastforce hence "?m + ?n = ?k" using card_Un_disjoint[OF _ _ 1] by simp hence "?m - ?n = 2 * ?m - ?k" by arith also have "?m \ ?j" using card_before_le_index[of x xs] card_mono[of ?bxxs, OF _ Int_lower1] by(auto intro: order_trans) finally show ?thesis by auto qed locale MTF_Off = fixes as :: "answer list" fixes rs :: "'a list" fixes s0 :: "'a list" assumes dist_s0[simp]: "distinct s0" assumes len_as: "length as = length rs" begin definition mtf_A :: "nat list" where "mtf_A = map fst as" definition sw_A :: "nat list list" where "sw_A = map snd as" fun s_A :: "nat \ 'a list" where "s_A 0 = s0" | "s_A(Suc n) = step (s_A n) (rs!n) (mtf_A!n, sw_A!n)" lemma length_s_A[simp]: "length(s_A n) = length s0" by (induction n) simp_all lemma dist_s_A[simp]: "distinct(s_A n)" by(induction n) (simp_all add: step_def) lemma set_s_A[simp]: "set(s_A n) = set s0" by(induction n) (simp_all add: step_def) fun s_mtf :: "nat \ 'a list" where "s_mtf 0 = s0" | "s_mtf (Suc n) = mtf (rs!n) (s_mtf n)" definition t_mtf :: "nat \ int" where "t_mtf n = index (s_mtf n) (rs!n) + 1" definition T_mtf :: "nat \ int" where "T_mtf n = (\i int" where "c_A n = index (swaps (sw_A!n) (s_A n)) (rs!n) + 1" definition f_A :: "nat \ int" where "f_A n = min (mtf_A!n) (index (swaps (sw_A!n) (s_A n)) (rs!n))" definition p_A :: "nat \ int" where "p_A n = size(sw_A!n)" definition t_A :: "nat \ int" where "t_A n = c_A n + p_A n" definition T_A :: "nat \ int" where "T_A n = (\i int" ("\") where "Phi n = card(Inv (s_A n) (s_mtf n))" lemma phi0: "Phi 0 = 0" by(simp add: Phi_def) lemma phi_pos: "Phi n \ 0" by(simp add: Phi_def) lemma mtf_ub: "t_mtf n + Phi (n+1) - Phi n \ 2 * c_A n - 1 + p_A n - f_A n" proof - let ?xs = "s_A n" let ?ys = "s_mtf n" let ?x = "rs!n" let ?xs' = "swaps (sw_A!n) ?xs" let ?ys' = "mtf ?x ?ys" show ?thesis proof cases assume xin: "?x \ set ?ys" let ?bb = "before ?x ?xs \ before ?x ?ys" let ?ab = "after ?x ?xs \ before ?x ?ys" have phi_mtf: "card(Inv ?xs' ?ys') - int(card (Inv ?xs' ?ys)) \ 2 * int(index ?xs' ?x) - int(card (before ?x ?ys))" using xin by(simp add: Inv_mtf phi_diff_aux amor_mtf_ub) have phi_sw: "card(Inv ?xs' ?ys) \ Phi n + length(sw_A!n)" proof - have "int(card (Inv ?xs' ?ys)) \ card(Inv ?xs' ?xs) + int(card(Inv ?xs ?ys))" using card_Inv_tri_ineq[of ?xs' ?xs ?ys] xin by (simp) also have "card(Inv ?xs' ?xs) = card(Inv ?xs ?xs')" by (rule card_Inv_sym) also have "card(Inv ?xs ?xs') \ size(sw_A!n)" by (metis card_Inv_swaps_le dist_s_A) finally show ?thesis by(fastforce simp: Phi_def) qed have phi_free: "card(Inv ?xs' ?ys') - Phi (Suc n) = f_A n" using xin by(simp add: Phi_def mtf2_def step_def card_Inv_mtf2 index_less_size_conv f_A_def) show ?thesis using xin phi_sw phi_mtf phi_free card_before[of "s_mtf n"] by(simp add: t_mtf_def c_A_def p_A_def) next assume notin: "?x \ set ?ys" have "int (card (Inv ?xs' ?ys)) - card (Inv ?xs ?ys) \ card(Inv ?xs ?xs')" using card_Inv_tri_ineq[OF _ dperm_inv, of ?xs' n] swaps_inv[of "sw_A!n" "s_A n"] by(simp add: card_Inv_sym) also have "\ \ size(sw_A!n)" by(simp add: card_Inv_swaps_le dperm_inv) finally show ?thesis using notin by(simp add: t_mtf_def step_def c_A_def p_A_def f_A_def Phi_def mtf2_def) qed qed theorem Sleator_Tarjan: "T_mtf n \ (\ii (\i = (\i = (\i 2*T_A n - n" proof - have "T_mtf n \ (\ii (\i \ 2* T_A n" by (simp add: sum_distrib_left T_A_def t_A_def) finally show "T_mtf n \ 2* T_A n - n" by auto qed lemma T_A_nneg: "0 \ T_A n" by(auto simp add: sum_nonneg T_A_def t_A_def c_A_def p_A_def) lemma T_mtf_ub: "\i set s0 \ T_mtf n \ n * size s0" proof(induction n) case 0 show ?case by(simp add: T_mtf_def) next case (Suc n) thus ?case using index_less_size_conv[of "s_mtf n" "rs!n"] by(simp add: T_mtf_def t_mtf_def less_Suc_eq del: index_less) qed corollary T_mtf_competitive: assumes "s0 \ []" and "\i set s0" shows "T_mtf n \ (2 - 1/(size s0)) * T_A n" proof cases assume 0: "real_of_int(T_A n) \ n * (size s0)" have "T_mtf n \ 2 * T_A n - n" proof - have "T_mtf n \ (\ii (\i \ 2 * T_A n" by (simp add: sum_distrib_left T_A_def t_A_def) finally show ?thesis by simp qed hence "real_of_int(T_mtf n) \ 2 * of_int(T_A n) - n" by simp also have "\ = 2 * of_int(T_A n) - (n * size s0) / size s0" using assms(1) by simp also have "\ \ 2 * real_of_int(T_A n) - T_A n / size s0" by(rule diff_left_mono[OF divide_right_mono[OF 0]]) simp also have "\ = (2 - 1 / size s0) * T_A n" by algebra finally show ?thesis . next assume 0: "\ real_of_int(T_A n) \ n * (size s0)" have "2 - 1 / size s0 \ 1" using assms(1) by (auto simp add: field_simps neq_Nil_conv) have "real_of_int (T_mtf n) \ n * size s0" using T_mtf_ub[OF assms(2)] by linarith also have "\ < of_int(T_A n)" using 0 by simp also have "\ \ (2 - 1 / size s0) * T_A n" using assms(1) T_A_nneg[of n] by(auto simp add: mult_le_cancel_right1 field_simps neq_Nil_conv) finally show ?thesis by linarith qed lemma t_A_t: "n < length rs \ t_A n = int (t (s_A n) (rs ! n) (as ! n))" by(simp add: t_A_def t_def c_A_def p_A_def sw_A_def len_as split: prod.split) lemma T_A_eq_lem: "(\i=0.. n < length rs" thus ?case by (simp add: len_as) qed qed lemma T_A_eq: "T_A (length rs) = T s0 rs as" using T_A_eq_lem by(simp add: T_A_def atLeast0LessThan) lemma nth_off_MTF: "n < length rs \ off2 MTF s rs ! n = (size(fst s) - 1,[])" by(induction rs arbitrary: s n)(auto simp add: MTF_def nth_Cons' Step_def) lemma t_mtf_MTF: "n < length rs \ t_mtf n = int (t (s_mtf n) (rs ! n) (off MTF s rs ! n))" by(simp add: t_mtf_def t_def nth_off_MTF split: prod.split) lemma mtf_MTF: "n < length rs \ length s = length s0 \ mtf (rs ! n) s = step s (rs ! n) (off MTF s0 rs ! n)" by(auto simp add: nth_off_MTF step_def mtf_eq_mtf2) lemma T_mtf_eq_lem: "(\i=0.. n < length rs" thus ?case by (simp add: len_as) qed qed lemma T_mtf_eq: "T_mtf (length rs) = T_on MTF s0 rs" using T_mtf_eq_lem by(simp add: T_mtf_def atLeast0LessThan) corollary MTF_competitive2: "s0 \ [] \ \i set s0 \ T_on MTF s0 rs \ (2 - 1/(size s0)) * T s0 rs as" by (metis T_mtf_competitive T_A_eq T_mtf_eq of_int_of_nat_eq) corollary MTF_competitive': "T_on MTF s0 rs \ 2 * T s0 rs as" using Sleator_Tarjan'[of "length rs"] T_A_eq T_mtf_eq by auto end theorem compet_MTF: assumes "s0 \ []" "distinct s0" "set rs \ set s0" shows "T_on MTF s0 rs \ (2 - 1/(size s0)) * T_opt s0 rs" proof- from assms(3) have 1: "\i < length rs. rs!i \ set s0" by auto { fix as :: "answer list" assume len: "length as = length rs" interpret MTF_Off as rs s0 proof qed (auto simp: assms(2) len) from MTF_competitive2[OF assms(1) 1] assms(1) have "T_on MTF s0 rs / (2 - 1 / (length s0)) \ of_int(T s0 rs as)" by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) } hence "T_on MTF s0 rs / (2 - 1/(size s0)) \ T_opt s0 rs" apply(simp add: T_opt_def Inf_nat_def) apply(rule LeastI2_wellorder) using length_replicate[of "length rs" undefined] apply fastforce apply auto done thus ?thesis using assms by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) qed theorem compet_MTF': assumes "distinct s0" shows "T_on MTF s0 rs \ (2::real) * T_opt s0 rs" proof- { fix as :: "answer list" assume len: "length as = length rs" interpret MTF_Off as rs s0 proof qed (auto simp: assms(1) len) from MTF_competitive' have "T_on MTF s0 rs / 2 \ of_int(T s0 rs as)" by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) } hence "T_on MTF s0 rs / 2 \ T_opt s0 rs" apply(simp add: T_opt_def Inf_nat_def) apply(rule LeastI2_wellorder) using length_replicate[of "length rs" undefined] apply fastforce apply auto done thus ?thesis using assms by(simp add: field_simps length_greater_0_conv[symmetric] del: length_greater_0_conv) qed theorem MTF_is_2_competitive: "compet MTF 2 {s . distinct s}" unfolding compet_def using compet_MTF' by fastforce subsection "Lower Bound for Competitiveness" text\This result is independent of MTF but is based on the list update problem defined in this theory.\ lemma rat_fun_lem: fixes l c :: real assumes [simp]: "F \ bot" assumes "0 < l" assumes ev: "eventually (\n. l \ f n / g n) F" "eventually (\n. (f n + c) / (g n + d) \ u) F" and g: "LIM n F. g n :> at_top" shows "l \ u" proof (rule dense_le_bounded[OF \0 < l\]) fix x assume x: "0 < x" "x < l" define m where "m = (x - l) / 2" define k where "k = l / (x - m)" have "x = l / k + m" "1 < k" "m < 0" unfolding k_def m_def using x by (auto simp: divide_simps) from \1 < k\ have "LIM n F. (k - 1) * g n :> at_top" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const _ g]) (simp add: field_simps) then have "eventually (\n. d \ (k - 1) * g n) F" by (simp add: filterlim_at_top) moreover have "eventually (\n. 1 \ g n) F" "eventually (\n. 1 - d \ g n) F" "eventually (\n. c / m - d \ g n) F" using g by (auto simp add: filterlim_at_top) ultimately have "eventually (\n. x \ u) F" using ev proof eventually_elim fix n assume d: "d \ (k - 1) * g n" "1 \ g n" "1 - d \ g n" "c / m - d \ g n" and l: "l \ f n / g n" and u: "(f n + c) / (g n + d) \ u" from d have "g n + d \ k * g n" by (simp add: field_simps) from d have g: "0 < g n" "0 < g n + d" by (auto simp: field_simps) with \0 < l\ l have "0 < f n" by (auto simp: field_simps intro: mult_pos_pos less_le_trans) note \x = l / k + m\ also have "l / k \ f n / (k * g n)" using l \1 < k\ by (simp add: field_simps) also have "\ \ f n / (g n + d)" using d \1 < k\ \0 < f n\ by (intro divide_left_mono mult_pos_pos) (auto simp: field_simps) also have "m \ c / (g n + d)" using \c / m - d \ g n\ \0 < g n\ \0 < g n + d\ \m < 0\ by (simp add: field_simps) also have "f n / (g n + d) + c / (g n + d) = (f n + c) / (g n + d)" using \0 < g n + d\ by (auto simp: add_divide_distrib) also note u finally show "x \ u" by simp qed then show "x \ u" by auto qed lemma compet_lb0: fixes a Aon Aoff cruel defines "f s0 rs == real(T_on Aon s0 rs)" defines "g s0 rs == real(T_off Aoff s0 rs)" assumes "\rs s0. size(Aoff s0 rs) = length rs" and "\n. cruel n \ []" assumes "compet Aon c S0" and "c\0" and "s0 \ S0" and l: "eventually (\n. f s0 (cruel n) / (g s0 (cruel n) + a) \ l) sequentially" and g: "LIM n sequentially. g s0 (cruel n) :> at_top" and "l > 0" and "\n. static s0 (cruel n)" shows "l \ c" proof- let ?h = "\b s0 rs. (f s0 rs - b) / g s0 rs" have g': "LIM n sequentially. g s0 (cruel n) + a :> at_top" using filterlim_tendsto_add_at_top[OF tendsto_const g] by (simp add: ac_simps) from competE[OF assms(5) \c\0\ _ \s0 \ S0\] assms(3) obtain b where "\rs. static s0 rs \ rs \ [] \ ?h b s0 rs \ c " by (fastforce simp del: neq0_conv simp: neq0_conv[symmetric] field_simps f_def g_def T_off_neq0[of Aoff, OF assms(3)]) hence "\n. (?h b s0 o cruel) n \ c" using assms(4,11) by simp with rat_fun_lem[OF sequentially_bot \l>0\ _ _ g', of "f s0 o cruel" "-b" "- a" c] assms(7) l show "l \ c" by (auto) qed text \Sorting\ fun ins_sws where "ins_sws k x [] = []" | "ins_sws k x (y#ys) = (if k x \ k y then [] else map Suc (ins_sws k x ys) @ [0])" fun sort_sws where "sort_sws k [] = []" | "sort_sws k (x#xs) = ins_sws k x (sort_key k xs) @ map Suc (sort_sws k xs)" lemma length_ins_sws: "length(ins_sws k x xs) \ length xs" by(induction xs) auto lemma length_sort_sws_le: "length(sort_sws k xs) \ length xs ^ 2" proof(induction xs) case (Cons x xs) thus ?case using length_ins_sws[of k x "sort_key k xs"] by (simp add: numeral_eq_Suc) qed simp lemma swaps_ins_sws: "swaps (ins_sws k x xs) (x#xs) = insort_key k x xs" by(induction xs)(auto simp: swap_def[of 0]) lemma swaps_sort_sws[simp]: "swaps (sort_sws k xs) xs = sort_key k xs" by(induction xs)(auto simp: swaps_ins_sws) text\The cruel adversary:\ fun cruel :: "('a,'is) alg_on \ 'a state * 'is \ nat \ 'a list" where "cruel A s 0 = []" | "cruel A s (Suc n) = last (fst s) # cruel A (Step A s (last (fst s))) n" definition adv :: "('a,'is) alg_on \ ('a::linorder) alg_off" where "adv A s rs = (if rs=[] then [] else let crs = cruel A (Step A (s, fst A s) (last s)) (size rs - 1) in (0,sort_sws (\x. size rs - 1 - count_list crs x) s) # replicate (size rs - 1) (0,[]))" lemma set_cruel: "s \ [] \ set(cruel A (s,is) n) \ set s" apply(induction n arbitrary: s "is") apply(auto simp: step_def Step_def split: prod.split) by (metis empty_iff swaps_inv last_in_set list.set(1) rev_subsetD set_mtf2) lemma static_cruel: "s \ [] \ static s (cruel A (s,is) n)" by(simp add: set_cruel static_def) (* Do not convert into structured proof - eta conversion screws it up! *) lemma T_cruel: "s \ [] \ distinct s \ T s (cruel A (s,is) n) (off2 A (s,is) (cruel A (s,is) n)) \ n*(length s)" apply(induction n arbitrary: s "is") apply(simp) apply(erule_tac x = "fst(Step A (s, is) (last s))" in meta_allE) apply(erule_tac x = "snd(Step A (s, is) (last s))" in meta_allE) apply(frule_tac sws = "snd(fst(snd A (s,is) (last s)))" in index_swaps_last_size) apply(simp add: distinct_step t_def split_def Step_def length_greater_0_conv[symmetric] del: length_greater_0_conv) done lemma length_cruel[simp]: "length (cruel A s n) = n" by (induction n arbitrary: s) (auto) lemma t_sort_sws: "t s r (mf, sort_sws k s) \ size s ^ 2 + size s + 1" using length_sort_sws_le[of k s] index_le_size[of "sort_key k s" r] by (simp add: t_def add_mono index_le_size algebra_simps) lemma T_noop: "n = length rs \ T s rs (replicate n (0, [])) = (\r\rs. index s r + 1)" by(induction rs arbitrary: s n)(auto simp: t_def step_def) lemma sorted_asc: "j\i \ i \x \ set ss. \y \ set ss. k(x) \ k(y) \ f y \ f x \ sorted (map k ss) \ f (ss ! i) \ f (ss ! j)" by (auto simp: sorted_iff_nth_mono) lemma sorted_weighted_gauss_Ico_div2: fixes f :: "nat \ nat" assumes "\i j. i \ j \ j < n \ f i \ f j" shows "(\i=0.. (n + 1) * sum f {0..i=0.. (\i=0..i=0..n. Suc i * f i)) \ 2 * (\i=0..n. Suc i) * sum f {0..n}" by (simp add: atLeastLessThanSuc_atLeastAtMost) also have "2 * (\i=0..n. Suc i) = Suc n * (n + 2)" using arith_series_nat [of 1 1 n] by simp finally have "2 * (\i=0..n. Suc i * f i) \ (n + 2) * sum f {0..n}" by (simp only: ac_simps Suc_mult_le_cancel1) with Suc show ?thesis by (simp only: atLeastLessThanSuc_atLeastAtMost) simp qed lemma T_adv: assumes "l \ 0" shows "T_off (adv A) [0.. l\<^sup>2 + l + 1 + (l + 1) * n div 2" (is "?l \ ?r") proof- let ?s = "[0..x. x < l \ ?c x \ n" by(simp) (metis count_le_length length_cruel) have "?l = t ?s (last ?s) (0, sort_sws ?k ?s) + (\x\set ?s'. ?c x * (index ?sort x + 1))" using assms apply(simp add: adv_def T_noop sum_list_map_eq_sum_count2[OF set_cruel] Step_def split: prod.split) apply(subst (3) step_def) apply(simp) done also have "(\x\set ?s'. ?c x * (index ?sort x + 1)) = (\x\{0.. = (\x\{0.. = (\x\{0.. \ (\x\{0.. \ (l+1) * (\x\{0..x\{0..x\{0.. = (\x\{0.. = length ?cr" using set_cruel[of ?s' A _ n] assms 1 by(auto simp add: sum_count_set Step_def split: prod.split) also have "\ = n" by simp also have "t ?s (last ?s) (0, sort_sws ?k ?s) \ (length ?s)^2 + length ?s + 1" by(rule t_sort_sws) also have "\ = l^2 + l + 1" by simp finally show "?l \ l\<^sup>2 + l + 1 + (l + 1) * n div 2" by auto qed text \The main theorem:\ theorem compet_lb2: assumes "compet A c {xs::nat list. size xs = l}" and "l \ 0" and "c \ 0" shows "c \ 2*l/(l+1)" proof (rule compet_lb0[OF _ _ assms(1) \c\0\]) let ?S0 = "{xs::nat list. size xs = l}" let ?s0 = "[0..s0 rs. length (adv A s0 rs) = length rs" by(simp add: adv_def) show "\n. ?cruel n \ []" by auto show "?s0 \ ?S0" by simp { fix Z::real and n::nat assume "n \ nat(ceiling Z)" have "?off n \ length(?cruel n)" by(rule T_ge_len) (simp add: adv_def) hence "?off n > n" by simp hence "Z \ ?off n" using \n \ nat(ceiling Z)\ by linarith } thus "LIM n sequentially. real (?off n) :> at_top" by(auto simp only: filterlim_at_top eventually_sequentially) let ?a = "- (l^2 + l + 1)" { fix n assume "n \ l^2 + l + 1" have "2*l/(l+1) = 2*l*(n+1) / ((l+1)*(n+1))" by (simp del: One_nat_def) also have "\ = 2*real(l*(n+1)) / ((l+1)*(n+1))" by simp also have "l * (n+1) \ ?on n" using T_cruel[of ?s0 "Suc n"] \l \ 0\ by (simp add: ac_simps) also have "2*real(?on n) / ((l+1)*(n+1)) \ 2*real(?on n)/(2*(?off n + ?a))" proof - have 0: "2*real(?on n) \ 0" by simp have 1: "0 < real ((l + 1) * (n + 1))" by (simp del: of_nat_Suc) have "?off n \ length(?cruel n)" by(rule T_ge_len) (simp add: adv_def) hence "?off n > n" by simp hence "?off n + ?a > 0" using \n \ l^2 + l + 1\ by linarith hence 2: "real_of_int(2*(?off n + ?a)) > 0" by(simp only: of_int_0_less_iff zero_less_mult_iff zero_less_numeral simp_thms) have "?off n + ?a \ (l+1)*(n) div 2" using T_adv[OF \l\0\, of A n] by (simp only: o_apply of_nat_add of_nat_le_iff) also have "\ \ (l+1)*(n+1) div 2" by (simp) finally have "2*(?off n + ?a) \ (l+1)*(n+1)" by (simp add: zdiv_int) hence "of_int(2*(?off n + ?a)) \ real((l+1)*(n+1))" by (simp only: of_int_le_iff) from divide_left_mono[OF this 0 mult_pos_pos[OF 1 2]] show ?thesis . qed also have "\ = ?on n / (?off n + ?a)" by (simp del: distrib_left_numeral One_nat_def cruel.simps) finally have "2*l/(l+1) \ ?on n / (real (?off n) + ?a)" by (auto simp: divide_right_mono) } thus "eventually (\n. (2 * l) / (l + 1) \ ?on n / (real(?off n) + ?a)) sequentially" by(auto simp add: filterlim_at_top eventually_sequentially) show "0 < 2*l / (l+1)" using \l \ 0\ by(simp) show "\n. static ?s0 (?cruel n)" using \l \ 0\ by(simp add: static_cruel del: cruel.simps) qed end diff --git a/thys/Locally-Nameless-Sigma/preliminary/ListPre.thy b/thys/Locally-Nameless-Sigma/preliminary/ListPre.thy --- a/thys/Locally-Nameless-Sigma/preliminary/ListPre.thy +++ b/thys/Locally-Nameless-Sigma/preliminary/ListPre.thy @@ -1,205 +1,199 @@ (* Title: ListPre.thy Author: Ludovic Henrio and Florian Kammuller, 2006 List lemmata and other general stuff as preparation for ASP. *) section \List features\ theory ListPre imports Main begin lemma drop_lem[rule_format]: fixes n :: nat and l :: "'a list" and g :: "'a list" assumes "drop n l = drop n g" and "length l = length g" and "n < length g" shows "l!n = g!n" proof - from assms(2-3) have "n < length l" by simp from Cons_nth_drop_Suc[OF this] Cons_nth_drop_Suc[OF assms(3)] assms(1) have "l!n # drop (Suc n) l = g!n # drop (Suc n) g" by simp thus ?thesis by simp qed lemma mem_append_lem': "x \ set (l @ [y]) \ x \ set l \ x = y" by auto lemma nth_last: "length l = n \ (l @ [x])!n = x" by auto lemma take_n: fixes n :: nat and l :: "'a list" and g :: "'a list" assumes "take n l = take n g" and "Suc n \ length g" and "length l = length g" shows "take (Suc n) (l[n := g!n]) = take (Suc n) g" proof - from assms(2) have ng: "n < length g" by simp with assms(3) have nlupd: "n < length (l[n := g!n])" by simp hence nl: "n < length l" by simp from sym[OF assms(1)] id_take_nth_drop[OF ng] take_Suc_conv_app_nth[OF nlupd] nth_list_update_eq[OF nl] take_Suc_conv_app_nth[OF ng] upd_conv_take_nth_drop[OF nl] assms(2-3) show ?thesis by simp qed lemma drop_n_lem: fixes n :: nat and l :: "'a list" assumes "Suc n \ length l" shows "drop (Suc n) (l[n := x]) = drop (Suc n) l" -proof - - from assms have "n < length l" by simp - from - upd_conv_take_nth_drop[OF this] drop_Suc[of n "l[n := x]"] - drop_Suc[of n l] assms - show ?thesis by (simp,(subst drop_tl)+, simp) -qed + using assms by simp lemma drop_n: fixes n :: nat and l :: "'a list" and g :: "'a list" assumes "drop n l = drop n g" and "Suc n \ length g" and "length l = length g" shows "drop (Suc n) (l[n := g!n]) = drop (Suc n) g" proof - from assms(2-3) have "Suc n \ length l" by simp from drop_n_lem[OF this] assms(1) show ?thesis by (simp, (subst drop_Suc)+, (subst drop_tl)+, simp) qed lemma nth_fst[rule_format]: "length l = n + 1 \ (l @ [x])!0 = l!0" by (induct l, simp_all) lemma nth_zero_app: fixes l :: "'a list" and x :: 'a and y :: 'a assumes "l \ []" and "l!0 = x" shows"(l @ [y])!0 = x" proof - have "l \ [] \ l!0 = x \ (l @ [y])!0 = x" by (induct l, simp_all) with assms show ?thesis by simp qed lemma rev_induct2[consumes 1]: fixes xs :: "'a list" and ys :: "'a list" and P :: "'a list \ 'a list \ bool" assumes "length xs = length ys" and "P [] []" and "\x xs y ys. \ length xs = length ys; P xs ys \ \ P (xs @ [x]) (ys @ [y])" shows "P xs ys" proof (simplesubst rev_rev_ident[symmetric]) from assms(1) have lrev: "length (rev xs) = length (rev ys)" by simp from assms have "P (rev (rev xs))(rev (rev ys))" by (induct rule: list_induct2[OF lrev], simp_all) thus "P xs (rev (rev ys))" by simp qed lemma list_induct3: (* Similar to induction for 2: see ML "thm \"list_induct2\""; *) "\ys zs. \ length xs = length ys; length zs = length xs; P [] [] []; \x xs y ys z zs. \ length xs = length ys; length zs = length xs; P xs ys zs \ \ P (x # xs)(y # ys)(z # zs) \ \ P xs ys zs" proof (induct xs, simp) case (Cons a xs ys zs) from \length (a#xs) = length ys\ \length zs = length (a#xs)\ have "ys \ [] \ zs \ []" by auto then obtain b ly c lz where "ys = b#ly" and "zs = c#lz" by (auto simp: neq_Nil_conv) with \length (a#xs) = length ys\ \length zs = length (a#xs)\ obtain "length xs = length ly" and "length lz = length xs" by auto from Cons(5)[OF this Cons(1)[OF this \P [] [] []\]] Cons(5) \ys = b#ly\ \zs = c#lz\ show ?case by simp qed primrec list_insert :: "'a list \ nat \ 'a \ 'a list" where "list_insert (ah#as) i a = (case i of 0 \ a#ah#as |Suc j \ ah#(list_insert as j a))" | "list_insert [] i a = [a]" lemma insert_eq[simp]: "\i\length l. (list_insert l i a)!i = a" by (induct l, simp, intro strip, simp split: nat.split) lemma insert_gt[simp]: "\i\length l. \j length l" and "j < Suc n" with Cons(1) show "(x#(list_insert l n a))!j = (x#l)!j" by (cases j) simp_all qed qed lemma insert_lt[simp]: "\j\length l. \i\j. (list_insert l i a)!Suc j = l!j" proof (induct l, simp) case (Cons x l) thus ?case proof (auto split: nat.split) fix n j assume "j \ Suc (length l)" and "Suc n \ j" with Cons(1) show "(list_insert l n a)!j = l!(j - Suc 0)" by (cases j) simp_all qed qed lemma insert_first[simp]: "list_insert l 0 b = b#l" by (induct l, simp_all) lemma insert_prepend[simp]: "i = Suc j \ list_insert (a#l) i b = a # list_insert l j b" by auto lemma insert_lt2[simp]: "\j. \i\j. (list_insert l i a)!Suc j = l!j" proof (induct l, simp) case (Cons x l) thus ?case proof (auto split: nat.split) fix n j assume "Suc n \ j" with Cons(1) show "(list_insert l n a)!j = l!(j - Suc 0)" by (cases j) simp_all qed qed lemma insert_commute[simp]: "\i\length l. (list_insert (list_insert l i b) 0 a) = (list_insert (list_insert l 0 a ) (Suc i) b)" by (induct l, auto split: nat.split) lemma insert_length': "\i x. length (list_insert l i x) = length (x#l)" by (induct l, auto split: nat.split) lemma insert_length[simp]: "length (list_insert l i b) = length (list_insert l j c)" by (simp add: insert_length') lemma insert_select[simp]: "the ((f(l \ t)) l) = t" by auto lemma dom_insert[simp]: "l \ dom f \ dom (f(l \ t)) = dom f" by auto lemma insert_select2[simp]: "l1 \ l2 \ ((f(l1 \ t)) l2) = (f l2)" by auto lemma the_insert_select[simp]: "\ l2 \ dom f; l1 \ l2 \ \ the ((f(l1 \ t)) l2) = the (f l2)" by auto lemma insert_dom_eq: "dom f = dom f' \ dom (f(l \ x)) = dom (f'(l \ x'))" by auto lemma insert_dom_less_eq: "\ x \ dom f; x \ dom f'; dom (f(x \ y)) = dom (f'(x \ y')) \ \ dom f = dom f'" by auto lemma one_more_dom[rule_format]: "\l\dom f . \f'. f = f'(l \ the(f l)) \ l \ dom f'" proof fix l assume "l \ dom f" hence "\la. f la = ((\la. if la = l then None else f la)(l \ the (f l))) la" by auto hence "f = (\la. if la = l then None else f la)(l \ the (f l))" by (rule ext) thus "\f'. f = f'(l \ the(f l) ) \ l \ dom f'" by auto qed end diff --git a/thys/Lp/Lp.thy b/thys/Lp/Lp.thy --- a/thys/Lp/Lp.thy +++ b/thys/Lp/Lp.thy @@ -1,2245 +1,2246 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) theory Lp imports Functional_Spaces begin text \The material in this file is essentially of analytic nature. However, one of the central proofs (the proof of Holder inequality below) uses a probability space, and Jensen's inequality there. Hence, we need to import \verb+Probability+. Moreover, we use several lemmas from \verb+SG_Library_Complement+.\ section \Conjugate exponents\ text \Two numbers $p$ and $q$ are \emph{conjugate} if $1/p + 1/q = 1$. This relation keeps appearing in the theory of $L^p$ spaces, as the dual of $L^p$ is $L^q$ where $q$ is the conjugate of $p$. This relation makes sense for real numbers, but also for ennreals (where the case $p=1$ and $q=\infty$ is most important). Unfortunately, manipulating the previous relation with ennreals is tedious as there is no good simproc involving addition and division there. To mitigate this difficulty, we prove once and for all most useful properties of such conjugates exponents in this paragraph.\ lemma Lp_cases_1_PInf: assumes "p \ (1::ennreal)" obtains (gr) p2 where "p = ennreal p2" "p2 > 1" "p > 1" | (one) "p = 1" | (PInf) "p = \" using assms by (metis (full_types) antisym_conv ennreal_cases ennreal_le_1 infinity_ennreal_def not_le) lemma Lp_cases: obtains (real_pos) p2 where "p = ennreal p2" "p2 > 0" "p > 0" | (zero) "p = 0" | (PInf) "p = \" by (metis enn2real_positive_iff ennreal_enn2real_if infinity_ennreal_def not_gr_zero top.not_eq_extremum) definition "conjugate_exponent p = 1 + 1/(p-1)" lemma conjugate_exponent_real: assumes "p > (1::real)" shows "1/p + 1/(conjugate_exponent p) = 1" "conjugate_exponent p > 1" "conjugate_exponent(conjugate_exponent p) = p" "(p-1) * conjugate_exponent p = p" "p - p / conjugate_exponent p = 1" unfolding conjugate_exponent_def using assms by (auto simp add: algebra_simps divide_simps) lemma conjugate_exponent_real_iff: assumes "p > (1::real)" shows "q = conjugate_exponent p \ (1/p + 1/q = 1)" unfolding conjugate_exponent_def using assms by (auto simp add: algebra_simps divide_simps) lemma conjugate_exponent_real_2 [simp]: "conjugate_exponent (2::real) = 2" unfolding conjugate_exponent_def by (auto simp add: algebra_simps divide_simps) lemma conjugate_exponent_realI: assumes "p > (0::real)" "q > 0" "1/p + 1/q = 1" shows "p > 1" "q = conjugate_exponent p" "q > 1" "p = conjugate_exponent q" unfolding conjugate_exponent_def using assms apply (auto simp add: algebra_simps divide_simps) apply (metis assms(3) divide_less_eq_1_pos less_add_same_cancel1 zero_less_divide_1_iff) using mult_less_cancel_left_pos by fastforce lemma conjugate_exponent_real_ennreal: assumes "p> (1::real)" shows "conjugate_exponent(ennreal p) = ennreal(conjugate_exponent p)" unfolding conjugate_exponent_def using assms by (auto, metis diff_gt_0_iff_gt divide_ennreal ennreal_1 ennreal_minus zero_le_one) lemma conjugate_exponent_ennreal_1_2_PInf [simp]: "conjugate_exponent (1::ennreal) = \" "conjugate_exponent (\::ennreal) = 1" "conjugate_exponent (\::ennreal) = 1" "conjugate_exponent (2::ennreal) = 2" using conjugate_exponent_real_ennreal[of 2] by (auto simp add: conjugate_exponent_def) lemma conjugate_exponent_ennreal: assumes "p \ (1::ennreal)" shows "1/p + 1/(conjugate_exponent p) = 1" "conjugate_exponent p \ 1" "conjugate_exponent(conjugate_exponent p) = p" proof - have "(1/p + 1/(conjugate_exponent p) = 1) \ (conjugate_exponent p \ 1) \ conjugate_exponent(conjugate_exponent p) = p" using \p \ 1\ proof (cases rule: Lp_cases_1_PInf) case (gr p2) then have *: "conjugate_exponent p = ennreal (conjugate_exponent p2)" using conjugate_exponent_real_ennreal[OF \p2 > 1\] by auto have a: "conjugate_exponent p \ 1" using * conjugate_exponent_real[OF \p2 > 1\] by auto have b: "conjugate_exponent(conjugate_exponent p) = p" using conjugate_exponent_real(3)[OF \p2 > 1\] conjugate_exponent_real_ennreal[OF \p2 > 1\] conjugate_exponent_real_ennreal[OF conjugate_exponent_real(2)[OF \p2 > 1\]] unfolding * \p = ennreal p2\ by auto have "1 / p + 1 / conjugate_exponent p = ennreal(1/p2 + 1/(conjugate_exponent p2))" unfolding * unfolding \p = ennreal p2\ using conjugate_exponent_real(2)[OF \p2 > 1\] \p2 > 1\ apply (subst ennreal_plus, auto) apply (subst divide_ennreal[symmetric], auto) using divide_ennreal_def inverse_ennreal inverse_eq_divide by auto then have c: "1 / p + 1 / conjugate_exponent p = 1" using conjugate_exponent_real[OF \p2 > 1\] by auto show ?thesis using a b c by simp qed (auto) then show "1/p + 1/(conjugate_exponent p) = 1" "conjugate_exponent p \ 1" "conjugate_exponent(conjugate_exponent p) = p" by auto qed lemma conjugate_exponent_ennreal_iff: assumes "p \ (1::ennreal)" shows "q = conjugate_exponent p \ (1/p + 1/q = 1)" using conjugate_exponent_ennreal[OF assms] by (auto, metis ennreal_add_diff_cancel_left ennreal_add_eq_top ennreal_top_neq_one one_divide_one_divide_ennreal) lemma conjugate_exponent_ennrealI: assumes "1/p + 1/q = (1::ennreal)" shows "p \ 1" "q \ 1" "p = conjugate_exponent q" "q = conjugate_exponent p" proof - have "1/p \ 1" using assms using le_iff_add by fastforce then show "p \ 1" by (metis assms divide_ennreal_def ennreal_add_eq_top ennreal_divide_self ennreal_divide_zero ennreal_le_epsilon ennreal_one_neq_top mult.left_neutral mult_left_le zero_le) then show "q = conjugate_exponent p" using conjugate_exponent_ennreal_iff assms by auto then show "q \ 1" using conjugate_exponent_ennreal[OF \p \ 1\] by auto show "p = conjugate_exponent q" using conjugate_exponent_ennreal_iff[OF \q\1\, of p] assms by (simp add: add.commute) qed section \Convexity inequalities and integration\ text \In this paragraph, we describe the basic inequalities relating the integral of a function and of its $p$-th power, for $p > 0$. These inequalities imply in particular that the $L^p$ norm satisfies the triangular inequality, a feature we will need when defining the $L^p$ spaces below. In particular, we prove the Hölder and Minkowski inequalities. The Hölder inequality, especially, is the basis of all further inequalities for $L^p$ spaces. \ lemma (in prob_space) bound_L1_Lp: assumes "p \ (1::real)" "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" shows "integrable M f" "abs(\x. f x \M) powr p \ (\x. \f x\ powr p \M)" "abs(\x. f x \M) \ (\x. \f x\ powr p \M) powr (1/p)" proof - have *: "norm x \ 1 + (norm x) powr p" for x::real apply (cases "norm x \ 1") apply (meson le_add_same_cancel1 order.trans powr_ge_pzero) apply (metis add_le_same_cancel2 assms(1) less_le_trans linear not_less not_one_le_zero powr_le_cancel_iff powr_one_gt_zero_iff) done show *: "integrable M f" apply (rule Bochner_Integration.integrable_bound[of _ "\x. 1 + \f x\ powr p"], auto simp add: assms) using * by auto show "abs(\x. f x \M) powr p \ (\x. \f x\ powr p \M)" by (rule jensens_inequality[OF * _ _ assms(3) convex_abs_powr[OF \p \ 1\]], auto) then have "(abs(\x. f x \M) powr p) powr (1/p) \ (\x. \f x\ powr p \M) powr (1/p)" using assms(1) powr_mono2 by auto then show "abs(\x. f x \M) \ (\x. \f x\ powr p \M) powr (1/p)" using \p \ 1\ by (auto simp add: powr_powr) qed theorem Holder_inequality: assumes "p > (0::real)" "q > 0" "1/p + 1/q = 1" and [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr q)" shows "integrable M (\x. f x * g x)" "(\x. \f x * g x\ \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" "abs(\x. f x * g x \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" proof - have "p > 1" using conjugate_exponent_realI(1)[OF \p>0\ \q>0\ \1/p+1/q=1\]. have *: "x * y \ x powr p + y powr q" if "x \ 0" "y \ 0" for x y proof - have "x * y = (x powr p) powr (1/p) * (y powr q) powr (1/q)" using \p > 0\ \q > 0\ powr_powr that(1) that(2) by auto also have "... \ (max (x powr p) (y powr q)) powr (1/p) * (max (x powr p) (y powr q)) powr (1/q)" apply (rule mult_mono, auto) using assms(1) assms(2) powr_mono2 by auto also have "... = max (x powr p) (y powr q)" by (metis max_def mult.right_neutral powr_add powr_powr assms(3)) also have "... \ x powr p + y powr q" by auto finally show ?thesis by simp qed show [simp]: "integrable M (\x. f x * g x)" apply (rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p + \g x\ powr q"], auto) by (rule Bochner_Integration.integrable_add, auto simp add: assms * abs_mult) text \The proof of the main inequality is done by applying the inequality $(\int |h| d\mu \leq \int |h|^p d\mu)^{1/p}$ to the right function $h$ in the right probability space. One should take $h = f \cdot |g|^{1-q}$, and $d\mu = |g|^q dM / I$, where $I = \int |g|^q$. This readily gives the result.\ show *: "(\x. \f x * g x\ \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" proof (cases "(\x. \g x\ powr q \M) = 0") case True then have "AE x in M. \g x\ powr q = 0" by (subst integral_nonneg_eq_0_iff_AE[symmetric], auto simp add: assms) then have *: "AE x in M. f x * g x = 0" using \q > 0\ by auto have "(\x. \f x * g x\ \M) = (\x. 0 \M)" apply (rule integral_cong_AE) using * by auto then show ?thesis by auto next case False moreover have "(\x. \g x\ powr q \M) \ (\x. 0 \M)" by (rule integral_mono, auto simp add: assms) ultimately have *: "(\x. \g x\ powr q \M) > 0" by (simp add: le_less) define I where "I = (\x. \g x\ powr q \M)" have [simp]: "I > 0" unfolding I_def using * by auto define M2 where "M2 = density M (\x. \g x\ powr q / I)" interpret prob_space M2 apply (standard, unfold M2_def, auto, subst emeasure_density, auto) apply (subst divide_ennreal[symmetric], auto, subst nn_integral_divide, auto) apply (subst nn_integral_eq_integral, auto simp add: assms, unfold I_def) using * by auto have [simp]: "p \ 1" "p \ 0" using \p > 1\ by auto have A: "q + (1 - q) * p = 0" using assms by (auto simp add: divide_simps algebra_simps) have B: "1 - 1/p = 1/q" using \1/p + 1/q = 1\ by auto define f2 where "f2 = (\x. f x * indicator {y\ space M. g y \ 0} x)" have [measurable]: "f2 \ borel_measurable M" unfolding f2_def by auto define h where "h = (\x. \f2 x\ * \g x\ powr (1-q))" have [measurable]: "h \ borel_measurable M" unfolding h_def by auto have [measurable]: "h \ borel_measurable M2" unfolding M2_def by auto have Eq: "(\g x\ powr q / I) *\<^sub>R \h x\ powr p = \f2 x\ powr p / I" for x apply (insert \I>0\, auto simp add: divide_simps, unfold h_def) apply (auto simp add: divide_nonneg_pos divide_simps powr_mult powr_powr powr_add[symmetric] A) unfolding f2_def by auto have "integrable M2 (\x. \h x\ powr p)" unfolding M2_def apply (subst integrable_density, simp, simp, simp add: divide_simps) apply (subst Eq, rule integrable_divide, rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p"], unfold f2_def) by (unfold indicator_def, auto simp add: \integrable M (\x. \f x\ powr p)\) then have "integrable M2 (\x. \h x\)" by (metis bound_L1_Lp(1) \random_variable borel h\ \p > 1\ integrable_abs le_less) have "(\x. \h x\ powr p \M2) = (\x. (\g x\ powr q / I) *\<^sub>R (\h x\ powr p) \M)" unfolding M2_def by (rule integral_density[of "\x. \h x\ powr p" M "\x. \g x\ powr q / I"], auto simp add: divide_simps) also have "... = (\x. \f2 x\ powr p / I \M)" apply (rule Bochner_Integration.integral_cong) using Eq by auto also have "... \ (\x. \f x\ powr p / I \M)" apply (rule integral_mono', rule integrable_divide[OF \integrable M (\x. \f x\ powr p)\]) unfolding f2_def indicator_def using \I > 0\ by (auto simp add: divide_simps) finally have C: "(\x. \h x\ powr p \M2) \ (\x. \f x\ powr p / I \M)" by simp have "(\x. \f x * g x\ \M) / I = (\x. \f x * g x\ / I \M)" by auto also have "... = (\x. \f2 x * g x\ / I \M)" by (auto simp add: divide_simps, rule Bochner_Integration.integral_cong, unfold f2_def indicator_def, auto) also have "... = (\x. \h x\ \M2)" apply (unfold M2_def, subst integral_density, simp, simp, simp add: divide_simps) by (rule Bochner_Integration.integral_cong, unfold h_def, auto simp add: divide_simps algebra_simps powr_add[symmetric] abs_mult) also have "... \ abs (\x. \h x\ \M2)" by auto also have "... \ (\x. abs(\h x\) powr p \M2) powr (1/p)" apply (rule bound_L1_Lp(3)[of p "\x. \h x\"]) by (auto simp add: \integrable M2 (\x. \h x\ powr p)\) also have "... \ (\x. \f x\ powr p / I \M) powr (1/p)" by (rule powr_mono2, insert C, auto) also have "... \ ((\x. \f x\ powr p \M) / I) powr (1/p)" apply (rule powr_mono2, auto simp add: divide_simps) using \p \ 0\ by auto also have "... = (\x. \f x\ powr p \M) powr (1/p) * I powr(-1/p)" by (auto simp add: less_imp_le powr_divide powr_minus_divide) finally have "(\x. \f x * g x\ \M) \ (\x. \f x\ powr p \M) powr (1/p) * I * I powr(-1/p)" by (auto simp add: divide_simps algebra_simps) also have "... = (\x. \f x\ powr p \M) powr (1/p) * I powr (1-1/p)" by (auto simp add: powr_mult_base less_imp_le) also have "... = (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" unfolding I_def using B by auto finally show ?thesis by simp qed have "abs(\x. f x * g x \M) \ (\x. \f x * g x\ \M)" by auto then show "abs(\x. f x * g x \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" using * by linarith qed theorem Minkowski_inequality: assumes "p \ (1::real)" and [measurable, simp]: "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr p)" shows "integrable M (\x. \f x + g x\ powr p)" "(\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p)" proof - have *: "\x + y\ powr p \ 2 powr p * (\x\ powr p + \y\ powr p)" for x y::real proof - have "\x + y\ \ \x\ + \y\" by auto also have "... \ (max \x\ \y\) + max \x\ \y\" by auto also have "... = 2 * max \x\ \y\" by auto finally have "\x + y\ powr p \ (2 * max \x\ \y\) powr p" using powr_mono2 \p \ 1\ by auto also have "... = 2 powr p * (max \x\ \y\) powr p" using powr_mult by auto also have "... \ 2 powr p * (\x\ powr p + \y\ powr p)" unfolding max_def by auto finally show ?thesis by simp qed show [simp]: "integrable M (\x. \f x + g x\ powr p)" by (rule Bochner_Integration.integrable_bound[of _ "\x. 2 powr p * (\f x\ powr p + \g x\ powr p)"], auto simp add: *) show "(\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p)" proof (cases "p=1") case True then show ?thesis apply (auto, subst Bochner_Integration.integral_add[symmetric], insert assms(4) assms(5), simp, simp) by (rule integral_mono', auto) next case False then have [simp]: "p > 1" "p \ 1" "p > 0" "p \ 0" using assms(1) by auto define q where "q = conjugate_exponent p" have [simp]: "q > 1" "q > 0" "1/p + 1/q = 1" "(p-1) * q = p" unfolding q_def using conjugate_exponent_real[OF \p>1\] by auto then have [simp]: "(z powr (p-1)) powr q = z powr p" for z by (simp add: powr_powr) have "(\x. \f x + g x\ powr p \M) = (\x. \f x + g x\ * \f x + g x\ powr (p-1) \M)" by (subst powr_mult_base, auto) also have "... \ (\x. \f x\ * \f x + g x\ powr (p-1) + \g x\ * \f x + g x\ powr (p-1) \M)" apply (rule integral_mono', rule Bochner_Integration.integrable_add) apply (rule Holder_inequality(1)[of p q], auto) apply (rule Holder_inequality(1)[of p q], auto) by (metis abs_ge_zero abs_triangle_ineq comm_semiring_class.distrib le_less mult_mono' powr_ge_pzero) also have "... = (\x. \f x\ * \f x + g x\ powr (p-1) \M) + (\x. \g x\ * \f x + g x\ powr (p-1) \M)" apply (rule Bochner_Integration.integral_add) by (rule Holder_inequality(1)[of p q], auto)+ also have "... \ abs (\x. \f x\ * \f x + g x\ powr (p-1) \M) + abs (\x. \g x\ * \f x + g x\ powr (p-1) \M)" by auto also have "... \ (\x. abs(\f x\) powr p \M) powr (1/p) * (\x. abs(\f x + g x\ powr (p-1)) powr q \M) powr (1/q) + (\x. abs(\g x\) powr p \M) powr (1/p) * (\x. abs(\f x + g x\ powr (p-1)) powr q \M) powr (1/q)" apply (rule add_mono) apply (rule Holder_inequality(3)[of p q], simp, simp, simp, simp, simp, simp, simp) apply (rule Holder_inequality(3)[of p q], simp, simp, simp, simp, simp, simp, simp) done also have "... = (\x. \f x + g x\ powr p \M) powr (1/q) * ((\x. abs(\f x\) powr p \M) powr (1/p) + (\x. abs(\g x\) powr p \M) powr (1/p))" by (auto simp add: algebra_simps) finally have *: "(\x. \f x + g x\ powr p \M) \ (\x. \f x + g x\ powr p \M) powr (1/q) * ((\x. abs(\f x\) powr p \M) powr (1/p) + (\x. abs(\g x\) powr p \M) powr (1/p))" by simp show ?thesis proof (cases "(\x. \f x + g x\ powr p \M) = 0") case True then show ?thesis by auto next case False then have **: "(\x. \f x + g x\ powr p \M) powr (1/q) > 0" by auto have "(\x. \f x + g x\ powr p \M) powr (1/q) * (\x. \f x + g x\ powr p \M) powr (1/p) = (\x. \f x + g x\ powr p \M)" by (auto simp add: powr_add[symmetric] add.commute) then have "(\x. \f x + g x\ powr p \M) powr (1/q) * (\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x + g x\ powr p \M) powr (1/q) * ((\x. abs(\f x\) powr p \M) powr (1/p) + (\x. abs(\g x\) powr p \M) powr (1/p))" using * by auto then show ?thesis using ** by auto qed qed qed text \When $p<1$, the function $x \mapsto |x|^p$ is not convex any more. Hence, the $L^p$ ``norm'' is not a norm any more, but a quasinorm. This is proved using a different convexity argument, as follows.\ theorem Minkowski_inequality_le_1: assumes "p > (0::real)" "p \ 1" and [measurable, simp]: "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr p)" shows "integrable M (\x. \f x + g x\ powr p)" "(\x. \f x + g x\ powr p \M) powr (1/p) \ 2 powr (1/p-1) * (\x. \f x\ powr p \M) powr (1/p) + 2 powr (1/p-1) * (\x. \g x\ powr p \M) powr (1/p)" proof - have *: "\a + b\ powr p \ \a\ powr p + \b\ powr p" for a b using x_plus_y_p_le_xp_plus_yp[OF \p > 0\ \p \ 1\, of "\a\" "\b\"] by (auto, meson abs_ge_zero abs_triangle_ineq assms(1) le_less order.trans powr_mono2) show "integrable M (\x. \f x + g x\ powr p)" by (rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p + \g x\ powr p"], auto simp add: *) have "(\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x\ powr p + \g x\ powr p \M) powr (1/p)" by (rule powr_mono2, simp add: \p > 0\ less_imp_le, simp, rule integral_mono', auto simp add: *) also have "... = 2 powr (1/p) * (((\x. \f x\ powr p \M) + (\x. \g x\ powr p \M)) / 2) powr (1/p)" by (auto simp add: powr_mult[symmetric] add_divide_distrib) also have "... \ 2 powr (1/p) * (((\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p)) / 2)" apply (rule mult_mono, simp, rule convex_on_mean_ineq[OF convex_powr[of "1/p"]]) using \p \ 1\ \p > 0\ by auto also have "... = 2 powr (1/p - 1) * ((\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p))" by (simp add: powr_diff) finally show "(\x. \f x + g x\ powr p \M) powr (1/p) \ 2 powr (1/p-1) * (\x. \f x\ powr p \M) powr (1/p) + 2 powr (1/p-1) * (\x. \g x\ powr p \M) powr (1/p)" by (auto simp add: algebra_simps) qed section \$L^p$ spaces\ text \We define $L^p$ spaces by giving their defining quasinorm. It is a norm for $p\in [1, \infty]$, and a quasinorm for $p \in (0,1)$. The construction of a quasinorm from a formula only makes sense if this formula is indeed a quasinorm, i.e., it is homogeneous and satisfies the triangular inequality with the given multiplicative defect. Thus, we have to show that this is indeed the case to be able to use the definition.\ definition Lp_space::"ennreal \ 'a measure \ ('a \ real) quasinorm" where "Lp_space p M = ( if p = 0 then quasinorm_of (1, (\f. if (f \ borel_measurable M) then 0 else \)) else if p < \ then quasinorm_of ( if p < 1 then 2 powr (1/enn2real p - 1) else 1, (\f. if (f \ borel_measurable M \ integrable M (\x. \f x\ powr (enn2real p))) then (\x. \f x\ powr (enn2real p) \M) powr (1/(enn2real p)) else (\::ennreal))) else quasinorm_of (1, (\f. if f \ borel_measurable M then esssup M (\x. ereal \f x\) else (\::ennreal))))" abbreviation "\ == Lp_space" subsection \$L^\infty$\ text \Let us check that, for $L^\infty$, the above definition makes sense.\ lemma L_infinity: "eNorm (\ \ M) f = (if f \ borel_measurable M then esssup M (\x. ereal \f x\) else (\::ennreal))" "defect (\ \ M) = 1" proof - have T: "esssup M (\x. ereal \(f + g) x\) \ e2ennreal (esssup M (\x. ereal \f x\)) + esssup M (\x. ereal \g x\)" if [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" for f g proof (cases "emeasure M (space M) = 0") case True then have "e2ennreal (esssup M (\x. ereal \(f + g) x\)) = 0" using esssup_zero_space[OF True] by (simp add: e2ennreal_neg) then show ?thesis by simp next case False have *: "esssup M (\x. \h x\) \ 0" for h::"'a \ real" proof - have "esssup M (\x. 0) \ esssup M (\x. \h x\)" by (rule esssup_mono, auto) then show ?thesis using esssup_const[OF False, of "0::ereal"] by simp qed have "esssup M (\x. ereal \(f + g) x\) \ esssup M (\x. ereal \f x\ + ereal \g x\)" by (rule esssup_mono, auto simp add: plus_fun_def) also have "... \ esssup M (\x. ereal \f x\) + esssup M (\x. ereal \g x\)" by (rule esssup_add) finally show ?thesis using * by (simp add: e2ennreal_mono eq_onp_def plus_ennreal.abs_eq) qed have H: "esssup M (\x. ereal \(c *\<^sub>R f) x\) \ ennreal \c\ * esssup M (\x. ereal \f x\)" if "c \ 0" for f c proof - have "abs c > 0" "ereal \c\ \ 0" using that by auto have *: "esssup M (\x. abs(c *\<^sub>R f x)) = abs c * esssup M (\x. \f x\)" apply (subst esssup_cmult[OF \abs c > 0\, of M "\x. ereal \f x\", symmetric]) using times_ereal.simps(1) by (auto simp add: abs_mult) show ?thesis unfolding e2ennreal_mult[OF \ereal \c\ \ 0\] * scaleR_fun_def - using ennreal.abs_eq ennreal.rep_eq by auto + by simp qed have "esssup M (\x. ereal 0) \ 0" using esssup_I by auto then have Z: "e2ennreal (esssup M (\x. ereal 0)) = 0" using e2ennreal_neg by auto have *: "quasinorm_on (borel_measurable M) 1 (\(f::'a\real). e2ennreal(esssup M (\x. ereal \f x\)))" apply (rule quasinorm_onI) using T H Z by auto have **: "quasinorm_on UNIV 1 (\(f::'a\real). if f \ borel_measurable M then e2ennreal(esssup M (\x. ereal \f x\)) else \)" by (rule extend_quasinorm[OF *]) show "eNorm (\ \ M) f = (if f \ borel_measurable M then e2ennreal(esssup M (\x. \f x\)) else \)" "defect (\ \ M) = 1" unfolding Lp_space_def using quasinorm_of[OF **] by auto qed lemma L_infinity_space: "space\<^sub>N (\ \ M) = {f \ borel_measurable M. \C. AE x in M. \f x\ \ C}" proof (auto simp del: infinity_ennreal_def) fix f assume H: "f \ space\<^sub>N (\ \ M)" then show "f \ borel_measurable M" unfolding space\<^sub>N_def using L_infinity(1)[of M] top.not_eq_extremum by fastforce then have *: "esssup M (\x. \f x\) < \" using H unfolding space\<^sub>N_def L_infinity(1)[of M] by (auto simp add: e2ennreal_infty) define C where "C = real_of_ereal(esssup M (\x. \f x\))" have "AE x in M. ereal \f x\ \ ereal C" proof (cases "emeasure M (space M) = 0") case True then show ?thesis using emeasure_0_AE by simp next case False then have "esssup M (\x. \f x\) \ 0" using esssup_mono[of "\x. 0" M "(\x. \f x\)"] esssup_const[OF False, of "0::ereal"] by auto then have "esssup M (\x. \f x\) = ereal C" unfolding C_def using * ereal_real by auto then show ?thesis using esssup_AE[of "(\x. ereal \f x\)" M] by simp qed then have "AE x in M. \f x\ \ C" by auto then show "\C. AE x in M. \f x\ \ C" by blast next fix f::"'a \ real" and C::real assume H: "f \ borel_measurable M" "AE x in M. \f x\ \ C" then have "esssup M (\x. \f x\) \ C" using esssup_I by auto then have "eNorm (\ \ M) f \ C" unfolding L_infinity(1) using H(1) - by (auto, metis antisym e2ennreal_mono enn2ereal_ennreal ennreal.abs_eq ennreal.rep_eq ennreal_eq_0_iff ereal_less_eq(4) linear order.trans zero_ereal_def) + by auto (metis e2ennreal_ereal e2ennreal_mono) then show "f \ space\<^sub>N (\ \ M)" using spaceN_iff le_less_trans by fastforce qed lemma L_infinity_zero_space: "zero_space\<^sub>N (\ \ M) = {f \ borel_measurable M. AE x in M. f x = 0}" proof (auto simp del: infinity_ennreal_def) fix f assume H: "f \ zero_space\<^sub>N (\ \ M)" then show "f \ borel_measurable M" unfolding zero_space\<^sub>N_def using L_infinity(1)[of M] top.not_eq_extremum by fastforce then have *: "e2ennreal(esssup M (\x. \f x\)) = 0" using H unfolding zero_space\<^sub>N_def using L_infinity(1)[of M] e2ennreal_infty by auto then have "esssup M (\x. \f x\) \ 0" by (metis e2ennreal_infty e2ennreal_mult ennreal_top_neq_zero ereal_mult_infty leI linear mult_zero_left) then have "f x = 0" if "ereal \f x\ \ esssup M (\x. \f x\)" for x using that order.trans by fastforce then show "AE x in M. f x = 0" using esssup_AE[of "\x. ereal \f x\" M] by auto next fix f::"'a \ real" assume H: "f \ borel_measurable M" "AE x in M. f x = 0" then have "esssup M (\x. \f x\) \ 0" using esssup_I by auto then have "eNorm (\ \ M) f = 0" unfolding L_infinity(1) using H(1) by (simp add: e2ennreal_neg) then show "f \ zero_space\<^sub>N (\ \ M)" using zero_spaceN_iff by auto qed lemma L_infinity_AE_ebound: "AE x in M. ennreal \f x\ \ eNorm (\ \ M) f" proof (cases "f \ borel_measurable M") case False then have "eNorm (\ \ M) f = \" unfolding L_infinity(1) by auto then show ?thesis by simp next case True then have "ennreal \f x\ \ eNorm (\ \ M) f" if "\f x\ \ esssup M (\x. \f x\)" for x - unfolding L_infinity(1) using that e2ennreal_mono ennreal.abs_eq ennreal.rep_eq by fastforce + unfolding L_infinity(1) using that e2ennreal_mono + by fastforce then show ?thesis using esssup_AE[of "\x. ereal \f x\"] by force qed lemma L_infinity_AE_bound: assumes "f \ space\<^sub>N (\ \ M)" shows "AE x in M. \f x\ \ Norm (\ \ M) f" using L_infinity_AE_ebound[of f M] unfolding eNorm_Norm[OF assms] by (simp) text \In the next lemma, the assumption $C \geq 0$ that might seem useless is in fact necessary for the second statement when the space has zero measure. Indeed, any function is then almost surely bounded by any constant!\ lemma L_infinity_I: assumes "f \ borel_measurable M" "AE x in M. \f x\ \ C" "C \ 0" shows "f \ space\<^sub>N (\ \ M)" "Norm (\ \ M) f \ C" proof - show "f \ space\<^sub>N (\ \ M)" using L_infinity_space assms(1) assms(2) by force have "esssup M (\x. \f x\) \ C" using assms(1) assms(2) esssup_I by auto then have "eNorm (\ \ M) f \ ereal C" unfolding L_infinity(1) using assms(1) e2ennreal_mono by force then have "ennreal (Norm (\ \ M) f) \ ennreal C" - using eNorm_Norm[OF \f \ space\<^sub>N (\ \ M)\] assms(3) ennreal.abs_eq ennreal.rep_eq by auto + using eNorm_Norm[OF \f \ space\<^sub>N (\ \ M)\] assms(3) by auto then show "Norm (\ \ M) f \ C" using assms(3) by auto qed lemma L_infinity_I': assumes [measurable]: "f \ borel_measurable M" and "AE x in M. ennreal \f x\ \ C" shows "eNorm (\ \ M) f \ C" proof - have "esssup M (\x. \f x\) \ enn2ereal C" apply (rule esssup_I, auto) using assms(2) less_eq_ennreal.rep_eq by auto then show ?thesis unfolding L_infinity using assms apply auto using e2ennreal_mono by fastforce qed lemma L_infinity_pos_measure: assumes [measurable]: "f \ borel_measurable M" and "eNorm (\ \ M) f > (C::real)" shows "emeasure M {x \ space M. \f x\ > C} > 0" proof - have *: "esssup M (\x. ereal(\f x\)) > ereal C" using \eNorm (\ \ M) f > C\ unfolding L_infinity proof (auto) assume a1: "ennreal C < e2ennreal (esssup M (\x. ereal \f x\))" have "\ e2ennreal (esssup M (\a. ereal \f a\)) \ e2ennreal (ereal C)" if "\ C < 0" using a1 that by (metis (no_types) e2ennreal_enn2ereal enn2ereal_ennreal leD leI) then have "e2ennreal (esssup M (\a. ereal \f a\)) \ e2ennreal (ereal C) \ (\e\esssup M (\a. ereal \f a\). ereal C < e)" using a1 e2ennreal_neg by fastforce then show ?thesis by (meson e2ennreal_mono leI less_le_trans) qed have "emeasure M {x \ space M. ereal(\f x\) > C} > 0" by (rule esssup_pos_measure[OF _ *], auto) then show ?thesis by auto qed lemma L_infinity_tendsto_AE: assumes "tendsto_in\<^sub>N (\ \ M) f g" "\n. f n \ space\<^sub>N (\ \ M)" "g \ space\<^sub>N (\ \ M)" shows "AE x in M. (\n. f n x) \ g x" proof - have *: "AE x in M. \(f n - g) x\ \ Norm (\ \ M) (f n - g)" for n apply (rule L_infinity_AE_bound) using assms spaceN_diff by blast have "AE x in M. \n. \(f n - g) x\ \ Norm (\ \ M) (f n - g)" apply (subst AE_all_countable) using * by auto moreover have "(\n. f n x) \ g x" if "\n. \(f n - g) x\ \ Norm (\ \ M) (f n - g)" for x proof - have "(\n. \(f n - g) x\) \ 0" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. Norm (\ \ M) (f n - g)"]) using that \tendsto_in\<^sub>N (\ \ M) f g\ unfolding tendsto_in\<^sub>N_def by auto then have "(\n. \f n x - g x\) \ 0" by auto then show ?thesis by (simp add: \(\n. \f n x - g x\) \ 0\ LIM_zero_cancel tendsto_rabs_zero_cancel) qed ultimately show ?thesis by auto qed text \As an illustration of the mechanism of spaces inclusion, let us show that bounded continuous functions belong to $L^\infty$.\ lemma bcontfun_subset_L_infinity: assumes "sets M = sets borel" shows "space\<^sub>N bcontfun\<^sub>N \ space\<^sub>N (\ \ M)" "\f. f \ space\<^sub>N bcontfun\<^sub>N \ Norm (\ \ M) f \ Norm bcontfun\<^sub>N f" "\f. eNorm (\ \ M) f \ eNorm bcontfun\<^sub>N f" "bcontfun\<^sub>N \\<^sub>N \ \ M" proof - have *: "f \ space\<^sub>N (\ \ M) \ Norm (\ \ M) f \ Norm bcontfun\<^sub>N f" if "f \ space\<^sub>N bcontfun\<^sub>N" for f proof - have H: "continuous_on UNIV f" "\x. abs(f x) \ Norm bcontfun\<^sub>N f" using bcontfun\<^sub>ND[OF \f \ space\<^sub>N bcontfun\<^sub>N\] by auto then have "f \ borel_measurable borel" using borel_measurable_continuous_onI by simp then have "f \ borel_measurable M" using assms by auto have *: "AE x in M. \f x\ \ Norm bcontfun\<^sub>N f" using H(2) by auto show ?thesis using L_infinity_I[OF \f \ borel_measurable M\ * Norm_nonneg] by auto qed show "space\<^sub>N bcontfun\<^sub>N \ space\<^sub>N (\ \ M)" "\f. f \ space\<^sub>N bcontfun\<^sub>N \ Norm (\ \ M) f \ Norm bcontfun\<^sub>N f" using * by auto show **: "bcontfun\<^sub>N \\<^sub>N \ \ M" apply (rule quasinorm_subsetI'[of _ _ 1]) using * by auto have "eNorm (\ \ M) f \ ennreal 1 * eNorm bcontfun\<^sub>N f" for f apply (rule quasinorm_subset_Norm_eNorm) using * ** by auto then show "eNorm (\ \ M) f \ eNorm bcontfun\<^sub>N f" for f by simp qed subsection \$L^p$ for $0 < p < \infty$\ lemma Lp: assumes "p \ (1::real)" shows "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 1" proof - define F where "F = {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" have *: "quasinorm_on F 1 (\(f::'a\real). (\x. \f x\ powr p \M) powr (1/p))" proof (rule quasinorm_onI) fix f g assume "f \ F" "g \ F" then show "f + g \ F" unfolding F_def plus_fun_def apply (auto) by (rule Minkowski_inequality(1), auto simp add: \p \ 1\) show "ennreal ((\x. \(f + g) x\ powr p \M) powr (1/p)) \ ennreal 1 * (\x. \f x\ powr p \M) powr (1/p) + ennreal 1 * (\x. \g x\ powr p \M) powr (1/p)" apply (auto, subst ennreal_plus[symmetric], simp, simp, rule ennreal_leI) unfolding plus_fun_def apply (rule Minkowski_inequality(2)[of p f M g], auto simp add: \p \ 1\) using \f \ F\ \g \ F\ unfolding F_def by auto next fix f and c::real assume "f \ F" show "c *\<^sub>R f \ F" using \f \ F\ unfolding scaleR_fun_def F_def by (auto simp add: abs_mult powr_mult) show "(\x. \(c *\<^sub>R f) x\ powr p \M) powr (1/p) \ ennreal(abs(c)) * (\x. \f x\ powr p \M) powr (1/p)" apply (rule eq_refl, subst ennreal_mult[symmetric], simp, simp, rule ennreal_cong) apply (unfold scaleR_fun_def, simp add: abs_mult powr_mult powr_powr) using \p \ 1\ by auto next show "0 \ F" unfolding zero_fun_def F_def by auto qed (auto) have "p \ 0" using \p \ 1\ by auto have **: "\ p M = quasinorm_of (1, (\f. if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal)))" unfolding Lp_space_def using enn2real_ennreal[OF \p \ 0\] \p \ 1\ apply auto using enn2real_ennreal[OF \p \ 0\] by presburger show "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 1" unfolding ** using quasinorm_of[OF extend_quasinorm[OF *]] unfolding F_def by auto qed lemma Lp_le_1: assumes "p > 0" "p \ (1::real)" shows "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 2 powr (1/p - 1)" proof - define F where "F = {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" have *: "quasinorm_on F (2 powr (1/p-1)) (\(f::'a\real). (\x. \f x\ powr p \M) powr (1/p))" proof (rule quasinorm_onI) fix f g assume "f \ F" "g \ F" then show "f + g \ F" unfolding F_def plus_fun_def apply (auto) by (rule Minkowski_inequality_le_1(1), auto simp add: \p > 0\ \p \ 1\) show "ennreal ((\x. \(f + g) x\ powr p \M) powr (1/p)) \ ennreal (2 powr (1/p-1)) * (\x. \f x\ powr p \M) powr (1/p) + ennreal (2 powr (1/p-1)) * (\x. \g x\ powr p \M) powr (1/p)" apply (subst ennreal_mult[symmetric], auto)+ apply (subst ennreal_plus[symmetric], simp, simp) apply (rule ennreal_leI) unfolding plus_fun_def apply (rule Minkowski_inequality_le_1(2)[of p f M g], auto simp add: \p > 0\ \p \ 1\) using \f \ F\ \g \ F\ unfolding F_def by auto next fix f and c::real assume "f \ F" show "c *\<^sub>R f \ F" using \f \ F\ unfolding scaleR_fun_def F_def by (auto simp add: abs_mult powr_mult) show "(\x. \(c *\<^sub>R f) x\ powr p \M) powr (1/p) \ ennreal(abs(c)) * (\x. \f x\ powr p \M) powr (1/p)" apply (rule eq_refl, subst ennreal_mult[symmetric], simp, simp, rule ennreal_cong) apply (unfold scaleR_fun_def, simp add: abs_mult powr_mult powr_powr) using \p > 0\ by auto next show "0 \ F" unfolding zero_fun_def F_def by auto show "1 \ 2 powr (1 / p - 1)" using \p > 0\ \p \ 1\ by (auto simp add: ge_one_powr_ge_zero) qed (auto) have "p \ 0" using \p > 0\ by auto have **: "\ p M = quasinorm_of (2 powr (1/p-1), (\f. if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal)))" unfolding Lp_space_def using \p > 0\ \p \ 1\ using enn2real_ennreal[OF \p \ 0\] apply auto by (insert enn2real_ennreal[OF \p \ 0\], presburger)+ show "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 2 powr (1/p-1)" unfolding ** using quasinorm_of[OF extend_quasinorm[OF *]] unfolding F_def by auto qed lemma Lp_space: assumes "p > (0::real)" shows "space\<^sub>N (\ p M) = {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" apply (auto simp add: spaceN_iff) using Lp(1) Lp_le_1(1) \p > 0\ apply (metis infinity_ennreal_def less_le not_less) using Lp(1) Lp_le_1(1) \p > 0\ apply (metis infinity_ennreal_def less_le not_less) using Lp(1) Lp_le_1(1) \p > 0\ by (metis ennreal_neq_top linear top.not_eq_extremum) lemma Lp_I: assumes "p > (0::real)" "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" shows "f \ space\<^sub>N (\ p M)" "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" proof - have *: "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" by (cases "p \ 1", insert assms, auto simp add: Lp_le_1(1) Lp(1)) then show **: "f \ space\<^sub>N (\ p M)" unfolding space\<^sub>N_def by auto show "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" using * unfolding Norm_def by auto then show "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" using eNorm_Norm[OF **] by auto qed lemma Lp_D: assumes "p>0" "f \ space\<^sub>N (\ p M)" shows "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" proof - show *: "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" using Lp_space[OF \p > 0\] assms(2) by auto then show "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" using Lp_I[OF \p > 0\] by auto qed lemma Lp_Norm: assumes "p > (0::real)" "f \ borel_measurable M" shows "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "(Norm (\ p M) f) powr p = (\x. \f x\ powr p \M)" proof - show *: "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" proof (cases "integrable M (\x. \f x\ powr p)") case True then show ?thesis using Lp_I[OF assms True] by auto next case False then have "f \ space\<^sub>N (\ p M)" using Lp_space[OF \p > 0\, of M] by auto then have *: "Norm (\ p M) f = 0" using eNorm_Norm' by auto have "(\x. \f x\ powr p \M) = 0" using False by (simp add: not_integrable_integral_eq) then have "(\x. \f x\ powr p \M) powr (1/p) = 0" by auto then show ?thesis using * by auto qed then show "(Norm (\ p M) f) powr p = (\x. \f x\ powr p \M)" unfolding * using powr_powr \p > 0\ by auto qed lemma Lp_zero_space: assumes "p > (0::real)" shows "zero_space\<^sub>N (\ p M) = {f \ borel_measurable M. AE x in M. f x = 0}" proof (auto) fix f assume H: "f \ zero_space\<^sub>N (\ p M)" then have *: "f \ {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" using Lp_space[OF assms] zero_spaceN_subset_spaceN by auto then show "f \ borel_measurable M" by auto have "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" by (cases "p \ 1", insert * \p > 0\, auto simp add: Lp_le_1(1) Lp(1)) then have "(\x. \f x\ powr p \M) = 0" using H unfolding zero_space\<^sub>N_def by auto then have "AE x in M. \f x\ powr p = 0" by (subst integral_nonneg_eq_0_iff_AE[symmetric], insert *, auto) then show "AE x in M. f x = 0" by auto next fix f::"'a \ real" assume H [measurable]: "f \ borel_measurable M" "AE x in M. f x = 0" then have *: "AE x in M. \f x\ powr p = 0" by auto have "integrable M (\x. \f x\ powr p)" using integrable_cong_AE[OF _ _ *] by auto have **: "(\x. \f x\ powr p \M) = 0" using integral_cong_AE[OF _ _ *] by auto have "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" by (cases "p \ 1", insert H(1) \integrable M (\x. \f x\ powr p)\ \p > 0\, auto simp add: Lp_le_1(1) Lp(1)) then have "eNorm (\ p M) f = 0" using ** by simp then show "f \ zero_space\<^sub>N (\ p M)" using zero_spaceN_iff by auto qed lemma Lp_tendsto_AE_subseq: assumes "p>(0::real)" "tendsto_in\<^sub>N (\ p M) f g" "\n. f n \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ p M)" shows "\r. strict_mono r \ (AE x in M. (\n. f (r n) x) \ g x)" proof - have "f n - g \ space\<^sub>N (\ p M)" for n using spaceN_diff[OF \\n. f n \ space\<^sub>N (\ p M)\ \g \ space\<^sub>N (\ p M)\] by simp have int: "integrable M (\x. \f n x - g x\ powr p)" for n using Lp_D(2)[OF \p > 0\ \f n - g \ space\<^sub>N (\ p M)\] by auto have "(\n. Norm (\ p M) (f n - g)) \ 0" using \tendsto_in\<^sub>N (\ p M) f g\ unfolding tendsto_in\<^sub>N_def by auto then have *: "(\n. (\x. \f n x - g x\ powr p \M) powr (1/p)) \ 0" using Lp_D(3)[OF \p > 0\ \\n. f n - g \ space\<^sub>N (\ p M)\] by auto have "(\n. ((\x. \f n x - g x\ powr p \M) powr (1/p)) powr p) \ 0" apply (rule tendsto_zero_powrI[of _ _ _ p]) using \p > 0\ * by auto then have **: "(\n. (\x. \f n x - g x\ powr p \M)) \ 0" using powr_powr \p > 0\ by auto have "\r. strict_mono r \ (AE x in M. (\n. \f (r n) x - g x\ powr p) \ 0)" apply (rule tendsto_L1_AE_subseq) using int ** by auto then obtain r where "strict_mono r" "AE x in M. (\n. \f (r n) x - g x\ powr p) \ 0" by blast moreover have "(\n. f (r n) x) \ g x" if "(\n. \f (r n) x - g x\ powr p) \ 0" for x proof - have "(\n. (\f (r n) x - g x\ powr p) powr (1/p)) \ 0" apply (rule tendsto_zero_powrI[of _ _ _ "1/p"]) using \p > 0\ that by auto then have "(\n. \f (r n) x - g x\) \ 0" using powr_powr \p > 0\ by auto show ?thesis by (simp add: \(\n. \f (r n) x - g x\) \ 0\ Limits.LIM_zero_cancel tendsto_rabs_zero_cancel) qed ultimately have "AE x in M. (\n. f (r n) x) \ g x" by auto then show ?thesis using \strict_mono r\ by auto qed subsection \Specialization to $L^1$\ lemma L1_space: "space\<^sub>N (\ 1 M) = {f. integrable M f}" unfolding one_ereal_def using Lp_space[of 1 M] integrable_abs_iff by auto lemma L1_I: assumes "integrable M f" shows "f \ space\<^sub>N (\ 1 M)" "Norm (\ 1 M) f = (\x. \f x\ \M)" "eNorm (\ 1 M) f = (\x. \f x\ \M)" unfolding one_ereal_def using Lp_I[of 1, OF _ borel_measurable_integrable[OF assms]] assms powr_to_1 by auto lemma L1_D: assumes "f \ space\<^sub>N (\ 1 M)" shows "f \ borel_measurable M" "integrable M f" "Norm (\ 1 M) f = (\x. \f x\ \M)" "eNorm (\ 1 M) f = (\x. \f x\ \M)" using assms by (auto simp add: L1_space L1_I) lemma L1_int_ineq: "abs(\x. f x \M) \ Norm (\ 1 M) f" proof (cases "integrable M f") case True then show ?thesis using L1_I(2)[OF True] by auto next case False then have "(\x. f x \M) = 0" by (simp add: not_integrable_integral_eq) then show ?thesis using Norm_nonneg by auto qed text \In $L^1$, one can give a direct formula for the eNorm of a measurable function, using a nonnegative integral. The same formula holds in $L^p$ for $p > 0$, with additional powers $p$ and $1/p$, but one can not write it down since \verb+powr+ is not defined on \verb+ennreal+.\ lemma L1_Norm: assumes [measurable]: "f \ borel_measurable M" shows "Norm (\ 1 M) f = (\x. \f x\ \M)" "eNorm (\ 1 M) f = (\\<^sup>+x. \f x\ \M)" proof - show *: "Norm (\ 1 M) f = (\x. \f x\ \M)" using Lp_Norm[of 1, OF _ assms] unfolding one_ereal_def by auto show "eNorm (\ 1 M) f = (\\<^sup>+x. \f x\ \M)" proof (cases "integrable M f") case True then have "f \ space\<^sub>N (\ 1 M)" using L1_space by auto then have "eNorm (\ 1 M) f = ennreal (Norm (\ 1 M) f)" using eNorm_Norm by auto then show ?thesis by (metis (mono_tags) * AE_I2 True abs_ge_zero integrable_abs nn_integral_eq_integral) next case False then have "eNorm (\ 1 M) f = \" using L1_space space\<^sub>N_def by (metis ennreal_add_eq_top infinity_ennreal_def le_iff_add le_less_linear mem_Collect_eq) moreover have "(\\<^sup>+x. \f x\ \M) = \" apply (rule nn_integral_nonneg_infinite) using False by (auto simp add: integrable_abs_iff) ultimately show ?thesis by simp qed qed lemma L1_indicator: assumes [measurable]: "A \ sets M" shows "eNorm (\ 1 M) (indicator A) = emeasure M A" by (subst L1_Norm, auto, metis assms ennreal_indicator nn_integral_cong nn_integral_indicator) lemma L1_indicator': assumes [measurable]: "A \ sets M" and "emeasure M A \ \" shows "indicator A \ space\<^sub>N (\ 1 M)" "Norm (\ 1 M) (indicator A) = measure M A" unfolding space\<^sub>N_def Norm_def using L1_indicator[OF \A \ sets M\] \emeasure M A \ \\ by (auto simp add: top.not_eq_extremum Sigma_Algebra.measure_def) subsection \$L^0$\ text \We have defined $L^p$ for all exponents $p$, although it does not really make sense for $p = 0$. We have chosen a definition in this case (the space of all measurable functions) so that many statements are true for all exponents. In this paragraph, we show the consistency of this definition.\ lemma L_zero: "eNorm (\ 0 M) f = (if f \ borel_measurable M then 0 else \)" "defect (\ 0 M) = 1" proof - have *: "quasinorm_on UNIV 1 (\(f::'a\real). (if f \ borel_measurable M then 0 else \))" by (rule extend_quasinorm, rule quasinorm_onI, auto) show "eNorm (\ 0 M) f = (if f \ borel_measurable M then 0 else \)" "defect (\ 0 M) = 1" using quasinorm_of[OF *] unfolding Lp_space_def by auto qed lemma L_zero_space [simp]: "space\<^sub>N (\ 0 M) = borel_measurable M" "zero_space\<^sub>N (\ 0 M) = borel_measurable M" apply (auto simp add: spaceN_iff zero_spaceN_iff L_zero(1)) using top.not_eq_extremum by force+ subsection \Basic results on $L^p$ for general $p$\ lemma Lp_measurable_subset: "space\<^sub>N (\ p M) \ borel_measurable M" proof (cases rule: Lp_cases[of p]) case zero then show ?thesis using L_zero_space by auto next case (real_pos p2) then show ?thesis using Lp_space[OF \p2 > 0\] by auto next case PInf then show ?thesis using L_infinity_space by auto qed lemma Lp_measurable: assumes "f \ space\<^sub>N (\ p M)" shows "f \ borel_measurable M" using assms Lp_measurable_subset by auto lemma Lp_infinity_zero_space: assumes "p > (0::ennreal)" shows "zero_space\<^sub>N (\ p M) = {f \ borel_measurable M. AE x in M. f x = 0}" proof (cases rule: Lp_cases[of p]) case PInf then show ?thesis using L_infinity_zero_space by auto next case (real_pos p2) then show ?thesis using Lp_zero_space[OF \p2 > 0\] unfolding \p = ennreal p2\ by auto next case zero then have False using assms by auto then show ?thesis by simp qed lemma (in prob_space) Lp_subset_Lq: assumes "p \ q" shows "\f. eNorm (\ p M) f \ eNorm (\ q M) f" "\ q M \\<^sub>N \ p M" "space\<^sub>N (\ q M) \ space\<^sub>N (\ p M)" "\f. f \ space\<^sub>N (\ q M) \ Norm (\ p M) f \ Norm (\ q M) f" proof - show "eNorm (\ p M) f \ eNorm (\ q M) f" for f proof (cases "eNorm (\ q M) f < \") case True then have "f \ space\<^sub>N (\ q M)" using spaceN_iff by auto then have f_meas [measurable]: "f \ borel_measurable M" using Lp_measurable by auto consider "p = 0" | "p = q" | "p > 0 \ p < \ \ q = \" | "p > 0 \ p < q \ q < \" using \p \ q\ apply (simp add: top.not_eq_extremum) using not_less_iff_gr_or_eq order.order_iff_strict by fastforce then show ?thesis proof (cases) case 1 then show ?thesis by (simp add: L_zero(1)) next case 2 then show ?thesis by auto next case 3 then have "q = \" by simp obtain p2 where "p = ennreal p2" "p2 > 0" using 3 enn2real_positive_iff[of p] by (cases p) auto have *: "AE x in M. \f x\ \ Norm (\ \ M) f" using L_infinity_AE_bound \f \ space\<^sub>N (\ q M)\ \q = \\ by auto have **: "integrable M (\x. \f x\ powr p2)" apply (rule Bochner_Integration.integrable_bound[of _ "\x. (Norm (\ \ M) f) powr p2"], auto) using * powr_mono2 \p2 > 0\ by force then have "eNorm (\ p2 M) f = (\x. \f x\ powr p2 \M) powr (1/p2)" using Lp_I(3)[OF \p2 > 0\ f_meas] by simp also have "... \ (\x. (Norm (\ \ M) f) powr p2 \M) powr (1/p2)" apply (rule ennreal_leI, rule powr_mono2, simp add: \p2 > 0\ less_imp_le, simp) apply (rule integral_mono_AE, auto simp add: **) using * powr_mono2 \p2 > 0\ by force also have "... = Norm (\ \ M) f" using \p2 > 0\ by (auto simp add: prob_space powr_powr) finally show ?thesis using \p = ennreal p2\ \q = \\ eNorm_Norm[OF \f \ space\<^sub>N (\ q M)\] by auto next case 4 then have "0 < p" "p < \" by auto then obtain p2 where "p = ennreal p2" "p2 > 0" using enn2real_positive_iff[of p] by (cases p) auto have "0 < q" "q < \" using 4 by auto then obtain q2 where "q = ennreal q2" "q2 > 0" using enn2real_positive_iff[of q] by (cases q) auto have "p2 < q2" using 4 \p = ennreal p2\ \q = ennreal q2\ using ennreal_less_iff by auto define r2 where "r2 = q2 / p2" have "r2 \ 1" unfolding r2_def using \p2 < q2\ \p2 > 0\ by auto have *: "abs (\z\ powr p2) powr r2 = \z\ powr q2" for z::real unfolding r2_def using \p2 > 0\ by (simp add: powr_powr) have I: "integrable M (\x. abs(\f x\ powr p2) powr r2)" unfolding * using \f \ space\<^sub>N (\ q M)\ \q = ennreal q2\ Lp_D(2)[OF \q2 > 0\] by auto have J: "integrable M (\x. \f x\ powr p2)" by (rule bound_L1_Lp(1)[OF \r2 \ 1\ _ I], auto) have "f \ space\<^sub>N (\ p2 M)" by (rule Lp_I(1)[OF \p2 > 0\ _ J], simp) have "(\x. \f x\ powr p2 \M) powr (1/p2) = abs(\x. \f x\ powr p2 \M) powr (1/p2)" by auto also have "... \ ((\x. abs (\f x\ powr p2) powr r2 \M) powr (1/r2)) powr (1/p2)" apply (subst powr_mono2, simp add: \p2 > 0\ less_imp_le, simp) apply (rule bound_L1_Lp, simp add: \r2 \ 1\, simp) unfolding * using \f \ space\<^sub>N (\ q M)\ \q = ennreal q2\ Lp_D(2)[OF \q2 > 0\] by auto also have "... = (\x. \f x\ powr q2 \M) powr (1/q2)" unfolding * using \p2 > 0\ by (simp add: powr_powr r2_def) finally show ?thesis using \f \ space\<^sub>N (\ q M)\ Lp_D(4)[OF \q2 > 0\] ennreal_leI unfolding \p = ennreal p2\ \q = ennreal q2\ Lp_D(4)[OF \p2 > 0\ \f \ space\<^sub>N (\ p2 M)\] by force qed next case False then have "eNorm (\ q M) f = \" using top.not_eq_extremum by fastforce then show ?thesis by auto qed then show "\ q M \\<^sub>N \ p M" using quasinorm_subsetI[of _ _ 1] by auto then show "space\<^sub>N (\ q M) \ space\<^sub>N (\ p M)" using quasinorm_subset_space by auto then show "Norm (\ p M) f \ Norm (\ q M) f" if "f \ space\<^sub>N (\ q M)" for f using eNorm_Norm that \eNorm (\ p M) f \ eNorm (\ q M) f\ ennreal_le_iff Norm_nonneg by (metis rev_subsetD) qed proposition Lp_domination: assumes [measurable]: "g \ borel_measurable M" and "f \ space\<^sub>N (\ p M)" "AE x in M. \g x\ \ \f x\" shows "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ Norm (\ p M) f" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable[OF assms(2)] by simp have "g \ space\<^sub>N (\ p M) \ Norm (\ p M) g \ Norm (\ p M) f" proof (cases rule: Lp_cases[of p]) case zero then have "Norm (\ p M) g = 0" unfolding Norm_def using L_zero(1)[of M] by auto then have "Norm (\ p M) g \ Norm (\ p M) f" using Norm_nonneg by auto then show ?thesis unfolding \p = 0\ L_zero_space by auto next case (real_pos p2) have *: "integrable M (\x. \f x\ powr p2)" using \f \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ using Lp_D(2) \p2 > 0\ by auto have **: "integrable M (\x. \g x\ powr p2)" apply (rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p2"]) using * apply auto using assms(3) powr_mono2 \p2 > 0\ by (auto simp add: less_imp_le) then have "g \ space\<^sub>N (\ p M)" unfolding \p = ennreal p2\ using Lp_space[OF \p2 > 0\, of M] by auto have "Norm (\ p M) g = (\x. \g x\ powr p2 \M) powr (1/p2)" unfolding \p = ennreal p2\ by (rule Lp_I(2)[OF \p2 > 0\ _ **], simp) also have "... \ (\x. \f x\ powr p2 \M) powr (1/p2)" apply (rule powr_mono2, simp add: \p2 > 0\ less_imp_le, simp) apply (rule integral_mono_AE, auto simp add: * **) using \p2 > 0\ less_imp_le powr_mono2 assms(3) by auto also have "... = Norm (\ p M) f" unfolding \p = ennreal p2\ by (rule Lp_I(2)[OF \p2 > 0\ _ *, symmetric], simp) finally show ?thesis using \g \ space\<^sub>N (\ p M)\ by auto next case PInf have "AE x in M. \f x\ \ Norm (\ p M) f" using \f \ space\<^sub>N (\ p M)\ L_infinity_AE_bound unfolding \p = \\ by auto then have *: "AE x in M. \g x\ \ Norm (\ p M) f" using assms(3) by auto show ?thesis using L_infinity_I[OF assms(1) *] Norm_nonneg \p = \\ by auto qed then show "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ Norm (\ p M) f" by auto qed lemma Lp_Banach_lattice: assumes "f \ space\<^sub>N (\ p M)" shows "(\x. \f x\) \ space\<^sub>N (\ p M)" "Norm (\ p M) (\x. \f x\) = Norm (\ p M) f" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable[OF assms] by simp show "(\x. \f x\) \ space\<^sub>N (\ p M)" by (rule Lp_domination(1)[OF _ assms], auto) have "Norm (\ p M) (\x. \f x\) \ Norm (\ p M) f" by (rule Lp_domination[OF _ assms], auto) moreover have "Norm (\ p M) f \ Norm (\ p M) (\x. \f x\)" by (rule Lp_domination[OF _ \(\x. \f x\) \ space\<^sub>N (\ p M)\], auto) finally show "Norm (\ p M) (\x. \f x\) = Norm (\ p M) f" by auto qed lemma Lp_bounded_bounded_support: assumes [measurable]: "f \ borel_measurable M" and "AE x in M. \f x\ \ C" "emeasure M {x \ space M. f x \ 0} \ \" shows "f \ space\<^sub>N(\ p M)" proof (cases rule: Lp_cases[of p]) case zero then show ?thesis using L_zero_space assms by blast next case PInf then show ?thesis using L_infinity_space assms by blast next case (real_pos p2) have *: "integrable M (\x. \f x\ powr p2)" apply (rule integrableI_bounded_set[of "{x \ space M. f x \ 0}" _ _ "C powr p2"]) using assms powr_mono2[OF less_imp_le[OF \p2 > 0\]] by (auto simp add: top.not_eq_extremum) show ?thesis unfolding \p = ennreal p2\ apply (rule Lp_I[OF \p2 > 0\]) using * by auto qed subsection \$L^p$ versions of the main theorems in integration theory\ text \The space $L^p$ is stable under almost sure convergence, for sequence with bounded norm. This is a version of Fatou's lemma (and it indeed follows from this lemma in the only nontrivial situation where $p \in (0, +\infty)$.\ proposition Lp_AE_limit: assumes [measurable]: "g \ borel_measurable M" and "AE x in M. (\n. f n x) \ g x" shows "eNorm (\ p M) g \ liminf (\n. eNorm (\ p M) (f n))" proof (cases "liminf (\n. eNorm (\ p M) (f n)) = \") case True then show ?thesis by auto next case False define le where "le = liminf (\n. eNorm (\ p M) (f n))" then have "le < \" using False by (simp add: top.not_eq_extremum) obtain r0 where r0: "strict_mono r0" "(\n. eNorm (\ p M) (f (r0 n))) \ le" using liminf_subseq_lim unfolding comp_def le_def by force then have "eventually (\n. eNorm (\ p M) (f (r0 n)) < \) sequentially" using False unfolding order_tendsto_iff le_def by (simp add: top.not_eq_extremum) then obtain N where N: "\n. n \ N \ eNorm (\ p M) (f (r0 n)) < \" unfolding eventually_sequentially by blast define r where "r = (\n. r0 (n + N))" have "strict_mono r" unfolding r_def using \strict_mono r0\ by (simp add: strict_mono_Suc_iff) have *: "(\n. eNorm (\ p M) (f (r n))) \ le" unfolding r_def using LIMSEQ_ignore_initial_segment[OF r0(2), of N]. have "f (r n) \ space\<^sub>N (\ p M)" for n using N spaceN_iff unfolding r_def by force then have [measurable]: "f (r n) \ borel_measurable M" for n using Lp_measurable by auto define l where "l = enn2real le" have "l \ 0" unfolding l_def by auto have "le = ennreal l" using \le < \\ unfolding l_def by auto have [tendsto_intros]: "(\n. Norm (\ p M) (f (r n))) \ l" apply (rule tendsto_ennrealD) using * \le < \\ unfolding eNorm_Norm[OF \\n. f (r n) \ space\<^sub>N (\ p M)\] l_def by auto show ?thesis proof (cases rule: Lp_cases[of p]) case zero then have "eNorm (\ p M) g = 0" using assms(1) by (simp add: L_zero(1)) then show ?thesis by auto next case (real_pos p2) then have "f (r n) \ space\<^sub>N (\ p2 M)" for n using \\n. f (r n) \ space\<^sub>N(\ p M)\ by auto have "liminf (\n. ennreal(\f (r n) x\ powr p2)) = \g x\ powr p2" if "(\n. f n x) \ g x" for x apply (rule lim_imp_Liminf, auto intro!: tendsto_intros simp add: \p2 > 0\) using LIMSEQ_subseq_LIMSEQ[OF that \strict_mono r\] unfolding comp_def by auto then have *: "AE x in M. liminf (\n. ennreal(\f (r n) x\ powr p2)) = \g x\ powr p2" using \AE x in M. (\n. f n x) \ g x\ by auto have "(\\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M) = ennreal((Norm (\ p M) (f (r n))) powr p2)" for n proof - have "(\\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M) = ennreal (\x. \f (r n) x\ powr p2 \M)" by (rule nn_integral_eq_integral, auto simp add: Lp_D(2)[OF \p2 > 0\ \f (r n) \ space\<^sub>N (\ p2 M)\]) also have "... = ennreal((Norm (\ p2 M) (f (r n))) powr p2)" unfolding Lp_D(3)[OF \p2 > 0\ \f (r n) \ space\<^sub>N (\ p2 M)\] using powr_powr \p2 > 0\ by auto finally show ?thesis using \p = ennreal p2\ by simp qed moreover have "(\n. ennreal((Norm (\ p M) (f (r n))) powr p2)) \ ennreal(l powr p2)" by (auto intro!:tendsto_intros simp add: \p2 > 0\) ultimately have **: "liminf (\n. (\\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M)) = ennreal(l powr p2)" using lim_imp_Liminf by force have "(\\<^sup>+x. \g x\ powr p2 \M) = (\\<^sup>+x. liminf (\n. ennreal(\f (r n) x\ powr p2)) \M)" apply (rule nn_integral_cong_AE) using * by auto also have "... \ liminf (\n. \\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M)" by (rule nn_integral_liminf, auto) finally have "(\\<^sup>+x. \g x\ powr p2 \M) \ ennreal(l powr p2)" using ** by auto then have "(\\<^sup>+x. \g x\ powr p2 \M) < \" using le_less_trans by fastforce then have intg: "integrable M (\x. \g x\ powr p2)" apply (intro integrableI_nonneg) by auto then have "g \ space\<^sub>N (\ p2 M)" using Lp_I(1)[OF \p2 > 0\, of _ M] by fastforce have "ennreal((Norm (\ p2 M) g) powr p2) = ennreal(\x. \g x\ powr p2 \M)" unfolding Lp_D(3)[OF \p2 > 0\ \g \ space\<^sub>N (\ p2 M)\] using powr_powr \p2 > 0\ by auto also have "... = (\\<^sup>+x. \g x\ powr p2 \M)" by (rule nn_integral_eq_integral[symmetric], auto simp add: intg) finally have "ennreal((Norm (\ p2 M) g) powr p2) \ ennreal(l powr p2)" using \(\\<^sup>+x. \g x\ powr p2 \M) \ ennreal(l powr p2)\ by auto then have "((Norm (\ p2 M) g) powr p2) powr (1/p2) \ (l powr p2) powr (1/p2)" using ennreal_le_iff \l \ 0\ \p2 > 0\ powr_mono2 by auto then have "Norm (\ p2 M) g \ l" using \p2 > 0\ \l \ 0\ by (auto simp add: powr_powr) then have "eNorm (\ p2 M) g \ le" unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ p2 M)\] \le = ennreal l\ using ennreal_leI by auto then show ?thesis unfolding le_def \p = ennreal p2\ by simp next case PInf then have "AE x in M. \n. \f (r n) x\ \ Norm (\ \ M) (f (r n))" apply (subst AE_all_countable) using L_infinity_AE_bound \\n. f (r n) \ space\<^sub>N (\ p M)\ by blast moreover have "\g x\ \ l" if "\n. \f (r n) x\ \ Norm (\ \ M) (f (r n))" "(\n. f n x) \ g x" for x proof - have "(\n. f (r n) x) \ g x" using that LIMSEQ_subseq_LIMSEQ[OF _ \strict_mono r\] unfolding comp_def by auto then have *: "(\n. \f (r n) x\) \ \g x\" by (auto intro!:tendsto_intros) show ?thesis apply (rule LIMSEQ_le[OF *]) using that(1) \(\n. Norm (\ p M) (f (r n))) \ l\ unfolding PInf by auto qed ultimately have "AE x in M. \g x\ \ l" using \AE x in M. (\n. f n x) \ g x\ by auto then have "g \ space\<^sub>N (\ \ M)" "Norm (\ \ M) g \ l" using L_infinity_I[OF \g \ borel_measurable M\ _ \l \ 0\] by auto then have "eNorm (\ \ M) g \ le" unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ \ M)\] \le = ennreal l\ using ennreal_leI by auto then show ?thesis unfolding le_def \p = \\ by simp qed qed lemma Lp_AE_limit': assumes "g \ borel_measurable M" "\n. f n \ space\<^sub>N (\ p M)" "AE x in M. (\n. f n x) \ g x" "(\n. Norm (\ p M) (f n)) \ l" shows "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ l" proof - have "l \ 0" by (rule LIMSEQ_le_const[OF \(\n. Norm (\ p M) (f n)) \ l\], auto) have "(\n. eNorm (\ p M) (f n)) \ ennreal l" unfolding eNorm_Norm[OF \\n. f n \ space\<^sub>N (\ p M)\] using \(\n. Norm (\ p M) (f n)) \ l\ by auto then have *: "ennreal l = liminf (\n. eNorm (\ p M) (f n))" using lim_imp_Liminf[symmetric] trivial_limit_sequentially by blast have "eNorm (\ p M) g \ ennreal l" unfolding * apply (rule Lp_AE_limit) using assms by auto then have "eNorm (\ p M) g < \" using le_less_trans by fastforce then show "g \ space\<^sub>N (\ p M)" using spaceN_iff by auto show "Norm (\ p M) g \ l" using \eNorm (\ p M) g \ ennreal l\ ennreal_le_iff[OF \l \ 0\] unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ p M)\] by auto qed lemma Lp_AE_limit'': assumes "g \ borel_measurable M" "\n. f n \ space\<^sub>N (\ p M)" "AE x in M. (\n. f n x) \ g x" "\n. Norm (\ p M) (f n) \ C" shows "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ C" proof - have "C \ 0" by (rule order_trans[OF Norm_nonneg[of "\ p M" "f 0"] \Norm (\ p M) (f 0) \ C\]) have *: "liminf (\n. ennreal C) = ennreal C" using Liminf_const trivial_limit_at_top_linorder by blast have "eNorm (\ p M) (f n) \ ennreal C" for n unfolding eNorm_Norm[OF \f n \ space\<^sub>N (\ p M)\] using \Norm (\ p M) (f n) \ C\ by (auto simp add: ennreal_leI) then have "liminf (\n. eNorm (\ p M) (f n)) \ ennreal C" using Liminf_mono[of "(\n. eNorm (\ p M) (f n))" "\_. C" sequentially] * by auto then have "eNorm (\ p M) g \ ennreal C" using Lp_AE_limit[OF \g \ borel_measurable M\ \AE x in M. (\n. f n x) \ g x\, of p] by auto then have "eNorm (\ p M) g < \" using le_less_trans by fastforce then show "g \ space\<^sub>N (\ p M)" using spaceN_iff by auto show "Norm (\ p M) g \ C" using \eNorm (\ p M) g \ ennreal C\ ennreal_le_iff[OF \C \ 0\] unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ p M)\] by auto qed text \We give the version of Lebesgue dominated convergence theorem in the setting of $L^p$ spaces.\ proposition Lp_domination_limit: fixes p::real assumes [measurable]: "g \ borel_measurable M" "\n. f n \ borel_measurable M" and "m \ space\<^sub>N (\ p M)" "AE x in M. (\n. f n x) \ g x" "\n. AE x in M. \f n x\ \ m x" shows "g \ space\<^sub>N (\ p M)" "tendsto_in\<^sub>N (\ p M) f g" proof - have [measurable]: "m \ borel_measurable M" using Lp_measurable[OF \m \ space\<^sub>N (\ p M)\] by auto have "f n \ space\<^sub>N(\ p M)" for n apply (rule Lp_domination[OF _ \m \ space\<^sub>N (\ p M)\]) using \AE x in M. \f n x\ \ m x\ by auto have "AE x in M. \n. \f n x\ \ m x" apply (subst AE_all_countable) using \\n. AE x in M. \f n x\ \ m x\ by auto moreover have "\g x\ \ m x" if "\n. \f n x\ \ m x" "(\n. f n x) \ g x" for x apply (rule LIMSEQ_le_const2[of "\n. \f n x\"]) using that by (auto intro!:tendsto_intros) ultimately have *: "AE x in M. \g x\ \ m x" using \AE x in M. (\n. f n x) \ g x\ by auto show "g \ space\<^sub>N(\ p M)" apply (rule Lp_domination[OF _ \m \ space\<^sub>N (\ p M)\]) using * by auto have "(\n. Norm (\ p M) (f n - g)) \ 0" proof (cases "p \ 0") case True then have "ennreal p = 0" by (simp add: ennreal_eq_0_iff) then show ?thesis unfolding Norm_def by (auto simp add: L_zero(1)) next case False then have "p > 0" by auto have "(\n. (\x. \f n x - g x\ powr p \M)) \ (\x. \0\ powr p \M)" proof (rule integral_dominated_convergence[of _ _ _ "(\x. \2 * m x\ powr p)"], auto) show "integrable M (\x. \2 * m x\ powr p)" unfolding abs_mult apply (subst powr_mult) using Lp_D(2)[OF \p > 0\ \m \ space\<^sub>N (\ p M)\] by auto have "(\n. \f n x - g x\ powr p) \ \0\ powr p" if "(\n. f n x) \ g x" for x apply (rule tendsto_powr') using \p > 0\ that apply (auto) using Lim_null tendsto_rabs_zero_iff by fastforce then show "AE x in M. (\n. \f n x - g x\ powr p) \ 0" using \AE x in M. (\n. f n x) \ g x\ by auto have "\f n x - g x\ powr p \ \2 * m x\ powr p" if "\f n x\ \ m x" "\g x\ \ m x" for n x using powr_mono2 \p > 0\ that by auto then show "AE x in M. \f n x - g x\ powr p \ \2 * m x\ powr p" for n using \AE x in M. \f n x\ \ m x\ \AE x in M. \g x\ \ m x\ by auto qed then have "(\n. (Norm (\ p M) (f n - g)) powr p) \ (Norm (\ p M) 0) powr p" unfolding Lp_D[OF \p > 0\ spaceN_diff[OF \\n. f n \ space\<^sub>N(\ p M)\ \g \ space\<^sub>N(\ p M)\]] using \p > 0\ by (auto simp add: powr_powr) then have "(\n. ((Norm (\ p M) (f n - g)) powr p) powr (1/p)) \ ((Norm (\ p M) 0) powr p) powr (1/p)" by (rule tendsto_powr', auto simp add: \p > 0\) then show ?thesis using powr_powr \p > 0\ by auto qed then show "tendsto_in\<^sub>N (\ p M) f g" unfolding tendsto_in\<^sub>N_def by auto qed text \We give the version of the monotone convergence theorem in the setting of $L^p$ spaces.\ proposition Lp_monotone_limit: fixes f::"nat \ 'a \ real" assumes "p > (0::ennreal)" "AE x in M. incseq (\n. f n x)" "\n. Norm (\ p M) (f n) \ C" "\n. f n \ space\<^sub>N (\ p M)" shows "AE x in M. convergent (\n. f n x)" "(\x. lim (\n. f n x)) \ space\<^sub>N (\ p M)" "Norm (\ p M) (\x. lim (\n. f n x)) \ C" proof - have [measurable]: "f n \ borel_measurable M" for n using Lp_measurable[OF assms(4)]. show "AE x in M. convergent (\n. f n x)" proof (cases rule: Lp_cases[of p]) case PInf have "AE x in M. \f n x\ \ C" for n using L_infinity_AE_bound[of "f n" M] \Norm (\ p M) (f n) \ C\ \f n \ space\<^sub>N (\ p M)\ unfolding \p=\\ by auto then have *: "AE x in M. \n. \f n x\ \ C" by (subst AE_all_countable, auto) have "(\n. f n x) \ (SUP n. f n x)" if "incseq (\n. f n x)" "\n. \f n x\ \ C" for x apply (rule LIMSEQ_incseq_SUP[OF _ \incseq (\n. f n x)\]) using that(2) abs_le_D1 by fastforce then have "convergent (\n. f n x)" if "incseq (\n. f n x)" "\n. \f n x\ \ C" for x unfolding convergent_def using that by auto then show ?thesis using \AE x in M. incseq (\n. f n x)\ * by auto next case (real_pos p2) define g where "g = (\n. f n - f 0)" have "AE x in M. incseq (\n. g n x)" unfolding g_def using \AE x in M. incseq (\n. f n x)\ by (simp add: incseq_def) have "g n \ space\<^sub>N (\ p2 M)" for n unfolding g_def using \\n. f n \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ by auto then have [measurable]: "g n \ borel_measurable M" for n using Lp_measurable by auto define D where "D = defect (\ p2 M) * C + defect (\ p2 M) * C" have "Norm (\ p2 M) (g n) \ D" for n proof - have "f n \ space\<^sub>N (\ p2 M)" using \f n \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ by auto have "Norm (\ p2 M) (g n) \ defect (\ p2 M) * Norm (\ p2 M) (f n) + defect (\ p2 M) * Norm (\ p2 M) (f 0)" unfolding g_def using Norm_triangular_ineq_diff[OF \f n \ space\<^sub>N (\ p2 M)\] by auto also have "... \ D" unfolding D_def apply(rule add_mono) using mult_left_mono defect_ge_1[of "\ p2 M"] \\n. Norm (\ p M) (f n) \ C\ unfolding \p = ennreal p2\ by auto finally show ?thesis by simp qed have g_bound: "(\\<^sup>+x. \g n x\ powr p2 \M) \ ennreal(D powr p2)" for n proof - have "(\\<^sup>+x. \g n x\ powr p2 \M) = ennreal(\x. \g n x\ powr p2 \M)" apply (rule nn_integral_eq_integral) using Lp_D(2)[OF \p2 > 0\ \g n \ space\<^sub>N (\ p2 M)\] by auto also have "... = ennreal((Norm (\ p2 M) (g n)) powr p2)" apply (subst Lp_Norm(2)[OF \p2 > 0\, of "g n", symmetric]) by auto also have "... \ ennreal(D powr p2)" by (auto intro!: powr_mono2 simp add: less_imp_le[OF \p2 > 0\] \Norm (\ p2 M) (g n) \ D\) finally show ?thesis by simp qed have "\n. g n x \ 0" if "incseq (\n. f n x)" for x unfolding g_def using that by (auto simp add: incseq_def) then have "AE x in M. \n. g n x \ 0" using \AE x in M. incseq (\n. f n x)\ by auto define h where "h = (\n x. ennreal(\g n x\ powr p2))" have [measurable]: "h n \ borel_measurable M" for n unfolding h_def by auto define H where "H = (\x. (SUP n. h n x))" have [measurable]: "H \ borel_measurable M" unfolding H_def by auto have "\n. h n x \ h (Suc n) x" if "\n. g n x \ 0" "incseq (\n. g n x)" for x unfolding h_def apply (auto intro!:powr_mono2) apply (auto simp add: less_imp_le[OF \p2 > 0\]) using that incseq_SucD by auto then have *: "AE x in M. h n x \ h (Suc n) x" for n using \AE x in M. \n. g n x \ 0\ \AE x in M. incseq (\n. g n x)\ by auto have "(\\<^sup>+x. H x \M) = (SUP n. \\<^sup>+x. h n x \M)" unfolding H_def by (rule nn_integral_monotone_convergence_SUP_AE, auto simp add: *) also have "... \ ennreal(D powr p2)" unfolding H_def h_def using g_bound by (simp add: SUP_least) finally have "(\\<^sup>+x. H x \M) < \" by (simp add: le_less_trans) then have "AE x in M. H x \ \" by (metis (mono_tags, lifting) \H \ borel_measurable M\ infinity_ennreal_def nn_integral_noteq_infinite top.not_eq_extremum) have "convergent (\n. f n x)" if "H x \ \" "incseq (\n. f n x)" for x proof - define A where "A = enn2real(H x)" then have "H x = ennreal A" using \H x \ \\ by (simp add: ennreal_enn2real_if) have "f n x \ f 0 x + A powr (1/p2)" for n proof - have "ennreal(\g n x\ powr p2) \ ennreal A" unfolding \H x = ennreal A\[symmetric] H_def h_def by (meson SUP_upper2 UNIV_I order_refl) then have "\g n x\ powr p2 \ A" by (subst ennreal_le_iff[symmetric], auto simp add: A_def) have "\g n x\ = (\g n x\ powr p2) powr (1/p2)" using \p2 > 0\ by (simp add: powr_powr) also have "... \ A powr (1/p2)" apply (rule powr_mono2) using \p2 > 0\ \\g n x\ powr p2 \ A\ by auto finally have "\g n x\ \ A powr (1/p2)" by simp then show ?thesis unfolding g_def by auto qed then show "convergent (\n. f n x)" using LIMSEQ_incseq_SUP[OF _ \incseq (\n. f n x)\] convergent_def by (metis bdd_aboveI2) qed then show "AE x in M. convergent (\n. f n x)" using \AE x in M. H x \ \\ \AE x in M. incseq (\n. f n x)\ by auto qed (insert \p>0\, simp) then have lim: "AE x in M. (\n. f n x) \ lim (\n. f n x)" using convergent_LIMSEQ_iff by auto show "(\x. lim (\n. f n x)) \ space\<^sub>N (\ p M)" apply (rule Lp_AE_limit''[of _ _ f, OF _ \\n. f n \ space\<^sub>N (\ p M)\ lim \\n. Norm (\ p M) (f n) \ C\]) by auto show "Norm (\ p M) (\x. lim (\n. f n x)) \ C" apply (rule Lp_AE_limit''[of _ _ f, OF _ \\n. f n \ space\<^sub>N (\ p M)\ lim \\n. Norm (\ p M) (f n) \ C\]) by auto qed subsection \Completeness of $L^p$\ text \We prove the completeness of $L^p$.\ theorem Lp_complete: "complete\<^sub>N (\ p M)" proof (cases rule: Lp_cases[of p]) case zero show ?thesis proof (rule complete\<^sub>N_I) fix u assume "\(n::nat). u n \ space\<^sub>N (\ p M)" then have "tendsto_in\<^sub>N (\ p M) u 0" unfolding tendsto_in\<^sub>N_def Norm_def \p = 0\ L_zero(1) L_zero_space by auto then show "\x\space\<^sub>N (\ p M). tendsto_in\<^sub>N (\ p M) u x" by auto qed next case (real_pos p2) show ?thesis proof (rule complete\<^sub>N_I'[of "\n. (1/2)^n * (1/(defect (\ p M))^(Suc n))"], unfold \p = ennreal p2\) show "0 < (1/2) ^ n * (1 / defect (\ (ennreal p2) M) ^ Suc n)" for n using defect_ge_1[of "\ (ennreal p2) M"] by (auto simp add: divide_simps) fix u assume "\(n::nat). u n \ space\<^sub>N (\ p2 M)" "\n. Norm (\ p2 M) (u n) \ (1/2)^n * (1/(defect (\ p2 M))^(Suc n))" then have H: "\n. u n \ space\<^sub>N (\ p2 M)" "\n. Norm (\ p2 M) (u n) \ (1/2) ^ n * (1/(defect (\ p2 M))^(Suc n))" unfolding \p = ennreal p2\ by auto have [measurable]: "u n \ borel_measurable M" for n using Lp_measurable[OF H(1)]. define w where "w = (\N x. (\n\{..u n x\))" have w2: "w = (\N. sum (\n x. \u n x\) {..N. w N x)" for x unfolding w2 by (rule incseq_SucI, auto) then have wN_inc: "AE x in M. incseq (\N. w N x)" by simp have abs_u_space: "(\x. \u n x\) \ space\<^sub>N (\ p2 M)" for n by (rule Lp_Banach_lattice[OF \u n \ space\<^sub>N (\ p2 M)\]) then have wN_space: "w N \ space\<^sub>N (\ p2 M)" for N unfolding w2 using H(1) by auto have abs_u_Norm: "Norm (\ p2 M) (\x. \u n x\) \ (1/2) ^ n * (1/(defect (\ p2 M))^(Suc n))" for n using Lp_Banach_lattice(2)[OF \u n \ space\<^sub>N (\ p2 M)\] H(2) by auto have wN_Norm: "Norm (\ p2 M) (w N) \ 2" for N proof - have *: "(defect (\ p2 M))^(Suc n) \ 0" "(defect (\ p2 M))^(Suc n) > 0" for n using defect_ge_1[of "\ p2 M"] by auto have "Norm (\ p2 M) (w N) \ (\n p2 M))^(Suc n) * Norm (\ p2 M) (\x. \u n x\))" unfolding w2 lessThan_Suc_atMost[symmetric] by (rule Norm_sum, simp add: abs_u_space) also have "... \ (\n p2 M))^(Suc n) * ((1/2) ^ n * (1/(defect (\ p2 M))^(Suc n))))" apply (rule sum_mono, rule mult_left_mono) using abs_u_Norm * by auto also have "... = (\n p2 M"] by (auto simp add: algebra_simps) also have "... \ (\n. (1/2) ^ n)" unfolding lessThan_Suc_atMost[symmetric] by (rule sum_le_suminf, rule summable_geometric[of "1/2"], auto) also have "... = 2" using suminf_geometric[of "1/2"] by auto finally show ?thesis by simp qed have "AE x in M. convergent (\N. w N x)" apply (rule Lp_monotone_limit[OF \p > 0\, of _ _ 2], unfold \p = ennreal p2\) using wN_inc wN_Norm wN_space by auto define m where "m = (\x. lim (\N. w N x))" have m_space: "m \ space\<^sub>N (\ p2 M)" unfolding m_def \p = ennreal p2\[symmetric] apply (rule Lp_monotone_limit[OF \p > 0\, of _ _ 2], unfold \p = ennreal p2\) using wN_inc wN_Norm wN_space by auto define v where "v = (\x. (\n. u n x))" have v_meas: "v \ borel_measurable M" unfolding v_def by auto have u_meas: "\n. (sum u {0.. borel_measurable M" by auto { fix x assume "convergent (\N. w N x)" then have S: "summable (\n. \u n x\)" unfolding w_def using summable_iff_convergent by auto then have "m x = (\n. \u n x\)" unfolding m_def w_def by (metis suminf_eq_lim) have "summable (\n. u n x)" using S by (rule summable_rabs_cancel) then have *: "(\n. (sum u {.. v x" unfolding v_def fun_sum_apply by (metis convergent_LIMSEQ_iff suminf_eq_lim summable_iff_convergent) have "\(sum u {.. \ m x" for n proof - have "\(sum u {.. \ (\i\{..u i x\)" unfolding fun_sum_apply by auto also have "... \ (\i. \u i x\)" apply (rule sum_le_suminf) using S by auto finally show ?thesis using \m x = (\n. \u n x\)\ by simp qed then have "(\n. \(sum u {0.. \ m x) \ (\n. (sum u {0.. v x" unfolding atLeast0LessThan using * by auto } then have m_bound: "\n. AE x in M. \(sum u {0.. \ m x" and u_conv: "AE x in M. (\n. (sum u {0.. v x" using \AE x in M. convergent (\N. w N x)\ by auto have "tendsto_in\<^sub>N (\ p2 M) (\n. sum u {0.. space\<^sub>N (\ p2 M)" by (rule Lp_domination_limit[OF v_meas u_meas m_space u_conv m_bound]) ultimately show "\v \ space\<^sub>N (\ p2 M). tendsto_in\<^sub>N (\ p2 M) (\n. sum u {0..N_I'[of "\n. (1/2)^n"]) fix u assume "\(n::nat). u n \ space\<^sub>N (\ p M)" "\n. Norm (\ p M) (u n) \ (1/2) ^ n" then have H: "\n. u n \ space\<^sub>N (\ \ M)" "\n. Norm (\ \ M) (u n) \ (1/2) ^ n" using PInf by auto have [measurable]: "u n \ borel_measurable M" for n using Lp_measurable[OF H(1)] by auto define v where "v = (\x. \n. u n x)" have [measurable]: "v \ borel_measurable M" unfolding v_def by auto define w where "w = (\N x. (\n\{0.. borel_measurable M" for N unfolding w_def by auto have "AE x in M. \u n x\ \ (1/2)^n" for n using L_infinity_AE_bound[OF H(1), of n] H(2)[of n] by auto then have "AE x in M. \n. \u n x\ \ (1/2)^n" by (subst AE_all_countable, auto) moreover have "\w N x - v x\ \ (1/2)^N * 2" if "\n. \u n x\ \ (1/2)^n" for N x proof - have *: "\n. \u n x\ \ (1/2)^n" using that by auto have **: "summable (\n. \u n x\)" apply (rule summable_norm_cancel, rule summable_comparison_test'[OF summable_geometric[of "1/2"]]) using * by auto have "\w N x - v x\ = \(\n. u (n + N) x)\" unfolding v_def w_def apply (subst suminf_split_initial_segment[OF summable_rabs_cancel[OF \summable (\n. \u n x\)\], of "N"]) by (simp add: lessThan_atLeast0) also have "... \ (\n. \u (n + N) x\)" apply (rule summable_rabs, subst summable_iff_shift) using ** by auto also have "... \ (\n. (1/2)^(n + N))" proof (rule suminf_le) show "\n. \u (n + N) x\ \ (1/2) ^ (n + N)" using *[of "_ + N"] by simp show "summable (\n. \u (n + N) x\)" using ** by (subst summable_iff_shift) simp show "summable (\n. (1/2::real) ^ (n + N))" using summable_geometric [of "1/2"] by (subst summable_iff_shift) simp qed also have "... = (1/2)^N * (\n. (1/2)^n)" by (subst power_add, subst suminf_mult2[symmetric], auto simp add: summable_geometric[of "1/2"]) also have "... = (1/2)^N * 2" by (subst suminf_geometric, auto) finally show ?thesis by simp qed ultimately have *: "AE x in M. \w N x - v x\ \ (1/2)^N * 2" for N by auto have **: "w N - v \ space\<^sub>N (\ \ M)" "Norm (\ \ M) (w N - v) \ (1/2)^N * 2" for N unfolding fun_diff_def using L_infinity_I[OF _ *] by auto have l: "(\N. ((1/2)^N) * (2::real)) \ 0 * 2" by (rule tendsto_mult, auto simp add: LIMSEQ_realpow_zero[of "1/2"]) have "tendsto_in\<^sub>N (\ \ M) w v" unfolding tendsto_in\<^sub>N_def apply (rule tendsto_sandwich[of "\_. 0" _ _ "\n. (1/2)^n * 2"]) using l **(2) by auto have "v = - (w 0 - v)" unfolding w_def by auto then have "v \ space\<^sub>N (\ \ M)" using **(1)[of 0] spaceN_add spaceN_diff by fastforce then show "\v \ space\<^sub>N (\ p M). tendsto_in\<^sub>N (\ p M) (\n. sum u {0..tendsto_in\<^sub>N (\ \ M) w v\ unfolding \p = \\ w_def fun_sum_apply[symmetric] by auto qed (simp) qed subsection \Multiplication of functions, duality\ text \The next theorem asserts that the multiplication of two functions in $L^p$ and $L^q$ belongs to $L^r$, where $r$ is determined by the equality $1/r = 1/p + 1/q$. This is essentially a case by case analysis, depending on the kind of $L^p$ space we are considering. The only nontrivial case is when $p$, $q$ (and $r$) are finite and nonzero. In this case, it reduces to H\"older inequality.\ theorem Lp_Lq_mult: fixes p q r::ennreal assumes "1/p + 1/q = 1/r" and "f \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ q M)" shows "(\x. f x * g x) \ space\<^sub>N (\ r M)" "Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" proof - have [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" using Lp_measurable assms by auto have "(\x. f x * g x) \ space\<^sub>N (\ r M) \ Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" proof (cases rule: Lp_cases[of r]) case zero have *: "(\x. f x * g x) \ borel_measurable M" by auto then have "Norm (\ r M) (\x. f x * g x) = 0" using L_zero[of M] unfolding Norm_def zero by auto then have "Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" using Norm_nonneg by auto then show ?thesis unfolding zero using * L_zero_space[of M] by auto next case (real_pos r2) have "p > 0" "q > 0" using \1/p + 1/q = 1/r\ \r > 0\ by (metis ennreal_add_eq_top ennreal_divide_eq_top_iff ennreal_top_neq_one gr_zeroI zero_neq_one)+ consider "p = \" | "q = \" | "p < \ \ q < \" using top.not_eq_extremum by force then show ?thesis proof (cases) case 1 then have "q = r" using \1/p + 1/q = 1/r\ by (metis ennreal_divide_top infinity_ennreal_def one_divide_one_divide_ennreal semiring_normalization_rules(5)) have "AE x in M. \f x\ \ Norm (\ p M) f" using \f \ space\<^sub>N (\ p M)\ L_infinity_AE_bound unfolding \p = \\ by auto then have *: "AE x in M. \f x * g x\ \ \Norm (\ p M) f * g x\" unfolding abs_mult using Norm_nonneg[of "\ p M" f] mult_right_mono by fastforce have **: "(\x. Norm (\ p M) f * g x) \ space\<^sub>N (\ r M)" using spaceN_cmult[OF \g \ space\<^sub>N (\ q M)\] unfolding \q = r\ scaleR_fun_def by simp have ***: "Norm (\ r M) (\x. Norm (\ p M) f * g x) = Norm (\ p M) f * Norm (\ q M) g" using Norm_cmult[of "\ r M"] unfolding \q = r\ scaleR_fun_def by auto then show ?thesis using Lp_domination[of "\x. f x * g x" M "\x. Norm (\ p M) f * g x" r] unfolding \q = r\ using * ** *** by auto next case 2 then have "p = r" using \1/p + 1/q = 1/r\ by (metis add.right_neutral ennreal_divide_top infinity_ennreal_def one_divide_one_divide_ennreal) have "AE x in M. \g x\ \ Norm (\ q M) g" using \g \ space\<^sub>N (\ q M)\ L_infinity_AE_bound unfolding \q = \\ by auto then have *: "AE x in M. \f x * g x\ \ \Norm (\ q M) g * f x\" apply (simp only: mult.commute[of "Norm (\ q M) g" _]) unfolding abs_mult using mult_left_mono Norm_nonneg[of "\ q M" g] by fastforce have **: "(\x. Norm (\ q M) g * f x) \ space\<^sub>N (\ r M)" using spaceN_cmult[OF \f \ space\<^sub>N (\ p M)\] unfolding \p = r\ scaleR_fun_def by simp have ***: "Norm (\ r M) (\x. Norm (\ q M) g * f x) = Norm (\ p M) f * Norm (\ q M) g" using Norm_cmult[of "\ r M"] unfolding \p = r\ scaleR_fun_def by auto then show ?thesis using Lp_domination[of "\x. f x * g x" M "\x. Norm (\ q M) g * f x" r] unfolding \p = r\ using * ** *** by auto next case 3 obtain p2 where "p = ennreal p2" "p2 > 0" using enn2real_positive_iff[of p] 3 \p > 0\ by (cases p) auto obtain q2 where "q = ennreal q2" "q2 > 0" using enn2real_positive_iff[of q] 3 \q > 0\ by (cases q) auto have "ennreal(1/r2) = 1/r" using \r = ennreal r2\ \r2 > 0\ divide_ennreal zero_le_one by fastforce also have "... = 1/p + 1/q" using assms by auto also have "... = ennreal(1/p2 + 1/q2)" using \p = ennreal p2\ \p2 > 0\ \q = ennreal q2\ \q2 > 0\ apply (simp only: divide_ennreal ennreal_1[symmetric]) using ennreal_plus[of "1/p2" "1/q2", symmetric] by auto finally have *: "1/r2 = 1/p2 + 1/q2" using ennreal_inj \p2 > 0\ \q2 > 0\ \r2 > 0\ by (metis divide_pos_pos ennreal_less_zero_iff le_less zero_less_one) define P where "P = p2 / r2" define Q where "Q = q2 / r2" have [simp]: "P > 0" "Q > 0" and "1/P + 1/Q = 1" using \p2 > 0\ \q2 > 0\ \r2 > 0\ * unfolding P_def Q_def by (auto simp add: divide_simps algebra_simps) have Pa: "(\z\ powr r2) powr P = \z\ powr p2" for z unfolding P_def powr_powr using \r2 > 0\ by auto have Qa: "(\z\ powr r2) powr Q = \z\ powr q2" for z unfolding Q_def powr_powr using \r2 > 0\ by auto have *: "integrable M (\x. \f x\ powr r2 * \g x\ powr r2)" apply (rule Holder_inequality[OF \P>0\ \Q>0\ \1/P + 1/Q = 1\], auto simp add: Pa Qa) using \f \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ using Lp_space[OF \p2 > 0\] apply auto using \g \ space\<^sub>N (\ q M)\ unfolding \q = ennreal q2\ using Lp_space[OF \q2 > 0\] by auto have "(\x. f x * g x) \ space\<^sub>N (\ r M)" unfolding \r = ennreal r2\ using Lp_space[OF \r2 > 0\, of M] by (auto simp add: * abs_mult powr_mult) have "Norm (\ r M) (\x. f x * g x) = (\x. \f x * g x\ powr r2 \M) powr (1/r2)" unfolding \r = ennreal r2\ using Lp_Norm[OF \r2 > 0\, of _ M] by auto also have "... = abs (\x. \f x\ powr r2 * \g x\ powr r2 \M) powr (1/r2)" by (auto simp add: powr_mult abs_mult) also have "... \ ((\x. \ \f x\ powr r2 \ powr P \M) powr (1/P) * (\x. \ \g x\ powr r2 \ powr Q \M) powr (1/Q)) powr (1/r2)" apply (rule powr_mono2, simp add: \r2 > 0\ less_imp_le, simp) apply (rule Holder_inequality[OF \P>0\ \Q>0\ \1/P + 1/Q = 1\], auto simp add: Pa Qa) using \f \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ using Lp_space[OF \p2 > 0\] apply auto using \g \ space\<^sub>N (\ q M)\ unfolding \q = ennreal q2\ using Lp_space[OF \q2 > 0\] by auto also have "... = (\x. \f x\ powr p2 \M) powr (1/p2) * (\x. \g x\ powr q2 \M) powr (1/q2)" apply (auto simp add: powr_mult powr_powr) unfolding P_def Q_def using \r2 > 0\ by auto also have "... = Norm (\ p M) f * Norm (\ q M) g" unfolding \p = ennreal p2\ \q = ennreal q2\ using Lp_Norm[OF \p2 > 0\, of _ M] Lp_Norm[OF \q2 > 0\, of _ M] by auto finally show ?thesis using \(\x. f x * g x) \ space\<^sub>N (\ r M)\ by auto qed next case PInf then have "p = \" "q = r" using \1/p + 1/q = 1/r\ by (metis add_eq_0_iff_both_eq_0 ennreal_divide_eq_0_iff infinity_ennreal_def not_one_le_zero order.order_iff_strict)+ have "AE x in M. \f x\ \ Norm (\ p M) f" using \f \ space\<^sub>N (\ p M)\ L_infinity_AE_bound unfolding \p = \\ by auto then have *: "AE x in M. \f x * g x\ \ \Norm (\ p M) f * g x\" unfolding abs_mult using Norm_nonneg[of "\ p M" f] mult_right_mono by fastforce have **: "(\x. Norm (\ p M) f * g x) \ space\<^sub>N (\ r M)" using spaceN_cmult[OF \g \ space\<^sub>N (\ q M)\] unfolding \q = r\ scaleR_fun_def by simp have ***: "Norm (\ r M) (\x. Norm (\ p M) f * g x) = Norm (\ p M) f * Norm (\ q M) g" using Norm_cmult[of "\ r M"] unfolding \q = r\ scaleR_fun_def by auto then show ?thesis using Lp_domination[of "\x. f x * g x" M "\x. Norm (\ p M) f * g x" r] unfolding \q = r\ using * ** *** by auto qed then show "(\x. f x * g x) \ space\<^sub>N (\ r M)" "Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" by auto qed text \The previous theorem admits an eNorm version in which one does not assume a priori that the functions under consideration belong to $L^p$ or $L^q$.\ theorem Lp_Lq_emult: fixes p q r::ennreal assumes "1/p + 1/q = 1/r" "f \ borel_measurable M" "g \ borel_measurable M" shows "eNorm (\ r M) (\x. f x * g x) \ eNorm (\ p M) f * eNorm (\ q M) g" proof (cases "r = 0") case True then have "eNorm (\ r M) (\x. f x * g x) = 0" using assms by (simp add: L_zero(1)) then show ?thesis by auto next case False then have "r > 0" using not_gr_zero by blast then have "p > 0" "q > 0" using \1/p + 1/q = 1/r\ by (metis ennreal_add_eq_top ennreal_divide_eq_top_iff ennreal_top_neq_one gr_zeroI zero_neq_one)+ then have Z: "zero_space\<^sub>N (\ p M) = {f \ borel_measurable M. AE x in M. f x = 0}" "zero_space\<^sub>N (\ q M) = {f \ borel_measurable M. AE x in M. f x = 0}" "zero_space\<^sub>N (\ r M) = {f \ borel_measurable M. AE x in M. f x = 0}" using \r > 0\ Lp_infinity_zero_space by auto have [measurable]: "(\x. f x * g x) \ borel_measurable M" using assms by auto consider "eNorm (\ p M) f = 0 \ eNorm (\ q M) g = 0" | "(eNorm (\ p M) f > 0 \ eNorm (\ q M) g = \) \ (eNorm (\ p M) f = \ \ eNorm (\ q M) g > 0)" | "eNorm (\ p M) f < \ \ eNorm (\ q M) g < \" using less_top by fastforce then show ?thesis proof (cases) case 1 then have "(AE x in M. f x = 0) \ (AE x in M. g x = 0)" using Z unfolding zero_space\<^sub>N_def by auto then have "AE x in M. f x * g x = 0" by auto then have "eNorm (\ r M) (\x. f x * g x) = 0" using Z unfolding zero_space\<^sub>N_def by auto then show ?thesis by simp next case 2 then have "eNorm (\ p M) f * eNorm (\ q M) g = \" using ennreal_mult_eq_top_iff by force then show ?thesis by auto next case 3 then have *: "f \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ q M)" unfolding space\<^sub>N_def by auto then have "(\x. f x * g x) \ space\<^sub>N (\ r M)" using Lp_Lq_mult(1)[OF assms(1)] by auto then show ?thesis using Lp_Lq_mult(2)[OF assms(1) *] by (simp add: eNorm_Norm * ennreal_mult'[symmetric]) qed qed lemma Lp_Lq_duality_bound: fixes p q::ennreal assumes "1/p + 1/q = 1" "f \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ q M)" shows "integrable M (\x. f x * g x)" "abs(\x. f x * g x \M) \ Norm (\ p M) f * Norm (\ q M) g" proof - have "(\x. f x * g x) \ space\<^sub>N (\ 1 M)" apply (rule Lp_Lq_mult[OF _ \f \ space\<^sub>N (\ p M)\ \g \ space\<^sub>N (\ q M)\]) using \1/p + 1/q = 1\ by auto then show "integrable M (\x. f x * g x)" using L1_space by auto have "abs(\x. f x * g x \M) \ Norm (\ 1 M) (\x. f x * g x)" using L1_int_ineq by auto also have "... \ Norm (\ p M) f * Norm (\ q M) g" apply (rule Lp_Lq_mult[OF _ \f \ space\<^sub>N (\ p M)\ \g \ space\<^sub>N (\ q M)\]) using \1/p + 1/q = 1\ by auto finally show "abs(\x. f x * g x \M) \ Norm (\ p M) f * Norm (\ q M) g" by simp qed text \The next theorem asserts that the norm of an $L^p$ function $f$ can be obtained by estimating the integrals of $fg$ over all $L^q$ functions $g$, where $1/p + 1/q = 1$. When $p = \infty$, it is necessary to assume that the space is sigma-finite: for instance, if the space is one single atom of infinite mass, then there is no nonzero $L^1$ function, so taking for $f$ the constant function equal to $1$, it has $L^\infty$ norm equal to $1$, but $\int fg = 0$ for all $L^1$ function $g$.\ theorem Lp_Lq_duality: fixes p q::ennreal assumes "f \ space\<^sub>N (\ p M)" "1/p + 1/q = 1" "p = \ \ sigma_finite_measure M" shows "bdd_above ((\g. (\x. f x * g x \M))`{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1})" "Norm (\ p M) f = (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable[OF assms(1)] by auto have B: "(\x. f x * g x \M) \ Norm (\ p M) f" if "g \ {g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}" for g proof - have g: "g \ space\<^sub>N (\ q M)" "Norm (\ q M) g \ 1" using that by auto have "(\x. f x * g x \M) \ abs(\x. f x * g x \M)" by auto also have "... \ Norm (\ p M) f * Norm (\ q M) g" using Lp_Lq_duality_bound(2)[OF \1/p + 1/q = 1\ \f \ space\<^sub>N (\ p M)\ g(1)] by auto also have "... \ Norm (\ p M) f" using g(2) Norm_nonneg[of "\ p M" f] mult_left_le by blast finally show "(\x. f x * g x \M) \ Norm (\ p M) f" by simp qed then show "bdd_above ((\g. (\x. f x * g x \M))`{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1})" by (meson bdd_aboveI2) show "Norm (\ p M) f = (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" proof (rule antisym) show "(SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. \x. f x * g x \M) \ Norm (\ p M) f" by (rule cSUP_least, auto, rule exI[of _ 0], auto simp add: B) have "p \ 1" using conjugate_exponent_ennrealI(1)[OF \1/p + 1/q = 1\] by simp show "Norm (\ p M) f \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" using \p \ 1\ proof (cases rule: Lp_cases_1_PInf) case PInf then have "f \ space\<^sub>N (\ \ M)" using \f \ space\<^sub>N(\ p M)\ by simp have "q = 1" using \1/p + 1/q = 1\ \p = \\ by (simp add: divide_eq_1_ennreal) have "c \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" if "c < Norm (\ p M) f" for c proof (cases "c < 0") case True then have "c \ (\x. f x * 0 x \M)" by auto also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" apply (rule cSUP_upper, auto simp add: zero_fun_def[symmetric]) using B by (meson bdd_aboveI2) finally show ?thesis by simp next case False then have "ennreal c < eNorm (\ \ M) f" using eNorm_Norm[OF \f \ space\<^sub>N (\ p M)\] that ennreal_less_iff unfolding \p = \\ by auto then have *: "emeasure M {x \ space M. \f x\ > c} > 0" using L_infinity_pos_measure[of f M c] by auto obtain A where [measurable]: "\(n::nat). A n \ sets M" and "(\i. A i) = space M" "\i. emeasure M (A i) \ \" using sigma_finite_measure.sigma_finite[OF \p = \ \ sigma_finite_measure M\[OF \p = \\]] by (metis UNIV_I sets_range) define Y where "Y = (\n::nat. {x \ A n. \f x\ > c})" have [measurable]: "Y n \ sets M" for n unfolding Y_def by auto have "{x \ space M. \f x\ > c} = (\n. Y n)" unfolding Y_def using \(\i. A i) = space M\ by auto then have "emeasure M (\n. Y n) > 0" using * by auto then obtain n where "emeasure M (Y n) > 0" using emeasure_pos_unionE[of Y, OF \\n. Y n \ sets M\] by auto have "emeasure M (Y n) \ emeasure M (A n)" apply (rule emeasure_mono) unfolding Y_def by auto then have "emeasure M (Y n) \ \" using \emeasure M (A n) \ \\ by (metis infinity_ennreal_def neq_top_trans) then have "measure M (Y n) > 0" using \emeasure M (Y n) > 0\ unfolding measure_def by (simp add: enn2real_positive_iff top.not_eq_extremum) have "\f x\ \ c" if "x \ Y n" for x using that less_imp_le unfolding Y_def by auto define g where "g = (\x. indicator (Y n) x * sgn(f x)) /\<^sub>R measure M (Y n)" have "g \ space\<^sub>N (\ 1 M)" apply (rule Lp_domination[of _ _ "indicator (Y n) /\<^sub>R measure M (Y n)"]) unfolding g_def using L1_indicator'[OF \Y n \ sets M\ \emeasure M (Y n) \ \\] by (auto simp add: abs_mult indicator_def abs_sgn_eq) have "Norm (\ 1 M) g = Norm (\ 1 M) (\x. indicator (Y n) x * sgn(f x)) / abs(measure M (Y n))" unfolding g_def Norm_cmult by (simp add: divide_inverse) also have "... \ Norm (\ 1 M) (indicator (Y n)) / abs(measure M (Y n))" using \measure M (Y n) > 0\ apply (auto simp add: divide_simps) apply (rule Lp_domination) using L1_indicator'[OF \Y n \ sets M\ \emeasure M (Y n) \ \\] by (auto simp add: abs_mult indicator_def abs_sgn_eq) also have "... = measure M (Y n) / abs(measure M (Y n))" using L1_indicator'[OF \Y n \ sets M\ \emeasure M (Y n) \ \\] by (auto simp add: abs_mult indicator_def abs_sgn_eq) also have "... = 1" using \measure M (Y n) > 0\ by auto finally have "Norm (\ 1 M) g \ 1" by simp have "c * measure M (Y n) = (\x. c * indicator (Y n) x \M)" using \measure M (Y n) > 0\ \emeasure M (Y n) \ \\ by auto also have "... \ (\x. \f x\ * indicator (Y n) x \M)" apply (rule integral_mono) using \emeasure M (Y n) \ \\ \0 < Sigma_Algebra.measure M (Y n)\ not_integrable_integral_eq apply fastforce apply (rule Bochner_Integration.integrable_bound[of _ "\x. Norm (\ \ M) f * indicator (Y n) x"]) using \emeasure M (Y n) \ \\ \0 < Sigma_Algebra.measure M (Y n)\ not_integrable_integral_eq apply fastforce using L_infinity_AE_bound[OF \f \ space\<^sub>N (\ \ M)\] by (auto simp add: indicator_def Y_def) finally have "c \ (\x. \f x\ * indicator (Y n) x \M) / measure M (Y n)" using \measure M (Y n) > 0\ by (auto simp add: divide_simps) also have "... = (\x. f x * indicator (Y n) x * sgn(f x) / measure M (Y n) \M)" using \measure M (Y n) > 0\ by (simp add: abs_sgn mult.commute mult.left_commute) also have "... = (\x. f x * g x \M)" unfolding divide_inverse g_def divideR_apply by (auto simp add: algebra_simps) also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" unfolding \q = 1\ apply (rule cSUP_upper, auto) using \g \ space\<^sub>N (\ 1 M)\ \Norm (\ 1 M) g \ 1\ apply auto using B \p = \\ \q = 1\ by (meson bdd_aboveI2) finally show ?thesis by simp qed then show ?thesis using dense_le by auto next case one then have "q = \" using \1/p + 1/q = 1\ by simp define g where "g = (\x. sgn (f x))" have [measurable]: "g \ space\<^sub>N (\ \ M)" apply (rule L_infinity_I[of g M 1]) unfolding g_def by (auto simp add: abs_sgn_eq) have "Norm (\ \ M) g \ 1" apply (rule L_infinity_I[of g M 1]) unfolding g_def by (auto simp add: abs_sgn_eq) have "Norm (\ p M) f = (\x. \f x\ \M)" unfolding \p = 1\ apply (rule L1_D(3)) using \f \ space\<^sub>N (\ p M)\ unfolding \p = 1\ by auto also have "... = (\x. f x * g x \M)" unfolding g_def by (simp add: abs_sgn) also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" unfolding \q = \\ apply (rule cSUP_upper, auto) using \g \ space\<^sub>N (\ \ M)\ \Norm (\ \ M) g \ 1\ apply auto using B \q = \\ by fastforce finally show ?thesis by simp next case (gr p2) then have "p2 > 0" by simp have "f \ space\<^sub>N (\ p2 M)" using \f \ space\<^sub>N (\ p M)\ \p = ennreal p2\ by auto define q2 where "q2 = conjugate_exponent p2" have "q2 > 1" "q2 > 0" using conjugate_exponent_real(2)[OF \p2 > 1\] unfolding q2_def by auto have "q = ennreal q2" unfolding q2_def conjugate_exponent_real_ennreal[OF \p2 > 1\, symmetric] \p = ennreal p2\[symmetric] using conjugate_exponent_ennreal_iff[OF \p \ 1\] \1/p + 1/q = 1\ by auto show ?thesis proof (cases "Norm (\ p M) f = 0") case True then have "Norm (\ p M) f \ (\x. f x * 0 x \M)" by auto also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" apply (rule cSUP_upper, auto simp add: zero_fun_def[symmetric]) using B by (meson bdd_aboveI2) finally show ?thesis by simp next case False then have "Norm (\ p2 M) f > 0" unfolding \p = ennreal p2\ using Norm_nonneg[of "\ p2 M" f] by linarith define h where "h = (\x. sgn(f x) * \f x\ powr (p2 - 1))" have [measurable]: "h \ borel_measurable M" unfolding h_def by auto have "(\\<^sup>+x. \h x\ powr q2 \M) = (\\<^sup>+x. (\f x\ powr (p2 - 1)) powr q2 \M)" unfolding h_def by (rule nn_integral_cong, auto simp add: abs_mult abs_sgn_eq) also have "... = (\\<^sup>+x. \f x\ powr p2 \M)" unfolding powr_powr q2_def using conjugate_exponent_real(4)[OF \p2 > 1\] by auto also have "... = (Norm (\ p2 M) f) powr p2" apply (subst Lp_Norm(2), auto simp add: \p2 > 0\) by (rule nn_integral_eq_integral, auto simp add: Lp_D(2)[OF \p2 > 0\ \f \ space\<^sub>N (\ p2 M)\]) finally have *: "(\\<^sup>+x. \h x\ powr q2 \M) = (Norm (\ p2 M) f) powr p2" by simp have "integrable M (\x. \h x\ powr q2)" apply (rule integrableI_bounded, auto) using * by auto then have "(\x. \h x\ powr q2 \M) = (\\<^sup>+x. \h x\ powr q2 \M)" by (rule nn_integral_eq_integral[symmetric], auto) then have **: "(\x. \h x\ powr q2 \M) = (Norm (\ p2 M) f) powr p2" using * by auto define g where "g = (\x. h x / (Norm (\ p2 M) f) powr (p2 / q2))" have [measurable]: "g \ borel_measurable M" unfolding g_def by auto have intg: "integrable M (\x. \g x\ powr q2)" unfolding g_def using \Norm (\ p2 M) f > 0\ \q2 > 1\ apply (simp add: abs_mult powr_divide powr_powr) using \integrable M (\x. \h x\ powr q2)\ integrable_divide_zero by blast have "g \ space\<^sub>N (\ q2 M)" by (rule Lp_I(1)[OF \q2 > 0\ _ intg], auto) have "(\x. \g x\ powr q2 \M) = 1" unfolding g_def using \Norm (\ p2 M) f > 0\ \q2 > 1\ by (simp add: abs_mult powr_divide powr_powr **) then have "Norm (\ q2 M) g = 1" apply (subst Lp_D[OF \q2 > 0\]) using \g \ space\<^sub>N (\ q2 M)\ by auto have "(\x. f x * g x \M) = (\x. f x * sgn(f x) * \f x\ powr (p2 - 1) / (Norm (\ p2 M) f) powr (p2 / q2) \M)" unfolding g_def h_def by (simp add: mult.assoc) also have "... = (\x. \f x\ * \f x\ powr (p2-1) \M) / (Norm (\ p2 M) f) powr (p2 / q2)" by (auto simp add: abs_sgn) also have "... = (\x. \f x\ powr p2 \M) / (Norm (\ p2 M) f) powr (p2 / q2)" by (subst powr_mult_base, auto) also have "... = (Norm (\ p2 M) f) powr p2 / (Norm (\ p2 M) f) powr (p2 / q2)" by (subst Lp_Norm(2)[OF \p2 > 0\], auto) also have "... = (Norm (\ p2 M) f) powr (p2 - p2/q2)" by (simp add: powr_diff [symmetric] ) also have "... = Norm (\ p2 M) f" unfolding q2_def using conjugate_exponent_real(5)[OF \p2 > 1\] by auto finally have "Norm (\ p M) f = (\x. f x * g x \M)" unfolding \p = ennreal p2\ by simp also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" unfolding \q = ennreal q2\ apply (rule cSUP_upper, auto) using \g \ space\<^sub>N (\ q2 M)\ \Norm (\ q2 M) g = 1\ apply auto using B \q = ennreal q2\ by fastforce finally show ?thesis by simp qed qed qed qed text \The previous theorem admits a version in which one does not assume a priori that the function under consideration belongs to $L^p$. This gives an efficient criterion to check if a function is indeed in $L^p$. In this case, it is always necessary to assume that the measure is sigma-finite. Note that, in the statement, the Bochner integral $\int fg$ vanishes by definition if $fg$ is not integrable. Hence, the statement really says that the eNorm can be estimated using functions $g$ for which $fg$ is integrable. It is precisely the construction of such functions $g$ that requires the space to be sigma-finite.\ theorem Lp_Lq_duality': fixes p q::ennreal assumes "1/p + 1/q = 1" "sigma_finite_measure M" and [measurable]: "f \ borel_measurable M" shows "eNorm (\ p M) f = (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. ennreal(\x. f x * g x \M))" proof (cases "eNorm (\ p M) f \ \") case True then have "f \ space\<^sub>N (\ p M)" unfolding space\<^sub>N_def by (simp add: top.not_eq_extremum) show ?thesis unfolding eNorm_Norm[OF \f \ space\<^sub>N (\ p M)\] Lp_Lq_duality[OF \f \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\] apply (rule SUP_real_ennreal[symmetric], auto, rule exI[of _ 0], auto) by (rule Lp_Lq_duality[OF \f \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\]) next case False have B: "\g \ {g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M) \ C" if "C < \" for C::ennreal proof - obtain Cr where "C = ennreal Cr" "Cr \ 0" using \C < \\ ennreal_cases less_irrefl by auto obtain A where A: "\n::nat. A n \ sets M" "incseq A" "(\n. A n) = space M" "\n. emeasure M (A n) \ \" using sigma_finite_measure.sigma_finite_incseq[OF \sigma_finite_measure M\] by (metis range_subsetD) define Y where "Y = (\n. {x \ A n. \f x\ \ n})" have [measurable]: "\n. Y n \ sets M" unfolding Y_def using \\n::nat. A n \ sets M\ by auto have "incseq Y" apply (rule incseq_SucI) unfolding Y_def using incseq_SucD[OF \incseq A\] by auto have *: "\N. \n \ N. f x * indicator (Y n) x = f x" if "x \ space M" for x proof - obtain n0 where n0: "x \ A n0" using \x \ space M\ \(\n. A n) = space M\ by auto obtain n1::nat where n1: "\f x\ \ n1" using real_arch_simple by blast have "x \ Y (max n0 n1)" unfolding Y_def using n1 apply auto using n0 \incseq A\ incseq_def max.cobounded1 by blast then have *: "x \ Y n" if "n \ max n0 n1" for n using \incseq Y\ that incseq_def by blast show ?thesis by (rule exI[of _ "max n0 n1"], auto simp add: *) qed have *: "(\n. f x * indicator (Y n) x) \ f x" if "x \ space M" for x using *[OF that] unfolding eventually_sequentially[symmetric] by (simp add: tendsto_eventually) have "liminf (\n. eNorm (\ p M) (\x. f x * indicator (Y n) x)) \ eNorm (\ p M) f" apply (rule Lp_AE_limit) using * by auto then have "liminf (\n. eNorm (\ p M) (\x. f x * indicator (Y n) x)) > Cr" using False neq_top_trans by force then have "limsup (\n. eNorm (\ p M) (\x. f x * indicator (Y n) x)) > Cr" using Liminf_le_Limsup less_le_trans trivial_limit_sequentially by blast then obtain n where n: "eNorm (\ p M) (\x. f x * indicator (Y n) x) > Cr" using Limsup_obtain by blast have "(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)" apply (rule Lp_bounded_bounded_support[of _ _ n], auto) unfolding Y_def indicator_def apply auto by (metis (mono_tags, lifting) A(1) A(4) emeasure_mono infinity_ennreal_def mem_Collect_eq neq_top_trans subsetI) have "Norm (\ p M) (\x. f x * indicator (Y n) x) > Cr" using n unfolding eNorm_Norm[OF \(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)\] by (meson ennreal_leI not_le) then have "(SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * indicator (Y n) x * g x \M)) > Cr" using Lp_Lq_duality(2)[OF \(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\] by auto then have "\g \ {g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * indicator (Y n) x * g x \M) > Cr" apply (subst less_cSUP_iff[symmetric]) using Lp_Lq_duality(1)[OF \(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\] apply auto by (rule exI[of _ 0], auto) then obtain g where g: "g \ space\<^sub>N (\ q M)" "Norm (\ q M) g \ 1" "(\x. f x * indicator (Y n) x * g x \M) > Cr" by auto then have [measurable]: "g \ borel_measurable M" using Lp_measurable by auto define h where "h = (\x. indicator (Y n) x * g x)" have "Norm (\ q M) h \ Norm (\ q M) g" apply (rule Lp_domination[of _ _ g]) unfolding h_def indicator_def using \g \ space\<^sub>N (\ q M)\ by auto then have a: "Norm (\ q M) h \ 1" using \Norm (\ q M) g \ 1\ by auto have b: "h \ space\<^sub>N (\ q M)" apply (rule Lp_domination[of _ _ g]) unfolding h_def indicator_def using \g \ space\<^sub>N (\ q M)\ by auto have "(\x. f x * h x \M) > Cr" unfolding h_def using g(3) by (auto simp add: mult.assoc) then have "(\x. f x * h x \M) > C" unfolding \C = ennreal Cr\ using \Cr \ 0\ by (simp add: ennreal_less_iff) then show ?thesis using a b by auto qed have "(SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. ennreal(\x. f x * g x \M)) \ \" apply (rule dense_le) using B by (meson SUP_upper2) then show ?thesis using False neq_top_trans by force qed subsection \Conditional expectations and $L^p$\ text \The $L^p$ space with respect to a subalgebra is included in the whole $L^p$ space.\ lemma Lp_subalgebra: assumes "subalgebra M F" shows "\f. eNorm (\ p M) f \ eNorm (\ p (restr_to_subalg M F)) f" "(\ p (restr_to_subalg M F)) \\<^sub>N \ p M" "space\<^sub>N ((\ p (restr_to_subalg M F))) \ space\<^sub>N (\ p M)" "\f. f \ space\<^sub>N ((\ p (restr_to_subalg M F))) \ Norm (\ p M) f = Norm (\ p (restr_to_subalg M F)) f" proof - have *: "f \ space\<^sub>N (\ p M) \ Norm (\ p M) f = Norm (\ p (restr_to_subalg M F)) f" if "f \ space\<^sub>N (\ p (restr_to_subalg M F))" for f proof - have [measurable]: "f \ borel_measurable (restr_to_subalg M F)" using that Lp_measurable by auto then have [measurable]: "f \ borel_measurable M" using assms measurable_from_subalg measurable_in_subalg' by blast show ?thesis proof (cases rule: Lp_cases[of p]) case zero then show ?thesis using that unfolding \p = 0\ L_zero_space Norm_def L_zero by auto next case PInf have [measurable]: "f \ borel_measurable (restr_to_subalg M F)" using that Lp_measurable by auto then have [measurable]: "f \ borel_measurable F" using assms measurable_in_subalg' by blast then have [measurable]: "f \ borel_measurable M" using assms measurable_from_subalg by blast have "AE x in (restr_to_subalg M F). \f x\ \ Norm (\ \ (restr_to_subalg M F)) f" using L_infinity_AE_bound that unfolding \p = \\ by auto then have a: "AE x in M. \f x\ \ Norm (\ \ (restr_to_subalg M F)) f" using assms AE_restr_to_subalg by blast have *: "f \ space\<^sub>N (\ \ M)" "Norm (\ \ M) f \ Norm (\ \ (restr_to_subalg M F)) f" using L_infinity_I[OF \f \ borel_measurable M\ a] by auto then have b: "AE x in M. \f x\ \ Norm (\ \ M) f" using L_infinity_AE_bound by auto have c: "AE x in (restr_to_subalg M F). \f x\ \ Norm (\ \ M) f" apply (rule AE_restr_to_subalg2[OF assms]) using b by auto have "Norm (\ \ (restr_to_subalg M F)) f \ Norm (\ \ M) f" using L_infinity_I[OF \f \ borel_measurable (restr_to_subalg M F)\ c] by auto then show ?thesis using * unfolding \p = \\ by auto next case (real_pos p2) then have a [measurable]: "f \ space\<^sub>N (\ p2 (restr_to_subalg M F))" using that unfolding \p = ennreal p2\ by auto then have b [measurable]: "f \ space\<^sub>N (\ p2 M)" unfolding Lp_space[OF \p2 > 0\] using integrable_from_subalg[OF assms] by auto show ?thesis unfolding \p = ennreal p2\ Lp_D[OF \p2 > 0\ a] Lp_D[OF \p2 > 0\ b] using integral_subalgebra2[OF assms, symmetric, of f] apply (auto simp add: b) by (metis (mono_tags, lifting) \integrable (restr_to_subalg M F) (\x. \f x\ powr p2)\ assms integrableD(1) integral_subalgebra2 measurable_in_subalg') qed qed show "space\<^sub>N ((\ p (restr_to_subalg M F))) \ space\<^sub>N (\ p M)" using * by auto show "Norm (\ p M) f = Norm (\ p (restr_to_subalg M F)) f" if "f \ space\<^sub>N ((\ p (restr_to_subalg M F)))" for f using * that by auto show "eNorm (\ p M) f \ eNorm (\ p (restr_to_subalg M F)) f" for f by (metis "*" eNorm_Norm eq_iff infinity_ennreal_def less_imp_le spaceN_iff top.not_eq_extremum) then show "(\ p (restr_to_subalg M F)) \\<^sub>N \ p M" by (metis ennreal_1 mult.left_neutral quasinorm_subsetI) qed text \For $p \geq 1$, the conditional expectation of an $L^p$ function still belongs to $L^p$, with an $L^p$ norm which is bounded by the norm of the original function. This is wrong for $p < 1$. One can prove this separating the cases and using the conditional version of Jensen's inequality, but it is much more efficient to do it with duality arguments, as follows.\ proposition Lp_real_cond_exp: assumes [simp]: "subalgebra M F" and "p \ (1::ennreal)" "sigma_finite_measure (restr_to_subalg M F)" "f \ space\<^sub>N (\ p M)" shows "real_cond_exp M F f \ space\<^sub>N (\ p (restr_to_subalg M F))" "Norm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ Norm (\ p M) f" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable assms by auto define q where "q = conjugate_exponent p" have "1/p + 1/q = 1" unfolding q_def using conjugate_exponent_ennreal[OF \p \ 1\] by simp have "eNorm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) = (SUP g\{g \ space\<^sub>N (\ q (restr_to_subalg M F)). Norm (\ q (restr_to_subalg M F)) g \ 1}. ennreal(\x. (real_cond_exp M F f) x * g x \(restr_to_subalg M F)))" by (rule Lp_Lq_duality'[OF \1/p + 1/q = 1\ \sigma_finite_measure (restr_to_subalg M F)\], simp) also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. ennreal(\x. f x * g x \M))" proof (rule SUP_mono, auto) fix g assume H: "g \ space\<^sub>N (\ q (restr_to_subalg M F))" "Norm (\ q (restr_to_subalg M F)) g \ 1" then have H2: "g \ space\<^sub>N (\ q M)" "Norm (\ q M) g \ 1" using Lp_subalgebra[OF \subalgebra M F\] by (auto simp add: subset_iff) have [measurable]: "g \ borel_measurable M" "g \ borel_measurable F" using Lp_measurable[OF H(1)] Lp_measurable[OF H2(1)] by auto have int: "integrable M (\x. f x * g x)" using Lp_Lq_duality_bound(1)[OF \1/p + 1/q = 1\ \f \ space\<^sub>N (\ p M)\ H2(1)]. have "(\x. (real_cond_exp M F f) x * g x \(restr_to_subalg M F)) = (\x. g x * (real_cond_exp M F f) x \M)" by (subst mult.commute, rule integral_subalgebra2[OF \subalgebra M F\], auto) also have "... = (\x. g x * f x \M)" apply (rule sigma_finite_subalgebra.real_cond_exp_intg, auto simp add: int mult.commute) unfolding sigma_finite_subalgebra_def using assms by auto finally have "ennreal (\x. (real_cond_exp M F f) x * g x \(restr_to_subalg M F)) \ ennreal (\x. f x * g x \M)" by (auto intro!: ennreal_leI simp add: mult.commute) then show "\m. m \ space\<^sub>N (\ q M) \ Norm (\ q M) m \ 1 \ ennreal (LINT x|restr_to_subalg M F. real_cond_exp M F f x * g x) \ ennreal (LINT x|M. f x * m x)" using H2 by blast qed also have "... = eNorm (\ p M) f" apply (rule Lp_Lq_duality'[OF \1/p + 1/q = 1\, symmetric], auto intro!: sigma_finite_subalgebra_is_sigma_finite[of _ F]) unfolding sigma_finite_subalgebra_def using assms by auto finally have *: "eNorm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ eNorm (\ p M) f" by simp then show a: "real_cond_exp M F f \ space\<^sub>N (\ p (restr_to_subalg M F))" apply (subst spaceN_iff) using \f \ space\<^sub>N (\ p M)\ by (simp add: space\<^sub>N_def) show "Norm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ Norm (\ p M) f" using * unfolding eNorm_Norm[OF \f \ space\<^sub>N (\ p M)\] eNorm_Norm[OF a] by simp qed lemma Lp_real_cond_exp_eNorm: assumes [simp]: "subalgebra M F" and "p \ (1::ennreal)" "sigma_finite_measure (restr_to_subalg M F)" shows "eNorm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ eNorm (\ p M) f" proof (cases "eNorm (\ p M) f = \") case False then have *: "f \ space\<^sub>N (\ p M)" unfolding spaceN_iff by (simp add: top.not_eq_extremum) show ?thesis using Lp_real_cond_exp[OF assms \f \ space\<^sub>N (\ p M)\] by (subst eNorm_Norm, auto simp: \f \ space\<^sub>N (\ p M)\)+ qed (simp) end diff --git a/thys/Noninterference_CSP/ClassicalNoninterference.thy b/thys/Noninterference_CSP/ClassicalNoninterference.thy --- a/thys/Noninterference_CSP/ClassicalNoninterference.thy +++ b/thys/Noninterference_CSP/ClassicalNoninterference.thy @@ -1,1860 +1,1860 @@ (* Title: Noninterference Security in Communicating Sequential Processes Author: Pasquale Noce Security Certification Specialist at Arjo Systems - Gep S.p.A. pasquale dot noce dot lavoro at gmail dot com pasquale dot noce at arjowiggins-it dot com *) section "CSP noninterference vs. classical noninterference" theory ClassicalNoninterference imports CSPNoninterference begin text \ \null The purpose of this section is to prove the equivalence of CSP noninterference security as defined previously to the classical notion of noninterference security as formulated in \cite{R3} in the case of processes representing deterministic state machines, henceforth briefly referred to as \emph{classical processes}. For clarity, all the constants and fact names defined in this section, with the possible exception of main theorems, contain prefix \c_\. \ subsection "Classical noninterference" text \ Here below are the formalizations of the functions \emph{sources} and \emph{ipurge} defined in \cite{R3}, as well as of the classical notion of noninterference security as stated ibid. for a deterministic state machine in the general case of a possibly intransitive noninterference policy. Observe that the function \emph{run} used in \emph{R3} is formalized as function \foldl step\, where \step\ is the state transition function of the machine. \null \ primrec c_sources :: "('d \ 'd) set \ ('a \ 'd) \ 'd \ 'a list \ 'd set" where "c_sources _ _ u [] = {u}" | "c_sources I D u (x # xs) = (if \v \ c_sources I D u xs. (D x, v) \ I then insert (D x) (c_sources I D u xs) else c_sources I D u xs)" primrec c_ipurge :: "('d \ 'd) set \ ('a \ 'd) \ 'd \ 'a list \ 'a list" where "c_ipurge _ _ _ [] = []" | "c_ipurge I D u (x # xs) = (if D x \ c_sources I D u (x # xs) then x # c_ipurge I D u xs else c_ipurge I D u xs)" definition c_secure :: "('s \ 'a \ 's) \ ('s \ 'a \ 'o) \ 's \ ('d \ 'd) set \ ('a \ 'd) \ bool" where "c_secure step out s\<^sub>0 I D \ \x xs. out (foldl step s\<^sub>0 xs) x = out (foldl step s\<^sub>0 (c_ipurge I D (D x) xs)) x" text \ \null In addition, the definitions are given of variants of functions \c_sources\ and \c_ipurge\ accepting in input a set of security domains rather than a single domain, and then some lemmas concerning them are demonstrated. These definitions and lemmas will turn out to be useful in subsequent proofs. \null \ primrec c_sources_aux :: "('d \ 'd) set \ ('a \ 'd) \ 'd set \ 'a list \ 'd set" where "c_sources_aux _ _ U [] = U" | "c_sources_aux I D U (x # xs) = (if \v \ c_sources_aux I D U xs. (D x, v) \ I then insert (D x) (c_sources_aux I D U xs) else c_sources_aux I D U xs)" primrec c_ipurge_aux :: "('d \ 'd) set \ ('a \ 'd) \ 'd set \ 'a list \ 'a list" where "c_ipurge_aux _ _ _ [] = []" | "c_ipurge_aux I D U (x # xs) = (if D x \ c_sources_aux I D U (x # xs) then x # c_ipurge_aux I D U xs else c_ipurge_aux I D U xs)" lemma c_sources_aux_singleton_1: "c_sources_aux I D {u} xs = c_sources I D u xs" by (induction xs, simp_all) lemma c_ipurge_aux_singleton: "c_ipurge_aux I D {u} xs = c_ipurge I D u xs" by (induction xs, simp_all add: c_sources_aux_singleton_1) lemma c_sources_aux_singleton_2: "D x \ c_sources_aux I D U [x] = (D x \ U \ (\v \ U. (D x, v) \ I))" by simp lemma c_sources_aux_append: "c_sources_aux I D U (xs @ [x]) = (if D x \ c_sources_aux I D U [x] then c_sources_aux I D (insert (D x) U) xs else c_sources_aux I D U xs)" by (induction xs, simp_all add: insert_absorb) lemma c_ipurge_aux_append: "c_ipurge_aux I D U (xs @ [x]) = (if D x \ c_sources_aux I D U [x] then c_ipurge_aux I D (insert (D x) U) xs @ [x] else c_ipurge_aux I D U xs)" by (induction xs, simp_all add: c_sources_aux_append) text \ \null In what follows, a few useful lemmas are proven about functions \c_sources\, \c_ipurge\ and their relationships with functions \sinks\, \ipurge_tr\. \null \ lemma c_sources_ipurge: "c_sources I D u (c_ipurge I D u xs) = c_sources I D u xs" by (induction xs, simp_all) lemma c_sources_append_1: "c_sources I D (D x) (xs @ [x]) = c_sources I D (D x) xs" by (induction xs, simp_all) lemma c_ipurge_append_1: "c_ipurge I D (D x) (xs @ [x]) = c_ipurge I D (D x) xs @ [x]" by (induction xs, simp_all add: c_sources_append_1) lemma c_sources_append_2: "(D x, u) \ I \ c_sources I D u (xs @ [x]) = c_sources I D u xs" by (induction xs, simp_all) lemma c_ipurge_append_2: "refl I \ (D x, u) \ I \ c_ipurge I D u (xs @ [x]) = c_ipurge I D u xs" proof (induction xs, simp_all add: refl_on_def c_sources_append_2) qed (rule notI, simp) lemma c_sources_mono: assumes A: "c_sources I D u ys \ c_sources I D u zs" shows "c_sources I D u (x # ys) \ c_sources I D u (x # zs)" proof (cases "\v \ c_sources I D u ys. (D x, v) \ I") assume B: "\v \ c_sources I D u ys. (D x, v) \ I" then obtain v where C: "v \ c_sources I D u ys" and D: "(D x, v) \ I" .. from A and C have "v \ c_sources I D u zs" .. with D have E: "\v \ c_sources I D u zs. (D x, v) \ I" .. have "insert (D x) (c_sources I D u ys) \ insert (D x) (c_sources I D u zs)" using A by (rule insert_mono) moreover have "c_sources I D u (x # ys) = insert (D x) (c_sources I D u ys)" using B by simp moreover have "c_sources I D u (x # zs) = insert (D x) (c_sources I D u zs)" using E by simp ultimately show "c_sources I D u (x # ys) \ c_sources I D u (x # zs)" by simp next assume "\ (\v \ c_sources I D u ys. (D x, v) \ I)" hence "c_sources I D u (x # ys) = c_sources I D u ys" by simp hence "c_sources I D u (x # ys) \ c_sources I D u zs" using A by simp moreover have "c_sources I D u zs \ c_sources I D u (x # zs)" by (simp add: subset_insertI) ultimately show "c_sources I D u (x # ys) \ c_sources I D u (x # zs)" by simp qed lemma c_sources_sinks [rule_format]: "D x \ c_sources I D u (x # xs) \ sinks I D (D x) (c_ipurge I D u xs) = {}" proof (induction xs, simp, rule impI) fix x' xs assume A: "D x \ c_sources I D u (x # xs) \ sinks I D (D x) (c_ipurge I D u xs) = {}" assume B: "D x \ c_sources I D u (x # x' # xs)" have "c_sources I D u xs \ c_sources I D u (x' # xs)" by (simp add: subset_insertI) hence "c_sources I D u (x # xs) \ c_sources I D u (x # x' # xs)" by (rule c_sources_mono) hence "D x \ c_sources I D u (x # xs)" using B by (rule contra_subsetD) with A have C: "sinks I D (D x) (c_ipurge I D u xs) = {}" .. show "sinks I D (D x) (c_ipurge I D u (x' # xs)) = {}" proof (cases "D x' \ c_sources I D u (x' # xs)", simp_all only: c_ipurge.simps if_True if_False) assume D: "D x' \ c_sources I D u (x' # xs)" have "(D x, D x') \ I" proof assume "(D x, D x') \ I" hence "\v \ c_sources I D u (x' # xs). (D x, v) \ I" using D .. hence "D x \ c_sources I D u (x # x' # xs)" by simp thus False using B by contradiction qed thus "sinks I D (D x) (x' # c_ipurge I D u xs) = {}" using C by (simp add: sinks_cons_nonint) next show "sinks I D (D x) (c_ipurge I D u xs) = {}" using C . qed qed lemmas c_ipurge_tr_ipurge = c_sources_sinks [THEN sinks_empty] lemma c_ipurge_aux_ipurge_tr [rule_format]: assumes R: "refl I" shows "\ (\v \ sinks I D u ys. \w \ U. (v, w) \ I) \ c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D U (xs @ ys)" proof (induction ys arbitrary: U rule: rev_induct, simp, rule impI) fix y ys U assume A: "\U. \ (\v \ sinks I D u ys. \w \ U. (v, w) \ I) \ c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D U (xs @ ys)" and B: "\ (\v \ sinks I D u (ys @ [y]). \w \ U. (v, w) \ I)" have C: "\ (\v \ sinks I D u ys. \w \ U. (v, w) \ I)" proof (rule notI, (erule bexE)+) fix v w assume "(v, w) \ I" and "w \ U" hence "\w \ U. (v, w) \ I" .. moreover assume "v \ sinks I D u ys" hence "v \ sinks I D u (ys @ [y])" by simp ultimately have "\v \ sinks I D u (ys @ [y]). \w \ U. (v, w) \ I" .. thus False using B by contradiction qed show "c_ipurge_aux I D U (xs @ ipurge_tr I D u (ys @ [y])) = c_ipurge_aux I D U (xs @ ys @ [y])" proof (cases "D y \ c_sources_aux I D U [y]", case_tac [!] "D y \ sinks I D u (ys @ [y])", simp_all (no_asm_simp) only: ipurge_tr.simps append_assoc [symmetric] c_ipurge_aux_append append_same_eq if_True if_False) assume D: "D y \ sinks I D u (ys @ [y])" assume "D y \ c_sources_aux I D U [y]" hence "D y \ U \ (\w \ U. (D y, w) \ I)" by (simp only: c_sources_aux_singleton_2) moreover { have "(D y, D y) \ I" using R by (simp add: refl_on_def) moreover assume "D y \ U" ultimately have "\w \ U. (D y, w) \ I" .. hence "\v \ sinks I D u (ys @ [y]). \w \ U. (v, w) \ I" using D .. } moreover { assume "\w \ U. (D y, w) \ I" hence "\v \ sinks I D u (ys @ [y]). \w \ U. (v, w) \ I" using D .. } ultimately have "\v \ sinks I D u (ys @ [y]). \w \ U. (v, w) \ I" by blast thus "c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D (insert (D y) U) (xs @ ys) @ [y]" using B by contradiction next assume D: "D y \ sinks I D u (ys @ [y])" have "\ (\v \ sinks I D u ys. \w \ insert (D y) U. (v, w) \ I) \ c_ipurge_aux I D (insert (D y) U) (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D (insert (D y) U) (xs @ ys)" using A . moreover have "\ (\v \ sinks I D u ys. \w \ insert (D y) U. (v, w) \ I)" proof (rule notI, (erule bexE)+, simp, erule disjE, simp) fix v assume "(v, D y) \ I" and "v \ sinks I D u ys" hence "\v \ sinks I D u ys. (v, D y) \ I" .. hence "D y \ sinks I D u (ys @ [y])" by simp thus False using D by contradiction next fix v w assume "(v, w) \ I" and "w \ U" hence "\w \ U. (v, w) \ I" .. moreover assume "v \ sinks I D u ys" ultimately have "\v \ sinks I D u ys. \w \ U. (v, w) \ I" .. thus False using C by contradiction qed ultimately show "c_ipurge_aux I D (insert (D y) U) (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D (insert (D y) U) (xs @ ys)" .. next have "\ (\v \ sinks I D u ys. \w \ U. (v, w) \ I) \ c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D U (xs @ ys)" using A . thus "c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D U (xs @ ys)" using C .. next have "\ (\v \ sinks I D u ys. \w \ U. (v, w) \ I) \ c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D U (xs @ ys)" using A . thus "c_ipurge_aux I D U (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D U (xs @ ys)" using C .. qed qed lemma c_ipurge_ipurge_tr: assumes R: "refl I" and D: "\ (\v \ sinks I D u ys. (v, u') \ I)" shows "c_ipurge I D u' (xs @ ipurge_tr I D u ys) = c_ipurge I D u' (xs @ ys)" proof - have "\ (\v \ sinks I D u ys. \w \ {u'}. (v, w) \ I)" using D by simp with R have "c_ipurge_aux I D {u'} (xs @ ipurge_tr I D u ys) = c_ipurge_aux I D {u'} (xs @ ys)" by (rule c_ipurge_aux_ipurge_tr) thus ?thesis by (simp add: c_ipurge_aux_singleton) qed subsection "Classical processes" text \ The deterministic state machines used as model of computation in the classical theory of noninterference security, as expounded in \cite{R3}, have the property that each action produces an output. Hence, it is natural to take as alphabet of a classical process the universe of the pairs \(x, p)\, where \x\ is an action and \p\ an output. For any state \s\, such an event \(x, p)\ may occur just in case \p\ matches the output produced by \x\ in \s\. Therefore, a trace of a classical process can be defined as an event list \xps\ such that for each item \(x, p)\, \p\ is equal to the output produced by \x\ in the state resulting from the previous actions in \xps\. Furthermore, for each trace \xps\, the refusals set associated to \xps\ is comprised of any set of pairs \(x, p)\ such that \p\ is different from the output produced by \x\ in the state resulting from the actions in \xps\. In accordance with the previous considerations, an inductive definition is formulated here below for the failures set \c_failures step out s\<^sub>0\ corresponding to the deterministic state machine with state transition function \step\, output function \out\, and initial state \s\<^sub>0\. Then, the classical process \c_process step out s\<^sub>0\ representing this machine is defined as the process having \c_failures step out s\<^sub>0\ as failures set and the empty set as divergences set. \null \ inductive_set c_failures :: "('s \ 'a \ 's) \ ('s \ 'a \ 'o) \ 's \ ('a \ 'o) failure set" for step :: "'s \ 'a \ 's" and out :: "'s \ 'a \ 'o" and s\<^sub>0 :: 's where R0: "([], {(x, p). p \ out s\<^sub>0 x}) \ c_failures step out s\<^sub>0" | R1: "\(xps, _) \ c_failures step out s\<^sub>0; s = foldl step s\<^sub>0 (map fst xps)\ \ (xps @ [(x, out s x)], {(y, p). p \ out (step s x) y}) \ c_failures step out s\<^sub>0" | R2: "\(xps, Y) \ c_failures step out s\<^sub>0; X \ Y\ \ (xps, X) \ c_failures step out s\<^sub>0" definition c_process :: "('s \ 'a \ 's) \ ('s \ 'a \ 'o) \ 's \ ('a \ 'o) process" where "c_process step out s\<^sub>0 \ Abs_process (c_failures step out s\<^sub>0, {})" text \ \null In what follows, the fact that classical processes are indeed processes is proven as a theorem. \null \ lemma c_process_prop_1 [simp]: "process_prop_1 (c_failures step out s\<^sub>0, {})" proof (simp add: process_prop_1_def) have "([], {(x, p). p \ out s\<^sub>0 x}) \ c_failures step out s\<^sub>0" by (rule R0) moreover have "{} \ {(x, p). p \ out s\<^sub>0 x}" .. ultimately show "([], {}) \ c_failures step out s\<^sub>0" by (rule R2) qed lemma c_process_prop_2 [simp]: "process_prop_2 (c_failures step out s\<^sub>0, {})" proof (simp only: process_prop_2_def fst_conv, (rule allI)+, rule impI) fix xps xp X assume "(xps @ [xp], X) \ c_failures step out s\<^sub>0" hence "(butlast (xps @ [xp]), {}) \ c_failures step out s\<^sub>0" proof (rule c_failures.induct [where P = "\xps X. (butlast xps, {}) \ c_failures step out s\<^sub>0"], simp_all) have "([], {(x, p). p \ out s\<^sub>0 x}) \ c_failures step out s\<^sub>0" by (rule R0) moreover have "{} \ {(x, p). p \ out s\<^sub>0 x}" .. ultimately show "([], {}) \ c_failures step out s\<^sub>0" by (rule R2) next fix xps' X' assume "(xps', X') \ c_failures step out s\<^sub>0" moreover have "{} \ X'" .. ultimately show "(xps', {}) \ c_failures step out s\<^sub>0" by (rule R2) qed thus "(xps, {}) \ c_failures step out s\<^sub>0" by simp qed lemma c_process_prop_3 [simp]: "process_prop_3 (c_failures step out s\<^sub>0, {})" by (simp only: process_prop_3_def fst_conv, (rule allI)+, rule impI, erule conjE, rule R2) lemma c_process_prop_4 [simp]: "process_prop_4 (c_failures step out s\<^sub>0, {})" proof (simp only: process_prop_4_def fst_conv, (rule allI)+, rule impI) fix xps xp X assume "(xps, X) \ c_failures step out s\<^sub>0" thus "(xps @ [xp], {}) \ c_failures step out s\<^sub>0 \ (xps, insert xp X) \ c_failures step out s\<^sub>0" proof (case_tac xp, rule c_failures.induct) fix x p assume A: "xp = (x, p)" have B: "([], {(x, p). p \ out s\<^sub>0 x}) \ c_failures step out s\<^sub>0" (is "(_, ?X) \ _") by (rule R0) show "([] @ [xp], {}) \ c_failures step out s\<^sub>0 \ ([], insert xp ?X) \ c_failures step out s\<^sub>0" proof (cases "p = out s\<^sub>0 x") assume C: "p = out s\<^sub>0 x" have "s\<^sub>0 = foldl step s\<^sub>0 (map fst [])" by simp with B have "([] @ [(x, out s\<^sub>0 x)], {(y, p). p \ out (step s\<^sub>0 x) y}) \ c_failures step out s\<^sub>0" (is "(_, ?Y) \ _") by (rule R1) hence "([] @ [xp], ?Y) \ c_failures step out s\<^sub>0" using A and C by simp moreover have "{} \ ?Y" .. ultimately have "([] @ [xp], {}) \ c_failures step out s\<^sub>0" by (rule R2) thus ?thesis .. next assume "p \ out s\<^sub>0 x" hence "xp \ ?X" using A by simp hence "insert xp ?X = ?X" by (rule insert_absorb) hence "([], insert xp ?X) \ c_failures step out s\<^sub>0" using B by simp thus ?thesis .. qed next fix x p xps' X' s x' let ?s = "step s x'" assume A: "xp = (x, p)" assume "(xps', X') \ c_failures step out s\<^sub>0" and S: "s = foldl step s\<^sub>0 (map fst xps')" hence B: "(xps' @ [(x', out s x')], {(y, p). p \ out ?s y}) \ c_failures step out s\<^sub>0" (is "(?xps, ?X) \ _") by (rule R1) show "(?xps @ [xp], {}) \ c_failures step out s\<^sub>0 \ (?xps, insert xp ?X) \ c_failures step out s\<^sub>0" proof (cases "p = out ?s x") assume C: "p = out ?s x" have "?s = foldl step s\<^sub>0 (map fst ?xps)" using S by simp with B have "(?xps @ [(x, out ?s x)], {(y, p). p \ out (step ?s x) y}) \ c_failures step out s\<^sub>0" (is "(_, ?Y) \ _") by (rule R1) hence "(?xps @ [xp], ?Y) \ c_failures step out s\<^sub>0" using A and C by simp moreover have "{} \ ?Y" .. ultimately have "(?xps @ [xp], {}) \ c_failures step out s\<^sub>0" by (rule R2) thus ?thesis .. next assume "p \ out ?s x" hence "xp \ ?X" using A by simp hence "insert xp ?X = ?X" by (rule insert_absorb) hence "(?xps, insert xp ?X) \ c_failures step out s\<^sub>0" using B by simp thus ?thesis .. qed next fix xps' X' Y assume "(xps' @ [xp], {}) \ c_failures step out s\<^sub>0 \ (xps', insert xp Y) \ c_failures step out s\<^sub>0" (is "?A \ ?B") and "X' \ Y" show "(xps' @ [xp], {}) \ c_failures step out s\<^sub>0 \ (xps', insert xp X') \ c_failures step out s\<^sub>0" using \?A \ ?B\ proof (rule disjE) assume ?A thus ?thesis .. next assume ?B moreover have "insert xp X' \ insert xp Y" using \X' \ Y\ by (rule insert_mono) ultimately have "(xps', insert xp X') \ c_failures step out s\<^sub>0" by (rule R2) thus ?thesis .. qed qed qed lemma c_process_prop_5 [simp]: "process_prop_5 (F, {})" by (simp add: process_prop_5_def) lemma c_process_prop_6 [simp]: "process_prop_6 (F, {})" by (simp add: process_prop_6_def) theorem c_process_process: "(c_failures step out s\<^sub>0, {}) \ process_set" by (simp add: process_set_def) text \ \null The continuation of this section is dedicated to the proof of a few lemmas on the properties of classical processes, particularly on the application to them of the generic functions acting on processes defined previously, and culminates in the theorem stating that classical processes are deterministic. Since they are intended to be a representation of deterministic state machines as processes, this result provides an essential confirmation of the correctness of such correspondence. \null \ lemma c_failures_last [rule_format]: "(xps, X) \ c_failures step out s\<^sub>0 \ xps \ [] \ snd (last xps) = out (foldl step s\<^sub>0 (butlast (map fst xps))) (last (map fst xps))" by (erule c_failures.induct, simp_all) lemma c_failures_ref: "(xps, X) \ c_failures step out s\<^sub>0 \ X \ {(x, p). p \ out (foldl step s\<^sub>0 (map fst xps)) x}" by (erule c_failures.induct, simp_all) lemma c_failures_failures: "failures (c_process step out s\<^sub>0) = c_failures step out s\<^sub>0" by (simp add: failures_def c_process_def c_process_process Abs_process_inverse) lemma c_futures_failures: "(yps, Y) \ futures (c_process step out s\<^sub>0) xps = ((xps @ yps, Y) \ c_failures step out s\<^sub>0)" by (simp add: futures_def failures_def c_process_def c_process_process Abs_process_inverse) lemma c_traces: "xps \ traces (c_process step out s\<^sub>0) = (\X. (xps, X) \ c_failures step out s\<^sub>0)" by (simp add: traces_def failures_def Domain_iff c_process_def c_process_process Abs_process_inverse) lemma c_refusals: "X \ refusals (c_process step out s\<^sub>0) xps = ((xps, X) \ c_failures step out s\<^sub>0)" by (simp add: refusals_def c_failures_failures) lemma c_next_events: "xp \ next_events (c_process step out s\<^sub>0) xps = (\X. (xps @ [xp], X) \ c_failures step out s\<^sub>0)" by (simp add: next_events_def c_traces) lemma c_traces_failures: "xps \ traces (c_process step out s\<^sub>0) \ (xps, {(x, p). p \ out (foldl step s\<^sub>0 (map fst xps)) x}) \ c_failures step out s\<^sub>0" proof (simp add: c_traces, erule exE, rule rev_cases [of xps], simp_all add: R0 split_paired_all) fix yps y p Y assume A: "(yps @ [(y, p)], Y) \ c_failures step out s\<^sub>0" let ?s = "foldl step s\<^sub>0 (map fst yps)" let ?ys' = "map fst (yps @ [(y, p)])" have "(yps @ [(y, p)], Y) \ failures (c_process step out s\<^sub>0)" using A by (simp add: c_failures_failures) hence "(yps, {}) \ failures (c_process step out s\<^sub>0)" by (rule process_rule_2) hence "(yps, {}) \ c_failures step out s\<^sub>0" by (simp add: c_failures_failures) moreover have "?s = foldl step s\<^sub>0 (map fst yps)" by simp ultimately have "(yps @ [(y, out ?s y)], {(x, p). p \ out (step ?s y) x}) \ c_failures step out s\<^sub>0" by (rule R1) moreover have "yps @ [(y, p)] \ []" by simp with A have "snd (last (yps @ [(y, p)])) = out (foldl step s\<^sub>0 (butlast ?ys')) (last ?ys')" by (rule c_failures_last) hence "p = out ?s y" by simp ultimately show "(yps @ [(y, p)], {(x, p). p \ out (step ?s y) x}) \ c_failures step out s\<^sub>0" by simp qed theorem c_process_deterministic: "deterministic (c_process step out s\<^sub>0)" proof (simp add: deterministic_def c_refusals c_next_events set_eq_iff, rule ballI, rule allI) fix xps X assume T: "xps \ traces (c_process step out s\<^sub>0)" let ?s = "foldl step s\<^sub>0 (map fst xps)" show "(xps, X) \ c_failures step out s\<^sub>0 = (\x p. (x, p) \ X \ (\X. (xps @ [(x, p)], X) \ c_failures step out s\<^sub>0))" (is "?P = ?Q") proof (rule iffI, (rule allI)+, rule impI, rule allI, rule notI) fix x p Y let ?xs' = "map fst (xps @ [(x, p)])" assume ?P hence "X \ {(x, p). p \ out ?s x}" (is "_ \ ?X'") by (rule c_failures_ref) moreover assume "(x, p) \ X" ultimately have "(x, p) \ ?X'" .. hence A: "p \ out ?s x" by simp assume "(xps @ [(x, p)], Y) \ c_failures step out s\<^sub>0" moreover have "xps @ [(x, p)] \ []" by simp ultimately have "snd (last (xps @ [(x, p)])) = out (foldl step s\<^sub>0 (butlast ?xs')) (last ?xs')" by (rule c_failures_last) hence "p = out ?s x" by simp thus False using A by contradiction next assume ?Q have A: "(xps, {(x, p). p \ out ?s x}) \ c_failures step out s\<^sub>0" using T by (rule c_traces_failures) moreover have "X \ {(x, p). p \ out ?s x}" proof (rule subsetI, simp add: split_paired_all, rule notI) fix x p assume "(x, p) \ X" and "p = out ?s x" hence "(xps @ [(x, out ?s x)], {(y, p). p \ out (step ?s x) y}) \ c_failures step out s\<^sub>0" using \?Q\ by simp moreover have "?s = foldl step s\<^sub>0 (map fst xps)" by simp with A have "(xps @ [(x, out ?s x)], {(y, p). p \ out (step ?s x) y}) \ c_failures step out s\<^sub>0" by (rule R1) ultimately show False by contradiction qed ultimately show ?P by (rule R2) qed qed subsection "Traces in classical processes" text \ Here below is the definition of function \c_tr\, where \c_tr step out s xs\ is the trace of classical process \c_process step out s\ corresponding to the trace \xs\ of the associated deterministic state machine. Moreover, some useful lemmas are proven about this function. \null \ function c_tr :: "('s \ 'a \ 's) \ ('s \ 'a \ 'o) \ 's \ 'a list \ ('a \ 'o) list" where "c_tr _ _ _ [] = []" | "c_tr step out s (xs @ [x]) = c_tr step out s xs @ [(x, out (foldl step s xs) x)]" proof (atomize_elim, simp_all add: split_paired_all) qed (rule rev_cases, rule disjI1, assumption, simp) termination by lexicographic_order lemma c_tr_length: "length (c_tr step out s xs) = length xs" by (rule rev_induct, simp_all) lemma c_tr_map: "map fst (c_tr step out s xs) = xs" by (rule rev_induct, simp_all) lemma c_tr_singleton: "c_tr step out s [x] = [(x, out s x)]" proof - have "c_tr step out s [x] = c_tr step out s ([] @ [x])" by simp also have "\ = c_tr step out s [] @ [(x, out (foldl step s []) x)]" by (rule c_tr.simps(2)) also have "\ = [(x, out s x)]" by simp finally show ?thesis . qed lemma c_tr_append: "c_tr step out s (xs @ ys) = c_tr step out s xs @ c_tr step out (foldl step s xs) ys" proof (rule_tac xs = ys in rev_induct, simp, subst append_assoc [symmetric]) qed (simp del: append_assoc) lemma c_tr_hd_tl: assumes A: "xs \ []" shows "c_tr step out s xs = (hd xs, out s (hd xs)) # c_tr step out (step s (hd xs)) (tl xs)" proof - let ?s = "foldl step s [hd xs]" have "c_tr step out s ([hd xs] @ tl xs) = c_tr step out s [hd xs] @ c_tr step out ?s (tl xs)" by (rule c_tr_append) moreover have "[hd xs] @ tl xs = xs" using A by simp ultimately have "c_tr step out s xs = c_tr step out s [hd xs] @ c_tr step out ?s (tl xs)" by simp moreover have "c_tr step out s [hd xs] = [(hd xs, out s (hd xs))]" by (simp add: c_tr_singleton) ultimately show ?thesis by simp qed lemma c_failures_tr: "(xps, X) \ c_failures step out s\<^sub>0 \ xps = c_tr step out s\<^sub>0 (map fst xps)" by (erule c_failures.induct, simp_all) lemma c_futures_tr: assumes A: "(yps, Y) \ futures (c_process step out s\<^sub>0) xps" shows "yps = c_tr step out (foldl step s\<^sub>0 (map fst xps)) (map fst yps)" proof - have B: "(xps @ yps, Y) \ c_failures step out s\<^sub>0" using A by (simp add: c_futures_failures) hence "xps @ yps = c_tr step out s\<^sub>0 (map fst (xps @ yps))" by (rule c_failures_tr) hence "xps @ yps = c_tr step out s\<^sub>0 (map fst xps) @ c_tr step out (foldl step s\<^sub>0 (map fst xps)) (map fst yps)" by (simp add: c_tr_append) moreover have "(xps @ yps, Y) \ failures (c_process step out s\<^sub>0)" using B by (simp add: c_failures_failures) hence "(xps, {}) \ failures (c_process step out s\<^sub>0)" by (rule process_rule_2_failures) hence "(xps, {}) \ c_failures step out s\<^sub>0" by (simp add: c_failures_failures) hence "xps = c_tr step out s\<^sub>0 (map fst xps)" by (rule c_failures_tr) ultimately show ?thesis by simp qed lemma c_tr_failures: "(c_tr step out s\<^sub>0 xs, {(x, p). p \ out (foldl step s\<^sub>0 xs) x}) \ c_failures step out s\<^sub>0" proof (rule rev_induct, simp_all, rule R0) fix x xs let ?s = "foldl step s\<^sub>0 (map fst (c_tr step out s\<^sub>0 xs))" assume "(c_tr step out s\<^sub>0 xs, {(x, p). p \ out (foldl step s\<^sub>0 xs) x}) \ c_failures step out s\<^sub>0" moreover have "?s = foldl step s\<^sub>0 (map fst (c_tr step out s\<^sub>0 xs))" by simp ultimately have "(c_tr step out s\<^sub>0 xs @ [(x, out ?s x)], {(y, p). p \ out (step ?s x) y}) \ c_failures step out s\<^sub>0" by (rule R1) moreover have "?s = foldl step s\<^sub>0 xs" by (simp add: c_tr_map) ultimately show "(c_tr step out s\<^sub>0 xs @ [(x, out (foldl step s\<^sub>0 xs) x)], {(y, p). p \ out (step (foldl step s\<^sub>0 xs) x) y}) \ c_failures step out s\<^sub>0" by simp qed lemma c_tr_futures: "(c_tr step out (foldl step s\<^sub>0 xs) ys, {(x, p). p \ out (foldl step (foldl step s\<^sub>0 xs) ys) x}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" proof (simp add: c_futures_failures) have "(c_tr step out s\<^sub>0 (xs @ ys), {(x, p). p \ out (foldl step s\<^sub>0 (xs @ ys)) x}) \ c_failures step out s\<^sub>0" by (rule c_tr_failures) moreover have "c_tr step out s\<^sub>0 (xs @ ys) = c_tr step out s\<^sub>0 xs @ c_tr step out (foldl step s\<^sub>0 xs) ys" by (rule c_tr_append) ultimately show "(c_tr step out s\<^sub>0 xs @ c_tr step out (foldl step s\<^sub>0 xs) ys, {(x, p). p \ out (foldl step (foldl step s\<^sub>0 xs) ys) x}) \ c_failures step out s\<^sub>0" by simp qed subsection "Noninterference in classical processes" text \ Given a mapping \D\ of the actions of a deterministic state machine into their security domains, it is natural to map each event \(x, p)\ of the corresponding classical process into the domain \D x\ of action \x\. Such mapping of events into domains, formalized as function \c_dom D\ in the continuation, ensures that the same noninterference policy applying to a deterministic state machine be applicable to the associated classical process as well. This is the simplest, and thus preferable way to construct a policy for the process such as to be isomorphic to the one assigned for the machine, as required in order to prove the equivalence of CSP noninterference security to the classical notion in the case of classical processes. In what follows, function \c_dom\ will be used in the proof of some useful lemmas concerning the application of functions \sinks\, \ipurge_tr\, \c_sources\, \c_ipurge\ from noninterference theory to the traces of classical processes, constructed by means of function \c_tr\. \null \ definition c_dom :: "('a \ 'd) \ ('a \ 'o) \ 'd" where "c_dom D xp \ D (fst xp)" lemma c_dom_sources: "c_sources I (c_dom D) u xps = c_sources I D u (map fst xps)" by (induction xps, simp_all add: c_dom_def) lemma c_dom_sinks: "sinks I (c_dom D) u xps = sinks I D u (map fst xps)" by (induction xps rule: rev_induct, simp_all add: c_dom_def) lemma c_tr_sources: "c_sources I (c_dom D) u (c_tr step out s xs) = c_sources I D u xs" by (simp add: c_dom_sources c_tr_map) lemma c_tr_sinks: "sinks I (c_dom D) u (c_tr step out s xs) = sinks I D u xs" by (simp add: c_dom_sinks c_tr_map) lemma c_tr_ipurge: "c_ipurge I (c_dom D) u (c_tr step out s (c_ipurge I D u xs)) = c_tr step out s (c_ipurge I D u xs)" proof (induction xs arbitrary: s, simp) fix x xs s assume A: "\s. c_ipurge I (c_dom D) u (c_tr step out s (c_ipurge I D u xs)) = c_tr step out s (c_ipurge I D u xs)" show "c_ipurge I (c_dom D) u (c_tr step out s (c_ipurge I D u (x # xs))) = c_tr step out s (c_ipurge I D u (x # xs))" proof (cases "D x \ c_sources I D u (x # xs)", simp_all del: c_sources.simps) have B: "c_tr step out s (x # c_ipurge I D u xs) = (x, out s x) # c_tr step out (step s x) (c_ipurge I D u xs)" by (simp add: c_tr_hd_tl) assume C: "D x \ c_sources I D u (x # xs)" hence "D x \ c_sources I D u (c_ipurge I D u (x # xs))" by (subst c_sources_ipurge) hence "D x \ c_sources I (c_dom D) u (c_tr step out s (x # c_ipurge I D u xs))" using C by (simp add: c_tr_sources) hence "c_dom D (x, out s x) \ c_sources I (c_dom D) u ((x, out s x) # c_tr step out (step s x) (c_ipurge I D u xs))" using B by (simp add: c_dom_def) hence "c_ipurge I (c_dom D) u (c_tr step out s (x # c_ipurge I D u xs)) = (x, out s x) # c_ipurge I (c_dom D) u (c_tr step out (step s x) (c_ipurge I D u xs))" using B by simp moreover have "c_ipurge I (c_dom D) u (c_tr step out (step s x) (c_ipurge I D u xs)) = c_tr step out (step s x) (c_ipurge I D u xs)" using A . ultimately show "c_ipurge I (c_dom D) u (c_tr step out s (x # c_ipurge I D u xs)) = c_tr step out s (x # c_ipurge I D u xs)" using B by simp next show "c_ipurge I (c_dom D) u (c_tr step out s (c_ipurge I D u xs)) = c_tr step out s (c_ipurge I D u xs)" using A . qed qed lemma c_tr_ipurge_tr_1 [rule_format]: "(\n \ {.. sinks I D u (take (Suc n) xs) \ out (foldl step s (ipurge_tr I D u (take n xs))) (xs ! n) = out (foldl step s (take n xs)) (xs ! n)) \ ipurge_tr I (c_dom D) u (c_tr step out s xs) = c_tr step out s (ipurge_tr I D u xs)" proof (induction xs rule: rev_induct, simp, rule impI) fix x xs assume "(\n \ {.. sinks I D u (take (Suc n) xs) \ out (foldl step s (ipurge_tr I D u (take n xs))) (xs ! n) = out (foldl step s (take n xs)) (xs ! n)) \ ipurge_tr I (c_dom D) u (c_tr step out s xs) = c_tr step out s (ipurge_tr I D u xs)" moreover assume A: "\n \ {.. sinks I D u (take (Suc n) (xs @ [x])) \ out (foldl step s (ipurge_tr I D u (take n (xs @ [x])))) ((xs @ [x]) ! n) = out (foldl step s (take n (xs @ [x]))) ((xs @ [x]) ! n)" have "\n \ {.. sinks I D u (take (Suc n) xs) \ out (foldl step s (ipurge_tr I D u (take n xs))) (xs ! n) = out (foldl step s (take n xs)) (xs ! n)" proof (rule ballI, rule impI) fix n assume B: "n \ {.. {.. sinks I D u (take (Suc n) (xs @ [x])) \ out (foldl step s (ipurge_tr I D u (take n (xs @ [x])))) ((xs @ [x]) ! n) = out (foldl step s (take n (xs @ [x]))) ((xs @ [x]) ! n)" .. hence "D (xs ! n) \ sinks I D u (take (Suc n) xs) \ out (foldl step s (ipurge_tr I D u (take n xs))) (xs ! n) = out (foldl step s (take n xs)) (xs ! n)" using B by (simp add: nth_append) moreover assume "D (xs ! n) \ sinks I D u (take (Suc n) xs)" ultimately show "out (foldl step s (ipurge_tr I D u (take n xs))) (xs ! n) = out (foldl step s (take n xs)) (xs ! n)" .. qed ultimately have C: "ipurge_tr I (c_dom D) u (c_tr step out s xs) = c_tr step out s (ipurge_tr I D u xs)" .. show "ipurge_tr I (c_dom D) u (c_tr step out s (xs @ [x])) = c_tr step out s (ipurge_tr I D u (xs @ [x]))" proof (cases "D x \ sinks I D u (xs @ [x])") case True then have "D x \ sinks I (c_dom D) u (c_tr step out s (xs @ [x]))" by (subst c_tr_sinks) hence "c_dom D (x, out (foldl step s xs) x) \ sinks I (c_dom D) u (c_tr step out s xs @ [(x, out (foldl step s xs) x)])" by (simp add: c_dom_def) with True show ?thesis using C by simp next case False then have "D x \ sinks I (c_dom D) u (c_tr step out s (xs @ [x]))" by (subst c_tr_sinks) hence "c_dom D (x, out (foldl step s xs) x) \ sinks I (c_dom D) u (c_tr step out s xs @ [(x, out (foldl step s xs) x)])" by (simp add: c_dom_def) with False show ?thesis proof (simp add: C) have "length xs \ {.. sinks I D u (take (Suc (length xs)) (xs @ [x])) \ out (foldl step s (ipurge_tr I D u (take (length xs) (xs @ [x])))) ((xs @ [x]) ! (length xs)) = out (foldl step s (take (length xs) (xs @ [x]))) ((xs @ [x]) ! (length xs))" .. hence "D x \ sinks I D u (xs @ [x]) \ out (foldl step s (ipurge_tr I D u xs)) x = out (foldl step s xs) x" by simp thus "out (foldl step s xs) x = out (foldl step s (ipurge_tr I D u xs)) x" using False by simp qed qed qed lemma c_tr_ipurge_tr_2 [rule_format]: assumes A: "\n \ {..length ys}. \Y. (ipurge_tr I (c_dom D) u (c_tr step out (foldl step s\<^sub>0 xs) (take n ys)), Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" shows "n \ {.. D (ys ! n) \ sinks I D u (take (Suc n) ys) \ out (foldl step (foldl step s\<^sub>0 xs) (ipurge_tr I D u (take n ys))) (ys ! n) = out (foldl step (foldl step s\<^sub>0 xs) (take n ys)) (ys ! n)" proof (rule nat_less_induct, (rule impI)+) fix n let ?s = "foldl step s\<^sub>0 xs" let ?yp = "(ys ! n, out (foldl step ?s (take n ys)) (ys ! n))" assume B: "\m < n. m \ {.. D (ys ! m) \ sinks I D u (take (Suc m) ys) \ out (foldl step ?s (ipurge_tr I D u (take m ys))) (ys ! m) = out (foldl step ?s (take m ys)) (ys ! m)" and C: "n \ {.. sinks I D u (take (Suc n) ys)" have "n < length ys" using C by simp hence E: "take (Suc n) ys = take n ys @ [ys ! n]" by (rule take_Suc_conv_app_nth) moreover have "Suc n \ {..length ys}" using C by simp with A have "\Y. (ipurge_tr I (c_dom D) u (c_tr step out ?s (take (Suc n) ys)), Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" .. then obtain Y where "(ipurge_tr I (c_dom D) u (c_tr step out ?s (take (Suc n) ys)), Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" .. ultimately have "(ipurge_tr I (c_dom D) u (c_tr step out ?s (take n ys) @ [?yp]), Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" by simp moreover have "c_dom D ?yp \ sinks I (c_dom D) u (c_tr step out ?s (take (Suc n) ys))" using D by (simp add: c_dom_def c_tr_sinks) hence "c_dom D ?yp \ sinks I (c_dom D) u (c_tr step out ?s (take n ys) @ [?yp])" using E by simp ultimately have "(ipurge_tr I (c_dom D) u (c_tr step out ?s (take n ys)) @ [?yp], Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" by simp moreover have "ipurge_tr I (c_dom D) u (c_tr step out ?s (take n ys)) = c_tr step out ?s (ipurge_tr I D u (take n ys))" - proof (rule c_tr_ipurge_tr_1, simp, erule conjE, simp add: min_def) + proof (rule c_tr_ipurge_tr_1, simp, erule conjE) fix m have "m < n \ m \ {.. D (ys ! m) \ sinks I D u (take (Suc m) ys) \ out (foldl step ?s (ipurge_tr I D u (take m ys))) (ys ! m) = out (foldl step ?s (take m ys)) (ys ! m)" using B .. moreover assume "m < n" ultimately have "m \ {.. D (ys ! m) \ sinks I D u (take (Suc m) ys) \ out (foldl step ?s (ipurge_tr I D u (take m ys))) (ys ! m) = out (foldl step ?s (take m ys)) (ys ! m)" .. moreover assume "m < length ys" hence "m \ {.. sinks I D u (take (Suc m) ys) \ out (foldl step ?s (ipurge_tr I D u (take m ys))) (ys ! m) = out (foldl step ?s (take m ys)) (ys ! m)" .. moreover assume "D (ys ! m) \ sinks I D u (take (Suc m) ys)" ultimately show "out (foldl step ?s (ipurge_tr I D u (take m ys))) (ys ! m) = out (foldl step ?s (take m ys)) (ys ! m)" .. qed ultimately have "(c_tr step out ?s (ipurge_tr I D u (take n ys)) @ [?yp], Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" by simp hence "(c_tr step out s\<^sub>0 (xs @ ipurge_tr I D u (take n ys)) @ [?yp], Y) \ c_failures step out s\<^sub>0" (is "(?xps, _) \ _") by (simp add: c_futures_failures c_tr_append) moreover have "?xps \ []" by simp ultimately have "snd (last ?xps) = out (foldl step s\<^sub>0 (butlast (map fst ?xps))) (last (map fst ?xps))" by (rule c_failures_last) thus "out (foldl step ?s (ipurge_tr I D u (take n ys))) (ys ! n) = out (foldl step ?s (take n ys)) (ys ! n)" by (simp add: c_tr_map butlast_append) qed lemma c_tr_ipurge_tr [rule_format]: assumes A: "\n \ {..length ys}. \Y. (ipurge_tr I (c_dom D) u (c_tr step out (foldl step s\<^sub>0 xs) (take n ys)), Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" shows "ipurge_tr I (c_dom D) u (c_tr step out (foldl step s\<^sub>0 xs) ys) = c_tr step out (foldl step s\<^sub>0 xs) (ipurge_tr I D u ys)" proof (rule c_tr_ipurge_tr_1) fix n have "\n. n \ {..length ys} \ \Y. (ipurge_tr I (c_dom D) u (c_tr step out (foldl step s\<^sub>0 xs) (take n ys)), Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" using A .. moreover assume "n \ {.. sinks I D u (take (Suc n) ys)" ultimately show "out (foldl step (foldl step s\<^sub>0 xs) (ipurge_tr I D u (take n ys))) (ys ! n) = out (foldl step (foldl step s\<^sub>0 xs) (take n ys)) (ys ! n)" by (rule c_tr_ipurge_tr_2) qed subsection "Equivalence between security properties" text \ The remainder of this section is dedicated to the proof of the equivalence between the CSP noninterference security of a classical process and the classical noninterference security of the corresponding deterministic state machine. In some detail, it will be proven that CSP noninterference security alone is a sufficient condition for classical noninterference security, whereas the latter security property entails the former for any reflexive noninterference policy. Therefore, the security properties under consideration turn out to be equivalent if the enforced noninterference policy is reflexive, which is the case for any policy of practical significance. \null \ lemma secure_implies_c_secure_aux: assumes S: "secure (c_process step out s\<^sub>0) I (c_dom D)" shows "out (foldl step (foldl step s\<^sub>0 xs) ys) x = out (foldl step (foldl step s\<^sub>0 xs) (c_ipurge I D (D x) ys)) x" proof (induction ys arbitrary: xs, simp) fix y ys xs assume "\xs. out (foldl step (foldl step s\<^sub>0 xs) ys) x = out (foldl step (foldl step s\<^sub>0 xs) (c_ipurge I D (D x) ys)) x" hence A: "out (foldl step (foldl step s\<^sub>0 (xs @ [y])) ys) x = out (foldl step (foldl step s\<^sub>0 (xs @ [y])) (c_ipurge I D (D x) ys)) x" . show "out (foldl step (foldl step s\<^sub>0 xs) (y # ys)) x = out (foldl step (foldl step s\<^sub>0 xs) (c_ipurge I D (D x) (y # ys))) x" proof (cases "D y \ c_sources I D (D x) (y # ys)") assume "D y \ c_sources I D (D x) (y # ys)" thus ?thesis using A by simp next let ?s = "foldl step s\<^sub>0 xs" let ?yp = "(y, out ?s y)" have "(c_tr step out ?s [y], {(x', p). p \ out (foldl step ?s [y]) x'}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" (is "(_, ?Y) \ _") by (rule c_tr_futures) hence "([?yp], ?Y) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" by (simp add: c_tr_hd_tl) moreover have "(c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])), {(x', p). p \ out (foldl step ?s (c_ipurge I D (D x) (ys @ [x]))) x'}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" (is "(_, ?Z) \ _") by (rule c_tr_futures) ultimately have "(?yp # ipurge_tr I (c_dom D) (c_dom D ?yp) (c_tr step out ?s (c_ipurge I D (D x) (ys @ [x]))), ipurge_ref I (c_dom D) (c_dom D ?yp) (c_tr step out ?s (c_ipurge I D (D x) (ys @ [x]))) ?Z) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" (is "(_, ?X) \ _") using S by (simp add: secure_def) hence C: "(?yp # ipurge_tr I (c_dom D) (c_dom D ?yp) (c_ipurge I (c_dom D) (D x) (c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])))), ?X) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" by (simp add: c_tr_ipurge) assume D: "D y \ c_sources I D (D x) (y # ys)" hence "D y \ c_sources I D (D x) ((y # ys) @ [x])" by (subst c_sources_append_1) hence "D y \ c_sources I D (D x) (y # ys @ [x])" by simp moreover have "c_sources I D (D x) (y # ys @ [x]) = c_sources I D (D x) (y # c_ipurge I D (D x) (ys @ [x]))" by (simp add: c_sources_ipurge) ultimately have "D y \ c_sources I D (D x) (y # c_ipurge I D (D x) (ys @ [x]))" by simp moreover have "map fst (?yp # c_tr step out ?s (c_ipurge I D (D x) (ys @ [x]))) = y # c_ipurge I D (D x) (ys @ [x])" by (simp add: c_tr_map) hence "c_sources I D (D x) (y # c_ipurge I D (D x) (ys @ [x])) = c_sources I (c_dom D) (D x) (?yp # c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])))" by (subst c_dom_sources, simp) ultimately have "c_dom D ?yp \ c_sources I (c_dom D) (D x) (?yp # c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])))" by (simp add: c_dom_def) hence "ipurge_tr I (c_dom D) (c_dom D ?yp) (c_ipurge I (c_dom D) (D x) (c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])))) = c_ipurge I (c_dom D) (D x) (c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])))" by (rule c_ipurge_tr_ipurge) hence "(?yp # c_tr step out ?s (c_ipurge I D (D x) (ys @ [x])), ?X) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 xs)" using C by (simp add: c_tr_ipurge) hence "(c_tr step out s\<^sub>0 xs @ ?yp # c_tr step out ?s (c_ipurge I D (D x) ys @ [x]), ?X) \ c_failures step out s\<^sub>0" (is "(?xps, _) \ _") by (simp add: c_futures_failures c_ipurge_append_1) moreover have "?xps \ []" by simp ultimately have "snd (last ?xps) = out (foldl step s\<^sub>0 (butlast (map fst ?xps))) (last (map fst ?xps))" by (rule c_failures_last) hence "snd (last ?xps) = out (foldl step (foldl step s\<^sub>0 (xs @ [y])) (c_ipurge I D (D x) ys)) x" by (simp add: c_tr_map butlast_append) moreover have "snd (last ?xps) = out (foldl step (foldl step s\<^sub>0 xs) (c_ipurge I D (D x) (y # ys))) x" using D by simp ultimately show ?thesis using A by simp qed qed theorem secure_implies_c_secure: assumes S: "secure (c_process step out s\<^sub>0) I (c_dom D)" shows "c_secure step out s\<^sub>0 I D" proof (simp add: c_secure_def, (rule allI)+) fix x xs have "out (foldl step (foldl step s\<^sub>0 []) xs) x = out (foldl step (foldl step s\<^sub>0 []) (c_ipurge I D (D x) xs)) x" using S by (rule secure_implies_c_secure_aux) thus "out (foldl step s\<^sub>0 xs) x = out (foldl step s\<^sub>0 (c_ipurge I D (D x) xs)) x" by simp qed lemma c_secure_futures_1: assumes R: "refl I" and S: "c_secure step out s\<^sub>0 I D" shows "(yps @ [yp], Y) \ futures (c_process step out s\<^sub>0) xps \ (yps, {x \ Y. (c_dom D yp, c_dom D x) \ I}) \ futures (c_process step out s\<^sub>0) xps" proof (simp add: c_futures_failures) let ?zs = "map fst (xps @ yps)" let ?y = "fst yp" assume A: "(xps @ yps @ [yp], Y) \ c_failures step out s\<^sub>0" hence "((xps @ yps) @ [yp], Y) \ failures (c_process step out s\<^sub>0)" by (simp add: c_failures_failures) hence "(xps @ yps, {}) \ failures (c_process step out s\<^sub>0)" by (rule process_rule_2_failures) hence "(xps @ yps, {}) \ c_failures step out s\<^sub>0" by (simp add: c_failures_failures) hence B: "xps @ yps = c_tr step out s\<^sub>0 ?zs" by (rule c_failures_tr) have "Y \ {(x, p). p \ out (foldl step s\<^sub>0 (map fst (xps @ yps @ [yp]))) x}" using A by (rule c_failures_ref) hence C: "Y \ {(x, p). p \ out (foldl step s\<^sub>0 (?zs @ [?y])) x}" (is "_ \ ?Y'") by simp have "(xps @ yps, {(x, p). p \ out (foldl step s\<^sub>0 ?zs) x}) \ c_failures step out s\<^sub>0" (is "(_, ?X') \ _") by (subst B, rule c_tr_failures) moreover have "{x \ Y. (c_dom D yp, c_dom D x) \ I} \ ?X'" (is "?X \ _") proof (rule subsetI, simp add: split_paired_all c_dom_def del: map_append, erule conjE) fix x p assume "(x, p) \ Y" with C have "(x, p) \ ?Y'" .. hence "p \ out (foldl step s\<^sub>0 (?zs @ [?y])) x" by simp moreover have "out (foldl step s\<^sub>0 (?zs @ [?y])) x = out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?zs @ [?y]))) x" using S by (simp add: c_secure_def) ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?zs @ [?y]))) x" by simp moreover assume "(D ?y, D x) \ I" with R have "c_ipurge I D (D x) (?zs @ [?y]) = c_ipurge I D (D x) ?zs" by (rule c_ipurge_append_2) ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) ?zs)) x" by simp moreover have "out (foldl step s\<^sub>0 (c_ipurge I D (D x) ?zs)) x = out (foldl step s\<^sub>0 ?zs) x" using S by (simp add: c_secure_def) ultimately show "p \ out (foldl step s\<^sub>0 ?zs) x" by simp qed ultimately show "(xps @ yps, ?X) \ c_failures step out s\<^sub>0" by (rule R2) qed lemma c_secure_implies_secure_aux_1 [rule_format]: assumes R: "refl I" and S: "c_secure step out s\<^sub>0 I D" shows "(yp # yps, Y) \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) yps, ipurge_ref I (c_dom D) (c_dom D yp) yps Y) \ futures (c_process step out s\<^sub>0) xps" proof (induction yps arbitrary: Y rule: length_induct, rule impI) fix yps Y assume A: "\yps'. length yps' < length yps \ (\Y'. (yp # yps', Y') \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) yps', ipurge_ref I (c_dom D) (c_dom D yp) yps' Y') \ futures (c_process step out s\<^sub>0) xps)" and B: "(yp # yps, Y) \ futures (c_process step out s\<^sub>0) xps" show "(ipurge_tr I (c_dom D) (c_dom D yp) yps, ipurge_ref I (c_dom D) (c_dom D yp) yps Y) \ futures (c_process step out s\<^sub>0) xps" proof (cases yps, simp add: ipurge_ref_def) case Nil hence "([] @ [yp], Y) \ futures (c_process step out s\<^sub>0) xps" using B by simp with R and S show "([], {x \ Y. (c_dom D yp, c_dom D x) \ I}) \ futures (c_process step out s\<^sub>0) xps" by (rule c_secure_futures_1) next case Cons have "\wps wp. yps = wps @ [wp]" by (rule rev_cases [of yps], simp_all add: Cons) then obtain wps and wp where C: "yps = wps @ [wp]" by blast have B': "((yp # wps) @ [wp], Y) \ futures (c_process step out s\<^sub>0) xps" using B and C by simp show ?thesis proof (simp only: C, cases "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])") let ?Y' = "{x \ Y. (c_dom D wp, c_dom D x) \ I}" have "length wps < length yps \ (\Y'. (yp # wps, Y') \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps Y') \ futures (c_process step out s\<^sub>0) xps)" using A .. moreover have "length wps < length yps" using C by simp ultimately have "\Y'. (yp # wps, Y') \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps Y') \ futures (c_process step out s\<^sub>0) xps" .. hence "(yp # wps, ?Y') \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps ?Y') \ futures (c_process step out s\<^sub>0) xps" .. moreover have "(yp # wps, ?Y') \ futures (c_process step out s\<^sub>0) xps" using R and S and B' by (rule c_secure_futures_1) ultimately have "(ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps ?Y') \ futures (c_process step out s\<^sub>0) xps" .. moreover assume D: "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])" hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = ipurge_tr I (c_dom D) (c_dom D yp) wps" by simp moreover have "ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Y = ipurge_ref I (c_dom D) (c_dom D yp) wps ?Y'" using D by (rule ipurge_ref_eq) ultimately show "(ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Y) \ futures (c_process step out s\<^sub>0) xps" by simp next let ?xs = "map fst xps" let ?y = "fst yp" let ?ws = "map fst wps" let ?w = "fst wp" let ?s = "foldl step s\<^sub>0 ?xs" have "(xps @ yp # wps @ [wp], Y) \ failures (c_process step out s\<^sub>0)" using B' by (simp add: c_futures_failures c_failures_failures) hence "(xps, {}) \ failures (c_process step out s\<^sub>0)" by (rule process_rule_2_failures) hence "(xps, {}) \ c_failures step out s\<^sub>0" by (simp add: c_failures_failures) hence X: "xps = c_tr step out s\<^sub>0 ?xs" by (rule c_failures_tr) have W: "(yp # wps, {}) \ futures (c_process step out s\<^sub>0) xps" using B' by (rule process_rule_2_futures) hence "yp # wps = c_tr step out ?s (map fst (yp # wps))" by (rule c_futures_tr) hence W': "yp # wps = c_tr step out ?s (?y # ?ws)" by simp assume D: "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])" hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = ipurge_tr I (c_dom D) (c_dom D yp) (yp # wps) @ [wp]" using R by (simp add: ipurge_tr_cons_same) hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = ipurge_tr I (c_dom D) (c_dom D yp) (c_tr step out ?s (?y # ?ws)) @ [wp]" using W' by simp also have "\ = c_tr step out ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws)) @ [wp]" proof (simp, rule c_tr_ipurge_tr) fix n show "\W. (ipurge_tr I (c_dom D) (c_dom D yp) (c_tr step out ?s (take n (?y # ?ws))), W) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" proof (cases n, simp_all add: c_tr_hd_tl) have "(c_tr step out ?s [], {(x, p). p \ out (foldl step ?s []) x}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" by (rule c_tr_futures) hence "([], {(x, p). p \ out ?s x}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" by simp thus "\W. ([], W) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" .. next case (Suc m) let ?wps' = "c_tr step out (step ?s ?y) (take m ?ws)" have "length ?wps' < length yps \ (\Y'. (yp # ?wps', Y') \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ipurge_ref I (c_dom D) (c_dom D yp) ?wps' Y') \ futures (c_process step out s\<^sub>0) xps)" using A .. moreover have "length ?wps' < length yps" using C by (simp add: c_tr_length) ultimately have "\Y'. (yp # ?wps', Y') \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ipurge_ref I (c_dom D) (c_dom D yp) ?wps' Y') \ futures (c_process step out s\<^sub>0) xps" .. hence "(yp # ?wps', {}) \ futures (c_process step out s\<^sub>0) xps \ (ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ipurge_ref I (c_dom D) (c_dom D yp) ?wps' {}) \ futures (c_process step out s\<^sub>0) xps" (is "_ \ (_, ?W') \ _") .. moreover have E: "yp # wps = (?y, out ?s ?y) # c_tr step out (step ?s ?y) (take m ?ws @ drop m ?ws)" using W' by (simp add: c_tr_hd_tl) hence F: "yp = (?y, out ?s ?y)" by simp hence "yp # wps = yp # ?wps' @ c_tr step out (foldl step (step ?s ?y) (take m ?ws)) (drop m ?ws)" using E by (simp only: c_tr_append) hence "((yp # ?wps') @ c_tr step out (foldl step (step ?s ?y) (take m ?ws)) (drop m ?ws), {}) \ futures (c_process step out s\<^sub>0) xps" using W by simp hence "(yp # ?wps', {}) \ futures (c_process step out s\<^sub>0) xps" by (rule process_rule_2_futures) ultimately have "(ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ?W') \ futures (c_process step out s\<^sub>0) xps" .. moreover have "ipurge_tr I (c_dom D) (c_dom D yp) ?wps' = ipurge_tr I (c_dom D) (c_dom D yp) ((?y, out ?s ?y) # ?wps')" using R and F by (simp add: ipurge_tr_cons_same) ultimately have "(ipurge_tr I (c_dom D) (c_dom D yp) ((?y, out ?s ?y) # ?wps'), ?W') \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" using X by simp thus "\W. (ipurge_tr I (c_dom D) (c_dom D yp) ((?y, out ?s ?y) # ?wps'), W) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" by (rule_tac x = ?W' in exI) qed qed finally have E: "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws)) @ [wp]" . have "(xps @ yp # wps @ [wp], Y) \ c_failures step out s\<^sub>0" (is "(?xps', _) \ _") using B' by (simp add: c_futures_failures) moreover have "?xps' \ []" by simp ultimately have "snd (last ?xps') = out (foldl step s\<^sub>0 (butlast (map fst ?xps'))) (last (map fst ?xps'))" by (rule c_failures_last) hence "snd wp = out (foldl step s\<^sub>0 (?xs @ ?y # ?ws)) ?w" by (simp add: butlast_append) hence "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ ?y # ?ws))) ?w" using S by (simp add: c_secure_def) moreover have F: "D ?w \ sinks I D (c_dom D yp) (?ws @ [?w])" using D by (simp only: c_dom_sinks, simp add: c_dom_def) have "\ (\v \ sinks I D (c_dom D yp) (?y # ?ws). (v, D ?w) \ I)" proof (rule notI, simp add: c_dom_def sinks_cons_same R, erule disjE) assume "(D ?y, D ?w) \ I" hence "D ?w \ sinks I D (c_dom D yp) (?ws @ [?w])" by (simp add: c_dom_def) thus False using F by contradiction next assume "\v \ sinks I D (D ?y) ?ws. (v, D ?w) \ I" hence "D ?w \ sinks I D (c_dom D yp) (?ws @ [?w])" by (simp add: c_dom_def) thus False using F by contradiction qed ultimately have "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ ipurge_tr I D (c_dom D yp) (?y # ?ws)))) ?w" using R by (simp add: c_ipurge_ipurge_tr) hence "snd wp = out (foldl step s\<^sub>0 (?xs @ ipurge_tr I D (c_dom D yp) (?y # ?ws))) ?w" using S by (simp add: c_secure_def) hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws)) @ [(?w, out (foldl step ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws))) ?w)]" using E by (cases wp, simp) hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws)) @ c_tr step out (foldl step ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws))) [?w]" by (simp add: c_tr_singleton) hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws) @ [?w])" by (simp add: c_tr_append) moreover have "(c_tr step out ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws) @ [?w]), {(x, p). p \ out (foldl step ?s (ipurge_tr I D (c_dom D yp) (?y # ?ws) @ [?w])) x}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" (is "(_, ?Y') \ _") by (rule c_tr_futures) ultimately have "(xps @ ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ?Y') \ c_failures step out s\<^sub>0" using X by (simp add: c_futures_failures) moreover have "ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Y \ ?Y'" proof (rule subsetI, simp add: split_paired_all ipurge_ref_def c_dom_def del: sinks.simps, (erule conjE)+) fix x p assume G: "\v \ sinks I (c_dom D) (D ?y) (wps @ [wp]). (v, D x) \ I" and H: "(D ?y, D x) \ I" have "(xps @ yp # wps @ [wp], Y) \ c_failures step out s\<^sub>0" using B' by (simp add: c_futures_failures) hence "Y \ {(x', p'). p' \ out (foldl step s\<^sub>0 (map fst (xps @ yp # wps @ [wp]))) x'}" by (rule c_failures_ref) hence "Y \ {(x', p'). p' \ out (foldl step s\<^sub>0 (?xs @ ?y # ?ws @ [?w])) x'}" by simp moreover assume "(x, p) \ Y" ultimately have "(x, p) \ {(x', p'). p' \ out (foldl step s\<^sub>0 (?xs @ ?y # ?ws @ [?w])) x'}" .. hence "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ ?y # ?ws @ [?w]))) x" using S by (simp add: c_secure_def) moreover have "\ (\v \ sinks I D (D ?y) (?y # ?ws @ [?w]). (v, D x) \ I)" proof assume "\v \ sinks I D (D ?y) (?y # ?ws @ [?w]). (v, D x) \ I" then obtain v where A: "v \ sinks I D (D ?y) (?y # ?ws @ [?w])" and B: "(v, D x) \ I" .. have "v = D ?y \ v \ sinks I D (D ?y) (?ws @ [?w])" using R and A by (simp add: sinks_cons_same) moreover { assume "v = D ?y" hence "(D ?y, D x) \ I" using B by simp hence False using H by contradiction } moreover { assume "v \ sinks I D (D ?y) (?ws @ [?w])" hence "v \ sinks I (c_dom D) (D ?y) (wps @ [wp])" by (simp only: c_dom_sinks, simp) with G have "(v, D x) \ I" .. hence False using B by contradiction } ultimately show False by blast qed ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ ipurge_tr I D (D ?y) (?y # ?ws @ [?w])))) x" using R by (simp add: c_ipurge_ipurge_tr) hence "p \ out (foldl step s\<^sub>0 (?xs @ ipurge_tr I D (D ?y) (?ws @ [?w]))) x" using R and S by (simp add: c_secure_def ipurge_tr_cons_same) hence "p \ out (foldl step s\<^sub>0 (?xs @ ipurge_tr I D (D ?y) ?ws @ [?w])) x" using F by (simp add: c_dom_def) thus "p \ out (step (foldl step ?s (ipurge_tr I D (D ?y) (?y # ?ws))) ?w) x" using R by (simp add: ipurge_tr_cons_same) qed ultimately have "(xps @ ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Y) \ c_failures step out s\<^sub>0" by (rule R2) thus "(ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Y) \ futures (c_process step out s\<^sub>0) xps" by (simp add: c_futures_failures) qed qed qed lemma c_secure_futures_2: assumes R: "refl I" and S: "c_secure step out s\<^sub>0 I D" shows "(yps @ [yp], A) \ futures (c_process step out s\<^sub>0) xps \ (yps, Y) \ futures (c_process step out s\<^sub>0) xps \ (yps @ [yp], {x \ Y. (c_dom D yp, c_dom D x) \ I}) \ futures (c_process step out s\<^sub>0) xps" proof (simp add: c_futures_failures) let ?zs = "map fst (xps @ yps)" let ?y = "fst yp" assume "(xps @ yps @ [yp], A) \ c_failures step out s\<^sub>0" hence "xps @ yps @ [yp] = c_tr step out s\<^sub>0 (map fst (xps @ yps @ [yp]))" by (rule c_failures_tr) hence A: "xps @ yps @ [yp] = c_tr step out s\<^sub>0 (?zs @ [?y])" by simp assume "(xps @ yps, Y) \ c_failures step out s\<^sub>0" hence B: "Y \ {(x, p). p \ out (foldl step s\<^sub>0 ?zs) x}" (is "_ \ ?Y'") by (rule c_failures_ref) have "(xps @ yps @ [yp], {(x, p). p \ out (foldl step s\<^sub>0 (?zs @ [?y])) x}) \ c_failures step out s\<^sub>0" (is "(_, ?X') \ _") by (subst A, rule c_tr_failures) moreover have "{x \ Y. (c_dom D yp, c_dom D x) \ I} \ ?X'" (is "?X \ _") proof (rule subsetI, simp add: split_paired_all c_dom_def del: map_append foldl_append, erule conjE) fix x p assume "(x, p) \ Y" with B have "(x, p) \ ?Y'" .. hence "p \ out (foldl step s\<^sub>0 ?zs) x" by simp moreover have "out (foldl step s\<^sub>0 ?zs) x = out (foldl step s\<^sub>0 (c_ipurge I D (D x) ?zs)) x" using S by (simp add: c_secure_def) ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) ?zs)) x" by simp moreover assume "(D ?y, D x) \ I" with R have "c_ipurge I D (D x) (?zs @ [?y]) = c_ipurge I D (D x) ?zs" by (rule c_ipurge_append_2) ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?zs @ [?y]))) x" by simp moreover have "out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?zs @ [?y]))) x = out (foldl step s\<^sub>0 (?zs @ [?y])) x" using S by (simp add: c_secure_def) ultimately show "p \ out (foldl step s\<^sub>0 (?zs @ [?y])) x" by simp qed ultimately show "(xps @ yps @ [yp], ?X) \ c_failures step out s\<^sub>0" by (rule R2) qed lemma c_secure_ipurge_tr: assumes R: "refl I" and S: "c_secure step out s\<^sub>0 I D" shows "ipurge_tr I (c_dom D) (D x) (c_tr step out (step (foldl step s\<^sub>0 xs) x) ys) = ipurge_tr I (c_dom D) (D x) (c_tr step out (foldl step s\<^sub>0 xs) ys)" proof (induction ys rule: rev_induct, simp, simp only: c_tr.simps) let ?s = "foldl step s\<^sub>0 xs" fix ys y assume A: "ipurge_tr I (c_dom D) (D x) (c_tr step out (step ?s x) ys) = ipurge_tr I (c_dom D) (D x) (c_tr step out ?s ys)" show "ipurge_tr I (c_dom D) (D x) (c_tr step out (step ?s x) ys @ [(y, out (foldl step (step ?s x) ys) y)]) = ipurge_tr I (c_dom D) (D x) (c_tr step out ?s ys @ [(y, out (foldl step ?s ys) y)])" (is "_ (_ @ [?yp']) = _ (_ @ [?yp])") proof (cases "D y \ sinks I D (D x) (ys @ [y])") assume D: "D y \ sinks I D (D x) (ys @ [y])" hence "c_dom D ?yp' \ sinks I (c_dom D) (D x) (c_tr step out (step ?s x) ys @ [?yp'])" using D by (simp only: c_dom_sinks, simp add: c_dom_def c_tr_map) hence "ipurge_tr I (c_dom D) (D x) (c_tr step out (step ?s x) ys @ [?yp']) = ipurge_tr I (c_dom D) (D x) (c_tr step out (step ?s x) ys)" by simp moreover have "c_dom D ?yp \ sinks I (c_dom D) (D x) (c_tr step out ?s ys @ [?yp])" using D by (simp only: c_dom_sinks, simp add: c_dom_def c_tr_map) hence "ipurge_tr I (c_dom D) (D x) (c_tr step out ?s ys @ [?yp]) = ipurge_tr I (c_dom D) (D x) (c_tr step out ?s ys)" by simp ultimately show ?thesis using A by simp next assume D: "D y \ sinks I D (D x) (ys @ [y])" hence "c_dom D ?yp' \ sinks I (c_dom D) (D x) (c_tr step out (step ?s x) ys @ [?yp'])" using D by (simp only: c_dom_sinks, simp add: c_dom_def c_tr_map) hence "ipurge_tr I (c_dom D) (D x) (c_tr step out (step ?s x) ys @ [?yp']) = ipurge_tr I (c_dom D) (D x) (c_tr step out (step ?s x) ys) @ [?yp']" by simp moreover have "c_dom D ?yp \ sinks I (c_dom D) (D x) (c_tr step out ?s ys @ [?yp])" using D by (simp only: c_dom_sinks, simp add: c_dom_def c_tr_map) hence "ipurge_tr I (c_dom D) (D x) (c_tr step out ?s ys @ [?yp]) = ipurge_tr I (c_dom D) (D x) (c_tr step out ?s ys) @ [?yp]" by simp ultimately show ?thesis proof (simp add: A) have B: "\ (\v \ sinks I D (D x) ys. (v, D y) \ I)" proof assume "\v \ sinks I D (D x) ys. (v, D y) \ I" hence "D y \ sinks I D (D x) (ys @ [y])" by simp thus False using D by contradiction qed have C: "\ (\v \ sinks I D (D x) (x # ys). (v, D y) \ I)" proof (rule notI, simp add: sinks_cons_same R B) assume "(D x, D y) \ I" hence "D y \ sinks I D (D x) (ys @ [y])" by simp thus False using D by contradiction qed have "out (foldl step (step ?s x) ys) y = out (foldl step s\<^sub>0 (xs @ x # ys)) y" by simp also have "\ = out (foldl step s\<^sub>0 (c_ipurge I D (D y) (xs @ x # ys))) y" using S by (simp add: c_secure_def) also have "\ = out (foldl step s\<^sub>0 (c_ipurge I D (D y) (xs @ ipurge_tr I D (D x) (x # ys)))) y" using R and C by (simp add: c_ipurge_ipurge_tr) also have "\ = out (foldl step s\<^sub>0 (c_ipurge I D (D y) (xs @ ipurge_tr I D (D x) ys))) y" using R by (simp add: ipurge_tr_cons_same) also have "\ = out (foldl step s\<^sub>0 (c_ipurge I D (D y) (xs @ ys))) y" using R and B by (simp add: c_ipurge_ipurge_tr) also have "\ = out (foldl step s\<^sub>0 (xs @ ys)) y" using S by (simp add: c_secure_def) also have "\ = out (foldl step ?s ys) y" by simp finally show "out (foldl step (step ?s x) ys) y = out (foldl step ?s ys) y" . qed qed qed lemma c_secure_implies_secure_aux_2 [rule_format]: assumes R: "refl I" and S: "c_secure step out s\<^sub>0 I D" and Y: "(yp # yps, Y) \ futures (c_process step out s\<^sub>0) xps" shows "(zps, Z) \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) zps, ipurge_ref I (c_dom D) (c_dom D yp) zps Z) \ futures (c_process step out s\<^sub>0) xps" proof (induction zps arbitrary: Z rule: length_induct, rule impI) fix zps Z assume A: "\zps'. length zps' < length zps \ (\Z'. (zps', Z') \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) zps', ipurge_ref I (c_dom D) (c_dom D yp) zps' Z') \ futures (c_process step out s\<^sub>0) xps)" and B: "(zps, Z) \ futures (c_process step out s\<^sub>0) xps" show "(yp # ipurge_tr I (c_dom D) (c_dom D yp) zps, ipurge_ref I (c_dom D) (c_dom D yp) zps Z) \ futures (c_process step out s\<^sub>0) xps" proof (cases zps, simp add: ipurge_ref_def) case Nil hence C: "([], Z) \ futures (c_process step out s\<^sub>0) xps" using B by simp have "(([] @ [yp]) @ yps, Y) \ futures (c_process step out s\<^sub>0) xps" using Y by simp hence "([] @ [yp], {}) \ futures (c_process step out s\<^sub>0) xps" by (rule process_rule_2_futures) with R and S have "([] @ [yp], {x \ Z. (c_dom D yp, c_dom D x) \ I}) \ futures (c_process step out s\<^sub>0) xps" using C by (rule c_secure_futures_2) thus "([yp], {x \ Z. (c_dom D yp, c_dom D x) \ I}) \ futures (c_process step out s\<^sub>0) xps" by simp next case Cons have "\wps wp. zps = wps @ [wp]" by (rule rev_cases [of zps], simp_all add: Cons) then obtain wps and wp where C: "zps = wps @ [wp]" by blast have B': "(wps @ [wp], Z) \ futures (c_process step out s\<^sub>0) xps" using B and C by simp show ?thesis proof (simp only: C, cases "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])") let ?Z' = "{x \ Z. (c_dom D wp, c_dom D x) \ I}" have "length wps < length zps \ (\Z'. (wps, Z') \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps Z') \ futures (c_process step out s\<^sub>0) xps)" using A .. moreover have "length wps < length zps" using C by simp ultimately have "\Z'. (wps, Z') \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps Z') \ futures (c_process step out s\<^sub>0) xps" .. hence "(wps, ?Z') \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps ?Z') \ futures (c_process step out s\<^sub>0) xps" .. moreover have "(wps, ?Z') \ futures (c_process step out s\<^sub>0) xps" using R and S and B' by (rule c_secure_futures_1) ultimately have "(yp # ipurge_tr I (c_dom D) (c_dom D yp) wps, ipurge_ref I (c_dom D) (c_dom D yp) wps ?Z') \ futures (c_process step out s\<^sub>0) xps" .. moreover assume D: "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])" hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = ipurge_tr I (c_dom D) (c_dom D yp) wps" by simp moreover have "ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Z = ipurge_ref I (c_dom D) (c_dom D yp) wps ?Z'" using D by (rule ipurge_ref_eq) ultimately show "(yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Z) \ futures (c_process step out s\<^sub>0) xps" by simp next let ?xs = "map fst xps" let ?y = "fst yp" let ?ws = "map fst wps" let ?w = "fst wp" let ?s = "foldl step s\<^sub>0 ?xs" let ?s' = "foldl step s\<^sub>0 (?xs @ [?y])" have "((xps @ [yp]) @ yps, Y) \ failures (c_process step out s\<^sub>0)" using Y by (simp add: c_futures_failures c_failures_failures) hence "(xps @ [yp], {}) \ failures (c_process step out s\<^sub>0)" by (rule process_rule_2_failures) hence "(xps @ [yp], {}) \ c_failures step out s\<^sub>0" by (simp add: c_failures_failures) hence "xps @ [yp] = c_tr step out s\<^sub>0 (map fst (xps @ [yp]))" by (rule c_failures_tr) hence XY: "xps @ [yp] = c_tr step out s\<^sub>0 (?xs @ [?y])" by simp hence X: "xps = c_tr step out s\<^sub>0 ?xs" by simp have "([yp] @ yps, Y) \ futures (c_process step out s\<^sub>0) xps" using Y by simp hence "([yp], {}) \ futures (c_process step out s\<^sub>0) xps" by (rule process_rule_2_futures) hence "[yp] = c_tr step out ?s (map fst [yp])" by (rule c_futures_tr) hence Y': "[yp] = c_tr step out ?s ([?y])" by simp have W: "(wps, {}) \ futures (c_process step out s\<^sub>0) xps" using B' by (rule process_rule_2_futures) hence W': "wps = c_tr step out (foldl step s\<^sub>0 ?xs) ?ws" by (rule c_futures_tr) assume D: "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])" hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = ipurge_tr I (c_dom D) (c_dom D yp) wps @ [wp]" by simp hence "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = ipurge_tr I (c_dom D) (c_dom D yp) (c_tr step out (foldl step s\<^sub>0 ?xs) ?ws) @ [wp]" using W' by simp also have "\ = ipurge_tr I (c_dom D) (c_dom D yp) (c_tr step out ?s' ?ws) @ [wp]" using R and S by (simp add: c_secure_ipurge_tr c_dom_def) also have "\ = c_tr step out ?s' (ipurge_tr I D (c_dom D yp) ?ws) @ [wp]" proof (simp del: foldl_append, rule c_tr_ipurge_tr) fix n let ?wps' = "c_tr step out ?s (take n ?ws)" have "length ?wps' < length zps \ (\Z'. (?wps', Z') \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ipurge_ref I (c_dom D) (c_dom D yp) ?wps' Z') \ futures (c_process step out s\<^sub>0) xps)" using A .. moreover have "length ?wps' < length zps" using C by (simp add: c_tr_length) ultimately have "\Z'. (?wps', Z') \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ipurge_ref I (c_dom D) (c_dom D yp) ?wps' Z') \ futures (c_process step out s\<^sub>0) xps" .. hence "(?wps', {}) \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ipurge_ref I (c_dom D) (c_dom D yp) ?wps' {}) \ futures (c_process step out s\<^sub>0) xps" (is "_ \ (_, ?W') \ _") .. moreover have "wps = c_tr step out ?s (take n ?ws @ drop n ?ws)" using W' by simp hence "wps = ?wps' @ c_tr step out (foldl step ?s (take n ?ws)) (drop n ?ws)" by (simp only: c_tr_append) hence "(?wps' @ c_tr step out (foldl step ?s (take n ?ws)) (drop n ?ws), {}) \ futures (c_process step out s\<^sub>0) xps" using W by simp hence "(?wps', {}) \ futures (c_process step out s\<^sub>0) xps" by (rule process_rule_2_futures) ultimately have "(yp # ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ?W') \ futures (c_process step out s\<^sub>0) xps" .. hence "(c_tr step out s\<^sub>0 (?xs @ [?y]) @ ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ?W') \ c_failures step out s\<^sub>0" using XY by (simp add: c_futures_failures) hence "(ipurge_tr I (c_dom D) (c_dom D yp) ?wps', ?W') \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 (?xs @ [?y]))" by (simp add: c_futures_failures) hence "(ipurge_tr I (c_dom D) (c_dom D yp) (c_tr step out ?s' (take n ?ws)), ?W') \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 (?xs @ [?y]))" using R and S by (simp add: c_dom_def c_secure_ipurge_tr) thus "\W. (ipurge_tr I (c_dom D) (c_dom D yp) (c_tr step out ?s' (take n ?ws)), W) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 (?xs @ [?y]))" by (rule_tac x = ?W' in exI) qed finally have E: "ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s' (ipurge_tr I D (c_dom D yp) ?ws) @ [wp]" . have "(xps @ wps @ [wp], Z) \ c_failures step out s\<^sub>0" (is "(?xps', _) \ _") using B' by (simp add: c_futures_failures) moreover have "?xps' \ []" by simp ultimately have "snd (last ?xps') = out (foldl step s\<^sub>0 (butlast (map fst ?xps'))) (last (map fst ?xps'))" by (rule c_failures_last) hence "snd wp = out (foldl step s\<^sub>0 (?xs @ ?ws)) ?w" by (simp add: butlast_append) hence "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ ?ws))) ?w" using S by (simp add: c_secure_def) moreover have F: "D ?w \ sinks I D (c_dom D yp) (?ws @ [?w])" using D by (simp only: c_dom_sinks, simp add: c_dom_def) have G: "\ (\v \ sinks I D (c_dom D yp) ?ws. (v, D ?w) \ I)" proof assume "\v \ sinks I D (c_dom D yp) ?ws. (v, D ?w) \ I" hence "D ?w \ sinks I D (c_dom D yp) (?ws @ [?w])" by simp thus False using F by contradiction qed ultimately have "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ ipurge_tr I D (c_dom D yp) ?ws))) ?w" using R by (simp add: c_ipurge_ipurge_tr) hence "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ ipurge_tr I D (c_dom D yp) (?y # ?ws)))) ?w" using R by (simp add: c_dom_def ipurge_tr_cons_same) moreover have "\ (\v \ sinks I D (c_dom D yp) (?y # ?ws). (v, D ?w) \ I)" proof (rule notI, simp add: sinks_cons_same c_dom_def R G [simplified]) assume "(D ?y, D ?w) \ I" hence "D ?w \ sinks I D (c_dom D yp) (?ws @ [?w])" by (simp add: c_dom_def) thus False using F by contradiction qed ultimately have "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ [?y] @ ?ws))) ?w" using R by (simp add: c_ipurge_ipurge_tr) moreover have "c_ipurge I D (D ?w) ((?xs @ [?y]) @ ipurge_tr I D (c_dom D yp) ?ws) = c_ipurge I D (D ?w) ((?xs @ [?y]) @ ?ws)" using R and G by (rule c_ipurge_ipurge_tr) ultimately have "snd wp = out (foldl step s\<^sub>0 (c_ipurge I D (D ?w) (?xs @ [?y] @ ipurge_tr I D (c_dom D yp) ?ws))) ?w" by simp hence "snd wp = out (foldl step s\<^sub>0 (?xs @ [?y] @ ipurge_tr I D (c_dom D yp) ?ws)) ?w" using S by (simp add: c_secure_def) hence "yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s ([?y]) @ c_tr step out ?s' (ipurge_tr I D (c_dom D yp) ?ws) @ [(?w, out (foldl step ?s' (ipurge_tr I D (c_dom D yp) ?ws)) ?w)]" using Y' and E by (cases wp, simp) hence "yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s ([?y]) @ c_tr step out ?s' (ipurge_tr I D (c_dom D yp) ?ws) @ c_tr step out (foldl step ?s' (ipurge_tr I D (c_dom D yp) ?ws)) [?w]" by (simp add: c_tr_singleton) hence "yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s ([?y]) @ c_tr step out (foldl step ?s [?y]) (ipurge_tr I D (c_dom D yp) ?ws @ [?w])" by (simp add: c_tr_append) hence "yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]) = c_tr step out ?s ([?y] @ ipurge_tr I D (c_dom D yp) ?ws @ [?w])" by (simp only: c_tr_append) moreover have "(c_tr step out ?s (?y # ipurge_tr I D (c_dom D yp) ?ws @ [?w]), {(x, p). p \ out (foldl step ?s (?y # ipurge_tr I D (c_dom D yp) ?ws @ [?w])) x}) \ futures (c_process step out s\<^sub>0) (c_tr step out s\<^sub>0 ?xs)" (is "(_, ?Z') \ _") by (rule c_tr_futures) ultimately have "(xps @ yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ?Z') \ c_failures step out s\<^sub>0" using X by (simp add: c_futures_failures) moreover have "ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Z \ ?Z'" proof (rule subsetI, simp add: split_paired_all ipurge_ref_def c_dom_def del: sinks.simps foldl.simps, (erule conjE)+) fix x p assume H: "\v \ sinks I (c_dom D) (D ?y) (wps @ [wp]). (v, D x) \ I" and I: "(D ?y, D x) \ I" have "(xps @ wps @ [wp], Z) \ c_failures step out s\<^sub>0" using B' by (simp add: c_futures_failures) hence "Z \ {(x', p'). p' \ out (foldl step s\<^sub>0 (map fst (xps @ wps @ [wp]))) x'}" by (rule c_failures_ref) hence "Z \ {(x', p'). p' \ out (foldl step s\<^sub>0 (?xs @ ?ws @ [?w])) x'}" by simp moreover assume "(x, p) \ Z" ultimately have "(x, p) \ {(x', p'). p' \ out (foldl step s\<^sub>0 (?xs @ ?ws @ [?w])) x'}" .. hence "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ ?ws @ [?w]))) x" using S by (simp add: c_secure_def) moreover have J: "\ (\v \ sinks I D (D ?y) (?ws @ [?w]). (v, D x) \ I)" proof (rule notI, cases "(D ?y, D ?w) \ I \ (\v \ sinks I D (D ?y) ?ws. (v, D ?w) \ I)", simp_all only: sinks.simps if_True if_False) case True hence "c_dom D wp \ sinks I (c_dom D) (c_dom D yp) (wps @ [wp])" by (simp only: c_dom_sinks, simp add: c_dom_def) thus False using D by contradiction next assume "\v \ sinks I D (D ?y) ?ws. (v, D x) \ I" then obtain v where A: "v \ sinks I D (D ?y) ?ws" and B: "(v, D x) \ I" .. have "v \ sinks I (c_dom D) (D ?y) (wps @ [wp])" using A by (simp add: c_dom_sinks) with H have "(v, D x) \ I" .. thus False using B by contradiction qed ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ ipurge_tr I D (D ?y) (?ws @ [?w])))) x" using R by (simp add: c_ipurge_ipurge_tr del: ipurge_tr.simps) hence "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ ipurge_tr I D (D ?y) (?y # ?ws @ [?w])))) x" using R by (simp add: ipurge_tr_cons_same) moreover have "\ (\v \ sinks I D (D ?y) (?y # ?ws @ [?w]). (v, D x) \ I)" proof assume "\v \ sinks I D (D ?y) (?y # ?ws @ [?w]). (v, D x) \ I" then obtain v where A: "v \ sinks I D (D ?y) (?y # ?ws @ [?w])" and B: "(v, D x) \ I" .. have "v = D ?y \ v \ sinks I D (D ?y) (?ws @ [?w])" using R and A by (simp add: sinks_cons_same) moreover { assume "v = D ?y" hence "(D ?y, D x) \ I" using B by simp hence False using I by contradiction } moreover { assume "v \ sinks I D (D ?y) (?ws @ [?w])" with B have "\v \ sinks I D (D ?y) (?ws @ [?w]). (v, D x) \ I" .. hence False using J by contradiction } ultimately show False by blast qed ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ [?y] @ ?ws @ [?w]))) x" using R by (simp add: c_ipurge_ipurge_tr del: ipurge_tr.simps) moreover have "c_ipurge I D (D x) ((?xs @ [?y]) @ ipurge_tr I D (D ?y) (?ws @ [?w])) = c_ipurge I D (D x) ((?xs @ [?y]) @ ?ws @ [?w])" using R and J by (rule c_ipurge_ipurge_tr) ultimately have "p \ out (foldl step s\<^sub>0 (c_ipurge I D (D x) (?xs @ ?y # ipurge_tr I D (D ?y) (?ws @ [?w])))) x" by simp hence "p \ out (foldl step s\<^sub>0 (?xs @ ?y # ipurge_tr I D (D ?y) (?ws @ [?w]))) x" using S by (simp add: c_secure_def) thus "p \ out (foldl step ?s (?y # ipurge_tr I D (D ?y) ?ws @ [?w])) x" using F by (simp add: c_dom_def) qed ultimately have "(xps @ yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Z) \ c_failures step out s\<^sub>0" by (rule R2) thus "(yp # ipurge_tr I (c_dom D) (c_dom D yp) (wps @ [wp]), ipurge_ref I (c_dom D) (c_dom D yp) (wps @ [wp]) Z) \ futures (c_process step out s\<^sub>0) xps" by (simp add: c_futures_failures) qed qed qed theorem c_secure_implies_secure: assumes R: "refl I" and S: "c_secure step out s\<^sub>0 I D" shows "secure (c_process step out s\<^sub>0) I (c_dom D)" proof (simp only: secure_def, (rule allI)+, rule impI, erule conjE) fix xps yp yps Y zps Z assume Y: "(yp # yps, Y) \ futures (c_process step out s\<^sub>0) xps" and Z: "(zps, Z) \ futures (c_process step out s\<^sub>0) xps" show "(ipurge_tr I (c_dom D) (c_dom D yp) yps, ipurge_ref I (c_dom D) (c_dom D yp) yps Y) \ futures (c_process step out s\<^sub>0) xps \ (yp # ipurge_tr I (c_dom D) (c_dom D yp) zps, ipurge_ref I (c_dom D) (c_dom D yp) zps Z) \ futures (c_process step out s\<^sub>0) xps" (is "?P \ ?Q") proof show ?P using R and S and Y by (rule c_secure_implies_secure_aux_1) next show ?Q using R and S and Y and Z by (rule c_secure_implies_secure_aux_2) qed qed theorem secure_equals_c_secure: "refl I \ secure (c_process step out s\<^sub>0) I (c_dom D) = c_secure step out s\<^sub>0 I D" by (rule iffI, rule secure_implies_c_secure, assumption, rule c_secure_implies_secure) end diff --git a/thys/Ordinal_Partitions/Erdos_Milner.thy b/thys/Ordinal_Partitions/Erdos_Milner.thy --- a/thys/Ordinal_Partitions/Erdos_Milner.thy +++ b/thys/Ordinal_Partitions/Erdos_Milner.thy @@ -1,1239 +1,1241 @@ theory Erdos_Milner imports Partitions begin subsection \Erdős-Milner theorem\ text \P. Erdős and E. C. Milner, A Theorem in the Partition Calculus. Canadian Math. Bull. 15:4 (1972), 501-505. Corrigendum, Canadian Math. Bull. 17:2 (1974), 305.\ text \The paper defines strong types as satisfying the criteria below. It remarks that ordinals satisfy those criteria. Here is a (too complicated) proof.\ proposition strong_ordertype_eq: assumes D: "D \ elts \" and "Ord \" obtains L where "\(List.set L) = D" "\X. X \ List.set L \ indecomposable (tp X)" and "\M. \M \ D; \X. X \ List.set L \ tp (M \ X) \ tp X\ \ tp M = tp D" proof - define \ where "\ \ inv_into D (ordermap D VWF)" then have bij_\: "bij_betw \ (elts (tp D)) D" using D bij_betw_inv_into down ordermap_bij by blast have \_cancel_left: "\d. d \ D \ \ (ordermap D VWF d) = d" by (metis D \_def bij_betw_inv_into_left down ordermap_bij total_on_VWF wf_VWF) have \_cancel_right: "\\. \ \ elts (tp D) \ ordermap D VWF (\ \) = \" by (metis \_def f_inv_into_f ordermap_surj subsetD) have "small D" "D \ ON" using assms down elts_subset_ON [of \] by auto then have \_less_iff: "\\ \. \\ \ elts (tp D); \ \ elts (tp D)\ \ \ \ < \ \ \ \ < \" by (metis \_def inv_ordermap_VWF_mono_iff inv_ordermap_mono_eq less_V_def) obtain \s where "List.set \s \ ON" and "\_dec \s" and tpD_eq: "tp D = \_sum \s" using Ord_ordertype \_nf_exists by blast \ \Cantor normal form of the ordertype\ have ord [simp]: "Ord (\s ! k)" "Ord (\_sum (take k \s))" if "k < length \s" for k using that \list.set \s \ ON\ by (auto simp: dual_order.trans set_take_subset elim: ON_imp_Ord) define E where "E \ \k. lift (\_sum (take k \s)) (\\(\s!k))" define L where "L \ map (\k. \ ` (elts (E k))) [0..s]" have lengthL [simp]: "length L = length \s" by (simp add: L_def) have in_elts_E_less: "elts (E k') \ elts (E k)" if "k's" for k k' \ \The ordinals have been partitioned into disjoint intervals\ proof - have ord\: "Ord (\ \ \s ! k')" using that by auto from that id_take_nth_drop [of k' "take k \s"] obtain l where "take k \s = take k' \s @ (\s!k') # l" by (simp add: min_def) then show ?thesis using that unfolding E_def lift_def less_sets_def by (auto dest!: OrdmemD [OF ord\] intro: less_le_trans) qed have elts_E: "elts (E k) \ elts (\_sum \s)" if "k < length \s" for k proof - have "\ \ (\s!k) \ \_sum (drop k \s)" by (metis that Cons_nth_drop_Suc \_sum_Cons add_le_cancel_left0) then have "(+) (\_sum (take k \s)) ` elts (\ \ (\s!k)) \ elts (\_sum (take k \s) + \_sum (drop k \s))" by blast also have "\ = elts (\_sum \s)" using \_sum_take_drop by auto finally show ?thesis by (simp add: lift_def E_def) qed have \_sum_in_tpD: "\_sum (take k \s) + \ \ elts (tp D)" if "\ \ elts (\ \ \s!k)" "k < length \s" for \ k using elts_E lift_def that tpD_eq by (auto simp: E_def) have Ord_\: "Ord (\ (\_sum (take k \s) + \))" if "\ \ elts (\ \ \s!k)" "k < length \s" for \ k using \_sum_in_tpD \D \ ON\ bij_\ bij_betw_imp_surj_on that by fastforce define \ where "\ \ \k. ((\y. odiff y (\_sum (take k \s))) \ ordermap D VWF)" \ \mapping the segments of @{term D} into some power of @{term \}\ have bij_\: "bij_betw (\ k) (\ ` elts (E k)) (elts (\ \ (\s!k)))" if "k < length \s" for k using that by (auto simp: bij_betw_def \_def E_def inj_on_def lift_def image_comp \_sum_in_tpD \_cancel_right that) have \_iff: "\ k x < \ k y \ (x,y) \ VWF" if "x \ \ ` elts (E k)" "y \ \ ` elts (E k)" "k < length \s" for k x y using that by (auto simp: \_def E_def lift_def \_sum_in_tpD \_cancel_right Ord_\ \_less_iff) have tp_E_eq [simp]: "tp (\ ` elts (E k)) = \\(\s!k)" if k: "k < length \s" for k by (metis Ord_\ Ord_oexp \_iff bij_\ ord(1) ordertype_VWF_eq_iff replacement small_elts that) have tp_L_eq [simp]: "tp (L!k) = \\(\s!k)" if "k < length \s" for k by (simp add: L_def that) have UL_eq_D: "\ (list.set L) = D" proof (rule antisym) show "\ (list.set L) \ D" by (force simp: L_def tpD_eq bij_betw_apply [OF bij_\] dest: elts_E) show "D \ \ (list.set L)" proof fix \ assume "\ \ D" then have "ordermap D VWF \ \ elts (\_sum \s)" by (metis \small D\ ordermap_in_ordertype tpD_eq) then show "\ \ \ (list.set L)" using \\ \ D\ \_cancel_left in_elts_\_sum by (fastforce simp: L_def E_def image_iff lift_def) qed qed show thesis proof show "indecomposable (tp X)" if "X \ list.set L" for X using that by (auto simp: L_def indecomposable_\_power) next fix M assume "M \ D" and *: "\X. X \ list.set L \ tp X \ tp (M \ X)" show "tp M = tp D" proof (rule antisym) show "tp M \ tp D" by (simp add: \M \ D\ \small D\ ordertype_VWF_mono) define \ where "\ \ \X. inv_into (M \ X) (ordermap (M \ X) VWF)" \ \The bijection from an @{term \}-power into the appropriate segment of @{term M}\ have bij_\: "bij_betw (\ X) (elts (tp (M \ X))) (M \ X)" for X unfolding \_def by (meson \M \ D\ \small D\ bij_betw_inv_into inf_le1 ordermap_bij smaller_than_small total_on_VWF wf_VWF) have Ord_\: "Ord (\ X \)" if "\ \ elts (tp (M \ X))" for \ X using \D \ ON\ \M \ D\ bij_betw_apply [OF bij_\] that by blast have \_cancel_right: "\\ X. \ \ elts (tp (M \ X)) \ ordermap (M \ X) VWF (\ X \) = \" by (metis \_def f_inv_into_f ordermap_surj subsetD) have smM: "small (M \ X)" for X by (meson \M \ D\ \small D\ inf_le1 subset_iff_less_eq_V) have "\k < length \s. \ \ elts (E k)" if \: "\ \ elts (tp D)" for \ proof - obtain X where "X \ set L" and X: "\ \ \ X" by (metis UL_eq_D \ Union_iff \_def in_mono inv_into_into ordermap_surj) then have "\k < length \s. \ \ elts (E k) \ X = \ ` elts (E k)" apply (clarsimp simp: L_def) by (metis \ \_cancel_right elts_E in_mono tpD_eq) then show ?thesis by blast qed then obtain K where K: "\\. \ \ elts (tp D) \ K \ < length \s \ \ \ elts (E (K \))" by metis \ \The index from @{term "tp D"} to the appropriate segment number\ define \ where "\ \ \d. \ (L ! K (ordermap D VWF d)) (\ (K (ordermap D VWF d)) d)" show "tp D \ tp M" proof (rule ordertype_inc_le) show "small D" "small M" using \M \ D\ \small D\ subset_iff_less_eq_V by auto next fix d' d assume xy: "d' \ D" "d \ D" and "(d',d) \ VWF" then have "d' < d" using ON_imp_Ord \D \ ON\ by auto let ?\' = "ordermap D VWF d'" let ?\ = "ordermap D VWF d" have len': "K ?\' < length \s" and elts': "?\' \ elts (E (K ?\'))" and len: "K ?\ < length \s" and elts: "?\ \ elts (E (K ?\))" using K \d' \ D\ \d \ D\ by (auto simp: \small D\ ordermap_in_ordertype) have Ord_\L: "Ord (\ (L!k) (\ k d))" if "d \ \ ` elts (E k)" "k < length \s" for k d by (metis (mono_tags) "*" Ord_\ bij_\ bij_betw_apply lengthL nth_mem that tp_L_eq vsubsetD) have "?\' < ?\" by (simp add: \(d', d) \ VWF\ \small D\ ordermap_mono_less xy) then have "K ?\' \ K ?\" using elts' elts in_elts_E_less by (meson leI len' less_asym less_sets_def) then consider (less) "K ?\' < K ?\" | (equal) "K ?\' = K ?\" by linarith then have "\ (L ! K ?\') (\ (K ?\') d') < \ (L ! K ?\) (\ (K ?\) d)" proof cases case less obtain \: "\ (L ! K ?\') (\ (K ?\') d') \ M \ L ! K ?\'" "\ (L ! K ?\) (\ (K ?\) d) \ M \ L ! K ?\" using elts' elts len' len * [THEN vsubsetD] by (metis lengthL \_cancel_left bij_\ bij_\ bij_betw_imp_surj_on imageI nth_mem tp_L_eq xy) then have "ordermap D VWF (\ (L ! K ?\') (\ (K ?\') d')) \ elts (E (K ?\'))" "ordermap D VWF (\ (L ! K ?\) (\ (K ?\) d)) \ elts (E (K ?\))" using len' len elts_E tpD_eq by (fastforce simp: L_def \_cancel_right)+ then have "ordermap D VWF (\ (L ! K ?\') (\ (K ?\') d')) < ordermap D VWF (\ (L ! K ?\) (\ (K ?\) d))" using in_elts_E_less len less by (meson less_sets_def) moreover have "\ (L ! K ?\') (\ (K ?\') d') \ D" "\ (L ! K ?\) (\ (K ?\) d) \ D" using \M \ D\ \ by auto ultimately show ?thesis by (metis \small D\ \_cancel_left \_less_iff ordermap_in_ordertype) next case equal have \_less: "\X \ \. \\ < \; \ \ elts (tp (M \ X)); \ \ elts (tp (M \ X))\ \ \ X \ < \ X \" by (metis \D \ ON\ \M \ D\ \_def dual_order.trans inv_ordermap_VWF_strict_mono_iff le_infI1 smM) have "\ (K ?\) d' < \ (K ?\) d" by (metis equal \(d', d) \ VWF\ \_cancel_left \_iff elts elts' imageI len xy) then show ?thesis unfolding equal by (metis * [THEN vsubsetD] lengthL \_cancel_left \_less bij_\ bij_betw_imp_surj_on elts elts' image_eqI len local.equal nth_mem tp_L_eq xy) qed moreover have "Ord (\ (L ! K ?\') (\ (K ?\') d'))" "Ord (\ (L ! K ?\) (\ (K ?\) d))" using Ord_\L \_cancel_left elts len elts' len' xy by fastforce+ ultimately show "(\ d', \ d) \ VWF" by (simp add: \_def) next show "\ ` D \ M" proof (clarsimp simp: \_def) fix d assume "d \ D" let ?\ = "ordermap D VWF d" have len: "K ?\ < length \s" and elts: "?\ \ elts (E (K ?\))" using K \d \ D\ by (auto simp: \small D\ ordermap_in_ordertype) have "\ (K ?\) d \ elts (tp (L! (K ?\)))" using bij_\ [OF len] \d \ D\ by (metis \_cancel_left bij_betw_apply elts imageI len tp_L_eq) then show "\ (L ! K (ordermap D VWF d)) (\ (K (ordermap D VWF d)) d) \ M" by (metis "*" lengthL Int_iff bij_\ bij_betw_apply len nth_mem vsubsetD) qed qed auto qed qed (simp add: UL_eq_D) qed text \The ``remark'' of Erdős and E. C. Milner, Canad. Math. Bull. Vol. 17 (2), 1974\ proposition indecomposable_imp_Ex_less_sets: assumes indec: "indecomposable \" and "\ \ \" and A: "tp A = \" "small A" "A \ ON" and "x \ A" and A1: "tp A1 = \" "A1 \ A" obtains A2 where "tp A2 = \" "A2 \ A1" "{x} \ A2" proof - have "Ord \" using indec indecomposable_imp_Ord by blast have "Limit \" by (meson \_gt1 assms indec indecomposable_imp_Limit less_le_trans) define \ where "\ \ inv_into A (ordermap A VWF)" then have bij_\: "bij_betw \ (elts \) A" using A bij_betw_inv_into down ordermap_bij by blast have bij_om: "bij_betw (ordermap A VWF) A (elts \)" using A down ordermap_bij by blast define \ where "\ \ ordermap A VWF x" have \: "\ \ elts \" unfolding \_def using A \x \ A\ down by auto then have "Ord \" using Ord_in_Ord \Ord \\ by blast define B where "B \ \ ` (elts (succ \))" have B: "{y \ A. ordermap A VWF y \ ordermap A VWF x} \ B" apply (clarsimp simp add: B_def \_def image_iff \_def) by (metis Ord_linear Ord_ordermap OrdmemD bij_betw_inv_into_left bij_om leD) show thesis proof have "small A1" by (meson \small A\ A1 smaller_than_small) then have "tp (A1 - B) \ tp A1" by (simp add: ordertype_VWF_mono) moreover have "tp (A1 - B) \ \" proof - have "\ (\ \ tp B)" unfolding B_def proof (subst ordertype_VWF_inc_eq) show "elts (succ \) \ ON" by (auto simp: \_def ordertype_VWF_inc_eq intro: Ord_in_Ord) have sub: "elts (succ \) \ elts \" using Ord_trans \ \Ord \\ by auto then show "\ ` elts (succ \) \ ON" using \A \ ON\ bij_\ bij_betw_imp_surj_on by blast have "succ \ \ elts \" using \ Limit_def \Limit \\ by blast with A sub show "\ u < \ v" if "u \ elts (succ \)" and "v \ elts (succ \)" and "u < v" for u v by (metis \_def inv_ordermap_VWF_strict_mono_iff subsetD that) show "\ \ \ tp (elts (succ \))" by (metis Limit_def Ord_succ \ \Limit \\ \Ord \\ mem_not_refl ordertype_eq_Ord vsubsetD) qed auto then show ?thesis using indecomposable_ordertype_ge [OF indec, of A1 B] \small A1\ A1 by (auto simp: B_def) qed ultimately show "tp (A1 - B) = \" using A1 by blast have "{x} \ (A - B)" using assms B apply (clarsimp simp: less_sets_def \_def subset_iff) by (metis Ord_not_le VWF_iff_Ord_less less_V_def order_refl ordermap_mono_less trans_VWF wf_VWF) with \A1 \ A\ show "{x} \ (A1 - B)" by auto qed auto qed text \the main theorem, from which they derive the headline result\ theorem Erdos_Milner_aux: assumes part: "partn_lst_VWF \ [k, \] 2" and indec: "indecomposable \" and "k > 1" "Ord \" and \: "\ \ elts \1" shows "partn_lst_VWF (\*\) [ord_of_nat (2*k), min \ (\*\)] 2" proof (cases "\\1 \ \=0") case True have "Ord \" using indec indecomposable_def by blast show ?thesis proof (cases "\=0") case True then show ?thesis by (simp add: partn_lst_triv0 [where i=1]) next case False then consider (0) "\=0" | (1) "\=1" by (metis Ord_0 OrdmemD True \Ord \\ mem_0_Ord one_V_def order.antisym succ_le_iff) then show ?thesis proof cases case 0 with part show ?thesis by (force simp add: partn_lst_def nsets_empty_iff less_Suc_eq) next case 1 then obtain "Ord \" by (meson ON_imp_Ord Ord_\1 True \ elts_subset_ON) then obtain i where "i < Suc (Suc 0)" "[ord_of_nat k, \] ! i \ \" using partn_lst_VWF_nontriv [OF part] 1 by auto then have "\ \ 1" using \\=1\ \k > 1\ by (fastforce simp: less_Suc_eq) then have "min \ (\*\) \ 1" by (metis Ord_1 Ord_\ Ord_linear_le Ord_mult \Ord \\ min_def order_trans) then show ?thesis using False by (auto simp: True \Ord \\ \\\0\ \\=1\ intro!: partn_lst_triv1 [where i=1]) qed qed next case False then have "\ \ \" by (meson Ord_1 Ord_not_less indec indecomposable_def indecomposable_imp_Limit omega_le_Limit) have "0 \ elts \" "\ \ 0" using False Ord_\1 Ord_in_Ord \ mem_0_Ord by blast+ show ?thesis unfolding partn_lst_def proof clarsimp fix f assume "f \ [elts (\*\)]\<^bsup>2\<^esup> \ {.. [elts (\*\)]\<^bsup>2\<^esup> \ {..<2::nat}" by (simp add: eval_nat_numeral) obtain ord [iff]: "Ord \" "Ord \" "Ord (\*\)" using Ord_\1 Ord_in_Ord \ indec indecomposable_imp_Ord Ord_mult by blast have *: False if i [rule_format]: "\H. tp H = ord_of_nat (2*k) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {0}" and ii [rule_format]: "\H. tp H = \ \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {1}" and iii [rule_format]: "\H. tp H = (\*\) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {1}" proof - have Ak0: "\X \ [A]\<^bsup>k\<^esup>. f ` [X]\<^bsup>2\<^esup> \ {0}" \\remark (8) about @{term"A \ S"}\ if A_\\: "A \ elts (\*\)" and ot: "tp A \ \" for A proof - let ?g = "inv_into A (ordermap A VWF)" have "small A" using down that by auto then have inj_g: "inj_on ?g (elts \)" by (meson inj_on_inv_into less_eq_V_def ordermap_surj ot subset_trans) have om_A_less: "\x y. \x \ A; y \ A; x < y\ \ ordermap A VWF x < ordermap A VWF y" using ord by (meson A_\\ Ord_in_Ord VWF_iff_Ord_less \small A\ in_mono ordermap_mono_less trans_VWF wf_VWF) have \_sub: "elts \ \ ordermap A VWF ` A" by (metis \small A\ elts_of_set less_eq_V_def ordertype_def ot replacement) have g: "?g \ elts \ \ elts (\*\)" by (meson A_\\ Pi_I' \_sub inv_into_into subset_eq) then have fg: "f \ (\X. ?g ` X) \ [elts \]\<^bsup>2\<^esup> \ {..<2}" by (rule nsets_compose_image_funcset [OF f _ inj_g]) obtain i H where "i < 2" "H \ elts \" and ot_eq: "tp H = [k,\]!i" "(f \ (\X. ?g ` X)) ` (nsets H 2) \ {i}" using ii partn_lst_E [OF part fg] by (auto simp: eval_nat_numeral) then consider (0) "i=0" | (1) "i=1" by linarith then show ?thesis proof cases case 0 then have "f ` [inv_into A (ordermap A VWF) ` H]\<^bsup>2\<^esup> \ {0}" using ot_eq \H \ elts \\ \_sub by (auto simp: nsets_def [of _ k] inj_on_inv_into elim!: nset_image_obtains) moreover have "finite H \ card H = k" using 0 ot_eq \H \ elts \\ down by (simp add: finite_ordertype_eq_card) then have "inv_into A (ordermap A VWF) ` H \ [A]\<^bsup>k\<^esup>" using \H \ elts \\ \_sub by (auto simp: nsets_def [of _ k] card_image inj_on_inv_into inv_into_into) ultimately show ?thesis by blast next case 1 have gH: "?g ` H \ elts (\*\)" by (metis A_\\ \_sub \H \ elts \\ image_subsetI inv_into_into subset_eq) have g_less: "?g x < ?g y" if "x < y" "x \ elts \" "y \ elts \" for x y using Pi_mem [OF g] ord that by (meson A_\\ Ord_in_Ord Ord_not_le \small A\ dual_order.trans elts_subset_ON inv_ordermap_VWF_mono_le ot vsubsetD) have [simp]: "tp (?g ` H) = tp H" by (meson \H \ elts \\ ord down dual_order.trans elts_subset_ON gH g_less ordertype_VWF_inc_eq subsetD) show ?thesis using ii [of "?g ` H"] subset_inj_on [OF inj_g \H \ elts \\] ot_eq 1 by (auto simp: gH elim!: nset_image_obtains) qed qed define K where "K \ \i x. {y \ elts (\*\). x \ y \ f{x,y} = i}" have small_K: "small (K i x)" for i x by (simp add: K_def) define KI where "KI \ \i X. (\x\X. K i x)" have KI_disj_self: "X \ KI i X = {}" for i X by (auto simp: KI_def K_def) define M where "M \ \D \ x. {\::V. \ \ D \ tp (K 1 x \ \ \) \ \}" have M_sub_D: "M D \ x \ D" for D \ x by (auto simp: M_def) have small_M [simp]: "small (M D \ x)" if "small D" for D \ x by (simp add: M_def that) have 9: "tp {x \ A. tp (M D \ x) \ tp D} \ \" (is "ordertype ?AD _ \ \") if inD: "indecomposable (tp D)" and D: "D \ elts \" and A: "A \ elts (\*\)" and tpA: "tp A = \" and \: "\ \ D \ {X. X \ elts (\*\) \ tp X = \}" for D A \ \\remark (9), assuming an indecomposable order type\ proof (rule ccontr) define A' where "A' \ {x \ A. \ tp (M D \ x) \ tp D}" have small [iff]: "small A" "small D" using A D down by (auto simp: M_def) have small_\: "small (\ \)" if "\ \ D" for \ using that \ by (auto simp: Pi_iff subset_iff_less_eq_V) assume not_\_le: "\ \ \ tp {x \ A. tp (M D \ x) \ tp D}" moreover obtain "small A" "small A'" "A' \ A" and A'_sub: "A' \ elts (\*\)" using A'_def A down by auto moreover have "A' = A - ?AD" by (force simp: A'_def) ultimately have A'_ge: "tp A' \ \" by (metis (no_types, lifting) dual_order.refl indec indecomposable_ordertype_eq mem_Collect_eq subsetI tpA) obtain X where "X \ A'" "finite X" "card X = k" and fX0: "f ` [X]\<^bsup>2\<^esup> \ {0}" using Ak0 [OF A'_sub A'_ge] by (auto simp: nsets_def [of _ k]) then have \: "\ tp (M D \ x) \ tp D" if "x \ X" for x using that by (auto simp: A'_def) obtain x where "x \ X" using \card X = k\ \k>1\ by fastforce have "\ D \ (\ x\X. M D \ x)" proof assume not: "D \ (\x\X. M D \ x)" have "\X\M D \ ` X. tp D \ tp X" proof (rule indecomposable_ordertype_finite_ge [OF inD]) show "M D \ ` X \ {}" using A'_def A'_ge not not_\_le by auto show "small (\ (M D \ ` X))" using \finite X\ by (simp add: finite_imp_small) qed (use \finite X\ not in auto) then show False by (simp add: \) qed then obtain \ where "\ \ D" and \: "\ \ (\ x\X. M D \ x)" by blast define \ where "\ \ {KI 0 X \ \ \, \x\X. K 1 x \ \ \, X \ \ \}" have \\: "X \ elts (\*\)" "\ \ \ elts (\*\)" using A'_sub \X \ A'\ \ \\ \ D\ by auto then have "KI 0 X \ (\x\X. K 1 x) \ X = elts (\*\)" using \x \ X\ f by (force simp: K_def KI_def Pi_iff less_2_cases_iff) with \\ have \\_\: "finite \" "\ \ \ \\" by (auto simp: \_def) then have "\ tp (K 1 x \ \ \) \ \" if "x \ X" for x using that \\ \ D\ \ \k > 1\ \card X = k\ by (fastforce simp: M_def) moreover have sm_K1: "small (\x\X. K 1 x \ \ \)" by (meson Finite_V Int_lower2 \\ \ D\ \finite X\ small_\ small_UN smaller_than_small) ultimately have not1: "\ tp (\x\X. K 1 x \ \ \) \ \" using \finite X\ \x \ X\ indecomposable_ordertype_finite_ge [OF indec, of "(\x. K 1 x \ \ \) ` X"] by blast moreover have "\ tp (X \ \ \) \ \" using \finite X\ \\ \ \\ by (meson finite_Int mem_not_refl ordertype_VWF_\ vsubsetD) moreover have "\ \ tp (\ \)" using \ \\ \ D\ small_\ by fastforce+ moreover have "small (\ \)" using \\ \ D\ small_\ by (fastforce simp: \_def intro: smaller_than_small sm_K1) ultimately have K0\_ge: "tp (KI 0 X \ \ \) \ \" using indecomposable_ordertype_finite_ge [OF indec \\_\] by (auto simp: \_def) have \\: "\ \ \ elts (\*\)" "tp (\ \) = \" using \\ \ D\ \ by blast+ then obtain Y where Ysub: "Y \ KI 0 X \ \ \" and "finite Y" "card Y = k" and fY0: "f ` [Y]\<^bsup>2\<^esup> \ {0}" using Ak0 [OF _ K0\_ge] by (auto simp: nsets_def [of _ k]) have \: "X \ Y = {}" using Ysub KI_disj_self by blast then have "card (X \ Y) = 2*k" by (simp add: \card X = k\ \card Y = k\ \finite X\ \finite Y\ card_Un_disjoint) moreover have "X \ Y \ elts (\*\)" using A'_sub \X \ A'\ \\(1) \Y \ KI 0 X \ \ \\ by auto moreover have "f ` [X \ Y]\<^bsup>2\<^esup> \ {0}" using fX0 fY0 Ysub by (auto simp: \ nsets_disjoint_2 image_Un image_UN KI_def K_def) ultimately show False using i \finite X\ \finite Y\ ordertype_VWF_finite_nat by auto qed have IX: "tp {x \ A. tp (M D \ x) \ tp D} \ \" if D: "D \ elts \" and A: "A \ elts (\*\)" and tpA: "tp A = \" and \: "\ \ D \ {X. X \ elts (\*\) \ tp X = \}" for D A \ \\remark (9) for any order type\ proof - obtain L where UL: "\(List.set L) \ D" and indL: "\X. X \ List.set L \ indecomposable (tp X)" and eqL: "\M. \M \ D; \X. X \ List.set L \ tp (M \ X) \ tp X\ \ tp M = tp D" using ord by (metis strong_ordertype_eq D order_refl) obtain A'' where A'': "A'' \ A" "tp A'' \ \" and "\x X. \x \ A''; X \ List.set L\ \ tp (M D \ x \ X) \ tp X" using UL indL proof (induction L arbitrary: thesis) case (Cons X L) then obtain A'' where A'': "A'' \ A" "tp A'' \ \" and "X \ D" and ge_X: "\x X. \x \ A''; X \ List.set L\ \ tp (M D \ x \ X) \ tp X" by auto then have tp_A'': "tp A'' = \" by (metis A antisym down ordertype_VWF_mono tpA) have ge_\: "tp {x \ A''. tp (M X \ x) \ tp X} \ \" by (rule 9) (use A A'' tp_A'' Cons.prems \D \ elts \\ \X \ D\ \ in auto) let ?A = "{x \ A''. tp (M D \ x \ X) \ tp X}" have M_eq: "M D \ x \ X = M X \ x" if "x \ A''" for x using that \X \ D\ by (auto simp: M_def) show thesis proof (rule Cons.prems) show "\ \ tp ?A" using ge_\ by (simp add: M_eq cong: conj_cong) show "tp Y \ tp (M D \ x \ Y)" if "x \ ?A" "Y \ list.set (X # L)" for x Y using that ge_X by force qed (use A'' in auto) qed (use tpA in auto) then have tp_M_ge: "tp (M D \ x) \ tp D" if "x \ A''" for x using eqL that by (auto simp: M_def) have "\ \ tp A''" by (simp add: A'') also have "\ \ tp {x \ A''. tp (M D \ x) \ tp D}" by (metis (mono_tags, lifting) tp_M_ge eq_iff mem_Collect_eq subsetI) also have "\ \ tp {x \ A. tp D \ tp (M D \ x)}" by (rule ordertype_mono) (use A'' A down in auto) finally show ?thesis . qed have IX': "tp {x \ A'. tp (K 1 x \ A) \ \} \ \" if A: "A \ elts (\*\)" "tp A = \" and A': "A' \ elts (\*\)" "tp A' = \" for A A' proof - have \: "\ \ tp (K 1 t \ A)" if "1 \ tp {\. \ = 0 \ \ \ tp (K 1 t \ A)}" for t using that by (metis Collect_empty_eq less_eq_V_0_iff ordertype_empty zero_neq_one) have "tp {x \ A'. 1 \ tp {\. \ = 0 \ \ \ tp (K 1 x \ A)}} \ tp {x \ A'. \ \ tp (K 1 x \ A)}" by (rule ordertype_mono) (use "\" A' in \auto simp: down\) then show ?thesis using IX [of "{0}" A' "\x. A"] that \0 \ elts \\ by (force simp: M_def) qed have 10: "\x0 \ A. \g \ elts \ \ elts \. strict_mono_on g (elts \) \ (\\ \ F. g \ = \) \ (\\ \ elts \. tp (K 1 x0 \ \ (g \)) \ \)" if F: "finite F" "F \ elts \" and A: "A \ elts (\*\)" "tp A = \" and \: "\ \ elts \ \ {X. X \ elts (\*\) \ tp X = \}" for F A \ proof - define p where "p \ card F" have "\ \ F" using that by auto then obtain \ :: "nat \ V" where bij\: "bij_betw \ {..p} (insert \ F)" and mon\: "strict_mono_on \ {..p}" using ZFC_Cardinals.ex_bij_betw_strict_mono_card [of "insert \ F"] elts_subset_ON \Ord \\ F by (simp add: p_def lessThan_Suc_atMost) blast have less_\_I: "\ k < \ l" if "k < l" "l \ p" for k l using mon\ that by (auto simp: strict_mono_on_def) then have less_\_D: "k < l" if "\ k < \ l" "k \ p" for k l by (metis less_asym linorder_neqE_nat that) have Ord_\: "Ord (\ k)" if "k \ p" for k by (metis (no_types, lifting) ON_imp_Ord atMost_iff insert_subset mem_Collect_eq order_trans \F \ elts \\ bij\ bij_betwE elts_subset_ON \Ord \\ that) have le_\0 [simp]: "\j. j \ p \ \ 0 \ \ j" by (metis eq_refl leI le_0_eq less_\_I less_imp_le) have le_\: "\ i \ \ (j - Suc 0)" if "i < j" "j \ p" for i j proof (cases i) case 0 then show ?thesis using le_\0 that by auto next case (Suc i') then show ?thesis by (metis (no_types, hide_lams) Suc_pred le_less less_Suc_eq less_Suc_eq_0_disj less_\_I not_less_eq that) qed have [simp]: "\ p = \" proof - obtain k where k: "\ k = \" "k \ p" by (meson atMost_iff bij\ bij_betw_iff_bijections insertI1) then have "k = p \ k < p" by linarith then show ?thesis using bij\ ord k that(2) by (metis OrdmemD atMost_iff bij_betw_iff_bijections insert_iff leD less_\_D order_refl subsetD) qed have F_imp_Ex: "\k < p. \ = \ k" if "\ \ F" for \ proof - obtain k where k: "k \ p" "\ = \ k" by (metis \\ \ F\ atMost_iff bij\ bij_betw_def imageE insert_iff) then show ?thesis using \\ \ F\ \\ p = \\ le_imp_less_or_eq that by blast qed have F_imp_ge: "\ \ \ 0" if "\ \ F" for \ using F_imp_Ex [OF that] by (metis dual_order.order_iff_strict le0 less_\_I) define D where "D \ \k. (if k=0 then {..<\ 0} else {\ (k-1)<..<\ k}) \ elts \" have D\: "D k \ elts \" for k by (auto simp: D_def) then have small_D [simp]: "small (D k)" for k by (meson down) have M_Int_D: "M (elts \) \ x \ D k = M (D k) \ x" if "k \ p" for x k using D\ by (auto simp: M_def) have \_le_if_D: "\ k \ \" if "\ \ D (Suc k)" for \ k using that by (simp add: D_def order.order_iff_strict) have mono_D: "D i \ D j" if "i < j" "j \ p" for i j proof (cases j) case (Suc j') with that show ?thesis apply (simp add: less_sets_def D_def Ball_def) by (metis One_nat_def diff_Suc_1 le_\ less_le_trans less_trans) qed (use that in auto) then have disjnt_DD: "disjnt (D i) (D j)" if "i \ j" "i \ p" "j \ p" for i j by (meson disjnt_sym less_linear less_sets_imp_disjnt that) have UN_D_eq: "(\l \ k. D l) = {..<\ k} \ (elts \ - F)" if "k \ p" for k using that proof (induction k) case 0 then show ?case by (auto simp: D_def F_imp_ge leD) next case (Suc k) have "D (Suc k) \ {..<\ k} \ (elts \ - F) = {..<\ (Suc k)} \ (elts \ - F)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using Suc.prems by (auto simp: D_def if_split_mem2 intro: less_\_I less_trans dest!: less_\_D F_imp_Ex) have "\x. \x < \ (Suc k); x \ elts \; x \ F; \ k \ x\ \ \ k < x" using Suc.prems \F \ elts \\ bij\ le_imp_less_or_eq by (fastforce simp: bij_betw_iff_bijections) then show "?rhs \ ?lhs" using Suc.prems by (auto simp: D_def Ord_not_less Ord_in_Ord [OF \Ord \\] Ord_\ if_split_mem2) qed then show ?case using Suc by (simp add: atMost_Suc) qed have \_decomp: "elts \ = F \ (\k \ p. D k)" using \F \ elts \\ OrdmemD [OF \Ord \\] by (auto simp: UN_D_eq) define \idx where "\idx \ \\. @k. \ \ D k \ k \ p" have \idx: "\ \ D (\idx \) \ \idx \ \ p" if "\ \ elts \ - F" for \ using that by (force simp: \idx_def \_decomp intro: someI_ex del: conjI) have any_imp_\idx: "k = \idx \" if "\ \ D k" "k \ p" for k \ proof (rule ccontr) assume non: "k \ \idx \" have "\ \ F" using that UN_D_eq by auto then show False using disjnt_DD [OF non] by (metis D\ Diff_iff \idx disjnt_iff subsetD that) qed have "\A'. A' \ A \ tp A' = \ \ (\x \ A'. F \ M (elts \) \ x)" using F proof induction case (insert \ F) then obtain A' where "A' \ A" and A': "A' \ elts (\*\)" "tp A' = \" and FN: "\x. x \ A' \ F \ M (elts \) \ x" using A(1) by auto define A'' where "A'' \ {x \ A'. \ \ tp (K 1 x \ \ \)}" have "\ \ elts \" "F \ elts \" using insert by auto note ordertype_eq_Ord [OF \Ord \\, simp] show ?case proof (intro exI conjI) show "A'' \ A" using \A' \ A\ by (auto simp: A''_def) show "tp A'' = \" proof (rule antisym) show "tp A'' \ \" using \A'' \ A\ down ordertype_VWF_mono A by blast have "\ \ \ elts (\*\)" "tp (\ \) = \" using \ \\ \ elts \\ by auto then show "\ \ tp A''" using IX' [OF _ _ A'] by (simp add: A''_def) qed show "\x\A''. insert \ F \ M (elts \) \ x" using A''_def FN M_def \\ \ elts \\ by blast qed qed (use A in auto) then obtain A' where A': "A' \ A" "tp A' = \" and FN: "\x. x \ A' \ F \ M (elts \) \ x" by metis have False if *: "\x0 g. \x0 \ A; g \ elts \ \ elts \; strict_mono_on g (elts \)\ \ (\\\F. g \ \ \) \ (\\\elts \. tp (K 1 x0 \ \ (g \)) < \)" proof - { fix x \ \construction of the monotone map @{term g} mentioned above\ assume "x \ A'" with A' have "x \ A" by blast have "\k. k \ p \ tp (M (D k) \ x) < tp (D k)" (is "?P") proof (rule ccontr) assume "\ ?P" then have le: "tp (D k) \ tp (M (D k) \ x)" if "k \ p" for k by (meson Ord_linear2 Ord_ordertype that) have "\f\D k \ M (D k) \ x. inj_on f (D k) \ (strict_mono_on f (D k))" if "k \ p" for k using le [OF that] that VWF_iff_Ord_less apply (clarsimp simp: ordertype_le_ordertype strict_mono_on_def) by (metis (full_types) D\ M_sub_D Ord_in_Ord PiE VWF_iff_Ord_less ord(2) subsetD) then obtain h where fun_h: "\k. k \ p \ h k \ D k \ M (D k) \ x" and inj_h: "\k. k \ p \ inj_on (h k) (D k)" and mono_h: "\k x y. k \ p \ strict_mono_on (h k) (D k)" by metis then have fun_hD: "\k. k \ p \ h k \ D k \ D k" by (auto simp: M_def) have h_increasing: "\ \ h k \" if "k \ p" and \: "\ \ D k" for k \ proof (rule Ord_mono_imp_increasing) show "h k \ D k \ D k" by (simp add: fun_hD that(1)) show "D k \ ON" using D\ elts_subset_ON ord(2) by blast qed (auto simp: that mono_h) define g where "g \ \\. if \ \ F then \ else h (\idx \) \" have [simp]: "g \ = \" if "\ \ F" for \ using that by (auto simp: g_def) have fun_g: "g \ elts \ \ elts \" proof (rule Pi_I) fix x assume "x \ elts \" then have "x \ D (\idx x)" "\idx x \ p" if "x \ F" using that by (auto simp: \idx) then show "g x \ elts \" using fun_h D\ M_sub_D \x \ elts \\ by (simp add: g_def) blast qed have h_in_D: "h (\idx \) \ \ D (\idx \)" if "\ \ F" "\ \ elts \" for \ using \idx fun_hD that by fastforce have 1: "\ k < h (\idx \) \" if "k < p" and \: "\ \ F" "\ \ elts \" and "\ k < \" for k \ using that h_in_D [OF \] \idx by (fastforce simp: D_def dest: h_increasing split: if_split_asm) moreover have 2: "h (\idx \) \ < \ k" if \: "\ \ F" "\ \ elts \" and "k < p" "\ < \ k" for \ k proof - have "\idx \ \ k" proof (rule ccontr) assume "\ \idx \ \ k" then have "k < \idx \" by linarith then show False using \_le_if_D \idx by (metis Diff_iff Suc_pred le0 leD le_\ le_less_trans \ \\ < \ k\) qed then show ?thesis using that h_in_D [OF \] by (smt (verit, best) Int_lower1 UN_D_eq UN_I atMost_iff lessThan_iff less_imp_le subset_eq) qed moreover have "h (\idx \) \ < h (\idx \) \" if \: "\ \ F" "\ \ elts \" and \: "\ \ F" "\ \ elts \" and "\ < \" for \ \ proof - have le: "\idx \ \ \idx \" if "\ (\idx \ - Suc 0) < h (\idx \) \" "h (\idx \) \ < \ (\idx \)" by (metis 2 that Diff_iff \idx \ \ \\ < \\ dual_order.strict_implies_order dual_order.strict_trans1 h_increasing leI le_\ less_asym) have "h 0 \ < h 0 \" if "\idx \ = 0" "\idx \ = 0" using that mono_h unfolding strict_mono_on_def by (metis Diff_iff \idx \ \ \\ < \\) moreover have "h 0 \ < h (\idx \) \" if "0 < \idx \" "h 0 \ < \ 0" and "\ (\idx \ - Suc 0) < h (\idx \) \" by (meson DiffI \idx \ le_\ le_less_trans less_le_not_le that) moreover have "\idx \ \ 0" if "0 < \idx \" "h 0 \ < \ 0" "\ (\idx \ - Suc 0) < h (\idx \) \" using le le_0_eq that by fastforce moreover have "h (\idx \) \ < h (\idx \) \" if "\ (\idx \ - Suc 0) < h (\idx \) \" "h (\idx \) \ < \ (\idx \)" "h (\idx \) \ < \ (\idx \)" "\ (\idx \ - Suc 0) < h (\idx \) \" using mono_h unfolding strict_mono_on_def by (metis le Diff_iff \idx \ \ \\ < \\ le_\ le_less le_less_trans that) ultimately show ?thesis using h_in_D [OF \] h_in_D [OF \] by (simp add: D_def split: if_split_asm) qed ultimately have sm_g: "strict_mono_on g (elts \)" by (auto simp: g_def strict_mono_on_def dest!: F_imp_Ex) have False if "\ \ elts \" and \: "tp (K 1 x \ \ (g \)) < \" for \ proof - have "F \ M (elts \) \ x" by (meson FN \x \ A'\) then have False if "tp (K (Suc 0) x \ \ \) < \" "\ \ F" using that by (auto simp: M_def) moreover have False if "tp (K (Suc 0) x \ \ (g \)) < \" "\ \ D k" "k \ p" "\ \ F" for k proof - have "h (\idx \) \ \ M (D (\idx \)) \ x" using fun_h \idx \\ \ elts \\ \\ \ F\ by auto then show False using that by (simp add: M_def g_def leD) qed ultimately show False using \\ \ elts \\ \ by (force simp: \_decomp) qed then show False using * [OF \x \ A\ fun_g sm_g] by auto qed then have "\l. l \ p \ tp (M (elts \) \ x \ D l) < tp (D l)" using M_Int_D by auto } then obtain l where lp: "\x. x \ A'\ l x \ p" and lless: "\x. x \ A'\ tp (M (elts \) \ x \ D (l x)) < tp (D (l x))" by metis obtain A'' L where "A'' \ A'" and A'': "A'' \ elts (\*\)" "tp A'' = \" and lL: "\x. x \ A'' \ l x = L" proof - have eq: "A' = (\i\p. {x \ A'. l x = i})" using lp by auto have "\X\(\i. {x \ A'. l x = i}) ` {..p}. \ \ tp X" proof (rule indecomposable_ordertype_finite_ge [OF indec]) show "small (\i\p. {x \ A'. l x = i})" by (metis A'(1) A(1) eq down smaller_than_small) qed (use A' eq in auto) then show thesis proof fix A'' assume A'': "A'' \ (\i. {x \ A'. l x = i}) ` {..p}" and "\ \ tp A''" then obtain L where L: "\x. x \ A'' \ l x = L" by auto have "A'' \ A'" using A'' by force then have "tp A'' \ tp A'" by (meson A' A down order_trans ordertype_VWF_mono) with \\ \ tp A''\ have "tp A'' = \" using A'(2) by auto moreover have "A'' \ elts (\*\)" using A' A \A'' \ A'\ by auto ultimately show thesis using L that [OF \A'' \ A'\] by blast qed qed have \D: "\ \ D L \ {X. X \ elts (\*\) \ tp X = \}" using \ D\ by blast have \: "\ \ tp {x \ A''. tp (D L) \ tp (M (D L) \ x)}" using IX [OF D\ A'' \D] by simp have "M (elts \) \ x \ D L = M (D L) \ x" for x using D\ by (auto simp: M_def) then have "tp (M (D L) \ x) < tp (D L)" if "x \ A''" for x using lless that \A'' \ A'\ lL by force then have [simp]: "{x \ A''. tp (D L) \ tp (M (D L) \ x)} = {}" using leD by blast show False using \ \\ \ \\ by simp qed then show ?thesis by (meson Ord_linear2 Ord_ordertype \Ord \\) qed let ?U = "UNIV :: nat set" define \ where "\ \ fst \ from_nat_into (elts \ \ ?U)" define q where "q \ to_nat_on (elts \ \ ?U)" have co_\U: "countable (elts \ \ ?U)" by (simp add: \ less_\1_imp_countable) moreover have "elts \ \ ?U \ {}" using \0 \ elts \\ by blast ultimately have "range (from_nat_into (elts \ \ ?U)) = (elts \ \ ?U)" by (metis range_from_nat_into) then have \_in_\ [simp]: "\ i \ elts \" for i by (metis SigmaE \_def comp_apply fst_conv range_eqI) then have Ord_\ [simp]: "Ord (\ i)" for i using Ord_in_Ord by blast have inf_\U: "infinite (elts \ \ ?U)" using \0 \ elts \\ finite_cartesian_productD2 by auto have 11 [simp]: "\ (q (\,n)) = \" if "\ \ elts \" for \ n by (simp add: \_def q_def that co_\U) have range_\ [simp]: "range \ = elts \" by (auto simp: image_iff) (metis 11) have [simp]: "KI i {} = UNIV" "KI i (insert a X) = K i a \ KI i X" for i a X by (auto simp: KI_def) define \ where "\ \ \n::nat. \\ x. (\\ \ elts \. \ \ \ elts (\*\) \ tp (\ \) = \) \ x ` {.. elts (\*\) \ (\\ \ elts \. \ \) \ KI 1 (x ` {.. strict_mono_sets (elts \) \" define \ where "\ \ \n::nat. \g \ \' xn. g \ elts \ \ elts \ \ strict_mono_on g (elts \) \ (\i\n. g (\ i) = \ i) \ (\\ \ elts \. \' \ \ K 1 xn \ \ (g \)) \ {xn} \ (\' (\ n)) \ xn \ \ (\ n)" let ?\0 = "\\. plus (\ * \) ` elts \" have base: "\ 0 ?\0 x" for x by (auto simp: \_def add_mult_less add_mult_less_add_mult ordertype_image_plus strict_mono_sets_def less_sets_def) have step: "Ex (\(g,\',xn). \ n g \ \' xn \ \ (Suc n) \' (x(n:=xn)))" if "\ n \ x" for n \ x proof - have \: "\\. \ \ elts \ \ \ \ \ elts (\*\) \ tp (\ \) = \" and x: "x ` {.. elts (\*\)" and sub: "\ (\ ` elts \) \ KI (Suc 0) (x ` {..) \" and \\: "\ ` {..n} \ elts \" and \sub: "\ (\ n) \ elts (\*\)" and \fun: "\ \ elts \ \ {X. X \ elts (\*\) \ tp X = \}" using that by (auto simp: \_def) then obtain xn g where xn: "xn \ \ (\ n)" and g: "g \ elts \ \ elts \" and sm_g: "strict_mono_on g (elts \)" and g_\: "\\ \ \`{..n}. g \ = \" and g_\: "\\ \ elts \. \ \ tp (K 1 xn \ \ (g \))" using 10 [OF _ \\ \sub _ \fun] by (auto simp: \) have tp1: "tp (K 1 xn \ \ (g \)) = \" if "\ \ elts \" for \ proof (rule antisym) have "tp (K 1 xn \ \ (g \)) \ tp (\ (g \))" by (metis Int_lower2 PiE \ down g ordertype_VWF_mono that) also have "\ = \" using \ g that by force finally show "tp (K 1 xn \ \ (g \)) \ \" . qed (use that g_\ in auto) have tp2: "tp (\ (\ n)) = \" by (auto simp: \) obtain "small (\ (\ n))" "\ (\ n) \ ON" by (meson \sub ord down elts_subset_ON subset_trans) then obtain A2 where A2: "tp A2 = \" "A2 \ K 1 xn \ \ (\ n)" "{xn} \ A2" using indecomposable_imp_Ex_less_sets [OF indec \\ \ \\ tp2] by (metis \_in_\ atMost_iff image_eqI inf_le2 le_refl xn tp1 g_\) then have A2_sub: "A2 \ \ (\ n)" by simp let ?\ = "\\. if \ = \ n then A2 else K 1 xn \ \ (g \)" have [simp]: "({.. {x. x \ n}) = ({.. (\x\elts \ \ {\. \ \ \ n}. \ (g x)) \ KI (Suc 0) (x ` {.. KI (Suc 0) (x ` {.. elts (\*\)" "xn \ elts (\*\)" using \sub sub A2 xn by fastforce+ moreover have "strict_mono_sets (elts \) ?\" using sm sm_g g g_\ A2_sub unfolding strict_mono_sets_def strict_mono_on_def less_sets_def Pi_iff subset_iff Ball_def Bex_def image_iff by (simp (no_asm_use) add: if_split_mem2) (smt order_refl) ultimately have "\ (Suc n) ?\ (x(n := xn))" using tp1 x A2 by (auto simp: \_def K_def) with A2 show ?thesis by (rule_tac x="(g,?\,xn)" in exI) (simp add: \_def g sm_g g_\ xn) qed define G where "G \ \n \ x. @(g,\',x'). \xn. \ n g \ \' xn \ x' = (x(n:=xn)) \ \ (Suc n) \' x'" have G\: "(\(g,\',x'). \ (Suc n) \' x') (G n \ x)" and G\: "(\(g,\',x'). \ n g \ \' (x' n)) (G n \ x)" if "\ n \ x" for n \ x using step [OF that] by (force simp: G_def dest: some_eq_imp)+ define H where "H \ rec_nat (id,?\0,undefined) (\n (g0,\,x0). G n \ x0)" have "(\(g,\,x). \ n \ x) (H n)" for n unfolding H_def by (induction n) (use G\ base in fastforce)+ then have H_imp_\: "\ n \ x" if "H n = (g,\,x)" for g \ x n by (metis case_prodD that) then have H_imp_\: "(\(g,\',x'). let (g0,\,x) = H n in \ n g \ \' (x' n)) (H (Suc n))" for n using G\ by (fastforce simp: H_def split: prod.split) define g where "g \ \n. (\(g,\,x). g) (H (Suc n))" have g: "g n \ elts \ \ elts \" and sm_g: "strict_mono_on (g n) (elts \)" and 13: "\i. i\n \ g n (\ i) = \ i" for n using H_imp_\ [of n] by (auto simp: g_def \_def) define \ where "\ \ \n. (\(g,\,x). \) (H n)" define x where "x \ \n. (\(g,\,x). x n) (H (Suc n))" have 14: "\ (Suc n) \ \ K 1 (x n) \ \ n (g n \)" if "\ \ elts \" for \ n using H_imp_\ [of n] that by (force simp: \_def \_def x_def g_def) then have x14: "\ (Suc n) \ \ \ n (g n \)" if "\ \ elts \" for \ n using that by blast have 15: "x n \ \ n (\ n)" and 16: "{x n} \ (\ (Suc n) (\ n))" for n using H_imp_\ [of n] by (force simp: \_def \_def x_def)+ have \_\\: "\ n \ \ elts (\*\)" if "\ \ elts \" for \ n using H_imp_\ [of n] that by (auto simp: \_def \_def split: prod.split) have 12: "strict_mono_sets (elts \) (\ n)" for n using H_imp_\ [of n] that by (auto simp: \_def \_def split: prod.split) have tp_\: "tp (\ n \) = \" if "\ \ elts \" for \ n using H_imp_\ [of n] that by (auto simp: \_def \_def split: prod.split) let ?Z = "range x" have S_dec: "\ (\ (m+k) ` elts \) \ \ (\ m ` elts \)" for k m by (induction k) (use 14 g in \fastforce+\) have "x n \ K 1 (x m)" if "m (\\ \ elts \. \ n \)" by (meson "15" UN_I \_in_\) also have "\ \ (\\ \ elts \. \ (Suc m) \)" using S_dec [of "Suc m"] less_iff_Suc_add that by auto also have "\ \ K 1 (x m)" using 14 by auto finally show ?thesis . qed then have "f{x m, x n} = 1" if "m2\<^esup> \ {1}" by (clarsimp simp: nsets_2_eq) (metis insert_commute less_linear) moreover have Z_sub: "?Z \ elts (\*\)" using "15" \_\\ \_in_\ by blast moreover have "tp ?Z \ \ * \" proof - define \ where "\ \ \i j x. wfrec (measure (\k. j-k)) (\\ k. if k (Suc k)) else x) i" have \: "\ i j x = (if i (Suc i) j x) else x)" for i j x by (simp add: \_def wfrec cut_apply) have 17: "\ k j (\ i) = \ i" if "i \ k" for i j k using wf_measure [of "\k. j-k"] that by (induction k rule: wf_induct_rule) (simp add: "13" \ le_imp_less_Suc) have \_in_\: "\ i j \ \ elts \" if "\ \ elts \" for i j \ using wf_measure [of "\k. j-k"] that proof (induction i rule: wf_induct_rule) case (less i) with g show ?case by (force simp: \ [of i]) qed then have \_fun: "\ i j \ elts \ \ elts \" for i j by simp have sm_\: "strict_mono_on (\ i j) (elts \)" for i j using wf_measure [of "\k. j-k"] proof (induction i rule: wf_induct_rule) case (less i) with sm_g show ?case by (auto simp: \ [of i] strict_mono_on_def \_in_\) qed have *: "\ j (\ j) \ \ i (\ i j (\ j))" if "i < j" for i j using wf_measure [of "\k. j-k"] that proof (induction i rule: wf_induct_rule) case (less i) then have "j - Suc i < j - i" by (metis (no_types) Suc_diff_Suc lessI) with less \_in_\ show ?case by (simp add: \ [of i]) (metis 17 Suc_lessI \_in_\ order_refl order_trans x14) qed have le_iff: "\ i j (\ j) \ \ i \ \ j \ \ i" for i j using sm_\ unfolding strict_mono_on_def by (metis "17" Ord_in_Ord Ord_linear2 \_in_\ leD le_refl less_V_def \Ord \\) then have less_iff: "\ i j (\ j) < \ i \ \ j < \ i" for i j by (metis (no_types, lifting) "17" \_in_\ less_V_def order_refl sm_\ strict_mono_on_def) have eq_iff: "\ i j (\ j) = \ i \ \ j = \ i" for i j by (metis eq_refl le_iff less_iff less_le) have \_if_ne: "\ m < \ n" if mn: "\ m (\ m) \ \ n (\ n)" "m \ n" for m n proof - have xmn: "x m < x n" using "15" less_setsD that(1) by blast have Ord\: "Ord (\ n m (\ m))" using Ord_in_Ord \_in_\ \_in_\ ord(2) by presburger have "\ \ m (\ m) \ \ n (\ n)" if "\ n = \ m" using "*" "15" eq_iff that unfolding less_sets_def by (metis in_mono less_irrefl not_less_iff_gr_or_eq) moreover have "\ n (\ n) \ \ m (\ m n (\ n)) \ \ m (\ m) \ \ n (\ n m (\ m))" using * mn by (meson antisym_conv3) then have False if "\ n < \ m" using strict_mono_setsD [OF 12] 15 xmn \_in_\ \_in_\ that by (smt (verit, best) Ord\ Ord_\ Ord_linear2 leD le_iff less_asym less_iff less_setsD subset_iff) ultimately show "\ m < \ n" by (meson that(1) Ord_\ Ord_linear_lt) qed have 18: "\ m (\ m) \ \ n (\ n) \ \ m < \ n" for m n proof (cases n m rule: linorder_cases) case less show ?thesis proof (intro iffI) assume "\ m < \ n" then have "\ n (\ n m (\ m)) \ \ n (\ n)" by (metis "12" \_in_\ \_in_\ eq_iff le_iff less_V_def strict_mono_sets_def) then show "\ m (\ m) \ \ n (\ n)" by (meson "*" less less_sets_weaken1) qed (use \_if_ne less in blast) next case equal with 15 show ?thesis by auto next case greater show ?thesis proof (intro iffI) assume "\ m < \ n" then have "\ m (\ m) \ (\ m (\ m n (\ n)))" by (meson 12 Ord_in_Ord Ord_linear2 \_in_\ \_in_\ le_iff leD ord(2) strict_mono_sets_def) then show "\ m (\ m) \ \ n (\ n)" by (meson "*" greater less_sets_weaken2) qed (use \_if_ne greater in blast) qed have \_increasing_\: "\ n (\ n) \ \ m (\ m)" if "m \ n" "\ m = \ n" for m n by (metis "*" "17" dual_order.order_iff_strict that) moreover have INF: "infinite {n. n \ m \ \ m = \ n}" for m proof - have "infinite (range (\n. q (\ m, n)))" unfolding q_def using to_nat_on_infinite [OF co_\U inf_\U] finite_image_iff by (simp add: finite_image_iff inj_on_def) moreover have "(range (\n. q (\ m, n))) \ {n. \ m = \ n}" using 11 [of "\ m"] by auto ultimately have "infinite {n. \ m = \ n}" using finite_subset by auto then have "infinite ({n. \ m = \ n} - {.. n" "\ p = \ n" "\ m = \ n" "n < p" with 16 [of n] show "x n < x p" by (metis "*" "15" "17" Suc_lessI insert_absorb insert_subset le_SucI less_sets_singleton1) qed then have inj_x: "inj_on x (?eqv m)" for m using strict_mono_on_imp_inj_on by blast define ZA where "ZA \ \m. ?Z \ \ m (\ m)" have small_ZA [simp]: "small (ZA m)" for m by (metis ZA_def inf_le1 small_image_nat smaller_than_small) have 19: "tp (ZA m) \ \" for m proof - have "x ` {n. m \ n \ \ m = \ n} \ ZA m" unfolding ZA_def using "15" \_increasing_\ by blast then have "infinite (ZA m)" using INF [of m] finite_image_iff [OF inj_x] by (meson finite_subset) then show ?thesis by (simp add: ordertype_infinite_ge_\) qed have "\f \ elts \ \ ZA m. strict_mono_on f (elts \)" for m proof - obtain Z where "Z \ ZA m" "tp Z = \" by (meson 19 Ord_\ le_ordertype_obtains_subset small_ZA) moreover have "ZA m \ ON" using Ord_in_Ord \_\\ \_in_\ unfolding ZA_def by blast ultimately show ?thesis by (metis strict_mono_on_ordertype Pi_mono small_ZA smaller_than_small subset_iff) qed then obtain \ where \: "\m. \ m \ elts \ \ ZA m" and sm_\: "\m. strict_mono_on (\ m) (elts \)" by metis have "Ex(\(m,\). \ \ elts \ \ \ = \ * \ + ord_of_nat m)" if "\ \ elts (\ * \)" for \ using that by (auto simp: mult [of \ \] lift_def elts_\) then obtain split where split: "\\. \ \ elts (\ * \) \ (\(m,\). \ \ elts \ \ \ = \ * \ + ord_of_nat m)(split \)" by meson have split_eq [simp]: "split (\ * \ + ord_of_nat m) = (m,\)" if "\ \ elts \" for \ m proof - have [simp]: "\ * \ + ord_of_nat m = \ * \ + ord_of_nat n \ \ = \ \ n = m" if "\ \ elts \" for \ n by (metis Ord_\ that Ord_mem_iff_less_TC mult_cancellation_lemma ord_of_nat_\ ord_of_nat_inject) show ?thesis using split [of "\*\ + m"] that by (auto simp: mult [of \ \] lift_def cong: conj_cong) qed define \ where "\ \ \\. (\(m,\). \ (q(\,0)) m)(split \)" have \_Pi: "\ \ elts (\ * \) \ (\m. ZA m)" using \ by (fastforce simp: \_def mult [of \ \] lift_def elts_\) moreover have "(\m. ZA m) \ ON" unfolding ZA_def using \_\\ \_in_\ elts_subset_ON by blast ultimately have Ord_\_Pi: "\ \ elts (\ * \) \ ON" by fastforce show "tp ?Z \ \ * \" proof - have \: "(\m. ZA m) = ?Z" using "15" by (force simp: ZA_def) moreover have "tp (elts (\ * \)) \ tp (\m. ZA m)" proof (rule ordertype_inc_le) show "\ ` elts (\ * \) \ (\m. ZA m)" using \_Pi by blast next fix u v assume x: "u \ elts (\ * \)" and y: "v \ elts (\ * \)" and "(u,v) \ VWF" then have "u Ord_in_Ord Ord_mult VWF_iff_Ord_less ord(2)) moreover obtain m \ n \ where ueq: "u = \ * \ + ord_of_nat m" and \: "\ \ elts \" and veq: "v = \ * \ + ord_of_nat n" and \: "\ \ elts \" using x y by (auto simp: mult [of \ \] lift_def elts_\) ultimately have "\ \ \" by (meson Ord_\ Ord_in_Ord Ord_linear2 \Ord \\ add_mult_less_add_mult less_asym ord_of_nat_\) consider (eq) "\ = \" | (lt) "\ < \" using \\ \ \\ le_neq_trans by blast then have "\ u < \ v" proof cases case eq then have "m < n" using ueq veq \u by simp then have "\ (q (\, 0)) m < \ (q (\, 0)) n" using sm_\ strict_mono_onD by blast then show ?thesis using eq ueq veq \ \m < n\ by (simp add: \_def) next case lt have "\ (q(\,0)) m \ \ (q(\,0)) (\(q(\,0)))" "\ (q (\,0)) n \ \ (q(\,0)) (\(q(\,0)))" using \ unfolding ZA_def by blast+ then show ?thesis using lt ueq veq \ \ 18 [of "q(\,0)" "q(\,0)"] by (simp add: \_def less_sets_def) qed then show "(\ u, \ v) \ VWF" using \_Pi by (metis Ord_\_Pi PiE VWF_iff_Ord_less x y mem_Collect_eq) qed (use \ in auto) ultimately show ?thesis by simp qed qed then obtain Z where "Z \ ?Z" "tp Z = \ * \" by (meson Ord_\ Ord_mult ord Z_sub down le_ordertype_obtains_subset) ultimately show False using iii [of Z] by (meson dual_order.trans image_mono nsets_mono) qed have False if 0: "\H. tp H = ord_of_nat (2*k) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {0}" and 1: "\H. tp H = min \ (\ * \) \ H \ elts (\*\) \ \ f ` [H]\<^bsup>2\<^esup> \ {1}" proof (cases "\*\ \ \") case True then have \: "\H'\H. tp H' = \ * \" if "tp H = \" "small H" for H by (metis Ord_\ Ord_\1 Ord_in_Ord Ord_mult \ le_ordertype_obtains_subset that) have [simp]: "min \ (\*\) = \*\" by (simp add: min_absorb2 that True) then show ?thesis using * [OF 0] 1 True by simp (meson \ down image_mono nsets_mono subset_trans) next case False then have \: "\H'\H. tp H' = \" if "tp H = \ * \" "small H" for H by (metis Ord_linear_le Ord_ordertype \Ord \\ le_ordertype_obtains_subset that) then have "\ \ \*\" by (meson Ord_\ Ord_\1 Ord_in_Ord Ord_linear_le Ord_mult \ \Ord \\ False) then have [simp]: "min \ (\*\) = \" by (simp add: min_absorb1) then show ?thesis using * [OF 0] 1 False by simp (meson \ down image_mono nsets_mono subset_trans) qed then show "\iH\elts (\*\). tp H = [ord_of_nat (2*k), min \ (\*\)] ! i \ f ` [H]\<^bsup>2\<^esup> \ {i}" by force qed qed theorem Erdos_Milner: assumes \: "\ \ elts \1" shows "partn_lst_VWF (\\(1 + \ * n)) [ord_of_nat (2^n), \\(1+\)] 2" proof (induction n) case 0 then show ?case using partn_lst_VWF_degenerate [of 1 2] by simp next case (Suc n) have "Ord \" using Ord_\1 Ord_in_Ord assms by blast have "1+\ \ \+1" by (simp add: \Ord \\ one_V_def plus_Ord_le) then have [simp]: "min (\ \ (1 + \)) (\ * \ \ \) = \ \ (1+\)" by (simp add: \Ord \\ oexp_add min_def) have ind: "indecomposable (\ \ (1 + \ * ord_of_nat n))" by (simp add: \Ord \\ indecomposable_\_power) show ?case proof (cases "n = 0") case True then show ?thesis using partn_lst_VWF_\_2 \Ord \\ one_V_def by auto next case False then have "Suc 0 < 2 ^ n" using less_2_cases not_less_eq by fastforce then have "partn_lst_VWF (\ \ (1 + \ * n) * \ \ \) [ord_of_nat (2 * 2 ^ n), \ \ (1 + \)] 2" using Erdos_Milner_aux [OF Suc ind, where \ = "\\\"] \Ord \\ \ by (auto simp: countable_oexp) then show ?thesis using \Ord \\ by (simp add: mult_succ mult.assoc oexp_add) qed qed corollary remark_3: "partn_lst_VWF (\\(Suc(4*k))) [4, \\(Suc(2*k))] 2" using Erdos_Milner [of "2*k" 2] apply (simp flip: ord_of_nat_mult ord_of_nat.simps) by (simp add: one_V_def) text \Theorem 3.2 of Jean A. Larson, ibid.\ corollary Theorem_3_2: fixes k n::nat shows "partn_lst_VWF (\\(n*k)) [\\n, ord_of_nat k] 2" proof (cases "n=0 \ k=0") case True then show ?thesis by (auto intro: partn_lst_triv0 [where i=1] partn_lst_triv1 [where i=0]) next case False then have "n > 0" "k > 0" by auto + from \k > 0\ less_exp [of \k - 1\] have \k \ 2 ^ (k - 1)\ + by (cases k) (simp_all add: less_eq_Suc_le) have PV: "partn_lst_VWF (\ \ (1 + ord_of_nat (n-1) * ord_of_nat (k-1))) [ord_of_nat (2 ^ (k-1)), \ \ (1 + ord_of_nat (n-1))] 2" using Erdos_Milner [of "ord_of_nat (n-1)" "k-1"] Ord_\1 Ord_mem_iff_lt less_imp_le by blast have "k+n \ Suc (k * n)" using False not0_implies_Suc by fastforce then have "1 + (n - 1) * (k - 1) \ n*k" using False by (auto simp: algebra_simps) then have "(1 + ord_of_nat (n - 1) * ord_of_nat (k - 1)) \ ord_of_nat(n*k)" by (metis (mono_tags, lifting) One_nat_def one_V_def ord_of_nat.simps ord_of_nat_add ord_of_nat_mono_iff ord_of_nat_mult) then have x: "\ \ (1 + ord_of_nat (n - 1) * ord_of_nat (k - 1)) \ \\(n*k)" by (simp add: oexp_mono_le) then have "partn_lst_VWF (\\(n*k)) [\ \ (1 + ord_of_nat (n-1)), ord_of_nat (2 ^ (k-1))] 2" by (metis PV partn_lst_two_swap Partitions.partn_lst_greater_resource less_eq_V_def) then have "partn_lst_VWF (\\(n*k)) [\ \ n, ord_of_nat (2 ^ (k-1))] 2" using ord_of_minus_1 [OF \n > 0\] by (simp add: one_V_def) then show ?thesis - using power_gt_expt [of 2 "k-1"] - by (force simp: less_Suc_eq intro: partn_lst_less) + using \k \ 2 ^ (k - 1)\ + by (auto elim!: partn_lst_less simp add: less_Suc_eq) qed end diff --git a/thys/Prime_Distribution_Elementary/More_Dirichlet_Misc.thy b/thys/Prime_Distribution_Elementary/More_Dirichlet_Misc.thy --- a/thys/Prime_Distribution_Elementary/More_Dirichlet_Misc.thy +++ b/thys/Prime_Distribution_Elementary/More_Dirichlet_Misc.thy @@ -1,431 +1,432 @@ (* File: More_Dirichlet_Misc.thy Author: Manuel Eberl, TU München Generalised Dirichlet products and Legendre's identity, and a weighted sum of the Möbius \ function *) section \Miscellaneous material\ theory More_Dirichlet_Misc imports Prime_Distribution_Elementary_Library Prime_Number_Theorem.Prime_Counting_Functions begin subsection \Generalised Dirichlet products\ (* TODO: Move to Dirichlet_Series *) definition dirichlet_prod' :: "(nat \ 'a :: comm_semiring_1) \ (real \ 'a) \ real \ 'a" where "dirichlet_prod' f g x = sum_upto (\m. f m * g (x / real m)) x" lemma dirichlet_prod'_one_left: "dirichlet_prod' (\n. if n = 1 then 1 else 0) f x = (if x \ 1 then f x else 0)" proof - have "dirichlet_prod' (\n. if n = 1 then 1 else 0) f x = (\i | 0 < i \ real i \ x. (if i = Suc 0 then 1 else 0) * f (x / real i))" by (simp add: dirichlet_prod'_def sum_upto_def) also have "\ = (\i\(if x \ 1 then {1::nat} else {}). f x)" by (intro sum.mono_neutral_cong_right) (auto split: if_splits) also have "\ = (if x \ 1 then f x else 0)" by simp finally show ?thesis . qed lemma dirichlet_prod'_cong: assumes "\n. n > 0 \ real n \ x \ f n = f' n" assumes "\y. y \ 1 \ y \ x \ g y = g' y" assumes "x = x'" shows "dirichlet_prod' f g x = dirichlet_prod' f' g' x'" unfolding dirichlet_prod'_def by (intro sum_upto_cong' assms, (subst assms | simp add: assms field_simps)+) (* 2.21 *) lemma dirichlet_prod'_assoc: "dirichlet_prod' f (\y. dirichlet_prod' g h y) x = dirichlet_prod' (dirichlet_prod f g) h x" proof - have "dirichlet_prod' f (\y. dirichlet_prod' g h y) x = (\m | m > 0 \ real m \ x. \n | n > 0 \ real n \ x / m. f m * g n * h (x / (m * n)))" by (simp add: algebra_simps dirichlet_prod'_def dirichlet_prod_def sum_upto_def sum_distrib_left sum_distrib_right) also have "\ = (\(m,n)\(SIGMA m:{m. m > 0 \ real m \ x}. {n. n > 0 \ real n \ x / m}). f m * g n * h (x / (m * n)))" by (subst sum.Sigma) auto also have "\ = (\(mn, m)\(SIGMA mn:{mn. mn > 0 \ real mn \ x}. {m. m dvd mn}). f m * g (mn div m) * h (x / mn))" by (rule sum.reindex_bij_witness[of _ "\(mn, m). (m, mn div m)" "\(m, n). (m * n, m)"]) (auto simp: case_prod_unfold field_simps dest: dvd_imp_le) also have "\ = dirichlet_prod' (dirichlet_prod f g) h x" by (subst sum.Sigma [symmetric]) (simp_all add: dirichlet_prod'_def dirichlet_prod_def sum_upto_def algebra_simps sum_distrib_left sum_distrib_right) finally show ?thesis . qed (* 2.22 *) lemma dirichlet_prod'_inversion1: assumes "\x\1. g x = dirichlet_prod' a f x" "x \ 1" "dirichlet_prod a ainv = (\n. if n = 1 then 1 else 0)" shows "f x = dirichlet_prod' ainv g x" proof - have "dirichlet_prod' ainv g x = dirichlet_prod' ainv (dirichlet_prod' a f) x" using assms by (intro dirichlet_prod'_cong) auto also have "\ = dirichlet_prod' (\n. if n = 1 then 1 else 0) f x" using assms by (simp add: dirichlet_prod'_assoc dirichlet_prod_commutes) also have "\ = f x" using assms by (subst dirichlet_prod'_one_left) auto finally show ?thesis .. qed lemma dirichlet_prod'_inversion2: assumes "\x\1. f x = dirichlet_prod' ainv g x" "x \ 1" "dirichlet_prod a ainv = (\n. if n = 1 then 1 else 0)" shows "g x = dirichlet_prod' a f x" proof - have "dirichlet_prod' a f x = dirichlet_prod' a (dirichlet_prod' ainv g) x" using assms by (intro dirichlet_prod'_cong) auto also have "\ = dirichlet_prod' (\n. if n = 1 then 1 else 0) g x" using assms by (simp add: dirichlet_prod'_assoc dirichlet_prod_commutes) also have "\ = g x" using assms by (subst dirichlet_prod'_one_left) auto finally show ?thesis .. qed lemma dirichlet_prod'_inversion: assumes "dirichlet_prod a ainv = (\n. if n = 1 then 1 else 0)" shows "(\x\1. g x = dirichlet_prod' a f x) \ (\x\1. f x = dirichlet_prod' ainv g x)" using dirichlet_prod'_inversion1[of g a f _ ainv] dirichlet_prod'_inversion2[of f ainv g _ a] assms by blast lemma dirichlet_prod'_inversion': assumes "a 1 * y = 1" defines "ainv \ dirichlet_inverse a y" shows "(\x\1. g x = dirichlet_prod' a f x) \ (\x\1. f x = dirichlet_prod' ainv g x)" unfolding ainv_def by (intro dirichlet_prod'_inversion dirichlet_prod_inverse assms) (* 3.11 *) lemma dirichlet_prod'_floor_conv_sum_upto: "dirichlet_prod' f (\x. real_of_int (floor x)) x = sum_upto (\n. sum_upto f (x / n)) x" proof - have [simp]: "sum_upto (\_. 1 :: real) x = real (nat \x\)" for x by (simp add: sum_upto_altdef) show ?thesis using sum_upto_dirichlet_prod[of "\n. 1::real" f] sum_upto_dirichlet_prod[of f "\n. 1::real"] by (simp add: dirichlet_prod'_def dirichlet_prod_commutes) qed lemma (in completely_multiplicative_function) dirichlet_prod_self: "dirichlet_prod f f n = f n * of_nat (divisor_count n)" proof (cases "n = 0") case False have "dirichlet_prod f f n = (\(r, d) | r * d = n. f (r * d))" by (simp add: dirichlet_prod_altdef2 mult) also have "\ = (\(r, d) | r * d = n. f n)" by (intro sum.cong) auto also have "\ = f n * of_nat (card {(r, d). r * d = n})" by (simp add: mult.commute) also have "bij_betw fst {(r, d). r * d = n} {r. r dvd n}" by (rule bij_betwI[of _ _ _ "\r. (r, n div r)"]) (use False in auto) hence "card {(r, d). r * d = n} = card {r. r dvd n}" by (rule bij_betw_same_card) also have "\ = divisor_count n" by (simp add: divisor_count_def) finally show ?thesis . qed auto lemma completely_multiplicative_imp_moebius_mu_inverse: fixes f :: "nat \ 'a :: {comm_ring_1}" assumes "completely_multiplicative_function f" shows "dirichlet_prod f (\n. moebius_mu n * f n) n = (if n = 1 then 1 else 0)" proof - interpret completely_multiplicative_function f by fact have [simp]: "fds f \ 0" by (auto simp: fds_eq_iff) have "dirichlet_prod f (\n. moebius_mu n * f n) n = (\(r, d) | r * d = n. moebius_mu r * f (r * d))" by (subst dirichlet_prod_commutes) (simp add: fds_eq_iff fds_nth_mult fds_nth_fds dirichlet_prod_altdef2 mult_ac mult) also have "\ = (\(r, d) | r * d = n. moebius_mu r * f n)" by (intro sum.cong) auto also have "\ = dirichlet_prod moebius_mu (\_. 1) n * f n" by (simp add: dirichlet_prod_altdef2 sum_distrib_right case_prod_unfold mult) also have "dirichlet_prod moebius_mu (\_. 1) n = fds_nth (fds moebius_mu * fds_zeta) n" by (simp add: fds_nth_mult) also have "fds moebius_mu * fds_zeta = 1" by (simp add: mult_ac fds_zeta_times_moebius_mu) also have "fds_nth 1 n * f n = fds_nth 1 n" by (auto simp: fds_eq_iff fds_nth_one) finally show ?thesis by (simp add: fds_nth_one) qed (* 2.23 *) lemma dirichlet_prod_inversion_completely_multiplicative: fixes a :: "nat \ 'a :: comm_ring_1" assumes "completely_multiplicative_function a" shows "(\x\1. g x = dirichlet_prod' a f x) \ (\x\1. f x = dirichlet_prod' (\n. moebius_mu n * a n) g x)" by (intro dirichlet_prod'_inversion ext completely_multiplicative_imp_moebius_mu_inverse assms) lemma divisor_sigma_conv_dirichlet_prod: "divisor_sigma x n = dirichlet_prod (\n. real n powr x) (\_. 1) n" proof (cases "n = 0") case False have "fds (divisor_sigma x) = fds_shift x fds_zeta * fds_zeta" using fds_divisor_sigma[of x] by (simp add: mult_ac) thus ?thesis using False by (auto simp: fds_eq_iff fds_nth_mult) qed simp_all subsection \Legendre's identity\ definition legendre_aux :: "real \ nat \ nat" where "legendre_aux x p = (if prime p then (\m | m > 0 \ real (p ^ m) \ x. nat \x / p ^ m\) else 0)" lemma legendre_aux_not_prime [simp]: "\prime p \ legendre_aux x p = 0" by (simp add: legendre_aux_def) lemma legendre_aux_eq_0: assumes "real p > x" shows "legendre_aux x p = 0" proof (cases "prime p") case True have [simp]: "\real p ^ m \ x" if "m > 0" for m proof - have "x < real p ^ 1" using assms by simp also have "\ \ real p ^ m" using prime_gt_1_nat[OF True] that by (intro power_increasing) auto finally show ?thesis by auto qed from assms have *: "{m. m > 0 \ real (p ^ m) \ x} = {}" using prime_gt_1_nat[OF True] by auto show ?thesis unfolding legendre_aux_def by (subst *) auto qed (auto simp: legendre_aux_def) lemma legendre_aux_posD: assumes "legendre_aux x p > 0" shows "prime p" "real p \ x" proof - show "real p \ x" using legendre_aux_eq_0[of x p] assms by (cases "real p \ x") auto qed (use assms in \auto simp: legendre_aux_def split: if_splits\) lemma exponents_le_finite: assumes "p > (1 :: nat)" "k > 0" shows "finite {i. real (p ^ (k * i + l)) \ x}" proof (rule finite_subset) show "{i. real (p ^ (k * i + l)) \ x} \ {..nat \x\}" proof safe fix i assume i: "real (p ^ (k * i + l)) \ x" - have "i < 2 ^ i" by (induction i) auto + have "i < 2 ^ i" + by (rule less_exp) also from assms have "i \ k * i + l" by (cases k) auto hence "2 ^ i \ (2 ^ (k * i + l) :: nat)" using assms by (intro power_increasing) auto also have "\ \ p ^ (k * i + l)" using assms by (intro power_mono) auto also have "real \ \ x" using i by simp finally show "i \ nat \x\" by linarith qed qed auto lemma finite_sum_legendre_aux: assumes "prime p" shows "finite {m. m > 0 \ real (p ^ m) \ x}" by (rule finite_subset[OF _ exponents_le_finite[where k = 1 and l = 0 and p = p]]) (use assms prime_gt_1_nat[of p] in auto) lemma legendre_aux_set_eq: assumes "prime p" "x \ 1" shows "{m. m > 0 \ real (p ^ m) \ x} = {0<..nat \log (real p) x\}" using prime_gt_1_nat[OF assms(1)] assms by (auto simp: le_nat_iff le_log_iff le_floor_iff powr_realpow) lemma legendre_aux_altdef1: "legendre_aux x p = (if prime p \ x \ 1 then (\m\{0<..nat \log (real p) x\}. nat \x / p ^ m\) else 0)" proof (cases "prime p \ x < 1") case False thus ?thesis using legendre_aux_set_eq[of p x] by (auto simp: legendre_aux_def) next case True have [simp]: "\(real p ^ m \ x)" for m proof - have "x < real 1" using True by simp also have "real 1 \ real (p ^ m)" unfolding of_nat_le_iff by (intro one_le_power) (use prime_gt_1_nat[of p] True in auto) finally show "\(real p ^ m \ x)" by auto qed have "{m. m > 0 \ real (p ^ m) \ x} = {}" by simp with True show ?thesis by (simp add: legendre_aux_def) qed lemma legendre_aux_altdef2: assumes "x \ 1" "prime p" "real p ^ Suc k > x" shows "legendre_aux x p = (\m\{0<..k}. nat \x / p ^ m\)" proof - have "legendre_aux x p = (\m | m > 0 \ real (p ^ m) \ x. nat \x / p ^ m\)" using assms by (simp add: legendre_aux_def) also have "\ = (\m\{0<..k}. nat \x / p ^ m\)" proof (intro sum.mono_neutral_left) show "{m. 0 < m \ real (p ^ m) \ x} \ {0<..k}" proof safe fix m assume "m > 0" "real (p ^ m) \ x" hence "real p ^ m \ x" by simp also note \x < real p ^ Suc k\ finally show "m \ {0<..k}" using \m > 0\ using prime_gt_1_nat[OF \prime p\] by (subst (asm) power_strict_increasing_iff) auto qed qed (use prime_gt_0_nat[of p] assms in \auto simp: field_simps\) finally show ?thesis . qed (* 3.14 *) theorem legendre_identity: "sum_upto ln x = prime_sum_upto (\p. legendre_aux x p * ln p) x" proof - define S where "S = (SIGMA p:{p. prime p \ real p \ x}. {i. i > 0 \ real (p ^ i) \ x})" have prime_power_leD: "real p \ x" if "real p ^ i \ x" "prime p" "i > 0" for p i proof - have "real p ^ 1 \ real p ^ i" using that prime_gt_1_nat[of p] by (intro power_increasing) auto also have "\ \ x" by fact finally show "real p \ x" by simp qed have "sum_upto ln x = sum_upto (\n. mangoldt n * real (nat \x / real n\)) x" by (rule sum_upto_ln_conv_sum_upto_mangoldt) also have "\ = (\(p, i) | prime p \ 0 < i \ real (p ^ i) \ x. ln p * real (nat \x / real (p ^ i)\))" by (subst sum_upto_primepows[where g = "\p i. ln p * real (nat \x / real (p ^ i)\)"]) (auto simp: mangoldt_non_primepow) also have "\ = (\(p,i)\S. ln p * real (nat \x / p ^ i\))" using prime_power_leD by (intro sum.cong refl) (auto simp: S_def) also have "\ = (\p | prime p \ real p \ x. \i | i > 0 \ real (p ^ i) \ x. ln p * real (nat \x / p ^ i\))" proof (unfold S_def, subst sum.Sigma) have "{p. prime p \ real p \ x} \ {..nat \x\}" by (auto simp: le_nat_iff le_floor_iff) thus "finite {p. prime p \ real p \ x}" by (rule finite_subset) auto next show "\p\{p. prime p \ real p \ x}. finite {i. 0 < i \ real (p ^ i) \ x}" by (intro ballI finite_sum_legendre_aux) auto qed auto also have "\ = (\p | prime p \ real p \ x. ln p * real (\i | i > 0 \ real (p ^ i) \ x. (nat \x / p ^ i\)))" by (simp add: sum_distrib_left) also have "\ = (\p | prime p \ real p \ x. ln p * real (legendre_aux x p))" by (intro sum.cong refl) (auto simp: legendre_aux_def) also have "\ = prime_sum_upto (\p. ln p * real (legendre_aux x p)) x" by (simp add: prime_sum_upto_def) finally show ?thesis by (simp add: mult_ac) qed lemma legendre_identity': "fact (nat \x\) = (\p | prime p \ real p \ x. p ^ legendre_aux x p)" proof - have fin: "finite {p. prime p \ real p \ x}" by (rule finite_subset[of _ "{..nat \x\}"]) (auto simp: le_nat_iff le_floor_iff) have "real (fact (nat \x\)) = exp (sum_upto ln x)" by (subst sum_upto_ln_conv_ln_fact) auto also have "sum_upto ln x = prime_sum_upto (\p. legendre_aux x p * ln p) x" by (rule legendre_identity) also have "exp \ = (\p | prime p \ real p \ x. exp (ln (real p) * legendre_aux x p))" unfolding prime_sum_upto_def using fin by (subst exp_sum) (auto simp: mult_ac) also have "\ = (\p | prime p \ real p \ x. real (p ^ legendre_aux x p))" proof (intro prod.cong refl) fix p assume "p \ {p. prime p \ real p \ x}" hence "p > 0" using prime_gt_0_nat[of p] by auto from \p > 0\ have "exp (ln (real p) * legendre_aux x p) = real p powr real (legendre_aux x p)" by (simp add: powr_def) also from \p > 0\ have "\ = real (p ^ legendre_aux x p)" by (subst powr_realpow) auto finally show "exp (ln (real p) * legendre_aux x p) = \" . qed also have "\ = real (\p | prime p \ real p \ x. p ^ legendre_aux x p)" by simp finally show ?thesis unfolding of_nat_eq_iff . qed subsection \A weighted sum of the Möbius \\\ function\ (* TODO: Move to Dirichlet_Series? *) (* 3.13 *) context fixes M :: "real \ real" defines "M \ (\x. sum_upto (\n. moebius_mu n / n) x)" begin lemma abs_sum_upto_moebius_mu_over_n_less: assumes x: "x \ 2" shows "\M x\ < 1" proof - have "x * sum_upto (\n. moebius_mu n / n) x - sum_upto (\n. moebius_mu n * frac (x / n)) x = sum_upto (\n. moebius_mu n * (x / n - frac (x / n))) x" by (subst mult.commute[of x]) (simp add: sum_upto_def sum_distrib_right sum_subtractf ring_distribs) also have "(\n. x / real n - frac (x / real n)) = (\n. of_int (floor (x / real n)))" by (simp add: frac_def) also have "sum_upto (\n. moebius_mu n * real_of_int \x / real n\) x = real_of_int (sum_upto (\n. moebius_mu n * \x / real n\) x)" by (simp add: sum_upto_def) also have "\ = 1" using x by (subst sum_upto_moebius_times_floor_linear) auto finally have eq: "x * M x = 1 + sum_upto (\n. moebius_mu n * frac (x / n)) x" by (simp add: M_def) have "x * \M x\ = \x * M x\" using x by (simp add: abs_mult) also note eq also have "\1 + sum_upto (\n. moebius_mu n * frac (x / n)) x\ \ 1 + \sum_upto (\n. moebius_mu n * frac (x / n)) x\" by linarith also have "\sum_upto (\n. moebius_mu n * frac (x / n)) x\ \ sum_upto (\n. \moebius_mu n * frac (x / n)\) x" unfolding sum_upto_def by (rule sum_abs) also have "\ \ sum_upto (\n. frac (x / n)) x" unfolding sum_upto_def by (intro sum_mono) (auto simp: moebius_mu_def abs_mult) also have "\ = (\n\{0<..nat \x\}. frac (x / n))" by (simp add: sum_upto_altdef) also have "{0<..nat \x\} = insert 1 {1<..nat \x\}" using x by (auto simp: le_nat_iff le_floor_iff) also have "(\n\\. frac (x / n)) = frac x + (\n\{1<..nat \x\}. frac (x / n))" by (subst sum.insert) auto also have "(\n\{1<..nat \x\}. frac (x / n)) < (\n\{1<..nat \x\}. 1)" using x by (intro sum_strict_mono frac_lt_1) auto also have "\ = nat \x\ - 1" by simp also have "1 + (frac x + real (nat \x\ - 1)) = x" using x by (subst of_nat_diff) (auto simp: le_nat_iff le_floor_iff frac_def) finally have "x * \M x\ < x * 1" by simp with x show "\M x\ < 1" by (subst (asm) mult_less_cancel_left_pos) auto qed lemma sum_upto_moebius_mu_over_n_eq: assumes "x < 2" shows "M x = (if x \ 1 then 1 else 0)" proof (cases "x \ 1") case True have "M x = (\n\{n. n > 0 \ real n \ x}. moebius_mu n / n)" by (simp add: M_def sum_upto_def) also from assms True have "{n. n > 0 \ real n \ x} = {1}" by auto thus ?thesis using True by (simp add: M_def sum_upto_def) next case False have "M x = (\n\{n. n > 0 \ real n \ x}. moebius_mu n / n)" by (simp add: M_def sum_upto_def) also from False have "{n. n > 0 \ real n \ x} = {}" by auto finally show ?thesis using False by (simp add: M_def) qed lemma abs_sum_upto_moebius_mu_over_n_le: "\M x\ \ 1" using sum_upto_moebius_mu_over_n_eq[of x] abs_sum_upto_moebius_mu_over_n_less[of x] by (cases "x < 2") auto end end \ No newline at end of file diff --git a/thys/Prime_Number_Theorem/Mertens_Theorems.thy b/thys/Prime_Number_Theorem/Mertens_Theorems.thy --- a/thys/Prime_Number_Theorem/Mertens_Theorems.thy +++ b/thys/Prime_Number_Theorem/Mertens_Theorems.thy @@ -1,755 +1,755 @@ (* File: Mertens_Theorems.thy Author: Manuel Eberl (TU München) *) section \Mertens' Theorems\ theory Mertens_Theorems imports Prime_Counting_Functions "Stirling_Formula.Stirling_Formula" begin (*<*) unbundle prime_counting_notation (*>*) text \ In this section, we will prove Mertens' First and Second Theorem. These are weaker results than the Prime Number Theorem, and we will derive them without using it. However, like Mertens himself, we will not only prove them \<^emph>\asymptotically\, but \<^emph>\absolutely\. This means that we will show that the remainder terms are not only ``Big-O'' of some bound, but we will give concrete (and reasonably tight) upper and lower bounds for them that hold on the entire domain. This makes the proofs a bit more tedious. \ subsection \Absolute Bounds for Mertens' First Theorem\ text \ We have already shown the asymptotic form of Mertens' first theorem, i.\,e.\ $\mathfrak{M}(n) = \ln n + O(1)$. We now want to obtain some absolute bounds on the $O(1)$ remainder term using a more careful derivation than before. The precise bounds we will show are $\mathfrak {M}(n) - \ln n \in (-1-\frac{9}{\pi^2}; \ln 4] \approx (-1.9119; 1.3863]$ for \n \ \\. First, we need a simple lemma on the finiteness of exponents to consider in a sum of all prime powers up to a certain point: \ lemma exponents_le_finite: assumes "p > (1 :: nat)" "k > 0" shows "finite {i. real (p ^ (k * i + l)) \ x}" proof (rule finite_subset) show "{i. real (p ^ (k * i + l)) \ x} \ {..nat \x\}" proof safe fix i assume i: "real (p ^ (k * i + l)) \ x" - have "i < 2 ^ i" by (induction i) auto + have "i < 2 ^ i" by (rule less_exp) also from assms have "i \ k * i + l" by (cases k) auto hence "2 ^ i \ (2 ^ (k * i + l) :: nat)" using assms by (intro power_increasing) auto also have "\ \ p ^ (k * i + l)" using assms by (intro power_mono) auto also have "real \ \ x" using i by simp finally show "i \ nat \x\" by linarith qed qed auto text \ Next, we need the following bound on $\zeta'(2)$: \ lemma deriv_zeta_2_bound: "Re (deriv zeta 2) > -1" proof - have "((\x::real. ln (x + 3) * (x + 3) powr -2) has_integral (ln 3 + 1) / 3) (interior {0..})" using ln_powr_has_integral_at_top[of 1 0 3 "-2"] by (simp add: interior_real_atLeast powr_minus) hence "((\x::real. ln (x + 3) * (x + 3) powr -2) has_integral (ln 3 + 1) / 3) {0..}" by (subst (asm) has_integral_interior) auto also have "?this \ ((\x::real. ln (x + 3) / (x + 3) ^ 2) has_integral (ln 3 + 1) / 3) {0..}" by (intro has_integral_cong) (auto simp: powr_minus field_simps) finally have int: \ . have "exp (1 / 2 :: real) ^ 2 \ 2 ^ 2" using exp_le by (subst exp_double [symmetric]) simp_all hence exp_half: "exp (1 / 2 :: real) \ 2" by (rule power2_le_imp_le) auto have mono: "ln x / x ^ 2 \ ln y / y ^ 2" if "y \ exp (1/2)" "x \ y" for x y :: real proof (rule DERIV_nonpos_imp_nonincreasing[of _ _ "\x. ln x / x ^ 2"]) fix t assume t: "t \ y" "t \ x" have "y > 0" by (rule less_le_trans[OF _ that(1)]) auto with t that have "ln t \ ln (exp (1 / 2))" by (subst ln_le_cancel_iff) auto hence "ln t \ 1 / 2" by (simp only: ln_exp) from t \y > 0\ have "((\x. ln x / x ^ 2) has_field_derivative ((1 - 2 * ln t) / t ^ 3)) (at t)" by (auto intro!: derivative_eq_intros simp: eval_nat_numeral field_simps) moreover have "(1 - 2 * ln t) / t ^ 3 \ 0" using t that \y > 0\ \ln t \ 1 / 2\ by (intro divide_nonpos_pos) auto ultimately show "\f'. ((\x. ln x / x ^ 2) has_field_derivative f') (at t) \ f' \ 0" by blast qed fact+ have "fds_converges (fds_deriv fds_zeta) (2 :: complex)" by (intro fds_converges_deriv) auto hence "(\n. of_real (-ln (real (Suc n)) / (of_nat (Suc n)) ^ 2)) sums deriv zeta 2" by (auto simp: fds_converges_altdef add_ac eval_fds_deriv_zeta fds_nth_deriv scaleR_conv_of_real simp del: of_nat_Suc) note * = sums_split_initial_segment[OF sums_minus[OF sums_Re[OF this]], of 3] have "(\n. ln (real (n+4)) / real (n+4) ^ 2) sums (-Re (deriv zeta 2) - (ln 2 / 4 + ln 3 / 9))" using * by (simp add: eval_nat_numeral) hence "-Re (deriv zeta 2) - (ln 2 / 4 + ln 3 / 9) = (\n. ln (real (Suc n) + 3) / (real (Suc n) + 3) ^ 2)" by (simp_all add: sums_iff algebra_simps) also have "\ \ (ln 3 + 1) / 3" using int exp_half by (intro decreasing_sum_le_integral divide_nonneg_pos mono) (auto simp: powr_minus field_simps) finally have "-Re (deriv zeta 2) \ (16 * ln 3 + 9 * ln 2 + 12) / 36" by simp also have "ln 3 \ (11 / 10 :: real)" using ln_approx_bounds[of 3 2] by (simp add: power_numeral_reduce numeral_2_eq_2) hence "(16 * ln 3 + 9 * ln 2 + 12) / 36 \ (16 * (11 / 10) + 9 * 25 / 36 + 12) / (36 :: real)" using ln2_le_25_over_36 by (intro add_mono mult_left_mono divide_right_mono) auto also have "\ < 1" by simp finally show ?thesis by simp qed text \ Using the logarithmic derivative of Euler's product formula for $\zeta(s)$ at $s = 2$ and the bound on $\zeta'(2)$ we have just derived, we can obtain the bound \[\sum_{p^i \leq x, i \geq 2} \frac{\ln p}{p^i} < \frac{9}{\pi^2}\ .\] \ lemma mertens_remainder_aux_bound: fixes x :: real defines "R \ (\(p,i) | prime p \ i > 1 \ real (p ^ i) \ x. ln (real p) / p ^ i)" shows "R < 9 / pi\<^sup>2" proof - define S' where "S' = {(p, i). prime p \ i > 1 \ real (p ^ i) \ x}" define S'' where "S'' = {(p, i). prime p \ i > 1 \ real (p ^ Suc i) \ x}" have finite_row: "finite {i. i > 1 \ real (p ^ (i + k)) \ x}" if p: "prime p" for p k proof (rule finite_subset) show "{i. i > 1 \ real (p ^ (i + k)) \ x} \ {..nat \x\}" proof safe fix i assume i: "i > 1" "real (p ^ (i + k)) \ x" have "i < 2 ^ (i + k)" by (induction i) auto also from p have "\ \ p ^ (i + k)" by (intro power_mono) (auto dest: prime_gt_1_nat) also have "real \ \ x" using i by simp finally show "i \ nat \x\" by linarith qed qed auto have "S'' \ S'" unfolding S''_def S'_def proof safe fix p i assume pi: "prime p" "real (p ^ Suc i) \ x" "i > 1" have "real (p ^ i) \ real (p ^ Suc i)" using pi unfolding of_nat_le_iff by (intro power_increasing) (auto dest: prime_gt_1_nat) also have "\ \ x" by fact finally show "real (p ^ i) \ x" . qed have S'_alt: "S' = (SIGMA p:{p. prime p \ real p \ x}. {i. i > 1 \ real (p ^ i) \ x})" unfolding S'_def proof safe fix p i assume "prime p" "real (p ^ i) \ x" "i > 1" hence "p ^ 1 \ p ^ i" by (intro power_increasing) (auto dest: prime_gt_1_nat) also have "real \ \ x" by fact finally show "real p \ x" by simp qed have finite: "finite {p. prime p \ real p \ x}" by (rule finite_subset[OF _ finite_Nats_le_real[of x]]) (auto dest: prime_gt_0_nat) have "finite S'" unfolding S'_alt using finite_row[of _ 0] by (intro finite_SigmaI finite) auto have "R \ 3 / 2 * (\(p, i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i))" proof - have "R = (\y\{0, 1}. \z | z \ S' \ snd z mod 2 = y. ln (real (fst z)) / real (fst z ^ snd z))" using \finite S'\ by (subst sum.group) (auto simp: case_prod_unfold R_def S'_def) also have "\ = (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i)) + (\(p,i) | (p, i) \ S' \ odd i. ln (real p) / real (p ^ i))" unfolding even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one by (simp add: case_prod_unfold) also have "(\(p,i) | (p, i) \ S' \ odd i. ln (real p) / real (p ^ i)) = (\(p,i) | (p, i) \ S'' \ even i. ln (real p) / real (p ^ Suc i))" by (intro sum.reindex_bij_witness[of _ "\(p,i). (p, Suc i)" "\(p,i). (p, i - 1)"]) (auto simp: case_prod_unfold S'_def S''_def elim: oddE simp del: power_Suc) also have "\ \ (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ Suc i))" using \S'' \ S'\ unfolding case_prod_unfold by (intro sum_mono2 divide_nonneg_pos ln_ge_zero finite_subset[OF _ \finite S'\]) (auto simp: S'_def S''_def case_prod_unfold dest: prime_gt_0_nat simp del: power_Suc) also have "\ \ (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (2 * p ^ i))" unfolding case_prod_unfold by (intro sum_mono divide_left_mono) (auto simp: S'_def dest!: prime_gt_1_nat) also have "\ = (1 / 2) * (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i))" by (subst sum_distrib_left) (auto simp: case_prod_unfold) also have "(\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i)) + \ = 3 / 2 * (\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i))" by simp finally show ?thesis by simp qed also have "(\(p,i) | (p, i) \ S' \ even i. ln (real p) / real (p ^ i)) = (\p | prime p \ real p \ x. ln (real p) * (\i | i > 0 \ even i \ real (p ^ i) \ x. (1 / real p) ^ i))" unfolding sum_distrib_left proof (subst sum.Sigma[OF _ ballI]) fix p assume p: "p \ {p. prime p \ real p \ x}" thus "finite {i. 0 < i \ even i \ real (p ^ i) \ x}" by (intro finite_subset[OF _ exponents_le_finite[of p 1 0 x]]) (auto dest: prime_gt_1_nat) qed (auto intro!: sum.cong finite_subset[OF _ finite_Nats_le_real[of x]] dest: prime_gt_0_nat simp: S'_alt power_divide) also have "\ \ (\p | prime p \ real p \ x. ln (real p) / (real p ^ 2 - 1))" proof (rule sum_mono) fix p assume p: "p \ {p. prime p \ real p \ x}" have "p > 1" using p by (auto dest: prime_gt_1_nat) have "(\i | i > 0 \ even i \ real (p ^ i) \ x. (1 / real p) ^ i) = (\i | real (p ^ (2 * i + 2)) \ x. (1 / real p) ^ (2 * i)) / real p ^ 2" (is "_ = ?S / _") unfolding sum_divide_distrib by (rule sum.reindex_bij_witness[of _ "\i. 2 * Suc i" "\i. (i - 2) div 2"]) (insert \p > 1\, auto simp: numeral_3_eq_3 power2_eq_square power_diff algebra_simps elim!: evenE) also have "?S = (\i | real (p ^ (2 * i + 2)) \ x. (1 / real p ^ 2) ^ i)" by (subst power_mult) (simp_all add: algebra_simps power_divide) also have "\ \ (\i. (1 / real p ^ 2) ^ i)" using exponents_le_finite[of p 2 2 x] \p > 1\ by (intro sum_le_suminf) (auto simp: summable_geometric_iff) also have "\ = real p ^ 2 / (real p ^ 2 - 1)" using \p > 1\ by (subst suminf_geometric) (auto simp: field_simps) also have "\ / real p ^ 2 = 1 / (real p ^ 2 - 1)" using \p > 1 \ by (simp add: divide_simps) finally have "(\i | 0 < i \ even i \ real (p ^ i) \ x. (1 / real p) ^ i) \ 1 / (real p ^ 2 - 1)" (is "?lhs \ ?rhs") using \p > 1\ by (simp add: divide_right_mono) thus "ln (real p) * ?lhs \ ln (real p) / (real p ^ 2 - 1)" using \p > 1\ by (simp add: divide_simps) qed also have "\ = (\\<^sub>a p | prime p \ real p \ x. ln (real p) / (real p ^ 2 - 1))" using finite by (intro infsetsum_finite [symmetric]) auto also have "\ \ (\\<^sub>a p | prime p. ln (real p) / (real p ^ 2 - 1))" using eval_fds_logderiv_zeta_real[of 2] finite by (intro infsetsum_mono_neutral_left divide_nonneg_pos) (auto simp: dest: prime_gt_1_nat) also have "\ = -Re (deriv zeta (of_real 2) / zeta (of_real 2))" by (subst eval_fds_logderiv_zeta_real) auto also have "\ = (-Re (deriv zeta 2)) * (6 / pi\<^sup>2)" by (simp add: zeta_even_numeral) also have "\ < 1 * (6 / pi\<^sup>2)" using deriv_zeta_2_bound by (intro mult_strict_right_mono) auto also have "3 / 2 * \ = 9 / pi\<^sup>2" by simp finally show ?thesis by simp qed text \ We now consider the equation \[\ln (n!) = \sum_{k\leq n} \Lambda(k) \left\lfloor\frac{n}{k}\right\rfloor\] and estimate both sides in different ways. The left-hand-side can be estimated using Stirling's formula, and we can simplify the right-hand side to \[\sum_{k\leq n} \Lambda(k) \left\lfloor\frac{n}{k}\right\rfloor = \sum_{p^i \leq x, i \geq 1} \ln p \left\lfloor\frac{n}{p^i}\right\rfloor\] and then split the sum into those $p^i$ with $i = 1$ and those with $i \geq 2$. Applying the bound we have just shown and some more routine estimates, we obtain the following reasonably strong version of Mertens' First Theorem on the naturals: $\mathfrak M(n) - \ln(n) \in (-1-\frac{9}{\pi^2}; \ln 4]$ \ theorem mertens_bound_strong: fixes n :: nat assumes n: "n > 0" shows "\ n - ln n \ {-1 - 9 / pi\<^sup>2<..ln 4}" proof (cases "n \ 3") case False with n consider "n = 1" | "n = 2" by force thus ?thesis proof cases assume [simp]: "n = 1" have "-1 + (-9 / pi\<^sup>2) < 0" by (intro add_neg_neg divide_neg_pos) auto thus ?thesis by simp next assume [simp]: "n = 2" have eq: "\ n - ln n = -ln 2 / 2" by (simp add: eval_\) have "-1 - 9 / pi ^ 2 + ln 2 / 2 \ -1 - 9 / 4 ^ 2 + 25 / 36 / 2" using pi_less_4 ln2_le_25_over_36 by (intro diff_mono add_mono divide_left_mono divide_right_mono power_mono) auto also have "\ < 0" by simp finally have "-ln 2 / 2 > -1 - 9 / pi\<^sup>2" by simp moreover { have "-ln 2 / 2 \ (0::real)" by (intro divide_nonpos_pos) auto also have "\ \ ln 4" by simp finally have "-ln 2 / 2 \ ln (4 :: real)" by simp } ultimately show ?thesis unfolding eq by simp qed next case True hence n: "n \ 3" by simp have finite: "finite {(p, i). prime p \ i \ 1 \ p ^ i \ n}" proof (rule finite_subset) show "{(p, i). prime p \ i \ 1 \ p ^ i \ n} \ {..nat \root 1 (real n)\} \ {..nat \log 2 (real n)\}" using primepows_le_subset[of "real n" 1] n unfolding of_nat_le_iff by auto qed auto define r where "r = prime_sum_upto (\p. ln (real p) * frac (real n / real p)) n" define R where "R = (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) * real (n div (p ^ i)))" define R' where "R' = (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) / p ^ i)" have [simp]: "ln (4 :: real) = 2 * ln 2" using ln_realpow[of 2 2] by simp from pi_less_4 have "ln pi \ ln 4" by (subst ln_le_cancel_iff) auto also have "\ = 2 * ln 2" by simp also have "\ \ 2 * (25 / 36)" by (intro mult_left_mono ln2_le_25_over_36) auto finally have ln_pi: "ln pi \ 25 / 18" by simp have "ln 3 \ ln (4::nat)" by (subst ln_le_cancel_iff) auto also have "\ = 2 * ln 2" by simp also have "\ \ 2 * (25 / 36)" by (intro mult_left_mono ln2_le_25_over_36) auto finally have ln_3: "ln (3::real) \ 25 / 18" by simp have "R / n = (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) * (real (n div (p ^ i)) / n))" by (simp add: R_def sum_divide_distrib field_simps case_prod_unfold) also have "\ \ (\(p,i) | prime p \ i > 1 \ p ^ i \ n. ln (real p) * (1 / p ^ i))" unfolding R'_def case_prod_unfold using n by (intro sum_mono mult_left_mono) (auto simp: field_simps real_of_nat_div dest: prime_gt_0_nat) also have "\ = R'" by (simp add: R'_def) also have "R' < 9 / pi\<^sup>2" unfolding R'_def using mertens_remainder_aux_bound[of n] by simp finally have "R / n < 9 / pi\<^sup>2" . moreover have "R \ 0" unfolding R_def by (intro sum_nonneg mult_nonneg_nonneg) (auto dest: prime_gt_0_nat) ultimately have R_bounds: "R / n \ {0..<9 / pi\<^sup>2}" by simp have "ln (fact n :: real) \ ln (2 * pi * n) / 2 + n * ln n - n + 1 / (12 * n)" using ln_fact_bounds(2)[of n] n by simp also have "\ / n - ln n = -1 + (ln 2 + ln pi) / (2 * n) + (ln n / n) / 2 + 1 / (12 * real n ^ 2)" using n by (simp add: power2_eq_square field_simps ln_mult) also have "\ \ -1 + (ln 2 + ln pi) / (2 * 3) + (ln 3 / 3) / 2 + 1 / (12 * 3\<^sup>2)" using exp_le n pi_gt3 by (intro add_mono divide_right_mono divide_left_mono mult_mono mult_pos_pos ln_x_over_x_mono power_mono) auto also have "\ \ -1 + (25 / 36 + 25 / 18) / (2 * 3) + (25 / 18 / 3) / 2 + 1 / (12 * 3\<^sup>2)" using ln_pi ln2_le_25_over_36 ln_3 by (intro add_mono divide_left_mono divide_right_mono) auto also have "\ \ 0" by simp finally have "ln n - ln (fact n) / n \ 0" using n by (simp add: divide_right_mono) have "-ln (fact n) \ -ln (2 * pi * n) / 2 - n * ln n + n" using ln_fact_bounds(1)[of n] n by simp also have "ln n + \ / n = -ln (2 * pi) / (2 * n) - (ln n / n) / 2 + 1" using n by (simp add: field_simps ln_mult) also have "\ \ 0 - 0 + 1" using pi_gt3 n by (intro add_mono diff_mono) auto finally have upper: "ln n - ln (fact n) / n \ 1" using n by (simp add: divide_right_mono) with \ln n - ln (fact n) / n \ 0\ have fact_bounds: "ln n - ln (fact n) / n \ {0..1}" by simp have "r \ prime_sum_upto (\p. ln p * 1) n" using less_imp_le[OF frac_lt_1] unfolding r_def \_def prime_sum_upto_def by (intro sum_mono mult_left_mono) (auto simp: dest: prime_gt_0_nat) also have "\ = \ n" by (simp add: \_def) also have "\ < ln 4 * n" using n by (intro \_upper_bound) auto finally have "r / n < ln 4" using n by (simp add: field_simps) moreover have "r \ 0" unfolding r_def prime_sum_upto_def by (intro sum_nonneg mult_nonneg_nonneg) (auto dest: prime_gt_0_nat) ultimately have r_bounds: "r / n \ {0..k. mangoldt k * real (n div k)) (real n)" by (simp add: ln_fact_conv_sum_upto_mangoldt) also have "\ = (\(p,i) | prime p \ i > 0 \ real (p ^ i) \ real n. ln (real p) * real (n div (p ^ i)))" by (intro sum_upto_primepows) (auto simp: mangoldt_non_primepow) also have "{(p, i). prime p \ i > 0 \ real (p ^ i) \ real n} = {(p, i). prime p \ i = 1 \ p \ n} \ {(p, i). prime p \ i > 1 \ (p ^ i) \ n}" unfolding of_nat_le_iff by (auto simp: not_less le_Suc_eq) also have "(\(p,i)\\. ln (real p) * real (n div (p ^ i))) = (\(p,i) | prime p \ i = 1 \ p \ n. ln (real p) * real (n div (p ^ i))) + R" (is "_ = ?S + _") by (subst sum.union_disjoint) (auto intro!: finite_subset[OF _ finite] simp: R_def) also have "?S = prime_sum_upto (\p. ln (real p) * real (n div p)) n" unfolding prime_sum_upto_def by (intro sum.reindex_bij_witness[of _ "\p. (p, 1)" fst]) auto also have "\ = prime_sum_upto (\p. ln (real p) * real n / real p) n - r" unfolding r_def prime_sum_upto_def sum_subtractf[symmetric] using n by (intro sum.cong) (auto simp: frac_def real_of_nat_div algebra_simps) also have "prime_sum_upto (\p. ln (real p) * real n / real p) n = n * \ n" by (simp add: primes_M_def sum_distrib_left sum_distrib_right prime_sum_upto_def field_simps) finally have "\ n - ln n = ln (fact n) / n - ln n + r / n - R / n" using n by (simp add: field_simps) hence "ln n - \ n = ln n - ln (fact n) / n - r / n + R / n" by simp with fact_bounds r_bounds R_bounds show "\ n - ln n \ {-1 - 9 / pi\<^sup>2<..ln 4}" by simp qed text \ As a simple corollary, we obtain a similar bound on the reals. \ lemma mertens_bound_real_strong: fixes x :: real assumes x: "x \ 1" shows "\ x - ln x \ {-1 - 9 / pi ^ 2 - ln (1 + frac x / real (nat \x\)) <.. ln 4}" proof - have "\ x - ln x \ \ (real (nat \x\)) - ln (real (nat \x\))" using assms by simp also have "\ \ ln 4" using mertens_bound_strong[of "nat \x\"] assms by simp finally have "\ x - ln x \ ln 4" . from assms have pos: "real_of_int \x\ \ 0" by linarith have "frac x / real (nat \x\) \ 0" using assms by (intro divide_nonneg_pos) auto moreover have "frac x / real (nat \x\) \ 1 / 1" using assms frac_lt_1[of x] by (intro frac_le) auto ultimately have *: "frac x / real (nat \x\) \ {0..1}" by auto have "ln x - ln (real (nat \x\)) = ln (x / real (nat \x\))" using assms by (subst ln_div) auto also have "x / real (nat \x\) = 1 + frac x / real (nat \x\)" using assms pos by (simp add: frac_def field_simps) finally have "\ x - ln x > -1-9/pi^2-ln (1 + frac x / real (nat \x\))" using mertens_bound_strong[of "nat \x\"] x by simp with \\ x - ln x \ ln 4\ show ?thesis by simp qed text \ We weaken this estimate a bit to obtain nicer bounds: \ lemma mertens_bound_real': fixes x :: real assumes x: "x \ 1" shows "\ x - ln x \ {-2<..25/18}" proof - have "\ x - ln x \ ln 4" using mertens_bound_real_strong[of x] x by simp also have "\ \ 25 / 18" using ln_realpow[of 2 2] ln2_le_25_over_36 by simp finally have "\ x - ln x \ 25 / 18" . have ln2: "ln (2 :: real) \ {2/3..25/36}" using ln_approx_bounds[of 2 1] by (simp add: eval_nat_numeral) have ln3: "ln (3::real) \ {1..10/9}" using ln_approx_bounds[of 3 1] by (simp add: eval_nat_numeral) have ln5: "ln (5::real) \ {4/3..76/45}" using ln_approx_bounds[of 5 1] by (simp add: eval_nat_numeral) have ln7: "ln (7::real) \ {3/2..15/7}" using ln_approx_bounds[of 7 1] by (simp add: eval_nat_numeral) have ln11: "ln (11::real) \ {5/3..290/99}" using ln_approx_bounds[of 11 1] by (simp add: eval_nat_numeral) \ \Choosing the lower bound -2 is somewhat arbitrary here; it is a trade-off between getting a reasonably tight bound and having to make lots of case distinctions. To get -2 as a lower bound, we have to show the cases up to \x = 11\ by case distinction,\ have "\ x - ln x > -2" proof (cases "x \ 11") case False hence "x \ {1..<2} \ x \ {2..<3} \ x \ {3..<5} \ x \ {5..<7} \ x \ {7..<11}" using x by force thus ?thesis proof (elim disjE) assume x: "x \ {1..<2}" hence "ln x - \ x \ ln 2 - 0" by (intro diff_mono) auto also have "\ < 2" using ln2_le_25_over_36 by simp finally show ?thesis by simp next assume x: "x \ {2..<3}" hence [simp]: "\x\ = 2" by (intro floor_unique) auto from x have "ln x - \ x \ ln 3 - ln 2 / 2" by (intro diff_mono) (auto simp: eval_\) also have "\ = ln (9 / 2) / 2" using ln_realpow[of 3 2] by (simp add: ln_div) also have "\ < 2" using ln_approx_bounds[of "9 / 2" 1] by (simp add: eval_nat_numeral) finally show ?thesis by simp next assume x: "x \ {3..<5}" hence "\ 3 = \ x" unfolding primes_M_def by (intro prime_sum_upto_eqI'[where a' = 3 and b' = 4]) (auto simp: nat_le_iff le_numeral_iff nat_eq_iff floor_eq_iff) also have "\ 3 = ln 2 / 2 + ln 3 / 3" by (simp add: eval_\ eval_nat_numeral mark_out_code) finally have [simp]: "\ x = ln 2 / 2 + ln 3 / 3" .. from x have "ln x - \ x \ ln 5 - (ln 2 / 2 + ln 3 / 3)" by (intro diff_mono) auto also have "\ < 2" using ln2 ln3 ln5 by simp finally show ?thesis by simp next assume x: "x \ {5..<7}" hence "\ 5 = \ x" unfolding primes_M_def by (intro prime_sum_upto_eqI'[where a' = 5 and b' = 6]) (auto simp: nat_le_iff le_numeral_iff nat_eq_iff floor_eq_iff) also have "\ 5 = ln 2 / 2 + ln 3 / 3 + ln 5 / 5" by (simp add: eval_\ eval_nat_numeral mark_out_code) finally have [simp]: "\ x = ln 2 / 2 + ln 3 / 3 + ln 5 / 5" .. from x have "ln x - \ x \ ln 7 - (ln 2 / 2 + ln 3 / 3 + ln 5 / 5)" by (intro diff_mono) auto also have "\ < 2" using ln2 ln3 ln5 ln7 by simp finally show ?thesis by simp next assume x: "x \ {7..<11}" hence "\ 7 = \ x" unfolding primes_M_def by (intro prime_sum_upto_eqI'[where a' = 7 and b' = 10]) (auto simp: nat_le_iff le_numeral_iff nat_eq_iff floor_eq_iff) also have "\ 7 = ln 2 / 2 + ln 3 / 3 + ln 5 / 5 + ln 7 / 7" by (simp add: eval_\ eval_nat_numeral mark_out_code) finally have [simp]: "\ x = ln 2 / 2 + ln 3 / 3 + ln 5 / 5 + ln 7 / 7" .. from x have "ln x - \ x \ ln 11 - (ln 2 / 2 + ln 3 / 3 + ln 5 / 5 + ln 7 / 7)" by (intro diff_mono) auto also have "\ < 2" using ln2 ln3 ln5 ln7 ln11 by simp finally show ?thesis by simp qed next case True have "ln x - \ x \ 1 + 9/pi^2 + ln (1 + frac x / real (nat \x\))" using mertens_bound_real_strong[of x] x by simp also have "1 + frac x / real (nat \x\) \ 1 + 1 / 11" using True frac_lt_1[of x] by (intro add_mono frac_le) auto hence "ln (1 + frac x / real (nat \x\)) \ ln (1 + 1 / 11)" using x by (subst ln_le_cancel_iff) (auto intro!: add_pos_nonneg) also have "\ = ln (12 / 11)" by simp also have "\ \ 1585 / 18216" using ln_approx_bounds[of "12 / 11" 1] by (simp add: eval_nat_numeral) also have "9 / pi ^ 2 \ 9 / 3.141592653588 ^ 2" using pi_approx by (intro divide_left_mono power_mono mult_pos_pos) auto also have "1 + \ + 1585 / 18216 < 2" by (simp add: power2_eq_square) finally show ?thesis by simp qed with \\ x - ln x \ 25 / 18\ show ?thesis by simp qed corollary mertens_first_theorem: fixes x :: real assumes x: "x \ 1" shows "\\ x - ln x\ < 2" using mertens_bound_real'[of x] x by (simp add: abs_if) subsection \Mertens' Second Theorem\ text \ Mertens' Second Theorem concerns the asymptotics of the Prime Harmonic Series, namely \[\sum\limits_{p \leq x} \frac{1}{p} = \ln\ln x + M + O\left(\frac{1}{\ln x}\right)\] where $M \approx 0.261497$ is the Meissel--Mertens constant. We define the constant in the following way: \ definition meissel_mertens where "meissel_mertens = 1 - ln (ln 2) + integral {2..} (\t. (\ t - ln t) / (t * ln t ^ 2))" text \ We will require the value of the integral $\int_a^\infty \frac{t}{\ln^2 t} \textrm{d}t = \frac{1}{\ln a}$ as an upper bound on the remainder term: \ lemma integral_one_over_x_ln_x_squared: assumes a: "(a::real) > 1" shows "set_integrable lborel {a<..} (\t. 1 / (t * ln t ^ 2))" (is ?th1) and "set_lebesgue_integral lborel {a<..} (\t. 1 / (t * ln t ^ 2)) = 1 / ln a" (is ?th2) and "((\t. 1 / (t * (ln t)\<^sup>2)) has_integral 1 / ln a) {a<..}" (is ?th3) proof - have cont: "isCont (\t. 1 / (t * (ln t)\<^sup>2)) x" if "x > a" for x using that a by (auto intro!: continuous_intros) have deriv: "((\x. -1 / ln x) has_real_derivative 1 / (x * (ln x)\<^sup>2)) (at x)" if "x > a" for x using that a by (auto intro!: derivative_eq_intros simp: power2_eq_square field_simps) have lim1: "(((\x. - 1 / ln x) \ real_of_ereal) \ -(1 / ln a)) (at_right (ereal a))" unfolding ereal_tendsto_simps using a by (real_asymp simp: field_simps) have lim2: "(((\x. - 1 / ln x) \ real_of_ereal) \ 0) (at_left \)" unfolding ereal_tendsto_simps using a by (real_asymp simp: field_simps) have "set_integrable lborel (einterval a \) (\t. 1 / (t * ln t ^ 2))" by (rule interval_integral_FTC_nonneg[OF _ deriv cont _ lim1 lim2]) (use a in auto) thus ?th1 by simp have "interval_lebesgue_integral lborel (ereal a) \ (\t. 1 / (t * ln t ^ 2)) = 0 - (-(1 / ln a))" by (rule interval_integral_FTC_nonneg[OF _ deriv cont _ lim1 lim2]) (use a in auto) thus ?th2 by (simp add: interval_integral_to_infinity_eq) have "((\t. 1 / (t * ln t ^ 2)) has_integral set_lebesgue_integral lebesgue {a<..} (\t. 1 / (t * ln t ^ 2))) {a<..}" using \?th1\ by (intro has_integral_set_lebesgue) (auto simp: set_integrable_def integrable_completion) also have "set_lebesgue_integral lebesgue {a<..} (\t. 1 / (t * ln t ^ 2)) = 1 / ln a" using \?th2\ unfolding set_lebesgue_integral_def by (subst integral_completion) auto finally show ?th3 . qed text \ We show that the integral in our definition of the Meissel--Mertens constant is well-defined and give an upper bound for its tails: \ lemma assumes "a > (1 :: real)" defines "r \ (\t. (\ t - ln t) / (t * ln t ^ 2))" shows integrable_meissel_mertens: "set_integrable lborel {a<..} r" and meissel_mertens_integral_le: "norm (integral {a<..} r) \ 2 / ln a" proof - have *: "((\t. 2 * (1 / (t * ln t ^ 2))) has_integral 2 * (1 / ln a)) {a<..}" using assms by (intro has_integral_mult_right integral_one_over_x_ln_x_squared) auto show "set_integrable lborel {a<..} r" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "integrable lborel (\t::real. indicator {a<..} t * (2 * (1 / (t * ln t ^ 2))))" using integrable_mult_right[of 2, OF integral_one_over_x_ln_x_squared(1)[of a, unfolded set_integrable_def]] assms by (simp add: algebra_simps) thus "integrable lborel (\t::real. indicator {a<..} t *\<^sub>R (2 / (t * ln t ^ 2)))" by simp fix x :: real show "norm (indicat_real {a<..} x *\<^sub>R r x) \ norm (indicat_real {a<..} x *\<^sub>R (2 / (x * ln x ^ 2)))" proof (cases "x > a") case True thus ?thesis unfolding norm_scaleR norm_mult r_def norm_divide using mertens_first_theorem[of x] assms by (intro mult_mono frac_le divide_nonneg_pos) (auto simp: indicator_def) qed (auto simp: indicator_def) qed (auto simp: r_def) hence "r integrable_on {a<..}" by (simp add: set_borel_integral_eq_integral(1)) hence "norm (integral {a<..} r) \ integral {a<..} (\x. 2 * (1 / (x * ln x ^ 2)))" proof (rule integral_norm_bound_integral) show "(\x. 2 * (1 / (x * (ln x)\<^sup>2))) integrable_on {a<..}" using * by (simp add: has_integral_iff) fix x assume "x \ {a<..}" hence "norm (r x) \ 2 / (x * (ln x)\<^sup>2)" unfolding r_def norm_divide using mertens_first_theorem[of x] assms by (intro mult_mono frac_le divide_nonneg_pos) (auto simp: indicator_def) thus "norm (r x) \ 2* (1 / (x * ln x ^ 2))" by simp qed also have "\ = 2 / ln a" using * by (simp add: has_integral_iff) finally show "norm (integral {a<..} r) \ 2 / ln a" . qed lemma integrable_on_meissel_mertens: assumes "A \ {1..}" "Inf A > 1" "A \ sets borel" shows "(\t. (\ t - ln t) / (t * ln t ^ 2)) integrable_on A" proof - from assms obtain x where x: "1 < x" "x < Inf A" using dense by blast from assms have "bdd_below A" by (intro bdd_belowI[of _ 1]) auto hence "A \ {Inf A..}" by (auto simp: cInf_lower) also have "\ \ {x<..}" using x by auto finally have *: "A \ {x<..}" . have "set_integrable lborel A (\t. (\ t - ln t) / (t * ln t ^ 2))" by (rule set_integrable_subset[OF integrable_meissel_mertens[of x]]) (use x * assms in auto) thus ?thesis by (simp add: set_borel_integral_eq_integral(1)) qed lemma meissel_mertens_bounds: "\meissel_mertens - 1 + ln (ln 2)\ \ 2 / ln 2" proof - have *: "{2..} - {2<..} = {2::real}" by auto also have "negligible \" by simp finally have "integral {2..} (\t. (\ t - ln t) / (t * (ln t)\<^sup>2)) = integral {2<..} (\t. (\ t - ln t) / (t * (ln t)\<^sup>2))" by (intro sym[OF integral_subset_negligible]) auto also have "norm \ \ 2 / ln 2" by (rule meissel_mertens_integral_le) auto finally show "\meissel_mertens - 1 + ln (ln 2)\ \ 2 / ln 2" by (simp add: meissel_mertens_def) qed text \ Finally, obtaining Mertens' second theorem from the first one is nothing but a routine summation by parts, followed by a use of the above bound: \ theorem mertens_second_theorem: defines "f \ prime_sum_upto (\p. 1 / p)" shows "\x. x \ 2 \ \f x - ln (ln x) - meissel_mertens\ \ 4 / ln x" and "(\x. f x - ln (ln x) - meissel_mertens) \ O(\x. 1 / ln x)" proof - define r where "r = (\t. (\ t - ln t) / (t * ln t ^ 2))" { fix x :: real assume x: "x > 2" have "((\t. \ t * (-1 / (t * ln t ^ 2))) has_integral \ x * (1 / ln x) - \ 2 * (1 / ln 2) - (\n\real -` {2<..x}. ind prime n * (ln n / real n) * (1 / ln n))) {2..x}" unfolding primes_M_def prime_sum_upto_altdef1 using x by (intro partial_summation_strong[of "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp: power2_eq_square simp flip: has_field_derivative_iff_has_vector_derivative) also have "\ x * (1 / ln x) - \ 2 * (1 / ln 2) - (\n\real -` {2<..x}. ind prime n * (ln n / n) * (1 / ln n)) = \ x / ln x - (\n\insert 2 (real -` {2<..x}). ind prime n * (ln n / n) * (1 / ln n))" (is "_ = _ - ?S") by (subst sum.insert) (auto simp: primes_M_def finite_vimage_real_of_nat_greaterThanAtMost eval_prime_sum_upto) also have "?S = f x" unfolding f_def prime_sum_upto_altdef1 sum_upto_def using x by (intro sum.mono_neutral_cong_left) (auto simp: not_less numeral_2_eq_2 le_Suc_eq) finally have "((\t. -\ t / (t * ln t ^ 2)) has_integral (\ x / ln x - f x)) {2..x}" by simp from has_integral_neg[OF this] have "((\t. \ t / (t * ln t ^ 2)) has_integral (f x - \ x / ln x)) {2..x}" by simp hence "((\t. \ t / (t * ln t ^ 2) - 1 / (t * ln t)) has_integral (f x - \ x / ln x - (ln (ln x) - ln (ln 2)))) {2..x}" using x by (intro has_integral_diff fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros) also have "?this \ (r has_integral (f x - \ x / ln x - (ln (ln x) - ln (ln 2)))) {2..x}" by (intro has_integral_cong) (auto simp: r_def field_simps power2_eq_square) finally have \ . } note integral = this define R\<^sub>\ where "R\<^sub>\ = (\x. \ x - ln x)" have \: "\ x = ln x + R\<^sub>\ x" for x by (simp add: R\<^sub>\_def) define I where "I = (\x. integral {x..} r)" define C where "C = (1 - ln (ln 2) + I 2)" have C_altdef: "C = meissel_mertens" by (simp add: I_def r_def C_def meissel_mertens_def) show bound: " \f x - ln (ln x) - meissel_mertens\ \ 4 / ln x" if x: "x \ 2" for x proof (cases "x = 2") case True hence "\f x - ln (ln x) - meissel_mertens\ = \1 / 2 - ln (ln 2) - meissel_mertens\" by (simp add: f_def eval_prime_sum_upto ) also have "\ \ 2 / ln 2 + 1 / 2" using meissel_mertens_bounds by linarith also have "\ \ 2 / ln 2 + 2 / ln 2" using ln2_le_25_over_36 by (intro add_mono divide_left_mono) auto finally show ?thesis using True by simp next case False hence x: "x > 2" using x by simp have "integral {2..x} r + I x = integral ({2..x} \ {x..}) r" unfolding I_def r_def using x by (intro integral_Un [symmetric] integrable_on_meissel_mertens) (auto simp: max_def r_def) also have "{2..x} \ {x..} = {2..}" using x by auto finally have *: "integral {2..x} r = I 2 - I x" unfolding I_def by simp have eq: "f x - ln (ln x) - C = R\<^sub>\ x / ln x - I x" using integral[OF x] x by (auto simp: C_def field_simps \ has_integral_iff *) also have "\\\ \ \R\<^sub>\ x / ln x\ + norm (I x)" unfolding real_norm_def by (rule abs_triangle_ineq4) also have "\R\<^sub>\ x / ln x\ \ 2 / \ln x\" unfolding R\<^sub>\_def abs_divide using mertens_first_theorem[of x] x by (intro divide_right_mono) auto also have "{x..} - {x<..} = {x}" and "{x<..} \ {x..}" by auto hence "I x = integral {x<..} r" unfolding I_def by (intro integral_subset_negligible [symmetric]) simp_all also have "norm \ \ 2 / ln x" using meissel_mertens_integral_le[of x] x by (simp add: r_def) finally show "\f x - ln (ln x) - meissel_mertens\ \ 4 / ln x" using x by (simp add: C_altdef) qed have "(\x. f x - ln (ln x) - C) \ O(\x. 1 / ln x)" proof (intro landau_o.bigI[of 4] eventually_mono[OF eventually_ge_at_top[of 2]]) fix x :: real assume x: "x \ 2" with bound[OF x] show "norm (f x - ln (ln x) - C) \ 4 * norm (1 / ln x)" by (simp add: C_altdef) qed (auto intro!: add_pos_nonneg) thus "(\x. f x - ln (ln x) - meissel_mertens) \ O(\x. 1 / ln x)" by (simp add: C_altdef) qed corollary prime_harmonic_asymp_equiv: "prime_sum_upto (\p. 1 / p) \[at_top] (\x. ln (ln x))" proof - define f where "f = prime_sum_upto (\p. 1 / p)" have "(\x. f x - ln (ln x) - meissel_mertens + meissel_mertens) \ o(\x. ln (ln x))" unfolding f_def by (rule sum_in_smallo[OF landau_o.big_small_trans[OF mertens_second_theorem(2)]]) real_asymp+ hence "(\x. f x - ln (ln x)) \ o(\x. ln (ln x))" by simp thus ?thesis unfolding f_def by (rule smallo_imp_asymp_equiv) qed text \ As a corollary, we get the divergence of the prime harmonic series. \ corollary prime_harmonic_diverges: "filterlim (prime_sum_upto (\p. 1 / p)) at_top at_top" using asymp_equiv_symI[OF prime_harmonic_asymp_equiv] by (rule asymp_equiv_at_top_transfer) real_asymp (*<*) unbundle no_prime_counting_notation (*>*) end \ No newline at end of file diff --git a/thys/Probabilistic_Noninterference/Language_Semantics.thy b/thys/Probabilistic_Noninterference/Language_Semantics.thy --- a/thys/Probabilistic_Noninterference/Language_Semantics.thy +++ b/thys/Probabilistic_Noninterference/Language_Semantics.thy @@ -1,1608 +1,1607 @@ section \Simple While Language with probabilistic choice and parallel execution\ theory Language_Semantics imports Interface begin subsection \Preliminaries\ text\Trivia\ declare zero_le_mult_iff[simp] declare split_mult_pos_le[simp] declare zero_le_divide_iff[simp] lemma in_minus_Un[simp]: assumes "i \ I" shows "I - {i} Un {i} = I" and "{i} Un (I - {i}) = I" apply(metis Un_commute assms insert_Diff_single insert_absorb insert_is_Un) by (metis assms insert_Diff_single insert_absorb insert_is_Un) lemma less_plus_cases[case_names Left Right]: assumes *: "(i::nat) < n1 \ phi" and **: "\ i2. i = n1 + i2 \ phi" shows phi proof(cases "i < n1") case True thus ?thesis using * by simp next case False hence "n1 \ i" by simp then obtain i2 where "i = n1 + i2" by (metis le_iff_add) thus ?thesis using ** by blast qed lemma less_plus_elim[elim!, consumes 1, case_names Left Right]: assumes i: "(i:: nat) < n1 + n2" and *: "i < n1 \ phi" and **: "\ i2. \i2 < n2; i = n1 + i2\ \ phi" shows phi apply(rule less_plus_cases[of i n1]) using assms by auto lemma nth_append_singl[simp]: "i < length al \ (al @ [a]) ! i = al ! i" by (auto simp add: nth_append) lemma take_append_singl[simp]: assumes "n < length al" shows "take n (al @ [a]) = take n al" using assms by (induct al rule: rev_induct) auto lemma length_unique_prefix: "al1 \ al \ al2 \ al \ length al1 = length al2 \ al1 = al2" by (metis not_equal_is_parallel parallelE prefix_same_cases less_eq_list_def) text\take:\ lemma take_length[simp]: "take (length al) al = al" using take_all by auto lemma take_le: assumes "n < length al" shows "take n al @ [al ! n] \ al" by (metis assms take_Suc_conv_app_nth take_is_prefix less_eq_list_def) lemma list_less_iff_prefix: "a < b \ strict_prefix a b" by (metis le_less less_eq_list_def less_irrefl prefix_order.le_less prefix_order.less_irrefl) lemma take_lt: "n < length al \ take n al < al" unfolding list_less_iff_prefix using prefix_order.le_less[of "take n al" al] by (simp add: take_is_prefix) lemma le_take: assumes "n1 \ n2" shows "take n1 al \ take n2 al" using assms proof(induct al arbitrary: n1 n2) case (Cons a al) thus ?case by (cases n1 n2 rule: nat.exhaust[case_product nat.exhaust]) auto qed auto lemma inj_take: assumes "n1 \ length al" and "n2 \ length al" shows "take n1 al = take n2 al \ n1 = n2" proof- {assume "take n1 al = take n2 al" hence "n1 = n2" using assms proof(induct al arbitrary: n1 n2) case (Cons a al) thus ?case by (cases n1 n2 rule: nat.exhaust[case_product nat.exhaust]) auto qed auto } thus ?thesis by auto qed lemma lt_take: "n1 < n2 \ n2 \ length al \ take n1 al < take n2 al" by (metis inj_take le_less_trans le_take not_less_iff_gr_or_eq order.not_eq_order_implies_strict order.strict_implies_order) text\lsum:\ definition lsum :: "('a \ nat) \ 'a list \ nat" where "lsum f al \ sum_list (map f al)" lemma lsum_simps[simp]: "lsum f [] = 0" "lsum f (al @ [a]) = lsum f al + f a" unfolding lsum_def by auto lemma lsum_append: "lsum f (al @ bl) = lsum f al + lsum f bl" unfolding lsum_def by auto lemma lsum_cong[fundef_cong]: assumes "\ a. a \ set al \ f a = g a" shows "lsum f al = lsum g al" using assms unfolding lsum_def by (metis map_eq_conv) lemma lsum_gt_0[simp]: assumes "al \ []" and "\ a. a \ set al \ 0 < f a" shows "0 < lsum f al" using assms by (induct rule: rev_induct) auto lemma lsum_mono[simp]: assumes "al \ bl" shows "lsum f al \ lsum f bl" proof- obtain cl where bl: "bl = al @ cl" using assms unfolding prefix_def less_eq_list_def by blast thus ?thesis unfolding bl lsum_append by simp qed lemma lsum_mono2[simp]: assumes f: "\ b. b \ set bl \ f b > 0" and le: "al < bl" shows "lsum f al < lsum f bl" proof- obtain cl where bl: "bl = al @ cl" and cl: "cl \ []" using assms unfolding list_less_iff_prefix prefix_def strict_prefix_def by blast have "lsum f al < lsum f al + lsum f cl" using f cl unfolding bl by simp also have "... = lsum f bl" unfolding bl lsum_append by simp finally show ?thesis . qed lemma lsum_take[simp]: "lsum f (take n al) \ lsum f al" by (metis lsum_mono take_is_prefix less_eq_list_def) lemma less_lsum_nchotomy: assumes f: "\ a. a \ set al \ 0 < f a" and i: "(i::nat) < lsum f al" shows "\ n j. n < length al \ j < f (al ! n) \ i = lsum f (take n al) + j" using assms proof(induct rule: rev_induct) case (snoc a al) hence i: "i < lsum f al + f a" and pos: "0 < f a" "\a'. a' \ set al \ 0 < f a'" by auto from i show ?case proof(cases rule: less_plus_elim) case Left then obtain n j where "n < length al \ j < f (al ! n) \ i = lsum f (take n al) + j" using pos snoc by auto thus ?thesis apply - apply(rule exI[of _ n]) apply(rule exI[of _ j]) by auto next case (Right j) thus ?thesis apply - apply(rule exI[of _ "length al"]) apply(rule exI[of _ j]) by simp qed qed auto lemma less_lsum_unique: assumes "\ a. a \ set al \ (0::nat) < f a" and "n1 < length al \ j1 < f (al ! n1)" and "n2 < length al \ j2 < f (al ! n2)" and "lsum f (take n1 al) + j1 = lsum f (take n2 al) + j2" shows "n1 = n2 \ j1 = j2" using assms proof(induct al arbitrary: n1 n2 j1 j2 rule: rev_induct) case (snoc a al) hence pos: "0 < f a" "\ a'. a' \ set al \ 0 < f a'" and n1: "n1 < length al + 1" and n2: "n2 < length al + 1" by auto from n1 show ?case proof(cases rule: less_plus_elim) case Left note n1 = Left hence j1: "j1 < f (al ! n1)" using snoc by auto obtain al' where al: "al = (take n1 al) @ ((al ! n1) # al')" using n1 by (metis append_take_drop_id Cons_nth_drop_Suc) have "j1 < lsum f ((al ! n1) # al')" using pos j1 unfolding lsum_def by simp hence "lsum f (take n1 al) + j1 < lsum f (take n1 al) + lsum f ((al ! n1) # al')" by simp also have "... = lsum f al" unfolding lsum_append[THEN sym] using al by simp finally have lsum1: "lsum f (take n1 al) + j1 < lsum f al" . from n2 show ?thesis proof(cases rule: less_plus_elim) case Left note n2 = Left hence j2: "j2 < f (al ! n2)" using snoc by auto show ?thesis apply(rule snoc(1)) using snoc using pos n1 j1 n2 j2 by auto next case Right hence n2: "n2 = length al" by simp hence j2: "j2 < f a" using snoc by simp have "lsum f (take n1 al) + j1 = lsum f al + j2" using n1 n2 snoc by simp hence False using lsum1 by auto thus ?thesis by simp qed next case Right hence n1: "n1 = length al" by simp hence j1: "j1 < f a" using snoc by simp from n2 show ?thesis proof(cases rule: less_plus_elim) case Left note n2 = Left hence j2: "j2 < f (al ! n2)" using snoc by auto obtain al' where al: "al = (take n2 al) @ ((al ! n2) # al')" using n2 by (metis append_take_drop_id Cons_nth_drop_Suc) have "j2 < lsum f ((al ! n2) # al')" using pos j2 unfolding lsum_def by simp hence "lsum f (take n2 al) + j2 < lsum f (take n2 al) + lsum f ((al ! n2) # al')" by simp also have "... = lsum f al" unfolding lsum_append[THEN sym] using al by simp finally have lsum2: "lsum f (take n2 al) + j2 < lsum f al" . have "lsum f al + j1 = lsum f (take n2 al) + j2" using n1 n2 snoc by simp hence False using lsum2 by auto thus ?thesis by simp next case Right hence n2: "n2 = length al" by simp have "j1 = j2" using n1 n2 snoc by simp thus ?thesis using n1 n2 by simp qed qed qed auto definition locate_pred where "locate_pred f al (i::nat) n_j \ fst n_j < length al \ snd n_j < f (al ! (fst n_j)) \ i = lsum f (take (fst n_j) al) + (snd n_j)" definition locate where "locate f al i \ SOME n_j. locate_pred f al i n_j" definition locate1 where "locate1 f al i \ fst (locate f al i)" definition locate2 where "locate2 f al i \ snd (locate f al i)" lemma locate_pred_ex: assumes "\ a. a \ set al \ 0 < f a" and "i < lsum f al" shows "\ n_j. locate_pred f al i n_j" using assms less_lsum_nchotomy unfolding locate_pred_def by force lemma locate_pred_unique: assumes "\ a. a \ set al \ 0 < f a" and "locate_pred f al i n1_j1" "locate_pred f al i n2_j2" shows "n1_j1 = n2_j2" using assms less_lsum_unique unfolding locate_pred_def apply(cases n1_j1, cases n2_j2) apply simp by blast lemma locate_locate_pred: assumes "\ a. a \ set al \ (0::nat) < f a" and "i < lsum f al" shows "locate_pred f al i (locate f al i)" proof- obtain n_j where "locate_pred f al i n_j" using assms locate_pred_ex[of al f i] by auto thus ?thesis unfolding locate_def apply(intro someI[of "locate_pred f al i"]) by blast qed lemma locate_locate_pred_unique: assumes "\ a. a \ set al \ (0::nat) < f a" and "locate_pred f al i n_j" shows "n_j = locate f al i" unfolding locate_def apply(rule sym, rule some_equality) using assms locate_locate_pred apply force using assms locate_pred_unique by blast lemma locate: assumes "\ a. a \ set al \ 0 < f a" and "i < lsum f al" shows "locate1 f al i < length al \ locate2 f al i < f (al ! (locate1 f al i)) \ i = lsum f (take (locate1 f al i) al) + (locate2 f al i)" using assms locate_locate_pred unfolding locate1_def locate2_def locate_pred_def by auto lemma locate_unique: assumes "\ a. a \ set al \ 0 < f a" and "n < length al" and "j < f (al ! n)" and "i = lsum f (take n al) + j" shows "n = locate1 f al i \ j = locate2 f al i" proof- define n_j where "n_j = (n,j)" have "locate_pred f al i n_j" using assms unfolding n_j_def locate_pred_def by auto hence "n_j = locate f al i" using assms locate_locate_pred_unique by blast thus ?thesis unfolding n_j_def locate1_def locate2_def by (metis fst_conv n_j_def snd_conv) qed text\sum:\ lemma sum_2[simp]: "sum f {..< 2} = f 0 + f (Suc 0)" proof- have "{..< 2} = {0, Suc 0}" by auto thus ?thesis by force qed lemma inj_Plus[simp]: "inj ((+) (a::nat))" unfolding inj_on_def by auto lemma inj_on_Plus[simp]: "inj_on ((+) (a::nat)) A" unfolding inj_on_def by auto lemma Plus_int[simp]: fixes a :: nat shows "(+) b ` {..< a} = {b ..< b + a}" proof safe fix x::nat assume "x \ {b..< b + a}" hence "b \ x" and x: "x < a + b" by auto then obtain y where xb: "x = b + y" by (metis le_iff_add) hence "y < a" using x by simp thus "x \ (+) b ` {.. n. n \ N \ finite (B1 n)" and int: "\ m n. {m, n} \ N \ m \ n \ B1 m \ B1 n = {}" and ss: "\ n. n \ N \ sum f1 (B1 n) = b2 n" shows "sum f1 A1 = a2" (is "?L = a2") proof- have "?L = sum (%n. sum f1 (B1 n)) N" unfolding A1 using sum.UNION_disjoint[of N B1 f1] fin int by simp also have "... = sum b2 N" using ss fin by auto also have "... = a2" using a2 by simp finally show ?thesis . qed lemma sum_UN_intro: assumes A1: "A1 = (UN n : N. B1 n)" and A2: "A2 = (UN n : N. B2 n)" and fin: "finite N" "\ n. n \ N \ finite (B1 n) \ finite (B2 n)" and int: "\ m n. {m, n} \ N \ m \ n \ B1 m \ B1 n = {}" "\ m n. {m, n} \ N \ B2 m \ B2 n = {}" and ss: "\ n. n \ N \ sum f1 (B1 n) = sum f2 (B2 n)" shows "sum f1 A1 = sum f2 A2" (is "?L = ?R") proof- have "?L = sum (%n. sum f1 (B1 n)) N" unfolding A1 using sum.UNION_disjoint[of N B1 f1] fin int by simp also have "... = sum (%n. sum f2 (B2 n)) N" using ss fin sum.mono_neutral_left by auto also have "... = ?R" unfolding A2 using sum.UNION_disjoint[of N B2 f2] fin int by simp finally show ?thesis . qed lemma sum_Minus_intro: fixes f1 :: "'a1 \ real" and f2 :: "'a2 \ real" assumes "B1 = A1 - {a1}" and "B2 = A2 - {a2}" and "a1 : A1" and "a2 : A2" and "finite A1" and "finite A2" "sum f1 A1 = sum f2 A2" and "f1 a1 = f2 a2" shows "sum f1 B1 = sum f2 B2" proof- have 1: "A1 = B1 Un {a1}" and 2: "A2 = B2 Un {a2}" using assms by blast+ from assms have "a1 \ B1" by simp with 1 \finite A1\ have "sum f1 A1 = sum f1 B1 + f1 a1" by simp hence 3: "sum f1 B1 = sum f1 A1 - f1 a1" by simp from assms have "a2 \ B2" by simp with 2 \finite A2 \have "sum f2 A2 = sum f2 B2 + f2 a2" by simp hence "sum f2 B2 = sum f2 A2 - f2 a2" by simp thus ?thesis using 3 assms by simp qed lemma sum_singl_intro: assumes "b = f a" and "finite A" and "a \ A" and "\ a'. \a' \ A; a' \ a\ \ f a' = 0" shows "sum f A = b" proof- define B where "B = A - {a}" have "A = B Un {a}" unfolding B_def using assms by blast moreover have "B Int {a} = {}" unfolding B_def by blast ultimately have "sum f A = sum f B + sum f {a}" using assms sum.union_disjoint by blast moreover have "\ b \ B. f b = 0" using assms unfolding B_def by auto ultimately show ?thesis using assms by auto qed lemma sum_all0_intro: assumes "b = 0" and "\ a. a \ A \ f a = 0" shows "sum f A = b" -apply(cases "finite A") -by(metis assms sum.neutral sum.infinite)+ + using assms by simp lemma sum_1: assumes I: "finite I" and ss: "sum f I = 1" and i: "i \ I - I1" and I1: "I1 \ I" and f: "\i. i \ I \ (0::real) \ f i" shows "f i \ 1 - sum f I1" proof- have "sum f I = sum f ({i} Un (I - {i}))" using i by (metis DiffE insert_Diff_single insert_absorb insert_is_Un) also have "... = sum f {i} + sum f (I - {i})" apply(rule sum.union_disjoint) using I by auto finally have "1 = f i + sum f (I - {i})" unfolding ss[THEN sym] by simp moreover have "sum f (I - {i}) \ sum f I1" apply(rule sum_mono2) using assms by auto ultimately have "1 \ f i + sum f I1" by auto thus ?thesis by auto qed subsection \Syntax\ datatype ('test, 'atom, 'choice) cmd = Done | Atm "'atom" | Seq "('test, 'atom, 'choice) cmd" "('test, 'atom, 'choice) cmd" ("_ ;; _" [60, 61] 60) | While "'test" "('test, 'atom, 'choice) cmd" | Ch 'choice "('test, 'atom, 'choice) cmd" "('test, 'atom, 'choice) cmd" | Par "('test, 'atom, 'choice) cmd list" | ParT "('test, 'atom, 'choice) cmd list" (* Commands containing no while loops: *) fun noWhile where "noWhile Done \ True" | "noWhile (Atm atm) \ True" | "noWhile (c1 ;; c2) \ noWhile c1 \ noWhile c2" | "noWhile (While tst c) \ False" | "noWhile (Ch ch c1 c2) \ noWhile c1 \ noWhile c2" | "noWhile (Par cl) \ (\ c \ set cl. noWhile c)" | "noWhile (ParT cl) \ (\ c \ set cl. noWhile c)" (* ``Finished" commands: *) fun finished where "finished Done \ True" | "finished (Atm atm) \ False" | "finished (c1 ;; c2) \ False" | "finished (While tst c) \ False" | "finished (Ch ch c1 c2) \ False" | "finished (Par cl) \ (\ c \ set cl. finished c)" | "finished (ParT cl) \ (\ c \ set cl. finished c)" definition noWhileL where "noWhileL cl \ \ c \ set cl. noWhile c" lemma fin_Par_noWhileL[simp]: "noWhile (Par cl) \ noWhileL cl" unfolding noWhileL_def by simp lemma fin_ParT_noWhileL[simp]: "noWhile (ParT cl) \ noWhileL cl" unfolding noWhileL_def by simp declare noWhile.simps(6) [simp del] declare noWhile.simps(7) [simp del] lemma noWhileL_intro[intro]: assumes "\ c. c \ set cl \ noWhile c" shows "noWhileL cl" using assms unfolding noWhileL_def by auto lemma noWhileL_fin[simp]: assumes "noWhileL cl" and "c \ set cl" shows "noWhile c" using assms unfolding noWhileL_def by simp lemma noWhileL_update[simp]: assumes cl: "noWhileL cl" and c': "noWhile c'" shows "noWhileL (cl[n := c'])" proof(cases "n < length cl") case True show ?thesis unfolding noWhileL_def proof safe fix c assume "c \ set (cl[n := c'])" hence "c \ insert c' (set cl)" using set_update_subset_insert by fastforce thus "noWhile c" using assms by (cases "c = c'") auto qed qed (insert cl, auto) definition finishedL where "finishedL cl \ \ c \ set cl. finished c" lemma finished_Par_finishedL[simp]: "finished (Par cl) \ finishedL cl" unfolding finishedL_def by simp lemma finished_ParT_finishedL[simp]: "finished (ParT cl) \ finishedL cl" unfolding finishedL_def by simp declare finished.simps(6) [simp del] declare finished.simps(7) [simp del] lemma finishedL_intro[intro]: assumes "\ c. c \ set cl \ finished c" shows "finishedL cl" using assms unfolding finishedL_def by auto lemma finishedL_finished[simp]: assumes "finishedL cl" and "c \ set cl" shows "finished c" using assms unfolding finishedL_def by simp lemma finishedL_update[simp]: assumes cl: "finishedL cl" and c': "finished c'" shows "finishedL (cl[n := c'])" proof(cases "n < length cl") case True show ?thesis unfolding finishedL_def proof safe fix c assume "c \ set (cl[n := c'])" hence "c \ insert c' (set cl)" using set_update_subset_insert by fastforce thus "finished c" using assms by (cases "c = c'") auto qed qed (insert cl, auto) lemma finished_fin[simp]: "finished c \ noWhile c" by(induct c) auto lemma finishedL_noWhileL[simp]: "finishedL cl \ noWhileL cl" unfolding finishedL_def noWhileL_def by auto locale PL = fixes aval :: "'atom \ 'state \ 'state" and tval :: "'test \ 'state \ bool" and cval :: "'choice \ 'state \ real" assumes properCh: "\ ch s. 0 \ cval ch s \ cval ch s \ 1" begin lemma [simp]: "(n::nat) < N \ 0 \ 1 / N" by auto lemma [simp]: "(n::nat) < N \ 1 / N \ 1" by auto lemma [simp]: "(n::nat) < N \ 0 \ 1 - 1 / N" by (simp add: divide_simps) lemma sum_equal: "0 < (N::nat) \ sum (\ n. 1/N) {..< N} = 1" unfolding sum_constant by auto fun proper where "proper Done \ True" | "proper (Atm x) \ True" | "proper (Seq c1 c2) \ proper c1 \ proper c2" | "proper (While tst c) \ proper c" | "proper (Ch ch c1 c2) \ proper c1 \ proper c2" | "proper (Par cl) \ cl \ [] \ (\ c \ set cl. proper c)" | "proper (ParT cl) \ cl \ [] \ (\ c \ set cl. proper c)" definition properL where "properL cl \ cl \ [] \ (\ c \ set cl. proper c)" lemma proper_Par_properL[simp]: "proper (Par cl) \ properL cl" unfolding properL_def by simp lemma proper_ParT_properL[simp]: "proper (ParT cl) \ properL cl" unfolding properL_def by simp declare proper.simps(6) [simp del] declare proper.simps(7) [simp del] lemma properL_intro[intro]: "\cl \ []; \ c. c \ set cl \ proper c\ \ properL cl" unfolding properL_def by auto lemma properL_notEmp[simp]: "properL cl \ cl \ []" unfolding properL_def by simp lemma properL_proper[simp]: "\properL cl; c \ set cl\ \ proper c" unfolding properL_def by simp lemma properL_update[simp]: assumes cl: "properL cl" and c': "proper c'" shows "properL (cl[n := c'])" proof(cases "n < length cl") case True show ?thesis unfolding properL_def proof safe fix c assume "c \ set (cl[n := c'])" hence "c \ insert c' (set cl)" using set_update_subset_insert by fastforce thus "proper c" using assms by (cases "c = c'") auto qed (insert cl, auto) qed (insert cl, auto) lemma proper_induct[consumes 1, case_names Done Atm Seq While Ch Par ParT]: assumes *: "proper c" and Done: "phi Done" and Atm: "\ atm. phi (Atm atm)" and Seq: "\ c1 c2. \phi c1; phi c2\ \ phi (c1 ;; c2)" and While: "\ tst c. phi c \ phi (While tst c)" and Ch: "\ ch c1 c2. \phi c1; phi c2\ \ phi (Ch ch c1 c2)" and Par: "\ cl. \properL cl; \ c. c \ set cl \ phi c\ \ phi (Par cl)" and ParT: "\ cl. \properL cl; \ c. c \ set cl \ phi c\ \ phi (ParT cl)" shows "phi c" using * apply(induct c) using assms unfolding properL_def by auto subsubsection \Operational Small-Step Semantics\ (* "The Finished Threads": The sublist of finished threads from a list of threads *) definition "theFT cl \ {n. n < length cl \ finished (cl!n)}" (* "The NopnFinished Threads": *) definition "theNFT cl \ {n. n < length cl \ \ finished (cl!n)}" lemma finite_theFT[simp]: "finite (theFT cl)" unfolding theFT_def by simp lemma theFT_length[simp]: "n \ theFT cl \ n < length cl" unfolding theFT_def by simp lemma theFT_finished[simp]: "n \ theFT cl \ finished (cl!n)" unfolding theFT_def by simp lemma finite_theNFT[simp]: "finite (theNFT cl)" unfolding theNFT_def by simp lemma theNFT_length[simp]: "n \ theNFT cl \ n < length cl" unfolding theNFT_def by simp lemma theNFT_notFinished[simp]: "n \ theNFT cl \ \ finished (cl!n)" unfolding theNFT_def by simp lemma theFT_Int_theNFT[simp]: "theFT cl Int theNFT cl = {}" and "theNFT cl Int theFT cl = {}" unfolding theFT_def theNFT_def by auto lemma theFT_Un_theNFT[simp]: "theFT cl Un theNFT cl = {..< length cl}" and "theNFT cl Un theFT cl = {..< length cl}" unfolding theFT_def theNFT_def by auto lemma in_theFT_theNFT[simp]: assumes "n1 \ theFT cl" and "n2 \ theNFT cl" shows "n1 \ n2" and "n2 \ n1" using assms theFT_Int_theNFT by blast+ (* The cumulated weight of the finished threads: *) definition "WtFT cl \ sum (\ (n::nat). 1/(length cl)) (theFT cl)" (* The cumulated weight of the non-finished threads: *) definition "WtNFT cl \ sum (\ (n::nat). 1/(length cl)) (theNFT cl)" lemma WtFT_WtNFT[simp]: assumes "0 < length cl" shows "WtFT cl + WtNFT cl = 1" (is "?A = 1") proof- let ?w = "\ n. 1 / length cl" let ?L = "theFT cl" let ?Lnot = "theNFT cl" have 1: "{..< length cl} = ?L Un ?Lnot" by auto have "?A = sum ?w ?L + sum ?w ?Lnot" unfolding WtFT_def WtNFT_def by simp also have "... = sum ?w {..< length cl}" unfolding 1 apply(rule sum.union_disjoint[THEN sym]) by auto also have "... = 1" unfolding sum_equal[OF assms] by auto finally show ?thesis . qed lemma WtNFT_1_WtFT: "0 < length cl \ WtNFT cl = 1 - WtFT cl" by (simp add: algebra_simps) lemma WtNFT_WtFT_1[simp]: assumes "0 < length cl" and "WtFT cl \ 1" shows "WtNFT cl / (1 - WtFT cl) = 1" (is "?A / ?B = 1") proof- have A: "?A = ?B" using assms WtNFT_1_WtFT by auto show ?thesis unfolding A using assms by auto qed lemma WtFT_ge_0[simp]: "WtFT cl \ 0" unfolding WtFT_def by (rule sum_nonneg) simp lemma WtFT_le_1[simp]: "WtFT cl \ 1" (is "?L \ 1") proof- let ?N = "length cl" have "?L \ sum (\ n::nat. 1/?N) {..< ?N}" unfolding WtFT_def apply(rule sum_mono2) by auto also have "... \ 1" by (metis div_by_0 le_cases neq0_conv not_one_le_zero of_nat_0 sum.neutral sum_equal) finally show ?thesis . qed lemma le_1_WtFT[simp]: "0 \ 1 - WtFT cl" (is "0 \ ?R") proof- have a: "-1 \ - WtFT cl" by simp have "(0 :: real) = 1 + (-1)" by simp also have "1 + (-1) \ 1 + (- WtFT cl)" using a by arith also have "... = ?R" by simp finally show ?thesis . qed lemma WtFT_lt_1[simp]: "WtFT cl \ 1 \ WtFT cl < 1" using WtFT_le_1 by (auto simp add: le_less) lemma lt_1_WtFT[simp]: "WtFT cl \ 1 \ 0 < 1 - WtFT cl" using le_1_WtFT by (metis le_1_WtFT eq_iff_diff_eq_0 less_eq_real_def) lemma notFinished_WtFT[simp]: assumes "n < length cl" and "\ finished (cl ! n)" shows "1 / length cl \ 1 - WtFT cl" proof- have "0 < length cl" using assms by auto thus ?thesis unfolding WtFT_def apply(intro sum_1[of "{..< length cl}"]) using assms by auto qed (* The branching of a command: *) fun brn :: "('test, 'atom, 'choice) cmd \ nat" where "brn Done = 1" | "brn (Atm atm) = 1" | "brn (c1 ;; c2) = brn c1" | "brn (While tst c) = 1" | "brn (Ch ch c1 c2) = 2" | "brn (Par cl) = lsum brn cl" | "brn (ParT cl) = lsum brn cl" lemma brn_gt_0: "proper c \ 0 < brn c" by (induct rule: proper_induct) auto lemma brn_gt_0_L: "\properL cl; c \ set cl\ \ 0 < brn c" by (metis brn_gt_0 properL_def) (* The locate-thread and locate-index operators. Given a thread list cl with n = length cl and i < (\ l < length cl . brn cl), (locateT cl i, locateI cl i) are (k, j) from the paper's Figure 1. *) definition "locateT \ locate1 brn" definition "locateI \ locate2 brn" definition "brnL cl n \ lsum brn (take n cl)" lemma brnL_lsum: "brnL cl (length cl) = lsum brn cl" unfolding brnL_def by simp lemma brnL_unique: assumes "properL cl" and "n1 < length cl \ j1 < brn (cl ! n1)" and "n2 < length cl \ j2 < brn (cl ! n2)" and "brnL cl n1 + j1 = brnL cl n2 + j2" shows "n1 = n2 \ j1 = j2" apply (rule less_lsum_unique) using assms brn_gt_0 unfolding brnL_def properL_def by auto lemma brn_Par_simp[simp]: "brn (Par cl) = brnL cl (length cl)" unfolding brnL_lsum by simp lemma brn_ParT_simp[simp]: "brn (ParT cl) = brnL cl (length cl)" unfolding brnL_lsum by simp declare brn.simps(6)[simp del] declare brn.simps(7)[simp del] lemma brnL_0[simp]: "brnL cl 0 = 0" unfolding brnL_def by auto lemma brnL_Suc[simp]: "n < length cl \ brnL cl (Suc n) = brnL cl n + brn (cl ! n)" unfolding brnL_def using take_Suc_conv_app_nth[of n cl] by simp lemma brnL_mono[simp]: "n1 \ n2 \ brnL cl n1 \ brnL cl n2" using le_take[of n1 n2 cl] unfolding brnL_def by simp lemma brnL_mono2[simp]: assumes p: "properL cl" and n: "n1 < n2" and l: "n2 \ length cl" shows "brnL cl n1 < brnL cl n2" (is "?L < ?R") proof- have 1: "\c. c \ set (take n2 cl) \ 0 < brn c" using p by (metis brn_gt_0 in_set_takeD properL_proper) have "take n1 cl < take n2 cl" using n l lt_take by auto hence "lsum brn (take n1 cl) < lsum brn (take n2 cl)" using lsum_mono2[of "take n2 cl" "%c. brn c" "take n1 cl"] 1 by simp thus ?thesis unfolding brnL_def . qed lemma brn_index[simp]: assumes n: "n < length cl" and i: "i < brn (cl ! n)" shows "brnL cl n + i < brnL cl (length cl)" (is "?L < ?R") proof- have "?L < brnL cl (Suc n)" using assms by simp also have "... \ ?R" using n brnL_mono[of "Suc n" "length cl" cl] by simp finally show ?thesis . qed lemma brnL_gt_0[simp]: "\properL cl; 0 < n\ \ 0 < brnL cl n" by (metis properL_def brnL_mono brnL_mono2 le_0_eq length_greater_0_conv nat_le_linear neq0_conv) lemma locateTI: assumes "properL cl" and "ii < brn (Par cl)" shows "locateT cl ii < length cl \ locateI cl ii < brn (cl ! (locateT cl ii)) \ ii = brnL cl (locateT cl ii) + locateI cl ii" -using assms locate[of cl brn] brn_gt_0 +using assms locate[of cl brn ii] brn_gt_0 unfolding locateT_def locateI_def brnL_def unfolding brnL_lsum[THEN sym] by auto lemma locateTI_unique: assumes "properL cl" and "n < length cl" and "i < brn (cl ! n)" and "ii = brnL cl n + i" shows "n = locateT cl ii \ i = locateI cl ii" using assms locate_unique[of cl brn] brn_gt_0 unfolding locateT_def locateI_def brnL_def unfolding brnL_lsum[THEN sym] by auto (* pickFT picks a finished thread (if there is any). It will be used to perform a dummy transition in case the cumulated weight of the finished threads is 0; specifically, one will assign probability 1 to that picked fresh. (Obviously, the particular choice does not matter.) *) definition pickFT_pred where "pickFT_pred cl n \ n < length cl \ finished (cl!n)" definition pickFT where "pickFT cl \ SOME n. pickFT_pred cl n" lemma pickFT_pred: assumes "WtFT cl = 1" shows "\ n. pickFT_pred cl n" proof(rule ccontr, unfold not_ex) assume "\n. \ pickFT_pred cl n" hence "\ n. n < length cl \ \ finished (cl!n)" unfolding pickFT_pred_def by auto hence "theFT cl = {}" unfolding theFT_def by auto hence "WtFT cl = 0" unfolding WtFT_def by simp thus False using assms by simp qed lemma pickFT_pred_pickFT: "WtFT cl = 1 \ pickFT_pred cl (pickFT cl)" unfolding pickFT_def by(auto intro: someI_ex pickFT_pred) lemma pickFT_length[simp]: "WtFT cl = 1 \ pickFT cl < length cl" using pickFT_pred_pickFT unfolding pickFT_pred_def by auto lemma pickFT_finished[simp]: "WtFT cl = 1 \ finished (cl ! (pickFT cl))" using pickFT_pred_pickFT unfolding pickFT_pred_def by auto lemma pickFT_theFT[simp]: "WtFT cl = 1 \ pickFT cl \ theFT cl" unfolding theFT_def by auto (* The weight, continuation and effect defined as a single operator: *) fun wt_cont_eff where "wt_cont_eff Done s i = (1, Done, s)" | "wt_cont_eff (Atm atm) s i = (1, Done, aval atm s)" | "wt_cont_eff (c1 ;; c2) s i = (case wt_cont_eff c1 s i of (x, c1', s') \ if finished c1' then (x, c2, s') else (x, c1' ;; c2, s'))" | "wt_cont_eff (While tst c) s i = (if tval tst s then (1, c ;; (While tst c), s) else (1, Done, s))" | "wt_cont_eff (Ch ch c1 c2) s i = (if i = 0 then cval ch s else 1 - cval ch s, if i = 0 then c1 else c2, s)" | "wt_cont_eff (Par cl) s ii = (if cl ! (locateT cl ii) \ set cl then (case wt_cont_eff (cl ! (locateT cl ii)) s (locateI cl ii) of (w, c', s') \ ((1 / (length cl)) * w, Par (cl [(locateT cl ii) := c']), s')) else undefined)" | "wt_cont_eff (ParT cl) s ii = (if cl ! (locateT cl ii) \ set cl then (case wt_cont_eff (cl ! (locateT cl ii)) s (locateI cl ii) of (w, c', s') \ (if WtFT cl = 1 then (if locateT cl ii = pickFT cl \ locateI cl ii = 0 then 1 else 0) else if finished (cl ! (locateT cl ii)) then 0 else (1 / (length cl)) / (1 - WtFT cl) * w, ParT (cl [(locateT cl ii) := c']), s')) else undefined)" (* weight, cont (transition) and effect: *) definition wt where "wt c s i = fst (wt_cont_eff c s i)" definition cont where "cont c s i = fst (snd (wt_cont_eff c s i))" definition eff where "eff c s i = snd (snd (wt_cont_eff c s i))" (* Their individual equations (corresponding to the paper's Figure 1): *) lemma wt_Done[simp]: "wt Done s i = 1" unfolding wt_def by simp lemma wt_Atm[simp]: "wt (Atm atm) s i = 1" unfolding wt_def by simp lemma wt_Seq[simp]: "wt (c1 ;; c2) s = wt c1 s" proof- {fix i have "wt (c1 ;; c2) s i = wt c1 s i" proof(cases "wt_cont_eff c1 s i") case (fields _ c1' _) thus ?thesis unfolding wt_def by(cases c1', auto) qed } thus ?thesis by auto qed lemma wt_While[simp]: "wt (While tst c) s i = 1" unfolding wt_def by simp lemma wt_Ch_L[simp]: "wt (Ch ch c1 c2) s 0 = cval ch s" unfolding wt_def by simp lemma wt_Ch_R[simp]: "wt (Ch ch c1 c2) s (Suc n) = 1 - cval ch s" unfolding wt_def by simp (* *) lemma cont_Done[simp]: "cont Done s i = Done" unfolding cont_def by simp lemma cont_Atm[simp]: "cont (Atm atm) s i = Done" unfolding cont_def by simp lemma cont_Seq_finished[simp]: "finished (cont c1 s i) \ cont (c1 ;; c2) s i = c2" unfolding cont_def by(cases "wt_cont_eff c1 s i") auto lemma cont_Seq_notFinished[simp]: assumes "\ finished (cont c1 s i)" shows "cont (c1 ;; c2) s i = (cont c1 s i) ;; c2" proof(cases "wt_cont_eff c1 s i") case (fields _ c1' _) thus ?thesis using assms unfolding cont_def by(cases c1') auto qed lemma cont_Seq_not_eq_finished[simp]: "\ finished c2 \ \ finished (cont (Seq c1 c2) s i)" by (cases "finished (cont c1 s i)") auto lemma cont_While_False[simp]: "tval tst s = False \ cont (While tst c) s i = Done" unfolding cont_def by simp lemma cont_While_True[simp]: "tval tst s = True \ cont (While tst c) s i = c ;; (While tst c)" unfolding cont_def by simp lemma cont_Ch_L[simp]: "cont (Ch ch c1 c2) s 0 = c1" unfolding cont_def by simp lemma cont_Ch_R[simp]: "cont (Ch ch c1 c2) s (Suc n) = c2" unfolding cont_def by simp (* *) lemma eff_Done[simp]: "eff Done s i = s" unfolding eff_def by simp lemma eff_Atm[simp]: "eff (Atm atm) s i = aval atm s" unfolding eff_def by simp lemma eff_Seq[simp]: "eff (c1 ;; c2) s = eff c1 s" proof- {fix i have "eff (c1 ;; c2) s i = eff c1 s i" proof(cases "wt_cont_eff c1 s i") case (fields _ c1' _) thus ?thesis unfolding eff_def by(cases c1') auto qed } thus ?thesis by auto qed lemma eff_While[simp]: "eff (While tst c) s i = s" unfolding eff_def by simp lemma eff_Ch[simp]: "eff (Ch ch c1 c2) s i = s" unfolding eff_def by simp (* wt-cont-eff simps for parallel composition *) lemma brnL_nchotomy: assumes "properL cl" and "ii < brnL cl (length cl)" shows "\ n i. n < length cl \ i < brn (cl ! n) \ ii = brnL cl n + i" unfolding brnL_def apply(rule less_lsum_nchotomy) using assms brn_gt_0 unfolding brnL_lsum[THEN sym] by auto corollary brnL_cases[consumes 2, case_names Local, elim]: assumes "properL cl" and "ii < brnL cl (length cl)" and "\ n i. \n < length cl; i < brn (cl ! n); ii = brnL cl n + i\ \ phi" shows phi using assms brnL_nchotomy by blast lemma wt_cont_eff_Par[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" shows "wt (Par cl) s (brnL cl n + i) = 1 / (length cl) * wt (cl ! n) s i" (is "?wL = ?wR") (* *) "cont (Par cl) s (brnL cl n + i) = Par (cl [n := cont (cl ! n) s i])" (is "?mL = ?mR") (* *) "eff (Par cl) s (brnL cl n + i) = eff (cl ! n) s i" (is "?eL = ?eR") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have "?wL = 1 / (length cl)* wt (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding wt_def by simp also have "... = ?wR" unfolding n_i by simp finally show "?wL = ?wR" . (* *) have "?mL = Par (cl [n1 := cont (cl ! n1) s i1])" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding cont_def by simp also have "... = ?mR" unfolding n_i by simp finally show "?mL = ?mR" . (* *) have "?eL = eff (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding eff_def by simp also have "eff (cl ! n1) s i1 = ?eR" unfolding n_i by simp finally show "?eL = ?eR" . qed lemma cont_eff_ParT[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" shows "cont (ParT cl) s (brnL cl n + i) = ParT (cl [n := cont (cl ! n) s i])" (is "?mL = ?mR") (* *) "eff (ParT cl) s (brnL cl n + i) = eff (cl ! n) s i" (is "?eL = ?eR") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have "?mL = ParT (cl [n1 := cont (cl ! n1) s i1])" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding cont_def by simp also have "... = ?mR" unfolding n_i by simp finally show "?mL = ?mR" . (* *) have "?eL = eff (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using set unfolding n1_def i1_def unfolding eff_def by simp also have "eff (cl ! n1) s i1 = ?eR" unfolding n_i by simp finally show "?eL = ?eR" . qed lemma wt_ParT_WtFT_pickFT_0[simp]: assumes p: "properL cl" and WtFT: "WtFT cl = 1" shows "wt (ParT cl) s (brnL cl (pickFT cl)) = 1" (is "?wL = 1") proof- define ii where "ii = brnL cl (pickFT cl)" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have ni: "pickFT cl < length cl" "0 < brn (cl! (pickFT cl))" using WtFT p brn_gt_0 by auto hence n_i: "pickFT cl = n1" "0 = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique[of cl "pickFT cl" 0 ii] by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using ni by (metis add.comm_neutral brn_index) hence "n1 < length cl" unfolding n1_def using ni p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have n1i1: "n1 = pickFT cl \ i1 = 0" using WtFT ni unfolding n_i by auto show "?wL = 1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT n1i1 set unfolding n1_def i1_def unfolding wt_def by simp qed lemma wt_ParT_WtFT_notPickFT_0[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" and WtFT: "WtFT cl = 1" and ni: "n = pickFT cl \ i \ 0" shows "wt (ParT cl) s (brnL cl n + i) = 0" (is "?wL = 0") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have n1i1: "n1 \ pickFT cl \ i1 \ 0" using WtFT ni unfolding n_i by auto show "?wL = 0" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT n1i1 set unfolding n1_def i1_def unfolding wt_def by force qed lemma wt_ParT_notWtFT_finished[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" and WtFT: "WtFT cl \ 1" and f: "finished (cl ! n)" shows "wt (ParT cl) s (brnL cl n + i) = 0" (is "?wL = 0") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have f: "finished (cl ! n1)" using f unfolding n_i by auto show "?wL = 0" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT f set unfolding n1_def i1_def unfolding wt_def by simp qed lemma wt_cont_eff_ParT_notWtFT_notFinished[simp]: assumes p: "properL cl" and n: "n < length cl" and i: "i < brn (cl ! n)" and WtFT: "WtFT cl \ 1" and nf: "\ finished (cl ! n)" shows "wt (ParT cl) s (brnL cl n + i) = (1 / (length cl)) / (1 - WtFT cl) * wt (cl ! n) s i" (is "?wL = ?wR") proof- define ii where "ii = brnL cl n + i" define n1 where "n1 = locateT cl ii" define i1 where "i1 = locateI cl ii" have n_i: "n = n1" "i = i1" using p unfolding n1_def i1_def using ii_def locateTI_unique n i by auto have lsum1: "ii < brnL cl (length cl)" unfolding ii_def using n i by simp hence "n1 < length cl" unfolding n1_def using i p locateTI[of cl ii] by simp hence set: "cl ! n1 \ set cl" by simp (* *) have nf: "\ finished (cl ! n1)" using nf unfolding n_i by auto have "?wL = (1 / (length cl)) / (1 - WtFT cl) * wt (cl ! n1) s i1" unfolding ii_def[THEN sym] apply (cases "wt_cont_eff (cl ! n1) s i1") using WtFT nf set unfolding n1_def i1_def unfolding wt_def by simp also have "... = ?wR" unfolding n_i by simp finally show "?wL = ?wR" . qed (* *) lemma wt_ge_0[simp]: assumes "proper c" and "i < brn c" shows "0 \ wt c s i" using assms proof (induct c arbitrary: i s rule: proper_induct) case (Ch ch c1 c2) thus ?case using properCh by (cases i) (auto simp: algebra_simps) next case (Par cl ii) have "properL cl" and "ii < brnL cl (length cl)" using Par by auto thus ?case apply (cases rule: brnL_cases) using Par by simp next case (ParT cl ii) have "properL cl" and "ii < brnL cl (length cl)" using ParT by auto thus ?case proof(cases rule: brnL_cases) case (Local n i) show ?thesis proof (cases "WtFT cl = 1") case True thus ?thesis using Local ParT by (cases "n = pickFT cl \ i = 0") auto next case False thus ?thesis using Local ParT by (cases "finished (cl ! n)") auto qed qed qed auto lemma wt_le_1[simp]: assumes "proper c" and "i < brn c" shows "wt c s i \ 1" using assms proof (induct c arbitrary: i s rule: proper_induct) case (Ch ch c1 c2) thus ?case using properCh by (cases i) auto next case (Par cl ii) hence "properL cl" and "ii < brnL cl (length cl)" by auto thus ?case apply (cases rule: brnL_cases) apply simp using Par apply auto by (metis add_increasing2 diff_is_0_eq gr0_conv_Suc less_imp_diff_less less_or_eq_imp_le nth_mem of_nat_0_le_iff of_nat_Suc) next case (ParT cl ii) have "properL cl" and "ii < brnL cl (length cl)" using ParT by auto thus ?case proof(cases rule: brnL_cases) case (Local n i) show ?thesis proof (cases "WtFT cl = 1") case True thus ?thesis using Local ParT by (cases "n = pickFT cl \ i = 0", auto) next case False note sch = False thus ?thesis using Local ParT proof (cases "finished (cl ! n)") case False note cln = False let ?L1 = "1 / (length cl)" let ?L2 = "wt (cl ! n) s i" let ?R = "WtFT cl" have "0 \ ?L1" and "0 \ ?L2" using ParT Local by auto moreover have "?L2 \ 1" using ParT Local by auto ultimately have "?L1 * ?L2 \ ?L1" by (metis mult_right_le_one_le) also have "?L1 \ 1 - ?R" using ParT Local cln by auto finally have "?L1 * ?L2 \ 1 - ?R" . thus ?thesis using Local ParT cln sch by (auto simp: pos_divide_le_eq mult.commute) qed (insert sch ParT Local, auto) qed qed qed auto abbreviation fromPlus ("(1{_..<+_})") where "{a ..<+ b} \ {a ..< a + b}" lemma brnL_UN: assumes "properL cl" shows "{..< brnL cl (length cl)} = (\ n < length cl. {brnL cl n ..<+ brn (cl!n)})" (is "?L = (\ n < length cl. ?R n)") proof safe fix ii assume ii: "ii < brnL cl (length cl)" from assms ii show "ii \ (\ n < length cl. ?R n)" proof(cases rule: brnL_cases) case (Local n i) hence "ii \ ?R n" by simp thus ?thesis using Local by force qed qed auto lemma brnL_Int_lt: assumes n12: "n1 < n2" and n2: "n2 < length cl" shows "{brnL cl n1 ..<+ brn (cl!n1)} \ {brnL cl n2 ..<+ brn (cl!n2)} = {}" proof- have "Suc n1 \ n2" using assms by simp hence "brnL cl (Suc n1) \ brnL cl n2" by simp thus ?thesis using assms by simp qed lemma brnL_Int: assumes "n1 \ n2" and "n1 < length cl" and "n2 < length cl" shows "{brnL cl n1 ..<+ brn (cl!n1)} \ {brnL cl n2 ..<+ brn (cl!n2)} = {}" proof(cases "n1 < n2") case True thus ?thesis using assms brnL_Int_lt by auto next case False hence "n2 < n1" using assms by simp thus ?thesis using brnL_Int_lt assms by fastforce qed lemma sum_wt_Par_sub[simp]: assumes cl: "properL cl" and n: "n < length cl" and I: "I \ {..< brn (cl ! n)}" shows "sum (wt (Par cl) s) ((+) (brnL cl n) ` I) = 1 /(length cl) * sum (wt (cl ! n) s) I" (is "?L = ?wSch * ?R") proof- have "?L = sum (%i. ?wSch * wt (cl ! n) s i) I" apply(rule sum.reindex_cong[of "(+) (brnL cl n)"]) using assms by auto also have "... = ?wSch * ?R" unfolding sum_distrib_left by simp finally show ?thesis . qed lemma sum_wt_Par[simp]: assumes cl: "properL cl" and n: "n < length cl" shows "sum (wt (Par cl) s) {brnL cl n ..<+ brn (cl!n)} = 1 /(length cl) * sum (wt (cl ! n) s) {..< brn (cl ! n)}" (is "?L = ?W * ?R") using assms by (simp add: sum_distrib_left) lemma sum_wt_ParT_sub_WtFT_pickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and I: "I \ {..< brn (cl ! (pickFT cl))}" "0 \ I" shows "sum (wt (ParT cl) s) ((+) (brnL cl (pickFT cl)) ` I) = 1" (is "?L = 1") proof- let ?n = "pickFT cl" let ?w = "%i. if i = 0 then (1::real) else 0" have n: "?n < length cl" using nf by simp have 0: "I = {0} Un (I - {0})" using I by auto have finI: "finite I" using I by (metis finite_nat_iff_bounded) have "?L = sum ?w I" proof (rule sum.reindex_cong [of "plus (brnL cl ?n)"]) fix i assume i: "i \ I" have "i < brn (cl ! ?n)" using i I by auto note i = this i show "wt (ParT cl) s (brnL cl ?n + i) = ?w i" using nf n i cl by (cases "i = 0") auto qed (insert assms, auto) also have "... = sum ?w ({0} Un (I - {0}))" using 0 by auto also have "... = sum ?w {0::real} + sum ?w (I - {0})" using sum.union_disjoint[of "{0}" "I - {0}" ?w] finI by auto also have "... = 1" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_WtFT_pickFT_0_2[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and II: "II \ {..< brnL cl (length cl)}" "brnL cl (pickFT cl) \ II" shows "sum (wt (ParT cl) s) II = 1" (is "?L = 1") proof- let ?n = "pickFT cl" let ?w = "%ii. if ii = brnL cl (pickFT cl) then (1::real) else 0" have n: "?n < length cl" using nf by simp have 0: "II = {brnL cl ?n} Un (II - {brnL cl ?n})" using II by auto have finI: "finite II" using II by (metis finite_nat_iff_bounded) have "?L = sum ?w II" proof(rule sum.cong) fix ii assume ii: "ii \ II" hence ii: "ii < brnL cl (length cl)" using II by auto from cl ii show "wt (ParT cl) s ii = ?w ii" proof(cases rule: brnL_cases) case (Local n i) show ?thesis proof(cases "ii = brnL cl (pickFT cl)") case True have "n = pickFT cl \ i = 0" apply(intro brnL_unique[of cl]) using Local cl nf brn_gt_0 unfolding True by auto thus ?thesis using cl nf True by simp next case False hence "n = pickFT cl \ i \ 0" unfolding Local by auto thus ?thesis using Local ii nf cl False by auto qed qed qed simp also have "... = sum ?w ({brnL cl ?n} Un (II - {brnL cl ?n}))" using 0 by simp also have "... = sum ?w {brnL cl ?n} + sum ?w (II - {brnL cl ?n})" apply(rule sum.union_disjoint) using finI by auto also have "... = 1" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_WtFT_notPickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and n: "n < length cl" and I: "I \ {..< brn (cl ! n)}" and nI: "n = pickFT cl \ 0 \ I" shows "sum (wt (ParT cl) s) ((+) (brnL cl n) ` I) = 0" (is "?L = 0") proof- have "?L = sum (%i. 0) I" proof (rule sum.reindex_cong [of "plus (brnL cl n)"]) fix i assume i: "i \ I" hence "n = pickFT cl \ i \ 0" using nI by metis moreover have "i < brn (cl ! n)" using i I by auto ultimately show "wt (ParT cl) s (brnL cl n + i) = 0" using nf n cl by simp qed (insert assms, auto) also have "... = 0" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_notWtFT_finished[simp]: assumes cl: "properL cl" and nf: "WtFT cl \ 1" and n: "n < length cl" and cln: "finished (cl!n)" and I: "I \ {..< brn (cl ! n)}" shows "sum (wt (ParT cl) s) ((+) (brnL cl n) ` I) = 0" (is "?L = 0") proof- have "?L = sum (%i. 0) I" apply(rule sum.reindex_cong[of "(+) (brnL cl n)"]) using assms by auto also have "... = 0" by simp finally show ?thesis . qed lemma sum_wt_ParT_sub_notWtFT_notFinished[simp]: assumes cl: "properL cl" and nf: "WtFT cl \ 1" and n: "n < length cl" and cln: "\ finished (cl!n)" and I: "I \ {..< brn (cl ! n)}" shows "sum (wt (ParT cl) s) ((+) (brnL cl n) ` I) = (1 / (length cl)) / (1 - WtFT cl) * sum (wt (cl ! n) s) I" (is "?L = ?w / (1 - ?wF) * ?R") proof- have "?L = sum (%i. ?w / (1 - ?wF) * wt (cl ! n) s i) I" apply(rule sum.reindex_cong[of "(+) (brnL cl n)"]) using assms by auto also have "... = ?w / (1 - ?wF) * ?R" unfolding sum_distrib_left by simp finally show ?thesis . qed lemma sum_wt_ParT_WtFT_pickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" shows "sum (wt (ParT cl) s) {brnL cl (pickFT cl) ..<+ brn (cl ! (pickFT cl))} = 1" proof- let ?n = "pickFT cl" have 1: "{brnL cl ?n ..<+ brn (cl ! ?n)} = (+) (brnL cl ?n) ` {..< brn (cl ! ?n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_WtFT_pickFT_0) using assms apply simp_all by (metis brn_gt_0_L nth_mem pickFT_length) qed lemma sum_wt_ParT_WtFT_notPickFT_0[simp]: assumes cl: "properL cl" and nf: "WtFT cl = 1" and n: "n < length cl" "n \ pickFT cl" shows "sum (wt (ParT cl) s) {brnL cl n ..<+ brn (cl!n)} = 0" proof- have 1: "{brnL cl n ..<+ brn (cl!n)} = (+) (brnL cl n) ` {..< brn (cl!n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_WtFT_notPickFT_0) using assms by auto qed lemma sum_wt_ParT_notWtFT_finished[simp]: assumes cl: "properL cl" and "WtFT cl \ 1" and n: "n < length cl" and cln: "finished (cl!n)" shows "sum (wt (ParT cl) s) {brnL cl n ..<+ brn (cl!n)} = 0" proof- have 1: "{brnL cl n ..<+ brn (cl!n)} = (+) (brnL cl n) ` {..< brn (cl!n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_notWtFT_finished) using assms by auto qed lemma sum_wt_ParT_notWtFT_notFinished[simp]: assumes cl: "properL cl" and nf: "WtFT cl \ 1" and n: "n < length cl" and cln: "\ finished (cl!n)" shows "sum (wt (ParT cl) s) {brnL cl n ..<+ brn (cl!n)} = (1 / (length cl)) / (1 - WtFT cl) * sum (wt (cl ! n) s) {..< brn (cl ! n)}" proof- have 1: "{brnL cl n ..<+ brn (cl!n)} = (+) (brnL cl n) ` {..< brn (cl!n)}" by simp show ?thesis unfolding 1 apply(rule sum_wt_ParT_sub_notWtFT_notFinished) using assms by auto qed lemma sum_wt[simp]: assumes "proper c" shows "sum (wt c s) {..< brn c} = 1" using assms proof (induct c arbitrary: s rule: proper_induct) case (Par cl) let ?w = "\ n. 1 / (length cl) * sum (wt (cl ! n) s) {..< brn (cl ! n)}" show ?case proof (rule sum_UN_introL [of _ "%n. {brnL cl n ..<+ brn (cl!n)}" "{..< length cl}" _ ?w]) have "1 = sum (\ n. 1 / (length cl)) {..< length cl}" using Par by simp also have "... = sum ?w {..< length cl}" using Par by simp finally show "1 = sum ?w {..< length cl}" . next fix m n assume "{m, n} \ {.. m \ n" thus "{brnL cl m ..<+ brn (cl!m)} \ {brnL cl n ..<+ brn (cl!n)} = {}" using brnL_Int by auto qed(insert Par brnL_UN sum_wt_Par, auto) next case (ParT cl) let ?v = "1/(length cl)" let ?wtF = "WtFT cl" let ?wtNF = "WtNFT cl" let ?w = "\n. if ?wtF = 1 then (if n = pickFT cl then 1 else 0) else (if finished (cl!n) then 0 else ?v / (1 - ?wtF) * sum (wt (cl ! n) s) {..< brn (cl ! n)})" define w where "w = ?w" have w: "\ n. ?wtF \ 1 \ n < length cl \ \ finished (cl!n) \ w n = ?v / (1 - ?wtF)" unfolding w_def using ParT by simp show ?case proof(cases "WtFT cl = 1") case True with ParT show ?thesis by simp next case False note nf = False show ?thesis proof (rule sum_UN_introL [of _ "%n. {brnL cl n ..<+ brn (cl!n)}" "{..< length cl}" _ w]) show "1 = sum w {..< length cl}" proof(cases "?wtF = 1") case True note sch = True let ?n = "pickFT cl" let ?L = "{?n}" let ?Lnot = "{.. {..iii {..< brn c}" shows "sum (wt c s) I \ 1" proof- define J where "J = {..< brn c}" have "I \ J" and "finite J" using ** unfolding J_def by auto moreover have "\ j \ J. wt c s j \ 0" unfolding J_def using * by simp ultimately have "sum (wt c s) I \ sum (wt c s) J" using sum_mono2[of J I "wt c s"] by auto also have "... = 1" using * unfolding J_def by simp finally show "sum (wt c s) I \ 1" unfolding J_def by simp qed lemma sum_le_1[simp]: assumes *: "proper c" and **: "i < brn c" shows "sum (wt c s) {..i} \ 1" proof- have "{..i} \ {..< brn c}" using ** by auto thus ?thesis using assms sum_subset_le_1[of c "{..i}" s] by blast qed subsubsection \Operations on configurations\ definition "cont_eff cf b = snd (wt_cont_eff (fst cf) (snd cf) b)" lemma cont_eff: "cont_eff cf b = (cont (fst cf) (snd cf) b, eff (fst cf) (snd cf) b)" unfolding cont_eff_def cont_def eff_def by simp end (* context PL *) end diff --git a/thys/RSAPSS/Word.thy b/thys/RSAPSS/Word.thy --- a/thys/RSAPSS/Word.thy +++ b/thys/RSAPSS/Word.thy @@ -1,2267 +1,2264 @@ (* Author: Sebastian Skalberg, TU Muenchen *) section \Binary Words\ theory Word imports Main begin subsection \Auxilary Lemmas\ lemma max_le [intro!]: "[| x \ z; y \ z |] ==> max x y \ z" by (simp add: max_def) lemma max_mono: fixes x :: "'a::linorder" assumes mf: "mono f" shows "max (f x) (f y) \ f (max x y)" proof - from mf and max.cobounded1 [of x y] have fx: "f x \ f (max x y)" by (rule monoD) from mf and max.cobounded2 [of y x] have fy: "f y \ f (max x y)" by (rule monoD) from fx and fy show "max (f x) (f y) \ f (max x y)" by auto qed declare zero_le_power [intro] and zero_less_power [intro] lemma int_nat_two_exp: "2 ^ k = int (2 ^ k)" by simp subsection \Bits\ datatype bit = Zero ("\") | One ("\") primrec bitval :: "bit => nat" where "bitval \ = 0" | "bitval \ = 1" primrec bitnot :: "bit => bit" ("\\<^sub>b _" [40] 40) where bitnot_zero: "(\\<^sub>b \) = \" | bitnot_one : "(\\<^sub>b \) = \" primrec bitand :: "bit => bit => bit" (infixr "\\<^sub>b" 35) where bitand_zero: "(\ \\<^sub>b y) = \" | bitand_one: "(\ \\<^sub>b y) = y" primrec bitor :: "bit => bit => bit" (infixr "\\<^sub>b" 30) where bitor_zero: "(\ \\<^sub>b y) = y" | bitor_one: "(\ \\<^sub>b y) = \" primrec bitxor :: "bit => bit => bit" (infixr "\\<^sub>b" 30) where bitxor_zero: "(\ \\<^sub>b y) = y" | bitxor_one: "(\ \\<^sub>b y) = (bitnot y)" lemma bitnot_bitnot [simp]: "(bitnot (bitnot b)) = b" by (cases b) simp_all lemma bitand_cancel [simp]: "(b \\<^sub>b b) = b" by (cases b) simp_all lemma bitor_cancel [simp]: "(b \\<^sub>b b) = b" by (cases b) simp_all lemma bitxor_cancel [simp]: "(b \\<^sub>b b) = \" by (cases b) simp_all subsection \Bit Vectors\ text \First, a couple of theorems expressing case analysis and induction principles for bit vectors.\ lemma bit_list_cases: assumes empty: "w = [] ==> P w" and zero: "!!bs. w = \ # bs ==> P w" and one: "!!bs. w = \ # bs ==> P w" shows "P w" proof (cases w) assume "w = []" thus ?thesis by (rule empty) next fix b bs assume [simp]: "w = b # bs" show "P w" proof (cases b) assume "b = \" hence "w = \ # bs" by simp thus ?thesis by (rule zero) next assume "b = \" hence "w = \ # bs" by simp thus ?thesis by (rule one) qed qed lemma bit_list_induct: assumes empty: "P []" and zero: "!!bs. P bs ==> P (\#bs)" and one: "!!bs. P bs ==> P (\#bs)" shows "P w" proof (induct w, simp_all add: empty) fix b bs assume "P bs" then show "P (b#bs)" by (cases b) (auto intro!: zero one) qed definition bv_msb :: "bit list => bit" where "bv_msb w = (if w = [] then \ else hd w)" definition bv_extend :: "[nat,bit,bit list]=>bit list" where "bv_extend i b w = (replicate (i - length w) b) @ w" definition bv_not :: "bit list => bit list" where "bv_not w = map bitnot w" lemma bv_length_extend [simp]: "length w \ i ==> length (bv_extend i b w) = i" by (simp add: bv_extend_def) lemma bv_not_Nil [simp]: "bv_not [] = []" by (simp add: bv_not_def) lemma bv_not_Cons [simp]: "bv_not (b#bs) = (bitnot b) # bv_not bs" by (simp add: bv_not_def) lemma bv_not_bv_not [simp]: "bv_not (bv_not w) = w" by (rule bit_list_induct [of _ w]) simp_all lemma bv_msb_Nil [simp]: "bv_msb [] = \" by (simp add: bv_msb_def) lemma bv_msb_Cons [simp]: "bv_msb (b#bs) = b" by (simp add: bv_msb_def) lemma bv_msb_bv_not [simp]: "0 < length w ==> bv_msb (bv_not w) = (bitnot (bv_msb w))" by (cases w) simp_all lemma bv_msb_one_length [simp,intro]: "bv_msb w = \ ==> 0 < length w" by (cases w) simp_all lemma length_bv_not [simp]: "length (bv_not w) = length w" by (induct w) simp_all definition bv_to_nat :: "bit list => nat" where "bv_to_nat = foldl (%bn b. 2 * bn + bitval b) 0" lemma bv_to_nat_Nil [simp]: "bv_to_nat [] = 0" by (simp add: bv_to_nat_def) lemma bv_to_nat_helper [simp]: "bv_to_nat (b # bs) = bitval b * 2 ^ length bs + bv_to_nat bs" proof - let ?bv_to_nat' = "foldl (\bn b. 2 * bn + bitval b)" have helper: "\base. ?bv_to_nat' base bs = base * 2 ^ length bs + ?bv_to_nat' 0 bs" proof (induct bs) case Nil show ?case by simp next case (Cons x xs base) show ?case apply (simp only: foldl_Cons) apply (subst Cons [of "2 * base + bitval x"]) apply simp apply (subst Cons [of "bitval x"]) apply (simp add: add_mult_distrib) done qed show ?thesis by (simp add: bv_to_nat_def) (rule helper) qed lemma bv_to_nat0 [simp]: "bv_to_nat (\#bs) = bv_to_nat bs" by simp lemma bv_to_nat1 [simp]: "bv_to_nat (\#bs) = 2 ^ length bs + bv_to_nat bs" by simp lemma bv_to_nat_upper_range: "bv_to_nat w < 2 ^ length w" proof (induct w, simp_all) fix b bs assume "bv_to_nat bs < 2 ^ length bs" show "bitval b * 2 ^ length bs + bv_to_nat bs < 2 * 2 ^ length bs" proof (cases b, simp_all) have "bv_to_nat bs < 2 ^ length bs" by fact also have "... < 2 * 2 ^ length bs" by auto finally show "bv_to_nat bs < 2 * 2 ^ length bs" by simp next have "bv_to_nat bs < 2 ^ length bs" by fact hence "2 ^ length bs + bv_to_nat bs < 2 ^ length bs + 2 ^ length bs" by arith also have "... = 2 * (2 ^ length bs)" by simp finally show "bv_to_nat bs < 2 ^ length bs" by simp qed qed lemma bv_extend_longer [simp]: assumes wn: "n \ length w" shows "bv_extend n b w = w" by (simp add: bv_extend_def wn) lemma bv_extend_shorter [simp]: assumes wn: "length w < n" shows "bv_extend n b w = bv_extend n b (b#w)" proof - from wn have s: "n - Suc (length w) + 1 = n - length w" by arith have "bv_extend n b w = replicate (n - length w) b @ w" by (simp add: bv_extend_def) also have "... = replicate (n - Suc (length w) + 1) b @ w" by (subst s) rule also have "... = (replicate (n - Suc (length w)) b @ replicate 1 b) @ w" by (subst replicate_add) rule also have "... = replicate (n - Suc (length w)) b @ b # w" by simp also have "... = bv_extend n b (b#w)" by (simp add: bv_extend_def) finally show "bv_extend n b w = bv_extend n b (b#w)" . qed primrec rem_initial :: "bit => bit list => bit list" where "rem_initial b [] = []" | "rem_initial b (x#xs) = (if b = x then rem_initial b xs else x#xs)" lemma rem_initial_length: "length (rem_initial b w) \ length w" by (rule bit_list_induct [of _ w],simp_all (no_asm),safe,simp_all) lemma rem_initial_equal: assumes p: "length (rem_initial b w) = length w" shows "rem_initial b w = w" proof - have "length (rem_initial b w) = length w --> rem_initial b w = w" proof (induct w, simp_all, clarify) fix xs assume "length (rem_initial b xs) = length xs --> rem_initial b xs = xs" assume f: "length (rem_initial b xs) = Suc (length xs)" with rem_initial_length [of b xs] show "rem_initial b xs = b#xs" by auto qed from this and p show ?thesis .. qed lemma bv_extend_rem_initial: "bv_extend (length w) b (rem_initial b w) = w" proof (induct w, simp_all, safe) fix xs assume ind: "bv_extend (length xs) b (rem_initial b xs) = xs" from rem_initial_length [of b xs] have [simp]: "Suc (length xs) - length (rem_initial b xs) = 1 + (length xs - length (rem_initial b xs))" by arith have "bv_extend (Suc (length xs)) b (rem_initial b xs) = replicate (Suc (length xs) - length (rem_initial b xs)) b @ (rem_initial b xs)" by (simp add: bv_extend_def) also have "... = replicate (1 + (length xs - length (rem_initial b xs))) b @ rem_initial b xs" by simp also have "... = (replicate 1 b @ replicate (length xs - length (rem_initial b xs)) b) @ rem_initial b xs" by (subst replicate_add) (rule refl) also have "... = b # bv_extend (length xs) b (rem_initial b xs)" by (auto simp add: bv_extend_def [symmetric]) also have "... = b # xs" by (simp add: ind) finally show "bv_extend (Suc (length xs)) b (rem_initial b xs) = b # xs" . qed lemma rem_initial_append1: assumes "rem_initial b xs ~= []" shows "rem_initial b (xs @ ys) = rem_initial b xs @ ys" using assms by (induct xs) auto lemma rem_initial_append2: assumes "rem_initial b xs = []" shows "rem_initial b (xs @ ys) = rem_initial b ys" using assms by (induct xs) auto definition norm_unsigned :: "bit list => bit list" where "norm_unsigned = rem_initial \" lemma norm_unsigned_Nil [simp]: "norm_unsigned [] = []" by (simp add: norm_unsigned_def) lemma norm_unsigned_Cons0 [simp]: "norm_unsigned (\#bs) = norm_unsigned bs" by (simp add: norm_unsigned_def) lemma norm_unsigned_Cons1 [simp]: "norm_unsigned (\#bs) = \#bs" by (simp add: norm_unsigned_def) lemma norm_unsigned_idem [simp]: "norm_unsigned (norm_unsigned w) = norm_unsigned w" by (rule bit_list_induct [of _ w],simp_all) fun nat_to_bv_helper :: "nat => bit list => bit list" where "nat_to_bv_helper n bs = (if n = 0 then bs else nat_to_bv_helper (n div 2) ((if n mod 2 = 0 then \ else \)#bs))" definition nat_to_bv :: "nat => bit list" where "nat_to_bv n = nat_to_bv_helper n []" lemma nat_to_bv0 [simp]: "nat_to_bv 0 = []" by (simp add: nat_to_bv_def) lemmas [simp del] = nat_to_bv_helper.simps lemma n_div_2_cases: assumes zero: "(n::nat) = 0 ==> R" and div : "[| n div 2 < n ; 0 < n |] ==> R" shows "R" proof (cases "n = 0") assume "n = 0" thus R by (rule zero) next assume "n ~= 0" hence "0 < n" by simp hence "n div 2 < n" by arith from this and \0 < n\ show R by (rule div) qed lemma int_wf_ge_induct: assumes ind : "!!i::int. (!!j. [| k \ j ; j < i |] ==> P j) ==> P i" shows "P i" proof (rule wf_induct_rule [OF wf_int_ge_less_than]) fix x assume ih: "(\y::int. (y, x) \ int_ge_less_than k \ P y)" thus "P x" by (rule ind) (simp add: int_ge_less_than_def) qed lemma unfold_nat_to_bv_helper: "nat_to_bv_helper b l = nat_to_bv_helper b [] @ l" proof - have "\l. nat_to_bv_helper b l = nat_to_bv_helper b [] @ l" proof (induct b rule: less_induct) fix n assume ind: "!!j. j < n \ \ l. nat_to_bv_helper j l = nat_to_bv_helper j [] @ l" show "\l. nat_to_bv_helper n l = nat_to_bv_helper n [] @ l" proof fix l show "nat_to_bv_helper n l = nat_to_bv_helper n [] @ l" proof (cases "n < 0") assume "n < 0" thus ?thesis by (simp add: nat_to_bv_helper.simps) next assume "~n < 0" show ?thesis proof (rule n_div_2_cases [of n]) assume [simp]: "n = 0" show ?thesis apply (simp only: nat_to_bv_helper.simps [of n]) apply simp done next assume n2n: "n div 2 < n" assume [simp]: "0 < n" hence n20: "0 \ n div 2" by arith from ind [of "n div 2"] and n2n n20 have ind': "\l. nat_to_bv_helper (n div 2) l = nat_to_bv_helper (n div 2) [] @ l" by blast show ?thesis apply (simp only: nat_to_bv_helper.simps [of n]) apply (cases "n=0") apply simp apply (simp only: if_False) apply simp apply (subst spec [OF ind',of "\#l"]) apply (subst spec [OF ind',of "\#l"]) apply (subst spec [OF ind',of "[\]"]) apply (subst spec [OF ind',of "[\]"]) apply simp done qed qed qed qed thus ?thesis .. qed lemma nat_to_bv_non0 [simp]: "n\0 ==> nat_to_bv n = nat_to_bv (n div 2) @ [if n mod 2 = 0 then \ else \]" proof - assume n: "n\0" show ?thesis apply (subst nat_to_bv_def [of n]) apply (simp only: nat_to_bv_helper.simps [of n]) apply (subst unfold_nat_to_bv_helper) apply (simp add: n) apply (subst nat_to_bv_def [of "n div 2"]) apply auto done qed lemma bv_to_nat_dist_append: "bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2" proof - have "\l2. bv_to_nat (l1 @ l2) = bv_to_nat l1 * 2 ^ length l2 + bv_to_nat l2" proof (induct l1, simp_all) fix x xs assume ind: "\l2. bv_to_nat (xs @ l2) = bv_to_nat xs * 2 ^ length l2 + bv_to_nat l2" show "\l2::bit list. bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2" proof fix l2 show "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2" proof - have "(2::nat) ^ (length xs + length l2) = 2 ^ length xs * 2 ^ length l2" by (induct ("length xs")) simp_all hence "bitval x * 2 ^ (length xs + length l2) + bv_to_nat xs * 2 ^ length l2 = bitval x * 2 ^ length xs * 2 ^ length l2 + bv_to_nat xs * 2 ^ length l2" by simp also have "... = (bitval x * 2 ^ length xs + bv_to_nat xs) * 2 ^ length l2" by (simp add: ring_distribs) finally show ?thesis by simp qed qed qed thus ?thesis .. qed lemma bv_nat_bv [simp]: "bv_to_nat (nat_to_bv n) = n" proof (induct n rule: less_induct) fix n assume ind: "!!j. j < n \ bv_to_nat (nat_to_bv j) = j" show "bv_to_nat (nat_to_bv n) = n" proof (rule n_div_2_cases [of n]) assume "n = 0" then show ?thesis by simp next assume nn: "n div 2 < n" assume n0: "0 < n" from ind and nn have ind': "bv_to_nat (nat_to_bv (n div 2)) = n div 2" by blast from n0 have n0': "n \ 0" by simp show ?thesis apply (subst nat_to_bv_def) apply (simp only: nat_to_bv_helper.simps [of n]) apply (simp only: n0' if_False) apply (subst unfold_nat_to_bv_helper) apply (subst bv_to_nat_dist_append) apply (fold nat_to_bv_def) apply (simp add: ind' split del: if_split) apply (cases "n mod 2 = 0") proof (simp_all) assume "n mod 2 = 0" with div_mult_mod_eq [of n 2] show "n div 2 * 2 = n" by simp next assume "n mod 2 = Suc 0" with div_mult_mod_eq [of n 2] show "Suc (n div 2 * 2) = n" by arith qed qed qed lemma bv_to_nat_type [simp]: "bv_to_nat (norm_unsigned w) = bv_to_nat w" by (rule bit_list_induct) simp_all lemma length_norm_unsigned_le [simp]: "length (norm_unsigned w) \ length w" by (rule bit_list_induct) simp_all lemma bv_to_nat_rew_msb: "bv_msb w = \ ==> bv_to_nat w = 2 ^ (length w - 1) + bv_to_nat (tl w)" by (rule bit_list_cases [of w]) simp_all lemma norm_unsigned_result: "norm_unsigned xs = [] \ bv_msb (norm_unsigned xs) = \" proof (rule length_induct [of _ xs]) fix xs :: "bit list" assume ind: "\ys. length ys < length xs --> norm_unsigned ys = [] \ bv_msb (norm_unsigned ys) = \" show "norm_unsigned xs = [] \ bv_msb (norm_unsigned xs) = \" proof (rule bit_list_cases [of xs],simp_all) fix bs assume [simp]: "xs = \#bs" from ind have "length bs < length xs --> norm_unsigned bs = [] \ bv_msb (norm_unsigned bs) = \" .. thus "norm_unsigned bs = [] \ bv_msb (norm_unsigned bs) = \" by simp qed qed lemma norm_empty_bv_to_nat_zero: assumes nw: "norm_unsigned w = []" shows "bv_to_nat w = 0" proof - have "bv_to_nat w = bv_to_nat (norm_unsigned w)" by simp also have "... = bv_to_nat []" by (subst nw) (rule refl) also have "... = 0" by simp finally show ?thesis . qed lemma bv_to_nat_lower_limit: assumes w0: "0 < bv_to_nat w" shows "2 ^ (length (norm_unsigned w) - 1) \ bv_to_nat w" proof - from w0 and norm_unsigned_result [of w] have msbw: "bv_msb (norm_unsigned w) = \" by (auto simp add: norm_empty_bv_to_nat_zero) have "2 ^ (length (norm_unsigned w) - 1) \ bv_to_nat (norm_unsigned w)" by (subst bv_to_nat_rew_msb [OF msbw],simp) thus ?thesis by simp qed lemmas [simp del] = nat_to_bv_non0 lemma norm_unsigned_length [intro!]: "length (norm_unsigned w) \ length w" by (subst norm_unsigned_def,rule rem_initial_length) lemma norm_unsigned_equal: "length (norm_unsigned w) = length w ==> norm_unsigned w = w" by (simp add: norm_unsigned_def,rule rem_initial_equal) lemma bv_extend_norm_unsigned: "bv_extend (length w) \ (norm_unsigned w) = w" by (simp add: norm_unsigned_def,rule bv_extend_rem_initial) lemma norm_unsigned_append1 [simp]: "norm_unsigned xs \ [] ==> norm_unsigned (xs @ ys) = norm_unsigned xs @ ys" by (simp add: norm_unsigned_def,rule rem_initial_append1) lemma norm_unsigned_append2 [simp]: "norm_unsigned xs = [] ==> norm_unsigned (xs @ ys) = norm_unsigned ys" by (simp add: norm_unsigned_def,rule rem_initial_append2) lemma bv_to_nat_zero_imp_empty: "bv_to_nat w = 0 \ norm_unsigned w = []" by (atomize (full), induct w rule: bit_list_induct) simp_all lemma bv_to_nat_nzero_imp_nempty: "bv_to_nat w \ 0 \ norm_unsigned w \ []" by (induct w rule: bit_list_induct) simp_all lemma nat_helper1: assumes ass: "nat_to_bv (bv_to_nat w) = norm_unsigned w" shows "nat_to_bv (2 * bv_to_nat w + bitval x) = norm_unsigned (w @ [x])" proof (cases x) assume [simp]: "x = \" have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\]" by (simp add: add.commute) also have "... = nat_to_bv (bv_to_nat w) @ [\]" by (subst div_add1_eq) simp also have "... = norm_unsigned w @ [\]" by (subst ass) (rule refl) also have "... = norm_unsigned (w @ [\])" by (cases "norm_unsigned w") simp_all finally have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" . then show ?thesis by (simp add: nat_to_bv_non0) next assume [simp]: "x = \" show ?thesis proof (cases "bv_to_nat w = 0") assume "bv_to_nat w = 0" thus ?thesis by (simp add: bv_to_nat_zero_imp_empty) next assume "bv_to_nat w \ 0" thus ?thesis apply simp apply (subst nat_to_bv_non0) apply simp apply auto apply (subst ass) apply (cases "norm_unsigned w") apply (simp_all add: norm_empty_bv_to_nat_zero) done qed qed lemma nat_helper2: "nat_to_bv (2 ^ length xs + bv_to_nat xs) = \ # xs" proof - have "\xs. nat_to_bv (2 ^ length (rev xs) + bv_to_nat (rev xs)) = \ # (rev xs)" (is "\xs. ?P xs") proof fix xs show "?P xs" proof (rule length_induct [of _ xs]) fix xs :: "bit list" assume ind: "\ys. length ys < length xs --> ?P ys" show "?P xs" proof (cases xs) assume "xs = []" then show ?thesis by (simp add: nat_to_bv_non0) next fix y ys assume [simp]: "xs = y # ys" show ?thesis apply simp apply (subst bv_to_nat_dist_append) apply simp proof - have "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) = nat_to_bv (2 * (2 ^ length ys + bv_to_nat (rev ys)) + bitval y)" by (simp add: ac_simps ac_simps) also have "... = nat_to_bv (2 * (bv_to_nat (\#rev ys)) + bitval y)" by simp also have "... = norm_unsigned (\#rev ys) @ [y]" proof - from ind have "nat_to_bv (2 ^ length (rev ys) + bv_to_nat (rev ys)) = \ # rev ys" by auto hence [simp]: "nat_to_bv (2 ^ length ys + bv_to_nat (rev ys)) = \ # rev ys" by simp show ?thesis apply (subst nat_helper1) apply simp_all done qed also have "... = (\#rev ys) @ [y]" by simp also have "... = \ # rev ys @ [y]" by simp finally show "nat_to_bv (2 * 2 ^ length ys + (bv_to_nat (rev ys) * 2 + bitval y)) = \ # rev ys @ [y]" . qed qed qed qed hence "nat_to_bv (2 ^ length (rev (rev xs)) + bv_to_nat (rev (rev xs))) = \ # rev (rev xs)" .. thus ?thesis by simp qed lemma nat_bv_nat [simp]: "nat_to_bv (bv_to_nat w) = norm_unsigned w" proof (rule bit_list_induct [of _ w],simp_all) fix xs assume "nat_to_bv (bv_to_nat xs) = norm_unsigned xs" have "bv_to_nat xs = bv_to_nat (norm_unsigned xs)" by simp have "bv_to_nat xs < 2 ^ length xs" by (rule bv_to_nat_upper_range) show "nat_to_bv (2 ^ length xs + bv_to_nat xs) = \ # xs" by (rule nat_helper2) qed lemma bv_to_nat_qinj: assumes one: "bv_to_nat xs = bv_to_nat ys" and len: "length xs = length ys" shows "xs = ys" proof - from one have "nat_to_bv (bv_to_nat xs) = nat_to_bv (bv_to_nat ys)" by simp hence xsys: "norm_unsigned xs = norm_unsigned ys" by simp have "xs = bv_extend (length xs) \ (norm_unsigned xs)" by (simp add: bv_extend_norm_unsigned) also have "... = bv_extend (length ys) \ (norm_unsigned ys)" by (simp add: xsys len) also have "... = ys" by (simp add: bv_extend_norm_unsigned) finally show ?thesis . qed lemma norm_unsigned_nat_to_bv [simp]: "norm_unsigned (nat_to_bv n) = nat_to_bv n" proof - have "norm_unsigned (nat_to_bv n) = nat_to_bv (bv_to_nat (norm_unsigned (nat_to_bv n)))" by (subst nat_bv_nat) simp also have "... = nat_to_bv n" by simp finally show ?thesis . qed lemma length_nat_to_bv_upper_limit: assumes nk: "n \ 2 ^ k - 1" shows "length (nat_to_bv n) \ k" proof (cases "n = 0") case True thus ?thesis by (simp add: nat_to_bv_def nat_to_bv_helper.simps) next case False hence n0: "0 < n" by simp show ?thesis proof (rule ccontr) assume "~ length (nat_to_bv n) \ k" hence "k < length (nat_to_bv n)" by simp hence "k \ length (nat_to_bv n) - 1" by arith hence "(2::nat) ^ k \ 2 ^ (length (nat_to_bv n) - 1)" by simp also have "... = 2 ^ (length (norm_unsigned (nat_to_bv n)) - 1)" by simp also have "... \ bv_to_nat (nat_to_bv n)" by (rule bv_to_nat_lower_limit) (simp add: n0) also have "... = n" by simp finally have "2 ^ k \ n" . with n0 have "2 ^ k - 1 < n" by arith with nk show False by simp qed qed lemma length_nat_to_bv_lower_limit: assumes nk: "2 ^ k \ n" shows "k < length (nat_to_bv n)" proof (rule ccontr) assume "~ k < length (nat_to_bv n)" hence lnk: "length (nat_to_bv n) \ k" by simp have "n = bv_to_nat (nat_to_bv n)" by simp also have "... < 2 ^ length (nat_to_bv n)" by (rule bv_to_nat_upper_range) also from lnk have "... \ 2 ^ k" by simp finally have "n < 2 ^ k" . with nk show False by simp qed subsection \Unsigned Arithmetic Operations\ definition bv_add :: "[bit list, bit list ] => bit list" where "bv_add w1 w2 = nat_to_bv (bv_to_nat w1 + bv_to_nat w2)" lemma bv_add_type1 [simp]: "bv_add (norm_unsigned w1) w2 = bv_add w1 w2" by (simp add: bv_add_def) lemma bv_add_type2 [simp]: "bv_add w1 (norm_unsigned w2) = bv_add w1 w2" by (simp add: bv_add_def) lemma bv_add_returntype [simp]: "norm_unsigned (bv_add w1 w2) = bv_add w1 w2" by (simp add: bv_add_def) lemma bv_add_length: "length (bv_add w1 w2) \ Suc (max (length w1) (length w2))" proof (unfold bv_add_def,rule length_nat_to_bv_upper_limit) from bv_to_nat_upper_range [of w1] and bv_to_nat_upper_range [of w2] have "bv_to_nat w1 + bv_to_nat w2 \ (2 ^ length w1 - 1) + (2 ^ length w2 - 1)" by arith also have "... \ max (2 ^ length w1 - 1) (2 ^ length w2 - 1) + max (2 ^ length w1 - 1) (2 ^ length w2 - 1)" by (rule add_mono,safe intro!: max.cobounded1 max.cobounded2) also have "... = 2 * max (2 ^ length w1 - 1) (2 ^ length w2 - 1)" by simp also have "... \ 2 ^ Suc (max (length w1) (length w2)) - 2" proof (cases "length w1 \ length w2") assume w1w2: "length w1 \ length w2" hence "(2::nat) ^ length w1 \ 2 ^ length w2" by simp hence "(2::nat) ^ length w1 - 1 \ 2 ^ length w2 - 1" by arith with w1w2 show ?thesis by (simp add: diff_mult_distrib2 split: split_max) next assume [simp]: "~ (length w1 \ length w2)" have "~ ((2::nat) ^ length w1 - 1 \ 2 ^ length w2 - 1)" proof assume "(2::nat) ^ length w1 - 1 \ 2 ^ length w2 - 1" hence "((2::nat) ^ length w1 - 1) + 1 \ (2 ^ length w2 - 1) + 1" by (rule add_right_mono) hence "(2::nat) ^ length w1 \ 2 ^ length w2" by simp hence "length w1 \ length w2" by simp thus False by simp qed thus ?thesis by (simp add: diff_mult_distrib2 split: split_max) qed finally show "bv_to_nat w1 + bv_to_nat w2 \ 2 ^ Suc (max (length w1) (length w2)) - 1" by arith qed definition bv_mult :: "[bit list, bit list ] => bit list" where "bv_mult w1 w2 = nat_to_bv (bv_to_nat w1 * bv_to_nat w2)" lemma bv_mult_type1 [simp]: "bv_mult (norm_unsigned w1) w2 = bv_mult w1 w2" by (simp add: bv_mult_def) lemma bv_mult_type2 [simp]: "bv_mult w1 (norm_unsigned w2) = bv_mult w1 w2" by (simp add: bv_mult_def) lemma bv_mult_returntype [simp]: "norm_unsigned (bv_mult w1 w2) = bv_mult w1 w2" by (simp add: bv_mult_def) lemma bv_mult_length: "length (bv_mult w1 w2) \ length w1 + length w2" proof (unfold bv_mult_def,rule length_nat_to_bv_upper_limit) from bv_to_nat_upper_range [of w1] and bv_to_nat_upper_range [of w2] have h: "bv_to_nat w1 \ 2 ^ length w1 - 1 \ bv_to_nat w2 \ 2 ^ length w2 - 1" by arith have "bv_to_nat w1 * bv_to_nat w2 \ (2 ^ length w1 - 1) * (2 ^ length w2 - 1)" apply (cut_tac h) apply (rule mult_mono) apply auto done also have "... < 2 ^ length w1 * 2 ^ length w2" by (rule mult_strict_mono,auto) also have "... = 2 ^ (length w1 + length w2)" by (simp add: power_add) finally show "bv_to_nat w1 * bv_to_nat w2 \ 2 ^ (length w1 + length w2) - 1" by arith qed subsection \Signed Vectors\ primrec norm_signed :: "bit list => bit list" where norm_signed_Nil: "norm_signed [] = []" | norm_signed_Cons: "norm_signed (b#bs) = (case b of \ => if norm_unsigned bs = [] then [] else b#norm_unsigned bs | \ => b#rem_initial b bs)" lemma norm_signed0 [simp]: "norm_signed [\] = []" by simp lemma norm_signed1 [simp]: "norm_signed [\] = [\]" by simp lemma norm_signed01 [simp]: "norm_signed (\#\#xs) = \#\#xs" by simp lemma norm_signed00 [simp]: "norm_signed (\#\#xs) = norm_signed (\#xs)" by simp lemma norm_signed10 [simp]: "norm_signed (\#\#xs) = \#\#xs" by simp lemma norm_signed11 [simp]: "norm_signed (\#\#xs) = norm_signed (\#xs)" by simp lemmas [simp del] = norm_signed_Cons definition int_to_bv :: "int => bit list" where "int_to_bv n = (if 0 \ n then norm_signed (\#nat_to_bv (nat n)) else norm_signed (bv_not (\#nat_to_bv (nat (-n- 1)))))" lemma int_to_bv_ge0 [simp]: "0 \ n ==> int_to_bv n = norm_signed (\ # nat_to_bv (nat n))" by (simp add: int_to_bv_def) lemma int_to_bv_lt0 [simp]: "n < 0 ==> int_to_bv n = norm_signed (bv_not (\#nat_to_bv (nat (-n- 1))))" by (simp add: int_to_bv_def) lemma norm_signed_idem [simp]: "norm_signed (norm_signed w) = norm_signed w" proof (rule bit_list_induct [of _ w], simp_all) fix xs assume eq: "norm_signed (norm_signed xs) = norm_signed xs" show "norm_signed (norm_signed (\#xs)) = norm_signed (\#xs)" proof (rule bit_list_cases [of xs],simp_all) fix ys assume "xs = \#ys" from this [symmetric] and eq show "norm_signed (norm_signed (\#ys)) = norm_signed (\#ys)" by simp qed next fix xs assume eq: "norm_signed (norm_signed xs) = norm_signed xs" show "norm_signed (norm_signed (\#xs)) = norm_signed (\#xs)" proof (rule bit_list_cases [of xs],simp_all) fix ys assume "xs = \#ys" from this [symmetric] and eq show "norm_signed (norm_signed (\#ys)) = norm_signed (\#ys)" by simp qed qed definition bv_to_int :: "bit list => int" where "bv_to_int w = (case bv_msb w of \ => int (bv_to_nat w) | \ => - int (bv_to_nat (bv_not w) + 1))" lemma bv_to_int_Nil [simp]: "bv_to_int [] = 0" by (simp add: bv_to_int_def) lemma bv_to_int_Cons0 [simp]: "bv_to_int (\#bs) = int (bv_to_nat bs)" by (simp add: bv_to_int_def) lemma bv_to_int_Cons1 [simp]: "bv_to_int (\#bs) = - int (bv_to_nat (bv_not bs) + 1)" by (simp add: bv_to_int_def) lemma bv_to_int_type [simp]: "bv_to_int (norm_signed w) = bv_to_int w" proof (rule bit_list_induct [of _ w], simp_all) fix xs assume ind: "bv_to_int (norm_signed xs) = bv_to_int xs" show "bv_to_int (norm_signed (\#xs)) = int (bv_to_nat xs)" proof (rule bit_list_cases [of xs], simp_all) fix ys assume [simp]: "xs = \#ys" from ind show "bv_to_int (norm_signed (\#ys)) = int (bv_to_nat ys)" by simp qed next fix xs assume ind: "bv_to_int (norm_signed xs) = bv_to_int xs" show "bv_to_int (norm_signed (\#xs)) = -1 - int (bv_to_nat (bv_not xs))" proof (rule bit_list_cases [of xs], simp_all) fix ys assume [simp]: "xs = \#ys" from ind show "bv_to_int (norm_signed (\#ys)) = -1 - int (bv_to_nat (bv_not ys))" by simp qed qed lemma bv_to_int_upper_range: "bv_to_int w < 2 ^ (length w - 1)" proof (rule bit_list_cases [of w],simp_all add: bv_to_nat_upper_range) fix bs have "-1 - int (bv_to_nat (bv_not bs)) \ 0" by simp also have "... < 2 ^ length bs" by (induct bs) simp_all finally show "-1 - int (bv_to_nat (bv_not bs)) < 2 ^ length bs" . qed lemma bv_to_int_lower_range: "- (2 ^ (length w - 1)) \ bv_to_int w" proof (rule bit_list_cases [of w],simp_all) fix bs :: "bit list" have "- (2 ^ length bs) \ (0::int)" by (induct bs) simp_all also have "... \ int (bv_to_nat bs)" by simp finally show "- (2 ^ length bs) \ int (bv_to_nat bs)" . next fix bs from bv_to_nat_upper_range [of "bv_not bs"] show "- (2 ^ length bs) \ -1 - int (bv_to_nat (bv_not bs))" apply (simp add: algebra_simps) by (metis of_nat_power add.commute not_less of_nat_numeral zle_add1_eq_le of_nat_le_iff) qed lemma int_bv_int [simp]: "int_to_bv (bv_to_int w) = norm_signed w" proof (rule bit_list_cases [of w],simp) fix xs assume [simp]: "w = \#xs" show ?thesis apply simp apply (subst norm_signed_Cons [of "\" "xs"]) apply simp using norm_unsigned_result [of xs] apply safe apply (rule bit_list_cases [of "norm_unsigned xs"]) apply simp_all done next fix xs assume [simp]: "w = \#xs" show ?thesis apply (simp del: int_to_bv_lt0) apply (rule bit_list_induct [of _ xs], simp) apply (subst int_to_bv_lt0) apply linarith apply simp apply (metis add.commute bitnot_zero bv_not_Cons bv_not_bv_not int_nat_two_exp length_bv_not nat_helper2 nat_int norm_signed10 of_nat_add) apply simp done qed lemma bv_int_bv [simp]: "bv_to_int (int_to_bv i) = i" by (cases "0 \ i") simp_all lemma bv_msb_norm [simp]: "bv_msb (norm_signed w) = bv_msb w" by (rule bit_list_cases [of w]) (simp_all add: norm_signed_Cons) lemma norm_signed_length: "length (norm_signed w) \ length w" apply (cases w, simp_all) apply (subst norm_signed_Cons) apply (case_tac a, simp_all) apply (rule rem_initial_length) done lemma norm_signed_equal: "length (norm_signed w) = length w ==> norm_signed w = w" proof (rule bit_list_cases [of w], simp_all) fix xs assume "length (norm_signed (\#xs)) = Suc (length xs)" thus "norm_signed (\#xs) = \#xs" by (simp add: norm_signed_Cons norm_unsigned_equal [THEN eqTrueI] split: if_split_asm) next fix xs assume "length (norm_signed (\#xs)) = Suc (length xs)" thus "norm_signed (\#xs) = \#xs" apply (simp add: norm_signed_Cons) apply (rule rem_initial_equal) apply assumption done qed lemma bv_extend_norm_signed: "bv_msb w = b ==> bv_extend (length w) b (norm_signed w) = w" proof (rule bit_list_cases [of w],simp_all) fix xs show "bv_extend (Suc (length xs)) \ (norm_signed (\#xs)) = \#xs" proof (simp add: norm_signed_def,auto) assume "norm_unsigned xs = []" hence xx: "rem_initial \ xs = []" by (simp add: norm_unsigned_def) have "bv_extend (Suc (length xs)) \ (\#rem_initial \ xs) = \#xs" apply (simp add: bv_extend_def replicate_app_Cons_same) apply (fold bv_extend_def) apply (rule bv_extend_rem_initial) done thus "bv_extend (Suc (length xs)) \ [\] = \#xs" by (simp add: xx) next show "bv_extend (Suc (length xs)) \ (\#norm_unsigned xs) = \#xs" apply (simp add: norm_unsigned_def) apply (simp add: bv_extend_def replicate_app_Cons_same) apply (fold bv_extend_def) apply (rule bv_extend_rem_initial) done qed next fix xs show "bv_extend (Suc (length xs)) \ (norm_signed (\#xs)) = \#xs" apply (simp add: norm_signed_Cons) apply (simp add: bv_extend_def replicate_app_Cons_same) apply (fold bv_extend_def) apply (rule bv_extend_rem_initial) done qed lemma bv_to_int_qinj: assumes one: "bv_to_int xs = bv_to_int ys" and len: "length xs = length ys" shows "xs = ys" proof - from one have "int_to_bv (bv_to_int xs) = int_to_bv (bv_to_int ys)" by simp hence xsys: "norm_signed xs = norm_signed ys" by simp hence xsys': "bv_msb xs = bv_msb ys" proof - have "bv_msb xs = bv_msb (norm_signed xs)" by simp also have "... = bv_msb (norm_signed ys)" by (simp add: xsys) also have "... = bv_msb ys" by simp finally show ?thesis . qed have "xs = bv_extend (length xs) (bv_msb xs) (norm_signed xs)" by (simp add: bv_extend_norm_signed) also have "... = bv_extend (length ys) (bv_msb ys) (norm_signed ys)" by (simp add: xsys xsys' len) also have "... = ys" by (simp add: bv_extend_norm_signed) finally show ?thesis . qed lemma int_to_bv_returntype [simp]: "norm_signed (int_to_bv w) = int_to_bv w" by (simp add: int_to_bv_def) lemma bv_to_int_msb0: "0 \ bv_to_int w1 ==> bv_msb w1 = \" by (rule bit_list_cases,simp_all) lemma bv_to_int_msb1: "bv_to_int w1 < 0 ==> bv_msb w1 = \" by (rule bit_list_cases,simp_all) lemma bv_to_int_lower_limit_gt0: assumes w0: "0 < bv_to_int w" shows "2 ^ (length (norm_signed w) - 2) \ bv_to_int w" proof - from w0 have "0 \ bv_to_int w" by simp hence [simp]: "bv_msb w = \" by (rule bv_to_int_msb0) have "2 ^ (length (norm_signed w) - 2) \ bv_to_int (norm_signed w)" proof (rule bit_list_cases [of w]) assume "w = []" with w0 show ?thesis by simp next fix w' assume weq: "w = \ # w'" thus ?thesis proof (simp add: norm_signed_Cons,safe) assume "norm_unsigned w' = []" with weq and w0 show False by (simp add: norm_empty_bv_to_nat_zero) next assume w'0: "norm_unsigned w' \ []" have "0 < bv_to_nat w'" proof (rule ccontr) assume "~ (0 < bv_to_nat w')" hence "bv_to_nat w' = 0" by arith hence "norm_unsigned w' = []" by (simp add: bv_to_nat_zero_imp_empty) with w'0 show False by simp qed with bv_to_nat_lower_limit [of w'] show "2 ^ (length (norm_unsigned w') - Suc 0) \ bv_to_nat w'" using One_nat_def int_nat_two_exp by presburger qed next fix w' assume weq: "w = \ # w'" from w0 have "bv_msb w = \" by simp with weq show ?thesis by simp qed also have "... = bv_to_int w" by simp finally show ?thesis . qed lemma norm_signed_result: "norm_signed w = [] \ norm_signed w = [\] \ bv_msb (norm_signed w) \ bv_msb (tl (norm_signed w))" apply (rule bit_list_cases [of w],simp_all) apply (case_tac "bs",simp_all) apply (case_tac "a",simp_all) apply (simp add: norm_signed_Cons) apply safe apply simp proof - fix l assume msb: "\ = bv_msb (norm_unsigned l)" assume "norm_unsigned l \ []" with norm_unsigned_result [of l] have "bv_msb (norm_unsigned l) = \" by simp with msb show False by simp next fix xs assume p: "\ = bv_msb (tl (norm_signed (\ # xs)))" have "\ \ bv_msb (tl (norm_signed (\ # xs)))" by (rule bit_list_induct [of _ xs],simp_all) with p show False by simp qed lemma bv_to_int_upper_limit_lem1: assumes w0: "bv_to_int w < -1" shows "bv_to_int w < - (2 ^ (length (norm_signed w) - 2))" proof - from w0 have "bv_to_int w < 0" by simp hence msbw [simp]: "bv_msb w = \" by (rule bv_to_int_msb1) have "bv_to_int w = bv_to_int (norm_signed w)" by simp also from norm_signed_result [of w] have "... < - (2 ^ (length (norm_signed w) - 2))" proof safe assume "norm_signed w = []" hence "bv_to_int (norm_signed w) = 0" by simp with w0 show ?thesis by simp next assume "norm_signed w = [\]" hence "bv_to_int (norm_signed w) = -1" by simp with w0 show ?thesis by simp next assume "bv_msb (norm_signed w) \ bv_msb (tl (norm_signed w))" hence msb_tl: "\ \ bv_msb (tl (norm_signed w))" by simp show "bv_to_int (norm_signed w) < - (2 ^ (length (norm_signed w) - 2))" proof (rule bit_list_cases [of "norm_signed w"]) assume "norm_signed w = []" hence "bv_to_int (norm_signed w) = 0" by simp with w0 show ?thesis by simp next fix w' assume nw: "norm_signed w = \ # w'" from msbw have "bv_msb (norm_signed w) = \" by simp with nw show ?thesis by simp next fix w' assume weq: "norm_signed w = \ # w'" show ?thesis proof (rule bit_list_cases [of w']) assume w'eq: "w' = []" from w0 have "bv_to_int (norm_signed w) < -1" by simp with w'eq and weq show ?thesis by simp next fix w'' assume w'eq: "w' = \ # w''" show ?thesis by (simp add: weq w'eq) next fix w'' assume w'eq: "w' = \ # w''" with weq and msb_tl show ?thesis by simp qed qed qed finally show ?thesis . qed lemma length_int_to_bv_upper_limit_gt0: assumes w0: "0 < i" and wk: "i \ 2 ^ (k - 1) - 1" shows "length (int_to_bv i) \ k" proof (rule ccontr) from w0 wk have k1: "1 < k" by (cases "k - 1",simp_all) assume "~ length (int_to_bv i) \ k" hence "k < length (int_to_bv i)" by simp hence "k \ length (int_to_bv i) - 1" by arith hence a: "k - 1 \ length (int_to_bv i) - 2" by arith hence "(2::int) ^ (k - 1) \ 2 ^ (length (int_to_bv i) - 2)" by simp also have "... \ i" proof - have "2 ^ (length (norm_signed (int_to_bv i)) - 2) \ bv_to_int (int_to_bv i)" proof (rule bv_to_int_lower_limit_gt0) from w0 show "0 < bv_to_int (int_to_bv i)" by simp qed thus ?thesis by simp qed finally have "2 ^ (k - 1) \ i" . with wk show False by simp qed lemma pos_length_pos: assumes i0: "0 < bv_to_int w" shows "0 < length w" proof - from norm_signed_result [of w] have "0 < length (norm_signed w)" proof (auto) assume ii: "norm_signed w = []" have "bv_to_int (norm_signed w) = 0" by (subst ii) simp hence "bv_to_int w = 0" by simp with i0 show False by simp next assume ii: "norm_signed w = []" assume jj: "bv_msb w \ \" have "\ = bv_msb (norm_signed w)" by (subst ii) simp also have "... \ \" by (simp add: jj) finally show False by simp qed also have "... \ length w" by (rule norm_signed_length) finally show ?thesis . qed lemma neg_length_pos: assumes i0: "bv_to_int w < -1" shows "0 < length w" proof - from norm_signed_result [of w] have "0 < length (norm_signed w)" proof (auto) assume ii: "norm_signed w = []" have "bv_to_int (norm_signed w) = 0" by (subst ii) simp hence "bv_to_int w = 0" by simp with i0 show False by simp next assume ii: "norm_signed w = []" assume jj: "bv_msb w \ \" have "\ = bv_msb (norm_signed w)" by (subst ii) simp also have "... \ \" by (simp add: jj) finally show False by simp qed also have "... \ length w" by (rule norm_signed_length) finally show ?thesis . qed lemma length_int_to_bv_lower_limit_gt0: assumes wk: "2 ^ (k - 1) \ i" shows "k < length (int_to_bv i)" proof (rule ccontr) have "0 < (2::int) ^ (k - 1)" by (rule zero_less_power) simp also have "... \ i" by (rule wk) finally have i0: "0 < i" . have lii0: "0 < length (int_to_bv i)" apply (rule pos_length_pos) apply (simp,rule i0) done assume "~ k < length (int_to_bv i)" hence "length (int_to_bv i) \ k" by simp with lii0 have a: "length (int_to_bv i) - 1 \ k - 1" by arith have "i < 2 ^ (length (int_to_bv i) - 1)" proof - have "i = bv_to_int (int_to_bv i)" by simp also have "... < 2 ^ (length (int_to_bv i) - 1)" by (rule bv_to_int_upper_range) finally show ?thesis . qed also have "(2::int) ^ (length (int_to_bv i) - 1) \ 2 ^ (k - 1)" using a by simp finally have "i < 2 ^ (k - 1)" . with wk show False by simp qed lemma length_int_to_bv_upper_limit_lem1: assumes w1: "i < -1" and wk: "- (2 ^ (k - 1)) \ i" shows "length (int_to_bv i) \ k" proof (rule ccontr) from w1 wk have k1: "1 < k" by (cases "k - 1") simp_all assume "~ length (int_to_bv i) \ k" hence "k < length (int_to_bv i)" by simp hence "k \ length (int_to_bv i) - 1" by arith hence a: "k - 1 \ length (int_to_bv i) - 2" by arith have "i < - (2 ^ (length (int_to_bv i) - 2))" proof - have "i = bv_to_int (int_to_bv i)" by simp also have "... < - (2 ^ (length (norm_signed (int_to_bv i)) - 2))" by (rule bv_to_int_upper_limit_lem1,simp,rule w1) finally show ?thesis by simp qed also have "... \ -(2 ^ (k - 1))" proof - have "(2::int) ^ (k - 1) \ 2 ^ (length (int_to_bv i) - 2)" using a by simp thus ?thesis by simp qed finally have "i < -(2 ^ (k - 1))" . with wk show False by simp qed lemma length_int_to_bv_lower_limit_lem1: assumes wk: "i < -(2 ^ (k - 1))" shows "k < length (int_to_bv i)" proof (rule ccontr) from wk have "i \ -(2 ^ (k - 1)) - 1" by simp also have "... < -1" proof - have "0 < (2::int) ^ (k - 1)" by (rule zero_less_power) simp hence "-((2::int) ^ (k - 1)) < 0" by simp thus ?thesis by simp qed finally have i1: "i < -1" . have lii0: "0 < length (int_to_bv i)" apply (rule neg_length_pos) apply (simp, rule i1) done assume "~ k < length (int_to_bv i)" hence "length (int_to_bv i) \ k" by simp with lii0 have a: "length (int_to_bv i) - 1 \ k - 1" by arith hence "(2::int) ^ (length (int_to_bv i) - 1) \ 2 ^ (k - 1)" by simp hence "-((2::int) ^ (k - 1)) \ - (2 ^ (length (int_to_bv i) - 1))" by simp also have "... \ i" proof - have "- (2 ^ (length (int_to_bv i) - 1)) \ bv_to_int (int_to_bv i)" by (rule bv_to_int_lower_range) also have "... = i" by simp finally show ?thesis . qed finally have "-(2 ^ (k - 1)) \ i" . with wk show False by simp qed subsection \Signed Arithmetic Operations\ subsubsection \Conversion from unsigned to signed\ definition utos :: "bit list => bit list" where "utos w = norm_signed (\ # w)" lemma utos_type [simp]: "utos (norm_unsigned w) = utos w" by (simp add: utos_def norm_signed_Cons) lemma utos_returntype [simp]: "norm_signed (utos w) = utos w" by (simp add: utos_def) lemma utos_length: "length (utos w) \ Suc (length w)" by (simp add: utos_def norm_signed_Cons) lemma bv_to_int_utos: "bv_to_int (utos w) = int (bv_to_nat w)" proof (simp add: utos_def norm_signed_Cons, safe) assume "norm_unsigned w = []" hence "bv_to_nat (norm_unsigned w) = 0" by simp thus "bv_to_nat w = 0" by simp qed subsubsection \Unary minus\ definition bv_uminus :: "bit list => bit list" where "bv_uminus w = int_to_bv (- bv_to_int w)" lemma bv_uminus_type [simp]: "bv_uminus (norm_signed w) = bv_uminus w" by (simp add: bv_uminus_def) lemma bv_uminus_returntype [simp]: "norm_signed (bv_uminus w) = bv_uminus w" by (simp add: bv_uminus_def) lemma bv_uminus_length: "length (bv_uminus w) \ Suc (length w)" proof - have "1 < -bv_to_int w \ -bv_to_int w = 1 \ -bv_to_int w = 0 \ -bv_to_int w = -1 \ -bv_to_int w < -1" by arith thus ?thesis proof safe assume p: "1 < - bv_to_int w" have lw: "0 < length w" apply (rule neg_length_pos) using p apply simp done show ?thesis proof (simp add: bv_uminus_def,rule length_int_to_bv_upper_limit_gt0,simp_all) from p show "bv_to_int w < 0" by simp next have "-(2^(length w - 1)) \ bv_to_int w" by (rule bv_to_int_lower_range) hence "- bv_to_int w \ 2^(length w - 1)" by simp also from lw have "... < 2 ^ length w" by simp finally show "- bv_to_int w < 2 ^ length w" by simp qed next assume p: "- bv_to_int w = 1" hence lw: "0 < length w" by (cases w) simp_all from p show ?thesis apply (simp add: bv_uminus_def) using lw apply (simp (no_asm) add: nat_to_bv_non0) done next assume "- bv_to_int w = 0" thus ?thesis by (simp add: bv_uminus_def) next assume p: "- bv_to_int w = -1" thus ?thesis by (simp add: bv_uminus_def) next assume p: "- bv_to_int w < -1" show ?thesis apply (simp add: bv_uminus_def) apply (rule length_int_to_bv_upper_limit_lem1) apply (rule p) apply simp proof - have "bv_to_int w < 2 ^ (length w - 1)" by (rule bv_to_int_upper_range) also have "... \ 2 ^ length w" by simp finally show "bv_to_int w \ 2 ^ length w" by simp qed qed qed lemma bv_uminus_length_utos: "length (bv_uminus (utos w)) \ Suc (length w)" proof - have "-bv_to_int (utos w) = 0 \ -bv_to_int (utos w) = -1 \ -bv_to_int (utos w) < -1" by (simp add: bv_to_int_utos, arith) thus ?thesis proof safe assume "-bv_to_int (utos w) = 0" thus ?thesis by (simp add: bv_uminus_def) next assume "-bv_to_int (utos w) = -1" thus ?thesis by (simp add: bv_uminus_def) next assume p: "-bv_to_int (utos w) < -1" show ?thesis apply (simp add: bv_uminus_def) apply (rule length_int_to_bv_upper_limit_lem1) apply (rule p) apply (simp add: bv_to_int_utos) using bv_to_nat_upper_range [of w] int_nat_two_exp apply presburger done qed qed definition bv_sadd :: "[bit list, bit list ] => bit list" where "bv_sadd w1 w2 = int_to_bv (bv_to_int w1 + bv_to_int w2)" lemma bv_sadd_type1 [simp]: "bv_sadd (norm_signed w1) w2 = bv_sadd w1 w2" by (simp add: bv_sadd_def) lemma bv_sadd_type2 [simp]: "bv_sadd w1 (norm_signed w2) = bv_sadd w1 w2" by (simp add: bv_sadd_def) lemma bv_sadd_returntype [simp]: "norm_signed (bv_sadd w1 w2) = bv_sadd w1 w2" by (simp add: bv_sadd_def) lemma adder_helper: assumes lw: "0 < max (length w1) (length w2)" shows "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \ 2 ^ max (length w1) (length w2)" proof - have "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \ 2 ^ (max (length w1) (length w2) - 1) + 2 ^ (max (length w1) (length w2) - 1)" by (auto simp:max_def) also have "... = 2 ^ max (length w1) (length w2)" proof - from lw show ?thesis apply simp apply (subst power_Suc [symmetric]) apply simp done qed finally show ?thesis . qed lemma bv_sadd_length: "length (bv_sadd w1 w2) \ Suc (max (length w1) (length w2))" proof - let ?Q = "bv_to_int w1 + bv_to_int w2" have helper: "?Q \ 0 ==> 0 < max (length w1) (length w2)" proof - assume p: "?Q \ 0" show "0 < max (length w1) (length w2)" proof (simp add: less_max_iff_disj,rule) assume [simp]: "w1 = []" show "w2 \ []" proof (rule ccontr,simp) assume [simp]: "w2 = []" from p show False by simp qed qed qed have "0 < ?Q \ ?Q = 0 \ ?Q = -1 \ ?Q < -1" by arith thus ?thesis proof safe assume "?Q = 0" thus ?thesis by (simp add: bv_sadd_def) next assume "?Q = -1" thus ?thesis by (simp add: bv_sadd_def) next assume p: "0 < ?Q" show ?thesis apply (simp add: bv_sadd_def) apply (rule length_int_to_bv_upper_limit_gt0) apply (rule p) proof simp from bv_to_int_upper_range [of w2] have "bv_to_int w2 \ 2 ^ (length w2 - 1)" by simp with bv_to_int_upper_range [of w1] have "bv_to_int w1 + bv_to_int w2 < (2 ^ (length w1 - 1)) + (2 ^ (length w2 - 1))" by (rule add_less_le_mono) also have "... \ 2 ^ max (length w1) (length w2)" apply (rule adder_helper) apply (rule helper) using p apply simp done finally show "?Q < 2 ^ max (length w1) (length w2)" . qed next assume p: "?Q < -1" show ?thesis apply (simp add: bv_sadd_def) apply (rule length_int_to_bv_upper_limit_lem1,simp_all) apply (rule p) proof - have "(2 ^ (length w1 - 1)) + 2 ^ (length w2 - 1) \ (2::int) ^ max (length w1) (length w2)" apply (rule adder_helper) apply (rule helper) using p apply simp done hence "-((2::int) ^ max (length w1) (length w2)) \ - (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1))" by simp also have "- (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1)) \ ?Q" apply (rule add_mono) apply (rule bv_to_int_lower_range [of w1]) apply (rule bv_to_int_lower_range [of w2]) done finally show "- (2^max (length w1) (length w2)) \ ?Q" . qed qed qed definition bv_sub :: "[bit list, bit list] => bit list" where "bv_sub w1 w2 = bv_sadd w1 (bv_uminus w2)" lemma bv_sub_type1 [simp]: "bv_sub (norm_signed w1) w2 = bv_sub w1 w2" by (simp add: bv_sub_def) lemma bv_sub_type2 [simp]: "bv_sub w1 (norm_signed w2) = bv_sub w1 w2" by (simp add: bv_sub_def) lemma bv_sub_returntype [simp]: "norm_signed (bv_sub w1 w2) = bv_sub w1 w2" by (simp add: bv_sub_def) lemma bv_sub_length: "length (bv_sub w1 w2) \ Suc (max (length w1) (length w2))" proof (cases "bv_to_int w2 = 0") assume p: "bv_to_int w2 = 0" show ?thesis proof (simp add: bv_sub_def bv_sadd_def bv_uminus_def p) have "length (norm_signed w1) \ length w1" by (rule norm_signed_length) also have "... \ max (length w1) (length w2)" by (rule max.cobounded1) also have "... \ Suc (max (length w1) (length w2))" by arith finally show "length (norm_signed w1) \ Suc (max (length w1) (length w2))" . qed next assume "bv_to_int w2 \ 0" hence "0 < length w2" by (cases w2,simp_all) hence lmw: "0 < max (length w1) (length w2)" by arith let ?Q = "bv_to_int w1 - bv_to_int w2" have "0 < ?Q \ ?Q = 0 \ ?Q = -1 \ ?Q < -1" by arith thus ?thesis proof safe assume "?Q = 0" thus ?thesis by (simp add: bv_sub_def bv_sadd_def bv_uminus_def) next assume "?Q = -1" thus ?thesis by (simp add: bv_sub_def bv_sadd_def bv_uminus_def) next assume p: "0 < ?Q" show ?thesis apply (simp add: bv_sub_def bv_sadd_def bv_uminus_def) apply (rule length_int_to_bv_upper_limit_gt0) apply (rule p) proof simp from bv_to_int_lower_range [of w2] have v2: "- bv_to_int w2 \ 2 ^ (length w2 - 1)" by simp have "bv_to_int w1 + - bv_to_int w2 < (2 ^ (length w1 - 1)) + (2 ^ (length w2 - 1))" apply (rule add_less_le_mono) apply (rule bv_to_int_upper_range [of w1]) apply (rule v2) done also have "... \ 2 ^ max (length w1) (length w2)" apply (rule adder_helper) apply (rule lmw) done finally show "?Q < 2 ^ max (length w1) (length w2)" by simp qed next assume p: "?Q < -1" show ?thesis apply (simp add: bv_sub_def bv_sadd_def bv_uminus_def) apply (rule length_int_to_bv_upper_limit_lem1) apply (rule p) proof simp have "(2 ^ (length w1 - 1)) + 2 ^ (length w2 - 1) \ (2::int) ^ max (length w1) (length w2)" apply (rule adder_helper) apply (rule lmw) done hence "-((2::int) ^ max (length w1) (length w2)) \ - (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1))" by simp also have "- (2 ^ (length w1 - 1)) + -(2 ^ (length w2 - 1)) \ bv_to_int w1 + -bv_to_int w2" apply (rule add_mono) apply (rule bv_to_int_lower_range [of w1]) using bv_to_int_upper_range [of w2] apply simp done finally show "- (2^max (length w1) (length w2)) \ ?Q" by simp qed qed qed definition bv_smult :: "[bit list, bit list] => bit list" where "bv_smult w1 w2 = int_to_bv (bv_to_int w1 * bv_to_int w2)" lemma bv_smult_type1 [simp]: "bv_smult (norm_signed w1) w2 = bv_smult w1 w2" by (simp add: bv_smult_def) lemma bv_smult_type2 [simp]: "bv_smult w1 (norm_signed w2) = bv_smult w1 w2" by (simp add: bv_smult_def) lemma bv_smult_returntype [simp]: "norm_signed (bv_smult w1 w2) = bv_smult w1 w2" by (simp add: bv_smult_def) lemma bv_smult_length: "length (bv_smult w1 w2) \ length w1 + length w2" proof - let ?Q = "bv_to_int w1 * bv_to_int w2" have lmw: "?Q \ 0 ==> 0 < length w1 \ 0 < length w2" by auto have "0 < ?Q \ ?Q = 0 \ ?Q = -1 \ ?Q < -1" by arith thus ?thesis proof (safe dest!: iffD1 [OF mult_eq_0_iff]) assume "bv_to_int w1 = 0" thus ?thesis by (simp add: bv_smult_def) next assume "bv_to_int w2 = 0" thus ?thesis by (simp add: bv_smult_def) next assume p: "?Q = -1" show ?thesis apply (simp add: bv_smult_def p) apply (cut_tac lmw) apply arith using p apply simp done next assume p: "0 < ?Q" thus ?thesis proof (simp add: zero_less_mult_iff,safe) assume bi1: "0 < bv_to_int w1" assume bi2: "0 < bv_to_int w2" show ?thesis apply (simp add: bv_smult_def) apply (rule length_int_to_bv_upper_limit_gt0) apply (rule p) proof simp have "?Q < 2 ^ (length w1 - 1) * 2 ^ (length w2 - 1)" apply (rule mult_strict_mono) apply (rule bv_to_int_upper_range) apply (rule bv_to_int_upper_range) apply (rule zero_less_power) apply simp using bi2 apply simp done also have "... \ 2 ^ (length w1 + length w2 - Suc 0)" apply simp apply (subst power_add [symmetric]) apply simp done finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" . qed next assume bi1: "bv_to_int w1 < 0" assume bi2: "bv_to_int w2 < 0" show ?thesis apply (simp add: bv_smult_def) apply (rule length_int_to_bv_upper_limit_gt0) apply (rule p) proof simp have "-bv_to_int w1 * -bv_to_int w2 \ 2 ^ (length w1 - 1) * 2 ^(length w2 - 1)" apply (rule mult_mono) using bv_to_int_lower_range [of w1] apply simp using bv_to_int_lower_range [of w2] apply simp apply (rule zero_le_power,simp) using bi2 apply simp done hence "?Q \ 2 ^ (length w1 - 1) * 2 ^(length w2 - 1)" by simp also have "... < 2 ^ (length w1 + length w2 - Suc 0)" apply simp apply (subst power_add [symmetric]) apply simp apply (cut_tac lmw) apply arith apply (cut_tac p) apply arith done finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" . qed qed next assume p: "?Q < -1" show ?thesis apply (subst bv_smult_def) apply (rule length_int_to_bv_upper_limit_lem1) apply (rule p) proof simp have "(2::int) ^ (length w1 - 1) * 2 ^(length w2 - 1) \ 2 ^ (length w1 + length w2 - Suc 0)" apply simp apply (subst power_add [symmetric]) apply simp done hence "-((2::int) ^ (length w1 + length w2 - Suc 0)) \ -(2^(length w1 - 1) * 2 ^ (length w2 - 1))" by simp also have "... \ ?Q" proof - from p have q: "bv_to_int w1 * bv_to_int w2 < 0" by simp thus ?thesis proof (simp add: mult_less_0_iff,safe) assume bi1: "0 < bv_to_int w1" assume bi2: "bv_to_int w2 < 0" have "-bv_to_int w2 * bv_to_int w1 \ ((2::int)^(length w2 - 1)) * (2 ^ (length w1 - 1))" apply (rule mult_mono) using bv_to_int_lower_range [of w2] apply simp using bv_to_int_upper_range [of w1] apply simp apply (rule zero_le_power,simp) using bi1 apply simp done hence "-?Q \ ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))" by (simp add: ac_simps) thus "-(((2::int)^(length w1 - Suc 0)) * (2 ^ (length w2 - Suc 0))) \ ?Q" by simp next assume bi1: "bv_to_int w1 < 0" assume bi2: "0 < bv_to_int w2" have "-bv_to_int w1 * bv_to_int w2 \ ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))" apply (rule mult_mono) using bv_to_int_lower_range [of w1] apply simp using bv_to_int_upper_range [of w2] apply simp apply (rule zero_le_power,simp) using bi2 apply simp done hence "-?Q \ ((2::int)^(length w1 - 1)) * (2 ^ (length w2 - 1))" by (simp add: ac_simps) thus "-(((2::int)^(length w1 - Suc 0)) * (2 ^ (length w2 - Suc 0))) \ ?Q" by simp qed qed finally show "-(2 ^ (length w1 + length w2 - Suc 0)) \ ?Q" . qed qed qed lemma bv_msb_one: "bv_msb w = \ ==> bv_to_nat w \ 0" by (cases w) simp_all lemma bv_smult_length_utos: "length (bv_smult (utos w1) w2) \ length w1 + length w2" proof - let ?Q = "bv_to_int (utos w1) * bv_to_int w2" have lmw: "?Q \ 0 ==> 0 < length (utos w1) \ 0 < length w2" by auto have "0 < ?Q \ ?Q = 0 \ ?Q = -1 \ ?Q < -1" by arith thus ?thesis proof (safe dest!: iffD1 [OF mult_eq_0_iff]) assume "bv_to_int (utos w1) = 0" thus ?thesis by (simp add: bv_smult_def) next assume "bv_to_int w2 = 0" thus ?thesis by (simp add: bv_smult_def) next assume p: "0 < ?Q" thus ?thesis proof (simp add: zero_less_mult_iff,safe) assume biw2: "0 < bv_to_int w2" show ?thesis apply (simp add: bv_smult_def) apply (rule length_int_to_bv_upper_limit_gt0) apply (rule p) proof simp have "?Q < 2 ^ length w1 * 2 ^ (length w2 - 1)" apply (rule mult_strict_mono) apply (simp add: bv_to_int_utos bv_to_nat_upper_range int_nat_two_exp del: of_nat_power) apply (rule bv_to_int_upper_range) apply (rule zero_less_power,simp) using biw2 apply simp done also have "... \ 2 ^ (length w1 + length w2 - Suc 0)" apply simp apply (subst power_add [symmetric]) apply simp apply (cut_tac lmw) apply arith using p apply auto done finally show "?Q < 2 ^ (length w1 + length w2 - Suc 0)" . qed next assume "bv_to_int (utos w1) < 0" thus ?thesis by (simp add: bv_to_int_utos) qed next assume p: "?Q = -1" thus ?thesis apply (simp add: bv_smult_def) apply (cut_tac lmw) apply arith apply simp done next assume p: "?Q < -1" show ?thesis apply (subst bv_smult_def) apply (rule length_int_to_bv_upper_limit_lem1) apply (rule p) proof simp have "(2::int) ^ length w1 * 2 ^(length w2 - 1) \ 2 ^ (length w1 + length w2 - Suc 0)" apply simp apply (subst power_add [symmetric]) apply simp apply (cut_tac lmw) apply arith apply (cut_tac p) apply arith done hence "-((2::int) ^ (length w1 + length w2 - Suc 0)) \ -(2^ length w1 * 2 ^ (length w2 - 1))" by simp also have "... \ ?Q" proof - from p have q: "bv_to_int (utos w1) * bv_to_int w2 < 0" by simp thus ?thesis proof (simp add: mult_less_0_iff,safe) assume bi1: "0 < bv_to_int (utos w1)" assume bi2: "bv_to_int w2 < 0" have "-bv_to_int w2 * bv_to_int (utos w1) \ ((2::int)^(length w2 - 1)) * (2 ^ length w1)" apply (rule mult_mono) using bv_to_int_lower_range [of w2] apply simp apply (simp add: bv_to_int_utos) using bv_to_nat_upper_range [of w1] apply (simp add: int_nat_two_exp del: of_nat_power) apply (rule zero_le_power,simp) using bi1 apply simp done hence "-?Q \ ((2::int)^length w1) * (2 ^ (length w2 - 1))" by (simp add: ac_simps) thus "-(((2::int)^length w1) * (2 ^ (length w2 - Suc 0))) \ ?Q" by simp next assume bi1: "bv_to_int (utos w1) < 0" thus "-(((2::int)^length w1) * (2 ^ (length w2 - Suc 0))) \ ?Q" by (simp add: bv_to_int_utos) qed qed finally show "-(2 ^ (length w1 + length w2 - Suc 0)) \ ?Q" . qed qed qed lemma bv_smult_sym: "bv_smult w1 w2 = bv_smult w2 w1" by (simp add: bv_smult_def ac_simps) subsection \Structural operations\ definition bv_select :: "[bit list,nat] => bit" where "bv_select w i = w ! (length w - 1 - i)" definition bv_chop :: "[bit list,nat] => bit list * bit list" where "bv_chop w i = (let len = length w in (take (len - i) w,drop (len - i) w))" definition bv_slice :: "[bit list,nat*nat] => bit list" where "bv_slice w = (\(b,e). fst (bv_chop (snd (bv_chop w (b+1))) e))" lemma bv_select_rev: assumes notnull: "n < length w" shows "bv_select w n = rev w ! n" proof - have "\n. n < length w --> bv_select w n = rev w ! n" proof (rule length_induct [of _ w],auto simp add: bv_select_def) fix xs :: "bit list" fix n assume ind: "\ys::bit list. length ys < length xs --> (\n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n)" assume notx: "n < length xs" show "xs ! (length xs - Suc n) = rev xs ! n" proof (cases xs) assume "xs = []" with notx show ?thesis by simp next fix y ys assume [simp]: "xs = y # ys" show ?thesis proof (auto simp add: nth_append) assume noty: "n < length ys" from spec [OF ind,of ys] have "\n. n < length ys --> ys ! (length ys - Suc n) = rev ys ! n" by simp hence "n < length ys --> ys ! (length ys - Suc n) = rev ys ! n" .. from this and noty show "ys ! (length ys - Suc n) = rev ys ! n" .. next assume "~ n < length ys" hence x: "length ys \ n" by simp from notx have "n < Suc (length ys)" by simp hence "n \ length ys" by simp with x have "length ys = n" by simp thus "y = [y] ! (n - length ys)" by simp qed qed qed then have "n < length w --> bv_select w n = rev w ! n" .. from this and notnull show ?thesis .. qed lemma bv_chop_append: "bv_chop (w1 @ w2) (length w2) = (w1,w2)" by (simp add: bv_chop_def Let_def) lemma append_bv_chop_id: "fst (bv_chop w l) @ snd (bv_chop w l) = w" by (simp add: bv_chop_def Let_def) lemma bv_chop_length_fst [simp]: "length (fst (bv_chop w i)) = length w - i" by (simp add: bv_chop_def Let_def) lemma bv_chop_length_snd [simp]: "length (snd (bv_chop w i)) = min i (length w)" by (simp add: bv_chop_def Let_def) lemma bv_slice_length [simp]: "[| j \ i; i < length w |] ==> length (bv_slice w (i,j)) = i - j + 1" by (auto simp add: bv_slice_def) definition length_nat :: "nat => nat" where [code del]: "length_nat x = (LEAST n. x < 2 ^ n)" lemma length_nat: "length (nat_to_bv n) = length_nat n" apply (simp add: length_nat_def) apply (rule Least_equality [symmetric]) prefer 2 apply (rule length_nat_to_bv_upper_limit) apply arith apply (rule ccontr) proof - assume "~ n < 2 ^ length (nat_to_bv n)" hence "2 ^ length (nat_to_bv n) \ n" by simp hence "length (nat_to_bv n) < length (nat_to_bv n)" by (rule length_nat_to_bv_lower_limit) thus False by simp qed lemma length_nat_0 [simp]: "length_nat 0 = 0" by (simp add: length_nat_def Least_equality) lemma length_nat_non0: assumes n0: "n \ 0" shows "length_nat n = Suc (length_nat (n div 2))" apply (simp add: length_nat [symmetric]) apply (subst nat_to_bv_non0 [of n]) apply (simp_all add: n0) done definition length_int :: "int => nat" where "length_int x = (if 0 < x then Suc (length_nat (nat x)) else if x = 0 then 0 else Suc (length_nat (nat (-x - 1))))" lemma length_int: "length (int_to_bv i) = length_int i" proof (cases "0 < i") assume i0: "0 < i" hence "length (int_to_bv i) = length (norm_signed (\ # norm_unsigned (nat_to_bv (nat i))))" by simp also from norm_unsigned_result [of "nat_to_bv (nat i)"] have "... = Suc (length_nat (nat i))" apply safe apply (simp del: norm_unsigned_nat_to_bv) apply (drule norm_empty_bv_to_nat_zero) using i0 apply simp apply (cases "norm_unsigned (nat_to_bv (nat i))") apply (drule norm_empty_bv_to_nat_zero [of "nat_to_bv (nat i)"]) using i0 apply simp apply (simp add: i0) using i0 apply (simp add: length_nat [symmetric]) done finally show ?thesis using i0 by (simp add: length_int_def) next assume "~ 0 < i" hence i0: "i \ 0" by simp show ?thesis proof (cases "i = 0") assume "i = 0" thus ?thesis by (simp add: length_int_def) next assume "i \ 0" with i0 have i0: "i < 0" by simp hence "length (int_to_bv i) = length (norm_signed (\ # bv_not (norm_unsigned (nat_to_bv (nat (- i) - 1)))))" by (simp add: int_to_bv_def nat_diff_distrib) also from norm_unsigned_result [of "nat_to_bv (nat (- i) - 1)"] have "... = Suc (length_nat (nat (- i) - 1))" apply safe apply (simp del: norm_unsigned_nat_to_bv) apply (drule norm_empty_bv_to_nat_zero [of "nat_to_bv (nat (-i) - Suc 0)"]) using i0 apply simp apply (cases "- i - 1 = 0") apply simp apply (simp add: length_nat [symmetric]) apply (cases "norm_unsigned (nat_to_bv (nat (- i) - 1))") apply simp apply simp done finally show ?thesis using i0 by (simp add: length_int_def nat_diff_distrib del: int_to_bv_lt0) qed qed lemma length_int_0 [simp]: "length_int 0 = 0" by (simp add: length_int_def) lemma length_int_gt0: "0 < i ==> length_int i = Suc (length_nat (nat i))" by (simp add: length_int_def) lemma length_int_lt0: "i < 0 ==> length_int i = Suc (length_nat (nat (- i) - 1))" by (simp add: length_int_def nat_diff_distrib) lemma bv_chopI: "[| w = w1 @ w2 ; i = length w2 |] ==> bv_chop w i = (w1,w2)" by (simp add: bv_chop_def Let_def) lemma bv_sliceI: "[| j \ i ; i < length w ; w = w1 @ w2 @ w3 ; Suc i = length w2 + j ; j = length w3 |] ==> bv_slice w (i,j) = w2" apply (simp add: bv_slice_def) apply (subst bv_chopI [of "w1 @ w2 @ w3" w1 "w2 @ w3"]) apply simp apply simp apply simp apply (subst bv_chopI [of "w2 @ w3" w2 w3],simp_all) done lemma bv_slice_bv_slice: assumes ki: "k \ i" and ij: "i \ j" and jl: "j \ l" and lw: "l < length w" shows "bv_slice w (j,i) = bv_slice (bv_slice w (l,k)) (j-k,i-k)" proof - define w1 w2 w3 w4 w5 where w_defs: "w1 = fst (bv_chop w (Suc l))" "w2 = fst (bv_chop (snd (bv_chop w (Suc l))) (Suc j))" "w3 = fst (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)" "w4 = fst (bv_chop (snd (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)) k)" "w5 = snd (bv_chop (snd (bv_chop (snd (bv_chop (snd (bv_chop w (Suc l))) (Suc j))) i)) k)" have w_def: "w = w1 @ w2 @ w3 @ w4 @ w5" by (simp add: w_defs append_bv_chop_id) from ki ij jl lw show ?thesis apply (subst bv_sliceI [where ?j = i and ?i = j and ?w = w and ?w1.0 = "w1 @ w2" and ?w2.0 = w3 and ?w3.0 = "w4 @ w5"]) apply simp_all apply (rule w_def) apply (simp add: w_defs) apply (simp add: w_defs) apply (subst bv_sliceI [where ?j = k and ?i = l and ?w = w and ?w1.0 = w1 and ?w2.0 = "w2 @ w3 @ w4" and ?w3.0 = w5]) apply simp_all apply (rule w_def) apply (simp add: w_defs) apply (simp add: w_defs) apply (subst bv_sliceI [where ?j = "i-k" and ?i = "j-k" and ?w = "w2 @ w3 @ w4" and ?w1.0 = w2 and ?w2.0 = w3 and ?w3.0 = w4]) apply simp_all apply (simp_all add: w_defs) done qed lemma bv_to_nat_extend [simp]: "bv_to_nat (bv_extend n \ w) = bv_to_nat w" apply (simp add: bv_extend_def) apply (subst bv_to_nat_dist_append) apply simp apply (induct ("n - length w")) apply simp_all done lemma bv_msb_extend_same [simp]: "bv_msb w = b ==> bv_msb (bv_extend n b w) = b" apply (simp add: bv_extend_def) apply (cases "n - length w") apply simp_all done lemma bv_to_int_extend [simp]: assumes a: "bv_msb w = b" shows "bv_to_int (bv_extend n b w) = bv_to_int w" proof (cases "bv_msb w") assume [simp]: "bv_msb w = \" with a have [simp]: "b = \" by simp show ?thesis by (simp add: bv_to_int_def) next assume [simp]: "bv_msb w = \" with a have [simp]: "b = \" by simp show ?thesis apply (simp add: bv_to_int_def) apply (simp add: bv_extend_def) apply (induct ("n - length w"), simp_all) done qed lemma length_nat_mono [simp]: "x \ y ==> length_nat x \ length_nat y" proof (rule ccontr) assume xy: "x \ y" assume "~ length_nat x \ length_nat y" hence lxly: "length_nat y < length_nat x" by simp hence "length_nat y < (LEAST n. x < 2 ^ n)" by (simp add: length_nat_def) hence "~ x < 2 ^ length_nat y" by (rule not_less_Least) hence xx: "2 ^ length_nat y \ x" by simp have yy: "y < 2 ^ length_nat y" apply (simp add: length_nat_def) - apply (rule LeastI) - apply (subgoal_tac "y < 2 ^ y",assumption) - apply (cases "0 \ y") - apply (induct y,simp_all) + using less_exp apply (rule LeastI) done with xx have "y < x" by simp with xy show False by simp qed lemma length_nat_mono_int: "x \ y ==> length_nat x \ length_nat y" by (rule length_nat_mono) arith lemma length_nat_pos [simp,intro!]: "0 < x ==> 0 < length_nat x" by (simp add: length_nat_non0) lemma length_int_mono_gt0: "[| 0 \ x ; x \ y |] ==> length_int x \ length_int y" by (cases "x = 0") (simp_all add: length_int_gt0 nat_le_eq_zle) lemma length_int_mono_lt0: "[| x \ y ; y \ 0 |] ==> length_int y \ length_int x" by (cases "y = 0") (simp_all add: length_int_lt0) lemmas [simp] = length_nat_non0 primrec fast_bv_to_nat_helper :: "[bit list, num] => num" where fast_bv_to_nat_Nil: "fast_bv_to_nat_helper [] k = k" | fast_bv_to_nat_Cons: "fast_bv_to_nat_helper (b#bs) k = fast_bv_to_nat_helper bs ((case_bit Num.Bit0 Num.Bit1 b) k)" declare fast_bv_to_nat_helper.simps [code del] lemma fast_bv_to_nat_Cons0: "fast_bv_to_nat_helper (\#bs) bin = fast_bv_to_nat_helper bs (Num.Bit0 bin)" by simp lemma fast_bv_to_nat_Cons1: "fast_bv_to_nat_helper (\#bs) bin = fast_bv_to_nat_helper bs (Num.Bit1 bin)" by simp lemma mult_Bit0_left: "Num.Bit0 m * n = Num.Bit0 (m * n)" by (simp add: num_eq_iff nat_of_num_mult distrib_right) lemma fast_bv_to_nat_def: "bv_to_nat (\ # bs) == numeral (fast_bv_to_nat_helper bs Num.One)" proof - have "\k. foldl (\bn b. 2 * bn + bitval b) (numeral k) bs = numeral (fast_bv_to_nat_helper bs k)" apply (induct bs, simp) apply (case_tac a, simp_all add: mult_Bit0_left) done thus "PROP ?thesis" by (simp add: bv_to_nat_def add: numeral_One [symmetric] del: numeral_One One_nat_def) qed declare fast_bv_to_nat_Cons [simp del] declare fast_bv_to_nat_Cons0 [simp] declare fast_bv_to_nat_Cons1 [simp] simproc_setup bv_to_nat ("bv_to_nat (x # xs)") = \ fn _ => fn ctxt => fn ct => let fun is_const_bool (Const(@{const_name True},_)) = true | is_const_bool (Const(@{const_name False},_)) = true | is_const_bool _ = false fun is_const_bit (Const(@{const_name Zero},_)) = true | is_const_bit (Const(@{const_name One},_)) = true | is_const_bit _ = false fun vec_is_usable (Const(@{const_name Nil},_)) = true | vec_is_usable (Const(@{const_name Cons},_) $ b $ bs) = vec_is_usable bs andalso is_const_bit b | vec_is_usable _ = false fun proc (Const(@{const_name bv_to_nat},_) $ (Const(@{const_name Cons},_) $ Const(@{const_name One},_) $ t)) = if vec_is_usable t then SOME (infer_instantiate ctxt [(("bs", 0), Thm.cterm_of ctxt t)] @{thm fast_bv_to_nat_def}) else NONE | proc _ = NONE in proc (Thm.term_of ct) end \ declare bv_to_nat1 [simp del] declare bv_to_nat_helper [simp del] definition bv_mapzip :: "[bit => bit => bit,bit list, bit list] => bit list" where "bv_mapzip f w1 w2 = (let g = bv_extend (max (length w1) (length w2)) \ in map (case_prod f) (zip (g w1) (g w2)))" lemma bv_length_bv_mapzip [simp]: "length (bv_mapzip f w1 w2) = max (length w1) (length w2)" by (simp add: bv_mapzip_def Let_def split: split_max) lemma bv_mapzip_Nil [simp]: "bv_mapzip f [] [] = []" by (simp add: bv_mapzip_def Let_def) lemma bv_mapzip_Cons [simp]: "length w1 = length w2 ==> bv_mapzip f (x#w1) (y#w2) = f x y # bv_mapzip f w1 w2" by (simp add: bv_mapzip_def Let_def) end diff --git a/thys/Recursion-Theory-I/PRecFinSet.thy b/thys/Recursion-Theory-I/PRecFinSet.thy --- a/thys/Recursion-Theory-I/PRecFinSet.thy +++ b/thys/Recursion-Theory-I/PRecFinSet.thy @@ -1,1063 +1,1066 @@ (* Title: Primitive recursive coding of finite sets Author: Michael Nedzelsky , 2008 Maintainer: Michael Nedzelsky *) section \Finite sets\ theory PRecFinSet imports PRecFun begin text \ We introduce a particular mapping \nat_to_set\ from natural numbers to finite sets of natural numbers and a particular mapping \set_to_nat\ from finite sets of natural numbers to natural numbers. See \cite{Rogers} and \cite{Soare} for more information. \ definition c_in :: "nat \ nat \ nat" where "c_in = (\ x u. (u div (2 ^ x)) mod 2)" lemma c_in_is_pr: "c_in \ PrimRec2" proof - from mod_is_pr power_is_pr div_is_pr have "(\ x u. (u div (2 ^ x)) mod 2) \ PrimRec2" by prec with c_in_def show ?thesis by auto qed definition nat_to_set :: "nat \ nat set" where "nat_to_set u \ {x. 2^x \ u \ c_in x u = 1}" lemma c_in_upper_bound: "c_in x u = 1 \ 2 ^ x \ u" proof - assume A: "c_in x u = 1" then have S1: "(u div (2^x)) mod 2 = 1" by (unfold c_in_def) then have S2: "u div (2^x) > 0" by arith show ?thesis proof (rule ccontr) assume "\ 2 ^ x \ u" then have "u < 2^x" by auto then have "u div (2^x) = 0" by (rule div_less) with S2 show False by auto qed qed lemma nat_to_set_upper_bound: "x \ nat_to_set u \ 2 ^ x \ u" by (simp add: nat_to_set_def) -lemma x_lt_2_x: "x < 2 ^ x" by (induct x) auto +lemma x_lt_2_x: "x < 2 ^ x" + by (rule less_exp) lemma nat_to_set_upper_bound1: "x \ nat_to_set u \ x < u" proof - assume "x \ nat_to_set u" then have S1: "2 ^ x \ u" by (simp add: nat_to_set_def) have S2: "x < 2 ^ x" by (rule x_lt_2_x) - from S1 S2 show ?thesis by auto + from S2 S1 show ?thesis + by (rule less_le_trans) qed lemma nat_to_set_upper_bound2: "nat_to_set u \ {i. i < u}" proof - from nat_to_set_upper_bound1 show ?thesis by blast qed lemma nat_to_set_is_finite: "finite (nat_to_set u)" proof - have S1: "finite {i. i {i. i nat_to_set u) = (c_in x u = 1)" by (auto simp add: nat_to_set_def c_in_upper_bound) definition log2 :: "nat \ nat" where "log2 = (\ x. Least(%z. x < 2^(z+1)))" lemma log2_at_0: "log2 0 = 0" proof - let ?v = "log2 0" have S1: "0 \ ?v" by auto have S2: "?v = Least(%(z::nat). (0::nat)<2^(z+1))" by (simp add: log2_def) have S3: "(0::nat)<2^(0+1)" by auto from S3 have S4: "Least(%(z::nat). (0::nat)<2^(z+1)) \ 0" by (rule Least_le) from S2 S4 have S5: "?v \ 0" by auto from S1 S5 have S6: "?v = 0" by auto thus ?thesis by auto qed lemma log2_at_1: "log2 1 = 0" proof - let ?v = "log2 1" have S1: "0 \ ?v" by auto have S2: "?v = Least(%(z::nat). (1::nat)<2^(z+1))" by (simp add: log2_def) have S3: "(1::nat)<2^(0+1)" by auto from S3 have S4: "Least(%(z::nat). (1::nat)<2^(z+1)) \ 0" by (rule Least_le) from S2 S4 have S5: "?v \ 0" by auto from S1 S5 have S6: "?v = 0" by auto thus ?thesis by auto qed lemma log2_le: "x > 0 \ 2 ^ log2 x \ x" proof - assume A: "x > 0" show ?thesis proof (cases) assume A1: "log2 x = 0" with A show ?thesis by auto next assume A1: "log2 x \ 0" then have S1: "log2 x > 0" by auto define y where "y = log2 x - 1" from S1 y_def have S2: "log2 x = y + 1" by auto then have S3: "y < log2 x" by auto have "2^(y+1) \ x" proof (rule ccontr) assume A2: "\ 2^(y+1) \ x" then have "x < 2^(y+1)" by auto then have "log2 x \ y" by (simp add: log2_def Least_le) with S3 show False by auto qed with S2 show ?thesis by auto qed qed lemma log2_gt: "x < 2 ^ (log2 x + 1)" proof - have "x < 2^x" by (rule x_lt_2_x) - then have S1: "x < 2^(x+1)" by simp + then have S1: "x < 2^(x+1)" + by (simp add: numeral_2_eq_2) define y where "y = x" from S1 y_def have S2: "x < 2^(y+1)" by auto let ?P = "\ z. x < 2^(z+1)" from S2 have S3: "?P y" by auto then have S4: "?P (Least ?P)" by (rule LeastI) from log2_def have S5: "log2 x = Least ?P" by (unfold log2_def, auto) from S4 S5 show ?thesis by auto qed lemma x_div_x: "x > 0 \ (x::nat) div x = 1" by auto lemma div_ge: "(k::nat) \ m div n \ n*k \ m" proof - assume A: "k \ m div n" have S1: "n * (m div n) + m mod n = m" by (rule mult_div_mod_eq) have S2: "0 \ m mod n" by auto from S1 S2 have S3: "n * (m div n) \ m" by arith from A have S4: "n * k \ n * (m div n)" by auto from S4 S3 show ?thesis by (rule order_trans) qed lemma div_lt: "m < n*k \ m div n < (k::nat)" proof - assume A: "m < n*k" show ?thesis proof (rule ccontr) assume "\ m div n < k" then have S1: "k \ m div n" by auto then have S2: "n*k \ m" by (rule div_ge) with A show False by auto qed qed lemma log2_lm1: "u > 0 \ u div 2 ^ (log2 u) = 1" proof - assume A: "u > 0" then have S1: "2^(log2 u) \ u" by (rule log2_le) have S2: "u < 2^(log2 u+1)" by (rule log2_gt) then have S3: "u < (2^log2 u)*2" by simp have "(2::nat) > 0" by simp then have "(2::nat)^log2 u > 0" by simp then have S4: "(2::nat)^log2 u div 2^log2 u = 1" by auto from S1 have S5: "(2::nat)^log2 u div 2^log2 u \ u div 2^log2 u" by (rule div_le_mono) with S4 have S6: "1 \ u div 2^log2 u" by auto from S3 have S7: "u div 2^log2 u < 2" by (rule div_lt) from S6 S7 show ?thesis by auto qed lemma log2_lm2: "u > 0 \ c_in (log2 u) u = 1" proof - assume A: "u > 0" then have S1: "u div 2 ^ (log2 u) = 1" by (rule log2_lm1) have "c_in (log2 u) u = (u div 2 ^ (log2 u)) mod 2" by (simp add: c_in_def) also from S1 have "\ = 1 mod 2" by simp also have "\ = 1" by auto finally show ?thesis by auto qed lemma log2_lm3: "log2 u < x \ c_in x u = 0" proof - assume A: "log2 u < x" then have S1: "(log2 u)+1 \ x" by auto have S2: "1 \ (2::nat)" by auto from S1 S2 have S3: "(2::nat)^ ((log2 u)+1) \ 2^x" by (rule power_increasing) have S4: "u < (2::nat)^ ((log2 u)+1)" by (rule log2_gt) from S3 S4 have S5: "u < 2^x" by auto then have S6: "u div 2^x = 0" by (rule div_less) have "c_in x u = (u div 2^x) mod 2" by (simp add: c_in_def) also from S6 have "\ = 0 mod 2" by simp also have "\ = 0" by auto finally have ?thesis by auto thus ?thesis by auto qed lemma log2_lm4: "c_in x u = 1 \ x \ log2 u" proof - assume A: "c_in x u = 1" show ?thesis proof (rule ccontr) assume "\ x \ log2 u" then have S1: "log2 u < x" by auto then have S2: "c_in x u = 0" by (rule log2_lm3) with A show False by auto qed qed lemma nat_to_set_lub: "x \ nat_to_set u \ x \ log2 u" proof - assume "x \ nat_to_set u" then have S1: "c_in x u = 1" by (simp add: x_in_u_eq) then show ?thesis by (rule log2_lm4) qed lemma log2_lm5: "u > 0 \ log2 u \ nat_to_set u" proof - assume A: "u > 0" then have "c_in (log2 u) u = 1" by (rule log2_lm2) then show ?thesis by (simp add: x_in_u_eq) qed lemma pos_imp_ne: "u > 0 \ nat_to_set u \ {}" proof - assume "u > 0" then have "log2 u \ nat_to_set u" by (rule log2_lm5) thus ?thesis by auto qed lemma empty_is_zero: "nat_to_set u = {} \ u = 0" proof (rule ccontr) assume A1: "nat_to_set u = {}" assume A2: "u \ 0" then have S1: "u > 0" by auto from S1 have "nat_to_set u \ {}" by (rule pos_imp_ne) with A1 show False by auto qed lemma log2_is_max: "u > 0 \ log2 u = Max (nat_to_set u)" proof - assume A: "u > 0" then have S1: "log2 u \ nat_to_set u" by (rule log2_lm5) define max where "max = Max (nat_to_set u)" from A have ne: "nat_to_set u \ {}" by (rule pos_imp_ne) have finite: "finite (nat_to_set u)" by (rule nat_to_set_is_finite) from max_def finite ne have max_in: "max \ nat_to_set u" by simp from max_in have S2: "c_in max u = 1" by (simp add: x_in_u_eq) then have S3: "max \ log2 u" by (rule log2_lm4) from finite ne S1 max_def have S4: "log2 u \ max" by simp from S3 S4 max_def show ?thesis by auto qed lemma zero_is_empty: "nat_to_set 0 = {}" proof - have S1: "{i. i<(0::nat)} = {}" by blast have S2: "nat_to_set 0 \ {i. i<0}" by (rule nat_to_set_upper_bound2) from S1 S2 show ?thesis by auto qed lemma ne_imp_pos: "nat_to_set u \ {} \ u > 0" proof (rule ccontr) assume A1: "nat_to_set u \ {}" assume "\ 0 < u" then have "u = 0" by auto then have "nat_to_set u = {}" by (simp add: zero_is_empty) with A1 show False by auto qed lemma div_mod_lm: "y < x \ ((u + (2::nat) ^ x) div (2::nat)^y) mod 2 = (u div (2::nat)^y) mod 2" proof - assume y_lt_x: "y < x" let ?n = "(2::nat)^y" have n_pos: "0 < ?n" by auto let ?s = "x-y" from y_lt_x have s_pos: "0 < ?s" by auto from y_lt_x have S3: "x = y + ?s" by auto from S3 have "(2::nat)^x = (2::nat)^(y + ?s)" by auto moreover have "(2::nat)^(y +?s) = (2::nat)^y * 2^ ?s" by (rule power_add) ultimately have "(2::nat)^x = 2^y * 2^?s" by auto then have S4: "u + (2::nat)^x = u + (2::nat)^y * 2^?s" by auto from n_pos have S5: "(u + (2::nat)^y * 2^?s) div 2^y = 2^?s + (u div 2^y)" by simp from S4 S5 have S6: "(u + (2::nat)^x) div 2^y = 2^?s + (u div 2^y)" by auto from s_pos have S8: "?s = (?s - 1) + 1" by auto have "(2::nat) ^ ((?s - (1::nat)) + (1::nat)) = (2::nat) ^ (?s - (1::nat)) * 2^1" by (rule power_add) with S8 have S9: "(2::nat) ^ ?s = (2::nat) ^ (?s - (1::nat)) * 2" by auto then have S10: "2^?s + (u div 2^y) = (u div 2^y) + (2::nat) ^ (?s - (1::nat)) * 2" by auto have S11: "((u div 2^y) + (2::nat) ^ (?s - (1::nat)) * 2) mod 2 = (u div 2^y) mod 2" by (rule mod_mult_self1) from S6 S10 S11 show ?thesis by auto qed lemma add_power: "u < 2^x \ nat_to_set (u + 2^x) = nat_to_set u \ {x}" proof - assume A: "u < 2^x" have log2_is_x: "log2 (u+2^x) = x" proof (unfold log2_def, rule Least_equality) from A show "u+2^x < 2^(x+1)" by auto next fix z assume A1: "u + 2^x < 2^(z+1)" show "x \ z" proof (rule ccontr) assume "\ x \ z" then have "z < x" by auto then have L1: "z+1 \ x" by auto have L2: "1 \ (2::nat)" by auto from L1 L2 have L3: "(2::nat)^(z+1) \ (2::nat)^x" by (rule power_increasing) with A1 show False by auto qed qed show ?thesis proof (rule subset_antisym) show "nat_to_set (u + 2 ^ x) \ nat_to_set u \ {x}" proof fix y assume A1: "y \ nat_to_set (u + 2 ^ x)" show "y \ nat_to_set u \ {x}" proof assume "y \ {x}" then have S1: "y \ x" by auto from A1 have "y \ log2 (u + 2 ^ x)" by (rule nat_to_set_lub) with log2_is_x have "y \ x" by auto with S1 have y_lt_x: "y < x" by auto from A1 have "c_in y (u + 2 ^ x) = 1" by (simp add: x_in_u_eq) then have S2: "((u + 2 ^ x) div 2^y) mod 2 = 1" by (unfold c_in_def) from y_lt_x have "((u + (2::nat) ^ x) div (2::nat)^y) mod 2 = (u div (2::nat)^y) mod 2" by (rule div_mod_lm) with S2 have "(u div 2^y) mod 2 = 1" by auto then have "c_in y u = 1" by (simp add: c_in_def) then show "y \ nat_to_set u" by (simp add: x_in_u_eq) qed qed next show "nat_to_set u \ {x} \ nat_to_set (u + 2 ^ x)" proof fix y assume A1: "y \ nat_to_set u \ {x}" show "y \ nat_to_set (u + 2 ^ x)" proof cases assume "y \ {x}" then have "y=x" by auto then have "y = log2 (u + 2 ^ x)" by (simp add: log2_is_x) then show ?thesis by (simp add: log2_lm5) next assume y_notin: "y \ {x}" then have y_ne_x: "y \ x" by auto from A1 y_notin have y_in: "y \ nat_to_set u" by auto have y_lt_x: "y < x" proof (rule ccontr) assume "\ y < x" with y_ne_x have y_gt_x: "x < y" by auto have "1 < (2::nat)" by auto from y_gt_x this have L1: "(2::nat)^x < 2^y" by (rule power_strict_increasing) from y_in have L2: "2^y \ u" by (rule nat_to_set_upper_bound) from L1 L2 have "(2::nat)^x < u" by arith with A show False by auto qed from y_in have "c_in y u = 1" by (simp add: x_in_u_eq) then have S2: "(u div 2^y) mod 2 = 1" by (unfold c_in_def) from y_lt_x have "((u + (2::nat) ^ x) div (2::nat)^y) mod 2 = (u div (2::nat)^y) mod 2" by (rule div_mod_lm) with S2 have "((u + (2::nat) ^ x)div 2^y) mod 2 = 1" by auto then have "c_in y (u + (2::nat) ^ x) = 1" by (simp add: c_in_def) then show "y \ nat_to_set (u + (2::nat) ^ x)" by (simp add: x_in_u_eq) qed qed qed qed theorem nat_to_set_inj: "nat_to_set u = nat_to_set v \ u = v" proof - assume A: "nat_to_set u = nat_to_set v" let ?P = "\ (n::nat). (\ (D::nat set). finite D \ card D \ n \ (\ u v. nat_to_set u = D \ nat_to_set v = D \ u = v))" have P_at_0: "?P 0" proof fix D show "finite D \ card D \ 0 \ (\u v. nat_to_set u = D \ nat_to_set v = D \ u = v)" proof (rule impI) assume A1: "finite D \ card D \ 0" from A1 have S1: "finite D" by auto from A1 have S2: "card D = 0" by auto from S1 S2 have S3: "D = {}" by auto show "(\u v. nat_to_set u = D \ nat_to_set v = D \ u = v)" proof (rule allI, rule allI) fix u v show "nat_to_set u = D \ nat_to_set v = D \ u = v" proof assume A2: " nat_to_set u = D \ nat_to_set v = D" from A2 have L1: "nat_to_set u = D" by auto from A2 have L2: "nat_to_set v = D" by auto from L1 S3 have "nat_to_set u = {}" by auto then have u_z: "u = 0" by (rule empty_is_zero) from L2 S3 have "nat_to_set v = {}" by auto then have v_z: "v = 0" by (rule empty_is_zero) from u_z v_z show "u=v" by auto qed qed qed qed have P_at_Suc: "\ n. ?P n \ ?P (Suc n)" proof - fix n assume A_n: "?P n" show "?P (Suc n)" proof fix D show "finite D \ card D \ Suc n \ (\u v. nat_to_set u = D \ nat_to_set v = D \ u = v)" proof (rule impI) assume A1: "finite D \ card D \ Suc n" from A1 have S1: "finite D" by auto from A1 have S2: "card D \ Suc n" by auto show "(\u v. nat_to_set u = D \ nat_to_set v = D \ u = v)" proof (rule allI, rule allI, rule impI) fix u v assume A2: " nat_to_set u = D \ nat_to_set v = D" from A2 have d_u_d: "nat_to_set u = D" by auto from A2 have d_v_d: "nat_to_set v = D" by auto show "u = v" proof (cases) assume A3: "D = {}" from A3 d_u_d have "nat_to_set u = {}" by auto then have u_z: "u = 0" by (rule empty_is_zero) from A3 d_v_d have "nat_to_set v = {}" by auto then have v_z: "v = 0" by (rule empty_is_zero) from u_z v_z show "u = v" by auto next assume A3: "D \ {}" from A3 d_u_d have "nat_to_set u \ {}" by auto then have u_pos: "u > 0" by (rule ne_imp_pos) from A3 d_v_d have "nat_to_set v \ {}" by auto then have v_pos: "v > 0" by (rule ne_imp_pos) define m where "m = Max D" from S1 m_def A3 have m_in: "m \ D" by auto from d_u_d m_def have m_u: "m = Max (nat_to_set u)" by auto from d_v_d m_def have m_v: "m = Max (nat_to_set v)" by auto from u_pos m_u log2_is_max have m_log_u: "m = log2 u" by auto from v_pos m_v log2_is_max have m_log_v: "m = log2 v" by auto define D1 where "D1 = D - {m}" define u1 where "u1 = u - 2^m" define v1 where "v1 = v - 2^m" have card_D1: "card D1 \ n" proof - from D1_def S1 m_in have "card D1 = (card D) - 1" by (simp add: card_Diff_singleton) with S2 show ?thesis by auto qed have u_u1: "u = u1 + 2^m" proof - from u_pos have L1: "2 ^ log2 u \ u" by (rule log2_le) with m_log_u have L2: "2 ^ m \ u" by auto with u1_def show ?thesis by auto qed have u1_d1: "nat_to_set u1 = D1" proof - from m_log_u log2_gt have "u < 2^(m+1)" by auto with u_u1 have u1_lt_2_m: "u1 < 2^m" by auto with u_u1 have L1: "nat_to_set u = nat_to_set u1 \ {m}" by (simp add: add_power) have m_notin: "m \ nat_to_set u1" proof (rule ccontr) assume "\ m \ nat_to_set u1" then have "m \ nat_to_set u1" by auto then have "2^m \ u1" by (rule nat_to_set_upper_bound) with u1_lt_2_m show False by auto qed from L1 m_notin have "nat_to_set u1 = nat_to_set u - {m}" by auto with d_u_d have "nat_to_set u1 = D - {m}" by auto with D1_def show ?thesis by auto qed have v_v1: "v = v1 + 2^m" proof - from v_pos have L1: "2 ^ log2 v \ v" by (rule log2_le) with m_log_v have L2: "2 ^ m \ v" by auto with v1_def show ?thesis by auto qed have v1_d1: "nat_to_set v1 = D1" proof - from m_log_v log2_gt have "v < 2^(m+1)" by auto with v_v1 have v1_lt_2_m: "v1 < 2^m" by auto with v_v1 have L1: "nat_to_set v = nat_to_set v1 \ {m}" by (simp add: add_power) have m_notin: "m \ nat_to_set v1" proof (rule ccontr) assume "\ m \ nat_to_set v1" then have "m \ nat_to_set v1" by auto then have "2^m \ v1" by (rule nat_to_set_upper_bound) with v1_lt_2_m show False by auto qed from L1 m_notin have "nat_to_set v1 = nat_to_set v - {m}" by auto with d_v_d have "nat_to_set v1 = D - {m}" by auto with D1_def show ?thesis by auto qed from S1 D1_def have P1: "finite D1" by auto with card_D1 have P2: "finite D1 \ card D1 \ n" by auto from A_n P2 have "(\u v. nat_to_set u = D1 \ nat_to_set v = D1 \ u = v)" by auto with u1_d1 v1_d1 have "u1 = v1" by auto with u_u1 v_v1 show "u = v" by auto qed qed qed qed qed from P_at_0 P_at_Suc have main: "\ n. ?P n" by (rule nat.induct) define D where "D = nat_to_set u" from D_def A have P1: "nat_to_set u = D" by auto from D_def A have P2: "nat_to_set v = D" by auto from D_def nat_to_set_is_finite have d_finite: "finite D" by auto define n where "n = card D" from n_def d_finite have card_le: "card D \ n" by auto from d_finite card_le have P3: "finite D \ card D \ n" by auto with main have P4: "\u v. nat_to_set u = D \ nat_to_set v = D \ u = v" by auto with P1 P2 show "u = v" by auto qed definition set_to_nat :: "nat set => nat" where "set_to_nat = (\ D. sum (\ x. 2 ^ x) D)" lemma two_power_sum: "sum (\ x. (2::nat) ^ x) {i. i< Suc m} = (2 ^ Suc m) - 1" proof (induct m) show "sum (\ x. (2::nat) ^ x) {i. i< Suc 0} = (2 ^ Suc 0) - 1" by auto next fix n assume A: "sum (\ x. (2::nat) ^ x) {i. i< Suc n} = (2 ^ Suc n) - 1" show "sum (\ x. (2::nat) ^ x) {i. i< Suc (Suc n)} = (2 ^ Suc (Suc n)) - 1" proof - let ?f = "\ x. (2::nat) ^ x" have S1: "{i. i< Suc (Suc n)} = {i. i \ Suc n}" by auto have S2: "{i. i \ Suc n} = {i. i < Suc n} \ { Suc n}" by auto from S1 S2 have S3: "{i. i< Suc (Suc n)} = {i. i < Suc n} \ { Suc n}" by auto have S4: "{i. i < Suc n} = (\ x. x) ` {i. i < Suc n}" by auto then have S5: "finite {i. i < Suc n}" by (rule nat_seg_image_imp_finite) have S6: "Suc n \ {i. i < Suc n}" by auto from S5 S6 sum.insert have S7: "sum ?f ({i. i< Suc n} \ {Suc n}) = 2 ^ Suc n + sum ?f {i. i< Suc n}" by auto from S3 have "sum ?f {i. i< Suc (Suc n)} = sum ?f ({i. i< Suc n} \ {Suc n})" by auto also from S7 have "\ = 2 ^ Suc n + sum ?f {i. i< Suc n}" by auto also from A have "\ = 2 ^ Suc n + (((2::nat) ^ Suc n)-(1::nat))" by auto also have "\ = (2 ^ Suc (Suc n)) - 1" by auto finally show ?thesis by auto qed qed lemma finite_interval: "finite {i. (i::nat) x. x) ` {i. i < m}" by auto then show ?thesis by (rule nat_seg_image_imp_finite) qed lemma set_to_nat_at_empty: "set_to_nat {} = 0" by (unfold set_to_nat_def, rule sum.empty) lemma set_to_nat_of_interval: "set_to_nat {i. (i::nat) finite B; A \ B\ \ set_to_nat A \ set_to_nat B" proof - assume b_finite: "finite B" assume a_le_b: "A \ B" let ?f = "\ (x::nat). (2::nat) ^ x" have S1: "set_to_nat A = sum ?f A" by (simp add: set_to_nat_def) have S2: "set_to_nat B = sum ?f B" by (simp add: set_to_nat_def) have S3: "\ x. x \ B - A \ 0 \ ?f x" by auto from b_finite a_le_b S3 have "sum ?f A \ sum ?f B" by (rule sum_mono2) with S1 S2 show ?thesis by auto qed theorem nat_to_set_srj: "finite (D::nat set) \ nat_to_set (set_to_nat D) = D" proof - assume A: "finite D" let ?P = "\ (n::nat). (\ (D::nat set). finite D \ card D = n \ nat_to_set (set_to_nat D) = D)" have P_at_0: "?P 0" proof (rule allI) fix D show "finite D \ card D = 0 \ nat_to_set (set_to_nat D) = D" proof assume A1: "finite D \ card D = 0" from A1 have S1: "finite D" by auto from A1 have S2: "card D = 0" by auto from S1 S2 have S3: "D = {}" by auto with set_to_nat_def have "set_to_nat D = sum (\ x. 2 ^ x) D" by simp with S3 sum.empty have "set_to_nat D = 0" by auto with zero_is_empty S3 show "nat_to_set (set_to_nat D) = D" by auto qed qed have P_at_Suc: "\ n. ?P n \ ?P (Suc n)" proof - fix n assume A_n: "?P n" show "?P (Suc n)" proof fix D show "finite D \ card D = Suc n \ nat_to_set (set_to_nat D) = D" proof assume A1: "finite D \ card D = Suc n" from A1 have S1: "finite D" by auto from A1 have S2: "card D = Suc n" by auto define m where "m = Max D" from S2 have D_ne: "D \ {}" by auto with S1 m_def have m_in: "m \ D" by auto define D1 where "D1 = D - {m}" from S1 D1_def have d1_finite: "finite D1" by auto from D1_def m_in S1 have "card D1 = card D - 1" by (simp add: card_Diff_singleton) with S2 have card_d1: "card D1 = n" by auto from d1_finite card_d1 have "finite D1 \ card D1 = n" by auto with A_n have S3: "nat_to_set (set_to_nat D1) = D1" by auto define u where "u = set_to_nat D" define u1 where "u1 = set_to_nat D1" from S1 m_in have "sum (\ (x::nat). (2::nat) ^ x) D = 2 ^ m + sum (\ x. 2 ^ x) (D - {m})" by (rule sum.remove) with set_to_nat_def have "set_to_nat D = 2 ^ m + set_to_nat (D - {m})" by auto with u_def u1_def D1_def have u_u1: "u = u1 + 2 ^ m" by auto from S3 u1_def have d1_u1: "nat_to_set u1 = D1" by auto have u1_lt: "u1 < 2 ^ m" proof - have L1: "D1 \ {i. i D1" show "x \ {i. i D" by auto from S1 D_ne L1_1 m_def have L1_2: "x \ m" by auto with A1 L1_1 D1_def have "x \ m" by auto with L1_2 show "x < m" by auto qed qed have L2: "finite {i. i set_to_nat {i. i set_to_nat {i. i 2 ^ m - 1" by auto have "0 < (2::nat) ^ m" by auto then have "(2::nat) ^ m - 1 < (2::nat) ^ m" by auto with L3 show ?thesis by arith qed from u_def have "nat_to_set (set_to_nat D) = nat_to_set u" by auto also from u_u1 have "\ = nat_to_set (u1 + 2 ^ m)" by auto also from u1_lt have "\ = nat_to_set u1 \ {m}" by (rule add_power) also from d1_u1 have "\ = D1 \ {m}" by auto also from D1_def m_in have "\ = D" by auto finally show "nat_to_set (set_to_nat D) = D" by auto qed qed qed from P_at_0 P_at_Suc have main: "\ n. ?P n" by (rule nat.induct) from A main show ?thesis by auto qed theorem nat_to_set_srj1: "finite (D::nat set) \ \ u. nat_to_set u = D" proof - assume A: "finite D" show " \ u. nat_to_set u = D" proof from A show "nat_to_set (set_to_nat D) = D" by (rule nat_to_set_srj) qed qed lemma sum_of_pr_is_pr: "g \ PrimRec1 \ (\ n. sum g {i. i PrimRec1" proof - assume g_is_pr: "g \ PrimRec1" define f where "f n = sum g {i. i PrimRec2" unfolding h_def by prec have f_at_Suc: "\ y. f (Suc y) = h y (f y)" proof fix y show "f (Suc y) = h y (f y)" proof - from f_def have S1: "f (Suc y) = sum g {i. i < Suc y}" by auto have S2: "{i. i < Suc y} = {i. i < y} \ {y}" by auto have S3: "finite {i. i < y}" by (rule finite_interval) have S4: "y \ {i. i < y}" by auto from S1 S2 have "f (Suc y) = sum g ({i. (i::nat) < y} \ {y})" by auto also from S3 S4 sum.insert have "\ = g y + sum g {i. i = g y + f y" by auto also from h_def have "\ = h y (f y)" by auto finally show ?thesis by auto qed qed from h_is_pr f_at_0 f_at_Suc have f_is_pr: "f \ PrimRec1" by (rule pr_rec1_scheme) with f_def [abs_def] show ?thesis by auto qed lemma sum_of_pr_is_pr2: "p \ PrimRec2 \ (\ n m. sum (\ x. p x m) {i. i PrimRec2" proof - assume p_is_pr: "p \ PrimRec2" define f where "f n m = sum (\ x. p x m) {i. i nat" where "g x = 0" for x have g_is_pr: "g \ PrimRec1" by (unfold g_def, rule const_is_pr [where ?n=0]) have f_at_0: "\ x. f 0 x = g x" proof fix x from f_def g_def show "f 0 x = g x" by auto qed define h where "h a b c = p a c + b" for a b c from p_is_pr have h_is_pr: "h \ PrimRec3" unfolding h_def by prec have f_at_Suc: "\ x y. f (Suc y) x = h y (f y x) x" proof (rule allI, rule allI) fix x y show "f (Suc y) x = h y (f y x) x" proof - from f_def have S1: "f (Suc y) x = sum (\ z. p z x) {i. i < Suc y}" by auto have S2: "{i. i < Suc y} = {i. i < y} \ {y}" by auto have S3: "finite {i. i < y}" by (rule finite_interval) have S4: "y \ {i. i < y}" by auto define g1 where "g1 z = p z x" for z from S1 S2 g1_def have "f (Suc y) x = sum g1 ({i. (i::nat) < y} \ {y})" by auto also from S3 S4 sum.insert have "\ = g1 y + sum g1 {i. i = g1 y + f y x" by auto also from h_def g1_def have "\ = h y (f y x) x" by auto finally show ?thesis by auto qed qed from g_is_pr h_is_pr f_at_0 f_at_Suc have f_is_pr: "f \ PrimRec2" by (rule pr_rec_scheme) with f_def [abs_def] show ?thesis by auto qed lemma sum_is_pr: "g \ PrimRec1 \ (\ u. sum g (nat_to_set u)) \ PrimRec1" proof - assume g_is_pr: "g \ PrimRec1" define g1 where "g1 x u = (if (c_in x u = 1) then (g x) else 0)" for x u have g1_is_pr: "g1 \ PrimRec2" proof (unfold g1_def, rule if_eq_is_pr2) show "c_in \ PrimRec2" by (rule c_in_is_pr) next show "(\x y. 1) \ PrimRec2" by (rule const_is_pr_2 [where ?n=1]) next from g_is_pr show "(\x y. g x) \ PrimRec2" by prec next show "(\x y. 0) \ PrimRec2" by (rule const_is_pr_2 [where ?n=0]) qed define f where "f u = sum (\ x. g1 x u) {i. (i::nat) < u}" for u define f1 where "f1 u v = sum (\ x. g1 x v) {i. (i::nat) < u}" for u v from g1_is_pr have "(\ (u::nat) v. sum (\ x. g1 x v) {i. (i::nat) < u}) \ PrimRec2" by (rule sum_of_pr_is_pr2) with f1_def [abs_def] have f1_is_pr: "f1 \ PrimRec2" by auto from f_def f1_def have f_f1: "f = (\ u. f1 u u)" by auto from f1_is_pr have "(\ u. f1 u u) \ PrimRec1" by prec with f_f1 have f_is_pr: "f \ PrimRec1" by auto have f_is_result: "f = (\ u. sum g (nat_to_set u))" proof fix u show "f u = sum g (nat_to_set u)" proof - define U where "U = {i. i < u}" define A where "A = {x \ U. c_in x u = 1}" define B where "B = {x \ U. c_in x u \ 1}" have U_finite: "finite U" by (unfold U_def, rule finite_interval) from A_def U_finite have A_finite: "finite A" by auto from B_def U_finite have B_finite: "finite B" by auto from U_def A_def B_def have U_A_B: "U = A \ B" by auto from U_def A_def B_def have A_B: "A \ B = {}" by auto from B_def g1_def have B_z: "sum (\ x. g1 x u) B = 0" by auto have u_in_U: "nat_to_set u \ U" by (unfold U_def, rule nat_to_set_upper_bound2) from u_in_U x_in_u_eq A_def have A_u: "A = nat_to_set u" by auto from A_u x_in_u_eq g1_def have A_res: "sum (\ x. g1 x u) A = sum g (nat_to_set u)" by auto from f_def have "f u = sum (\ x. g1 x u) {i. (i::nat) < u}" by auto also from U_def have "\ = sum (\ x. g1 x u) U" by auto also from U_A_B have "\ = sum (\ x. g1 x u) (A \ B)" by auto also from A_finite B_finite A_B have "\ = sum (\ x. g1 x u) A + sum (\ x. g1 x u) B" by (rule sum.union_disjoint) also from B_z have "\ = sum (\ x. g1 x u) A" by auto also from A_res have "\ = sum g (nat_to_set u)" by auto finally show ?thesis by auto qed qed with f_is_pr show ?thesis by auto qed definition c_card :: "nat \ nat" where "c_card = (\ u. card (nat_to_set u))" theorem c_card_is_pr: "c_card \ PrimRec1" proof - define g :: "nat \ nat" where "g x = 1" for x have g_is_pr: "g \ PrimRec1" by (unfold g_def, rule const_is_pr) have "c_card = (\ u. sum g (nat_to_set u))" proof fix u show "c_card u = sum g (nat_to_set u)" by (unfold c_card_def, unfold g_def, rule card_eq_sum) qed moreover from g_is_pr have "(\ u. sum g (nat_to_set u)) \ PrimRec1" by (rule sum_is_pr) ultimately show ?thesis by auto qed definition c_insert :: "nat \ nat \ nat" where "c_insert = (\ x u. if c_in x u = 1 then u else u + 2^x)" lemma c_insert_is_pr: "c_insert \ PrimRec2" proof (unfold c_insert_def, rule if_eq_is_pr2) show "c_in \ PrimRec2" by (rule c_in_is_pr) next show "(\x y. 1) \ PrimRec2" by (rule const_is_pr_2) next show "(\x y. y) \ PrimRec2" by (rule pr_id2_2) next from power_is_pr show "(\x y. y + 2 ^ x) \ PrimRec2" by prec qed lemma [simp]: "set_to_nat (nat_to_set u) = u" proof - define D where "D = nat_to_set u" from D_def nat_to_set_is_finite have D_finite: "finite D" by auto then have "nat_to_set (set_to_nat D) = D" by (rule nat_to_set_srj) with D_def have "nat_to_set (set_to_nat D) = nat_to_set u" by auto then have "set_to_nat D = u" by (rule nat_to_set_inj) with D_def show ?thesis by auto qed lemma insert_lemma: "x \ nat_to_set u \ set_to_nat (nat_to_set u \ {x}) = u + 2 ^ x" proof - assume A: "x \ nat_to_set u" define D where "D = nat_to_set u" from A D_def have S1: "x \ D" by auto have "finite (nat_to_set u)" by (rule nat_to_set_is_finite) with D_def have D_finite: "finite D" by auto let ?f = "\ (x::nat). (2::nat)^x" from set_to_nat_def have "set_to_nat (D \ {x}) = sum ?f (D \ {x})" by auto also from D_finite S1 have "\ = ?f x + sum ?f D" by simp also from set_to_nat_def have "\ = 2 ^ x + set_to_nat D" by auto finally have "set_to_nat (D \ {x}) = set_to_nat D + 2 ^ x" by auto with D_def show ?thesis by auto qed lemma c_insert_df: "c_insert = (\ x u. set_to_nat ((nat_to_set u) \ {x}))" proof (rule ext, rule ext) fix x u show "c_insert x u = set_to_nat (nat_to_set u \ {x})" proof (cases) assume A: "x \ nat_to_set u" then have "nat_to_set u \ {x} = nat_to_set u" by auto then have S1: "set_to_nat (nat_to_set u \ {x}) = u" by auto from A have "c_in x u = 1" by (simp add: x_in_u_eq) then have "c_insert x u = u" by (unfold c_insert_def, simp) with S1 show ?thesis by auto next assume A: "x \ nat_to_set u" then have S1: "c_in x u \ 1" by (simp add: x_in_u_eq) then have S2: "c_insert x u = u + 2 ^ x" by (unfold c_insert_def, simp) from A have "set_to_nat (nat_to_set u \ {x}) = u + 2 ^ x" by (rule insert_lemma) with S2 show ?thesis by auto qed qed definition c_remove :: "nat \ nat \ nat" where "c_remove = (\ x u. if c_in x u = 0 then u else u - 2^x)" lemma c_remove_is_pr: "c_remove \ PrimRec2" proof (unfold c_remove_def, rule if_eq_is_pr2) show "c_in \ PrimRec2" by (rule c_in_is_pr) next show "(\x y. 0) \ PrimRec2" by (rule const_is_pr_2) next show "(\x y. y) \ PrimRec2" by (rule pr_id2_2) next from power_is_pr show "(\x y. y - 2 ^ x) \ PrimRec2" by prec qed lemma remove_lemma: "x \ nat_to_set u \ set_to_nat (nat_to_set u - {x}) = u - 2 ^ x" proof - assume A: "x \ nat_to_set u" define D where "D = nat_to_set u - {x}" from A D_def have S1: "x \ D" by auto have "finite (nat_to_set u)" by (rule nat_to_set_is_finite) with D_def have D_finite: "finite D" by auto let ?f = "\ (x::nat). (2::nat)^x" from set_to_nat_def have "set_to_nat (D \ {x}) = sum ?f (D \ {x})" by auto also from D_finite S1 have "\ = ?f x + sum ?f D" by simp also from set_to_nat_def have "\ = 2 ^ x + set_to_nat D" by auto finally have S2: "set_to_nat (D \ {x}) = set_to_nat D + 2 ^ x" by auto from A D_def have "D \ {x} = nat_to_set u" by auto with S2 have S3: "u = set_to_nat D + 2 ^ x" by auto from A have S4: "2 ^ x \ u" by (rule nat_to_set_upper_bound) with S3 D_def show ?thesis by auto qed lemma c_remove_df: "c_remove = (\ x u. set_to_nat ((nat_to_set u) - {x}))" proof (rule ext, rule ext) fix x u show "c_remove x u = set_to_nat (nat_to_set u - {x})" proof (cases) assume A: "x \ nat_to_set u" then have S1: "c_in x u = 1" by (simp add: x_in_u_eq) then have S2: "c_remove x u = u - 2^x" by (simp add: c_remove_def) from A have "set_to_nat (nat_to_set u - {x}) = u - 2 ^ x" by (rule remove_lemma) with S2 show ?thesis by auto next assume A: "x \ nat_to_set u" then have S1: "c_in x u \ 1" by (simp add: x_in_u_eq) then have S2: "c_remove x u = u" by (simp add: c_remove_def c_in_def) from A have "nat_to_set u - {x} = nat_to_set u" by auto with S2 show ?thesis by auto qed qed definition c_union :: "nat \ nat \ nat" where "c_union = (\ u v. set_to_nat (nat_to_set u \ nat_to_set v))" theorem c_union_is_pr: "c_union \ PrimRec2" proof - define f where "f y x = set_to_nat ((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y})" for y x have f_is_pr: "f \ PrimRec2" proof - define g where "g = c_fst" from c_fst_is_pr g_def have g_is_pr: "g \ PrimRec1" by auto define h where "h a b c = (if c_in a (c_snd c) = 1 then c_insert a b else b)" for a b c from c_in_is_pr c_insert_is_pr have h_is_pr: "h \ PrimRec3" unfolding h_def by prec have f_at_0: "\ x. f 0 x = g x" proof fix x show "f 0 x = g x" by (unfold f_def, unfold g_def, simp) qed have f_at_Suc: "\ x y. f (Suc y) x = h y (f y x) x" proof (rule allI, rule allI) fix x y show "f (Suc y) x = h y (f y x) x" proof (cases) assume A: "c_in y (c_snd x) = 1" then have S1: "y \ (nat_to_set (c_snd x))" by (simp add: x_in_u_eq) from A h_def have S2: "h y (f y x) x = c_insert y (f y x)" by auto from S1 have S3: "{z \ nat_to_set (c_snd x). z < Suc y} = {z \ nat_to_set (c_snd x). z < y} \ {y}" by auto from nat_to_set_is_finite have S4: "finite ((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y})" by auto with nat_to_set_srj f_def have S5: "nat_to_set (f y x) = (nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y}" by auto from f_def have S6: "f (Suc y) x = set_to_nat ((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < Suc y})" by simp also from S3 have "\ = set_to_nat (((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y}) \ {y})" by auto also from S5 have "\ = set_to_nat (nat_to_set (f y x) \ {y})" by auto also have "\ = c_insert y (f y x)" by (simp add: c_insert_df) finally show ?thesis by (simp add: S2) next assume A: "\ c_in y (c_snd x) = 1" then have S1: "y \ (nat_to_set (c_snd x))" by (simp add: x_in_u_eq) from A h_def have S2: "h y (f y x) x = f y x" by auto have S3: "{z \ nat_to_set (c_snd x). z < Suc y} = {z \ nat_to_set (c_snd x). z < y}" proof - have "{z \ nat_to_set (c_snd x). z < Suc y} = {z \ nat_to_set (c_snd x). z < y} \ {z \ nat_to_set (c_snd x). z = y}" by auto with S1 show ?thesis by auto qed from nat_to_set_is_finite have S4: "finite ((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y})" by auto with nat_to_set_srj f_def have S5: "nat_to_set (f y x) = (nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y}" by auto from f_def have S6: "f (Suc y) x = set_to_nat ((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < Suc y})" by simp also from S3 have "\ = set_to_nat (((nat_to_set (c_fst x)) \ {z \ nat_to_set (c_snd x). z < y}))" by auto also from S5 have "\ = set_to_nat (nat_to_set (f y x))" by auto also have "\ = f y x" by simp finally show ?thesis by (simp add: S2) qed qed from g_is_pr h_is_pr f_at_0 f_at_Suc show ?thesis by (rule pr_rec_scheme) qed define union where "union u v = f v (c_pair u v)" for u v from f_is_pr have union_is_pr: "union \ PrimRec2" unfolding union_def by prec have "\ u v. union u v = set_to_nat (nat_to_set u \ nat_to_set v)" proof - fix u v show "union u v = set_to_nat (nat_to_set u \ nat_to_set v)" proof - from nat_to_set_upper_bound1 have "{z \ nat_to_set v. z < v} = nat_to_set v" by auto with union_def f_def show ?thesis by auto qed qed then have "union = (\ u v. set_to_nat (nat_to_set u \ nat_to_set v))" by (simp add: ext) with c_union_def have "c_union = union" by simp with union_is_pr show ?thesis by simp qed definition c_diff :: "nat \ nat \ nat" where "c_diff = (\ u v. set_to_nat (nat_to_set u - nat_to_set v))" theorem c_diff_is_pr: "c_diff \ PrimRec2" proof - define f where "f y x = set_to_nat ((nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < y})" for y x have f_is_pr: "f \ PrimRec2" proof - define g where "g = c_fst" from c_fst_is_pr g_def have g_is_pr: "g \ PrimRec1" by auto define h where "h a b c = (if c_in a (c_snd c) = 1 then c_remove a b else b)" for a b c from c_in_is_pr c_remove_is_pr have h_is_pr: "h \ PrimRec3" unfolding h_def by prec have f_at_0: "\ x. f 0 x = g x" proof fix x show "f 0 x = g x" by (unfold f_def, unfold g_def, simp) qed have f_at_Suc: "\ x y. f (Suc y) x = h y (f y x) x" proof (rule allI, rule allI) fix x y show "f (Suc y) x = h y (f y x) x" proof (cases) assume A: "c_in y (c_snd x) = 1" then have S1: "y \ (nat_to_set (c_snd x))" by (simp add: x_in_u_eq) from A h_def have S2: "h y (f y x) x = c_remove y (f y x)" by auto have "(nat_to_set (c_fst x)) - ({z \ nat_to_set (c_snd x). z < y} \ {y}) = ((nat_to_set (c_fst x)) - ({z \ nat_to_set (c_snd x). z < y}) - {y})" by auto then have lm1: "set_to_nat (nat_to_set (c_fst x) - ({z \ nat_to_set (c_snd x). z < y} \ {y})) = set_to_nat (nat_to_set (c_fst x) - {z \ nat_to_set (c_snd x). z < y} - {y})" by auto from S1 have S3: "{z \ nat_to_set (c_snd x). z < Suc y} = {z \ nat_to_set (c_snd x). z < y} \ {y}" by auto from nat_to_set_is_finite have S4: "finite ((nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < y})" by auto with nat_to_set_srj f_def have S5: "nat_to_set (f y x) = (nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < y}" by auto from f_def have S6: "f (Suc y) x = set_to_nat ((nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < Suc y})" by simp also from S3 have "\ = set_to_nat ((nat_to_set (c_fst x)) - ({z \ nat_to_set (c_snd x). z < y} \ {y}))" by auto also have "\ = set_to_nat (((nat_to_set (c_fst x)) - ({z \ nat_to_set (c_snd x). z < y}) - {y}))" by (rule lm1) also from S5 have "\ = set_to_nat (nat_to_set (f y x) - {y})" by auto also have "\ = c_remove y (f y x)" by (simp add: c_remove_df) finally show ?thesis by (simp add: S2) next assume A: "\ c_in y (c_snd x) = 1" then have S1: "y \ (nat_to_set (c_snd x))" by (simp add: x_in_u_eq) from A h_def have S2: "h y (f y x) x = f y x" by auto have S3: "{z \ nat_to_set (c_snd x). z < Suc y} = {z \ nat_to_set (c_snd x). z < y}" proof - have "{z \ nat_to_set (c_snd x). z < Suc y} = {z \ nat_to_set (c_snd x). z < y} \ {z \ nat_to_set (c_snd x). z = y}" by auto with S1 show ?thesis by auto qed from nat_to_set_is_finite have S4: "finite ((nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < y})" by auto with nat_to_set_srj f_def have S5: "nat_to_set (f y x) = (nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < y}" by auto from f_def have S6: "f (Suc y) x = set_to_nat ((nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < Suc y})" by simp also from S3 have "\ = set_to_nat (((nat_to_set (c_fst x)) - {z \ nat_to_set (c_snd x). z < y}))" by auto also from S5 have "\ = set_to_nat (nat_to_set (f y x))" by auto also have "\ = f y x" by simp finally show ?thesis by (simp add: S2) qed qed from g_is_pr h_is_pr f_at_0 f_at_Suc show ?thesis by (rule pr_rec_scheme) qed define diff where "diff u v = f v (c_pair u v)" for u v from f_is_pr have diff_is_pr: "diff \ PrimRec2" unfolding diff_def by prec have "\ u v. diff u v = set_to_nat (nat_to_set u - nat_to_set v)" proof - fix u v show "diff u v = set_to_nat (nat_to_set u - nat_to_set v)" proof - from nat_to_set_upper_bound1 have "{z \ nat_to_set v. z < v} = nat_to_set v" by auto with diff_def f_def show ?thesis by auto qed qed then have "diff = (\ u v. set_to_nat (nat_to_set u - nat_to_set v))" by (simp add: ext) with c_diff_def have "c_diff = diff" by simp with diff_is_pr show ?thesis by simp qed definition c_intersect :: "nat \ nat \ nat" where "c_intersect = (\ u v. set_to_nat (nat_to_set u \ nat_to_set v))" theorem c_intersect_is_pr: "c_intersect \ PrimRec2" proof - define f where "f u v = c_diff (c_union u v) (c_union (c_diff u v) (c_diff v u))" for u v from c_diff_is_pr c_union_is_pr have f_is_pr: "f \ PrimRec2" unfolding f_def by prec have "\ u v. f u v = c_intersect u v" proof - fix u v show "f u v = c_intersect u v" proof - let ?A = "nat_to_set u" let ?B = "nat_to_set v" have A_fin: "finite ?A" by (rule nat_to_set_is_finite) have B_fin: "finite ?B" by (rule nat_to_set_is_finite) have S1: "c_union u v = set_to_nat (?A \ ?B)" by (simp add: c_union_def) have S2: "c_diff u v = set_to_nat (?A - ?B)" by (simp add: c_diff_def) have S3: "c_diff v u = set_to_nat (?B - ?A)" by (simp add: c_diff_def) from S2 A_fin B_fin have S4: "nat_to_set (c_diff u v) = ?A - ?B" by (simp add: nat_to_set_srj) from S3 A_fin B_fin have S5: "nat_to_set (c_diff v u) = ?B - ?A" by (simp add: nat_to_set_srj) from S4 S5 have S6: "c_union (c_diff u v) (c_diff v u) = set_to_nat ((?A - ?B) \ (?B - ?A))" by (simp add: c_union_def) from S1 A_fin B_fin have S7: "nat_to_set (c_union u v) = ?A \ ?B" by (simp add: nat_to_set_srj) from S6 A_fin B_fin have S8: "nat_to_set (c_union (c_diff u v) (c_diff v u)) = (?A - ?B) \ (?B - ?A)" by (simp add: nat_to_set_srj) from S7 S8 have S9: "f u v = set_to_nat ((?A \ ?B) - ((?A - ?B) \ (?B - ?A)))" by (simp add: c_diff_def f_def) have S10: "?A \ ?B = (?A \ ?B) - ((?A - ?B) \ (?B - ?A))" by auto with S9 have S11: "f u v = set_to_nat (?A \ ?B)" by auto have "c_intersect u v = set_to_nat (?A \ ?B)" by (simp add: c_intersect_def) with S11 show ?thesis by auto qed qed then have "f = c_intersect" by (simp add: ext) with f_is_pr show ?thesis by auto qed end diff --git a/thys/Simple_Firewall/Primitives/Iface.thy b/thys/Simple_Firewall/Primitives/Iface.thy --- a/thys/Simple_Firewall/Primitives/Iface.thy +++ b/thys/Simple_Firewall/Primitives/Iface.thy @@ -1,650 +1,648 @@ section\Network Interfaces\ theory Iface imports "HOL-Library.Char_ord" (*WARNING: importing char ord*) begin text\Network interfaces, e.g. \texttt{eth0}, \texttt{wlan1}, ... iptables supports wildcard matching, e.g. \texttt{eth+} will match \texttt{eth}, \texttt{eth1}, \texttt{ethFOO}, ... The character `+' is only a wildcard if it appears at the end. \ datatype iface = Iface (iface_sel: "string") \ \no negation supported, but wildcards\ text\Just a normal lexicographical ordering on the interface strings. Used only for optimizing code. WARNING: not a semantic ordering.\ (*We cannot use List_lexord because it clashed with the Word_Lib imported ordering!*) instantiation iface :: linorder begin function (sequential) less_eq_iface :: "iface \ iface \ bool" where "(Iface []) \ (Iface _) \ True" | "(Iface _) \ (Iface []) \ False" | "(Iface (a#as)) \ (Iface (b#bs)) \ (if a = b then Iface as \ Iface bs else a \ b)" by(pat_completeness) auto termination "less_eq :: iface \ _ \ bool" apply(relation "measure (\is. size (iface_sel (fst is)) + size (iface_sel (snd is)))") apply(rule wf_measure, unfold in_measure comp_def) apply(simp) done lemma Iface_less_eq_empty: "Iface x \ Iface [] \ x = []" by(induction "Iface x" "Iface []" rule: less_eq_iface.induct) auto lemma less_eq_empty: "Iface [] \ q" by(induction "Iface []" q rule: less_eq_iface.induct) auto lemma iface_cons_less_eq_i: "Iface (b # bs) \ i \ \ q qs. i=Iface (q#qs) \ (b < q \ (Iface bs) \ (Iface qs))" apply(induction "Iface (b # bs)" i rule: less_eq_iface.induct) apply(simp_all split: if_split_asm) apply(clarify) apply(simp) done function (sequential) less_iface :: "iface \ iface \ bool" where "(Iface []) < (Iface []) \ False" | "(Iface []) < (Iface _) \ True" | "(Iface _) < (Iface []) \ False" | "(Iface (a#as)) < (Iface (b#bs)) \ (if a = b then Iface as < Iface bs else a < b)" by(pat_completeness) auto termination "less :: iface \ _ \ bool" apply(relation "measure (\is. size (iface_sel (fst is)) + size (iface_sel (snd is)))") apply(rule wf_measure, unfold in_measure comp_def) apply(simp) done instance proof fix n m :: iface show "n < m \ n \ m \ \ m \ n" proof(induction rule: less_iface.induct) case 4 thus ?case by simp fastforce qed(simp+) next fix n :: iface have "n = m \ n \ m" for m by(induction n m rule: less_eq_iface.induct) simp+ thus "n \ n" by simp next fix n m :: iface show "n \ m \ m \ n \ n = m" proof(induction n m rule: less_eq_iface.induct) case 1 thus ?case using Iface_less_eq_empty by blast next case 3 thus ?case by (simp split: if_split_asm) qed(simp)+ next fix n m q :: iface show "n \ m \ m \ q \ n \ q" proof(induction n q arbitrary: m rule: less_eq_iface.induct) case 1 thus ?case by simp next case 2 thus ?case apply simp apply(drule iface_cons_less_eq_i) apply(elim exE conjE disjE) apply(simp; fail) by fastforce next case 3 thus ?case apply simp apply(frule iface_cons_less_eq_i) by(auto split: if_split_asm) qed next fix n m :: iface show "n \ m \ m \ n" apply(induction n m rule: less_eq_iface.induct) apply(simp_all) by fastforce qed end definition ifaceAny :: iface where "ifaceAny \ Iface ''+''" (* there is no IfaceFalse, proof below *) text\If the interface name ends in a ``+'', then any interface which begins with this name will match. (man iptables) Here is how iptables handles this wildcard on my system. A packet for the loopback interface \texttt{lo} is matched by the following expressions \<^item> lo \<^item> lo+ \<^item> l+ \<^item> + It is not matched by the following expressions \<^item> lo++ \<^item> lo+++ \<^item> lo1+ \<^item> lo1 By the way: \texttt{Warning: weird characters in interface ` ' ('/' and ' ' are not allowed by the kernel).} However, happy snowman and shell colors are fine. \ context begin subsection\Helpers for the interface name (@{typ string})\ (*Do not use outside this thy! Type is really misleading.*) text\ argument 1: interface as in firewall rule - Wildcard support argument 2: interface a packet came from - No wildcard support\ fun internal_iface_name_match :: "string \ string \ bool" where "internal_iface_name_match [] [] \ True" | "internal_iface_name_match (i#is) [] \ (i = CHR ''+'' \ is = [])" | "internal_iface_name_match [] (_#_) \ False" | "internal_iface_name_match (i#is) (p_i#p_is) \ (if (i = CHR ''+'' \ is = []) then True else ( (p_i = i) \ internal_iface_name_match is p_is ))" (*<*) \ \Examples\ lemma "internal_iface_name_match ''lo'' ''lo''" by eval lemma "internal_iface_name_match ''lo+'' ''lo''" by eval lemma "internal_iface_name_match ''l+'' ''lo''" by eval lemma "internal_iface_name_match ''+'' ''lo''" by eval lemma "\ internal_iface_name_match ''lo++'' ''lo''" by eval lemma "\ internal_iface_name_match ''lo+++'' ''lo''" by eval lemma "\ internal_iface_name_match ''lo1+'' ''lo''" by eval lemma "\ internal_iface_name_match ''lo1'' ''lo''" by eval text\The wildcard interface name\ lemma "internal_iface_name_match ''+'' ''''" by eval (*>*) fun iface_name_is_wildcard :: "string \ bool" where "iface_name_is_wildcard [] \ False" | "iface_name_is_wildcard [s] \ s = CHR ''+''" | "iface_name_is_wildcard (_#ss) \ iface_name_is_wildcard ss" private lemma iface_name_is_wildcard_alt: "iface_name_is_wildcard eth \ eth \ [] \ last eth = CHR ''+''" proof(induction eth rule: iface_name_is_wildcard.induct) qed(simp_all) private lemma iface_name_is_wildcard_alt': "iface_name_is_wildcard eth \ eth \ [] \ hd (rev eth) = CHR ''+''" unfolding iface_name_is_wildcard_alt by (simp add: hd_rev) private lemma iface_name_is_wildcard_fst: "iface_name_is_wildcard (i # is) \ is \ [] \ iface_name_is_wildcard is" by(simp add: iface_name_is_wildcard_alt) private fun internal_iface_name_to_set :: "string \ string set" where "internal_iface_name_to_set i = (if \ iface_name_is_wildcard i then {i} else {(butlast i)@cs | cs. True})" private lemma "{(butlast i)@cs | cs. True} = (\s. (butlast i)@s) ` (UNIV::string set)" by fastforce private lemma internal_iface_name_to_set: "internal_iface_name_match i p_iface \ p_iface \ internal_iface_name_to_set i" proof(induction i p_iface rule: internal_iface_name_match.induct) case 4 thus ?case apply(simp) apply(safe) apply(simp_all add: iface_name_is_wildcard_fst) apply (metis (full_types) iface_name_is_wildcard.simps(3) list.exhaust) by (metis append_butlast_last_id) qed(simp_all) private lemma internal_iface_name_to_set2: "internal_iface_name_to_set ifce = {i. internal_iface_name_match ifce i}" by (simp add: internal_iface_name_to_set) private lemma internal_iface_name_match_refl: "internal_iface_name_match i i" proof - { fix i j have "i=j \ internal_iface_name_match i j" by(induction i j rule: internal_iface_name_match.induct)(simp_all) } thus ?thesis by simp qed subsection\Matching\ fun match_iface :: "iface \ string \ bool" where "match_iface (Iface i) p_iface \ internal_iface_name_match i p_iface" \ \Examples\ lemma " match_iface (Iface ''lo'') ''lo''" " match_iface (Iface ''lo+'') ''lo''" " match_iface (Iface ''l+'') ''lo''" " match_iface (Iface ''+'') ''lo''" "\ match_iface (Iface ''lo++'') ''lo''" "\ match_iface (Iface ''lo+++'') ''lo''" "\ match_iface (Iface ''lo1+'') ''lo''" "\ match_iface (Iface ''lo1'') ''lo''" " match_iface (Iface ''+'') ''eth0''" " match_iface (Iface ''+'') ''eth0''" " match_iface (Iface ''eth+'') ''eth0''" "\ match_iface (Iface ''lo+'') ''eth0''" " match_iface (Iface ''lo+'') ''loX''" "\ match_iface (Iface '''') ''loX''" (*<*)by eval+(*>*) lemma match_ifaceAny: "match_iface ifaceAny i" by(cases i, simp_all add: ifaceAny_def) lemma match_IfaceFalse: "\(\ IfaceFalse. (\i. \ match_iface IfaceFalse i))" apply(simp) apply(intro allI, rename_tac IfaceFalse) apply(case_tac IfaceFalse, rename_tac name) apply(rule_tac x="name" in exI) by(simp add: internal_iface_name_match_refl) \ \@{const match_iface} explained by the individual cases\ lemma match_iface_case_nowildcard: "\ iface_name_is_wildcard i \ match_iface (Iface i) p_i \ i = p_i" proof(induction i p_i rule: internal_iface_name_match.induct) qed(auto simp add: iface_name_is_wildcard_alt split: if_split_asm) lemma match_iface_case_wildcard_prefix: "iface_name_is_wildcard i \ match_iface (Iface i) p_i \ butlast i = take (length i - 1) p_i" apply(induction i p_i rule: internal_iface_name_match.induct) apply(simp; fail) apply(simp add: iface_name_is_wildcard_alt split: if_split_asm; fail) apply(simp; fail) apply(simp) apply(intro conjI) apply(simp add: iface_name_is_wildcard_alt split: if_split_asm; fail) apply(simp add: iface_name_is_wildcard_fst) by (metis One_nat_def length_0_conv list.sel(1) list.sel(3) take_Cons') lemma match_iface_case_wildcard_length: "iface_name_is_wildcard i \ match_iface (Iface i) p_i \ length p_i \ (length i - 1)" proof(induction i p_i rule: internal_iface_name_match.induct) qed(simp_all add: iface_name_is_wildcard_alt split: if_split_asm) corollary match_iface_case_wildcard: "iface_name_is_wildcard i \ match_iface (Iface i) p_i \ butlast i = take (length i - 1) p_i \ length p_i \ (length i - 1)" using match_iface_case_wildcard_length match_iface_case_wildcard_prefix by blast lemma match_iface_set: "match_iface (Iface i) p_iface \ p_iface \ internal_iface_name_to_set i" using internal_iface_name_to_set by simp private definition internal_iface_name_wildcard_longest :: "string \ string \ string option" where "internal_iface_name_wildcard_longest i1 i2 = ( if take (min (length i1 - 1) (length i2 - 1)) i1 = take (min (length i1 - 1) (length i2 - 1)) i2 then Some (if length i1 \ length i2 then i2 else i1) else None)" private lemma "internal_iface_name_wildcard_longest ''eth+'' ''eth3+'' = Some ''eth3+''" by eval private lemma "internal_iface_name_wildcard_longest ''eth+'' ''e+'' = Some ''eth+''" by eval private lemma "internal_iface_name_wildcard_longest ''eth+'' ''lo'' = None" by eval private lemma internal_iface_name_wildcard_longest_commute: "iface_name_is_wildcard i1 \ iface_name_is_wildcard i2 \ internal_iface_name_wildcard_longest i1 i2 = internal_iface_name_wildcard_longest i2 i1" - apply(simp add: internal_iface_name_wildcard_longest_def) - apply(safe) - apply(simp_all add: iface_name_is_wildcard_alt) - apply (metis One_nat_def append_butlast_last_id butlast_conv_take) - by (metis min.commute)+ + by (cases i1 rule: rev_cases; cases i2 rule: rev_cases) + (simp_all add: internal_iface_name_wildcard_longest_def iface_name_is_wildcard_alt) + private lemma internal_iface_name_wildcard_longest_refl: "internal_iface_name_wildcard_longest i i = Some i" by(simp add: internal_iface_name_wildcard_longest_def) private lemma internal_iface_name_wildcard_longest_correct: "iface_name_is_wildcard i1 \ iface_name_is_wildcard i2 \ match_iface (Iface i1) p_i \ match_iface (Iface i2) p_i \ (case internal_iface_name_wildcard_longest i1 i2 of None \ False | Some x \ match_iface (Iface x) p_i)" proof - assume assm1: "iface_name_is_wildcard i1" and assm2: "iface_name_is_wildcard i2" { assume assm3: "internal_iface_name_wildcard_longest i1 i2 = None" have "\ (internal_iface_name_match i1 p_i \ internal_iface_name_match i2 p_i)" proof - from match_iface_case_wildcard_prefix[OF assm1] have 1: "internal_iface_name_match i1 p_i = (take (length i1 - 1) i1 = take (length i1 - 1) p_i)" by(simp add: butlast_conv_take) from match_iface_case_wildcard_prefix[OF assm2] have 2: "internal_iface_name_match i2 p_i = (take (length i2 - 1) i2 = take (length i2 - 1) p_i)" by(simp add: butlast_conv_take) from assm3 have 3: "take (min (length i1 - 1) (length i2 - 1)) i1 \ take (min (length i1 - 1) (length i2 - 1)) i2" by(simp add: internal_iface_name_wildcard_longest_def split: if_split_asm) from 3 show ?thesis using 1 2 min.commute take_take by metis qed } note internal_iface_name_wildcard_longest_correct_None=this { fix X assume assm3: "internal_iface_name_wildcard_longest i1 i2 = Some X" have "(internal_iface_name_match i1 p_i \ internal_iface_name_match i2 p_i) \ internal_iface_name_match X p_i" proof - from assm3 have assm3': "take (min (length i1 - 1) (length i2 - 1)) i1 = take (min (length i1 - 1) (length i2 - 1)) i2" unfolding internal_iface_name_wildcard_longest_def by(simp split: if_split_asm) { fix i1 i2 assume iw1: "iface_name_is_wildcard i1" and iw2: "iface_name_is_wildcard i2" and len: "length i1 \ length i2" and take_i1i2: "take (length i1 - 1) i1 = take (length i1 - 1) i2" from len have len': "length i1 - 1 \ length i2 - 1" by fastforce { fix x::string from len' have "take (length i1 - 1) x = take (length i1 - 1) (take (length i2 - 1) x)" by(simp add: min_def) } note takei1=this { fix m::nat and n::nat and a::string and b c have "m \ n \ take n a = take n b \ take m a = take m c \ take m c = take m b" by (metis min_absorb1 take_take) } note takesmaller=this from match_iface_case_wildcard_prefix[OF iw1, simplified] have 1: "internal_iface_name_match i1 p_i \ take (length i1 - 1) i1 = take (length i1 - 1) p_i" by(simp add: butlast_conv_take) also have "\ \ take (length i1 - 1) (take (length i2 - 1) i1) = take (length i1 - 1) (take (length i2 - 1) p_i)" using takei1 by simp finally have "internal_iface_name_match i1 p_i = (take (length i1 - 1) (take (length i2 - 1) i1) = take (length i1 - 1) (take (length i2 - 1) p_i))" . from match_iface_case_wildcard_prefix[OF iw2, simplified] have 2: "internal_iface_name_match i2 p_i \ take (length i2 - 1) i2 = take (length i2 - 1) p_i" by(simp add: butlast_conv_take) have "internal_iface_name_match i2 p_i \ internal_iface_name_match i1 p_i" unfolding 1 2 apply(rule takesmaller[of "(length i1 - 1)" "(length i2 - 1)" i2 p_i]) using len' apply (simp; fail) apply (simp; fail) using take_i1i2 by simp } note longer_iface_imp_shorter=this show ?thesis proof(cases "length i1 \ length i2") case True with assm3 have "X = i2" unfolding internal_iface_name_wildcard_longest_def by(simp split: if_split_asm) from True assm3' have take_i1i2: "take (length i1 - 1) i1 = take (length i1 - 1) i2" by linarith from longer_iface_imp_shorter[OF assm1 assm2 True take_i1i2] \X = i2\ show "(internal_iface_name_match i1 p_i \ internal_iface_name_match i2 p_i) \ internal_iface_name_match X p_i" by fastforce next case False with assm3 have "X = i1" unfolding internal_iface_name_wildcard_longest_def by(simp split: if_split_asm) from False assm3' have take_i1i2: "take (length i2 - 1) i2 = take (length i2 - 1) i1" by (metis min_def min_diff) from longer_iface_imp_shorter[OF assm2 assm1 _ take_i1i2] False \X = i1\ show "(internal_iface_name_match i1 p_i \ internal_iface_name_match i2 p_i) \ internal_iface_name_match X p_i" by auto qed qed } note internal_iface_name_wildcard_longest_correct_Some=this from internal_iface_name_wildcard_longest_correct_None internal_iface_name_wildcard_longest_correct_Some show ?thesis by(simp split:option.split) qed fun iface_conjunct :: "iface \ iface \ iface option" where "iface_conjunct (Iface i1) (Iface i2) = (case (iface_name_is_wildcard i1, iface_name_is_wildcard i2) of (True, True) \ map_option Iface (internal_iface_name_wildcard_longest i1 i2) | (True, False) \ (if match_iface (Iface i1) i2 then Some (Iface i2) else None) | (False, True) \ (if match_iface (Iface i2) i1 then Some (Iface i1) else None) | (False, False) \ (if i1 = i2 then Some (Iface i1) else None))" lemma iface_conjunct_Some: "iface_conjunct i1 i2 = Some x \ match_iface x p_i \ match_iface i1 p_i \ match_iface i2 p_i" apply(cases i1, cases i2, rename_tac i1name i2name) apply(simp) apply(case_tac "iface_name_is_wildcard i1name") apply(case_tac [!] "iface_name_is_wildcard i2name") apply(simp_all) using internal_iface_name_wildcard_longest_correct apply auto[1] apply (metis match_iface.simps match_iface_case_nowildcard option.distinct(1) option.sel) apply (metis match_iface.simps match_iface_case_nowildcard option.distinct(1) option.sel) by (metis match_iface.simps option.distinct(1) option.inject) lemma iface_conjunct_None: "iface_conjunct i1 i2 = None \ \ (match_iface i1 p_i \ match_iface i2 p_i)" apply(cases i1, cases i2, rename_tac i1name i2name) apply(simp split: bool.split_asm if_split_asm) using internal_iface_name_wildcard_longest_correct apply fastforce apply (metis match_iface.simps match_iface_case_nowildcard)+ done lemma iface_conjunct: "match_iface i1 p_i \ match_iface i2 p_i \ (case iface_conjunct i1 i2 of None \ False | Some x \ match_iface x p_i)" apply(simp split: option.split) by(blast dest: iface_conjunct_Some iface_conjunct_None) lemma match_iface_refl: "match_iface (Iface x) x" by (simp add: internal_iface_name_match_refl) lemma match_iface_eqI: assumes "x = Iface y" shows "match_iface x y" unfolding assms using match_iface_refl . lemma iface_conjunct_ifaceAny: "iface_conjunct ifaceAny i = Some i" apply(simp add: ifaceAny_def) apply(case_tac i, rename_tac iname) apply(simp) apply(case_tac "iface_name_is_wildcard iname") apply(simp add: internal_iface_name_wildcard_longest_def iface_name_is_wildcard_alt Suc_leI; fail) apply(simp) using internal_iface_name_match.elims(3) by fastforce lemma iface_conjunct_commute: "iface_conjunct i1 i2 = iface_conjunct i2 i1" apply(induction i1 i2 rule: iface_conjunct.induct) apply(rename_tac i1 i2, simp) apply(case_tac "iface_name_is_wildcard i1") apply(case_tac [!] "iface_name_is_wildcard i2") apply(simp_all) by (simp add: internal_iface_name_wildcard_longest_commute) private definition internal_iface_name_subset :: "string \ string \ bool" where "internal_iface_name_subset i1 i2 = (case (iface_name_is_wildcard i1, iface_name_is_wildcard i2) of (True, True) \ length i1 \ length i2 \ take ((length i2) - 1) i1 = butlast i2 | (True, False) \ False | (False, True) \ take (length i2 - 1) i1 = butlast i2 | (False, False) \ i1 = i2 )" private lemma butlast_take_length_helper: fixes x ::"char list" assumes a1: "length i2 \ length i1" assumes a2: "take (length i2 - Suc 0) i1 = butlast i2" assumes a3: "butlast i1 = take (length i1 - Suc 0) x" shows "butlast i2 = take (length i2 - Suc 0) x" proof - (*sledgehammer spass Isar proof*) have f4: "List.gen_length 0 i2 \ List.gen_length 0 i1" using a1 by (simp add: length_code) have f5: "\cs. List.gen_length 0 (cs::char list) - Suc 0 = List.gen_length 0 (tl cs)" by (metis (no_types) One_nat_def length_code length_tl) obtain nn :: "(nat \ nat) \ nat" where "\f. \ f (nn f) \ f (Suc (nn f)) \ f (List.gen_length 0 i2) \ f (List.gen_length 0 i1)" using f4 by (meson lift_Suc_mono_le) hence "\ nn (\n. n - Suc 0) - Suc 0 \ nn (\n. n - Suc 0) \ List.gen_length 0 (tl i2) \ List.gen_length 0 (tl i1)" using f5 by (metis (lifting) diff_Suc_Suc diff_zero) hence f6: "min (List.gen_length 0 (tl i2)) (List.gen_length 0 (tl i1)) = List.gen_length 0 (tl i2)" using diff_le_self min.absorb1 by blast { assume "take (List.gen_length 0 (tl i2)) i1 \ take (List.gen_length 0 (tl i2)) x" have "List.gen_length 0 (tl i2) = 0 \ take (List.gen_length 0 (tl i2)) i1 = take (List.gen_length 0 (tl i2)) x" using f6 f5 a3 by (metis (lifting) One_nat_def butlast_conv_take length_code take_take) hence "take (List.gen_length 0 (tl i2)) i1 = take (List.gen_length 0 (tl i2)) x" by force } thus "butlast i2 = take (length i2 - Suc 0) x" using f5 a2 by (metis (full_types) length_code) qed private lemma internal_iface_name_subset: "internal_iface_name_subset i1 i2 \ {i. internal_iface_name_match i1 i} \ {i. internal_iface_name_match i2 i}" unfolding internal_iface_name_subset_def proof(cases "iface_name_is_wildcard i1", case_tac [!] "iface_name_is_wildcard i2", simp_all) assume a1: "iface_name_is_wildcard i1" assume a2: "iface_name_is_wildcard i2" show "(length i2 \ length i1 \ take (length i2 - Suc 0) i1 = butlast i2) \ ({i. internal_iface_name_match i1 i} \ {i. internal_iface_name_match i2 i})" (is "?l \ ?r") proof(rule iffI) assume ?l with a1 a2 show ?r apply(clarify, rename_tac x) apply(drule_tac p_i=x in match_iface_case_wildcard_prefix)+ apply(simp) using butlast_take_length_helper by blast next assume ?r hence r': "internal_iface_name_to_set i1 \ internal_iface_name_to_set i2 " apply - apply(subst(asm) internal_iface_name_to_set2[symmetric])+ by assumption have hlp1: "\i1 i2. {x. \cs. x = i1 @ cs} \ {x. \cs. x = i2 @ cs} \ length i2 \ length i1" apply(simp add: Set.Collect_mono_iff) by force have hlp2: "\i1 i2. {x. \cs. x = i1 @ cs} \ {x. \cs. x = i2 @ cs} \ take (length i2) i1 = i2" apply(simp add: Set.Collect_mono_iff) by force from r' a1 a2 show ?l apply(simp add: internal_iface_name_to_set) apply(safe) apply(drule hlp1) apply(simp) apply (metis One_nat_def Suc_pred diff_Suc_eq_diff_pred diff_is_0_eq iface_name_is_wildcard.simps(1) length_greater_0_conv) apply(drule hlp2) apply(simp) by (metis One_nat_def butlast_conv_take length_butlast length_take take_take) qed next show "iface_name_is_wildcard i1 \ \ iface_name_is_wildcard i2 \ \ Collect (internal_iface_name_match i1) \ Collect (internal_iface_name_match i2)" using internal_iface_name_match_refl match_iface_case_nowildcard by fastforce next show "\ iface_name_is_wildcard i1 \ iface_name_is_wildcard i2 \ (take (length i2 - Suc 0) i1 = butlast i2) \ ({i. internal_iface_name_match i1 i} \ {i. internal_iface_name_match i2 i})" using match_iface_case_nowildcard match_iface_case_wildcard_prefix by force next show " \ iface_name_is_wildcard i1 \ \ iface_name_is_wildcard i2 \ (i1 = i2) \ ({i. internal_iface_name_match i1 i} \ {i. internal_iface_name_match i2 i})" using match_iface_case_nowildcard by force qed definition iface_subset :: "iface \ iface \ bool" where "iface_subset i1 i2 \ internal_iface_name_subset (iface_sel i1) (iface_sel i2)" lemma iface_subset: "iface_subset i1 i2 \ {i. match_iface i1 i} \ {i. match_iface i2 i}" unfolding iface_subset_def apply(cases i1, cases i2) by(simp add: internal_iface_name_subset) definition iface_is_wildcard :: "iface \ bool" where "iface_is_wildcard ifce \ iface_name_is_wildcard (iface_sel ifce)" lemma iface_is_wildcard_ifaceAny: "iface_is_wildcard ifaceAny" by(simp add: iface_is_wildcard_def ifaceAny_def) subsection\Enumerating Interfaces\ private definition all_chars :: "char list" where "all_chars \ Enum.enum" private lemma all_chars: "set all_chars = (UNIV::char set)" by(simp add: all_chars_def enum_UNIV) text\we can compute this, but its horribly inefficient!\ (*valid chars in an interface are NOT limited to the printable ones!*) private lemma strings_of_length_n: "set (List.n_lists n all_chars) = {s::string. length s = n}" apply(induction n) apply(simp; fail) apply(simp add: all_chars) apply(safe) apply(simp; fail) apply(simp) apply(rename_tac n x) apply(rule_tac x="drop 1 x" in exI) apply(simp) apply(case_tac x) apply(simp_all) done text\Non-wildacrd interfaces of length @{term n}\ private definition non_wildcard_ifaces :: "nat \ string list" where "non_wildcard_ifaces n \ filter (\i. \ iface_name_is_wildcard i) (List.n_lists n all_chars)" text\Example: (any number higher than zero are probably too inefficient)\ private lemma "non_wildcard_ifaces 0 = ['''']" by eval private lemma non_wildcard_ifaces: "set (non_wildcard_ifaces n) = {s::string. length s = n \ \ iface_name_is_wildcard s}" using strings_of_length_n non_wildcard_ifaces_def by auto private lemma "(\ i \ set (non_wildcard_ifaces n). internal_iface_name_to_set i) = {s::string. length s = n \ \ iface_name_is_wildcard s}" by(simp add: non_wildcard_ifaces) text\Non-wildacrd interfaces up to length @{term n}\ private fun non_wildcard_ifaces_upto :: "nat \ string list" where "non_wildcard_ifaces_upto 0 = [[]]" | "non_wildcard_ifaces_upto (Suc n) = (non_wildcard_ifaces (Suc n)) @ non_wildcard_ifaces_upto n" private lemma non_wildcard_ifaces_upto: "set (non_wildcard_ifaces_upto n) = {s::string. length s \ n \ \ iface_name_is_wildcard s}" apply(induction n) apply fastforce using non_wildcard_ifaces by fastforce subsection\Negating Interfaces\ private lemma inv_iface_name_set: "- (internal_iface_name_to_set i) = ( if iface_name_is_wildcard i then {c |c. length c < length (butlast i)} \ {c @ cs |c cs. length c = length (butlast i) \ c \ butlast i} else {c | c. length c < length i} \ {c@cs | c cs. length c \ length i \ c \ i} )" proof - { fix i::string have inv_i_wildcard: "- {i@cs | cs. True} = {c | c. length c < length i} \ {c@cs | c cs. length c = length i \ c \ i}" apply(rule Set.equalityI) prefer 2 apply(safe)[1] apply(simp;fail) apply(simp;fail) apply(simp) apply(rule Compl_anti_mono[where B="{i @ cs |cs. True}" and A="- ({c | c. length c < length i} \ {c@cs | c cs. length c = length i \ c \ i})", simplified]) apply(safe) apply(simp) apply(case_tac "(length i) = length x") apply(erule_tac x=x in allE, simp) apply(erule_tac x="take (length i) x" in allE) apply(simp add: min_def) by (metis append_take_drop_id) } note inv_i_wildcard=this { fix i::string have inv_i_nowildcard: "- {i::string} = {c | c. length c < length i} \ {c@cs | c cs. length c \ length i \ c \ i}" proof - have x: "{c | c. length c = length i \ c \ i} \ {c | c. length c > length i} = {c@cs | c cs. length c \ length i \ c \ i}" apply(safe) apply force+ done have "- {i::string} = {c |c . c \ i}" by(safe, simp) also have "\ = {c | c. length c < length i} \ {c | c. length c = length i \ c \ i} \ {c | c. length c > length i}" by(auto) finally show ?thesis using x by auto qed } note inv_i_nowildcard=this show ?thesis proof(cases "iface_name_is_wildcard i") case True with inv_i_wildcard show ?thesis by force next case False with inv_i_nowildcard show ?thesis by force qed qed text\Negating is really not intuitive. The Interface @{term "''et''"} is in the negated set of @{term "''eth+''"}. And the Interface @{term "''et+''"} is also in this set! This is because @{term "''+''"} is a normal interface character and not a wildcard here! In contrast, the set described by @{term "''et+''"} (with @{term "''+''"} a wildcard) is not a subset of the previously negated set.\ lemma "''et'' \ - (internal_iface_name_to_set ''eth+'')" by(simp) lemma "''et+'' \ - (internal_iface_name_to_set ''eth+'')" by(simp) lemma "''+'' \ - (internal_iface_name_to_set ''eth+'')" by(simp) lemma "\ {i. match_iface (Iface ''et+'') i} \ - (internal_iface_name_to_set ''eth+'')" by force text\Because @{term "''+''"} can appear as interface wildcard and normal interface character, we cannot take negate an @{term "Iface i"} such that we get back @{typ "iface list"} which describe the negated interface.\ lemma "''+'' \ - (internal_iface_name_to_set ''eth+'')" by(simp) fun compress_pos_interfaces :: "iface list \ iface option" where "compress_pos_interfaces [] = Some ifaceAny" | "compress_pos_interfaces [i] = Some i" | "compress_pos_interfaces (i1#i2#is) = (case iface_conjunct i1 i2 of None \ None | Some i \ compress_pos_interfaces (i#is))" lemma compress_pos_interfaces_Some: "compress_pos_interfaces ifces = Some ifce \ match_iface ifce p_i \ (\ i\ set ifces. match_iface i p_i)" proof(induction ifces rule: compress_pos_interfaces.induct) case 1 thus ?case by (simp add: match_ifaceAny) next case 2 thus ?case by simp next case (3 i1 i2) thus ?case apply(simp) apply(case_tac "iface_conjunct i1 i2") apply(simp; fail) apply(simp) using iface_conjunct_Some by presburger qed lemma compress_pos_interfaces_None: "compress_pos_interfaces ifces = None \ \ (\ i\ set ifces. match_iface i p_i)" proof(induction ifces rule: compress_pos_interfaces.induct) case 1 thus ?case by (simp add: match_ifaceAny) next case 2 thus ?case by simp next case (3 i1 i2) thus ?case apply(cases "iface_conjunct i1 i2", simp_all) apply (blast dest: iface_conjunct_None) by (blast dest: iface_conjunct_Some) qed end end diff --git a/thys/Taylor_Models/Polynomial_Expression_Additional.thy b/thys/Taylor_Models/Polynomial_Expression_Additional.thy --- a/thys/Taylor_Models/Polynomial_Expression_Additional.thy +++ b/thys/Taylor_Models/Polynomial_Expression_Additional.thy @@ -1,407 +1,415 @@ theory Polynomial_Expression_Additional imports "Polynomial_Expression" "HOL-Decision_Procs.Approximation" begin lemma real_of_float_eq_zero_iff[simp]: "real_of_float x = 0 \ x = 0" by (simp add: real_of_float_eq) text \Theory @{theory Taylor_Models.Polynomial_Expression} contains a, more or less, 1:1 generalization of theory \Multivariate_Polynomial\. Any additions belong here.\ declare [[coercion_map map_poly]] declare [[coercion "interval_of::float\float interval"]] text \Apply float interval arguments to a float poly.\ value "Ipoly [Ivl (Float 4 (-6)) (Float 10 6)] (poly.Add (poly.C (Float 3 5)) (poly.Bound 0))" text \@{term map_poly} for homomorphisms\ lemma map_poly_homo_polyadd_eq_zero_iff: "map_poly f (p +\<^sub>p q) = 0\<^sub>p \ p +\<^sub>p q = 0\<^sub>p" if [symmetric, simp]: "\x y. f (x + y) = f x + f y" "\x. f x = 0 \ x = 0" by (induction p q rule: polyadd.induct) auto lemma zero_iffD: "(\x. f x = 0 \ x = 0) \ f 0 = 0" by auto lemma map_poly_homo_polyadd: "map_poly f (p1 +\<^sub>p p2) = map_poly f p1 +\<^sub>p map_poly f p2" if [simp]: "\x y. f (x + y) = f x + f y" "\x. f x = 0 \ x = 0" by (induction p1 p2 rule: polyadd.induct) (auto simp: zero_iffD[OF that(2)] Let_def map_poly_homo_polyadd_eq_zero_iff) lemma map_poly_homo_polyneg: "map_poly f (~\<^sub>p p1) = ~\<^sub>p (map_poly f p1)" if [simp]: "\x y. f (- x) = - f x" by (induction p1) (auto simp: Let_def map_poly_homo_polyadd_eq_zero_iff) lemma map_poly_homo_polysub: "map_poly f (p1 -\<^sub>p p2) = map_poly f p1 -\<^sub>p map_poly f p2" if [simp]: "\x y. f (x + y) = f x + f y" "\x. f x = 0 \ x = 0" "\x y. f (- x) = - f x" by (auto simp: polysub_def map_poly_homo_polyadd map_poly_homo_polyneg) lemma map_poly_homo_polymul: "map_poly f (p1 *\<^sub>p p2) = map_poly f p1 *\<^sub>p map_poly f p2" if [simp]: "\x y. f (x + y) = f x + f y" "\x. f x = 0 \ x = 0" "\x y. f (x * y) = f x * f y" by (induction p1 p2 rule: polymul.induct) (auto simp: zero_iffD[OF that(2)] map_poly_homo_polyadd) lemma map_poly_homo_polypow: "map_poly f (p1 ^\<^sub>p n) = map_poly f p1 ^\<^sub>p n" if [simp]: "\x y. f (x + y) = f x + f y" "\x. f x = 0 \ x = 0" "\x y. f (x * y) = f x * f y" "f 1 = 1" proof(induction n rule: nat_less_induct) case (1 n) then show ?case apply (cases n) apply (auto simp: map_poly_homo_polyadd map_poly_homo_polymul) by (smt Suc_less_eq div2_less_self even_Suc odd_Suc_div_two map_poly_homo_polymul that) qed lemmas map_poly_homo_polyarith = map_poly_homo_polyadd map_poly_homo_polyneg map_poly_homo_polysub map_poly_homo_polymul map_poly_homo_polypow text \Count the number of parameters of a polynomial.\ fun num_params :: "'a poly \ nat" where "num_params (poly.C c) = 0" | "num_params (poly.Bound n) = Suc n" | "num_params (poly.Add a b) = max (num_params a) (num_params b)" | "num_params (poly.Sub a b) = max (num_params a) (num_params b)" | "num_params (poly.Mul a b) = max (num_params a) (num_params b)" | "num_params (poly.Neg a) = num_params a" | "num_params (poly.Pw a n) = num_params a" | "num_params (poly.CN a n b) = max (max (num_params a) (num_params b)) (Suc n)" lemma num_params_map_poly[simp]: shows "num_params (map_poly f p) = num_params p" by (induction p, simp_all) lemma num_params_polyadd: shows "num_params (p1 +\<^sub>p p2) \ max (num_params p1) (num_params p2)" proof (induction p1 p2 rule: polyadd.induct) case (4 c n p c' n' p') then show ?case - by auto (auto simp: max_def Let_def split: if_splits) + apply (simp only: num_params.simps polyadd.simps ac_simps not_less Let_def le_max_iff_disj max.bounded_iff split: if_split) + apply simp + apply (smt (verit, ccfv_SIG) dual_order.trans le_cases3) + done qed auto lemma num_params_polyneg: shows "num_params (~\<^sub>p p) = num_params p" by (induction p rule: polyneg.induct) simp_all lemma num_params_polymul: shows "num_params (p1 *\<^sub>p p2) \ max (num_params p1) (num_params p2)" proof (induction p1 p2 rule: polymul.induct) case (4 c n p c' n' p') then show ?case - by auto (auto simp: max_def Let_def split: if_splits - intro!: num_params_polyadd[THEN order_trans]) + apply (cases n n' rule: linorder_cases) + apply (simp_all only: num_params.simps polyadd.simps polymul.simps ac_simps not_less Let_def le_max_iff_disj max.bounded_iff split: if_split) + apply simp_all + apply (smt (verit, best) le_cases3 order_trans) + apply (smt (verit, del_insts) le_max_iff_disj max_0L max_def num_params.simps(1) num_params.simps(8) num_params_polyadd) + apply (smt (z3) dual_order.trans nle_le) + done qed auto lemma num_params_polypow: shows "num_params (p ^\<^sub>p n) \ num_params p" apply (induction n rule: polypow.induct) unfolding polypow.simps by (auto intro!: order_trans[OF num_params_polymul] simp: Let_def simp del: polypow.simps) lemma num_params_polynate: shows "num_params (polynate p) \ num_params p" proof(induction p rule: polynate.induct) case (2 l r) thus ?case using num_params_polyadd[of "polynate l" "polynate r"] by simp next case (3 l r) thus ?case using num_params_polyadd[of "polynate l" "~\<^sub>p (polynate r)"] by (simp add: polysub_def num_params_polyneg) next case (4 l r) thus ?case using num_params_polymul[of "polynate l" "polynate r"] by simp next case (5 p) thus ?case by (simp add: num_params_polyneg) next case (6 p n) thus ?case using num_params_polypow[of n "polynate p"] by simp qed simp_all lemma polynate_map_poly_real[simp]: fixes p :: "float poly" shows "map_poly real_of_float (polynate p) = polynate (map_poly real_of_float p)" by (induction p) (simp_all add: map_poly_homo_polyarith) text \Evaluating a float poly is equivalent to evaluating the corresponding real poly with the float parameters converted to reals.\ lemma Ipoly_real_float_eqiv: fixes p::"float poly" and xs::"float list" assumes "num_params p \ length xs" shows "Ipoly xs (p::real poly) = Ipoly xs p" using assms by (induction p, simp_all) text \Evaluating an \'a poly\ with \'a interval\ arguments is monotone.\ lemma Ipoly_interval_args_mono: fixes p::"'a::linordered_idom poly" and x::"'a list" and xs::"'a interval list" assumes "x all_in\<^sub>i xs" assumes "num_params p \ length xs" shows "Ipoly x p \ set_of (Ipoly xs (map_poly interval_of p))" using assms by (induction p) (auto simp: all_in_i_def plus_in_intervalI minus_in_intervalI times_in_intervalI uminus_in_intervalI set_of_power_mono) lemma Ipoly_interval_args_inc_mono: fixes p::"'a::{real_normed_algebra, linear_continuum_topology, linordered_idom} poly" and I::"'a interval list" and J::"'a interval list" assumes "num_params p \ length I" assumes "I all_subset J" shows "set_of (Ipoly I (map_poly interval_of p)) \ set_of (Ipoly J (map_poly interval_of p))" using assms by (induction p) (simp_all add: set_of_add_inc set_of_sub_inc set_of_mul_inc set_of_neg_inc set_of_pow_inc) section \Splitting polynomials to reduce floating point precision\ text \TODO: Move this! Definitions regarding floating point numbers should not be in a theory about polynomials.\ fun float_prec :: "float \ int" where "float_prec f = (let p=exponent f in if p \ 0 then 0 else -p)" fun float_round :: "nat \ float \ float" where "float_round prec f = ( let d = float_down prec f; u = float_up prec f in if f - d < u - f then d else u)" text \Splits any polynomial \p\ into two polynomials \l\, \r\, such that \\x::real. p(x) = l(x) + r(x)\ and all floating point coefficients in \p\ are rounded to precision \prec\. Not all cases need to give good results. Polynomials normalized with polynate only contain \poly.C\ and \poly.CN\ constructors.\ fun split_by_prec :: "nat \ float poly \ float poly * float poly" where "split_by_prec prec (poly.C f) = (let r = float_round prec f in (poly.C r, poly.C (f - r)))" | "split_by_prec prec (poly.Bound n) = (poly.Bound n, poly.C 0)" | "split_by_prec prec (poly.Add l r) = (let (ll, lr) = split_by_prec prec l; (rl, rr) = split_by_prec prec r in (poly.Add ll rl, poly.Add lr rr))" | "split_by_prec prec (poly.Sub l r) = (let (ll, lr) = split_by_prec prec l; (rl, rr) = split_by_prec prec r in (poly.Sub ll rl, poly.Sub lr rr))" | "split_by_prec prec (poly.Mul l r) = (let (ll, lr) = split_by_prec prec l; (rl, rr) = split_by_prec prec r in (poly.Mul ll rl, poly.Add (poly.Add (poly.Mul lr rl) (poly.Mul ll rr)) (poly.Mul lr rr)))" | "split_by_prec prec (poly.Neg p) = (let (l, r) = split_by_prec prec p in (poly.Neg l, poly.Neg r))" | "split_by_prec prec (poly.Pw p 0) = (poly.C 1, poly.C 0)" | "split_by_prec prec (poly.Pw p (Suc n)) = (let (l, r) = split_by_prec prec p in (poly.Pw l n, poly.Sub (poly.Pw p (Suc n)) (poly.Pw l n)))" | "split_by_prec prec (poly.CN c n p) = (let (cl, cr) = split_by_prec prec c; (pl, pr) = split_by_prec prec p in (poly.CN cl n pl, poly.CN cr n pr))" text \TODO: Prove precision constraint on \l\.\ lemma split_by_prec_correct: fixes args :: "real list" assumes "(l, r) = split_by_prec prec p" shows "Ipoly args p = Ipoly args l + Ipoly args r" (is ?P1) and "num_params l \ num_params p" (is ?P2) and "num_params r \ num_params p" (is ?P3) unfolding atomize_conj using assms proof(induction p arbitrary: l r) case (Add p1 p2 l r) thus ?case apply(simp add: Add(1,2)[OF prod.collapse] split_beta) using max.coboundedI1 max.coboundedI2 prod.collapse by metis next case (Sub p1 p2 l r) thus ?case apply(simp add: Sub(1,2)[OF prod.collapse] split_beta) using max.coboundedI1 max.coboundedI2 prod.collapse by metis next case (Mul p1 p2 l r) thus ?case apply(simp add: Mul(1,2)[OF prod.collapse] split_beta algebra_simps) using max.coboundedI1 max.coboundedI2 prod.collapse by metis next case (Neg p l r) thus ?case by (simp add: Neg(1)[OF prod.collapse] split_beta) next case (Pw p n l r) thus ?case by (cases n) (simp_all add: Pw(1)[OF prod.collapse] split_beta) next case (CN c n p2) thus ?case apply(simp add: CN(1,2)[OF prod.collapse] split_beta algebra_simps) by (meson le_max_iff_disj prod.collapse) qed (simp_all add: Let_def) section \Splitting polynomials by degree\ fun maxdegree :: "('a::zero) poly \ nat" where "maxdegree (poly.C c) = 0" | "maxdegree (poly.Bound n) = 1" | "maxdegree (poly.Add l r) = max (maxdegree l) (maxdegree r)" | "maxdegree (poly.Sub l r) = max (maxdegree l) (maxdegree r)" | "maxdegree (poly.Mul l r) = maxdegree l + maxdegree r" | "maxdegree (poly.Neg p) = maxdegree p" | "maxdegree (poly.Pw p n) = n * maxdegree p" | "maxdegree (poly.CN c n p) = max (maxdegree c) (1 + maxdegree p)" fun split_by_degree :: "nat \ 'a::zero poly \ 'a poly * 'a poly" where "split_by_degree n (poly.C c) = (poly.C c, poly.C 0)" | "split_by_degree 0 p = (poly.C 0, p)" | "split_by_degree (Suc n) (poly.CN c v p) = ( let (cl, cr) = split_by_degree (Suc n) c; (pl, pr) = split_by_degree n p in (poly.CN cl v pl, poly.CN cr v pr))" \ \This function is only intended for use on polynomials in normal form. Hence most cases never get executed.\ | "split_by_degree n p = (poly.C 0, p)" lemma split_by_degree_correct: fixes x :: "real list" and p :: "float poly" assumes "(l, r) = split_by_degree ord p" shows "maxdegree l \ ord" (is ?P1) and "Ipoly x p = Ipoly x l + Ipoly x r" (is ?P2) and "num_params l \ num_params p" (is ?P3) and "num_params r \ num_params p" (is ?P4) unfolding atomize_conj using assms proof(induction p arbitrary: l r ord) case (C c l r ord) thus ?case by simp next case (Bound v l r ord) thus ?case by (cases ord) simp_all next case (Add p1 p2 l r ord) thus ?case by (cases ord) simp_all next case (Sub p1 p2 l r ord) thus ?case by (cases ord) simp_all next case (Mul p1 p2 l r ord) thus ?case by (cases ord) simp_all next case (Neg p l r ord) thus ?case by (cases ord) simp_all next case (Pw p k l r ord) thus ?case by (cases ord) simp_all next case (CN c v p l r ord) then show ?case proof(cases ord) case (Suc m) obtain cl cr where cl_cr_def: "(cl, cr) = split_by_degree (Suc m) c" by (cases "split_by_degree (Suc m) c", simp) obtain pl pr where pl_pr_def: "(pl, pr) = split_by_degree m p" by (cases "split_by_degree m p", simp) have [simp]: "Ipoly x p = Ipoly x pl + Ipoly x pr" using CN(2)[OF pl_pr_def] by (cases ord) simp_all from CN(3) have l_decomp: "l = CN cl v pl" and r_decomp: "r = CN cr v pr" by (simp_all add: Suc cl_cr_def[symmetric] pl_pr_def[symmetric]) show ?thesis using CN(1)[OF cl_cr_def] CN(2)[OF pl_pr_def] unfolding l_decomp by (cases p) (auto simp add: l_decomp r_decomp algebra_simps Suc) qed simp qed text \Operations on lists.\ lemma length_map2[simp]: "length (map2 f a b) = min (length a) (length b)" proof(induction "map2 f a b" arbitrary: a b) case (Nil a b) hence "a = [] | b = []" by(cases a, simp, cases b, simp_all) then show ?case by auto next case (Cons x c a b) have "0 < length a \ 0 < length b" using Cons(2) by (cases a, simp, cases b, simp_all) then obtain xa ar xb br where a_decomp[simp]: "a = xa # ar" and b_decomp[simp]: "b = xb # br" by (cases a, simp_all, cases b, simp_all) show ?case using Cons by simp qed lemma map2_nth[simp]: assumes "n < length a" assumes "n < length b" shows "(map2 f a b)!n = f (a!n) (b!n)" using assms proof(induction n arbitrary: a b) case (0 a b) have "0 < length a" and "0 < length b" using 0 by simp_all thus ?case using 0 by simp next case (Suc n a b) from Suc.prems have "0 < length a" "0 < length b" "n < length (tl a)" "n < length (tl b)" using Suc.prems by auto have "map2 f a b = map2 f (hd a # tl a) (hd b # tl b)" using \0 < length a\ \0 < length b\ by simp also have "\ ! Suc n = map2 f (tl a) (tl b) ! n" by simp also have "\ = f (tl a ! n) (tl b ! n)" using \n < length (tl a)\ \n < length (tl b)\ by (rule Suc.IH) also have "tl a ! n = (hd a # tl a) ! Suc n" by simp also have "(hd a # tl a) = a" using \0 < length a\ by simp also have "tl b ! n = (hd b # tl b) ! Suc n" by simp also have "(hd b # tl b) = b" using \0 < length b\ by simp finally show ?case . qed text \Translating a polynomial by a vector.\ fun poly_translate :: "'a list \ 'a poly \ 'a poly" where "poly_translate vs (poly.C c) = poly.C c" | "poly_translate vs (poly.Bound n) = poly.Add (poly.Bound n) (poly.C (vs ! n))" | "poly_translate vs (poly.Add l r) = poly.Add (poly_translate vs l) (poly_translate vs r)" | "poly_translate vs (poly.Sub l r) = poly.Sub (poly_translate vs l) (poly_translate vs r)" | "poly_translate vs (poly.Mul l r) = poly.Mul (poly_translate vs l) (poly_translate vs r)" | "poly_translate vs (poly.Neg p) = poly.Neg (poly_translate vs p)" | "poly_translate vs (poly.Pw p n) = poly.Pw (poly_translate vs p) n" | "poly_translate vs (poly.CN c n p) = poly.Add (poly_translate vs c) (poly.Mul (poly.Add (poly.Bound n) (poly.C (vs ! n))) (poly_translate vs p))" text \Translating a polynomial is equivalent to translating its argument.\ lemma poly_translate_correct: assumes "num_params p \ length x" assumes "length x = length v" shows "Ipoly x (poly_translate v p) = Ipoly (map2 (+) x v) p" using assms by (induction p, simp_all) lemma real_poly_translate: assumes "num_params p \ length v" shows "Ipoly x (map_poly real_of_float (poly_translate v p)) = Ipoly x (poly_translate v (map_poly real_of_float p))" using assms by (induction p, simp_all) lemma num_params_poly_translate[simp]: shows "num_params (poly_translate v p) = num_params p" by (induction p, simp_all) end diff --git a/thys/Timed_Automata/Approx_Beta.thy b/thys/Timed_Automata/Approx_Beta.thy --- a/thys/Timed_Automata/Approx_Beta.thy +++ b/thys/Timed_Automata/Approx_Beta.thy @@ -1,2222 +1,2225 @@ theory Approx_Beta imports DBM_Zone_Semantics Regions_Beta Closure begin chapter \Correctness of \\\-approximation from \\\-regions\ text \Instantiating real\ instantiation real :: linordered_ab_monoid_add begin definition neutral_real: "\ = (0 :: real)" instance by standard (auto simp: neutral_real) end text \Merging the locales for the two types of regions\ locale Regions = fixes X and k :: "'c \ nat" and v :: "'c \ nat" and n :: nat and not_in_X assumes finite: "finite X" assumes clock_numbering: "clock_numbering' v n" "\k\n. k > 0 \ (\c \ X. v c = k)" "\ c \ X. v c \ n" assumes not_in_X: "not_in_X \ X" assumes non_empty: "X \ {}" begin definition \_def: "\ \ {Regions.region X I r | I r. Regions.valid_region X k I r}" definition \\<^sub>\_def: "\\<^sub>\ \ {Regions_Beta.region X I J r | I J r. Regions_Beta.valid_region X k I J r}" definition V_def: "V \ {v . \ x \ X. v x \ 0}" sublocale alpha_interp: AlphaClosure X k \ V by (unfold_locales) (auto simp: finite \_def V_def) sublocale beta_interp: Beta_Regions' X k \\<^sub>\ V v n not_in_X using finite non_empty clock_numbering not_in_X by (unfold_locales) (auto simp: \\<^sub>\_def V_def) abbreviation "Approx\<^sub>\ \ beta_interp.Approx\<^sub>\" section \Preparing Bouyer's Theorem\ lemma region_dbm: assumes "R \ \" defines "v' \ \ i. THE c. c \ X \ v c = i" obtains M where"[M]\<^bsub>v,n\<^esub> = R" and "\ i \ n. \ j \ n. M i 0 = \ \ j > 0 \ i \ j\ M i j = \ \ M j i = \" and "\ i \ n. M i i = Le 0" and "\ i \ n. \ j \ n. i > 0 \ j > 0 \ M i 0 \ \ \ M j 0 \ \ \ (\ d :: int. (- k (v' j) \ d \ d \ k (v' i) \ M i j = Le d \ M j i = Le (-d)) \ (- k (v' j) \ d - 1 \ d \ k (v' i) \ M i j = Lt d \ M j i = Lt (-d + 1)))" and "\ i \ n. i > 0 \ M i 0 \ \ \ (\ d :: int. d \ k (v' i) \ d \ 0 \ (M i 0 = Le d \ M 0 i = Le (-d) \ M i 0 = Lt d \ M 0 i = Lt (-d + 1)))" and "\ i \ n. i > 0 \ (\ d :: int. - k (v' i) \ d \ d \ 0 \ (M 0 i = Le d \ M 0 i = Lt d))" and "\ i. \ j. M i j \ \ \ get_const (M i j) \ \" and "\ i \ n. \ j \ n. M i j \ \ \ i > 0 \ j > 0 \ (\ d:: int. (M i j = Le d \ M i j = Lt d) \ (- k (v' j)) \ d \ d \ k (v' i))" proof - from assms obtain I r where R: "R = region X I r" "valid_region X k I r" unfolding \_def by blast let ?X\<^sub>0 = "{x \ X. \d. I x = Regions.intv.Intv d}" define f where "f x = (if isIntv (I x) then Lt (intv_const (I x) + 1) else if isConst (I x) then Le (intv_const (I x)) else \)" for x define g where "g x = (if isIntv (I x) then Lt (- intv_const (I x)) else if isConst (I x) then Le (- intv_const (I x)) else Lt (- k x))" for x define h where "h x y = (if isIntv (I x) \ isIntv (I y) then if (y, x) \ r \ (x, y) \ r then Lt (int (intv_const (I x)) - intv_const (I y) + 1) else if (x, y) \ r \ (y, x) \ r then Lt (int (intv_const (I x)) - intv_const (I y)) else Le (int (intv_const (I x)) - intv_const (I y)) else if isConst (I x) \ isConst (I y) then Le (int (intv_const (I x)) - intv_const (I y)) else if isIntv (I x) \ isConst (I y) then Lt (int (intv_const (I x)) + 1 - intv_const (I y)) else if isConst (I x) \ isIntv (I y) then Lt (int (intv_const (I x)) - intv_const (I y)) else \)" for x y let ?M = "\ i j. if i = 0 then if j = 0 then Le 0 else g (v' j) else if j = 0 then f (v' i) else if i = j then Le 0 else h (v' i) (v' j)" have "[?M]\<^bsub>v,n\<^esub> \ R" proof fix u assume u: "u \ [?M]\<^bsub>v,n\<^esub>" show "u \ R" unfolding R proof (standard, goal_cases) case 1 show ?case proof fix c assume c: "c \ X" with clock_numbering have c2: "v c \ n" "v c > 0" "v' (v c) = c" unfolding v'_def by auto with u have "dbm_entry_val u None (Some c) (g c)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto then show "0 \ u c" by (cases "isIntv (I c)"; cases "isConst (I c)") (auto simp: g_def) qed next case 2 show ?case proof fix c assume c: "c \ X" with clock_numbering have c2: "v c \ n" "v c > 0" "v' (v c) = c" unfolding v'_def by auto with u have *: "dbm_entry_val u None (Some c) (g c)" "dbm_entry_val u (Some c) None (f c)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto show "intv_elem c u (I c)" proof (cases "I c") case (Const d) then have "\ isIntv (I c)" "isConst (I c)" by auto with * Const show ?thesis unfolding g_def f_def using Const by auto next case (Intv d) then have "isIntv (I c)" "\ isConst (I c)" by auto with * Intv show ?thesis unfolding g_def f_def by auto next case (Greater d) then have "\ isIntv (I c)" "\ isConst (I c)" by auto with * Greater R(2) c show ?thesis unfolding g_def f_def by fastforce qed qed next show "?X\<^sub>0 = ?X\<^sub>0" .. show "\x \ ?X\<^sub>0. \ y \ ?X\<^sub>0. (x, y) \ r \ frac (u x) \ frac (u y)" proof (standard, standard) fix x y assume A: "x \ ?X\<^sub>0" "y \ ?X\<^sub>0" show "(x, y) \ r \ frac (u x) \ frac (u y)" proof (cases "x = y") case True have "refl_on ?X\<^sub>0 r" using R(2) by auto with A True show ?thesis unfolding refl_on_def by auto next case False from A obtain d d' where AA: "I x = Intv d" "I y = Intv d'" "isIntv (I x)" "isIntv (I y)" "\ isConst (I x)" "\ isConst (I y)" by auto from A False clock_numbering have B: "v x \ n" "v x > 0" "v' (v x) = x" "v y \ n" "v y > 0" "v' (v y) = y" "v x \ v y" unfolding v'_def by auto with u have *: "dbm_entry_val u (Some x) (Some y) (h x y)" "dbm_entry_val u (Some y) (Some x) (h y x)" "dbm_entry_val u None (Some x) (g x)" "dbm_entry_val u (Some x) None (f x)" "dbm_entry_val u None (Some y) (g y)" "dbm_entry_val u (Some y) None (f y)" unfolding DBM_zone_repr_def DBM_val_bounded_def by force+ show "(x, y) \ r \ frac (u x) \ frac (u y)" proof assume C: "(x, y) \ r" show "frac (u x) \ frac (u y)" proof (cases "(y, x) \ r") case False with * AA C have **: "u x - u y < int d - d'" "d < u x" "u x < d + 1" "d' < u y" "u y < d' + 1" unfolding f_def g_def h_def by auto from nat_intv_frac_decomp[OF **(2,3)] nat_intv_frac_decomp[OF **(4,5)] **(1) show "frac (u x) \ frac (u y)" by simp next case True with * AA C have **: "u x - u y \ int d - d'" "d < u x" "u x < d + 1" "d' < u y" "u y < d' + 1" unfolding f_def g_def h_def by auto from nat_intv_frac_decomp[OF **(2,3)] nat_intv_frac_decomp[OF **(4,5)] **(1) show "frac (u x) \ frac (u y)" by simp qed next assume "frac (u x) \ frac (u y)" show "(x, y) \ r" proof (rule ccontr) assume C: "(x,y) \ r" moreover from R(2) have "total_on ?X\<^sub>0 r" by auto ultimately have "(y, x) \ r" using False A unfolding total_on_def by auto with *(2-) AA C have **: "u y - u x < int d' - d" "d < u x" "u x < d + 1" "d' < u y" "u y < d' + 1" unfolding f_def g_def h_def by auto from nat_intv_frac_decomp[OF **(2,3)] nat_intv_frac_decomp[OF **(4,5)] **(1) have "frac (u y) < frac (u x)" by simp with \frac _ \ _\ show False by auto qed qed qed qed qed qed moreover have "R \ [?M]\<^bsub>v,n\<^esub>" proof fix u assume u: "u \ R" show "u \ [?M]\<^bsub>v,n\<^esub>" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (safe, goal_cases) case 1 then show ?case by auto next case (2 c) with clock_numbering have "c \ X" by metis with clock_numbering have *: "c \ X" "v c > 0" "v' (v c) = c" unfolding v'_def by auto with R u have "intv_elem c u (I c)" "valid_intv (k c) (I c)" by auto then have "dbm_entry_val u None (Some c) (g c)" unfolding g_def by (cases "I c") auto with * show ?case by auto next case (3 c) with clock_numbering have "c \ X" by metis with clock_numbering have *: "c \ X" "v c > 0" "v' (v c) = c" unfolding v'_def by auto with R u have "intv_elem c u (I c)" "valid_intv (k c) (I c)" by auto then have "dbm_entry_val u (Some c) None (f c)" unfolding f_def by (cases "I c") auto with * show ?case by auto next case (4 c1 c2) with clock_numbering have "c1 \ X" "c2 \ X" by metis+ with clock_numbering have *: "c1 \ X" "v c1 > 0" "v' (v c1) = c1" "c2 \ X" "v c2 > 0" "v' (v c2) = c2" unfolding v'_def by auto with R u have "intv_elem c1 u (I c1)" "valid_intv (k c1) (I c1)" "intv_elem c2 u (I c2)" "valid_intv (k c2) (I c2)" by auto then have "dbm_entry_val u (Some c1) (Some c2) (h c1 c2)" unfolding h_def proof(cases "I c1", cases "I c2", fastforce+, cases "I c2", fastforce, goal_cases) case (1 d d') then show ?case proof (cases "(c2, c1) \ r", goal_cases) case 1 show ?case proof (cases "(c1, c2) \ r") case True with 1 *(1,4) R(1) u have "frac (u c1) = frac (u c2)" by auto with 1 have "u c1 - u c2 = real d - d'" by (fastforce dest: nat_intv_frac_decomp) with 1 show ?thesis by auto next case False with 1 show ?thesis by auto qed next case 2 show ?case proof (cases "c1 = c2") case True then show ?thesis by auto next case False with 2 R(2) *(1,4) have "(c1, c2) \ r" by (fastforce simp: total_on_def) with 2 *(1,4) R(1) u have "frac (u c1) < frac (u c2)" by auto with 2 have "u c1 - u c2 < real d - d'" by (fastforce dest: nat_intv_frac_decomp) with 2 show ?thesis by auto qed qed qed fastforce+ then show ?case proof (cases "v c1 = v c2", goal_cases) case True with * clock_numbering have "c1 = c2" by auto then show ?thesis by auto next case 2 with * show ?case by auto qed qed qed ultimately have "[?M]\<^bsub>v,n\<^esub> = R" by blast moreover have "\ i \ n. \ j \ n. ?M i 0 = \ \ j > 0 \ i \ j \ ?M i j = \ \ ?M j i = \" unfolding f_def h_def by auto moreover have "\ i \ n. ?M i i = Le 0" by auto moreover { fix i j assume A: "i \ n" "j \ n" "i > 0" "j > 0" "?M i 0 \ \" "?M j 0 \ \" with clock_numbering(2) obtain c1 c2 where B: "v c1 = i" "v c2 = j" "c1 \ X" "c2 \ X" by meson with clock_numbering(1) A have C: "v' i = c1" "v' j = c2" unfolding v'_def by force+ from R(2) B have valid: "valid_intv (k c1) (I c1)" "valid_intv (k c2) (I c2)" by auto have "\ d :: int. (- k (v' j) \ d \ d \ k (v' i) \ ?M i j = Le d \ ?M j i = Le (-d) \ (- k (v' j) \ d - 1 \ d \ k (v' i) \ ?M i j = Lt d \ ?M j i = Lt (-d + 1)))" proof (cases "i = j") case True then show ?thesis by auto next case False then show ?thesis proof (cases "I c1", goal_cases) case 1 then show ?case proof (cases "I c2") case Const let ?d = "int (intv_const (I c1)) - int (intv_const (I c2))" from Const 1 have "isConst (I c1)" "isConst (I c2)" by auto with A(1-4) C valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case Intv let ?d = "int(intv_const (I c1)) - int (intv_const (I c2))" from Intv 1 have "isConst (I c1)" "isIntv (I c2)" by auto with A(1-4) C valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case Greater then have "\ isIntv (I c2)" "\ isConst (I c2)" by auto with A 1(1) C have False unfolding f_def by simp then show ?thesis by fast qed next case 2 then show ?case proof (cases "I c2") case Const let ?d = "int (intv_const (I c1)) + 1 - int (intv_const (I c2))" from Const 2 have "isIntv (I c1)" "isConst (I c2)" by auto with A(1-4) C valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case Intv with 2 have *: "isIntv (I c1)" "isIntv (I c2)" by auto from Intv A(1-4) C show ?thesis apply simp proof (standard, goal_cases) case 1 show ?case proof (cases "(c2, c1) \ r") case True note T = this show ?thesis proof (cases "(c1, c2) \ r") case True let ?d = "int (intv_const (I c1)) - int (intv_const (I c2))" from True T * valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case False let ?d = "int (intv_const (I c1)) - int (intv_const (I c2)) + 1" from False T * valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto qed next case False let ?d = "int (intv_const (I c1)) - int (intv_const (I c2))" from False * valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto qed qed next case Greater then have "\ isIntv (I c2)" "\ isConst (I c2)" by auto with A 2(1) C have False unfolding f_def by simp then show ?thesis by fast qed next case 3 then have "\ isIntv (I c1)" "\ isConst (I c1)" by auto with A 3(1) C have False unfolding f_def by simp then show ?thesis by fast qed qed } moreover { fix i assume A: "i \ n" "i > 0" "?M i 0 \ \" with clock_numbering(2) obtain c1 where B: "v c1 = i" "c1 \ X" by meson with clock_numbering(1) A have C: "v' i = c1" unfolding v'_def by force+ from R(2) B have valid: "valid_intv (k c1) (I c1)" by auto have "\ d :: int. d \ k (v' i) \ d \ 0 \ (?M i 0 = Le d \ ?M 0 i = Le (-d) \ ?M i 0 = Lt d \ ?M 0 i = Lt (-d + 1))" proof (cases "i = 0") case True then show ?thesis by auto next case False then show ?thesis proof (cases "I c1", goal_cases) case 1 let ?d = "int (intv_const (I c1))" from 1 have "isConst (I c1)" "\ isIntv (I c1)" by auto with A C valid show ?thesis unfolding f_def g_def by (intro exI[where x = ?d]) auto next case 2 let ?d = "int (intv_const (I c1)) + 1" from 2 have "isIntv(I c1)" "\ isConst (I c1)" by auto with A C valid show ?thesis unfolding f_def g_def by (intro exI[where x = ?d]) auto next case 3 then have "\ isIntv (I c1)" "\ isConst (I c1)" by auto with A 3(1) C have False unfolding f_def by simp then show ?thesis by fast qed qed } moreover { fix i assume A: "i \ n" "i > 0" with clock_numbering(2) obtain c1 where B: "v c1 = i" "c1 \ X" by meson with clock_numbering(1) A have C: "v' i = c1" unfolding v'_def by force+ from R(2) B have valid: "valid_intv (k c1) (I c1)" by auto have "\ d :: int. - k (v' i) \ d \ d \ 0 \ (?M 0 i = Le d \ ?M 0 i = Lt d)" proof (cases "i = 0") case True then show ?thesis by auto next case False then show ?thesis proof (cases "I c1", goal_cases) case 1 let ?d = "- int (intv_const (I c1))" from 1 have "isConst (I c1)" "\ isIntv (I c1)" by auto with A C valid show ?thesis unfolding f_def g_def by (intro exI[where x = ?d]) auto next case 2 let ?d = "- int (intv_const (I c1))" from 2 have "isIntv(I c1)" "\ isConst (I c1)" by auto with A C valid show ?thesis unfolding f_def g_def by (intro exI[where x = ?d]) auto next case 3 let ?d = "- (k c1)" from 3 have "\ isIntv (I c1)" "\ isConst (I c1)" by auto with A C show ?thesis unfolding g_def by (intro exI[where x = ?d]) auto qed qed } moreover have "\ i. \ j. ?M i j \ \ \ get_const (?M i j) \ \" unfolding f_def g_def h_def by auto moreover have "\ i \ n. \ j \ n. i > 0 \ j > 0 \ ?M i j \ \ \ (\ d:: int. (?M i j = Le d \ ?M i j = Lt d) \ (- k (v' j)) \ d \ d \ k (v' i))" proof (auto, goal_cases) case A: (1 i j) with clock_numbering(2) obtain c1 c2 where B: "v c1 = i" "c1 \ X" "v c2 = j" "c2 \ X" by meson with clock_numbering(1) A have C: "v' i = c1" "v' j = c2" unfolding v'_def by force+ from R(2) B have valid: "valid_intv (k c1) (I c1)" "valid_intv (k c2) (I c2)" by auto with A B C show ?case proof (simp, goal_cases) case 1 show ?case proof (cases "I c1", goal_cases) case 1 then show ?case proof (cases "I c2") case Const let ?d = "int (intv_const (I c1)) - int (intv_const (I c2))" from Const 1 have "isConst (I c1)" "isConst (I c2)" by auto with A(1-4) C valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case Intv let ?d = "int(intv_const (I c1)) - int (intv_const (I c2))" from Intv 1 have "isConst (I c1)" "isIntv (I c2)" by auto with A(1-4) C valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case Greater then have "\ isIntv (I c2)" "\ isConst (I c2)" by auto with A 1(1) C show ?thesis unfolding h_def by simp qed next case 2 then show ?case proof (cases "I c2") case Const let ?d = "int (intv_const (I c1)) + 1 - int (intv_const (I c2))" from Const 2 have "isIntv (I c1)" "isConst (I c2)" by auto with A(1-4) C valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case Intv with 2 have *: "isIntv (I c1)" "isIntv (I c2)" by auto from Intv A(1-4) C show ?thesis proof goal_cases case 1 show ?case proof (cases "(c2, c1) \ r") case True note T = this show ?thesis proof (cases "(c1, c2) \ r") case True let ?d = "int (intv_const (I c1)) - int (intv_const (I c2))" from True T * valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto next case False let ?d = "int (intv_const (I c1)) - int (intv_const (I c2)) + 1" from False T * valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto qed next case False let ?d = "int (intv_const (I c1)) - int (intv_const (I c2))" from False * valid show ?thesis unfolding h_def by (intro exI[where x = ?d]) auto qed qed next case Greater then have "\ isIntv (I c2)" "\ isConst (I c2)" by auto with A 2(1) C show ?thesis unfolding h_def by simp qed next case 3 then have "\ isIntv (I c1)" "\ isConst (I c1)" by auto with A 3(1) C show ?thesis unfolding h_def by simp qed qed qed moreover show ?thesis apply (rule that) apply (rule calculation(1)) apply (rule calculation(2)) apply (rule calculation(3)) apply (blast intro: calculation)+ apply (rule calculation(7)) using calculation(8) apply blast done qed lemma len_inf_elem: "(a, b) \ set (arcs i j xs) \ M a b = \ \ len M i j xs = \" apply (induction rule: arcs.induct) apply (auto simp: mult) apply (rename_tac a' b' x xs) apply (case_tac "M a' x") by auto lemma dbm_add_strict_right_mono_neutral: "a < Le d \ a + Le (-d) < Le 0" unfolding less mult by (cases a) (auto elim!: dbm_lt.cases) lemma dbm_lt_not_inf_less[intro]: "A \ \ \ A \ \" by (cases A) auto lemma add_inf[simp]: "a + \ = \" "\ + a = \" unfolding mult by (cases a) auto lemma inf_lt[simp,dest!]: "\ < x \ False" by (cases x) (auto simp: less) lemma zone_diag_lt: assumes "a \ n" "b \ n" and C: "v c1 = a" "v c2 = b" and not0: "a > 0" "b > 0" shows "[(\ i j. if i = a \ j = b then Lt d else \)]\<^bsub>v,n\<^esub> = {u. u c1 - u c2 < d}" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (standard, goal_cases) case 1 then show ?case using \a \ n\ \b \ n\ C by fastforce next case 2 then show ?case proof (safe, goal_cases) case 1 from not0 show ?case unfolding dbm_le_def by auto next case 2 with not0 show ?case by auto next case 3 with not0 show ?case by auto next case (4 u' y z) show ?case proof (cases "v y = a \ v z = b") case True with 4 clock_numbering C \a \ n\ \b \ n\ have "u' y - u' z < d" by metis with True show ?thesis by auto next case False then show ?thesis by auto qed qed qed lemma zone_diag_le: assumes "a \ n" "b \ n" and C: "v c1 = a" "v c2 = b" and not0: "a > 0" "b > 0" shows "[(\ i j. if i = a \ j = b then Le d else \)]\<^bsub>v,n\<^esub> = {u. u c1 - u c2 \ d}" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (rule, goal_cases) case 1 then show ?case using \a \ n\ \b \ n\ C by fastforce next case 2 then show ?case proof (safe, goal_cases) case 1 from not0 show ?case unfolding dbm_le_def by auto next case 2 with not0 show ?case by auto next case 3 with not0 show ?case by auto next case (4 u' y z) show ?case proof (cases "v y = a \ v z = b") case True with 4 clock_numbering C \a \ n\ \b \ n\ have "u' y - u' z \ d" by metis with True show ?thesis by auto next case False then show ?thesis by auto qed qed qed lemma zone_diag_lt_2: assumes "a \ n" and C: "v c = a" and not0: "a > 0" shows "[(\ i j. if i = a \ j = 0 then Lt d else \)]\<^bsub>v,n\<^esub> = {u. u c < d}" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (rule, goal_cases) case 1 then show ?case using \a \ n\ C by fastforce next case 2 then show ?case proof (safe, goal_cases) case 1 from not0 show ?case unfolding dbm_le_def by auto next case 2 with not0 show ?case by auto next case (3 u c) show ?case proof (cases "v c = a") case False then show ?thesis by auto next case True with 3 clock_numbering C \a \ n\ have "u c < d" by metis with C show ?thesis by auto qed next case (4 u' y z) from clock_numbering(1) have "0 < v z" by auto then show ?case by auto qed qed lemma zone_diag_le_2: assumes "a \ n" and C: "v c = a" and not0: "a > 0" shows "[(\ i j. if i = a \ j = 0 then Le d else \)]\<^bsub>v,n\<^esub> = {u. u c \ d}" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (rule, goal_cases) case 1 then show ?case using \a \ n\ C by fastforce next case 2 then show ?case proof (safe, goal_cases) case 1 from not0 show ?case unfolding dbm_le_def by auto next case 2 with not0 show ?case by auto next case (3 u c) show ?case proof (cases "v c = a") case False then show ?thesis by auto next case True with 3 clock_numbering C \a \ n\ have "u c \ d" by metis with C show ?thesis by auto qed next case (4 u' y z) from clock_numbering(1) have "0 < v z" by auto then show ?case by auto qed qed lemma zone_diag_lt_3: assumes "a \ n" and C: "v c = a" and not0: "a > 0" shows "[(\ i j. if i = 0 \ j = a then Lt d else \)]\<^bsub>v,n\<^esub> = {u. - u c < d}" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (rule, goal_cases) case 1 then show ?case using \a \ n\ C by fastforce next case 2 then show ?case proof (safe, goal_cases) case 1 from not0 show ?case unfolding dbm_le_def by auto next case (2 u c) show ?case proof (cases "v c = a", goal_cases) case False then show ?thesis by auto next case True with 2 clock_numbering C \a \ n\ have "- u c < d" by metis with C show ?thesis by auto qed next case (3 u) with not0 show ?case by auto next case (4 u' y z) from clock_numbering(1) have "0 < v y" by auto then show ?case by auto qed qed lemma len_int_closed: "\ i j. (M i j :: real) \ \ \ len M i j xs \ \" by (induction xs arbitrary: i) auto lemma get_const_distr: "a \ \ \ b \ \ \ get_const (a + b) = get_const a + get_const b" by (cases a) (cases b, auto simp: mult)+ lemma len_int_dbm_closed: "\ (i, j) \ set (arcs i j xs). (get_const (M i j) :: real) \ \ \ M i j \ \ \ get_const (len M i j xs) \ \ \ len M i j xs \ \" by (induction xs arbitrary: i) (auto simp: get_const_distr, simp add: dbm_add_not_inf mult) lemma zone_diag_le_3: assumes "a \ n" and C: "v c = a" and not0: "a > 0" shows "[(\ i j. if i = 0 \ j = a then Le d else \)]\<^bsub>v,n\<^esub> = {u. - u c \ d}" unfolding DBM_zone_repr_def DBM_val_bounded_def proof (rule, goal_cases) case 1 then show ?case using \a \ n\ C by fastforce next case 2 then show ?case proof (safe, goal_cases) case 1 from not0 show ?case unfolding dbm_le_def by auto next case (2 u c) show ?case proof (cases "v c = a") case False then show ?thesis by auto next case True with 2 clock_numbering C \a \ n\ have "- u c \ d" by metis with C show ?thesis by auto qed next case (3 u) with not0 show ?case by auto next case (4 u' y z) from clock_numbering(1) have "0 < v y" by auto then show ?case by auto qed qed lemma dbm_lt': assumes "[M]\<^bsub>v,n\<^esub> \ V" "M a b \ Lt d" "a \ n" "b \ n" "v c1 = a" "v c2 = b" "a > 0" "b > 0" shows "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 < d}" proof - from assms have "[M]\<^bsub>v,n\<^esub> \ [(\ i j. if i = a \ j = b then Lt d else \)]\<^bsub>v,n\<^esub>" apply safe apply (rule DBM_le_subset) unfolding less_eq dbm_le_def by auto moreover from zone_diag_lt[OF \a \ n\ \b \ n\ assms(5-)] have "[(\ i j. if i = a \ j = b then Lt d else \)]\<^bsub>v,n\<^esub> = {u. u c1 - u c2 < d}" by blast moreover from assms have "[M]\<^bsub>v,n\<^esub> \ V" by auto ultimately show ?thesis by auto qed lemma dbm_lt'2: assumes "[M]\<^bsub>v,n\<^esub> \ V" "M a 0 \ Lt d" "a \ n" "v c1 = a" "a > 0" shows "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 < d}" proof - from assms(2) have "[M]\<^bsub>v,n\<^esub> \ [(\ i j. if i = a \ j = 0 then Lt d else \)]\<^bsub>v,n\<^esub>" apply safe apply (rule DBM_le_subset) unfolding less_eq dbm_le_def by auto moreover from zone_diag_lt_2[OF \a \ n\ assms(4,5)] have "[(\ i j. if i = a \ j = 0 then Lt d else \)]\<^bsub>v,n\<^esub> = {u. u c1 < d}" by blast ultimately show ?thesis using assms(1) by auto qed lemma dbm_lt'3: assumes "[M]\<^bsub>v,n\<^esub> \ V" "M 0 a \ Lt d" "a \ n" "v c1 = a" "a > 0" shows "[M]\<^bsub>v,n\<^esub> \ {u \ V. - u c1 < d}" proof - from assms(2) have "[M]\<^bsub>v,n\<^esub> \ [(\ i j. if i = 0 \ j = a then Lt d else \)]\<^bsub>v,n\<^esub>" apply safe apply (rule DBM_le_subset) unfolding less_eq dbm_le_def by auto moreover from zone_diag_lt_3[OF \a \ n\ assms(4,5)] have "[(\ i j. if i = 0 \ j = a then Lt d else \)]\<^bsub>v,n\<^esub> = {u. - u c1 < d}" by blast ultimately show ?thesis using assms(1) by auto qed lemma dbm_le': assumes "[M]\<^bsub>v,n\<^esub> \ V" "M a b \ Le d" "a \ n" "b \ n" "v c1 = a" "v c2 = b" "a > 0" "b > 0" shows "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 \ d}" proof - from assms have "[M]\<^bsub>v,n\<^esub> \ [(\ i j. if i = a \ j = b then Le d else \)]\<^bsub>v,n\<^esub>" apply safe apply (rule DBM_le_subset) unfolding less_eq dbm_le_def by auto moreover from zone_diag_le[OF \a \ n\ \b \ n\ assms(5-)] have "[(\ i j. if i = a \ j = b then Le d else \)]\<^bsub>v,n\<^esub> = {u. u c1 - u c2 \ d}" by blast moreover from assms have "[M]\<^bsub>v,n\<^esub> \ V" by auto ultimately show ?thesis by auto qed lemma dbm_le'2: assumes "[M]\<^bsub>v,n\<^esub> \ V" "M a 0 \ Le d" "a \ n" "v c1 = a" "a > 0" shows "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 \ d}" proof - from assms(2) have "[M]\<^bsub>v,n\<^esub> \ [(\ i j. if i = a \ j = 0 then Le d else \)]\<^bsub>v,n\<^esub>" apply safe apply (rule DBM_le_subset) unfolding less_eq dbm_le_def by auto moreover from zone_diag_le_2[OF \a \ n\ assms(4,5)] have "[(\ i j. if i = a \ j = 0 then Le d else \)]\<^bsub>v,n\<^esub> = {u. u c1 \ d}" by blast ultimately show ?thesis using assms(1) by auto qed lemma dbm_le'3: assumes "[M]\<^bsub>v,n\<^esub> \ V" "M 0 a \ Le d" "a \ n" "v c1 = a" "a > 0" shows "[M]\<^bsub>v,n\<^esub> \ {u \ V. - u c1 \ d}" proof - from assms(2) have "[M]\<^bsub>v,n\<^esub> \ [(\ i j. if i = 0 \ j = a then Le d else \)]\<^bsub>v,n\<^esub>" apply safe apply (rule DBM_le_subset) unfolding less_eq dbm_le_def by auto moreover from zone_diag_le_3[OF \a \ n\ assms(4,5)] have "[(\ i j. if i = 0 \ j = a then Le d else \)]\<^bsub>v,n\<^esub> = {u. - u c1 \ d}" by blast ultimately show ?thesis using assms(1) by auto qed lemma int_zone_dbm: assumes "\ (_,d) \ collect_clock_pairs cc. d \ \" "\ c \ collect_clks cc. v c \ n" obtains M where "{u. u \ cc} = [M]\<^bsub>v,n\<^esub>" and "dbm_int M n" using int_zone_dbm[OF _ assms] clock_numbering(1) by auto lemma non_empty_dbm_diag_set': assumes "clock_numbering' v n" "\i\n. \j\n. M i j \ \ \ get_const (M i j) \ \" "[M]\<^bsub>v,n\<^esub> \ {}" obtains M' where "[M]\<^bsub>v,n\<^esub> = [M']\<^bsub>v,n\<^esub> \ (\i\n. \j\n. M' i j \ \ \ get_const (M' i j) \ \) \ (\ i \ n. M' i i = \)" proof - let ?M = "\i j. if i = j then \ else M i j" from non_empty_dbm_diag_set[OF assms(1,3)] have "[M]\<^bsub>v,n\<^esub> = [?M]\<^bsub>v,n\<^esub>" by auto moreover from assms(2) have "\i\n. \j\n. ?M i j \ \ \ get_const (?M i j) \ \" unfolding neutral by auto moreover have "\ i \ n. ?M i i = \" by auto ultimately show ?thesis by (auto intro: that) qed lemma dbm_entry_int: "x \ \ \ get_const x \ \ \ \ d :: int. x = Le d \ x = Lt d" apply (cases x) using Ints_cases by auto abbreviation "vabstr \ beta_interp.vabstr" section \Bouyer's Main Theorem\ theorem region_zone_intersect_empty_approx_correct: assumes "R \ \" "Z \ V" "R \ Z = {}" "vabstr Z M" shows "R \ Approx\<^sub>\ Z = {}" proof - define v' where "v' i = (THE c. c \ X \ v c = i)" for i from region_dbm[OF assms(1)] obtain M\<^sub>R where M\<^sub>R: "[M\<^sub>R]\<^bsub>v,n\<^esub> = R" "\i\n. \j\n. M\<^sub>R i 0 = \ \ 0 < j \ i \ j \ M\<^sub>R i j = \ \ M\<^sub>R j i = \" "\i\n. M\<^sub>R i i = Le 0" "\i\n. \j\n. 0 < i \ 0 < j \ M\<^sub>R i 0 \ \ \ M\<^sub>R j 0 \ \ \ (\d. - int (k (THE c. c \ X \ v c = j)) \ d \ d \ int (k (THE c. c \ X \ v c = i)) \ M\<^sub>R i j = Le d \ M\<^sub>R j i = Le (real_of_int (- d)) \ - int (k (THE c. c \ X \ v c = j)) \ d - 1 \ d \ int (k (THE c. c \ X \ v c = i)) \ M\<^sub>R i j = Lt d \ M\<^sub>R j i = Lt (real_of_int (- d + 1)))" "\i\n. 0 < i \ M\<^sub>R i 0 \ \ \ (\d\int (k (THE c. c \ X \ v c = i)). d \ 0 \ (M\<^sub>R i 0 = Le d \ M\<^sub>R 0 i = Le (real_of_int (- d)) \ M\<^sub>R i 0 = Lt d \ M\<^sub>R 0 i = Lt (real_of_int (- d + 1))))" "\i\n. 0 < i \ (\d\- int (k (THE c. c \ X \ v c = i)). d \ 0 \ (M\<^sub>R 0 i = Le d \ M\<^sub>R 0 i = Lt d))" "\i j. M\<^sub>R i j \ \ \ get_const (M\<^sub>R i j) \ \" "\i\n. \j\n. M\<^sub>R i j \ \ \ 0 < i \ 0 < j \ (\d. (M\<^sub>R i j = Le d \ M\<^sub>R i j = Lt d) \ - int (k (THE c. c \ X \ v c = j)) \ d \ d \ int (k (THE c. c \ X \ v c = i)))" . show ?thesis proof (cases "R = {}") case True then show ?thesis by auto next case False from clock_numbering(2) have cn_weak: "\k\n. 0 < k \ (\ c. v c = k)" by auto show ?thesis proof (cases "Z = {}") case True then show ?thesis using beta_interp.apx_empty by blast next case False from assms(4) have "Z = [M]\<^bsub>v,n\<^esub>" "\ i\n. \ j\n. M i j \ \ \ get_const (M i j) \ \" by auto from this(1) non_empty_dbm_diag_set'[OF clock_numbering(1) this(2)] \Z \ {}\ obtain M where M: "Z = [M]\<^bsub>v,n\<^esub> \ (\i\n. \j\n. M i j \ \ \ get_const (M i j) \ \) \ (\i\n. M i i = \)" by auto with not_empty_cyc_free[OF cn_weak] False have "cyc_free M n" by auto then have "cycle_free M n" using cycle_free_diag_equiv by auto from M have "Z = [FW M n]\<^bsub>v,n\<^esub>" unfolding neutral by (auto intro!: FW_zone_equiv[OF cn_weak]) moreover from fw_canonical[OF \cycle_free M _\] M have "canonical (FW M n) n" unfolding neutral by auto moreover from FW_int_preservation M have "\i\n. \j\n. FW M n i j \ \ \ get_const (FW M n i j) \ \" by auto ultimately obtain M where M: "[M]\<^bsub>v,n\<^esub> = Z" "canonical M n" "\i\n. \j\n. M i j \ \ \ get_const (M i j) \ \" by blast let ?M = "\ i j. min (M i j) (M\<^sub>R i j)" from M(1) M\<^sub>R(1) assms have "[M]\<^bsub>v,n\<^esub> \ [M\<^sub>R]\<^bsub>v,n\<^esub> = {}" by auto moreover from DBM_le_subset[folded less_eq, of n ?M M] have "[?M]\<^bsub>v,n\<^esub> \ [M]\<^bsub>v,n\<^esub>" by auto moreover from DBM_le_subset[folded less_eq, of n ?M M\<^sub>R] have "[?M]\<^bsub>v,n\<^esub> \ [M\<^sub>R]\<^bsub>v,n\<^esub>" by auto ultimately have "[?M]\<^bsub>v,n\<^esub> = {}" by blast then have "\ cyc_free ?M n" using cyc_free_not_empty[of n ?M v] clock_numbering(1) by auto then obtain i xs where xs: "i \ n" "set xs \ {0..n}" "len ?M i i xs < \" by auto from this(1,2) canonical_shorten_rotate_neg_cycle[OF M(2) this(2,1,3)] obtain i ys where ys: "len ?M i i ys < \" "set ys \ {0..n}" "successive (\(a, b). ?M a b = M a b) (arcs i i ys)" "i \ n" and distinct: "distinct ys" "i \ set ys" and cycle_closes: "ys \ [] \ ?M i (hd ys) \ M i (hd ys) \ ?M (last ys) i \ M (last ys) i" by fastforce have one_M_aux: "len ?M i j ys = len M\<^sub>R i j ys" if "\ (a,b) \ set (arcs i j ys). M a b \ M\<^sub>R a b" for j using that by (induction ys arbitrary: i) (auto simp: min_def) have one_M: "\ (a,b) \ set (arcs i i ys). M a b < M\<^sub>R a b" proof (rule ccontr, goal_cases) case 1 then have "\(a, b)\set (arcs i i ys). M\<^sub>R a b \ M a b" by auto from one_M_aux[OF this] have "len ?M i i ys = len M\<^sub>R i i ys" . with Nil ys(1) xs(3) have "len M\<^sub>R i i ys < \" by simp from DBM_val_bounded_neg_cycle[OF _ \i \ n\ \set ys \ _\ this cn_weak] have "[M\<^sub>R]\<^bsub>v,n\<^esub> = {}" unfolding DBM_zone_repr_def by auto with \R \ {}\ M\<^sub>R(1) show False by auto qed have one_M_R_aux: "len ?M i j ys = len M i j ys" if "\ (a,b) \ set (arcs i j ys). M a b \ M\<^sub>R a b" for j using that by (induction ys arbitrary: i) (auto simp: min_def) have one_M_R: "\ (a,b) \ set (arcs i i ys). M a b > M\<^sub>R a b" proof (rule ccontr, goal_cases) case 1 then have "\(a, b)\set (arcs i i ys). M\<^sub>R a b \ M a b" by auto from one_M_R_aux[OF this] have "len ?M i i ys = len M i i ys" . with Nil ys(1) xs(3) have "len M i i ys < \" by simp from DBM_val_bounded_neg_cycle[OF _ \i \ n\ \set ys \ _\ this cn_weak] have "[M]\<^bsub>v,n\<^esub> = {}" unfolding DBM_zone_repr_def by auto with \Z \ {}\ M(1) show False by auto qed have 0: "(0,0) \ set (arcs i i ys)" proof (cases "ys = []") case False with distinct show ?thesis using arcs_distinct1 by blast next case True with ys(1) have "?M i i < \" by auto then have "M i i < \ \ M\<^sub>R i i < \" by (simp add: min_less_iff_disj) from one_M one_M_R True show ?thesis by auto qed { fix a b assume A: "(a,b) \ set (arcs i i ys)" assume not0: "a > 0" from aux1[OF ys(4,4,2) A] have C2: "a \ n" by auto then obtain c1 where C: "v c1 = a" "c1 \ X" using clock_numbering(2) not0 unfolding v'_def by meson then have "v' a = c1" using clock_numbering C2 not0 unfolding v'_def by fastforce with C C2 have "\ c \ X. v c = a \ v' a = c" "a \ n" by auto } note clock_dest_1 = this { fix a b assume A: "(a,b) \ set (arcs i i ys)" assume not0: "b > 0" from aux1[OF ys(4,4,2) A] have C2: "b \ n" by auto then obtain c2 where C: "v c2 = b" "c2 \ X" using clock_numbering(2) not0 unfolding v'_def by meson then have "v' b = c2" using clock_numbering C2 not0 unfolding v'_def by fastforce with C C2 have "\ c \ X. v c = b \ v' b = c" "b \ n" by auto } note clock_dest_2 = this have clock_dest: "\ a b. (a,b) \ set (arcs i i ys) \ a > 0 \ b > 0 \ \ c1 \ X. \ c2 \ X. v c1 = a \ v c2 = b \ v' a = c1 \ v' b = c2 &&& a \ n &&& b \ n" using clock_dest_1 clock_dest_2 by (auto) presburger { fix a assume A: "(a,0) \ set (arcs i i ys)" assume not0: "a > 0" assume bounded: "M\<^sub>R a 0 \ \" assume lt: "M a 0 < M\<^sub>R a 0" from clock_dest_1[OF A not0] obtain c1 where C: "v c1 = a" "c1 \ X" "v' a = c1" and C2: "a \ n" by blast from C2 not0 bounded M\<^sub>R(5) obtain d :: int where *: "d \ int (k (v' a))" "M\<^sub>R a 0 = Le d \ M\<^sub>R 0 a = Le (- d) \ M\<^sub>R a 0 = Lt d \ M\<^sub>R 0 a = Lt (- d + 1)" unfolding v'_def by auto with C have **: "d \ int (k c1)" by auto from *(2) have ?thesis proof (standard, goal_cases) case 1 with lt have "M a 0 < Le d" by auto then have "M a 0 \ Lt d" unfolding less less_eq dbm_le_def by (fastforce elim!: dbm_lt.cases) from dbm_lt'2[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 < d}" by auto from beta_interp.\_boundedness_lt'[OF ** C(2) this] have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 < d}" . moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) None (M\<^sub>R a 0)" "dbm_entry_val u None (Some c1) (M\<^sub>R 0 a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto then have "u c1 = d" using 1 by auto then have "u \ {u \ V. u c1 < d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto next case 2 from 2 lt have "M a 0 \ \" by auto with dbm_entry_int[OF this] M(3) \a \ n\ obtain d' :: int where d': "M a 0 = Le d' \ M a 0 = Lt d'" by auto then have "M a 0 \ Le (d - 1)" using lt 2 apply (auto simp: less_eq dbm_le_def less) apply (cases rule: dbm_lt.cases) apply auto apply rule apply (cases rule: dbm_lt.cases) by auto with lt have "M a 0 \ Le (d - 1)" by auto from dbm_le'2[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 \ d - 1}" by auto from beta_interp.\_boundedness_le'[OF _ C(2) this] ** have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 \ d - 1}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u None (Some c1) (M\<^sub>R 0 a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto then have "u c1 > d - 1" using 2 by auto then have "u \ {u \ V. u c1 \ d - 1}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed } note bounded_zero_1 = this { fix a assume A: "(0,a) \ set (arcs i i ys)" assume not0: "a > 0" assume bounded: "M\<^sub>R a 0 \ \" assume lt: "M 0 a < M\<^sub>R 0 a" from clock_dest_2[OF A not0] obtain c1 where C: "v c1 = a" "c1 \ X" "v' a = c1" and C2: "a \ n" by blast from C2 not0 bounded M\<^sub>R(5) obtain d :: int where *: "d \ int (k (v' a))" "M\<^sub>R a 0 = Le d \ M\<^sub>R 0 a = Le (- d) \ M\<^sub>R a 0 = Lt d \ M\<^sub>R 0 a = Lt (- d + 1)" unfolding v'_def by auto with C have **: "- int (k c1) \ - d" by auto from *(2) have ?thesis proof (standard, goal_cases) case 1 with lt have "M 0 a < Le (-d)" by auto then have "M 0 a \ Lt (-d)" unfolding less less_eq dbm_le_def by (fastforce elim!: dbm_lt.cases) from dbm_lt'3[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. d < u c1}" by auto from beta_interp.\_boundedness_gt'[OF _ C(2) this] ** have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. - u c1 < -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) None (M\<^sub>R a 0)" "dbm_entry_val u None (Some c1) (M\<^sub>R 0 a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with 1 have "u \ {u \ V. - u c1 < -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto next case 2 from 2 lt have "M 0 a \ \" by auto with dbm_entry_int[OF this] M(3) \a \ n\ obtain d' :: int where d': "M 0 a = Le d' \ M 0 a = Lt d'" by auto then have "M 0 a \ Le (-d)" using lt 2 apply (auto simp: less_eq dbm_le_def less) apply (cases rule: dbm_lt.cases) apply auto apply rule apply (metis get_const.simps(2) 2 of_int_less_iff of_int_minus zless_add1_eq) apply (cases rule: dbm_lt.cases) apply auto apply (rule dbm_lt.intros(5)) by (simp add: int_lt_Suc_le) from dbm_le'3[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. d \ u c1}" by auto from beta_interp.\_boundedness_ge'[OF _ C(2) this] ** have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. - u c1 \ -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) None (M\<^sub>R a 0)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with 2 have "u \ {u \ V. - u c1 \ -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed } note bounded_zero_2 = this { fix a b c c1 c2 assume A: "(a,b) \ set (arcs i i ys)" assume not0: "a > 0" "b > 0" assume lt: "M a b = Lt c" assume neg: "M a b + M\<^sub>R b a < \" assume C: "v c1 = a" "v c2 = b" "c1 \ X" "c2 \ X" and C2: "a \ n" "b \ n" assume valid: "-k c2 \ -get_const (M\<^sub>R b a)" "-get_const (M\<^sub>R b a) \ k c1" from neg have "M\<^sub>R b a \ \" by auto then obtain d where *: "M\<^sub>R b a = Le d \ M\<^sub>R b a = Lt d" by (cases "M\<^sub>R b a", auto)+ with M\<^sub>R(7) \_ _ _ \ \\ have "d \ \" by fastforce with * obtain d :: int where *: "M\<^sub>R b a = Le d \ M\<^sub>R b a = Lt d" using Ints_cases by auto with valid have valid: "- k c2 \ -d" "-d \ k c1" by auto from * neg lt have "M a b \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_lt'[OF assms(2)[folded M(1)] this C2 C(1,2) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 < - d}" . from beta_interp.\_boundedness_diag_lt'[OF valid C(3,4) this] have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 - u c2 < -d}" . moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c2) (Some c1) (M\<^sub>R b a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with * have "u \ {u \ V. u c1 - u c2 < -d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note neg_sum_lt = this { fix a b assume A: "(a,b) \ set (arcs i i ys)" assume not0: "a > 0" "b > 0" assume neg: "M a b + M\<^sub>R b a < \" from clock_dest[OF A not0] obtain c1 c2 where C: "v c1 = a" "v c2 = b" "c1 \ X" "c2 \ X" and C2: "a \ n" "b \ n" by blast then have C3: "v' a = c1" "v' b = c2" unfolding v'_def using clock_numbering(1) by auto from neg have inf: "M a b \ \" "M\<^sub>R b a \ \" by auto from M\<^sub>R(8) inf not0 C(3,4) C2 C3 obtain d :: int where d: "M\<^sub>R b a = Le d \ M\<^sub>R b a = Lt d" "- int (k c1) \ d" "d \ int (k c2)" unfolding v'_def by auto from inf obtain c where c: "M a b = Le c \ M a b = Lt c" by (cases "M a b") auto { assume **: "M a b \ Lt (-d)" from dbm_lt'[OF assms(2)[folded M(1)] this C2 C(1,2) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 < (- d)}" . from beta_interp.\_boundedness_diag_lt'[OF _ _ C(3,4) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 - u c2 < -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c2) (Some c1) (M\<^sub>R b a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with d have "u \ {u \ V. u c1 - u c2 < -d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note aux = this from c have ?thesis proof (standard, goal_cases) case 2 with neg d have "M a b \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 1 note A = this from d(1) show ?thesis proof (standard, goal_cases) case 1 with A neg d have "M a b \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 2 with A neg d have "M a b \ Le (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_le'[OF assms(2)[folded M(1)] this C2 C(1,2) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 \ - d}" . from beta_interp.\_boundedness_diag_le'[OF _ _ C(3,4) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 - u c2 \ -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c2) (Some c1) (M\<^sub>R b a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with A 2 have "u \ {u \ V. u c1 - u c2 \ -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed qed } note neg_sum_1 = this { fix a b assume A: "(a,0) \ set (arcs i i ys)" assume not0: "a > 0" assume neg: "M a 0 + M\<^sub>R 0 a < \" from clock_dest_1[OF A not0] obtain c1 where C: "v c1 = a" "c1 \ X" and C2: "a \ n" by blast with clock_numbering(1) have C3: "v' a = c1" unfolding v'_def by auto from neg have inf: "M a 0 \ \" "M\<^sub>R 0 a \ \" by auto from M\<^sub>R(6) not0 C2 C3 obtain d :: int where d: "M\<^sub>R 0 a = Le d \ M\<^sub>R 0 a = Lt d" "- int (k c1) \ d" "d \ 0" unfolding v'_def by auto from inf obtain c where c: "M a 0 = Le c \ M a 0 = Lt c" by (cases "M a 0") auto { assume "M a 0 \ Lt (-d)" from dbm_lt'2[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 < - d}" . from beta_interp.\_boundedness_lt'[OF _ C(2) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 < -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u None (Some c1) (M\<^sub>R 0 a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with d have "u \ {u \ V. u c1 < -d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note aux = this from c have ?thesis proof (standard, goal_cases) case 2 with neg d have "M a 0 \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 1 note A = this from d(1) show ?thesis proof (standard, goal_cases) case 1 with A neg d have "M a 0 \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 2 with A neg d have "M a 0 \ Le (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_le'2[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 \ - d}" . from beta_interp.\_boundedness_le'[OF _ C(2) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 \ -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u None (Some c1) (M\<^sub>R 0 a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with A 2 have "u \ {u \ V. u c1 \ -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed qed } note neg_sum_1' = this { fix a b assume A: "(0,b) \ set (arcs i i ys)" assume not0: "b > 0" assume neg: "M 0 b + M\<^sub>R b 0 < \" from clock_dest_2[OF A not0] obtain c2 where C: "v c2 = b" "c2 \ X" and C2: "b \ n" by blast with clock_numbering(1) have C3: "v' b = c2" unfolding v'_def by auto from neg have "M 0 b \ \" "M\<^sub>R b 0 \ \" by auto with M\<^sub>R(5) not0 C2 C3 obtain d :: int where d: "M\<^sub>R b 0 = Le d \ M\<^sub>R b 0 = Lt d" "d \ k c2" unfolding v'_def by fastforce from \M 0 b \ \\ obtain c where c: "M 0 b = Le c \ M 0 b = Lt c" by (cases "M 0 b") auto { assume "M 0 b \ Lt (-d)" from dbm_lt'3[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c2 > d}" by simp from beta_interp.\_boundedness_gt'[OF _ C(2) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. - u c2 < -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c2) None (M\<^sub>R b 0)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with d have "u \ {u \ V. - u c2 < -d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note aux = this from c have ?thesis proof (standard, goal_cases) case 2 with neg d have "M 0 b \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case A: 1 from d(1) show ?thesis proof (standard, goal_cases) case 1 with A neg have "M 0 b \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 2 with A neg c have "M 0 b \ Le (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_le'3[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c2 \ d}" by simp from beta_interp.\_boundedness_ge'[OF _ C(2) this] d(2) have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. - u c2 \ -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c2) None (M\<^sub>R b 0)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with A 2 have "u \ {u \ V. - u c2 \ -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed qed } note neg_sum_1'' = this { fix a b assume A: "(a,b) \ set (arcs i i ys)" assume not0: "b > 0" "a > 0" assume neg: "M\<^sub>R a b + M b a < \" from clock_dest[OF A not0(2,1)] obtain c1 c2 where C: "v c1 = a" "v c2 = b" "c1 \ X" "c2 \ X" and C2: "a \ n" "b \ n" by blast then have C3: "v' a = c1" "v' b = c2" unfolding v'_def using clock_numbering(1) by auto from neg have inf: "M b a \ \" "M\<^sub>R a b \ \" by auto with M\<^sub>R(8) not0 C(3,4) C2 C3 obtain d :: int where d: "M\<^sub>R a b = Le d \ M\<^sub>R a b = Lt d" "d \ -int (k c2)" "d \ int (k c1)" unfolding v'_def by blast from inf obtain c where c: "M b a = Le c \ M b a = Lt c" by (cases "M b a") auto { assume "M b a \ Lt (-d)" from dbm_lt'[OF assms(2)[folded M(1)] this C2(2,1) C(2,1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c2 - u c1 < - d}" . from beta_interp.\_boundedness_diag_lt'[OF _ _ C(4,3) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c2 - u c1 < -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) (Some c2) (M\<^sub>R a b)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with d have "u \ {u \ V. u c2 - u c1 < -d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note aux = this from c have ?thesis proof (standard, goal_cases) case 2 with neg d have "M b a \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case A: 1 from d(1) show ?thesis proof (standard, goal_cases) case 1 with A neg d have "M b a \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 2 with A neg d have "M b a \ Le (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_le'[OF assms(2)[folded M(1)] this C2(2,1) C(2,1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c2 - u c1 \ - d}" . from beta_interp.\_boundedness_diag_le'[OF _ _ C(4,3) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c2 - u c1 \ -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) (Some c2) (M\<^sub>R a b)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with A 2 have "u \ {u \ V. u c2 - u c1 \ -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed qed } note neg_sum_2 = this { fix a b assume A: "(a,0) \ set (arcs i i ys)" assume not0: "a > 0" assume neg: "M\<^sub>R a 0 + M 0 a < \" from clock_dest_1[OF A not0] obtain c1 where C: "v c1 = a" "c1 \ X" and C2: "a \ n" by blast with clock_numbering(1) have C3: "v' a = c1" unfolding v'_def by auto from neg have inf: "M 0 a \ \" "M\<^sub>R a 0 \ \" by auto with M\<^sub>R(5) not0 C2 C3 obtain d :: int where d: "M\<^sub>R a 0 = Le d \ M\<^sub>R a 0 = Lt d" "d \ int (k c1)" "d \ 0" unfolding v'_def by auto from inf obtain c where c: "M 0 a = Le c \ M 0 a = Lt c" by (cases "M 0 a") auto { assume "M 0 a \ Lt (-d)" from dbm_lt'3[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 > d}" by simp from beta_interp.\_boundedness_gt'[OF _ C(2) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 > d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) None (M\<^sub>R a 0)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with d have "u \ {u \ V. u c1 > d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note aux = this from c have ?thesis proof (standard, goal_cases) case 2 with neg d have "M 0 a \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case A: 1 from d(1) show ?thesis proof (standard, goal_cases) case 1 with A neg d have "M 0 a \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 2 with A neg d have "M 0 a \ Le (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_le'3[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 \ d}" by simp from beta_interp.\_boundedness_ge'[OF _ C(2) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 \ d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) None (M\<^sub>R a 0)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with A 2 have "u \ {u \ V. u c1 \ d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed qed } note neg_sum_2' = this { fix a b assume A: "(0,b) \ set (arcs i i ys)" assume not0: "b > 0" assume neg: "M\<^sub>R 0 b + M b 0 < \" from clock_dest_2[OF A not0] obtain c2 where C: "v c2 = b" "c2 \ X" and C2: "b \ n" by blast with clock_numbering(1) have C3: "v' b = c2" unfolding v'_def by auto from neg have "M b 0 \ \" "M\<^sub>R 0 b \ \" by auto with M\<^sub>R(6) not0 C2 C3 obtain d :: int where d: "M\<^sub>R 0 b = Le d \ M\<^sub>R 0 b = Lt d" "-d \ k c2" unfolding v'_def by fastforce from \M b 0 \ \\ obtain c where c: "M b 0 = Le c \ M b 0 = Lt c" by (cases "M b 0") auto { assume "M b 0 \ Lt (-d)" from dbm_lt'2[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c2 < - d}" by simp from beta_interp.\_boundedness_lt'[OF _ C(2) this] d have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c2 < -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u None (Some c2) (M\<^sub>R 0 b)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with d have "u \ {u \ V. u c2 < -d}" by auto } ultimately have ?thesis using M\<^sub>R(1) M(1) by auto } note aux = this from c have ?thesis proof (standard, goal_cases) case 2 with neg d have "M b 0 \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 1 note A = this from d(1) show ?thesis proof (standard, goal_cases) case 1 with A neg have "M b 0 \ Lt (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) with aux show ?thesis . next case 2 with A neg c have "M b 0 \ Le (-d)" unfolding less_eq dbm_le_def mult neutral less by (auto elim!: dbm_lt.cases) from dbm_le'2[OF assms(2)[folded M(1)] this C2 C(1) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c2 \ - d}" by simp from beta_interp.\_boundedness_le'[OF _ C(2) this] d(2) have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c2 \ -d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u None (Some c2) (M\<^sub>R 0 b)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with A 2 have "u \ {u \ V. u c2 \ -d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed qed } note neg_sum_2'' = this { fix a b assume A: "(a,b) \ set (arcs i i ys)" assume not0: "a > 0" "b > 0" assume bounded: "M\<^sub>R a 0 \ \" "M\<^sub>R b 0 \ \" assume lt: "M a b < M\<^sub>R a b" from clock_dest[OF A not0] obtain c1 c2 where C: "v c1 = a" "v c2 = b" "c1 \ X" "c2 \ X" and C2: "a \ n" "b \ n" by blast from C C2 clock_numbering(1,3) have C3: "v' b = c2" "v' a = c1" unfolding v'_def by blast+ with C C2 not0 bounded M\<^sub>R(4) obtain d :: int where *: "- int (k c2) \ d \ d \ int (k c1) \ M\<^sub>R a b = Le d \ M\<^sub>R b a = Le (- d) \ - int (k c2) \ d - 1 \ d \ int (k c1) \ M\<^sub>R a b = Lt d \ M\<^sub>R b a = Lt (- d + 1)" unfolding v'_def by force from * have ?thesis proof (standard, goal_cases) case 1 with lt have "M a b < Le d" by auto then have "M a b \ Lt d" unfolding less less_eq dbm_le_def by (fastforce elim!: dbm_lt.cases) from dbm_lt'[OF assms(2)[folded M(1)] this C2 C(1,2) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 < d}" . from beta_interp.\_boundedness_diag_lt'[OF _ _ C(3,4) this] 1 have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 - u c2 < d}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c1) (Some c2) (M\<^sub>R a b)" "dbm_entry_val u (Some c2) (Some c1) (M\<^sub>R b a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with 1 have "u \ {u \ V. u c1 - u c2 < d}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto next case 2 with lt have "M a b \ \" by auto with dbm_entry_int[OF this] M(3) \a \ n\ \b \ n\ obtain d' :: int where d': "M a b = Le d' \ M a b = Lt d'" by auto then have "M a b \ Le (d - 1)" using lt 2 apply (auto simp: less_eq dbm_le_def less) apply (cases rule: dbm_lt.cases) apply auto apply (rule dbm_lt.intros) apply (cases rule: dbm_lt.cases) by auto with lt have "M a b \ Le (d - 1)" by auto from dbm_le'[OF assms(2)[folded M(1)] this C2 C(1,2) not0] have "[M]\<^bsub>v,n\<^esub> \ {u \ V. u c1 - u c2 \ d - 1}" . from beta_interp.\_boundedness_diag_le'[OF _ _ C(3,4) this] 2 have "Approx\<^sub>\ ([M]\<^bsub>v,n\<^esub>) \ {u \ V. u c1 - u c2 \ d - 1}" by auto moreover { fix u assume u: "u \ [M\<^sub>R]\<^bsub>v,n\<^esub>" with C C2 have "dbm_entry_val u (Some c2) (Some c1) (M\<^sub>R b a)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with 2 have "u \ {u \ V. u c1 - u c2 \ d - 1}" by auto } ultimately show ?thesis using M\<^sub>R(1) M(1) by auto qed } note bounded = this { assume not_bounded: "\ (a,b) \ set (arcs i i ys). M a b < M\<^sub>R a b \ M\<^sub>R a 0 = \ \ M\<^sub>R b 0 = \" have "\ y z zs. set zs \ {0, y, z} = set (i # ys) \ len ?M 0 0 (y # z # zs) < Le 0 \ (\ (a,b) \ set (arcs 0 0 (y # z # zs)). M a b < M\<^sub>R a b \ a = y \ b = z) \ M y z < M\<^sub>R y z \ distinct (0 # y # z # zs) \ ?thesis" proof (cases ys) case Nil show ?thesis proof (cases "M i i < M\<^sub>R i i") case True - then have "?M i i = M i i" by (simp add: min.strict_order_iff) + then have "?M i i = M i i" + by simp with Nil ys(1) xs(3) have *: "M i i < \" by simp with neg_cycle_empty[OF cn_weak _ \i \ n\, of "[]" M] have "[M]\<^bsub>v,n\<^esub> = {}" by auto with \Z \ {}\ M(1) show ?thesis by auto next case False then have "?M i i = M\<^sub>R i i" by (simp add: min_absorb2) with Nil ys(1) xs(3) have "M\<^sub>R i i < \" by simp with neg_cycle_empty[OF cn_weak _ \i \ n\, of "[]" M\<^sub>R] have "[M\<^sub>R]\<^bsub>v,n\<^esub> = {}" by auto with \R \ {}\ M\<^sub>R(1) show ?thesis by auto qed next case (Cons w ws) note ws = this show ?thesis proof (cases ws) case Nil with ws ys xs(3) have *: "?M i w + ?M w i < \" "?M w i = M w i \ ?M i w \ M i w" "(i, w) \ set (arcs i i ys)" by auto have "R \ Approx\<^sub>\ Z = {}" proof (cases "?M w i = M w i") case True with *(2) have "?M i w = M\<^sub>R i w" unfolding min_def by auto with *(1) True have neg: "M\<^sub>R i w + M w i < \" by auto show ?thesis proof (cases "i = 0") case True show ?thesis proof (cases "w = 0") case True with 0 \i = 0\ *(3) show ?thesis by auto next case False with \i = 0\ neg_sum_2'' *(3) neg show ?thesis by blast qed next case False show ?thesis proof (cases "w = 0") case True with \i \ 0\ neg_sum_2' *(3) neg show ?thesis by blast next case False with \i \ 0\ neg_sum_2 *(3) neg show ?thesis by blast qed qed next case False have "M\<^sub>R w i < M w i" proof (rule ccontr, goal_cases) case 1 then have "M\<^sub>R w i \ M w i" by auto with False show False unfolding min_def by auto qed with one_M ws Nil have "M i w < M\<^sub>R i w" by auto then have "?M i w = M i w" unfolding min_def by auto moreover from False *(2) have "?M w i = M\<^sub>R w i" unfolding min_def by auto ultimately have neg: "M i w + M\<^sub>R w i < \" using *(1) by auto show ?thesis proof (cases "i = 0") case True show ?thesis proof (cases "w = 0") case True with 0 \i = 0\ *(3) show ?thesis by auto next case False with \i = 0\ neg_sum_1'' *(3) neg show ?thesis by blast qed next case False show ?thesis proof (cases "w = 0") case True with \i \ 0\ neg_sum_1' *(3) neg show ?thesis by blast next case False with \i \ 0\ neg_sum_1 *(3) neg show ?thesis by blast qed qed qed then show ?thesis by simp next case zs: (Cons z zs) from one_M obtain a b where *: "(a,b) \ set (arcs i i ys)" "M a b < M\<^sub>R a b" by fastforce from cycle_rotate_3'[OF _ *(1) ys(3)] ws cycle_closes obtain ws' where ws': "len ?M i i ys = len ?M a a (b # ws')" "set (a # b # ws') = set (i # ys)" "1 + length ws' = length ys" "set (arcs i i ys) = set (arcs a a (b # ws'))" and successive: "successive (\(a, b). ?M a b = M a b) (arcs a a (b # ws') @ [(a, b)])" by blast from successive have successive_arcs: "successive (\(a, b). ?M a b = M a b) (arcs a b (b # ws' @ [a]))" using arcs_decomp_tail by auto from ws'(4) one_M_R *(2) obtain c d where **: "(c,d) \ set (arcs a a (b # ws'))" "M c d > M\<^sub>R c d" "(a,b) \ (c,d)" by fastforce from card_distinct[of "a # b # ws'"] distinct_card[of "i # ys"] ws'(2,3) distinct have distinct: "distinct (a # b # ws')" by simp from ws zs ws'(3) have "ws' \ []" by auto then obtain z zs where z: "ws' = zs @ [z]" by (metis append_butlast_last_id) then have "b # ws' = (b # zs) @ [z]" by simp with len_decomp[OF this, of ?M a a] arcs_decomp_tail have rotated: "len ?M a a (b # ws') = len ?M z z (a # b # zs)" "set (arcs a a (b # ws')) = set (arcs z z (a # b # zs))" by (auto simp add: comm) from ys(1) xs(3) ws'(1) have "len ?M a a (b # ws') < \" by auto from ws'(2) ys(2) \i \ n\ z have n_bounds: "a \ n" "b \ n" "set ws' \ {0..n}" "z \ n" by auto - from * have a_b: "?M a b = M a b" by (simp add: min.strict_order_iff) + from * have a_b: "?M a b = M a b" + by simp from successive successive_split[of _ "arcs a z (b # zs)" "[(z,a), (a,b)]"] have first: "successive (\(a, b). ?M a b = M a b) (arcs a z (b # zs))" and last_two: "successive (\(a, b). ?M a b = M a b) [(z, a), (a, b)]" using arcs_decomp_tail z by auto from * not_bounded have not_bounded': "M\<^sub>R a 0 = \ \ M\<^sub>R b 0 = \" by auto from this(1) have "z = 0" proof assume inf: "M\<^sub>R b 0 = \" from a_b successive obtain z where z: "(b,z) \ set (arcs b a ws')" "?M b z \ M b z" by (cases ws') auto then have "?M b z = M\<^sub>R b z" by (meson min_def) from arcs_distinct2[OF _ _ _ _ z(1)] distinct have "b \ z" by auto from z n_bounds have "z \ n" apply (induction ws' arbitrary: b) apply auto[] apply (rename_tac ws' b) apply (case_tac ws') apply auto done have "M\<^sub>R b z = \" proof (cases "z = 0") case True with inf show ?thesis by auto next case False with inf M\<^sub>R(2) \b \ z\ \z \ n\ \b \ n\ show ?thesis by blast qed with \?M b z = M\<^sub>R b z\ have "len ?M b a ws' = \" by (auto intro: len_inf_elem[OF z(1)]) then have "\ = len ?M a a (b # ws')" by simp with \len ?M a a _ < \\ show ?thesis by auto next assume inf: "M\<^sub>R a 0 = \" show "z = 0" proof (rule ccontr) assume "z \ 0" with last_two a_b have "?M z a = M\<^sub>R z a" by (auto simp: min_def) from distinct z have "a \ z" by auto with \z \ 0\ \a \ n\ \z \ n\ M\<^sub>R(2) inf have "M\<^sub>R z a = \" by blast with \?M z a = M\<^sub>R z a\ have "len ?M z z (a # b # zs) = \" by (auto intro: len_inf_elem) with \len ?M a a _ < \\ rotated show False by auto qed qed { fix c d assume A: "(c, d) \ set (arcs 0 0 (a # b # zs))" "M c d < M\<^sub>R c d" - then have *: "?M c d = M c d" by (simp add: min.strict_order_iff) + then have *: "?M c d = M c d" + by simp from rotated(2) A \z = 0\ not_bounded ws'(4) have **: "M\<^sub>R c 0 = \ \ M\<^sub>R d 0 = \" by auto { assume inf: "M\<^sub>R c 0 = \" fix x assume x: "(x, c) \ set (arcs a 0 (b # zs))" "?M x c \ M x c" from x(2) have "?M x c = M\<^sub>R x c" unfolding min_def by auto from arcs_elem[OF x(1)] z \z = 0\ have "x \ set (a # b # ws')" "c \ set (a # b # ws')" by auto with n_bounds have "x \ n" "c \ n" by auto have "x = 0" proof (rule ccontr) assume "x \ 0" from distinct z arcs_distinct1[OF _ _ _ _ x(1)] \z = 0\have "x \ c" by auto with \x \ 0\ \c \ n\ \x \ n\ M\<^sub>R(2) inf have "M\<^sub>R x c = \" by blast with \?M x c = M\<^sub>R x c\ have "len ?M a 0 (b # zs) = \" by (fastforce intro: len_inf_elem[OF x(1)]) with \z = 0\ have "len ?M z z (a # b # zs) = \" by auto with \len ?M a a _ < \\ rotated show False by auto qed with arcs_distinct_dest1[OF _ x(1), of z] z distinct x \z = 0\ have False by auto } note c_0_inf = this have "a = c \ b = d" proof (cases "(c, d) = (0, a)") case True with last_two \z = 0\ * a_b have False by auto then show ?thesis by simp next case False show ?thesis proof (rule ccontr, goal_cases) case 1 with False A(1) have ***: "(c, d) \ set (arcs b 0 zs)" by auto from successive z \z = 0\ have "successive (\(a, b). ?M a b = M a b) ([(a, b)] @ arcs b 0 zs @ [(0, a), (a, b)])" by (simp add: arcs_decomp) then have ****: "successive (\(a, b). ?M a b = M a b) (arcs b 0 zs)" using successive_split[of _ "[(a, b)]" "arcs b 0 zs @ [(0, a), (a, b)]"] successive_split[of _ "arcs b 0 zs" "[(0, a), (a, b)]"] by auto from successive_predecessor[OF *** _ this] successive z obtain x where x: "(x, c) \ set (arcs a 0 (b # zs))" "?M x c \ M x c" proof (cases "c = b") case False then have "zs \ []" using *** by auto from successive_predecessor[OF *** False **** _ this] * obtain x where x: "(zs = [c] \ x = b \ (\ys. zs = c # d # ys \ x = b) \ (\ys. zs = ys @ [x, c] \ d = 0) \ (\ys ws. zs = ys @ x # c # d # ws))" "?M x c \ M x c" by blast+ from this(1) have "(x, c) \ set (arcs a 0 (b # zs))" using arcs_decomp by auto with x(2) show ?thesis by (auto intro: that) next case True have ****: "successive (\(a, b). ?M a b = M a b) (arcs a 0 (b # zs))" using first \z = 0\ arcs_decomp successive_arcs z by auto show ?thesis proof (cases zs) case Nil with **** True *** * show ?thesis by (auto intro: that) next case (Cons u us) with *** True distinct z \z = 0\ have "distinct (b # u # us @ [0])" by auto from arcs_distinct_fix[OF this] *** True Cons have "d = u" by auto with **** * Cons True show ?thesis by (auto intro: that) qed qed show False proof (cases "d = 0") case True from ** show False proof assume "M\<^sub>R c 0 = \" from c_0_inf[OF this x] show False . next assume "M\<^sub>R d 0 = \" with \d = 0\ M\<^sub>R(3) show False by auto qed next case False with *** have "zs \ []" by auto from successive_successor[OF \(c,d) \ set (arcs b 0 zs)\ False **** _ this] * obtain e where "(zs = [d] \ e = 0 \ (\ys. zs = d # e # ys) \ (\ys. zs = ys @ [c, d] \ e = 0) \ (\ys ws. zs = ys @ c # d # e # ws))" "?M d e \ M d e" by blast then have e: "(d, e) \ set (arcs b 0 zs)" "?M d e \ M d e" using arcs_decomp by auto from ** show False proof assume inf: "M\<^sub>R d 0 = \" from e have "?M d e = M\<^sub>R d e" by (meson min_def) from arcs_distinct2[OF _ _ _ _ e(1)] z \z = 0\ distinct have "d \ e" by auto from z n_bounds have "set zs \ {0..n}" by auto with e have "e \ n" apply (induction zs arbitrary: d) apply auto apply (case_tac zs) apply auto done from n_bounds z arcs_elem(2)[OF A(1)] have "d \ n" by auto have "M\<^sub>R d e = \" proof (cases "e = 0") case True with inf show ?thesis by auto next case False with inf M\<^sub>R(2) \d \ e\ \e \ n\ \d \ n\ show ?thesis by blast qed with \?M d e = M\<^sub>R d e\ have "len ?M b 0 zs = \" by (auto intro: len_inf_elem[OF e(1)]) with \z = 0\ rotated have "\ = len ?M a a (b # ws')" by simp with \len ?M a a _ < \\ show ?thesis by auto next assume "M\<^sub>R c 0 = \" from c_0_inf[OF this x] show False . qed qed qed qed } then have "\(c, d)\set (arcs 0 0 (a # b # zs)). M c d < M\<^sub>R c d \ c = a \ d = b" by blast moreover from ys(1) xs(3) have "len ?M i i ys < Le 0" unfolding neutral by auto moreover with rotated ws'(1) have "len ?M z z (a # b # zs) < Le 0" by auto moreover from \z = 0\ z ws'(2) have "set zs \ {0, a, b} = set (i # ys)" by auto moreover from \z = 0\ distinct z have "distinct (0 # a # b # zs)" by auto ultimately show ?thesis using \z = 0\ \M a b < M\<^sub>R a b\ by blast qed qed note * = this { assume "\ ?thesis" with * obtain y z zs where *: "set zs \ {0, y, z} = set (i # ys)" "len ?M 0 0 (y # z # zs) < Le 0" "\(a, b)\set (arcs 0 0 (y # z # zs)). M a b < M\<^sub>R a b \ a = y \ b = z" "M y z < M\<^sub>R y z" and distinct': "distinct (0 # y # z # zs)" by blast then have "y \ 0" "z \ 0" by auto let ?r = "len M\<^sub>R z 0 zs" have "\(a, b)\set (arcs z 0 zs). ?M a b = M\<^sub>R a b" proof (safe, goal_cases) case A: (1 a b) have "M\<^sub>R a b \ M a b" proof (rule ccontr, goal_cases) case 1 with *(3) A have "a = y" "b = z" by auto with A distinct' arcs_distinct3[OF _ A, of y] show False by auto qed then show ?case by (simp add: min_def) qed then have r: "len ?M z 0 zs = ?r" by (induction zs arbitrary: z) auto with *(2) have **: "?M 0 y + (?M y z + ?r) < Le 0" by simp from M\<^sub>R(1) \R \ {}\ obtain u where u: "DBM_val_bounded v u M\<^sub>R n" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto from *(1) \i \ n\ \set ys \ _\ have "y \ n" "z \ n" by fastforce+ from *(1) ys(2,4) have "set zs \ {0 ..n}" by auto from \y \ n\ \z \ n\ clock_numbering(2) \y \ 0\ \z \ 0\ obtain c1 c2 where C: "c1 \ X" "c2 \ X" "v c1 = y" "v c2 = z" by blast+ with clock_numbering(1,3) have C2: "v' y = c1" "v' z = c2" unfolding v'_def by auto with C have "v (v' z) = z" by auto with DBM_val_bounded_len'1[OF u, of zs "v' z"] have "dbm_entry_val u (Some (v' z)) None ?r" using \z \ n\ clock_numbering(2) \set zs \ _\ distinct' by force from len_inf_elem ** have tl_not_inf: "\(a, b)\set (arcs z 0 zs). M\<^sub>R a b \ \" by fastforce with M\<^sub>R(7) len_int_dbm_closed have "get_const ?r \ \ \ ?r \ \" by blast then obtain r :: int where r': "?r = Le r \ ?r = Lt r" using Ints_cases by (cases ?r) auto from r' \dbm_entry_val _ _ _ _\ C C2 have le: "u (v' z) \ r" by fastforce from arcs_ex_head obtain z' where "(z, z') \ set (arcs z 0 zs)" by blast then have z': "(z, z') \ set (arcs 0 0 (y # z # zs))" "(z, z') \ set (arcs z 0 zs)" by auto have "M\<^sub>R z 0 \ \" proof (rule ccontr, goal_cases) case 1 then have inf: "M\<^sub>R z 0 = \" by auto have "M\<^sub>R z z' = \" proof (cases "z' = 0") case True with 1 show ?thesis by auto next case False from arcs_elem[OF z'(1)] *(1) \i \ n\ \set ys \ _\ have "z' \ n" by fastforce moreover from distinct' *(1) arcs_distinct1[OF _ _ _ _ z'(1)] have "z \ z'" by auto ultimately show ?thesis using M\<^sub>R(2) \z \ n\ False inf by blast qed with tl_not_inf z'(2) show False by auto qed with M\<^sub>R(5) \z \ 0\ \z \ n\ obtain d :: int where d: "M\<^sub>R z 0 = Le d \ M\<^sub>R 0 z = Le (-d) \ M\<^sub>R z 0 = Lt d \ M\<^sub>R 0 z = Lt (-d + 1)" "d \ k (v' z)" "0 \ d" unfolding v'_def by auto text \Needs property that len of integral dbm entries is integral and definition of \M_R\\ from this (1) have rr: "?r \ M\<^sub>R z 0" proof (standard, goal_cases) case A: 1 with u \z \ n\ C C2 have *: "- u (v' z) \ -d" unfolding DBM_val_bounded_def by fastforce from r' show ?case proof (standard, goal_cases) case 1 with le * A show ?case unfolding less_eq dbm_le_def by fastforce next case 2 with \dbm_entry_val _ _ _ _\ C C2 have "u (v' z) < r" by fastforce with * have "r > d" by auto with A 2 show ?case unfolding less_eq dbm_le_def by fastforce qed next case A: 2 with u \z \ n\ C C2 have *: "- u (v' z) < -d + 1" unfolding DBM_val_bounded_def by fastforce from r' show ?case proof (standard, goal_cases) case 1 with le * A show ?case unfolding less_eq dbm_le_def by fastforce next case 2 with \dbm_entry_val _ _ _ _\ C C2 have "u (v' z) \ r" by fastforce with * have "r \ d" by auto with A 2 show ?case unfolding less_eq dbm_le_def by fastforce qed qed with *(3) \y \ 0\ have "M 0 y \ M\<^sub>R 0 y" by fastforce then have "?M 0 y = M\<^sub>R 0 y" by (simp add: min.absorb2) moreover from *(4) have "?M y z = M y z" unfolding min_def by auto ultimately have **: "M\<^sub>R 0 y + (M y z + M\<^sub>R z 0) < Le 0" using ** add_mono_right[OF add_mono_right[OF rr], of "M\<^sub>R 0 y" "M y z"] by simp from ** have not_inf: "M\<^sub>R 0 y \ \" "M y z \ \" "M\<^sub>R z 0 \ \" by auto from M\<^sub>R(6) \y \ 0\ \y \ n\ obtain c :: int where c: "M\<^sub>R 0 y = Le c \ M\<^sub>R 0 y = Lt c" "- k (v' y) \ c" "c \ 0" unfolding v'_def by auto have ?thesis proof (cases "M\<^sub>R 0 y + M\<^sub>R z 0 = Lt (c + d)") case True from ** have "(M\<^sub>R 0 y + M\<^sub>R z 0) + M y z < Le 0" using comm assoc by metis with True have **: "Lt (c + d) + M y z < Le 0" by simp then have "M y z \ Le (- (c + d))" unfolding less less_eq dbm_le_def mult by (cases "M y z") (fastforce elim!: dbm_lt.cases)+ from dbm_le'[OF assms(2)[folded M(1)] this \y \ n\ \z \ n\ C(3,4)] \y \ 0\ \z \ 0\ M have subs: "Z \ {u \ V. u c1 - u c2 \ - (c + d)}" by blast with c d have "- k (v' z) \ - (c + d)" "- (c + d) \ k (v' y)" by auto with beta_interp.\_boundedness_diag_le'[OF _ _ C(1,2) subs] C2 have "Approx\<^sub>\ Z \ {u \ V. u c1 - u c2 \ - (c + d)}" by auto moreover { fix u assume u: "u \ R" with C \y \ n\ \z \ n\ M\<^sub>R(1) have "dbm_entry_val u (Some c2) None (M\<^sub>R z 0)" "dbm_entry_val u None (Some c1) (M\<^sub>R 0 y)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with True c d(1) have "u \ {u \ V. u c1 - u c2 \ - (c + d)}" unfolding mult by auto } ultimately show ?thesis by blast next case False with c d have "M\<^sub>R 0 y + M\<^sub>R z 0 = Le (c + d)" unfolding mult by fastforce moreover from ** have "(M\<^sub>R 0 y + M\<^sub>R z 0) + M y z < Le 0" using comm assoc by metis ultimately have **: "Le (c + d) + M y z < Le 0" by simp then have "M y z \ Lt (- (c + d))" unfolding less less_eq dbm_le_def mult by (cases "M y z") (fastforce elim!: dbm_lt.cases)+ from dbm_lt'[OF assms(2)[folded M(1)] this \y \ n\ \z \ n\ C(3,4)] \y \ 0\ \z \ 0\ M have subs: "Z \ {u \ V. u c1 - u c2 < - (c + d)}" by auto from c d(2-) C2 have "- k c2 \ - (c + d)" "- (c + d) \ k c1" by auto from beta_interp.\_boundedness_diag_lt'[OF this C(1,2) subs] have "Approx\<^sub>\ Z \ {u \ V. u c1 - u c2 < - (c + d)}" . moreover { fix u assume u: "u \ R" with C \y \ n\ \z \ n\ M\<^sub>R(1) have "dbm_entry_val u (Some c2) None (M\<^sub>R z 0)" "dbm_entry_val u None (Some c1) (M\<^sub>R 0 y)" unfolding DBM_zone_repr_def DBM_val_bounded_def by auto with c d(1) have "u \ {u \ V. u c1 - u c2 < - (c + d)}" by auto } ultimately show ?thesis by auto qed } then have ?thesis by auto } with bounded 0 bounded_zero_1 bounded_zero_2 show ?thesis by blast qed qed qed section \Nice Corollaries of Bouyer's Theorem\ lemma \_V: "\ \ = V" unfolding V_def \_def using region_cover[of X _ k] by auto lemma regions_beta_V: "R \ \\<^sub>\ \ R \ V" unfolding V_def \\<^sub>\_def by auto lemma apx_V: "Z \ V \ Approx\<^sub>\ Z \ V" proof (goal_cases) case 1 from beta_interp.apx_in[OF 1] obtain U where "Approx\<^sub>\ Z = \U" "U \ \\<^sub>\" by auto with regions_beta_V show ?thesis by auto qed corollary approx_\_closure_\: assumes "Z \ V" "vabstr Z M" shows "Approx\<^sub>\ Z \ Closure\<^sub>\ Z" proof - note T = region_zone_intersect_empty_approx_correct[OF _ assms(1) _ assms(2-)] have "- \{R \ \. R \ Z \ {}} = \{R \ \. R \ Z = {}} \ - V" proof (safe, goal_cases) case 1 with \_V show False by fast next case 2 then show ?case using alpha_interp.valid_regions_distinct_spec by fastforce next case 3 then show ?case using \_V unfolding V_def by blast qed with T apx_V[OF assms(1)] have "Approx\<^sub>\ Z \ - \{R \ \. R \ Z \ {}} = {}" by auto then show ?thesis unfolding alpha_interp.cla_def by blast qed definition "V' \ {Z. Z \ V \ (\ M. vabstr Z M)}" corollary approx_\_closure_\': "Z \ V' \ Approx\<^sub>\ Z \ Closure\<^sub>\ Z" using approx_\_closure_\ unfolding V'_def by auto text \We could prove this more directly too (without using \Closure\<^sub>\ Z\), obviously\ lemma apx_empty_iff: assumes "Z \ V" "vabstr Z M" shows "Z = {} \ Approx\<^sub>\ Z = {}" using alpha_interp.cla_empty_iff[OF assms(1)] approx_\_closure_\[OF assms] beta_interp.apx_subset by auto lemma apx_empty_iff': assumes "Z \ V'" shows "Z = {} \ Approx\<^sub>\ Z = {}" using apx_empty_iff assms unfolding V'_def by force lemma apx_V': assumes "Z \ V" shows "Approx\<^sub>\ Z \ V'" proof (cases "Z = {}") case True with beta_interp.apx_empty beta_interp.empty_zone_dbm show ?thesis unfolding V'_def neutral by auto next case False then have non_empty: "Approx\<^sub>\ Z \ {}" using beta_interp.apx_subset by blast from beta_interp.apx_in[OF assms] obtain U M where *: "Approx\<^sub>\ Z = \U" "U \ \\<^sub>\" "Z \ Approx\<^sub>\ Z" "vabstr (Approx\<^sub>\ Z) M" by blast moreover from * beta_interp.\_union have "\ U \ V" by blast ultimately show ?thesis using *(1,4) unfolding V'_def by auto qed section \A New Zone Semantics Abstracting with \Approx\<^sub>\\\ lemma step_z_V': assumes "A \ \l,Z\ \ \l',Z'\" "valid_abstraction A X k" "\c\clk_set A. v c \ n" "Z \ V'" shows "Z' \ V'" proof - from assms(3) clock_numbering have numbering: "global_clock_numbering A v n" by metis from assms(4) obtain M where M: "Z \ V" "Z = [M]\<^bsub>v,n\<^esub>" "dbm_int M n" unfolding V'_def by auto from alpha_interp.step_z_V[OF assms(1) M(1)] M(2) assms(1) step_z_dbm_DBM[OF _ numbering] step_z_dbm_preserves_int[OF _ numbering assms(2) M(3)] obtain M' where M': "Z' \ V" "Z' = [M']\<^bsub>v,n\<^esub>" "dbm_int M' n" by metis then show ?thesis unfolding V'_def by blast qed lemma steps_z_V': "A \ \l,Z\ \* \l',Z'\ \ valid_abstraction A X k \ \c\clk_set A. v c \ n \ Z \ V' \ Z' \ V'" by (induction rule: steps_z.induct) (auto intro: step_z_V') subsection \Single Step\ inductive step_z_beta :: "('a, 'c, t, 's) ta \ 's \ ('c, t) zone \ 's \ ('c, t) zone \ bool" ("_ \ \_, _\ \\<^sub>\ \_, _\" [61,61,61] 61) where step_beta: "A \ \l, Z\ \ \l', Z'\ \ A \ \l, Z\ \\<^sub>\ \l', Approx\<^sub>\ Z'\" inductive_cases[elim!]: "A \ \l, u\ \\<^sub>\ \l',u'\" declare step_z_beta.intros[intro] lemma step_z_alpha_sound: "A \ \l, Z\ \\<^sub>\ \l',Z'\ \ valid_abstraction A X k \ \c\clk_set A. v c \ n \ Z \ V' \ Z' \ {} \ \ Z''. A \ \l, Z\ \ \l',Z''\ \ Z'' \ {}" apply (induction rule: step_z_beta.induct) apply (frule step_z_V') apply assumption+ apply (rotate_tac 4) apply (drule apx_empty_iff') by blast lemma step_z_alpha_complete: "A \ \l, Z\ \ \l',Z'\ \ valid_abstraction A X k \ \c\clk_set A. v c \ n \ Z \ V' \ Z' \ {} \ \ Z''. A \ \l, Z\ \\<^sub>\ \l', Z''\ \ Z'' \ {}" apply (frule step_z_V') apply assumption+ apply (rotate_tac 4) apply (drule apx_empty_iff') by blast subsection \Multi step\ inductive steps_z_beta :: "('a, 'c, t, 's) ta \ 's \ ('c, t) zone \ 's \ ('c, t) zone \ bool" ("_ \ \_, _\ \\<^sub>\* \_, _\" [61,61,61] 61) where refl: "A \ \l, Z\ \\<^sub>\* \l, Z\" | step: "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ A \ \l', Z'\ \\<^sub>\ \l'', Z''\ \ A \ \l, Z\ \\<^sub>\* \l'', Z''\" declare steps_z_beta.intros[intro] lemma V'_V: "Z \ V' \ Z \ V" unfolding V'_def by auto lemma steps_z_beta_V': "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ valid_abstraction A X k \\c\clk_set A. v c \ n \ Z \ V' \ Z' \ V'" proof (induction rule: steps_z_beta.induct) case refl then show ?case by fast next case (step A l Z l' Z' l'' Z'') from this(2) obtain Z''' where Z''': "A \ \l', Z'\ \ \l'',Z'''\" "Z'' = Approx\<^sub>\ Z'''" by auto from step_z_V'[OF this(1)] step have "Z''' \ V'" by auto from apx_V'[OF V'_V, OF this] Z'''(2) show ?case by auto qed lemma alpha_beta_step: "A \ \l, Z\ \\<^sub>\ \l', Z'\ \ valid_abstraction A X k \ \c\clk_set A. v c \ n \ Z \ V' \ \ Z''. A \ \l, Z\ \\<^sub>\ \l', Z''\ \ Z' \ Z''" apply (induction rule: step_z_beta.induct) apply (frule step_z_V') apply assumption+ apply (rotate_tac 4) apply (drule approx_\_closure_\') apply auto done subsubsection \Soundness\ lemma alpha_beta_step': "A \ \l, Z\ \\<^sub>\ \l', Z'\ \ valid_abstraction A X k \ \c\clk_set A. v c \ n \ Z \ V' \ W \ V \ Z \ W \ \ W'. A \ \l, W\ \\<^sub>\ \l', W'\ \ Z' \ W'" proof (induction rule: step_z_beta.induct) case (step_beta A l Z l' Z') from alpha_interp.step_z_mono[OF step_beta(1,6)] obtain W' where W': "A \ \l, W\ \ \l',W'\" "Z' \ W'" by blast from approx_\_closure_\'[OF step_z_V'[OF step_beta(1-4)]] alpha_interp.cla_mono[OF this(2)] this(1) show ?case by auto qed lemma alpha_beta_steps: "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ valid_abstraction A X k \ \c\clk_set A. v c \ n \ Z \ V' \ \ Z''. A \ \l, Z\ \\<^sub>\* \l', Z''\ \ Z' \ Z''" proof (induction rule: steps_z_beta.induct) case refl then show ?case by auto next case (step A l Z l' Z' l'' Z'') then obtain Z''' where *: "A \ \l, Z\ \\<^sub>\* \l',Z'''\" "Z' \ Z'''" by auto from alpha_beta_step'[OF step.hyps(2) step.prems(1,2) steps_z_beta_V'[OF step.hyps(1) step.prems] alpha_interp.steps_z_alpha_V[OF this(1) V'_V] this(2)] step.prems obtain W' where "A \ \l', Z'''\ \\<^sub>\ \l'',W'\" "Z'' \ W'" by blast with * show ?case by auto qed corollary steps_z_beta_sound: "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ \c\clk_set A. v c \ n \ valid_abstraction A X k \ Z \ V' \ Z' \ {} \ \ Z''. A \ \l, Z\ \* \l', Z''\ \ Z'' \ {}" proof (goal_cases) case 1 then have "Z \ V" unfolding V'_def by auto from alpha_beta_steps[OF 1(1,3,2,4)] obtain Z''' where *: "A \ \l, Z\ \\<^sub>\* \l',Z'''\" "Z' \ Z'''" by blast from alpha_interp.steps_z_alpha_closure_involutive[OF *(1) 1(3) \Z \ V\] obtain Z'' where Z'': "A \ \l, Z\ \* \l',Z''\" "Closure\<^sub>\ Z''' \ Closure\<^sub>\ Z''" "Z'' \ Z'''" by blast with alpha_interp.closure_subs[OF alpha_interp.steps_z_alpha_V[OF *(1) \Z \ V\]] 1(5) alpha_interp.cla_empty_iff[OF alpha_interp.steps_z_V, OF this(1) \Z \ V\] *(2) have "Z'' \ {}" by auto with Z'' show ?thesis by auto qed subsubsection \Completeness\ lemma apx_mono: "Z' \ V \ Z \ Z' \ Approx\<^sub>\ Z \ Approx\<^sub>\ Z'" proof (goal_cases) case 1 with beta_interp.apx_in have "Approx\<^sub>\ Z' \ {S. \U M. S = \U \ U \ \\<^sub>\ \ Z' \ S \ beta_interp.vabstr S M \ beta_interp.normalized M}" by auto with 1 obtain U M where "Approx\<^sub>\ Z' = \U" "U \ \\<^sub>\" "Z \ Approx\<^sub>\ Z'" "beta_interp.vabstr (Approx\<^sub>\ Z') M" "beta_interp.normalized M" by auto with beta_interp.apx_min show ?thesis by auto qed lemma step_z_beta_mono: "A \ \l, Z\ \\<^sub>\ \l', Z'\ \ Z \ W \ W \ V \ \ W'. A \ \l, W\ \\<^sub>\ \l', W'\ \ Z' \ W'" proof (goal_cases) case 1 then obtain Z'' where *: "A \ \l, Z\ \ \l',Z''\" "Z' = Approx\<^sub>\ Z''" by auto from alpha_interp.step_z_mono[OF this(1) 1(2)] obtain W' where "A \ \l, W\ \ \l',W'\" "Z'' \ W'" by auto moreover with *(2) apx_mono[OF alpha_interp.step_z_V] \W \ V\ have "Z' \ Approx\<^sub>\ W'" by metis ultimately show ?case by blast qed lemma steps_z_beta_V: "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ Z \ V \ Z' \ V" proof (induction rule: steps_z_beta.induct) case refl then show ?case by blast next case (step A l Z l' Z' l'' Z'') then obtain Z''' where "A \ \l', Z'\ \ \l'',Z'''\" "Z'' = Approx\<^sub>\ Z'''" by auto with alpha_interp.step_z_V[OF this(1)] apx_V step(3,4) show "Z'' \ V" by auto qed lemma steps_z_beta_mono: "A \ \l, Z\ \\<^sub>\* \l', Z'\ \ Z \ W \ W \ V \ \ W'. A \ \l, W\ \\<^sub>\* \l', W'\ \ Z' \ W'" proof (induction rule: steps_z_beta.induct) case refl then show ?case by auto next case (step A l Z l' Z' l'' Z'') then obtain W' where "A \ \l, W\ \\<^sub>\* \l',W'\" "Z' \ W'" by auto with step_z_beta_mono[OF step(2) this(2) steps_z_beta_V[OF this(1) step(5)]] show ?case by blast qed lemma steps_z_beta_alt: "A \ \l, Z\ \\<^sub>\ \l', Z'\ \ A \ \l', Z'\ \\<^sub>\* \l'', Z''\ \ A \ \l, Z\ \\<^sub>\* \l'', Z''\" by (rotate_tac, induction rule: steps_z_beta.induct) blast+ lemma steps_z_beta_complete: "A \ \l, Z\ \* \l', Z'\ \ valid_abstraction A X k \ Z \ V \ \ Z''. A \ \l, Z\ \\<^sub>\* \l',Z''\ \ Z' \ Z''" proof (induction rule: steps_z.induct) case refl with apx_empty_iff show ?case by blast next case (step A l Z l' Z' l'' Z'') with alpha_interp.step_z_V[OF this(1,5)] obtain Z''' where "A \ \l', Z'\ \\<^sub>\* \l'',Z'''\" "Z'' \ Z'''" by blast with steps_z_beta_mono[OF this(1) beta_interp.apx_subset apx_V[OF alpha_interp.step_z_V[OF step(1,5)]]] obtain W' where "A \ \l', Approx\<^sub>\ Z'\ \\<^sub>\* \l'', W'\" " Z'' \ W'" by auto moreover with step(1) have "A \ \l, Z\ \\<^sub>\* \l'',W'\" by (auto intro: steps_z_beta_alt) ultimately show ?case by auto qed lemma steps_z_beta_complete': "A \ \l, Z\ \* \l',Z'\ \ valid_abstraction A X k \ Z \ V \ Z' \ {} \ \ Z''. A \ \l, Z\ \\<^sub>\* \l',Z''\ \ Z'' \ {}" using steps_z_beta_complete by fast end end diff --git a/thys/Timed_Automata/Paths_Cycles.thy b/thys/Timed_Automata/Paths_Cycles.thy --- a/thys/Timed_Automata/Paths_Cycles.thy +++ b/thys/Timed_Automata/Paths_Cycles.thy @@ -1,1514 +1,1509 @@ theory Paths_Cycles imports Floyd_Warshall Timed_Automata begin chapter \Library for Paths, Arcs and Lengths\ lemma length_eq_distinct: assumes "set xs = set ys" "distinct xs" "length xs = length ys" shows "distinct ys" using assms card_distinct distinct_card by fastforce section \Arcs\ fun arcs :: "nat \ nat \ nat list \ (nat * nat) list" where "arcs a b [] = [(a,b)]" | "arcs a b (x # xs) = (a, x) # arcs x b xs" definition arcs' :: "nat list \ (nat * nat) set" where "arcs' xs = set (arcs (hd xs) (last xs) (butlast (tl xs)))" lemma arcs'_decomp: "length xs > 1 \ (i, j) \ arcs' xs \ \ zs ys. xs = zs @ i # j # ys" proof (induction xs) case Nil thus ?case by auto next case (Cons x xs) then have "length xs > 0" by auto then obtain y ys where xs: "xs = y # ys" by (metis Suc_length_conv less_imp_Suc_add) show ?case proof (cases "(i, j) = (x, y)") case True with xs have "x # xs = [] @ i # j # ys" by simp then show ?thesis by auto next case False then show ?thesis proof (cases "length ys > 0", goal_cases) case 2 then have "ys = []" by auto then have "arcs' (x#xs) = {(x,y)}" using xs by (auto simp add: arcs'_def) with Cons.prems(2) 2(1) show ?case by auto next case True with xs Cons.prems(2) False have "(i, j) \ arcs' xs" by (auto simp add: arcs'_def) with Cons.IH[OF _ this] True xs obtain zs ys where "xs = zs @ i # j # ys" by auto then have "x # xs = (x # zs) @ i # j # ys" by simp then show ?thesis by blast qed qed qed lemma arcs_decomp_tail: "arcs j l (ys @ [i]) = arcs j i ys @ [(i, l)]" by (induction ys arbitrary: j) auto lemma arcs_decomp: "xs = ys @ y # zs \ arcs x z xs = arcs x y ys @ arcs y z zs" by (induction ys arbitrary: x xs) simp+ lemma distinct_arcs_ex: "distinct xs \ i \ set xs \ xs \ [] \ \ a b. a \ x \ (a,b) \ set (arcs i j xs)" apply (induction xs arbitrary: i) apply simp apply (rename_tac a xs i) apply (case_tac xs) apply simp apply metis by auto lemma cycle_rotate_2_aux: "(i, j) \ set (arcs a b (xs @ [c])) \ (i,j) \ (c,b) \ (i, j) \ set (arcs a c xs)" by (induction xs arbitrary: a) auto lemma arcs_set_elem1: assumes "j \ k" "k \ set (i # xs)" shows "\ l. (k, l) \ set (arcs i j xs)" using assms by (induction xs arbitrary: i) auto lemma arcs_set_elem2: assumes "i \ k" "k \ set (j # xs)" shows "\ l. (l, k) \ set (arcs i j xs)" using assms proof (induction xs arbitrary: i) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases "k = x") auto qed section \Length of Paths\ lemmas (in linordered_ab_monoid_add) comm = add.commute lemma len_add: fixes M :: "('a :: linordered_ab_monoid_add) mat" shows "len M i j xs + len M i j xs = len (\i j. M i j + M i j) i j xs" proof (induction xs arbitrary: i j) case Nil thus ?case by auto next case (Cons x xs) have "M i x + len M x j xs + (M i x + len M x j xs) = M i x + (len M x j xs + M i x) + len M x j xs" by (simp add: assoc) also have "\ = M i x + (M i x + len M x j xs) + len M x j xs" by (simp add: comm) also have "\ = (M i x + M i x) + (len M x j xs + len M x j xs)" by (simp add: assoc) finally have "M i x + len M x j xs + (M i x + len M x j xs) = (M i x + M i x) + len (\i j. M i j + M i j) x j xs" using Cons by simp thus ?case by simp qed section \Canonical Matrices\ abbreviation "canonical M n \ \ i j k. i \ n \ j \ n \ k \ n \ M i k \ M i j + M j k" lemma fw_canonical: "cycle_free m n \ canonical (fw m n n n n) n" proof (clarify, goal_cases) case 1 with fw_shortest[OF \cycle_free m n\] show ?case by auto qed lemma canonical_len: "canonical M n \ i \ n \ j \ n \ set xs \ {0..n} \ M i j \ len M i j xs" proof (induction xs arbitrary: i) case Nil thus ?case by auto next case (Cons x xs) then have "M x j \ len M x j xs" by auto from Cons.prems \canonical M n\ have "M i j \ M i x + M x j" by simp also with Cons have "\ \ M i x + len M x j xs" by (simp add: add_mono) finally show ?case by simp qed section \Cycle Rotation\ lemma cycle_rotate: fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "length xs > 1" "(i, j) \ arcs' xs" shows "\ ys zs. len M a a xs = len M i i (j # ys @ a # zs) \ xs = zs @ i # j # ys" using assms proof - assume A: "length xs > 1" "(i, j) \ arcs' xs" from arcs'_decomp[OF this] obtain ys zs where xs: "xs = zs @ i # j # ys" by blast from len_decomp[OF this, of M a a] have "len M a a xs = len M a i zs + len M i a (j # ys)" . also have "\ = len M i a (j # ys) + len M a i zs" by (simp add: comm) also from len_comp[of M i i "j # ys" a zs] have "\ = len M i i (j # ys @ a # zs)" by auto finally show ?thesis using xs by auto qed lemma cycle_rotate_2: fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "xs \ []" "(i, j) \ set (arcs a a xs)" shows "\ ys. len M a a xs = len M i i (j # ys) \ set ys \ set (a # xs) \ length ys < length xs" using assms proof - assume A:"xs \ []" "(i, j) \ set (arcs a a xs)" { fix ys assume A:"a = i" "xs = j # ys" then have ?thesis by auto } note * = this { fix b ys assume A:"a = j" "xs = ys @ [i]" then have ?thesis proof (auto, goal_cases) case 1 have "len M j j (ys @ [i]) = M i j + len M j i ys" using len_decomp[of "ys @ [i]" ys i "[]" M j j] by (auto simp: comm) thus ?case by blast qed } note ** = this { assume "length xs = 1" then obtain b where xs: "xs = [b]" by (metis One_nat_def length_0_conv length_Suc_conv) with A(2) have "a = i \ b = j \ a = j \ b = i" by auto then have ?thesis using * ** xs by auto } note *** = this show ?thesis proof (cases "length xs = 0") case True with A show ?thesis by auto next case False thus ?thesis proof (cases "length xs = 1", goal_cases) case True with *** show ?thesis by auto next case 2 hence "length xs > 1" by linarith then obtain b c ys where ys:"xs = b # ys @ [c]" by (metis One_nat_def assms(1) 2(2) length_0_conv length_Cons list.exhaust rev_exhaust) thus ?thesis proof (cases "(i,j) = (a,b)", goal_cases) case True with ys * show ?thesis by auto next case False then show ?thesis proof (cases "(i,j) = (c,a)", goal_cases) case True with ys ** show ?thesis by auto next case 2 with A(2) ys have "(i, j) \ arcs' xs" using cycle_rotate_2_aux by (auto simp: arcs'_def) (* slow *) from cycle_rotate[OF \length xs > 1\ this, of M a] show ?thesis by auto qed qed qed qed qed lemma cycle_rotate_len_arcs: fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "length xs > 1" "(i, j) \ arcs' xs" shows "\ ys zs. len M a a xs = len M i i (j # ys @ a # zs) \ set (arcs a a xs) = set (arcs i i (j # ys @ a # zs)) \ xs = zs @ i # j # ys" using assms proof - assume A: "length xs > 1" "(i, j) \ arcs' xs" from arcs'_decomp[OF this] obtain ys zs where xs: "xs = zs @ i # j # ys" by blast from len_decomp[OF this, of M a a] have "len M a a xs = len M a i zs + len M i a (j # ys)" . also have "\ = len M i a (j # ys) + len M a i zs" by (simp add: comm) also from len_comp[of M i i "j # ys" a zs] have "\ = len M i i (j # ys @ a # zs)" by auto finally show ?thesis using xs arcs_decomp[OF xs, of a a] arcs_decomp[of "j # ys @ a # zs" "j # ys" a zs i i] by force qed lemma cycle_rotate_2': fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "xs \ []" "(i, j) \ set (arcs a a xs)" shows "\ ys. len M a a xs = len M i i (j # ys) \ set (i # j # ys) = set (a # xs) \ 1 + length ys = length xs \ set (arcs a a xs) = set (arcs i i (j # ys))" proof - note A = assms { fix ys assume A:"a = i" "xs = j # ys" then have ?thesis by auto } note * = this { fix b ys assume A:"a = j" "xs = ys @ [i]" then have ?thesis proof (auto, goal_cases) case 1 have "len M j j (ys @ [i]) = M i j + len M j i ys" using len_decomp[of "ys @ [i]" ys i "[]" M j j] by (auto simp: comm) moreover have "arcs j j (ys @ [i]) = arcs j i ys @ [(i, j)]" using arcs_decomp_tail by auto ultimately show ?case by auto qed } note ** = this { assume "length xs = 1" then obtain b where xs: "xs = [b]" by (metis One_nat_def length_0_conv length_Suc_conv) with A(2) have "a = i \ b = j \ a = j \ b = i" by auto then have ?thesis using * ** xs by auto } note *** = this show ?thesis proof (cases "length xs = 0") case True with A show ?thesis by auto next case False thus ?thesis proof (cases "length xs = 1", goal_cases) case True with *** show ?thesis by auto next case 2 hence "length xs > 1" by linarith then obtain b c ys where ys:"xs = b # ys @ [c]" by (metis One_nat_def assms(1) 2(2) length_0_conv length_Cons list.exhaust rev_exhaust) thus ?thesis proof (cases "(i,j) = (a,b)") case True with ys * show ?thesis by blast next case False then show ?thesis proof (cases "(i,j) = (c,a)", goal_cases) case True with ys ** show ?thesis by force next case 2 with A(2) ys have "(i, j) \ arcs' xs" using cycle_rotate_2_aux by (auto simp add: arcs'_def) (* slow *) from cycle_rotate_len_arcs[OF \length xs > 1\ this, of M a] show ?thesis by auto qed qed qed qed qed section \Equivalent Characterizations of Cycle-Freeness\ lemma negative_cycle_dest_diag: "\ cycle_free M n \ \ i xs. i \ n \ set xs \ {0..n} \ len M i i xs < \" proof (auto simp: cycle_free_def, goal_cases) case (1 i xs j) from this(4) have "len M i j xs < len M i j (rem_cycles i j xs)" by auto from negative_cycle_dest[OF this] obtain i' ys where *:"len M i' i' ys < \" "set (i' # ys) \ set (i # j # xs)" by auto from this(2) 1(1-3) have "set ys \ {0..n}" "i' \ n" by auto with * show ?case by auto next case 2 then show ?case by fastforce qed abbreviation cyc_free :: "('a::linordered_ab_monoid_add) mat \ nat \ bool" where "cyc_free m n \ \ i xs. i \ n \ set xs \ {0..n} \ len m i i xs \ \" lemma cycle_free_diag_intro: "cyc_free M n \ cycle_free M n" using negative_cycle_dest_diag by force lemma cycle_free_diag_equiv: "cyc_free M n \ cycle_free M n" using negative_cycle_dest_diag by (force simp: cycle_free_def) lemma cycle_free_diag_dest: "cycle_free M n \ cyc_free M n" using cycle_free_diag_equiv by blast lemma cyc_free_diag_dest: assumes "cyc_free M n" "i \ n" "set xs \ {0..n}" shows "len M i i xs \ \" using assms by auto lemma cycle_free_0_0: fixes M :: "('a::linordered_ab_monoid_add) mat" assumes "cycle_free M n" shows "M 0 0 \ \" using cyc_free_diag_dest[OF cycle_free_diag_dest[OF assms], of 0 "[]"] by auto section \More Theorems Related to Floyd-Warshall\ lemma D_cycle_free_len_dest: "cycle_free m n \ \ i \ n. \ j \ n. D m i j n = m' i j \ i \ n \ j \ n \ set xs \ {0..n} \ \ ys. set ys \ {0..n} \ len m' i j xs = len m i j ys" proof (induction xs arbitrary: i) case Nil with Nil have "m' i j = D m i j n" by simp from D_dest''[OF this] obtain ys where "set ys \ {0..n}" "len m' i j [] = len m i j ys" by auto then show ?case by auto next case (Cons y ys) from Cons.IH[OF Cons.prems(1,2) _ \j \ n\, of y] Cons.prems(5) obtain zs where zs:"set zs \ {0..n}" "len m' y j ys = len m y j zs" by auto with Cons have "m' i y = D m i y n" by simp from D_dest''[OF this] obtain ws where ws:"set ws \ {0..n}" "m' i y = len m i y ws" by auto with len_comp[of m i j ws y zs] zs Cons.prems(5) have "len m' i j (y # ys) = len m i j (ws @ y # zs)" "set (ws @ y # zs) \ {0..n}" by auto then show ?case by blast qed lemma D_cyc_free_preservation: "cyc_free m n \ \ i \ n. \ j \ n. D m i j n = m' i j \ cyc_free m' n" proof (auto, goal_cases) case (1 i xs) from D_cycle_free_len_dest[OF _ 1(2,3,3,4)] 1(1) cycle_free_diag_equiv obtain ys where "set ys \ {0..n} \ len m' i i xs = len m i i ys" by fast with 1(1,3) show ?case by auto qed abbreviation "FW m n \ fw m n n n n" lemma FW_cyc_free_preservation: "cyc_free m n \ cyc_free (FW m n) n" apply (rule D_cyc_free_preservation) apply assumption apply safe apply (rule fw_shortest_path) using cycle_free_diag_equiv by auto lemma cyc_free_diag_dest': "cyc_free m n \ i \ n \ m i i \ \" proof goal_cases case 1 then have "i \ n \ set [] \ {0..n}" by auto with 1(1) have "\ \ len m i i []" by blast then show ?case by auto qed lemma FW_diag_neutral_preservation: "\ i \ n. M i i = \ \ cyc_free M n \ \ i\n. (FW M n) i i = \" proof (auto, goal_cases) case (1 i) from this(3) have "(FW M n) i i \ M i i" by (auto intro: fw_mono) with 1 have "(FW M n) i i \ \" by auto with cyc_free_diag_dest'[OF FW_cyc_free_preservation[OF 1(2)] \i \ n\] show "FW M n i i = \" by auto qed lemma FW_fixed_preservation: fixes M :: "('a::linordered_ab_monoid_add) mat" assumes A: "i \ n" "M 0 i + M i 0 = \" "canonical (FW M n) n" "cyc_free (FW M n) n" shows "FW M n 0 i + FW M n i 0 = \" using assms proof - let ?M' = "FW M n" assume A: "i \ n" "M 0 i + M i 0 = \" "canonical ?M' n" "cyc_free ?M' n" from \i \ n\ have "?M' 0 i + ?M' i 0 \ M 0 i + M i 0" by (auto intro: fw_mono add_mono) with A(2) have "?M' 0 i + ?M' i 0 \ \" by auto moreover from \canonical ?M' n\ \i \ n\ have "?M' 0 0 \ ?M' 0 i + ?M' i 0" by auto moreover from cyc_free_diag_dest'[OF \cyc_free ?M' n\] have "\ \ ?M' 0 0" by simp ultimately show "?M' 0 i + ?M' i 0 = \" by (auto simp: add_mono) qed lemma diag_cyc_free_neutral: "cyc_free M n \ \k\n. M k k \ \ \ \i\n. M i i = \" proof (clarify, goal_cases) case (1 i) note A = this then have "i \ n \ set [] \ {0..n}" by auto with A(1) have "\ \ M i i" by fastforce with A(2) \i \ n\ show "M i i = \" by auto qed lemma fw_upd_canonical_id: "canonical M n \ i \ n \ j \ n \ k \ n \ fw_upd M k i j = M" -proof (auto simp: fw_upd_def upd_def less_eq[symmetric] min.coboundedI2, goal_cases) - case 1 - then have "M i j \ M i k + M k j" by auto - then have "min (M i j) (M i k + M k j) = M i j" by (simp split: split_min) - thus ?case by force -qed + by (auto simp: fw_upd_def upd_def) lemma fw_canonical_id: "canonical M n \ i \ n \ j \ n \ k \ n \ fw M n k i j = M" proof (induction k arbitrary: i j) case 0 then show ?case proof (induction i arbitrary: j) case 0 then show ?case proof (induction j) case 0 thus ?case by (auto intro: fw_upd_canonical_id) next case Suc thus ?case by (auto intro: fw_upd_canonical_id) qed next case Suc then show ?case proof (induction j) case 0 thus ?case by (auto intro: fw_upd_canonical_id) next case Suc thus ?case by (auto intro: fw_upd_canonical_id) qed qed next case Suc then show ?case proof (induction i arbitrary: j) case 0 then show ?case proof (induction j) case 0 thus ?case by (auto intro: fw_upd_canonical_id) next case Suc thus ?case by (auto intro: fw_upd_canonical_id) qed next case Suc then show ?case proof (induction j) case 0 thus ?case by (auto intro: fw_upd_canonical_id) next case Suc thus ?case by (auto intro: fw_upd_canonical_id) qed qed qed lemmas FW_canonical_id = fw_canonical_id[OF _ order.refl order.refl order.refl] section \Helper Lemmas for Bouyer's Theorem on Approximation\ lemma aux1: "i \ n \ j \ n \ set xs \ {0..n} \ (a,b) \ set (arcs i j xs) \ a \ n \ b \ n" by (induction xs arbitrary: i) auto lemma arcs_distinct1: "i \ set xs \ j \ set xs \ distinct xs \ xs \ [] \ (a,b) \ set (arcs i j xs) \ a \ b" apply (induction xs arbitrary: i) apply fastforce apply (rename_tac a' xs i) apply (case_tac xs) apply auto done lemma arcs_distinct2: "i \ set xs \ j \ set xs \ distinct xs \ i \ j \ (a,b) \ set (arcs i j xs) \ a \ b" by (induction xs arbitrary: i) auto lemma arcs_distinct3: "distinct (a # b # c # xs) \ (i,j) \ set (arcs a b xs) \ i \ c \ j \ c" by (induction xs arbitrary: a) force+ lemma arcs_elem: assumes "(a, b) \ set (arcs i j xs)" shows "a \ set (i # xs)" "b \ set (j # xs)" using assms by (induction xs arbitrary: i) auto lemma arcs_distinct_dest1: "distinct (i # a # zs) \ (b,c) \ set (arcs a j zs) \ b \ i" using arcs_elem by fastforce lemma arcs_distinct_fix: "distinct (a # x # xs @ [b]) \ (a,c) \ set (arcs a b (x # xs)) \ c = x" using arcs_elem(1) by fastforce lemma disjE3: "A \ B \ C \ (A \ G) \ (B \ G) \ (C \ G) \ G" by auto lemma arcs_predecessor: assumes "(a, b) \ set (arcs i j xs)" "a \ i" shows "\ c. (c, a) \ set (arcs i j xs)" using assms by (induction xs arbitrary: i) auto lemma arcs_successor: assumes "(a, b) \ set (arcs i j xs)" "b \ j" shows "\ c. (b,c) \ set (arcs i j xs)" using assms apply (induction xs arbitrary: i) apply simp apply (rename_tac aa xs i) apply (case_tac xs) by auto lemma arcs_predecessor': assumes "(a, b) \ set (arcs i j (x # xs))" "(a,b) \ (i, x)" shows "\ c. (c, a) \ set (arcs i j (x # xs))" using assms by (induction xs arbitrary: i x) auto lemma arcs_cases: assumes "(a, b) \ set (arcs i j xs)" "xs \ []" shows "(\ ys. xs = b # ys \ a = i) \ (\ ys. xs = ys @ [a] \ b = j) \ (\ c d ys. (a,b) \ set (arcs c d ys) \ xs = c # ys @ [d])" using assms proof (induction xs arbitrary: i) case Nil then show ?case by auto next case (Cons x xs) show ?case proof (cases "(a, b) = (i, x)") case True with Cons.prems show ?thesis by auto next case False note F = this show ?thesis proof (cases "xs = []") case True with F Cons.prems show ?thesis by auto next case False from F Cons.prems have "(a, b) \ set (arcs x j xs)" by auto from Cons.IH[OF this False] have "(\ys. xs = b # ys \ a = x) \ (\ys. xs = ys @ [a] \ b = j) \ (\c d ys. (a, b) \ set (arcs c d ys) \ xs = c # ys @ [d])" . then show ?thesis proof (rule disjE3, goal_cases) case 1 from 1 obtain ys where *: "xs = b # ys \ a = x" by auto show ?thesis proof (cases "ys = []") case True with * show ?thesis by auto next case False then obtain z zs where zs: "ys = zs @ [z]" by (metis append_butlast_last_id) with * show ?thesis by auto qed next case 2 then show ?case by auto next case 3 with False show ?case by auto qed qed qed qed lemma arcs_cases': assumes "(a, b) \ set (arcs i j xs)" "xs \ []" shows "(\ ys. xs = b # ys \ a = i) \ (\ ys. xs = ys @ [a] \ b = j) \ (\ ys zs. xs = ys @ a # b # zs)" using assms proof (induction xs arbitrary: i) case Nil then show ?case by auto next case (Cons x xs) show ?case proof (cases "(a, b) = (i, x)") case True with Cons.prems show ?thesis by auto next case False note F = this show ?thesis proof (cases "xs = []") case True with F Cons.prems show ?thesis by auto next case False from F Cons.prems have "(a, b) \ set (arcs x j xs)" by auto from Cons.IH[OF this False] have "(\ys. xs = b # ys \ a = x) \ (\ys. xs = ys @ [a] \ b = j) \ (\ys zs. xs = ys @ a # b # zs)" . then show ?thesis proof (rule disjE3, goal_cases) case 1 from 1 obtain ys where *: "xs = b # ys \ a = x" by auto show ?thesis proof (cases "ys = []") case True with * show ?thesis by auto next case False then obtain z zs where zs: "ys = zs @ [z]" by (metis append_butlast_last_id) with * show ?thesis by auto qed next case 2 then show ?case by auto next case 3 then obtain ys zs where "xs = ys @ a # b # zs" by auto then have "x # xs = (x # ys) @ a # b # zs" by auto then show ?thesis by blast qed qed qed qed lemma arcs_successor': assumes "(a, b) \ set (arcs i j xs)" "b \ j" shows "\ c. xs = [b] \ a = i \ (\ ys. xs = b # c # ys \ a = i) \ (\ ys. xs = ys @ [a,b] \ c = j) \ (\ ys zs. xs = ys @ a # b # c # zs)" using assms proof (induction xs arbitrary: i) case Nil then show ?case by auto next case (Cons x xs) show ?case proof (cases "(a, b) = (i, x)") case True with Cons.prems show ?thesis by (cases xs) auto next case False note F = this show ?thesis proof (cases "xs = []") case True with F Cons.prems show ?thesis by auto next case False from F Cons.prems have "(a, b) \ set (arcs x j xs)" by auto from Cons.IH[OF this \b \ j\] obtain c where c: "xs = [b] \ a = x \ (\ys. xs = b # c # ys \ a = x) \ (\ys. xs = ys @ [a, b] \ c = j) \ (\ys zs. xs = ys @ a # b # c # zs)" .. then show ?thesis proof (standard, goal_cases) case 1 with Cons.prems show ?case by auto next case 2 then show ?thesis proof (rule disjE3, goal_cases) case 1 from 1 obtain ys where *: "xs = b # ys \ a = x" by auto show ?thesis proof (cases "ys = []") case True with * show ?thesis by auto next case False then obtain z zs where zs: "ys = z # zs" by (cases ys) auto with * show ?thesis by fastforce qed next case 2 then show ?case by auto next case 3 then obtain ys zs where "xs = ys @ a # b # c # zs" by auto then have "x # xs = (x # ys) @ a # b # c # zs" by auto then show ?thesis by blast qed qed qed qed qed lemma list_last: "xs = [] \ (\ y ys. xs = ys @ [y])" by (induction xs) auto lemma arcs_predecessor'': assumes "(a, b) \ set (arcs i j xs)" "a \ i" shows "\ c. xs = [a] \ (\ ys. xs = a # b # ys) \ (\ ys. xs = ys @ [c,a] \ b = j) \ (\ ys zs. xs = ys @ c # a # b # zs)" using assms proof (induction xs arbitrary: i) case Nil then show ?case by auto next case (Cons x xs) show ?case proof (cases "(a, b) = (i, x)") case True with Cons.prems show ?thesis by (cases xs) auto next case False note F = this show ?thesis proof (cases "xs = []") case True with F Cons.prems show ?thesis by auto next case False from F Cons.prems have *: "(a, b) \ set (arcs x j xs)" by auto from False obtain y ys where xs: "xs = y # ys" by (cases xs) auto show ?thesis proof (cases "(a,b) = (x,y)") case True with * xs show ?thesis by auto next case False with * xs have **: "(a, b) \ set (arcs y j ys)" by auto show ?thesis proof (cases "ys = []") case True with ** xs show ?thesis by force next case False from arcs_cases'[OF ** this] obtain ws zs where ***: "ys = b # ws \ a = y \ ys = ws @ [a] \ b = j \ ys = ws @ a # b # zs" by auto then show ?thesis apply rule using xs apply blast apply safe using xs list_last[of ws] apply - apply (rotate_tac 3) apply (case_tac "ws = []") apply auto[] apply (subgoal_tac "\y ys. ws = ys @ [y]") apply fastforce apply simp apply (case_tac "ws = []") apply (subgoal_tac "x # xs = [x] @ y # a # b # zs") apply (rule exI[where x = y]) apply blast apply simp subgoal proof goal_cases case 1 then obtain u us where "ws = us @ [u]" by auto with 1(1,2) have "x # xs = (x # y # us) @ u # a # b # zs" by auto then show ?case by blast qed done qed qed qed qed qed lemma arcs_ex_middle: "\ b. (a, b) \ set (arcs i j (ys @ a # xs))" by (induction xs arbitrary: i ys) (auto simp: arcs_decomp) lemma arcs_ex_head: "\ b. (i, b) \ set (arcs i j xs)" by (cases xs) auto subsection \Successive\ fun successive where "successive _ [] = True" | "successive P [_] = True" | "successive P (x # y # xs) \ \ P y \ successive P xs \ \ P x \ successive P (y # xs)" lemma "\ successive (\ x. x > (0 :: nat)) [Suc 0, Suc 0]" by simp lemma "successive (\ x. x > (0 :: nat)) [Suc 0]" by simp lemma "successive (\ x. x > (0 :: nat)) [Suc 0, 0, Suc 0]" by simp lemma "\ successive (\ x. x > (0 :: nat)) [Suc 0, 0, Suc 0, Suc 0]" by simp lemma "\ successive (\ x. x > (0 :: nat)) [Suc 0, 0, 0, Suc 0, Suc 0]" by simp lemma "successive (\ x. x > (0 :: nat)) [Suc 0, 0, Suc 0, 0, Suc 0]" by simp lemma "\ successive (\ x. x > (0 :: nat)) [Suc 0, Suc 0, 0, Suc 0]" by simp lemma "successive (\ x. x > (0 :: nat)) [0, 0, Suc 0, 0]" by simp lemma successive_step: "successive P (x # xs) \ \ P x \ successive P xs" apply (cases xs) apply simp apply (rename_tac y ys) apply (case_tac ys) apply auto done lemma successive_step_2: "successive P (x # y # xs) \ \ P y \ successive P xs" apply (cases xs) apply simp apply (rename_tac z zs) apply (case_tac zs) apply auto done lemma successive_stepI: "successive P xs \ \ P x \ successive P (x # xs)" by (cases xs) auto theorem list_two_induct[case_names Nil Single Cons]: fixes P :: "'a list \ bool" and list :: "'a list" assumes Nil: "P []" assumes Single: "\ x. P [x]" and Cons: "\x1 x2 xs. P xs \ P (x2 # xs) \ P (x1 # x2 # xs)" shows "P xs" using assms apply (induction "length xs" arbitrary: xs rule: less_induct) apply (rename_tac xs) apply (case_tac xs) apply simp apply (rename_tac ys) apply (case_tac ys) apply simp apply (rename_tac zs) apply (case_tac zs) by auto lemma successive_end_1: "successive P xs \ \ P x \ successive P (xs @ [x])" by (induction _ xs rule: list_two_induct) auto lemma successive_ends_1: "successive P xs \ \ P x \ successive P ys \ successive P (xs @ x # ys)" by (induction _ xs rule: list_two_induct) (fastforce intro: successive_stepI)+ lemma successive_ends_1': "successive P xs \ \ P x \ P y \ \ P z \ successive P ys \ successive P (xs @ x # y # z # ys)" by (induction _ xs rule: list_two_induct) (fastforce intro: successive_stepI)+ lemma successive_end_2: "successive P xs \ \ P x \ successive P (xs @ [x,y])" by (induction _ xs rule: list_two_induct) auto lemma successive_end_2': "successive P (xs @ [x]) \ \ P x \ successive P (xs @ [x,y])" by (induction _ xs rule: list_two_induct) auto lemma successive_end_3: "successive P (xs @ [x]) \ \ P x \ P y \ \ P z \ successive P (xs @ [x,y,z])" by (induction _ xs rule: list_two_induct) auto lemma successive_step_rev: "successive P (xs @ [x]) \ \ P x \ successive P xs" by (induction _ xs rule: list_two_induct) auto lemma successive_glue: "successive P (zs @ [z]) \ successive P (x # xs) \ \ P z \ \ P x \ successive P (zs @ [z] @ x # xs)" proof goal_cases case A: 1 show ?thesis proof (cases "P x") case False with A(1,2) successive_ends_1 successive_step show ?thesis by fastforce next case True with A(1,3) successive_step_rev have "\ P z" "successive P zs" by fastforce+ with A(2) successive_ends_1 show ?thesis by fastforce qed qed lemma successive_glue': "successive P (zs @ [y]) \ \ P z \ successive P zs \ \ P y \ successive P (x # xs) \ \ P w \ successive P xs \ \ P x \ \ P z \ \ P w \ successive P (zs @ y # z # w # x # xs)" by (metis append_Cons append_Nil successive.simps(3) successive_ends_1 successive_glue successive_stepI) lemma successive_dest_head: "xs = w # x # ys \ successive P xs \ successive P (x # xs) \ \ P w \ successive P xs \ \ P x" by auto lemma successive_dest_tail: "xs = zs @ [y,z] \ successive P xs \ successive P (zs @ [y]) \ \ P z \ successive P zs \ \ P y" apply (induction _ xs arbitrary: zs rule: list_two_induct) apply simp+ apply (rename_tac zs) apply (case_tac zs) apply simp apply (rename_tac ws) apply (case_tac ws) apply force+ done lemma successive_split: "xs = ys @ zs \ successive P xs \ successive P ys \ successive P zs" apply (induction _ xs arbitrary: ys rule: list_two_induct) apply simp apply (rename_tac ys, case_tac ys) apply simp apply simp apply (rename_tac ys, case_tac ys) apply simp apply (rename_tac list, case_tac list) apply (auto intro: successive_stepI) done lemma successive_decomp: "xs = x # ys @ zs @ [z] \ successive P xs \ \ P x \ \ P z \ successive P (zs @ [z] @ (x # ys))" by (metis Cons_eq_appendI successive_glue successive_split) lemma successive_predecessor: assumes "(a, b) \ set (arcs i j xs)" "a \ i" "successive P (arcs i j xs)" "P (a,b)" "xs \ []" shows "\ c. (xs = [a] \ c = i \ (\ ys. xs = a # b # ys \ c = i) \ (\ ys. xs = ys @ [c,a] \ b = j) \ (\ ys zs. xs = ys @ c # a # b # zs)) \ \ P (c,a)" proof - from arcs_predecessor''[OF assms(1,2)] obtain c where c: "xs = [a] \ (\ys. xs = a # b # ys) \ (\ys. xs = ys @ [c, a] \ b = j) \ (\ys zs. xs = ys @ c # a # b # zs)" by auto then show ?thesis proof (safe, goal_cases) case 1 with assms have "arcs i j xs = [(i, a), (a, j)]" by auto with assms have "\ P (i, a)" by auto with 1 show ?case by simp next case 2 with assms have "\ P (i, a)" by fastforce with 2 show ?case by auto next case 3 with assms have "\ P (c, a)" using arcs_decomp successive_dest_tail by fastforce with 3 show ?case by auto next case 4 with assms(3,4) have "\ P (c, a)" using arcs_decomp successive_split by fastforce with 4 show ?case by auto qed qed thm arcs_successor' lemma successive_successor: assumes "(a, b) \ set (arcs i j xs)" "b \ j" "successive P (arcs i j xs)" "P (a,b)" "xs \ []" shows "\ c. (xs = [b] \ c = j \ (\ ys. xs = b # c # ys) \ (\ ys. xs = ys @ [a,b] \ c = j) \ (\ ys zs. xs = ys @ a # b # c # zs)) \ \ P (b,c)" proof - from arcs_successor'[OF assms(1,2)] obtain c where c: "xs = [b] \ a = i \ (\ys. xs = b # c # ys \ a = i) \ (\ys. xs = ys @ [a, b] \ c = j) \ (\ys zs. xs = ys @ a # b # c # zs)" .. then show ?thesis proof (safe, goal_cases) case 1 with assms(1,2) have "arcs i j xs = [(a,b), (b,j)]" by auto with assms have "\ P (b,j)" by auto with 1 show ?case by simp next case 2 with assms have "\ P (b, c)" by fastforce with 2 show ?case by auto next case 3 with assms have "\ P (b, c)" using arcs_decomp successive_dest_tail by fastforce with 3 show ?case by auto next case 4 with assms(3,4) have "\ P (b, c)" using arcs_decomp successive_split by fastforce with 4 show ?case by auto qed qed lemmas add_mono_right = add_mono[OF order_refl] lemmas add_mono_left = add_mono[OF _ order_refl] subsubsection \Obtaining successive and distinct paths\ lemma canonical_successive: fixes A B defines "M \ \ i j. min (A i j) (B i j)" assumes "canonical A n" assumes "set xs \ {0..n}" assumes "i \ n" "j \ n" shows "\ ys. len M i j ys \ len M i j xs \ set ys \ {0..n} \ successive (\ (a, b). M a b = A a b) (arcs i j ys)" using assms proof (induction xs arbitrary: i rule: list_two_induct) case Nil show ?case by fastforce next case 2: (Single x i) show ?case proof (cases "M i x = A i x \ M x j = A x j") case False then have "successive (\(a, b). M a b = A a b) (arcs i j [x])" by auto with 2 show ?thesis by blast next case True with 2 have "M i j \ M i x + M x j" unfolding min_def by fastforce with 2(3-) show ?thesis apply simp apply (rule exI[where x = "[]"]) by auto qed next case 3: (Cons x y xs i) show ?case proof (cases "M i y \ M i x + M x y", goal_cases) case 1 from 3 obtain ys where "len M i j ys \ len M i j (y # xs) \ set ys \ {0..n} \ successive (\a. case a of (a, b) \ M a b = A a b) (arcs i j ys)" by fastforce moreover from 1 have "len M i j (y # xs) \ len M i j (x # y # xs)" using add_mono by (auto simp: assoc[symmetric]) ultimately show ?case by force next case False { assume "M i x = A i x" "M x y = A x y" with 3(3-) have "A i y \ M i x + M x y" by auto then have "M i y \ M i x + M x y" unfolding M_def min_def by auto } note * = this with False have "M i x \ A i x \ M x y \ A x y" by auto then show ?thesis proof (standard, goal_cases) case 1 from 3 obtain ys where ys: "len M x j ys \ len M x j (y # xs)" "set ys \ {0..n}" "successive (\a. case a of (a, b) \ M a b = A a b) (arcs x j ys)" by force+ from 1 successive_stepI[OF ys(3), of "(i, x)"] have "successive (\a. case a of (a, b) \ M a b = A a b) (arcs i j (x # ys))" by auto moreover have "len M i j (x # ys) \ len M i j (x # y # xs)" using add_mono_right[OF ys(1)] by auto ultimately show ?case using 3(5) ys(2) by fastforce next case 2 from 3 obtain ys where ys: "len M y j ys \ len M y j xs" "set ys \ {0..n}" "successive (\a. case a of (a, b) \ M a b = A a b) (arcs y j ys)" by force+ from this(3) 2 have "successive (\a. case a of (a, b) \ M a b = A a b) (arcs i j (x # y # ys))" by simp moreover from add_mono_right[OF ys(1)] have "len M i j (x # y # ys) \ len M i j (x # y # xs)" by (auto simp: assoc[symmetric]) ultimately show ?thesis using ys(2) 3(5) by fastforce qed qed qed lemma canonical_successive_distinct: fixes A B defines "M \ \ i j. min (A i j) (B i j)" assumes "canonical A n" assumes "set xs \ {0..n}" assumes "i \ n" "j \ n" assumes "distinct xs" "i \ set xs" "j \ set xs" shows "\ ys. len M i j ys \ len M i j xs \ set ys \ set xs \ successive (\ (a, b). M a b = A a b) (arcs i j ys) \ distinct ys \ i \ set ys \ j \ set ys" using assms proof (induction xs arbitrary: i rule: list_two_induct) case Nil show ?case by fastforce next case 2: (Single x i) show ?case proof (cases "M i x = A i x \ M x j = A x j") case False then have "successive (\(a, b). M a b = A a b) (arcs i j [x])" by auto with 2 show ?thesis by blast next case True with 2 have "M i j \ M i x + M x j" unfolding min_def by fastforce with 2(3-) show ?thesis apply simp apply (rule exI[where x = "[]"]) by auto qed next case 3: (Cons x y xs i) show ?case proof (cases "M i y \ M i x + M x y") case 1: True from 3(2)[OF 3(3,4)] 3(5-10) obtain ys where ys: "len M i j ys \ len M i j (y # xs)" "set ys \ set (x # y # xs)" "successive (\a. case a of (a, b) \ M a b = A a b) (arcs i j ys)" "distinct ys \ i \ set ys \ j \ set ys" by fastforce moreover from 1 have "len M i j (y # xs) \ len M i j (x # y # xs)" using add_mono by (auto simp: assoc[symmetric]) ultimately have "len M i j ys \ len M i j (x # y # xs)" by auto then show ?thesis using ys(2-) by blast next case False { assume "M i x = A i x" "M x y = A x y" with 3(3-) have "A i y \ M i x + M x y" by auto then have "M i y \ M i x + M x y" unfolding M_def min_def by auto } note * = this with False have "M i x \ A i x \ M x y \ A x y" by auto then show ?thesis proof (standard, goal_cases) case 1 from 3(2)[OF 3(3,4)] 3(5-10) obtain ys where ys: "len M x j ys \ len M x j (y # xs)" "set ys \ set (y # xs)" "successive (\a. case a of (a, b) \ M a b = A a b) (arcs x j ys)" "distinct ys" "i \ set ys" "x \ set ys" "j \ set ys" by fastforce from 1 successive_stepI[OF ys(3), of "(i, x)"] have "successive (\a. case a of (a, b) \ M a b = A a b) (arcs i j (x # ys))" by auto moreover have "len M i j (x # ys) \ len M i j (x # y # xs)" using add_mono_right[OF ys(1)] by auto moreover have "distinct (x # ys)" "i \ set (x # ys)" "j \ set (x # ys)" using ys(4-) 3(8-) by auto moreover from ys(2) have "set (x # ys) \ set (x # y # xs)" by auto ultimately show ?case by fastforce next case 2 from 3(1)[OF 3(3,4)] 3(5-) obtain ys where ys: "len M y j ys \ len M y j xs" "set ys \ set xs" "successive (\a. case a of (a, b) \ M a b = A a b) (arcs y j ys)" "distinct ys" "j \ set ys" "y \ set ys" "i \ set ys" "x \ set ys" by fastforce from this(3) 2 have "successive (\a. case a of (a, b) \ M a b = A a b) (arcs i j (x # y # ys))" by simp moreover from add_mono_right[OF ys(1)] have "len M i j (x # y # ys) \ len M i j (x # y # xs)" by (auto simp: assoc[symmetric]) moreover have "distinct (x # y # ys)" "i \ set (x # y # ys)" "j \ set (x # y # ys)" using ys(4-) 3(8-) by auto ultimately show ?thesis using ys(2) by fastforce qed qed qed lemma successive_snd_last: "successive P (xs @ [x, y]) \ P y \ \ P x" by (induction _ xs rule: list_two_induct) auto lemma canonical_shorten_rotate_neg_cycle: fixes A B defines "M \ \ i j. min (A i j) (B i j)" assumes "canonical A n" assumes "set xs \ {0..n}" assumes "i \ n" assumes "len M i i xs < \" shows "\ j ys. len M j j ys < \ \ set (j # ys) \ set (i # xs) \ successive (\ (a, b). M a b = A a b) (arcs j j ys) \ distinct ys \ j \ set ys \ (ys \ [] \ M j (hd ys) \ A j (hd ys) \ M (last ys) j \ A (last ys) j)" using assms proof - note A = assms from negative_len_shortest[OF _ A(5)] obtain j ys where ys: "distinct (j # ys)" "len M j j ys < \" "j \ set (i # xs)" "set ys \ set xs" by blast from this(1,3) canonical_successive_distinct[OF A(2) subset_trans[OF this(4) A(3)], of j j B] A(3,4) obtain zs where zs: "len M j j zs \ len M j j ys" "set zs \ set ys" "successive (\(a, b). M a b = A a b) (arcs j j zs)" "distinct zs" "j \ set zs" by (force simp: M_def) show ?thesis proof (cases "zs = []") assume "zs \ []" then obtain w ws where ws: "zs = w # ws" by (cases zs) auto show ?thesis proof (cases "ws = []") case False then obtain u us where us: "ws = us @ [u]" by (induction ws) auto show ?thesis proof (cases "M j w = A j w \ M u j = A u j") case True have "u \ n" "j \ n" "w \ n" using us ws zs(2) ys(3,4) A(3,4) by auto with A(2) True have "M u w \ M u j + M j w" unfolding M_def min_def by fastforce then have "len M u u (w # us) \ len M j j zs" using ws us by (simp add: len_comp comm) (auto intro: add_mono simp: assoc[symmetric]) moreover have "set (u # w # us) \ set (i # xs)" using ws us zs(2) ys(3,4) by auto moreover have "distinct (w # us)" "u \ set (w # us)" using ws us zs(4) by auto moreover have "successive (\(a, b). M a b = A a b) (arcs u u (w # us))" proof (cases us) case Nil with zs(3) ws us True show ?thesis by auto next case (Cons v vs) with zs(3) ws us True have "M w v \ A w v" by auto with ws us Cons zs(3) True arcs_decomp_tail successive_split show ?thesis by (simp, blast) qed moreover have "M (last (w # us)) u \ A (last (w # us)) u" proof (cases "us = []") case T: True with zs(3) ws us True show ?thesis by auto next case False then obtain v vs where vs: "us = vs @ [v]" by (induction us) auto with ws us have "arcs j j zs = arcs j v (w # vs) @ [(v, u), (u,j)]" by (simp add: arcs_decomp) with zs(3) True have "M v u \ A v u" using successive_snd_last[of "\(a, b). M a b = A a b" "arcs j v (w # vs)"] by auto with vs show ?thesis by simp qed ultimately show ?thesis using zs(1) ys(2) by (intro exI[where x = u], intro exI[where x = "w # us"]) fastforce next case False with zs ws us ys show ?thesis by (intro exI[where x = j], intro exI[where x = "zs"]) auto qed next case True with True ws zs ys show ?thesis by (intro exI[where x = j], intro exI[where x = "zs"]) fastforce qed next case True with ys zs show ?thesis by (intro exI[where x = j], intro exI[where x = "zs"]) fastforce qed qed (* Generated by sledgehammer/z3 *) lemma successive_arcs_extend_last: "successive P (arcs i j xs) \ \ P (i, hd xs) \ \ P (last xs, j) \ xs \ [] \ successive P (arcs i j xs @ [(i, hd xs)])" proof - assume a1: "\ P (i, hd xs) \ \ P (last xs, j)" assume a2: "successive P (arcs i j xs)" assume a3: "xs \ []" then have f4: "\ P (last xs, j) \ successive P (arcs i (last xs) (butlast xs))" using a2 by (metis (no_types) append_butlast_last_id arcs_decomp_tail successive_step_rev) have f5: "arcs i j xs = arcs i (last xs) (butlast xs) @ [(last xs, j)]" using a3 by (metis (no_types) append_butlast_last_id arcs_decomp_tail) have "([] @ arcs i j xs @ [(i, hd xs)]) @ [(i, hd xs)] = arcs i j xs @ [(i, hd xs), (i, hd xs)]" by simp then have "P (last xs, j) \ successive P (arcs i j xs @ [(i, hd xs)])" using a2 a1 by (metis (no_types) self_append_conv2 successive_end_2 successive_step_rev) then show ?thesis using f5 f4 successive_end_2 by fastforce qed lemma cycle_rotate_arcs: fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "length xs > 1" "(i, j) \ arcs' xs" shows "\ ys zs. set (arcs a a xs) = set (arcs i i (j # ys @ a # zs)) \ xs = zs @ i # j # ys" using assms proof - assume A: "length xs > 1" "(i, j) \ arcs' xs" from arcs'_decomp[OF this] obtain ys zs where xs: "xs = zs @ i # j # ys" by blast with arcs_decomp[OF this, of a a] arcs_decomp[of "j # ys @ a # zs" "j # ys" a zs i i] show ?thesis by force qed lemma cycle_rotate_len_arcs_successive: fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "length xs > 1" "(i, j) \ arcs' xs" "successive P (arcs a a xs)" "\ P (a, hd xs) \ \ P (last xs, a)" shows "\ ys zs. len M a a xs = len M i i (j # ys @ a # zs) \ set (arcs a a xs) = set (arcs i i (j # ys @ a # zs)) \ xs = zs @ i # j # ys \ successive P (arcs i i (j # ys @ a # zs))" using assms proof - note A = assms from arcs'_decomp[OF A(1,2)] obtain ys zs where xs: "xs = zs @ i # j # ys" by blast note arcs1 = arcs_decomp[OF xs, of a a] note arcs2 = arcs_decomp[of "j # ys @ a # zs" "j # ys" a zs i i] have *:"successive P (arcs i i (j # ys @ a # zs))" proof (cases "ys = []") case True show ?thesis proof (cases zs) case Nil with A(3,4) xs True show ?thesis by auto next case (Cons z zs') with True arcs2 A(3,4) xs show ?thesis apply simp by (metis arcs.simps(1,2) arcs1 successive.simps(3) successive_split successive_step) qed next case False then obtain y ys' where ys: "ys = ys' @ [y]" by (metis append_butlast_last_id) show ?thesis proof (cases zs) case Nil with A(3,4) xs ys have "\ P (a, i) \ \ P (y, a)" "successive P (arcs a a (i # j # ys' @ [y]))" by simp+ from successive_decomp[OF _ this(2,1)] show ?thesis using ys Nil arcs_decomp by fastforce next case (Cons z zs') with A(3,4) xs ys have "\ P (a, z) \ \ P (y, a)" "successive P (arcs a a (z # zs' @ i # j # ys' @ [y]))" by simp+ from successive_decomp[OF _ this(2,1)] show ?thesis using ys Cons arcs_decomp by fastforce qed qed from len_decomp[OF xs, of M a a] have "len M a a xs = len M a i zs + len M i a (j # ys)" . also have "\ = len M i a (j # ys) + len M a i zs" by (simp add: comm) also from len_comp[of M i i "j # ys" a zs] have "\ = len M i i (j # ys @ a # zs)" by auto finally show ?thesis using * xs arcs_decomp[OF xs, of a a] arcs_decomp[of "j # ys @ a # zs" "j # ys" a zs i i] by force qed lemma successive_successors: "xs = ys @ a # b # c # zs \ successive P (arcs i j xs) \ \ P (a,b) \ \ P (b, c)" apply (induction _ xs arbitrary: i ys rule: list_two_induct) apply fastforce apply fastforce apply (rename_tac ys, case_tac ys) apply fastforce apply (rename_tac list, case_tac list) apply fastforce+ done lemma successive_successors': "xs = ys @ a # b # zs \ successive P xs \ \ P a \ \ P b" using successive_split by fastforce lemma cycle_rotate_len_arcs_successive': fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "length xs > 1" "(i, j) \ arcs' xs" "successive P (arcs a a xs)" "\ P (a, hd xs) \ \ P (last xs, a)" shows "\ ys zs. len M a a xs = len M i i (j # ys @ a # zs) \ set (arcs a a xs) = set (arcs i i (j # ys @ a # zs)) \ xs = zs @ i # j # ys \ successive P (arcs i i (j # ys @ a # zs) @ [(i,j)])" using assms proof - note A = assms from arcs'_decomp[OF A(1,2)] obtain ys zs where xs: "xs = zs @ i # j # ys" by blast note arcs1 = arcs_decomp[OF xs, of a a] note arcs2 = arcs_decomp[of "j # ys @ a # zs" "j # ys" a zs i i] have *:"successive P (arcs i i (j # ys @ a # zs) @ [(i,j)])" proof (cases "ys = []") case True show ?thesis proof (cases zs) case Nil with A(3,4) xs True show ?thesis by auto next case (Cons z zs') with True arcs2 A(3,4) xs show ?thesis apply simp apply (cases "P (a, z)") apply (simp add: arcs_decomp) apply (simp only: append_Cons[symmetric]) using successive_split[of "((a, z) # arcs z i zs') @ [(i, j), (j, a)]" _ "[(j, a)]" P] apply auto[] subgoal (* Generated by sledgehammer *) proof simp assume a1: "successive P ((a, z) # arcs z a (zs' @ [i, j]))" assume a2: "\ P (a, z)" assume a3: "zs = z # zs'" assume a4: "ys = []" assume a5: "xs = z # zs' @ [i, j]" have f6: "\p pa ps. \ successive p ((pa::nat \ nat) # ps) \ p pa \ successive p ps" by (meson successive_step) have "(a, z) # arcs z i zs' @ (i, j) # arcs j a [] = arcs a a xs" using a4 a3 by (simp add: arcs1) then have "arcs z a (zs' @ [i, j]) = arcs z i zs' @ [(i, j), (j, a)]" using a5 by simp then show "\ P (j, a) \ successive P ((a, z) # arcs z i zs' @ [(i, j)]) \ \ P (i, j) \ (successive P (arcs z i zs' @ [(i, j)]) \ \ P (j, a) \ successive P ((a, z) # arcs z i zs' @ [(i, j)]))" using f6 a2 a1 by (metis successive.simps(1) successive_dest_tail successive_ends_1 successive_stepI) qed done qed next case False then obtain y ys' where ys: "ys = ys' @ [y]" by (metis append_butlast_last_id) show ?thesis proof (cases zs) case Nil with A(3,4) xs ys have *: "\ P (a, i) \ \ P (y, a)" "successive P (arcs a a (i # j # ys' @ [y]))" by simp+ from successive_decomp[OF _ this(2,1)] ys Nil arcs_decomp have "successive P (arcs i i (j # ys @ a # zs))" by fastforce moreover from * have "\ P (a, i) \ \ P (i, j)" by auto ultimately show ?thesis by (metis append_Cons last_snoc list.distinct(1) list.sel(1) Nil successive_arcs_extend_last) next case (Cons z zs') with A(3,4) xs ys have *: "\ P (a, z) \ \ P (y, a)" "successive P (arcs a a (z # zs' @ i # j # ys' @ [y]))" by simp_all from successive_decomp[OF _ this(2,1)] ys Cons arcs_decomp have **: "successive P (arcs i i (j # ys @ a # zs))" by fastforce from Cons have "zs \ []" by auto then obtain w ws where ws: "zs = ws @ [w]" by (induction zs) auto with A(3,4) xs ys have *: "successive P (arcs a a (ws @ [w] @ i # j # ys' @ [y]))" by simp from successive_successors[OF _ *] have "\ P (w, i) \ \ P (i, j)" by auto with * show ?thesis by (metis ** append_is_Nil_conv last.simps last_append list.distinct(2) list.sel(1) successive_arcs_extend_last ws) qed qed from len_decomp[OF xs, of M a a] have "len M a a xs = len M a i zs + len M i a (j # ys)" . also have "\ = len M i a (j # ys) + len M a i zs" by (simp add: comm) also from len_comp[of M i i "j # ys" a zs] have "\ = len M i i (j # ys @ a # zs)" by auto finally show ?thesis using * xs arcs_decomp[OF xs, of a a] arcs_decomp[of "j # ys @ a # zs" "j # ys" a zs i i] by force qed lemma cycle_rotate_3: fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "xs \ []" "(i, j) \ set (arcs a a xs)" "successive P (arcs a a xs)" "\ P (a, hd xs) \ \ P (last xs, a)" shows "\ ys. len M a a xs = len M i i (j # ys) \ set (i # j # ys) = set (a # xs) \ 1 + length ys = length xs \ set (arcs a a xs) = set (arcs i i (j # ys)) \ successive P (arcs i i (j # ys))" proof - note A = assms { fix ys assume A:"a = i" "xs = j # ys" with assms(3) have ?thesis by auto } note * = this have **: ?thesis if A: "a = j" "xs = ys @ [i]" for ys using A proof (safe, goal_cases) case 1 have "len M j j (ys @ [i]) = M i j + len M j i ys" using len_decomp[of "ys @ [i]" ys i "[]" M j j] by (auto simp: comm) moreover have "arcs j j (ys @ [i]) = arcs j i ys @ [(i, j)]" using arcs_decomp_tail by auto moreover with assms(3,4) A have "successive P ((i,j) # arcs j i ys)" apply simp apply (case_tac ys) apply simp apply simp by (metis arcs.simps(2) calculation(2) 1(1) successive_split successive_step) ultimately show ?case by auto qed { assume "length xs = 1" then obtain b where xs: "xs = [b]" by (metis One_nat_def length_0_conv length_Suc_conv) with A(2) have "a = i \ b = j \ a = j \ b = i" by auto then have ?thesis using * ** xs by auto } note *** = this show ?thesis proof (cases "length xs = 0") case True with A show ?thesis by auto next case False thus ?thesis proof (cases "length xs = 1", goal_cases) case True with *** show ?thesis by auto next case 2 hence "length xs > 1" by linarith then obtain b c ys where ys:"xs = b # ys @ [c]" by (metis One_nat_def assms(1) 2(2) length_0_conv length_Cons list.exhaust rev_exhaust) thus ?thesis proof (cases "(i,j) = (a,b)") case True with ys * show ?thesis by blast next case False then show ?thesis proof (cases "(i,j) = (c,a)", goal_cases) case True with ys ** show ?thesis by force next case 2 with A(2) ys have "(i, j) \ arcs' xs" using cycle_rotate_2_aux by (auto simp add: arcs'_def) (* slow *) from cycle_rotate_len_arcs_successive[OF \length xs > 1\ this A(3,4), of M] show ?thesis by auto qed qed qed qed qed lemma cycle_rotate_3': fixes M :: "('a :: linordered_ab_monoid_add) mat" assumes "xs \ []" "(i, j) \ set (arcs a a xs)" "successive P (arcs a a xs)" "\ P (a, hd xs) \ \ P (last xs, a)" shows "\ ys. len M a a xs = len M i i (j # ys) \ set (i # j # ys) = set (a # xs) \ 1 + length ys = length xs \ set (arcs a a xs) = set (arcs i i (j # ys)) \ successive P (arcs i i (j # ys) @ [(i, j)])" proof - note A = assms have *: ?thesis if "a = i" "xs = j # ys" for ys using that assms(3) successive_arcs_extend_last[OF assms(3,4)] by auto have **: ?thesis if A:"a = j" "xs = ys @ [i]" for ys using A proof (auto, goal_cases) case 1 have "len M j j (ys @ [i]) = M i j + len M j i ys" using len_decomp[of "ys @ [i]" ys i "[]" M j j] by (auto simp: comm) moreover have "arcs j j (ys @ [i]) = arcs j i ys @ [(i, j)]" using arcs_decomp_tail by auto moreover with assms(3,4) A have "successive P ((i,j) # arcs j i ys @ [(i, j)])" apply simp apply (case_tac ys) apply simp apply simp by (metis successive_step) ultimately show ?case by auto qed { assume "length xs = 1" then obtain b where xs: "xs = [b]" by (metis One_nat_def length_0_conv length_Suc_conv) with A(2) have "a = i \ b = j \ a = j \ b = i" by auto then have ?thesis using * ** xs by auto } note *** = this show ?thesis proof (cases "length xs = 0") case True with A show ?thesis by auto next case False thus ?thesis proof (cases "length xs = 1", goal_cases) case True with *** show ?thesis by auto next case 2 hence "length xs > 1" by linarith then obtain b c ys where ys:"xs = b # ys @ [c]" by (metis One_nat_def assms(1) 2(2) length_0_conv length_Cons list.exhaust rev_exhaust) thus ?thesis proof (cases "(i,j) = (a,b)") case True with ys * show ?thesis by blast next case False then show ?thesis proof (cases "(i,j) = (c,a)", goal_cases) case True with ys ** show ?thesis by force next case 2 with A(2) ys have "(i, j) \ arcs' xs" using cycle_rotate_2_aux by (auto simp add: arcs'_def) (* slow *) from cycle_rotate_len_arcs_successive'[OF \length xs > 1\ this A(3,4), of M] show ?thesis by auto qed qed qed qed qed end diff --git a/thys/Universal_Turing_Machine/Recursive.thy b/thys/Universal_Turing_Machine/Recursive.thy --- a/thys/Universal_Turing_Machine/Recursive.thy +++ b/thys/Universal_Turing_Machine/Recursive.thy @@ -1,2830 +1,2843 @@ (* Title: thys/Recursive.thy Author: Jian Xu, Xingyuan Zhang, and Christian Urban Modifications: Sebastiaan Joosten *) theory Recursive imports Abacus Rec_Def Abacus_Hoare begin fun addition :: "nat \ nat \ nat \ abc_prog" where "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, Inc m, Goto 4]" fun mv_box :: "nat \ nat \ abc_prog" where "mv_box m n = [Dec m 3, Inc n, Goto 0]" text \The compilation of \z\-operator.\ definition rec_ci_z :: "abc_inst list" where "rec_ci_z \ [Goto 1]" text \The compilation of \s\-operator.\ definition rec_ci_s :: "abc_inst list" where "rec_ci_s \ (addition 0 1 2 [+] [Inc 1])" text \The compilation of \id i j\-operator\ fun rec_ci_id :: "nat \ nat \ abc_inst list" where "rec_ci_id i j = addition j i (i + 1)" fun mv_boxes :: "nat \ nat \ nat \ abc_inst list" where "mv_boxes ab bb 0 = []" | "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] mv_box (ab + n) (bb + n)" fun empty_boxes :: "nat \ abc_inst list" where "empty_boxes 0 = []" | "empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]" fun cn_merge_gs :: "(abc_inst list \ nat \ nat) list \ nat \ abc_inst list" where "cn_merge_gs [] p = []" | "cn_merge_gs (g # gs) p = (let (gprog, gpara, gn) = g in gprog [+] mv_box gpara p [+] cn_merge_gs gs (Suc p))" text \ The compiler of recursive functions, where \rec_ci recf\ return \(ap, arity, fp)\, where \ap\ is the Abacus program, \arity\ is the arity of the recursive function \recf\, \fp\ is the amount of memory which is going to be used by \ap\ for its execution. \ fun rec_ci :: "recf \ abc_inst list \ nat \ nat" where "rec_ci z = (rec_ci_z, 1, 2)" | "rec_ci s = (rec_ci_s, 1, 3)" | "rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" | "rec_ci (Cn n f gs) = (let cied_gs = map (\ g. rec_ci g) gs in let (fprog, fpara, fn) = rec_ci f in let pstr = Max (set (Suc n # fn # (map (\ (aprog, p, n). n) cied_gs))) in let qstr = pstr + Suc (length gs) in (cn_merge_gs cied_gs pstr [+] mv_boxes 0 qstr n [+] mv_boxes pstr 0 (length gs) [+] fprog [+] mv_box fpara pstr [+] empty_boxes (length gs) [+] mv_box pstr n [+] mv_boxes qstr 0 n, n, qstr + n))" | "rec_ci (Pr n f g) = (let (fprog, fpara, fn) = rec_ci f in let (gprog, gpara, gn) = rec_ci g in let p = Max (set ([n + 3, fn, gn])) in let e = length gprog + 7 in (mv_box n p [+] fprog [+] mv_box n (Suc n) [+] (([Dec p e] [+] gprog [+] [Inc n, Dec (Suc n) 3, Goto 1]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]), Suc n, p + 1))" | "rec_ci (Mn n f) = (let (fprog, fpara, fn) = rec_ci f in let len = length (fprog) in (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3), Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn))" declare rec_ci.simps [simp del] rec_ci_s_def[simp del] rec_ci_z_def[simp del] rec_ci_id.simps[simp del] mv_boxes.simps[simp del] mv_box.simps[simp del] addition.simps[simp del] declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] abc_step_l.simps[simp del] inductive_cases terminate_pr_reverse: "terminate (Pr n f g) xs" inductive_cases terminate_z_reverse[elim!]: "terminate z xs" inductive_cases terminate_s_reverse[elim!]: "terminate s xs" inductive_cases terminate_id_reverse[elim!]: "terminate (id m n) xs" inductive_cases terminate_cn_reverse[elim!]: "terminate (Cn n f gs) xs" inductive_cases terminate_mn_reverse[elim!]:"terminate (Mn n f) xs" fun addition_inv :: "nat \ nat list \ nat \ nat \ nat \ nat list \ bool" where "addition_inv (as, lm') m n p lm = (let sn = lm ! n in let sm = lm ! m in lm ! p = 0 \ (if as = 0 then \ x. x \ lm ! m \ lm' = lm[m := x, n := (sn + sm - x), p := (sm - x)] else if as = 1 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm - x - 1), p := (sm - x - 1)] else if as = 2 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm - x), p := (sm - x - 1)] else if as = 3 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm - x), p := (sm - x)] else if as = 4 then \ x. x \ lm ! m \ lm' = lm[m := x, n := (sn + sm), p := (sm - x)] else if as = 5 then \ x. x < lm ! m \ lm' = lm[m := x, n := (sn + sm), p := (sm - x - 1)] else if as = 6 then \ x. x < lm ! m \ lm' = lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)] else if as = 7 then lm' = lm[m := sm, n := (sn + sm)] else False))" fun addition_stage1 :: "nat \ nat list \ nat \ nat \ nat" where "addition_stage1 (as, lm) m p = (if as = 0 \ as = 1 \ as = 2 \ as = 3 then 2 else if as = 4 \ as = 5 \ as = 6 then 1 else 0)" fun addition_stage2 :: "nat \ nat list \ nat \ nat \ nat" where "addition_stage2 (as, lm) m p = (if 0 \ as \ as \ 3 then lm ! m else if 4 \ as \ as \ 6 then lm ! p else 0)" fun addition_stage3 :: "nat \ nat list \ nat \ nat \ nat" where "addition_stage3 (as, lm) m p = (if as = 1 then 4 else if as = 2 then 3 else if as = 3 then 2 else if as = 0 then 1 else if as = 5 then 2 else if as = 6 then 1 else if as = 4 then 0 else 0)" fun addition_measure :: "((nat \ nat list) \ nat \ nat) \ (nat \ nat \ nat)" where "addition_measure ((as, lm), m, p) = (addition_stage1 (as, lm) m p, addition_stage2 (as, lm) m p, addition_stage3 (as, lm) m p)" definition addition_LE :: "(((nat \ nat list) \ nat \ nat) \ ((nat \ nat list) \ nat \ nat)) set" where "addition_LE \ (inv_image lex_triple addition_measure)" lemma wf_additon_LE[simp]: "wf addition_LE" by(auto simp: addition_LE_def lex_triple_def lex_pair_def) declare addition_inv.simps[simp del] lemma update_zero_to_zero[simp]: "\am ! n = (0::nat); n < length am\ \ am[n := 0] = am" apply(simp add: list_update_same_conv) done lemma addition_inv_init: "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ addition_inv (0, lm) m n p lm" apply(simp add: addition_inv.simps Let_def ) apply(rule_tac x = "lm ! m" in exI, simp) done lemma abs_fetch[simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)" "abc_fetch (Suc 0) (addition m n p) = Some (Inc n)" "abc_fetch 2 (addition m n p) = Some (Inc p)" "abc_fetch 3 (addition m n p) = Some (Goto 0)" "abc_fetch 4 (addition m n p) = Some (Dec p 7)" "abc_fetch 5 (addition m n p) = Some (Inc m)" "abc_fetch 6 (addition m n p) = Some (Goto 4)" by(simp_all add: abc_fetch.simps addition.simps) lemma exists_small_list_elem1[simp]: "\m \ n; p < length lm; lm ! p = 0; m < p; n < p; x \ lm ! m; 0 < x\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xam \ n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\ \ \xa\lm ! m. lm[m := x, n := lm ! n + lm ! m - x, p := lm ! m - x] = lm[m := xa, n := lm ! n + lm ! m - xa, p := lm ! m - xa]" apply(rule_tac x = x in exI, simp) done lemma exists_small_list_elem5[simp]: "\m \ n; p < length lm; lm ! p = 0; m < p; n < p; x \ lm ! m; lm ! m \ x\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xam \ n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\ \ \xa\lm ! m. lm[m := Suc x, n := lm ! n + lm ! m, p := lm ! m - Suc x] = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]" apply(rule_tac x = "Suc x" in exI, simp) done lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm" apply(cases asm, simp add: abc_steps_l.simps) done lemma list_double_update_2: "lm[a := x, b := y, a := z] = lm[b := y, a:=z]" by (metis list_update_overwrite list_update_swap) declare Let_def[simp] lemma addition_halt_lemma: "\m \ n; max m n < p; length lm > p\ \ \na. \ (\(as, lm') (m, p). as = 7) (abc_steps_l (0, lm) (addition m n p) na) (m, p) \ addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm \ addition_inv (abc_steps_l (0, lm) (addition m n p) (Suc na)) m n p lm \ ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), abc_steps_l (0, lm) (addition m n p) na, m, p) \ addition_LE" proof - assume assms:"m\n" "max m n < p" "length lm > p" { fix na obtain a b where ab:"abc_steps_l (0, lm) (addition m n p) na = (a, b)" by force assume assms2: "\ (\(as, lm') (m, p). as = 7) (abc_steps_l (0, lm) (addition m n p) na) (m, p)" "addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm" have r1:"addition_inv (abc_steps_l (0, lm) (addition m n p) (Suc na)) m n p lm" using assms(1-3) assms2 unfolding abc_step_red2 ab abc_step_l.simps abc_lm_v.simps abc_lm_s.simps addition_inv.simps by (auto split:if_splits simp add: addition_inv.simps Suc_diff_Suc) have r2:"((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), abc_steps_l (0, lm) (addition m n p) na, m, p) \ addition_LE" using assms(1-3) assms2 unfolding abc_step_red2 ab apply(auto split:if_splits simp add: addition_inv.simps abc_steps_zero) by(auto simp add: addition_LE_def lex_triple_def lex_pair_def abc_step_l.simps abc_lm_v.simps abc_lm_s.simps split: if_splits) note r1 r2 } thus ?thesis by auto qed lemma addition_correct': "\m \ n; max m n < p; length lm > p; lm ! p = 0\ \ \ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" apply(insert halt_lemma2[of addition_LE "\ ((as, lm'), m, p). addition_inv (as, lm') m n p lm" "\ stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)" "\ ((as, lm'), m, p). as = 7"], simp add: abc_steps_zero addition_inv_init) apply(drule_tac addition_halt_lemma,force,force) apply (simp,erule_tac exE) apply(rename_tac na) apply(rule_tac x = na in exI) apply(auto) done lemma length_addition[simp]: "length (addition a b c) = 7" by(auto simp: addition.simps) lemma addition_correct: assumes "m \ n" "max m n < p" "length lm > p" "lm ! p = 0" shows "{\ a. a = lm} (addition m n p) {\ nl. addition_inv (7, nl) m n p lm}" using assms proof(rule_tac abc_Hoare_haltI, simp) fix lma assume "m \ n" "m < p \ n < p" "p < length lm" "lm ! p = 0" then have "\ stp. (\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" by(rule_tac addition_correct', auto simp: addition_inv.simps) then obtain stp where "(\ (as, lm'). as = 7 \ addition_inv (as, lm') m n p lm) (abc_steps_l (0, lm) (addition m n p) stp)" using exE by presburger thus "\na. abc_final (abc_steps_l (0, lm) (addition m n p) na) (addition m n p) \ (\nl. addition_inv (7, nl) m n p lm) abc_holds_for abc_steps_l (0, lm) (addition m n p) na" by(auto intro:exI[of _ stp]) qed lemma compile_s_correct': "{\nl. nl = n # 0 \ 2 @ anything} addition 0 (Suc 0) 2 [+] [Inc (Suc 0)] {\nl. nl = n # Suc n # 0 # anything}" proof(rule_tac abc_Hoare_plus_halt) show "{\nl. nl = n # 0 \ 2 @ anything} addition 0 (Suc 0) 2 {\ nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \ 2 @ anything)}" by(rule_tac addition_correct, auto simp: numeral_2_eq_2) next show "{\nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \ 2 @ anything)} [Inc (Suc 0)] {\nl. nl = n # Suc n # 0 # anything}" by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps) qed declare rec_exec.simps[simp del] lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)" apply(auto simp: abc_comp.simps abc_shift.simps) apply(rename_tac x) apply(case_tac x, auto) done lemma compile_z_correct: "\rec_ci z = (ap, arity, fp); rec_exec z [n] = r\ \ {\nl. nl = n # 0 \ (fp - arity) @ anything} ap {\nl. nl = n # r # 0 \ (fp - Suc arity) @ anything}" apply(rule_tac abc_Hoare_haltI) apply(rule_tac x = 1 in exI) apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_exec.simps) done lemma compile_s_correct: "\rec_ci s = (ap, arity, fp); rec_exec s [n] = r\ \ {\nl. nl = n # 0 \ (fp - arity) @ anything} ap {\nl. nl = n # r # 0 \ (fp - Suc arity) @ anything}" apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_exec.simps) done lemma compile_id_correct': assumes "n < length args" shows "{\nl. nl = args @ 0 \ 2 @ anything} addition n (length args) (Suc (length args)) {\nl. nl = args @ rec_exec (recf.id (length args) n) args # 0 # anything}" proof - have "{\nl. nl = args @ 0 \ 2 @ anything} addition n (length args) (Suc (length args)) {\nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \ 2 @ anything)}" using assms by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append) thus "?thesis" using assms by(simp add: addition_inv.simps rec_exec.simps nth_append numeral_2_eq_2 list_update_append) qed lemma compile_id_correct: "\n < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_exec (recf.id m n) xs = r\ \ {\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ r # 0 \ (fp - Suc arity) @ anything}" apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct') done lemma cn_merge_gs_tl_app: "cn_merge_gs (gs @ [g]) pstr = cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)" apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, auto) apply(simp add: abc_comp_commute) done lemma footprint_ge: "rec_ci a = (p, arity, fp) \ arity < fp" proof(induct a) case (Cn x1 a x3) then show ?case by(cases "rec_ci a", auto simp:rec_ci.simps) next case (Pr x1 a1 a2) then show ?case by(cases "rec_ci a1";cases "rec_ci a2", auto simp:rec_ci.simps) next case (Mn x1 a) then show ?case by(cases "rec_ci a", auto simp:rec_ci.simps) qed (auto simp: rec_ci.simps) lemma param_pattern: "\terminate f xs; rec_ci f = (p, arity, fp)\ \ length xs = arity" proof(induct arbitrary: p arity fp rule: terminate.induct) case (termi_cn f xs gs n) thus ?case by(cases "rec_ci f", (auto simp: rec_ci.simps)) next case (termi_pr x g xs n f) thus ?case by (cases "rec_ci f", cases "rec_ci g", auto simp: rec_ci.simps) next case (termi_mn xs n f r) thus ?case by (cases "rec_ci f", auto simp: rec_ci.simps) qed (auto simp: rec_ci.simps) lemma replicate_merge_anywhere: "x\a @ x\b @ ys = x\(a+b) @ ys" by(simp add:replicate_add) fun mv_box_inv :: "nat \ nat list \ nat \ nat \ nat list \ bool" where "mv_box_inv (as, lm) m n initlm = (let plus = initlm ! m + initlm ! n in length initlm > max m n \ m \ n \ (if as = 0 then \ k l. lm = initlm[m := k, n := l] \ k + l = plus \ k \ initlm ! m else if as = 1 then \ k l. lm = initlm[m := k, n := l] \ k + l + 1 = plus \ k < initlm ! m else if as = 2 then \ k l. lm = initlm[m := k, n := l] \ k + l = plus \ k \ initlm ! m else if as = 3 then lm = initlm[m := 0, n := plus] else False))" fun mv_box_stage1 :: "nat \ nat list \ nat \ nat" where "mv_box_stage1 (as, lm) m = (if as = 3 then 0 else 1)" fun mv_box_stage2 :: "nat \ nat list \ nat \ nat" where "mv_box_stage2 (as, lm) m = (lm ! m)" fun mv_box_stage3 :: "nat \ nat list \ nat \ nat" where "mv_box_stage3 (as, lm) m = (if as = 1 then 3 else if as = 2 then 2 else if as = 0 then 1 else 0)" fun mv_box_measure :: "((nat \ nat list) \ nat) \ (nat \ nat \ nat)" where "mv_box_measure ((as, lm), m) = (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m, mv_box_stage3 (as, lm) m)" definition lex_pair :: "((nat \ nat) \ nat \ nat) set" where "lex_pair = less_than <*lex*> less_than" definition lex_triple :: "((nat \ (nat \ nat)) \ (nat \ (nat \ nat))) set" where "lex_triple \ less_than <*lex*> lex_pair" definition mv_box_LE :: "(((nat \ nat list) \ nat) \ ((nat \ nat list) \ nat)) set" where "mv_box_LE \ (inv_image lex_triple mv_box_measure)" lemma wf_lex_triple: "wf lex_triple" by (auto simp:lex_triple_def lex_pair_def) lemma wf_mv_box_le[intro]: "wf mv_box_LE" by(auto intro:wf_lex_triple simp: mv_box_LE_def) declare mv_box_inv.simps[simp del] lemma mv_box_inv_init: "\m < length initlm; n < length initlm; m \ n\ \ mv_box_inv (0, initlm) m n initlm" apply(simp add: abc_steps_l.simps mv_box_inv.simps) apply(rule_tac x = "initlm ! m" in exI, rule_tac x = "initlm ! n" in exI, simp) done lemma abc_fetch[simp]: "abc_fetch 0 (mv_box m n) = Some (Dec m 3)" "abc_fetch (Suc 0) (mv_box m n) = Some (Inc n)" "abc_fetch 2 (mv_box m n) = Some (Goto 0)" "abc_fetch 3 (mv_box m n) = None" apply(simp_all add: mv_box.simps abc_fetch.simps) done lemma replicate_Suc_iff_anywhere: "x # x\b @ ys = x\(Suc b) @ ys" by simp lemma exists_smaller_in_list0[simp]: "\m \ n; m < length initlm; n < length initlm; k + l = initlm ! m + initlm ! n; k \ initlm ! m; 0 < k\ \ \ka la. initlm[m := k, n := l, m := k - Suc 0] = initlm[m := ka, n := la] \ Suc (ka + la) = initlm ! m + initlm ! n \ ka < initlm ! m" apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, auto) apply(subgoal_tac "initlm[m := k, n := l, m := k - Suc 0] = initlm[n := l, m := k, m := k - Suc 0]",force intro:list_update_swap) by(simp add: list_update_swap) lemma exists_smaller_in_list1[simp]: "\m \ n; m < length initlm; n < length initlm; Suc (k + l) = initlm ! m + initlm ! n; k < initlm ! m\ \ \ka la. initlm[m := k, n := l, n := Suc l] = initlm[m := ka, n := la] \ ka + la = initlm ! m + initlm ! n \ ka \ initlm ! m" apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto) done lemma abc_steps_prop[simp]: "\length initlm > max m n; m \ n\ \ \ (\(as, lm) m. as = 3) (abc_steps_l (0, initlm) (mv_box m n) na) m \ mv_box_inv (abc_steps_l (0, initlm) (mv_box m n) na) m n initlm \ mv_box_inv (abc_steps_l (0, initlm) (mv_box m n) (Suc na)) m n initlm \ ((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m), abc_steps_l (0, initlm) (mv_box m n) na, m) \ mv_box_LE" apply(rule impI, simp add: abc_step_red2) apply(cases "(abc_steps_l (0, initlm) (mv_box m n) na)", simp) apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps) apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def abc_step_l.simps abc_steps_l.simps mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps split: if_splits ) apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp) done lemma mv_box_inv_halt: "\length initlm > max m n; m \ n\ \ \ stp. (\ (as, lm). as = 3 \ mv_box_inv (as, lm) m n initlm) (abc_steps_l (0::nat, initlm) (mv_box m n) stp)" apply(insert halt_lemma2[of mv_box_LE "\ ((as, lm), m). mv_box_inv (as, lm) m n initlm" "\ stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)" "\ ((as, lm), m). as = (3::nat)" ]) apply(insert wf_mv_box_le) apply(simp add: mv_box_inv_init abc_steps_zero) apply(erule_tac exE) by (metis (no_types, lifting) case_prodE' case_prodI) lemma mv_box_halt_cond: "\m \ n; mv_box_inv (a, b) m n lm; a = 3\ \ b = lm[n := lm ! m + lm ! n, m := 0]" apply(simp add: mv_box_inv.simps, auto) apply(simp add: list_update_swap) done lemma mv_box_correct': "\length lm > max m n; m \ n\ \ \ stp. abc_steps_l (0::nat, lm) (mv_box m n) stp = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])" by(drule mv_box_inv_halt, auto dest:mv_box_halt_cond) lemma length_mvbox[simp]: "length (mv_box m n) = 3" by(simp add: mv_box.simps) lemma mv_box_correct: "\length lm > max m n; m\n\ \ {\ nl. nl = lm} mv_box m n {\ nl. nl = lm[n := (lm ! m + lm ! n), m:=0]}" apply(drule_tac mv_box_correct', simp) apply(auto simp: abc_Hoare_halt_def) by (metis abc_final.simps abc_holds_for.simps length_mvbox) declare list_update.simps(2)[simp del] lemma zero_case_rec_exec[simp]: "\length xs < gf; gf \ ft; n < length gs\ \ (rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) [ft + n - length xs := rec_exec (gs ! n) xs, 0 := 0] = 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything" using list_update_append[of "rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs)" "0 \ (length gs - n) @ 0 # 0 \ length xs @ anything" "ft + n - length xs" "rec_exec (gs ! n) xs"] apply(auto) apply(cases "length gs - n", simp, simp add: list_update.simps replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc) apply(simp add: list_update.simps) done lemma compile_cn_gs_correct': assumes g_cond: "\g\set (take n gs). terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" shows "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0\(length gs - n) @ 0 \ Suc (length xs) @ anything}" using g_cond proof(induct n) case 0 have "ft > length xs" using ft by simp thus "?case" apply(rule_tac abc_Hoare_haltI) apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps replicate_add[THEN sym] replicate_Suc[THEN sym] del: replicate_Suc) done next case (Suc n) have ind': "\g\set (take n gs). terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc})) \ {\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything}" by fact have g_newcond: "\g\set (take (Suc n) gs). terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" by fact from g_newcond have ind: "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything}" apply(rule_tac ind', rule_tac ballI, erule_tac x = g in ballE, simp_all add: take_Suc) by(cases "n < length gs", simp add:take_Suc_conv_app_nth, simp) show "?case" proof(cases "n < length gs") case True have h: "n < length gs" by fact thus "?thesis" - proof(simp add: take_Suc_conv_app_nth cn_merge_gs_tl_app) + proof (simp add: take_Suc_conv_app_nth cn_merge_gs_tl_app) obtain gp ga gf where a: "rec_ci (gs!n) = (gp, ga, gf)" by (metis prod_cases3) moreover have "min (length gs) n = n" using h by simp moreover have "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft [+] (gp [+] mv_box ga (ft + n)) {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" proof(rule_tac abc_Hoare_plus_halt) show "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything}" using ind by simp next have x: "gs!n \ set (take (Suc n) gs)" using h by(simp add: take_Suc_conv_app_nth) have b: "terminate (gs!n) xs" using a g_newcond h x by(erule_tac x = "gs!n" in ballE, simp_all) hence c: "length xs = ga" using a param_pattern by metis have d: "gf > ga" using footprint_ge a by simp have e: "ft \ gf" using ft a h Max_ge image_eqI by(simp, rule_tac max.coboundedI2, rule_tac Max_ge, simp, rule_tac insertI2, rule_tac f = "(\(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp, rule_tac x = "gs!n" in image_eqI, simp, simp) show "{\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything} gp [+] mv_box ga (ft + n) {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" proof(rule_tac abc_Hoare_plus_halt) show "{\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything} gp {\nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything}" proof - have "({\nl. nl = xs @ 0 \ (gf - ga) @ 0\(ft - gf)@map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything} gp {\nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \ (gf - Suc ga) @ 0\(ft - gf)@map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 \ Suc (length xs) @ anything})" using a g_newcond h x apply(erule_tac x = "gs!n" in ballE) apply(simp, simp) done thus "?thesis" using a b c d e by(simp add: replicate_merge_anywhere) qed next show "{\nl. nl = xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything} mv_box ga (ft + n) {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" proof - have "{\nl. nl = xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything} mv_box ga (ft + n) {\nl. nl = (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! ga + (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! (ft + n), ga := 0]}" using a c d e h apply(rule_tac mv_box_correct) - apply(simp, arith, arith) + apply simp_all done moreover have "(xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! ga + (xs @ rec_exec (gs ! n) xs # 0 \ (ft - Suc (length xs)) @ map (\i. rec_exec i xs) (take n gs) @ 0 \ (length gs - n) @ 0 # 0 \ length xs @ anything) ! (ft + n), ga := 0]= xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything" using a c d e h by(simp add: list_update_append nth_append length_replicate split: if_splits del: list_update.simps(2), auto) ultimately show "?thesis" by(simp) qed qed qed ultimately show - "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} - cn_merge_gs (map rec_ci (take n gs)) ft [+] (case rec_ci (gs ! n) of (gprog, gpara, gn) \ - gprog [+] mv_box gpara (ft + min (length gs) n)) - {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \ (length gs - Suc n) @ 0 # 0 \ length xs @ anything}" + "{\nl. nl = + xs @ + 0 # 0 \ (ft + length gs) @ anything} + cn_merge_gs (map rec_ci (take n gs)) ft [+] + (case rec_ci (gs ! n) of + (gprog, gpara, gn) \ + gprog [+] mv_box gpara (ft + n)) + {\nl. nl = + xs @ + 0 \ (ft - length xs) @ + map (\i. rec_exec i xs) (take n gs) @ + rec_exec (gs ! n) xs # + 0 \ (length gs - Suc n) @ + 0 # 0 \ length xs @ anything}" by simp qed next case False have h: "\ n < length gs" by fact hence ind': "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci gs) ft {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything}" using ind by simp thus "?thesis" using h by(simp) qed qed lemma compile_cn_gs_correct: assumes g_cond: "\g\set gs. terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" shows "{\nl. nl = xs @ 0 # 0 \ (ft + length gs) @ anything} cn_merge_gs (map rec_ci gs) ft {\nl. nl = xs @ 0 \ (ft - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything}" using assms using compile_cn_gs_correct'[of "length gs" gs xs ft ffp anything ] apply(auto) done lemma length_mvboxes[simp]: "length (mv_boxes aa ba n) = 3*n" by(induct n, auto simp: mv_boxes.simps) lemma exp_suc: "a\Suc b = a\b @ [a]" by(simp add: exp_ind del: replicate.simps) lemma last_0[simp]: "\Suc n \ ba - aa; length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ \ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)" proof - assume h: "Suc n \ ba - aa" and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)" from h and g have k: "ba - aa = Suc (length lm3 + n)" by arith from k show "(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0" apply(simp, insert g) apply(simp add: nth_append) done qed lemma butlast_last[simp]: "length lm1 = aa \ (lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2" apply(simp add: nth_append) done lemma arith_as_simp[simp]: "\Suc n \ ba - aa; aa < ba\ \ (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False" apply arith done lemma butlast_elem[simp]: "\Suc n \ ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ \ (lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0" using nth_append[of "lm1 @ (0::'a)\n @ last lm2 # lm3 @ butlast lm2" "(0::'a) # lm4" "ba + n"] apply(simp) done lemma update_butlast_eq0[simp]: "\Suc n \ ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\ \ (lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4) [ba + n := last lm2, aa + n := 0] = lm1 @ 0 # 0\n @ lm3 @ lm2 @ lm4" using list_update_append[of "lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2" "0 # lm4" "ba + n" "last lm2"] apply(simp add: list_update_append list_update.simps(2-) replicate_Suc_iff_anywhere exp_suc del: replicate_Suc) apply(cases lm2, simp, simp) done lemma update_butlast_eq1[simp]: "\Suc (length lm1 + n) \ ba; length lm2 = Suc n; length lm3 = ba - Suc (length lm1 + n); \ ba - Suc (length lm1) < ba - Suc (length lm1 + n); \ ba + n - length lm1 < n\ \ (0::nat) \ n @ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4)[ba - length lm1 := last lm2, 0 := 0] = 0 # 0 \ n @ lm3 @ lm2 @ lm4" apply(subgoal_tac "ba - length lm1 = Suc n + length lm3", simp add: list_update.simps(2-) list_update_append) apply(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc) apply(cases lm2, simp, simp) apply(auto) done lemma mv_boxes_correct: "\aa + n \ ba; ba > aa; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\ \ {\ nl. nl = lm1 @ lm2 @ lm3 @ 0\n @ lm4} (mv_boxes aa ba n) {\ nl. nl = lm1 @ 0\n @ lm3 @ lm2 @ lm4}" proof(induct n arbitrary: lm2 lm3 lm4) case 0 thus "?case" by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) next case (Suc n) have ind: "\lm2 lm3 lm4. \aa + n \ ba; aa < ba; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\ \ {\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ n @ lm4} mv_boxes aa ba n {\nl. nl = lm1 @ 0 \ n @ lm3 @ lm2 @ lm4}" by fact have h1: "aa + Suc n \ ba" by fact have h2: "aa < ba" by fact have h3: "length lm1 = aa" by fact have h4: "length lm2 = Suc n" by fact have h5: "length lm3 = ba - aa - Suc n" by fact have "{\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) {\nl. nl = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4}" proof(rule_tac abc_Hoare_plus_halt) have "{\nl. nl = lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \ n @ (0 # lm4)} mv_boxes aa ba n {\ nl. nl = lm1 @ 0\n @ (last lm2 # lm3) @ butlast lm2 @ (0 # lm4)}" using h1 h2 h3 h4 h5 by(rule_tac ind, simp_all) moreover have " lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \ n @ (0 # lm4) = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4" using h4 by(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc, cases lm2, simp_all) ultimately show "{\nl. nl = lm1 @ lm2 @ lm3 @ 0 \ Suc n @ lm4} mv_boxes aa ba n {\ nl. nl = lm1 @ 0\n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4}" by (metis append_Cons) next let ?lm = "lm1 @ 0 \ n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4" have "{\nl. nl = ?lm} mv_box (aa + n) (ba + n) {\ nl. nl = ?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]}" using h1 h2 h3 h4 h5 by(rule_tac mv_box_correct, simp_all) moreover have "?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0] = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4" using h1 h2 h3 h4 h5 by(auto simp: nth_append list_update_append split: if_splits) ultimately show "{\nl. nl = lm1 @ 0 \ n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4} mv_box (aa + n) (ba + n) {\nl. nl = lm1 @ 0 \ Suc n @ lm3 @ lm2 @ lm4}" by simp qed thus "?case" by(simp add: mv_boxes.simps) qed lemma update_butlast_eq2[simp]: "\Suc n \ aa - length lm1; length lm1 < aa; length lm2 = aa - Suc (length lm1 + n); length lm3 = Suc n; \ aa - Suc (length lm1) < aa - Suc (length lm1 + n); \ aa + n - length lm1 < n\ \ butlast lm3 @ ((0::nat) # lm2 @ 0 \ n @ last lm3 # lm4)[0 := last lm3, aa - length lm1 := 0] = lm3 @ lm2 @ 0 # 0 \ n @ lm4" apply(subgoal_tac "aa - length lm1 = length lm2 + Suc n") apply(simp add: list_update.simps list_update_append) apply(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc) apply(cases lm3, simp, simp) apply(auto) done lemma mv_boxes_correct2: "\n \ aa - ba; ba < aa; length (lm1::nat list) = ba; length (lm2::nat list) = aa - ba - n; length (lm3::nat list) = n\ \{\ nl. nl = lm1 @ 0\n @ lm2 @ lm3 @ lm4} (mv_boxes aa ba n) {\ nl. nl = lm1 @ lm3 @ lm2 @ 0\n @ lm4}" proof(induct n arbitrary: lm2 lm3 lm4) case 0 thus "?case" by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) next case (Suc n) have ind: "\lm2 lm3 lm4. \n \ aa - ba; ba < aa; length lm1 = ba; length lm2 = aa - ba - n; length lm3 = n\ \ {\nl. nl = lm1 @ 0 \ n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n {\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ n @ lm4}" by fact have h1: "Suc n \ aa - ba" by fact have h2: "ba < aa" by fact have h3: "length lm1 = ba" by fact have h4: "length lm2 = aa - ba - Suc n" by fact have h5: "length lm3 = Suc n" by fact have "{\nl. nl = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) {\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4}" proof(rule_tac abc_Hoare_plus_halt) have "{\ nl. nl = lm1 @ 0 \ n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4)} mv_boxes aa ba n {\ nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\n @ (last lm3 # lm4)}" using h1 h2 h3 h4 h5 by(rule_tac ind, simp_all) moreover have "lm1 @ 0 \ n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4) = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4" using h5 by(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc, cases lm3, simp_all) ultimately show "{\nl. nl = lm1 @ 0 \ Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n {\ nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\n @ (last lm3 # lm4)}" by metis next thm mv_box_correct let ?lm = "lm1 @ butlast lm3 @ (0 # lm2) @ 0 \ n @ last lm3 # lm4" have "{\nl. nl = ?lm} mv_box (aa + n) (ba + n) {\nl. nl = ?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]}" using h1 h2 h3 h4 h5 by(rule_tac mv_box_correct, simp_all) moreover have "?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0] = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4" using h1 h2 h3 h4 h5 by(auto simp: nth_append list_update_append split: if_splits) ultimately show "{\nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0 \ n @ last lm3 # lm4} mv_box (aa + n) (ba + n) {\nl. nl = lm1 @ lm3 @ lm2 @ 0 \ Suc n @ lm4}" by simp qed thus "?case" by(simp add: mv_boxes.simps) qed lemma save_paras: "{\nl. nl = xs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ map (\i. rec_exec i xs) gs @ 0 \ Suc (length xs) @ anything} mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) {\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything}" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" have "{\nl. nl = [] @ xs @ (0\(?ft - length xs) @ map (\i. rec_exec i xs) gs @ [0]) @ 0 \ (length xs) @ anything} mv_boxes 0 (Suc ?ft + length gs) (length xs) {\nl. nl = [] @ 0 \ (length xs) @ (0\(?ft - length xs) @ map (\i. rec_exec i xs) gs @ [0]) @ xs @ anything}" by(rule_tac mv_boxes_correct, auto) thus "?thesis" by(simp add: replicate_merge_anywhere) qed lemma length_le_max_insert_rec_ci[intro]: "length gs \ ffp \ length gs \ max x1 (Max (insert ffp (x2 ` x3 ` set gs)))" apply(rule_tac max.coboundedI2) apply(simp add: Max_ge_iff) done lemma restore_new_paras: "ffp \ length gs \ {\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything} mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) {\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ 0 # xs @ anything}" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" assume j: "ffp \ length gs" hence "{\ nl. nl = [] @ 0\length gs @ 0\(?ft - length gs) @ map (\i. rec_exec i xs) gs @ ((0 # xs) @ anything)} mv_boxes ?ft 0 (length gs) {\ nl. nl = [] @ map (\i. rec_exec i xs) gs @ 0\(?ft - length gs) @ 0\length gs @ ((0 # xs) @ anything)}" by(rule_tac mv_boxes_correct2, auto) moreover have "?ft \ length gs" using j by(auto) ultimately show "?thesis" using j by(simp add: replicate_merge_anywhere le_add_diff_inverse) qed lemma le_max_insert[intro]: "ffp \ max x0 (Max (insert ffp (x1 ` x2 ` set gs)))" by (rule max.coboundedI2) auto declare max_less_iff_conj[simp del] lemma save_rs: "\far = length gs; ffp \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))); far < ffp\ \ {\nl. nl = map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything} mv_box far (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) {\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything}" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" thm mv_box_correct let ?lm= " map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ ?ft @ xs @ anything" assume h: "far = length gs" "ffp \ ?ft" "far < ffp" hence "{\ nl. nl = ?lm} mv_box far ?ft {\ nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]}" apply(rule_tac mv_box_correct) by( auto) moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0] = map (\i. rec_exec i xs) gs @ 0 \ (?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" using h apply(simp add: nth_append) using list_update_length[of "map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - Suc (length gs))" 0 "0 \ length gs @ xs @ anything" "rec_exec (Cn (length xs) f gs) xs"] apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc) by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc) ultimately show "?thesis" by(simp) qed lemma length_empty_boxes[simp]: "length (empty_boxes n) = 2*n" apply(induct n, simp, simp) done lemma empty_one_box_correct: "{\nl. nl = 0 \ n @ x # lm} [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ lm}" proof(induct x) case 0 thus "?case" by(simp add: abc_Hoare_halt_def, rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps replicate_Suc[THEN sym] exp_suc del: replicate_Suc) next case (Suc x) have "{\nl. nl = 0 \ n @ x # lm} [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ lm}" by fact then obtain stp where "abc_steps_l (0, 0 \ n @ x # lm) [Dec n 2, Goto 0] stp = (Suc (Suc 0), 0 # 0 \ n @ lm)" apply(auto simp: abc_Hoare_halt_def) by (smt abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3)) moreover have "abc_steps_l (0, 0\n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0)) = (0, 0 \ n @ x # lm)" by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps list_update.simps list_update_append) ultimately have "abc_steps_l (0, 0\n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0) + stp) = (Suc (Suc 0), 0 # 0\n @ lm)" by(simp only: abc_steps_add) thus "?case" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = "Suc (Suc stp)" in exI, simp) done qed lemma empty_boxes_correct: "length lm \ n \ {\ nl. nl = lm} empty_boxes n {\ nl. nl = 0\n @ drop n lm}" proof(induct n) case 0 thus "?case" by(simp add: empty_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps) next case (Suc n) have ind: "n \ length lm \ {\nl. nl = lm} empty_boxes n {\nl. nl = 0 \ n @ drop n lm}" by fact have h: "Suc n \ length lm" by fact have "{\nl. nl = lm} empty_boxes n [+] [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ drop (Suc n) lm}" proof(rule_tac abc_Hoare_plus_halt) show "{\nl. nl = lm} empty_boxes n {\nl. nl = 0 \ n @ drop n lm}" using h by(rule_tac ind, simp) next show "{\nl. nl = 0 \ n @ drop n lm} [Dec n 2, Goto 0] {\nl. nl = 0 # 0 \ n @ drop (Suc n) lm}" using empty_one_box_correct[of n "lm ! n" "drop (Suc n) lm"] using h by(simp add: Cons_nth_drop_Suc) qed thus "?case" by(simp add: empty_boxes.simps) qed lemma insert_dominated[simp]: "length gs \ ffp \ length gs + (max xs (Max (insert ffp (x1 ` x2 ` set gs))) - length gs) = max xs (Max (insert ffp (x1 ` x2 ` set gs)))" apply(rule_tac le_add_diff_inverse) apply(rule_tac max.coboundedI2) apply(simp add: Max_ge_iff) done lemma clean_paras: "ffp \ length gs \ {\nl. nl = map (\i. rec_exec i xs) gs @ 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything} empty_boxes (length gs) {\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything}" proof- let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" assume h: "length gs \ ffp" let ?lm = "map (\i. rec_exec i xs) gs @ 0 \ (?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" have "{\ nl. nl = ?lm} empty_boxes (length gs) {\ nl. nl = 0\length gs @ drop (length gs) ?lm}" by(rule_tac empty_boxes_correct, simp) moreover have "0\length gs @ drop (length gs) ?lm = 0 \ ?ft @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" using h by(simp add: replicate_merge_anywhere) ultimately show "?thesis" by metis qed lemma restore_rs: "{\nl. nl = 0 \ max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything} mv_box (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) {\nl. nl = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @ 0 \ length gs @ xs @ anything}" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" let ?lm = "0\(length xs) @ 0\(?ft - (length xs)) @ rec_exec (Cn (length xs) f gs) xs # 0 \ length gs @ xs @ anything" thm mv_box_correct have "{\ nl. nl = ?lm} mv_box ?ft (length xs) {\ nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]}" by(rule_tac mv_box_correct, simp, simp) moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0] = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - (length xs)) @ 0 \ length gs @ xs @ anything" apply(auto simp: list_update_append nth_append) (* ~ 7 sec *) apply(cases ?ft, simp_all add: Suc_diff_le list_update.simps) apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc) done ultimately show "?thesis" by(simp add: replicate_merge_anywhere) qed lemma restore_orgin_paras: "{\nl. nl = 0 \ length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ 0 \ length gs @ xs @ anything} mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs) {\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" thm mv_boxes_correct2 have "{\ nl. nl = [] @ 0\(length xs) @ (rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - length xs) @ 0 \ length gs) @ xs @ anything} mv_boxes (Suc ?ft + length gs) 0 (length xs) {\ nl. nl = [] @ xs @ (rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft - length xs) @ 0 \ length gs) @ 0\length xs @ anything}" by(rule_tac mv_boxes_correct2, auto) thus "?thesis" by(simp add: replicate_merge_anywhere) qed lemma compile_cn_correct': assumes f_ind: "\ anything r. rec_exec f (map (\g. rec_exec g xs) gs) = rec_exec (Cn (length xs) f gs) xs \ {\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ anything} fap {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (ffp - Suc far) @ anything}" and compile: "rec_ci f = (fap, far, ffp)" and term_f: "terminate f (map (\g. rec_exec g xs) gs)" and g_cond: "\g\set gs. terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" shows "{\nl. nl = xs @ 0 # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything} cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) [+] (mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+] (mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+] (fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) [+] (empty_boxes (length gs) [+] (mv_box (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+] mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs))))))) {\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" proof - let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\(aprog, p, n). n) ` rec_ci ` set gs)))" let ?A = "cn_merge_gs (map rec_ci gs) ?ft" let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)" let ?C = "mv_boxes ?ft 0 (length gs)" let ?D = fap let ?E = "mv_box far ?ft" let ?F = "empty_boxes (length gs)" let ?G = "mv_box ?ft (length xs)" let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)" let ?P1 = "\nl. nl = xs @ 0 # 0 \ (?ft + length gs) @ anything" let ?S = "\nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (?ft + length gs) @ anything" let ?Q1 = "\ nl. nl = xs @ 0\(?ft - length xs) @ map (\ i. rec_exec i xs) gs @ 0\(Suc (length xs)) @ anything" show "{?P1} (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) {?S}" proof(rule_tac abc_Hoare_plus_halt) show "{?P1} ?A {?Q1}" using g_cond by(rule_tac compile_cn_gs_correct, auto) next let ?Q2 = "\nl. nl = 0 \ ?ft @ map (\i. rec_exec i xs) gs @ 0 # xs @ anything" show "{?Q1} (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) {?S}" proof(rule_tac abc_Hoare_plus_halt) show "{?Q1} ?B {?Q2}" by(rule_tac save_paras) next let ?Q3 = "\ nl. nl = map (\i. rec_exec i xs) gs @ 0\?ft @ 0 # xs @ anything" show "{?Q2} (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) {?S}" proof(rule_tac abc_Hoare_plus_halt) have "ffp \ length gs" using compile term_f apply(subgoal_tac "length gs = far") apply(drule_tac footprint_ge, simp) by(drule_tac param_pattern, auto) thus "{?Q2} ?C {?Q3}" by(erule_tac restore_new_paras) next let ?Q4 = "\ nl. nl = map (\i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0\?ft @ xs @ anything" have a: "far = length gs" using compile term_f by(drule_tac param_pattern, auto) have b:"?ft \ ffp" by auto have c: "ffp > far" using compile by(erule_tac footprint_ge) show "{?Q3} (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) {?S}" proof(rule_tac abc_Hoare_plus_halt) have "{\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ 0\(?ft - ffp + far) @ 0 # xs @ anything} fap {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \ (ffp - Suc far) @ 0\(?ft - ffp + far) @ 0 # xs @ anything}" by(rule_tac f_ind, simp add: rec_exec.simps) thus "{?Q3} fap {?Q4}" using a b c by(simp add: replicate_merge_anywhere, cases "?ft", simp_all add: exp_suc del: replicate_Suc) next let ?Q5 = "\nl. nl = map (\i. rec_exec i xs) gs @ 0\(?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0\(length gs)@ xs @ anything" show "{?Q4} (?E [+] (?F [+] (?G [+] ?H))) {?S}" proof(rule_tac abc_Hoare_plus_halt) from a b c show "{?Q4} ?E {?Q5}" by(erule_tac save_rs, simp_all) next let ?Q6 = "\nl. nl = 0\?ft @ rec_exec (Cn (length xs) f gs) xs # 0\(length gs)@ xs @ anything" show "{?Q5} (?F [+] (?G [+] ?H)) {?S}" proof(rule_tac abc_Hoare_plus_halt) have "length gs \ ffp" using a b c by simp thus "{?Q5} ?F {?Q6}" by(erule_tac clean_paras) next let ?Q7 = "\nl. nl = 0\length xs @ rec_exec (Cn (length xs) f gs) xs # 0\(?ft - (length xs)) @ 0\(length gs)@ xs @ anything" show "{?Q6} (?G [+] ?H) {?S}" proof(rule_tac abc_Hoare_plus_halt) show "{?Q6} ?G {?Q7}" by(rule_tac restore_rs) next show "{?Q7} ?H {?S}" by(rule_tac restore_orgin_paras) qed qed qed qed qed qed qed qed lemma compile_cn_correct: assumes termi_f: "terminate f (map (\g. rec_exec g xs) gs)" and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ {\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec f (map (\g. rec_exec g xs) gs) # 0 \ (fp - Suc arity) @ anything}" and g_cond: "\g\set gs. terminate g xs \ (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ rec_exec g xs # 0 \ (xb - Suc xa) @ xc}))" and compile: "rec_ci (Cn n f gs) = (ap, arity, fp)" and len: "length xs = n" shows "{\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \ (fp - Suc arity) @ anything}" proof(cases "rec_ci f") fix fap far ffp assume h: "rec_ci f = (fap, far, ffp)" then have f_newind: "\ anything .{\nl. nl = map (\g. rec_exec g xs) gs @ 0 \ (ffp - far) @ anything} fap {\nl. nl = map (\g. rec_exec g xs) gs @ rec_exec f (map (\g. rec_exec g xs) gs) # 0 \ (ffp - Suc far) @ anything}" by(rule_tac f_ind, simp_all) thus "{\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \ (fp - Suc arity) @ anything}" using compile len h termi_f g_cond apply(auto simp: rec_ci.simps abc_comp_commute) apply(rule_tac compile_cn_correct', simp_all) done qed lemma mv_box_correct_simp[simp]: "\length xs = n; ft = max (n+3) (max fft gft)\ \ {\nl. nl = xs @ 0 # 0 \ (ft - n) @ anything} mv_box n ft {\nl. nl = xs @ 0 # 0 \ (ft - n) @ anything}" using mv_box_correct[of n ft "xs @ 0 # 0 \ (ft - n) @ anything"] by(auto) lemma length_under_max[simp]: "length xs < max (length xs + 3) fft" by auto lemma save_init_rs: "\length xs = n; ft = max (n+3) (max fft gft)\ \ {\nl. nl = xs @ rec_exec f xs # 0 \ (ft - n) @ anything} mv_box n (Suc n) {\nl. nl = xs @ 0 # rec_exec f xs # 0 \ (ft - Suc n) @ anything}" using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (ft - n) @ anything"] apply(auto simp: list_update_append list_update.simps nth_append split: if_splits) apply(cases "(max (length xs + 3) (max fft gft))", simp_all add: list_update.simps Suc_diff_le) done lemma less_then_max_plus2[simp]: "n + (2::nat) < max (n + 3) x" by auto lemma less_then_max_plus3[simp]: "n < max (n + (3::nat)) x" by auto lemma mv_box_max_plus_3_correct[simp]: "length xs = n \ {\nl. nl = xs @ x # 0 \ (max (n + (3::nat)) (max fft gft) - n) @ anything} mv_box n (max (n + 3) (max fft gft)) {\nl. nl = xs @ 0 \ (max (n + 3) (max fft gft) - n) @ x # anything}" proof - assume h: "length xs = n" let ?ft = "max (n+3) (max fft gft)" let ?lm = "xs @ x # 0\(?ft - Suc n) @ 0 # anything" have g: "?ft > n + 2" by simp thm mv_box_correct have a: "{\ nl. nl = ?lm} mv_box n ?ft {\ nl. nl = ?lm[?ft := ?lm!n + ?lm!?ft, n := 0]}" using h by(rule_tac mv_box_correct, auto) have b:"?lm = xs @ x # 0 \ (max (n + 3) (max fft gft) - n) @ anything" by(cases ?ft, simp_all add: Suc_diff_le exp_suc del: replicate_Suc) have c: "?lm[?ft := ?lm!n + ?lm!?ft, n := 0] = xs @ 0 \ (max (n + 3) (max fft gft) - n) @ x # anything" using h g apply(auto simp: nth_append list_update_append split: if_splits) using list_update_append[of "x # 0 \ (max (length xs + 3) (max fft gft) - Suc (length xs))" "0 # anything" "max (length xs + 3) (max fft gft) - length xs" "x"] apply(auto simp: if_splits) apply(simp add: list_update.simps replicate_Suc[THEN sym] del: replicate_Suc) done from a c show "?thesis" using h apply(simp) using b by simp qed lemma max_less_suc_suc[simp]: "max n (Suc n) < Suc (Suc (max (n + 3) x + anything - Suc 0))" by arith lemma suc_less_plus_3[simp]: "Suc n < max (n + 3) x" by arith lemma mv_box_ok_suc_simp[simp]: "length xs = n \ {\nl. nl = xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything} mv_box n (Suc n) {\nl. nl = xs @ 0 # rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ x # anything}" using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything"] apply(simp add: nth_append list_update_append list_update.simps) apply(cases "max (n + 3) (max fft gft)", simp_all) apply(cases "max (n + 3) (max fft gft) - 1", simp_all add: Suc_diff_le list_update.simps(2)) done lemma abc_append_frist_steps_eq_pre: assumes notfinal: "abc_notfinal (abc_steps_l (0, lm) A n) A" and notnull: "A \ []" shows "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" using notfinal proof(induct n) case 0 thus "?case" by(simp add: abc_steps_l.simps) next case (Suc n) have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \ abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" by fact have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A" by(simp add: notfinal_Suc) then have b: "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n" using ind by simp obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')" by (metis prod.exhaust) then have d: "s < length A \ abc_steps_l (0, lm) (A @ B) n = (s, lm')" using a b by simp thus "?case" using c by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append) qed lemma abc_append_first_step_eq_pre: "st < length A \ abc_step_l (st, lm) (abc_fetch st (A @ B)) = abc_step_l (st, lm) (abc_fetch st A)" by(simp add: abc_step_l.simps abc_fetch.simps nth_append) lemma abc_append_frist_steps_halt_eq': assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" and notnull: "A \ []" shows "\ n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')" proof - have "\ n'. abc_notfinal (abc_steps_l (0, lm) A n') A \ abc_final (abc_steps_l (0, lm) A (Suc n')) A" using assms by(rule_tac n = n in abc_before_final, simp_all) then obtain na where a: "abc_notfinal (abc_steps_l (0, lm) A na) A \ abc_final (abc_steps_l (0, lm) A (Suc na)) A" .. obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)" by (metis prod.exhaust) then have c: "abc_steps_l (0, lm) (A @ B) na = (sa, lma)" using a abc_append_frist_steps_eq_pre[of lm A na B] assms by simp have d: "sa < length A" using b a by simp then have e: "abc_step_l (sa, lma) (abc_fetch sa (A @ B)) = abc_step_l (sa, lma) (abc_fetch sa A)" by(rule_tac abc_append_first_step_eq_pre) from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')" using final equal_when_halt by(cases "abc_steps_l (0, lm) A (Suc na)" , simp) then have "abc_steps_l (0, lm) (A @ B) (Suc na) = (length A, lm')" using a b c e by(simp add: abc_step_red2) thus "?thesis" by blast qed lemma abc_append_frist_steps_halt_eq: assumes final: "abc_steps_l (0, lm) A n = (length A, lm')" shows "\ n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')" using final apply(cases "A = []") apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null) apply(rule_tac abc_append_frist_steps_halt_eq', simp_all) done lemma suc_suc_max_simp[simp]: "Suc (Suc (max (xs + 3) fft - Suc (Suc ( xs)))) = max ( xs + 3) fft - ( xs)" by arith lemma contract_dec_ft_length_plus_7[simp]: "\ft = max (n + 3) (max fft gft); length xs = n\ \ {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything} [Dec ft (length gap + 7)] {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything}" - apply(simp add: abc_Hoare_halt_def) - apply(rule_tac x = 1 in exI) + apply (simp add: abc_Hoare_halt_def) + apply (rule_tac x = 1 in exI) + apply (auto simp add: max_def) apply(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps list_update_append) using list_update_length [of "(x - Suc y) # rec_exec (Pr (length xs) f g) (xs @ [x - Suc y]) # 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" "Suc y" anything y] - apply(simp) + apply (auto simp add: Suc_diff_Suc) + using numeral_3_eq_3 apply presburger + using numeral_3_eq_3 apply presburger done lemma adjust_paras': "length xs = n \ {\nl. nl = xs @ x # y # anything} [Inc n] [+] [Dec (Suc n) 2, Goto 0] {\nl. nl = xs @ Suc x # 0 # anything}" proof(rule_tac abc_Hoare_plus_halt) assume "length xs = n" thus "{\nl. nl = xs @ x # y # anything} [Inc n] {\ nl. nl = xs @ Suc x # y # anything}" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = 1 in exI, force simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2)) done next assume h: "length xs = n" thus "{\nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\nl. nl = xs @ Suc x # 0 # anything}" proof(induct y) case 0 thus "?case" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2)) done next case (Suc y) have "length xs = n \ {\nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\nl. nl = xs @ Suc x # 0 # anything}" by fact then obtain stp where "abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc n) 2, Goto 0] stp = (2, xs @ Suc x # 0 # anything)" using h apply(auto simp: abc_Hoare_halt_def numeral_2_eq_2) by (metis (mono_tags, lifting) abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3)) moreover have "abc_steps_l (0, xs @ Suc x # Suc y # anything) [Dec (Suc n) 2, Goto 0] 2 = (0, xs @ Suc x # y # anything)" using h by(simp add: abc_steps_l.simps numeral_2_eq_2 abc_step_l.simps abc_fetch.simps abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2)) ultimately show "?case" apply(simp add: abc_Hoare_halt_def) by(rule exI[of _ "2 + stp"], simp only: abc_steps_add, simp) qed qed lemma adjust_paras: "length xs = n \ {\nl. nl = xs @ x # y # anything} [Inc n, Dec (Suc n) 3, Goto (Suc 0)] {\nl. nl = xs @ Suc x # 0 # anything}" using adjust_paras'[of xs n x y anything] by(simp add: abc_comp.simps abc_shift.simps numeral_2_eq_2 numeral_3_eq_3) lemma rec_ci_SucSuc_n[simp]: "\rec_ci g = (gap, gar, gft); \yx\ \ gar = Suc (Suc n)" by(auto dest:param_pattern elim!:allE[of _ y]) lemma loop_back': assumes h: "length A = length gap + 4" "length xs = n" and le: "y \ x" shows "\ stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (length A, xs @ m # y # 0 # anything)" using le proof(induct x) case 0 thus "?case" using h by(rule_tac x = 0 in exI, auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_s.simps abc_lm_v.simps) next case (Suc x) have "x \ y \ \stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (length A, xs @ m # y # 0 # anything)" by fact moreover have "Suc x \ y" by fact moreover then have "\ stp. abc_steps_l (length A, xs @ m # (y - Suc x) # Suc x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (length A, xs @ m # (y - x) # x # anything)" using h apply(rule_tac x = 3 in exI) by(simp add: abc_steps_l.simps numeral_3_eq_3 abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps list_update_append list_update.simps(2)) ultimately show "?case" apply(auto simp add: abc_steps_add) by (metis abc_steps_add) qed lemma loop_back: assumes h: "length A = length gap + 4" "length xs = n" shows "\ stp. abc_steps_l (length A, xs @ m # 0 # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp = (0, xs @ m # x # 0 # anything)" using loop_back'[of A gap xs n x x m anything] assms apply(auto) apply(rename_tac stp) apply(rule_tac x = "stp + 1" in exI) apply(simp only: abc_steps_add, simp) apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps) done lemma rec_exec_pr_0_simps: "rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" by(simp add: rec_exec.simps) lemma rec_exec_pr_Suc_simps: "rec_exec (Pr n f g) (xs @ [Suc y]) = rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" apply(induct y) apply(simp add: rec_exec.simps) apply(simp add: rec_exec.simps) done lemma suc_max_simp[simp]: "Suc (max (n + 3) fft - Suc (Suc (Suc n))) = max (n + 3) fft - Suc (Suc n)" by arith lemma pr_loop: assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" and len: "length xs = n" and g_ind: "\ yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" and compile_g: "rec_ci g = (gap, gar, gft)" and termi_g: "\ y x" shows "\stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) code stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (ft - Suc (Suc n)) @ y # anything)" proof - let ?A = "[Dec ft (length gap + 7)]" let ?B = "gap" let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]" let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" have "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" proof - have "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" proof - have "{\ nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything} (?A [+] (?B [+] ?C)) {\ nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" proof(rule_tac abc_Hoare_plus_halt) show "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything} [Dec ft (length gap + 7)] {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything}" using ft len by(simp) next show "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything} ?B [+] ?C {\nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" proof(rule_tac abc_Hoare_plus_halt) have a: "gar = Suc (Suc n)" using compile_g termi_g len less by simp have b: "gft > gar" using compile_g by(erule_tac footprint_ge) show "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ y # anything} gap {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" proof - have "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (gft - gar) @ 0\(ft - gft) @ y # anything} gap {\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # rec_exec g (xs @ [(x - Suc y), rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (gft - Suc gar) @ 0\(ft - gft) @ y # anything}" using g_ind less by simp thus "?thesis" using a b ft by(simp add: replicate_merge_anywhere numeral_3_eq_3) qed next show "{\nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything} [Inc n, Dec (Suc n) 3, Goto (Suc 0)] {\nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything}" using len less using adjust_paras[of xs n "x - Suc y" " rec_exec (Pr n f g) (xs @ [x - Suc y])" " rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything"] by(simp add: Suc_diff_Suc) qed qed thus "?thesis" apply(simp add: abc_Hoare_halt_def, auto) apply(rename_tac na) apply(rule_tac x = na in exI, case_tac "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp) done qed then obtain stpa where "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (ft - Suc (Suc n)) @ Suc y # anything) ((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" .. thus "?thesis" by(erule_tac abc_append_frist_steps_halt_eq) qed moreover have "\ stp. abc_steps_l (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything) ((?A [+] (?B [+] ?C)) @ ?D) stp = (0, xs @ (x - y) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 # 0 \ (ft - Suc (Suc (Suc n))) @ y # anything)" using len by(rule_tac loop_back, simp_all) moreover have "rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) = rec_exec (Pr n f g) (xs @ [x - y])" using less apply(cases "x - y", simp_all add: rec_exec_pr_Suc_simps) apply(rename_tac nat) by(subgoal_tac "nat = x - Suc y", simp, arith) ultimately show "?thesis" using code ft apply (auto simp add: abc_steps_add replicate_Suc_iff_anywhere) apply(rename_tac stp stpa) apply(rule_tac x = "stp + stpa" in exI) by (simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc) qed lemma abc_lm_s_simp0[simp]: "length xs = n \ abc_lm_s (xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything" apply(simp add: abc_lm_s.simps) using list_update_length[of "xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n))" 0 anything 0] apply(auto simp: Suc_diff_Suc) apply(simp add: exp_suc[THEN sym] Suc_diff_Suc del: replicate_Suc) done lemma index_at_zero_elem[simp]: "(xs @ x # re # 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs))) @ 0 # anything) ! max (length xs + 3) (max fft gft) = 0" using nth_append_length[of "xs @ x # re # 0 \ (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0 anything] by(simp) lemma pr_loop_correct: assumes less: "y \ x" and len: "length xs = n" and compile_g: "rec_ci g = (gap, gar, gft)" and termi_g: "\ y yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" shows "{\nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (max (n + 3) (max fft gft) - Suc (Suc n)) @ y # anything} ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)] {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything}" using less proof(induct y) case 0 thus "?case" using len apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = 1 in exI) by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_comp.simps abc_shift.simps, simp add: abc_lm_v.simps) next case (Suc y) let ?ft = "max (n + 3) (max fft gft)" let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" have ind: "y \ x \ {\nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything} ?C {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything}" by fact have less: "Suc y \ x" by fact have stp1: "\ stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (?ft - Suc (Suc n)) @ Suc y # anything) ?C stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything)" using assms less by(rule_tac pr_loop, auto) then obtain stp1 where a: "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \ (?ft - Suc (Suc n)) @ Suc y # anything) ?C stp1 = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything)" .. moreover have "\ stp. abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C stp = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything)" using ind less apply(auto simp: abc_Hoare_halt_def) apply(rename_tac na,case_tac "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI) by simp then obtain stp2 where b: "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \ (?ft - Suc (Suc n)) @ y # anything) ?C stp2 = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything)" .. from a b show "?case" apply(simp add: abc_Hoare_halt_def) apply(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add). qed lemma compile_pr_correct': assumes termi_g: "\ y yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" and termi_f: "terminate f xs" and f_ind: "\ anything. {\nl. nl = xs @ 0 \ (fft - far) @ anything} fap {\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ anything}" and len: "length xs = n" and compile1: "rec_ci f = (fap, far, fft)" and compile2: "rec_ci g = (gap, gar, gft)" shows "{\nl. nl = xs @ x # 0 \ (max (n + 3) (max fft gft) - n) @ anything} mv_box n (max (n + 3) (max fft gft)) [+] (fap [+] (mv_box n (Suc n) [+] ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]))) {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (max (n + 3) (max fft gft) - Suc n) @ anything}" proof - let ?ft = "max (n+3) (max fft gft)" let ?A = "mv_box n ?ft" let ?B = "fap" let ?C = "mv_box n (Suc n)" let ?D = "[Dec ?ft (length gap + 7)]" let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]" let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" let ?P = "\nl. nl = xs @ x # 0 \ (?ft - n) @ anything" let ?S = "\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (?ft - Suc n) @ anything" let ?Q1 = "\nl. nl = xs @ 0 \ (?ft - n) @ x # anything" show "{?P} (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) {?S}" proof(rule_tac abc_Hoare_plus_halt) show "{?P} ?A {?Q1}" using len by simp next let ?Q2 = "\nl. nl = xs @ rec_exec f xs # 0 \ (?ft - Suc n) @ x # anything" have a: "?ft \ fft" by arith have b: "far = n" using compile1 termi_f len by(drule_tac param_pattern, auto) have c: "fft > far" using compile1 by(simp add: footprint_ge) show "{?Q1} (?B [+] (?C [+] (?D [+] ?E @ ?F))) {?S}" proof(rule_tac abc_Hoare_plus_halt) have "{\nl. nl = xs @ 0 \ (fft - far) @ 0\(?ft - fft) @ x # anything} fap {\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ 0\(?ft - fft) @ x # anything}" by(rule_tac f_ind) moreover have "fft - far + ?ft - fft = ?ft - far" using a b c by arith moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n" using a b c by arith ultimately show "{?Q1} ?B {?Q2}" using b by(simp add: replicate_merge_anywhere) next let ?Q3 = "\ nl. nl = xs @ 0 # rec_exec f xs # 0\(?ft - Suc (Suc n)) @ x # anything" show "{?Q2} (?C [+] (?D [+] ?E @ ?F)) {?S}" proof(rule_tac abc_Hoare_plus_halt) show "{?Q2} (?C) {?Q3}" using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \ (max (n + 3) (max fft gft) - Suc n) @ x # anything"] using len by(auto) next show "{?Q3} (?D [+] ?E @ ?F) {?S}" using pr_loop_correct[of x x xs n g gap gar gft f fft anything] assms by(simp add: rec_exec_pr_0_simps) qed qed qed qed lemma compile_pr_correct: assumes g_ind: "\y (\x xa xb. rec_ci g = (x, xa, xb) \ (\xc. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (xb - Suc xa) @ xc}))" and termi_f: "terminate f xs" and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ {\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec f xs # 0 \ (fp - Suc arity) @ anything}" and len: "length xs = n" and compile: "rec_ci (Pr n f g) = (ap, arity, fp)" shows "{\nl. nl = xs @ x # 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \ (fp - Suc arity) @ anything}" proof(cases "rec_ci f", cases "rec_ci g") fix fap far fft gap gar gft assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)" have g: "\y (\anything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything}))" using g_ind h by(auto) hence termi_g: "\ y yanything. {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \ (gft - gar) @ anything} gap {\nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \ (gft - Suc gar) @ anything})" by auto have f_newind: "\ anything. {\nl. nl = xs @ 0 \ (fft - far) @ anything} fap {\nl. nl = xs @ rec_exec f xs # 0 \ (fft - Suc far) @ anything}" using h by(rule_tac f_ind, simp) show "?thesis" using termi_f termi_g h compile apply(simp add: rec_ci.simps abc_comp_commute, auto) using g_newind f_newind len by(rule_tac compile_pr_correct', simp_all) qed fun mn_ind_inv :: "nat \ nat list \ nat \ nat \ nat \ nat list \ nat list \ bool" where "mn_ind_inv (as, lm') ss x rsx suf_lm lm = (if as = ss then lm' = lm @ x # rsx # suf_lm else if as = ss + 1 then \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx else if as = ss + 2 then \y. (lm' = lm @ x # y # suf_lm) \ y \ rsx else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm else False )" fun mn_stage1 :: "nat \ nat list \ nat \ nat \ nat" where "mn_stage1 (as, lm) ss n = (if as = 0 then 0 else if as = ss + 4 then 1 else if as = ss + 3 then 2 else if as = ss + 2 \ as = ss + 1 then 3 else if as = ss then 4 else 0 )" fun mn_stage2 :: "nat \ nat list \ nat \ nat \ nat" where "mn_stage2 (as, lm) ss n = (if as = ss + 1 \ as = ss + 2 then (lm ! (Suc n)) else 0)" fun mn_stage3 :: "nat \ nat list \ nat \ nat \ nat" where "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)" fun mn_measure :: "((nat \ nat list) \ nat \ nat) \ (nat \ nat \ nat)" where "mn_measure ((as, lm), ss, n) = (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n, mn_stage3 (as, lm) ss n)" definition mn_LE :: "(((nat \ nat list) \ nat \ nat) \ ((nat \ nat list) \ nat \ nat)) set" where "mn_LE \ (inv_image lex_triple mn_measure)" lemma wf_mn_le[intro]: "wf mn_LE" by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def) declare mn_ind_inv.simps[simp del] lemma put_in_tape_small_enough0[simp]: "0 < rsx \ \y. (xs @ x # rsx # anything)[Suc (length xs) := rsx - Suc 0] = xs @ x # y # anything \ y \ rsx" apply(rule_tac x = "rsx - 1" in exI) apply(simp add: list_update_append list_update.simps) done lemma put_in_tape_small_enough1[simp]: "\y \ rsx; 0 < y\ \ \ya. (xs @ x # y # anything)[Suc (length xs) := y - Suc 0] = xs @ x # ya # anything \ ya \ rsx" apply(rule_tac x = "y - 1" in exI) apply(simp add: list_update_append list_update.simps) done lemma abc_comp_null[simp]: "(A [+] B = []) = (A = [] \ B = [])" by(auto simp: abc_comp.simps abc_shift.simps) lemma rec_ci_not_null[simp]: "(rec_ci f \ ([], a, b))" proof(cases f) case (Cn x41 x42 x43) then show ?thesis by(cases "rec_ci x42", auto simp: mv_box.simps rec_ci.simps rec_ci_id.simps) next case (Pr x51 x52 x53) then show ?thesis apply(cases "rec_ci x52", cases "rec_ci x53") by (auto simp: mv_box.simps rec_ci.simps rec_ci_id.simps) next case (Mn x61 x62) then show ?thesis by(cases "rec_ci x62") (auto simp: rec_ci.simps rec_ci_id.simps) qed (auto simp: rec_ci_z_def rec_ci_s_def rec_ci.simps addition.simps rec_ci_id.simps) lemma mn_correct: assumes compile: "rec_ci rf = (fap, far, fft)" and ge: "0 < rsx" and len: "length xs = arity" and B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and f: "f = (\ stp. (abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp, (length fap), arity)) " and P: "P =(\ ((as, lm), ss, arity). as = 0)" and Q: "Q = (\ ((as, lm), ss, arity). mn_ind_inv (as, lm) (length fap) x rsx anything xs)" shows "\ stp. P (f stp) \ Q (f stp)" proof(rule_tac halt_lemma2) show "wf mn_LE" using wf_mn_le by simp next show "Q (f 0)" by(auto simp: Q f abc_steps_l.simps mn_ind_inv.simps) next have "fap \ []" using compile by auto thus "\ P (f 0)" by(auto simp: f P abc_steps_l.simps) next have "fap \ []" using compile by auto then have "\\ P (f stp); Q (f stp)\ \ Q (f (Suc stp)) \ (f (Suc stp), f stp) \ mn_LE" for stp using ge len apply(cases "(abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp)") apply(simp add: abc_step_red2 B f P Q) apply(auto split:if_splits simp add:abc_steps_l.simps mn_ind_inv.simps abc_steps_zero B abc_fetch.simps nth_append) by(auto simp: mn_LE_def lex_triple_def lex_pair_def abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps abc_lm_v.simps abc_lm_s.simps nth_append abc_fetch.simps split: if_splits) thus "\stp. \ P (f stp) \ Q (f stp) \ Q (f (Suc stp)) \ (f (Suc stp), f stp) \ mn_LE" by(auto) qed lemma abc_Hoare_haltE: "{\ nl. nl = lm1} p {\ nl. nl = lm2} \ \ stp. abc_steps_l (0, lm1) p stp = (length p, lm2)" by(auto simp:abc_Hoare_halt_def elim!: abc_holds_for.elims) lemma mn_loop: assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and far: "far = Suc arity" and ind: " (\xc. ({\nl. nl = xs @ x # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \ (fft - Suc far) @ xc}))" and exec_less: "rec_exec f (xs @ [x]) > 0" and compile: "rec_ci f = (fap, far, fft)" shows "\ stp > 0. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" proof - have "\ stp. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "\ stp. abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "{\nl. nl = xs @ x # 0 \ (fft - far) @ 0\(ft - fft) @ anything} fap {\nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \ (fft - Suc far) @ 0\(ft - fft) @ anything}" using ind by simp moreover have "fft > far" using compile by(erule_tac footprint_ge) ultimately show "?thesis" using ft far apply(drule_tac abc_Hoare_haltE) by(simp add: replicate_merge_anywhere max_absorb2) qed then obtain stp where "abc_steps_l (0, xs @ x # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything)" .. thus "?thesis" by(erule_tac abc_append_frist_steps_halt_eq) qed moreover have "\ stp > 0. abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 # 0 \ (ft - Suc (Suc arity)) @ anything)" using mn_correct[of f fap far fft "rec_exec f (xs @ [x])" xs arity B "(\stp. (abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp, length fap, arity))" x "0 \ (ft - Suc (Suc arity)) @ anything" "(\((as, lm), ss, arity). as = 0)" "(\((as, lm), ss, aritya). mn_ind_inv (as, lm) (length fap) x (rec_exec f (xs @ [x])) (0 \ (ft - Suc (Suc arity)) @ anything) xs) "] B compile exec_less len apply(subgoal_tac "fap \ []", auto) apply(rename_tac stp y) apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps) by(case_tac "stp", simp_all add: abc_steps_l.simps) moreover have "fft > far" using compile by(erule_tac footprint_ge) ultimately show "?thesis" using ft far apply(auto) apply(rename_tac stp1 stp2) by(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add replicate_Suc[THEN sym] diff_Suc_Suc del: replicate_Suc) qed lemma mn_loop_correct': assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and ind_all: "\i\x. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" and exec_ge: "\ i\x. rec_exec f (xs @ [i]) > 0" and compile: "rec_ci f = (fap, far, fft)" and far: "far = Suc arity" shows "\ stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using ind_all exec_ge proof(induct x) case 0 thus "?case" using assms by(rule_tac mn_loop, simp_all) next case (Suc x) have ind': "\\i\x. \xc. {\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}; \i\x. 0 < rec_exec f (xs @ [i])\ \ \stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" by fact have exec_ge: "\i\Suc x. 0 < rec_exec f (xs @ [i])" by fact have ind_all: "\i\Suc x. \xc. {\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}" by fact have ind: "\stp > x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using ind' exec_ge ind_all by simp have stp: "\ stp > 0. abc_steps_l (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc (Suc x) # 0 \ (ft - Suc arity) @ anything)" using ind_all exec_ge B ft len far compile by(rule_tac mn_loop, simp_all) from ind stp show "?case" apply(auto simp add: abc_steps_add) apply(rename_tac stp1 stp2) by(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add) qed lemma mn_loop_correct: assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and ind_all: "\i\x. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" and exec_ge: "\ i\x. rec_exec f (xs @ [i]) > 0" and compile: "rec_ci f = (fap, far, fft)" and far: "far = Suc arity" shows "\ stp. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" proof - have "\stp>x. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \ (ft - Suc arity) @ anything)" using assms by(rule_tac mn_loop_correct', simp_all) thus "?thesis" by(auto) qed lemma compile_mn_correct': assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" and ft: "ft = max (Suc arity) fft" and len: "length xs = arity" and termi_f: "terminate f (xs @ [r])" and f_ind: "\anything. {\nl. nl = xs @ r # 0 \ (fft - far) @ anything} fap {\nl. nl = xs @ r # 0 # 0 \ (fft - Suc far) @ anything}" and ind_all: "\i < r. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" and exec_less: "\ i 0" and exec: "rec_exec f (xs @ [r]) = 0" and compile: "rec_ci f = (fap, far, fft)" shows "{\nl. nl = xs @ 0 \ (max (Suc arity) fft - arity) @ anything} fap @ B {\nl. nl = xs @ rec_exec (Mn arity f) xs # 0 \ (max (Suc arity) fft - Suc arity) @ anything}" proof - have a: "far = Suc arity" using len compile termi_f by(drule_tac param_pattern, auto) have b: "fft > far" using compile by(erule_tac footprint_ge) have "\ stp. abc_steps_l (0, xs @ 0 # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ r # 0 \ (ft - Suc arity) @ anything)" using assms a apply(cases r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp) by(rule_tac mn_loop_correct, auto) moreover have "\ stp. abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) (fap @ B) stp = (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "\ stp. abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" proof - have "{\nl. nl = xs @ r # 0 \ (fft - far) @ 0\(ft - fft) @ anything} fap {\nl. nl = xs @ r # rec_exec f (xs @ [r]) # 0 \ (fft - Suc far) @ 0\(ft - fft) @ anything}" using f_ind exec by simp thus "?thesis" using ft a b apply(drule_tac abc_Hoare_haltE) by(simp add: replicate_merge_anywhere max_absorb2) qed then obtain stp where "abc_steps_l (0, xs @ r # 0 \ (ft - Suc arity) @ anything) fap stp = (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" .. thus "?thesis" by(erule_tac abc_append_frist_steps_halt_eq) qed moreover have "\ stp. abc_steps_l (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = (length fap + 5, xs @ r # rec_exec f (xs @ [r]) # 0 \ (ft - Suc (Suc arity)) @ anything)" using ft a b len B exec apply(rule_tac x = 1 in exI, auto) by(auto simp: abc_steps_l.simps B abc_step_l.simps abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps) moreover have "rec_exec (Mn (length xs) f) xs = r" using exec exec_less apply(auto simp: rec_exec.simps Least_def) thm the_equality apply(rule_tac the_equality, auto) apply(metis exec_less less_not_refl3 linorder_not_less) by (metis le_neq_implies_less less_not_refl3) ultimately show "?thesis" using ft a b len B exec apply(auto simp: abc_Hoare_halt_def) apply(rename_tac stp1 stp2 stp3) apply(rule_tac x = "stp1 + stp2 + stp3" in exI) by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc) qed lemma compile_mn_correct: assumes len: "length xs = n" and termi_f: "terminate f (xs @ [r])" and f_ind: "\ap arity fp anything. rec_ci f = (ap, arity, fp) \ {\nl. nl = xs @ r # 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ r # 0 # 0 \ (fp - Suc arity) @ anything}" and exec: "rec_exec f (xs @ [r]) = 0" and ind_all: "\i (\x xa xb. rec_ci f = (x, xa, xb) \ (\xc. {\nl. nl = xs @ i # 0 \ (xb - xa) @ xc} x {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (xb - Suc xa) @ xc})) \ 0 < rec_exec f (xs @ [i])" and compile: "rec_ci (Mn n f) = (ap, arity, fp)" shows "{\nl. nl = xs @ 0 \ (fp - arity) @ anything} ap {\nl. nl = xs @ rec_exec (Mn n f) xs # 0 \ (fp - Suc arity) @ anything}" proof(cases "rec_ci f") fix fap far fft assume h: "rec_ci f = (fap, far, fft)" hence f_newind: "\anything. {\nl. nl = xs @ r # 0 \ (fft - far) @ anything} fap {\nl. nl = xs @ r # 0 # 0 \ (fft - Suc far) @ anything}" by(rule_tac f_ind, simp) have newind_all: "\i < r. (\xc. ({\nl. nl = xs @ i # 0 \ (fft - far) @ xc} fap {\nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \ (fft - Suc far) @ xc}))" using ind_all h by(auto) have all_less: "\ i 0" using ind_all by auto show "?thesis" using h compile f_newind newind_all all_less len termi_f exec apply(auto simp: rec_ci.simps) by(rule_tac compile_mn_correct', auto) qed lemma recursive_compile_correct: "\terminate recf args; rec_ci recf = (ap, arity, fp)\ \ {\ nl. nl = args @ 0\(fp - arity) @ anything} ap {\ nl. nl = args@ rec_exec recf args # 0\(fp - Suc arity) @ anything}" apply(induct arbitrary: ap arity fp anything rule: terminate.induct) apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct compile_cn_correct compile_pr_correct compile_mn_correct) done definition dummy_abc :: "nat \ abc_inst list" where "dummy_abc k = [Inc k, Dec k 0, Goto 3]" definition abc_list_crsp:: "nat list \ nat list \ bool" where "abc_list_crsp xs ys = (\ n. xs = ys @ 0\n \ ys = xs @ 0\n)" lemma abc_list_crsp_simp1[intro]: "abc_list_crsp (lm @ 0\m) lm" by(auto simp: abc_list_crsp_def) lemma abc_list_crsp_lm_v: "abc_list_crsp lma lmb \ abc_lm_v lma n = abc_lm_v lmb n" by(auto simp: abc_list_crsp_def abc_lm_v.simps nth_append) lemma abc_list_crsp_elim: "\abc_list_crsp lma lmb; \ n. lma = lmb @ 0\n \ lmb = lma @ 0\n \ P \ \ P" by(auto simp: abc_list_crsp_def) lemma abc_list_crsp_simp[simp]: "\abc_list_crsp lma lmb; m < length lma; m < length lmb\ \ abc_list_crsp (lma[m := n]) (lmb[m := n])" by(auto simp: abc_list_crsp_def list_update_append) lemma abc_list_crsp_simp2[simp]: "\abc_list_crsp lma lmb; m < length lma; \ m < length lmb\ \ abc_list_crsp (lma[m := n]) (lmb @ 0 \ (m - length lmb) @ [n])" apply(auto simp: abc_list_crsp_def list_update_append) apply(rename_tac N) apply(rule_tac x = "N + length lmb - Suc m" in exI) apply(rule_tac disjI1) apply(simp add: upd_conv_take_nth_drop min_absorb1) done lemma abc_list_crsp_simp3[simp]: "\abc_list_crsp lma lmb; \ m < length lma; m < length lmb\ \ abc_list_crsp (lma @ 0 \ (m - length lma) @ [n]) (lmb[m := n])" apply(auto simp: abc_list_crsp_def list_update_append) apply(rename_tac N) apply(rule_tac x = "N + length lma - Suc m" in exI) apply(rule_tac disjI2) apply(simp add: upd_conv_take_nth_drop min_absorb1) done lemma abc_list_crsp_simp4[simp]: "\abc_list_crsp lma lmb; \ m < length lma; \ m < length lmb\ \ abc_list_crsp (lma @ 0 \ (m - length lma) @ [n]) (lmb @ 0 \ (m - length lmb) @ [n])" by(auto simp: abc_list_crsp_def list_update_append replicate_merge_anywhere) lemma abc_list_crsp_lm_s: "abc_list_crsp lma lmb \ abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)" by(auto simp: abc_lm_s.simps) lemma abc_list_crsp_step: "\abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); abc_step_l (aa, lmb) i = (a', lmb')\ \ a' = a \ abc_list_crsp lma' lmb'" apply(cases i, auto simp: abc_step_l.simps abc_list_crsp_lm_s abc_list_crsp_lm_v split: abc_inst.splits if_splits) done lemma abc_list_crsp_steps: "\abc_steps_l (0, lm @ 0\m) aprog stp = (a, lm'); aprog \ []\ \ \ lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ abc_list_crsp lm' lma" proof(induct stp arbitrary: a lm') case (Suc stp) then show ?case using [[simproc del: defined_all]] apply(cases "abc_steps_l (0, lm @ 0\m) aprog stp", simp add: abc_step_red) proof - fix stp a lm' aa b assume ind: "\a lm'. aa = a \ b = lm' \ \lma. abc_steps_l (0, lm) aprog stp = (a, lma) \ abc_list_crsp lm' lma" and h: "abc_step_l (aa, b) (abc_fetch aa aprog) = (a, lm')" "abc_steps_l (0, lm @ 0\m) aprog stp = (aa, b)" "aprog \ []" have "\lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \ abc_list_crsp b lma" apply(rule_tac ind, simp) done from this obtain lma where g2: "abc_steps_l (0, lm) aprog stp = (aa, lma) \ abc_list_crsp b lma" .. hence g3: "abc_steps_l (0, lm) aprog (Suc stp) = abc_step_l (aa, lma) (abc_fetch aa aprog)" apply(rule_tac abc_step_red, simp) done show "\lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \ abc_list_crsp lm' lma" using g2 g3 h apply(auto) apply(cases "abc_step_l (aa, b) (abc_fetch aa aprog)", cases "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp) apply(rule_tac abc_list_crsp_step, auto) done qed qed (force simp add: abc_steps_l.simps) lemma list_crsp_simp2: "abc_list_crsp (lm1 @ 0\n) lm2 \ abc_list_crsp lm1 lm2" proof(induct n) case 0 thus "?case" by(auto simp: abc_list_crsp_def) next case (Suc n) have ind: "abc_list_crsp (lm1 @ 0 \ n) lm2 \ abc_list_crsp lm1 lm2" by fact have h: "abc_list_crsp (lm1 @ 0 \ Suc n) lm2" by fact then have "abc_list_crsp (lm1 @ 0 \ n) lm2" apply(auto simp only: exp_suc abc_list_crsp_def del: replicate_Suc) apply (metis Suc_pred append_eq_append_conv append_eq_append_conv2 butlast_append butlast_snoc length_replicate list.distinct(1) neq0_conv replicate_Suc replicate_Suc_iff_anywhere replicate_app_Cons_same replicate_empty self_append_conv self_append_conv2) apply (auto,metis replicate_Suc) . thus "?case" using ind by auto qed lemma recursive_compile_correct_norm': "\rec_ci f = (ap, arity, ft); terminate f args\ \ \ stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \ abc_list_crsp (args @ [rec_exec f args]) rl" using recursive_compile_correct[of f args ap arity ft "[]"] apply(auto simp: abc_Hoare_halt_def) apply(rename_tac n) apply(rule_tac x = n in exI) apply(case_tac "abc_steps_l (0, args @ 0 \ (ft - arity)) ap n", auto) apply(drule_tac abc_list_crsp_steps, auto) apply(rule_tac list_crsp_simp2, auto) done lemma find_exponent_rec_exec[simp]: assumes a: "args @ [rec_exec f args] = lm @ 0 \ n" and b: "length args < length lm" shows "\m. lm = args @ rec_exec f args # 0 \ m" using assms apply(cases n, simp) apply(rule_tac x = 0 in exI, simp) apply(drule_tac length_equal, simp) done lemma find_exponent_complex[simp]: "\args @ [rec_exec f args] = lm @ 0 \ n; \ length args < length lm\ \ \m. (lm @ 0 \ (length args - length lm) @ [Suc 0])[length args := 0] = args @ rec_exec f args # 0 \ m" apply(cases n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc) apply(subgoal_tac "length args = Suc (length lm)", simp) apply(rule_tac x = "Suc (Suc 0)" in exI, simp) apply(drule_tac length_equal, simp, auto) done lemma compile_append_dummy_correct: assumes compile: "rec_ci f = (ap, ary, fp)" and termi: "terminate f args" shows "{\ nl. nl = args} (ap [+] dummy_abc (length args)) {\ nl. (\ m. nl = args @ rec_exec f args # 0\m)}" proof(rule_tac abc_Hoare_plus_halt) show "{\nl. nl = args} ap {\ nl. abc_list_crsp (args @ [rec_exec f args]) nl}" using compile termi recursive_compile_correct_norm'[of f ap ary fp args] apply(auto simp: abc_Hoare_halt_def) by (metis abc_final.simps abc_holds_for.simps) next show "{abc_list_crsp (args @ [rec_exec f args])} dummy_abc (length args) {\nl. \m. nl = args @ rec_exec f args # 0 \ m}" apply(auto simp: dummy_abc_def abc_Hoare_halt_def) apply(rule_tac x = 3 in exI) by(force simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps) qed lemma cn_merge_gs_split: "\i < length gs; rec_ci (gs!i) = (ga, gb, gc)\ \ cn_merge_gs (map rec_ci gs) p = cn_merge_gs (map rec_ci (take i gs)) p [+] (ga [+] mv_box gb (p + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)" proof(induct i arbitrary: gs p) case 0 then show ?case by(cases gs; simp) next case (Suc i) then show ?case by(cases gs, simp, cases "rec_ci (hd gs)", simp add: abc_comp_commute[THEN sym]) qed lemma cn_unhalt_case: assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \ length args = ar" and g: "i < length gs" and compile2: "rec_ci (gs!i) = (gap, gar, gft) \ gar = length args" and g_unhalt: "\ anything. {\ nl. nl = args @ 0\(gft - gar) @ anything} gap \" and g_ind: "\ apj arj ftj j anything. \j < i; rec_ci (gs!j) = (apj, arj, ftj)\ \ {\ nl. nl = args @ 0\(ftj - arj) @ anything} apj {\ nl. nl = args @ rec_exec (gs!j) args # 0\(ftj - Suc arj) @ anything}" and all_termi: "\ j nl. nl = args @ 0\(ft - ar) @ anything} ap \" using compile1 apply(cases "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) proof(rule_tac abc_Hoare_plus_unhalt1) fix fap far fft let ?ft = "max (Suc (length args)) (Max (insert fft ((\(aprog, p, n). n) ` rec_ci ` set gs)))" let ?Q = "\nl. nl = args @ 0\ (?ft - length args) @ map (\i. rec_exec i args) (take i gs) @ 0\(length gs - i) @ 0\ Suc (length args) @ anything" have "cn_merge_gs (map rec_ci gs) ?ft = cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+] mv_box gar (?ft + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)" using g compile2 cn_merge_gs_split by simp thus "{\nl. nl = args @ 0 # 0 \ (?ft + length gs) @ anything} (cn_merge_gs (map rec_ci gs) ?ft) \" proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2, rule_tac abc_Hoare_plus_unhalt1) let ?Q_tmp = "\nl. nl = args @ 0\ (gft - gar) @ 0\(?ft - (length args) - (gft -gar)) @ map (\i. rec_exec i args) (take i gs) @ 0\(length gs - i) @ 0\ Suc (length args) @ anything" have a: "{?Q_tmp} gap \" using g_unhalt[of "0 \ (?ft - (length args) - (gft - gar)) @ map (\i. rec_exec i args) (take i gs) @ 0 \ (length gs - i) @ 0 \ Suc (length args) @ anything"] by simp moreover have "?ft \ gft" using g compile2 apply(rule_tac max.coboundedI2, rule_tac Max_ge, simp, rule_tac insertI2) apply(rule_tac x = "rec_ci (gs ! i)" in image_eqI, simp) by(rule_tac x = "gs!i" in image_eqI, simp, simp) then have b:"?Q_tmp = ?Q" using compile2 apply(rule_tac arg_cong) by(simp add: replicate_merge_anywhere) thus "{?Q} gap \" using a by simp next show "{\nl. nl = args @ 0 # 0 \ (?ft + length gs) @ anything} cn_merge_gs (map rec_ci (take i gs)) ?ft {\nl. nl = args @ 0 \ (?ft - length args) @ map (\i. rec_exec i args) (take i gs) @ 0 \ (length gs - i) @ 0 \ Suc (length args) @ anything}" using all_termi by(rule_tac compile_cn_gs_correct', auto simp: set_conv_nth intro:g_ind) qed qed lemma mn_unhalt_case': assumes compile: "rec_ci f = (a, b, c)" and all_termi: "\i. terminate f (args @ [i]) \ 0 < rec_exec f (args @ [i])" and B: "B = [Dec (Suc (length args)) (length a + 5), Dec (Suc (length args)) (length a + 3), Goto (Suc (length a)), Inc (length args), Goto 0]" shows "{\nl. nl = args @ 0 \ (max (Suc (length args)) c - length args) @ anything} a @ B \" proof(rule_tac abc_Hoare_unhaltI, auto) fix n have a: "b = Suc (length args)" using all_termi compile apply(erule_tac x = 0 in allE) by(auto, drule_tac param_pattern,auto) moreover have b: "c > b" using compile by(elim footprint_ge) ultimately have c: "max (Suc (length args)) c = c" by arith have "\ stp > n. abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) stp = (0, args @ Suc n # 0\(c - Suc (length args)) @ anything)" using assms a b c proof(rule_tac mn_loop_correct', auto) fix i xc show "{\nl. nl = args @ i # 0 \ (c - Suc (length args)) @ xc} a {\nl. nl = args @ i # rec_exec f (args @ [i]) # 0 \ (c - Suc (Suc (length args))) @ xc}" using all_termi recursive_compile_correct[of f "args @ [i]" a b c xc] compile a by(simp) qed then obtain stp where d: "stp > n \ abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) stp = (0, args @ Suc n # 0\(c - Suc (length args)) @ anything)" .. then obtain d where e: "stp = n + Suc d" by (metis add_Suc_right less_iff_Suc_add) obtain s nl where f: "abc_steps_l (0, args @ 0 # 0\(c - Suc (length args)) @ anything) (a @ B) n = (s, nl)" by (metis prod.exhaust) have g: "s < length (a @ B)" using d e f apply(rule_tac classical, simp only: abc_steps_add) by(simp add: halt_steps2 leI) from f g show "abc_notfinal (abc_steps_l (0, args @ 0 \ (max (Suc (length args)) c - length args) @ anything) (a @ B) n) (a @ B)" using c b a by(simp add: replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc) qed lemma mn_unhalt_case: assumes compile: "rec_ci (Mn n f) = (ap, ar, ft) \ length args = ar" and all_term: "\ i. terminate f (args @ [i]) \ rec_exec f (args @ [i]) > 0" shows "{\ nl. nl = args @ 0\(ft - ar) @ anything} ap \ " using assms apply(cases "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) by(rule_tac mn_unhalt_case', simp_all) fun tm_of_rec :: "recf \ instr list" where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in let tp = tm_of (ap [+] dummy_abc k) in tp @ (shift (mopup k) (length tp div 2)))" lemma recursive_compile_to_tm_correct1: assumes compile: "rec_ci recf = (ap, ary, fp)" and termi: " terminate recf args" and tp: "tp = tm_of (ap [+] dummy_abc (length args))" shows "\ stp m l. steps0 (Suc 0, Bk # Bk # ires, @ Bk\rn) (tp @ shift (mopup (length args)) (length tp div 2)) stp = (0, Bk\m @ Bk # Bk # ires, Oc\Suc (rec_exec recf args) @ Bk\l)" proof - have "{\nl. nl = args} ap [+] dummy_abc (length args) {\nl. \m. nl = args @ rec_exec recf args # 0 \ m}" using compile termi compile by(rule_tac compile_append_dummy_correct, auto) then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp = (length (ap [+] dummy_abc (length args)), args @ rec_exec recf args # 0\m) " apply(simp add: abc_Hoare_halt_def, auto) apply(rename_tac n) by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto) thus "?thesis" using assms tp compile_correct_halt[OF refl refl _ h _ _ refl] by(auto simp: crsp.simps start_of.simps abc_lm_v.simps) qed lemma recursive_compile_to_tm_correct2: assumes termi: " terminate recf args" shows "\ stp m l. steps0 (Suc 0, [Bk, Bk], ) (tm_of_rec recf) stp = (0, Bk\Suc (Suc m), Oc\Suc (rec_exec recf args) @ Bk\l)" proof(cases "rec_ci recf", simp add: tm_of_rec.simps) fix ap ar fp assume "rec_ci recf = (ap, ar, fp)" thus "\stp m l. steps0 (Suc 0, [Bk, Bk], ) (tm_of (ap [+] dummy_abc ar) @ shift (mopup ar) (sum_list (layout_of (ap [+] dummy_abc ar)))) stp = (0, Bk # Bk # Bk \ m, Oc # Oc \ rec_exec recf args @ Bk \ l)" using recursive_compile_to_tm_correct1[of recf ap ar fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0] assms param_pattern[of recf args ap ar fp] by(simp add: replicate_Suc[THEN sym] replicate_Suc_iff_anywhere del: replicate_Suc, simp add: exp_suc del: replicate_Suc) qed lemma recursive_compile_to_tm_correct3: assumes termi: "terminate recf args" shows "{\ tp. tp =([Bk, Bk], )} (tm_of_rec recf) {\ tp. \ k l. tp = (Bk\ k, @ Bk \ l)}" using recursive_compile_to_tm_correct2[OF assms] apply(auto simp add: Hoare_halt_def ) apply(rename_tac stp M l) apply(rule_tac x = stp in exI) apply(auto simp add: tape_of_nat_def) apply(rule_tac x = "Suc (Suc M)" in exI) apply(simp) done lemma list_all_suc_many[simp]: "list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \ list_all (\(acn, s). s \ Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs" proof(induct xs) case (Cons a xs) then show ?case by(cases a, simp) qed simp lemma shift_append: "shift (xs @ ys) n = shift xs n @ shift ys n" apply(simp add: shift.simps) done lemma length_shift_mopup[simp]: "length (shift (mopup n) ss) = 4 * n + 12" apply(auto simp: mopup.simps shift_append mopup_b_def) done lemma length_tm_even[intro]: "length (tm_of ap) mod 2 = 0" apply(simp add: tm_of.simps) done lemma tms_of_at_index[simp]: "k < length ap \ tms_of ap ! k = ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)" apply(simp add: tms_of.simps tpairs_of.simps) done lemma start_of_suc_inc: "\k < length ap; ap ! k = Inc n\ \ start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2 * n + 9" apply(rule_tac start_of_Suc1, auto simp: abc_fetch.simps) done lemma start_of_suc_dec: "\k < length ap; ap ! k = (Dec n e)\ \ start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2 * n + 16" apply(rule_tac start_of_Suc2, auto simp: abc_fetch.simps) done lemma inc_state_all_le: "\k < length ap; ap ! k = Inc n; (a, b) \ set (shift (shift tinc_b (2 * n)) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") apply(arith) apply(cases "Suc k = length ap", simp) apply(rule_tac start_of_less, simp) apply(auto simp: tinc_b_def shift.simps start_of_suc_inc length_of.simps ) done lemma findnth_le[elim]: "(a, b) \ set (shift (findnth n) (start_of (layout_of ap) k - Suc 0)) \ b \ Suc (start_of (layout_of ap) k + 2 * n)" apply(induct n, force simp add: shift.simps) apply(simp add: shift_append, auto) apply(auto simp: shift.simps) done lemma findnth_state_all_le1: "\k < length ap; ap ! k = Inc n; (a, b) \ set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") apply(arith) apply(cases "Suc k = length ap", simp) apply(rule_tac start_of_less, simp) apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k)", auto) apply(auto simp: tinc_b_def shift.simps length_of.simps start_of_suc_inc) done lemma start_of_eq: "length ap < as \ start_of (layout_of ap) as = start_of (layout_of ap) (length ap)" proof(induct as) case (Suc as) then show ?case apply(cases "length ap < as", simp add: start_of.simps) apply(subgoal_tac "as = length ap") apply(simp add: start_of.simps) apply arith done qed simp lemma start_of_all_le: "start_of (layout_of ap) as \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "as > length ap \ as = length ap \ as < length ap", auto simp: start_of_eq start_of_less) done lemma findnth_state_all_le2: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap)", auto) apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16", simp) apply(simp add: start_of_suc_dec) apply(rule_tac start_of_all_le) done lemma dec_state_all_le[simp]: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (shift (shift tdec_b (2 * n)) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \ start_of (layout_of ap) (length ap) \ start_of (layout_of ap) k > 0") prefer 2 apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16 \ start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap)") apply(simp, rule_tac conjI) apply(simp add: start_of_suc_dec) apply(rule_tac start_of_all_le) apply(auto simp: tdec_b_def shift.simps) done lemma tms_any_less: "\k < length ap; (a, b) \ set (tms_of ap ! k)\ \ b \ start_of (layout_of ap) (length ap)" apply(cases "ap!k", auto simp: tms_of.simps tpairs_of.simps ci.simps shift_append adjust.simps) apply(erule_tac findnth_state_all_le1, simp_all) apply(erule_tac inc_state_all_le, simp_all) apply(erule_tac findnth_state_all_le2, simp_all) apply(rule_tac start_of_all_le) apply(rule_tac start_of_all_le) done lemma concat_in: "i < length (concat xs) \ \k < length xs. concat xs ! i \ set (xs ! k)" proof(induct xs rule: rev_induct) case (snoc x xs) then show ?case apply(cases "i < length (concat xs)", simp) apply(erule_tac exE, rule_tac x = k in exI) apply(simp add: nth_append) apply(rule_tac x = "length xs" in exI, simp) apply(simp add: nth_append) done qed auto declare length_concat[simp] lemma in_tms: "i < length (tm_of ap) \ \ k < length ap. (tm_of ap ! i) \ set (tms_of ap ! k)" apply(simp only: tm_of.simps) using concat_in[of i "tms_of ap"] apply(auto) done lemma all_le_start_of: "list_all (\(acn, s). s \ start_of (layout_of ap) (length ap)) (tm_of ap)" apply(simp only: list_all_length) apply(rule_tac allI, rule_tac impI) apply(drule_tac in_tms, auto elim: tms_any_less) done lemma length_ci: "\k < length ap; length (ci ly y (ap ! k)) = 2 * qa\ \ layout_of ap ! k = qa" apply(cases "ap ! k") apply(auto simp: layout_of.simps ci.simps length_of.simps tinc_b_def tdec_b_def length_findnth adjust.simps) done lemma ci_even[intro]: "length (ci ly y i) mod 2 = 0" apply(cases i, auto simp: ci.simps length_findnth tinc_b_def adjust.simps tdec_b_def) done lemma sum_list_ci_even[intro]: "sum_list (map (length \ (\(x, y). ci ly x y)) zs) mod 2 = 0" proof(induct zs rule: rev_induct) case (snoc x xs) then show ?case apply(cases x, simp) apply(subgoal_tac "length (ci ly (fst x) (snd x)) mod 2 = 0") apply(auto) done qed (simp) lemma zip_pre: "(length ys) \ length ap \ zip ys ap = zip ys (take (length ys) (ap::'a list))" proof(induct ys arbitrary: ap) case (Cons a ys) from Cons(2) have z:"ap = aa # list \ zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)" for aa list using Cons(1)[of list] by simp thus ?case by (cases ap;simp) qed simp lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap) div 2)" using tpa_states[of "tm_of ap" "length ap" ap] by(simp add: tm_of.simps) lemma list_all_add_6E[elim]: "list_all (\(acn, s). s \ Suc q) xs \ list_all (\(acn, s). s \ q + (2 * n + 6)) xs" by(auto simp: list_all_length) lemma mopup_b_12[simp]: "length mopup_b = 12" by(simp add: mopup_b_def) lemma mp_up_all_le: "list_all (\(acn, s). s \ q + (2 * n + 6)) [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))), (W0, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q), (L, 6 + 2 * n + q), (R, 0), (L, 6 + 2 * n + q)]" by(auto) lemma mopup_le6[simp]: "(a, b) \ set (mopup_a n) \ b \ 2 * n + 6" by(induct n, auto simp: mopup_a.simps) lemma shift_le2[simp]: "(a, b) \ set (shift (mopup n) x) \ b \ (2 * x + length (mopup n)) div 2" apply(auto simp: mopup.simps shift_append shift.simps) apply(auto simp: mopup_b_def) done lemma mopup_ge2[intro]: " 2 \ x + length (mopup n)" apply(simp add: mopup.simps) done lemma mopup_even[intro]: " (2 * x + length (mopup n)) mod 2 = 0" by(auto simp: mopup.simps) lemma mopup_div_2[simp]: "b \ Suc x \ b \ (2 * x + length (mopup n)) div 2" by(auto simp: mopup.simps) lemma wf_tm_from_abacus: assumes "tp = tm_of ap" shows "tm_wf0 (tp @ shift (mopup n) (length tp div 2))" proof - have "is_even (length (mopup n))" for n using tm_wf.simps by blast moreover have "(aa, ba) \ set (mopup n) \ ba \ length (mopup n) div 2" for aa ba by (metis (no_types, lifting) add_cancel_left_right case_prodD tm_wf.simps wf_mopup) moreover have "(\x\set (tm_of ap). case x of (acn, s) \ s \ Suc (sum_list (layout_of ap))) \ (a, b) \ set (tm_of ap) \ b \ sum_list (layout_of ap) + length (mopup n) div 2" for a b s by (metis (no_types, lifting) add_Suc add_cancel_left_right case_prodD div_mult_mod_eq le_SucE mult_2_right not_numeral_le_zero tm_wf.simps trans_le_add1 wf_mopup) ultimately show ?thesis unfolding assms using length_start_of_tm[of ap] all_le_start_of[of ap] tm_wf.simps by(auto simp: List.list_all_iff shift.simps) qed lemma wf_tm_from_recf: assumes compile: "tp = tm_of_rec recf" shows "tm_wf0 tp" proof - obtain a b c where "rec_ci recf = (a, b, c)" by (metis prod_cases3) thus "?thesis" using compile using wf_tm_from_abacus[of "tm_of (a [+] dummy_abc b)" "(a [+] dummy_abc b)" b] by simp qed end \ No newline at end of file diff --git a/thys/Word_Lib/More_Arithmetic.thy b/thys/Word_Lib/More_Arithmetic.thy --- a/thys/Word_Lib/More_Arithmetic.thy +++ b/thys/Word_Lib/More_Arithmetic.thy @@ -1,141 +1,139 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Arithmetic lemmas\ theory More_Arithmetic imports Main "HOL-Library.Type_Length" "HOL-Library.Bit_Operations" begin declare iszero_0 [intro] -declare min.absorb1 [simp] min.absorb2 [simp] \ \TODO: consider for move to distro; and the same also for max\ - -lemma n_less_equal_power_2 [simp]: +lemma n_less_equal_power_2: "n < 2 ^ n" by (fact less_exp) lemma min_pm [simp]: "min a b + (a - b) = a" for a b :: nat by arith lemma min_pm1 [simp]: "a - b + min a b = a" for a b :: nat by arith lemma rev_min_pm [simp]: "min b a + (a - b) = a" for a b :: nat by arith lemma rev_min_pm1 [simp]: "a - b + min b a = a" for a b :: nat by arith lemma min_minus [simp]: "min m (m - k) = m - k" for m k :: nat by arith lemma min_minus' [simp]: "min (m - k) m = m - k" for m k :: nat by arith lemma nat_less_power_trans: fixes n :: nat assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" shows "2 ^ k * n < 2 ^ m" proof (rule order_less_le_trans) show "2 ^ k * n < 2 ^ k * 2 ^ (m - k)" by (rule mult_less_mono2 [OF nv zero_less_power]) simp show "(2::nat) ^ k * 2 ^ (m - k) \ 2 ^ m" using nv kv by (subst power_add [symmetric]) simp qed lemma nat_le_power_trans: fixes n :: nat shows "\n \ 2 ^ (m - k); k \ m\ \ 2 ^ k * n \ 2 ^ m" by (metis le_add_diff_inverse mult_le_mono2 semiring_normalization_rules(26)) lemma nat_add_offset_less: fixes x :: nat assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" and "0 < qy" by (auto dest: less_imp_add_positive) have "x * 2 ^ n + y < x * 2 ^ n + 2 ^ n" by simp fact+ also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y < 2 ^ (m + n)" . qed lemma nat_power_less_diff: assumes lt: "(2::nat) ^ n * q < 2 ^ m" shows "q < 2 ^ (m - n)" using lt proof (induct n arbitrary: m) case 0 then show ?case by simp next case (Suc n) have ih: "\m. 2 ^ n * q < 2 ^ m \ q < 2 ^ (m - n)" and prem: "2 ^ Suc n * q < 2 ^ m" by fact+ show ?case proof (cases m) case 0 then show ?thesis using Suc by simp next case (Suc m') then show ?thesis using prem by (simp add: ac_simps ih) qed qed lemma power_2_mult_step_le: "\n' \ n; 2 ^ n' * k' < 2 ^ n * k\ \ 2 ^ n' * (k' + 1) \ 2 ^ n * (k::nat)" apply (cases "n'=n", simp) apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7)) apply (drule (1) le_neq_trans) apply clarsimp apply (subgoal_tac "\m. n = n' + m") prefer 2 apply (simp add: le_Suc_ex) apply (clarsimp simp: power_add) apply (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) done lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma diff_diff_less: "(i < m - (m - (n :: nat))) = (i < m \ i < n)" by auto lemma small_powers_of_2: \x < 2 ^ (x - 1)\ if \x \ 3\ for x :: nat proof - define m where \m = x - 3\ with that have \x = m + 3\ by simp moreover have \m + 3 < 4 * 2 ^ m\ by (induction m) simp_all ultimately show ?thesis by simp qed end diff --git a/thys/Zeta_Function/Zeta_Library.thy b/thys/Zeta_Function/Zeta_Library.thy --- a/thys/Zeta_Function/Zeta_Library.thy +++ b/thys/Zeta_Function/Zeta_Library.thy @@ -1,898 +1,898 @@ section \Various preliminary material\ theory Zeta_Library imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" "Dirichlet_Series.Dirichlet_Series_Analysis" begin subsection \Facts about limits\ lemma at_within_altdef: "at x within A = (INF S\{S. open S \ x \ S}. principal (S \ (A - {x})))" unfolding at_within_def nhds_def inf_principal [symmetric] by (subst INF_inf_distrib [symmetric]) (auto simp: INF_constant) lemma tendsto_at_left_realI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_left c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_left c)" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \dx\{d<.. A m" proof (rule ccontr) assume "\ (\m. \dx\{d<.. A m)" then obtain m where **: "\d. d < c \ \x\{d<.. A m" by auto have "\X. \n. (f (X n) \ A m \ X n < c) \ X (Suc n) > c - max 0 ((c - X n) / 2)" proof (intro dependent_nat_choice, goal_cases) case 1 from **[of "c - 1"] show ?case by auto next case (2 x n) with **[of "c - max 0 (c - x) / 2"] show ?case by force qed then obtain X where X: "\n. f (X n) \ A m" "\n. X n < c" "\n. X (Suc n) > c - max 0 ((c - X n) / 2)" - by auto + by auto (metis diff_gt_0_iff_gt half_gt_zero_iff max.absorb3 max.commute) have X_ge: "X n \ c - (c - X 0) / 2 ^ n" for n proof (induction n) case (Suc n) have "c - (c - X 0) / 2 ^ Suc n = c - (c - (c - (c - X 0) / 2 ^ n)) / 2" by simp also have "c - (c - (c - (c - X 0) / 2 ^ n)) / 2 \ c - (c - X n) / 2" by (intro diff_left_mono divide_right_mono Suc diff_right_mono) auto also have "\ = c - max 0 ((c - X n) / 2)" using X[of n] by (simp add: max_def) also have "\ < X (Suc n)" using X[of n] by simp finally show ?case by linarith qed auto have "X \ c" proof (rule tendsto_sandwich) show "eventually (\n. X n \ c) sequentially" using X by (intro always_eventually) (auto intro!: less_imp_le) show "eventually (\n. X n \ c - (c - X 0) / 2 ^ n) sequentially" using X_ge by (intro always_eventually) auto qed real_asymp+ hence "filterlim X (at_left c) sequentially" by (rule tendsto_imp_filterlim_at_left) (use X in \auto intro!: always_eventually less_imp_le\) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain d where d: "d m < c" "x \ {d m<.. f x \ A m" for m x by metis have ***: "at_left c = (INF S\{S. open S \ c \ S}. principal (S \ {..m. {d m<..}"]) auto qed lemma shows at_right_PInf [simp]: "at_right (\ :: ereal) = bot" and at_left_MInf [simp]: "at_left (-\ :: ereal) = bot" proof - have "{(\::ereal)<..} = {}" "{..<-(\::ereal)} = {}" by auto thus "at_right (\ :: ereal) = bot" "at_left (-\ :: ereal) = bot" by (simp_all add: at_within_def) qed lemma tendsto_at_left_erealI_sequentially: fixes f :: "ereal \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_left c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_left c)" proof (cases c) case [simp]: PInf have "((\x. f (ereal x)) \ y) at_top" using assms by (intro tendsto_at_topI_sequentially assms) (simp_all flip: ereal_tendsto_simps add: o_def filterlim_at) thus ?thesis by (simp add: at_left_PInf filterlim_filtermap) next case [simp]: MInf thus ?thesis by auto next case [simp]: (real c') have "((\x. f (ereal x)) \ y) (at_left c')" proof (intro tendsto_at_left_realI_sequentially assms) fix X assume *: "filterlim X (at_left c') sequentially" show "filterlim (\n. ereal (X n)) (at_left c) sequentially" by (rule filterlim_compose[OF _ *]) (simp add: sequentially_imp_eventually_within tendsto_imp_filterlim_at_left) qed thus ?thesis by (simp add: at_left_ereal filterlim_filtermap) qed lemma tendsto_at_right_realI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_right c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_right c)" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \d>c. \x\{c<.. A m" proof (rule ccontr) assume "\ (\m. \d>c. \x\{c<.. A m)" then obtain m where **: "\d. d > c \ \x\{c<.. A m" by auto have "\X. \n. (f (X n) \ A m \ X n > c) \ X (Suc n) < c + max 0 ((X n - c) / 2)" proof (intro dependent_nat_choice, goal_cases) case 1 from **[of "c + 1"] show ?case by auto next case (2 x n) with **[of "c + max 0 (x - c) / 2"] show ?case by force qed then obtain X where X: "\n. f (X n) \ A m" "\n. X n > c" "\n. X (Suc n) < c + max 0 ((X n - c) / 2)" - by auto + by auto (metis add.left_neutral half_gt_zero_iff less_diff_eq max.absorb4) have X_le: "X n \ c + (X 0 - c) / 2 ^ n" for n proof (induction n) case (Suc n) have "X (Suc n) < c + max 0 ((X n - c) / 2)" by (intro X) also have "\ = c + (X n - c) / 2" using X[of n] by (simp add: field_simps max_def) also have "\ \ c + (c + (X 0 - c) / 2 ^ n - c) / 2" by (intro add_left_mono divide_right_mono Suc diff_right_mono) auto also have "\ = c + (X 0 - c) / 2 ^ Suc n" by simp finally show ?case by linarith qed auto have "X \ c" proof (rule tendsto_sandwich) show "eventually (\n. X n \ c) sequentially" using X by (intro always_eventually) (auto intro!: less_imp_le) show "eventually (\n. X n \ c + (X 0 - c) / 2 ^ n) sequentially" using X_le by (intro always_eventually) auto qed real_asymp+ hence "filterlim X (at_right c) sequentially" by (rule tendsto_imp_filterlim_at_right) (use X in \auto intro!: always_eventually less_imp_le\) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain d where d: "d m > c" "x \ {c<.. f x \ A m" for m x by metis have ***: "at_right c = (INF S\{S. open S \ c \ S}. principal (S \ {c<..}))" by (simp add: at_within_altdef) from d show ?thesis unfolding *** A using A(1,2) by (intro filterlim_base[of _ "\m. {.. 'b::first_countable_topology" assumes *: "\X. filterlim X (at_right c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_right c)" proof (cases c) case [simp]: MInf have "((\x. f (-ereal x)) \ y) at_top" using assms by (intro tendsto_at_topI_sequentially assms) (simp_all flip: uminus_ereal.simps ereal_tendsto_simps add: o_def filterlim_at) thus ?thesis by (simp add: at_right_MInf filterlim_filtermap at_top_mirror) next case [simp]: PInf thus ?thesis by auto next case [simp]: (real c') have "((\x. f (ereal x)) \ y) (at_right c')" proof (intro tendsto_at_right_realI_sequentially assms) fix X assume *: "filterlim X (at_right c') sequentially" show "filterlim (\n. ereal (X n)) (at_right c) sequentially" by (rule filterlim_compose[OF _ *]) (simp add: sequentially_imp_eventually_within tendsto_imp_filterlim_at_right) qed thus ?thesis by (simp add: at_right_ereal filterlim_filtermap) qed proposition analytic_continuation': assumes hol: "f holomorphic_on S" "g holomorphic_on S" and "open S" and "connected S" and "U \ S" and "\ \ S" and "\ islimpt U" and fU0 [simp]: "\z. z \ U \ f z = g z" and "w \ S" shows "f w = g w" using analytic_continuation[OF holomorphic_on_diff[OF hol] assms(3-7) _ assms(9)] assms(8) by simp subsection \Various facts about integrals\ lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" assumes "set_integrable M A f" "set_borel_measurable M A g" assumes "AE x in M. x \ A \ norm (g x) \ norm (f x)" shows "set_integrable M A g" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) from assms(1) show "integrable M (\x. indicator A x *\<^sub>R f x)" by (simp add: set_integrable_def) from assms(2) show "(\x. indicat_real A x *\<^sub>R g x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(3) show "AE x in M. norm (indicat_real A x *\<^sub>R g x) \ norm (indicat_real A x *\<^sub>R f x)" by eventually_elim (simp add: indicator_def) qed (* TODO replace version in library. Better name? *) lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" using nonneg by (simp add: indicator_def of_bool_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto also have "(\x. abs (indicator \ x * f x)) = (\x. indicator \ x * f x)" using nonneg by (auto simp: indicator_def fun_eq_iff) finally show ?thesis . qed subsection \Uniform convergence of integrals\ lemma has_absolute_integral_change_of_variables_1': fixes f :: "real \ real" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "(\x. \g' x\ *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1) \ (\x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \ integral (g ` S) (\x. vec (f x)) = (vec b :: real ^ 1)" using assms unfolding has_field_derivative_iff_has_vector_derivative by (intro has_absolute_integral_change_of_variables_1 assms) auto thus ?thesis by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq) qed lemma set_nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes "set_borel_measurable borel A f" assumes "\x. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \lborel) < \" shows "(\\<^sup>+x\A. f x \lborel) = integral A f" proof - have eq: "(\\<^sup>+x\A. f x \lborel) = (\\<^sup>+x. ennreal (indicator A x * f x) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = integral UNIV (\x. indicator A x * f x)" using assms eq by (intro nn_integral_lborel_eq_integral) (auto simp: indicator_def set_borel_measurable_def) also have "integral UNIV (\x. indicator A x * f x) = integral A (\x. indicator A x * f x)" by (rule integral_spike_set) (auto intro: empty_imp_negligible) also have "\ = integral A f" by (rule integral_cong) (auto simp: indicator_def) finally show ?thesis . qed lemma nn_integral_has_integral_lebesgue': fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. ennreal (f x) * indicator \ x) = ennreal I" proof - have "integral\<^sup>N lborel (\x. ennreal (f x) * indicator \ x) = integral\<^sup>N lborel (\x. ennreal (indicator \ x * f x))" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal I" using assms by (intro nn_integral_has_integral_lebesgue) finally show ?thesis . qed lemma uniform_limit_set_lebesgue_integral: fixes f :: "'a \ 'b :: euclidean_space \ 'c :: {banach, second_countable_topology}" assumes "set_integrable lborel X' g" assumes [measurable]: "X' \ sets borel" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel X' (f y)" assumes "\y. y \ Y \ (AE t\X' in lborel. norm (f y t) \ g t)" assumes "eventually (\x. X x \ sets borel \ X x \ X') F" assumes "filterlim (\x. set_lebesgue_integral lborel (X x) g) (nhds (set_lebesgue_integral lborel X' g)) F" shows "uniform_limit Y (\x y. set_lebesgue_integral lborel (X x) (f y)) (\y. set_lebesgue_integral lborel X' (f y)) F" proof (rule uniform_limitI, goal_cases) case (1 \) have integrable_g: "set_integrable lborel U g" if "U \ sets borel" "U \ X'" for U by (rule set_integrable_subset[OF assms(1)]) (use that in auto) have "eventually (\x. dist (set_lebesgue_integral lborel (X x) g) (set_lebesgue_integral lborel X' g) < \) F" using \\ > 0\ assms by (auto simp: tendsto_iff) from this show ?case using \eventually (\_. _ \ _) F\ proof eventually_elim case (elim x) hence [measurable]:"X x \ sets borel" and "X x \ X'" by auto have integrable: "set_integrable lborel U (f y)" if "y \ Y" "U \ sets borel" "U \ X'" for y U apply (rule set_integrable_subset) apply (rule set_integrable_bound[OF assms(1)]) apply (use assms(3) that in \simp add: set_borel_measurable_def\) using assms(4)[OF \y \ Y\] apply eventually_elim apply force using that apply simp_all done show ?case proof fix y assume "y \ Y" have "dist (set_lebesgue_integral lborel (X x) (f y)) (set_lebesgue_integral lborel X' (f y)) = norm (set_lebesgue_integral lborel X' (f y) - set_lebesgue_integral lborel (X x) (f y))" by (simp add: dist_norm norm_minus_commute) also have "set_lebesgue_integral lborel X' (f y) - set_lebesgue_integral lborel (X x) (f y) = set_lebesgue_integral lborel (X' - X x) (f y)" unfolding set_lebesgue_integral_def apply (subst Bochner_Integration.integral_diff [symmetric]) unfolding set_integrable_def [symmetric] apply (rule integrable; (fact | simp)) apply (rule integrable; fact) apply (intro Bochner_Integration.integral_cong) apply (use \X x \ X'\ in \auto simp: indicator_def\) done also have "norm \ \ (\t\X'-X x. norm (f y t) \lborel)" by (intro set_integral_norm_bound integrable) (fact | simp)+ also have "AE t\X' - X x in lborel. norm (f y t) \ g t" using assms(4)[OF \y \ Y\] by eventually_elim auto with \y \ Y\ have "(\t\X'-X x. norm (f y t) \lborel) \ (\t\X'-X x. g t \lborel)" by (intro set_integral_mono_AE set_integrable_norm integrable integrable_g) auto also have "\ = (\t\X'. g t \lborel) - (\t\X x. g t \lborel)" unfolding set_lebesgue_integral_def apply (subst Bochner_Integration.integral_diff [symmetric]) unfolding set_integrable_def [symmetric] apply (rule integrable_g; (fact | simp)) apply (rule integrable_g; fact) apply (intro Bochner_Integration.integral_cong) apply (use \X x \ X'\ in \auto simp: indicator_def\) done also have "\ \ dist (\t\X x. g t \lborel) (\t\X'. g t \lborel)" by (simp add: dist_norm) also have "\ < \" by fact finally show "dist (set_lebesgue_integral lborel (X x) (f y)) (set_lebesgue_integral lborel X' (f y)) < \" . qed qed qed lemma integral_dominated_convergence_at_right: fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M and c :: real assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) (at_right c)" assumes bound: "\\<^sub>F i in at_right c. AE x in M. norm (s i x) \ w x" shows "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) (at_right c)" proof (rule tendsto_at_right_realI_sequentially) fix X :: "nat \ real" assume X: "filterlim X (at_right c) sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) (at_right c)" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma integral_dominated_convergence_at_left: fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M and c :: real assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) (at_left c)" assumes bound: "\\<^sub>F i in at_left c. AE x in M. norm (s i x) \ w x" shows "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) (at_left c)" proof (rule tendsto_at_left_realI_sequentially) fix X :: "nat \ real" assume X: "filterlim X (at_left c) sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) (at_left c)" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma uniform_limit_interval_integral_right: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes "a < b" shows "uniform_limit Y (\b' y. LBINT t=a..b'. f y t) (\y. LBINT t=a..b. f y t) (at_left b)" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\b'. b' \ {a<..a < b\ by (intro eventually_at_leftI) with \a < b\ have "?thesis \ uniform_limit Y (\b' y. \t\einterval a (min b b'). f y t \lborel) (\y. \t\einterval a b. f y t \lborel) (at_left b)" by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff min_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F b' in at_left b. einterval a (min b b') \ sets borel \ einterval a (min b b') \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\b'. set_lebesgue_integral lborel (einterval a (min b b')) g) \ set_lebesgue_integral lborel (einterval a b) g) (at_left b)" unfolding set_lebesgue_integral_def proof (intro tendsto_at_left_erealI_sequentially integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix X :: "nat \ ereal" and n :: nat have "set_borel_measurable borel (einterval a (min b (X n))) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next fix X :: "nat \ ereal" assume X: "filterlim X (at_left b) sequentially" show "AE x in lborel. (\n. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval a (min b (X t))) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval a (min b (X t))"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {max a x<..t. X t \ {max a x<..t. indicator (einterval a (min b (X t))) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval a (min b (X n))) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma uniform_limit_interval_integral_left: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes "a < b" shows "uniform_limit Y (\a' y. LBINT t=a'..b. f y t) (\y. LBINT t=a..b. f y t) (at_right a)" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\b'. b' \ {a<..a < b\ by (intro eventually_at_rightI) with \a < b\ have "?thesis \ uniform_limit Y (\a' y. \t\einterval (max a a') b. f y t \lborel) (\y. \t\einterval a b. f y t \lborel) (at_right a)" by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff max_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F a' in at_right a. einterval (max a a') b \ sets borel \ einterval (max a a') b \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\a'. set_lebesgue_integral lborel (einterval (max a a') b) g) \ set_lebesgue_integral lborel (einterval a b) g) (at_right a)" unfolding set_lebesgue_integral_def proof (intro tendsto_at_right_erealI_sequentially integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix X :: "nat \ ereal" and n :: nat have "set_borel_measurable borel (einterval (max a (X n)) b) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next fix X :: "nat \ ereal" assume X: "filterlim X (at_right a) sequentially" show "AE x in lborel. (\n. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval (max a (X t)) b) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval (max a (X t)) b"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {a<..t. X t \ {a<..t. indicator (einterval (max a (X t)) b) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma uniform_limit_interval_integral_sequentially: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes a': "filterlim a' (at_right a) sequentially" assumes b': "filterlim b' (at_left b) sequentially" assumes "a < b" shows "uniform_limit Y (\n y. LBINT t=a' n..b' n. f y t) (\y. LBINT t=a..b. f y t) sequentially" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\n. a < a' n \ a' n < b' n \ b' n < b) sequentially" proof - from ereal_dense2[OF \a < b\] obtain t where t: "a < ereal t" "ereal t < b" by blast from t have "eventually (\n. a' n \ {a<..n. b' n \ {t<..n. a < a' n \ a' n < b' n \ b' n < b) sequentially" by eventually_elim auto qed have "?thesis \ uniform_limit Y (\n y. \t\einterval (max a (a' n)) (min b (b' n)). f y t \lborel) (\y. \t\einterval a b. f y t \lborel) sequentially" using \a < b\ by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff min_def max_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F n in sequentially. einterval (max a (a' n)) (min b (b' n)) \ sets borel \ einterval (max a (a' n)) (min b (b' n)) \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\n. set_lebesgue_integral lborel (einterval (max a (a' n)) (min b (b' n))) g) \ set_lebesgue_integral lborel (einterval a b) g) sequentially" unfolding set_lebesgue_integral_def proof (intro integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix n :: nat have "set_borel_measurable borel (einterval (max a (a' n)) (min b (b' n))) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next show "AE x in lborel. (\n. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval (max a (a' t)) (min b (b' t))) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval (max a (a' t)) (min b (b' t))"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {a<..n. x \ {a' n<..n. a' n \ {a<..n. b' n \ {x<..n. x \ {a' n<..t. indicator (einterval (max a (a' t)) (min b (b' t))) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma interval_lebesgue_integrable_combine: assumes "interval_lebesgue_integrable lborel A B f" assumes "interval_lebesgue_integrable lborel B C f" assumes "set_borel_measurable borel (einterval A C) f" assumes "A \ B" "B \ C" shows "interval_lebesgue_integrable lborel A C f" proof - have meas: "set_borel_measurable borel (einterval A B \ einterval B C) f" by (rule set_borel_measurable_subset[OF assms(3)]) (use assms in \auto simp: einterval_def\) have "set_integrable lborel (einterval A B \ einterval B C) f" using assms by (intro set_integrable_Un) (auto simp: interval_lebesgue_integrable_def) also have "?this \ set_integrable lborel (einterval A C) f" proof (cases "B \ {\, -\}") case True with assms have "einterval A B \ einterval B C = einterval A C" by (auto simp: einterval_def) thus ?thesis by simp next case False then obtain B' where [simp]: "B = ereal B'" by (cases B) auto have "indicator (einterval A C) x = (indicator (einterval A B \ einterval B C) x :: real)" if "x \ B'" for x using assms(4,5) that by (cases A; cases C) (auto simp: einterval_def indicator_def) hence "{x \ space lborel. indicat_real (einterval A B \ einterval B C) x *\<^sub>R f x \ indicat_real (einterval A C) x *\<^sub>R f x} \ {B'}" by force thus ?thesis unfolding set_integrable_def using meas assms by (intro integrable_cong_AE AE_I[of _ _ "{B'}"]) (simp_all add: set_borel_measurable_def) qed also have "\ \ ?thesis" using order.trans[OF assms(4,5)] by (simp add: interval_lebesgue_integrable_def) finally show ?thesis . qed lemma interval_lebesgue_integrable_bigo_right: fixes A B :: real fixes f :: "real \ real" assumes "f \ O[at_left B](g)" assumes cont: "continuous_on {A.. 0" "eventually (\x. norm (f x) \ c * norm (g x)) (at_left B)" by (elim landau_o.bigE) then obtain B' where B': "B' < B" "\x. x \ {B'<.. norm (f x) \ c * norm (g x)" using \A < B\ by (auto simp: Topological_Spaces.eventually_at_left[of A]) show ?thesis proof (rule interval_lebesgue_integrable_combine) show "interval_lebesgue_integrable lborel A (max A B') f" using B' assms by (intro interval_integrable_continuous_on continuous_on_subset[OF cont]) auto show "set_borel_measurable borel (einterval (ereal A) (ereal B)) f" using assms by simp have meas': "set_borel_measurable borel {max A B'<..x. c * g x)" using assms by (simp add: interval_lebesgue_integrable_def) thus "set_integrable lborel {max A B'<..x. c * g x)" by (rule set_integrable_subset) auto next fix x assume "x \ {max A B'<.. c * norm (g x)" by (intro B') auto also have "\ \ norm (c * g x)" unfolding norm_mult by (intro mult_right_mono) auto finally show "norm (f x) \ norm (c * g x)" . qed (use meas' in \simp_all add: set_borel_measurable_def\) thus "interval_lebesgue_integrable lborel (ereal (max A B')) (ereal B) f" unfolding interval_lebesgue_integrable_def einterval_eq_Icc using \B' < B\ assms by simp qed (use B' assms in auto) qed lemma interval_lebesgue_integrable_bigo_left: fixes A B :: real fixes f :: "real \ real" assumes "f \ O[at_right A](g)" assumes cont: "continuous_on {A<..B} f" assumes meas: "set_borel_measurable borel {A<.. 0" "eventually (\x. norm (f x) \ c * norm (g x)) (at_right A)" by (elim landau_o.bigE) then obtain A' where A': "A' > A" "\x. x \ {A<.. norm (f x) \ c * norm (g x)" using \A < B\ by (auto simp: Topological_Spaces.eventually_at_right[of A]) show ?thesis proof (rule interval_lebesgue_integrable_combine) show "interval_lebesgue_integrable lborel (min B A') B f" using A' assms by (intro interval_integrable_continuous_on continuous_on_subset[OF cont]) auto show "set_borel_measurable borel (einterval (ereal A) (ereal B)) f" using assms by simp have meas': "set_borel_measurable borel {A<..x. c * g x)" using assms by (simp add: interval_lebesgue_integrable_def) thus "set_integrable lborel {A<..x. c * g x)" by (rule set_integrable_subset) auto next fix x assume "x \ {A<.. c * norm (g x)" by (intro A') auto also have "\ \ norm (c * g x)" unfolding norm_mult by (intro mult_right_mono) auto finally show "norm (f x) \ norm (c * g x)" . qed (use meas' in \simp_all add: set_borel_measurable_def\) thus "interval_lebesgue_integrable lborel (ereal A) (ereal (min B A')) f" unfolding interval_lebesgue_integrable_def einterval_eq_Icc using \A' > A\ assms by simp qed (use A' assms in auto) qed subsection \Other material\ (* TODO: Library *) lemma summable_comparison_test_bigo: fixes f :: "nat \ real" assumes "summable (\n. norm (g n))" "f \ O(g)" shows "summable f" proof - from \f \ O(g)\ obtain C where C: "eventually (\x. norm (f x) \ C * norm (g x)) at_top" by (auto elim: landau_o.bigE) thus ?thesis by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult) qed lemma fps_expansion_cong: assumes "eventually (\x. g x = h x) (nhds x)" shows "fps_expansion g x = fps_expansion h x" proof - have "(deriv ^^ n) g x = (deriv ^^ n) h x" for n by (intro higher_deriv_cong_ev assms refl) thus ?thesis by (simp add: fps_expansion_def) qed lemma fps_expansion_eq_zero_iff: assumes "g holomorphic_on ball z r" "r > 0" shows "fps_expansion g z = 0 \ (\z\ball z r. g z = 0)" proof assume *: "\z\ball z r. g z = 0" have "eventually (\w. w \ ball z r) (nhds z)" using assms by (intro eventually_nhds_in_open) auto hence "eventually (\z. g z = 0) (nhds z)" by eventually_elim (use * in auto) hence "fps_expansion g z = fps_expansion (\_. 0) z" by (intro fps_expansion_cong) thus "fps_expansion g z = 0" by (simp add: fps_expansion_def fps_zero_def) next assume *: "fps_expansion g z = 0" have "g w = 0" if "w \ ball z r" for w by (rule holomorphic_fun_eq_0_on_ball[OF assms(1) that]) (use * in \auto simp: fps_expansion_def fps_eq_iff\) thus "\w\ball z r. g w = 0" by blast qed lemma fds_nth_higher_deriv: "fds_nth ((fds_deriv ^^ k) F) = (\n. (-1) ^ k * of_real (ln n) ^ k * fds_nth F n)" by (induction k) (auto simp: fds_nth_deriv fun_eq_iff simp flip: scaleR_conv_of_real) lemma binomial_n_n_minus_one [simp]: "n > 0 \ n choose (n - Suc 0) = n" by (cases n) auto lemma has_field_derivative_complex_powr_right: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z within A)" by (rule DERIV_subset, rule has_field_derivative_powr_right) auto lemmas has_field_derivative_complex_powr_right' = has_field_derivative_complex_powr_right[THEN DERIV_chain2] end \ No newline at end of file diff --git a/thys/pGCL/Tutorial/Primitives.thy b/thys/pGCL/Tutorial/Primitives.thy --- a/thys/pGCL/Tutorial/Primitives.thy +++ b/thys/pGCL/Tutorial/Primitives.thy @@ -1,165 +1,165 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section "Language Primitives" theory Primitives imports "../pGCL" begin text \Programs in pGCL are probabilistic automata. They can do anything a traditional program can, plus, they may make truly probabilistic choices.\ subsection \The Basics\ text \Imagine flipping a pair of fair coins: @{term a} and @{term b}. Using a record type for the state allows a number of syntactic niceties, which we describe shortly:\ datatype coin = Heads | Tails record coins = a :: coin b :: coin text \The primitive state operation is @{term Apply}, which takes a state transformer as an argument, constructs the pGCL equivalent. Thus @{term "Apply (\s. s\ a := Heads \)"} sets the value of coin @{term a} to @{term Heads}. As records are so common as state types, we introduce syntax to make these update neater: The same program may be defined more simply as @{term "a := (\s. Heads)"} (note that the syntax translation involved does not apply to Latex output, and thus this lemma appears trivial): \ lemma "Apply (\s. s \ a := Heads \) = (a := (\s. Heads))" by(simp) text \We can treat the record's fields as the names of \emph{variables}. Note that the right-hand side of an assignment is always a function of the current state. Thus we may use a record accessor directly, for example @{term "a := b"}, which updates @{term a} with the current value of @{term b}. If we wish to formally establish that the previous statement is correct i.e. that in the final state, @{term a} really will have whatever value @{term b} had in the initial state, we must first introduce the assertion language.\ subsection \Assertion and Annotation\ text \Assertions in pGCL are real-valued functions of the state, which are often interpreted as a probability distribution over possible outcomes. These functions are termed \emph{expectations}, for reasons which shortly be clear. Initially, however, we need only consider \emph{standard} expectations: those derived from a binary predicate. A predicate @{term_type "P::'s \ bool"} is embedded as @{term_type "\P::'s \ bool\"}, such that @{term "P s \ \P\ s = 1 \ \ P s \ \P\ s = 0"}. An annotation consists of an assertion on the initial state and one on the final state, which for standard expectations may be interpreted as `if @{term P} holds in the initial state, then @{term Q} will hold in the final state'. These are in weakest-precondition form: we assert that the precondition implies the \emph{weakest precondition}: the weakest assertion on the initial state, which implies that the postcondition must hold on the final state. So far, this is identical to the standard approach. Remember, however, that we are working with \emph{real-valued} assertions. For standard expectations, the logic is nevertheless identical, if the implication @{term "\s. P s \ Q s"} is substituted with the equivalent expectation entailment @{term "\P\ \ \Q\"}, @{thm entails_implies}. Thus a valid specification of @{term "a := b"} is:\ lemma "\x. \\s. b s = x\ \ wp (a := b) \\s. a s = x\" by(pvcg, simp add:o_def) text \Any ordinary computation and its associated annotation can be expressed in this form.\ subsection \Probability\ text \Next, we introduce the syntax @{term "x ;; y"} for the sequential composition of @{term x} and @{term y}, and also demonstrate that one can operate directly on a real-valued (and thus infinite) state space:\ lemma "\\s::real. s \ 0\ \ wp (Apply ((*) 2) ;; Apply (\s. s / s)) \\s. s = 1\" by(pvcg, simp add:o_def) text \So far, we haven't done anything that required probabilities, or expectations other than 0 and 1. As an example of both, we show that a single coin toss is fair. We introduce the syntax @{term "x \<^bsub>p\<^esub>\ y"} for a probabilistic choice between @{term x} and @{term y}. This program behaves as @{term x} with probability @{term p}, and as @{term y} with probability @{term "1-p"}. The probability may depend on the state, and is therefore of type @{typ "'s \ real"}. The following annotation states that the probability of heads is exactly 1/2:\ definition flip_a :: "real \ coins prog" where "flip_a p = a := (\_. Heads) \<^bsub>(\s. p)\<^esub>\ a := (\_. Tails)" lemma "(\s. 1/2) = wp (flip_a (1/2)) \\s. a s = Heads\" unfolding flip_a_def txt \Sufficiently small problems can be handled by the simplifier, by symbolic evaluation.\ by(simp add:wp_eval o_def) subsection \Nondeterminism\ text \We can also under-specify a program, using the \emph{nondeterministic choice} operator, @{term "x \ y"}. This is interpreted demonically, giving the pointwise \emph{minimum} of the pre-expectations for @{term x} and @{term y}: the chance of seeing heads, if your opponent is allowed choose between a pair of coins, one biased 2/3 heads and one 2/3 tails, and then flips it, is \emph{at least} 1/3, but we can make no stronger statement:\ lemma "\s. 1/3 \ wp (flip_a (2/3) \ flip_a (1/3)) \\s. a s = Heads\" unfolding flip_a_def - by(pvcg, simp add:o_def le_funI) + by pvcg subsection \Properties of Expectations\ text \The probabilities of independent events combine as usual, by multiplying: The chance of getting heads on two separate coins is @{term "1/4"}.\ definition flip_b :: "real \ coins prog" where "flip_b p = b := (\_. Heads) \<^bsub>(\s. p)\<^esub>\ b := (\_. Tails)" lemma "(\s. 1/4) = wp (flip_a (1/2) ;; flip_b (1/2)) \\s. a s = Heads \ b s = Heads\" unfolding flip_a_def flip_b_def by(simp add:wp_eval o_def) text \If, rather than two coins, we use two dice, we can make some slightly more involved calculations. We see that the weakest pre-expectation of the value on the face of the die after rolling is its \emph{expected value} in the initial state, which justifies the use of the term expectation. \ record dice = red :: nat blue :: nat definition Puniform :: "'a set \ ('a \ real)" where "Puniform S = (\x. if x \ S then 1 / card S else 0)" lemma Puniform_in: "x \ S \ Puniform S x = 1 / card S" by(simp add:Puniform_def) lemma Puniform_out: "x \ S \ Puniform S x = 0" by(simp add:Puniform_def) lemma supp_Puniform: "finite S \ supp (Puniform S) = S" by(auto simp:Puniform_def supp_def) text \The expected value of a roll of a six-sided die is @{term "7/2"}:\ lemma "(\s. 7/2) = wp (bind v at (\s. Puniform {1..6} v) in red := (\_. v)) red" by(simp add:wp_eval supp_Puniform sum.atLeast_Suc_atMost Puniform_in) text \The expectations of independent variables add:\ lemma "(\s. 7) = wp ((bind v at (\s. Puniform {1..6} v) in red := (\s. v)) ;; (bind v at (\s. Puniform {1..6} v) in blue := (\s. v))) (\s. red s + blue s)" by(simp add:wp_eval supp_Puniform sum.atLeast_Suc_atMost Puniform_in) end