diff --git a/thys/CAVA_Automata/Automata.thy b/thys/CAVA_Automata/Automata.thy --- a/thys/CAVA_Automata/Automata.thy +++ b/thys/CAVA_Automata/Automata.thy @@ -1,1431 +1,1433 @@ section \Automata\ (* Author: Peter Lammich *) theory Automata imports Digraph begin text \ In this theory, we define Generalized Buchi Automata and Buchi Automata based on directed graphs \ hide_const (open) prod subsection "Generalized Buchi Graphs" text \ A generalized Buchi graph is a graph where each node belongs to a set of acceptance classes. An infinite run on this graph is accepted, iff it visits nodes from each acceptance class infinitely often. The standard encoding of acceptance classes is as a set of sets of nodes, each inner set representing one acceptance class. \ record 'Q gb_graph_rec = "'Q graph_rec" + gbg_F :: "'Q set set" locale gb_graph = graph G for G :: "('Q,'more) gb_graph_rec_scheme" + assumes finite_F[simp, intro!]: "finite (gbg_F G)" assumes F_ss: "gbg_F G \ Pow V" begin abbreviation "F \ gbg_F G" lemma is_gb_graph: "gb_graph G" by unfold_locales definition is_acc :: "'Q word \ bool" where "is_acc r \ (\A\F. \\<^sub>\i. r i \ A)" definition "is_acc_run r \ is_run r \ is_acc r" (* For presentation in paper *) lemma "is_acc_run r \ is_run r \ (\A\F. \\<^sub>\i. r i \ A)" unfolding is_acc_run_def is_acc_def . lemma acc_run_run: "is_acc_run r \ is_run r" unfolding is_acc_run_def by simp lemmas acc_run_reachable = run_reachable[OF acc_run_run] lemma acc_eq_limit: assumes FIN: "finite (range r)" shows "is_acc r \ (\A\F. limit r \ A \ {})" proof assume "\A\F. limit r \ A \ {}" thus "is_acc r" unfolding is_acc_def by (metis limit_inter_INF) next from FIN have FIN': "\A. finite (A \ range r)" by simp assume "is_acc r" hence AUX: "\A\F. \\<^sub>\i. r i \ (A \ range r)" unfolding is_acc_def by auto have "\A\F. limit r \ (A \ range r) \ {}" apply (rule ballI) apply (drule bspec[OF AUX]) apply (subst (asm) fin_ex_inf_eq_limit[OF FIN']) . thus "\A\F. limit r \ A \ {}" by auto qed lemma is_acc_run_limit_alt: assumes "finite (E\<^sup>* `` V0)" shows "is_acc_run r \ is_run r \ (\A\F. limit r \ A \ {})" using assms acc_eq_limit[symmetric] unfolding is_acc_run_def by (auto dest: run_reachable finite_subset) lemma is_acc_suffix[simp]: "is_acc (suffix i r) \ is_acc r" unfolding is_acc_def suffix_def apply (clarsimp simp: INFM_nat) apply (rule iffI) apply (metis trans_less_add2) by (metis add_lessD1 less_imp_add_positive nat_add_left_cancel_less) lemma finite_V_Fe: assumes "finite V" "A \ F" shows "finite A" using assms by (metis Pow_iff infinite_super rev_subsetD F_ss) end definition "gb_rename_ecnv ecnv f G \ \ gbg_F = { f`A | A. A\gbg_F G }, \ = ecnv G \" abbreviation "gb_rename_ext ecnv f \ fr_rename_ext (gb_rename_ecnv ecnv f) f" locale gb_rename_precond = gb_graph G + g_rename_precond G f "gb_rename_ecnv ecnv f" for G :: "('u,'more) gb_graph_rec_scheme" and f :: "'u \ 'v" and ecnv begin lemma G'_gb_fields: "gbg_F G' = { f`A | A. A\F }" unfolding gb_rename_ecnv_def fr_rename_ext_def by simp sublocale G': gb_graph G' apply unfold_locales apply (simp_all add: G'_fields G'_gb_fields) using F_ss by auto lemma acc_sim1: "is_acc r \ G'.is_acc (f o r)" unfolding is_acc_def G'.is_acc_def G'_gb_fields by (fastforce intro: imageI simp: INFM_nat) lemma acc_sim2: assumes "G'.is_acc r" shows "is_acc (fi o r)" proof - from assms have 1: "\A m. A \ gbg_F G \ \i>m. r i \ f`A" unfolding G'.is_acc_def G'_gb_fields by (auto simp: INFM_nat) { fix A m assume 2: "A \ gbg_F G" from 1[OF this, of m] have "\i>m. fi (r i) \ A" using F_ss apply clarsimp by (metis Pow_iff 2 fi_f in_mono) } thus ?thesis unfolding is_acc_def by (auto simp: INFM_nat) qed lemma acc_run_sim1: "is_acc_run r \ G'.is_acc_run (f o r)" using acc_sim1 run_sim1 unfolding G'.is_acc_run_def is_acc_run_def by auto lemma acc_run_sim2: "G'.is_acc_run r \ is_acc_run (fi o r)" using acc_sim2 run_sim2 unfolding G'.is_acc_run_def is_acc_run_def by auto end subsection "Generalized Buchi Automata" text \ A GBA is obtained from a GBG by adding a labeling function, that associates each state with a set of labels. A word is accepted if there is an accepting run that can be labeld with this word. \ record ('Q,'L) gba_rec = "'Q gb_graph_rec" + gba_L :: "'Q \ 'L \ bool" locale gba = gb_graph G for G :: "('Q,'L,'more) gba_rec_scheme" + assumes L_ss: "gba_L G q l \ q \ V" begin abbreviation "L \ gba_L G" lemma is_gba: "gba G" by unfold_locales definition "accept w \ \r. is_acc_run r \ (\i. L (r i) (w i))" lemma acceptI[intro?]: "\is_acc_run r; \i. L (r i) (w i)\ \ accept w" by (auto simp: accept_def) definition "lang \ Collect (accept)" lemma langI[intro?]: "accept w \ w\lang" by (auto simp: lang_def) end definition "gba_rename_ecnv ecnv f G \ \ gba_L = \q l. if q\f`g_V G then gba_L G (the_inv_into (g_V G) f q) l else False, \ = ecnv G \" abbreviation "gba_rename_ext ecnv f \ gb_rename_ext (gba_rename_ecnv ecnv f) f" locale gba_rename_precond = gb_rename_precond G f "gba_rename_ecnv ecnv f" + gba G for G :: "('u,'L,'more) gba_rec_scheme" and f :: "'u \ 'v" and ecnv begin lemma G'_gba_fields: "gba_L G' = (\q l. if q\f`V then L (fi q) l else False)" unfolding gb_rename_ecnv_def gba_rename_ecnv_def fr_rename_ext_def fi_def by simp sublocale G': gba G' apply unfold_locales apply (auto simp add: G'_gba_fields G'_fields split: if_split_asm) done lemma L_sim1: "\range r \ V; L (r i) l\ \ G'.L (f (r i)) l" by (auto simp: G'_gba_fields fi_def[symmetric] fi_f dest: inj_onD[OF INJ] dest!: rev_subsetD[OF rangeI[of _ i]]) lemma L_sim2: "\ range r \ f`V; G'.L (r i) l \ \ L (fi (r i)) l" by (auto simp: G'_gba_fields fi_def[symmetric] f_fi dest!: rev_subsetD[OF rangeI[of _ i]]) lemma accept_eq[simp]: "G'.accept = accept" apply (rule ext) unfolding accept_def G'.accept_def proof safe fix w r assume R: "G'.is_acc_run r" assume L: "\i. G'.L (r i) (w i)" from R have RAN: "range r \ f`V" using G'.run_reachable[OF G'.acc_run_run[OF R]] G'.reachable_V unfolding G'_fields by simp from L show "\r. is_acc_run r \ (\i. L (r i) (w i))" using acc_run_sim2[OF R] L_sim2[OF RAN] by auto next fix w r assume R: "is_acc_run r" assume L: "\i. L (r i) (w i)" from R have RAN: "range r \ V" using run_reachable[OF acc_run_run[OF R]] reachable_V by simp from L show "\r. G'.is_acc_run r \ (\i. G'.L (r i) (w i))" using acc_run_sim1[OF R] L_sim1[OF RAN] by auto qed lemma lang_eq[simp]: "G'.lang = lang" unfolding G'.lang_def lang_def by simp lemma finite_G'_V: assumes "finite V" shows "finite G'.V" using assms by (auto simp add: G'_gba_fields G'_fields split: if_split_asm) end abbreviation "gba_rename \ gba_rename_ext (\_. ())" lemma gba_rename_correct: fixes G :: "('v,'l,'m) gba_rec_scheme" assumes "gba G" assumes INJ: "inj_on f (g_V G)" defines "G' \ gba_rename f G" shows "gba G'" and "finite (g_V G) \ finite (g_V G')" and "gba.accept G' = gba.accept G" and "gba.lang G' = gba.lang G" unfolding G'_def proof - let ?G' = "gba_rename f G" interpret gba G by fact from INJ interpret gba_rename_precond G f "\_. ()" by unfold_locales simp_all show "gba ?G'" by (rule G'.is_gba) show "finite (g_V G) \ finite (g_V ?G')" by (fact finite_G'_V) show "G'.accept = accept" by simp show "G'.lang = lang" by simp qed subsection "Buchi Graphs" text \A Buchi graph has exactly one acceptance class\ record 'Q b_graph_rec = "'Q graph_rec" + bg_F :: "'Q set" locale b_graph = graph G for G :: "('Q,'more) b_graph_rec_scheme" + assumes F_ss: "bg_F G \ V" begin abbreviation F where "F \ bg_F G" lemma is_b_graph: "b_graph G" by unfold_locales definition "to_gbg_ext m \ \ g_V = V, g_E=E, g_V0=V0, gbg_F = if F=UNIV then {} else {F}, \ = m \" abbreviation "to_gbg \ to_gbg_ext ()" sublocale gbg: gb_graph "to_gbg_ext m" apply unfold_locales using V0_ss E_ss F_ss apply (auto simp: to_gbg_ext_def split: if_split_asm) done definition is_acc :: "'Q word \ bool" where "is_acc r \ (\\<^sub>\i. r i \ F)" definition is_acc_run where "is_acc_run r \ is_run r \ is_acc r" lemma to_gbg_alt: "gbg.V T m = V" "gbg.E T m = E" "gbg.V0 T m = V0" "gbg.F T m = (if F=UNIV then {} else {F})" "gbg.is_run T m = is_run" "gbg.is_acc T m = is_acc" "gbg.is_acc_run T m = is_acc_run" unfolding is_run_def[abs_def] gbg.is_run_def[abs_def] is_acc_def[abs_def] gbg.is_acc_def[abs_def] is_acc_run_def[abs_def] gbg.is_acc_run_def[abs_def] by (auto simp: to_gbg_ext_def) end subsection "Buchi Automata" text \Buchi automata are labeled Buchi graphs\ record ('Q,'L) ba_rec = "'Q b_graph_rec" + ba_L :: "'Q \ 'L \ bool" locale ba = bg?: b_graph G for G :: "('Q,'L,'more) ba_rec_scheme" + assumes L_ss: "ba_L G q l \ q \ V" begin abbreviation L where "L == ba_L G" lemma is_ba: "ba G" by unfold_locales abbreviation "to_gba_ext m \ to_gbg_ext \ gba_L = L, \=m \" abbreviation "to_gba \ to_gba_ext ()" sublocale gba: gba "to_gba_ext m" apply unfold_locales unfolding to_gbg_ext_def using L_ss apply auto [] done lemma ba_acc_simps[simp]: "gba.L T m = L" by (simp add: to_gbg_ext_def) definition "accept w \ (\r. is_acc_run r \ (\i. L (r i) (w i)))" definition "lang \ Collect accept" lemma to_gba_alt_accept: "gba.accept T m = accept" apply (intro ext) unfolding accept_def gba.accept_def apply (simp_all add: to_gbg_alt) done lemma to_gba_alt_lang: "gba.lang T m = lang" unfolding lang_def gba.lang_def apply (simp_all add: to_gba_alt_accept) done lemmas to_gba_alt = to_gbg_alt to_gba_alt_accept to_gba_alt_lang end subsection "Indexed acceptance classes" record 'Q igb_graph_rec = "'Q graph_rec" + igbg_num_acc :: nat igbg_acc :: "'Q \ nat set" locale igb_graph = graph G for G :: "('Q,'more) igb_graph_rec_scheme" + assumes acc_bound: "\(range (igbg_acc G)) \ {0..<(igbg_num_acc G)}" assumes acc_ss: "igbg_acc G q \ {} \ q\V" begin abbreviation num_acc where "num_acc \ igbg_num_acc G" abbreviation acc where "acc \ igbg_acc G" lemma is_igb_graph: "igb_graph G" by unfold_locales lemma acc_boundI[simp, intro]: "x\acc q \ x {q . i\acc q}" definition "F \ { accn i | i. i \ g_V = V, g_E = E, g_V0 = V0, gbg_F = F, \=m \" sublocale gbg: gb_graph "to_gbg_ext m" apply unfold_locales using V0_ss E_ss acc_ss apply (auto simp: to_gbg_ext_def F_def accn_def) done lemma to_gbg_alt1: "gbg.E T m = E" "gbg.V0 T m = V0" "gbg.F T m = F" by (simp_all add: to_gbg_ext_def) lemma F_fin[simp,intro!]: "finite F" unfolding F_def by auto definition is_acc :: "'Q word \ bool" where "is_acc r \ (\n\<^sub>\i. n \ acc (r i))" definition "is_acc_run r \ is_run r \ is_acc r" lemma is_run_gbg: "gbg.is_run T m = is_run" unfolding is_run_def[abs_def] is_acc_run_def[abs_def] gbg.is_run_def[abs_def] gbg.is_acc_run_def[abs_def] by (simp_all add: to_gbg_ext_def) lemma is_acc_gbg: "gbg.is_acc T m = is_acc" apply (intro ext) unfolding gbg.is_acc_def is_acc_def apply (simp add: to_gbg_alt1 is_run_gbg) unfolding F_def accn_def apply (blast intro: INFM_mono) done lemma is_acc_run_gbg: "gbg.is_acc_run T m = is_acc_run" apply (intro ext) unfolding gbg.is_acc_run_def is_acc_run_def by (simp_all add: to_gbg_alt1 is_run_gbg is_acc_gbg) lemmas to_gbg_alt = to_gbg_alt1 is_run_gbg is_acc_gbg is_acc_run_gbg lemma acc_limit_alt: assumes FIN: "finite (range r)" shows "is_acc r \ (\n accn n \ {})" proof assume "\n accn n \ {}" thus "is_acc r" unfolding is_acc_def accn_def by (auto dest!: limit_inter_INF) next from FIN have FIN': "\A. finite (A \ range r)" by simp assume "is_acc r" hence "\n (accn n \ range r) \ {}" unfolding is_acc_def accn_def by (auto simp: fin_ex_inf_eq_limit[OF FIN', symmetric]) thus "\n accn n \ {}" by auto qed lemma acc_limit_alt': "finite (range r) \ is_acc r \ (\(acc ` limit r) = {0.. 'L \ bool" locale igba = igbg?: igb_graph G for G :: "('Q,'L,'more) igba_rec_scheme" + assumes L_ss: "igba_L G q l \ q \ V" begin abbreviation L where "L \ igba_L G" lemma is_igba: "igba G" by unfold_locales abbreviation "to_gba_ext m \ to_gbg_ext \ gba_L = igba_L G, \=m \" sublocale gba: gba "to_gba_ext m" apply unfold_locales unfolding to_gbg_ext_def using L_ss apply auto done lemma to_gba_alt_L: "gba.L T m = L" by (auto simp: to_gbg_ext_def) definition "accept w \ \r. is_acc_run r \ (\i. L (r i) (w i))" definition "lang \ Collect accept" lemma accept_gba_alt: "gba.accept T m = accept" apply (intro ext) unfolding accept_def gba.accept_def apply (simp add: to_gbg_alt to_gba_alt_L) done lemma lang_gba_alt: "gba.lang T m = lang" unfolding lang_def gba.lang_def apply (simp add: accept_gba_alt) done lemmas to_gba_alt = to_gbg_alt to_gba_alt_L accept_gba_alt lang_gba_alt end subsubsection \Indexing Conversion\ definition F_to_idx :: "'Q set set \ (nat \ ('Q \ nat set)) nres" where "F_to_idx F \ do { Flist \ SPEC (\Flist. distinct Flist \ set Flist = F); let num_acc = length Flist; let acc = (\v. {i . i v\Flist!i}); RETURN (num_acc,acc) }" lemma F_to_idx_correct: shows "F_to_idx F \ SPEC (\(num_acc,acc). F = { {q. i\acc q} | i. i \(range acc) \ {0.. do { let acc = Map.empty; (_,acc) \ nfoldli Flist (\_. True) (\A (i,acc). do { acc \ FOREACHi (\it acc'. acc' = (\v. if v\A-it then Some (insert i (the_default {} (acc v))) else acc v ) ) A (\v acc. RETURN (acc(v\insert i (the_default {} (acc v))))) acc; RETURN (Suc i,acc) }) (0,acc); RETURN (\x. the_default {} (acc x)) }" lemma mk_acc_impl_correct: assumes F: "(Flist',Flist)\Id" assumes FIN: "\A\set Flist. finite A" shows "mk_acc_impl Flist' \ \Id (RETURN (\v. {i . i v\Flist!i}))" using F apply simp unfolding mk_acc_impl_def apply (refine_rcg nfoldli_rule[where I="\l1 l2 (i,res). i=length l1 \ the_default {} o res = (\v. {j . j v\Flist!j})" ] refine_vcg ) using FIN apply (simp_all) apply (rule ext) apply auto [] apply (rule ext) apply (auto split: if_split_asm simp: nth_append nth_Cons') [] apply (rule ext) apply (auto split: if_split_asm simp: nth_append nth_Cons' fun_comp_eq_conv) [] apply (rule ext) apply (auto simp: fun_comp_eq_conv) [] done definition F_to_idx_impl :: "'Q set set \ (nat \ ('Q \ nat set)) nres" where "F_to_idx_impl F \ do { Flist \ SPEC (\Flist. distinct Flist \ set Flist = F); let num_acc = length Flist; acc \ mk_acc_impl Flist; RETURN (num_acc,acc) }" lemma F_to_idx_refine: assumes FIN: "\A\F. finite A" shows "F_to_idx_impl F \ \Id (F_to_idx F)" using assms unfolding F_to_idx_impl_def F_to_idx_def apply (refine_rcg bind_Let_refine2[OF mk_acc_impl_correct]) apply auto done definition gbg_to_idx_ext :: "_ \ ('a, 'more) gb_graph_rec_scheme \ ('a, 'more') igb_graph_rec_scheme nres" where "gbg_to_idx_ext ecnv A = do { (num_acc,acc) \ F_to_idx_impl (gbg_F A); RETURN \ g_V = g_V A, g_E = g_E A, g_V0=g_V0 A, igbg_num_acc = num_acc, igbg_acc = acc, \ = ecnv A \ }" lemma (in gb_graph) gbg_to_idx_ext_correct: assumes [simp, intro]: "\ A. A \ F \ finite A" shows "gbg_to_idx_ext ecnv G \ SPEC (\G'. igb_graph.is_acc_run G' = is_acc_run \ g_V G' = V \ g_E G' = E \ g_V0 G' = V0 \ igb_graph_rec.more G' = ecnv G \ igb_graph G' )" proof - note F_to_idx_refine[of F] also note F_to_idx_correct finally have R: "F_to_idx_impl F \ SPEC (\(num_acc, acc). F = {{q. i \ acc q} |i. i < num_acc} \ \(range acc) \ {0..a b c. (b\c) \ (a&b \ a&c)" by simp { fix acc :: "'Q \ nat set" and num_acc r have "(\A. (\i. A = {q. i \ acc q} \ i < num_acc) \ (limit r \ A \ {})) \ (\iq\limit r. i\acc q)" by blast } note aux1=this { fix acc :: "'Q \ nat set" and num_acc i assume FE: "F = {{q. i \ acc q} |i. i < num_acc}" assume INR: "(\x. acc x) \ {0.. acc q}" proof (cases "i acc q} = {}" using INR by force thus ?thesis by simp qed } note aux2=this { fix acc :: "'Q \ nat set" and num_acc q assume FE: "F = {{q. i \ acc q} |i. i < num_acc}" and INR: "(\x. acc x) \ {0.. {}" then obtain i where "i\acc q" by auto moreover with INR have "i\F" by (auto simp: FE) with F_ss have "q\V" by auto } note aux3=this show ?thesis unfolding gbg_to_idx_ext_def apply (refine_rcg order_trans[OF R] refine_vcg) proof clarsimp_all fix acc and num_acc :: nat assume FE[simp]: "F = {{q. i \ acc q} |i. i < num_acc}" and BOUND: "(\x. acc x) \ {0.. 'q igb_graph_rec nres" where "gbg_to_idx \ gbg_to_idx_ext (\_. ())" definition ti_Lcnv where "ti_Lcnv ecnv A \ \ igba_L = gba_L A, \=ecnv A \" abbreviation "gba_to_idx_ext ecnv \ gbg_to_idx_ext (ti_Lcnv ecnv)" abbreviation "gba_to_idx \ gba_to_idx_ext (\_. ())" lemma (in gba) gba_to_idx_ext_correct: assumes [simp, intro]: "\ A. A \ F \ finite A" shows "gba_to_idx_ext ecnv G \ SPEC (\G'. igba.accept G' = accept \ g_V G' = V \ g_E G' = E \ g_V0 G' = V0 \ igba_rec.more G' = ecnv G \ igba G')" apply (rule order_trans[OF gbg_to_idx_ext_correct]) apply (rule, assumption) apply (rule SPEC_rule) apply (elim conjE, intro conjI) proof - fix G' assume ARUN: "igb_graph.is_acc_run G' = is_acc_run" and MORE: "igb_graph_rec.more G' = ti_Lcnv ecnv G" and LOC: "igb_graph G'" and FIELDS: "g_V G' = V" "g_E G' = E" "g_V0 G' = V0" from LOC interpret igb: igb_graph G' . interpret igb: igba G' apply unfold_locales using MORE FIELDS L_ss unfolding ti_Lcnv_def apply (cases G') apply simp done show "igba.accept G' = accept" and "igba_rec.more G' = ecnv G" using ARUN MORE unfolding accept_def[abs_def] igb.accept_def[abs_def] ti_Lcnv_def apply (cases G', (auto) []) + done show "igba G'" by unfold_locales qed corollary (in gba) gba_to_idx_ext_lang_correct: assumes [simp, intro]: "\ A. A \ F \ finite A" shows "gba_to_idx_ext ecnv G \ SPEC (\G'. igba.lang G' = lang \ igba_rec.more G' = ecnv G \ igba G')" apply (rule order_trans[OF gba_to_idx_ext_correct]) apply (rule, assumption) apply (rule SPEC_rule) unfolding lang_def[abs_def] apply (subst igba.lang_def) apply auto done subsubsection \Degeneralization\ context igb_graph begin definition degeneralize_ext :: "_ \ ('Q \ nat, _) b_graph_rec_scheme" where "degeneralize_ext ecnv \ if num_acc = 0 then \ g_V = V \ {0}, g_E = {((q,0),(q',0)) | q q'. (q,q')\E}, g_V0 = V0\{0}, bg_F = V \ {0}, \ = ecnv G \ else \ g_V = V \ {0.. (q,q')\E \ i' = (if i\acc q then (i+1) mod num_acc else i) }, g_V0 = V0 \ {0}, bg_F = {(q,0)| q. 0\acc q}, \ = ecnv G \" abbreviation degeneralize where "degeneralize \ degeneralize_ext (\_. ())" lemma degen_more[simp]: "b_graph_rec.more (degeneralize_ext ecnv) = ecnv G" unfolding degeneralize_ext_def by auto lemma degen_invar: "b_graph (degeneralize_ext ecnv)" proof let ?G' = "degeneralize_ext ecnv" show "g_V0 ?G' \ g_V ?G'" unfolding degeneralize_ext_def using V0_ss by auto show "g_E ?G' \ g_V ?G' \ g_V ?G'" unfolding degeneralize_ext_def using E_ss by auto show "bg_F ?G' \ g_V ?G'" unfolding degeneralize_ext_def using acc_ss by auto qed sublocale degen: b_graph "degeneralize_ext m" using degen_invar . lemma degen_finite_reachable: assumes [simp, intro]: "finite (E\<^sup>* `` V0)" shows "finite ((g_E (degeneralize_ext ecnv))\<^sup>* `` g_V0 (degeneralize_ext ecnv))" proof - let ?G' = "degeneralize_ext ecnv" have "((g_E ?G')\<^sup>* `` g_V0 ?G') \ E\<^sup>*``V0 \ {0..num_acc}" proof - { fix q n q' n' assume "((q,n),(q',n'))\(g_E ?G')\<^sup>*" and 0: "(q,n)\g_V0 ?G'" hence G1: "(q,q')\E\<^sup>* \ n'\num_acc" apply (induction rule: rtrancl_induct2) by (auto simp: degeneralize_ext_def split: if_split_asm) from 0 have G2: "q\V0 \ n\num_acc" by (auto simp: degeneralize_ext_def split: if_split_asm) note G1 G2 } thus ?thesis by fastforce qed also have "finite \" by auto finally (finite_subset) show "finite ((g_E ?G')\<^sup>* `` g_V0 ?G')" . qed lemma degen_is_run_sound: "degen.is_run T m r \ is_run (fst o r)" unfolding degen.is_run_def is_run_def unfolding degeneralize_ext_def apply (clarsimp split: if_split_asm simp: ipath_def) apply (metis fst_conv)+ done lemma degen_path_sound: assumes "path (degen.E T m) u p v" shows "path E (fst u) (map fst p) (fst v)" using assms by induction (auto simp: degeneralize_ext_def path_simps split: if_split_asm) lemma degen_V0_sound: assumes "u \ degen.V0 T m" shows "fst u \ V0" using assms by (auto simp: degeneralize_ext_def path_simps split: if_split_asm) lemma degen_visit_acc: assumes "path (degen.E T m) (q,n) p (q',n')" assumes "n\n'" shows "\qa. (qa,n)\set p \ n\acc qa" using assms proof (induction _ "(q,n)" p "(q',n')" arbitrary: q rule: path.induct) case (path_prepend qnh p) then obtain qh nh where [simp]: "qnh=(qh,nh)" by (cases qnh) from \((q,n),qnh) \ degen.E T m\ have "nh=n \ (nh=(n+1) mod num_acc \ n\acc q)" by (auto simp: degeneralize_ext_def split: if_split_asm) thus ?case proof assume [simp]: "nh=n" from path_prepend obtain qa where "(qa, n) \ set p" and "n \ acc qa" by auto thus ?case by auto next assume "(nh=(n+1) mod num_acc \ n\acc q)" thus ?case by auto qed qed simp lemma degen_run_complete0: assumes [simp]: "num_acc = 0" assumes R: "is_run r" shows "degen.is_run T m (\i. (r i,0))" using R unfolding degen.is_run_def is_run_def unfolding ipath_def degeneralize_ext_def by auto lemma degen_acc_run_complete0: assumes [simp]: "num_acc = 0" assumes R: "is_acc_run r" shows "degen.is_acc_run T m (\i. (r i,0))" using R unfolding degen.is_acc_run_def is_acc_run_def is_acc_def degen.is_acc_def apply (simp add: degen_run_complete0) unfolding degeneralize_ext_def using run_reachable[of r] reachable_V by (auto simp: INFM_nat) lemma degen_run_complete: assumes [simp]: "num_acc \ 0" assumes R: "is_run r" shows "\r'. degen.is_run T m r' \ r = fst o r'" using R unfolding degen.is_run_def is_run_def ipath_def apply (elim conjE) proof - assume R0: "r 0 \ V0" and RS: "\i. (r i, r (Suc i)) \ E" define r' where "r' = rec_nat (r 0,0) (\i (q,n). (r (Suc i), if n \ acc q then (n+1) mod num_acc else n))" have [simp]: "r' 0 = (r 0,0)" "\i. r' (Suc i) = ( let (q,n)=r' i in (r (Suc i), if n \ acc q then (n+1) mod num_acc else n) )" unfolding r'_def by auto have R0': "r' 0 \ degen.V0 T m" using R0 unfolding degeneralize_ext_def by auto have MAP: "r = fst o r'" proof (rule ext) fix i show "r i = (fst o r') i" by (cases i) (auto simp: split: prod.split) qed have [simp]: "0i. snd (r' i) < num_acc" proof - fix i show "snd (r' i) < num_acc" by (induction i) (auto split: prod.split) qed have RS': "\i. (r' i, r' (Suc i)) \ degen.E T m" proof fix i obtain n where [simp]: "r' i = (r i,n)" apply (cases i) apply (force) apply (force split: prod.split) done from SND_LESS[of i] have [simp]: "n degen.E T m" using RS by (auto simp: degeneralize_ext_def) qed from R0' RS' MAP show "\r'. (r' 0 \ degen.V0 T m \ (\i. (r' i, r' (Suc i)) \ degen.E T m)) \ r = fst \ r'" by blast qed lemma degen_run_bound: assumes [simp]: "num_acc \ 0" assumes R: "degen.is_run T m r" shows "snd (r i) < num_acc" apply (induction i) using R unfolding degen.is_run_def is_run_def unfolding degeneralize_ext_def ipath_def apply - apply auto [] apply clarsimp by (metis snd_conv) lemma degen_acc_run_complete_aux1: assumes NN0[simp]: "num_acc \ 0" assumes R: "degen.is_run T m r" assumes EXJ: "\j\i. n \ acc (fst (r j))" assumes RI: "r i = (q,n)" shows "\j\i. \q'. r j = (q',n) \ n \ acc q'" proof - define j where "j = (LEAST j. j\i \ n \ acc (fst (r j)))" from RI have "ni" "n \ acc (fst (r j))" "\k\i. n \ acc (fst (r k)) \ j\k" using LeastI_ex[OF EXJ] unfolding j_def apply (auto) [2] apply (metis (lifting) Least_le) done hence "\k\i. k n \ acc (fst (r k))" by auto have "\k. k\i \ k\j \ (snd (r k) = n)" proof (clarify) fix k assume "i\k" "k\j" thus "snd (r k) = n" proof (induction k rule: less_induct) case (less k) show ?case proof (cases "k=i") case True thus ?thesis using RI by simp next case False with less.prems have "k - 1 < k" "i \ k - 1" "k - 1\j" by auto from less.IH[OF this] have "snd (r (k - 1)) = n" . moreover from R have "(r (k - 1), r k) \ degen.E T m" unfolding degen.is_run_def is_run_def ipath_def by clarsimp (metis One_nat_def Suc_diff_1 \k - 1 < k\ less_nat_zero_code neq0_conv) moreover have "n \ acc (fst (r (k - 1)))" using \\k\i. k < j \ n \ acc (fst (r k))\ \i \ k - 1\ \k - 1 < k\ dual_order.strict_trans1 less.prems(2) by blast ultimately show ?thesis by (auto simp: degeneralize_ext_def) qed qed qed thus ?thesis by (metis \i \ j\ \n \ local.acc (fst (r j))\ order_refl surjective_pairing) qed lemma degen_acc_run_complete_aux1': assumes NN0[simp]: "num_acc \ 0" assumes R: "degen.is_run T m r" assumes ACC: "\n\<^sub>\i. n \ acc (fst (r i))" assumes RI: "r i = (q,n)" shows "\j\i. \q'. r j = (q',n) \ n \ acc q'" proof - from RI have "nj\i. n \ acc (fst (r j))" unfolding INFM_nat_le by blast from degen_acc_run_complete_aux1[OF NN0 R EXJ RI] show ?thesis . qed lemma degen_acc_run_complete_aux2: assumes NN0[simp]: "num_acc \ 0" assumes R: "degen.is_run T m r" assumes ACC: "\n\<^sub>\i. n \ acc (fst (r i))" assumes RI: "r i = (q,n)" and OFS: "ofsj\i. \q'. r j = (q',(n + ofs) mod num_acc) \ (n + ofs) mod num_acc \ acc q'" using RI OFS proof (induction ofs arbitrary: q n i) case 0 from degen_run_bound[OF NN0 R, of i] \r i = (q, n)\ have NLE: "nr i = (q, n)\] show ?case by auto next case (Suc ofs) from Suc.IH[OF Suc.prems(1)] Suc.prems(2) obtain j q' where "j\i" and RJ: "r j = (q',(n+ofs) mod num_acc)" and A: "(n+ofs) mod num_acc \ acc q'" by auto from R have "(r j, r (Suc j)) \ degen.E T m" by (auto simp: degen.is_run_def is_run_def ipath_def) with RJ A obtain q2 where RSJ: "r (Suc j) = (q2,(n+Suc ofs) mod num_acc)" by (auto simp: degeneralize_ext_def mod_simps) have aux: "\j'. i\j \ Suc j \ j' \ i\j'" by auto from degen_acc_run_complete_aux1'[OF NN0 R ACC RSJ] \j\i\ show ?case by (auto dest: aux) qed lemma degen_acc_run_complete: assumes AR: "is_acc_run r" obtains r' where "degen.is_acc_run T m r'" and "r = fst o r'" proof (cases "num_acc = 0") case True with AR degen_acc_run_complete0 show ?thesis by (auto intro!: that[of "(\i. (r i, 0))"]) next case False assume NN0[simp]: "num_acc \ 0" from AR have R: "is_run r" and ACC: "\n\<^sub>\i. n \ acc (r i)" unfolding is_acc_run_def is_acc_def by auto from degen_run_complete[OF NN0 R] obtain r' where R': "degen.is_run T m r'" and [simp]: "r = fst \ r'" by auto from ACC have ACC': "\n\<^sub>\i. n \ acc (fst (r' i))" by simp have "\i. \j>i. r' j \ degen.F T m" proof fix i obtain q n where RI: "r' (Suc i) = (q,n)" by (cases "r' (Suc i)") have "(n + (num_acc - n mod num_acc)) mod num_acc = 0" - by (metis NN0 R' \r' (Suc i) = (q, n)\ add_diff_cancel_left' - degen_run_bound less_imp_add_positive mod_self nat_mod_eq' snd_conv) + apply (rule dvd_imp_mod_0) + apply (metis (mono_tags, lifting) NN0 add_diff_inverse mod_0_imp_dvd + mod_add_left_eq mod_less_divisor mod_self nat_diff_split not_gr_zero zero_less_diff) + done then obtain ofs where OFS_LESS: "ofsi" "r' j = (q',0)" and "0\acc q'" by (auto simp: less_eq_Suc_le) thus "\j>i. r' j \ degen.F T m" by (auto simp: degeneralize_ext_def) qed hence "\\<^sub>\i. r' i \ degen.F T m" by (auto simp: INFM_nat) have "degen.is_acc_run T m r'" unfolding degen.is_acc_run_def degen.is_acc_def by rule fact+ thus ?thesis by (auto intro: that) qed lemma degen_run_find_change: assumes NN0[simp]: "num_acc \ 0" assumes R: "degen.is_run T m r" assumes A: "i\j" "r i = (q,n)" "r j = (q',n')" "n\n'" obtains k qk where "i\k" "k acc qk" proof - from degen_run_bound[OF NN0 R] A have "n snd (r k) \ n)" have "i n" by (metis (lifting, mono_tags) LeastI_ex A k_def leD less_linear snd_conv)+ from Least_le[where P="\k. i snd (r k) \ n", folded k_def] have LEK_EQN: "\k'. i\k' \ k' snd (r k') = n" using \r i = (q,n)\ by clarsimp (metis le_neq_implies_less not_le snd_conv) hence SND_RKMO: "snd (r (k - 1)) = n" using \i by auto moreover from R have "(r (k - 1), r k) \ degen.E T m" unfolding degen.is_run_def ipath_def using \i by clarsimp (metis Suc_pred gr_implies_not0 neq0_conv) moreover note \snd (r k) \ n\ ultimately have "n \ acc (fst (r (k - 1)))" by (auto simp: degeneralize_ext_def split: if_split_asm) moreover have "k - 1 < j" using A LEK_EQN apply (rule_tac ccontr) apply clarsimp by (metis One_nat_def \snd (r (k - 1)) = n\ less_Suc_eq less_imp_diff_less not_less_eq snd_conv) ultimately show thesis apply - apply (rule that[of "k - 1" "fst (r (k - 1))"]) using \i SND_RKMO by auto qed lemma degen_run_find_acc_aux: assumes NN0[simp]: "num_acc \ 0" assumes AR: "degen.is_acc_run T m r" assumes A: "r i = (q,0)" "0 \ acc q" "nj qj. i\j \ r j = (qj,n) \ n \ acc qj" proof - from AR have R: "degen.is_run T m r" and ACC: "\\<^sub>\i. r i \ degen.F T m" (*and ACC: "limit r \ bg_F (degeneralize_ext ecnv) \ {}"*) unfolding degen.is_acc_run_def degen.is_acc_def by auto from ACC have ACC': "\i. \j>i. r j \ degen.F T m" by (auto simp: INFM_nat) show ?thesis using \n proof (induction n) case 0 thus ?case using A by auto next case (Suc n) then obtain j qj where "i\j" "r j = (qj,n)" "n\acc qj" by auto moreover from R have "(r j, r (Suc j)) \ degen.E T m" unfolding degen.is_run_def ipath_def by auto ultimately obtain qsj where RSJ: "r (Suc j) = (qsj,Suc n)" unfolding degeneralize_ext_def using \Suc n by auto from ACC' obtain k q0 where "Suc j \ k" "r k = (q0, 0)" unfolding degeneralize_ext_def apply auto by (metis less_imp_le_nat) from degen_run_find_change[OF NN0 R \Suc j \ k\ RSJ \r k = (q0, 0)\] obtain l ql where "Suc j \ l" "l < k" "r l = (ql, Suc n)" "Suc n \ acc ql" by blast thus ?case using \i \ j\ by (intro exI[where x=l] exI[where x=ql]) auto qed qed lemma degen_acc_run_sound: assumes A: "degen.is_acc_run T m r" shows "is_acc_run (fst o r)" proof - from A have R: "degen.is_run T m r" and ACC: "\\<^sub>\i. r i \ degen.F T m" unfolding degen.is_acc_run_def degen.is_acc_def by auto from degen_is_run_sound[OF R] have R': "is_run (fst o r)" . show ?thesis proof (cases "num_acc = 0") case NN0[simp]: False from ACC have ACC': "\i. \j>i. r j \ degen.F T m" by (auto simp: INFM_nat) have "\ni. \j>i. n \ acc (fst (r j))" proof (intro allI impI) fix n i obtain j qj where "j>i" and RJ: "r j = (qj,0)" and ACCJ: "0 \ acc (qj)" using ACC' unfolding degeneralize_ext_def by fastforce assume NLESS: "nj>i. n \ acc (fst (r j))" proof (cases n) case 0 thus ?thesis using \j>i\ RJ ACCJ by auto next case [simp]: (Suc n') from degen_run_find_acc_aux[OF NN0 A RJ ACCJ NLESS] obtain k qk where "j\k" "r k = (qk,n)" "n \ acc qk" by auto thus ?thesis by (metis \i < j\ dual_order.strict_trans1 fst_conv) qed qed hence "\n\<^sub>\i. n \ acc (fst (r i))" by (auto simp: INFM_nat) with R' show ?thesis unfolding is_acc_run_def is_acc_def by auto next case [simp]: True with R' show ?thesis unfolding is_acc_run_def is_acc_def by auto qed qed lemma degen_acc_run_iff: "is_acc_run r \ (\r'. fst o r' = r \ degen.is_acc_run T m r')" using degen_acc_run_complete degen_acc_run_sound by blast end subsection "System Automata" text \ System automata are (finite) rooted graphs with a labeling function. They are used to describe the model (system) to be checked. \ record ('Q,'L) sa_rec = "'Q graph_rec" + sa_L :: "'Q \ 'L" locale sa = g?: graph G for G :: "('Q, 'L, 'more) sa_rec_scheme" begin abbreviation L where "L \ sa_L G" definition "accept w \ \r. is_run r \ w = L o r" lemma acceptI[intro?]: "\is_run r; w = L o r\ \ accept w" by (auto simp: accept_def) definition "lang \ Collect accept" lemma langI[intro?]: "accept w \ w\lang" by (auto simp: lang_def) end subsubsection "Product Construction" text \ In this section we formalize the product construction between a GBA and a system automaton. The result is a GBG and a projection function, such that projected runs of the GBG correspond to words accepted by the GBA and the system. \ locale igba_sys_prod_precond = igba: igba G + sa: sa S for G :: "('q,'l,'moreG) igba_rec_scheme" and S :: "('s,'l,'moreS) sa_rec_scheme" begin definition "prod \ \ g_V = igba.V \ sa.V, g_E = { ((q,s),(q',s')). igba.L q (sa.L s) \ (q,q') \ igba.E \ (s,s') \ sa.E }, g_V0 = igba.V0 \ sa.V0, igbg_num_acc = igba.num_acc, igbg_acc = (\(q,s). if s\sa.V then igba.acc q else {} ) \" lemma prod_invar: "igb_graph prod" apply unfold_locales using igba.V0_ss sa.V0_ss apply (auto simp: prod_def) [] using igba.E_ss sa.E_ss apply (auto simp: prod_def) [] using igba.acc_bound apply (auto simp: prod_def split: if_split_asm) [] using igba.acc_ss apply (fastforce simp: prod_def split: if_split_asm) [] done sublocale prod: igb_graph prod using prod_invar . lemma prod_finite_reachable: assumes "finite (igba.E\<^sup>* `` igba.V0)" "finite (sa.E\<^sup>* `` sa.V0)" shows "finite ((g_E prod)\<^sup>* `` g_V0 prod)" proof - { fix q s q' s' assume "((q,s),(q',s')) \ (g_E prod)\<^sup>*" hence "(q,q') \ (igba.E)\<^sup>* \ (s,s') \ (sa.E)\<^sup>*" apply (induction rule: rtrancl_induct2) apply (auto simp: prod_def) done } note gsp_reach=this have [simp]: "\q s. (q,s) \ g_V0 prod \ q \ igba.V0 \ s \ sa.V0" by (auto simp: prod_def) have reachSS: "((g_E prod)\<^sup>* `` g_V0 prod) \ ((igba.E)\<^sup>* `` igba.V0) \ (sa.E\<^sup>* `` sa.V0)" by (auto dest: gsp_reach) show ?thesis apply (rule finite_subset[OF reachSS]) using assms by simp qed lemma prod_fields: "prod.V = igba.V \ sa.V" "prod.E = { ((q,s),(q',s')). igba.L q (sa.L s) \ (q,q') \ igba.E \ (s,s') \ sa.E }" "prod.V0 = igba.V0 \ sa.V0" "prod.num_acc = igba.num_acc" "prod.acc = (\(q,s). if s\sa.V then igba.acc q else {} )" unfolding prod_def apply simp_all done lemma prod_run: "prod.is_run r \ igba.is_run (fst o r) \ sa.is_run (snd o r) \ (\i. igba.L (fst (r i)) (sa.L (snd (r i))))" (is "?L=?R") apply rule unfolding igba.is_run_def sa.is_run_def prod.is_run_def unfolding prod_def ipath_def apply (auto split: prod.split_asm intro: in_prod_fst_sndI) apply (metis surjective_pairing) apply (metis surjective_pairing) apply (metis fst_conv snd_conv) apply (metis fst_conv snd_conv) apply (metis fst_conv snd_conv) done lemma prod_acc: assumes A: "range (snd o r) \ sa.V" shows "prod.is_acc r \ igba.is_acc (fst o r)" proof - { fix i from A have "prod.acc (r i) = igba.acc (fst (r i))" unfolding prod_fields apply safe apply (clarsimp_all split: if_split_asm) by (metis UNIV_I comp_apply imageI snd_conv subsetD) } note [simp] = this show ?thesis unfolding prod.is_acc_def igba.is_acc_def by (simp add: prod_fields(4)) qed lemma gsp_correct1: assumes A: "prod.is_acc_run r" shows "sa.is_run (snd o r) \ (sa.L o snd o r \ igba.lang)" proof - from A have PR: "prod.is_run r" and PA: "prod.is_acc r" unfolding prod.is_acc_run_def by auto have PRR: "range r \ prod.V" using prod.run_reachable prod.reachable_V PR by auto have RSR: "range (snd o r) \ sa.V" using PRR unfolding prod_fields by auto show ?thesis using PR PA unfolding igba.is_acc_run_def igba.lang_def igba.accept_def[abs_def] apply (auto simp: prod_run prod_acc[OF RSR]) done qed lemma gsp_correct2: assumes A: "sa.is_run r" "sa.L o r \ igba.lang" shows "\r'. r = snd o r' \ prod.is_acc_run r'" proof - have [simp]: "\r r'. fst o (\i. (r i, r' i)) = r" "\r r'. snd o (\i. (r i, r' i)) = r'" by auto from A show ?thesis unfolding prod.is_acc_run_def igba.lang_def igba.accept_def[abs_def] igba.is_acc_run_def apply (clarsimp simp: prod_run) apply (rename_tac ra) apply (rule_tac x="\i. (ra i, r i)" in exI) apply clarsimp apply (subst prod_acc) using order_trans[OF sa.run_reachable sa.reachable_V] apply auto [] apply auto [] done qed end end diff --git a/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy b/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy --- a/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy @@ -1,1416 +1,1413 @@ theory SM_Ample_Impl imports SM_Datastructures "../Analysis/SM_POR" "HOL-Library.RBT_Mapping" begin text \ Given a sticky program, we implement a valid ample-function by picking the transitions from the first PID that qualifies as an ample-set\ context visible_prog begin definition ga_gample :: "(cmdc\cmdc) set \ pid_global_config \ global_action set" where "ga_gample sticky_E gc \ let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc; as = find_min_idx_f (\lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c) in if (\(a,_)\as. read_globals a = {} \ write_globals a = {}) then let as = {(a,c')\as. la_en' (ls,gs) a} in if as={} \ (\(a,c')\as. (c,c')\sticky_E) then None else Some (c,as) else None ) lcs in case as of Some (pid,c,as) \ (\(a,c'). (pid,c,a,c'))`as | None \ ga_gen gc" definition "ga_ample sticky_E \ stutter_extend_en (ga_gample sticky_E)" lemma ga_gample_subset: "ga_gample sticky_E gc \ ga_gen gc" unfolding ga_gample_def ga_gen_def by (auto simp: comp.astep_impl_def split: option.splits if_split_asm dest!: find_min_idx_f_SomeD) lemma ga_gample_empty_iff: "ga_gample sticky_E gc = {} \ ga_gen gc = {}" unfolding ga_gample_def ga_gen_def by (fastforce simp: comp.astep_impl_def split: option.splits if_split_asm dest!: find_min_idx_f_SomeD) lemma ga_ample_None_iff: "None \ ga_ample sticky_E gc \ None \ ga_en gc" by (auto simp: ga_ample_def ga_en_def stutter_extend_en_def ga_gample_empty_iff) lemma ga_ample_neq_en_eq: assumes "ga_ample sticky_E gc \ ga_en gc" shows "ga_ample sticky_E gc = Some`ga_gample sticky_E gc \ ga_en gc = Some`ga_gen gc" using assms ga_gample_subset unfolding ga_ample_def ga_en_def stutter_extend_en_def by (auto split: if_split_asm simp: ga_gample_empty_iff) lemma pid_pac_alt: "pid_pac pid = insert None (Some`{(pid',cac). pid'=pid})" apply (auto simp: pid_pac_def split: option.splits) apply (case_tac x) apply auto done end context sticky_prog begin sublocale ample_prog prog is_vis_var sticky_E "ga_ample sticky_E" proof - have aux: "Some ` A = Some ` B \ pid_pac pid \ A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" for A B pid proof assume 1: "Some ` A = Some ` B \ pid_pac pid" show "A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" using 1 apply (auto simp: pid_pac_alt) unfolding image_def apply auto using 1 by blast next show "Some ` A = Some ` B \ pid_pac pid" if "A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" using that by (auto simp: pid_pac_alt) qed show "ample_prog prog is_vis_var sticky_E (ga_ample sticky_E)" apply unfold_locales apply (frule ga_ample_neq_en_eq, clarsimp) apply (intro conjI) apply (auto simp: ga_gample_def sticky_ga_def dest!: find_min_idx_f_SomeD split: option.splits if_split_asm) [] apply (fastforce simp: ga_gample_def sticky_ga_def dest!: find_min_idx_f_SomeD split: option.splits if_split_asm) [] unfolding aux (* TODO: Clean up! *) apply (simp add: inj_image_eq_iff pid_pac_def) apply (thin_tac "ga_ample c a = b" for a b c) apply (simp add: ga_gample_def split: option.splits prod.splits) apply (thin_tac "a \ ga_gen b" for a b) apply (drule find_min_idx_f_SomeD) apply clarsimp apply (rename_tac gc pid c as) apply (rule_tac x=pid in exI) apply (auto split: if_split_asm prod.splits simp: ga_gen_def comp.astep_impl_def) unfolding is_ample_static_def comp.astep_impl_def apply force done qed end text \ The following locale expresses a correct ample-reduction on the level of subset and stuttering-equivalences of traces \ locale hl_reduction = visible_prog + ample: sa "\ g_V = UNIV, g_E = step, g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" for step + assumes ample_accept_subset: "ample.accept w \ lv.sa.accept w" assumes ample_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ ample.accept w'" locale hl_ample_reduction = visible_prog + ample: sa "\ g_V = UNIV, g_E = rel_of_enex (ample, ga_ex), g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" for ample + assumes ample_accept_subset: "ample.accept w \ lv.sa.accept w" assumes ample_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ ample.accept w'" begin sublocale hl_reduction prog is_vis_var "rel_of_enex (ample, ga_ex)" apply unfold_locales using ample_accept_subset ample_accept_stuttering by auto end context sticky_prog begin sublocale hl_ample: hl_ample_reduction prog is_vis_var "ga_ample sticky_E" apply unfold_locales unfolding ample_rel_def[symmetric] unfolding reduced_automaton_def[symmetric] apply simp apply simp using ample_accept_subset apply simp apply (erule ample_accept_stuttering) apply blast done end text \Next, we implement the selection of a set of sticky edges\ context visible_prog begin definition "cr_ample \ do { sticky_E \ find_fas_init cfgc_G vis_edges; RETURN (ga_ample (sticky_E)) }" lemma cr_ample_reduction: "cr_ample \ SPEC (hl_ample_reduction prog is_vis_var)" unfolding cr_ample_def apply (refine_vcg order_trans[OF find_fas_init_correct]) apply unfold_locales[] apply simp proof clarify fix sticky_E assume "is_fas cfgc_G sticky_E" "vis_edges \ sticky_E" then interpret sticky_prog prog is_vis_var sticky_E by unfold_locales show "hl_ample_reduction prog is_vis_var (ga_ample sticky_E)" by unfold_locales qed end text \We derive an implementation for finding the feedback arcs, and an efficient implementation for ga_gample.\ text \We replace finding the sticky edges by an efficient algorithm\ context visible_prog begin definition ga_gample_m :: "_ \ _ \ pid_global_config \ global_action set" where "ga_gample_m mem sticky_E gc \ let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc; as = find_min_idx_f (\lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c) in if (\(a,_)\as. read_globals a = {} \ write_globals a = {}) then let as = {(a,c')\as. la_en' (ls,gs) a} in if as={} \ (\(a,c')\as. mem (c,c') sticky_E) then None else Some (c,as) else None ) lcs in case as of Some (pid,c,as) \ (\(a,c'). (pid,c,a,c'))`as | None \ ga_gen gc" lemma ga_gample_m_refine: fixes Rs :: "('si\'s) set" shows "(ga_gample_m, ga_gample_m) \ (Id\\<^sub>rId \ Rs \ bool_rel) \ Rs \ Id \ \Id\set_rel" proof (intro fun_relI, simp) fix m :: "cmdc\cmdc \ 's \ bool" and m' :: "cmdc\cmdc \ 'si \ bool" and se se' gc assume "(m',m)\Id \ Rs \ bool_rel" and "(se',se)\Rs" hence "\c c'. (m' (c,c') se', m (c,c') se)\Id" by parametricity simp_all hence A: "\c c'. m' (c,c') se' = m (c,c') se" by simp show "ga_gample_m m' se' gc = ga_gample_m m se gc" unfolding ga_gample_m_def by (subst A) (rule refl) qed lemma ga_gample_eq_gample_m: "ga_gample = ga_gample_m (\)" unfolding ga_gample_def[abs_def] ga_gample_m_def[abs_def] by auto lemma stutter_extend_en_refine: "(stutter_extend_en, stutter_extend_en) \ (R \ \Id\set_rel) \ R \ \\Id\option_rel\set_rel" unfolding stutter_extend_en_def[abs_def] apply parametricity by auto lemma is_vis_var_refine: "(is_vis_var, is_vis_var) \ Id \ bool_rel" by simp lemma vis_edges_refine: "(\x. x\vis_edges,vis_edges)\\Id\\<^sub>rId\fun_set_rel" by (simp add: fun_set_rel_def br_def) lemma [autoref_op_pat_def]: "vis_edges \ Autoref_Tagging.OP vis_edges" "ga_gample_m \ Autoref_Tagging.OP ga_gample_m" "cfgc_G \ Autoref_Tagging.OP cfgc_G" by simp_all schematic_goal cr_ample_impl1_aux: notes [autoref_rules] = ga_gample_m_refine stutter_extend_en_refine is_vis_var_refine cfgc_G_impl_refine vis_edges_refine IdI[of prog] shows "(RETURN (?c::?'c),cr_ample)\?R" unfolding cr_ample_def ga_ample_def ga_gample_eq_gample_m using [[autoref_trace_failed_id, autoref_trace_intf_unif]] by (autoref_monadic (plain, trace)) concrete_definition cr_ample_impl1 uses cr_ample_impl1_aux lemma cr_ample_impl1_reduction: "hl_ample_reduction prog is_vis_var cr_ample_impl1" proof - note cr_ample_impl1.refine[THEN nres_relD] also note cr_ample_reduction finally show ?thesis by simp qed end text \The efficient implementation of @{const visible_prog.ga_gample} uses the @{const collecti_index} function\ lemma collecti_index_cong: assumes C: "\i. i f i (l!i) = f' i (l'!i)" assumes [simp]: "l=l'" shows "collecti_index f l = collecti_index f' l'" proof - { fix i0 s0 assume "\iix. x\set l \ f x = f' x" assumes "l=l'" shows "find_min_idx_f f l = find_min_idx_f f' l'" unfolding assms(2)[symmetric] using assms(1) apply (induction l) apply (auto split: option.split) done lemma find_min_idx_impl_collecti: "( case find_min_idx_f f l of Some (i,r) \ {i} \ s1 i r | None \ {(i,r) | i r. i r\s2 i (l!i)} ) = ( collecti_index (\i x. case f x of Some r \ (True, s1 i r) | None \ (False, s2 i x) ) l )" (is "?l=?r" is "_ = collecti_index ?fc l") proof (cases "\i None") define i where "i \ LEAST i. i < length l \ f (l ! i) \ None" have C_ALT: "(\i None) \ (\i. i < length l \ fst (?fc i (l ! i)))" by (auto split: option.split) have i_alt: "i = (LEAST i. i < length l \ fst (?fc i (l ! i)))" unfolding i_def apply (fo_rule arg_cong) by (auto split: option.split) { case True from LeastI_ex[OF True] obtain r where [simp]: "i s1 i r" by simp from collecti_index_collect[of ?fc l, folded i_alt] True C_ALT have REQ: "?r = {i} \ snd (?fc i (l!i))" by simp show "?l=?r" unfolding LEQ unfolding REQ by auto } { case False from False find_min_idx_f_LEAST_eq[of f l] have "find_min_idx_f f l = None" by simp hence LEQ: "?l = {(i,r) | i r. i r\s2 i (l!i)}" by simp from False C_ALT collecti_index_collect[of ?fc l] have REQ: "?r = {(i, x) |i x. i < length l \ x \ snd (?fc i (l ! i))}" by auto show "?l=?r" unfolding LEQ unfolding REQ using False by auto } qed lemma find_min_idx_impl_collecti_scheme: assumes "\x. s11 x = (case x of (i,r) \ {i} \ iss i r)" assumes "s12 = {(i,r) | i r. i r\isn i (l!i)}" assumes "\i. i f2 i (l!i) = (case f (l!i) of Some r \ (True, iss i r) | None \ (False, isn i (l!i)))" shows "(case find_min_idx_f f l of Some x \ s11 x | None \ s12) = (collecti_index f2 l)" apply (simp only: assms(1,2,3) cong: collecti_index_cong) apply (subst find_min_idx_impl_collecti) apply simp done schematic_goal stutter_extend_en_list_aux: "(?c, stutter_extend_en) \ (R \ \S\list_set_rel) \ R \ \\S\option_rel\list_set_rel" unfolding stutter_extend_en_def[abs_def] apply (autoref) done concrete_definition stutter_extend_en_list uses stutter_extend_en_list_aux lemma succ_of_enex_impl_aux: "{s'. \a\en s. s'=ex a s} = (\a. ex a s)`en s" by auto schematic_goal succ_of_enex_list_aux: "(?c, succ_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ Id \ \Id\list_set_rel" unfolding succ_of_enex_def[abs_def] unfolding succ_of_enex_impl_aux by (autoref (trace, keep_goal)) concrete_definition succ_of_enex_list uses succ_of_enex_list_aux schematic_goal pred_of_enex_list_aux: "(?c, pred_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ Id \ Id \ bool_rel" unfolding pred_of_enex_def[abs_def] by (autoref (trace, keep_goal)) concrete_definition pred_of_enex_list uses pred_of_enex_list_aux (* TODO: can this be defined using autoref? *) definition "rel_of_enex_list enex \ rel_of_pred (pred_of_enex_list enex)" lemma rel_of_enex_list_refine: "(rel_of_enex_list, rel_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ \Id \\<^sub>r Id\ set_rel" proof - have 1: "rel_of_enex_list = rel_of_pred \ pred_of_enex_list" using rel_of_enex_list_def by force have 2: "rel_of_enex = rel_of_pred \ pred_of_enex" by auto have 3: "(rel_of_pred, rel_of_pred) \ (Id \ Id \ bool_rel) \ \Id \\<^sub>r Id\set_rel" by auto show ?thesis unfolding 1 2 by (parametricity add: 3 pred_of_enex_list.refine) qed context visible_prog begin definition "ga_gample_mi2 mem sticky_E (gc::pid_global_config) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c); s={(a,c')\as. la_en' (ls,gs) a}; ample = s\{} \ (\(a,_)\as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\s. \ mem (c,c') sticky_E) in (ample,{c}\s) ) lcs )" lemma ga_gample_mi2_impl: "ga_gample_m mem sticky_E gc = ga_gample_mi2 mem sticky_E gc" unfolding ga_gample_m_def ga_gample_mi2_def apply simp apply (rule find_min_idx_impl_collecti_scheme[where iss="\pid (c,S). (\(a,c'). (c,a,c'))`S" and isn="\pid lc. {(c,a,c'). cfgc c a c' \ c = cmd_of_pid gc pid \ la_en' (local_config.state lc, pid_global_config.state gc) a}"] ) apply auto [] unfolding ga_gen_def apply auto [] apply (auto simp: comp.astep_impl_def) done definition "cr_ample_impl2 \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (\x. x \ vis_edges) in stutter_extend_en (ga_gample_mi2 (\x f. f x) sticky)" definition "ample_step_impl2 \ rel_of_enex (cr_ample_impl2,ga_ex)" lemma cr_ample_impl2_reduction: "hl_reduction prog is_vis_var ample_step_impl2" proof - from cr_ample_impl1_reduction interpret hl_ample_reduction prog is_vis_var cr_ample_impl1 . have 1: "ample_step_impl2 = rel_of_enex (cr_ample_impl1, ga_ex)" unfolding ample_step_impl2_def cr_ample_impl1_def cr_ample_impl2_def unfolding ga_gample_mi2_impl[abs_def] by simp show ?thesis apply unfold_locales unfolding 1 apply simp apply simp apply (blast intro: ample_accept_subset ample_accept_stuttering)+ done qed end text \Implementing the State\ type_synonym valuation_impl = "(ident,uint32) mapping" definition "vi_\ \ map_option Rep_uint32 \\ Mapping.rep" abbreviation "vi_rel \ br vi_\ (\_. True)" abbreviation "val_rel \ br Rep_uint32 (\_. True)" lemma vi_rel_rew: "(x = vi_\ y) \ (y,x)\vi_rel" "(vi_\ y = x) \ (y,x)\vi_rel" by (auto simp: br_def) export_code Mapping.lookup Mapping.update Mapping.empty checking SML lift_definition vi_empty :: "valuation_impl" is "Map.empty::valuation" . lemma [code]: "vi_empty = Mapping.empty" by transfer simp lemma vi_empty_correct: "vi_\ vi_empty = Map.empty" unfolding vi_\_def by transfer auto lift_definition vi_lookup :: "valuation_impl \ ident \ uint32" is "\vi::valuation. \x. vi x" . lemma [code]: "vi_lookup = Mapping.lookup" apply (intro ext) unfolding vi_lookup_def Mapping.lookup_def map_fun_def[abs_def] o_def id_apply option.map_comp Rep_uint32_inverse option.map_ident by simp lemma vi_lookup_correct: "vi_lookup v x = map_option Abs_uint32 (vi_\ v x)" unfolding vi_\_def apply transfer apply (simp add: option.map_comp o_def option.map_ident) done lift_definition vi_update :: "ident \ uint32 \ valuation_impl \ valuation_impl" is "\x v. \vi::valuation. vi(x\v)" . lemma [code]: "vi_update = Mapping.update" by transfer simp lemma vi_update_correct: "vi_\ (vi_update k v m) = vi_\ m (k\Rep_uint32 v)" unfolding vi_\_def by transfer simp export_code vi_lookup vi_update vi_empty checking SML record local_state_impl = variables :: valuation_impl record global_state_impl = variables :: valuation_impl type_synonym focused_state_impl = "local_state_impl \ global_state_impl" type_synonym local_config_impl = "(cmdc,local_state_impl) Gen_Scheduler.local_config" type_synonym pid_global_config_impl = "(cmdc,local_state_impl,global_state_impl) Pid_Scheduler.pid_global_config" definition "local_state_rel \ { (\ local_state_impl.variables = vi \, \ local_state.variables = v \) | vi v. v = vi_\ vi }" definition "global_state_rel \ { (\ global_state_impl.variables = vi \, \ global_state.variables = v \) | vi v. v = vi_\ vi }" abbreviation "focused_state_rel \ local_state_rel \\<^sub>r global_state_rel" definition "local_config_rel \ { (\ local_config.command = ci, local_config.state=lsi \, \ local_config.command = c, local_config.state=ls \) | ci c lsi ls. (ci::cmdc,c)\Id \ (lsi,ls)\local_state_rel }" definition "global_config_rel \ { (\ pid_global_config.processes = psi, pid_global_config.state = gsi \, \ pid_global_config.processes = ps, pid_global_config.state = gs \) | psi ps gsi gs. (psi,ps)\\local_config_rel\list_rel \ (gsi,gs)\global_state_rel }" definition "init_valuation_impl vd \ fold (\x v. vi_update x 0 v) vd (vi_empty)" lemma init_valuation_impl: "(init_valuation_impl, init_valuation) \ \Id\list_rel \ vi_rel" proof - { fix vd m have "m ++ init_valuation vd = fold (\x v. v(x\0)) vd m" apply (rule sym) apply (induction vd arbitrary: m) unfolding init_valuation_def[abs_def] apply (auto intro!: ext simp: Map.map_add_def) done } note aux=this have E1: "\vd. init_valuation vd = fold (\x v. v(x\0)) vd Map.empty" by (subst aux[symmetric]) simp have E2: "\vdi vd. (vdi,vd)\\Id\list_rel \ (init_valuation_impl vdi, fold (\x v. v(x\0)) vd Map.empty) \ vi_rel" unfolding init_valuation_impl_def apply parametricity apply (auto simp: vi_update_correct vi_empty_correct br_def zero_uint32.rep_eq) done show ?thesis apply (intro fun_relI) apply (subst E1) apply (rule E2) by auto qed context cprog begin definition init_pc_impl :: "proc_decl \ local_config_impl" where "init_pc_impl pd \ \ local_config.command = comp.\ (proc_decl.body pd), local_config.state = \ local_state_impl.variables = init_valuation_impl (proc_decl.local_vars pd) \ \" lemma init_pc_impl: "(init_pc_impl, init_pc) \ Id \ local_config_rel" apply (rule fun_relI) unfolding init_pc_impl_def init_pc_def local_config_rel_def local_state_rel_def apply (simp del: pair_in_Id_conv, intro conjI) apply simp apply (simp only: vi_rel_rew) apply (parametricity add: init_valuation_impl) apply simp done definition pid_init_gc_impl :: "pid_global_config_impl" where "pid_init_gc_impl \ \ pid_global_config.processes = (map init_pc_impl (program.processes prog)), pid_global_config.state = \ global_state_impl.variables = init_valuation_impl (program.global_vars prog) \ \" lemma pid_init_gc_impl: "(pid_init_gc_impl, pid_init_gc) \ global_config_rel" unfolding pid_init_gc_impl_def pid_init_gc_def global_config_rel_def global_state_rel_def apply (simp del: pair_in_Id_conv, intro conjI) apply (parametricity add: init_pc_impl) apply simp apply (simp only: vi_rel_rew) apply (parametricity add: init_valuation_impl) apply simp done end lift_definition val_of_bool_impl :: "bool \ uint32" is val_of_bool . lemma [code]: "val_of_bool_impl b = (if b then 1 else 0)" by transfer auto lift_definition bool_of_val_impl :: "uint32 \ bool" is bool_of_val . lemma [code]: "bool_of_val_impl v = (v\0)" by transfer (auto simp: bool_of_val_def) definition "set_of_rel R \ {(a,b) . R b a}" definition "rel_of_set S \ \b a. (a,b)\S" lemma [simp]: "(a,b)\set_of_rel R \ R b a" "rel_of_set S b a \ (a,b)\S" unfolding rel_of_set_def set_of_rel_def by auto lemma [simp]: "set_of_rel (rel_of_set S) = S" "rel_of_set (set_of_rel R) = R" unfolding rel_of_set_def set_of_rel_def by auto lemma rel_fun_2set: "rel_fun A B = rel_of_set (set_of_rel A \ set_of_rel B)" unfolding rel_fun_def fun_rel_def unfolding rel_of_set_def set_of_rel_def by (auto) lemma [simp]: "set_of_rel cr_uint32 = val_rel" "set_of_rel (=) = Id" by (auto simp: br_def cr_uint32_def) lemma bool_of_val_impl: "(bool_of_val_impl, bool_of_val) \ val_rel \ bool_rel" using bool_of_val_impl.transfer by (simp add: rel_fun_2set) lemma smod_by_div_abs: "a smod b = a - a sdiv b * b" - using sdiv_smod_id[of a b] - by (metis add_diff_cancel2) + by (subst (2) sdiv_smod_id [of a b, symmetric]) simp lift_definition sdiv_impl :: "uint32 \ uint32 \ uint32" is "(sdiv)" . lift_definition smod_impl :: "uint32 \ uint32 \ uint32" is "(smod)" . (* TODO: Is there a more efficient way? *) lemma [code]: "sdiv_impl x y = Abs_uint32' (Rep_uint32' x sdiv Rep_uint32' y)" apply transfer by simp lemma [code]: "smod_impl a b = a - sdiv_impl a b * b" - apply transfer - using sdiv_smod_id - by (metis add_diff_cancel2) + by transfer (simp add: smod_by_div_abs) primrec eval_bin_op_impl_aux :: "bin_op \ uint32 \ uint32 \ uint32" where "eval_bin_op_impl_aux bo_plus v1 v2 = v1+v2" | "eval_bin_op_impl_aux bo_minus v1 v2 = v1-v2" | "eval_bin_op_impl_aux bo_mul v1 v2 = v1*v2" | "eval_bin_op_impl_aux bo_div v1 v2 = sdiv_impl v1 v2" | "eval_bin_op_impl_aux bo_mod v1 v2 = smod_impl v1 v2" | "eval_bin_op_impl_aux bo_less v1 v2 = val_of_bool_impl (v1 < v2)" | "eval_bin_op_impl_aux bo_less_eq v1 v2 = val_of_bool_impl (v1 \ v2)" | "eval_bin_op_impl_aux bo_eq v1 v2 = val_of_bool_impl (v1 = v2)" | "eval_bin_op_impl_aux bo_and v1 v2 = v1 AND v2" | "eval_bin_op_impl_aux bo_or v1 v2 = v1 OR v2" | "eval_bin_op_impl_aux bo_xor v1 v2 = v1 XOR v2" lift_definition eval_bin_op_impl :: "bin_op \ uint32 \ uint32 \ uint32" is eval_bin_op . print_theorems lemma [code]: "eval_bin_op_impl bop v1 v2 = eval_bin_op_impl_aux bop v1 v2" apply (cases bop) apply simp_all apply (transfer, simp)+ done primrec eval_un_op_impl_aux :: "un_op \ uint32 \ uint32" where "eval_un_op_impl_aux uo_minus v = -v" | "eval_un_op_impl_aux uo_not v = NOT v" lift_definition eval_un_op_impl :: "un_op \ uint32 \ uint32" is eval_un_op . lemma [code]: "eval_un_op_impl uop v = eval_un_op_impl_aux uop v" apply (cases uop) apply simp_all apply (transfer, simp)+ done fun eval_exp_impl :: "exp \ focused_state_impl \ uint32" where "eval_exp_impl (e_var x) (ls,gs) = do { let lv = local_state_impl.variables ls; let gv = global_state_impl.variables gs; case vi_lookup lv x of Some v \ Some v | None \ (case vi_lookup gv x of Some v \ Some v | None \ None) }" | "eval_exp_impl (e_localvar x) (ls,gs) = vi_lookup (local_state_impl.variables ls) x" | "eval_exp_impl (e_globalvar x) (ls,gs) = vi_lookup (global_state_impl.variables gs) x" | "eval_exp_impl (e_const n) fs = do { assert_option (n\min_signed \ n\max_signed); Some (uint32_of_int n) }" | "eval_exp_impl (e_bin bop e1 e2) fs = do { v1\eval_exp_impl e1 fs; v2\eval_exp_impl e2 fs; Some (eval_bin_op_impl bop v1 v2) }" | "eval_exp_impl (e_un uop e) fs = do { v\eval_exp_impl e fs; Some (eval_un_op_impl uop v) }" (* TODO: Move *) lemma map_option_bind: "map_option f (m\g) = m \ (map_option f o g)" by (auto split: Option.bind_split) lemma val_option_rel_as_eq: "(a,b)\\val_rel\option_rel \ map_option Rep_uint32 a = b" unfolding br_def apply (cases a) apply (cases b, simp_all) apply (cases b, auto) done lemma eval_exp_impl: "(eval_exp_impl, eval_exp) \ Id \ focused_state_rel \ \val_rel\option_rel" proof - { fix e ls gs lsi gsi assume "(lsi,ls)\local_state_rel" "(gsi,gs)\global_state_rel" hence "map_option Rep_uint32 (eval_exp_impl e (lsi,gsi)) = eval_exp e (ls,gs)" apply (induction e) apply (auto simp: local_state_rel_def global_state_rel_def br_def simp: vi_lookup_correct vi_update_correct simp: Abs_uint32_inverse[simplified] uint32_of_int.rep_eq signed_of_int_def simp: option.map_comp option.map_ident o_def map_option_bind simp: eval_bin_op_impl.rep_eq eval_un_op_impl.rep_eq split: option.splits) apply (drule sym) apply (drule sym) apply (simp split: Option.bind_split) apply (drule sym) apply (simp split: Option.bind_split) done } thus ?thesis by (auto simp: val_option_rel_as_eq) qed text \Enabledness and effects of actions\ primrec la_en_impl :: "focused_state_impl \ action \ bool" where "la_en_impl fs (AAssign e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (AAssign_local e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (AAssign_global e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (ATest e) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl _ (ASkip) = Some True" lemma param_rec_action[param]: "(rec_action, rec_action) \ (Id \ Id \ Id \ A) \ (Id \ Id \ Id \ A) \ (Id \ Id \ Id \ A) \ (Id \ A) \ A \ Id \ A" apply (intro fun_relI) subgoal for fi1 f1 fi2 f2 fi3 f3 fi4 f4 fi5 f5 ai a apply simp apply (induction a) apply simp_all apply (parametricity, simp_all?)+ done done (* TODO: Move *) lemma param_option_bind[param]: "(Option.bind, Option.bind) \ \A\option_rel \ (A \ \B\option_rel) \ \B\option_rel" unfolding Option.bind_def by parametricity lemma la_en_impl: "(la_en_impl,la_en) \ focused_state_rel \ Id \ \Id\option_rel" unfolding la_en_impl_def la_en_def by (parametricity add: eval_exp_impl bool_of_val_impl) definition "la_en'_impl fs a \ case la_en_impl fs a of Some b \ b | None \ False" lemma la_en'_impl: "(la_en'_impl,la_en') \ focused_state_rel \ Id \ bool_rel" unfolding la_en'_impl_def[abs_def] la_en'_def[abs_def] by (parametricity add: la_en_impl) abbreviation "exists_var_impl v x \ vi_lookup v x \ None" fun la_ex_impl :: "focused_state_impl \ action \ focused_state_impl" where "la_ex_impl (ls,gs) (AAssign ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); if exists_var_impl (local_state_impl.variables ls) x then do { let ls = local_state_impl.variables_update (vi_update x v) ls; Some (ls,gs) } else do { assert_option (exists_var_impl (global_state_impl.variables gs) x); let gs = global_state_impl.variables_update (vi_update x v) gs; Some (ls,gs) } }" | "la_ex_impl (ls,gs) (AAssign_local ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); assert_option (exists_var_impl (local_state_impl.variables ls) x); let ls = local_state_impl.variables_update (vi_update x v) ls; Some (ls,gs) }" | "la_ex_impl (ls,gs) (AAssign_global ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); assert_option (exists_var_impl (global_state_impl.variables gs) x); let gs = global_state_impl.variables_update (vi_update x v) gs; Some (ls,gs) }" | "la_ex_impl fs (ATest e) = do { v \ eval_exp_impl e fs; assert_option (bool_of_val_impl v); Some fs }" | "la_ex_impl fs ASkip = Some fs" (* TODO: Move *) lemma param_assert_option[param]: "(assert_option, assert_option) \ bool_rel \ \unit_rel\option_rel" unfolding assert_option_def by parametricity lemma la_ex_impl: "(la_ex_impl, la_ex) \ focused_state_rel \ Id \ \focused_state_rel\option_rel" proof (intro fun_relI) fix fsi fs and ai a :: action assume "(fsi,fs)\focused_state_rel" "(ai,a)\Id" thus "(la_ex_impl fsi ai, la_ex fs a)\\focused_state_rel\option_rel" apply (cases fs, cases fsi, clarsimp) apply (cases a) apply simp_all apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct simp: br_def intro: domI) [7] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct br_def intro: domI) [9] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct br_def intro: domI) [4] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply simp done qed definition "la_ex'_impl fs a \ case la_ex_impl fs a of Some fs' \ fs' | None \ fs" lemma la_ex'_impl: "(la_ex'_impl,la_ex') \ focused_state_rel \ Id \ focused_state_rel" unfolding la_ex'_impl_def[abs_def] la_ex'_def[abs_def] by (parametricity add: la_ex_impl) lemma param_collecti_index: "(collecti_index, collecti_index) \ (nat_rel \ B \ bool_rel \\<^sub>r \Id\set_rel) \ \B\list_rel \ \nat_rel \\<^sub>r Id\set_rel" unfolding collecti_index'_def apply parametricity apply auto done export_code la_ex_impl la_en_impl checking SML context visible_prog begin definition "ga_gample_mi3 mem sticky_E (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c); s={(a,c')\as. la_en'_impl (ls,gs) a}; ample = s\{} \ (\(a,_)\as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\s. \ mem (c,c') sticky_E) in (ample,{c}\s) ) lcs )" lemma ga_gample_mi3_refine: "(ga_gample_mi3, ga_gample_mi2)\ Id \ Id \ global_config_rel \ \nat_rel \\<^sub>r Id\set_rel" proof - have [param]: "(comp.succ_impl,comp.succ_impl)\Id\\Id\\<^sub>rId\list_rel" by simp from la_en'_impl have aux1: "\fsi fs a. (fsi,fs)\focused_state_rel \ la_en'_impl fsi a = la_en' fs a" apply (rule_tac IdD) apply parametricity by auto show ?thesis using [[goals_limit=1]] apply (intro fun_relI) unfolding ga_gample_mi3_def ga_gample_mi2_def apply (parametricity add: param_collecti_index la_en'_impl) apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply simp using aux1 apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply simp_all done qed definition (in cprog) "is_vis_edge is_vis_var \ \(c,c'). \(a,cc')\set (comp.succ_impl c). c'=cc' \ (\v\write_globals a. is_vis_var v)" lemma is_vis_edge_correct: "is_vis_edge is_vis_var = (\x. x\vis_edges)" unfolding is_vis_edge_def vis_edges_def by (auto intro!: ext simp: comp.astep_impl_def) definition "cr_ample_impl3 \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (\x. x \ vis_edges) in stutter_extend_en (ga_gample_mi3 (\x f. f x) sticky)" lemma cr_ample_impl3_refine: "(cr_ample_impl3, cr_ample_impl2) \ global_config_rel \ \\nat_rel\\<^sub>rId\option_rel\set_rel" proof - have [param]: "\R S. S=Id \ (stutter_extend_en, stutter_extend_en) \ (R \ \S\set_rel) \ R \ \\S\option_rel\set_rel" apply (simp only: ) by (rule stutter_extend_en_refine) show ?thesis unfolding cr_ample_impl3_def cr_ample_impl2_def is_vis_edge_correct apply (parametricity add: stutter_extend_en_refine ga_gample_mi3_refine) apply (rule IdI) by auto qed (* TODO: The current Pid_Scheduler-locale is focused on refinement. We cannot use it to get an instantiation of ga_gex_impl.*) definition (in -) ga_gex_impl :: "global_action \ pid_global_config_impl \ pid_global_config_impl" where "ga_gex_impl ga gc \ let (pid,c,a,c') = ga; fs = fs_of_pid gc pid; (ls',gs') = la_ex'_impl fs a; gc' = \ pid_global_config.processes = (pid_global_config.processes gc) [pid := \ local_config.command = c', local_config.state = ls'\ ], pid_global_config.state = gs'\ in gc'" lemma lc_of_pid_refine: assumes V: "pid_valid gc pid" assumes [param]: "(pidi,pid)\nat_rel" assumes GCI: "(gci,gc)\global_config_rel" shows "(lc_of_pid gci pidi, lc_of_pid gc pid)\local_config_rel" apply parametricity apply fact using GCI by (auto simp: global_config_rel_def) lemma ls_of_pid_refine: assumes "pid_valid gc pid" assumes "(pidi,pid)\nat_rel" assumes "(gci,gc)\global_config_rel" shows "(ls_of_pid gci pidi, ls_of_pid gc pid)\local_state_rel" using lc_of_pid_refine[OF assms] unfolding local_config_rel_def by auto lemma ga_gex_impl: assumes V: "pid_valid gc pid" assumes [param]: "(pidi,pid)\Id" "(ci,c)\Id" "(ai,a)\Id" and CI'I: "(ci',c')\Id" and GCI[param]: "(gci,gc)\global_config_rel" shows "(ga_gex_impl (pidi,ci,ai,ci') gci, ga_gex (pid,c,a,c') gc) \ global_config_rel" unfolding ga_gex_impl_def[abs_def] ga_gex_def[abs_def] apply simp apply (parametricity add: la_ex'_impl ls_of_pid_refine[OF V]) using GCI CI'I apply (auto simp: global_config_rel_def local_config_rel_def) [] apply parametricity apply auto [] using GCI apply (auto simp: global_config_rel_def) [] done definition (in -) "ga_ex_impl \ stutter_extend_ex ga_gex_impl" fun ga_valid where "ga_valid gc None = True" | "ga_valid gc (Some (pid,c,a,c')) = pid_valid gc pid" lemma ga_ex_impl: assumes V: "ga_valid gc ga" assumes P: "(gai,ga)\Id" "(gci,gc)\global_config_rel" shows "(ga_ex_impl gai gci, ga_ex ga gc) \ global_config_rel" using assms apply (cases "(gc,ga)" rule: ga_valid.cases) using P apply (simp add: ga_ex_impl_def) apply (simp add: ga_ex_impl_def ga_ex_def) apply (parametricity add: ga_gex_impl) by auto definition "ample_step_impl3 \ rel_of_enex (cr_ample_impl3,ga_ex_impl)" definition (in -) pid_interp_gc_impl :: "(ident \ bool) \ pid_global_config_impl \ exp set" where "pid_interp_gc_impl is_vis_var gci \ sm_props ( vi_\ (global_state_impl.variables (pid_global_config.state gci)) |` Collect is_vis_var )" lemma pid_interp_gc_impl: "(pid_interp_gc_impl is_vis_var, pid_interp_gc) \ global_config_rel \ Id" apply rule unfolding pid_interp_gc_impl_def pid_interp_gc_def apply (auto simp: global_config_rel_def global_state_rel_def interp_gs_def sm_props_def br_def) done lemma cr_ample_impl3_sim_aux: assumes GCI: "(gci, gc) \ global_config_rel" assumes GA: "ga \ cr_ample_impl3 gci" shows "\gc'. ga\cr_ample_impl2 gc \ gc'=ga_ex ga gc \ (ga_ex_impl ga gci, gc')\global_config_rel" proof (clarsimp, safe) show G1: "ga\cr_ample_impl2 gc" using cr_ample_impl3_refine[param_fo, OF GCI] GA by auto { (* TODO: We are re-proving the validity of the action, due to the lack of a clean refinement argumentation *) fix mem sticky_E gc pid c assume "(pid,c)\ga_gample_m mem sticky_E gc" hence "pid_valid gc pid" unfolding ga_gample_m_def apply (auto split: option.splits) apply (auto simp: ga_gen_def) [] apply (auto dest: find_min_idx_f_SomeD) done } note aux=this from G1 have "ga_valid gc ga" unfolding cr_ample_impl2_def stutter_extend_en_def apply (simp split: if_split_asm) unfolding ga_gample_mi2_impl[symmetric] apply (auto dest: aux) done thus "(ga_ex_impl ga gci, ga_ex ga gc)\global_config_rel" by (parametricity add: ga_ex_impl GCI IdI[of ga]) qed lemma cr_ample_impl3_sim_aux2: assumes GCI: "(gci, gc) \ global_config_rel" assumes GA: "ga \ cr_ample_impl2 gc" shows "ga\cr_ample_impl3 gci \ (ga_ex_impl ga gci, ga_ex ga gc) \ global_config_rel" proof show G1: "ga\cr_ample_impl3 gci" using cr_ample_impl3_refine[param_fo, OF GCI] GA by auto from cr_ample_impl3_sim_aux[OF GCI G1] show "(ga_ex_impl ga gci, ga_ex ga gc) \ global_config_rel" by auto qed sublocale impl3: sa "\ g_V = UNIV, g_E = ample_step_impl3, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" by unfold_locales auto sublocale impl3_sim: lbisimulation global_config_rel "\ g_V = UNIV, g_E = ample_step_impl3, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" "\ g_V = UNIV, g_E = ample_step_impl2, g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" apply unfold_locales apply simp apply simp apply simp using pid_init_gc_impl apply (auto dest: fun_relD) [] unfolding ample_step_impl2_def ample_step_impl3_def apply (auto dest: cr_ample_impl3_sim_aux) [] using pid_interp_gc_impl apply (auto dest: fun_relD) [] apply simp using pid_init_gc_impl apply (auto dest: fun_relD) [] defer using pid_interp_gc_impl apply (auto dest: fun_relD) [] apply auto apply (auto dest: cr_ample_impl3_sim_aux2) [] done text \Correctness of impl3\ (* TODO: Locale hl_reduction seems very special, and we cannot use it here*) lemma impl3_accept_subset: "impl3.accept w \ lv.sa.accept w" using impl3_sim.accept_bisim cr_ample_impl2_reduction hl_reduction.ample_accept_subset by simp lemma impl3_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ impl3.accept w'" using impl3_sim.accept_bisim cr_ample_impl2_reduction hl_reduction.ample_accept_stuttering by simp text \Next, we go from sets of actions to lists of actions\ definition (in cprog) "ga_gample_mil4 mem sticky_Ei (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index_list (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = comp.succ_impl c; s = filter (la_en'_impl (ls,gs) o fst) as; ample = s\[] \ (\(a,_)\set as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\set s. \mem (c,c') sticky_Ei) in (ample,map (Pair c) s) ) lcs )" lemma ga_gample_mil4_refine: "(ga_gample_mil4, ga_gample_mi3) \ (Id\\<^sub>rId \ Id \ bool_rel) \ Id \ Id \ \nat_rel \\<^sub>r Id\list_set_rel" apply (intro fun_relI) unfolding ga_gample_mil4_def ga_gample_mi3_def apply (parametricity) apply simp apply (rule IdI) apply simp apply (rule IdI) apply (intro fun_relI) apply (rule collecti_index_list_refine[param_fo]) apply (intro fun_relI) apply (simp add: list_set_rel_def br_def distinct_map comp.succ_impl_invar) apply (auto simp: filter_empty_conv) apply force done definition (in cprog) "cr_ample_impl4 is_vis_var \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (is_vis_edge is_vis_var) in stutter_extend_en_list (ga_gample_mil4 (\x f. f x) sticky)" lemma cr_ample_impl4_refine: "(cr_ample_impl4 is_vis_var, cr_ample_impl3)\Id \ \\nat_rel \\<^sub>r Id\option_rel\list_set_rel" unfolding cr_ample_impl4_def cr_ample_impl3_def apply (parametricity add: stutter_extend_en_list.refine ga_gample_mil4_refine) apply (simp_all add: is_vis_edge_correct) done text \Finally, we combine the ample-implementation and the executable implementation to get a step function\ definition (in cprog) "ample_step_impl4 is_vis_var \ rel_of_enex_list (cr_ample_impl4 is_vis_var,ga_ex_impl)" lemma ample_step_impl4: "(ample_step_impl4 is_vis_var, ample_step_impl3) \ \Id \\<^sub>r Id\ set_rel" unfolding ample_step_impl4_def ample_step_impl3_def rel_of_pred_def apply (parametricity add: rel_of_enex_list_refine cr_ample_impl4_refine[simplified]) by auto sublocale impl4: sa "\ g_V = UNIV, g_E = ample_step_impl4 is_vis_var, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" by unfold_locales auto lemma impl4_accept_subset: "impl4.accept w \ lv.sa.accept w" using impl3_accept_subset ample_step_impl4 by simp lemma impl4_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ impl4.accept w'" using impl3_accept_stuttering ample_step_impl4 by simp lemma (in -) pred_of_succ_of_enex_list: fixes enex :: "('c \ 'a list) \ ('a \ 'c \ 'c)" shows "pred_of_succ (set o succ_of_enex_list enex) = pred_of_enex_list enex" proof - { fix x and l have "glist_member (=) x l \ x\set l" by (induction l) auto } note [simp] = this { fix l :: "'a list" and g and l0 have "set (foldli l (\_. True) (\x. glist_insert (=) (g x)) l0) = set l0 \ g`set l" by (induction l arbitrary: l0) (auto simp: glist_insert_def) } note [simp] = this { fix l::"'a list" and P i have "foldli l Not (\x _. P x) i \ i \ (\x\set l. P x)" by (induction l arbitrary: i) auto } note [simp] = this show ?thesis unfolding succ_of_enex_list_def[abs_def] pred_of_enex_list_def[abs_def] by (auto simp: pred_of_succ_def gen_image_def gen_bex_def intro!: ext) qed lemma (in -) rel_of_succ_of_enex_list: fixes enex :: "('c \ 'a list) \ ('a \ 'c \ 'c)" shows "rel_of_succ (set o succ_of_enex_list enex) = rel_of_enex_list enex" unfolding rel_of_enex_list_def unfolding pred_of_succ_of_enex_list[symmetric] by simp definition (in cprog) "impl4_succ is_vis_var \ succ_of_enex_list (cr_ample_impl4 is_vis_var,ga_ex_impl)" lemma impl4_succ_pred: "rel_of_succ (set o (impl4_succ is_vis_var)) = ample_step_impl4 is_vis_var" unfolding impl4_succ_def ample_step_impl4_def by (simp add: rel_of_succ_of_enex_list) end export_code ga_ex_impl checking SML definition "ccfg_E_succ_impl mst \ remdups o map snd o ccfg_succ_impl mst" lemma (in cprog) ccfg_E_succ_impl: "ccfg_E_succ_impl (comp_info_of prog) = cfgc_E_succ" apply (intro ext) unfolding ccfg_E_succ_impl_def cfgc_E_succ_def by (simp add: ccfg_succ_impl) definition init_pc_impl_impl :: "comp_info \ proc_decl \ local_config_impl" where "init_pc_impl_impl mst pd \ \ local_config.command = comp_\_impl mst (proc_decl.body pd), local_config.state = \ local_state_impl.variables = init_valuation_impl (proc_decl.local_vars pd) \ \" lemma (in cprog) init_pc_impl_impl: "init_pc_impl_impl (comp_info_of prog) = init_pc_impl" apply (intro ext) unfolding init_pc_impl_impl_def init_pc_impl_def by (simp add: comp_\_impl) definition pid_init_gc_impl_impl :: "program \ comp_info \ pid_global_config_impl" where "pid_init_gc_impl_impl prog mst \ \ pid_global_config.processes = (map (init_pc_impl_impl mst) (program.processes prog)), pid_global_config.state = \ global_state_impl.variables = init_valuation_impl (program.global_vars prog) \ \" lemma (in cprog) pid_init_gc_impl_impl: "pid_init_gc_impl_impl prog (comp_info_of prog) = pid_init_gc_impl" unfolding pid_init_gc_impl_impl_def pid_init_gc_impl_def by (simp add: init_pc_impl_impl) definition "ccfg_V0_impl prog mst \ map (comp_\_impl mst) (cfg_V0_list prog)" lemma (in cprog) ccfg_V0_impl: "ccfg_V0_impl prog (comp_info_of prog) = cfgc_V0_list" unfolding ccfg_V0_impl_def cfgc_V0_list_def by (simp add: comp_\_impl) definition "ccfg_G_impl_impl prog mst \ \ gi_V = \ _. True, gi_E = ccfg_E_succ_impl mst, gi_V0 = ccfg_V0_impl prog mst \" lemma (in cprog) ccfg_G_impl_impl: "ccfg_G_impl_impl prog (comp_info_of prog) = cfgc_G_impl" unfolding ccfg_G_impl_impl_def cfgc_G_impl_def unfolding ccfg_E_succ_impl ccfg_V0_impl by rule definition "is_vis_edge_impl mst is_vis_var \ \(c,c'). \(a,cc')\set (ccfg_succ_impl mst c). c'=cc' \ (\v\write_globals a. is_vis_var v)" lemma (in cprog) is_vis_edge_impl: "is_vis_edge_impl (comp_info_of prog) = is_vis_edge" apply (intro ext) unfolding is_vis_edge_impl_def is_vis_edge_def by (simp add: ccfg_succ_impl) definition "ga_gample_mil4_impl mst mem sticky_Ei (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index_list (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = ccfg_succ_impl mst c; s = filter (la_en'_impl (ls,gs) o fst) as; ample = s\[] \ (\(a,_)\set as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\set s. \mem (c,c') sticky_Ei) in (ample,map (Pair c) s) ) lcs )" lemma (in cprog) ga_gample_mil4_impl: "ga_gample_mil4_impl (comp_info_of prog) = ga_gample_mil4" apply (intro ext) unfolding ga_gample_mil4_impl_def ga_gample_mil4_def by (simp add: ccfg_succ_impl) definition "cr_ample_impl4_impl prog mst is_vis_var \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) (ccfg_G_impl_impl prog mst) (is_vis_edge_impl mst is_vis_var) in stutter_extend_en_list (ga_gample_mil4_impl mst (\x f. f x) sticky)" lemma (in cprog) cr_ample_impl4_impl: "cr_ample_impl4_impl prog (comp_info_of prog) = cr_ample_impl4" apply (intro ext) unfolding cr_ample_impl4_impl_def cr_ample_impl4_def by (simp add: ga_gample_mil4_impl ccfg_G_impl_impl is_vis_edge_impl) definition "impl4_succ_impl prog mst is_vis_var \ succ_of_enex_list (cr_ample_impl4_impl prog mst is_vis_var,ga_ex_impl)" lemma (in cprog) impl4_succ_impl: "impl4_succ_impl prog (comp_info_of prog) = impl4_succ" apply (intro ext) unfolding impl4_succ_impl_def impl4_succ_def by (simp add: cr_ample_impl4_impl) definition "ample_step_impl4_impl prog mst is_vis_var \ rel_of_enex_list (cr_ample_impl4_impl prog mst is_vis_var,ga_ex_impl)" lemma (in cprog) ample_step_impl4_impl: "ample_step_impl4_impl prog (comp_info_of prog) = ample_step_impl4" apply (intro ext) unfolding ample_step_impl4_impl_def ample_step_impl4_def by (simp add: cr_ample_impl4_impl) export_code impl4_succ_impl pid_init_gc_impl_impl comp_info_of checking SML end diff --git a/thys/CAVA_LTL_Modelchecker/SM/Refine/Rename_Cfg.thy b/thys/CAVA_LTL_Modelchecker/SM/Refine/Rename_Cfg.thy --- a/thys/CAVA_LTL_Modelchecker/SM/Refine/Rename_Cfg.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/Refine/Rename_Cfg.thy @@ -1,737 +1,737 @@ theory Rename_Cfg imports "../RefPoint/SM_Cfg" Type_System SM_Finite_Reachable Gen_Cfg_Bisim "HOL-Library.IArray" begin (* TODO: Move *) lemma set_list_bind[simp]: "set (List.bind m f) = Set.bind (set m) (\x. set (f x))" by (auto simp: List.bind_def) lemma rel_mset_of_mapI: assumes "\x\set l. R (f x) (g x)" shows "rel_mset R (mset (map f l)) (mset (map g l))" using assms apply (induction l) apply (auto simp: rel_mset_Zero rel_mset_Plus) done (* TODO: Move? *) lemma finite_max1_elemI: assumes "(\x y. x\S \ y\S \ x=y)" shows "finite S" proof (cases "S={}") case False with assms obtain x where "S={x}" by auto thus ?thesis by auto qed auto lemma left_unique_finite_sngI: assumes "left_unique R" shows "finite {a. R a b}" proof (rule finite_subset) { fix x y assume "x\{a. R a b}" "y\{a. R a b}" hence "x=y" using assms by (auto dest: left_uniqueD) } thus "finite {a. R a b}" by (blast intro: finite_max1_elemI) qed auto function approx_reachable_list_aux :: "cmd \ cmd list" where "approx_reachable_list_aux (Assign c x e) = [Assign c x e, Skip]" | "approx_reachable_list_aux (Assign_local c x e) = [Assign_local c x e, Skip]" | "approx_reachable_list_aux (Assign_global c x e) = [Assign_global c x e, Skip]" | "approx_reachable_list_aux (Test e) = [Test e, Skip]" | "approx_reachable_list_aux (Skip) = [Skip]" | "approx_reachable_list_aux (Seq c1 c2) = (approx_reachable_list_aux c2) @ (do { c1\approx_reachable_list_aux c1; [Seq c1 c2] })" | "approx_reachable_list_aux (Or c1 c2) = (Or c1 c2) # (approx_reachable_list_aux c1 @ approx_reachable_list_aux c2)" | "approx_reachable_list_aux (Break) = [Break,Invalid]" | "approx_reachable_list_aux (Continue) = [Continue,Invalid]" | "approx_reachable_list_aux (Iterate c1 c2) = Skip # (do { c1\approx_reachable_list_aux c1; [Iterate c1 c2] }) @ (do { c2'\approx_reachable_list_aux c2; [Iterate c2' c2] })" | "approx_reachable_list_aux Invalid = [Invalid]" by pat_completeness auto termination apply (relation "reachable_term_order_aux <*mlex*> measure size") apply (metis wf_measure wf_mlex) apply (auto simp: mlex_prod_def) done lemma approx_reachable_list_aux_refine: "set (approx_reachable_list_aux c) = approx_reachable c" apply (induction c rule: approx_reachable.induct) apply (simp_all) done definition "cfg_V0 prog \ proc_decl.body ` set (program.processes prog)" definition "cfg_V0_list prog \ remdups (map proc_decl.body (program.processes prog))" lemma cfg_V0_list_invar: "distinct (cfg_V0_list prog)" unfolding cfg_V0_list_def by auto lemma cfg_V0_list_refine: "set (cfg_V0_list prog) = cfg_V0 prog" unfolding cfg_V0_list_def cfg_V0_def by auto definition "approx_reachable_list prog \ remdups (concat (map approx_reachable_list_aux (cfg_V0_list prog)))" definition "approx_reachable_prog prog \ \(approx_reachable`(cfg_V0 prog))" lemma approx_reachable_list_invar: "distinct (approx_reachable_list prog)" unfolding approx_reachable_list_def by auto lemma approx_reachable_list_refine: "set (approx_reachable_list prog) = approx_reachable_prog prog" unfolding approx_reachable_list_def approx_reachable_prog_def by (auto simp: approx_reachable_list_aux_refine cfg_V0_list_refine) text \CFG successors\ primrec cfg_succ :: "cmd \ ((action+brk_ctd) \ cmd) set" where "cfg_succ (Assign c x e) = {((Inl (AAssign c x e)), Skip)}" | "cfg_succ (Assign_local c x e) = {((Inl (AAssign_local c x e)), Skip)}" | "cfg_succ (Assign_global c x e) = {((Inl (AAssign_global c x e)), Skip)}" | "cfg_succ (Test e) = {((Inl (ATest e)), Skip)}" | "cfg_succ (Seq c1 c2) = ( if c1=Skip then {((Inl ASkip), c2)} else { (e,if c1'=Skip then c2 else Seq c1' c2) | c1' e. (e,c1')\cfg_succ c1} )" | "cfg_succ (Or c1 c2) = cfg_succ c1 \ cfg_succ c2" | "cfg_succ (Break) = {(Inr AIBreak, Invalid)}" | "cfg_succ (Continue) = {(Inr AIContinue, Invalid)}" | cfg_succ_iterate: "cfg_succ (Iterate c cb) = (if c=Skip then {(Inl ASkip, Iterate cb cb)} else {}) \ ( \ (Inl e,c') \ (Inl e, if c'=Skip then Iterate cb cb else Iterate c' cb) | (Inr AIBreak,_) \ (Inl ASkip,Skip) | (Inr AIContinue,_) \ (Inl ASkip, Iterate cb cb) )`cfg_succ c" | "cfg_succ Skip = {}" | "cfg_succ Invalid = {}" declare cfg_succ_iterate[simp del] lemma cfg_succ_iterate'[simp]: "cfg_succ (Iterate c cb) = (if c=Skip then {(Inl ASkip, Iterate cb cb)} else {}) \ { (Inl e, if c'=Skip then Iterate cb cb else Iterate c' cb) | e c'. (Inl e, c') \ cfg_succ c } \ (if \c'. (Inr AIBreak,c')\ cfg_succ c then { (Inl ASkip, Skip)} else {}) \ (if \c'. (Inr AIContinue,c')\ cfg_succ c then { (Inl ASkip, Iterate cb cb)} else {})" unfolding cfg_succ_iterate apply (rule sym) apply (cases "c=Skip") apply simp apply safe apply simp_all apply force apply (force split: if_split_asm) apply (force split: if_split_asm) apply (auto split: if_split_asm sum.splits brk_ctd.splits) done lemma cfg_succ_iff_cfg: "(a,c')\cfg_succ c \ cfg c a c'" apply rule apply (induction c arbitrary: a c') apply ((auto intro: cfg.intros split: if_split_asm) [])+ apply (induction rule: cfg.induct) apply auto done theorem cfg_succ_correct: "cfg_succ c = {(a,c'). cfg c a c'}" using cfg_succ_iff_cfg by auto lemma cfg_succ_finite[simp, intro!]: "finite (cfg_succ c)" apply (induction c) using [[simproc add: finite_Collect]] apply (simp_all del: cfg_succ_iterate' add: cfg_succ_iterate) done primrec cfg_succ_list :: "cmd \ ((action+brk_ctd) \ cmd) list" where "cfg_succ_list (Assign ge x e) = [((Inl (AAssign ge x e)), Skip)]" | "cfg_succ_list (Assign_local ge x e) = [((Inl (AAssign_local ge x e)), Skip)]" | "cfg_succ_list (Assign_global ge x e) = [((Inl (AAssign_global ge x e)), Skip)]" | "cfg_succ_list (Test e) = [((Inl (ATest e)), Skip)]" | "cfg_succ_list (Seq c1 c2) = ( if c1=Skip then [((Inl ASkip), c2)] else map (\(e,c1'). (e,if c1'=Skip then c2 else Seq c1' c2)) (cfg_succ_list c1) )" | "cfg_succ_list (Or c1 c2) = remdups (cfg_succ_list c1 @ cfg_succ_list c2)" | "cfg_succ_list (Break) = [(Inr AIBreak, Invalid)]" | "cfg_succ_list (Continue) = [(Inr AIContinue, Invalid)]" | cfg_succ_list_iterate: "cfg_succ_list (Iterate c cb) = (if c=Skip then [(Inl ASkip, Iterate cb cb)] else []) @ remdups (map ( \ (Inl e,c') \ (Inl e, if c'=Skip then Iterate cb cb else Iterate c' cb) | (Inr AIBreak,_) \ (Inl ASkip,Skip) | (Inr AIContinue,_) \ (Inl ASkip, Iterate cb cb) ) (cfg_succ_list c))" | "cfg_succ_list Skip = []" | "cfg_succ_list Invalid = []" (* FIXME: Strange, the simplifier should know such lemmas!? *) lemma Seq_ne_structural[simp]: "Seq a b \ a" "Seq a b \ b" "a \ Seq a b" "b \ Seq a b" proof - show G1: "Seq a b \ a" by (induction a arbitrary: b) auto show G2: "Seq a b \ b" by (induction b arbitrary: a) auto show "a \ Seq a b" using G1 by simp show "b \ Seq a b" using G2 by simp qed lemma cfg_succ_list_invar: "distinct (cfg_succ_list c)" apply (induction c) apply (auto simp: distinct_map) apply (rule inj_onI) apply (auto split: if_split_asm ) [] done lemma cfg_succ_list_refine: "set (cfg_succ_list c) = cfg_succ c" apply (induction c) apply (simp_all only: cfg_succ.simps cfg_succ_list.simps - if_distrib[of set] if_same_eq set_map) + if_distrib[of set] set_map) apply (simp_all) apply (auto, force+) [] done definition "cfg'_succ_list c \ map (map_prod projl id) (filter (isl o fst) (cfg_succ_list c))" lemma cfg'_succ_list_invar: "distinct (cfg'_succ_list c)" unfolding cfg'_succ_list_def apply (auto simp: distinct_map cfg_succ_list_invar intro!: inj_onI) apply (rename_tac a a' foo) apply (case_tac a, simp_all) done lemma cfg'_succ_list_refine: "set (cfg'_succ_list c) = {(a,c'). cfg' c a c'}" unfolding cfg'_succ_list_def by (force simp: cfg_succ_list_refine cfg'_def cfg_succ_correct) definition "index_of l x \ if \i l!i=x) else None" lemma index_of_Some_conv: assumes "distinct l" shows "index_of l x = Some i \ i < length l \ l ! i = x" proof safe show "index_of l x = Some i \ i < length l" using assms unfolding index_of_def distinct_conv_nth by (metis (mono_tags, lifting) option.distinct(1) option.inject theI) show "index_of l x = Some i \ l ! i = x" using assms unfolding index_of_def distinct_conv_nth by (metis (mono_tags, lifting) option.distinct(1) option.inject theI) show "i < length l \ x = l ! i \ index_of l (l ! i) = Some i" using assms unfolding index_of_def distinct_conv_nth by auto qed lemma dom_index_of: "dom (index_of l) = set l" by (auto simp: index_of_def in_set_conv_nth split: if_split_asm) lemma index_of_None_conv: "index_of l x = None \ x\set l" using dom_index_of[of l] by auto lemma index_of_Nil: "index_of [] x = None" by (auto simp: index_of_None_conv) lemma index_of_Cons: "distinct (a#l) \ index_of (a#l) x = (if x=a then Some 0 else map_option Suc (index_of l x))" apply (auto simp: index_of_Some_conv) apply (cases "index_of (a # l) x") apply (simp_all only: ) apply (cases "index_of l x") apply (simp_all only: ) [2] apply (auto simp: index_of_Some_conv index_of_None_conv) [2] apply (cases "index_of l x") apply (simp_all only: ) apply (auto simp: index_of_Some_conv index_of_None_conv in_set_conv_nth nth_Cons nth_eq_iff_index_eq split: nat.splits) done primrec cr_index_of where "cr_index_of [] x = None" | "cr_index_of (a#l) x = (if x=a then Some 0 else map_option Suc (cr_index_of l x))" lemma cr_index_of_eq: "distinct l \ cr_index_of l = index_of l" apply (induction l) apply (auto simp: index_of_Nil index_of_Cons) done locale sl_graph = asystem + fixes rlist and succ_list assumes rl_distinct: "distinct rlist" assumes rl_reachable: "E\<^sup>* `` V0 \ set rlist" assumes rl_closed: "\c c'. \c\set rlist; (c, c') \ step\ \ c'\set rlist" assumes sl_distinct: "c\set rlist \ distinct (succ_list c)" assumes sl_succs: "set (succ_list c) = {(a,c'). astep c a c'}" begin definition "succs_tab \ map succ_list rlist" lemma astep_succs_tab_conv: assumes R: "c\set rlist" shows "astep c a c' \ (\i (a,c')\set (succs_tab!i))" proof - from R obtain i where "ij. \c = rlist!j; j \ i=j" proof - fix j assume "c = rlist!j" "ji \j] show "i=j" by simp qed from 1 show ?thesis using sl_succs by (auto dest: AUX) qed lemma succs_tab_invar: "i distinct (succs_tab!i)" unfolding succs_tab_def using rl_reachable sl_distinct by auto lemma succs_tab_rl: "\iset (succs_tab!i)\ \ c' \ set rlist" unfolding succs_tab_def apply (simp add: sl_succs) apply (rule rl_closed[rotated]) unfolding step_def apply blast apply simp done definition "mapped_succs_tab \ IArray ( map (map (\(a,c'). (a,the (cr_index_of rlist c')))) succs_tab)" lemma [simp]: "IArray.length mapped_succs_tab = length succs_tab" "length succs_tab = length rlist" unfolding mapped_succs_tab_def succs_tab_def by auto definition "\ c \ the (cr_index_of rlist c)" definition "\ i \ rlist!i" definition "succ_impl ci \ if ci (a,c')\set (succ_impl c)" lemma astep_succ_finite[simp, intro!]: "finite {(a,c'). astep_impl c a c'}" unfolding astep_impl_def by auto lemma astep_succ_finite_c[simp, intro!]: "finite ({c'. \a. astep_impl c a c'})" proof - have "{c'. \a. astep_impl c a c'} = snd`{(a,c'). astep_impl c a c'}" by auto also have "finite \" by auto finally show ?thesis . qed lemma \_inj: "inj_on \ {0.._def nth_eq_iff_index_eq[OF rl_distinct]) lemma \_inj: "inj_on \ (set rlist)" using rl_reachable apply (auto intro!: inj_onI simp: \_def rl_distinct cr_index_of_eq) by (metis index_of_None_conv index_of_Some_conv not_None_eq option.sel rl_distinct) lemma \_\_inverse: "i \ (\ i) = i" unfolding \_def \_def apply (auto simp: cr_index_of_eq[OF rl_distinct]) apply (cases "index_of rlist (rlist ! i)", simp_all) apply (auto simp: index_of_None_conv index_of_Some_conv[OF rl_distinct] simp: nth_eq_iff_index_eq[OF rl_distinct]) done lemma \_\_inverse: "c\set rlist \ \ (\ c) = c" unfolding \_def \_def apply (auto simp: cr_index_of_eq[OF rl_distinct]) apply (cases "index_of rlist c", simp_all) apply (auto simp: index_of_None_conv index_of_Some_conv[OF rl_distinct] simp: nth_eq_iff_index_eq[OF rl_distinct]) done lemma \_invar: "c\set rlist \ \ c < length rlist" apply (auto simp: \_def cr_index_of_eq rl_distinct) by (metis index_of_None_conv index_of_Some_conv not_None_eq option.sel rl_distinct) lemma \_invar: "i \ i \ set rlist" unfolding \_def by simp definition "invar i \ i brp \ invar" definition "reli \ brp \ (\x. x\set rlist)" lemma bi_unique_rel: "bi_unique rel" apply (intro bi_uniqueI left_uniqueI right_uniqueI) by (auto simp: rel_def build_relp_def invar_def inj_onD[OF \_inj]) lemma bi_unique_reli: "bi_unique reli" apply (intro bi_uniqueI left_uniqueI right_uniqueI) by (auto simp: reli_def build_relp_def inj_onD[OF \_inj]) lemma rel_inv_eq_reli: "rel\\ = reli" by (auto intro!: ext simp: rel_def reli_def build_relp_def invar_def simp: \_\_inverse \_\_inverse \_invar \_invar ) corollary reli_inv_eq_rel: "reli\\ = rel" using arg_cong[where f="\r. r\\", OF rel_inv_eq_reli, simplified] by simp lemma rel_unfold1: "c\set rlist \ rel i c \ i = \ c" unfolding reli_inv_eq_rel[symmetric] by (simp add: reli_def build_relp_def) lemma is_reachable_criE: assumes R: "c \ E\<^sup>* `` V0" obtains i where "cr_index_of rlist c = Some i" proof - from R have IS: "c\set rlist" using rl_reachable by auto show ?thesis using IS by (auto simp: cr_index_of_eq[OF rl_distinct] index_of_Some_conv[OF rl_distinct] simp: in_set_conv_nth intro: that) qed lemma sim1i: "\ astep c1 a c2; reli c1 c1' \ \ \c2'. reli c2 c2' \ astep_impl c1' a c2'" apply (auto simp: reli_def build_relp_def) apply (rule rl_closed, auto simp: step_def) [] proof - assume R: "c1\set rlist" and S: "astep c1 a c2" then obtain i where L: "iset (succs_tab!i)" by (auto simp: astep_succs_tab_conv) have [simp]: "i c1 = i" unfolding \_def apply (cases "cr_index_of rlist c1") apply simp_all apply (auto simp: cr_index_of_eq[OF rl_distinct] index_of_Some_conv[OF rl_distinct] index_of_None_conv nth_eq_iff_index_eq[OF rl_distinct]) done show "astep_impl (\ c1) a (\ c2)" using AC' apply (auto simp: astep_impl_def succ_impl_def mapped_succs_tab_def) apply (auto simp: \_def) done qed lemma sim2i: "\ astep_impl c1' a c2'; reli c1 c1' \ \ \c2. reli c2 c2' \ astep c1 a c2" proof - assume "reli c1 c1'" hence R: "c1\set rlist" and [simp]: "c1'=\ c1" by (auto simp: reli_def build_relp_def) hence [simp]: "\ c1 < length rlist" unfolding \_def using rl_reachable apply (cases "cr_index_of rlist c1", simp_all) apply (auto simp: \_def cr_index_of_eq[OF rl_distinct] index_of_None_conv index_of_Some_conv[OF rl_distinct]) done have [simp]: "rlist!\ c1 = c1" unfolding \_def[symmetric] by (simp add: \_\_inverse R) assume "astep_impl c1' a c2'" then obtain c2 where "(a, c2) \ set (succs_tab ! \ c1)" and [simp]: "c2' = \ c2" apply (auto simp: astep_impl_def succ_impl_def mapped_succs_tab_def) apply (auto simp: \_def) done with astep_succs_tab_conv[THEN iffD2, OF R, OF exI, of c1' a c2] have "astep c1 a c2" by auto moreover hence "c2\set rlist" using rl_closed[OF R, of c2] by (auto simp: step_def) hence "reli c2 c2'" by (auto simp: reli_def build_relp_def) ultimately show "\c2. reli c2 c2' \ astep c1 a c2" by blast qed corollary sim1: "\ astep_impl c1' a c2'; rel c1' c1 \ \ \c2. rel c2' c2 \ astep c1 a c2" using sim2i unfolding reli_inv_eq_rel[symmetric] by simp corollary sim2: "\ astep c1 a c2; rel c1' c1 \ \ \c2'. rel c2' c2 \ astep_impl c1' a c2'" using sim1i unfolding reli_inv_eq_rel[symmetric] by simp lemma succ_impl_invar: "distinct (succ_impl ci)" unfolding succ_impl_def mapped_succs_tab_def invar_def using succs_tab_invar[of ci] apply (clarsimp simp: distinct_map) apply (fold \_def) apply (rule inj_onI) apply clarsimp (* TODO: This proof is not particular clean *) apply (drule (1) succs_tab_rl[rotated]) apply (drule (1) succs_tab_rl[rotated]) by (metis \_\_inverse local.\_def) lemma succ_impl_invarc: "\invar c; (a,c')\set (succ_impl c)\ \ invar c'" unfolding succ_impl_def mapped_succs_tab_def invar_def apply (clarsimp) by (metis \_def \_invar succs_tab_rl) lemma astep_impl_invarc: "\invar c; astep_impl c a c'\ \ invar c'" unfolding astep_impl_def by (auto intro: succ_impl_invarc) end lemma path'_path_conv: "li'.cs.path c p c' \ cfg_lts.path c (map Inl p) c'" by (induction p arbitrary: c) (auto simp: cfg'_def) lemma approx_reachable_approx': "li'.cs.reachable c c' \ c' \ approx_reachable c" using approx_reachable_approx[of c c'] by (auto simp: LTS.reachable_def path'_path_conv) (* TODO: Move *) lemma approx_reachable_closed': assumes "c\approx_reachable c0" assumes "cfg' c a c'" shows "c'\approx_reachable c0" using assms unfolding cfg'_def using approx_reachable_closed by blast type_synonym cmdc = nat type_synonym local_config = "(cmdc, local_state) Gen_Scheduler.local_config" type_synonym global_config = "(cmdc, local_state, global_state) Gen_Scheduler.global_config" locale cprog = fixes prog :: program begin sublocale comp: asystem "\x. x\cfg_V0 prog" cfg' . sublocale comp: sl_graph "\x. x\cfg_V0 prog" cfg' "approx_reachable_list prog" "cfg'_succ_list" apply unfold_locales apply (simp_all add: approx_reachable_list_invar approx_reachable_list_refine add: cfg'_succ_list_invar cfg'_succ_list_refine add: approx_reachable_prog_def ) using approx_reachable_approx' apply (auto simp: comp.step_path_conv LTS.reachable_def) [] using approx_reachable_closed' unfolding comp.step_def by blast abbreviation "cfgc \ comp.astep_impl" definition init_pc :: "proc_decl \ local_config" where "init_pc pd \ \ local_config.command = comp.\ (proc_decl.body pd), local_config.state = \ local_state.variables = init_valuation (proc_decl.local_vars pd) \ \" definition init_gc :: "global_config" where "init_gc \ \ global_config.processes = mset (map init_pc (program.processes prog)), global_config.state = \ global_state.variables = init_valuation (program.global_vars prog) \ \" sublocale lih': Gen_Scheduler'_linit cfgc la_en' la_ex' "{init_gc}" global_config.state . sublocale hc_bisim: Gen_Cfg_Bisim_Pre comp.rel cfgc cfg' la_en' la_ex' apply unfold_locales apply (blast dest: comp.sim1) apply (blast dest: comp.sim2) done lemma approx_reachable_prog0: "p \ set (program.processes prog) \ body p \ approx_reachable_prog prog" unfolding approx_reachable_prog_def by (auto simp: cfg_V0_def intro!: bexI) sublocale hc_bisim: Gen_Cfg_LBisim comp.rel cfgc "{init_gc}" global_config.state cfg' "{SM_Semantics.init_gc prog}" global_config.state la_en' la_ex' apply unfold_locales apply (auto simp: init_gc_def SM_Semantics.init_gc_def simp: init_pc_def SM_Semantics.init_pc_def simp: hc_bisim.rel_gc_def hc_bisim.rel_lc_def simp del: mset_map simp: mset_map[symmetric] intro!: rel_mset_of_mapI) apply (subst comp.rel_unfold1) apply (unfold approx_reachable_list_refine) apply (simp add: approx_reachable_prog0) apply simp apply (subst comp.rel_unfold1) apply (unfold approx_reachable_list_refine) apply (simp add: approx_reachable_prog0) apply simp done lemma lih'_accept_eq_li': "lih'.sa.accept = li'.sa.accept prog" by (rule hc_bisim.accept_bisim) definition "lc_\ lc \ \ local_config.command = comp.\ (local_config.command lc), local_config.state = local_config.state lc \" definition "lc_invar lc \ comp.invar (local_config.command lc)" lemma hc_bisim_rel_lc_brp_conv: "hc_bisim.rel_lc = brp lc_\ lc_invar" apply (intro ext) unfolding hc_bisim.rel_lc_def unfolding comp.rel_def lc_\_def lc_invar_def build_relp_def by auto definition "gc_\ gc \ \ global_config.processes = image_mset lc_\ (global_config.processes gc), global_config.state = global_config.state gc \" definition "gc_invar gc \ \lc. lc \# global_config.processes gc \ lc_invar lc" lemma hc_bisim_rel_gc_brp_conv: "hc_bisim.rel_gc = brp gc_\ gc_invar" apply (intro ext) unfolding hc_bisim.rel_gc_def hc_bisim_rel_lc_brp_conv unfolding comp.rel_def gc_\_def gc_invar_def apply (auto simp: rel_mset_brp) apply (auto simp: build_relp_def) done lemmas bi_unique_rel_gc = hc_bisim.bi_unique_rel_gcI[OF comp.bi_unique_rel] lemmas bi_unique_rel_lc = hc_bisim.bi_unique_rel_lcI[OF comp.bi_unique_rel] lemma run_sim_lih'_li: "lih'.sa.is_run r \ li'.sa.is_run prog (gc_\ o r)" apply (rule hc_bisim.s1.br_run_sim) unfolding hc_bisim_rel_gc_brp_conv by auto lemmas lih'_reachable_finite_sim = hc_bisim.s1.reachable_finite_sim end context well_typed_prog begin sublocale cprog . lemma lih'_accept_eq: "lih'.sa.accept = ref_accept prog" by (simp add: lih'_accept_eq_li' li'.accept_eq) lemma lih'_is_run_sim: "lih'.sa.is_run r \ ref_is_run prog (Some o gc_\ o r)" apply (drule run_sim_lih'_li) unfolding li'.is_run_conv by (auto simp: o_def) lemma lih'_finite_reachable: "finite ((g_E lih'.system_automaton)\<^sup>* `` g_V0 lih'.system_automaton)" apply (rule lih'_reachable_finite_sim) apply (rule li'_finite_reachable) unfolding rel_of_pred_def apply simp apply (rule left_unique_finite_sngI) using bi_unique_rel_gc apply (simp add: bi_unique_alt_def) done end (* Executable *) type_synonym comp_info = "(cmd list \ (action \ nat) list iarray)" (* TODO: Using a list is slow. Could use a hashmap to map commands to their indices *) definition comp_info_of :: "program \ comp_info" where "comp_info_of prog \ let rlist = approx_reachable_list prog; succs_tab = map cfg'_succ_list rlist; mapped_succs_tab = IArray ( map (map (\(a,c'). (a,the (cr_index_of rlist c')))) succs_tab) in (rlist,mapped_succs_tab)" definition ccfg_succ_impl :: "comp_info \ nat \ (action \ nat) list" where "ccfg_succ_impl cinf c \ let mst=snd cinf in if c_impl :: "comp_info \ cmd \ cmdc" where "comp_\_impl cinf c \ the (cr_index_of (fst cinf) c)" lemma (in cprog) comp_\_impl: "comp_\_impl (comp_info_of prog) = comp.\" apply (intro ext) unfolding comp.\_def unfolding comp_\_impl_def comp_info_of_def by simp export_code comp_info_of ccfg_succ_impl comp_\_impl checking SML end diff --git a/thys/IP_Addresses/NumberWang_IPv4.thy b/thys/IP_Addresses/NumberWang_IPv4.thy --- a/thys/IP_Addresses/NumberWang_IPv4.thy +++ b/thys/IP_Addresses/NumberWang_IPv4.thy @@ -1,60 +1,65 @@ theory NumberWang_IPv4 imports Main "HOL-Word.Word" begin +lemma zdiv_mult_self: + \m \ 0 \ (a + m * n) div m = a div m + n\ + for a m n :: int + by simp + section\Helper Lemmas for Low-Level Operations on Machine Words\ text\Needed for IPv4 Syntax\ lemma mod256: "((d::nat) + 256 * c + 65536 * b + 16777216 * a) mod 256 = d mod 256" proof - from mod_mult_self2[where a="d + 256 * c + 65536 * b" and b="256" and c="65536 * a"] have "(d + 256 * c + 65536 * b + 256 * 65536 * a) mod 256 = (d + 256 * c + 65536 * b) mod 256" by simp also have "\ = (d + 256 * c) mod 256" using mod_mult_self2[where a="d + 256 * c" and b="256" and c="256 * b"] by simp also have "\ = d mod 256" using mod_mult_self2 by blast finally show ?thesis by simp qed lemma div65536: assumes "a < 256" "b < 256" "c < 256" "d < 256" shows "nat ((int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 65536 mod 256) = b" proof - from zdiv_mult_self[where a="int d + int (256 * c) + int (65536 * b)" and m=65536 and n="256 * (int a)"] have "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 65536 = ((int d + int (256 * c) + int (65536 * b)) div 65536) + 256 * int a" by linarith also from zdiv_mult_self[where a="int d + int (256 * c)" and m="65536" and n="int b"] have "\ = (int d + int (256 * c)) div 65536 + int b + 256 * int a" by linarith also from assms have "\ = int b + 256 * int a" by simp finally have helper: "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 65536 = int b + 256 * int a" . with assms show ?thesis by simp qed lemma div256: assumes "a < 256" "b < 256" "c < 256" "d < 256" shows "nat ((int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 256 mod 256) = c" proof - from zdiv_mult_self[where a="int d + int (256 * c) + int (65536 * b)" and m=256 and n="65536 * (int a)"] have "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 256 = ((int d + int (256 * c) + int (65536 * b)) div 256) + 65536 * int a" by linarith also from zdiv_mult_self[where a="int d + int (256 * c)" and m="256" and n="256 * int b"] have "\ = (int d + int (256 * c)) div 256 + 256 * int b + 65536 * int a" by linarith also from zdiv_mult_self[where a="int d" and m="256" and n="int c"] have "\ = (int d) div 256 + int c + 256 * int b + 65536 * int a" by linarith also from assms have "\ = int c + 256 * int b + 65536 * int a" by simp finally have helper1: "(int d + int (256 * c) + int (65536 * b) + int (16777216 * a)) div 256 = int c + 256 * int b + 65536 * int a" . from mod_mult_self2[where a="int c + 256 * int b" and c="256 * int a" and b=256] have "(int c + 256 * int b + 65536 * int a) mod 256 = (int c + 256 * int b) mod 256" by simp also have "\ = int c mod 256" using mod_mult_self2[where a="int c" and b=256 and c="int b"] by simp also have "\ = int c" using assms apply(subst mod_pos_pos_trivial) by(simp_all) finally have helper2: "(int c + 256 * int b + 65536 * int a) mod 256 = int c" . from helper1 helper2 show ?thesis by simp qed end diff --git a/thys/JinjaThreads/MM/JMM_Spec.thy b/thys/JinjaThreads/MM/JMM_Spec.thy --- a/thys/JinjaThreads/MM/JMM_Spec.thy +++ b/thys/JinjaThreads/MM/JMM_Spec.thy @@ -1,2069 +1,2069 @@ (* Title: JinjaThreads/MM/JMM_Spec.thy Author: Andreas Lochbihler *) section \Axiomatic specification of the JMM\ theory JMM_Spec imports Orders "../Common/Observable_Events" Coinductive.Coinductive_List begin subsection \Definitions\ type_synonym JMM_action = nat type_synonym ('addr, 'thread_id) execution = "('thread_id \ ('addr, 'thread_id) obs_event action) llist" definition "actions" :: "('addr, 'thread_id) execution \ JMM_action set" where "actions E = {n. enat n < llength E}" definition action_tid :: "('addr, 'thread_id) execution \ JMM_action \ 'thread_id" where "action_tid E a = fst (lnth E a)" definition action_obs :: "('addr, 'thread_id) execution \ JMM_action \ ('addr, 'thread_id) obs_event action" where "action_obs E a = snd (lnth E a)" definition tactions :: "('addr, 'thread_id) execution \ 'thread_id \ JMM_action set" where "tactions E t = {a. a \ actions E \ action_tid E a = t}" inductive is_new_action :: "('addr, 'thread_id) obs_event action \ bool" where NewHeapElem: "is_new_action (NormalAction (NewHeapElem a hT))" inductive is_write_action :: "('addr, 'thread_id) obs_event action \ bool" where NewHeapElem: "is_write_action (NormalAction (NewHeapElem ad hT))" | WriteMem: "is_write_action (NormalAction (WriteMem ad al v))" text \ Initialisation actions are synchronisation actions iff they initialize volatile fields -- cf. JMM mailing list, message no. 62 (5 Nov. 2006). However, intuitively correct programs might not be correctly synchronized: \begin{verbatim} x = 0 --------------- r1 = x | r2 = x \end{verbatim} Here, if x is not volatile, the initial write can belong to at most one thread. Hence, it is happens-before to either r1 = x or r2 = x, but not both. In any sequentially consistent execution, both reads must read from the initilisation action x = 0, but it is not happens-before ordered to one of them. Moreover, if only volatile initialisations synchronize-with all thread-start actions, this breaks the proof of the DRF guarantee since it assumes that the happens-before relation $<=hb$ a for an action $a$ in a topologically sorted action sequence depends only on the actions before it. Counter example: (y is volatile) [(t1, start), (t1, init x), (t2, start), (t1, init y), ... Here, (t1, init x) $<=hb$ (t2, start) via: (t1, init x) $<=po$ (t1, init y) $<=sw$ (t2, start), but in [(t1, start), (t1, init x), (t2, start)], not (t1, init x) $<=hb$ (t2, start). Sevcik speculated that one might add an initialisation thread which performs all initialisation actions. All normal threads' start action would then synchronize on the final action of the initialisation thread. However, this contradicts the memory chain condition in the final field extension to the JMM (threads must read addresses of objects that they have not created themselves before they can access the fields of the object at that address) -- not modelled here. Instead, we leave every initialisation action in the thread it belongs to, but order it explicitly before the thread's start action and add synchronizes-with edges from \emph{all} initialisation actions to \emph{all} thread start actions. \ inductive saction :: "'m prog \ ('addr, 'thread_id) obs_event action \ bool" for P :: "'m prog" where NewHeapElem: "saction P (NormalAction (NewHeapElem a hT))" | Read: "is_volatile P al \ saction P (NormalAction (ReadMem a al v))" | Write: "is_volatile P al \ saction P (NormalAction (WriteMem a al v))" | ThreadStart: "saction P (NormalAction (ThreadStart t))" | ThreadJoin: "saction P (NormalAction (ThreadJoin t))" | SyncLock: "saction P (NormalAction (SyncLock a))" | SyncUnlock: "saction P (NormalAction (SyncUnlock a))" | ObsInterrupt: "saction P (NormalAction (ObsInterrupt t))" | ObsInterrupted: "saction P (NormalAction (ObsInterrupted t))" | InitialThreadAction: "saction P InitialThreadAction" | ThreadFinishAction: "saction P ThreadFinishAction" definition sactions :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action set" where "sactions P E = {a. a \ actions E \ saction P (action_obs E a)}" inductive_set write_actions :: "('addr, 'thread_id) execution \ JMM_action set" for E :: "('addr, 'thread_id) execution" where write_actionsI: "\ a \ actions E; is_write_action (action_obs E a) \ \ a \ write_actions E" text \ @{term NewObj} and @{term NewArr} actions only initialize those fields and array cells that are in fact in the object or array. Hence, they are not a write for - reads from addresses for which no object/array is created during the whole execution - reads from fields/cells that are not part of the object/array at the specified address. \ primrec addr_locs :: "'m prog \ htype \ addr_loc set" where "addr_locs P (Class_type C) = {CField D F|D F. \fm T. P \ C has F:T (fm) in D}" | "addr_locs P (Array_type T n) = ({ACell n'|n'. n' < n} \ {CField Object F|F. \fm T. P \ Object has F:T (fm) in Object})" text \ \action_loc_aux\ would naturally be an inductive set, but inductive\_set does not allow to pattern match on parameters. Hence, specify it using function and derive the setup manually. \ fun action_loc_aux :: "'m prog \ ('addr, 'thread_id) obs_event action \ ('addr \ addr_loc) set" where "action_loc_aux P (NormalAction (NewHeapElem ad (Class_type C))) = {(ad, CField D F)|D F T fm. P \ C has F:T (fm) in D}" | "action_loc_aux P (NormalAction (NewHeapElem ad (Array_type T n'))) = {(ad, ACell n)|n. n < n'} \ {(ad, CField D F)|D F T fm. P \ Object has F:T (fm) in D}" | "action_loc_aux P (NormalAction (WriteMem ad al v)) = {(ad, al)}" | "action_loc_aux P (NormalAction (ReadMem ad al v)) = {(ad, al)}" | "action_loc_aux _ _ = {}" lemma action_loc_aux_intros [intro?]: "P \ class_type_of hT has F:T (fm) in D \ (ad, CField D F) \ action_loc_aux P (NormalAction (NewHeapElem ad hT))" "n < n' \ (ad, ACell n) \ action_loc_aux P (NormalAction (NewHeapElem ad (Array_type T n')))" "(ad, al) \ action_loc_aux P (NormalAction (WriteMem ad al v))" "(ad, al) \ action_loc_aux P (NormalAction (ReadMem ad al v))" by(cases hT) auto lemma action_loc_aux_cases [elim?, cases set: action_loc_aux]: assumes "adal \ action_loc_aux P obs" obtains (NewHeapElem) hT F T fm D ad where "obs = NormalAction (NewHeapElem ad hT)" "adal = (ad, CField D F)" "P \ class_type_of hT has F:T (fm) in D" | (NewArr) n n' ad T where "obs = NormalAction (NewHeapElem ad (Array_type T n'))" "adal = (ad, ACell n)" "n < n'" | (WriteMem) ad al v where "obs = NormalAction (WriteMem ad al v)" "adal = (ad, al)" | (ReadMem) ad al v where "obs = NormalAction (ReadMem ad al v)" "adal = (ad, al)" using assms by(cases "(P, obs)" rule: action_loc_aux.cases) fastforce+ lemma action_loc_aux_simps [simp]: "(ad', al') \ action_loc_aux P (NormalAction (NewHeapElem ad hT)) \ (\D F T fm. ad = ad' \ al' = CField D F \ P \ class_type_of hT has F:T (fm) in D) \ (\n T n'. ad = ad' \ al' = ACell n \ hT = Array_type T n' \ n < n')" "(ad', al') \ action_loc_aux P (NormalAction (WriteMem ad al v)) \ ad = ad' \ al = al'" "(ad', al') \ action_loc_aux P (NormalAction (ReadMem ad al v)) \ ad = ad' \ al = al'" "(ad', al') \ action_loc_aux P InitialThreadAction" "(ad', al') \ action_loc_aux P ThreadFinishAction" "(ad', al') \ action_loc_aux P (NormalAction (ExternalCall a m vs v))" "(ad', al') \ action_loc_aux P (NormalAction (ThreadStart t))" "(ad', al') \ action_loc_aux P (NormalAction (ThreadJoin t))" "(ad', al') \ action_loc_aux P (NormalAction (SyncLock a))" "(ad', al') \ action_loc_aux P (NormalAction (SyncUnlock a))" "(ad', al') \ action_loc_aux P (NormalAction (ObsInterrupt t))" "(ad', al') \ action_loc_aux P (NormalAction (ObsInterrupted t))" by(cases hT) auto declare action_loc_aux.simps [simp del] abbreviation action_loc :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ ('addr \ addr_loc) set" where "action_loc P E a \ action_loc_aux P (action_obs E a)" inductive_set read_actions :: "('addr, 'thread_id) execution \ JMM_action set" for E :: "('addr, 'thread_id) execution" where ReadMem: "\ a \ actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ a \ read_actions E" fun addr_loc_default :: "'m prog \ htype \ addr_loc \ 'addr val" where "addr_loc_default P (Class_type C) (CField D F) = default_val (fst (the (map_of (fields P C) (F, D))))" | "addr_loc_default P (Array_type T n) (ACell n') = default_val T" | addr_loc_default_Array_CField: "addr_loc_default P (Array_type T n) (CField D F) = default_val (fst (the (map_of (fields P Object) (F, Object))))" | "addr_loc_default P _ _ = undefined" definition new_actions_for :: "'m prog \ ('addr, 'thread_id) execution \ ('addr \ addr_loc) \ JMM_action set" where "new_actions_for P E adal = {a. a \ actions E \ adal \ action_loc P E a \ is_new_action (action_obs E a)}" inductive_set external_actions :: "('addr, 'thread_id) execution \ JMM_action set" for E :: "('addr, 'thread_id) execution" where "\ a \ actions E; action_obs E a = NormalAction (ExternalCall ad M vs v) \ \ a \ external_actions E" fun value_written_aux :: "'m prog \ ('addr, 'thread_id) obs_event action \ addr_loc \ 'addr val" where "value_written_aux P (NormalAction (NewHeapElem ad' hT)) al = addr_loc_default P hT al" | value_written_aux_WriteMem': "value_written_aux P (NormalAction (WriteMem ad al' v)) al = (if al = al' then v else undefined)" | value_written_aux_undefined: "value_written_aux P _ al = undefined" primrec value_written :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ ('addr \ addr_loc) \ 'addr val" where "value_written P E a (ad, al) = value_written_aux P (action_obs E a) al" definition value_read :: "('addr, 'thread_id) execution \ JMM_action \ 'addr val" where "value_read E a = (case action_obs E a of NormalAction obs \ (case obs of ReadMem ad al v \ v | _ \ undefined) | _ \ undefined)" definition action_order :: "('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_ \ _ \a _" [51,0,50] 50) where "E \ a \a a' \ a \ actions E \ a' \ actions E \ (if is_new_action (action_obs E a) then is_new_action (action_obs E a') \ a \ a' else \ is_new_action (action_obs E a') \ a \ a')" definition program_order :: "('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_ \ _ \po _" [51,0,50] 50) where "E \ a \po a' \ E \ a \a a' \ action_tid E a = action_tid E a'" inductive synchronizes_with :: "'m prog \ ('thread_id \ ('addr, 'thread_id) obs_event action) \ ('thread_id \ ('addr, 'thread_id) obs_event action) \ bool" ("_ \ _ \sw _" [51, 51, 51] 50) for P :: "'m prog" where ThreadStart: "P \ (t, NormalAction (ThreadStart t')) \sw (t', InitialThreadAction)" | ThreadFinish: "P \ (t, ThreadFinishAction) \sw (t', NormalAction (ThreadJoin t))" | UnlockLock: "P \ (t, NormalAction (SyncUnlock a)) \sw (t', NormalAction (SyncLock a))" | \ \Only volatile writes synchronize with volatile reads. We could check volatility of @{term "al"} here, but this is checked by @{term "sactions"} in @{text sync_with} anyway.\ Volatile: "P \ (t, NormalAction (WriteMem a al v)) \sw (t', NormalAction (ReadMem a al v'))" | VolatileNew: " al \ addr_locs P hT \ P \ (t, NormalAction (NewHeapElem a hT)) \sw (t', NormalAction (ReadMem a al v))" | NewHeapElem: "P \ (t, NormalAction (NewHeapElem a hT)) \sw (t', InitialThreadAction)" | Interrupt: "P \ (t, NormalAction (ObsInterrupt t')) \sw (t'', NormalAction (ObsInterrupted t'))" definition sync_order :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \so _" [51,0,0,50] 50) where "P,E \ a \so a' \ a \ sactions P E \ a' \ sactions P E \ E \ a \a a'" definition sync_with :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \sw _" [51, 0, 0, 50] 50) where "P,E \ a \sw a' \ P,E \ a \so a' \ P \ (action_tid E a, action_obs E a) \sw (action_tid E a', action_obs E a')" definition po_sw :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" where "po_sw P E a a' \ E \ a \po a' \ P,E \ a \sw a'" abbreviation happens_before :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \hb _" [51, 0, 0, 50] 50) where "happens_before P E \ (po_sw P E)^++" type_synonym write_seen = "JMM_action \ JMM_action" definition is_write_seen :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ bool" where "is_write_seen P E ws \ (\a \ read_actions E. \ad al v. action_obs E a = NormalAction (ReadMem ad al v) \ ws a \ write_actions E \ (ad, al) \ action_loc P E (ws a) \ value_written P E (ws a) (ad, al) = v \ \ P,E \ a \hb ws a \ (is_volatile P al \ \ P,E \ a \so ws a) \ (\w' \ write_actions E. (ad, al) \ action_loc P E w' \ (P,E \ ws a \hb w' \ P,E \ w' \hb a \ is_volatile P al \ P,E \ ws a \so w' \ P,E \ w' \so a) \ w' = ws a))" definition thread_start_actions_ok :: "('addr, 'thread_id) execution \ bool" where "thread_start_actions_ok E \ (\a \ actions E. \ is_new_action (action_obs E a) \ (\i. i \ a \ action_obs E i = InitialThreadAction \ action_tid E i = action_tid E a))" primrec wf_exec :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ bool" ("_ \ _ \" [51, 50] 51) where "P \ (E, ws) \ \ is_write_seen P E ws \ thread_start_actions_ok E" inductive most_recent_write_for :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \ _ \mrw _" [50, 0, 51] 51) for P :: "'m prog" and E :: "('addr, 'thread_id) execution" and ra :: JMM_action and wa :: JMM_action where "\ ra \ read_actions E; adal \ action_loc P E ra; E \ wa \a ra; wa \ write_actions E; adal \ action_loc P E wa; \wa'. \ wa' \ write_actions E; adal \ action_loc P E wa' \ \ E \ wa' \a wa \ E \ ra \a wa' \ \ P,E \ ra \mrw wa" primrec sequentially_consistent :: "'m prog \ (('addr, 'thread_id) execution \ write_seen) \ bool" where "sequentially_consistent P (E, ws) \ (\r \ read_actions E. P,E \ r \mrw ws r)" subsection \Actions\ inductive_cases is_new_action_cases [elim!]: "is_new_action (NormalAction (ExternalCall a M vs v))" "is_new_action (NormalAction (ReadMem a al v))" "is_new_action (NormalAction (WriteMem a al v))" "is_new_action (NormalAction (NewHeapElem a hT))" "is_new_action (NormalAction (ThreadStart t))" "is_new_action (NormalAction (ThreadJoin t))" "is_new_action (NormalAction (SyncLock a))" "is_new_action (NormalAction (SyncUnlock a))" "is_new_action (NormalAction (ObsInterrupt t))" "is_new_action (NormalAction (ObsInterrupted t))" "is_new_action InitialThreadAction" "is_new_action ThreadFinishAction" inductive_simps is_new_action_simps [simp]: "is_new_action (NormalAction (NewHeapElem a hT))" "is_new_action (NormalAction (ExternalCall a M vs v))" "is_new_action (NormalAction (ReadMem a al v))" "is_new_action (NormalAction (WriteMem a al v))" "is_new_action (NormalAction (ThreadStart t))" "is_new_action (NormalAction (ThreadJoin t))" "is_new_action (NormalAction (SyncLock a))" "is_new_action (NormalAction (SyncUnlock a))" "is_new_action (NormalAction (ObsInterrupt t))" "is_new_action (NormalAction (ObsInterrupted t))" "is_new_action InitialThreadAction" "is_new_action ThreadFinishAction" lemmas is_new_action_iff = is_new_action.simps inductive_simps is_write_action_simps [simp]: "is_write_action InitialThreadAction" "is_write_action ThreadFinishAction" "is_write_action (NormalAction (ExternalCall a m vs v))" "is_write_action (NormalAction (ReadMem a al v))" "is_write_action (NormalAction (WriteMem a al v))" "is_write_action (NormalAction (NewHeapElem a hT))" "is_write_action (NormalAction (ThreadStart t))" "is_write_action (NormalAction (ThreadJoin t))" "is_write_action (NormalAction (SyncLock a))" "is_write_action (NormalAction (SyncUnlock a))" "is_write_action (NormalAction (ObsInterrupt t))" "is_write_action (NormalAction (ObsInterrupted t))" declare saction.intros [intro!] inductive_cases saction_cases [elim!]: "saction P (NormalAction (ExternalCall a M vs v))" "saction P (NormalAction (ReadMem a al v))" "saction P (NormalAction (WriteMem a al v))" "saction P (NormalAction (NewHeapElem a hT))" "saction P (NormalAction (ThreadStart t))" "saction P (NormalAction (ThreadJoin t))" "saction P (NormalAction (SyncLock a))" "saction P (NormalAction (SyncUnlock a))" "saction P (NormalAction (ObsInterrupt t))" "saction P (NormalAction (ObsInterrupted t))" "saction P InitialThreadAction" "saction P ThreadFinishAction" inductive_simps saction_simps [simp]: "saction P (NormalAction (ExternalCall a M vs v))" "saction P (NormalAction (ReadMem a al v))" "saction P (NormalAction (WriteMem a al v))" "saction P (NormalAction (NewHeapElem a hT))" "saction P (NormalAction (ThreadStart t))" "saction P (NormalAction (ThreadJoin t))" "saction P (NormalAction (SyncLock a))" "saction P (NormalAction (SyncUnlock a))" "saction P (NormalAction (ObsInterrupt t))" "saction P (NormalAction (ObsInterrupted t))" "saction P InitialThreadAction" "saction P ThreadFinishAction" lemma new_action_saction [simp, intro]: "is_new_action a \ saction P a" by(blast elim: is_new_action.cases) lemmas saction_iff = saction.simps lemma actionsD: "a \ actions E \ enat a < llength E" unfolding actions_def by blast lemma actionsE: assumes "a \ actions E" obtains "enat a < llength E" using assms unfolding actions_def by blast lemma actions_lappend: "llength xs = enat n \ actions (lappend xs ys) = actions xs \ ((+) n) ` actions ys" unfolding actions_def apply safe apply(erule contrapos_np) apply(rule_tac x="x - n" in image_eqI) apply simp_all apply(case_tac [!] "llength ys") apply simp_all done lemma tactionsE: assumes "a \ tactions E t" obtains obs where "a \ actions E" "action_tid E a = t" "action_obs E a = obs" using assms by(cases "lnth E a")(auto simp add: tactions_def action_tid_def action_obs_def) lemma sactionsI: "\ a \ actions E; saction P (action_obs E a) \ \ a \ sactions P E" unfolding sactions_def by blast lemma sactionsE: assumes "a \ sactions P E" obtains "a \ actions E" "saction P (action_obs E a)" using assms unfolding sactions_def by blast lemma sactions_actions [simp]: "a \ sactions P E \ a \ actions E" by(rule sactionsE) lemma value_written_aux_WriteMem [simp]: "value_written_aux P (NormalAction (WriteMem ad al v)) al = v" by simp declare value_written_aux_undefined [simp del] declare value_written_aux_WriteMem' [simp del] inductive_simps is_write_action_iff: "is_write_action a" inductive_simps write_actions_iff: "a \ write_actions E" lemma write_actions_actions [simp]: "a \ write_actions E \ a \ actions E" by(rule write_actions.induct) inductive_simps read_actions_iff: "a \ read_actions E" lemma read_actions_actions [simp]: "a \ read_actions E \ a \ actions E" by(rule read_actions.induct) lemma read_action_action_locE: assumes "r \ read_actions E" obtains ad al where "(ad, al) \ action_loc P E r" using assms by cases auto lemma read_actions_not_write_actions: "\ a \ read_actions E; a \ write_actions E \ \ False" by(auto elim!: read_actions.cases write_actions.cases) lemma read_actions_Int_write_actions [simp]: "read_actions E \ write_actions E = {}" "write_actions E \ read_actions E = {}" by(blast dest: read_actions_not_write_actions)+ lemma action_loc_addr_fun: "\ (ad, al) \ action_loc P E a; (ad', al') \ action_loc P E a \ \ ad = ad'" by(auto elim!: action_loc_aux_cases) lemma value_written_cong [cong]: "\ P = P'; a = a'; action_obs E a' = action_obs E' a' \ \ value_written P E a = value_written P' E' a'" by(rule ext)(auto split: action.splits) declare value_written.simps [simp del] lemma new_actionsI: "\ a \ actions E; adal \ action_loc P E a; is_new_action (action_obs E a) \ \ a \ new_actions_for P E adal" unfolding new_actions_for_def by blast lemma new_actionsE: assumes "a \ new_actions_for P E adal" obtains "a \ actions E" "adal \ action_loc P E a" "is_new_action (action_obs E a)" using assms unfolding new_actions_for_def by blast lemma action_loc_read_action_singleton: "\ r \ read_actions E; adal \ action_loc P E r; adal' \ action_loc P E r \ \ adal = adal'" by(cases adal, cases adal')(fastforce elim: read_actions.cases action_loc_aux_cases) lemma addr_locsI: "P \ class_type_of hT has F:T (fm) in D \ CField D F \ addr_locs P hT" "\ hT = Array_type T n; n' < n \ \ ACell n' \ addr_locs P hT" by(cases hT)(auto dest: has_field_decl_above) subsection \Orders\ subsection \Action order\ lemma action_orderI: assumes "a \ actions E" "a' \ actions E" and "\ is_new_action (action_obs E a); is_new_action (action_obs E a') \ \ a \ a'" and "\ is_new_action (action_obs E a) \ \ is_new_action (action_obs E a') \ a \ a'" shows "E \ a \a a'" using assms unfolding action_order_def by simp lemma action_orderE: assumes "E \ a \a a'" obtains "a \ actions E" "a' \ actions E" "is_new_action (action_obs E a)" "is_new_action (action_obs E a') \ a \ a'" | "a \ actions E" "a' \ actions E" "\ is_new_action (action_obs E a)" "\ is_new_action (action_obs E a')" "a \ a'" using assms unfolding action_order_def by(simp split: if_split_asm) lemma refl_action_order: "refl_onP (actions E) (action_order E)" by(rule refl_onPI)(auto elim: action_orderE intro: action_orderI) lemma antisym_action_order: "antisymp (action_order E)" by(rule antisympI)(auto elim!: action_orderE) lemma trans_action_order: "transp (action_order E)" by(rule transpI)(auto elim!: action_orderE intro: action_orderI) lemma porder_action_order: "porder_on (actions E) (action_order E)" by(blast intro: porder_onI refl_action_order antisym_action_order trans_action_order) lemma total_action_order: "total_onP (actions E) (action_order E)" by(rule total_onPI)(auto simp add: action_order_def) lemma torder_action_order: "torder_on (actions E) (action_order E)" by(blast intro: torder_onI total_action_order porder_action_order) lemma wf_action_order: "wfP (action_order E)\<^sup>\\<^sup>\" unfolding wfP_eq_minimal proof(intro strip) fix Q and x :: JMM_action assume "x \ Q" show "\z \ Q. \y. (action_order E)\<^sup>\\<^sup>\ y z \ y \ Q" proof(cases "\a \ Q. a \ actions E \ is_new_action (action_obs E a)") case True then obtain a where a: "a \ actions E \ is_new_action (action_obs E a) \ a \ Q" by blast define a' where "a' = (LEAST a'. a' \ actions E \ is_new_action (action_obs E a') \ a' \ Q)" from a have a': "a' \ actions E \ is_new_action (action_obs E a') \ a' \ Q" unfolding a'_def by(rule LeastI) { fix y assume y_le_a': "(action_order E)\<^sup>\\<^sup>\ y a'" have "y \ Q" proof assume "y \ Q" with y_le_a' a' have y: "y \ actions E \ is_new_action (action_obs E y) \ y \ Q" by(auto elim: action_orderE) hence "a' \ y" unfolding a'_def by(rule Least_le) with y_le_a' a' show False by(auto elim: action_orderE) qed } with a' show ?thesis by blast next case False hence not_new: "\a. \ a \ Q; a \ actions E \ \ \ is_new_action (action_obs E a)" by blast show ?thesis proof(cases "Q \ actions E = {}") case True with \x \ Q\ show ?thesis by(auto elim: action_orderE) next case False define a' where "a' = (LEAST a'. a' \ Q \ a' \ actions E \ \ is_new_action (action_obs E a'))" from False obtain a where "a \ Q" "a \ actions E" by blast with not_new[OF this] have "a \ Q \ a \ actions E \ \ is_new_action (action_obs E a)" by blast hence a': "a' \ Q \ a' \ actions E \ \ is_new_action (action_obs E a')" unfolding a'_def by(rule LeastI) { fix y assume y_le_a': "(action_order E)\<^sup>\\<^sup>\ y a'" hence "y \ actions E" by(auto elim: action_orderE) have "y \ Q" proof assume "y \ Q" hence y_not_new: "\ is_new_action (action_obs E y)" using \y \ actions E\ by(rule not_new) with \y \ Q\ \y \ actions E\ have "a' \ y" unfolding a'_def by -(rule Least_le, blast) with y_le_a' y_not_new show False by(auto elim: action_orderE) qed } with a' show ?thesis by blast qed qed qed lemma action_order_is_new_actionD: "\ E \ a \a a'; is_new_action (action_obs E a') \ \ is_new_action (action_obs E a)" by(auto elim: action_orderE) subsection \Program order\ lemma program_orderI: assumes "E \ a \a a'" and "action_tid E a = action_tid E a'" shows "E \ a \po a'" using assms unfolding program_order_def by auto lemma program_orderE: assumes "E \ a \po a'" obtains t obs obs' where "E \ a \a a'" and "action_tid E a = t" "action_obs E a = obs" and "action_tid E a' = t" "action_obs E a' = obs'" using assms unfolding program_order_def by(cases "lnth E a")(cases "lnth E a'", auto simp add: action_obs_def action_tid_def) lemma refl_on_program_order: "refl_onP (actions E) (program_order E)" by(rule refl_onPI)(auto elim: action_orderE program_orderE intro: program_orderI refl_onPD[OF refl_action_order]) lemma antisym_program_order: "antisymp (program_order E)" using antisympD[OF antisym_action_order] by(auto intro: antisympI elim!: program_orderE) lemma trans_program_order: "transp (program_order E)" by(rule transpI)(auto elim!: program_orderE intro: program_orderI dest: transPD[OF trans_action_order]) lemma porder_program_order: "porder_on (actions E) (program_order E)" by(blast intro: porder_onI refl_on_program_order antisym_program_order trans_program_order) lemma total_program_order_on_tactions: "total_onP (tactions E t) (program_order E)" by(rule total_onPI)(auto elim: tactionsE simp add: program_order_def dest: total_onD[OF total_action_order]) subsection \Synchronization order\ lemma sync_orderI: "\ E \ a \a a'; a \ sactions P E; a' \ sactions P E \ \ P,E \ a \so a'" unfolding sync_order_def by blast lemma sync_orderE: assumes "P,E \ a \so a'" obtains "a \ sactions P E" "a' \ sactions P E" "E \ a \a a'" using assms unfolding sync_order_def by blast lemma refl_on_sync_order: "refl_onP (sactions P E) (sync_order P E)" by(rule refl_onPI)(fastforce elim: sync_orderE intro: sync_orderI refl_onPD[OF refl_action_order])+ lemma antisym_sync_order: "antisymp (sync_order P E)" using antisympD[OF antisym_action_order] by(rule antisympI)(auto elim!: sync_orderE) lemma trans_sync_order: "transp (sync_order P E)" by(rule transpI)(auto elim!: sync_orderE intro: sync_orderI dest: transPD[OF trans_action_order]) lemma porder_sync_order: "porder_on (sactions P E) (sync_order P E)" by(blast intro: porder_onI refl_on_sync_order antisym_sync_order trans_sync_order) lemma total_sync_order: "total_onP (sactions P E) (sync_order P E)" apply(rule total_onPI) apply(simp add: sync_order_def) apply(rule total_onPD[OF total_action_order]) apply simp_all done lemma torder_sync_order: "torder_on (sactions P E) (sync_order P E)" by(blast intro: torder_onI porder_sync_order total_sync_order) subsection \Synchronizes with\ lemma sync_withI: "\ P,E \ a \so a'; P \ (action_tid E a, action_obs E a) \sw (action_tid E a', action_obs E a') \ \ P,E \ a \sw a'" unfolding sync_with_def by blast lemma sync_withE: assumes "P,E \ a \sw a'" obtains "P,E \ a \so a'" "P \ (action_tid E a, action_obs E a) \sw (action_tid E a', action_obs E a')" using assms unfolding sync_with_def by blast lemma irrefl_synchronizes_with: "irreflP (synchronizes_with P)" by(rule irreflPI)(auto elim: synchronizes_with.cases) lemma irrefl_sync_with: "irreflP (sync_with P E)" by(rule irreflPI)(auto elim: sync_withE intro: irreflPD[OF irrefl_synchronizes_with]) lemma anitsym_sync_with: "antisymp (sync_with P E)" using antisymPD[OF antisym_sync_order, of P E] by -(rule antisymPI, auto elim: sync_withE) lemma consistent_program_order_sync_order: "order_consistent (program_order E) (sync_order P E)" apply(rule order_consistent_subset) apply(rule antisym_order_consistent_self[OF antisym_action_order[of E]]) apply(blast elim: program_orderE sync_orderE)+ done lemma consistent_program_order_sync_with: "order_consistent (program_order E) (sync_with P E)" by(rule order_consistent_subset[OF consistent_program_order_sync_order])(blast elim: sync_withE)+ subsection \Happens before\ lemma porder_happens_before: "porder_on (actions E) (happens_before P E)" unfolding po_sw_def [abs_def] by(rule porder_on_sub_torder_on_tranclp_porder_onI[OF porder_program_order torder_sync_order consistent_program_order_sync_order])(auto elim: sync_withE) lemma porder_tranclp_po_so: "porder_on (actions E) (\a a'. program_order E a a' \ sync_order P E a a')^++" by(rule porder_on_torder_on_tranclp_porder_onI[OF porder_program_order torder_sync_order consistent_program_order_sync_order]) auto lemma happens_before_refl: assumes "a \ actions E" shows "P,E \ a \hb a" using porder_happens_before[of E P] by(rule porder_onE)(erule refl_onPD[OF _ assms]) lemma happens_before_into_po_so_tranclp: assumes "P,E \ a \hb a'" shows "(\a a'. E \ a \po a' \ P,E \ a \so a')^++ a a'" using assms unfolding po_sw_def [abs_def] by(induct)(blast elim: sync_withE intro: tranclp.trancl_into_trancl)+ lemma po_sw_into_action_order: "po_sw P E a a' \ E \ a \a a'" by(auto elim: program_orderE sync_withE sync_orderE simp add: po_sw_def) lemma happens_before_into_action_order: assumes "P,E \ a \hb a'" shows "E \ a \a a'" using assms by induct(blast intro: po_sw_into_action_order transPD[OF trans_action_order])+ lemma action_order_consistent_with_happens_before: "order_consistent (action_order E) (happens_before P E)" by(blast intro: order_consistent_subset antisym_order_consistent_self antisym_action_order happens_before_into_action_order) lemma happens_before_new_actionD: assumes hb: "P,E \ a \hb a'" and new: "is_new_action (action_obs E a')" shows "is_new_action (action_obs E a)" "action_tid E a = action_tid E a'" "a \ a'" using hb proof(induct rule: converse_tranclp_induct) case (base a) case 1 from new base show ?case by(auto dest: po_sw_into_action_order elim: action_orderE) case 2 from new base show ?case by(auto simp add: po_sw_def elim!: sync_withE elim: program_orderE synchronizes_with.cases) case 3 from new base show ?case by(auto dest: po_sw_into_action_order elim: action_orderE) next case (step a a'') note po_sw = \po_sw P E a a''\ and new = \is_new_action (action_obs E a'')\ and tid = \action_tid E a'' = action_tid E a'\ case 1 from new po_sw show ?case by(auto dest: po_sw_into_action_order elim: action_orderE) case 2 from new po_sw tid show ?case by(auto simp add: po_sw_def elim!: sync_withE elim: program_orderE synchronizes_with.cases) case 3 from new po_sw \a'' \ a'\ show ?case by(auto dest!: po_sw_into_action_order elim!: action_orderE) qed lemma external_actions_not_new: "\ a \ external_actions E; is_new_action (action_obs E a) \ \ False" by(erule external_actions.cases)(simp) subsection \Most recent writes and sequential consistency\ lemma most_recent_write_for_fun: "\ P,E \ ra \mrw wa; P,E \ ra \mrw wa' \ \ wa = wa'" apply(erule most_recent_write_for.cases)+ apply clarsimp apply(erule meta_allE)+ apply(erule meta_impE) apply(rotate_tac 3) apply assumption apply(erule (1) meta_impE) apply(frule (1) action_loc_read_action_singleton) apply(rotate_tac 1) apply assumption apply(fastforce dest: antisymPD[OF antisym_action_order] elim: write_actions.cases read_actions.cases) done lemma THE_most_recent_writeI: "P,E \ r \mrw w \ (THE w. P,E \ r \mrw w) = w" by(blast dest: most_recent_write_for_fun)+ lemma most_recent_write_for_write_actionsD: assumes "P,E \ ra \mrw wa" shows "wa \ write_actions E" using assms by cases lemma most_recent_write_recent: "\ P,E \ r \mrw w; adal \ action_loc P E r; w' \ write_actions E; adal \ action_loc P E w' \ \ E \ w' \a w \ E \ r \a w'" apply(erule most_recent_write_for.cases) apply(drule (1) action_loc_read_action_singleton) apply(rotate_tac 1) apply assumption apply clarsimp done lemma is_write_seenI: "\ \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ ws a \ write_actions E; \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ (ad, al) \ action_loc P E (ws a); \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ value_written P E (ws a) (ad, al) = v; \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ \ P,E \ a \hb ws a; \a ad al v. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v); is_volatile P al \ \ \ P,E \ a \so ws a; \a ad al v a'. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v); a' \ write_actions E; (ad, al) \ action_loc P E a'; P,E \ ws a \hb a'; P,E \ a' \hb a \ \ a' = ws a; \a ad al v a'. \ a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v); a' \ write_actions E; (ad, al) \ action_loc P E a'; is_volatile P al; P,E \ ws a \so a'; P,E \ a' \so a \ \ a' = ws a \ \ is_write_seen P E ws" unfolding is_write_seen_def by(blast 30) lemma is_write_seenD: "\ is_write_seen P E ws; a \ read_actions E; action_obs E a = NormalAction (ReadMem ad al v) \ \ ws a \ write_actions E \ (ad, al) \ action_loc P E (ws a) \ value_written P E (ws a) (ad, al) = v \ \ P,E \ a \hb ws a \ (is_volatile P al \ \ P,E \ a \so ws a) \ (\a' \ write_actions E. (ad, al) \ action_loc P E a' \ (P,E \ ws a \hb a' \ P,E \ a' \hb a \ is_volatile P al \ P,E \ ws a \so a' \ P,E \ a' \so a) \ a' = ws a)" unfolding is_write_seen_def by blast lemma thread_start_actions_okI: "(\a. \ a \ actions E; \ is_new_action (action_obs E a) \ \ \i. i \ a \ action_obs E i = InitialThreadAction \ action_tid E i = action_tid E a) \ thread_start_actions_ok E" unfolding thread_start_actions_ok_def by blast lemma thread_start_actions_okD: "\ thread_start_actions_ok E; a \ actions E; \ is_new_action (action_obs E a) \ \ \i. i \ a \ action_obs E i = InitialThreadAction \ action_tid E i = action_tid E a" unfolding thread_start_actions_ok_def by blast lemma thread_start_actions_ok_prefix: "\ thread_start_actions_ok E'; lprefix E E' \ \ thread_start_actions_ok E" -apply(clarsimp simp add: lprefix_conv_lappend) -apply(rule thread_start_actions_okI) -apply(drule_tac a=a in thread_start_actions_okD) - apply(simp add: actions_def) - apply(metis Suc_ile_eq enat_le_plus_same(1) xtr6) -apply(auto simp add: action_obs_def lnth_lappend1 actions_def action_tid_def le_less_trans[where y="enat a" for a]) -done + apply(clarsimp simp add: lprefix_conv_lappend) + apply(rule thread_start_actions_okI) + apply(drule_tac a=a in thread_start_actions_okD) + apply(simp add: actions_def) + apply(auto simp add: action_obs_def lnth_lappend1 actions_def action_tid_def le_less_trans[where y="enat a" for a]) + apply (metis add.right_neutral add_strict_mono not_gr_zero) + done lemma wf_execI [intro?]: "\ is_write_seen P E ws; thread_start_actions_ok E \ \ P \ (E, ws) \" by simp lemma wf_exec_is_write_seenD: "P \ (E, ws) \ \ is_write_seen P E ws" by simp lemma wf_exec_thread_start_actions_okD: "P \ (E, ws) \ \ thread_start_actions_ok E" by simp lemma sequentially_consistentI: "(\r. r \ read_actions E \ P,E \ r \mrw ws r) \ sequentially_consistent P (E, ws)" by simp lemma sequentially_consistentE: assumes "sequentially_consistent P (E, ws)" "a \ read_actions E" obtains "P,E \ a \mrw ws a" using assms by simp declare sequentially_consistent.simps [simp del] subsection \Similar actions\ text \Similar actions differ only in the values written/read.\ inductive sim_action :: "('addr, 'thread_id) obs_event action \ ('addr, 'thread_id) obs_event action \ bool" ("_ \ _" [50, 50] 51) where InitialThreadAction: "InitialThreadAction \ InitialThreadAction" | ThreadFinishAction: "ThreadFinishAction \ ThreadFinishAction" | NewHeapElem: "NormalAction (NewHeapElem a hT) \ NormalAction (NewHeapElem a hT)" | ReadMem: "NormalAction (ReadMem ad al v) \ NormalAction (ReadMem ad al v')" | WriteMem: "NormalAction (WriteMem ad al v) \ NormalAction (WriteMem ad al v')" | ThreadStart: "NormalAction (ThreadStart t) \ NormalAction (ThreadStart t)" | ThreadJoin: "NormalAction (ThreadJoin t) \ NormalAction (ThreadJoin t)" | SyncLock: "NormalAction (SyncLock a) \ NormalAction (SyncLock a)" | SyncUnlock: "NormalAction (SyncUnlock a) \ NormalAction (SyncUnlock a)" | ExternalCall: "NormalAction (ExternalCall a M vs v) \ NormalAction (ExternalCall a M vs v)" | ObsInterrupt: "NormalAction (ObsInterrupt t) \ NormalAction (ObsInterrupt t)" | ObsInterrupted: "NormalAction (ObsInterrupted t) \ NormalAction (ObsInterrupted t)" definition sim_actions :: "('addr, 'thread_id) execution \ ('addr, 'thread_id) execution \ bool" ("_ [\] _" [51, 50] 51) where "sim_actions = llist_all2 (\(t, a) (t', a'). t = t' \ a \ a')" lemma sim_action_refl [intro!, simp]: "obs \ obs" apply(cases obs) apply(rename_tac obs') apply(case_tac "obs'") apply(auto intro: sim_action.intros) done inductive_cases sim_action_cases [elim!]: "InitialThreadAction \ obs" "ThreadFinishAction \ obs" "NormalAction (NewHeapElem a hT) \ obs" "NormalAction (ReadMem ad al v) \ obs" "NormalAction (WriteMem ad al v) \ obs" "NormalAction (ThreadStart t) \ obs" "NormalAction (ThreadJoin t) \ obs" "NormalAction (SyncLock a) \ obs" "NormalAction (SyncUnlock a) \ obs" "NormalAction (ObsInterrupt t) \ obs" "NormalAction (ObsInterrupted t) \ obs" "NormalAction (ExternalCall a M vs v) \ obs" "obs \ InitialThreadAction" "obs \ ThreadFinishAction" "obs \ NormalAction (NewHeapElem a hT)" "obs \ NormalAction (ReadMem ad al v')" "obs \ NormalAction (WriteMem ad al v')" "obs \ NormalAction (ThreadStart t)" "obs \ NormalAction (ThreadJoin t)" "obs \ NormalAction (SyncLock a)" "obs \ NormalAction (SyncUnlock a)" "obs \ NormalAction (ObsInterrupt t)" "obs \ NormalAction (ObsInterrupted t)" "obs \ NormalAction (ExternalCall a M vs v)" inductive_simps sim_action_simps [simp]: "InitialThreadAction \ obs" "ThreadFinishAction \ obs" "NormalAction (NewHeapElem a hT) \ obs" "NormalAction (ReadMem ad al v) \ obs" "NormalAction (WriteMem ad al v) \ obs" "NormalAction (ThreadStart t) \ obs" "NormalAction (ThreadJoin t) \ obs" "NormalAction (SyncLock a) \ obs" "NormalAction (SyncUnlock a) \ obs" "NormalAction (ObsInterrupt t) \ obs" "NormalAction (ObsInterrupted t) \ obs" "NormalAction (ExternalCall a M vs v) \ obs" "obs \ InitialThreadAction" "obs \ ThreadFinishAction" "obs \ NormalAction (NewHeapElem a hT)" "obs \ NormalAction (ReadMem ad al v')" "obs \ NormalAction (WriteMem ad al v')" "obs \ NormalAction (ThreadStart t)" "obs \ NormalAction (ThreadJoin t)" "obs \ NormalAction (SyncLock a)" "obs \ NormalAction (SyncUnlock a)" "obs \ NormalAction (ObsInterrupt t)" "obs \ NormalAction (ObsInterrupted t)" "obs \ NormalAction (ExternalCall a M vs v)" lemma sim_action_trans [trans]: "\ obs \ obs'; obs' \ obs'' \ \ obs \ obs''" by(erule sim_action.cases) auto lemma sim_action_sym [sym]: assumes "obs \ obs'" shows "obs' \ obs" using assms by cases simp_all lemma sim_actions_sym [sym]: "E [\] E' \ E' [\] E" unfolding sim_actions_def by(auto simp add: llist_all2_conv_all_lnth split_beta intro: sim_action_sym) lemma sim_actions_action_obsD: "E [\] E' \ action_obs E a \ action_obs E' a" unfolding sim_actions_def action_obs_def by(cases "enat a < llength E")(auto dest: llist_all2_lnthD llist_all2_llengthD simp add: split_beta lnth_beyond split: enat.split) lemma sim_actions_action_tidD: "E [\] E' \ action_tid E a = action_tid E' a" unfolding sim_actions_def action_tid_def by(cases "enat a < llength E")(auto dest: llist_all2_lnthD llist_all2_llengthD simp add: lnth_beyond split: enat.split) lemma action_loc_aux_sim_action: "a \ a' \ action_loc_aux P a = action_loc_aux P a'" by(auto elim!: action_loc_aux_cases intro: action_loc_aux_intros) lemma eq_into_sim_actions: assumes "E = E'" shows "E [\] E'" unfolding sim_actions_def assms by(rule llist_all2_reflI)(auto) subsection \Well-formedness conditions for execution sets\ locale executions_base = fixes \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" locale drf = executions_base \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" + assumes \_new_actions_for_fun: "\ E \ \; a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" and \_sequential_completion: "\ E \ \; P \ (E, ws) \; \a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a \ \ \E' \ \. \ws'. P \ (E', ws') \ \ ltake (enat r) E = ltake (enat r) E' \ sequentially_consistent P (E', ws') \ action_tid E r = action_tid E' r \ action_obs E r \ action_obs E' r \ (r \ actions E \ r \ actions E')" locale executions_aux = executions_base \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" + assumes init_before_read: "\ E \ \; P \ (E, ws) \; r \ read_actions E; adal \ action_loc P E r; \a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a \ \ \i new_actions_for P E adal" and \_new_actions_for_fun: "\ E \ \; a \ new_actions_for P E adal; a' \ new_actions_for P E adal \ \ a = a'" locale sc_legal = executions_aux \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" + assumes \_hb_completion: "\ E \ \; P \ (E, ws) \; \a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a \ \ \E' \ \. \ws'. P \ (E', ws') \ \ ltake (enat r) E = ltake (enat r) E' \ (\a \ read_actions E'. if a < r then ws' a = ws a else P,E' \ ws' a \hb a) \ action_tid E' r = action_tid E r \ (if r \ read_actions E then sim_action else (=)) (action_obs E' r) (action_obs E r) \ (r \ actions E \ r \ actions E')" locale jmm_consistent = drf?: drf \ P + sc_legal \ P for \ :: "('addr, 'thread_id) execution set" and P :: "'m prog" subsection \Legal executions\ record ('addr, 'thread_id) pre_justifying_execution = committed :: "JMM_action set" justifying_exec :: "('addr, 'thread_id) execution" justifying_ws :: "write_seen" record ('addr, 'thread_id) justifying_execution = "('addr, 'thread_id) pre_justifying_execution" + action_translation :: "JMM_action \ JMM_action" type_synonym ('addr, 'thread_id) justification = "nat \ ('addr, 'thread_id) justifying_execution" definition wf_action_translation_on :: "('addr, 'thread_id) execution \ ('addr, 'thread_id) execution \ JMM_action set \ (JMM_action \ JMM_action) \ bool" where "wf_action_translation_on E E' A f \ inj_on f (actions E) \ (\a \ A. action_tid E a = action_tid E' (f a) \ action_obs E a \ action_obs E' (f a))" abbreviation wf_action_translation :: "('addr, 'thread_id) execution \ ('addr, 'thread_id) justifying_execution \ bool" where "wf_action_translation E J \ wf_action_translation_on (justifying_exec J) E (committed J) (action_translation J)" context fixes P :: "'m prog" and E :: "('addr, 'thread_id) execution" and ws :: "write_seen" and J :: "('addr, 'thread_id) justification" begin text \ This context defines the causality constraints for the JMM. The weak versions are for the fixed JMM as presented by Sevcik and Aspinall at ECOOP 2008. \ text \Committed actions are an ascending chain with all actions of E as a limit\ definition is_commit_sequence :: bool where "is_commit_sequence \ committed (J 0) = {} \ (\n. action_translation (J n) ` committed (J n) \ action_translation (J (Suc n)) ` committed (J (Suc n))) \ actions E = (\n. action_translation (J n) ` committed (J n))" definition justification_well_formed :: bool where "justification_well_formed \ (\n. P \ (justifying_exec (J n), justifying_ws (J n)) \)" definition committed_subset_actions :: bool where \ \JMM constraint 1\ "committed_subset_actions \ (\n. committed (J n) \ actions (justifying_exec (J n)))" definition happens_before_committed :: bool where \ \JMM constraint 2\ "happens_before_committed \ (\n. happens_before P (justifying_exec (J n)) |` committed (J n) = inv_imageP (happens_before P E) (action_translation (J n)) |` committed (J n))" definition happens_before_committed_weak :: bool where \ \relaxed JMM constraint\ "happens_before_committed_weak \ (\n. \r \ read_actions (justifying_exec (J n)) \ committed (J n). let r' = action_translation (J n) r; w' = ws r'; w = inv_into (actions (justifying_exec (J n))) (action_translation (J n)) w' in (P,E \ w' \hb r' \ P,justifying_exec (J n) \ w \hb r) \ \ P,justifying_exec (J n) \ r \hb w)" definition sync_order_committed :: bool where \ \JMM constraint 3\ "sync_order_committed \ (\n. sync_order P (justifying_exec (J n)) |` committed (J n) = inv_imageP (sync_order P E) (action_translation (J n)) |` committed (J n))" definition value_written_committed :: bool where \ \JMM constraint 4\ "value_written_committed \ (\n. \w \ write_actions (justifying_exec (J n)) \ committed (J n). let w' = action_translation (J n) w in (\adal \ action_loc P E w'. value_written P (justifying_exec (J n)) w adal = value_written P E w' adal))" definition write_seen_committed :: bool where \ \JMM constraint 5\ "write_seen_committed \ (\n. \r' \ read_actions (justifying_exec (J n)) \ committed (J n). let r = action_translation (J n) r'; r'' = inv_into (actions (justifying_exec (J (Suc n)))) (action_translation (J (Suc n))) r in action_translation (J (Suc n)) (justifying_ws (J (Suc n)) r'') = ws r)" text \uncommited reads see writes that happen before them -- JMM constraint 6\ (* this constraint does not apply to the 0th justifying execution, so reads may see arbitrary writes, but this cannot be exploited because reads cannot be committed in the 1st justifying execution since no writes are committed in the 0th execution *) definition uncommitted_reads_see_hb :: bool where "uncommitted_reads_see_hb \ (\n. \r' \ read_actions (justifying_exec (J (Suc n))). action_translation (J (Suc n)) r' \ action_translation (J n) ` committed (J n) \ P,justifying_exec (J (Suc n)) \ justifying_ws (J (Suc n)) r' \hb r')" text \ newly committed reads see already committed writes and write-seen relationship must not change any more -- JMM constraint 7 \ definition committed_reads_see_committed_writes :: bool where "committed_reads_see_committed_writes \ (\n. \r' \ read_actions (justifying_exec (J (Suc n))) \ committed (J (Suc n)). let r = action_translation (J (Suc n)) r'; committed_n = action_translation (J n) ` committed (J n) in r \ committed_n \ (action_translation (J (Suc n)) (justifying_ws (J (Suc n)) r') \ committed_n \ ws r \ committed_n))" definition committed_reads_see_committed_writes_weak :: bool where "committed_reads_see_committed_writes_weak \ (\n. \r' \ read_actions (justifying_exec (J (Suc n))) \ committed (J (Suc n)). let r = action_translation (J (Suc n)) r'; committed_n = action_translation (J n) ` committed (J n) in r \ committed_n \ ws r \ committed_n)" text \external actions must be committed as soon as hb-subsequent actions are committed -- JMM constraint 9\ definition external_actions_committed :: bool where "external_actions_committed \ (\n. \a \ external_actions (justifying_exec (J n)). \a' \ committed (J n). P,justifying_exec (J n) \ a \hb a' \ a \ committed (J n))" text \well-formedness conditions for action translations\ definition wf_action_translations :: bool where "wf_action_translations \ (\n. wf_action_translation_on (justifying_exec (J n)) E (committed (J n)) (action_translation (J n)))" end text \ Rule 8 of the justification for the JMM is incorrect because there might be no transitive reduction of the happens-before relation for an infinite execution, if infinitely many initialisation actions have to be ordered before the start action of every thread. Hence, \is_justified_by\ omits this constraint. \ primrec is_justified_by :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ ('addr, 'thread_id) justification \ bool" ("_ \ _ justified'_by _" [51, 50, 50] 50) where "P \ (E, ws) justified_by J \ is_commit_sequence E J \ justification_well_formed P J \ committed_subset_actions J \ happens_before_committed P E J \ sync_order_committed P E J \ value_written_committed P E J \ write_seen_committed ws J \ uncommitted_reads_see_hb P J \ committed_reads_see_committed_writes ws J \ external_actions_committed P J \ wf_action_translations E J" text \ Sevcik requires in the fixed JMM that external actions may only be committed when everything that happens before has already been committed. On the level of legality, this constraint is vacuous because it is always possible to delay committing external actions, so we omit it here. \ primrec is_weakly_justified_by :: "'m prog \ ('addr, 'thread_id) execution \ write_seen \ ('addr, 'thread_id) justification \ bool" ("_ \ _ weakly'_justified'_by _" [51, 50, 50] 50) where "P \ (E, ws) weakly_justified_by J \ is_commit_sequence E J \ justification_well_formed P J \ committed_subset_actions J \ happens_before_committed_weak P E ws J \ \ \no \sync_order\ constraint\ value_written_committed P E J \ write_seen_committed ws J \ uncommitted_reads_see_hb P J \ committed_reads_see_committed_writes_weak ws J \ wf_action_translations E J" text \ Notion of conflict is strengthened to explicitly exclude volatile locations. Otherwise, the following program is not correctly synchronised: \begin{verbatim} volatile x = 0; --------------- r = x; | x = 1; \end{verbatim} because in the SC execution [Init x 0, (t1, Read x 0), (t2, Write x 1)], the read and write are unrelated in hb, because synchronises-with is asymmetric for volatiles. The JLS considers conflicting volatiles for data races, but this is only a remark on the DRF guarantee. See JMM mailing list posts \#2477 to 2488. \ (* Problem already exists in Sevcik's formalisation *) definition non_volatile_conflict :: "'m prog \ ('addr, 'thread_id) execution \ JMM_action \ JMM_action \ bool" ("_,_ \/(_)\(_)" [51,50,50,50] 51) where "P,E \ a \ a' \ (a \ read_actions E \ a' \ write_actions E \ a \ write_actions E \ a' \ read_actions E \ a \ write_actions E \ a' \ write_actions E) \ (\ad al. (ad, al) \ action_loc P E a \ action_loc P E a' \ \ is_volatile P al)" definition correctly_synchronized :: "'m prog \ ('addr, 'thread_id) execution set \ bool" where "correctly_synchronized P \ \ (\E \ \. \ws. P \ (E, ws) \ \ sequentially_consistent P (E, ws) \ (\a \ actions E. \a' \ actions E. P,E \ a \ a' \ P,E \ a \hb a' \ P,E \ a' \hb a))" primrec gen_legal_execution :: "('m prog \ ('addr, 'thread_id) execution \ write_seen \ ('addr, 'thread_id) justification \ bool) \ 'm prog \ ('addr, 'thread_id) execution set \ ('addr, 'thread_id) execution \ write_seen \ bool" where "gen_legal_execution is_justification P \ (E, ws) \ E \ \ \ P \ (E, ws) \ \ (\J. is_justification P (E, ws) J \ range (justifying_exec \ J) \ \)" abbreviation legal_execution :: "'m prog \ ('addr, 'thread_id) execution set \ ('addr, 'thread_id) execution \ write_seen \ bool" where "legal_execution \ gen_legal_execution is_justified_by" abbreviation weakly_legal_execution :: "'m prog \ ('addr, 'thread_id) execution set \ ('addr, 'thread_id) execution \ write_seen \ bool" where "weakly_legal_execution \ gen_legal_execution is_weakly_justified_by" declare gen_legal_execution.simps [simp del] lemma sym_non_volatile_conflict: "symP (non_volatile_conflict P E)" unfolding non_volatile_conflict_def by(rule symPI) blast lemma legal_executionI: "\ E \ \; P \ (E, ws) \; is_justification P (E, ws) J; range (justifying_exec \ J) \ \ \ \ gen_legal_execution is_justification P \ (E, ws)" unfolding gen_legal_execution.simps by blast lemma legal_executionE: assumes "gen_legal_execution is_justification P \ (E, ws)" obtains J where "E \ \" "P \ (E, ws) \" "is_justification P (E, ws) J" "range (justifying_exec \ J) \ \" using assms unfolding gen_legal_execution.simps by blast lemma legal_\D: "gen_legal_execution is_justification P \ (E, ws) \ E \ \" by(erule legal_executionE) lemma legal_wf_execD: "gen_legal_execution is_justification P \ Ews \ P \ Ews \" by(cases Ews)(auto elim: legal_executionE) lemma correctly_synchronizedD: "\ correctly_synchronized P \; E \ \; P \ (E, ws) \; sequentially_consistent P (E, ws) \ \ \a a'. a \ actions E \ a' \ actions E \ P,E \ a \ a' \ P,E \ a \hb a' \ P,E \ a' \hb a" unfolding correctly_synchronized_def by blast lemma wf_action_translation_on_actionD: "\ wf_action_translation_on E E' A f; a \ A \ \ action_tid E a = action_tid E' (f a) \ action_obs E a \ action_obs E' (f a)" unfolding wf_action_translation_on_def by blast lemma wf_action_translation_on_inj_onD: "wf_action_translation_on E E' A f \ inj_on f (actions E)" unfolding wf_action_translation_on_def by simp lemma wf_action_translation_on_action_locD: "\ wf_action_translation_on E E' A f; a \ A \ \ action_loc P E a = action_loc P E' (f a)" apply(drule (1) wf_action_translation_on_actionD) apply(cases "(P, action_obs E a)" rule: action_loc_aux.cases) apply auto done lemma weakly_justified_write_seen_hb_read_committed: assumes J: "P \ (E, ws) weakly_justified_by J" and r: "r \ read_actions (justifying_exec (J n))" "r \ committed (J n)" shows "ws (action_translation (J n) r) \ action_translation (J n) ` committed (J n)" using r proof(induct n arbitrary: r) case 0 from J have [simp]: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) with 0 show ?case by simp next case (Suc n) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" note r = \r \ read_actions (?E (Suc n))\ hence "r \ actions (?E (Suc n))" by simp from J have wfan: "wf_action_translation_on (?E n) E (?C n) (?\ n)" and wfaSn: "wf_action_translation_on (?E (Suc n)) E (?C (Suc n)) (?\ (Suc n))" by(simp_all add: wf_action_translations_def) from wfaSn have injSn: "inj_on (?\ (Suc n)) (actions (?E (Suc n)))" by(rule wf_action_translation_on_inj_onD) from J have C_sub_A: "?C (Suc n) \ actions (?E (Suc n))" by(simp add: committed_subset_actions_def) from J have CnCSn: "?\ n ` ?C n \ ?\ (Suc n) ` ?C (Suc n)" by(simp add: is_commit_sequence_def) from J have wsSn: "is_write_seen P (?E (Suc n)) (?ws (Suc n))" by(simp add: justification_well_formed_def) from r obtain ad al v where "action_obs (?E (Suc n)) r = NormalAction (ReadMem ad al v)" by cases from is_write_seenD[OF wsSn r this] have wsSn: "?ws (Suc n) r \ actions (?E (Suc n))" by simp show ?case proof(cases "?\ (Suc n) r \ ?\ n ` ?C n") case True then obtain r' where r': "r' \ ?C n" and r_r': "?\ (Suc n) r = ?\ n r'" by(auto) from r' wfan have "action_tid (?E n) r' = action_tid E (?\ n r')" and "action_obs (?E n) r' \ action_obs E (?\ n r')" by(blast dest: wf_action_translation_on_actionD)+ moreover from r' CnCSn have "?\ (Suc n) r \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto hence "r \ ?C (Suc n)" unfolding inj_on_image_mem_iff[OF injSn C_sub_A \r \ actions (?E (Suc n))\] . with wfaSn have "action_tid (?E (Suc n)) r = action_tid E (?\ (Suc n) r)" and "action_obs (?E (Suc n)) r \ action_obs E (?\ (Suc n) r)" by(blast dest: wf_action_translation_on_actionD)+ ultimately have tid: "action_tid (?E n) r' = action_tid (?E (Suc n)) r" and obs: "action_obs (?E n) r' \ action_obs (?E (Suc n)) r" unfolding r_r' by(auto intro: sim_action_trans sim_action_sym) from J have "?C n \ actions (?E n)" by(simp add: committed_subset_actions_def) with r' have "r' \ actions (?E n)" by blast with r obs have "r' \ read_actions (?E n)" by cases(auto intro: read_actions.intros) hence ws: "ws (?\ n r') \ ?\ n ` ?C n" using r' by(rule Suc) have r_conv_inv: "r = inv_into (actions (?E (Suc n))) (?\ (Suc n)) (?\ n r')" using \r \ actions (?E (Suc n))\ unfolding r_r'[symmetric] by(simp add: inv_into_f_f[OF injSn]) with \r' \ ?C n\ r J \r' \ read_actions (?E n)\ have ws_eq: "?\ (Suc n) (?ws (Suc n) r) = ws (?\ n r')" by(simp add: Let_def write_seen_committed_def) with ws CnCSn have "?\ (Suc n) (?ws (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" by auto hence "?ws (Suc n) r \ ?C (Suc n)" by(subst (asm) inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) moreover from ws CnCSn have "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto ultimately show ?thesis by simp next case False with r \r \ ?C (Suc n)\ J have "ws (?\ (Suc n) r) \ ?\ n ` ?C n" unfolding is_weakly_justified_by.simps Let_def committed_reads_see_committed_writes_weak_def by blast hence "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" using CnCSn by blast+ thus ?thesis by(simp add: inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) qed qed lemma justified_write_seen_hb_read_committed: assumes J: "P \ (E, ws) justified_by J" and r: "r \ read_actions (justifying_exec (J n))" "r \ committed (J n)" shows "justifying_ws (J n) r \ committed (J n)" (is ?thesis1) and "ws (action_translation (J n) r) \ action_translation (J n) ` committed (J n)" (is ?thesis2) proof - have "?thesis1 \ ?thesis2" using r proof(induct n arbitrary: r) case 0 from J have [simp]: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) with 0 show ?case by simp next case (Suc n) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" note r = \r \ read_actions (?E (Suc n))\ hence "r \ actions (?E (Suc n))" by simp from J have wfan: "wf_action_translation_on (?E n) E (?C n) (?\ n)" and wfaSn: "wf_action_translation_on (?E (Suc n)) E (?C (Suc n)) (?\ (Suc n))" by(simp_all add: wf_action_translations_def) from wfaSn have injSn: "inj_on (?\ (Suc n)) (actions (?E (Suc n)))" by(rule wf_action_translation_on_inj_onD) from J have C_sub_A: "?C (Suc n) \ actions (?E (Suc n))" by(simp add: committed_subset_actions_def) from J have CnCSn: "?\ n ` ?C n \ ?\ (Suc n) ` ?C (Suc n)" by(simp add: is_commit_sequence_def) from J have wsSn: "is_write_seen P (?E (Suc n)) (?ws (Suc n))" by(simp add: justification_well_formed_def) from r obtain ad al v where "action_obs (?E (Suc n)) r = NormalAction (ReadMem ad al v)" by cases from is_write_seenD[OF wsSn r this] have wsSn: "?ws (Suc n) r \ actions (?E (Suc n))" by simp show ?case proof(cases "?\ (Suc n) r \ ?\ n ` ?C n") case True then obtain r' where r': "r' \ ?C n" and r_r': "?\ (Suc n) r = ?\ n r'" by(auto) from r' wfan have "action_tid (?E n) r' = action_tid E (?\ n r')" and "action_obs (?E n) r' \ action_obs E (?\ n r')" by(blast dest: wf_action_translation_on_actionD)+ moreover from r' CnCSn have "?\ (Suc n) r \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto hence "r \ ?C (Suc n)" unfolding inj_on_image_mem_iff[OF injSn C_sub_A \r \ actions (?E (Suc n))\] . with wfaSn have "action_tid (?E (Suc n)) r = action_tid E (?\ (Suc n) r)" and "action_obs (?E (Suc n)) r \ action_obs E (?\ (Suc n) r)" by(blast dest: wf_action_translation_on_actionD)+ ultimately have tid: "action_tid (?E n) r' = action_tid (?E (Suc n)) r" and obs: "action_obs (?E n) r' \ action_obs (?E (Suc n)) r" unfolding r_r' by(auto intro: sim_action_trans sim_action_sym) from J have "?C n \ actions (?E n)" by(simp add: committed_subset_actions_def) with r' have "r' \ actions (?E n)" by blast with r obs have "r' \ read_actions (?E n)" by cases(auto intro: read_actions.intros) hence "?ws n r' \ ?C n \ ws (?\ n r') \ ?\ n ` ?C n" using r' by(rule Suc) then obtain ws: "ws (?\ n r') \ ?\ n ` ?C n" .. have r_conv_inv: "r = inv_into (actions (?E (Suc n))) (?\ (Suc n)) (?\ n r')" using \r \ actions (?E (Suc n))\ unfolding r_r'[symmetric] by(simp add: inv_into_f_f[OF injSn]) with \r' \ ?C n\ r J \r' \ read_actions (?E n)\ have ws_eq: "?\ (Suc n) (?ws (Suc n) r) = ws (?\ n r')" by(simp add: Let_def write_seen_committed_def) with ws CnCSn have "?\ (Suc n) (?ws (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" by auto hence "?ws (Suc n) r \ ?C (Suc n)" by(subst (asm) inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) moreover from ws CnCSn have "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" unfolding r_r' by auto ultimately show ?thesis by simp next case False with r \r \ ?C (Suc n)\ J have "?\ (Suc n) (?ws (Suc n) r) \ ?\ n ` ?C n" and "ws (?\ (Suc n) r) \ ?\ n ` ?C n" unfolding is_justified_by.simps Let_def committed_reads_see_committed_writes_def by blast+ hence "?\ (Suc n) (?ws (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" and "ws (?\ (Suc n) r) \ ?\ (Suc n) ` ?C (Suc n)" using CnCSn by blast+ thus ?thesis by(simp add: inj_on_image_mem_iff[OF injSn C_sub_A wsSn]) qed qed thus ?thesis1 ?thesis2 by simp_all qed lemma is_justified_by_imp_is_weakly_justified_by: assumes justified: "P \ (E, ws) justified_by J" and wf: "P \ (E, ws) \" shows "P \ (E, ws) weakly_justified_by J" unfolding is_weakly_justified_by.simps proof(intro conjI) let ?E = "\n. justifying_exec (J n)" and ?ws = "\n. justifying_ws (J n)" and ?C = "\n. committed (J n)" and ?\ = "\n. action_translation (J n)" from justified show "is_commit_sequence E J" "justification_well_formed P J" "committed_subset_actions J" "value_written_committed P E J" "write_seen_committed ws J" "uncommitted_reads_see_hb P J" "wf_action_translations E J" by(simp_all) show "happens_before_committed_weak P E ws J" unfolding happens_before_committed_weak_def Let_def proof(intro strip conjI) fix n r assume "r \ read_actions (?E n) \ ?C n" hence read: "r \ read_actions (?E n)" and committed: "r \ ?C n" by simp_all with justified have committed_ws: "?ws n r \ ?C n" and committed_ws': "ws (?\ n r) \ ?\ n ` ?C n" by(rule justified_write_seen_hb_read_committed)+ from committed_ws' obtain w where w: "ws (?\ n r) = ?\ n w" and committed_w: "w \ ?C n" by blast from committed_w justified have "w \ actions (?E n)" by(auto simp add: committed_subset_actions_def) moreover from justified have "inj_on (?\ n) (actions (?E n))" by(auto simp add: wf_action_translations_def dest: wf_action_translation_on_inj_onD) ultimately have w_def: "w = inv_into (actions (?E n)) (?\ n) (ws (?\ n r))" by(simp_all add: w) from committed committed_w have "P,?E n \ w \hb r \ (happens_before P (?E n) |` ?C n) w r" by auto also have "\ \ (inv_imageP (happens_before P E) (?\ n) |` ?C n) w r" using justified by(simp add: happens_before_committed_def) also have "\ \ P,E \ ?\ n w \hb ?\ n r" using committed committed_w by auto finally show "P,E \ ws (?\ n r) \hb ?\ n r \ P,?E n \ inv_into (actions (?E n)) (?\ n) (ws (?\ n r)) \hb r" unfolding w[symmetric] unfolding w_def .. have "P,?E n\ r \hb w \ (happens_before P (?E n) |` ?C n) r w" using committed committed_w by auto also have "\ \ (inv_imageP (happens_before P E) (?\ n) |` ?C n) r w" using justified by(simp add: happens_before_committed_def) also have "\ \ P,E \ ?\ n r \hb ws (?\ n r)" using w committed committed_w by auto also { from read obtain ad al v where "action_obs (?E n) r = NormalAction (ReadMem ad al v)" by cases auto with justified committed obtain v' where obs': "action_obs E (?\ n r) = NormalAction (ReadMem ad al v')" by(fastforce simp add: wf_action_translations_def dest!: wf_action_translation_on_actionD) moreover from committed justified have "?\ n r \ actions E" by(auto simp add: is_commit_sequence_def) ultimately have read': "?\ n r \ read_actions E" by(auto intro: read_actions.intros) from wf have "is_write_seen P E ws" by(rule wf_exec_is_write_seenD) from is_write_seenD[OF this read' obs'] have "\ P,E \ ?\ n r \hb ws (?\ n r)" by simp } ultimately show "\ P,?E n \ r \hb inv_into (actions (?E n)) (?\ n) (ws (?\ n r))" unfolding w_def by simp qed from justified have "committed_reads_see_committed_writes ws J" by simp thus "committed_reads_see_committed_writes_weak ws J" by(auto simp add: committed_reads_see_committed_writes_def committed_reads_see_committed_writes_weak_def) qed corollary legal_imp_weakly_legal_execution: "legal_execution P \ Ews \ weakly_legal_execution P \ Ews" by(cases Ews)(auto 4 4 simp add: gen_legal_execution.simps simp del: is_justified_by.simps is_weakly_justified_by.simps intro: is_justified_by_imp_is_weakly_justified_by) lemma drop_0th_justifying_exec: assumes "P \ (E, ws) justified_by J" and wf: "P \ (E', ws') \" shows "P \ (E, ws) justified_by (J(0 := \committed = {}, justifying_exec = E', justifying_ws = ws', action_translation = id\))" (is "_ \ _ justified_by ?J") using assms unfolding is_justified_by.simps is_commit_sequence_def justification_well_formed_def committed_subset_actions_def happens_before_committed_def sync_order_committed_def value_written_committed_def write_seen_committed_def uncommitted_reads_see_hb_def committed_reads_see_committed_writes_def external_actions_committed_def wf_action_translations_def proof(intro conjI strip) let ?E = "\n. justifying_exec (?J n)" and ?\ = "\n. action_translation (?J n)" and ?C = "\n. committed (?J n)" and ?ws = "\n. justifying_ws (?J n)" show "?C 0 = {}" by simp from assms have C_0: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) hence "(\n. ?\ n ` ?C n) = (\n. action_translation (J n) ` committed (J n))" by -(rule SUP_cong, simp_all) also have "\ = actions E" using assms by(simp add: is_commit_sequence_def) finally show "actions E = (\n. ?\ n ` ?C n)" .. fix n { fix r' assume "r' \ read_actions (?E (Suc n))" thus "?\ (Suc n) r' \ ?\ n ` ?C n \ P,?E (Suc n) \ ?ws (Suc n) r' \hb r'" using assms by(auto dest!: bspec simp add: uncommitted_reads_see_hb_def is_commit_sequence_def) } { fix r' assume r': "r' \ read_actions (?E (Suc n)) \ ?C (Suc n)" have "n \ 0" proof assume "n = 0" hence "r' \ read_actions (justifying_exec (J 1)) \ committed (J 1)" using r' by simp hence "action_translation (J 1) r' \ action_translation (J 0) ` committed (J 0) \ ws (action_translation (J 1) r') \ action_translation (J 0) ` committed (J 0)" using assms unfolding One_nat_def is_justified_by.simps Let_def committed_reads_see_committed_writes_def by(metis (lifting)) thus False unfolding C_0 by simp qed thus "let r = ?\ (Suc n) r'; committed_n = ?\ n ` ?C n in r \ committed_n \ (?\ (Suc n) (?ws (Suc n) r') \ committed_n \ ws r \ committed_n)" using assms r' by(simp add: committed_reads_see_committed_writes_def) } { fix a a' assume "a \ external_actions (?E n)" and "a' \ ?C n" "P,?E n \ a \hb a'" moreover hence "n > 0" by(simp split: if_split_asm) ultimately show "a \ ?C n" using assms by(simp add: external_actions_committed_def) blast } from assms have "wf_action_translation E (?J 0)" by(simp add: wf_action_translations_def wf_action_translation_on_def) thus "wf_action_translation E (?J n)" using assms by(simp add: wf_action_translations_def) qed auto lemma drop_0th_weakly_justifying_exec: assumes "P \ (E, ws) weakly_justified_by J" and wf: "P \ (E', ws') \" shows "P \ (E, ws) weakly_justified_by (J(0 := \committed = {}, justifying_exec = E', justifying_ws = ws', action_translation = id\))" (is "_ \ _ weakly_justified_by ?J") using assms unfolding is_weakly_justified_by.simps is_commit_sequence_def justification_well_formed_def committed_subset_actions_def happens_before_committed_weak_def value_written_committed_def write_seen_committed_def uncommitted_reads_see_hb_def committed_reads_see_committed_writes_weak_def external_actions_committed_def wf_action_translations_def proof(intro conjI strip) let ?E = "\n. justifying_exec (?J n)" and ?\ = "\n. action_translation (?J n)" and ?C = "\n. committed (?J n)" and ?ws = "\n. justifying_ws (?J n)" show "?C 0 = {}" by simp from assms have C_0: "committed (J 0) = {}" by(simp add: is_commit_sequence_def) hence "(\n. ?\ n ` ?C n) = (\n. action_translation (J n) ` committed (J n))" by -(rule SUP_cong, simp_all) also have "\ = actions E" using assms by(simp add: is_commit_sequence_def) finally show "actions E = (\n. ?\ n ` ?C n)" .. fix n { fix r' assume "r' \ read_actions (?E (Suc n))" thus "?\ (Suc n) r' \ ?\ n ` ?C n \ P,?E (Suc n) \ ?ws (Suc n) r' \hb r'" using assms by(auto dest!: bspec simp add: uncommitted_reads_see_hb_def is_commit_sequence_def) } { fix r' assume r': "r' \ read_actions (?E (Suc n)) \ ?C (Suc n)" have "n \ 0" proof assume "n = 0" hence "r' \ read_actions (justifying_exec (J 1)) \ committed (J 1)" using r' by simp hence "action_translation (J 1) r' \ action_translation (J 0) ` committed (J 0) \ ws (action_translation (J 1) r') \ action_translation (J 0) ` committed (J 0)" using assms unfolding One_nat_def is_weakly_justified_by.simps Let_def committed_reads_see_committed_writes_weak_def by(metis (lifting)) thus False unfolding C_0 by simp qed thus "let r = ?\ (Suc n) r'; committed_n = ?\ n ` ?C n in r \ committed_n \ ws r \ committed_n" using assms r' by(simp add: committed_reads_see_committed_writes_weak_def) } from assms have "wf_action_translation E (?J 0)" by(simp add: wf_action_translations_def wf_action_translation_on_def) thus "wf_action_translation E (?J n)" using assms by(simp add: wf_action_translations_def) qed auto subsection \Executions with common prefix\ lemma actions_change_prefix: assumes read: "a \ actions E" and prefix: "ltake n E [\] ltake n E'" and rn: "enat a < n" shows "a \ actions E'" using llist_all2_llengthD[OF prefix[unfolded sim_actions_def]] read rn by(simp add: actions_def min_def split: if_split_asm) lemma action_obs_change_prefix: assumes prefix: "ltake n E [\] ltake n E'" and rn: "enat a < n" shows "action_obs E a \ action_obs E' a" proof - from rn have "action_obs E a = action_obs (ltake n E) a" by(simp add: action_obs_def lnth_ltake) also from prefix have "\ \ action_obs (ltake n E') a" by(rule sim_actions_action_obsD) also have "\ = action_obs E' a" using rn by(simp add: action_obs_def lnth_ltake) finally show ?thesis . qed lemma action_obs_change_prefix_eq: assumes prefix: "ltake n E = ltake n E'" and rn: "enat a < n" shows "action_obs E a = action_obs E' a" proof - from rn have "action_obs E a = action_obs (ltake n E) a" by(simp add: action_obs_def lnth_ltake) also from prefix have "\ = action_obs (ltake n E') a" by(simp add: action_obs_def) also have "\ = action_obs E' a" using rn by(simp add: action_obs_def lnth_ltake) finally show ?thesis . qed lemma read_actions_change_prefix: assumes read: "r \ read_actions E" and prefix: "ltake n E [\] ltake n E'" "enat r < n" shows "r \ read_actions E'" using read action_obs_change_prefix[OF prefix] actions_change_prefix[OF _ prefix] by(cases)(auto intro: read_actions.intros) lemma sim_action_is_write_action_eq: assumes "obs \ obs'" shows "is_write_action obs \ is_write_action obs'" using assms by cases simp_all lemma write_actions_change_prefix: assumes "write": "w \ write_actions E" and prefix: "ltake n E [\] ltake n E'" "enat w < n" shows "w \ write_actions E'" using "write" action_obs_change_prefix[OF prefix] actions_change_prefix[OF _ prefix] by(cases)(auto intro: write_actions.intros dest: sim_action_is_write_action_eq) lemma action_loc_change_prefix: assumes "ltake n E [\] ltake n E'" "enat a < n" shows "action_loc P E a = action_loc P E' a" using action_obs_change_prefix[OF assms] by(fastforce elim!: action_loc_aux_cases intro: action_loc_aux_intros) lemma sim_action_is_new_action_eq: assumes "obs \ obs'" shows "is_new_action obs = is_new_action obs'" using assms by cases auto lemma action_order_change_prefix: assumes ao: "E \ a \a a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "E' \ a \a a'" using ao actions_change_prefix[OF _ prefix an] actions_change_prefix[OF _ prefix a'n] action_obs_change_prefix[OF prefix an] action_obs_change_prefix[OF prefix a'n] by(auto simp add: action_order_def split: if_split_asm dest: sim_action_is_new_action_eq) lemma value_written_change_prefix: assumes eq: "ltake n E = ltake n E'" and an: "enat a < n" shows "value_written P E a = value_written P E' a" using action_obs_change_prefix_eq[OF eq an] by(simp add: value_written_def fun_eq_iff) lemma action_tid_change_prefix: assumes prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" shows "action_tid E a = action_tid E' a" proof - from an have "action_tid E a = action_tid (ltake n E) a" by(simp add: action_tid_def lnth_ltake) also from prefix have "\ = action_tid (ltake n E') a" by(rule sim_actions_action_tidD) also from an have "\ = action_tid E' a" by(simp add: action_tid_def lnth_ltake) finally show ?thesis . qed lemma program_order_change_prefix: assumes po: "E \ a \po a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "E' \ a \po a'" using po action_order_change_prefix[OF _ prefix an a'n] action_tid_change_prefix[OF prefix an] action_tid_change_prefix[OF prefix a'n] by(auto elim!: program_orderE intro: program_orderI) lemma sim_action_sactionD: assumes "obs \ obs'" shows "saction P obs \ saction P obs'" using assms by cases simp_all lemma sactions_change_prefix: assumes sync: "a \ sactions P E" and prefix: "ltake n E [\] ltake n E'" and rn: "enat a < n" shows "a \ sactions P E'" using sync action_obs_change_prefix[OF prefix rn] actions_change_prefix[OF _ prefix rn] unfolding sactions_def by(simp add: sim_action_sactionD) lemma sync_order_change_prefix: assumes so: "P,E \ a \so a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "P,E' \ a \so a'" using so action_order_change_prefix[OF _ prefix an a'n] sactions_change_prefix[OF _ prefix an, of P] sactions_change_prefix[OF _ prefix a'n, of P] by(simp add: sync_order_def) lemma sim_action_synchronizes_withD: assumes "obs \ obs'" "obs'' \ obs'''" shows "P \ (t, obs) \sw (t', obs'') \ P \ (t, obs') \sw (t', obs''')" using assms by(auto elim!: sim_action.cases synchronizes_with.cases intro: synchronizes_with.intros) lemma sync_with_change_prefix: assumes sw: "P,E \ a \sw a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "P,E' \ a \sw a'" using sw sync_order_change_prefix[OF _ prefix an a'n, of P] action_tid_change_prefix[OF prefix an] action_tid_change_prefix[OF prefix a'n] action_obs_change_prefix[OF prefix an] action_obs_change_prefix[OF prefix a'n] by(auto simp add: sync_with_def dest: sim_action_synchronizes_withD) lemma po_sw_change_prefix: assumes posw: "po_sw P E a a'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "po_sw P E' a a'" using posw sync_with_change_prefix[OF _ prefix an a'n, of P] program_order_change_prefix[OF _ prefix an a'n] by(auto simp add: po_sw_def) lemma happens_before_new_not_new: assumes tsa_ok: "thread_start_actions_ok E" and a: "a \ actions E" and a': "a' \ actions E" and new_a: "is_new_action (action_obs E a)" and new_a': "\ is_new_action (action_obs E a')" shows "P,E \ a \hb a'" proof - from thread_start_actions_okD[OF tsa_ok a' new_a'] obtain i where "i \ a'" and obs_i: "action_obs E i = InitialThreadAction" and "action_tid E i = action_tid E a'" by auto from \i \ a'\ a' have "i \ actions E" by(auto simp add: actions_def le_less_trans[where y="enat a'"]) with \i \ a'\ obs_i a' new_a' have "E \ i \a a'" by(simp add: action_order_def) hence "E \ i \po a'" using \action_tid E i = action_tid E a'\ by(rule program_orderI) moreover { from \i \ actions E\ obs_i have "i \ sactions P E" by(auto intro: sactionsI) from a \i \ actions E\ new_a obs_i have "E \ a \a i" by(simp add: action_order_def) moreover from a new_a have "a \ sactions P E" by(auto intro: sactionsI) ultimately have "P,E \ a \so i" using \i \ sactions P E\ by(rule sync_orderI) moreover from new_a obs_i have "P \ (action_tid E a, action_obs E a) \sw (action_tid E i, action_obs E i)" by cases(auto intro: synchronizes_with.intros) ultimately have "P,E \ a \sw i" by(rule sync_withI) } ultimately show ?thesis unfolding po_sw_def [abs_def] by(blast intro: tranclp.r_into_trancl tranclp_trans) qed lemma happens_before_change_prefix: assumes hb: "P,E \ a \hb a'" and tsa_ok: "thread_start_actions_ok E'" and prefix: "ltake n E [\] ltake n E'" and an: "enat a < n" and a'n: "enat a' < n" shows "P,E' \ a \hb a'" using hb an a'n proof induct case (base a') thus ?case by(rule tranclp.r_into_trancl[where r="po_sw P E'", OF po_sw_change_prefix[OF _ prefix]]) next case (step a' a'') show ?case proof(cases "is_new_action (action_obs E a') \ \ is_new_action (action_obs E a'')") case False from \po_sw P E a' a''\ have "E \ a' \a a''" by(rule po_sw_into_action_order) with \enat a'' < n\ False have "enat a' < n" by(safe elim!: action_orderE)(metis Suc_leI Suc_n_not_le_n enat_ord_simps(2) le_trans nat_neq_iff xtrans(10))+ with \enat a < n\ have "P,E' \ a \hb a'" by(rule step) moreover from \po_sw P E a' a''\ prefix \enat a' < n\ \enat a'' < n\ have "po_sw P E' a' a''" by(rule po_sw_change_prefix) ultimately show ?thesis .. next case True then obtain new_a': "is_new_action (action_obs E a')" and "\ is_new_action (action_obs E a'')" .. from \P,E \ a \hb a'\ new_a' have new_a: "is_new_action (action_obs E a)" and tid: "action_tid E a = action_tid E a'" and "a \ a'" by(rule happens_before_new_actionD)+ note tsa_ok moreover from porder_happens_before[of E P] have "a \ actions E" by(rule porder_onE)(erule refl_onPD1, rule \P,E \ a \hb a'\) hence "a \ actions E'" using an by(rule actions_change_prefix[OF _ prefix]) moreover from \po_sw P E a' a''\ refl_on_program_order[of E] refl_on_sync_order[of P E] have "a'' \ actions E" unfolding po_sw_def by(auto dest: refl_onPD2 elim!: sync_withE) hence "a'' \ actions E'" using \enat a'' < n\ by(rule actions_change_prefix[OF _ prefix]) moreover from new_a action_obs_change_prefix[OF prefix an] have "is_new_action (action_obs E' a)" by(cases) auto moreover from \\ is_new_action (action_obs E a'')\ action_obs_change_prefix[OF prefix \enat a'' < n\] have "\ is_new_action (action_obs E' a'')" by(auto elim: is_new_action.cases) ultimately show "P,E' \ a \hb a''" by(rule happens_before_new_not_new) qed qed lemma thread_start_actions_ok_change: assumes tsa: "thread_start_actions_ok E" and sim: "E [\] E'" shows "thread_start_actions_ok E'" proof(rule thread_start_actions_okI) fix a assume "a \ actions E'" "\ is_new_action (action_obs E' a)" from sim have len_eq: "llength E = llength E'" by(simp add: sim_actions_def)(rule llist_all2_llengthD) with sim have sim': "ltake (llength E) E [\] ltake (llength E) E'" by(simp add: ltake_all) from \a \ actions E'\ len_eq have "enat a < llength E" by(simp add: actions_def) with \a \ actions E'\ sim'[symmetric] have "a \ actions E" by(rule actions_change_prefix) moreover have "\ is_new_action (action_obs E a)" using action_obs_change_prefix[OF sim' \enat a < llength E\] \\ is_new_action (action_obs E' a)\ by(auto elim!: is_new_action.cases) ultimately obtain i where "i \ a" "action_obs E i = InitialThreadAction" "action_tid E i = action_tid E a" by(blast dest: thread_start_actions_okD[OF tsa]) thus "\i \ a. action_obs E' i = InitialThreadAction \ action_tid E' i = action_tid E' a" using action_tid_change_prefix[OF sim', of i] action_tid_change_prefix[OF sim', of a] \enat a < llength E\ action_obs_change_prefix[OF sim', of i] by(cases "llength E")(auto intro!: exI[where x=i]) qed context executions_aux begin lemma \_new_same_addr_singleton: assumes E: "E \ \" shows "\a. new_actions_for P E adal \ {a}" by(blast dest: \_new_actions_for_fun[OF E]) lemma new_action_before_read: assumes E: "E \ \" and wf: "P \ (E, ws) \" and ra: "ra \ read_actions E" and adal: "adal \ action_loc P E ra" and new: "wa \ new_actions_for P E adal" and sc: "\a. \ a < ra; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "wa < ra" using \_new_same_addr_singleton[OF E, of adal] init_before_read[OF E wf ra adal sc] new by auto lemma mrw_before: assumes E: "E \ \" and wf: "P \ (E, ws) \" and mrw: "P,E \ r \mrw w" and sc: "\a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "w < r" using mrw read_actions_not_write_actions[of r E] apply cases apply(erule action_orderE) apply(erule (1) new_action_before_read[OF E wf]) apply(simp add: new_actions_for_def) apply(erule (1) sc) apply(cases "w = r") apply auto done lemma mrw_change_prefix: assumes E': "E' \ \" and mrw: "P,E \ r \mrw w" and tsa_ok: "thread_start_actions_ok E'" and prefix: "ltake n E [\] ltake n E'" and an: "enat r < n" and a'n: "enat w < n" shows "P,E' \ r \mrw w" using mrw proof cases fix adal assume r: "r \ read_actions E" and adal_r: "adal \ action_loc P E r" and war: "E \ w \a r" and w: "w \ write_actions E" and adal_w: "adal \ action_loc P E w" and mrw: "\wa'. \wa' \ write_actions E; adal \ action_loc P E wa'\ \ E \ wa' \a w \ E \ r \a wa'" show ?thesis proof(rule most_recent_write_for.intros) from r prefix an show r': "r \ read_actions E'" by(rule read_actions_change_prefix) from adal_r show "adal \ action_loc P E' r" by(simp add: action_loc_change_prefix[OF prefix[symmetric] an]) from war prefix a'n an show "E' \ w \a r" by(rule action_order_change_prefix) from w prefix a'n show w': "w \ write_actions E'" by(rule write_actions_change_prefix) from adal_w show adal_w': "adal \ action_loc P E' w" by(simp add: action_loc_change_prefix[OF prefix[symmetric] a'n]) fix wa' assume wa': "wa' \ write_actions E'" and adal_wa': "adal \ action_loc P E' wa'" show "E' \ wa' \a w \ E' \ r \a wa'" proof(cases "enat wa' < n") case True note wa'n = this with wa' prefix[symmetric] have "wa' \ write_actions E" by(rule write_actions_change_prefix) moreover from adal_wa' have "adal \ action_loc P E wa'" by(simp add: action_loc_change_prefix[OF prefix wa'n]) ultimately have "E \ wa' \a w \ E \ r \a wa'" by(rule mrw) thus ?thesis proof assume "E \ wa' \a w" hence "E' \ wa' \a w" using prefix wa'n a'n by(rule action_order_change_prefix) thus ?thesis .. next assume "E \ r \a wa'" hence "E' \ r \a wa'" using prefix an wa'n by(rule action_order_change_prefix) thus ?thesis .. qed next case False note wa'n = this show ?thesis proof(cases "is_new_action (action_obs E' wa')") case False hence "E' \ r \a wa'" using wa'n r' wa' an by(auto intro!: action_orderI) (metis enat_ord_code(1) linorder_le_cases order_le_less_trans) thus ?thesis .. next case True with wa' adal_wa' have new: "wa' \ new_actions_for P E' adal" by(simp add: new_actions_for_def) show ?thesis proof(cases "is_new_action (action_obs E' w)") case True with adal_w' a'n w' have "w \ new_actions_for P E' adal" by(simp add: new_actions_for_def) with E' new have "wa' = w" by(rule \_new_actions_for_fun) thus ?thesis using w' by(auto intro: refl_onPD[OF refl_action_order]) next case False with True wa' w' show ?thesis by(auto intro!: action_orderI) qed qed qed qed qed lemma action_order_read_before_write: assumes E: "E \ \" "P \ (E, ws) \" and ao: "E \ w \a r" and r: "r \ read_actions E" and w: "w \ write_actions E" and adal: "adal \ action_loc P E r" "adal \ action_loc P E w" and sc: "\a. \ a < r; a \ read_actions E \ \ P,E \ a \mrw ws a" shows "w < r" using ao proof(cases rule: action_orderE) case 1 from init_before_read[OF E r adal(1) sc] obtain i where "i < r" "i \ new_actions_for P E adal" by blast moreover from \is_new_action (action_obs E w)\ adal(2) \w \ actions E\ have "w \ new_actions_for P E adal" by(simp add: new_actions_for_def) ultimately show "w < r" using E by(auto dest: \_new_actions_for_fun) next case 2 with r w show ?thesis by(cases "w = r")(auto dest: read_actions_not_write_actions) qed end end diff --git a/thys/Knuth_Morris_Pratt/KMP.thy b/thys/Knuth_Morris_Pratt/KMP.thy --- a/thys/Knuth_Morris_Pratt/KMP.thy +++ b/thys/Knuth_Morris_Pratt/KMP.thy @@ -1,938 +1,939 @@ theory KMP imports Refine_Imperative_HOL.IICF "HOL-Library.Sublist" begin declare len_greater_imp_nonempty[simp del] min_absorb2[simp] no_notation Ref.update ("_ := _" 62) section\Specification\text_raw\\label{sec:spec}\ subsection\Sublist-predicate with a position check\ subsubsection\Definition\ text \One could define\ definition "sublist_at' xs ys i \ take (length xs) (drop i ys) = xs" text\However, this doesn't handle out-of-bound indexes uniformly:\ value[nbe] "sublist_at' [] [a] 5" value[nbe] "sublist_at' [a] [a] 5" value[nbe] "sublist_at' [] [] 5" text\Instead, we use a recursive definition:\ fun sublist_at :: "'a list \ 'a list \ nat \ bool" where "sublist_at (x#xs) (y#ys) 0 \ x=y \ sublist_at xs ys 0" | "sublist_at xs (y#ys) (Suc i) \ sublist_at xs ys i" | "sublist_at [] ys 0 \ True" | "sublist_at _ [] _ \ False" text\In the relevant cases, both definitions agree:\ lemma "i \ length ys \ sublist_at xs ys i \ sublist_at' xs ys i" unfolding sublist_at'_def by (induction xs ys i rule: sublist_at.induct) auto text\However, the new definition has some reasonable properties:\ subsubsection\Properties\ lemma sublist_lengths: "sublist_at xs ys i \ i + length xs \ length ys" by (induction xs ys i rule: sublist_at.induct) auto lemma Nil_is_sublist: "sublist_at ([] :: 'x list) ys i \ i \ length ys" by (induction "[] :: 'x list" ys i rule: sublist_at.induct) auto text\Furthermore, we need:\ lemma sublist_step[intro]: "\i + length xs < length ys; sublist_at xs ys i; ys!(i + length xs) = x\ \ sublist_at (xs@[x]) ys i" apply (induction xs ys i rule: sublist_at.induct) apply auto using sublist_at.elims(3) by fastforce lemma all_positions_sublist: "\i + length xs \ length ys; \jj \ sublist_at xs ys i" proof (induction xs rule: rev_induct) case Nil then show ?case by (simp add: Nil_is_sublist) next case (snoc x xs) from \i + length (xs @ [x]) \ length ys\ have "i + length xs \ length ys" by simp moreover have "\jj \jjIt also connects well to theory @{theory "HOL-Library.Sublist"} (compare @{thm[source] sublist_def}):\ lemma sublist_at_altdef: "sublist_at xs ys i \ (\ps ss. ys = ps@xs@ss \ i = length ps)" proof (induction xs ys i rule: sublist_at.induct) case (2 ss t ts i) show "sublist_at ss (t#ts) (Suc i) \ (\xs ys. t#ts = xs@ss@ys \ Suc i = length xs)" (is "?lhs \ ?rhs") proof assume ?lhs then have "sublist_at ss ts i" by simp with "2.IH" obtain xs where "\ys. ts = xs@ss@ys \ i = length xs" by auto then have "\ys. t#ts = (t#xs)@ss@ys \ Suc i = length (t#xs)" by simp then show ?rhs by blast next assume ?rhs - then obtain xs where "\ys. t#ts = xs@ss@ys \ Suc i = length xs" by blast + then obtain xs where "\ys. t#ts = xs@ss@ys \ length xs = Suc i" + by (blast dest: sym) then have "\ys. ts = (tl xs)@ss@ys \ i = length (tl xs)" - by (metis hd_Cons_tl length_0_conv list.sel(3) nat.simps(3) size_Cons_lem_eq tl_append2) + by (auto simp add: length_Suc_conv) then have "\xs ys. ts = xs@ss@ys \ i = length xs" by blast with "2.IH" show ?lhs by simp qed qed auto corollary sublist_iff_sublist_at: "Sublist.sublist xs ys \ (\i. sublist_at xs ys i)" by (simp add: sublist_at_altdef Sublist.sublist_def) subsection\Sublist-check algorithms\ text\ We use the Isabelle Refinement Framework (Theory @{theory Refine_Monadic.Refine_Monadic}) to phrase the specification and the algorithm. \ text\@{term s} for "searchword" / "searchlist", @{term t} for "text"\ definition "kmp_SPEC s t = SPEC (\ None \ \i. sublist_at s t i | Some i \ sublist_at s t i \ (\iisublist_at s t ii))" lemma is_arg_min_id: "is_arg_min id P i \ P i \ (\iiP ii)" unfolding is_arg_min_def by auto lemma kmp_result: "kmp_SPEC s t = RETURN (if sublist s t then Some (LEAST i. sublist_at s t i) else None)" unfolding kmp_SPEC_def sublist_iff_sublist_at apply (auto intro: LeastI dest: not_less_Least split: option.splits) by (meson LeastI nat_neq_iff not_less_Least) corollary weak_kmp_SPEC: "kmp_SPEC s t \ SPEC (\pos. pos\None \ Sublist.sublist s t)" by (simp add: kmp_result) lemmas kmp_SPEC_altdefs = kmp_SPEC_def[folded is_arg_min_id] kmp_SPEC_def[folded sublist_iff_sublist_at] kmp_result section\Naive algorithm\ text\Since KMP is a direct advancement of the naive "test-all-starting-positions" approach, we provide it here for comparison:\ subsection\Invariants\ definition "I_out_na s t \ \(i,j,pos). (\iisublist_at s t ii) \ (case pos of None \ j = 0 | Some p \ p=i \ sublist_at s t i)" definition "I_in_na s t i \ \(j,pos). case pos of None \ j < length s \ (\jj sublist_at s t i" subsection\Algorithm\ (*Algorithm is common knowledge \ remove citation here, move explanations to KMP below?*) text\The following definition is taken from Helmut Seidl's lecture on algorithms and data structures@{cite GAD} except that we \<^item> output the identified position @{term \pos :: nat option\} instead of just @{const True} \<^item> use @{term \pos :: nat option\} as break-flag to support the abort within the loops \<^item> rewrite @{prop \i \ length t - length s\} in the first while-condition to @{prop \i + length s \ length t\} to avoid having to use @{typ int} for list indexes (or the additional precondition @{prop \length s \ length t\}) \ definition "naive_algorithm s t \ do { let i=0; let j=0; let pos=None; (_,_,pos) \ WHILEIT (I_out_na s t) (\(i,_,pos). i + length s \ length t \ pos=None) (\(i,j,pos). do { (_,pos) \ WHILEIT (I_in_na s t i) (\(j,pos). t!(i+j) = s!j \ pos=None) (\(j,_). do { let j=j+1; if j=length s then RETURN (j,Some i) else RETURN (j,None) }) (j,pos); if pos=None then do { let i = i + 1; let j = 0; RETURN (i,j,None) } else RETURN (i,j,Some i) }) (i,j,pos); RETURN pos }" subsection\Correctness\ text\The basic lemmas on @{const sublist_at} from the previous chapter together with @{theory Refine_Monadic.Refine_Monadic}'s verification condition generator / solver suffice:\ lemma "s \ [] \ naive_algorithm s t \ kmp_SPEC s t" unfolding naive_algorithm_def kmp_SPEC_def I_out_na_def I_in_na_def apply (refine_vcg WHILEIT_rule[where R="measure (\(i,_,pos). length t - i + (if pos = None then 1 else 0))"] WHILEIT_rule[where R="measure (\(j,_::nat option). length s - j)"] ) apply (vc_solve solve: asm_rl) subgoal by (metis add_Suc_right all_positions_sublist less_antisym) subgoal using less_Suc_eq by blast subgoal by (metis less_SucE sublist_all_positions) subgoal by (auto split: option.splits) (metis sublist_lengths add_less_cancel_right leI le_less_trans) done text\Note that the precondition cannot be removed without an extra branch: If @{prop \s = []\}, the inner while-condition accesses out-of-bound memory. This will apply to KMP, too.\ section\Knuth--Morris--Pratt algorithm\ text\Just like our templates@{cite KMP77}@{cite GAD}, we first verify the main routine and discuss the computation of the auxiliary values @{term \\ s\} only in a later section.\ subsection\Preliminaries: Borders of lists\ definition "border xs ys \ prefix xs ys \ suffix xs ys" definition "strict_border xs ys \ border xs ys \ length xs < length ys" definition "intrinsic_border ls \ ARG_MAX length b. strict_border b ls" subsubsection\Properties\ interpretation border_order: order border strict_border by standard (auto simp: border_def suffix_def strict_border_def) interpretation border_bot: order_bot Nil border strict_border by standard (simp add: border_def) lemma borderE[elim]: fixes xs ys :: "'a list" assumes "border xs ys" obtains "prefix xs ys" and "suffix xs ys" using assms unfolding border_def by blast lemma strict_borderE[elim]: fixes xs ys :: "'a list" assumes "strict_border xs ys" obtains "border xs ys" and "length xs < length ys" using assms unfolding strict_border_def by blast lemma strict_border_simps[simp]: "strict_border xs [] \ False" "strict_border [] (x # xs) \ True" by (simp_all add: strict_border_def) lemma strict_border_prefix: "strict_border xs ys \ strict_prefix xs ys" and strict_border_suffix: "strict_border xs ys \ strict_suffix xs ys" and strict_border_imp_nonempty: "strict_border xs ys \ ys \ []" and strict_border_prefix_suffix: "strict_border xs ys \ strict_prefix xs ys \ strict_suffix xs ys" by (auto simp: border_order.order.strict_iff_order border_def) lemma border_length_le: "border xs ys \ length xs \ length ys" unfolding border_def by (simp add: prefix_length_le) lemma border_length_r_less (*rm*): "\xs. strict_border xs ys \ length xs < length ys" using strict_borderE by auto lemma border_positions: "border xs ys \ \ii \ length w; i \ length x; \j \ drop (length w - i) w = take i x" by (cases "i = length x") (auto intro: nth_equalityI) lemma all_positions_suffix_take: "\i \ length w; i \ length x; \j \ suffix (take i x) w" by (metis all_positions_drop_length_take suffix_drop) lemma suffix_butlast: "suffix xs ys \ suffix (butlast xs) (butlast ys)" unfolding suffix_def by (metis append_Nil2 butlast.simps(1) butlast_append) lemma positions_border: "\j border (take l w) w" by (cases "l < length w") (simp_all add: border_def all_positions_suffix_take take_is_prefix) lemma positions_strict_border: "l < length w \ \j strict_border (take l w) w" by (simp add: positions_border strict_border_def) lemmas intrinsic_borderI = arg_max_natI[OF _ border_length_r_less, folded intrinsic_border_def] lemmas intrinsic_borderI' = border_bot.bot.not_eq_extremum[THEN iffD1, THEN intrinsic_borderI] lemmas intrinsic_border_max = arg_max_nat_le[OF _ border_length_r_less, folded intrinsic_border_def] lemma nonempty_is_arg_max_ib: "ys \ [] \ is_arg_max length (\xs. strict_border xs ys) (intrinsic_border ys)" by (simp add: intrinsic_borderI' intrinsic_border_max is_arg_max_linorder) lemma intrinsic_border_less: "w \ [] \ length (intrinsic_border w) < length w" using intrinsic_borderI[of w] border_length_r_less intrinsic_borderI' by blast lemma intrinsic_border_take_less: "j > 0 \ w \ [] \ length (intrinsic_border (take j w)) < length w" by (metis intrinsic_border_less length_take less_not_refl2 min_less_iff_conj take_eq_Nil) subsubsection\Examples\ lemma border_example: "{b. border b ''aabaabaa''} = {'''', ''a'', ''aa'', ''aabaa'', ''aabaabaa''}" (is "{b. border b ?l} = {?take0, ?take1, ?take2, ?take5, ?l}") proof show "{?take0, ?take1, ?take2, ?take5, ?l} \ {b. border b ?l}" by simp eval have "\border ''aab'' ?l" "\border ''aaba'' ?l" "\border ''aabaab'' ?l" "\border ''aabaaba'' ?l" by eval+ moreover have "{b. border b ?l} \ set (prefixes ?l)" using border_def in_set_prefixes by blast ultimately show "{b. border b ?l} \ {?take0, ?take1, ?take2, ?take5, ?l}" by auto qed corollary strict_border_example: "{b. strict_border b ''aabaabaa''} = {'''', ''a'', ''aa'', ''aabaa''}" (is "?l = ?r") proof have "?l \ {b. border b ''aabaabaa''}" by auto also have "\ = {'''', ''a'', ''aa'', ''aabaa'', ''aabaabaa''}" by (fact border_example) finally show "?l \ ?r" by auto show "?r \ ?l" by simp eval qed corollary "intrinsic_border ''aabaabaa'' = ''aabaa''" proof - \ \We later obtain a fast algorithm for that.\ have exhaust: "strict_border b ''aabaabaa'' \ b \ {'''', ''a'', ''aa'', ''aabaa''}" for b using strict_border_example by auto then have "\is_arg_max length (\b. strict_border b ''aabaabaa'') ''''" "\is_arg_max length (\b. strict_border b ''aabaabaa'') ''a''" "\is_arg_max length (\b. strict_border b ''aabaabaa'') ''aa''" "is_arg_max length (\b. strict_border b ''aabaabaa'') ''aabaa''" unfolding is_arg_max_linorder by auto moreover have "strict_border (intrinsic_border ''aabaabaa'') ''aabaabaa''" using intrinsic_borderI' by blast note this[unfolded exhaust] ultimately show ?thesis by simp (metis list.discI nonempty_is_arg_max_ib) qed subsection\Main routine\ text\The following is Seidl's "border"-table@{cite GAD} (values shifted by 1 so we don't need @{typ int}), or equivalently, "f" from Knuth's, Morris' and Pratt's paper@{cite KMP77} (with indexes starting at 0).\ fun \ :: "'a list \ nat \ nat" where "\ s 0 = 0" \ \This increments the compare position while @{prop \j=(0::nat)\}\ | "\ s j = length (intrinsic_border (take j s)) + 1" text\Note that we use their "next" only implicitly.\ subsubsection\Invariants\ definition "I_outer s t \ \(i,j,pos). (\iisublist_at s t ii) \ (case pos of None \ (\jj j < length s | Some p \ p=i \ sublist_at s t i)" text\For the inner loop, we can reuse @{const I_in_na}.\ subsubsection\Algorithm\ text\First, we use the non-evaluable function @{const \} directly:\ definition "kmp s t \ do { ASSERT (s \ []); let i=0; let j=0; let pos=None; (_,_,pos) \ WHILEIT (I_outer s t) (\(i,j,pos). i + length s \ length t \ pos=None) (\(i,j,pos). do { ASSERT (i + length s \ length t); (j,pos) \ WHILEIT (I_in_na s t i) (\(j,pos). t!(i+j) = s!j \ pos=None) (\(j,pos). do { let j=j+1; if j=length s then RETURN (j,Some i) else RETURN (j,None) }) (j,pos); if pos=None then do { ASSERT (j < length s); let i = i + (j - \ s j + 1); let j = max 0 (\ s j - 1); \ \\max\ not necessary\ RETURN (i,j,None) } else RETURN (i,j,Some i) }) (i,j,pos); RETURN pos }" subsubsection\Correctness\ lemma \_eq_0_iff_j_eq_0[simp]: "\ s j = 0 \ j = 0" by (cases j) simp_all lemma j_le_\_le: "j \ length s \ \ s j \ j" apply (cases j) apply simp_all by (metis Suc_leI intrinsic_border_less length_take list.size(3) min.absorb2 nat.simps(3) not_less) lemma j_le_\_le': "0 < j \ j \ length s \ \ s j - 1 < j" by (metis diff_less j_le_\_le le_eq_less_or_eq less_imp_diff_less less_one) lemma \_le: "s \ [] \ \ s j - 1 < length s" by (cases j) (simp_all add: intrinsic_border_take_less) (* Only needed for run-time analysis lemma "p576 et seq": assumes "j \ length s" and assignments: "i' = i + (j + 1 - \ s j)" "j' = max 0 (\ s j - 1)" shows sum_no_decrease: "i' + j' \ i + j" and i_increase: "i' > i" using assignments by (simp_all add: j_le_\_le[OF assms(1), THEN le_imp_less_Suc]) *) lemma reuse_matches: assumes j_le: "j \ length s" and old_matches: "\jjjj<\ s j - 1. t ! (i + (j - \ s j + 1) + jj) = s ! jj" (is "\jj0") assume "j>0" have \_le: "\ s j \ j" by (simp add: j_le j_le_\_le) with old_matches have 1: "\jj s j + 1 + jj)" by (metis ab_semigroup_add_class.add.commute add.assoc diff_diff_cancel less_diff_conv) have [simp]: "length (take j s) = j" "length (intrinsic_border (take j s)) = ?j'" by (simp add: j_le) (metis \0 < j\ diff_add_inverse2 \.elims nat_neq_iff) then have "\jj s j - 1) + jj)" by (metis intrinsic_borderI' \0 < j\ border_positions length_greater_0_conv strict_border_def) then have "\jj s j + 1 + jj)" by (simp add: \_le) then have 2: "\jj s j + 1 + jj) = s ! jj" using \_le by simp from 1 2 show ?thesis by simp qed simp theorem shift_safe: assumes "\iisublist_at s t ii" "t!(i+j) \ s!j" and [simp]: "j < length s" and matches: "\jj i + (j - \ s j + 1)" shows "\iisublist_at s t ii" proof (standard, standard) fix ii assume "ii < i'" then consider \ \The position falls into one of three categories:\ (old) "ii < i" | (current) "ii = i" | (skipped) "ii > i" by linarith then show "\sublist_at s t ii" proof cases case old \ \Old position, use invariant.\ with \\iisublist_at s t ii\ show ?thesis by simp next case current \ \The mismatch occurred while testing this alignment.\ with \t!(i+j) \ s!j\ show ?thesis using sublist_all_positions[of s t i] by auto next case skipped \ \The skipped positions.\ then have "0ii < i'\ assignment by linarith then have less_j[simp]: "j + i - ii < j" and le_s: "j + i - ii \ length s" using \ii < i'\ assms(3) skipped by linarith+ note \_le[simp] = j_le_\_le[OF assms(3)[THEN less_imp_le]] have "0 < \ s j" using \0 < j\ \_eq_0_iff_j_eq_0 neq0_conv by blast then have "j + i - ii > \ s j - 1" using \ii < i'\ assignment \_le by linarith then have contradiction_goal: "j + i - ii > length (intrinsic_border (take j s))" by (metis \.elims \0 < j\ add_diff_cancel_right' not_gr_zero) show ?thesis proof assume "sublist_at s t ii" note sublist_all_positions[OF this] with le_s have a: "\jj < j+i-ii. t!(ii+jj) = s!jj" by simp have ff1: "\ ii < i" by (metis not_less_iff_gr_or_eq skipped) then have "i + (ii - i + jj) = ii + jj" for jj by (metis add.assoc add_diff_inverse_nat) then have "\ jj < j + i - ii \ t ! (ii + jj) = s ! (ii - i + jj)" if "ii - i + jj < j" for jj using that ff1 by (metis matches) then have "\ jj < j + i - ii \ t ! (ii + jj) = s ! (ii - i + jj)" for jj using ff1 by auto with matches have "\jj < j+i-ii. t!(ii+jj) = s!(ii-i+jj)" by metis then have "\jj < j+i-ii. s!jj = s!(ii-i+jj)" using a by auto then have "\jj < j+i-ii. (take j s)!jj = (take j s)!(ii-i+jj)" using \i by auto with positions_strict_border[of "j+i-ii" "take j s", simplified] have "strict_border (take (j+i-ii) s) (take j s)". note intrinsic_border_max[OF this] also note contradiction_goal also have "j+i-ii \ length s" by (fact le_s) ultimately show False by simp qed qed qed lemma kmp_correct: "s \ [] \ kmp s t \ kmp_SPEC s t" unfolding kmp_def kmp_SPEC_def I_outer_def I_in_na_def apply (refine_vcg WHILEIT_rule[where R="measure (\(i,_,pos). length t - i + (if pos = None then 1 else 0))"] WHILEIT_rule[where R="measure (\(j,_::nat option). length s - j)"] ) apply (vc_solve solve: asm_rl) subgoal by (metis add_Suc_right all_positions_sublist less_antisym) subgoal using less_antisym by blast subgoal for i jout j using shift_safe[of i s t j] by fastforce subgoal for i jout j using reuse_matches[of j s t i] \_le by simp subgoal by (auto split: option.splits) (metis sublist_lengths add_less_cancel_right leI le_less_trans) done subsubsection\Storing the @{const \}-values\ text\We refine the algorithm to compute the @{const \}-values only once at the start:\ definition compute_\s_SPEC :: "'a list \ nat list nres" where "compute_\s_SPEC s \ SPEC (\\s. length \s = length s + 1 \ (\j\length s. \s!j = \ s j))" definition "kmp1 s t \ do { ASSERT (s \ []); let i=0; let j=0; let pos=None; \s \ compute_\s_SPEC (butlast s); \ \At the last char, we abort instead.\ (_,_,pos) \ WHILEIT (I_outer s t) (\(i,j,pos). i + length s \ length t \ pos=None) (\(i,j,pos). do { ASSERT (i + length s \ length t); (j,pos) \ WHILEIT (I_in_na s t i) (\(j,pos). t!(i+j) = s!j \ pos=None) (\(j,pos). do { let j=j+1; if j=length s then RETURN (j,Some i) else RETURN (j,None) }) (j,pos); if pos=None then do { ASSERT (j < length \s); let i = i + (j - \s!j + 1); let j = max 0 (\s!j - 1); \ \\max\ not necessary\ RETURN (i,j,None) } else RETURN (i,j,Some i) }) (i,j,pos); RETURN pos }" lemma \_butlast[simp]: "j < length s \ \ (butlast s) j = \ s j" by (cases j) (simp_all add: take_butlast) lemma kmp1_refine: "kmp1 s t \ kmp s t" apply (rule refine_IdD) unfolding kmp1_def kmp_def Let_def compute_\s_SPEC_def nres_monad_laws apply (intro ASSERT_refine_right ASSERT_refine_left) apply simp apply (rule Refine_Basic.intro_spec_refine) apply refine_rcg apply refine_dref_type apply vc_solve done text\Next, an algorithm that satisfies @{const compute_\s_SPEC}:\ subsection\Computing @{const \}\ subsubsection\Invariants\ definition "I_out_cb s \ \(\s,i,j). length s + 1 = length \s \ (\jjs!jj = \ s jj) \ \s!(j-1) = i \ 0 < j" definition "I_in_cb s j \ \i. if j=1 then i=0 \ \first iteration\ else strict_border (take (i-1) s) (take (j-1) s) \ \ s j \ i + 1" subsubsection\Algorithm\ text\Again, we follow Seidl@{cite GAD}, p.582. Apart from the +1-shift, we make another modification: Instead of directly setting @{term \\s!1\}, we let the first loop-iteration (if there is one) do that for us. This allows us to remove the precondition @{prop \s \ []\}, as the index bounds are respected even in that corner case.\ definition compute_\s :: "'a list \ nat list nres" where "compute_\s s = do { let \s=replicate (length s + 1) 0; \ \only the first 0 is needed\ let i=0; let j=1; (\s,_,_) \ WHILEIT (I_out_cb s) (\(\s,_,j). j < length \s) (\(\s,i,j). do { i \ WHILEIT (I_in_cb s j) (\i. i>0 \ s!(i-1) \ s!(j-1)) (\i. do { ASSERT (i-1 < length \s); let i=\s!(i-1); RETURN i }) i; let i=i+1; ASSERT (j < length \s); let \s=\s[j:=i]; let j=j+1; RETURN (\s,i,j) }) (\s,i,j); RETURN \s }" subsubsection\Correctness\ lemma take_length_ib[simp]: assumes "0 < j" "j \ length s" shows "take (length (intrinsic_border (take j s))) s = intrinsic_border (take j s)" proof - from assms have "prefix (intrinsic_border (take j s)) (take j s)" by (metis intrinsic_borderI' border_def list.size(3) neq0_conv not_less strict_border_def take_eq_Nil) also have "prefix (take j s) s" by (simp add: \j \ length s\ take_is_prefix) finally show ?thesis by (metis append_eq_conv_conj prefixE) qed lemma ib_singleton[simp]: "intrinsic_border [z] = []" by (metis intrinsic_border_less length_Cons length_greater_0_conv less_Suc0 list.size(3)) lemma border_butlast: "border xs ys \ border (butlast xs) (butlast ys)" apply (auto simp: border_def) apply (metis butlast_append prefixE prefix_order.eq_refl prefix_prefix prefixeq_butlast) apply (metis Sublist.suffix_def append.right_neutral butlast.simps(1) butlast_append) done corollary strict_border_butlast: "xs \ [] \ strict_border xs ys \ strict_border (butlast xs) (butlast ys)" unfolding strict_border_def by (simp add: border_butlast less_diff_conv) lemma border_take_lengths: "i \ length s \ border (take i s) (take j s) \ i \ j" using border_length_le by fastforce lemma border_step: "border xs ys \ border (xs@[ys!length xs]) (ys@[ys!length xs])" apply (auto simp: border_def suffix_def) using append_one_prefix prefixE apply fastforce using append_prefixD apply blast done corollary strict_border_step: "strict_border xs ys \ strict_border (xs@[ys!length xs]) (ys@[ys!length xs])" unfolding strict_border_def using border_step by auto lemma ib_butlast: "length w \ 2 \ length (intrinsic_border w) \ length (intrinsic_border (butlast w)) + 1" proof - assume "length w \ 2" then have "w \ []" by auto then have "strict_border (intrinsic_border w) w" by (fact intrinsic_borderI') with \2 \ length w\ have "strict_border (butlast (intrinsic_border w)) (butlast w)" by (metis One_nat_def border_bot.bot.not_eq_extremum butlast.simps(1) len_greater_imp_nonempty length_butlast lessI less_le_trans numerals(2) strict_border_butlast zero_less_diff) then have "length (butlast (intrinsic_border w)) \ length (intrinsic_border (butlast w))" using intrinsic_border_max by blast then show ?thesis by simp qed corollary \_Suc(*rm*): "Suc i \ length w \ \ w (Suc i) \ \ w i + 1" apply (cases i) apply (simp_all add: take_Suc0) by (metis One_nat_def Suc_eq_plus1 Suc_to_right butlast_take diff_is_0_eq ib_butlast length_take min.absorb2 nat.simps(3) not_less_eq_eq numerals(2)) lemma \_step_bound(*rm*): assumes "j \ length w" shows "\ w j \ \ w (j-1) + 1" using assms[THEN j_le_\_le] \_Suc assms by (metis One_nat_def Suc_pred le_SucI not_gr_zero trans_le_add2) lemma border_take_\: "border (take (\ s i - 1) s ) (take i s)" apply (cases i, simp_all) by (metis intrinsic_borderI' border_order.eq_iff border_order.less_imp_le border_positions nat.simps(3) nat_le_linear positions_border take_all take_eq_Nil take_length_ib zero_less_Suc) corollary \_strict_borderI: "y = \ s (i-1) \ strict_border (take (i-1) s) (take (j-1) s) \ strict_border (take (y-1) s) (take (j-1) s)" using border_order.less_le_not_le border_order.order.trans border_take_\ by blast corollary strict_border_take_\: "0 < i \ i \ length s \ strict_border (take (\ s i - 1) s ) (take i s)" by (meson border_order.less_le_not_le border_take_\ border_take_lengths j_le_\_le' leD) lemma \_is_max: "j \ length s \ strict_border b (take j s) \ \ s j \ length b + 1" by (metis \.elims add_le_cancel_right add_less_same_cancel2 border_length_r_less intrinsic_border_max length_take min_absorb2 not_add_less2) theorem skipping_ok: assumes j_bounds[simp]: "1 < j" "j \ length s" and mismatch: "s!(i-1) \ s!(j-1)" and greater_checked: "\ s j \ i + 1" and "strict_border (take (i-1) s) (take (j-1) s)" shows "\ s j \ \ s (i-1) + 1" proof (rule ccontr) assume "\\ s j \ \ s (i-1) + 1" then have i_bounds: "0 < i" "i \ length s" using greater_checked assms(5) take_Nil by fastforce+ then have i_less_j: "i < j" using assms(5) border_length_r_less nz_le_conv_less by auto from \\\ s j \ \ s (i-1) + 1\ greater_checked consider (tested) "\ s j = i + 1" \ \This contradicts @{thm mismatch}\ | (skipped) "\ s (i-1) + 1 < \ s j" "\ s j \ i" \ \This contradicts @{thm \_is_max[of "i-1" s]}\ by linarith then show False proof cases case tested then have "\ s j - 1 = i" by simp moreover note border_positions[OF border_take_\[of s j, unfolded this]] ultimately have "take j s ! (i-1) = s!(j-1)" using i_bounds i_less_j by simp with \i < j\ have "s!(i-1) = s!(j-1)" by (simp add: less_imp_diff_less) with mismatch show False.. next case skipped let ?border = "take (i-1) s" \ \This border of @{term \take (j-1) s\} could not be extended to a border of @{term \take j s\} due to the mismatch.\ let ?impossible = "take (\ s j - 2) s" \ \A strict border longer than @{term \intrinsic_border ?border\}, a contradiction.\ have "length (take j s) = j" by simp have "\ s j - 2 < i - 1" using skipped by linarith then have less_s: "\ s j - 2 < length s" "i - 1 < length s" using \i < j\ j_bounds(2) by linarith+ then have strict: "length ?impossible < length ?border" using \\ s j - 2 < i - 1\ by auto moreover { have "prefix ?impossible (take j s)" using prefix_length_prefix take_is_prefix by (metis (no_types, lifting) \length (take j s) = j\ j_bounds(2) diff_le_self j_le_\_le length_take less_s(1) min_simps(2) order_trans) moreover have "prefix ?border (take j s)" by (metis (no_types, lifting) \length (take j s) = j\ diff_le_self i_less_j le_trans length_take less_or_eq_imp_le less_s(2) min_simps(2) prefix_length_prefix take_is_prefix) ultimately have "prefix ?impossible ?border" using strict less_imp_le_nat prefix_length_prefix by blast } moreover { have "suffix (take (\ s j - 1) s) (take j s)" using border_take_\ by (auto simp: border_def) note suffix_butlast[OF this] then have "suffix ?impossible (take (j-1) s)" by (metis One_nat_def j_bounds(2) butlast_take diff_diff_left \_le len_greater_imp_nonempty less_or_eq_imp_le less_s(2) one_add_one) then have "suffix ?impossible (take (j-1) s)" "suffix ?border (take (j-1) s)" using assms(5) by auto from suffix_length_suffix[OF this strict[THEN less_imp_le]] have "suffix ?impossible ?border". } ultimately have "strict_border ?impossible ?border" unfolding strict_border_def[unfolded border_def] by blast note \_is_max[of "i-1" s, OF _ this] then have "length (take (\ s j - 2) s) + 1 \ \ s (i-1)" using less_imp_le_nat less_s(2) by blast then have "\ s j - 1 \ \ s (i-1)" by (simp add: less_s(1)) then have "\ s j \ \ s (i-1) + 1" using le_diff_conv by blast with skipped(1) show False by linarith qed qed lemma extend_border: assumes "j \ length s" assumes "s!(i-1) = s!(j-1)" assumes "strict_border (take (i-1) s) (take (j-1) s)" assumes "\ s j \ i + 1" shows "\ s j = i + 1" proof - from assms(3) have pos_in_range: "i - 1 < length s " "length (take (i-1) s) = i - 1" using border_length_r_less min_less_iff_conj by auto with strict_border_step[THEN iffD1, OF assms(3)] have "strict_border (take (i-1) s @ [s!(i-1)]) (take (j-1) s @ [s!(i-1)])" by (metis assms(3) border_length_r_less length_take min_less_iff_conj nth_take) with pos_in_range have "strict_border (take i s) (take (j-1) s @ [s!(i-1)])" by (metis Suc_eq_plus1 Suc_pred add.left_neutral border_bot.bot.not_eq_extremum border_order.less_asym neq0_conv take_0 take_Suc_conv_app_nth) then have "strict_border (take i s) (take (j-1) s @ [s!(j-1)])" by (simp only: \s!(i-1) = s!(j-1)\) then have "strict_border (take i s) (take j s)" by (metis One_nat_def Suc_pred assms(1,3) diff_le_self less_le_trans neq0_conv nz_le_conv_less strict_border_imp_nonempty take_Suc_conv_app_nth take_eq_Nil) with \_is_max[OF assms(1) this] have "\ s j \ i + 1" using Suc_leI by fastforce with \\ s j \ i + 1\ show ?thesis using le_antisym by presburger qed lemma compute_\s_correct: "compute_\s s \ compute_\s_SPEC s" unfolding compute_\s_SPEC_def compute_\s_def I_out_cb_def I_in_cb_def apply (simp, refine_vcg WHILEIT_rule[where R="measure (\(\s,i,j). length s + 1 - j)"] WHILEIT_rule[where R="measure id"] \ \@{term \i::nat\} decreases with every iteration.\ ) apply (vc_solve, fold One_nat_def) subgoal for b j by (rule strict_border_take_\, auto) subgoal by (metis Suc_eq_plus1 \_step_bound less_Suc_eq_le) subgoal by fastforce subgoal by (metis (no_types, lifting) One_nat_def Suc_lessD Suc_pred border_length_r_less \_strict_borderI length_take less_Suc_eq less_Suc_eq_le min.absorb2) subgoal for b j i by (metis (no_types, lifting) One_nat_def Suc_diff_1 Suc_eq_plus1 Suc_leI border_take_lengths less_Suc_eq_le less_antisym skipping_ok strict_border_def) subgoal by (metis Suc_diff_1 border_take_lengths j_le_\_le less_Suc_eq_le strict_border_def) subgoal for b j i jj by (metis Suc_eq_plus1 Suc_eq_plus1_left add.right_neutral extend_border \_eq_0_iff_j_eq_0 j_le_\_le le_zero_eq less_Suc_eq less_Suc_eq_le nth_list_update_eq nth_list_update_neq) subgoal by linarith done subsubsection\Index shift\ text\To avoid inefficiencies, we refine @{const compute_\s} to take @{term s} instead of @{term \butlast s\} (it still only uses @{term \butlast s\}).\ definition compute_butlast_\s :: "'a list \ nat list nres" where "compute_butlast_\s s = do { let \s=replicate (length s) 0; let i=0; let j=1; (\s,_,_) \ WHILEIT (I_out_cb (butlast s)) (\(b,i,j). j < length b) (\(\s,i,j). do { ASSERT (j < length \s); i \ WHILEIT (I_in_cb (butlast s) j) (\i. i>0 \ s!(i-1) \ s!(j-1)) (\i. do { ASSERT (i-1 < length \s); let i=\s!(i-1); RETURN i }) i; let i=i+1; ASSERT (j < length \s); let \s=\s[j:=i]; let j=j+1; RETURN (\s,i,j) }) (\s,i,j); RETURN \s }" lemma compute_\s_inner_bounds: assumes "I_out_cb s (\s,ix,j)" assumes "j < length \s" assumes "I_in_cb s j i" shows "i-1 < length s" "j-1 < length s" using assms by (auto simp: I_out_cb_def I_in_cb_def split: if_splits) lemma compute_butlast_\s_refine[refine]: assumes "(s,s') \ br butlast ((\) [])" shows "compute_butlast_\s s \ \ Id (compute_\s_SPEC s')" proof - have "compute_butlast_\s s \ \ Id (compute_\s s')" unfolding compute_butlast_\s_def compute_\s_def apply (refine_rcg) apply (refine_dref_type) using assms apply (vc_solve simp: in_br_conv) apply (metis Suc_pred length_greater_0_conv replicate_Suc) by (metis One_nat_def compute_\s_inner_bounds nth_butlast) also note compute_\s_correct finally show ?thesis by simp qed subsection\Conflation\ text\We replace @{const compute_\s_SPEC} with @{const compute_butlast_\s}\ definition "kmp2 s t \ do { ASSERT (s \ []); let i=0; let j=0; let pos=None; \s \ compute_butlast_\s s; (_,_,pos) \ WHILEIT (I_outer s t) (\(i,j,pos). i + length s \ length t \ pos=None) (\(i,j,pos). do { ASSERT (i + length s \ length t \ pos=None); (j,pos) \ WHILEIT (I_in_na s t i) (\(j,pos). t!(i+j) = s!j \ pos=None) (\(j,pos). do { let j=j+1; if j=length s then RETURN (j,Some i) else RETURN (j,None) }) (j,pos); if pos=None then do { ASSERT (j < length \s); let i = i + (j - \s!j + 1); let j = max 0 (\s!j - 1); \ \\max\ not necessary\ RETURN (i,j,None) } else RETURN (i,j,Some i) }) (i,j,pos); RETURN pos }" text\Using @{thm [source] compute_butlast_\s_refine} (it has attribute @{attribute refine}), the proof is trivial:\ lemma kmp2_refine: "kmp2 s t \ kmp1 s t" apply (rule refine_IdD) unfolding kmp2_def kmp1_def apply refine_rcg apply refine_dref_type apply (vc_solve simp: in_br_conv) done lemma kmp2_correct: "s \ [] \ kmp2 s t \ kmp_SPEC s t" proof - assume "s \ []" have "kmp2 s t \ kmp1 s t" by (fact kmp2_refine) also have "... \ kmp s t" by (fact kmp1_refine) also have "... \ kmp_SPEC s t" by (fact kmp_correct[OF \s \ []\]) finally show ?thesis. qed text\For convenience, we also remove the precondition:\ definition "kmp3 s t \ do { if s=[] then RETURN (Some 0) else kmp2 s t }" lemma kmp3_correct: "kmp3 s t \ kmp_SPEC s t" unfolding kmp3_def by (simp add: kmp2_correct) (simp add: kmp_SPEC_def) section \Refinement to Imperative/HOL\ lemma eq_id_param: "((=), (=)) \ Id \ Id \ Id" by simp lemmas in_bounds_aux = compute_\s_inner_bounds[of "butlast s" for s, simplified] sepref_definition compute_butlast_\s_impl is compute_butlast_\s :: "(arl_assn id_assn)\<^sup>k \\<^sub>a array_assn nat_assn" unfolding compute_butlast_\s_def supply in_bounds_aux[dest] supply eq_id_param[where 'a='a, sepref_import_param] apply (rewrite array_fold_custom_replicate) by sepref declare compute_butlast_\s_impl.refine[sepref_fr_rules] sepref_register compute_\s lemma kmp_inner_in_bound: assumes "i + length s \ length t" assumes "I_in_na s t i (j,None)" shows "i + j < length t" "j < length s" using assms by (auto simp: I_in_na_def) sepref_definition kmp_impl is "uncurry kmp3" :: "(arl_assn id_assn)\<^sup>k *\<^sub>a (arl_assn id_assn)\<^sup>k \\<^sub>a option_assn nat_assn" unfolding kmp3_def kmp2_def apply (simp only: max_0L) \ \Avoid the unneeded @{const max}\ apply (rewrite in "WHILEIT (I_in_na _ _ _) \" conj_commute) apply (rewrite in "WHILEIT (I_in_na _ _ _) \" short_circuit_conv) supply kmp_inner_in_bound[dest] supply option.splits[split] supply eq_id_param[where 'a='a, sepref_import_param] by sepref export_code kmp_impl in SML_imp module_name KMP lemma kmp3_correct': "(uncurry kmp3, uncurry kmp_SPEC) \ Id \\<^sub>r Id \\<^sub>f \Id\nres_rel" apply (intro frefI nres_relI; clarsimp) apply (fact kmp3_correct) done lemmas kmp_impl_correct' = kmp_impl.refine[FCOMP kmp3_correct'] subsection \Overall Correctness Theorem\ text \The following theorem relates the final Imperative HOL algorithm to its specification, using, beyond basic HOL concepts \<^item> Hoare triples for Imperative/HOL, provided by the Separation Logic Framework for Imperative/HOL (Theory @{theory Separation_Logic_Imperative_HOL.Sep_Main}); \<^item> The assertion @{const arl_assn} to specify array-lists, which we use to represent the input strings of the algorithm; \<^item> The @{const sublist_at} function that we defined in section \ref{sec:spec}. \ theorem kmp_impl_correct: "< arl_assn id_assn s si * arl_assn id_assn t ti > kmp_impl si ti <\r. arl_assn id_assn s si * arl_assn id_assn t ti * \( case r of None \ \i. sublist_at s t i | Some i \ sublist_at s t i \ (\ii sublist_at s t ii) )>\<^sub>t" by (sep_auto simp: pure_def kmp_SPEC_def split: option.split heap: kmp_impl_correct'[THEN hfrefD, THEN hn_refineD, of "(s,t)" "(si,ti)", simplified]) definition "kmp_string_impl \ kmp_impl :: (char array \ nat) \ _" section \Tests of Generated ML-Code\ ML_val \ fun str2arl s = (Array.fromList (@{code String.explode} s), @{code nat_of_integer} (String.size s)) fun kmp s t = map_option @{code integer_of_nat} (@{code kmp_string_impl} (str2arl s) (str2arl t) ()) val test1 = kmp "anas" "bananas" val test2 = kmp "" "bananas" val test3 = kmp "hide_fact" (File.read @{file \~~/src/HOL/Main.thy\}) val test4 = kmp "sorry" (File.read @{file \~~/src/HOL/HOL.thy\}) \ end diff --git a/thys/LLL_Basis_Reduction/LLL_Impl.thy b/thys/LLL_Basis_Reduction/LLL_Impl.thy --- a/thys/LLL_Basis_Reduction/LLL_Impl.thy +++ b/thys/LLL_Basis_Reduction/LLL_Impl.thy @@ -1,1143 +1,1143 @@ (* Authors: René Thiemann License: BSD *) subsection \Integer LLL Implementation which Stores Multiples of the $\mu$-Values\ text \In this part we aim to update the integer values $d\,(j+1) * \mu_{i,j}$ as well as the Gramian determinants $d\,i$. \ theory LLL_Impl imports LLL List_Representation Gram_Schmidt_Int begin subsubsection \Updates of the integer values for Swap, Add, etc.\ text \We provide equations how to implement the LLL-algorithm by storing the integer values $d\, (j+1) * \mu_{i,j}$ and all $d\ i$ in addition to the vectors in $f$. Moreover, we show how to check condition like the one on norms via the integer values.\ definition round_num_denom :: "int \ int \ int" where "round_num_denom n d = ((2 * n + d) div (2 * d))" lemma round_num_denom: "round_num_denom num denom = round (of_int num / rat_of_int denom)" proof (cases "denom = 0") case False have "denom \ 0 \ ?thesis" unfolding round_def round_num_denom_def unfolding floor_divide_of_int_eq[where ?'a = rat, symmetric] by (rule arg_cong[of _ _ floor], simp add: add_divide_distrib) with False show ?thesis by auto next case True show ?thesis unfolding True round_num_denom_def by auto qed context fs_int_indpt begin lemma round_num_denom_d\_d: assumes j: "j \ i" and i: "i < m" shows "round_num_denom (d\ i j) (d fs (Suc j)) = round (gs.\ i j)" proof - from j i have sj: "Suc j \ m" by auto show ?thesis unfolding round_num_denom by (rule arg_cong[of _ _ round], subst d\[OF _ i], insert j i fs_int_d_pos[OF sj], auto) qed lemma d_sq_norm_comparison: assumes quot: "quotient_of \ = (num,denom)" and i: "i < m" and i0: "i \ 0" shows "(d fs i * d fs i * denom \ num * d fs (i - 1) * d fs (Suc i)) = (sq_norm (gs.gso (i - 1)) \ \ * sq_norm (gs.gso i))" proof - let ?r = "rat_of_int" let ?x = "sq_norm (gs.gso (i - 1))" let ?y = "\ * sq_norm (gs.gso i)" from i have le: "i - 1 \ m" " i \ m" "Suc i \ m" by auto note pos = fs_int_d_pos[OF le(1)] fs_int_d_pos[OF le(2)] quotient_of_denom_pos[OF quot] have "(d fs i * d fs i * denom \ num * d fs (i - 1) * d fs (Suc i)) = (?r (d fs i * d fs i * denom) \ ?r (num * d fs (i - 1) * d fs (Suc i)))" (is "?cond = _") by presburger also have "\ = (?r (d fs i) * ?r (d fs i) * ?r denom \ ?r num * ?r (d fs (i - 1)) * ?r (d fs (Suc i)))" by simp also have "\ = (?r (d fs i) * ?r (d fs i) \ \ * ?r (d fs (i - 1)) * ?r (d fs (Suc i)))" using pos unfolding quotient_of_div[OF quot] by (auto simp: field_simps) also have "\ = (?r (d fs i) / ?r (d fs (i - 1)) \ \ * (?r (d fs (Suc i)) / ?r (d fs i)))" using pos by (auto simp: field_simps) also have "?r (d fs i) / ?r (d fs (i - 1)) = ?x" using fs_int_d_Suc[of "i - 1"] pos i i0 by (auto simp: field_simps) also have "\ * (?r (d fs (Suc i)) / ?r (d fs i)) = ?y" using fs_int_d_Suc[OF i] pos i i0 by (auto simp: field_simps) finally show "?cond = (?x \ ?y)" . qed end context LLL begin lemma d_d\_add_row: assumes Linv: "LLL_invariant True i fs" and i: "i < m" and j: "j < i" and fs': "fs' = fs[ i := fs ! i - c \\<^sub>v fs ! j]" shows (* d-updates: none *) "\ ii. ii \ m \ d fs' ii = d fs ii" (* d\-updates: *) "\ i' j'. i' < m \ j' < i' \ d\ fs' i' j' = ( if i' = i \ j' < j then d\ fs i' j' - c * d\ fs j j' else if i' = i \ j' = j then d\ fs i' j' - c * d fs (Suc j) else d\ fs i' j')" (is "\ i' j'. _ \ _ \ _ = ?new_mu i' j'") proof - interpret fs: fs_int' n m fs_init \ True i fs by standard (use Linv in auto) note add = basis_reduction_add_row_main[OF Linv i j fs'] interpret fs': fs_int' n m fs_init \ True i fs' by standard (use add in auto) show d: "\ ii. ii \ m \ d fs' ii = d fs ii" by fact fix i' j' assume i': "i' < m" and j': "j' < i'" hence j'm: "j' < m" and j'': "j' \ i'" by auto note updates = add(5)[OF i' j'm] show "d\ fs' i' j' = ?new_mu i' j'" proof (cases "i' = i") case False thus ?thesis using d i' j' unfolding d\_def updates by auto next case True have id': "d fs' (Suc j') = d fs (Suc j')" by (rule d, insert i' j', auto) note fs'.d\[] have *: "rat_of_int (d\ fs' i' j') = rat_of_int (d fs' (Suc j')) * fs'.gs.\ i' j'" unfolding d\_def d_def apply(rule fs'.d\[unfolded fs'.d\_def fs'.d_def]) using j' i' LLL_invD[OF add(1)] by (auto) have **: "rat_of_int (d\ fs i' j') = rat_of_int (d fs (Suc j')) * fs.gs.\ i' j'" unfolding d\_def d_def apply(rule fs.d\[unfolded fs.d\_def fs.d_def]) using j' i' LLL_invD[OF Linv] by (auto) have ***: "rat_of_int (d\ fs j j') = rat_of_int (d fs (Suc j')) * fs.gs.\ j j'" if "j' < j" unfolding d\_def d_def apply(rule fs.d\[unfolded fs.d\_def fs.d_def]) using that j i LLL_invD[OF Linv] by (auto) show ?thesis apply(intro int_via_rat_eqI) apply(unfold if_distrib[of rat_of_int] of_int_diff of_int_mult ** * updates id' ring_distribs) apply(insert True i' j' i j) by(auto simp: fs.gs.\.simps algebra_simps ***) qed qed end context LLL_with_assms begin lemma d_d\_swap: assumes inv: "LLL_invariant False k fs" and k: "k < m" and k0: "k \ 0" and norm_ineq: "sq_norm (gso fs (k - 1)) > \ * sq_norm (gso fs k)" and fs'_def: "fs' = fs[k := fs ! (k - 1), k - 1 := fs ! k]" shows (* d-updates *) "\ i. i \ m \ d fs' i = ( if i = k then (d fs (Suc k) * d fs (k - 1) + d\ fs k (k - 1) * d\ fs k (k - 1)) div d fs k else d fs i)" and (* d\-updates *) "\ i j. i < m \ j < i \ d\ fs' i j = ( if i = k - 1 then d\ fs k j else if i = k \ j \ k - 1 then d\ fs (k - 1) j else if i > k \ j = k then (d fs (Suc k) * d\ fs i (k - 1) - d\ fs k (k - 1) * d\ fs i j) div d fs k else if i > k \ j = k - 1 then (d\ fs k (k - 1) * d\ fs i j + d\ fs i k * d fs (k - 1)) div d fs k else d\ fs i j)" (is "\ i j. _ \ _ \ _ = ?new_mu i j") proof - note swap = basis_reduction_swap_main[OF inv k k0 norm_ineq fs'_def] from k k0 have kk: "k - 1 < k" and le_m: "k - 1 \ m" "k \ m" "Suc k \ m" by auto from LLL_invD[OF inv] have len: "length fs = m" by auto interpret fs: fs_int' n m fs_init \ False k fs by standard (use inv in auto) interpret fs': fs_int' n m fs_init \ False "k - 1" fs' by standard (use swap(1) in auto) let ?r = rat_of_int let ?n = "\ i. sq_norm (gso fs i)" let ?n' = "\ i. sq_norm (gso fs' i)" let ?dn = "\ i. ?r (d fs i * d fs i) * ?n i" let ?dn' = "\ i. ?r (d fs' i * d fs' i) * ?n' i" let ?dmu = "\ i j. ?r (d fs (Suc j)) * \ fs i j" let ?dmu' = "\ i j. ?r (d fs' (Suc j)) * \ fs' i j" note dmu = fs.d\ note dmu' = fs'.d\ note inv' = LLL_invD[OF inv] have nim1: "?n k + square_rat (\ fs k (k - 1)) * ?n (k - 1) = ?n' (k - 1)" by (subst swap(4), insert k, auto) have ni: "?n k * (?n (k - 1) / ?n' (k - 1)) = ?n' k" by (subst swap(4)[of k], insert k k0, auto) have mu': "\ fs k (k - 1) * (?n (k - 1) / ?n' (k - 1)) = \ fs' k (k - 1)" by (subst swap(5), insert k k0, auto) have fi: "fs ! (k - 1) = fs' ! k" "fs ! k = fs' ! (k - 1)" unfolding fs'_def using inv'(6) k k0 by auto let ?d'i = "(d fs (Suc k) * d fs (k - 1) + d\ fs k (k - 1) * d\ fs k (k - 1)) div (d fs k)" have rat': "i < m \ j < i \ ?r (d\ fs' i j) = ?dmu' i j" for i j using dmu'[of j i] LLL_invD[OF swap(1)] unfolding d\_def fs'.d\_def d_def fs'.d_def by auto have rat: "i < m \ j < i \ ?r (d\ fs i j) = ?dmu i j" for i j using dmu[of j i] LLL_invD[OF inv] unfolding d\_def fs.d\_def d_def fs.d_def by auto from k k0 have sim1: "Suc (k - 1) = k" and km1: "k - 1 < m" by auto from LLL_d_Suc[OF inv km1, unfolded sim1] have dn_km1: "?dn (k - 1) = ?r (d fs k) * ?r (d fs (k - 1))" by simp note pos = Gramian_determinant[OF inv le_refl] from pos(2) have "?r (gs.Gramian_determinant fs m) \ 0" by auto from this[unfolded pos(1)] have nzero: "i < m \ ?n i \ 0" for i by auto note pos = Gramian_determinant[OF swap(1) le_refl] from pos(2) have "?r (gs.Gramian_determinant fs' m) \ 0" by auto from this[unfolded pos(1)] have nzero': "i < m \ ?n' i \ 0" for i by auto have dzero: "i \ m \ d fs i \ 0" for i using LLL_d_pos[OF inv, of i] by auto have dzero': "i \ m \ d fs' i \ 0" for i using LLL_d_pos[OF swap(1), of i] by auto { define start where "start = ?dmu' k (k - 1)" have "start = (?n' (k - 1) / ?n (k - 1) * ?r (d fs k)) * \ fs' k (k - 1)" using start_def swap(6)[of k] k k0 by simp also have "\ fs' k (k - 1) = \ fs k (k - 1) * (?n (k - 1) / ?n' (k - 1))" using mu' by simp also have "(?n' (k - 1) / ?n (k - 1) * ?r (d fs k)) * \ = ?r (d fs k) * \ fs k (k - 1)" using nzero[OF km1] nzero'[OF km1] by simp also have "\ = ?dmu k (k - 1)" using k0 by simp finally have "?dmu' k (k - 1) = ?dmu k (k - 1)" unfolding start_def . } note dmu_i_im1 = this { (* d updates *) fix j assume j: "j \ m" define start where "start = d fs' j" { assume jj: "j \ k" have "?r start = ?r (d fs' j)" unfolding start_def .. also have "?r (d fs' j) = ?r (d fs j)" by (subst swap(6), insert j jj, auto) finally have "start = d fs j" by simp } note d_j = this { assume jj: "j = k" have "?r start = ?r (d fs' k)" unfolding start_def unfolding jj by simp also have "\ = ?n' (k - 1) / ?n (k - 1) * ?r (d fs k)" by (subst swap(6), insert k, auto) also have "?n' (k - 1) = (?r (d fs k) / ?r (d fs k)) * (?r (d fs k) / ?r (d fs k)) * (?n k + \ fs k (k - 1) * \ fs k (k - 1) * ?n (k - 1))" by (subst swap(4)[OF km1], insert dzero[of k], insert k, simp) also have "?n (k - 1) = ?r (d fs k) / ?r (d fs (k - 1))" unfolding LLL_d_Suc[OF inv km1, unfolded sim1] using dzero[of "k - 1"] k k0 by simp finally have "?r start = ((?r (d fs k) * ?n k) * ?r (d fs (k - 1)) + ?dmu k (k - 1) * ?dmu k (k - 1)) / (?r (d fs k))" using k k0 dzero[of k] dzero[of "k - 1"] by (simp add: ring_distribs) also have "?r (d fs k) * ?n k = ?r (d fs (Suc k))" unfolding LLL_d_Suc[OF inv k] by simp also have "?dmu k (k - 1) = ?r (d\ fs k (k - 1))" by (subst rat, insert k k0, auto) finally have "?r start = (?r (d fs (Suc k) * d fs (k - 1) + d\ fs k (k - 1) * d\ fs k (k - 1))) / (?r (d fs k))" by simp from division_to_div[OF this] have "start = ?d'i" . } note d_i = this from d_j d_i show "d fs' j = (if j = k then ?d'i else d fs j)" unfolding start_def by auto } have "length fs' = m" using fs'_def inv'(6) by auto { fix i j assume i: "i < m" and j: "j < i" from j i have sj: "Suc j \ m" by auto note swaps = swap(5)[OF i j] swap(6)[OF sj] show "d\ fs' i j = ?new_mu i j" proof (cases "i < k - 1") case small: True hence id: "?new_mu i j = d\ fs i j" by auto show ?thesis using swaps small i j k k0 by (auto simp: d\_def) next case False from j i have sj: "Suc j \ m" by auto let ?start = "d\ fs' i j" define start where "start = ?start" note rat'[OF i j] note rat_i_j = rat[OF i j] from False consider (i_k) "i = k" "j = k - 1" | (i_small) "i = k" "j \ k - 1" | (i_km1) "i = k - 1" | (i_large) "i > k" by linarith thus ?thesis proof cases case *: i_small show ?thesis unfolding swaps d\_def using * i k k0 by auto next case *: i_k show ?thesis using dmu_i_im1 rat_i_j * k0 by (auto simp: d\_def) next case *: i_km1 show ?thesis unfolding swaps d\_def using * i j k k0 by auto next case *: i_large consider (jj) "j \ k - 1" "j \ k" | (ji) "j = k" | (jim1) "j = k - 1" by linarith thus ?thesis proof cases case jj show ?thesis unfolding swaps d\_def using * i j jj k k0 by auto next case ji have "?r start = ?dmu' i j" unfolding start_def by fact also have "?r (d fs' (Suc j)) = ?r (d fs (Suc k))" unfolding swaps unfolding ji by simp also have "\ fs' i j = \ fs i (k - 1) - \ fs k (k - 1) * \ fs i k" unfolding swaps unfolding ji using k0 * by auto also have "?r (d fs (Suc k)) * \ = ?r (d fs (Suc k)) * ?r (d fs k) / ?r (d fs k) * \" using dzero[of k] k by auto also have "\ = (?r (d fs (Suc k)) * ?dmu i (k - 1) - ?dmu k (k - 1) * ?dmu i k) / ?r (d fs k)" using k0 by (simp add: field_simps) also have "\ = (?r (d fs (Suc k)) * ?r (d\ fs i (k - 1)) - ?r (d\ fs k (k - 1)) * ?r (d\ fs i k)) / ?r (d fs k)" by (subst (1 2 3) rat, insert k k0 i *, auto) also have "\ = ?r (d fs (Suc k) * d\ fs i (k - 1) - d\ fs k (k - 1) * d\ fs i k) / ?r (d fs k)" (is "_ = of_int ?x / _") by simp finally have "?r start = ?r ?x / ?r (d fs k)" . from division_to_div[OF this] have id: "?start = (d fs (Suc k) * d\ fs i (k - 1) - d\ fs k (k - 1) * d\ fs i j) div d fs k" unfolding start_def ji . show ?thesis unfolding id using * ji by simp next case jim1 hence id'': "(j = k - 1) = True" "(j = k) = False" using k0 by auto have "?r (start) = ?dmu' i j" unfolding start_def by fact also have "\ fs' i j = \ fs i (k - 1) * \ fs' k (k - 1) + \ fs i k * ?n k / ?n' (k - 1)" (is "_ = ?x1 + ?x2") unfolding swaps unfolding jim1 using k0 * by auto also have "?r (d fs' (Suc j)) * (?x1 + ?x2) = ?r (d fs' (Suc j)) * ?x1 + ?r (d fs' (Suc j)) * ?x2" by (simp add: ring_distribs) also have "?r (d fs' (Suc j)) * ?x1 = ?dmu' k (k - 1) * (?r (d fs k) * \ fs i (k - 1)) / ?r (d fs k)" unfolding jim1 using k0 dzero[of k] k by simp also have "?dmu' k (k - 1) = ?dmu k (k - 1)" by fact also have "?r (d fs k) * \ fs i (k - 1) = ?dmu i (k - 1)" using k0 by simp also have "?r (d fs' (Suc j)) = ?n' (k - 1) * ?r (d fs k) / ?n (k - 1)" unfolding swaps unfolding jim1 using k k0 by simp also have "\ * ?x2 = (?n k * ?r (d fs k)) / ?n (k - 1) * \ fs i k" using k k0 nzero'[of "k - 1"] by simp also have "?n k * ?r (d fs k) = ?r (d fs (Suc k))" unfolding LLL_d_Suc[OF inv k] .. also have "?r (d fs (Suc k)) / ?n (k - 1) * \ fs i k = ?dmu i k / ?n (k - 1)" by simp also have "\ = ?dmu i k * ?r (d fs (k - 1) * d fs (k - 1)) / ?dn (k - 1)" using dzero[of "k - 1"] k by simp finally have "?r start = (?dmu k (k - 1) * ?dmu i j * ?dn (k - 1) + ?dmu i k * (?r (d fs (k - 1) * d fs (k - 1) * d fs k))) / (?r (d fs k) * ?dn (k - 1))" unfolding add_divide_distrib of_int_mult jim1 using dzero[of "k - 1"] nzero[of "k - 1"] k dzero[of k] by auto also have "\ = (?r (d\ fs k (k - 1)) * ?r (d\ fs i j) * (?r (d fs k) * ?r (d fs (k - 1))) + ?r (d\ fs i k) * (?r (d fs (k - 1) * d fs (k - 1) * d fs k))) / (?r (d fs k) * (?r (d fs k) * ?r (d fs (k - 1))))" unfolding dn_km1 by (subst (1 2 3) rat, insert k k0 i * j, auto) also have "\ = (?r (d\ fs k (k - 1)) * ?r (d\ fs i j) + ?r (d\ fs i k) * ?r (d fs (k - 1))) / ?r (d fs k)" unfolding of_int_mult using dzero[of k] dzero[of "k - 1"] k k0 by (auto simp: field_simps) also have "\ = ?r (d\ fs k (k - 1) * d\ fs i j + d\ fs i k * d fs (k - 1)) / ?r (d fs k)" (is "_ = of_int ?x / _" ) by simp finally have "?r start = ?r ?x / ?r (d fs k)" . from division_to_div[OF this] have id: "?start = (d\ fs k (k - 1) * d\ fs i j + d\ fs i k * d fs (k - 1)) div (d fs k)" unfolding start_def . show ?thesis unfolding id using * jim1 k0 by auto qed qed qed } qed end subsubsection \Implementation of LLL via Integer Operations and Arrays\ hide_fact (open) Word.inc_i type_synonym LLL_dmu_d_state = "int vec list_repr \ int iarray iarray \ int iarray" fun fi_state :: "LLL_dmu_d_state \ int vec" where "fi_state (f,mu,d) = get_nth_i f" fun fim1_state :: "LLL_dmu_d_state \ int vec" where "fim1_state (f,mu,d) = get_nth_im1 f" fun d_state :: "LLL_dmu_d_state \ nat \ int" where "d_state (f,mu,d) i = d !! i" fun fs_state :: "LLL_dmu_d_state \ int vec list" where "fs_state (f,mu,d) = of_list_repr f" fun upd_fi_mu_state :: "LLL_dmu_d_state \ nat \ int vec \ int iarray \ LLL_dmu_d_state" where "upd_fi_mu_state (f,mu,d) i fi mu_i = (update_i f fi, iarray_update mu i mu_i,d)" fun small_fs_state :: "LLL_dmu_d_state \ int vec list" where "small_fs_state (f,_) = fst f" fun dmu_ij_state :: "LLL_dmu_d_state \ nat \ nat \ int" where "dmu_ij_state (f,mu,_) i j = mu !! i !! j" fun inc_state :: "LLL_dmu_d_state \ LLL_dmu_d_state" where "inc_state (f,mu,d) = (inc_i f, mu, d)" fun basis_reduction_add_rows_loop where "basis_reduction_add_rows_loop n state i j [] = state" | "basis_reduction_add_rows_loop n state i sj (fj # fjs) = ( let fi = fi_state state; dsj = d_state state sj; j = sj - 1; c = round_num_denom (dmu_ij_state state i j) dsj; state' = (if c = 0 then state else upd_fi_mu_state state i (vec n (\ i. fi $ i - c * fj $ i)) (IArray.of_fun (\ jj. let mu = dmu_ij_state state i jj in if jj < j then mu - c * dmu_ij_state state j jj else if jj = j then mu - dsj * c else mu) i)) in basis_reduction_add_rows_loop n state' i j fjs)" text \More efficient code which breaks abstraction of state.\ lemma basis_reduction_add_rows_loop_code: "basis_reduction_add_rows_loop n state i sj (fj # fjs) = ( case state of ((f1,f2),mus,ds) \ let fi = hd f2; j = sj - 1; dsj = ds !! sj; mui = mus !! i; c = round_num_denom (mui !! j) dsj in (if c = 0 then basis_reduction_add_rows_loop n state i j fjs else let muj = mus !! j in basis_reduction_add_rows_loop n ((f1, vec n (\ i. fi $ i - c * fj $ i) # tl f2), iarray_update mus i (IArray.of_fun (\ jj. let mu = mui !! jj in if jj < j then mu - c * muj !! jj else if jj = j then mu - dsj * c else mu) i), ds) i j fjs))" proof - obtain f1 f2 mus ds where state: "state = ((f1,f2),mus, ds)" by (cases state, auto) show ?thesis unfolding basis_reduction_add_rows_loop.simps Let_def state split dmu_ij_state.simps fi_state.simps get_nth_i_def update_i_def upd_fi_mu_state.simps d_state.simps by simp qed lemmas basis_reduction_add_rows_loop_code_equations = basis_reduction_add_rows_loop.simps(1) basis_reduction_add_rows_loop_code declare basis_reduction_add_rows_loop_code_equations[code] definition basis_reduction_add_rows where "basis_reduction_add_rows n upw i state = (if upw then basis_reduction_add_rows_loop n state i i (small_fs_state state) else state)" context fixes \ :: rat and n m :: nat and fs_init :: "int vec list" begin definition swap_mu :: "int iarray iarray \ nat \ int \ int \ int \ int \ int iarray iarray" where "swap_mu dmu i dmu_i_im1 dim1 di dsi = (let im1 = i - 1 in IArray.of_fun (\ ii. if ii < im1 then dmu !! ii else if ii > i then let dmu_ii = dmu !! ii in IArray.of_fun (\ j. let dmu_ii_j = dmu_ii !! j in if j = i then (dsi * dmu_ii !! im1 - dmu_i_im1 * dmu_ii_j) div di else if j = im1 then (dmu_i_im1 * dmu_ii_j + dmu_ii !! i * dim1) div di else dmu_ii_j) ii else if ii = i then let mu_im1 = dmu !! im1 in IArray.of_fun (\ j. if j = im1 then dmu_i_im1 else mu_im1 !! j) ii else IArray.of_fun (\ j. dmu !! i !! j) ii) \ \ii = i - 1\ m)" definition basis_reduction_swap where "basis_reduction_swap i state = (let di = d_state state i; dsi = d_state state (Suc i); dim1 = d_state state (i - 1); fi = fi_state state; fim1 = fim1_state state; dmu_i_im1 = dmu_ij_state state i (i - 1); fi' = fim1; fim1' = fi in (case state of (f,dmus,djs) \ (False, i - 1, (dec_i (update_im1 (update_i f fi') fim1'), swap_mu dmus i dmu_i_im1 dim1 di dsi, iarray_update djs i ((dsi * dim1 + dmu_i_im1 * dmu_i_im1) div di)))))" text \More efficient code which breaks abstraction of state.\ lemma basis_reduction_swap_code[code]: "basis_reduction_swap i ((f1,f2), dmus, ds) = (let di = ds !! i; dsi = ds !! (Suc i); im1 = i - 1; dim1 = ds !! im1; fi = hd f2; fim1 = hd f1; dmu_i_im1 = dmus !! i !! im1; fi' = fim1; fim1' = fi in (False, im1, ((tl f1,fim1' # fi' # tl f2), swap_mu dmus i dmu_i_im1 dim1 di dsi, iarray_update ds i ((dsi * dim1 + dmu_i_im1 * dmu_i_im1) div di))))" proof - show ?thesis unfolding basis_reduction_swap_def split Let_def fi_state.simps fim1_state.simps d_state.simps get_nth_im1_def get_nth_i_def update_i_def update_im1_def dec_i_def by simp qed definition basis_reduction_step where "basis_reduction_step upw i state = (if i = 0 then (True, Suc i, inc_state state) else let state' = basis_reduction_add_rows n upw i state; di = d_state state' i; dsi = d_state state' (Suc i); dim1 = d_state state' (i - 1); (num,denom) = quotient_of \ in if di * di * denom \ num * dim1 * dsi then (True, Suc i, inc_state state') else basis_reduction_swap i state')" partial_function (tailrec) basis_reduction_main where [code]: "basis_reduction_main upw i state = (if i < m then case basis_reduction_step upw i state of (upw',i',state') \ basis_reduction_main upw' i' state' else state)" definition "initial_state = (let dmus = d\_impl fs_init; ds = IArray.of_fun (\ i. if i = 0 then 1 else let i1 = i - 1 in dmus !! i1 !! i1) (Suc m); dmus' = IArray.of_fun (\ i. let row_i = dmus !! i in IArray.of_fun (\ j. row_i !! j) i) m in (([], fs_init), dmus', ds) :: LLL_dmu_d_state)" end definition "basis_reduction \ n fs = (let m = length fs in basis_reduction_main \ n m True 0 (initial_state m fs))" definition "reduce_basis \ fs = (case fs of Nil \ fs | Cons f _ \ fs_state (basis_reduction \ (dim_vec f) fs))" definition "short_vector \ fs = hd (reduce_basis \ fs)" lemma map_rev_Suc: "map f (rev [0.. int vec list \ bool" where "mu_repr mu fs = (mu = IArray.of_fun (\ i. IArray.of_fun (d\ fs i) i) m)" definition d_repr :: "int iarray \ int vec list \ bool" where "d_repr ds fs = (ds = IArray.of_fun (d fs) (Suc m))" fun LLL_impl_inv :: "LLL_dmu_d_state \ nat \ int vec list \ bool" where "LLL_impl_inv (f,mu,ds) i fs = (list_repr i f (map (\ j. fs ! j) [0.. d_repr ds fs \ mu_repr mu fs)" context fixes state i fs upw f mu ds assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant upw i fs" and state: "state = (f,mu,ds)" begin lemma to_list_repr: "list_repr i f (map ((!) fs) [0.. fs ii j" unfolding to_mu_repr[unfolded mu_repr_def] state using ii j by auto lemma fi_state: "i < m \ fi_state state = fs ! i" using get_nth_i[OF to_list_repr(1)] unfolding state by auto lemma fim1_state: "i < m \ i \ 0 \ fim1_state state = fs ! (i - 1)" using get_nth_im1[OF to_list_repr(1)] unfolding state by auto lemma d_state: "ii \ m \ d_state state ii = d fs ii" using to_d_repr[unfolded d_repr_def] state unfolding state by (auto simp: nth_append) lemma fs_state: "length fs = m \ fs_state state = fs" using of_list_repr[OF to_list_repr(1)] unfolding state by (auto simp: o_def intro!: nth_equalityI) lemma LLL_state_inc_state: assumes i: "i < m" shows "LLL_impl_inv (inc_state state) (Suc i) fs" "fs_state (inc_state state) = fs_state state" proof - from LLL_invD[OF inv] have len: "length fs = m" by auto note inc = inc_i[OF to_list_repr(1)] from inc i impl show "LLL_impl_inv (inc_state state) (Suc i) fs" unfolding state by auto from of_list_repr[OF inc(1)] of_list_repr[OF to_list_repr(1)] i show "fs_state (inc_state state) = fs_state state" unfolding state by auto qed end end context LLL_with_assms begin lemma basis_reduction_add_rows_loop_impl: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant True i fs" and mu_small: "\_small_row i fs j" and res: "LLL_Impl.basis_reduction_add_rows_loop n state i j (map ((!) fs) (rev [0 ..< j])) = state'" (is "LLL_Impl.basis_reduction_add_rows_loop n state i j (?mapf fs j) = _") and j: "j \ i" and i: "i < m" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i fs'" "basis_reduction_add_rows_loop i fs j = fs'" proof (atomize(full), insert assms(1-6), induct j arbitrary: fs state) case (0 fs state) from LLL_invD[OF 0(2)] have len: "length fs = m" by auto from fs_state[OF 0(1-2) _ len] have "fs_state state = fs" by (cases state, auto) thus ?case using 0 i fs' by auto next case (Suc j fs state) hence j: "j < i" and jj: "j \ i" and id: "(j < i) = True" by auto obtain f mu ds where state: "state = (f,mu,ds)" by (cases state, auto) note Linv = Suc(3) note inv = LLL_invD[OF Linv] note impl = Suc(2) from fi_state[OF impl Linv state i] have fi: "fi_state state = fs ! i" by auto have id: "Suc j - 1 = j" by simp note mu = dmu_ij_state[OF impl Linv state j i] let ?c = "round (\ fs i j)" interpret fs: fs_int' n m fs_init \ True i fs by standard (use Linv in auto) have floor: "round_num_denom (d\ fs i j) (d fs (Suc j)) = round (fs.gs.\ i j)" using jj i inv unfolding d\_def d_def by (intro fs.round_num_denom_d\_d[unfolded fs.d\_def fs.d_def]) auto from LLL_d_pos[OF Linv] j i have dj: "d fs (Suc j) > 0" by auto note updates = d_d\_add_row[OF Linv i j refl] note d_state = d_state[OF impl Linv state] from d_state[of "Suc j"] j i have djs: "d_state state (Suc j) = d fs (Suc j)" by auto note res = Suc(5)[unfolded floor map_rev_Suc djs append.simps LLL_Impl.basis_reduction_add_rows_loop.simps fi Let_def mu id int_times_rat_def] show ?case proof (cases "?c = 0") case True from res[unfolded True] have res: "LLL_Impl.basis_reduction_add_rows_loop n state i j (?mapf fs j) = state'" by simp note step = Linv basis_reduction_add_row_main_0[OF Linv i j True Suc(4)] show ?thesis using Suc(1)[OF impl step(1-2) res _ i] j True by auto next case False hence id: "(?c = 0) = False" by auto from i j have jm: "j < m" by auto have idd: "vec n (\ia. fs ! i $ ia - ?c * fs ! j $ ia) = fs ! i - ?c \\<^sub>v fs ! j" by (intro eq_vecI, insert inv(4)[OF i] inv(4)[OF jm], auto) define fi' where "fi' = fs ! i - ?c \\<^sub>v fs ! j" obtain fs'' where fs'': "fs[i := fs ! i - ?c \\<^sub>v fs ! j] = fs''" by auto note step = basis_reduction_add_row_main[OF Linv i j fs''[symmetric]] note updates = updates[where c = ?c, unfolded fs''] have map_id_f: "?mapf fs j = ?mapf fs'' j" by (rule nth_equalityI, insert j i, auto simp: rev_nth fs''[symmetric]) have nth_id: "[0.. fs'' i) i" (is "_ = ?mu'i") proof (rule iarray_cong', goal_cases) case (1 jj) from 1 j i have jm: "j < m" by auto show ?case unfolding dmu_ij_state[OF impl Linv state 1 i] using dmu_ij_state[OF impl Linv state _ jm] by (subst updates(2)[OF i 1], auto) qed { fix ii assume ii: "ii < m" "ii \ i" hence "(IArray.of_fun (\i. IArray.of_fun (d\ fs i) i) m) !! ii = IArray.of_fun (d\ fs ii) ii" by auto also have "\ = IArray.of_fun (d\ fs'' ii) ii" proof (rule iarray_cong', goal_cases) case (1 j) with ii have j: "Suc j \ m" by auto show ?case unfolding updates(2)[OF ii(1) 1] using ii by auto qed finally have "(IArray.of_fun (\i. IArray.of_fun (d\ fs i) i) m) !! ii = IArray.of_fun (d\ fs'' ii) ii" by auto } note ii = this let ?mu'' = "iarray_update mu i (IArray.of_fun (d\ fs'' i) i)" have new_array: "?mu'' = IArray.of_fun (\ i. IArray.of_fun (d\ fs'' i) i) m" unfolding iarray_update_of_fun to_mu_repr[OF impl Linv state, unfolded mu_repr_def] by (rule iarray_cong', insert ii, auto) have d': "(map (?d fs) (rev [0..\<^sub>v fs ! j = fs'' ! i" unfolding fs''[symmetric] using inv(6) i by auto note res = res[unfolded mu' mu d'] show ?thesis unfolding basis_reduction_add_rows_loop.simps Let_def id if_False fs'' proof (rule Suc(1)[OF _ step(1,2) res _ i]) note list_repr = to_list_repr[OF impl Linv state] from i have ii: "i < length [0.._small_row i fs j" and res: "LLL_Impl.basis_reduction_add_rows_loop n state i j (map ((!) fs) (rev [0 ..< j])) = state'" (is "LLL_Impl.basis_reduction_add_rows_loop n state i j (?mapf fs j) = _") and j: "j \ i" and i: "i < m" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i fs'" "LLL_invariant False i fs'" "LLL_measure i fs' = LLL_measure i fs" "basis_reduction_add_rows_loop i fs j = fs'" using basis_reduction_add_rows_loop_impl[OF assms] basis_reduction_add_rows_loop[OF inv mu_small _ i j] by blast+ lemma basis_reduction_add_rows_impl: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant upw i fs" and res: "LLL_Impl.basis_reduction_add_rows n upw i state = state'" and i: "i < m" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i fs'" "basis_reduction_add_rows upw i fs = fs'" proof (atomize(full), goal_cases) case 1 obtain f mu ds where state: "state = (f,mu,ds)" by (cases state, auto) note def = LLL_Impl.basis_reduction_add_rows_def basis_reduction_add_rows_def show ?case proof (cases upw) case False from LLL_invD[OF inv] have len: "length fs = m" by auto from fs_state[OF impl inv state len] have "fs_state state = fs" by auto with assms False show ?thesis by (auto simp: def) next case True with inv have "LLL_invariant True i fs" by auto note start = this \_small_row_refl[of i fs] have id: "small_fs_state state = map (\ i. fs ! i) (rev [0.. \ * sq_norm (gso fs i)" and i: "i < m" and i0: "i \ 0" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i' fs'" (is ?g1) "basis_reduction_swap i fs = (upw',i',fs')" (is ?g2) proof - from i i0 have ii: "i - 1 < i" and le_m: "i - 1 \ m" "i \ m" "Suc i \ m" by auto obtain f mu ds where state: "state = (f,mu,ds)" by (cases state, auto) note dmu_ij_state = dmu_ij_state[OF impl inv state] note d_state = d_state[OF impl inv state] note res = res[unfolded LLL_Impl.basis_reduction_swap_def Let_def split state, folded state, unfolded fi_state[OF impl inv state i] fim1_state[OF impl inv state i i0]] note state_id = dmu_ij_state[OF ii i] note d_state_i = d_state[OF le_m(1)] d_state[OF le_m(2)] d_state[OF le_m(3)] from LLL_invD[OF inv] have len: "length fs = m" by auto from fs_state[OF impl inv state len] have fs: "fs_state state = fs" by auto obtain fs'' where fs'': "fs[i := fs ! (i - 1), i - 1 := fs ! i] = fs''" by auto let ?r = rat_of_int let ?d = "d fs" let ?d' = "d fs''" let ?dmus = "dmu_ij_state state" let ?ds = "d_state state" note swap = basis_reduction_swap_main[OF inv i i0 cond refl, unfolded fs''] interpret fs: fs_int' n m fs_init \ False i fs by standard (use inv in auto) interpret fs'': fs_int' n m fs_init \ False "i - 1" fs'' by standard (use swap in auto) note dmu = fs.d\ note dmu' = fs''.d\ note inv' = LLL_invD[OF inv] have fi: "fs ! (i - 1) = fs'' ! i" "fs ! i = fs'' ! (i - 1)" unfolding fs''[symmetric] using inv'(6) i i0 by auto from res have upw': "upw' = False" "i' = i - 1" by auto let ?dmu_repr' = "swap_mu m mu i (?dmus i (i - 1)) (?d (i - 1)) (?d i) (?d (Suc i))" let ?d'i = "(?d (Suc i) * ?d (i - 1) + ?dmus i (i - 1) * ?dmus i (i - 1)) div (?d i)" from res[unfolded fi d_state_i] have res: "upw' = False" "i' = i - 1" "state' = (dec_i (update_im1 (update_i f (fs'' ! i)) (fs'' ! (i - 1))), ?dmu_repr', iarray_update ds i ?d'i)" by auto from i have ii: "i < length [0.. {i-1,i}", auto) finally have f_repr: "list_repr (i - 1) ?fr (map ((!) fs'') [0.._swap[OF inv i i0 cond fs''[symmetric]] note dmu_ii = dmu_ij_state[OF \i - 1 < i\ i] show ?g1 unfolding fs_id LLL_impl_inv.simps res proof (intro conjI f_repr) show "d_repr (iarray_update ds i ?d'i) fs''" unfolding d_repr[unfolded d_repr_def] d_repr_def iarray_update_of_fun dmu_ii by (rule iarray_cong', subst updates(1), auto simp: nth_append intro: arg_cong) show "mu_repr ?dmu_repr' fs''" unfolding mu_repr_def swap_mu_def Let_def dmu_ii proof (rule iarray_cong', goal_cases) case ii: (1 ii) show ?case proof (cases "ii < i - 1") case small: True hence id: "(ii = i) = False" "(ii = i - 1) = False" "(i < ii) = False" "(ii < i - 1) = True" by auto have mu: "mu !! ii = IArray.of_fun (d\ fs ii) ii" using ii unfolding mu_def by auto show ?thesis unfolding id if_True if_False mu by (rule iarray_cong', insert small ii i i0, subst updates(2), simp_all, linarith) next case False hence iFalse: "(ii < i - 1) = False" by auto show ?thesis unfolding iFalse if_False if_distrib[of "\ f. IArray.of_fun f ii", symmetric] dmu_ij_state.simps[of f mu ds, folded state, symmetric] proof (rule iarray_cong', goal_cases) case j: (1 j) note upd = updates(2)[OF ii j] dmu_ii dmu_ij_state[OF j ii] if_distrib[of "\ x. x j"] note simps = dmu_ij_state[OF _ ii] dmu_ij_state[OF _ im1] dmu_ij_state[OF _ i] from False consider (I) "ii = i" "j = i - 1" | (Is) "ii = i" "j \ i - 1" | (Im1) "ii = i - 1" | (large) "ii > i" by linarith thus ?case proof (cases) case (I) show ?thesis unfolding upd using I by auto next case (Is) show ?thesis unfolding upd using Is j simps by auto next case (Im1) hence id: "(i < ii) = False" "(ii = i) = False" "(ii = i - 1) = True" using i0 by auto show ?thesis unfolding upd unfolding id if_False if_True by (rule simps, insert j Im1, auto) next case (large) hence "i - 1 < ii" "i < ii" by auto note simps = simps(1)[OF this(1)] simps(1)[OF this(2)] from large have id: "(i < ii) = True" "(ii = i - 1) = False" "\ x. (ii = i \ x) = False" by auto show ?thesis unfolding id if_True if_False upd simps by auto qed qed qed qed qed show ?g2 unfolding fs_id fs''[symmetric] basis_reduction_swap_def unfolding res .. qed lemma basis_reduction_swap: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant False i fs" and res: "LLL_Impl.basis_reduction_swap m i state = (upw',i',state')" and cond: "sq_norm (gso fs (i - 1)) > \ * sq_norm (gso fs i)" and i: "i < m" and i0: "i \ 0" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i' fs'" "LLL_invariant upw' i' fs'" "LLL_measure i' fs' < LLL_measure i fs" "basis_reduction_swap i fs = (upw',i',fs')" using basis_reduction_swap_impl[OF assms] basis_reduction_swap[OF inv _ cond i i0] by blast+ lemma basis_reduction_step_impl: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant upw i fs" and res: "LLL_Impl.basis_reduction_step \ n m upw i state = (upw',i',state')" and i: "i < m" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i' fs'" "basis_reduction_step upw i fs = (upw',i',fs')" proof (atomize(full), goal_cases) case 1 obtain f mu ds where state: "state = (f,mu,ds)" by (cases state, auto) note def = LLL_Impl.basis_reduction_step_def basis_reduction_step_def from LLL_invD[OF inv] have len: "length fs = m" by auto from fs_state[OF impl inv state len] have fs: "fs_state state = fs" by auto show ?case proof (cases "i = 0") case True from LLL_state_inc_state[OF impl inv state i] i assms increase_i[OF inv i True] True res fs' fs show ?thesis by (auto simp: def) next case False hence id: "(i = 0) = False" by auto obtain state'' where state'': "LLL_Impl.basis_reduction_add_rows n upw i state = state''" by auto define fs'' where fs'': "fs'' = fs_state state''" obtain f mu ds where state: "state'' = (f,mu,ds)" by (cases state'', auto) from basis_reduction_add_rows[OF impl inv state'' i fs''] have inv: "LLL_invariant False i fs''" and meas: "LLL_measure i fs = LLL_measure i fs''" and impl: "LLL_impl_inv state'' i fs''" and impl': "basis_reduction_add_rows upw i fs = fs''" by auto obtain num denom where quot: "quotient_of \ = (num,denom)" by force note d_state = d_state[OF impl inv state] from i have le: "i - 1 \ m" " i \ m" "Suc i \ m" by auto note d_state = d_state[OF le(1)] d_state[OF le(2)] d_state[OF le(3)] interpret fs'': fs_int' n m fs_init \ False i fs'' by standard (use inv in auto) have "i < length fs''" using LLL_invD[OF inv] i by auto note d_sq_norm_comparison = fs''.d_sq_norm_comparison[OF quot this False] note res = res[unfolded def id if_False Let_def state'' quot split d_state this] note pos = LLL_d_pos[OF inv le(1)] LLL_d_pos[OF inv le(2)] quotient_of_denom_pos[OF quot] from False have sim1: "Suc (i - 1) = i" by simp let ?r = "rat_of_int" let ?x = "sq_norm (gso fs'' (i - 1))" let ?y = "\ * sq_norm (gso fs'' i)" show ?thesis proof (cases "?x \ ?y") case True from increase_i[OF inv i _ True] True res meas LLL_state_inc_state[OF impl inv state i] fs' fs'' d_def d_sq_norm_comparison fs''.d_def impl' False show ?thesis by (auto simp: def) next case F: False hence gt: "?x > ?y" and id: "(?x \ ?y) = False" by auto from res[unfolded id if_False] d_def d_sq_norm_comparison fs''.d_def id have "LLL_Impl.basis_reduction_swap m i state'' = (upw', i', state')" by auto from basis_reduction_swap[OF impl inv this gt i False fs'] show ?thesis using meas F False by (auto simp: def Let_def impl') qed qed qed lemma basis_reduction_step: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant upw i fs" and res: "LLL_Impl.basis_reduction_step \ n m upw i state = (upw',i',state')" and i: "i < m" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' i' fs'" "LLL_invariant upw' i' fs'" "LLL_measure i' fs' < LLL_measure i fs" "basis_reduction_step upw i fs = (upw',i',fs')" using basis_reduction_step_impl[OF assms] basis_reduction_step[OF inv _ i] by blast+ lemma basis_reduction_main_impl: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant upw i fs" and res: "LLL_Impl.basis_reduction_main \ n m upw i state = state'" and fs': "fs' = fs_state state'" shows "LLL_impl_inv state' m fs'" "basis_reduction_main (upw,i,fs) = fs'" proof (atomize(full), insert assms(1-3), induct "LLL_measure i fs" arbitrary: i fs upw state rule: less_induct) case (less i fs upw) have id: "LLL_invariant upw i fs = True" using less by auto note res = less(4)[unfolded LLL_Impl.basis_reduction_main.simps[of _ _ _ upw]] note inv = less(3) note impl = less(2) note IH = less(1) show ?case proof (cases "i < m") case i: True obtain i'' state'' upw'' where step: "LLL_Impl.basis_reduction_step \ n m upw i state = (upw'',i'',state'')" (is "?step = _") by (cases ?step, auto) with res i have res: "LLL_Impl.basis_reduction_main \ n m upw'' i'' state'' = state'" by auto note main = basis_reduction_step[OF impl inv step i refl] from IH[OF main(3,1,2) res] main(4) step res show ?thesis by (simp add: i inv basis_reduction_main.simps) next case False from LLL_invD[OF inv] have len: "length fs = m" by auto obtain f mu ds where state: "state = (f,mu,ds)" by (cases state, auto) from fs_state[OF impl inv state len] have fs: "fs_state state = fs" by auto from False fs res fs' have fs_id: "fs = fs'" by simp from False LLL_invD[OF inv] have i: "i = m" by auto with False res inv impl fs have "LLL_invariant upw m fs' \ LLL_impl_inv state' m fs'" by (auto simp: fs') thus ?thesis unfolding basis_reduction_main.simps[of upw i fs] using False by (auto simp: LLL_invariant_def fs_id) qed qed lemma basis_reduction_main: assumes impl: "LLL_impl_inv state i fs" and inv: "LLL_invariant upw i fs" and res: "LLL_Impl.basis_reduction_main \ n m upw i state = state'" and fs': "fs' = fs_state state'" shows "LLL_invariant True m fs'" "LLL_impl_inv state' m fs'" "basis_reduction_main (upw,i,fs) = fs'" using basis_reduction_main_impl[OF assms] basis_reduction_main[OF inv] by blast+ lemma initial_state: "LLL_impl_inv (initial_state m fs_init) 0 fs_init" (is ?g1) "fs_state (initial_state m fs_init) = fs_init" (is ?g2) proof - have f_repr: "list_repr 0 ([], fs_init) (map ((!) fs_init) [0.. Rn" by auto have 1: "1 = d fs_init 0" unfolding d_def by simp define j where "j = m" have jm: "j \ m" unfolding j_def by auto have 0: "0 = m - j" unfolding j_def by auto interpret fs_init: fs_int_indpt n fs_init by (standard) (use lin_dep in auto) have mu_repr: "mu_repr (IArray.of_fun (\i. IArray.of_fun ((!!) (d\_impl fs_init !! i)) i) m) fs_init" unfolding fs_init.d\_impl mu_repr_def fs_init.d\_def d\_def fs_init.d_def d_def apply(rule iarray_cong') unfolding len[symmetric] by (auto simp add: nth_append) have d_repr: "d_repr (IArray.of_fun (\i. if i = 0 then 1 else d\_impl fs_init !! (i - 1) !! (i - 1)) (Suc m)) fs_init" unfolding fs_init.d\_impl d_repr_def proof (intro iarray_cong', goal_cases) case (1 i) show ?case proof (cases "i = 0") case False hence le: "i - 1 < length fs_init" "i - 1 < i" and id: "(i = 0) = False" "Suc (i - 1) = i" using 1 len by auto show ?thesis unfolding of_fun_nth[OF le(1)] of_fun_nth[OF le(2)] id if_False d\_def fs_init.d\_def fs_init.d_def d_def by (auto simp add: gs.\.simps ) next case True have "d fs_init 0 = 1" unfolding d_def gs.Gramian_determinant_0 by simp thus ?thesis unfolding True by simp qed qed show ?g1 unfolding initial_state_def Let_def LLL_impl_inv.simps id by (intro conjI f_repr mu_repr d_repr) from fs_state[OF this LLL_inv_initial_state] show ?g2 unfolding initial_state_def Let_def by (simp add: of_list_repr_def) qed lemma basis_reduction: assumes res: "basis_reduction \ n fs_init = state" and fs: "fs = fs_state state" shows "LLL_invariant True m fs" "LLL_impl_inv state m fs" "basis_reduction_main (True, 0, fs_init) = fs" using basis_reduction_main[OF initial_state(1) LLL_inv_initial_state res[unfolded basis_reduction_def len Let_def] fs] by auto lemma reduce_basis_impl: "LLL_Impl.reduce_basis \ fs_init = reduce_basis" proof - obtain fs where res: "LLL_Impl.reduce_basis \ fs_init = fs" by blast have "reduce_basis = fs" proof (cases fs_init) case (Cons f) from fs_init[unfolded Cons] have "dim_vec f = n" by auto from res[unfolded LLL_Impl.reduce_basis_def Cons list.simps this, folded Cons] have "fs_state (LLL_Impl.basis_reduction \ n fs_init) = fs" by auto from basis_reduction(3)[OF refl refl, unfolded this] show "reduce_basis = fs" unfolding reduce_basis_def . next case Nil with len have m0: "m = 0" by auto show ?thesis using res unfolding reduce_basis_def LLL_Impl.reduce_basis_def basis_reduction_main.simps using Nil m0 by simp qed with res show ?thesis by simp qed lemma reduce_basis: assumes "LLL_Impl.reduce_basis \ fs_init = fs" shows "lattice_of fs = L" "reduced fs m" "lin_indep fs" "length fs = m" "LLL_invariant True m fs" using reduce_basis_impl assms reduce_basis reduce_basis_inv by metis+ lemma short_vector_impl: "LLL_Impl.short_vector \ fs_init = short_vector" using reduce_basis_impl unfolding LLL_Impl.short_vector_def short_vector_def by simp lemma short_vector: assumes res: "LLL_Impl.short_vector \ fs_init = v" and m0: "m \ 0" shows "v \ carrier_vec n" "v \ L - {0\<^sub>v n}" "h \ L - {0\<^sub>v n} \ rat_of_int (sq_norm v) \ \ ^ (m - 1) * rat_of_int (sq_norm h)" "v \ 0\<^sub>v j" using short_vector[OF assms[unfolded short_vector_impl]] by metis+ end end diff --git a/thys/LLL_Basis_Reduction/Missing_Lemmas.thy b/thys/LLL_Basis_Reduction/Missing_Lemmas.thy --- a/thys/LLL_Basis_Reduction/Missing_Lemmas.thy +++ b/thys/LLL_Basis_Reduction/Missing_Lemmas.thy @@ -1,835 +1,835 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Missing lemmas\ text \This theory contains many results that are important but not specific for our development. They could be moved to the stardard library and some other AFP entries.\ theory Missing_Lemmas imports Berlekamp_Zassenhaus.Sublist_Iteration (* for thm upt_append *) Berlekamp_Zassenhaus.Square_Free_Int_To_Square_Free_GFp (* for thm large_mod_0 *) Algebraic_Numbers.Resultant Jordan_Normal_Form.Conjugate Jordan_Normal_Form.Missing_VectorSpace Jordan_Normal_Form.VS_Connect Berlekamp_Zassenhaus.Finite_Field_Factorization_Record_Based (* for transfer rules for thm vec_of_list_Nil *) Berlekamp_Zassenhaus.Berlekamp_Hensel (* for unique_factorization_m_factor *) begin no_notation test_bit (infixl "!!" 100) hide_const(open) module.smult up_ring.monom up_ring.coeff (**** Could be merged to HOL/Rings.thy ****) lemma (in zero_less_one) zero_le_one [simp]: "0 \ 1" by (rule less_imp_le, simp) subclass (in zero_less_one) zero_neq_one by (unfold_locales, simp add: less_imp_neq) class ordered_semiring_1 = Rings.ordered_semiring_0 + monoid_mult + zero_less_one begin subclass semiring_1.. lemma of_nat_ge_zero[intro!]: "of_nat n \ 0" using add_right_mono[of _ _ 1] by (induct n, auto) (* Following lemmas are moved from @{class ordered_idom}. *) lemma zero_le_power [simp]: "0 \ a \ 0 \ a ^ n" by (induct n) simp_all lemma power_mono: "a \ b \ 0 \ a \ a ^ n \ b ^ n" by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" using power_mono [of 1 a n] by simp lemma power_le_one: "0 \ a \ a \ 1 \ a ^ n \ 1" using power_mono [of a 1 n] by simp lemma power_gt1_lemma: assumes gt1: "1 < a" shows "1 < a * a ^ n" proof - from gt1 have "0 \ a" by (fact order_trans [OF zero_le_one less_imp_le]) from gt1 have "1 * 1 < a * 1" by simp also from gt1 have "\ \ a * a ^ n" by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed lemma power_gt1: "1 < a \ 1 < a ^ Suc n" by (simp add: power_gt1_lemma) lemma one_less_power [simp]: "1 < a \ 0 < n \ 1 < a ^ n" by (cases n) (simp_all add: power_gt1_lemma) lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" proof (induction N) case (Suc N) then have "a * a^N \ 1 * a^n" if "n \ N" using that by (intro mult_mono) auto then show ?case using Suc by (auto simp add: le_Suc_eq) qed (auto) lemma power_increasing: "n \ N \ 1 \ a \ a ^ n \ a ^ N" proof (induction N) case (Suc N) then have "1 * a^n \ a * a^N" if "n \ N" using that by (intro mult_mono) (auto simp add: order_trans[OF zero_le_one]) then show ?case using Suc by (auto simp add: le_Suc_eq) qed (auto) lemma power_Suc_le_self: "0 \ a \ a \ 1 \ a ^ Suc n \ a" using power_decreasing [of 1 "Suc n" a] by simp end lemma prod_list_nonneg: "(\ x. (x :: 'a :: ordered_semiring_1) \ set xs \ x \ 0) \ prod_list xs \ 0" by (induct xs, auto) subclass (in ordered_idom) ordered_semiring_1 by unfold_locales auto (**** End of lemmas that could be moved to HOL/Rings.thy ****) (* missing lemma on logarithms *) lemma log_prod: assumes "0 < a" "a \ 1" "\ x. x \ X \ 0 < f x" shows "log a (prod f X) = sum (log a o f) X" using assms(3) proof (induct X rule: infinite_finite_induct) case (insert x F) have "log a (prod f (insert x F)) = log a (f x * prod f F)" using insert by simp also have "\ = log a (f x) + log a (prod f F)" by (rule log_mult[OF assms(1-2) insert(4) prod_pos], insert insert, auto) finally show ?case using insert by auto qed auto (* TODO: Jordan_Normal_Form/Missing_Ring.ordered_idom should be redefined *) subclass (in ordered_idom) zero_less_one by (unfold_locales, auto) hide_fact Missing_Ring.zero_less_one (**** The following lemmas could be part of the standard library ****) instance real :: ordered_semiring_strict by (intro_classes, auto) instance real :: linordered_idom.. (*This is a generalisation of thm less_1_mult*) lemma less_1_mult': fixes a::"'a::linordered_semidom" shows "1 < a \ 1 \ b \ 1 < a * b" by (metis le_less less_1_mult mult.right_neutral) lemma upt_minus_eq_append: "i\j \ i\j-k \ [i.. [0.. A" and ff: "\a. a \ A \ f (f a) = a" shows "bij_betw f A A" by (intro bij_betwI[OF f f], simp_all add: ff) lemma range_subsetI: assumes "\x. f x = g (h x)" shows "range f \ range g" using assms by auto lemma Gcd_uminus: fixes A::"int set" assumes "finite A" shows "Gcd A = Gcd (uminus ` A)" using assms by (induct A, auto) lemma aux_abs_int: fixes c :: int assumes "c \ 0" shows "\x\ \ \x * c\" proof - have "abs x = abs x * 1" by simp also have "\ \ abs x * abs c" by (rule mult_left_mono, insert assms, auto) finally show ?thesis unfolding abs_mult by auto qed lemma mod_0_abs_less_imp_0: fixes a::int assumes a1: "[a = 0] (mod m)" and a2: "abs(a)0" using assms by auto thus ?thesis using assms unfolding cong_def using int_mod_pos_eq large_mod_0 zless_imp_add1_zle by (metis abs_of_nonneg le_less not_less zabs_less_one_iff zmod_trival_iff) qed (* an intro version of sum_list_0 *) lemma sum_list_zero: assumes "set xs \ {0}" shows "sum_list xs = 0" using assms by (induct xs, auto) (* About @{const max} *) lemma max_idem [simp]: shows "max a a = a" by (simp add: max_def) lemma hom_max: assumes "a \ b \ f a \ f b" shows "f (max a b) = max (f a) (f b)" using assms by (auto simp: max_def) lemma le_max_self: fixes a b :: "'a :: preorder" assumes "a \ b \ b \ a" shows "a \ max a b" and "b \ max a b" using assms by (auto simp: max_def) lemma le_max: fixes a b :: "'a :: preorder" assumes "c \ a \ c \ b" and "a \ b \ b \ a" shows "c \ max a b" using assms(1) le_max_self[OF assms(2)] by (auto dest: order_trans) fun max_list where "max_list [] = (THE x. False)" (* more convenient than "undefined" *) | "max_list [x] = x" | "max_list (x # y # xs) = max x (max_list (y # xs))" declare max_list.simps(1) [simp del] declare max_list.simps(2-3)[code] lemma max_list_Cons: "max_list (x#xs) = (if xs = [] then x else max x (max_list xs))" by (cases xs, auto) lemma max_list_mem: "xs \ [] \ max_list xs \ set xs" by (induct xs, auto simp: max_list_Cons max_def) lemma mem_set_imp_le_max_list: fixes xs :: "'a :: preorder list" assumes "\a b. a \ set xs \ b \ set xs \ a \ b \ b \ a" and "a \ set xs" shows "a \ max_list xs" proof (insert assms, induct xs arbitrary:a) case Nil with assms show ?case by auto next case (Cons x xs) show ?case proof (cases "xs = []") case False have "x \ max_list xs \ max_list xs \ x" apply (rule Cons(2)) using max_list_mem[of xs] False by auto note 1 = le_max_self[OF this] from Cons have "a = x \ a \ set xs" by auto then show ?thesis proof (elim disjE) assume a: "a = x" show ?thesis by (unfold a max_list_Cons, auto simp: False intro!: 1) next assume "a \ set xs" then have "a \ max_list xs" by (intro Cons, auto) with 1 have "a \ max x (max_list xs)" by (auto dest: order_trans) then show ?thesis by (unfold max_list_Cons, auto simp: False) qed qed (insert Cons, auto) qed lemma le_max_list: fixes xs :: "'a :: preorder list" assumes ord: "\a b. a \ set xs \ b \ set xs \ a \ b \ b \ a" and ab: "a \ b" and b: "b \ set xs" shows "a \ max_list xs" proof- note ab also have "b \ max_list xs" by (rule mem_set_imp_le_max_list, fact ord, fact b) finally show ?thesis. qed lemma max_list_le: fixes xs :: "'a :: preorder list" assumes a: "\x. x \ set xs \ x \ a" and xs: "xs \ []" shows "max_list xs \ a" using max_list_mem[OF xs] a by auto lemma max_list_as_Greatest: assumes "\x y. x \ set xs \ y \ set xs \ x \ y \ y \ x" shows "max_list xs = (GREATEST a. a \ set xs)" proof (cases "xs = []") case True then show ?thesis by (unfold Greatest_def, auto simp: max_list.simps(1)) next case False from assms have 1: "x \ set xs \ x \ max_list xs" for x by (auto intro: le_max_list) have 2: "max_list xs \ set xs" by (fact max_list_mem[OF False]) have "\!x. x \ set xs \ (\y. y \ set xs \ y \ x)" (is "\!x. ?P x") proof (intro ex1I) from 1 2 show "?P (max_list xs)" by auto next fix x assume 3: "?P x" with 1 have "x \ max_list xs" by auto moreover from 2 3 have "max_list xs \ x" by auto ultimately show "x = max_list xs" by auto qed note 3 = theI_unique[OF this,symmetric] from 1 2 show ?thesis by (unfold Greatest_def Cons 3, auto) qed lemma hom_max_list_commute: assumes "xs \ []" and "\x y. x \ set xs \ y \ set xs \ h (max x y) = max (h x) (h y)" shows "h (max_list xs) = max_list (map h xs)" by (insert assms, induct xs, auto simp: max_list_Cons max_list_mem) (*Efficient rev [i.. nat \ nat list" ("(1[_>.._])") where rev_upt_0: "[0>..j] = []" | rev_upt_Suc: "[(Suc i)>..j] = (if i \ j then i # [i>..j] else [])" lemma rev_upt_rec: "[i>..j] = (if i>j then [i>..Suc j] @ [j] else [])" by (induct i, auto) definition rev_upt_aux :: "nat \ nat \ nat list \ nat list" where "rev_upt_aux i j js = [i>..j] @ js" lemma upt_aux_rec [code]: "rev_upt_aux i j js = (if j\i then js else rev_upt_aux i (Suc j) (j#js))" by (induct j, auto simp add: rev_upt_aux_def rev_upt_rec) lemma rev_upt_code[code]: "[i>..j] = rev_upt_aux i j []" by(simp add: rev_upt_aux_def) lemma upt_rev_upt: "rev [j>..i] = [i....i]" by (induct j, auto) lemma length_rev_upt [simp]: "length [i>..j] = i - j" by (induct i) (auto simp add: Suc_diff_le) lemma nth_rev_upt [simp]: "j + k < i \ [i>..j] ! k = i - 1 - k" proof - assume jk_i: "j + k < i" have "[i>..j] = rev [j....n]) ! i = f (m - 1 - i)" proof - have "(map f [m>..n]) ! i = f ([m>..n] ! i)" by (rule nth_map, auto simp add: i) also have "... = f (m - 1 - i)" proof (rule arg_cong[of _ _ f], rule nth_rev_upt) show "n + i < m" using i by linarith qed finally show ?thesis . qed lemma coeff_mult_monom: "coeff (p * monom a d) i = (if d \ i then a * coeff p (i - d) else 0)" using coeff_monom_mult[of a d p] by (simp add: ac_simps) (**** End of the lemmas which may be part of the standard library ****) (**** The following lemmas could be moved to Algebraic_Numbers/Resultant.thy ****) lemma vec_of_poly_0 [simp]: "vec_of_poly 0 = 0\<^sub>v 1" by (auto simp: vec_of_poly_def) lemma vec_index_vec_of_poly [simp]: "i \ degree p \ vec_of_poly p $ i = coeff p (degree p - i)" by (simp add: vec_of_poly_def Let_def) lemma poly_of_vec_vec: "poly_of_vec (vec n f) = Poly (rev (map f [0.. Suc) [0..) = Poly (rev (map (f \ Suc) [0.. = poly_of_vec (vec n (f \ Suc)) + monom (f 0) n" by (fold Suc, simp) also have "\ = poly_of_vec (vec (Suc n) f)" apply (unfold poly_of_vec_def Let_def dim_vec sum.lessThan_Suc) by (auto simp add: Suc_diff_Suc) finally show ?case.. qed lemma sum_list_map_dropWhile0: assumes f0: "f 0 = 0" shows "sum_list (map f (dropWhile ((=) 0) xs)) = sum_list (map f xs)" by (induct xs, auto simp add: f0) lemma coeffs_poly_of_vec: "coeffs (poly_of_vec v) = rev (dropWhile ((=) 0) (list_of_vec v))" proof- obtain n f where v: "v = vec n f" by transfer auto show ?thesis by (simp add: v poly_of_vec_vec) qed lemma poly_of_vec_vCons: "poly_of_vec (vCons a v) = monom a (dim_vec v) + poly_of_vec v" (is "?l = ?r") by (auto intro: poly_eqI simp: coeff_poly_of_vec vec_index_vCons) lemma poly_of_vec_as_Poly: "poly_of_vec v = Poly (rev (list_of_vec v))" by (induct v, auto simp:poly_of_vec_vCons Poly_snoc ac_simps) lemma poly_of_vec_add: assumes "dim_vec a = dim_vec b" shows "poly_of_vec (a + b) = poly_of_vec a + poly_of_vec b" using assms by (auto simp add: poly_eq_iff coeff_poly_of_vec) (*TODO: replace the one in Resultant.thy*) lemma degree_poly_of_vec_less: assumes "0 < dim_vec v" and "dim_vec v \ n" shows "degree (poly_of_vec v) < n" using degree_poly_of_vec_less assms by (auto dest: less_le_trans) lemma (in vec_module) poly_of_vec_finsum: assumes "f \ X \ carrier_vec n" shows "poly_of_vec (finsum V f X) = (\i\X. poly_of_vec (f i))" proof (cases "finite X") case False then show ?thesis by auto next case True show ?thesis proof (insert True assms, induct X rule: finite_induct) case IH: (insert a X) have [simp]: "f x \ carrier_vec n" if x: "x \ X" for x using x IH.prems unfolding Pi_def by auto have [simp]: "f a \ carrier_vec n" using IH.prems unfolding Pi_def by auto have [simp]: "dim_vec (finsum V f X) = n" by simp have [simp]: "dim_vec (f a) = n" by simp show ?case proof (cases "a \ X") case True then show ?thesis by (auto simp: insert_absorb IH) next case False then have "(finsum V f (insert a X)) = f a + (finsum V f X)" by (auto intro: finsum_insert IH) also have "poly_of_vec ... = poly_of_vec (f a) + poly_of_vec (finsum V f X)" by (rule poly_of_vec_add, simp) also have "... = (\i\insert a X. poly_of_vec (f i))" using IH False by (subst sum.insert, auto) finally show ?thesis . qed qed auto qed (*This function transforms a polynomial to a vector of dimension n*) definition "vec_of_poly_n p n = vec n (\i. if i < n - degree p - 1 then 0 else coeff p (n - i - 1))" (* TODO: make it abbreviation? *) lemma vec_of_poly_as: "vec_of_poly_n p (Suc (degree p)) = vec_of_poly p" by (induct p, auto simp: vec_of_poly_def vec_of_poly_n_def) lemma vec_of_poly_n_0 [simp]: "vec_of_poly_n p 0 = vNil" by (auto simp: vec_of_poly_n_def) lemma vec_dim_vec_of_poly_n [simp]: "dim_vec (vec_of_poly_n p n) = n" "vec_of_poly_n p n \ carrier_vec n" unfolding vec_of_poly_n_def by auto lemma dim_vec_of_poly [simp]: "dim_vec (vec_of_poly f) = degree f + 1" by (simp add: vec_of_poly_as[symmetric]) lemma vec_index_of_poly_n: assumes "i < n" shows "vec_of_poly_n p n $ i = (if i < n - Suc (degree p) then 0 else coeff p (n - i - 1))" using assms by (auto simp: vec_of_poly_n_def Let_def) lemma vec_of_poly_n_pCons[simp]: shows "vec_of_poly_n (pCons a p) (Suc n) = vec_of_poly_n p n @\<^sub>v vec_of_list [a]" (is "?l = ?r") proof (unfold vec_eq_iff, intro conjI allI impI) show "dim_vec ?l = dim_vec ?r" by auto show "i < dim_vec ?r \ ?l $ i = ?r $ i" for i by (cases "n - i", auto simp: coeff_pCons less_Suc_eq_le vec_index_of_poly_n) qed lemma vec_of_poly_pCons: shows "vec_of_poly (pCons a p) = (if p = 0 then vec_of_list [a] else vec_of_poly p @\<^sub>v vec_of_list [a])" by (cases "degree p", auto simp: vec_of_poly_as[symmetric]) lemma list_of_vec_of_poly [simp]: "list_of_vec (vec_of_poly p) = (if p = 0 then [0] else rev (coeffs p))" by (induct p, auto simp: vec_of_poly_pCons) lemma poly_of_vec_of_poly_n: assumes p: "degree p n" for i by (rule coeff_eq_0, insert i2 p, simp) ultimately show ?thesis using assms unfolding poly_eq_iff unfolding coeff_poly_of_vec by auto qed lemma vec_of_poly_n0[simp]: "vec_of_poly_n 0 n = 0\<^sub>v n" unfolding vec_of_poly_n_def by auto lemma vec_of_poly_n_add: "vec_of_poly_n (a + b) n = vec_of_poly_n a n + vec_of_poly_n b n" proof (induct n arbitrary: a b) case 0 then show ?case by auto next case (Suc n) then show ?case by (cases a, cases b, auto) qed lemma vec_of_poly_n_poly_of_vec: assumes n: "dim_vec g = n" shows "vec_of_poly_n (poly_of_vec g) n = g" proof (auto simp add: poly_of_vec_def vec_of_poly_n_def assms vec_eq_iff Let_def) have d: "degree (\ii degree (poly_of_vec g)" using n by linarith then show "g $ i = 0" using i1 i2 i3 by (metis (no_types, lifting) Suc_diff_Suc coeff_poly_of_vec diff_Suc_less diff_diff_cancel leD le_degree less_imp_le_nat n neq0_conv) next fix i assume "i < n" thus "coeff (\i\<^sub>v (vec_of_poly_n b n)) = smult a b" using assms by (auto simp add: poly_eq_iff coeff_poly_of_vec vec_of_poly_n_def coeff_eq_0) (*TODO: replace the one in Resultant.thy*) definition vec_of_poly_rev_shifted where "vec_of_poly_rev_shifted p n s j \ vec n (\i. if i \ j \ j \ s + i then coeff p (s + i - j) else 0)" lemma vec_of_poly_rev_shifted_dim[simp]: "dim_vec (vec_of_poly_rev_shifted p n s j) = n" unfolding vec_of_poly_rev_shifted_def by auto lemma col_sylvester_sub: (* TODO: from this directly derive col_sylvester *) assumes j: "j < m + n" shows "col (sylvester_mat_sub m n p q) j = vec_of_poly_rev_shifted p n m j @\<^sub>v vec_of_poly_rev_shifted q m n j" (is "?l = ?r") proof show "dim_vec ?l = dim_vec ?r" by simp fix i assume "i < dim_vec ?r" then have i: "i < m+n" by auto show "?l $ i = ?r $ i" unfolding vec_of_poly_rev_shifted_def apply (subst index_col) using i apply simp using j apply simp apply (subst sylvester_mat_sub_index) using i apply simp using j apply simp apply (cases "i < n") using i apply force using i apply (auto simp: not_less not_le intro!: coeff_eq_0) done qed lemma vec_of_poly_rev_shifted_scalar_prod: fixes p v defines "q \ poly_of_vec v" assumes m: "degree p \ m" and n: "dim_vec v = n" assumes j: "j < m+n" shows "vec_of_poly_rev_shifted p n m (n+m-Suc j) \ v = coeff (p * q) j" (is "?l = ?r") proof - have id1: "\ i. m + i - (n + m - Suc j) = i + Suc j - n" using j by auto let ?g = "\ i. if i \ n + m - Suc j \ n - Suc j \ i then coeff p (i + Suc j - n) * v $ i else 0" have "?thesis = ((\i = 0..i\j. coeff p i * (if j - i < n then v $ (n - Suc (j - i)) else 0)))" (is "_ = (?l = ?r)") unfolding vec_of_poly_rev_shifted_def coeff_mult m scalar_prod_def n q_def coeff_poly_of_vec by (subst sum.cong, insert id1, auto) also have "..." proof - have "?r = (\i\j. (if j - i < n then coeff p i * v $ (n - Suc (j - i)) else 0))" (is "_ = sum ?f _") by (rule sum.cong, auto) also have "sum ?f {..j} = sum ?f ({i. i \ j \ j - i < n} \ {i. i \ j \ \ j - i < n})" (is "_ = sum _ (?R1 \ ?R2)") by (rule sum.cong, auto) also have "\ = sum ?f ?R1 + sum ?f ?R2" by (subst sum.union_disjoint, auto) also have "sum ?f ?R2 = 0" by (rule sum.neutral, auto) also have "sum ?f ?R1 + 0 = sum (\ i. coeff p i * v $ (i + n - Suc j)) ?R1" (is "_ = sum ?F _") by (subst sum.cong, auto simp: ac_simps) also have "\ = sum ?F ((?R1 \ {..m}) \ (?R1 - {..m}))" (is "_ = sum _ (?R \ ?R')") by (rule sum.cong, auto) also have "\ = sum ?F ?R + sum ?F ?R'" by (subst sum.union_disjoint, auto) also have "sum ?F ?R' = 0" proof - { fix x assume "x > m" with m have "?F x = 0" by (subst coeff_eq_0, auto) } thus ?thesis by (subst sum.neutral, auto) qed finally have r: "?r = sum ?F ?R" by simp have "?l = sum ?g ({i. i < n \ i \ n + m - Suc j \ n - Suc j \ i} \ {i. i < n \ \ (i \ n + m - Suc j \ n - Suc j \ i)})" (is "_ = sum _ (?L1 \ ?L2)") by (rule sum.cong, auto) also have "\ = sum ?g ?L1 + sum ?g ?L2" by (subst sum.union_disjoint, auto) also have "sum ?g ?L2 = 0" by (rule sum.neutral, auto) also have "sum ?g ?L1 + 0 = sum (\ i. coeff p (i + Suc j - n) * v $ i) ?L1" (is "_ = sum ?G _") by (subst sum.cong, auto) also have "\ = sum ?G (?L1 \ {i. i + Suc j - n \ m} \ (?L1 - {i. i + Suc j - n \ m}))" (is "_ = sum _ (?L \ ?L')") by (subst sum.cong, auto) also have "\ = sum ?G ?L + sum ?G ?L'" by (subst sum.union_disjoint, auto) also have "sum ?G ?L' = 0" proof - { fix x assume "x + Suc j - n > m" with m have "?G x = 0" by (subst coeff_eq_0, auto) } thus ?thesis by (subst sum.neutral, auto) qed finally have l: "?l = sum ?G ?L" by simp let ?bij = "\ i. i + n - Suc j" { fix x assume x: "j < m + n" "Suc (x + j) - n \ m" "x < n" "n - Suc j \ x" define y where "y = x + Suc j - n" from x have "x + Suc j \ n" by auto with x have xy: "x = ?bij y" unfolding y_def by auto from x have y: "y \ ?R" unfolding y_def by auto have "x \ ?bij ` ?R" unfolding xy using y by blast } note tedious = this show ?thesis unfolding l r by (rule sum.reindex_cong[of ?bij], insert j, auto simp: inj_on_def tedious) qed finally show ?thesis by simp qed lemma sylvester_sub_poly: fixes p q :: "'a :: comm_semiring_0 poly" assumes m: "degree p \ m" assumes n: "degree q \ n" assumes v: "v \ carrier_vec (m+n)" shows "poly_of_vec ((sylvester_mat_sub m n p q)\<^sup>T *\<^sub>v v) = poly_of_vec (vec_first v n) * p + poly_of_vec (vec_last v m) * q" (is "?l = ?r") proof (rule poly_eqI) fix i let ?Tv = "(sylvester_mat_sub m n p q)\<^sup>T *\<^sub>v v" have dim: "dim_vec (vec_first v n) = n" "dim_vec (vec_last v m) = m" "dim_vec ?Tv = n + m" using v by auto have if_distrib: "\ x y z. (if x then y else (0 :: 'a)) * z = (if x then y * z else 0)" by auto show "coeff ?l i = coeff ?r i" proof (cases "i < m+n") case False hence i_mn: "i \ m+n" and i_n: "\x. x \ i \ x < n \ x < n" and i_m: "\x. x \ i \ x < m \ x < m" by auto have "coeff ?r i = (\ x < n. vec_first v n $ (n - Suc x) * coeff p (i - x)) + (\ x < m. vec_last v m $ (m - Suc x) * coeff q (i - x))" (is "_ = sum ?f _ + sum ?g _") unfolding coeff_add coeff_mult Let_def unfolding coeff_poly_of_vec dim if_distrib unfolding atMost_def apply(subst sum.inter_filter[symmetric],simp) apply(subst sum.inter_filter[symmetric],simp) unfolding mem_Collect_eq unfolding i_n i_m unfolding lessThan_def by simp also { fix x assume x: "x < n" have "coeff p (i-x) = 0" apply(rule coeff_eq_0) using i_mn x m by auto hence "?f x = 0" by auto } hence "sum ?f {..T *\<^sub>v v) $ (n + m - Suc i)" unfolding coeff_poly_of_vec dim sum.distrib[symmetric] by auto also have "... = coeff (p * poly_of_vec (vec_first v n) + q * poly_of_vec (vec_last v m)) i" apply(subst index_mult_mat_vec) using True apply simp apply(subst row_transpose) using True apply simp apply(subst col_sylvester_sub) using True apply simp apply(subst vec_first_last_append[of v n m,symmetric]) using v apply(simp add: add.commute) apply(subst scalar_prod_append) apply (rule carrier_vecI,simp)+ apply (subst vec_of_poly_rev_shifted_scalar_prod[OF m],simp) using True apply simp apply (subst add.commute[of n m]) apply (subst vec_of_poly_rev_shifted_scalar_prod[OF n]) apply simp using True apply simp by simp also have "... = (\x\i. (if x < n then vec_first v n $ (n - Suc x) else 0) * coeff p (i - x)) + (\x\i. (if x < m then vec_last v m $ (m - Suc x) else 0) * coeff q (i - x))" unfolding coeff_poly_of_vec[of "vec_first v n",unfolded dim_vec_first,symmetric] unfolding coeff_poly_of_vec[of "vec_last v m",unfolded dim_vec_last,symmetric] unfolding coeff_mult[symmetric] by (simp add: mult.commute) also have "... = coeff ?r i" unfolding coeff_add coeff_mult Let_def unfolding coeff_poly_of_vec dim.. finally show ?thesis. qed qed (**** End of the lemmas which could be moved to Algebraic_Numbers/Resultant.thy ****) (**** The following lemmas could be moved to Computational_Algebra/Polynomial.thy ****) lemma normalize_field [simp]: "normalize (a :: 'a :: {field, semiring_gcd}) = (if a = 0 then 0 else 1)" using unit_factor_normalize by fastforce lemma content_field [simp]: "content (p :: 'a :: {field,semiring_gcd} poly) = (if p = 0 then 0 else 1)" by (induct p, auto simp: content_def) lemma primitive_part_field [simp]: "primitive_part (p :: 'a :: {field,semiring_gcd} poly) = p" by (cases "p = 0", auto intro!: primitive_part_prim) lemma primitive_part_dvd: "primitive_part a dvd a" by (metis content_times_primitive_part dvd_def dvd_refl mult_smult_right) lemma degree_abs [simp]: "degree \p\ = degree p" by (auto simp: abs_poly_def) lemma degree_gcd1: assumes a_not0: "a \ 0" shows "degree (gcd a b) \ degree a" proof - let ?g = "gcd a b" have gcd_dvd_b: "?g dvd a" by simp from this obtain c where a_gc: "a = ?g * c" unfolding dvd_def by auto have g_not0: "?g \0" using a_not0 a_gc by auto have c0: "c \ 0" using a_not0 a_gc by auto have "degree ?g \ degree (?g * c)" by (rule degree_mult_right_le[OF c0]) also have "... = degree a" using a_gc by auto finally show ?thesis . qed lemma primitive_part_neg [simp]: fixes a::"'a :: {factorial_ring_gcd,factorial_semiring_multiplicative} poly" shows "primitive_part (-a) = - primitive_part a" proof - have "primitive_part (-a) = primitive_part (smult (-1) a)" by auto then show ?thesis unfolding primitive_part_smult by (simp add: is_unit_unit_factor) qed lemma content_uminus[simp]: fixes f::"int poly" shows "content (-f) = content f" proof - have "-f = - (smult 1 f)" by auto also have "... = smult (-1) f" using smult_minus_left by auto finally have "content (-f) = content (smult (-1) f)" by auto also have "... = normalize (- 1) * content f" unfolding content_smult .. finally show ?thesis by auto qed lemma pseudo_mod_monic: fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" defines "r \ pseudo_mod f g" assumes monic_g: "monic g" shows "\q. f = g * q + r" "r = 0 \ degree r < degree g" proof - let ?cg = "coeff g (degree g)" let ?cge = "?cg ^ (Suc (degree f) - degree g)" define a where "a = ?cge" from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)" by (cases "pseudo_divmod f g") auto have g: "g \ 0" using monic_g by auto from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" by (auto simp: a_def) have a1: "a = 1" unfolding a_def using monic_g by auto hence id2: "f = g * q + r" using id by auto show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id2 show "\q. f = g * q + r" by auto qed lemma monic_imp_div_mod_int_poly_degree: fixes p :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" assumes m: "monic u" shows "\q r. p = q*u + r \ (r = 0 \ degree r < degree u)" using pseudo_mod_monic[OF m] using mult.commute by metis corollary monic_imp_div_mod_int_poly_degree2: fixes p :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" assumes m: "monic u" and deg_u: "degree u > 0" shows "\q r. p = q*u + r \ (degree r < degree u)" proof - obtain q r where "p = q * u + r" and r: "(r = 0 \ degree r < degree u)" using monic_imp_div_mod_int_poly_degree[OF m, of p] by auto moreover have "degree r < degree u" using deg_u r by auto ultimately show ?thesis by auto qed (**** End of the lemmas that could be moved to Computational_Algebra/Polynomial.thy ****) (* To be categorized *) lemma (in zero_hom) hom_upper_triangular: "A \ carrier_mat n n \ upper_triangular A \ upper_triangular (map_mat hom A)" by (auto simp: upper_triangular_def) end diff --git a/thys/Native_Word/Code_Symbolic_Bits_Int.thy b/thys/Native_Word/Code_Symbolic_Bits_Int.thy --- a/thys/Native_Word/Code_Symbolic_Bits_Int.thy +++ b/thys/Native_Word/Code_Symbolic_Bits_Int.thy @@ -1,143 +1,143 @@ (* Title: Code_Symbolic_Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Symbolic implementation of bit operations on int\ theory Code_Symbolic_Bits_Int imports More_Bits_Int begin section \Implementations of bit operations on \<^typ>\int\ operating on symbolic representation\ lemma not_minus_numeral_inc_eq: \NOT (- numeral (Num.inc n)) = (numeral n :: int)\ - by (simp add: not_int_def sub_inc_One) + by (simp add: not_int_def sub_inc_One_eq) lemma [code_abbrev]: \test_bit = (bit :: int \ nat \ bool)\ by (simp add: fun_eq_iff) lemma test_bit_int_code [code]: "test_bit (0::int) n = False" "test_bit (Int.Neg num.One) n = True" "test_bit (Int.Pos num.One) 0 = True" "test_bit (Int.Pos (num.Bit0 m)) 0 = False" "test_bit (Int.Pos (num.Bit1 m)) 0 = True" "test_bit (Int.Neg (num.Bit0 m)) 0 = False" "test_bit (Int.Neg (num.Bit1 m)) 0 = True" "test_bit (Int.Pos num.One) (Suc n) = False" "test_bit (Int.Pos (num.Bit0 m)) (Suc n) = test_bit (Int.Pos m) n" "test_bit (Int.Pos (num.Bit1 m)) (Suc n) = test_bit (Int.Pos m) n" "test_bit (Int.Neg (num.Bit0 m)) (Suc n) = test_bit (Int.Neg m) n" "test_bit (Int.Neg (num.Bit1 m)) (Suc n) = test_bit (Int.Neg (Num.inc m)) n" by (simp_all add: Num.add_One bit_Suc) lemma int_not_code [code]: "NOT (0 :: int) = -1" "NOT (Int.Pos n) = Int.Neg (Num.inc n)" "NOT (Int.Neg n) = Num.sub n num.One" by(simp_all add: Num.add_One int_not_def) lemma int_and_code [code]: fixes i j :: int shows "0 AND j = 0" "i AND 0 = 0" "Int.Pos n AND Int.Pos m = (case bitAND_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n AND Int.Neg m = NOT (Num.sub n num.One OR Num.sub m num.One)" "Int.Pos n AND Int.Neg num.One = Int.Pos n" "Int.Pos n AND Int.Neg (num.Bit0 m) = Num.sub (bitORN_num (Num.BitM m) n) num.One" "Int.Pos n AND Int.Neg (num.Bit1 m) = Num.sub (bitORN_num (num.Bit0 m) n) num.One" "Int.Neg num.One AND Int.Pos m = Int.Pos m" "Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (bitORN_num (Num.BitM n) m) num.One" "Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (bitORN_num (num.Bit0 n) m) num.One" apply (simp_all add: int_numeral_bitAND_num Num.add_One - sub_inc_One inc_BitM not_minus_numeral_inc_eq + sub_inc_One_eq inc_BitM_eq not_minus_numeral_inc_eq flip: int_not_neg_numeral int_or_not_bitORN_num split: option.split) apply (simp_all add: ac_simps) done lemma int_or_code [code]: fixes i j :: int shows "0 OR j = j" "i OR 0 = i" "Int.Pos n OR Int.Pos m = Int.Pos (bitOR_num n m)" "Int.Neg n OR Int.Neg m = NOT (Num.sub n num.One AND Num.sub m num.One)" "Int.Pos n OR Int.Neg num.One = Int.Neg num.One" "Int.Pos n OR Int.Neg (num.Bit0 m) = (case bitANDN_num (Num.BitM m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Pos n OR Int.Neg (num.Bit1 m) = (case bitANDN_num (num.Bit0 m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg num.One OR Int.Pos m = Int.Neg num.One" "Int.Neg (num.Bit0 n) OR Int.Pos m = (case bitANDN_num (Num.BitM n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg (num.Bit1 n) OR Int.Pos m = (case bitANDN_num (num.Bit0 n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" apply (simp_all add: int_numeral_bitOR_num flip: int_not_neg_numeral) apply (simp_all add: or_int_def int_and_comm int_not_and_bitANDN_num del: int_not_simps(4) split: option.split) apply (simp_all add: Num.add_One) done lemma int_xor_code [code]: fixes i j :: int shows "0 XOR j = j" "i XOR 0 = i" "Int.Pos n XOR Int.Pos m = (case bitXOR_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n XOR Int.Neg m = Num.sub n num.One XOR Num.sub m num.One" "Int.Neg n XOR Int.Pos m = NOT (Num.sub n num.One XOR Int.Pos m)" "Int.Pos n XOR Int.Neg m = NOT (Int.Pos n XOR Num.sub m num.One)" by(fold int_not_neg_numeral)(simp_all add: int_numeral_bitXOR_num int_xor_not cong: option.case_cong) lemma bin_rest_code: "bin_rest i = i >> 1" by (simp add: shiftr_int_def) lemma sbintrunc_code [code]: "sbintrunc n bin = (let bin' = bin AND bin_mask (n + 1) in if bin' !! n then bin' - (2 << n) else bin')" proof (induction n arbitrary: bin) case 0 then show ?case by (simp add: mod_2_eq_odd and_one_eq) next case (Suc n) have *: \(4 * 2 ^ n - 1) div 2 = 2 * 2 ^ n - (1::int)\ by simp from Suc [of \bin div 2\] show ?case by (auto simp add: Let_def bin_mask_conv_pow2 shiftl_int_def algebra_simps * and_int_rec [of \_ + _ * 2\] and_int_rec [of \_ * 2\] bit_double_iff even_bit_succ_iff elim!: evenE oddE) qed lemma set_bits_code [code]: "set_bits = Code.abort (STR ''set_bits is unsupported on type int'') (\_. set_bits :: _ \ int)" by simp lemma fixes i :: int shows int_set_bit_True_conv_OR [code]: "set_bit i n True = i OR (1 << n)" and int_set_bit_False_conv_NAND [code]: "set_bit i n False = i AND NOT (1 << n)" and int_set_bit_conv_ops: "set_bit i n b = (if b then i OR (1 << n) else i AND NOT (1 << n))" by(simp_all add: set_bit_int_def bin_set_conv_OR bin_clr_conv_NAND) lemma int_shiftr_code [code]: fixes i :: int shows "i >> 0 = i" "0 >> Suc n = (0 :: int)" "Int.Pos num.One >> Suc n = 0" "Int.Pos (num.Bit0 m) >> Suc n = Int.Pos m >> n" "Int.Pos (num.Bit1 m) >> Suc n = Int.Pos m >> n" "Int.Neg num.One >> Suc n = -1" "Int.Neg (num.Bit0 m) >> Suc n = Int.Neg m >> n" "Int.Neg (num.Bit1 m) >> Suc n = Int.Neg (Num.inc m) >> n" by (simp_all add: shiftr_eq_drop_bit drop_bit_Suc add_One) lemma int_shiftl_code [code]: "i << 0 = i" "i << Suc n = Int.dup i << n" by (simp_all add: shiftl_int_def) lemma int_lsb_code [code]: "lsb (0 :: int) = False" "lsb (Int.Pos num.One) = True" "lsb (Int.Pos (num.Bit0 w)) = False" "lsb (Int.Pos (num.Bit1 w)) = True" "lsb (Int.Neg num.One) = True" "lsb (Int.Neg (num.Bit0 w)) = False" "lsb (Int.Neg (num.Bit1 w)) = True" by simp_all end diff --git a/thys/Native_Word/Code_Target_Word_Base.thy b/thys/Native_Word/Code_Target_Word_Base.thy --- a/thys/Native_Word/Code_Target_Word_Base.thy +++ b/thys/Native_Word/Code_Target_Word_Base.thy @@ -1,442 +1,443 @@ (* Title: Code_Target_Word_Base.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Common base for target language implementations of word types\ theory Code_Target_Word_Base imports "HOL-Word.Word" Bits_Integer begin text \More lemmas\ -lemma nat_div_eq_Suc_0_iff: "n div m = Suc 0 \ n \ m \ n < 2 * m" - apply (auto simp add: sdl) - using not_less apply fastforce - apply (metis One_nat_def Suc_1 div_eq_0_iff lessI neq0_conv td_gal_lt) +lemma nat_div_eq_Suc_0_iff: "n div m = Suc 0 \ m \ n \ n < 2 * m" + apply auto + using div_greater_zero_iff apply fastforce + apply (metis One_nat_def div_greater_zero_iff dividend_less_div_times mult.right_neutral mult_Suc mult_numeral_1 numeral_2_eq_2 zero_less_numeral) + apply (simp add: div_nat_eqI) done lemma Suc_0_lt_2p_len_of: "Suc 0 < 2 ^ LENGTH('a :: len)" by (metis One_nat_def len_gt_0 lessI numeral_2_eq_2 one_less_power) lemma div_half_nat: fixes x y :: nat assumes "y \ 0" shows "(x div y, x mod y) = (let q = 2 * (x div 2 div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - let ?q = "2 * (x div 2 div y)" have q: "?q = x div y - x div y mod 2" by(metis div_mult2_eq mult.commute minus_mod_eq_mult_div [symmetric]) let ?r = "x - ?q * y" have r: "?r = x mod y + x div y mod 2 * y" by(simp add: q diff_mult_distrib minus_mod_eq_div_mult [symmetric])(metis diff_diff_cancel mod_less_eq_dividend mod_mult2_eq add.commute mult.commute) show ?thesis proof(cases "y \ x - ?q * y") case True with assms q have "x div y mod 2 \ 0" unfolding r by (metis Nat.add_0_right diff_0_eq_0 diff_Suc_1 le_div_geq mod2_gr_0 mod_div_trivial mult_0 neq0_conv numeral_1_eq_Suc_0 numerals(1)) hence "x div y = ?q + 1" unfolding q by simp moreover hence "x mod y = ?r - y" by simp(metis minus_div_mult_eq_mod [symmetric] diff_commute diff_diff_left mult_Suc) ultimately show ?thesis using True by(simp add: Let_def) next case False hence "x div y mod 2 = 0" unfolding r by(simp add: not_le)(metis Nat.add_0_right assms div_less div_mult_self2 mod_div_trivial mult.commute) hence "x div y = ?q" unfolding q by simp moreover hence "x mod y = ?r" by (metis minus_div_mult_eq_mod [symmetric]) ultimately show ?thesis using False by(simp add: Let_def) qed qed lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" proof(induct n) case 0 thus ?case by simp next case (Suc n) then obtain n' where "LENGTH('a) = Suc n'" by(cases "LENGTH('a)") simp_all with Suc show ?case by (simp add: unat_word_ariths bintrunc_mod2p) qed lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by (simp add: word_eq_iff word_less_def word_test_bit_def uint_div) lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply(simp only: word_arith_nat_defs word_le_nat_alt nat_div_eq_Suc_0_iff[symmetric]) apply(rule word_unat.Abs_inject) apply(simp only: unat_div[symmetric] word_unat.Rep) apply(simp add: unats_def Suc_0_lt_2p_len_of) done lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = (x >> 1) div y << 1; r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (cases x) moreover obtain m where m: "y = of_nat m" "m < 2 ^ LENGTH('a)" by (cases y) ultimately have [simp]: \unat (of_nat n :: 'a word) = n\ \unat (of_nat m :: 'a word) = m\ by (simp add: unat_def, transfer, simp_all add: take_bit_of_nat take_bit_eq_self)+ let ?q = "(x >> 1) div y << 1" let ?q' = "2 * (n div 2 div m)" have "n div 2 div m < 2 ^ LENGTH('a)" using n by (metis of_nat_inverse unat_lt2p uno_simps(2)) hence q: "?q = of_nat ?q'" using n m apply (auto simp add: shiftr_word_eq drop_bit_eq_div shiftl_t2n word_arith_nat_div uno_simps) apply (metis \unat (of_nat n) = n\ uno_simps(2)) done from assms have "m \ 0" using m by -(rule notI, simp) from n have "2 * (n div 2 div m) < 2 ^ LENGTH('a)" by(metis mult.commute div_mult2_eq minus_mod_eq_mult_div [symmetric] less_imp_diff_less of_nat_inverse unat_lt2p uno_simps(2)) moreover have "2 * (n div 2 div m) * m < 2 ^ LENGTH('a)" using n unfolding div_mult2_eq[symmetric] by(subst (2) mult.commute)(simp add: minus_mod_eq_div_mult [symmetric] diff_mult_distrib minus_mod_eq_mult_div [symmetric] div_mult2_eq) moreover have "2 * (n div 2 div m) * m \ n" - by(metis div_mult2_eq div_mult_le mult.assoc mult.commute) + by (metis div_mult2_eq dtle mult.assoc mult.left_commute) ultimately have r: "x - ?q * y = of_nat (n - ?q' * m)" and "y \ x - ?q * y \ of_nat (n - ?q' * m) - y = of_nat (n - ?q' * m - m)" using n m unfolding q by (simp_all add: word_sub_wi word_mult_def uint_nat unat_of_nat of_nat_mult [symmetric] word_of_nat[symmetric] of_nat_diff word_le_nat_alt del: of_nat_mult) (metis diff_diff_left less_imp_diff_less of_nat_diff of_nat_inverse word_of_nat) thus ?thesis using n m div_half_nat[OF \m \ 0\, of n] unfolding q by(simp add: word_le_nat_alt word_div_def word_mod_def uint_nat unat_of_nat zmod_int[symmetric] zdiv_int[symmetric] word_of_nat[symmetric])(simp add: Let_def split del: if_split split: if_split_asm) qed lemma word_test_bit_set_bits: "(BITS n. f n :: 'a :: len word) !! n \ n < LENGTH('a) \ f n" by(auto simp add: word_set_bits_def test_bit_bl word_bl.Abs_inverse word_size) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. i !! n)" by (rule word_eqI) (auto simp add: word_test_bit_set_bits test_bit.eq_norm) lemma word_and_mask_or_conv_and_mask: "n !! index \ (n AND mask index) OR (1 << index) = n AND mask (index + 1)" by(rule word_eqI)(auto simp add: word_ao_nth word_size nth_shiftl simp del: shiftl_1) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "n !! (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = 1 << LENGTH('a) - 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (1 << LENGTH('a) - 1 :: 'a word)" using assms by (simp add: uint_shiftl word_size bintrunc_shiftl del: shiftl_1) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (1 << LENGTH('a) - 1 :: 'a word))" by(simp add: uint_or) also have "\ = uint (n AND mask (LENGTH('a) - 1 + 1))" using assms by(simp only: word_and_mask_or_conv_and_mask) also have "\ = uint n" by simp finally show ?thesis . qed text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ definition word_sdiv :: "'a :: len word \ 'a word \ 'a word" (infixl "sdiv" 70) where [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" definition word_smod :: "'a :: len word \ 'a word \ 'a word" (infixl "smod" 70) where [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" lemma sdiv_smod_id: "(a sdiv b) * b + (a smod b) = a" proof - note [simp] = word_sdiv_def word_smod_def have F5: "\u::'a word. - (- u) = u" by (metis minus_minus) have F7: "\v u::'a word. u + v = v + u" by(metis add.left_commute add_0_right) have F8: "\(w::'a word) (v::int) u::int. word_of_int u + word_of_int v * w = word_of_int (u + v * sint w)" by (metis word_sint.Rep_inverse wi_hom_syms(1) wi_hom_syms(3)) have "\u. u = - sint b \ word_of_int (sint a mod u + - (- u * (sint a div u))) = a" using F5 by (metis minus_minus word_sint.Rep_inverse' mult_minus_left add.commute mult_div_mod_eq [symmetric]) hence "word_of_int (sint a mod - sint b + - (sint b * (sint a div - sint b))) = a" by (metis equation_minus_iff) hence "word_of_int (sint a mod - sint b) + word_of_int (- (sint a div - sint b)) * b = a" using F8 by(metis mult.commute mult_minus_left) hence eq: "word_of_int (- (sint a div - sint b)) * b + word_of_int (sint a mod - sint b) = a" using F7 by metis show ?thesis proof(cases "sint a < 0") case True note a = this show ?thesis proof(cases "sint b < 0") case True with a show ?thesis by simp (metis F7 F8 eq minus_equation_iff minus_mult_minus mod_div_mult_eq) next case False from eq have "word_of_int (- (- sint a div sint b)) * b + word_of_int (- (- sint a mod sint b)) = a" by (metis div_minus_right mod_minus_right) with a False show ?thesis by simp qed next case False note a = this show ?thesis proof(cases "sint b < 0") case True with a eq show ?thesis by simp next case False with a show ?thesis by simp (metis wi_hom_add wi_hom_mult add.commute mult.commute word_sint.Rep_inverse add.commute mult_div_mod_eq [symmetric]) qed qed qed text \ This algorithm implements unsigned division in terms of signed division. Taken from Hacker's Delight. \ lemma divmod_via_sdivmod: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (if 1 << (LENGTH('a) - 1) \ y then if x < y then (0, x) else (1, x - y) else let q = ((x >> 1) sdiv y) << 1; r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "1 << (LENGTH('a) - 1) \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by (cases x, cases y) (simp add: word_less_def word_mod_def) thus ?thesis using True y by(simp add: word_div_lt_eq_0) next case False obtain n where n: "y = of_nat n" "n < 2 ^ LENGTH('a)" by(cases y) have "unat x < 2 ^ LENGTH('a)" by(rule unat_lt2p) also have "\ = 2 * 2 ^ (LENGTH('a) - 1)" by(metis Suc_pred len_gt_0 power_Suc One_nat_def) also have "\ \ 2 * n" using y n by(simp add: word_le_nat_alt unat_of_nat unat_p2) finally have div: "x div of_nat n = 1" using False n by(simp add: word_div_eq_1_iff not_less word_le_nat_alt unat_of_nat) moreover have "x mod y = x - x div y * y" - by (metis add_diff_cancel2 word_mod_div_equality) + by (simp add: minus_div_mult_eq_mod) with div n have "x mod y = x - y" by simp ultimately show ?thesis using False y n by simp qed next case False note y = this obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by(cases x) hence "int n div 2 + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" by (cases "LENGTH('a)") (simp_all, simp only: of_nat_numeral [where ?'a = int, symmetric] zdiv_int [symmetric] of_nat_power [symmetric]) with y n have "sint (x >> 1) = uint (x >> 1)" by (simp add: sint_uint sbintrunc_mod2p shiftr_div_2n) (simp add: uint_nat unat_of_nat) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases "LENGTH('a)") (simp_all add: not_le word_2p_lem word_size) then have "sint y = uint y" by (simp add: sint_uint sbintrunc_mod2p) ultimately show ?thesis using y by(subst div_half_word[OF assms])(simp add: word_sdiv_def uint_div[symmetric]) qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting begin lift_definition word_of_integer :: "integer \ 'a :: len word" is word_of_int . lemma word_of_integer_code [code]: "word_of_integer n = word_of_int (int_of_integer n)" by(simp add: word_of_integer.rep_eq) end lemma word_of_int_code [code abstract]: "uint (word_of_int x :: 'a word) = x AND bin_mask (LENGTH('a :: len))" by(simp add: uint_word_of_int and_bin_mask_conv_mod) context fixes f :: "nat \ bool" begin lemma bit_set_bits_word_iff: \bit (set_bits f :: 'a word) n \ n < LENGTH('a::len) \ f n\ by (auto simp add: of_nth_def test_bit_of_bl simp flip: test_bit_word_eq) definition set_bits_aux :: \'a word \ nat \ 'a :: len word\ where \set_bits_aux w n = push_bit n w OR take_bit n (set_bits f)\ lemma set_bits_aux_conv: \set_bits_aux w n = (w << n) OR (set_bits f AND mask n)\ by (rule bit_word_eqI) (auto simp add: set_bits_aux_def shiftl_word_eq mask_eq_mask bit_and_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_mask_iff bit_set_bits_word_iff exp_eq_zero_iff) corollary set_bits_conv_set_bits_aux: \set_bits f = (set_bits_aux 0 (LENGTH('a)) :: 'a :: len word)\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_0 [simp]: \set_bits_aux w 0 = w\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_Suc [simp]: \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by (simp add: set_bits_aux_def shiftl_word_eq bit_eq_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_set_bits_word_iff) (auto simp add: bit_exp_iff not_less bit_1_iff less_Suc_eq_le exp_eq_zero_iff) lemma set_bits_aux_simps [code]: \set_bits_aux w 0 = w\ \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by simp_all end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = bin_mask (LENGTH('a))" and shift_def: "shift = 1 << LENGTH('a)" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = 1 << (LENGTH('a) - 1)" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if i' !! index then if i' - shift < least \ overflow \ i' - shift then arbitrary1 i' else word_of_int (i' - shift) else if i' < least \ overflow \ i' then arbitrary2 i' else word_of_int i')" proof - define i' where "i' = i AND mask" have "shift = mask + 1" unfolding assms by(simp add: bin_mask_p1_conv_shift) hence "i' < shift" by(simp add: mask_def i'_def int_and_le) show ?thesis proof(cases "i' !! index") case True hence unf: "i' = overflow OR i'" unfolding assms i'_def by(auto intro!: bin_eqI simp add: bin_nth_ops) have "overflow \ i'" by(subst unf)(rule le_int_or, simp add: bin_sign_and assms i'_def) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ by(auto intro!: word_eqI simp add: i'_def shift_def mask_def bin_nth_ops bin_nth_minus_p2 bin_sign_and) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False hence "i' = i AND bin_mask (LENGTH('a) - 1)" unfolding assms i'_def by(clarsimp simp add: i'_def bin_nth_ops intro!: bin_eqI)(cases "LENGTH('a)", auto simp add: less_Suc_eq) also have "\ \ bin_mask (LENGTH('a) - 1)" by(rule int_and_le) simp also have "\ < overflow" unfolding overflow_def by(simp add: bin_mask_p1_conv_shift[symmetric]) also have "least \ 0" unfolding least_def overflow_def by simp have "0 \ i'" by(simp add: i'_def mask_def bin_mask_ge0) hence "least \ i'" using \least \ 0\ by simp moreover have "word_of_int i' = (word_of_int i :: 'a word)" by(rule word_eqI)(auto simp add: i'_def bin_nth_ops mask_def) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed text \Quickcheck conversion functions\ notation scomp (infixl "\\" 60) definition qc_random_cnv :: "(natural \ 'a::term_of) \ natural \ Random.seed \ ('a \ (unit \ Code_Evaluation.term)) \ Random.seed" where "qc_random_cnv a_of_natural i = Random.range (i + 1) \\ (\k. Pair ( let n = a_of_natural k in (n, \_. Code_Evaluation.term_of n)))" no_notation scomp (infixl "\\" 60) definition qc_exhaustive_cnv :: "(natural \ 'a) \ ('a \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.exhaustive (%x. f (a_of_natural x)) d" definition qc_full_exhaustive_cnv :: "(natural \ ('a::term_of)) \ ('a \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_full_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.full_exhaustive (%(x, xt). f (a_of_natural x, %_. Code_Evaluation.term_of (a_of_natural x))) d" declare [[quickcheck_narrowing_ghc_options = "-XTypeSynonymInstances"]] definition qc_narrowing_drawn_from :: "'a list \ integer \ _" where "qc_narrowing_drawn_from xs = foldr Quickcheck_Narrowing.sum (map Quickcheck_Narrowing.cons (butlast xs)) (Quickcheck_Narrowing.cons (last xs))" locale quickcheck_narrowing_samples = fixes a_of_integer :: "integer \ 'a \ 'a :: {partial_term_of, term_of}" and zero :: "'a" and tr :: "typerep" begin function narrowing_samples :: "integer \ 'a list" where "narrowing_samples i = (if i > 0 then let (a, a') = a_of_integer i in narrowing_samples (i - 1) @ [a, a'] else [zero])" by pat_completeness auto termination including integer.lifting proof(relation "measure nat_of_integer") fix i :: integer assume "0 < i" thus "(i - 1, i) \ measure nat_of_integer" by simp(transfer, simp) qed simp definition partial_term_of_sample :: "integer \ 'a" where "partial_term_of_sample i = (if i < 0 then undefined else if i = 0 then zero else if i mod 2 = 0 then snd (a_of_integer (i div 2)) else fst (a_of_integer (i div 2 + 1)))" lemma partial_term_of_code: "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_variable p t) \ Code_Evaluation.Free (STR ''_'') tr" "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_constructor i []) \ Code_Evaluation.term_of (partial_term_of_sample i)" by (rule partial_term_of_anything)+ end lemmas [code] = quickcheck_narrowing_samples.narrowing_samples.simps quickcheck_narrowing_samples.partial_term_of_sample_def text \ The separate code target \SML_word\ collects setups for the code generator that PolyML does not provide. \ setup \Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])\ code_identifier code_module Code_Target_Word_Base \ (SML) Word and (Haskell) Word and (OCaml) Word and (Scala) Word export_code sbintrunc bin_mask in SML module_name Code end diff --git a/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy b/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy --- a/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy +++ b/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy @@ -1,3045 +1,3050 @@ section"Example: Lorenz attractor" theory Lorenz_Approximation imports "HOL-ODE-Numerics.ODE_Numerics" Result_File_Coarse begin + +lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" + by (simp add: numeral_eq_Suc) + + text \\label{sec:lorenz}\ text \TODO: move to isabelle? \ lifting_update blinfun.lifting lifting_forget blinfun.lifting lemma eventually_uniformly_on: "(\\<^sub>F x in uniformly_on T l. P x) = (\e>0. \f. (\x\T. dist (f x) (l x) < e) \ P f)" unfolding uniformly_on_def apply (subst eventually_INF) apply safe subgoal for E apply (cases "E = {}") subgoal by (auto intro!: exI[where x=1]) subgoal apply (auto simp: INF_principal_finite eventually_principal elim!: ) proof goal_cases case (1 x) from 1 have "0 < Min E" apply (subst Min_gr_iff) apply force apply force apply force done have *: "(\e\E. {f. \t\T. dist (f t) (l t) < e}) = {f. \t\T. dist (f t) (l t) < Min E}" using 1 apply (auto simp: ) apply (subst Min_gr_iff) apply force apply force apply force apply (drule bspec, assumption) apply (rule less_le_trans, assumption) apply auto done from 1 have "\f. (\x\T. dist (f x) (l x) < Min E) \ P f" unfolding * by simp then show ?case using 1(4)[rule_format, OF \0 < Min E\] by auto qed done subgoal for e apply (rule exI[where x="{e}"]) by (auto simp: eventually_principal) done lemma op_cast_image_impl[autoref_rules]: "(\x. x, op_cast_image::'a::executable_euclidean_space set \ 'b::executable_euclidean_space set) \ aform.appr_rel \ aform.appr_rel" if "DIM('a) = DIM('b)" using that apply (auto simp: aform.appr_rel_def intro!: relcompI) unfolding lv_rel_def set_rel_br by (force simp: intro!: brI dest!: brD) lemma cast_bl_blinfun_of_list[simp]: "cast_bl (blinfun_of_list xs::'a \\<^sub>L 'a) = (blinfun_of_list xs::'b\\<^sub>L'b)" if "DIM('a::executable_euclidean_space) = DIM('b::executable_euclidean_space)" using that apply (auto simp: cast_bl_rep intro!: blinfun_eqI) by (auto simp: blinfun_of_list_def blinfun_of_matrix_apply linear_sum linear.scaleR sum_Basis_sum_nth_Basis_list linear_cast) lemma cast_idem[simp]: "cast x = x" by (auto simp: cast_def) lemma cast_bl_idem[simp]: "cast_bl x = x" by (auto simp: cast_bl_rep intro!: blinfun_eqI) lemma op_cast_eucl1_image_impl[autoref_rules]: "(\x. x, op_cast_eucl1_image::'a::executable_euclidean_space c1_info set \ 'b::executable_euclidean_space c1_info set) \ aform.appr1_rel \ aform.appr1_rel" if "DIM_precond TYPE('a) D" "DIM_precond TYPE('b) D" using that proof (auto, goal_cases) case (1 a b a') then show ?case apply (auto simp: aform.appr1_rel_br set_rel_br br_def) subgoal for w x y z apply (auto simp: aform.c1_info_of_appr_def cast_eucl1_def aform.c1_info_invar_def split: option.splits) apply (rule image_eqI) apply (rule cast_eucl_of_list, force, force simp: Joints_imp_length_eq, force) subgoal for s t apply (rule image_eqI[where x="t"]) supply [simp del] = eucl_of_list_take_DIM apply (auto simp: flow1_of_list_def) apply (subst cast_eucl_of_list) subgoal by simp subgoal by (auto dest!: Joints_imp_length_eq simp: power2_eq_square flow1_of_list_def[abs_def]) subgoal by simp done done subgoal for w x apply (rule image_eqI[where x="cast_eucl1 (w, x)"]) apply (auto simp: aform.c1_info_of_appr_def cast_eucl1_def aform.c1_info_invar_def split: option.splits) apply (rule image_eqI) apply (rule cast_eucl_of_list, force, force simp: Joints_imp_length_eq, force) subgoal for s t apply (rule image_eqI[where x="t"]) supply [simp del] = eucl_of_list_take_DIM apply (auto simp: flow1_of_list_def) apply (subst cast_eucl_of_list) subgoal by simp subgoal by (auto dest!: Joints_imp_length_eq simp: power2_eq_square flow1_of_list_def[abs_def]) subgoal by simp done done done qed lemma less_3_iff: "i < (3::nat) \ i = 0 \ i = 1 \ i = 2" by arith definition mat3_of_vec::"R3 \ real^3^3" where "mat3_of_vec x = (let xs = list_of_eucl x in eucl_of_list [xs!0,0,0, xs!1,0,0, xs!2,0,0])" lemma ll3: "{..<3} = {0,1,2::nat}" by (auto simp: less_3_iff) lemma mat3_of_vec: "cast (mat3_of_vec x *v eucl_of_list [1, 0, 0]) = x" by (auto simp: mat3_of_vec_def eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list linear_sum linear_cast linear.scaleR ll3 linear_add Basis_list_R3 inner_prod_def prod_eq_iff) primrec bisect_form where "bisect_form p f xs l u 0 = (l, u)" | "bisect_form p f xs l u (Suc n) = (let m = (l + u)/2 in if approx_form_aform p (f m) xs then bisect_form p f xs l m n else bisect_form p f xs m u n)" text \This should prove that the expansion estimates are sufficient.\ lemma expansion_main: "expansion_main (coarse_results) = Some True" by eval context includes floatarith_notation begin definition "matrix_of_degrees2\<^sub>e = (let e = Var 2; ur = Rad_of (Var 0); vr = Rad_of (Var 1); x1 = Cos ur; x2 = Cos vr; y1 = Sin ur; y2 = Sin vr in [x1 + (e * (x2 - x1)), 0, 0, y1 + (e * (y2 - y1)), 0, 0, 0, 0, 0])" definition "matrix_of_degrees2 u v = approx_floatariths 30 matrix_of_degrees2\<^sub>e (aforms_of_ivls [u, v, 0] [u, v, 1])" text \following \vector_field.h\ / \vector_field.cc\\ abbreviation "S \ 10::real" abbreviation "B \ 8/3::real" abbreviation "TEMP \ sqrt((S + 1) * (S + 1) + 4 * S * (28 - 1))" abbreviation "K1 \ S / TEMP" abbreviation "K2 \ (S - 1 + TEMP) / (2 * S)" abbreviation "K3 \ (S - 1 - TEMP) / (2 * S)" abbreviation "E1 \ (- (S + 1) + TEMP) / 2" abbreviation "E2 \ (- (S + 1) - TEMP) / 2" abbreviation "E3 \ - B" abbreviation "C1 \ \X. X ! 0 + X ! 1" abbreviation "C2 \ \X. K1 * C1 X * X ! 2" schematic_goal lorenz_fas: "[E1 * X!0 - C2 X, E2 * X!1 + C2 X, E3 * X!2 + C1 X * (K2 * X!0 + K3 * X!1)] = interpret_floatariths ?fas X" by (reify_floatariths) concrete_definition lorenz_fas uses lorenz_fas end interpretation lorenz: ode_interpretation true_form UNIV lorenz_fas "\(X0, X1, X2). (E1 * X0 - K1 * (X0 + X1) * X2, E2 * X1 + K1 * (X0 + X1) * X2, E3 * X2 + (X0 + X1) * (K2 * X0 + K3 * X1))::real*real*real" "d::3" for d by standard (auto simp: lorenz_fas_def less_Suc_eq_0_disj nth_Basis_list_prod Basis_list_real_def mk_ode_ops_def eucl_of_list_prod power2_eq_square inverse_eq_divide intro!: isFDERIV_I) value [code] "length (slp_of_fas lorenz_fas)" definition "mig_aform p x = mig_componentwise (Inf_aform' p x) (Sup_aform' p x)" context includes floatarith_notation begin definition "mig_aforms p x = real_of_float ((lower o the) ((approx p (Norm (map (Num o float_of o (mig_aform p)) x))) []))" definition "column_of_c1_info x N j = (map (\i. the (snd x) ! i) (map (\i. i * N + j) [0..e b))) (rotate_z_fa (Rad_of (R\<^sub>e a)))) xs)" definition "perspective_projection_aforms xs = the (approx_slp_outer 30 3 (rotate_zx_slp (-30) (-60) (map Var [0..<3])) xs)" definition "print_lorenz_aform print_fun cx cy cz ci cd1 cd2 = (\a b. let (s1, n) = ((-6), False); _ = print_fun (String.implode (''# gen(''@ show a@''): ''@ shows_aforms_hr (b) '''' @ ''\'')); _ = print_fun (String.implode (''# box(''@ show a@''): ''@ shows_box_of_aforms_hr (b) '''' @ ''\'')); ((x0, y0, z0), (x1, y1, z1)) = case (map (Inf_aform' 30) (take 3 b), map (Sup_aform' 30) (take 3 b)) of ([x0, y0, z0], [x1, y1, z1]) \ ((x0, y0, z0), (x1, y1, z1)); _ = print_fun (String.implode (shows_segments_of_aform 0 1 b ((shows cx o shows_space o shows z0 o shows_space o shows z1)'''') ''\'')); _ = print_fun (String.implode (shows_segments_of_aform 0 2 b ((shows cy o shows_space o shows y0 o shows_space o shows y1)'''') ''\'')); _ = print_fun (String.implode (shows_segments_of_aform 1 2 b ((shows cz o shows_space o shows x0 o shows_space o shows x1)'''') ''\'')); PS = perspective_projection_aforms b; _ = print_fun (String.implode (shows_segments_of_aform 0 1 PS ((shows ci o shows_space o shows (fst (PS ! 2)) o shows_space o shows (fst (b ! 2))) '''') ''\'')) in if \ a \ length b > 10 then print_fun (String.implode (shows_aforms_vareq 3 [(0, 1), (0, 2), (1, 2)] [0..<3] cd1 cd2 (FloatR 1 s1 * (if n then real_divl 30 1 (max (mig_aforms 30 (map (\i. b ! i) [3,6,9])) (mig_aforms 30 (map (\i. b ! i) [4,7,10]))) else 1)) \ \always length \2^s!\\ ''# no C1 info'' b '''')) else ())" definition "print_lorenz_aform_std print_fun = print_lorenz_aform print_fun ''0x000001'' ''0x000002'' ''0x000012'' ''0x000003'' [''0xa66f00'', ''0x06266f'', ''0xc60000''] [''0xffaa00'', ''0x1240ab'', ''0xc60000'']" definition "lorenz_optns print_funo = (let pf = the_default (\_ _. ()) (map_option print_lorenz_aform_std print_funo); tf = the_default (\_ _. ()) (map_option (\print_fun a b. let _ = print_fun (String.implode (''# '' @ a @ ''\'')) in case b of Some b \ (print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\''))) | None \ ()) print_funo) in \ precision = 30, adaptive_atol = FloatR 1 (-30), adaptive_rtol = FloatR 1 (-30), method_id = 2, start_stepsize = FloatR 1 (- 8), iterations = 40, halve_stepsizes = 10, widening_mod = 40, rk2_param = FloatR 1 0, default_reduce = correct_girard 30 50 25, printing_fun = pf, tracing_fun = tf \)" definition lorenz_optns' where "lorenz_optns' pf m N rk2p a = lorenz_optns pf \ default_reduce := correct_girard 30 m N, rk2_param := rk2p, adaptive_atol := a, adaptive_rtol := a \" definition mirror_irects where "mirror_irects = map (\irect. case irect of [i, j, k] \ [if j < 0 then - i else i , abs j, k] | irect \ irect)" definition "print_irects irects = (let _ = map (\is. let _ = map (\j. let _ = print (String.implode (show j)) in print (STR '' '')) is in print (STR ''\'')) irects in ())" abbreviation "aforms_of_ivl \ \x. aforms_of_ivls (fst x) (snd x)" definition "conefield_propagation\<^sub>e = ([Deg_of (Arctan (Var (6) / Var (3))), Deg_of (Arctan (Var (7) / Var (4))), Min (Norm [Var(3), Var (6), Var (9)]) (Norm [Var(4), Var (7), Var (10)])])" definition "conefield_propagation DX = approx_floatariths 30 conefield_propagation\<^sub>e DX" definition "conefield_propagation_slp = slp_of_fas conefield_propagation\<^sub>e" lemma [simp]: "length conefield_propagation_slp = 51" by eval definition op_with_unit_matrix :: "'a::real_normed_vector \ 'a \ 'a \\<^sub>L 'a" where "op_with_unit_matrix X = (X, 1\<^sub>L)" context includes blinfun.lifting begin lemma matrix_vector_mult_blinfun_works[simp]: "matrix e *v g = e g" for e::"(real^'n) \\<^sub>L (real^'m)" by transfer (simp add: bounded_linear_def matrix_works) end lemma length_conefield_propagation\<^sub>e[simp]: "length conefield_propagation\<^sub>e = 3" by (simp add: conefield_propagation\<^sub>e_def) lemma interpret_floatariths_conefield_propagation: "interpret_floatariths conefield_propagation\<^sub>e (list_of_eucl (vec1_of_flow1 (xDX::(real^3) \ ((real^3)\\<^sub>L(real^3))))) = (let DX = snd xDX; DXu = DX (eucl_of_list [1, 0, 0]); DXv = DX (eucl_of_list [0, 1, 0]) in [deg_of (arctan (vec_nth DXu 1 / vec_nth DXu 0)), deg_of (arctan (vec_nth DXv 1 / vec_nth DXv 0)), min (norm DXu) (norm DXv)] )" apply (auto simp: conefield_propagation\<^sub>e_def Let_def interpret_mvmult_nth[where 'n=3 and 'm=3] inverse_eq_divide vec1_of_flow1_def nth_append) apply (auto simp: matrix_inner_Basis_list ) apply (auto simp: interpret_floatarith_norm[where 'a="real ^ 3"] einterpret_mvmult_fa[where 'n=3 and 'm=3] matrix_inner_Basis_list nth_append) by (auto simp: matrix_def axis_eq_eucl_of_list eucl_of_list_012) definition "conefield_bounds_form l u = (fold Conj [ Less (-90) (N\<^sub>r l), LessEqual (N\<^sub>r l) (N\<^sub>r u), LessEqual (Var 9) (0), LessEqual 0 (Var 9), Less (N\<^sub>r u) (90), Less 0 (Var 3), AtLeastAtMost (Var 6) (Tan (Rad_of (N\<^sub>r l)) * Var 3) (Tan (Rad_of (N\<^sub>r u)) * Var 3)] true_form)" definition "contract_angles X i = (snd (bisect_form 30 (\x. conefield_bounds_form x (89)) X 89 (-89) i), snd (bisect_form 30 (conefield_bounds_form (-89)) X (-89) 89 i))" definition "approx_conefield_bounds (DX::(R3 \ (R3 \\<^sub>L R3)) set) l u = do { let DX = (cast_eucl1 ` DX::3 eucl1 set); DXo \ aform.vec1rep DX; DX \ (case DXo of None \ do { let _ = aform.print_msg (''# approx_conefield_bounds failed DXo...''); SUCCEED } | Some DX \ RETURN DX); let _ = aform.trace_set (''# approx_conefield_bounds DX: '') (Some DX); approx_form_spec (conefield_bounds_form l u) (list_of_eucl ` DX) }" lemma [autoref_rules]: includes autoref_syntax shows "(conefield_bounds_form, conefield_bounds_form) \ Id \ Id \ Id " by auto lemma [autoref_rules_raw]: "DIM_precond TYPE((real, 3) vec \ ((real, 3) vec, 3) vec) 12" "DIM_precond TYPE(R3) 3" "DIM_precond TYPE((real, 3) vec) 3" by auto schematic_goal approx_conefield_bounds_impl: includes autoref_syntax fixes optns::"real aform numeric_options" assumes [autoref_rules]: "(DXi, DX) \ aform.appr1_rel" assumes [autoref_rules]: "(li, l) \ Id" assumes [autoref_rules]: "(ui, u) \ Id" notes [autoref_rules] = aform.print_msg_impl[where optns = optns] aform.ivl_rep_of_set_autoref[where optns = optns] aform.transfer_operations(12)[where optns = optns] aform.approx_euclarithform[where optns=optns] aform.trace_set_impl[of optns] shows "(nres_of ?r, approx_conefield_bounds $ DX $ l $ u) \ \bool_rel\nres_rel" unfolding autoref_tag_defs unfolding approx_conefield_bounds_def including art by autoref_monadic concrete_definition approx_conefield_bounds_impl for optns li ui DXi uses approx_conefield_bounds_impl lemmas [autoref_rules] = approx_conefield_bounds_impl.refine context includes autoref_syntax begin lemma [autoref_rules]: "(real_of_ereal, real_of_ereal) \ ereal_rel \ rnv_rel" "(\, \) \ ereal_rel" by auto end lemma interpret_form_true_form[simp]: "interpret_form true_form \ \_. True" by (force simp: true_form_def intro!: eq_reflection) lemma interpret_form_conefield_bounds_form_list: "interpret_form (conefield_bounds_form L U) [x, y, z, ux, vx, wx, uy, vy, wy, uz, vz, wz] \ (0 < ux \ -90 < L \ L \ U \ U < 90 \ uz = 0 \ uy \ tan (rad_of U) * ux \ tan (rad_of L) * ux \ uy)" if "U \ float" "L \ float" "e \ float" "em \ float" using that by (auto simp: conefield_bounds_form_def L2_set_def) lemma list_of_eucl_eucl1_3: includes vec_syntax shows "(list_of_eucl (vec1_of_flow1 (xDX::(real^3) \ ((real^3)\\<^sub>L(real^3))))) = (let (x, DX) = xDX; DXu = DX (eucl_of_list [1, 0, 0]); DXv = DX (eucl_of_list [0, 1, 0]); DXw = DX (eucl_of_list [0, 0, 1]) in [x $ 0, x $ 1, x $ 2, vec_nth DXu 0, vec_nth DXv 0, vec_nth DXw 0, vec_nth DXu 1, vec_nth DXv 1, vec_nth DXw 1, vec_nth DXu 2, vec_nth DXv 2, vec_nth DXw 2])" apply (auto simp: matrix_inner_Basis_list Let_def vec1_of_flow1_def concat_map_map_index less_Suc_eq_0_disj list_of_eucl_matrix eval_nat_numeral aform.inner_Basis_eq_vec_nth intro!: nth_equalityI split: prod.splits) by (auto simp: matrix_def axis_eq_eucl_of_list eucl_of_list_012) lemma interpret_form_conefield_bounds_form: "interpret_form (conefield_bounds_form L U) (list_of_eucl (vec1_of_flow1 (xDX::(real^3) \ ((real^3)\\<^sub>L(real^3))))) = (let DX = snd xDX; DXu = DX (eucl_of_list [1, 0, 0]); DXv = DX (eucl_of_list [0, 1, 0]); uz = vec_nth DXu 2; uy = vec_nth DXu 1; ux = vec_nth DXu 0; vy = vec_nth DXv 1; vx = vec_nth DXv 0 in ux > 0 \ -90 < L \ L \ U \ U < 90 \ uz = 0 \ (uy / ux) \ {tan (rad_of L) .. tan (rad_of U)} )" if "L \ float" "U \ float" using that unfolding list_of_eucl_eucl1_3 Let_def by (auto split: prod.splits simp: interpret_form_conefield_bounds_form_list divide_simps) lemma approx_conefield_bounds_cast: "approx_conefield_bounds DX L U \ SPEC (\b. b \ (\(x, dx) \ cast_eucl1 ` DX::3 eucl1 set. let u' = dx (eucl_of_list [1, 0, 0]) in vec_nth u' 1 / vec_nth u' 0 \ {tan (rad_of L) .. tan (rad_of U)} \ vec_nth u' 2 = 0 \ vec_nth u' 0 > 0 \ -90 < L \ L \ U \ U < 90) )" if "L \ float" "U \ float" unfolding approx_conefield_bounds_def apply refine_vcg apply (auto simp: env_len_def ) subgoal for a b c apply (drule bspec, assumption) unfolding interpret_form_conefield_bounds_form[OF that] by (auto simp: Let_def divide_simps) done lemma approx_conefield_bounds[le, refine_vcg]: "approx_conefield_bounds DX l u \ SPEC (\b. b \ (\(x, dx) \ DX::R3 c1_info set. let (u'1, u'2, u'3) = dx ((1, 0, 0)) in u'2 / u'1 \ {tan (rad_of l) .. tan (rad_of u)} \ u'3 = 0 \ u'1 > 0 \ -90 < l \ l \ u \ u < 90) )" if "l \ float" "u \ float" apply (rule approx_conefield_bounds_cast[le, OF that]) apply (auto dest!: bspec simp: Let_def split: prod.splits) by (auto simp: cast_eucl1_def cast_def) schematic_goal MU\<^sub>e: "-E2 / E1 = interpret_floatarith ?fas []" by (reify_floatariths) concrete_definition MU\<^sub>e uses MU\<^sub>e schematic_goal NU\<^sub>e: "-E3 / E1 = interpret_floatarith ?fas []" by (reify_floatariths) concrete_definition NU\<^sub>e uses NU\<^sub>e definition "approx_ivls p fas xs = do { let xs = ivls_of_aforms p xs; res \ those (map (\f. approx p f xs) fas); Some (map (real_of_float o lower) res, map (real_of_float o upper) res) }" definition "deform p t exit XDX = (case XDX of (lu, (X, DX)) \ let d = ldec 0.1; d = if exit then (1 - real_of_float (lb_sqrt 30 (1 - 2 * float_of (d)))) else d; sd = real_of_float ((float_of (d*d))); C0Deform = aforms_of_ivls [-sd,-sd,-sd] [sd, sd, sd]; result = msum_aforms' X C0Deform in (lu, case DX of None \ (result, None) | Some DX \ let C1_norm = 2 * d; C1_norm = if exit then real_divr 30 C1_norm (1 - C1_norm) else C1_norm; l = -C1_norm; u = C1_norm; D_M = aforms_of_ivls [1 + l,0,0, 0,1 + l,0, 0,0,1 + l] [1 + u,0,0, 0,1 + u,0, 0,0,1 + u]; (ri, ru) = the (approx_ivls p (mmult_fa 3 3 3 (map Var [0..<9]) (map Var [9..<18])) (D_M @ DX)); Dresult = aforms_of_ivls ri ru; resultDresult = product_aforms result Dresult in (take 3 resultDresult, Some (drop 3 resultDresult))))" definition "ivls_of_aforms' p r = (map (Inf_aform' p) r, map (Sup_aform' p) r)" definition "compute_half_exit p t XDX = (case XDX of ((l, u::ereal), (X, DX)) \ let \ \ASSERTING that \Y\ straddles zero\ (x0, y0, _) = case map (Inf_aform' p) X of [x,y,z] \ (x, y, z); (x1, y1, _) = case map (Sup_aform' p) X of [x,y,z] \ (x, y, z); splitting = x0 = 0 \ x1 = 0; sign_x = if (x0 + x1) / 2 > 0 then 1 else -1; mag_x = max (abs x0) (abs x1); sign_x\<^sub>e = N\<^sub>r sign_x; exit_rad\<^sub>e = N\<^sub>r (ldec 0.1); X\<^sub>e = Var (0); Y\<^sub>e = Var (1); Z\<^sub>e = Var (2); max_x_over_r\<^sub>e = N\<^sub>r mag_x / exit_rad\<^sub>e; abs_x_over_r\<^sub>e = (Abs X\<^sub>e) / exit_rad\<^sub>e; result = (if splitting then let result = (the (approx_floatariths p [sign_x\<^sub>e * exit_rad\<^sub>e, Y\<^sub>e * Powr (max_x_over_r\<^sub>e) MU\<^sub>e, Z\<^sub>e * Powr (max_x_over_r\<^sub>e) NU\<^sub>e] X)); (ir, sr) = ivls_of_aforms' p result in aforms_of_ivls (ir[1:=min 0 (ir!1), 2:=min 0 (ir!2)]) (sr[1:=max 0 (sr!1), 2:=max 0 (sr!2)]) else the (approx_floatariths p [sign_x\<^sub>e * exit_rad\<^sub>e, Y\<^sub>e * Powr (abs_x_over_r\<^sub>e) MU\<^sub>e, Z\<^sub>e * Powr (abs_x_over_r\<^sub>e) NU\<^sub>e] X)); _ = () in ((l::ereal, \::ereal), (case DX of None \ (result, None) | Some DX \ let ux\<^sub>e = Var (3); uy\<^sub>e = Var (6); P21\<^sub>e = if splitting then (MU\<^sub>e / exit_rad\<^sub>e) * Y\<^sub>e * sign_x\<^sub>e * Powr (max_x_over_r\<^sub>e) (MU\<^sub>e - 1) \ \No need for \Hull(0)\ because \y\ straddles zero\ else (MU\<^sub>e / exit_rad\<^sub>e) * Y\<^sub>e * sign_x\<^sub>e * Powr (abs_x_over_r\<^sub>e) (MU\<^sub>e - 1); P22\<^sub>e = if splitting then Powr (max_x_over_r\<^sub>e) MU\<^sub>e else Powr (abs_x_over_r\<^sub>e) MU\<^sub>e; P31\<^sub>e = if splitting then sign_x\<^sub>e * (NU\<^sub>e / exit_rad\<^sub>e) * Z\<^sub>e * Powr (max_x_over_r\<^sub>e) (NU\<^sub>e - 1) \ \No need for \Hull(\)\ because scaling afterwards\ else sign_x\<^sub>e * (NU\<^sub>e / exit_rad\<^sub>e) * Z\<^sub>e * Powr (abs_x_over_r\<^sub>e) (NU\<^sub>e - 1); ry = (P21\<^sub>e * ux\<^sub>e) + (P22\<^sub>e * uy\<^sub>e); rz = P31\<^sub>e * ux\<^sub>e; (iDr, sDr) = the (approx_ivls p ([0, 0, 0, ry, 0, 0, rz, 0, 0]) (X @ DX)); Dresult = aforms_of_ivls (iDr[3:=min 0 (iDr!3)]) (sDr[3:=max 0 (sDr!3)]); resultDresult = product_aforms result Dresult in (take 3 resultDresult, Some (drop 3 resultDresult)) )))" fun list3 where "list3 [a,b,c] = (a, b, c)" definition "split_x n x0 y0 z0 x1 y1 z1 = (let elem = (\(x0, x1). aforms_of_ivls [x0, y0, z0] [x1, y1, z1]); coord = (\x0 n i. i * x0 * FloatR 1 (-int n)); us = map (coord x0 n) (rev [0..<(2^n)]) @ map (coord x1 n) [Suc 0..X'. ((l, u), (X', DX'))) X's; Xes = map (compute_half_exit p t) XDX's; Xlumpies = map (deform p t True) Xes in Xlumpies)" definition "cube_enteri = (map ldec [-0.1, -0.00015, 0.1, 0.8,0,0, 0.0005,0,0, 0,0,0], map udec [ 0.1, 0.00015, 0.1, 1.7,0,0, 0.002,0,0, 0,0,0])" definition "cube_enter = set_of_ivl (pairself eucl_of_list cube_enteri)" value [code] "println ((show) (map (ivls_of_aforms' 100 o list_of_appr1e_aform) (compute_cube_exit 30 (FloatR 1 (-10)) ((ereal 1, ereal 1), (aforms_of_ivls (take 3 (fst cube_enteri)) (take 3 (snd cube_enteri)), Some (aforms_of_ivls (drop 3 (fst cube_enteri)) (drop 3 (snd cube_enteri))))))))" definition "cube_exiti = [(aforms_of_ivls (map ldec [-0.12, -0.024, -0.012]) (map udec [-0.088, 0.024, 0.13]), Some (aforms_of_ivls (map ldec [0,0,0, -0.56,0,0, -0.6,0,0]) (map udec [0,0,0, 0.56,0,0, -0.08,0,0]))), (aforms_of_ivls (map ldec [ 0.088, -0.024, -0.012]) (map udec [ 0.12, 0.024, 0.13]), Some (aforms_of_ivls (map ldec [0,0,0, -0.53,0,0, 0.08,0,0]) (map udec [0,0,0, 0.56,0,0, 0.6,0,0])))]" definition "cube_exitv = aform.c1_info_of_apprs cube_exiti" lemma cube_enteri[autoref_rules]: "(cube_enteri, cube_enter::'a set) \ lvivl_rel" if "DIM_precond TYPE('a::executable_euclidean_space) 12" using that by (auto simp: cube_enteri_def cube_enter_def set_of_ivl_def intro!: brI lv_relivl_relI) lemma cube_exiti[autoref_rules]: "(cube_exiti, cube_exitv::'n eucl1 set) \ clw_rel aform.appr1_rel" if "DIM_precond TYPE('n::enum rvec) 3" unfolding cube_exitv_def cube_exiti_def apply (rule aform.clw_rel_appr1_relI) using that by (auto simp: aform.c1_info_invar_def power2_eq_square) definition "lorenz_interrupt (optns::real aform numeric_options) b (eX::3 eucl1 set) = do { ((el, eu), X) \ scaleR2_rep eX; let fX = fst ` X; fentry \ op_image_fst_ivl (cube_enter::3 vec1 set); interrupt \ aform.op_subset (fX:::aform.appr_rel) fentry; (ol, ou) \ ivl_rep fentry; aform.CHECKs (ST ''asdf'') (0 < el \ ol \ ou); let _ = (if b then aform.trace_set (ST ''Potential Interrupt: '') (Some fX) else ()); let _ = (if b then aform.trace_set (ST ''With: '') (Some ({ol .. ou::3 rvec}:::aform.appr_rel)) else ()); if \b \ \interrupt then RETURN (op_empty_coll, mk_coll eX) else do { vX \ aform.vec1rep X; let _ = (if b then aform.trace_set1e (ST ''Actual Interrupt: '') (Some eX) else ()); let l = (eucl_of_list [-1/2/2,-1/2/2,-1/2/2]::3 rvec); let u = eucl_of_list [1/2/2, 1/2/2, 1/2/2]; ASSERT (l \ u); let CX = mk_coll ({l .. u}:::aform.appr_rel); (C0::3 eucl1 set) \ scaleRe_ivl_coll_spec el eu (fst ` cube_exitv \ UNIV); (C1::3 eucl1 set) \ scaleRe_ivl_coll_spec el eu (cube_exitv); case vX of None \ RETURN (CX, C0) | Some vX \ do { b \ aform.op_subset vX cube_enter; aform.CHECKs (ST ''FAILED: TANGENT VECTORs are not contained'') b; RETURN (CX, C1) } } }" lemma [autoref_rules]: includes autoref_syntax shows "(real_of_int, real_of_int) \ int_rel \ rnv_rel" "(ldec, ldec) \ Id \ rnv_rel" "(udec, udec) \ Id \ rnv_rel" by auto schematic_goal lorenz_interrupti: includes autoref_syntax assumes[autoref_rules]: "(bi, b) \ bool_rel" "(Xi, X::3 eucl1 set) \ aform.appr1e_rel" "(optnsi, optns) \ Id" shows "(nres_of ?r, lorenz_interrupt optns b X) \ \clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel\nres_rel" unfolding lorenz_interrupt_def including art by autoref_monadic concrete_definition lorenz_interrupti for optnsi1 bi Xi uses lorenz_interrupti[where optnsi = "optnsi" and optnsa = "\_ _ _ _ _ _ _ _. optnsi" and optnsb = "\_ _ _ _ _ _ _ _ _. optnsi" and optnsc = "\_ _ _ _ _ _ _ _ _ _ _. optnsi" and optnsd = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. optnsi" and optnse = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. optnsi" and optnsf = "\_ _ _ _ _ _ _ _ _. optnsi" and optns = "\_ _ _ _ _. optnsi" for optnsi] lemma lorenz_interrupti_refine[autoref_rules]: includes autoref_syntax shows "(\optnsi bi Xi. (lorenz_interrupti optnsi bi Xi), lorenz_interrupt) \ num_optns_rel \ bool_rel \ aform.appr1e_rel \ \clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel\dres_nres_rel" using lorenz_interrupti.refine by (auto simp: nres_rel_def dres_nres_rel_def) definition "(large_cube::R3 set) = {-1/4 .. 1/4} \ {-1/4 .. 1/4} \ {-1/4 .. 1/4}" definition "cube_entry = (cast_eucl1 ` (flow1_of_vec1 ` cube_enter::3 eucl1 set)::R3 c1_info set)" definition "cube_exit = (cast_eucl1 ` (cube_exitv::3 eucl1 set)::R3 c1_info set)" text \protect locale parameters\ lemma flow0_cong[cong]: "auto_ll_on_open.flow0 ode X = auto_ll_on_open.flow0 ode X" by (auto simp:) lemma existence_ivl0_cong[cong]: "auto_ll_on_open.existence_ivl0 ode X = auto_ll_on_open.existence_ivl0 ode X" by (auto simp:) lemma Dflow_cong[cong]: "c1_on_open_euclidean.Dflow ode ode_d X = c1_on_open_euclidean.Dflow ode ode_d X" by (auto simp:) lemma flowsto_cong[cong]: "c1_on_open_euclidean.flowsto ode ode_d D = c1_on_open_euclidean.flowsto ode ode_d D" by (auto simp:) lemma poincare_mapsto_cong[cong]: "c1_on_open_euclidean.poincare_mapsto ode X = c1_on_open_euclidean.poincare_mapsto ode X" by (auto simp:) lemma returns_to_cong[cong]: "auto_ll_on_open.returns_to ode X = auto_ll_on_open.returns_to ode X" by (auto simp:) lemma return_time_cong[cong]: "auto_ll_on_open.return_time ode X = auto_ll_on_open.return_time ode X" by (auto simp: ) lemma poincare_map_cong[cong]: "auto_ll_on_open.poincare_map ode X = auto_ll_on_open.poincare_map ode X" by (auto simp: ) lemma eq_nth_iff_index: "distinct xs \ n < length xs \ i = xs ! n \ index xs i = n" using index_nth_id by fastforce lemma cast_in_BasisI: "(cast i::'a) \ Basis" if "(i::'c) \ Basis""DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" using that by (auto simp: cast_def eucl_of_list_nth inner_Basis if_distrib if_distribR eq_nth_iff_index cong: if_cong) lemma cast_le_iff: "(cast (x::'a)::'c) \ y \ x \ cast y" if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" apply (auto simp: eucl_le[where 'a='a] eucl_le[where 'a='c] dest!: bspec intro!: ) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) apply (metis cast_eqI2 cast_inner that) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) apply (metis cast_eqI2 cast_inner that) done lemma cast_le_cast_iff: "(cast (x::'a)::'c) \ cast y \ x \ y" if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" apply (auto simp: eucl_le[where 'a='a] eucl_le[where 'a='c] dest!: bspec intro!: ) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) apply (rule cast_in_BasisI, assumption) apply (auto simp: that) by (metis cast_eqI2 cast_inner that) lemma cast_image_Icc[simp]: "cast ` {a .. b::'c} = {cast a .. cast b::'a}" if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)" using that apply (auto simp: cast_le_iff dest!:) subgoal for x apply (rule image_eqI[where x="cast x"]) by (auto simp: cast_le_iff) done lemma cast_eucl1_image_scaleR2: "cast_eucl1 ` scaleR2 l u X = scaleR2 l u (cast_eucl1 ` (X::'b c1_info set)::'a c1_info set)" if "DIM('a::executable_euclidean_space) = DIM('b::executable_euclidean_space)" using that apply (auto simp: scaleR2_def image_def vimage_def cast_eucl1_def linear.scaleR linear_cast_bl) apply force+ apply (rule exI conjI)+ apply assumption apply (rule exI conjI)+ apply assumption apply (rule bexI) prefer 2 apply assumption apply force by (auto simp: linear.scaleR linear_cast_bl) lemma scaleR2_diff_prod2: "scaleR2 d e (X) - Y \ UNIV = scaleR2 d e (X - Y \ UNIV)" by (force simp: scaleR2_def vimage_def image_def) end lemma (in c1_on_open_euclidean) flowsto_scaleR2I: "flowsto (scaleR2 d e X0) T (CX \ UNIV) (scaleR2 d e Y)" if "flowsto (X0) T (CX \ UNIV) (Y)" using that apply (auto simp: flowsto_def scaleR2_def) apply (drule bspec, assumption) apply auto apply (rule bexI) prefer 2 apply assumption apply auto subgoal for x a b h by (auto intro!: image_eqI[where x="(x, (flow0 a h, Dflow a h o\<^sub>L b))"] blinfun_eqI simp: blinfun.bilinear_simps) done definition "aforms_of_resultrect x0 x1 y0 y1 = aforms_of_ivl (ivl_of_resultrect x0 x1 y0 y1)" definition "flatten_varveq x = fst x @ the_default [] (snd x)" (* LessEqual (N\<^sub>r e) (N\<^sub>r em *\<^sub>e floatarith.Min (norm\<^sub>e [(3)\<^sub>e, (6)\<^sub>e, (9)\<^sub>e]) (norm\<^sub>e [(4)\<^sub>e, (7)\<^sub>e, (10)\<^sub>e])) *) derive "show" ereal definition \::"(real*real*real) set" where "\ = {(-6, -6, 27) .. (6, 6, 27)}" definition \\<^sub>l\<^sub>e::"(real*real*real) set" where "\\<^sub>l\<^sub>e = {(x, y, z). z \ 27}" definition "results = symmetrize coarse_results" definition "results_at x = {res \ set results. x \ source_of_res res}" text \a part of the stable manifold (up to the (backward) first intersection with \\\)\ definition \::"(real*real*real) set" where "\ = {x. {0..} \ lorenz.existence_ivl0 x \ (\t>0. lorenz.flow0 x t \ \) \ (lorenz.flow0 x \ 0) at_top}" definition "\\<^sub>i intr = (if intr then \ else {})" definition "\\<^sub>i\<^sub>v intr = cast ` (\\<^sub>i intr)" definition "sourcei_of_res res = source_of_res res - (\\<^sub>i (invoke_nf res))" definition "resultsi_at x = {res \ set results. x \ sourcei_of_res res}" definition "N = \(source_of_res ` (set results))" definition "\ x = \(conefield_of_res ` (results_at x))" definition "R = lorenz.poincare_map \" definition "DR x = frechet_derivative (lorenz.poincare_map \) (at x within \\<^sub>l\<^sub>e)" definition "\ x = Min (expansion ` results_at x)" definition "\\<^sub>p x = Min (preexpansion ` results_at x)" abbreviation returns_to (infixl "returns'_to" 50) where "(x returns_to P) \ lorenz.returns_to P x" lemma closed_\[intro, simp]: "closed \" by (auto simp: \_def) lemma \_stable: "lorenz.stable_on (- \) \" unfolding lorenz.stable_on_def proof (intro allI impI) fix t x0 assume outside: "\s\{0<..t}. lorenz.flow0 x0 s \ - \" assume assms: "lorenz.flow0 x0 t \ \" "t \ lorenz.existence_ivl0 x0" "0 < t" from assms have *: "{0..} \ lorenz.existence_ivl0 (lorenz.flow0 x0 t)" "(lorenz.flow0 (lorenz.flow0 x0 t) \ 0) at_top" by (auto simp: \_def) have nonneg_exivl: "s \ lorenz.existence_ivl0 x0" if "s \ 0" for s proof (cases "s \ t") case True then show ?thesis using \0 \ s\ assms(2) lorenz.ivl_subset_existence_ivl[of t x0] by auto next case False define u where "u = s - t" with False have "u > 0" "s = t + u" by auto note this(2) also have "t + u \ lorenz.existence_ivl0 x0" apply (rule lorenz.existence_ivl_trans) apply fact using * \u > 0\ by auto finally show ?thesis . qed show "x0 \ \" unfolding \_def proof (safe intro!: nonneg_exivl) have "\\<^sub>F s in at_top. (s::real) \ 0" using eventually_ge_at_top by blast then have "\\<^sub>F s in at_top. lorenz.flow0 (lorenz.flow0 x0 t) s = lorenz.flow0 x0 (s + t)" proof eventually_elim case (elim s) then have "s \ lorenz.existence_ivl0 x0" using nonneg_exivl[OF \0 \ s\] by simp then have "lorenz.flow0 (lorenz.flow0 x0 t) s = lorenz.flow0 x0 (t + s)" apply (subst lorenz.flow_trans) using assms * elim by auto then show ?case by (simp add: ac_simps) qed then have "((\s. (lorenz.flow0 x0 (s + t))) \ 0) at_top" by (blast intro: * Lim_transform_eventually) then show "(lorenz.flow0 x0 \ 0) at_top" unfolding aform.tendsto_at_top_translate_iff . next fix s::real assume s: "0 < s" "lorenz.flow0 x0 s \ \" show False proof (cases "s \ t") case True then show ?thesis using outside s by (auto simp: \_def) next case False then obtain u where u: "u = s - t" "u > 0" by auto then have "lorenz.flow0 x0 (s) = lorenz.flow0 x0 (t + u)" by (simp add: algebra_simps) also have "\ = lorenz.flow0 (lorenz.flow0 x0 t) u" apply (subst lorenz.flow_trans) subgoal by fact subgoal unfolding u apply (rule lorenz.diff_existence_ivl_trans) apply fact+ apply (rule nonneg_exivl) using \0 < s\ by simp subgoal by simp done also from assms(1) \u > 0\ have "\ \ \" by (auto simp: \_def) finally show ?thesis using s by auto qed qed qed lemma (in auto_ll_on_open) stable_on_empty[intro, simp]: "stable_on asdf {}" by (auto simp: stable_on_def) lemma \\<^sub>i_stable: "lorenz.stable_on (- \) (\\<^sub>i b)" using \_stable unfolding \\<^sub>i_def apply (cases b) subgoal by auto subgoal using lorenz.stable_on_empty by (auto simp: \\<^sub>i_def) done definition "\\<^sub>v = (cast ` \)" definition "NF = lorenz.flowsto (cube_entry - \ \ UNIV) {0..} (large_cube \ UNIV) cube_exit" lemma NF0: "lorenz.flowsto ((fst ` cube_entry - \) \ UNIV) {0..} (large_cube \ UNIV) (fst ` cube_exit \ UNIV)" if NF using that unfolding NF_def lorenz.flowsto_def apply (auto simp: NF_def) apply (drule bspec, force) by auto lemma [autoref_rules]: includes autoref_syntax shows "(\_. (), \\<^sub>i\<^sub>v) \ bool_rel \ ghost_rel" by (auto simp: ghost_rel_def) lemma lorenz_interrupt[le, refine_vcg]: "lorenz_interrupt optns b X \ SPEC (\(CX, R). lorenz.flowsto ((cast_eucl1 ` X::R3 c1_info set) - (\\<^sub>i b \ UNIV)) {0..} (cast ` CX \ UNIV) (cast_eucl1 ` R))" if NF unfolding lorenz_interrupt_def apply refine_vcg subgoal by (rule lorenz.flowsto_self) auto subgoal by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def) subgoal for a b c d e f g h i j k l m n p apply (auto) apply (simp add: cast_eucl1_image_scaleR2 scaleR2_diff_prod2) apply (erule make_neg_goal) apply (rule lorenz.flowsto_scaleR2I) using NF0[OF that] apply (rule lorenz.flowsto_subset) subgoal for q apply (auto simp: scaleR2_def cast_eucl1_def) apply (auto simp: linear_cast_bl linear.scaleR cube_entry_def cast_eucl1_def image_image) subgoal premises prems for r s t u v proof - from prems have "fst ` c \ fst ` (cube_enter::3 vec1 set)" by auto with \(u, v) \ c\ obtain w where "((u, w)::3 vec1) \ cube_enter" by auto from _ this have "cast u \ (\x. cast (fst (x::3 vec1))) ` cube_enter" by (rule image_eqI) auto then show ?thesis using prems by blast qed subgoal by (auto simp: \\<^sub>i_def) done subgoal by simp subgoal premises _ by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def eucl_of_list_def large_cube_def) subgoal premises prems for x supply [[simproc del: defined_all]] apply (auto simp:) subgoal for A B C D E apply (rule image_eqI[where x="cast_eucl1 (((B, C, D), A))"]) apply (auto simp: cast_eucl1_def) subgoal premises prems using prems(1) apply (auto simp: cube_exit_def aform.c1_info_of_apprs_def cube_exiti_def cast_eucl1_def aform.c1_info_of_appr_def) done done done done subgoal for a b c d e f g h i j k l m n p apply (auto) apply (erule make_neg_goal, thin_tac "\ _") apply (simp add: cast_eucl1_image_scaleR2 scaleR2_diff_prod2) apply (rule lorenz.flowsto_scaleR2I) using that[unfolded NF_def] apply (rule lorenz.flowsto_subset) subgoal for q apply (auto simp: scaleR2_def cast_eucl1_def ) apply (auto simp: linear_cast_bl linear.scaleR cube_entry_def cast_eucl1_def image_image) subgoal premises prems for r s t u v proof - from prems have \vec1_of_flow1 (u, v) \ cube_enter\ by auto from _ this have "(cast u, cast_bl v) \ (\x. (cast (fst (x::3 vec1)), cast_bl (snd (flow1_of_vec1 x)))) ` cube_enter" by (rule image_eqI) (auto simp: ) then show ?thesis using prems by blast qed subgoal by (auto simp: \\<^sub>i_def) done subgoal by simp subgoal premises _ by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def eucl_of_list_def large_cube_def) subgoal by (simp add: cube_exit_def) done done definition "lorenz_S X = (case X of (x, y, z) \ (-x, -y, z))" lemma lorenz_symI: "((\t. lorenz_S (f t)) has_vderiv_on lf') T" if "(f has_vderiv_on f') T" "\t. t \ T \ lf' t = lorenz_S (f' t)" using that by (auto simp: has_vderiv_on_def lorenz_S_def split_beta' has_vector_derivative_def intro!: derivative_eq_intros) lemma lorenz_S: "t \ lorenz.existence_ivl0 (lorenz_S X)" (is ?th1) "lorenz.flow0 (lorenz_S X) t = lorenz_S (lorenz.flow0 X t)" (is ?th2) if "t \ lorenz.existence_ivl0 X" proof - have 1: "((\t. lorenz_S (lorenz.flow0 X t)) solves_ode (\_ (X0, X1, X2). (E1 * X0 - K1 * (X0 + X1) * X2, E2 * X1 + K1 * (X0 + X1) * X2, E3 * X2 + (X0 + X1) * (K2 * X0 + K3 * X1)))) {0--t} UNIV" apply (rule solves_odeI) apply (rule lorenz_symI) apply (rule lorenz.flow_has_vderiv_on_compose) apply simp apply simp apply (rule derivative_intros) apply (rule refl) using that apply (rule lorenz.in_existence_between_zeroI) apply assumption apply (rule refl) unfolding lorenz_S_def apply (split prod.splits)+ apply (simp add: field_simps) apply simp done have "lorenz.flow0 X 0 = X" unfolding lorenz.flow_initial_time_if by simp then have 2: "lorenz_S (lorenz.flow0 X 0) = lorenz_S X" "is_interval {0--t}" "0 \ {0--t}" "{0--t} \ UNIV" by auto from lorenz.maximal_existence_flow[OF 1 2] show ?th1 ?th2 by fast+ qed lemma \\<^sub>l\<^sub>e_impl[autoref_rules]: "(Sctn [0, 0, 1] 27, \\<^sub>l\<^sub>e) \ \lv_rel\below_rel" apply (auto simp: below_rel_def \\<^sub>l\<^sub>e_def below_halfspace_def sctn_rel_def intro!: relcompI[where b="Sctn (0, 0, 1) 27"] brI lv_relI) subgoal unfolding lv_rel_def by (auto intro!: brI) unfolding le_halfspace_def by (auto intro!: brI) lemma [autoref_rules]: "((), \\<^sub>v) \ ghost_rel" by (auto intro!: ghost_relI) no_notation vec_nth (infixl "$" 90) and vec_lambda (binder "\" 10) abbreviation "guards_rel \ \clw_rel (\\lv_rel\ivl_rel, \lv_rel\plane_rel\inter_rel) \\<^sub>r aform.reach_optns_rel\list_rel" definition "aform_poincare_onto_from optns = aform.poincare_onto_from" lemma aform_poincare_onto_from[autoref_rules]: includes autoref_syntax shows "DIM_precond TYPE('b rvec) E \ (XSi, XS::'b::enum eucl1 set) \ clw_rel aform.appr1e_rel \ (sctni, sctn) \ \lv_rel\sctn_rel \ (ivli, ivl) \ \lv_rel\ivl_rel \ (Si, Sa) \ \lv_rel\halfspaces_rel \ (guardsi, guards) \ guards_rel \ (symstartd, symstart) \ aform.appr1e_rel \ \clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel\dres_nres_rel \ ((), trap) \ ghost_rel \ (roi, roptn) \ aform.reach_optns_rel \ (odoi, odo) \ ode_ops_rel \ (optnsi, optns) \ num_optns_rel \ (nres_of (solve_poincare_map_aform optnsi E odoi symstartd Si guardsi ivli sctni roi XSi), aform_poincare_onto_from $ optns $ odo $ symstart $ trap $ Sa $ guards $ ivl $ sctn $ roptn $ XS) \ \clw_rel aform.appr1e_rel\nres_rel" unfolding autoref_tag_defs aform_poincare_onto_from_def using aform.poincare_onto_from_impl.refine[OF _ aform_ncc aform_ncc, where 'a='b, of E odoi odo XSi XS Si Sa guardsi guards ivli ivl sctni sctn roi roptn "(\x. nres_of (symstartd x))" symstart symstartd trap optnsi, unfolded autoref_tag_defs, OF _ _ _ _ _ _ _ _ _ order_refl] by (auto simp: dest!: aform.dres_nres_rel_nres_relD) definition "lorenz_odo_impl = init_ode_ops True True lorenz.odo" interpretation autoref_op_pat_def lorenz.odo . lemma lorenz_odo_impl[autoref_rules]: "(lorenz_odo_impl, lorenz.odo) \ ode_ops_rel" by (auto simp: ode_ops_rel_def lorenz_odo_impl_def) definition lorenz_poincare where "lorenz_poincare optns interrupt guards roptn XS0 = aform_poincare_onto_from optns lorenz.odo (lorenz_interrupt optns interrupt) (\\<^sub>i\<^sub>v interrupt:::ghost_rel) ((below_halfspaces {Sctn (eucl_of_list [0, 0, 1]) 27}::(real^3) set):::\lv_rel\halfspaces_rel) guards (op_atLeastAtMost_ivl (eucl_of_list [-6, -6, 27]:::lv_rel) (eucl_of_list [6, 6, 27]:::lv_rel):::lvivl_rel::(real^3) set) (Sctn (eucl_of_list [0, 0, -1]) (- 27)::(real^3) sctn) roptn XS0" lemma [autoref_rules_raw]: includes autoref_syntax shows "((), (OP \\<^sub>i\<^sub>v ::: bool_rel \ ghost_rel) $ (OP intr ::: bool_rel)) \ ghost_rel" by (auto simp: ghost_rel_def) schematic_goal lorenz_poincare_impl[autoref_rules]: includes autoref_syntax assumes [autoref_rules]: "(XSi, XS) \ clw_rel aform.appr1e_rel" "(intri, intr) \ bool_rel" "(guardsi, guards) \ guards_rel" "(roi, roptn) \ aform.reach_optns_rel" "(optnsi, optns) \ num_optns_rel" shows "(nres_of ?r, lorenz_poincare $ optns $ intr $ guards $ roptn $ XS) \ \clw_rel aform.appr1e_rel\nres_rel" unfolding autoref_tag_defs unfolding lorenz_poincare_def including art supply [autoref_rules_raw] = ghost_relI by autoref_monadic lemma cast_image_eqI: "cast ` X = Y" if "DIM('a) = DIM('b)" "(X::'a::executable_euclidean_space set) = cast ` (Y::'b::executable_euclidean_space set)" using that by (auto simp: image_image) lemma transfer_\[transfer_rule]: "(rel_set lorenz.rel_ve) \\<^sub>v \" unfolding \\<^sub>v_def by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma transfer_\\<^sub>l\<^sub>e[transfer_rule]: "(rel_set lorenz.rel_ve) (cast ` \\<^sub>l\<^sub>e) \\<^sub>l\<^sub>e" by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma transfer_\\<^sub>i[transfer_rule]: "(rel_fun (=) (rel_set lorenz.rel_ve)) \\<^sub>i\<^sub>v \\<^sub>i" unfolding \\<^sub>i\<^sub>v_def by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma transfer_\[transfer_rule]: "(rel_set lorenz.rel_ve) (cast ` \) \" by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI) lemma len_fas: "length lorenz_fas = 3" by (auto simp: lorenz_fas_def) lemma lorenz_poincare[le, refine_vcg]: "lorenz_poincare optns intr guards roptn XS \ SPEC (\R. aform.poincare_mapsto lorenz.odo (cast ` \) (XS - (\\<^sub>i\<^sub>v intr \ UNIV)) (cast ` \\<^sub>l\<^sub>e) UNIV R)" if [refine_vcg]: NF unfolding lorenz_poincare_def aform_poincare_onto_from_def apply (refine_vcg) subgoal by (simp add: aform.wd_def aform.ode_e_def len_fas) subgoal for a b c d apply (auto simp: lorenz.flowsto_eq[symmetric]) proof goal_cases case 1 from 1(2)[unfolded lorenz.flowsto_eq[symmetric]] show ?case by transfer (simp add: lorenz.avflowsto_eq) qed subgoal unfolding aform.stable_on_def unfolding autoref_tag_defs proof (intro allI impI, goal_cases) case (1 t x0) from 1 have t: "t \ lorenz.v.existence_ivl0 x0" using lorenz.vex_ivl_eq by simp from 1 have f: "lorenz.v.flow0 x0 t \ \\<^sub>i\<^sub>v intr" using lorenz.vflow_eq[OF t] by simp from 1 have "lorenz.v.flow0 x0 s \ - cast ` \" if "s\{0<..t}" for s proof - from that t have s: "s \ lorenz.v.existence_ivl0 x0" by (auto dest!: lorenz.a.v.closed_segment_subset_existence_ivl simp: closed_segment_eq_real_ivl) have "lorenz.v.flow0 x0 s \ aform.Csafe lorenz.odo - op_atLeastAtMost_ivl (eucl_of_list [- 6, - 6, 27]) (eucl_of_list [6, 6, 27]) \ plane_of (Sctn (eucl_of_list [0, 0, - 1]) (- 27))" using 1(4)[rule_format, OF that] unfolding lorenz.vflow_eq[OF s] by auto also have "\ \ - cast ` \" by (auto simp: eucl_le[where 'a="real^3"] eucl_of_list_inner axis_eq_axis cast_def Basis_list_real_def Basis_list_vec3 \_def plane_of_def eucl_of_list_inner_eq inner_lv_rel_def) finally show "lorenz.v.flow0 x0 s \ - cast ` \" . qed show "x0 \ \\<^sub>i\<^sub>v intr" by (rule \\<^sub>i_stable[Transfer.untransferred, unfolded lorenz.v.stable_on_def, rule_format]) fact+ qed subgoal for R proof (clarsimp, goal_cases) case 1 note 1(2) also have "({eucl_of_list [- 6, - 6, 27]..eucl_of_list [6, 6, 27]::3 rvec} \ plane_of (Sctn (eucl_of_list [0, 0, - 1]) (- 27))) = cast ` \" apply auto apply (auto simp: \_def o_def plane_of_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) subgoal by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) subgoal by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) subgoal apply (auto simp: cast_le_iff[symmetric]) by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def list_of_eucl_def inner_simps inner_axis_axis) subgoal apply (auto simp: cast_le_iff) by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def list_of_eucl_def inner_simps inner_axis_axis) subgoal by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def list_of_eucl_def inner_simps inner_axis_axis) done also have "(below_halfspaces {Sctn (eucl_of_list [0, 0, 1]::3 rvec) 27}) = (cast ` \\<^sub>l\<^sub>e)" apply (auto simp: \\<^sub>l\<^sub>e_def below_halfspaces_def below_halfspace_def le_halfspace_def eucl_of_list_def Basis_list_R3 Basis_list_vec3) apply (rule image_eqI) apply (rule cast_cast[symmetric]) by (auto simp: cast_def list_of_eucl_def o_def plane_of_def inner_simps inner_axis_axis eucl_of_list_def Basis_list_R3 Basis_list_vec3) finally show ?case by (rule aform.poincare_mapsto_subset) (force simp: lorenz.vX_def intro: cast_cast[symmetric])+ qed done context includes floatarith_notation begin definition "mat1\<^sub>e = [Var 0, Var 1, Var 2, Var 3, 0, 0, Var 4, 0, 0, Var 5, 0, 0]" definition mat1_nres::"3 rvec set \ 3 rvec set \ 3 eucl1 set nres" where "mat1_nres X v = do { Xv \ aform.approx_slp_appr mat1\<^sub>e (slp_of_fas mat1\<^sub>e) (concat ` listset [list_of_eucl ` X, list_of_eucl ` v]); RETURN (flow1_of_vec1 ` Xv) }" lemma [simp]: "(x, x') \ aform.appr_rel \ aform.ncc x'" using aform_ncc[where 'a='a] by (auto simp: aform.ncc_precond_def) lemma mat1e_autoref[autoref_rules]: "(mat1\<^sub>e, mat1\<^sub>e) \ \Id\list_rel" by auto schematic_goal mat1_impl: includes autoref_syntax assumes [autoref_rules]: "(Xi, X) \ aform.appr_rel" "(vi, v) \ aform.appr_rel" shows "(nres_of ?r, mat1_nres $ X $ v) \ \aform.appr1_rel\nres_rel" unfolding mat1_nres_def including art by autoref_monadic concrete_definition mat1_impl for Xi vi uses mat1_impl lemmas [autoref_rules] = mat1_impl.refine lemma mat_nres[le, refine_vcg]: "mat1_nres X v \ SPEC (\M. X \ v \ (\x. (fst x, blinfun_apply (snd x) (eucl_of_list [1, 0, 0]))) ` M)" unfolding mat1_nres_def apply refine_vcg apply (auto simp: dest!: bspec) apply (auto simp: mat1\<^sub>e_def image_image ) subgoal for x a b apply (rule image_eqI[where x="eucl_of_list [(list_of_eucl a @ list_of_eucl b) ! 0, (list_of_eucl a @ list_of_eucl b) ! Suc 0, (list_of_eucl a @ list_of_eucl b) ! 2, (list_of_eucl a @ list_of_eucl b) ! 3, 0, 0, (list_of_eucl a @ list_of_eucl b) ! 4, 0, 0, (list_of_eucl a @ list_of_eucl b) ! 5, 0, 0]"]) apply (auto simp: eucl_of_list_prod eucl_of_list_inner nth_append Basis_list_vec3 intro!: euclidean_eqI[where 'a="3 rvec"]) unfolding Basis_list[symmetric] Basis_list_vec3 by (auto simp: flow1_of_vec1_def blinfun_of_vmatrix.rep_eq Basis_list_vec3 inner_simps matrix_vector_mult_eq_list_of_eucl_nth ll3 inner_axis_axis) done definition [simp]: "op_image_cast_eucl1e = (`) cast_eucl1" definition [simp]: "op_image_cast_eucl1e_coll = (`) cast_eucl1" lemma prod_relI'': "\(fst ab, a')\R1; (snd ab, b')\R2\ \ (ab,(a', b'))\\R1,R2\prod_rel" by (auto simp: prod_rel_def) lemma strange_aux_lemma: "(b, b') \ A \ (b, snd (a'a, b')) \ A" by auto lemma [autoref_rules]: includes autoref_syntax assumes "DIM_precond TYPE('a::executable_euclidean_space) D" "DIM_precond TYPE('b::executable_euclidean_space) D" shows "(\x. x, (op_image_cast_eucl1e::('a::executable_euclidean_space c1_info set \ 'b::executable_euclidean_space c1_info set))) \ aform.appr1e_rel \ aform.appr1e_rel" (is ?th1) and "(\x. x, op_image_cast_eucl1e_coll::'a::executable_euclidean_space c1_info set \ 'b::executable_euclidean_space c1_info set) \ clw_rel aform.appr1e_rel \ clw_rel aform.appr1e_rel" (is ?th2) proof - show 1: ?th1 unfolding scaleR2_rel_def apply (rule subsetD) apply (rule fun_rel_comp_dist) apply (rule relcompI) apply (rule fun_relI) apply (erule prod_relE) apply simp apply (rule prod_relI) apply simp apply (rule fst_conv[symmetric]) apply (rule op_cast_eucl1_image_impl[OF assms, param_fo]) apply (rule strange_aux_lemma) apply (auto simp: br_def scaleR2_def image_def vimage_def cast_eucl1_def) subgoal for a b c d e f g apply (rule exI[where x=e] conjI)+ apply assumption apply (rule exI conjI)+ apply assumption apply (rule exI conjI)+ apply (rule bexI) prefer 2 apply assumption apply (rule conjI) apply force apply (rule refl) using assms by (auto simp: blinfun.bilinear_simps linear_cast linear.scaleR intro!: blinfun_eqI) subgoal for a b c d e f g apply (rule exI[where x=e] exI conjI)+ apply assumption apply (rule exI conjI)+ apply assumption apply (rule bexI) prefer 2 apply assumption apply force using assms by (auto simp: blinfun.bilinear_simps linear_cast linear.scaleR intro!: blinfun_eqI) done have id_map: "(\x. x) = (map (\x. x))" by simp show ?th2 apply (subst id_map) apply (rule lift_clw_rel_map) apply (rule relator_props)+ subgoal using 1 by auto subgoal by auto done qed definition "lorenz_poincare_tangents optns interrupt guards roptn c1 (X0::R3 set) (tangents::R3 set) = do { X0tanmat \ (if c1 then do { R \ mat1_nres (cast ` X0) (cast ` tangents); RETURN (R::3 eucl1 set) } else RETURN (cast ` X0 \ UNIV)); XDX0 \ scaleRe_ivl_spec 1 1 (X0tanmat); let _ = aform.trace_set1e ''START'' (Some XDX0); let _ = aform.print_set1e False (XDX0); P \ lorenz_poincare optns interrupt guards roptn ((mk_coll XDX0:::clw_rel aform.appr1e_rel)); RETURN (op_image_cast_eucl1e_coll P::R3 c1_info set) }" lemma [autoref_rules_raw]: "DIM(real \ real \ real) = DIM((real, 3) vec)" by auto schematic_goal lorenz_poincare_tangents_impl: includes autoref_syntax assumes [autoref_rules]: "(optnsi, optns) \ Id" "(intrri, intr) \ bool_rel" "(guardsi, guards) \ guards_rel" "(roi, roptn) \ aform.reach_optns_rel" "(c1i, c1) \ bool_rel" "(X0i, X0) \ aform.appr_rel" "(tangentsi, tangents) \ aform.appr_rel" shows "(nres_of ?r, lorenz_poincare_tangents $ optns $ intr $ guards $ roptn $ c1 $ (X0::R3 set) $ tangents) \ \clw_rel aform.appr1e_rel\nres_rel" unfolding lorenz_poincare_tangents_def including art by (autoref_monadic) concrete_definition lorenz_poincare_tangents_impl uses lorenz_poincare_tangents_impl[where optnsa = "\_ _ _ _ _ _ _ _. optns" and optnsb = "\_ _ _ _ _ _ _ _ _. optns" and optnsi = "optns" and optnsc = "optns" and optns = "\_ _ _ _ _ _ _. optns" for optns optnsc] lemma lorenz_poincare_tangents_impl_refine[autoref_rules]: includes autoref_syntax shows "(\optnsi intrri guardsi roi c1i X0i tangentsi. nres_of (lorenz_poincare_tangents_impl optnsi intrri guardsi roi c1i X0i tangentsi), lorenz_poincare_tangents) \ num_optns_rel \ bool_rel \ guards_rel \ aform.reach_optns_rel \ bool_rel \ aform.appr_rel \ aform.appr_rel \ \clw_rel aform.appr1e_rel\nres_rel" using lorenz_poincare_tangents_impl.refine by force lemma transfer_UNIV_rel_blinfun[transfer_rule]: "(rel_set (rel_blinfun lorenz.rel_ve lorenz.rel_ve)) UNIV UNIV" apply (auto intro!: rel_setI simp: rel_blinfun_def) subgoal for x apply (rule exI[where x="cast_bl x"]) by (auto intro!: rel_funI simp: lorenz.rel_ve_cast) subgoal for x apply (rule exI[where x="cast_bl x"]) by (auto intro!: rel_funI simp: lorenz.rel_ve_cast) done lemma lorenz_vX[simp]: "lorenz.vX = (UNIV::3 rvec set)" by (force simp: lorenz.vX_def intro!: cast_cast[symmetric]) lemma closed_cast_\[intro, simp]: "closed (cast ` \::3 rvec set)" by (auto simp: \_def ) lemma blinfun_apply_transfer[transfer_rule]: "(rel_fun (rel_blinfun lorenz.rel_ve lorenz.rel_ve) (rel_fun (rel_prod (=) (rel_prod (=) (=))) lorenz.rel_ve)) (blinfun_apply o cast_bl) blinfun_apply" by (auto intro!: rel_funI simp: rel_blinfun_def lorenz.rel_ve_cast dest!: rel_funD) lemma lorenz_poincare_tangents[le, refine_vcg]: "lorenz_poincare_tangents optns intr guards roptn c1 (X0::R3 set) tangents \ SPEC (\x. (if c1 then \tans. X0 \ tangents \ (\(x, y). (x, blinfun_apply y (1, 0, 0))) ` tans \ lorenz.poincare_mapsto \ (tans - \\<^sub>i intr \ UNIV) (\\<^sub>l\<^sub>e) UNIV x else lorenz.poincare_mapsto \ ((X0 - \\<^sub>i intr) \ UNIV) (\\<^sub>l\<^sub>e) UNIV x))" if [refine_vcg]: NF unfolding lorenz_poincare_tangents_def apply refine_vcg apply auto apply (subst lorenz.poincare_mapsto_eq[symmetric]) apply simp proof goal_cases case (2 R) then show ?case apply transfer apply (subst lorenz.vpoincare_mapsto_eq[symmetric]) apply (auto simp: ) apply (rule aform.poincare_mapsto_subset, assumption) by (force simp: scaleR2_def )+ next case (1 tans R) show ?case apply (rule exI[where x="cast_eucl1 ` tans"]) apply (rule conjI) subgoal including lifting_syntax using 1(2) by transfer (force simp: cast_def o_def) subgoal using 1(3) apply transfer apply (subst lorenz.avpoincare_mapsto_eq[symmetric]) by (auto simp: ) done qed definition of_mat1_image::"R3 c1_info set \ R3 set nres" where [refine_vcg_def]: "of_mat1_image X = SPEC (\R. R = (\x. blinfun_apply (snd x) (1, 0, 0)) ` X)" lemma of_mat1_image_impl[autoref_rules]: "(\x. (case x of (_, Some xs) \ RETURN [xs ! 0, xs ! 3, xs ! 6] | (_, None) \ SUCCEED), of_mat1_image) \ aform.appr1_rel \ \aform.appr_rel\nres_rel" apply (auto simp: of_mat1_image_def RETURN_RES_refine_iff nres_rel_def aform.appr1_rel_internal aform.appr_rel_def intro!: relcompI split: option.splits) unfolding aforms_rel_def apply (rule brI) apply (auto simp: ) unfolding lv_rel_def set_rel_br apply (rule brI) prefer 2 apply (force simp: Joints_imp_length_eq) apply (auto elim!: mem_Joints_appendE simp: flow1_of_list_def Joints_imp_length_eq) subgoal for a b c d e f g h i j apply (rule image_eqI[where x="[j ! 0, j ! 3, j! 6]"]) apply (auto simp: blinfun_of_list_def blinfun_of_matrix_apply Basis_prod_def Basis_list_R3 Basis_list_vec3 eval_nat_numeral zero_prod_def) apply (rule map_nth_Joints'[of _ _ "[0, Suc (Suc (Suc 0)), Suc (Suc (Suc (Suc (Suc (Suc 0)))))]", simplified]) apply auto done subgoal for a b c d e f unfolding image_image apply (auto simp: Joints_def valuate_def) subgoal for g apply (rule image_eqI) prefer 2 apply (rule image_eqI[where x=g]) apply (rule refl) apply (auto simp: ) apply (auto simp: blinfun_of_list_def blinfun_of_matrix_apply flow1_of_list_def Basis_prod_def Basis_list_R3 Basis_list_vec3 eval_nat_numeral zero_prod_def) done done done definition [refine_vcg_def]: "floatdegs res = SPEC (\_::unit. min_deg res \ float \ max_deg res \ float)" definition [simp]: "isinfloat x \ x \ float" lemma [code]: "isinfloat (real_of_float x) = True" by (auto) lemma floatdegs_impl[autoref_rules]: includes autoref_syntax shows "(\res. (if isinfloat (min_deg res) \ isinfloat (max_deg res) then RETURN () else SUCCEED), floatdegs) \ Id \ \unit_rel\nres_rel" by (auto simp: nres_rel_def floatdegs_def) definition "check_c1_entry optns em P (res0::result) (res::result) = do { uv_ret \ of_mat1_image P; nuv \ aform.mig_set 3 uv_ret; floatdegs res0; floatdegs res; let e' = em * ereal nuv; b1 \ approx_conefield_bounds P (min_deg res) (max_deg res); let b2 = e' \ preexpansion res; let b3 = e' \ expansion res0; let _ = aform.print_msg ((shows ''# check_c1_entry: '' o shows_list [b1, b2, b3] o shows_space o shows_list [e', preexpansion res, expansion res0]) ''''); RETURN (em \ 0 \ b1 \ b2 \ b3) }" lemma [autoref_itype]: "shows_prec ::\<^sub>i i_nat \\<^sub>i A \\<^sub>i i_string \\<^sub>i i_string" by auto lemma [autoref_rules]: includes autoref_syntax shows "PREFER_id A \ (shows_list, shows_list) \ \A\list_rel \ string_rel \ string_rel" "(shows_prec, shows_prec) \ nat_rel \ string_rel \ string_rel \ string_rel" "(shows_prec, shows_prec) \ nat_rel \ ereal_rel \ string_rel \ string_rel" "(shows_prec, shows_prec::_\result \_) \ nat_rel \ Id \ string_rel \ string_rel" "(shows_space, shows_space) \ string_rel \ string_rel" by (auto simp: string_rel_def) lemma [autoref_rules]: includes autoref_syntax shows "(expansion, expansion) \ Id \ rnv_rel" "(preexpansion, preexpansion) \ Id \ rnv_rel" "(min_deg, min_deg) \ Id \ rnv_rel" "(max_deg, max_deg) \ Id \ rnv_rel" by auto interpretation autoref_op_pat_def aform.mig_set . lemma [autoref_rules_raw]: "DIM_precond TYPE(real \ real \ real) (OP 3 ::: nat_rel)" by simp schematic_goal check_c1_entry_impl: includes autoref_syntax assumes [autoref_rules]: "(optnsi, optns) \ Id" "(res0i, res0) \ Id" "(resi, res) \ Id" "(emi, em) \ ereal_rel" "(Pei, P) \ aform.appr1_rel" shows "(nres_of ?r, check_c1_entry optns em P res0 res) \ \bool_rel\nres_rel" unfolding check_c1_entry_def including art by autoref_monadic concrete_definition check_c1_entry_impl uses check_c1_entry_impl[ where optns = "\_ . optnsi" and optnsi="optnsi" and optnsc=optns and optnsa="\_ _ _ _ _. optnsi" and optnsb="\_ _ _ _ _ _ _ _ . optnsi" and optns="\_. optnsi" for optns optnsi] lemmas check_c1_entry_impl_refine[autoref_rules] = check_c1_entry_impl.refine[autoref_higher_order_rule] definition "c1_entry_correct (em::ereal) (P::R3 c1_info set) res0 res = (\(_, d)\P. case (d (1, 0, 0)) of (dx, dy, dz) \ dz = 0 \ dx > 0 \ -90 < min_deg res \ min_deg res \ max_deg res \ max_deg res < 90 \ ereal (preexpansion res) \ em * (norm (dx, dy, dz)) \ ereal (expansion res0) \ em * (norm (dx, dy, dz)) \ dy / dx \ {tan (rad_of (min_deg res)) .. tan (rad_of (max_deg res))})" lemma check_c1_entry[le, refine_vcg]: "check_c1_entry optns em P res0 res \ SPEC (\b. b \ c1_entry_correct em P res0 res)" unfolding check_c1_entry_def c1_entry_correct_def apply refine_vcg apply (auto dest!: bspec simp:) apply (rule order_trans, assumption, rule ereal_mult_left_mono, force, force) apply (rule order_trans, assumption, rule ereal_mult_left_mono, force, force) done subsection \options for the lorenz system\ definition aform_numeric_optns::"_ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ _ \ real aform numeric_options" where "aform_numeric_optns = numeric_options.fields" fun zbucket::"real \ real \ real \ real \ real \ real \ real \ ((real list \ real list) \ real list sctn) list" where "zbucket d (x0,x1) (y0, y1) (z0, z1) = [zsec' (x0 - d, x0 + d) (y0 - d, y1 + d) z0, \ \bottom\ xsec' x0 (y0 - d, y1 + d) (z0 - d, z1), \ \left\ xsec x1 (y0 - d, y1 + d) (z0 - d, z1), \ \right\ ysec' (x0 - d, x1 + d) y0 (z0 - d, z1), \ \backno\ ysec (x0 - d, x1 + d) y1 (z0 - d, z1)] \ \front\" subsubsection \Hybridizations\ definition "reduce_weak_params c1 = (if c1 then (12::nat, 0::nat) else (3, 0))" definition "reduce_hard_params c1 = (if c1 then (0::nat, 100::nat) else (0, 100))" definition "ro_split_weak c1 w = (case reduce_weak_params c1 of (m, n) \ ro (w + 1) m n w w (-5))" definition "ro_split_weak' c1 w = (case reduce_weak_params c1 of (m, n) \ ro w m n w w (-5))" definition "ro_split_weak'' c1 w = (case reduce_weak_params c1 of (m, n) \ ro (w + 2) m n w w (-5))" definition "ro_split_weak4' c1 w = (case reduce_weak_params c1 of (m, n) \ ro (w + 4) m n w w (-5))" definition "ro_split_weak2 c1 w w2 = (case reduce_weak_params c1 of (m, n) \ ro (w + 1) m n w w2 (-5))" definition "ro_split_weak2' c1 w w2 = (case reduce_weak_params c1 of (m, n) \ ro (w) m n w w2 (-5))" definition "ro_split_hard c1 w0 w1 = (case reduce_hard_params c1 of (m, n) \ ro (w0 + 1) m n w0 w1 (-5))" definition "ro_split_hard'' c1 w0 w1 = (case reduce_hard_params c1 of (m, n) \ ro (w0 + 2) m n w0 w1 (-5))" definition "ro_split_not c1 w = (case reduce_weak_params c1 of (m, n) \ ro 0 m n w w (-5))" definition "ro_split_not2 c1 w w2 = (case reduce_weak_params c1 of (m, n) \ ro 0 m n w w2 (-5))" definition "xsecs x y z = [xsec' (-x) (-y, y) (0, z), xsec x (-y, y) (0, z)]" type_synonym run_options = "(nat \ nat) \ int \ (((real list \ real list) \ real list sctn) list \ real aform reach_options) list \ real aform reach_options" abbreviation "p1 \ ldec 0.1" definition mode_middle::"_ \ run_options" where "mode_middle c1 = (reduce_weak_params c1, -14, [([zsec' (-2, 2) (-1, 1) 10], ro_split_weak' c1 (-3)), (xsecs (5 * p1) 10 10 @ xsecs p1 10 (6) @ [zsec' (-p1, p1) (-p1, p1) p1], ro_split_hard c1 (-5) (-2)), (xsecs (3/2/2) 10 (10), ro_split_not2 c1 0 (-2)), \ \To collect after interrupt\ ([zsec (-30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-1) (-3)) ], ro_split_weak4' c1 (-5))" definition mode_inner3::"bool\bool\run_options" where "mode_inner3 c1 very_inner = (reduce_weak_params c1, -15, (if very_inner then [([zsec' (-2, 2) (-1, 1) 10], ro_split_weak' c1 (-2))] else [])@ [(xsecs (3/2) 15 27@xsecs 1 (10) (11/2), ro_split_weak2 c1 (-2) (-1)), ([ zsec (-30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-1) (-3)) ], if very_inner then ro_split_weak4' c1 (-5) else ro_split_weak'' c1 (-5))" definition mode_inner2::"bool \ real \ run_options" where "mode_inner2 c1 x = (reduce_weak_params c1, -14, [(xsecs x 10 27, ro_split_weak2' c1 (-2) (-1)), ([zsec ( -30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-3) (-3))], ro_split_not c1 (-6))" definition "ro_outer c1 w = (case reduce_weak_params c1 of (m, n) \ ro w m n (-6) (-6) (-5))" definition mode_outer::"bool\_\_\run_options" where "mode_outer c1 w i = (reduce_weak_params c1, (-i), [([zsec (-30, -6) (-10, 10) 27, zsec (6, 30) (-10, 10) 27], ro_split_not2 c1 (-3) (-4))], ro_outer c1 w)" definition lookup_mode::"bool \ result \ _" where "lookup_mode c1 i = (if gridx0 i \ - 1024 then mode_outer c1 (-3) 16 else if gridx0 i \ - 120 then mode_outer c1 (-3) 14 else if gridx0 i \ 107 then mode_inner2 c1 (4) else if gridx0 i \ 169 then mode_inner3 c1 False else if gridx0 i \ 196 then mode_inner3 c1 True else if gridx0 i \ 201 then mode_middle c1 else if gridx0 i \ 235 then mode_inner3 c1 True else if gridx0 i \ 290 then mode_inner3 c1 False else if gridx0 i \ 450 then mode_inner2 c1 4 else mode_outer c1 (-3) 14)" definition mode_ro_spec::"bool \ result \ ((nat \ nat) \ int \ ((real, 3) vec set \ unit) list \ unit) nres" where [refine_vcg_def]: "mode_ro_spec c1 res = SPEC (\_. True)" lemma reach_options_rel_br: "reach_options_rel TYPE('ty) = br (\_. ()) (\_. True)" by (auto simp: reach_options_rel_def br_def) lemma mode_ro_spec_impl[autoref_rules]: includes autoref_syntax shows "(\b x. RETURN (lookup_mode b x), mode_ro_spec) \ bool_rel \ Id \ \(nat_rel \\<^sub>r nat_rel) \\<^sub>r int_rel \\<^sub>r guards_rel \\<^sub>r aform.reach_optns_rel\nres_rel" supply [simp del] = prod_rel_id_simp apply (rule fun_relI) apply (rule fun_relI) apply (rule nres_relI) unfolding mode_ro_spec_def apply (rule RETURN_SPEC_refine) apply (auto simp: mode_ro_spec_def nres_rel_def RETURN_RES_refine_iff) apply (rule exI)+ apply (rule prod_relI'' IdI)+ unfolding lv_rel_def ivl_rel_def br_rel_prod br_chain plane_rel_br inter_rel_br clw_rel_br br_list_rel Id_br prod_eq_iff reach_options_rel_br apply (rule brI refl)+ defer apply (rule brI) apply (rule refl) apply auto apply (auto simp: lookup_mode_def mode_outer_def mode_inner2_def mode_inner3_def xsecs_def mode_middle_def) done lemma [autoref_rules]: includes autoref_syntax shows "(ivl_of_res, ivl_of_res) \ Id \ \rnv_rel\list_rel \\<^sub>r \rnv_rel\list_rel" by auto lemma [autoref_rules]: includes autoref_syntax shows "(Polygon.pairself, Polygon.pairself) \ (A \ C) \ (A \\<^sub>r A) \ (C \\<^sub>r C)" by (auto dest: fun_relD) lemma set_of_ivl_impl[autoref_rules]: includes autoref_syntax shows "(\x. x, set_of_ivl) \ (A \\<^sub>r A) \ \A\ivl_rel" by (auto simp: ivl_rel_def br_def) lemma eucl_of_list_pad: includes autoref_syntax shows "DIM_precond TYPE('a::executable_euclidean_space) D \ (\xs. take D xs @ replicate (D - length xs) 0, eucl_of_list::_\'a) \ rl_rel \ lv_rel" unfolding lv_rel_def by (auto simp: intro!: brI) concrete_definition eucl_of_list_pad uses eucl_of_list_pad lemmas [autoref_rules] = eucl_of_list_pad.refine lemma source_of_res_impl[autoref_rules]: includes autoref_syntax shows "(ivl_of_res, source_of_res) \ Id \ \lv_rel\ivl_rel" unfolding source_of_res_def apply (auto simp: ivl_rel_def intro!: relcompI brI) subgoal for a apply (auto simp: ivl_of_res_def ivl_of_resultrect_def intro!: lv_relI) unfolding lv_rel_def apply (auto intro!: brI) done done definition tangent_seg_of_res :: "real aform numeric_options \ result \ R3 set nres" where "tangent_seg_of_res optns res0 = do { let fas = map (OP (nth matrix_of_degrees2\<^sub>e)) [0, 3, 6]; let u = min_deg res0; let v = max_deg res0; aform.approx_slp_appr fas (slp_of_fas fas) (lv_ivl [u, v, 0] [u, v, 1]) }" lemmas [refine_vcg_def] = tangent_seg_of_res_spec_def lemma tangent_seg_of_res[le, refine_vcg]: "tangent_seg_of_res optns res \ tangent_seg_of_res_spec res" unfolding tangent_seg_of_res_def tangent_seg_of_res_spec_def apply refine_vcg apply (auto simp: matrix_of_degrees2\<^sub>e_def Let_def in_segment) subgoal for x a b c u by (drule bspec[where x="[min_deg res, max_deg res, u]"]) (auto simp: tangent_of_deg_def lv_ivl_def algebra_simps intro!:) done lemma [autoref_rules]: includes autoref_syntax shows "(nth matrix_of_degrees2\<^sub>e, nth matrix_of_degrees2\<^sub>e) \ nat_rel \ Id" by auto schematic_goal tangent_seg_of_res_impl: includes autoref_syntax assumes [autoref_rules]: "(resi, res) \ Id" "(optnsi, optns) \ num_optns_rel" shows "(nres_of ?r, tangent_seg_of_res optns res) \ \aform.appr_rel\nres_rel" unfolding tangent_seg_of_res_def including art by autoref_monadic concrete_definition tangent_seg_of_res_impl uses tangent_seg_of_res_impl lemmas [autoref_rules] = tangent_seg_of_res_impl.refine[where optnsi = optnsi and optnsa=optns and optns="\_ _ _. optnsi" for optns optnsi, autoref_higher_order_rule] lemma return_of_res_impl: includes autoref_syntax shows "(\results res. (get_results (inf_retx res) (inf_rety res) (sup_retx res) (sup_rety res) results), return_of_res) \ \Id\list_rel \ Id \ \Id\list_wset_rel" by (auto simp: return_of_res_def list_wset_rel_def intro!: brI) concrete_definition return_of_res_impl uses return_of_res_impl lemmas [autoref_rules] = return_of_res_impl.refine lemma lorenz_optns'_impl[autoref_rules]: includes autoref_syntax shows "(lorenz_optns', lorenz_optns') \ \Id \ unit_rel\option_rel \ nat_rel \ nat_rel \ rnv_rel \ rnv_rel \ num_optns_rel" by auto lemma [autoref_rules]: includes autoref_syntax shows "(results, results) \ \Id\list_rel" "(invoke_nf, invoke_nf) \ Id \ bool_rel" by auto definition "check_line_nres print_fun m0 n0 c1 res0 = do { let X0 = source_of_res res0; (X0l, X0u) \ ivl_rep X0; ((m::nat, n::nat), a::int, modes, ro) \ mode_ro_spec c1 res0; let interrupt = invoke_nf res0; let optns = lorenz_optns' print_fun (the_default m m0) (the_default n n0) 1 (FloatR 1 a); tangents \ tangent_seg_of_res optns res0; aform.CHECKs (ST ''check_line_nres le'') (X0l \ X0u); sp \ aform.subset_spec_plane X0 (Sctn (eucl_of_list [0, 0, 1]) 27); aform.CHECKs (ST ''check_line_nres le'') sp; ASSERT (X0l \ X0u); Pe \ lorenz_poincare_tangents optns interrupt modes ro c1 ({X0l .. X0u}) tangents; PeS \ sets_of_coll Pe; let RETs = (return_of_res results res0); let RET = \((mk_coll ` (source_of_res ` RETs:::\lvivl_rel\list_wset_rel):::\clw_rel lvivl_rel\list_wset_rel)); every \ WEAK_ALL\<^bsup>\Pe. \P em eM Rivls. em > 0 \ Pe = scaleR2 em eM P \ fst ` P \ \Rivls \ (\Rivl \ Rivls. (\res\RETs. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res)))\<^esup> PeS (\Pe. do { let _ = aform.trace_set1e (ST ''# Return Element: '') (Some Pe); ((em, eM), P) \ scaleR2_rep Pe; aform.CHECKs (ST ''check_line_nres pos'') (0 < em); let R = (fst ` P:::aform.appr_rel); (Ri, Rs) \ op_ivl_rep_of_set R; let Rivl = (op_atLeastAtMost_ivl Ri Rs); Rivls \ aform.split_along_ivls2 3 (mk_coll Rivl) RET; Rivlss \ sets_of_coll Rivls; WEAK_ALL\<^bsup>\Rivl. \res\RETs. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res)\<^esup> Rivlss (\Rivl. do { b \ WEAK_EX\<^bsup>\res. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res)\<^esup> RETs (\res. do { let src = (source_of_res res:::lvivl_rel); let subs = Rivl \ src; cones \ if \(c1 \ subs) then RETURN True else check_c1_entry optns em P res0 res; RETURN (subs \ cones) }); let _ = aform.print_msg ((shows (ST ''# return of '') o shows res0 o shows (if b then ST '' OK'' else ST '' FAILED''))''''); RETURN b }) }); RETURN (every, Pe, RET) }" definition "aform_subset_spec_plane optns = aform.subset_spec_plane" lemma aform_subset_spec_plane_impl[autoref_rules]: includes autoref_syntax shows "DIM_precond TYPE('a::executable_euclidean_space) D \ (Xi, X::'a set) \ \lv_rel\ivl_rel \ (sctni, sctn) \ \lv_rel\sctn_rel \ (optnsi, optns) \ num_optns_rel \ (nres_of (subset_spec_plane_impl_aform optnsi D Xi sctni), aform_subset_spec_plane $ optns $ X $ sctn) \ \bool_rel\nres_rel" using aform.subset_spec_plane_impl.refine[where 'a='a, of D Xi X sctni sctn optnsi] by (force simp: aform_subset_spec_plane_def) schematic_goal check_line_impl: includes autoref_syntax assumes [autoref_rules]: "(pfi, pf) \ \Id \ unit_rel\option_rel" "(c1i, c1) \ bool_rel" "(res0i, res0) \ Id" "(m0i, m0) \ \nat_rel\option_rel" "(n0i, n0) \ \nat_rel\option_rel" shows "(nres_of ?r, check_line_nres $ pf $ m0 $ n0 $ c1 $ res0) \ \bool_rel \\<^sub>r clw_rel aform.appr1e_rel \\<^sub>r clw_rel lvivl_rel\nres_rel" unfolding check_line_nres_def including art by autoref_monadic concrete_definition check_line_impl uses check_line_impl[where optns = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ . lorenz_optns pfi" and optnsa = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsb = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsc = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsd = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnse = "\_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and optnsf = "\ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi" and pfi = pfi for pfi] lemmas [autoref_rules] = check_line_impl.refine lemma check_line_nres: "check_line_nres pf m0 n0 c1 res0 \ SPEC (\(every, Pe, RET). \P. Pe = \P \ (if c1 then \tans. source_of_res res0 \ {tangent_of_deg (min_deg res0)--tangent_of_deg (max_deg res0)} \ (\(x, y). (x, blinfun_apply y (1, 0, 0))) ` tans \ lorenz.poincare_mapsto \ (tans - \\<^sub>i (invoke_nf res0) \ UNIV) \\<^sub>l\<^sub>e UNIV (Pe) else lorenz.poincare_mapsto \ ((sourcei_of_res res0) \ UNIV) \\<^sub>l\<^sub>e UNIV (\P)) \ source_of_res res0 \ plane_of (Sctn (0, 0, 1) 27) \ (every \ (\x\P. \P em. em > 0 \ (\eM. x = scaleR2 em eM P) \ (\Rivls. fst ` P \ \Rivls \ (\Rivl\Rivls. \res\return_of_res results res0. Rivl \ source_of_res res \ (c1 \ c1_entry_correct em P res0 res))))))" if [refine_vcg]: NF unfolding check_line_nres_def sourcei_of_res_def apply (refine_vcg, clarsimp_all) using [[goals_limit=1]] subgoal for a0 a b c d e f g h i j k l m n p q r s t apply (rule exI[where x=i]) apply (rule exI[where x=j]) apply (rule conjI) apply force apply (rule conjI) apply (rule exI[where x=k]) apply force apply (rule exI[where x=s]) apply (rule conjI) defer apply force apply blast done subgoal for x0 y0 z0 x1 y1 z1 tangents s R every apply (rule exI[where x=R]) apply auto subgoal for tans apply (rule exI[where x=tans]) by auto subgoal for tans apply (rule exI[where x=tans]) by auto done done definition "print_sets_color (print_fun::String.literal \ unit) (c::string) (X::'a::executable_euclidean_space set) = ()" definition "print_lorenz_color print_fun cx cy cz ci cd1 cd2 P = ()" definition "print_aforms print_fun c aforms = fold (\a _. print_fun (String.implode (shows_segments_of_aform 0 1 a c ''\''))) aforms ()" lemma print_sets_color_impl[autoref_rules]: includes autoref_syntax shows "(\print_fun c X. print_aforms print_fun c X, print_sets_color) \ (Id \ unit_rel) \ string_rel \ clw_rel aform.appr_rel \ unit_rel" by auto lemma print_lorenz_color_impl[autoref_rules]: includes autoref_syntax shows "(\print_fun cx cy cz ci cd1 cd2 P. fold (\(_, x) b. print_lorenz_aform print_fun cx cy cz ci cd1 cd2 False (fst x @ the_default [] (snd x)) ) P (), print_lorenz_color) \ (Id \ unit_rel) \ string_rel \ string_rel \ string_rel \ string_rel \ \\string_rel\list_rel, \\string_rel\list_rel, (\clw_rel aform.appr1e_rel, unit_rel\fun_rel)\fun_rel\fun_rel" by auto definition check_line_core where "check_line_core print_funo m0 n0 c1 i = do { let print_fun = the_default (\_. ()) print_funo; CHECK (\_. print_fun (STR ''Hey, out of bounds!'')) (i < length results); let res = ((results:::\Id\list_rel) ! (i:::nat_rel)); (r, P, B) \ check_line_nres print_funo m0 n0 c1 res; let _ = print_sets_color print_fun (ST ''0x007f00'') (aform.sets_of_ivls B); (_, Pu) \ scaleR2_rep_coll P; let _ = print_sets_color print_fun (ST ''0x7f0000'') (aform.op_image_fst_coll (Pu:::clw_rel aform.appr1_rel):::clw_rel aform.appr_rel); let _ = print_lorenz_color print_fun (ST ''0x7f0000'') (ST ''0x7f0001'') (ST ''0x7f0002'') (ST ''0x7f0003'') [(ST ''0xc60000''), (ST ''0xc60000''), (ST ''0xc60000'')] [(ST ''0xf60000''), (ST ''0xf60000''), (ST ''0xf60000'')] P; let _ = (if r then print_fun (String.implode ((show (ST ''# VERIFIED '') @ show i @ show (ST ''\'')))) else print_fun (String.implode ((show (ST ''# Failed to verify '') @ show i @ show (ST ''\'')) ))); RETURN r }" lemma [autoref_rules]: includes autoref_syntax shows "(shows_prec, shows_prec) \ nat_rel \ nat_rel \ string_rel \ string_rel" "(shows_prec, shows_prec) \ nat_rel \ string_rel \ string_rel \ string_rel" "(String.implode, String.implode) \ string_rel \ Id" by (auto simp: string_rel_def) schematic_goal check_line_core_impl: includes autoref_syntax assumes [autoref_rules]: "(pfi, pf) \ \Id \ unit_rel\option_rel" "(c1i, c1) \ bool_rel" "(ii, i) \ nat_rel" "(m0i, m0) \ \nat_rel\option_rel" "(n0i, n0) \ \nat_rel\option_rel" shows "(nres_of ?f, check_line_core $ pf $ m0 $ n0 $ c1 $ i) \ \bool_rel\nres_rel" unfolding check_line_core_def including art by autoref_monadic concrete_definition check_line_core_impl for pfi m0i n0i c1i ii uses check_line_core_impl lemmas [autoref_rules] = check_line_core_impl.refine definition "c1i_of_res res = sourcei_of_res res \ conefield_of_res res" definition "correct_res res = ((\(x, dx) \ c1i_of_res res. x \ plane_of (Sctn (0, 0, 1) 27) \ dx \ plane_of (Sctn (0, 0, 1) 0) \ ((lorenz.returns_to \ x \ lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e \ (\D. (lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e) \ norm (D dx) \ expansion res * norm dx \ (\res2 \ return_of_res results res. (lorenz.poincare_map \ x, D dx) \ c1_of_res res2 \ norm (D dx) \ preexpansion res2 * norm dx))))))" lemma check_line_nres_c0_correct: "check_line_nres pf m0 n0 c1 res \ SPEC (\(every, Pe, RET). every \ (\x \ sourcei_of_res res. lorenz.poincare_map \ x \ \(source_of_res ` return_of_res results res)))" if NF apply (rule check_line_nres[OF \NF\, le]) apply (auto simp: c1i_of_res_def lorenz.poincare_mapsto_def) subgoal apply (drule bspec, force) apply (auto dest!: spec[where x="1\<^sub>L"]) apply (drule bspec, force) apply (force simp: scaleR2_def image_def vimage_def) done subgoal premises prems for a b c d e tans proof - obtain t where "((c, d, e), t) \ tans - \\<^sub>i (invoke_nf res) \ UNIV" "((c, d, e), tangent_of_deg (min_deg res)) = (\(x, y). (x, blinfun_apply y (1, 0, 0))) ((c, d, e), t)" using prems by (auto simp: sourcei_of_res_def) with prems(6)[rule_format, of "((c, d, e), t)"] prems(3) obtain D x where "tangent_of_deg (min_deg res) = blinfun_apply t (1, 0, 0)" "(c, d, e) returns_to \" "fst ` (tans - \\<^sub>i (invoke_nf res) \ UNIV) \ \\<^sub>l\<^sub>e" "lorenz.return_time \ differentiable at (c, d, e) within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative blinfun_apply D) (at (c, d, e) within \\<^sub>l\<^sub>e)" "x \ b" "(lorenz.poincare_map \ (c, d, e), D o\<^sub>L t) \ x" by auto with prems show ?thesis subgoal apply (auto dest!: bspec[OF _ \x \ b\]) apply (auto simp: scaleR2_def image_def vimage_def) apply (auto simp: subset_iff) by fastforce \\slow\ done qed subgoal for a b c d e f apply (drule bspec[where A="sourcei_of_res res"]) apply force apply (auto dest!: spec[where x="1\<^sub>L"]) apply (drule bspec, force) apply auto apply (auto simp: scaleR2_def image_def vimage_def) apply (auto simp: subset_iff) by fastforce\ \slow\ done lemma cone_conefield[intro, simp]: "cone (conefield a b)" unfolding conefield_alt_def by (rule cone_cone_hull) lemma in_segment_norm_bound: "c \ {a -- b} \ norm c \ max (norm a) (norm b)" apply (auto simp: in_segment max_def intro!: norm_triangle_le) apply (auto simp: algebra_simps intro: add_mono mult_left_mono mult_right_mono) using affine_ineq by blast lemma norm_tangent_of_deg[simp]: "norm (tangent_of_deg d) = 1" by (auto simp: tangent_of_deg_def norm_prod_def) lemma check_line_nres_c1_correct: "check_line_nres pf m0 n0 True res \ SPEC (\(correct, Pe, RET). correct \ correct_res res)" if NF proof (rule check_line_nres[OF \NF\, le], clarsimp, goal_cases) case P: (1 a P tans) let ?tans = "{tangent_of_deg (min_deg res)--tangent_of_deg (max_deg res)}" have tans_plane: "?tans \ UNIV \ UNIV \ {0}" by (auto simp: in_segment tangent_of_deg_def) from P have *: "x \ plane_of (Sctn (0, 0, 1) 27)" if "x \ sourcei_of_res res" for x using that by (auto simp: that sourcei_of_res_def) from tans_plane P have **: "dx \ plane_of (Sctn (0, 0, 1) 0)" if "x \ sourcei_of_res res" "dx \ conefield_of_res res" for x dx proof - from tans_plane that obtain c dx' dy' where c: "dx = c *\<^sub>R (dx', dy', 0)" "(dx', dy', 0) \ ?tans" "c \ 0" unfolding conefield_of_res_def conefield_alt_def cone_hull_expl by auto then show ?thesis by (auto simp: plane_of_def) qed show ?case unfolding correct_res_def proof (intro ballI conjI, clarsimp_all simp add: * ** c1i_of_res_def c1_of_res_def sourcei_of_res_def, goal_cases) case source: (1 x y z dx dy dz) from tans_plane source obtain c dx' dy' where c: "(dx, dy, dz) = c *\<^sub>R (dx', dy', 0)" "(dx', dy', 0) \ ?tans" "c \ 0" unfolding conefield_of_res_def conefield_alt_def cone_hull_expl by auto from c source P obtain t where tans: "((x, y, z), t) \ tans" "blinfun_apply t (1, 0, 0) = (dx', dy', 0)" by auto from P(3) tans(1) source(3) obtain Re D where Re: "(x, y, z) returns_to \" "fst ` (tans - \\<^sub>i (invoke_nf res) \ UNIV) \ \\<^sub>l\<^sub>e" "lorenz.return_time \ differentiable at (x, y, z) within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative blinfun_apply D) (at (x, y, z) within \\<^sub>l\<^sub>e)" "Re \ P" "(lorenz.poincare_map \ (x, y, z), D o\<^sub>L t) \ Re" by (auto simp: lorenz.poincare_mapsto_def dest!: bspec) from P(5)[rule_format, OF \Re \ P\] obtain R em eM Rivls where R: "Re = scaleR2 em eM R" "em > 0" "fst ` R \ \Rivls" "\Rivl. Rivl\Rivls \ \resa\return_of_res results res. Rivl \ source_of_res resa \ c1_entry_correct em R res resa" by auto have "lorenz.poincare_map \ (x, y, z) \ fst ` R" and s2: "(lorenz.poincare_map \ (x, y, z), D o\<^sub>L t) \ scaleR2 em eM R" using Re R by (auto simp: scaleR2_def) then obtain Rivl res' where Rivl: "lorenz.poincare_map \ (x, y, z) \ Rivl" "Rivl \ Rivls" "res' \ return_of_res results res" "Rivl \ source_of_res res'" and c1: "c1_entry_correct em R res res'" using R by force from s2 obtain ed Dt where Dt: "em \ ereal ed" "ereal ed \ eM" "D o\<^sub>L t = ed *\<^sub>R Dt" "(lorenz.poincare_map \ (x, y, z), Dt) \ R" by (force simp: scaleR2_def) then have Dt_simp[simp]: "Dt = inverse ed *\<^sub>R (D o\<^sub>L t)" using \0 < em\ by (cases em) (auto simp: intro!: simp: blinfun.bilinear_simps inverse_eq_divide) from c1[unfolded c1_entry_correct_def, rule_format, OF Dt(4)] obtain dxr dyr where dxrdyr: "blinfun_apply D (dx', dy', 0) /\<^sub>R ed = (dxr, dyr, 0)" "ereal (preexpansion res') \ em * ereal (norm (dxr, dyr, 0::real))" "ereal (expansion res) \ em * ereal (norm (dxr, dyr, 0::real))" "-90 < min_deg res'" "min_deg res' \ max_deg res'" "tan (rad_of (min_deg res')) \ (dyr / dxr)" "(dyr / dxr) \ tan (rad_of (max_deg res'))" "max_deg res' < 90" "0 < dxr" by (auto simp: blinfun.bilinear_simps tans) then obtain emr where emr: "em = ereal emr" "0 < emr" "emr \ ed" "(preexpansion res') \ emr * (norm (dxr, dyr, 0::real))" "(expansion res) \ emr * (norm (dxr, dyr, 0::real))" using \0 < em\ Dt by (cases em) (auto simp: simp: blinfun.bilinear_simps divide_simps prod_eq_iff) from dxrdyr have Ddx'dy': "D (dx', dy', 0) = ed *\<^sub>R (dxr, dyr, 0)" using \0 < em\ Dt by (cases em) (auto simp: simp: blinfun.bilinear_simps divide_simps prod_eq_iff) note \(x, y, z) returns_to \\ moreover note \lorenz.return_time \ differentiable at (x, y, z) within \\<^sub>l\<^sub>e\ moreover note \(lorenz.poincare_map \ has_derivative D) (at (x, y, z) within \\<^sub>l\<^sub>e)\ moreover note \res' \ return_of_res results res\ moreover have "lorenz.poincare_map \ (x, y, z) \ source_of_res res'" using Rivl by force moreover have \0 \ ed\ using Dt \0 < em\ by (cases em) auto have \D (dx, dy, dz) \ conefield_of_res res'\ unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' apply (intro mem_cone, simp_all add: \0 \ ed\ \0 \ c\ tangent_of_deg_def) apply (rule conefield_prod3I) unfolding fun_cong[OF tan_def, symmetric] subgoal by fact subgoal using dxrdyr apply (intro cos_gt_zero_pi) unfolding rad_of_lt_iff rad_of_gt_iff by (auto simp: deg_of_def) subgoal using dxrdyr apply (intro cos_gt_zero_pi) unfolding rad_of_lt_iff rad_of_gt_iff by (auto simp: deg_of_def) subgoal by fact subgoal by fact done moreover have norms_le: "emr * norm (dx', dy', 0::real) * (\c\ * norm (dxr, dyr, 0::real)) \ \ed\ * (\c\ * norm (dxr, dyr, 0::real))" proof - from c(2)[THEN in_segment_norm_bound] have "norm (dx', dy', 0::real) \ 1" by auto also have "\ \ ed / emr" using dxrdyr emr unfolding Ddx'dy' by auto finally show ?thesis using emr by (intro mult_right_mono) (auto simp: divide_simps ac_simps) qed then have "expansion res * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' norm_scaleR apply - apply (rule order_trans) apply (rule mult_right_mono) apply (rule emr) by (auto simp: ac_simps) moreover have "preexpansion res' * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" using norms_le unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' norm_scaleR apply - apply (rule order_trans) apply (rule mult_right_mono) apply (rule emr) by (auto simp: ac_simps) ultimately show ?case by blast qed qed lemma conefield_ne_empyt[simp]: "conefield a b \ {}" by (auto simp: conefield_def conesegment_def cone_hull_empty_iff[symmetric]) lemma in_return_of_resD: "res' \ return_of_res results res \ res' \ set results" by (auto simp: return_of_res_def get_results_def) lemma finite_results_at[intro, simp]: "finite (results_at x)" by (auto simp: results_at_def) lemma lorenz_bounds_lemma: "x returns_to \" "R x \ N" "(R has_derivative DR x) (at x within \\<^sub>l\<^sub>e)" "\c. c \ \ x \ DR x c \ \ (R x)" "\c. c \ \ x \ norm (DR x c) \ \ x * norm c" "\c. c \ \ x \ norm (DR x c) \ \\<^sub>p (R x) * norm c" if "x \ N - \" NF "\res. res \ set results \ correct_res res" proof - from \x \ N - \\ obtain res where res: "res \ set results" "x \ sourcei_of_res res" by (auto simp: N_def sourcei_of_res_def \\<^sub>i_def) then have ne: "c1i_of_res res \ {}" by (auto simp: c1i_of_res_def conefield_of_res_def) from res this obtain dx where dx: "(x, dx) \ c1i_of_res res" by (auto simp: c1i_of_res_def) from that(3)[OF \res \ set _\] have "correct_res res" by simp from this[unfolded correct_res_def, rule_format, OF dx] res obtain res' D where res': "x returns_to \" "lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" "expansion res * norm dx \ norm (D dx)" "res' \ return_of_res results res" "(lorenz.poincare_map \ x, D dx) \ c1_of_res res'" "preexpansion res' * norm dx \ norm (D dx)" by auto show "x returns_to \" by fact show "R x \ N" using res' by (auto simp: R_def N_def N_def c1i_of_res_def c1_of_res_def in_return_of_resD sourcei_of_res_def) show "(R has_derivative DR x) (at x within \\<^sub>l\<^sub>e)" apply (auto simp: R_def DR_def N_def c1_of_res_def in_return_of_resD) apply (subst frechet_derivative_works[symmetric]) apply (rule differentiableI) by fact next fix dx assume "dx \ \ x" then obtain res where res: "res \ set results" and dx: "(x, dx) \ c1_of_res res" by (auto simp: \_def results_at_def c1_of_res_def ) then have dx: "(x, dx) \ c1i_of_res res" using \x \ N - _\ by (auto simp: c1i_of_res_def sourcei_of_res_def c1_of_res_def \\<^sub>i_def) from res dx have ne: "c1i_of_res res \ {}" by (auto simp: c1_of_res_def conefield_of_res_def) from that(3)[OF \res \ set _\] have "correct_res res" by simp from that this[unfolded correct_res_def, rule_format, OF dx] res obtain res' D where res': "x returns_to \" "x \ plane_of (Sctn (0, 0, 1) 27)" "lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" "expansion res * norm dx \ norm (D dx)" "res' \ return_of_res results res" "(lorenz.poincare_map \ x, D dx) \ c1_of_res res'" "preexpansion res' * norm dx \ norm (D dx)" by auto have DRD: "DR x = D" unfolding DR_def apply (rule frechet_derivative_unique_within) apply (subst frechet_derivative_works[symmetric]) apply (rule differentiableI) apply fact apply fact using \x \ plane_of _\ apply safe subgoal for _ _ _ e by (auto simp: \\<^sub>l\<^sub>e_def Basis_prod_def prod_eq_iff plane_of_def prod_eq_iff inner_prod_def intro!: exI[where x="-e/2"]) done have [intro, simp]: "res' \ results_at (lorenz.poincare_map \ x)" using res' by (auto simp: c1_of_res_def results_at_def in_return_of_resD R_def intro!: exI[where x=res']) have [intro, simp]: "res \ results_at x" using res dx by (auto simp: c1i_of_res_def results_at_def sourcei_of_res_def) show "DR x dx \ \ (R x)" unfolding DRD \_def using res' by (auto simp: c1_of_res_def R_def) have "\ x * norm dx \ expansion res * norm dx" by (rule mult_right_mono) (auto simp: \_def) also have "\ \ norm (DR x dx)" unfolding DRD by fact finally show "\ x * norm dx \ norm (DR x dx)" . have "\\<^sub>p (R x) * norm dx \ preexpansion res' * norm dx" by (rule mult_right_mono) (auto simp: \\<^sub>p_def R_def) also have "\ \ norm (DR x dx)" unfolding DRD by fact finally show "\\<^sub>p (R x) * norm dx \ norm (DR x dx)" . qed lemma check_line_core_correct: "check_line_core pf m0 n0 True i \ SPEC (\correct. correct \ correct_res (results ! i))" if [refine_vcg]: NF unfolding check_line_core_def supply [refine_vcg] = check_line_nres_c1_correct[le] by refine_vcg text \The symmetric reduction\ lemma source_of_res_mirror: "(x, y, z) \ source_of_res (mirror_result res) \ (-x, -y, z) \ source_of_res res" by (cases res) (auto simp: source_of_res_def ivl_of_res_def ivl_of_resultrect_def set_of_ivl_def) lemma conefield_of_res_mirror[simp]: "(x, y, z) \ conefield_of_res (mirror_result res) \ (x, y, z) \ conefield_of_res res" by (cases res) (auto simp: conefield_of_res_def ivl_of_res_def) lemma c1_of_res_mirror: "((x, y, z), dx, dy, dz) \ c1_of_res (mirror_result res) \ ((-x, -y, z), dx, dy, dz) \ c1_of_res res" by (auto simp: c1_of_res_def source_of_res_mirror) lemmas [simp] = lorenz_S(2) lemma lorenz_S_idem[simp]: "lorenz_S (lorenz_S x) = (x::R3)" by (auto simp: lorenz_S_def split_beta') lemma lorenz_S_exivl[simp]: "lorenz.existence_ivl0 (lorenz_S X) = lorenz.existence_ivl0 X" using lorenz_S(1)[of _ X] using lorenz_S(1)[of _ "lorenz_S X"] by auto lemma lorenz_S_zero[simp]: "lorenz_S x = 0 \ (x::R3) = 0" by (auto simp: lorenz_S_def split_beta' prod_eq_iff) lemma lorenz_S_returns_toI[simp]: "x returns_to (lorenz_S ` P) \ lorenz_S x returns_to P" apply (auto simp: lorenz.returns_to_def) subgoal premises prems for t proof - have " \\<^sub>F s in at_right 0. s < t" using tendsto_ident_at \0 < t\ by (rule order_tendstoD) then have " \\<^sub>F s in at_right 0. s \ lorenz.existence_ivl0 x" unfolding eventually_at_filter apply eventually_elim using \0 < t\ lorenz.closed_segment_subset_existence_ivl[OF prems(3)] by (auto simp: closed_segment_eq_real_ivl subset_iff) then show ?thesis using prems(1) by eventually_elim force qed done lemma lorenz_S_returns_to[simp]: "lorenz_S x returns_to P \ x returns_to (lorenz_S ` P)" using lorenz_S_returns_toI[of P x] lorenz_S_returns_toI[of "lorenz_S ` P" "lorenz_S x"] by (auto simp: image_image) lemma lorenz_S_image_Sigma[simp]: "lorenz_S ` \ = \" apply (auto simp: \_def lorenz_S_def) apply (rule image_eqI) apply (rule lorenz_S_idem[symmetric]) apply (auto simp: \_def lorenz_S_def) done lemma linear_lorenz_S: "linear lorenz_S" by unfold_locales (auto simp: lorenz_S_def) lemma inj_lorenz_S: "inj_on (lorenz_S::R3 \ _) G" by (rule inj_onI) (auto simp: lorenz_S_def prod_eq_iff) lemma lorenz_S_return_time: "lorenz.return_time P (lorenz_S x) = lorenz.return_time (lorenz_S ` P) x" if "x returns_to (lorenz_S ` P)" "closed P" proof - from lorenz.returns_toE[OF that(1)] obtain t0 t1 where f: "0 < t0" "t0 \ t1" " t1 \ lorenz.existence_ivl0 x" "lorenz.flow0 x t1 \ lorenz_S ` P" "\t. 0 < t \ t < t0 \ lorenz.flow0 x t \ lorenz_S ` P" by auto have [simp]: "lorenz.return_time (lorenz_S ` P) x \ lorenz.existence_ivl0 x" by (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl inj_lorenz_S) have c': "closed (lorenz_S ` P)" by (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl lorenz.return_time_pos inj_lorenz_S) show ?thesis using f(1-4) using lorenz.return_time_returns[OF that(1) c'] apply (intro lorenz.return_time_eqI) apply (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl lorenz.return_time_pos c' inj_lorenz_S) subgoal premises prems for a b c d e f g proof - have [simp]: "a \ lorenz.existence_ivl0 x" using _ that(1) apply (rule lorenz.less_return_time_imp_exivl) using prems that(2) c' by auto have "lorenz.return_time (lorenz_S ` P) x \ a" apply (rule lorenz.return_time_le) using prems apply (auto intro!: that closed_injective_linear_image linear_lorenz_S lorenz.return_time_exivl lorenz.return_time_pos c' inj_lorenz_S) apply (rule image_eqI) apply (rule lorenz_S_idem[symmetric]) by auto then show ?thesis using prems by simp qed done qed lemma lorenz_S_poincare_map: "lorenz.poincare_map P (lorenz_S x) = lorenz_S (lorenz.poincare_map (lorenz_S ` P) x)" if "x returns_to (lorenz_S ` P)" "closed P" using that unfolding lorenz.poincare_map_def apply (auto simp: lorenz_S_return_time) apply (subst lorenz_S) by (auto intro!: lorenz.return_time_exivl that closed_injective_linear_image linear_lorenz_S inj_lorenz_S) lemma [continuous_intros]: "isCont (lorenz_S::_\R3) x" "continuous_on (G::R3 set) lorenz_S" by (auto simp:lorenz_S_def[abs_def] split_beta' continuous_intros) lemma filtermap_lorenz_S_le: "filtermap lorenz_S (at x within lorenz_S ` P) \(at (lorenz_S x::R3) within P)"\ \TODO: generalize!\ unfolding at_within_def apply (auto simp: intro!: antisym filtermap_inf[le] filtermap_inf[ge]) apply (rule inf.coboundedI1) apply (subst filtermap_nhds_open_map) apply (auto simp: intro!: invariance_of_domain inj_lorenz_S continuous_intros) apply (rule inf.coboundedI2) apply (auto simp: image_image ) apply (auto simp: lorenz_S_def split_beta')[] done lemma filtermap_lorenz_S_eq: "filtermap lorenz_S (at (x::R3) within lorenz_S ` P) = (at (lorenz_S x::R3) within P)" apply (rule antisym) using filtermap_lorenz_S_le[of "x" P] apply simp subgoal proof - have "filtermap lorenz_S (at (lorenz_S x) within P) \ filtermap lorenz_S (filtermap lorenz_S (at x within lorenz_S ` P))" using filtermap_lorenz_S_le[of "lorenz_S x" "lorenz_S ` P"] by (auto simp: image_image filtermap_filtermap) then show ?thesis apply (subst (asm) filtermap_mono_strong) by (auto simp: inj_lorenz_S) qed done lemma norm_lorenz_S[simp]: "norm (lorenz_S x) = norm x" by (auto simp: lorenz_S_def norm_prod_def split_beta') lemma bl_lorenz_S: "bounded_linear (lorenz_S)" by unfold_locales (auto simp: lorenz_S_def norm_prod_def intro!: exI[where x=1]) lemma filtermap_lorenz_S_eq_bot[simp]: "filtermap (lorenz_S::R3\_) F = bot \ F = bot" apply (auto simp: ) apply (subst (asm) filtermap_bot[symmetric]) apply (subst (asm) filtermap_eq_strong) by (auto simp: inj_lorenz_S) lemma netlimit_filtermap[simp]: "at x within X \ bot \ netlimit (filtermap lorenz_S (at x within X)) = lorenz_S (x::R3)" apply (rule tendsto_Lim) unfolding filterlim_filtermap apply simp by (auto intro!: tendsto_eq_intros simp: split_beta' lorenz_S_def[abs_def]) lemma lorenz_S_halfspace [simp]: "lorenz_S ` \\<^sub>l\<^sub>e = \\<^sub>l\<^sub>e" apply (auto simp: \\<^sub>l\<^sub>e_def lorenz_S_def[abs_def]) apply (rule image_eqI) apply auto apply (rule sym) apply (rule minus_minus) apply (rule minus_minus[symmetric]) done lemma closure_Sigma_le_eq: "closure \\<^sub>l\<^sub>e = \\<^sub>l\<^sub>e" proof (rule closure_closed) have "\\<^sub>l\<^sub>e = {x. x \ (0, 0, 1) \ 27}" by (auto simp: \\<^sub>l\<^sub>e_def ) also have "closed \" by (rule closed_halfspace_component_le) finally show "closed \\<^sub>l\<^sub>e" . qed lemma closure_Sigma_le[simp]: "closure (\\<^sub>l\<^sub>e - {x}) = \\<^sub>l\<^sub>e" proof (cases "x \ \\<^sub>l\<^sub>e") case that: True have "closure \\<^sub>l\<^sub>e \ closure (insert x (\\<^sub>l\<^sub>e - {x}))" by (rule closure_mono) auto also have "\ = insert x (closure (\\<^sub>l\<^sub>e - {x}))" apply (subst closure_insert) by simp also have "x \ closure (\\<^sub>l\<^sub>e - {x})" apply (rule closed_sequentially[where f="\n. x - (0, 0, inverse (Suc n))"]) apply (rule closed_closure) subgoal apply (auto simp: ) apply (rule subsetD) apply (rule closure_subset) using that apply (auto simp: \\<^sub>l\<^sub>e_def prod_eq_iff) apply (rule order_trans) apply (rule diff_right_mono) apply (assumption) apply simp done subgoal apply (rule tendsto_eq_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule tendsto_intros) apply (rule LIMSEQ_inverse_real_of_nat) by (auto simp: prod_eq_iff) done then have "insert x (closure (\\<^sub>l\<^sub>e - {x})) \ closure (\\<^sub>l\<^sub>e - {x})" by auto finally have "closure \\<^sub>l\<^sub>e \ closure (\\<^sub>l\<^sub>e - {x})" . moreover have "closure (\\<^sub>l\<^sub>e - {x}) \ closure (\\<^sub>l\<^sub>e)" by (rule closure_mono) auto ultimately have "closure (\\<^sub>l\<^sub>e - {x}) = closure (\\<^sub>l\<^sub>e)" by simp also have "\ = \\<^sub>l\<^sub>e" by (rule closure_Sigma_le_eq) finally show ?thesis . next case False then show ?thesis apply simp apply (rule closure_Sigma_le_eq) done qed lemma lorenz_S_return_time_has_derivative: assumes "(lorenz.return_time \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" and "lorenz.returns_to \ x" and "x \ \\<^sub>l\<^sub>e" shows "(lorenz.return_time \ has_derivative D o lorenz_S) (at (lorenz_S x) within \\<^sub>l\<^sub>e)" proof - have [simp]: "\trivial_limit (at x within \\<^sub>l\<^sub>e)" unfolding at_within_eq_bot_iff using assms by simp interpret bounded_linear "lorenz_S::R3\_" by (rule bl_lorenz_S) have "\\<^sub>F x in at x within \\<^sub>l\<^sub>e. (x::R3) returns_to \" by (blast intro: lorenz.eventually_returns_to_continuousI has_derivative_continuous assms) then have "\\<^sub>F y in at x within \\<^sub>l\<^sub>e. inverse (norm (y - x)) * (lorenz.return_time \ y - lorenz.return_time \ x - D (y - x)) = inverse (norm (lorenz_S y - lorenz_S x)) * (lorenz.return_time \ (lorenz_S y) - lorenz.return_time \ (lorenz_S x) - D (y - x))" by eventually_elim (auto simp: lorenz_S_return_time assms diff[symmetric]) then show ?thesis using assms apply (subst filtermap_lorenz_S_eq[symmetric]) apply (auto simp: has_derivative_def filterlim_filtermap) unfolding o_def apply (rule bounded_linear_compose, assumption, rule bl_lorenz_S) unfolding diff lorenz_S_idem apply (auto simp: Lim_ident_at) apply (blast intro: Lim_transform_eventually) done qed lemma lorenz_S_return_time_differentiable: "lorenz.return_time \ differentiable at (lorenz_S x) within \\<^sub>l\<^sub>e" if "lorenz.return_time \ differentiable at x within \\<^sub>l\<^sub>e" "lorenz.returns_to \ x" "x \ \\<^sub>l\<^sub>e" proof - from that obtain D where "(lorenz.return_time \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" by (auto simp: differentiable_def) then have "(lorenz.return_time \ has_derivative D o lorenz_S) (at (lorenz_S x) within \\<^sub>l\<^sub>e)" by (rule lorenz_S_return_time_has_derivative) fact+ then show ?thesis by (auto simp: differentiable_def) qed lemma lorenz_S_has_derivative: "(lorenz_S has_derivative lorenz_S) (at (x::R3) within X)" by (auto simp: lorenz_S_def[abs_def] split_beta' intro!: derivative_eq_intros) lemma lorenz_S_poincare_map_has_derivative: assumes "(lorenz.poincare_map \ has_derivative D) (at x within \\<^sub>l\<^sub>e)" "(lorenz.return_time \ has_derivative Dr) (at x within \\<^sub>l\<^sub>e)" "lorenz.returns_to \ x" "x \ \\<^sub>l\<^sub>e" shows "(lorenz.poincare_map \ has_derivative lorenz_S o D o lorenz_S) (at (lorenz_S x) within \\<^sub>l\<^sub>e)" proof - have [simp]: "\trivial_limit (at x within \\<^sub>l\<^sub>e)" unfolding at_within_eq_bot_iff using assms by simp interpret bounded_linear "lorenz_S::R3\_" by (rule bl_lorenz_S) have "\\<^sub>F x in at x within \\<^sub>l\<^sub>e. (x::R3) returns_to \" by (blast intro: lorenz.eventually_returns_to_continuousI has_derivative_continuous assms) then have "\\<^sub>F y in at x within \\<^sub>l\<^sub>e. (lorenz_S (lorenz.poincare_map \ y) - lorenz_S (lorenz.poincare_map \ x) - lorenz_S (D (y - x))) /\<^sub>R norm (y - x) = (lorenz.poincare_map \ (lorenz_S y) - lorenz.poincare_map \ (lorenz_S x) - lorenz_S (D (y - x))) /\<^sub>R norm (lorenz_S y - lorenz_S x)" by eventually_elim (auto simp: lorenz_S_return_time lorenz_S_poincare_map assms diff[symmetric]) then show ?thesis using has_derivative_compose[OF assms(1) lorenz_S_has_derivative] assms apply (subst filtermap_lorenz_S_eq[symmetric]) apply (auto simp: has_derivative_def filterlim_filtermap) unfolding o_def apply (rule bounded_linear_compose, rule bl_lorenz_S) apply (rule bounded_linear_compose, assumption, rule bl_lorenz_S) unfolding diff lorenz_S_idem apply (auto simp: Lim_ident_at) apply (blast intro: Lim_transform_eventually) done qed lemma [simp]: "expansion (mirror_result res) = expansion res" by (cases res) auto lemma lorenz_S_on_plane: "lorenz_S (dx, dy, 0::real) = - (dx, dy, 0)" by (auto simp: lorenz_S_def ) lemma mirror_result_idem[simp]: "mirror_result (mirror_result x) = x" by (cases x) (auto simp: mirror_result_def) lemma mirror_in_set: "x \ set results \ mirror_result x \ set results" by (auto simp: results_def symmetrize_def) lemma mirror_result_in: "mirror_result res2 \ return_of_res results (mirror_result res)" if "res2 \ return_of_res results res" proof - from that have "res2 \ set results" by (rule in_return_of_resD) from mirror_in_set[OF this] have "mirror_result res2 \ set results" . then show ?thesis apply (cases res2; cases res) using that by (auto simp: return_of_res_def get_results_def) qed lemma in_source_of_res_mirrorI: "(x::R3) \ source_of_res (mirror_result (r))" if "lorenz_S x \ source_of_res r" using that apply (cases r; cases x) by (auto simp: source_of_res_def set_of_ivl_def ivl_of_res_def lorenz_S_def less_eq_prod_def ivl_of_resultrect_def) lemma conefield_of_res_mirror_simp[simp]: "conefield_of_res (mirror_result res2) = conefield_of_res res2" by (cases res2) (auto simp: conefield_of_res_def) lemma lorenz_minus_planeI: "lorenz_S (- x) = x" if "snd (snd (x::R3)) = 0" using that by (auto simp: lorenz_S_def split_beta' prod_eq_iff) lemma preexpansion_mirror_result[simp]: "preexpansion (mirror_result res2) = preexpansion res2" by (cases res2) (auto simp: ) lemma lorenz_S_tendsto_0I: "(lorenz.flow0 (lorenz_S x) \ 0) at_top" if "{0..} \ lorenz.existence_ivl0 x" "(lorenz.flow0 x \ 0) at_top" proof (rule Lim_transform_eventually) have "\\<^sub>F s in at_top. (s::real) \ 0" using eventually_ge_at_top by blast then show "\\<^sub>F s in at_top. lorenz_S (lorenz.flow0 x s) = lorenz.flow0 (lorenz_S x) s" by eventually_elim (use that in auto) show "((\s. lorenz_S (lorenz.flow0 x s)) \ 0) at_top" unfolding Zfun_def[symmetric] by (rule bounded_linear.tendsto_zero[OF bl_lorenz_S that(2)]) qed lemma lorenz_S_tendsto_0_iff: "(lorenz.flow0 (lorenz_S x) \ 0) at_top \ (lorenz.flow0 x \ 0) at_top" if "{0..} \ lorenz.existence_ivl0 x" using lorenz_S_tendsto_0I[of x, OF that] lorenz_S_tendsto_0I[of "lorenz_S x"] that by auto lemma lorenz_S_eq_iff[simp]: "lorenz_S y = lorenz_S x \ y = x" for x y::"real*real*real" by (auto simp: lorenz_S_def split: prod.splits) lemma lorenz_S_\: "lorenz_S x \ \ \ x \ \" apply (auto simp: \_def lorenz_S_tendsto_0_iff ) subgoal for t apply (auto simp: dest!: spec[where x=t]) apply (subst (asm) lorenz_S) apply auto apply (subst (asm) (2) lorenz_S_image_Sigma[symmetric]) by (simp del: lorenz_S_image_Sigma) subgoal for t apply (auto simp: dest!: spec[where x=t]) apply (subst (asm) lorenz_S) apply auto apply (subst (asm) lorenz_S_image_Sigma[symmetric]) apply (auto simp del: lorenz_S_image_Sigma) done done lemma sourcei_of_res_mirror: "(x, y, z) \ sourcei_of_res (mirror_result res) \ (-x, -y, z) \ sourcei_of_res res" using lorenz_S_\[of "(x, y, z)"] by (cases res) (auto simp: source_of_res_def sourcei_of_res_def ivl_of_res_def ivl_of_resultrect_def set_of_ivl_def \\<^sub>i_def lorenz_S_def) lemma c1i_of_res_mirror: "((x, y, z), dx, dy, dz) \ c1i_of_res (mirror_result res) \ ((-x, -y, z), dx, dy, dz) \ c1i_of_res res" by (auto simp: c1i_of_res_def sourcei_of_res_mirror) lemma correct_res_mirror_result: "correct_res (mirror_result res)" if "correct_res res" unfolding correct_res_def proof (clarsimp simp add: c1i_of_res_mirror, goal_cases) case (1 x y z dx dy dz) then have 1: "(lorenz_S (x, y, z), dx, dy, dz) \ c1i_of_res res" by (auto simp: lorenz_S_def) from that[unfolded correct_res_def, rule_format, OF 1, simplified] have "(lorenz_S (x, y, z)) \ plane_of (Sctn (0, 0, 1) 27)" "(dx, dy, dz) \ plane_of (Sctn (0, 0, 1) 0)" by auto then have plane: "(x, y, z) \ plane_of (Sctn (0, 0, 1) 27)" "(dx, dy, dz) \ plane_of (Sctn (0, 0, 1) 0)" by (auto simp: plane_of_def lorenz_S_def) then show ?case proof (clarsimp, goal_cases) case mem: 1 with that[unfolded correct_res_def, rule_format, OF 1, simplified] obtain D res2 where D: "lorenz_S (x, y, z) returns_to \" "lorenz.return_time \ differentiable at (lorenz_S (x, y, z)) within \\<^sub>l\<^sub>e" "(lorenz.poincare_map \ has_derivative D) (at (lorenz_S (x, y, z)) within \\<^sub>l\<^sub>e)" "expansion res * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" "res2 \ return_of_res results res" "(lorenz.poincare_map \ (lorenz_S (x, y, z)), D (dx, dy, dz)) \ c1_of_res res2" "preexpansion res2 * norm (dx, dy, dz) \ norm (D (dx, dy, dz))" by auto from plane have S_le: "lorenz_S (x, y, z) \ \\<^sub>l\<^sub>e" by (auto simp: \\<^sub>l\<^sub>e_def plane_of_def lorenz_S_def) interpret linear D by (rule has_derivative_linear; fact) have ret: "(x, y, z) returns_to \" using D(1) lorenz_S_returns_to by simp moreover have "lorenz.return_time \ differentiable at (x, y, z) within \\<^sub>l\<^sub>e" using lorenz_S_return_time_differentiable[OF D(2) D(1) S_le] by auto moreover from D obtain Dr where Dr: "(lorenz.return_time \ has_derivative Dr) (at (lorenz_S (x, y, z)) within \\<^sub>l\<^sub>e)" by (auto simp: differentiable_def) let ?D = "lorenz_S \ D \ lorenz_S" have "(lorenz.poincare_map \ has_derivative ?D) (at (x, y, z) within \\<^sub>l\<^sub>e)" using lorenz_S_poincare_map_has_derivative[OF D(3) Dr D(1) S_le] by auto moreover from plane have [simp]: "dz = 0" by (auto simp: plane_of_def) have "expansion (mirror_result res) * norm (dx, dy, dz) \ norm (?D (dx, dy, dz))" using D(4) apply (auto simp: ) unfolding lorenz_S_on_plane neg by simp moreover have \mirror_result res2 \ return_of_res results (mirror_result res)\ using D(5) by (rule mirror_result_in) moreover have "(lorenz.poincare_map \ (x, y, z), ?D (dx, dy, dz)) \ c1_of_res (mirror_result res2)" using D(6) apply (subst (asm) lorenz_S_poincare_map) apply auto apply fact apply (auto simp: c1_of_res_def in_source_of_res_mirrorI) unfolding lorenz_S_on_plane neg apply (subst lorenz_minus_planeI) apply (auto simp: conefield_of_res_def conefield_alt_def cone_hull_expl in_segment tangent_of_deg_def) done moreover have "preexpansion (mirror_result res2) * norm (dx, dy, dz) \ norm (?D (dx, dy, dz))" using D(7) apply (auto simp: ) unfolding lorenz_S_on_plane neg by simp ultimately show ?case by (force intro!: exI[where x = ?D] bexI[where x="mirror_result res2"]) qed qed lemma reduce_lorenz_symmetry: "Ball (set results) correct_res" if "Ball (set coarse_results) correct_res" using that by (auto simp: results_def symmetrize_def intro!: correct_res_mirror_result) end subsection \Code Generation\ definition [code_abbrev]: "my_divide_integer (i::integer) (j::integer) = i div j" code_printing constant my_divide_integer \ (SML) "IntInf.div/ (_,/ _)" subsection \Tuning code equations\ definition mult_twopow_int::"int \ int \ int" where "mult_twopow_int x n = x * (power_int 2 n)" definition div_twopow_int :: "int \ int \ int" where "div_twopow_int x n = x div (power_int 2 n)" context includes integer.lifting begin lift_definition mult_twopow_integer :: "integer \ integer \ integer" is mult_twopow_int . lift_definition div_twopow_integer :: "integer \ integer \ integer" is div_twopow_int . end lemma compute_float_round_down[code]: "float_round_down prec (Float m e) = (let d = bitlen \m\ - int prec - 1 in if 0 < d then Float (div_twopow_int m d) (e + d) else Float m e)" including float.lifting using Float.compute_float_down[of "Suc prec - bitlen \m\ - e" m e, symmetric] by transfer (auto simp add: field_simps abs_mult log_mult bitlen_alt_def truncate_down_def div_twopow_int_def power_int_def cong del: if_weak_cong) lemma compute_float_plus_down[code]: fixes p::nat and m1 e1 m2 e2::int shows "float_plus_down p (Float m1 e1) (Float m2 e2) = (if m1 = 0 then float_round_down p (Float m2 e2) else if m2 = 0 then float_round_down p (Float m1 e1) else (if e1 \ e2 then (let k1 = Suc p - nat (bitlen \m1\) in if bitlen \m2\ > e1 - e2 - k1 - 2 then float_round_down p ((Float m1 e1) + (Float m2 e2)) else float_round_down p (Float (mult_twopow_int m1 (int k1 + 2) + sgn m2) (e1 - int k1 - 2))) else float_plus_down p (Float m2 e2) (Float m1 e1)))" using Float.compute_float_plus_down[of p m1 e1 m2 e2] by (auto simp: mult_twopow_int_def Let_def power_int_def nat_add_distrib) subsection \Codegen\ definition "is_dRETURN_True x = (case x of dRETURN b \ b | _ \ False)" definition "file_output_option s f = (case s of None \ f None | Some s \ file_output (String.implode s) (\pf. f (Some pf)))" definition "check_line_lookup_out s m0 n0 c1 i = is_dRETURN_True (file_output_option s (\pf. check_line_core_impl pf m0 n0 c1 i))" fun alternating where "alternating [] xs = xs" | "alternating xs [] = xs" | "alternating (x#xs) (y#ys) = x#y#alternating xs ys" definition "ordered_lines = alternating (rev [0..<222]) ([222..<400])" \ \the hard ones ``first'', potentially useless due to nondeterministic \Parallel.map\\ definition "parallel_check filenameo m n c1 ns = Parallel.forall (\i. let _ = print (String.implode (''# Starting '' @ show i @ ''\'')); b = check_line_lookup_out (map_option (\f. f @ show i) filenameo) (Some m) (Some n) c1 i; _ = if b then print (String.implode (''# Success: '' @ show i @ ''\'')) else print (String.implode (''# Failed: '' @ show i @ ''\'')) in b ) ns" ML \val check_line = @{computation_check terms: Trueprop parallel_check ordered_lines check_line_core_impl check_line_lookup_out (* bool *) True False (* num *) Num.One Num.Bit0 Num.Bit1 (* nat *) Suc "0::nat" "1::nat" "numeral::num\nat" (* int / integer*) "numeral::num\int" "numeral::num\integer" "uminus::_\int" "uminus::_\integer" int_of_integer integer_of_int "0::int" "1::int" (* Pairs *) "Pair::_ \ _\ (real list \ real list)" "Pair::_\_\(real list \ real list) \ real list sctn" "Pair::_\_\((real list \ real list) \ real list sctn) list \ real aform reach_options" (* Option *) "None::nat option" "Some::_\nat option" "None::string option" "Some::_\string option" (* Lists *) "Nil::real list" "Cons::_\_\real list" "Nil::nat list" "Cons::_\_\nat list" "Nil::real aform list" "Cons::_\_\real aform list" "Nil::((real list \ real list) \ real list sctn) list" "Cons::_\_\((real list \ real list) \ real list sctn) list" "Nil::(((real list \ real list) \ real list sctn) list \ real aform reach_options)list" "Cons::_\_\(((real list \ real list) \ real list sctn) list \ real aform reach_options)list" (* String *) String.Char String.implode "Cons::char \ char list \ char list" "Nil::char list" (* float *) Float float_of_int float_of_nat (* real *) "numeral::num\real" "real_of_float" "(/)::real\real\real" "uminus::real\_" real_divl real_divr real_of_int (* section *) "Sctn::_\_\real list sctn" (* aform *) "aforms_of_ivls::_\_\real aform list" (* input *) coarse_results (* modes *) xsec xsec' ysec ysec' zsec zsec' zbucket lookup_mode ro ro_outer mode_outer (* unit *) "()" }\ lemma is_dRETURN_True_iff[simp]: "is_dRETURN_True x \ (x = dRETURN True)" by (auto simp: is_dRETURN_True_def split: dres.splits) lemma check_line_core_impl_True: "check_line_core_impl pfo m n True i = dRETURN True \ NF \ correct_res (results ! i)" apply (cases "check_line_core_impl pfo m n True i") using check_line_core_correct[of pfo m n i] check_line_core_impl.refine[of pfo pfo True True i i m m n n] apply (auto simp: nres_rel_def) apply (drule order_trans[where y="check_line_core pfo m n True i"]) apply assumption by auto lemma check_line_lookup_out: "correct_res (results ! i)" if "\s m n. check_line_lookup_out s m n True i" NF using that by (auto simp: check_line_lookup_out_def file_output_iff check_line_core_impl_True file_output_option_def split: dres.splits option.splits) definition "check_lines c1 ns = list_all (\i. \s m n. check_line_lookup_out s m n c1 i) ns" lemma check_linesI: "check_lines c1 ns" if "parallel_check s m n c1 ns" using that by (auto simp: parallel_check_def check_lines_def list_all_iff) subsection \Automate generation of lemmas\ lemma length_coarse_results[simp]: "length coarse_results = 400" by (simp add: coarse_results_def) lemma correct_res_coarse_resultsI: "correct_res (results ! i) \ i < 400 \ correct_res (coarse_results ! i)" by (auto simp: results_def symmetrize_def nth_append) lemma Ball_coarseI: "Ball (set coarse_results) correct_res" if NF "check_lines True xs" "set xs = {..<400}" using that by (force simp: check_lines_def list_all_iff in_set_conv_nth intro!: correct_res_coarse_resultsI check_line_lookup_out) ML \map_option (using_master_directory_term @{context}) (SOME "a")\ ML \ fun mk_optionT ty = Type (@{type_name "option"}, [ty]) fun mk_None ty = Const (@{const_name "None"}, mk_optionT ty) fun mk_Some ty x = Const (@{const_name "Some"}, ty --> mk_optionT ty) $ x fun mk_option ty _ NONE = mk_None ty | mk_option ty f (SOME x) = mk_Some ty (f x) fun check_lines_tac' s m n ctxt = resolve_tac ctxt [Thm.instantiate ([], [("s", @{typ "string option"}, mk_option @{typ string} (using_master_directory_term ctxt) s), ("m", @{typ nat}, HOLogic.mk_nat m), ("n", @{typ nat}, HOLogic.mk_nat n)] |> map (fn (s, ty, t) => (((s, 0), ty), Thm.cterm_of ctxt t))) @{thm check_linesI}] THEN' CONVERSION (check_line ctxt) THEN' resolve_tac ctxt @{thms TrueI} \ method_setup parallel_check = \ Scan.lift (Parse.maybe Parse.string) -- Scan.lift Parse.nat -- Scan.lift Parse.nat >> (fn ((s, m), n) => fn ctxt => SIMPLE_METHOD' (check_lines_tac' s m n ctxt)) \ lemma lorenz_bounds_lemma_asym: "\x \ N - \. x returns_to \" "R ` (N - \) \ N" "\x \ N - \. (R has_derivative DR x) (at x within \\<^sub>l\<^sub>e)" "\x \ N - \. DR x ` \ x \ \ (R x)" "\x \ N - \. \c \ \ x. norm (DR x c) \ \ x * norm c" "\x \ N - \. \c \ \ x. norm (DR x c) \ \\<^sub>p (R x) * norm c" if NF "Ball (set results) correct_res" using that by (auto intro!: lorenz_bounds_lemma) end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy b/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy @@ -1,2707 +1,2708 @@ theory Example_Utilities imports Init_ODE_Solver begin definition "true_form = Less (floatarith.Num 0) (floatarith.Num 1)" lemma open_true_form[intro, simp]: "open_form true_form" by (auto simp: true_form_def) lemma max_Var_form_true_form[simp]: "max_Var_form true_form = 0" by (auto simp: true_form_def) lemma interpret_form_true_form[simp]: "interpret_form true_form = (\_. True)" by (auto simp: true_form_def) lemmas [simp] = length_aforms_of_ivls declare INF_cong_simp [cong] SUP_cong_simp [cong] image_cong_simp [cong del] declare [[ cd_patterns "_ = interpret_floatariths ?fas _" "_ = interpret_floatarith ?fa _"]] concrete_definition reify_example for i j k uses reify_example hide_const (open) Print.file_output definition "file_output s f = (if s = STR '''' then f (\_. ()) else if s = STR ''-'' then f print else Print.file_output s f)" definition "aforms_of_point xs = aforms_of_ivls xs xs" definition "unit_matrix_list D = concat (map (\i. map (\j. if i = j then 1 else 0::real) [0..) l x \ list_all2 (\) x u}" context includes lifting_syntax begin lemma list_interval_transfer[transfer_rule]: "((list_all2 A) ===> (list_all2 A) ===> rel_set (list_all2 A)) list_interval list_interval" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding list_interval_def by transfer_prover end lemma in_list_interval_lengthD: "x \ list_interval a b \ length x = length a" by (auto simp: list_interval_def list_all2_lengthD) context includes floatarith_notation begin definition "varvec_fas' D C = ((map Var [0..b. (map (\i. (Num (C i)) + Var (D + D * D) * (mvmult_fa D D (map Var [D..i. (map (\j. (Num (C i)) + Var (D + D * D) * Var (D + D * i + j)) [0.. \for illustration\ assumes[simp]: "D=3" "rf = real_of_float" shows "interpret_floatariths (varvec_fas D (\i. [a, b, c] ! i)) [a, b, c, d11, d12, d13, d21, d22, d23, d31, d32, d33, 2] = [rf a, rf b, rf c, rf a + 2 * rf d11, rf a + 2 * rf d12, rf a + 2 * rf d13, rf b + 2 * rf d21, rf b + 2 * rf d22, rf b + 2 * rf d23, rf c + 2 * rf d31, rf c + 2 * rf d32, rf c + 2 * rf d33]" by (simp add: varvec_fas_def mvmult_fa_def eval_nat_numeral) definition "vareq_projections n \ \dimension\ ps \ \pairs of coordinates to project onto\ ds \ \partial derivatives w.r.t. which variables\ cs \ \(color) coding for partial derivatives\ = [(i + n * (x + 1)::nat, i + n * (y + 1), c). (i, c) \ zip ds cs, (x, y) \ ps]" definition "varvec_aforms_line D X line = approx_floatariths 30 (varvec_fas D (\i. float_of (fst (X ! i)))) (take (D + D*D) X @ line)" definition "varvec_aforms_head D X s = varvec_aforms_line D X (aforms_of_point [s])" definition "varvec_aforms_vec D X s = varvec_aforms_line D (map (\x. (fst x, zero_pdevs)) X) [aform_of_ivl 0 s]" definition "shows_aforms_vareq n \ \dimension\ ps \ \pairs of coordinates to project onto\ ds \ \partial derivatives w.r.t. which variables\ csl \ \color coding for partial derivatives ('arrow' heads)\ csh \ \color coding for partial derivatives (lines)\ s \ \scale vectors for partial derivatives\ (no_str::string) \ \default string if no C1 info is present\ X \ \affine form with C1 info\ = (case (varvec_aforms_head n X s, varvec_aforms_vec n X s) of (Some X, Some Y) \ shows_sep (\(x, y, c). shows_segments_of_aform x y X c) shows_nl (vareq_projections n ps ds csl) o shows_nl o shows_sep (\(x, y, c). shows_segments_of_aform x y Y c) shows_nl (vareq_projections n ps ds csh) o shows_nl | _ \ shows_string no_str o shows_nl)" abbreviation "print_string s \ print (String.implode s)" abbreviation "print_show s \ print_string (s '''')" value [code] "print_show (shows_aforms_vareq 3 [(x, y). x \ [0..<3], y \ [0..<3], x < y] [0..<3] [''0x0000ff'', ''0x00ff00'', ''0xff0000''] [''0x0000ff'', ''0x00ff00'', ''0xff0000''] (FloatR 1 (-1)) ''# no C1 info'' ((((\(a, b). aforms_of_ivls a b) (with_unit_matrix 3 ([10, 20, 30], [12, 22, 32]))))))" method_setup guess_rhs = \ let fun compute ctxt var lhs = let val lhs' = Code_Evaluation.dynamic_value_strict ctxt lhs; val clhs' = Thm.cterm_of ctxt lhs'; val inst = Thm.instantiate ([], [(var, clhs')]); in PRIMITIVE inst end; fun eval_schematic_rhs ctxt t = (case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) t of SOME (lhs, Var var) => compute ctxt var lhs | _ => no_tac); in Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (SUBGOAL (fn (t, _) => eval_schematic_rhs ctxt t)))) end \ lemma length_aforms_of_point[simp]: "length (aforms_of_point xs) = length xs" by (auto simp: aforms_of_point_def) definition "aform2d_plot_segments x y a = shows_segments_of_aform x y a ''0x000000''" lemma list_of_eucl_prod[simp]: "list_of_eucl (x, y) = list_of_eucl x @ list_of_eucl y" by (auto simp: list_of_eucl_def Basis_list_prod_def intro!: nth_equalityI) lemma list_of_eucl_real[simp]: "list_of_eucl (x::real) = [x]" by (auto simp: list_of_eucl_def Basis_list_real_def) lemma Joints_aforms_of_ivls_self[simp]: "xs \ Joints (aforms_of_ivls xs xs)" by (auto intro!: aforms_of_ivls) lemma Joints_aforms_of_ivls_self_eq[simp]: "Joints (aforms_of_ivls xs xs) = {xs}" apply (auto ) by (auto simp: aforms_of_ivls_def Joints_def valuate_def aform_val_def intro!: nth_equalityI) lemma local_lipschitz_c1_euclideanI: fixes T::"real set" and X::"'a::euclidean_space set" and f::"real \ 'a \ 'a" assumes f': "\t x. t \ T \ x \ X \ (f t has_derivative f' t x) (at x)" assumes cont_f': "\i. i \ Basis \ continuous_on (T \ X) (\(t, x). f' t x i)" assumes "open T" assumes "open X" shows "local_lipschitz T X f" using assms apply (intro c1_implies_local_lipschitz[where f'="\(t, x). Blinfun (f' t x)"]) apply (auto simp: bounded_linear_Blinfun_apply has_derivative_bounded_linear split_beta' intro!: has_derivative_Blinfun continuous_on_blinfun_componentwise) apply (subst continuous_on_cong[OF refl]) defer apply assumption apply auto apply (subst bounded_linear_Blinfun_apply) apply (rule has_derivative_bounded_linear) by auto definition "list_aform_of_aform (x::real aform) = (fst x, list_of_pdevs (snd x))" primrec split_aforms_list:: "(real aform) list list \ nat \ nat \ (real aform) list list" where "split_aforms_list Xs i 0 = Xs" | "split_aforms_list Xs i (Suc n) = split_aforms_list (concat (map (\x. (let (a, b) = split_aforms x i in [a, b])) Xs)) i n" definition "shows_aforms x y c X = shows_lines (map (\b. (shows_segments_of_aform x y b c ''\'')) X)" end definition "the_odo ode_fas safe_form = the(mk_ode_ops ode_fas safe_form)" locale ode_interpretation = fixes safe_form::form and safe_set::"'a::executable_euclidean_space set" and ode_fas::"floatarith list" and ode::"'a \ 'a" and finite::"'n::enum" assumes dims: "DIM('a) = CARD('n)" assumes len: "length ode_fas = CARD('n)" assumes safe_set_form: "safe_set = {x. interpret_form safe_form (list_of_eucl x)}" assumes interpret_fas: "\x. x \ safe_set \ einterpret ode_fas (list_of_eucl x) = ode x" assumes odo: "mk_ode_ops ode_fas safe_form \ None" assumes isFDERIV: "\xs. interpret_form safe_form xs \ isFDERIV (length ode_fas) [0.. the_odo ode_fas safe_form" lemmas odo_def = the_odo_def lemma odo_simps[simp]: "ode_expression odo = ode_fas" "safe_form_expr odo = safe_form" using odo by (auto simp: odo_def ode_expression_mk_ode_ops safe_form_expr_mk_ode_ops) lemma safe_set: "safe_set = aform.Csafe odo" using odo dims safe_set_form isFDERIV unfolding aform.Csafe_def aform.safe_def aform.safe_form_def aform.ode_e_def by (auto simp: mk_ode_ops_def safe_set_form len split: if_splits) lemma ode: "\x. x \ safe_set \ ode x = aform.ode odo x" by (auto simp: aform.ode_def aform.ode_e_def interpret_fas) sublocale auto_ll_on_open ode safe_set by (rule aform.auto_ll_on_open_congI[OF safe_set[symmetric] ode[symmetric]]) lemma ode_has_derivative_ode_d1: "(ode has_derivative blinfun_apply (aform.ode_d1 odo x)) (at x)" if "x \ safe_set" for x proof - from aform.fderiv[OF that[unfolded safe_set]] have "(aform.ode odo has_derivative blinfun_apply (aform.ode_d1 odo x)) (at x)" by simp moreover from topological_tendstoD[OF tendsto_ident_at open_domain(2) that] have "\\<^sub>F x' in at x. x' \ safe_set" . then have "\\<^sub>F x' in at x. aform.ode odo x' = ode x'" by eventually_elim (auto simp: ode) ultimately show ?thesis by (rule has_derivative_transform_eventually) (auto simp: ode that) qed sublocale c1_on_open_euclidean ode "aform.ode_d1 odo" safe_set apply unfold_locales subgoal by simp subgoal by (simp add: ode_has_derivative_ode_d1) subgoal by (rule aform.cont_fderiv') (auto intro!: continuous_intros simp: safe_set) done sublocale transfer_eucl_vec for a::'a and n::'n by unfold_locales (simp add: dims) lemma flow_eq: "t \ existence_ivl0 x \ aform.flow0 odo x t = flow0 x t" and Dflow_eq: "t \ existence_ivl0 x \ aform.Dflow odo x t = Dflow x t" and ex_ivl_eq: "t \ aform.existence_ivl0 odo x \ aform.existence_ivl0 odo x = existence_ivl0 x" and poincare_mapsto_eq: "closed a \ aform.poincare_mapsto odo a b c d e = poincare_mapsto a b c d e" and flowsto_eq: "aform.flowsto odo = flowsto" apply - subgoal by (rule flow0_cong[symmetric]) (auto simp: safe_set ode) subgoal by (rule Dflow_cong[symmetric]) (auto simp: safe_set ode) subgoal by (rule existence_ivl0_cong[symmetric]) (auto simp: safe_set ode) subgoal apply (subst aform.poincare_mapsto_cong[OF safe_set[symmetric]]) by (auto simp: ode) subgoal apply (intro ext) apply (subst flowsto_congI[OF safe_set ode]) by (auto simp: safe_set) done definition "avf \ \x::'n rvec. cast (aform.ode odo (cast x)::'a)::'n rvec" context includes lifting_syntax begin lemma aform_ode_transfer[transfer_rule]: "((=) ===> rel_ve ===> rel_ve) aform.ode aform.ode" unfolding aform.ode_def by transfer_prover lemma cast_aform_ode: "cast (aform.ode odo (cast (x::'n rvec))::'a) = aform.ode odo x" by transfer simp lemma aform_safe_transfer[transfer_rule]: "((=) ===> rel_ve ===> (=)) aform.safe aform.safe" unfolding aform.safe_def by transfer_prover lemma aform_Csafe_transfer[transfer_rule]: "((=) ===> rel_set rel_ve) aform.Csafe aform.Csafe" unfolding aform.Csafe_def by transfer_prover lemma cast_safe_set: "(cast ` safe_set::'n rvec set) = aform.Csafe odo" unfolding safe_set by transfer simp lemma aform_ode_d_raw_transfer[transfer_rule]: "((=) ===> (=) ===> rel_ve ===> rel_ve ===> rel_ve ===> rel_ve) aform.ode_d_raw aform.ode_d_raw" unfolding aform.ode_d_raw_def by transfer_prover lemma aform_ode_d_raw_aux_transfer: "((=) ===> rel_ve ===> rel_ve ===> rel_ve) (\x xb xa. if xb \ aform.Csafe x then aform.ode_d_raw x 0 xb 0 xa else 0) (\x xb xa. if xb \ aform.Csafe x then aform.ode_d_raw x 0 xb 0 xa else 0)" by transfer_prover lemma aform_ode_d1_transfer[transfer_rule]: "((=) ===> rel_ve ===> rel_blinfun rel_ve rel_ve) aform.ode_d1 aform.ode_d1" apply (auto simp: rel_blinfun_def aform.ode_d1_def intro!: rel_funI) unfolding aform.ode_d.rep_eq using aform_ode_d_raw_aux_transfer apply - apply (drule rel_funD, rule refl) apply (drule rel_funD, assumption) apply (drule rel_funD; assumption) done lemma cast_bl_transfer[transfer_rule]: "(rel_blinfun (=) (=) ===> rel_blinfun rel_ve rel_ve) id_blinfun cast_bl" by (auto simp: rel_ve_cast rel_blinfun_def intro!: rel_funI dest!: rel_funD) lemma cast_bl_transfer'[transfer_rule]: "(rel_blinfun rel_ve rel_ve ===> rel_blinfun (=) (=)) id_blinfun cast_bl" apply (auto simp: rel_ve_cast rel_blinfun_def cast_cast intro!: rel_funI dest!: rel_funD) by (subst cast_cast) auto lemma rel_blinfun_eq[relator_eq]: "rel_blinfun (=) (=) = (=)" by (auto simp: Rel_def rel_blinfun_def blinfun_ext rel_fun_eq intro!: rel_funI ext) lemma cast_aform_ode_D1: "cast_bl (aform.ode_d1 odo (cast (x::'n rvec))::'a\\<^sub>L'a) = (aform.ode_d1 odo x::'n rvec \\<^sub>L 'n rvec)" by transfer simp end definition "vf \ \x. cast (ode (cast x))" definition "vf' \ \x::'n rvec. cast_bl (aform.ode_d1 odo (cast x::'a)) ::'n rvec \\<^sub>L 'n rvec" definition "vX \ cast ` safe_set" sublocale a?: transfer_c1_on_open_euclidean a n ode "aform.ode_d1 odo" safe_set vf vf' vX for a::'a and n::'n by unfold_locales (simp_all add: dims vf_def vf'_def vX_def) sublocale av: transfer_c1_on_open_euclidean a n "aform.ode odo" "aform.ode_d1 odo" "(aform.Csafe odo)" avf vf' vX for a::'a and n::'n apply unfold_locales unfolding vX_def by (simp_all add: dims avf_def safe_set) lemma vflow_eq: "t \ v.existence_ivl0 x \ aform.flow0 odo x t = v.flow0 x t" thm flow_eq[of t "cast x"] flow_eq[of t "cast x", untransferred] apply (subst flow_eq[of t "cast x", untransferred, symmetric]) apply simp unfolding avf_def vX_def cast_aform_ode cast_safe_set .. lemma vf'_eq: "vf' = aform.ode_d1 odo" unfolding vf'_def cast_aform_ode_D1 .. lemma vDflow_eq: "t \ v.existence_ivl0 x \ aform.Dflow odo x t = v.Dflow x t" apply (subst Dflow_eq[of t "cast x", untransferred, symmetric]) apply simp unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq .. lemma vex_ivl_eq: "t \ aform.existence_ivl0 odo x \ aform.existence_ivl0 odo x = v.existence_ivl0 x" apply (subst ex_ivl_eq[of t "cast x", untransferred, symmetric]) unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto context includes lifting_syntax begin lemma id_cast_eucl1_transfer_eq: "(\x. x) = (\x. (fst x, 1\<^sub>L o\<^sub>L snd x o\<^sub>L 1\<^sub>L))" by auto lemma cast_eucl1_transfer[transfer_rule]: "(rel_prod (=) (rel_blinfun (=) (=)) ===> rel_prod rel_ve (rel_blinfun rel_ve rel_ve)) (\x. x) cast_eucl1" unfolding cast_eucl1_def id_cast_eucl1_transfer_eq apply transfer_prover_start apply (transfer_step) apply (transfer_step) apply (transfer_step) apply (transfer_step) apply (transfer_step) apply simp done end lemma avpoincare_mapsto_eq: "aform.poincare_mapsto odo a (b::'n eucl1 set) c d e = av.v.poincare_mapsto a b c d e" if "closed a" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto lemma vpoincare_mapsto_eq: "aform.poincare_mapsto odo a (b::'n eucl1 set) c d e = v.poincare_mapsto a b c d e" if "closed a" proof - have "closed (cast ` a::'a set)" using that by transfer auto from poincare_mapsto_eq[of "cast ` a::'a set" "cast_eucl1 ` b::('a \ 'a \\<^sub>L 'a) set" "cast ` c::'a set" "cast ` d::'a set" "cast_eucl1 ` e::('a \ 'a \\<^sub>L 'a) set", OF this, untransferred] have "v.poincare_mapsto a b c d e = av.v.poincare_mapsto a b c d e" by auto also have "\ = aform.poincare_mapsto odo a (b::'n eucl1 set) c d e" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto finally show ?thesis by simp qed lemma avflowsto_eq: "aform.flowsto odo = (av.v.flowsto::'n eucl1 set \ _)" proof (intro ext, goal_cases) case (1 a b c d) have "av.v.flowsto a b c d = aform.flowsto odo a b c d" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto then show ?case by simp qed lemma vflowsto_eq: "aform.flowsto odo = (v.flowsto::'n eucl1 set \ _)" proof (intro ext, goal_cases) case (1 a b c d) have "aform.flowsto odo (cast_eucl1 ` a::'a c1_info set) b (cast_eucl1 ` c) (cast_eucl1 ` d) = flowsto (cast_eucl1 ` a::'a c1_info set) b (cast_eucl1 ` c) (cast_eucl1 ` d)" by (subst flowsto_eq) auto from this[untransferred] have "v.flowsto a b c d = av.v.flowsto a b c d" by auto also have "\ = aform.flowsto odo a b c d" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto finally show ?case by simp qed context includes lifting_syntax begin lemma flow1_of_list_transfer[transfer_rule]: "(list_all2 (=) ===> rel_prod rel_ve (rel_blinfun rel_ve rel_ve)) flow1_of_list flow1_of_list" unfolding flow1_of_list_def blinfun_of_list_def o_def flow1_of_vec1_def by transfer_prover lemma c1_info_of_appr_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (rel_option (list_all2 (=))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appr aform.c1_info_of_appr" unfolding aform.c1_info_of_appr_def by transfer_prover lemma c0_info_of_appr_transfer[transfer_rule]: "((list_all2 (=)) ===> rel_set rel_ve) aform.c0_info_of_appr aform.c0_info_of_appr" unfolding aform.c0_info_of_appr_def by transfer_prover lemma aform_scaleR2_transfer[transfer_rule]: "((=) ===> (=) ===> rel_set (rel_prod A B) ===> rel_set (rel_prod A B)) scaleR2 scaleR2" if [unfolded Rel_def, transfer_rule]: "((=) ===> B ===> B) (*\<^sub>R) (*\<^sub>R)" unfolding scaleR2_def by transfer_prover lemma scaleR_rel_blinfun_transfer[transfer_rule]: "((=) ===> rel_blinfun rel_ve rel_ve ===> rel_blinfun rel_ve rel_ve) (*\<^sub>R) (*\<^sub>R)" apply (auto intro!: rel_funI simp: rel_blinfun_def blinfun.bilinear_simps) apply (drule rel_funD) apply assumption apply (rule scaleR_transfer[THEN rel_funD, THEN rel_funD]) apply auto done lemma c1_info_of_appre_transfer[transfer_rule]: "(rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=)))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appre aform.c1_info_of_appre" unfolding aform.c1_info_of_appre_def by transfer_prover lemma c1_info_of_apprs_transfer[transfer_rule]: "((=) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_apprs aform.c1_info_of_apprs" unfolding aform.c1_info_of_apprs_def by transfer_prover lemma c1_info_of_appr'_transfer[transfer_rule]: "(rel_option (list_all2 (=)) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appr' aform.c1_info_of_appr'" unfolding aform.c1_info_of_appr'_def by transfer_prover lemma c0_info_of_apprs_transfer[transfer_rule]: "((=) ===> rel_set rel_ve) aform.c0_info_of_apprs aform.c0_info_of_apprs" unfolding aform.c0_info_of_apprs_def by transfer_prover lemma c0_info_of_appr'_transfer[transfer_rule]: "(rel_option (list_all2 (=)) ===> rel_set rel_ve) aform.c0_info_of_appr' aform.c0_info_of_appr'" unfolding aform.c0_info_of_appr'_def by transfer_prover lemma aform_Csafe_vX[simp]: "aform.Csafe odo = (vX::'n rvec set)" by (simp add: vX_def cast_safe_set) definition blinfuns_of_lvivl::"real list \ real list \ ('b \\<^sub>L 'b::executable_euclidean_space) set" where "blinfuns_of_lvivl x = blinfun_of_list ` list_interval (fst x) (snd x)" lemma blinfun_of_list_transfer[transfer_rule]: "(list_all2 (=) ===> rel_blinfun rel_ve rel_ve) blinfun_of_list blinfun_of_list" unfolding blinfun_of_list_def by transfer_prover lemma blinfuns_of_lvivl_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (list_all2 (=)) ===> rel_set (rel_blinfun rel_ve rel_ve)) blinfuns_of_lvivl blinfuns_of_lvivl" unfolding blinfuns_of_lvivl_def by transfer_prover definition "blinfuns_of_lvivl' x = (case x of None \ UNIV | Some x \ blinfuns_of_lvivl x)" lemma blinfuns_of_lvivl'_transfer[transfer_rule]: "(rel_option (rel_prod (list_all2 (=)) (list_all2 (=))) ===> rel_set (rel_blinfun rel_ve rel_ve)) blinfuns_of_lvivl' blinfuns_of_lvivl'" unfolding blinfuns_of_lvivl'_def by transfer_prover lemma atLeastAtMost_transfer[transfer_rule]: "(A ===> A ===> rel_set A) atLeastAtMost atLeastAtMost" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" "bi_unique A" unfolding atLeastAtMost_def atLeast_def atMost_def by transfer_prover lemma set_of_ivl_transfer[transfer_rule]: "(rel_prod A A ===> rel_set A) set_of_ivl set_of_ivl" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" "bi_unique A" unfolding set_of_ivl_def by transfer_prover lemma set_of_lvivl_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (list_all2 (=)) ===> rel_set rel_ve) set_of_lvivl set_of_lvivl" unfolding set_of_lvivl_def by transfer_prover lemma set_of_lvivl_eq: "set_of_lvivl I = (eucl_of_list ` list_interval (fst I) (snd I)::'b::executable_euclidean_space set)" if [simp]: "length (fst I) = DIM('b)" "length (snd I) = DIM('b)" proof (auto simp: set_of_lvivl_def list_interval_def set_of_ivl_def, goal_cases) case (1 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "fst I" "list_of_eucl x"] lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "list_of_eucl x" "snd I"] show ?case by force next case (2 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "fst I" "x"] show ?case by (auto simp: list_all2_lengthD) next case (3 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "x" "snd I"] show ?case by (auto simp: list_all2_lengthD) qed lemma bounded_linear_matrix_vector_mul[THEN bounded_linear_compose, bounded_linear_intros]: "bounded_linear ((*v) x)" for x::"real^'x^'y" unfolding linear_linear by (rule matrix_vector_mul_linear) lemma blinfun_of_list_eq: "blinfun_of_list x = blinfun_of_vmatrix (eucl_of_list x::((real, 'b) vec, 'b) vec)" if "length x = CARD('b::enum)*CARD('b)" unfolding blinfun_of_list_def apply (transfer fixing: x) apply (rule linear_eq_stdbasis) unfolding linear_conv_bounded_linear apply (auto intro!: bounded_linear_intros) proof goal_cases case (1 b) have "(eucl_of_list x::((real, 'b) vec, 'b) vec) *v b = (eucl_of_list x::((real, 'b) vec, 'b) vec) *v eucl_of_list (list_of_eucl b)" by simp also have "\ = (\ij Basis_list ! j)) *\<^sub>R Basis_list ! i)" by (subst eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list) (auto simp: that) also have "\ = (\i\Basis. \j\Basis. (b \ j * x ! (index Basis_list i * CARD('b) + index Basis_list j)) *\<^sub>R i)" apply (subst sum_list_Basis_list[symmetric])+ apply (subst sum_list_sum_nth)+ by (auto simp add: atLeast0LessThan scaleR_sum_left intro!: sum.cong) finally show ?case by simp qed lemma blinfuns_of_lvivl_eq: "blinfuns_of_lvivl x = (blinfun_of_vmatrix ` set_of_lvivl x::((real, 'b) vec \\<^sub>L (real, 'b) vec) set)" if "length (fst x) = CARD('b::enum)*CARD('b)" "length (snd x) = CARD('b)*CARD('b)" apply (subst set_of_lvivl_eq) subgoal by (simp add: that) subgoal by (simp add: that) unfolding blinfuns_of_lvivl_def image_image by (auto simp: that blinfun_of_list_eq[symmetric] in_list_interval_lengthD cong: image_cong) lemma range_blinfun_of_vmatrix[simp]: "range blinfun_of_vmatrix = UNIV" apply auto apply transfer subgoal for x by (rule image_eqI[where x="matrix x"]) auto done lemma blinfun_of_vmatrix_image: "blinfun_of_vmatrix ` aform.set_of_lvivl' x = (blinfuns_of_lvivl' x::((real, 'b) vec \\<^sub>L (real, 'b) vec) set)" if "aform.lvivl'_invar (CARD('b)*CARD('b::enum)) x" using that by (auto simp: aform.set_of_lvivl'_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_eq aform.lvivl'_invar_def split: option.splits) lemma one_step_result123: "solves_one_step_until_time_aform optns odo X0i t1 t2 E dE \ (x0, d0) \ aform.c1_info_of_appre X0i \ t \ {t1 .. t2} \ set_of_lvivl E \ S \ blinfuns_of_lvivl' dE \ dS \ length (fst E) = CARD('n) \ length (snd E) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dE \ aform.c1_info_invare DIM('a) X0i \ aform.D odo = DIM('a) \ (t \ existence_ivl0 (x0::'a) \ flow0 x0 t \ S) \ Dflow x0 t o\<^sub>L d0 \ dS" apply (transfer fixing: optns X0i t1 t2 t E dE) subgoal premises prems for x0 d0 S dS proof - have "t \ aform.existence_ivl0 odo x0 \ aform.flow0 odo x0 t \ S \ aform.Dflow odo x0 t o\<^sub>L d0 \ dS" apply (rule one_step_in_ivl[of t t1 t2 x0 d0 X0i "fst E" "snd E" S dE dS odo optns]) using prems by (auto simp: eucl_of_list_prod set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image aform.D_def solves_one_step_until_time_aform_def) with vflow_eq[of t x0] vDflow_eq[of t x0] vex_ivl_eq[symmetric, of t x0] show ?thesis by simp qed done lemmas one_step_result12 = one_step_result123[THEN conjunct1] and one_step_result3 = one_step_result123[THEN conjunct2] lemmas one_step_result1 = one_step_result12[THEN conjunct1] and one_step_result2 = one_step_result12[THEN conjunct2] lemma plane_of_transfer[transfer_rule]: "(rel_sctn A ===> rel_set A) plane_of plane_of" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding plane_of_def by transfer_prover lemma below_halfspace_transfer[transfer_rule]: "(rel_sctn A ===> rel_set A) below_halfspace below_halfspace" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding below_halfspace_def le_halfspace_def by transfer_prover definition "rel_nres A a b \ (a, b) \ \{(a, b). A a b}\nres_rel" lemma FAILi_transfer[transfer_rule]: "(rel_nres B) FAILi FAILi" by (auto simp: rel_nres_def nres_rel_def) lemma RES_transfer[transfer_rule]: "(rel_set B ===> rel_nres B) RES RES" by (auto simp: rel_nres_def nres_rel_def rel_set_def intro!: rel_funI RES_refine) context includes autoref_syntax begin lemma RETURN_dres_nres_relI: "(fi, f) \ A \ B \ (\x. dRETURN (fi x), (\x. RETURN (f x))) \ A \ \B\dres_nres_rel" by (auto simp: dres_nres_rel_def dest: fun_relD) end lemma br_transfer[transfer_rule]: "((B ===> C) ===> (B ===> (=)) ===> rel_set (rel_prod B C)) br br" if [transfer_rule]: "bi_total B" "bi_unique C" "bi_total C" unfolding br_def by transfer_prover lemma aform_appr_rel_transfer[transfer_rule]: "(rel_set (rel_prod (list_all2 (=)) (rel_set rel_ve))) aform.appr_rel aform.appr_rel" unfolding aform.appr_rel_br by (transfer_prover) lemma appr1_rel_transfer[transfer_rule]: "(rel_set (rel_prod (rel_prod (list_all2 (=)) (rel_option (list_all2 (=)))) (rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))))) aform.appr1_rel aform.appr1_rel" unfolding aform.appr1_rel_internal by transfer_prover lemma relAPP_transfer[transfer_rule]: "((rel_set (rel_prod B C) ===> D) ===> rel_set (rel_prod B C) ===> D) Relators.relAPP Relators.relAPP" unfolding relAPP_def by transfer_prover lemma prod_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod D E) ===> rel_set (rel_prod (rel_prod B D) (rel_prod C E))) prod_rel prod_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" "bi_unique E" "bi_total E" unfolding prod_rel_def_internal by transfer_prover lemma Domain_transfer[transfer_rule]: "(rel_set (rel_prod A B) ===> rel_set A) Domain Domain" if [transfer_rule]: "bi_total A" "bi_unique A" "bi_total B" "bi_unique B" unfolding Domain_unfold by transfer_prover lemma set_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod (rel_set B) (rel_set C))) set_rel set_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" unfolding set_rel_def_internal by transfer_prover lemma relcomp_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod C D) ===> rel_set (rel_prod B D)) relcomp relcomp" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" unfolding relcomp_unfold by transfer_prover lemma Union_rel_transfer[transfer_rule]: "(rel_set (rel_prod B (rel_set C)) ===> rel_set (rel_prod C (rel_set D)) ===> rel_set (rel_prod B (rel_set D))) Union_rel Union_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" unfolding Union_rel_internal top_fun_def top_bool_def by transfer_prover lemma fun_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod D E) ===> rel_set (rel_prod (B ===> D) (C ===> E))) Relators.fun_rel Relators.fun_rel" if [transfer_rule]: "bi_unique B" "bi_unique C" "bi_unique D" "bi_total D" "bi_unique E" "bi_total E" unfolding fun_rel_def_internal by transfer_prover lemma c1_info_of_apprse_transfer[transfer_rule]: "(list_all2 (rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=))))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_apprse aform.c1_info_of_apprse" unfolding aform.c1_info_of_apprse_def by transfer_prover term scaleR2_rel (* "scaleR2_rel" :: "('b \ ('c \ 'd) set) set \ (((ereal \ ereal) \ 'b) \ ('c \ 'd) set) set" *) lemma scaleR2_rel_transfer[transfer_rule]: "(rel_set (rel_prod (=) (rel_set (rel_prod (=) (=)))) ===> rel_set (rel_prod (rel_prod (rel_prod (=) (=)) (=)) (rel_set (rel_prod (=) (=))))) scaleR2_rel scaleR2_rel" unfolding scaleR2_rel_internal by transfer_prover lemma appr1_rele_transfer[transfer_rule]: "(rel_set (rel_prod (rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=))))) (rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))))) aform.appr1e_rel aform.appr1e_rel" unfolding scaleR2_rel_internal by transfer_prover lemma flow1_of_vec1_times: "flow1_of_vec1 ` (A \ B) = A \ blinfun_of_vmatrix ` B" by (auto simp: flow1_of_vec1_def vec1_of_flow1_def) lemma stable_on_transfer[transfer_rule]: "(rel_set rel_ve ===> rel_set rel_ve ===> (=)) v.stable_on stable_on" unfolding stable_on_def v.stable_on_def by transfer_prover theorem solves_poincare_map_aform: "solves_poincare_map_aform optns odo (\x. dRETURN (symstart x)) [S] guards ivl sctn roi XS RET dRET \ (symstart, symstarta) \ fun_rel (aform.appr1e_rel) (clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel) \ (\X0. (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) X) (symstarta X0)) \ stable_on (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) trap \ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ length (normal S) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS - trap \ UNIV) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns symstart S guards ivl sctn roi XS RET dRET) subgoal premises prems for symstarta trap proof - have "aform.poincare_mapsto odo (set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS - trap \ UNIV) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (flow1_of_vec1 ` ({eucl_of_list (fst RET)..eucl_of_list (snd RET)} \ aform.set_of_lvivl' dRET))" apply (rule solves_poincare_map[OF _ RETURN_dres_nres_relI RETURN_rule, of optns odo symstart S guards ivl sctn roi XS "fst RET" "snd RET" dRET symstarta trap]) subgoal using prems(1) by (simp add: solves_poincare_map_aform_def) subgoal using prems(2) by (auto simp: fun_rel_def_internal) subgoal for X0 using prems(3)[of X0] vflowsto_eq by auto subgoal unfolding aform.stable_on_def proof (safe, goal_cases) case (1 t x0) from 1 have a: "t \ v.existence_ivl0 x0" using vex_ivl_eq by blast with 1 have b: "v.flow0 x0 t \ trap" using vflow_eq by simp have c: "v.flow0 x0 s \ vX - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)" if s: "s \ {0<..t}" for s using 1(4)[rule_format, OF s] apply (subst (asm) vflow_eq) unfolding aform_Csafe_vX[symmetric] using s a by (auto dest!: a.v.ivl_subset_existence_ivl) from a b c show ?case using prems(4)[unfolded v.stable_on_def, rule_format, OF b a 1(3) c] by simp qed subgoal using prems by auto subgoal using prems by (auto simp: aform.D_def) subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto done then show ?thesis using vflow_eq vex_ivl_eq vflowsto_eq prems apply (subst vpoincare_mapsto_eq[symmetric]) by (auto simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times) qed done theorem solves_poincare_map_aform': "solves_poincare_map_aform' optns odo S guards ivl sctn roi XS RET dRET\ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ length (normal S) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns S guards ivl sctn roi XS RET dRET) subgoal using solves_poincare_map'[of optns odo S guards ivl sctn roi XS "fst RET" "snd RET" dRET] using vflow_eq vex_ivl_eq vflowsto_eq apply (subst vpoincare_mapsto_eq[symmetric]) by (auto intro!: closed_Int simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times aform.D_def solves_poincare_map_aform'_def) done theorem solves_poincare_map_onto_aform: "solves_poincare_map_onto_aform optns odo guards ivl sctn roi XS RET dRET\ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS) UNIV (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns guards ivl sctn roi XS RET dRET) subgoal using solves_poincare_map_onto[of optns odo guards ivl sctn roi XS "fst RET" "snd RET" dRET, where 'n='n, unfolded aform.poincare_maps_onto_def] using vflow_eq vex_ivl_eq vflowsto_eq apply (subst vpoincare_mapsto_eq[symmetric]) by (auto intro!: closed_Int simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times aform.D_def solves_poincare_map_onto_aform_def) done end end subsection \Example Utilities!\ hide_const floatarith.Max floatarith.Min lemma degree_sum_pdevs_scaleR_Basis: "degree (sum_pdevs (\i. pdevs_scaleR (a i) i) (Basis::'b::euclidean_space set)) = Max ((\i. degree (a i)) ` Basis)" apply (rule antisym) subgoal apply (rule degree_le) by (auto ) subgoal apply (rule Max.boundedI) apply simp apply simp apply (auto simp: intro!: degree_leI) by (auto simp: euclidean_eq_iff[where 'a='b]) done lemma Inf_aform_eucl_of_list_aform: assumes "length a = DIM('b::executable_euclidean_space)" shows "Inf_aform (eucl_of_list_aform a::'b aform) = eucl_of_list (map Inf_aform a)" using assms apply (auto simp: eucl_of_list_aform_def Inf_aform_def[abs_def] algebra_simps eucl_of_list_inner inner_sum_left intro!: euclidean_eqI[where 'a='b]) apply (auto simp: tdev_def inner_sum_left abs_inner inner_Basis if_distrib cong: if_cong) apply (rule sum.mono_neutral_cong_left) apply simp by (auto simp: degree_sum_pdevs_scaleR_Basis) lemma Sup_aform_eucl_of_list_aform: assumes "length a = DIM('b::executable_euclidean_space)" shows "Sup_aform (eucl_of_list_aform a::'b aform) = eucl_of_list (map Sup_aform a)" using assms apply (auto simp: eucl_of_list_aform_def Sup_aform_def[abs_def] algebra_simps eucl_of_list_inner inner_sum_left intro!: euclidean_eqI[where 'a='b]) apply (auto simp: tdev_def inner_sum_left abs_inner inner_Basis if_distrib cong: if_cong) apply (rule sum.mono_neutral_cong_right) apply simp by (auto simp: degree_sum_pdevs_scaleR_Basis) lemma eucl_of_list_map_Inf_aform_leI: assumes "x \ Affine (eucl_of_list_aform a::'b::executable_euclidean_space aform)" assumes "length a = DIM('b)" shows "eucl_of_list (map Inf_aform a) \ x" using Inf_aform_le_Affine[OF assms(1)] assms(2) by (auto simp: Inf_aform_eucl_of_list_aform) lemma eucl_of_list_map_Sup_aform_geI: assumes "x \ Affine (eucl_of_list_aform a::'b::executable_euclidean_space aform)" assumes "length a = DIM('b)" shows "x \ eucl_of_list (map Sup_aform a)" using Sup_aform_ge_Affine[OF assms(1)] assms(2) by (auto simp: Sup_aform_eucl_of_list_aform) lemma mem_Joints_appendE: assumes "x \ Joints (xs @ ys)" obtains x1 x2 where "x = x1 @ x2" "x1 \ Joints xs" "x2 \ Joints ys" using assms by (auto simp: Joints_def valuate_def) lemma c1_info_of_appr_subsetI1: fixes X1::"'b::executable_euclidean_space set" assumes subset: "{eucl_of_list (map Inf_aform (fst R)) .. eucl_of_list (map Sup_aform (fst R))} \ X1" assumes len: "length (fst R) = DIM('b)" shows "aform.c1_info_of_appr R \ X1 \ UNIV" using len apply (auto simp: aform.c1_info_of_appr_def flow1_of_list_def split: option.splits intro!: subsetD[OF subset] elim!: mem_Joints_appendE) subgoal by (auto intro!: eucl_of_list_mem_eucl_of_list_aform eucl_of_list_map_Inf_aform_leI eucl_of_list_map_Sup_aform_geI) subgoal by (auto intro!: eucl_of_list_mem_eucl_of_list_aform eucl_of_list_map_Inf_aform_leI eucl_of_list_map_Sup_aform_geI) subgoal apply (subst (2) eucl_of_list_take_DIM[symmetric, OF refl]) apply (auto simp: min_def) apply (simp add: Joints_imp_length_eq eucl_of_list_map_Inf_aform_leI eucl_of_list_mem_eucl_of_list_aform) apply (simp add: Joints_imp_length_eq eucl_of_list_map_Inf_aform_leI eucl_of_list_mem_eucl_of_list_aform) done subgoal apply (subst (2) eucl_of_list_take_DIM[symmetric, OF refl]) apply (auto simp: min_def) by (simp add: Joints_imp_length_eq eucl_of_list_map_Sup_aform_geI eucl_of_list_mem_eucl_of_list_aform) done lemmas [simp] = compute_tdev syntax product_aforms::"(real aform) list \ (real aform) list \ (real aform) list" (infixr "\\<^sub>a" 70) lemma matrix_inner_Basis_list: includes vec_syntax assumes "k < CARD('n) * CARD('m)" shows "(f::(('n::enum rvec, 'm::enum) vec)) \ Basis_list ! k = vec_nth (vec_nth f (enum_class.enum ! (k div CARD('n)))) (enum_class.enum ! (k mod CARD('n)))" proof - have "f \ Basis_list ! k = (\x\UNIV. \xa\UNIV. if enum_class.enum ! (k mod CARD('n)) = xa \ enum_class.enum ! (k div CARD('n)) = x then f $ x $ xa else 0)" using assms unfolding inner_vec_def apply (auto simp: Basis_list_vec_def concat_map_map_index) apply (subst (2) sum.cong[OF refl]) apply (subst sum.cong[OF refl]) apply (subst (2) vec_nth_Basis2) apply (force simp add: Basis_vec_def Basis_list_real_def) apply (rule refl) apply (rule refl) apply (auto simp: if_distribR if_distrib axis_eq_axis Basis_list_real_def cong: if_cong) done also have "\ = f $ enum_class.enum ! (k div CARD('n)) $ enum_class.enum ! (k mod CARD('n))" apply (subst if_conn) apply (subst sum.delta') apply simp by (simp add: sum.delta') finally show ?thesis by simp qed lemma list_of_eucl_matrix: includes vec_syntax shows "(list_of_eucl (M::(('n::enum rvec, 'm::enum) vec))) = concat (map (\i. map (\j. M $ (enum_class.enum ! i)$ (enum_class.enum ! j) ) [0.. real_of_float (lapprox_rat 20 i j))" definition "udec x = (case quotient_of x of (i, j) \ real_of_float (rapprox_rat 20 i j))" lemma ldec: "ldec x \ real_of_rat x" and udec: "real_of_rat x \ udec x" apply (auto simp: ldec_def udec_def split: prod.splits intro!: lapprox_rat[le] rapprox_rat[ge]) apply (metis Fract_of_int_quotient less_eq_real_def less_int_code(1) of_rat_rat quotient_of_denom_pos quotient_of_div) apply (metis Fract_of_int_quotient less_eq_real_def less_int_code(1) of_rat_rat quotient_of_denom_pos quotient_of_div) done context includes floatarith_notation begin definition "matrix_of_degrees\<^sub>e = (let ur = Rad_of (Var 0); vr = Rad_of (Var 1) in [Cos ur, Cos vr, 0, Sin ur, Sin vr, 0, 0, 0, 0])" definition "matrix_of_degrees u v = approx_floatariths 30 matrix_of_degrees\<^sub>e (aforms_of_point ([u, v, 0]))" lemma interpret_floatariths_matrix_of_degrees: "interpret_floatariths matrix_of_degrees\<^sub>e (([u::real, v::real, 0])) = [cos (rad_of u), cos (rad_of v), 0, sin (rad_of u), sin (rad_of v), 0, 0, 0, 0]" by (auto simp: matrix_of_degrees\<^sub>e_def Let_def inverse_eq_divide) definition "num_options p sstep m N a projs print_fun = \ precision = p, adaptive_atol = FloatR 1 (- a), adaptive_rtol = FloatR 1 (- a), method_id = 2, start_stepsize = FloatR 1 (- sstep), iterations = 40, halve_stepsizes = 40, widening_mod = 10, rk2_param = FloatR 1 0, default_reduce = correct_girard (p) (m) (N), printing_fun = (\a b. let _ = fold (\(x, y, c) _. print_fun (String.implode (shows_segments_of_aform (x) (y) b c ''\''))) projs (); _ = print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\'')) in () ), tracing_fun = (\a b. let _ = print_fun (String.implode (''# '' @ a @ ''\'')) in case b of Some b \ (let _ = () in print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\''))) | None \ ()) \" definition "num_options_c1 p sstep m N a projs dcolors print_fun = (let no = num_options p sstep m N a (map (\(x, y, c, ds). (x, y, c)) projs) print_fun; D = length dcolors in no \printing_fun:= (\a b. let _ = printing_fun no a b in if a then () else fold (\(x, y, c, ds) _. print_fun (String.implode (shows_aforms_vareq D [(x, y)] ds dcolors dcolors (FloatR 1 (-1)) ''# no C1 info'' b ''''))) projs () ) \)" definition "num_options_code p sstep m N a projs print_fun = num_options (nat_of_integer p) (int_of_integer sstep) (nat_of_integer m) (nat_of_integer N) (int_of_integer a) (map (\(i, j, k). (nat_of_integer i, nat_of_integer j, k)) projs) print_fun" definition "ro s n M g0 g1 inter_step = \max_tdev_thres = FloatR 1 s, pre_split_reduce = correct_girard 30 n M, pre_inter_granularity = FloatR 1 g0, post_inter_granularity = (FloatR 1 g1), pre_collect_granularity = FloatR 1 g0, max_intersection_step = FloatR 1 inter_step\" definition "code_ro s n m g0 g1 inter_step = ro (int_of_integer s) (nat_of_integer n) (nat_of_integer m) (int_of_integer g0) (int_of_integer g1) (int_of_integer inter_step)" fun xsec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec x (y0, y1) (z0, z1) = (([x, y0, z0], [x, y1, z1]), Sctn [1,0,0] x)" fun xsec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec' x (y0, y1) (z0, z1) = (([x, y0, z0], [x, y1, z1]), Sctn [-1,0,0] (-x))" fun ysec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec (x0, x1) y (z0, z1) = (([x0, y, z0], [x1, y, z1]), Sctn [0, 1,0] y)" fun ysec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec' (x0, x1) y (z0, z1) = (([x0, y, z0], [x1, y, z1]), Sctn [0, -1,0] (-y))" fun zsec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "zsec (x0, x1) (y0, y1) z = (([x0, y0, z], [x1, y1, z]), Sctn [0, 0, 1] z)" fun zsec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "zsec' (x0, x1) (y0, y1) z = (([x0, y0, z], [x1, y1, z]), Sctn [0, 0, -1] (-z))" fun xsec2:: "real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec2 x (y0, y1) = (([x, y0], [x, y1]), Sctn [1,0] x)" fun xsec2':: "real \ real \ real \(real list \ real list) \ real list sctn" where "xsec2' x (y0, y1) = (([x, y0], [x, y1]), Sctn [-1,0] (-x))" fun ysec2:: "real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec2 (x0, x1) y = (([x0, y], [x1, y]), Sctn [0, 1] y)" fun ysec2':: "real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec2' (x0, x1) y = (([x0, y], [x1, y]), Sctn [0, -1] (-y))" fun ysec4:: "real \ real \ real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec4 (x0, x1) y (z0, z1) (m0, m1) = (([x0, y, z0, m0], [x1, y, z1, m1]), Sctn [0, 1,0, 0] (y))" fun ysec4':: "real \ real \ real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec4' (x0, x1) y (z0, z1) (m0, m1) = (([x0, y, z0, m0], [x1, y, z1, m1]), Sctn [0, -1,0, 0] (-y))" definition "code_sctn N n c = Sctn ((replicate (nat_of_integer N) (0::real))[nat_of_integer n := 1]) c" definition "code_sctn' N n c = Sctn ((replicate (nat_of_integer N) (0::real))[nat_of_integer n := -1]) (-c)" definition "lrat i j = real_of_float (lapprox_rat 30 (int_of_integer i) (int_of_integer j))" definition "urat i j = real_of_float (lapprox_rat 30 (int_of_integer i) (int_of_integer j))" definition [simp]: "TAG_optns (optns::string \ ((String.literal \ unit) \ (real aform) numeric_options)) = True" lemma TAG_optns: "P \ (TAG_optns optns \ P)" by (auto simp: ) definition [simp]: "TAG_reach_optns (roi::real aform reach_options) = True" lemma TAG_reach_optns: "P \ (TAG_reach_optns optns \ P)" by (auto simp: ) definition [simp]: "TAG_sctn (b::bool) = True" lemma TAG_sctn: "P \ (TAG_sctn optns \ P)" by (auto simp: ) subsection \Integrals and Computation\ lemma has_vderiv_on_PairD: assumes "((\t. (f t, g t)) has_vderiv_on fg') T" shows "(f has_vderiv_on (\t. fst (fg' t))) T" "(g has_vderiv_on (\t. snd (fg' t))) T" proof - from assms have "((\x. (f x, g x)) has_derivative (\xa. xa *\<^sub>R fg' t)) (at t within T)" if "t \ T" for t by (auto simp: has_vderiv_on_def has_vector_derivative_def that intro!: derivative_eq_intros) from diff_chain_within[OF this has_derivative_fst[OF has_derivative_ident]] diff_chain_within[OF this has_derivative_snd[OF has_derivative_ident]] show "(f has_vderiv_on (\t. fst (fg' t))) T" "(g has_vderiv_on (\t. snd (fg' t))) T" by (auto simp: has_vderiv_on_def has_vector_derivative_def o_def) qed lemma solves_autonomous_odeI: assumes "((\t. (t, phi t)) solves_ode (\t x. (1, f (fst x) (snd x)))) S (T \ X)" shows "(phi solves_ode f) S X" proof (rule solves_odeI) from solves_odeD[OF assms] have "((\t. (t, phi t)) has_vderiv_on (\t. (1, f (fst (t, phi t)) (snd (t, phi t))))) S" "\t. t \ S \ (phi t) \ X" by (auto simp: ) from has_vderiv_on_PairD(2)[OF this(1)] this(2) show "(phi has_vderiv_on (\t. f t (phi t))) S" "\t. t \ S \ phi t \ X" by auto qed lemma integral_solves_autonomous_odeI: fixes f::"real \ 'a::executable_euclidean_space" assumes "(phi solves_ode (\t _. f t)) {a .. b} X" "phi a = 0" assumes "a \ b" shows "(f has_integral phi b) {a .. b}" proof - have "(f has_integral phi b - phi a) {a..b}" apply (rule fundamental_theorem_of_calculus[of a b phi f]) unfolding has_vderiv_on_def[symmetric] apply fact using solves_odeD[OF assms(1)] by (simp add: has_vderiv_on_def) then show ?thesis by (simp add: assms) qed lemma zero_eq_eucl_of_list_rep_DIM: "(0::'a::executable_euclidean_space) = eucl_of_list (replicate (DIM('a)) 0)" by (auto intro!: euclidean_eqI[where 'a='a] simp: eucl_of_list_inner) lemma zero_eq_eucl_of_list_rep: "(0::'a::executable_euclidean_space) = eucl_of_list (replicate D 0)" if "D \ DIM('a)" proof - from that have "replicate D (0::real) = replicate (DIM('a)) 0 @ replicate (D - DIM('a)) 0" by (auto simp: replicate_add[symmetric]) also have "eucl_of_list \ = (eucl_of_list (replicate DIM('a) 0)::'a)" by (rule eucl_of_list_append_zeroes) also have "\ = 0" by (rule zero_eq_eucl_of_list_rep_DIM[symmetric]) finally show ?thesis by simp qed lemma one_has_ivl_integral: "((\x. 1::real) has_ivl_integral b - a) a b" using has_integral_const_real[of "1::real" a b] has_integral_const_real[of "1::real" b a] by (auto simp: has_ivl_integral_def split: if_splits) lemma Joints_aforms_of_point_self[simp]: "xs \ Joints (aforms_of_point xs)" by (simp add: aforms_of_point_def) lemma bind_eq_dRETURN_conv: "(f \ g = dRETURN S) \ (\R. f = dRETURN R \ g R = dRETURN S)" by (cases f) auto end lemma list_of_eucl_memI: "list_of_eucl (x::'x::executable_euclidean_space) \ S" if "x \ eucl_of_list ` S" "\x. x \ S \ length x = DIM('x)" using that by auto lemma Joints_aforms_of_ivls_append_point: "Joints (xs @ aforms_of_ivls p p) = (\x. x @ p) ` Joints xs" using aform.set_of_appr_of_ivl_append_point[unfolded aform_ops_def approximate_set_ops.simps] . context ode_interpretation begin theorem solves_one_step_ivl: assumes T: "T \ {t1 .. t2}" assumes X: "X \ {eucl_of_list lx .. eucl_of_list ux}" "length lx = DIM('a)" "length ux = DIM('a)" assumes S: "{eucl_of_list ls::'a .. eucl_of_list us} \ S" assumes lens: "length ls = DIM('a)" "length us = DIM('a)" \ \TODO: this could be verified\ assumes [simp]: "aform.D odo = DIM('a)" assumes r: "solves_one_step_until_time_aform optns odo ((1,1), aforms_of_ivls lx ux, None) t1 t2 (ls, us) None" shows "t \ T \ x0 \ X \ t \ existence_ivl0 x0 \ flow0 x0 t \ S" proof (intro impI) assume t: "t \ T" and x0: "x0 \ X" from S have S: "set_of_lvivl (ls, us) \ S" by (auto simp: set_of_lvivl_def set_of_ivl_def) have lens: "length (fst (ls, us)) = CARD('n)" "length (snd (ls, us)) = CARD('n)" by (auto simp: lens) have x0: "list_of_eucl x0 \ Joints (aforms_of_ivls lx ux)" apply (rule aforms_of_ivls) subgoal by (simp add: X) subgoal by (simp add: X) using subsetD[OF X(1) x0] apply (auto simp: eucl_le[where 'a='a] X) apply (metis assms(3) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) apply (metis assms(4) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) done from t T have t: "t \ {t1 .. t2}" by auto show "t \ existence_ivl0 x0 \ flow0 x0 t \ S" by (rule one_step_result12[OF r aform.c1_info_of_appre_c0_I[OF x0] t S subset_UNIV lens]) (auto simp: aform.c1_info_invare_None lens X) qed theorem solves_one_step_ivl': assumes T: "T \ {t1 .. t2}" assumes X: "X \ {eucl_of_list lx .. eucl_of_list ux}" "length lx = DIM('a)" "length ux = DIM('a)" assumes DS: "list_interval lds uds \ list_interval ld ud" and lends: "length lds = DIM('a)*DIM('a)" "length uds = DIM('a)*DIM('a)" assumes S: "{eucl_of_list ls::'a .. eucl_of_list us} \ S" assumes lens0: "length ls = DIM('a)" "length us = DIM('a)" \ \TODO: this could be verified\ "length dx0s = DIM('a)*DIM('a)" assumes [simp]: "aform.D odo = DIM('a)" assumes r: "solves_one_step_until_time_aform optns odo ((1,1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s)) t1 t2 (ls, us) (Some (lds, uds))" shows "t \ T \ x0 \ X \ t \ existence_ivl0 x0 \ flow0 x0 t \ S \ Dflow x0 t o\<^sub>L blinfun_of_list dx0s \ blinfuns_of_lvivl (ld, ud)" proof (intro impI) assume t: "t \ T" and x0: "x0 \ X" from S have S: "set_of_lvivl (ls, us) \ S" by (auto simp: set_of_lvivl_def set_of_ivl_def) have lens: "length (fst (ls, us)) = CARD('n)" "length (snd (ls, us)) = CARD('n)" by (auto simp: lens0) have x0: "list_of_eucl x0 \ Joints (aforms_of_ivls lx ux)" apply (rule aforms_of_ivls) subgoal by (simp add: X) subgoal by (simp add: X) using subsetD[OF X(1) x0] apply (auto simp: eucl_le[where 'a='a] X) apply (metis assms(3) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) apply (metis assms(4) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) done have x0dx0: "(x0, blinfun_of_list dx0s) \ aform.c1_info_of_appre ((1, 1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s))" apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI[where x="list_of_eucl x0@dx0s"]) using lens0 apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) apply (rule x0) done from t T have t: "t \ {t1 .. t2}" by auto have DS: "blinfuns_of_lvivl' (Some (lds, uds)) \ blinfun_of_list ` list_interval ld ud" using DS by (auto simp: blinfuns_of_lvivl'_def blinfuns_of_lvivl_def) have inv: "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lds, uds))" "aform.c1_info_invare DIM('a) ((1::ereal, 1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s))" by (auto simp: aform.lvivl'_invar_def lends aform.c1_info_invare_def X lens0 power2_eq_square aform.c1_info_invar_def) from one_step_result123[OF r x0dx0 t S DS lens inv \aform.D _ = _\] show "t \ existence_ivl0 x0 \ flow0 x0 t \ S \ Dflow x0 t o\<^sub>L blinfun_of_list dx0s \ blinfuns_of_lvivl (ld, ud)" by (auto simp: blinfuns_of_lvivl_def) qed end definition "zero_aforms D = map (\_. (0, zero_pdevs)) [0..pf. solves_one_step_until_time_aform (snd soptns pf) a b c d e f)" definition "solves_poincare_map_aform'_fo soptns a b c d e f g h i = file_output (String.implode (fst soptns)) (\pf. solves_poincare_map_aform' (snd soptns pf) a b c d e f g h i)" definition "solves_poincare_map_onto_aform_fo soptns a b c d e f g h = file_output (String.implode (fst soptns)) (\pf. solves_poincare_map_onto_aform (snd soptns pf) a b c d e f g h)" lemma solves_one_step_until_time_aform_foI: "solves_one_step_until_time_aform (snd optns (\_. ())) a b c d e f" if "solves_one_step_until_time_aform_fo optns a b c d e f" using that by (auto simp: solves_one_step_until_time_aform_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) lemma solves_poincare_map_aform'_foI: "solves_poincare_map_aform' (snd optns (\_. ())) a b c d e f g h i" if "solves_poincare_map_aform'_fo optns a b c d e f g h i" using that by (auto simp: solves_poincare_map_aform'_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) lemma solves_poincare_map_onto_aform_foI: "solves_poincare_map_onto_aform (snd optns (\_. ())) a b c d e f g h" if "solves_poincare_map_onto_aform_fo optns a b c d e f g h" using that by (auto simp: solves_poincare_map_onto_aform_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) definition "can_mk_ode_ops fas safe_form \ mk_ode_ops fas safe_form \ None" theorem solve_one_step_until_time_aform_integral_bounds: fixes f::"real \ 'a::executable_euclidean_space" assumes "a \ b" assumes ba: "b - a \ {t1 .. t2}" assumes a: "a \ {a1 .. a2}" assumes ls_us_subset: "{eucl_of_list ls .. eucl_of_list us} \ {l .. u}" assumes fas: "\xs::real list. length xs > 0 \ interpret_form safe_form xs \ (1::real, f (xs ! 0)) = einterpret fas xs" assumes D: "D = DIM('a) + 1" "D = CARD('i::enum)" assumes lenlu: "length ls + 1 = D" "length us + 1 = D" assumes lfas: "length fas = D" assumes mv: "can_mk_ode_ops fas safe_form" assumes FDERIV: "\xs. interpret_form safe_form xs \ isFDERIV (length fas) [0.. {l .. u}" proof - have lens0: "length ((x::real) # replicate (D - 1) 0) = DIM(real \ 'a)" for x using assms by auto have a0: "(a, 0) \ {eucl_of_list (a1 # replicate (D - 1) 0)..eucl_of_list (a2 # replicate (D - 1) 0)}" using assms by (auto simp: eucl_of_list_prod) let ?U = "{x::real\'a. interpret_form safe_form (list_of_eucl x)}" interpret ode_interpretation safe_form ?U fas "\x. (1::real, f (fst x))" "undefined::'i" apply unfold_locales subgoal using assms by simp subgoal using assms by simp subgoal using mv by (simp add: D lfas) subgoal for x apply (cases x) by (rule HOL.trans[OF fas[symmetric]]) (auto simp: fas) subgoal using mv by (simp add: can_mk_ode_ops_def) subgoal by (rule FDERIV) done have lens: "length (0 # ls) = DIM(real \ 'a)" "length (t2 # us) = DIM(real \ 'a)" "aform.D odo = DIM(real \ 'a)" using lenlu by (simp_all add: lfas aform.D_def D aform.ode_e_def ) have D_odo: "aform.D odo = DIM(real \ 'a)" by (auto simp: aform.D_def aform.ode_e_def lfas D) from solves_one_step_ivl[rule_format, OF order_refl order_refl lens0 lens0 order_refl lens(1,2) D_odo, unfolded odo_def, OF sos ba a0] have lsus: "flow0 (a, 0) (b - a) \ {eucl_of_list (0#ls)..eucl_of_list (t2#us)}" and exivl: "b - a \ existence_ivl0 (a, 0)" by auto have flow: "flow0 (a, 0) (b - a) \ UNIV \ {l..u}" using lsus apply (rule rev_subsetD) using ls_us_subset by (auto simp: eucl_of_list_prod) from ivl_subset_existence_ivl[OF exivl] \a \ b\ exivl have "0 \ existence_ivl0 (a, 0)" by (auto simp del: existence_ivl_initial_time_iff) from mem_existence_ivl_iv_defined(2)[OF this] have safe: "(a, 0::'a) \ ?U" by simp from flow_solves_ode[OF UNIV_I this] have fs: "((\t. (fst (flow0 (a, 0) t), snd (flow0 (a, 0) t))) solves_ode (\_ x. (1, f (fst x)))) (existence_ivl0 (a, 0)) ?U" by simp with solves_odeD[OF fs] have vdp: "((\t. (fst (flow0 (a, 0) t), snd (flow0 (a, 0) t))) has_vderiv_on (\t. (1, f (fst (flow0 (a, 0) t))))) (existence_ivl0 (a, 0))" by simp have "fst (flow0 (a, 0) t) = a + t" if "t \ existence_ivl0 (a, 0)" for t proof - have "fst (flow0 (a, 0) 0) = a" using safe by (auto simp: ) have "((\t. fst (flow0 (a, 0) t)) has_vderiv_on (\x. 1)) (existence_ivl0 (a, 0))" using has_vderiv_on_PairD[OF vdp] by auto then have "((\t. fst (flow0 (a, 0) t)) has_vderiv_on (\x. 1)) {0--t}" apply (rule has_vderiv_on_subset) using closed_segment_subset_existence_ivl[OF that] by auto from fundamental_theorem_of_calculus_ivl_integral[OF this, THEN ivl_integral_unique] one_has_ivl_integral[of t 0, THEN ivl_integral_unique] safe show ?thesis by auto qed with vdp have "((\t. (t, snd (flow0 (a, 0) t))) solves_ode (\t x. (1, f (a + fst x)))) (existence_ivl0 (a, 0)) ((existence_ivl0 (a, 0)) \ UNIV)" apply (intro solves_odeI) apply auto apply (auto simp: has_vderiv_on_def has_vector_derivative_def) proof goal_cases case (1 x) have r: "((\(x, y). (x - a, y::'a)) has_derivative (\x. x)) (at x within t)" for x t by (auto intro!: derivative_eq_intros) from 1 have "((\x. (a + x, snd (flow0 (a, 0) x))) has_derivative (\xa. (xa, xa *\<^sub>R f (a + x)))) (at x within existence_ivl0 (a, 0))" by auto from has_derivative_in_compose2[OF r subset_UNIV _ this, simplified] 1 have "((\x. (x, snd (flow0 (a, 0) x))) has_derivative (\y. (y, y *\<^sub>R f (a + x)))) (at x within existence_ivl0 (a, 0))" by auto then show ?case by simp qed from solves_autonomous_odeI[OF this] have "((\t. snd (flow0 (a, 0) t)) solves_ode (\b c. f (a + b))) (existence_ivl0 (a, 0)) UNIV" by simp \ \TODO: do non-autonomous -- autonomous conversion automatically!\ then have "((\t. snd (flow0 (a, 0) t)) solves_ode (\b c. f (a + b))) {0 .. b - a} UNIV" apply (rule solves_ode_on_subset) using exivl by (rule ivl_subset_existence_ivl) (rule order_refl) from integral_solves_autonomous_odeI[OF this] have "((\b. f (a + b)) has_integral snd (flow0 (a, 0) (b - a))) (cbox 0 (b - a))" using \a \ b\ safe by auto from has_integral_affinity[OF this, where m=1 and c="-a"] have "(f has_integral snd (flow0 (a, 0) (b - a))) {a..b}" by auto then have "integral {a..b} f = snd (flow0 (a, 0) (b - a))" by blast also have "\ \ {l .. u}" using flow by auto finally show ?thesis by simp qed lemma [code_computation_unfold]: "numeral x = real_of_int (Code_Target_Int.positive x)" by simp lemma [code_computation_unfold]: "numeral x \ Float (Code_Target_Int.positive x) 0" by (simp add: Float_def float_of_numeral) definition no_print::"String.literal \ unit" where "no_print x = ()" lemmas [approximation_preproc] = list_of_eucl_real list_of_eucl_prod append.simps named_theorems DIM_simps lemmas [DIM_simps] = DIM_real DIM_prod length_nth_simps add_numeral_special add_numeral_special card_sum card_prod card_bit0 card_bit1 card_num0 card_num1 numeral_times_numeral numeral_mult mult_1_right mult_1 aform.D_def lemma numeral_refl: "numeral x = numeral x" by simp lemma plain_floatarith_approx_eq_SomeD: "approx prec fa [] = Some (the (approx prec fa []))" if "plain_floatarith 0 fa" using that by (auto dest!: plain_floatarith_approx_not_None[where p=prec and XS=Nil]) definition [simp]: "approx1 p f xs = real_of_float (lower (the (approx p f xs)))" definition [simp]: "approx2 p f xs = real_of_float (upper (the (approx p f xs)))" definition [simp]: "approx_defined p f xs \ approx p f xs \ None" definition "approxs p fs xs = those (map (\f. approx p f xs) fs)" definition [simp]: "approxs1 p f xs = (case approxs p f xs of Some y \ (map (real_of_float o lower) y) | None \ replicate (length f) 0)" definition [simp]: "approxs2 p f xs = (case approxs p f xs of Some y \ (map (real_of_float o upper) y) | None \ replicate (length f) 0)" definition "approxs_defined p fs xs \ (those (map (\f. approx p f xs) fs) \ None)" lemma length_approxs: "length (approxs1 p f xs) = length f" "length (approxs2 p f xs) = length f" by (auto simp: approxs_def dest!: those_eq_SomeD split: option.splits) lemma real_in_approxI: "x \ {(approx1 prec fa []) .. (approx2 prec fa [])}" if "x = interpret_floatarith fa []" "approx_defined prec fa []" using that by (force dest: approx_emptyD simp: set_of_eq) lemma real_subset_approxI: "{a .. b} \ {(approx1 prec fa []) .. (approx2 prec fb [])}" if "a = interpret_floatarith fa []" "b = interpret_floatarith fb []" "approx_defined prec fa []" "approx_defined prec fb []" using that by (force dest: approx_emptyD simp: set_of_eq) lemma approxs_eq_Some_lengthD: "length ys = length fas" if "approxs prec fas XS = Some ys" using that by (auto simp: approxs_def dest!: those_eq_SomeD) lemma approxs_pointwise: "interpret_floatarith (fas ! ia) xs \ {real_of_float (lower (ys ! ia)) .. (upper (ys ! ia))}" if "approxs prec fas XS = Some ys" "bounded_by xs XS" "ia < length fas" proof - from those_eq_SomeD[OF that(1)[unfolded approxs_def]] have ys: "ys = map (the \ (\f. approx prec f XS)) fas" and ex: "\y. i < length fas \ approx prec (fas ! i) XS = Some y" for i by auto from ex[of ia] that obtain ivl where ivl: "approx prec (fas ! ia) XS = Some ivl" by auto from approx[OF that(2) this] have "interpret_floatarith (fas ! ia) xs \\<^sub>r ivl" by auto moreover have "ys ! ia = ivl" unfolding ys apply (auto simp: o_def) apply (subst nth_map) apply (simp add: that) using ivl by simp ultimately show ?thesis using that by (auto simp: approxs_eq_Some_lengthD set_of_eq split: prod.splits) qed lemmas approxs_pointwise_le = approxs_pointwise[simplified, THEN conjunct1] and approxs_pointwise_ge = approxs_pointwise[simplified, THEN conjunct2] lemma approxs_eucl: "eucl_of_list (interpret_floatariths fas xs) \ {eucl_of_list (map lower ys) .. eucl_of_list (map upper ys)::'a::executable_euclidean_space}" if "approxs prec fas XS = Some ys" "length fas = DIM('a)" "bounded_by xs XS" using that by (auto simp: eucl_le[where 'a='a] eucl_of_list_inner o_def approxs_eq_Some_lengthD intro!: approxs_pointwise_le approxs_pointwise_ge) lemma plain_floatariths_approx_eq_SomeD: "approxs prec fas [] = Some (the (approxs prec fas []))" if "list_all (plain_floatarith 0) fas" using that apply (induction fas) apply (auto simp: approxs_def split: option.splits dest!: plain_floatarith_approx_eq_SomeD) subgoal for a fas aa apply (cases "those (map (\f. approx prec f []) fas)") by auto done lemma approxs_definedD: "approxs prec fas xs = Some (the (approxs prec fas xs))" if "approxs_defined prec fas xs" using that by (auto simp: approxs_defined_def approxs_def) lemma approxs_defined_ne_None[simp]: "approxs prec fas xs \ None" if "approxs_defined prec fas xs" using that by (auto simp: approxs_defined_def approxs_def) lemma approx_subset_euclI: "{eucl_of_list (approxs2 prec fals [])::'a .. eucl_of_list (approxs1 prec faus [])} \ {l .. u}" if "list_of_eucl l = interpret_floatariths fals []" and "list_of_eucl u = interpret_floatariths faus []" and "length fals = DIM('a::executable_euclidean_space)" and "length faus = DIM('a::executable_euclidean_space)" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that by (auto intro!: bounded_by_Nil dest!: approxs_eucl[where 'a='a] list_of_eucl_eqD plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) lemma eucl_subset_approxI: "{l .. u} \ {eucl_of_list (approxs1 prec fals [])::'a .. eucl_of_list (approxs2 prec faus [])}" if "list_of_eucl l = interpret_floatariths fals []" and "list_of_eucl u = interpret_floatariths faus []" and "length fals = DIM('a::executable_euclidean_space)" and "length faus = DIM('a::executable_euclidean_space)" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that by (auto intro!: bounded_by_Nil dest!: approxs_eucl[where 'a='a] list_of_eucl_eqD plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) lemma approx_subset_listI: "list_interval (approxs2 prec fals []) (approxs1 prec faus []) \ list_interval l u" if "l = interpret_floatariths fals []" and "u = interpret_floatariths faus []" and "length fals = length l" and "length faus = length u" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that apply (auto simp: list_interval_def list_all2_conv_all_nth dest: approxs_eq_Some_lengthD intro!: bounded_by_Nil dest!: plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) apply (rule order_trans) apply (rule approxs_pointwise_ge) apply assumption apply (rule bounded_by_Nil) apply (auto dest: approxs_eq_Some_lengthD) apply (subst interpret_floatariths_nth) apply (auto dest: approxs_eq_Some_lengthD) apply (rule approxs_pointwise_le[ge]) apply assumption apply (rule bounded_by_Nil) apply (auto dest: approxs_eq_Some_lengthD) done definition "unit_list D n = (replicate D 0)[n:=1]" definition "mirror_sctn (sctn::real list sctn) = Sctn (map uminus (normal sctn)) (- pstn sctn)" definition "mirrored_sctn b (sctn::real list sctn) = (if b then mirror_sctn sctn else sctn)" lemma mirror_sctn_simps[simp]: "pstn (mirror_sctn sctn) = - pstn sctn" "normal (mirror_sctn sctn) = map uminus (normal sctn)" by (cases sctn) (auto simp: mirror_sctn_def) lemma length_unit_list[simp]: "length (unit_list D n) = D" by (auto simp: unit_list_def) lemma eucl_of_list_unit_list[simp]: "(eucl_of_list (unit_list D n)::'a::executable_euclidean_space) = Basis_list ! n" if "D = DIM('a)" "n < D" using that by (auto simp: unit_list_def eucl_of_list_inner inner_Basis nth_list_update' intro!: euclidean_eqI[where 'a='a]) lemma le_eucl_of_list_iff: "(t::'a::executable_euclidean_space) \ eucl_of_list uX0 \ (\i. i < DIM('a) \ t \ Basis_list ! i \ uX0 ! i)" if "length uX0 = DIM('a)" using that apply (auto simp: eucl_le[where 'a='a] eucl_of_list_inner) using nth_Basis_list_in_Basis apply force by (metis Basis_list in_Basis_index_Basis_list index_less_size_conv length_Basis_list) lemma eucl_of_list_le_iff: "eucl_of_list uX0 \ (t::'a::executable_euclidean_space) \ (\i. i < DIM('a) \ uX0 ! i \ t \ Basis_list ! i)" if "length uX0 = DIM('a)" using that apply (auto simp: eucl_le[where 'a='a] eucl_of_list_inner) using nth_Basis_list_in_Basis apply force by (metis Basis_list in_Basis_index_Basis_list index_less_size_conv length_Basis_list) lemma Joints_aforms_of_ivls: "Joints (aforms_of_ivls lX0 uX0) = list_interval lX0 uX0" if "list_all2 (\) lX0 uX0" using that apply (auto simp: list_interval_def dest: Joints_aforms_of_ivlsD1[OF _ that] Joints_aforms_of_ivlsD2[OF _ that] list_all2_lengthD intro!: aforms_of_ivls) by (auto simp: list_all2_conv_all_nth) lemma list_of_eucl_in_list_interval_iff: "list_of_eucl x0 \ list_interval lX0 uX0 \ x0 \ {eucl_of_list lX0 .. eucl_of_list uX0::'a}" if "length lX0 = DIM('a::executable_euclidean_space)" "length uX0 = DIM('a::executable_euclidean_space)" using that by (auto simp: list_interval_def eucl_of_list_le_iff le_eucl_of_list_iff list_all2_conv_all_nth) text \TODO: make a tactic out of this?!\ lemma file_output_iff: "file_output s f = f (\_. ())" by (auto simp: file_output_def print_def[abs_def] Print.file_output_def) context ode_interpretation begin lemma poincare_mapsto_subset: "poincare_mapsto P X0 SS CX R" if "poincare_mapsto P' Y0 RR CZ S" "X0 \ Y0" "CZ \ CX" "S \ R" "RR = SS" "P = P'" using that by (force simp: poincare_mapsto_def) theorem solves_poincare_map_aform'_derivI: assumes solves: "solves_poincare_map_aform'_fo optns odo (Sctn (unit_list D n) (lP ! n)) guards (lP, uP) (Sctn (unit_list D n) (lP ! n)) roi [((1,1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] (lR, uR) (Some (lDR, uDR))" and D: "D = DIM('a)" assumes DS: "list_interval lDR uDR \ list_interval lDS uDS" and dim: "aform.D odo = DIM('a)" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" "length DX0 = DIM('a)*DIM('a)" "length lDR = CARD('n) * CARD('n)" "length uDR = CARD('n) * CARD('n)" and guards: "(\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a))" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and SS: "SS = {x::'a. x \ Basis_list ! n \ lP ! n}" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" shows "\x\X0. returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" proof (rule ballI) fix x assume "x \ X0" then have la2: "list_all2 (\) lX0 uX0" using X0 by (force simp: subset_iff eucl_of_list_le_iff le_eucl_of_list_iff lens list_all2_conv_all_nth) have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens power2_eq_square) have 2: "length (normal (Sctn (unit_list D n) (lP ! n))) = DIM('a)" by (auto simp: D) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (Sctn (unit_list D n) (lP ! n))) = DIM('a)" by (auto simp: D) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lDR, uDR))" by (auto simp: lens aform.lvivl'_invar_def) note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))]) (below_halfspace (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' (Some (lDR, uDR)))" by (rule solves_poincare_map_aform'[OF solves, OF 1 dim 2 3 4 guards 5]) auto then have "poincare_mapsto P (X0 \ {blinfun_of_list DX0}::('a \ ('a \\<^sub>L 'a)) set) SS UNIV (R \ blinfuns_of_lvivl (lDS, uDS))" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def aform.c1_info_of_apprse_def) subgoal for x0 apply (rule image_eqI[where x="list_of_eucl x0@DX0"]) using lens apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) using X0 by (auto simp: Joints_aforms_of_ivls la2 list_of_eucl_in_list_interval_iff) done subgoal by simp subgoal using R DS by (auto simp: set_of_lvivl_def set_of_ivl_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_def lens) subgoal using assms by (simp add: below_halfspace_def le_halfspace_def[abs_def]) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff) done then show "returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" using \x \ X0\ by (auto simp: poincare_mapsto_def) qed definition guards_invar::"nat \ (((real list \ real list) \ real list sctn) list \ (real \ real pdevs) reach_options) list \ bool" where "guards_invar D guards = (\(xs, ro) \ set guards. \((a, b), ba) \ set xs. length a = D \ length b = D \ length (normal ba) = D)" theorem solves_poincare_map_aform'I: assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" assumes D: "D = DIM('a)" assumes guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and dim: "aform.D odo = DIM('a)" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" and solves: "solves_poincare_map_aform'_fo optns odo (Sctn (unit_list D n) (lP ! n)) guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, None)] (lR, uR) None" shows "\x\X0. returns_to P x \ poincare_map P x \ R" proof - note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, None)] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens) have 2: "length (normal ((mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (((Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) from guards have guards: "(xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)" for xs ro a b ba by (auto simp: guards_invar_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) None" by (auto simp: lens) have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, None)]) (below_halfspace (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' None)" by (rule solves_poincare_map_aform'[OF solves, OF 1 dim 2 3 4 guards 5]) then have "poincare_mapsto P (X0 \ UNIV::('a \ ('a \\<^sub>L 'a)) set) (below_halfspace (map_sctn eucl_of_list (((Sctn (unit_list D n) (lP ! n)))))) UNIV (R \ UNIV)" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_apprse_def aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI) apply (rule eucl_of_list_list_of_eucl[symmetric]) apply (rule aforms_of_ivls) by (auto simp add: lens subset_iff le_eucl_of_list_iff eucl_of_list_le_iff) subgoal by simp subgoal using R by (auto simp: set_of_lvivl_def set_of_ivl_def) subgoal using assms by (simp add: below_halfspace_def le_halfspace_def[abs_def]) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "\x\X0. returns_to P x \ poincare_map P x \ R" by (auto simp: poincare_mapsto_def) qed definition "poincare_map_from_outside = poincare_map" theorem poincare_maps_onto_aformI: assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" assumes D: "D = DIM('a)" assumes guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and dim: "aform.D odo = DIM('a)" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" and solves: "solves_poincare_map_onto_aform_fo optns odo guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, None)] (lR, uR) None" shows "\x\X0. returns_to P x \ poincare_map_from_outside P x \ R" proof - note solves = solves[unfolded solves_poincare_map_onto_aform_fo_def file_output_iff] have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, None)] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens) have 2: "length (normal ((mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) from guards have guards: "(xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)" for xs ro a b ba by (auto simp: guards_invar_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) None" by (auto simp: lens) have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, None)]) UNIV (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' None)" by (rule solves_poincare_map_onto_aform[OF solves, OF 1 dim 2 3 guards 5]) then have "poincare_mapsto P (X0 \ UNIV::('a \ ('a \\<^sub>L 'a)) set) UNIV UNIV (R \ UNIV)" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_apprse_def aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI) apply (rule eucl_of_list_list_of_eucl[symmetric]) apply (rule aforms_of_ivls) by (auto simp add: lens subset_iff le_eucl_of_list_iff eucl_of_list_le_iff) subgoal by simp subgoal using R by (auto simp: set_of_lvivl_def set_of_ivl_def) subgoal by simp subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "\x\X0. returns_to P x \ poincare_map_from_outside P x \ R" by (auto simp: poincare_mapsto_def poincare_map_from_outside_def) qed end lemmas [simp] = length_approxs context includes ode_ops.lifting begin lift_definition empty_ode_ops::"ode_ops" is "([], true_form)" by (auto simp: ) end ML \val ode_numerics_conv = @{computation_check terms: Trueprop Not solves_one_step_until_time_aform_fo solves_poincare_map_aform'_fo solves_poincare_map_onto_aform_fo num_options num_options_c1 ro (* nat *) Suc "0::nat" "1::nat" "(+)::nat \ nat \ nat" "(-) ::nat \ nat \ nat" "(=)::nat\nat\bool" "(^)::nat\nat\nat" (* int / integer*) "(=)::int\int\bool" "(+)::int\int\int" "uminus::_\int" "uminus::_\integer" int_of_integer integer_of_int "0::int" "1::int" "(^)::int\nat\int" (* real *) "(=)::real\real\bool" "real_of_float" "(/)::real\real\real" "(^)::real\nat\real" "uminus::real\_" "(+)::real\real\real" "(-)::real\real\real" "(*)::real\real\real" real_divl real_divr real_of_int "0::real" "1::real" (* rat *) Fract "0::rat" "1::rat" "(+)::rat\rat\rat" "(-)::rat\rat\rat" "(*)::rat\rat\rat" "uminus::rat\_" "(/)::rat\rat\rat" "(^)::rat\nat\rat" (* ereal *) "1::ereal" (* lists: *) "replicate::_\real\_" "unit_list::_\_\real list" "Nil::(nat \ nat \ string) list" "Cons::_\_\(nat \ nat \ string) list" "Nil::(nat \ nat \ string \ nat list) list" "Cons::_\_\(nat \ nat \ string \ nat list) list" "Nil::real list" "Cons::_\_\real list" "Nil::nat list" "Cons::_\_\nat list" "Nil::string list" "Cons::_\_\string list" "Nil::real aform list" "Cons::_\_\real aform list" "Nil::(float interval) option list" "Cons::_\_\(float interval) option list" "nth::_\_\real" "upt" (* products: *) "Pair::_\_\(nat \ string)" "Pair::_\_\(nat \ nat \ string)" "Pair::_\_\char list \ nat list" "Pair::_\_\nat \ char list \ nat list" "Pair::_\_\nat \ nat \ char list \ nat list" "Pair::_\_\char list \ ((String.literal \ unit) \ (real \ real pdevs) numeric_options)" "Pair::_\_\ereal\ereal" "Pair::_\_\real aform list \ real aform list option" "Pair::_\_\(ereal \ ereal) \ real aform list \ real aform list option" "Pair::_\_\real aform" "Pair::_\_\real list \ real list" "Nil::(((real list \ real list) \ real list sctn) list \ (real aform) reach_options) list" "Cons::_\_\(((real list \ real list) \ real list sctn) list \ (real aform) reach_options) list" "Nil::((real list \ real list) \ real list sctn) list" "Cons::_\_\((real list \ real list) \ real list sctn) list" "Pair::_\_\((real list \ real list) \ real list sctn) list \ real aform reach_options" "Nil::((ereal \ ereal) \ (real aform) list \ (real aform) list option) list" "Cons::_\_\((ereal \ ereal) \ (real aform) list \ (real aform) list option) list" (* option *) "None::(real aform) list option" "Some::_\(real aform) list option" "None::(real list \ real list) option" "Some::_\(real list \ real list) option" (* aforms *) "aform_of_ivl::_\_\real aform" aforms_of_ivls aforms_of_point (* pdevs*) "zero_pdevs::real pdevs" "zero_aforms::_ \ real aform list" (* ode_ops *) mk_ode_ops init_ode_ops empty_ode_ops can_mk_ode_ops "the::ode_ops option \ ode_ops" the_odo (* Characters/Strings *) String.Char String.implode "Nil::string" "Cons::_\_\string" (* float *) "(=)::float\float\bool" "(+)::float\float\float" "uminus::_\float" "(-)::_\_\float" Float float_of_int float_of_nat (* approximation... *) approx approx1 approx2 approxs1 approxs2 approx_defined approxs_defined (* floatarith *) "0::floatarith" "1::floatarith" "(+)::_\_\floatarith" "(-)::_\_\floatarith" "(*)::_\_\floatarith" "(/)::_\_\floatarith" "inverse::_\floatarith" "uminus::_\floatarith" "Sum\<^sub>e::_\nat list\floatarith" Sin Half Tan R\<^sub>e Norm (* form *) true_form Half (* Printing *) print no_print (* sections *) xsec xsec' ysec ysec' zsec zsec' xsec2 xsec2' ysec2 ysec2' ysec4 ysec4' mirrored_sctn Code_Target_Nat.natural Code_Target_Int.positive Code_Target_Int.negative Code_Numeral.positive Code_Numeral.negative datatypes: bool num floatarith "floatarith list" form "real list sctn" "real \ real" } \ ML \fun ode_numerics_tac ctxt = CONVERSION (ode_numerics_conv ctxt) THEN' (resolve_tac ctxt [TrueI])\ lemma eq_einterpretI: assumes "list_of_eucl (VS::'a::executable_euclidean_space) = interpret_floatariths fas xs" assumes "length fas = DIM('a)" shows "VS = eucl_of_list (interpret_floatariths fas xs)" using assms apply (subst (asm) list_of_eucl_eucl_of_list[symmetric]) apply (auto intro!: ) by (metis eucl_of_list_list_of_eucl) lemma one_add_square_ne_zero[simp]: "1 + t * t \ 0" for t::real by (metis semiring_normalization_rules(12) sum_squares_eq_zero_iff zero_neq_one) lemma abs_rat_bound: "abs (x - y) \ e / f" if "y \ {yl .. yu}" "x \ {yu - real_divl p e f.. yl + real_divl p e f}" for x y e::real proof - note \x \ _\ also have "{yu - real_divl p e f.. yl + real_divl p e f} \ {yu - e / f .. yl + e / f}" by (auto intro!: diff_mono real_divl) finally show ?thesis using that unfolding abs_diff_le_iff by auto qed lemma in_ivl_selfI: "a \ {a .. a}" for a::real by auto lemma pi4_bnds: "pi / 4 \ {real_divl 80 (lb_pi 80) 4 .. real_divr 80 (ub_pi 80) 4}" using pi_boundaries[of 80] unfolding atLeastAtMost_iff by (intro conjI real_divl[le] real_divr[ge] divide_right_mono) auto lemma abs_minus_leI: "\x - x'\ \ e" if "x \ {x' - e .. x' + e}" for x e::real using that by (auto simp: ) lemmas [DIM_simps] = Suc_numeral One_nat_def[symmetric] TrueI Suc_1 length_approxs arith_simps lemma (in ode_interpretation) length_ode_e[DIM_simps]: "length (ode_expression odo) = DIM('a)" by (auto simp: len) named_theorems solves_one_step_ivl_thms context ode_interpretation begin lemmas [solves_one_step_ivl_thms] = TAG_optns[OF solves_one_step_ivl[OF _ _ _ _ _ _ _ _ solves_one_step_until_time_aform_foI], rotated -1, of optns _ _ _ _ _ _ _ _ _ optns for optns] lemmas [solves_one_step_ivl_thms] = TAG_optns[OF solves_one_step_ivl'[OF _ _ _ _ _ _ _ _ _ _ _ _ solves_one_step_until_time_aform_foI], rotated -1, of optns _ _ _ _ _ _ _ _ _ _ _ _ _ _ optns for optns] lemmas [solves_one_step_ivl_thms] = solves_poincare_map_aform'I poincare_maps_onto_aformI end lemma TAG_optnsI: "TAG_optns optns" by simp named_theorems poincare_tac_theorems lemmas [DIM_simps] = one_less_numeral_iff rel_simps abbreviation "point_ivl a \ {a .. a}" lemma isFDERIV_compute: "isFDERIV D vs fas xs \ (list_all (\i. list_all (\j. isDERIV (vs ! i) (fas ! j) xs) [0.. length fas = D \ length vs = D" unfolding isFDERIV_def by (auto simp: list.pred_set) theorem (in ode_interpretation) solves_poincare_map_aform'_derivI'[solves_one_step_ivl_thms]: \ \TODO: replace @{thm solves_poincare_map_aform'_derivI}\ assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" and D: "D = DIM('a)" assumes DS: "list_interval lDR uDR \ list_interval lDS uDS" and ode_fas: "aform.D odo = DIM('a)" and guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" "length DX0 = DIM('a)*DIM('a)" "length lDR = CARD('n) * CARD('n)" "length uDR = CARD('n) * CARD('n)" and SS: "SS = {x::'a. if mirrored then x \ Basis_list ! n \ lP ! n else x \ Basis_list ! n \ lP ! n}" assumes solves: "solves_poincare_map_aform'_fo optns odo (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n))) guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] (lR, uR) (Some (lDR, uDR))" shows "\x\X0. returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" proof (rule ballI) fix x assume "x \ X0" then have la2: "list_all2 (\) lX0 uX0" using X0 by (force simp: subset_iff eucl_of_list_le_iff le_eucl_of_list_iff lens list_all2_conv_all_nth) have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens power2_eq_square) have 2: "length (normal (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n)))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n)))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lDR, uDR))" by (auto simp: lens aform.lvivl'_invar_def) note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))]) (below_halfspace (map_sctn eucl_of_list (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n))))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' (Some (lDR, uDR)))" by (rule solves_poincare_map_aform'[OF solves, OF 1 ode_fas 4 3 2 _ 5]) (use guards in \auto simp: guards_invar_def\) then have "poincare_mapsto P (X0 \ {blinfun_of_list DX0}::('a \ ('a \\<^sub>L 'a)) set) SS UNIV (R \ blinfuns_of_lvivl (lDS, uDS))" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def aform.c1_info_of_apprse_def) subgoal for x0 apply (rule image_eqI[where x="list_of_eucl x0@DX0"]) using lens apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) using X0 by (auto simp: Joints_aforms_of_ivls la2 list_of_eucl_in_list_interval_iff) done subgoal by simp subgoal using R DS by (auto simp: set_of_lvivl_def set_of_ivl_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_def lens) subgoal using assms by (auto simp: below_halfspace_def le_halfspace_def[abs_def] mirrored_sctn_def mirror_sctn_def) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" using \x \ X0\ by (auto simp: poincare_mapsto_def) qed lemmas [DIM_simps] = aform.ode_e_def ML \ structure ODE_Numerics_Tac = struct fun mk_nat n = HOLogic.mk_number @{typ nat} n fun mk_int n = HOLogic.mk_number @{typ int} n fun mk_integer n = @{term integer_of_int} $ (HOLogic.mk_number @{typ int} n) fun mk_bool b = if b then @{term True} else @{term False} fun mk_numeralT n = let fun mk_bit 0 ty = Type (@{type_name bit0}, [ty]) | mk_bit 1 ty = Type (@{type_name bit1}, [ty]); fun bin_of n = if n = 1 then @{typ num1} else if n = 0 then @{typ num0} else if n = ~1 then raise TERM ("negative type numeral", []) else let val (q, r) = Integer.div_mod n 2; in mk_bit r (bin_of q) end; in bin_of n end; fun print_tac' ctxt s = K (print_tac ctxt s) val using_master_directory = File.full_path o Resources.master_directory o Proof_Context.theory_of; fun using_master_directory_term ctxt s = (if s = "-" orelse s = "" then s else Path.explode s |> using_master_directory ctxt |> Path.implode) |> HOLogic.mk_string fun real_in_approx_tac ctxt p = let val inst_approx = ([], [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) val approx_thm = Thm.instantiate inst_approx @{thm real_in_approxI} in resolve_tac ctxt [approx_thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' ode_numerics_tac ctxt end fun real_subset_approx_tac ctxt p = let val inst_approx = ([], [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) val approx_thm = Thm.instantiate inst_approx @{thm real_subset_approxI} in resolve_tac ctxt [approx_thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' ode_numerics_tac ctxt THEN' ode_numerics_tac ctxt end fun basic_nt_ss ctxt nt = put_simpset HOL_basic_ss ctxt addsimps Named_Theorems.get ctxt nt fun DIM_tac defs ctxt = (Simplifier.simp_tac (basic_nt_ss ctxt @{named_theorems DIM_simps} addsimps defs)) fun subset_approx_preconds_tac ctxt p thm = let val inst_approx = ([], [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) in resolve_tac ctxt [Thm.instantiate inst_approx thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (DIM_tac [] ctxt) THEN' SOLVED' (DIM_tac [] ctxt) THEN' SOLVED' (ode_numerics_tac ctxt) THEN' SOLVED' (ode_numerics_tac ctxt) end val cfg_trace = Attrib.setup_config_bool @{binding ode_numerics_trace} (K false) fun tracing_tac ctxt = if Config.get ctxt cfg_trace then print_tac ctxt else K all_tac fun tracing_tac' ctxt = fn s => K (tracing_tac ctxt s) fun eucl_subset_approx_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm eucl_subset_approxI} fun approx_subset_eucl_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm approx_subset_euclI} fun approx_subset_list_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm approx_subset_listI} val static_simpset = Simplifier.simpset_of @{context} fun nth_tac ctxt = Simplifier.simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms nth_Cons_0 nth_Cons_Suc numeral_nat}) fun nth_list_eq_tac ctxt n = Subgoal.FOCUS_PARAMS (fn {context, concl, ...} => case try (Thm.term_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) concl of SOME (@{const List.nth(real)} $ xs $ Var _, @{const List.nth(real)} $ ys $ Var _) => let val i = find_index (op=) (HOLogic.dest_list xs ~~ HOLogic.dest_list ys) val thm = Goal.prove context [] [] (HOLogic.mk_eq (@{const List.nth(real)} $ xs $ HOLogic.mk_number @{typ nat} i, @{const List.nth(real)} $ ys $ HOLogic.mk_number @{typ nat} i) |> HOLogic.mk_Trueprop) (fn {context, ...} => HEADGOAL (nth_tac context)) in SOLVE (HEADGOAL (resolve_tac context [thm])) end | _ => no_tac ) ctxt n fun numeric_precond_step_tac defs thms p = Subgoal.FOCUS_PARAMS (fn {context, concl, ...} => let val prems = Logic.strip_imp_prems (Thm.term_of concl) val conclusion = Logic.strip_imp_concl (Thm.term_of concl) in (case conclusion |> HOLogic.dest_Trueprop of @{const Set.member(real)} $ _ $ _ => tracing_tac context "numeric_precond_step: real in approx" THEN HEADGOAL (real_in_approx_tac context p) | Const(@{const_name less_eq}, _) $ (Const (@{const_name atLeastAtMost}, _) $ _ $ _) $ (Const (@{const_name atLeastAtMost}, _) $ Var _ $ Var _) => tracing_tac context "numeric_precond_step: approx subset eucl" THEN HEADGOAL (real_subset_approx_tac context p) | Const (@{const_name less_eq}, _) $ (Const (@{const_name atLeastAtMost}, _) $ (Const (@{const_name eucl_of_list}, _) $ Var _) $ _) $ _ => tracing_tac context "numeric_precond_step: approx subset eucl" THEN HEADGOAL (approx_subset_eucl_tac context p) | Const (@{const_name less_eq}, _) $ _ $ (Const (@{const_name atLeastAtMost}, _) $ (Const (@{const_name eucl_of_list}, _) $ Var _) $ _) => tracing_tac context "numeric_precond_step: eucl subset approx" THEN HEADGOAL (eucl_subset_approx_tac context p) | Const (@{const_name less_eq}, _) $ (@{const list_interval(real)} $ _ $ _) $ (@{const list_interval(real)} $ _ $ _) => tracing_tac context "numeric_precond_step: approx subset list" THEN HEADGOAL (approx_subset_list_tac context p) | @{const HOL.eq(nat)} $ _ $ _ => tracing_tac context "numeric_precond_step: DIM_tac" THEN HEADGOAL (SOLVED' (DIM_tac [] context)) | @{const less(nat)} $ _ $ _ => tracing_tac context "numeric_precond_step: DIM_tac" THEN HEADGOAL (SOLVED' (DIM_tac [] context)) | @{const HOL.eq(real)} $ (@{const nth(real)} $ _ $ _) $ (@{const nth(real)} $ _ $ _) => tracing_tac context "numeric_precond_step: nth_list_eq_tac" THEN HEADGOAL (SOLVED' (nth_list_eq_tac context)) | Const (@{const_name "HOL.eq"}, _) $ _ $ (Const (@{const_name eucl_of_list}, _) $ (@{const interpret_floatariths} $ _ $ _)) => tracing_tac context "numeric_precond_step: reify floatariths" THEN HEADGOAL (resolve_tac context @{thms eq_einterpretI} THEN' reify_floatariths_tac context) | t as _ $ _ => let val (c, args) = strip_comb t in if member (op=) [@{const "solves_one_step_until_time_aform_fo"}, @{const "solves_poincare_map_aform'_fo"}, @{const "solves_poincare_map_onto_aform_fo"}, @{const "can_mk_ode_ops"} ] c then tracing_tac context "numeric_precond_step: ode_numerics_tac" THEN HEADGOAL ( CONVERSION (Simplifier.rewrite (put_simpset HOL_basic_ss context addsimps defs)) THEN' tracing_tac' context "numeric_precond_step: ode_numerics_tac (unfolded)" THEN' ode_numerics_tac context) else if member (op=) [@{const "isFDERIV"}] c then tracing_tac context "numeric_precond_step: isFDERIV" THEN HEADGOAL (SOLVED'(Simplifier.asm_full_simp_tac (put_simpset static_simpset context addsimps (@{thms isFDERIV_def less_Suc_eq_0_disj isDERIV_Power_iff} @ thms @ defs)) THEN' tracing_tac' context "numeric_precond_step: simplified isFDERIV" )) else tracing_tac context "numeric_precond_step: boolean, try thms" THEN HEADGOAL (SOLVED' (resolve_tac context thms)) end | _ => tracing_tac context "numeric_precond_step: boolean constant" THEN no_tac ) end) fun integral_bnds_tac_gen_start sstep d p m N atol filename ctxt i = let val insts = ([((("'i", 0), @{sort "{enum}"}), mk_numeralT (d + 1) |> Thm.ctyp_of ctxt)], [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, (@{term num_options} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ @{term "[(0::nat, 1::nat, ''0x000000'')]"})) |> Thm.cterm_of ctxt), ((("safe_form", 0), @{typ form}), @{cterm true_form}) ]) in resolve_tac ctxt [Thm.instantiate insts @{thm solve_one_step_until_time_aform_integral_bounds}] i THEN (Lin_Arith.tac ctxt i ORELSE Simplifier.simp_tac ctxt i) end fun integral_bnds_tac_gen sstep d p m N atol filename thms ctxt = integral_bnds_tac_gen_start sstep d p m N atol filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac [] thms p ctxt) val integral_bnds_tac = integral_bnds_tac_gen 5 fun mk_proj (m, n, s) = HOLogic.mk_tuple [mk_nat m, mk_nat n, HOLogic.mk_string s] fun mk_projs projs = HOLogic.mk_list @{typ "nat \ nat \ string"} (map mk_proj projs) fun mk_string_list ds = HOLogic.mk_list @{typ "string"} (map HOLogic.mk_string ds) fun mk_nat_list ds = HOLogic.mk_list @{typ "nat"} (map mk_nat ds) fun mk_proj_c1 (m, n, s, ds) = HOLogic.mk_tuple [mk_nat m, mk_nat n, HOLogic.mk_string s, mk_nat_list ds] fun mk_projs_c1 projs = HOLogic.mk_list @{typ "nat \ nat \ string \ nat list"} (map mk_proj_c1 projs) fun TAG_optns_thm p sstep m N atol projs filename ctxt = Thm.instantiate ([], [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, @{term num_options} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ mk_projs projs) |> Thm.cterm_of ctxt)]) @{thm TAG_optnsI} fun TAG_optns_c1_thm p sstep m N atol projs ds filename ctxt = Thm.instantiate ([], [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, @{term num_options_c1} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ mk_projs_c1 projs $ mk_string_list ds) |> Thm.cterm_of ctxt)]) @{thm TAG_optnsI} fun ode_bnds_tac_gen_start sstep p m N atol projs filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun ode_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = ode_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac ode_defs [] p ctxt) val ode_bnds_tac = ode_bnds_tac_gen 5 fun ode'_bnds_tac_gen_start sstep p m N atol projs ds filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_c1_thm p sstep m N atol projs ds filename ctxt] fun ode'_bnds_tac_gen sstep ode_defs p m N atol projs ds filename ctxt = ode'_bnds_tac_gen_start sstep p m N atol projs ds filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac ode_defs [] p ctxt) val ode'_bnds_tac = ode'_bnds_tac_gen 5 fun poincare_bnds_tac_gen_start sstep p m N atol projs filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun poincare_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = poincare_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD ( numeric_precond_step_tac ode_defs (Named_Theorems.get ctxt @{named_theorems poincare_tac_theorems}) p ctxt) val poincare_bnds_tac = poincare_bnds_tac_gen 5 fun poincare'_bnds_tac_gen_start sstep p m N atol projs filename ctxt = resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun poincare'_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = poincare'_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD ( numeric_precond_step_tac ode_defs (Named_Theorems.get ctxt @{named_theorems poincare_tac_theorems}) p ctxt) val poincare'_bnds_tac = poincare'_bnds_tac_gen 5 end \ lemma (in auto_ll_on_open) Poincare_Banach_fixed_pointI: assumes "convex S" and c: "complete S" "S \ {}" and "S \ T" assumes derivative_bounded: "\x\S. poincare_map \ x \ S \ (\D. (poincare_map \ has_derivative D) (at x within T) \ onorm D \ B)" assumes B: "B < 1" shows "\!x. x \ S \ poincare_map \ x = x" using c _ B proof (rule banach_fix) from derivative_bounded c show "0 \ B" by (auto dest!: has_derivative_bounded_linear onorm_pos_le) from derivative_bounded show "poincare_map \ ` S \ S" by auto obtain D where D: "\x \ S. (poincare_map \ has_derivative D x) (at x within T) \ onorm (D x) \ B" apply atomize_elim apply (rule bchoice) using derivative_bounded by auto with \S \ T\ have "(\x. x \ S \ (poincare_map \ has_derivative D x) (at x within S))" by (auto intro: has_derivative_within_subset) from bounded_derivative_imp_lipschitz[of S "poincare_map \" D B, OF this] \convex S\ D c \0 \ B\ have "B-lipschitz_on S (poincare_map \)" by auto then show "\x\S. \y\S. dist (poincare_map \ x) (poincare_map \ y) \ B * dist x y" by (auto simp: lipschitz_on_def) qed ML \open ODE_Numerics_Tac\ lemma isFDERIV_product: "isFDERIV n xs fas vs \ length fas = n \ length xs = n \ list_all (\(x, f). isDERIV x f vs) (List.product xs fas)" apply (auto simp: isFDERIV_def list_all2_iff in_set_zip list_all_length product_nth) - apply (metis gt_or_eq_0 less_mult_imp_div_less mod_less_divisor not_less0) - by auto + apply auto + apply (metis gr_implies_not_zero gr_zeroI less_mult_imp_div_less pos_mod_bound) + done end diff --git a/thys/VerifyThis2018/Challenge2.thy b/thys/VerifyThis2018/Challenge2.thy --- a/thys/VerifyThis2018/Challenge2.thy +++ b/thys/VerifyThis2018/Challenge2.thy @@ -1,597 +1,601 @@ chapter \Colored Tiles\ section \Challenge\ text_raw \{\upshape This problem is based on Project Euler problem \#114. Alice and Bob are decorating their kitchen, and they want to add a single row of fifty tiles on the edge of the kitchen counter. Tiles can be either red or black, and for aesthetic reasons, Alice and Bob insist that red tiles come by blocks of at least three consecutive tiles. Before starting, they wish to know how many ways there are of doing this. They come up with the following algorithm: \begin{lstlisting}[language=C,morekeywords={procedure,function,end,to,in,var,then,not,mod}] var count[51] // count[i] is the number of valid rows of size i count[0] := 1 // [] count[1] := 1 // [B] - cannot have a single red tile count[2] := 1 // [BB] - cannot have one or two red tiles count[3] := 2 // [BBB] or [RRR] for n = 4 to 50 do count[n] := count[n-1] // either the row starts with a black tile for k = 3 to n-1 do // or it starts with a block of k red tiles count[n] := count[n] + count[n-k-1] // followed by a black one end-for count[n] := count[n]+1 // or the entire row is red end-for \end{lstlisting} \paragraph{Verification tasks.} You should verify that at the end, \texttt{count[50]} will contain the right number. \bigskip \noindent\emph{Hint:} Since the algorithm works by enumerating the valid colorings, we expect you to give a nice specification of a valid coloring and to prove the following properties: \begin{enumerate} \item Each coloring counted by the algorithm is valid. \item No coloring is counted twice. \item No valid coloring is missed. \end{enumerate} } \clearpage \ section \Solution\ theory Challenge2 imports "lib/VTcomp" begin text \ The algorithm describes a dynamic programming scheme. Instead of proving the 3 properties stated in the challenge separately, we approach the problem by \<^enum> Giving a natural specification of a valid tiling as a grammar \<^enum> Deriving a recursion equation for the number of valid tilings \<^enum> Verifying that the program returns the correct number (which obviously implies all three properties stated in the challenge) \ subsection \Problem Specification\ subsubsection \Colors\ datatype color = R | B subsubsection \Direct Natural Definition of a Valid Line\ inductive valid where "valid []" | "valid xs \ valid (B # xs)" | "valid xs \ n \ 3 \ valid (replicate n R @ xs)" definition "lcount n = card {l. length l=n \ valid l}" subsection \Derivation of Recursion Equations\ text \This alternative variant helps us to prove the split lemma below.\ inductive valid' where "valid' []" | "n \ 3 \ valid' (replicate n R)" | "valid' xs \ valid' (B # xs)" | "valid' xs \ n \ 3 \ valid' (replicate n R @ B # xs)" lemma valid_valid': "valid l \ valid' l" by (induction rule: valid.induct) (auto 4 4 intro: valid'.intros elim: valid'.cases simp: replicate_add[symmetric] append_assoc[symmetric] ) lemmas valid_red = valid.intros(3)[OF valid.intros(1), simplified] lemma valid'_valid: "valid' l \ valid l" by (induction rule: valid'.induct) (auto intro: valid.intros valid_red) lemma valid_eq_valid': "valid' l = valid l" using valid_valid' valid'_valid by metis subsubsection \Additional Facts on Replicate\ lemma replicate_iff: "(\i (\ n. l = replicate n R)" by auto (metis (full_types) in_set_conv_nth replicate_eqI) lemma replicate_iff2: "(\i (\ l'. l = replicate n R @ l')" if "n < length l" using that by (auto simp: list_eq_iff_nth_eq nth_append intro: exI[where x = "drop n l"]) lemma replicate_Cons_eq: "replicate n x = y # ys \ (\ n'. n = Suc n' \ x = y \ replicate n' x = ys)" by (cases n) auto subsubsection \Main Case Analysis on \@term valid\\ lemma valid_split: "valid l \ l = [] \ (l!0 = B \ valid (tl l)) \ length l \ 3 \ (\ i < length l. l ! i = R) \ (\ j < length l. j \ 3 \ (\ i < j. l ! i = R) \ l ! j = B \ valid (drop (j + 1) l))" unfolding valid_eq_valid'[symmetric] apply standard subgoal by (erule valid'.cases) (auto simp: nth_append nth_Cons split: nat.splits) subgoal by (auto intro: valid'.intros simp: replicate_iff elim!: disjE1) (fastforce intro: valid'.intros simp: neq_Nil_conv replicate_iff2 nth_append)+ done subsubsection \Base cases\ lemma lc0_aux: "{l. l = [] \ valid l} = {[]}" by (auto intro: valid.intros) lemma lc0: "lcount 0 = 1" by (auto simp: lc0_aux lcount_def) lemma lc1aux: "{l. length l=1 \ valid l} = {[B]}" by (auto intro: valid.intros elim: valid.cases simp: replicate_Cons_eq) lemma lc2aux: "{l. length l=2 \ valid l} = {[B,B]}" by (auto 4 3 intro: valid.intros elim: valid.cases simp: replicate_Cons_eq) + lemma valid_3R: \valid [R, R, R]\ + using valid.intros(3) [of \[]\ 3] by (simp add: numeral_eq_Suc valid.intros) + lemma lc3_aux: "{l. length l=3 \ valid l} = {[B,B,B], [R,R,R]}" - by (auto 4 4 intro: valid.intros valid_red[of 3, simplified] elim: valid.cases - simp: replicate_Cons_eq) + by (auto 4 4 intro: valid.intros valid_3R elim: valid.cases + simp: replicate_Cons_eq) + lemma lcounts_init: "lcount 0 = 1" "lcount 1 = 1" "lcount 2 = 1" "lcount 3 = 2" using lc0 lc1aux lc2aux lc3_aux unfolding lcount_def by simp_all subsubsection \The Recursion Case\ lemma finite_valid_length: "finite {l. length l = n \ valid l}" (is "finite ?S") proof - have "?S \ lists {R, B} \ {l. length l = n}" by (auto intro: color.exhaust) moreover have "finite \" by (auto intro: lists_of_len_fin1) ultimately show ?thesis by (rule finite_subset) qed lemma valid_line_just_B: "valid (replicate n B)" by (induction n) (auto intro: valid.intros) lemma valid_line_aux: "{l. length l = n \ valid l} \ {}" (is "?S \ {}") using valid_line_just_B[of n] by force lemma replicate_unequal_aux: "replicate x R @ B # l \ replicate y R @ B # l'" (is "?l \ ?r") if \x < y\ for l l' proof - have "?l ! x = B" "?r ! x = R" using that by (auto simp: nth_append) then show ?thesis by auto qed lemma valid_prepend_B_iff: "valid (B # xs) \ valid xs" by (auto intro: valid.intros elim: valid.cases simp: Cons_replicate_eq Cons_eq_append_conv) lemma lcrec: "lcount n = lcount (n-1) + 1 + (\i=3..n>3\ proof - have "{l. length l = n \ valid l} = {l. length l = n \ valid (tl l) \ l!0=B} \ {l. length l = n \ (\ i. i < n \ i \ 3 \ (\ k < i. l!k = R) \ l!i = B \ valid (drop (i + 1) l))} \ {l. length l = n \ (\i ?D \ ?C") using \n > 3\ by (subst valid_split) auto let ?B1 = "((#) B) ` {l. length l = n - Suc 0 \ valid l}" from \n > 3\ have "?B = ?B1" apply safe subgoal for l by (cases l) (auto simp: valid_prepend_B_iff) by auto have 1: "card ?B1 = lcount (n-1)" unfolding lcount_def by (auto intro: card_image) have "?C = {replicate n R}" by (auto simp: nth_equalityI) have 2: "card {replicate n R} = 1" by auto let ?D1="(\ i \ {3.. l. replicate i R @ B # l)` {l. length l = n - i - 1 \ valid l})" have "?D = (\i \ {3.. (\ k < i. l!k = R) \ l!i = B \ valid (drop (i + 1) l)})" by auto have "{l. length l = n \ (\ k < i. l!k = R) \ l!i = B \ valid (drop (i + 1) l)} = (\ l. replicate i R @ B # l)` {l. length l = n - i - 1 \ valid l}" if "i < n" "2 < i" for i apply safe subgoal for l apply (rule image_eqI[where x = "drop (i + 1) l"]) apply (rule nth_equalityI) using that apply (simp_all split: nat.split add: nth_Cons nth_append) using add_diff_inverse_nat apply fastforce done using that by (simp add: nth_append; fail)+ then have D_eq: "?D = ?D1" unfolding \?D = _\ by auto have inj: "inj_on (\l. replicate x R @ B # l) {l. length l = n - Suc x \ valid l}" for x unfolding inj_on_def by auto have *: "(\l. replicate x R @ B # l) ` {l. length l = n - Suc x \ valid l} \ (\l. replicate y R @ B # l) ` {l. length l = n - Suc y \ valid l} = {}" if "3 \ x" "x < y" "y < n" for x y using that replicate_unequal_aux[OF \x < y\] by auto have 3: "card ?D1 = (\i=3..?A = _\ \?B = _\ \?C = _\ D_eq apply (subst card_Un_disjoint) (* Finiteness *) apply (blast intro: finite_subset[OF _ finite_valid_length])+ (* Disjointness *) subgoal using Cons_replicate_eq[of B _ n R] replicate_unequal_aux by fastforce apply (subst card_Un_disjoint) (* Finiteness *) apply (blast intro: finite_subset[OF _ finite_valid_length])+ (* Disjointness & final rewriting *) unfolding 1 2 3 by (auto simp: Cons_replicate_eq Cons_eq_append_conv) qed subsection \Verification of Program\ subsubsection \Inner Loop: Summation\ definition "sum_prog \ l u f \ nfoldli [l.._. True) (\i s. doN { ASSERT (\ i); RETURN (s+f i) }) 0" lemma sum_spec[THEN SPEC_trans, refine_vcg]: assumes "l\u" assumes "\i. l\i \ i \ i" shows "sum_prog \ l u f \ SPEC (\r. r=(\i=l..Main Program\ definition "icount M \ doN { ASSERT (M>2); let c = op_array_replicate (M+1) 0; let c = c[0:=1, 1:=1, 2:=1, 3:=2]; ASSERT (\i<4. c!i = lcount i); c\nfoldli [4.._. True) (\n c. doN { \<^cancel>\\let sum = (\i=3..\ sum \ sum_prog (\i. n-i-1 < length c) 3 n (\i. c!(n-i-1)); ASSERT (n-1 ni\M. c!i = lcount i); ASSERT (M < length c); RETURN (c!M) }" subsubsection \Abstract Correctness Statement\ theorem icount_correct: "M>2 \ icount M \ SPEC (\r. r=lcount M)" unfolding icount_def thm nfoldli_upt_rule supply nfoldli_upt_rule[where I="\n c. length c = M+1 \ ( \i{0,1,2,3}") using lcounts_init by (auto) subgoal for i c j apply (cases "jRefinement to Imperative Code\ sepref_definition icount_impl is "icount" :: "nat_assn\<^sup>k \\<^sub>a nat_assn" unfolding icount_def sum_prog_def by sepref subsubsection \Main Correctness Statement\ text \ As the main theorem, we prove the following Hoare triple, stating: starting from the empty heap, our program will compute the correct result (@{term "lcount M"}). \ theorem icount_impl_correct: "M>2 \ icount_impl M <\r. \(r = lcount M)>\<^sub>t" proof - note A = icount_impl.refine[to_hnr, THEN hn_refineD] note A = A[unfolded autoref_tag_defs] note A = A[unfolded hn_ctxt_def pure_def, of M M, simplified] note [sep_heap_rules] = A assume "M>2" show ?thesis using icount_correct[OF \M>2\] by (sep_auto simp: refine_pw_simps pw_le_iff) qed subsubsection \Code Export\ export_code icount_impl in SML_imp module_name Tiling export_code icount_impl in OCaml_imp module_name Tiling export_code icount_impl in Haskell module_name Tiling export_code icount_impl in Scala_imp module_name Tiling subsection \Alternative Problem Specification\ text \Alternative definition of a valid line that we used in the competition\ context fixes l :: "color list" begin inductive valid_point where "\i+2 \ valid_point i" | "\1\i;i+1 \ valid_point i" | "\2\i; i \ valid_point i" | "\ i \ valid_point i" definition "valid_line = (\i i. i < length l \ valid_point l i" shows "valid_line l" using assms unfolding valid_line_def by auto lemma valid_B_first: "valid_point xs i \ i < length xs \ valid_point (B # xs) (i + 1)" by (auto intro: valid_point.intros simp: numeral_2_eq_2 elim!: valid_point.cases) lemma valid_line_prepend_B: "valid_line (B # xs)" if "valid_line xs" using that apply - apply (rule valid_lineI) subgoal for i by (cases i) (auto intro: valid_B_first[simplified] valid_point.intros simp: valid_line_def) done lemma valid_drop_B: "valid_point xs (i - 1)" if "valid_point (B # xs) i" "i > 0" using that apply cases apply (fastforce intro: valid_point.intros) subgoal by (cases "i = 1") (auto intro: valid_point.intros(2)) subgoal unfolding numeral_nat by (cases "i = 2") (auto intro: valid_point.intros(3)) apply (fastforce intro: valid_point.intros) done lemma valid_line_drop_B: "valid_line xs" if "valid_line (B # xs)" using that unfolding valid_line_def proof (safe, goal_cases) case (1 i) with valid_drop_B[of xs "i + 1"] show ?case by auto qed lemma valid_line_prepend_B_iff: "valid_line (B # xs) \ valid_line xs" using valid_line_prepend_B valid_line_drop_B by metis lemma cases_valid_line: assumes "l = [] \ (l!0 = B \ valid_line (tl l)) \ length l \ 3 \ (\ i < length l. l ! i = R) \ (\ j < length l. j \ 3 \ (\ i < j. l ! i = R) \ l ! j = B \ valid_line (drop (j + 1) l))" (is "?a \ ?b \ ?c \ ?d") shows "valid_line l" proof - from assms consider (empty) ?a | (B) "\ ?a \ ?b" | (all_red) ?c | (R_B) ?d by blast then show ?thesis proof cases case empty then show ?thesis by (simp add: valid_line_def) next case B then show ?thesis by (cases l) (auto simp: valid_line_prepend_B_iff) next case prems: all_red show ?thesis proof (rule valid_lineI) fix i assume "i < length l" consider "i = 0" | "i = 1" | "i > 1" by atomize_elim auto then show "valid_point l i" using \i < _\ prems by cases (auto 4 4 intro: valid_point.intros) qed next case R_B then obtain j where j: "j j" "(\ij \ 3\ consider "i \ j - 3" | "i = j - 2" | "i = j - 1" | "i = j" | "i > j" by atomize_elim auto then show "valid_point l i" proof cases case 5 with \valid_line _\ \i < length l\ have "valid_point (drop (j + 1) l) (i - j - 1)" unfolding valid_line_def by auto then show ?thesis using \i > j\ by cases (auto intro: valid_point.intros) qed (use j in \auto intro: valid_point.intros\) qed qed qed lemma valid_line_cases: "l = [] \ (l!0 = B \ valid_line (tl l)) \ length l \ 3 \ (\ i < length l. l ! i = R) \ (\ j < length l. j \ 3 \ (\ i < j. l ! i = R) \ l ! j = B \ valid_line (drop (j + 1) l))" if "valid_line l" proof (cases "l = []") case True then show ?thesis by (simp add: valid_line_def) next case False show ?thesis proof (cases "l!0 = B") case True with \l \ []\ have "l = B # tl l" by (cases l) auto with \valid_line l\ True show ?thesis by (metis valid_line_prepend_B_iff) next case False from \valid_line l\ \l \ []\ have "valid_point l 0" unfolding valid_line_def by auto with False have red_start: "length l \ 3" "l!0 = R" "l!1 = R" "l!2 = R" by (auto elim!: valid_point.cases simp: numeral_2_eq_2) show ?thesis proof (cases "\i < length l. l ! i = R") case True with \length l \ 3\ show ?thesis by auto next case False let ?S = "{j. j < length l \ j \ 3 \ l ! j = B}" let ?j = "Min ?S" have B_ge_3: "i \ 3" if "l ! i = B" for i proof - consider "i = 0" | "i = 1" | "i = 2" | "i \ 3" by atomize_elim auto then show "i \ 3" using red_start \l ! i = B\ by cases auto qed from False obtain i where "l ! i = B" "i < length l" "i \ 3" by (auto intro: B_ge_3 color.exhaust) then have "?j \ ?S" by - (rule Min_in, auto) have "\i < ?j. l ! i = R" proof - { fix i assume "i < ?j" "l ! i = B" then have "i \ 3" by (auto intro: B_ge_3) with \i < ?j\ \l ! i = B\ red_start \?j \ ?S\ have "i \ ?S" by auto then have "?j \ i" by (auto intro: Min_le) with \i < ?j\ have False by simp } then show ?thesis by (auto intro: color.exhaust) qed with \?j \ ?S\ obtain j where j: "j < length l" "j \ 3" "\i < j. l ! i = R" "l ! j = B" by blast moreover have "valid_line (drop (j + 1) l)" proof (rule valid_lineI) fix i assume "i < length (drop (j + 1) l)" with j \valid_line l\ have "valid_point l (j + i + 1)" unfolding valid_line_def by auto then show "valid_point (drop (j + 1) l) i" proof cases case 2 then show ?thesis using j by (cases i) (auto intro: valid_point.intros) next case prems: 3 consider "i = 0" | "i = 1" | "i > 1" by atomize_elim auto then show ?thesis using j prems by cases (auto intro: valid_point.intros) qed (auto intro: valid_point.intros) qed ultimately show ?thesis by auto qed qed qed lemma valid_line_split: "valid_line l \ l = [] \ (l!0 = B \ valid_line (tl l)) \ length l \ 3 \ (\ i < length l. l ! i = R) \ (\ j < length l. j \ 3 \ (\ i < j. l ! i = R) \ l ! j = B \ valid_line (drop (j + 1) l))" using valid_line_cases cases_valid_line by blast text \Connection to the easier definition given above\ lemma valid_valid_line: "valid l \ valid_line l" by (induction l rule: length_induct, subst valid_line_split, subst valid_split, auto) end diff --git a/thys/VerifyThis2018/Challenge3.thy b/thys/VerifyThis2018/Challenge3.thy --- a/thys/VerifyThis2018/Challenge3.thy +++ b/thys/VerifyThis2018/Challenge3.thy @@ -1,1315 +1,1312 @@ chapter \Array-Based Queuing Lock\ section \Challenge\ text_raw \{\upshape Array-Based Queuing Lock (ABQL) is a variation of the Ticket Lock algorithm with a bounded number of concurrent threads and improved scalability due to better cache behaviour. %(\url{https://en.wikipedia.org/wiki/Array_Based_Queuing_Locks}) We assume that there are \texttt{N} threads and we allocate a shared Boolean array \texttt{pass[]} of length \texttt{N}. We also allocate a shared integer value \texttt{next}. In practice, \texttt{next} is an unsigned bounded integer that wraps to 0 on overflow, and we assume that the maximal value of \texttt{next} is of the form $k\mathtt{N} - 1$. Finally, we assume at our disposal an atomic \texttt{fetch\_and\_add} instruction, such that \texttt{fetch\_and\_add(next,1)} increments the value of \texttt{next} by 1 and returns the original value of \texttt{next}. The elements of \texttt{pass[]} are spinlocks, assigned individually to each thread in the waiting queue. Initially, each element of \texttt{pass[]} is set to \texttt{false}, except \texttt{pass[0]} which is set to \texttt{true}, allowing the first coming thread to acquire the lock. Variable \texttt{next} contains the number of the first available place in the waiting queue and is initialized to 0. Here is an implementation of the locking algorithm in pseudocode: \begin{lstlisting}[language=C,morekeywords={procedure,function,end,to,in,var,then,not,mod}] procedure abql_init() for i = 1 to N - 1 do pass[i] := false end-for pass[0] := true next := 0 end-procedure function abql_acquire() var my_ticket := fetch_and_add(next,1) mod N while not pass[my_ticket] do end-while return my_ticket end-function procedure abql_release(my_ticket) pass[my_ticket] := false pass[(my_ticket + 1) mod N] := true end-procedure \end{lstlisting} Each thread that acquires the lock must eventually release it by calling \texttt{abql\_release(my\_ticket)}, where \texttt{my\_ticket} is the return value of the earlier call of \texttt{abql\_acquire()}. We assume that no thread tries to re-acquire the lock while already holding it, neither it attempts to release the lock which it does not possess. Notice that the first assignment in \texttt{abql\_release()} can be moved at the end of \texttt{abql\_acquire()}. \bigskip \noindent\textbf{Verification task~1.}~Verify the safety of ABQL under the given assumptions. Specifically, you should prove that no two threads can hold the lock at any given time. \bigskip \noindent\textbf{Verification task~2.}~Verify the fairness, namely that the threads acquire the lock in order of request. \bigskip \noindent\textbf{Verification task~3.}~Verify the liveness under a fair scheduler, namely that each thread requesting the lock will eventually acquire it. \bigskip You have liberty of adapting the implementation and specification of the concurrent setting as best suited for your verification tool. In particular, solutions with a fixed value of \texttt{N} are acceptable. We expect, however, that the general idea of the algorithm and the non-deterministic behaviour of the scheduler shall be preserved. } \clearpage \ section \Solution\ theory Challenge3 imports "lib/VTcomp" "lib/DF_System" begin text \The Isabelle Refinement Framework does not support concurrency. However, Isabelle is a general purpose theorem prover, thus we can model the problem as a state machine, and prove properties over runs. For this polished solution, we make use of a small library for transition systems and simulations: @{theory VerifyThis2018.DF_System}. Note, however, that our definitions are still quite ad-hoc, and there are lots of opportunities to define libraries that make similar proofs simpler and more canonical. We approach the final ABQL with three refinement steps: \<^enum> We model a ticket lock with unbounded counters, and prove safety, fairness, and liveness. \<^enum> We bound the counters by \mod N\ and \mod (k*N) respectively\ \<^enum> We implement the current counter by an array, yielding exactly the algorithm described in the challenge. With a simulation argument, we transfer the properties of the abstract system over the refinements. The final theorems proving safety, fairness, and liveness can be found at the end of this chapter, in Subsection~\ref{sec:main_theorems}. \ subsection \General Definitions\ text \We fix a positive number \N\ of threads\ consts N :: nat specification (N) N_not0[simp, intro!]: "N\0" by auto lemma N_gt0[simp, intro!]: "0A thread's state, representing the sequence points in the given algorithm. This will not change over the refinements.\ datatype thread = INIT | is_WAIT: WAIT (ticket: nat) | is_HOLD: HOLD (ticket: nat) | is_REL: REL (ticket: nat) subsection \Refinement 1: Ticket Lock with Unbounded Counters\ text \System's state: Current ticket, next ticket, thread states\ type_synonym astate = "nat \ nat \ (nat \ thread)" abbreviation "cc \ fst" abbreviation "nn s \ fst (snd s)" abbreviation "tts s \ snd (snd s)" text \The step relation of a single thread\ inductive astep_sng where enter_wait: "astep_sng (c,n,INIT) (c,(n+1),WAIT n)" | loop_wait: "c\k \ astep_sng (c,n,WAIT k) (c,n,WAIT k)" | exit_wait: "astep_sng (c,n,WAIT c) (c,n,HOLD c)" | start_release: "astep_sng (c,n,HOLD k) (c,n,REL k)" | release: "astep_sng (c,n,REL k) (k+1,n,INIT)" text \The step relation of the system\ inductive alstep for t where "\ t \ alstep t (c,n,ts) (c',n',ts(t:=s'))" text \Initial state of the system\ definition "as\<^sub>0 \ (0, 0, \_. INIT)" interpretation A: system as\<^sub>0 alstep . text \In our system, each thread can always perform a step\ lemma never_blocked: "A.can_step l s \ lThus, our system is in particular deadlock free\ interpretation A: df_system as\<^sub>0 alstep apply unfold_locales subgoal for s using never_blocked[of 0 s] unfolding A.can_step_def by auto done subsubsection \Safety: Mutual Exclusion\ text \Predicates to express that a thread uses or holds a ticket\ definition "has_ticket s k \ s=WAIT k \ s=HOLD k \ s=REL k" lemma has_ticket_simps[simp]: "\has_ticket INIT k" "has_ticket (WAIT k) k'\ k'=k" "has_ticket (HOLD k) k'\ k'=k" "has_ticket (REL k) k'\ k'=k" unfolding has_ticket_def by auto definition "locks_ticket s k \ s=HOLD k \ s=REL k" lemma locks_ticket_simps[simp]: "\locks_ticket INIT k" "\locks_ticket (WAIT k) k'" "locks_ticket (HOLD k) k'\ k'=k" "locks_ticket (REL k) k'\ k'=k" unfolding locks_ticket_def by auto lemma holds_imp_uses: "locks_ticket s k \ has_ticket s k" unfolding locks_ticket_def by auto text \We show the following invariant. Intuitively, it can be read as follows: \<^item> Current lock is less than or equal next lock \<^item> For all threads that use a ticket (i.e., are waiting, holding, or releasing): \<^item> The ticket is in between current and next \<^item> No other thread has the same ticket \<^item> Only the current ticket can be held (or released) \ definition "invar1 \ \(c,n,ts). c \ n \ (\t k. t has_ticket (ts t) k \ c \ k \ k < n \ (\t' k'. t' has_ticket (ts t') k' \ t\t' \ k\k') \ (\k. k\c \ \locks_ticket (ts t) k) ) " lemma is_invar1: "A.is_invar invar1" apply rule subgoal by (auto simp: invar1_def as\<^sub>0_def) subgoal for s s' apply (clarify) apply (erule alstep.cases) apply (erule astep_sng.cases) apply (clarsimp_all simp: invar1_def) apply fastforce apply fastforce apply fastforce apply fastforce by (metis Suc_le_eq holds_imp_uses locks_ticket_def le_neq_implies_less) done text \From the above invariant, it's straightforward to show mutual exclusion\ theorem mutual_exclusion: "\A.reachable s; tt'; is_HOLD (tts s t); is_HOLD (tts s t') \ \ False" apply (cases "tts s t"; simp) apply (cases "tts s t'"; simp) using A.invar_reachable[OF is_invar1, of s] apply (auto simp: invar1_def) by (metis locks_ticket_simps(3) has_ticket_simps(3)) lemma mutual_exclusion': "\A.reachable s; tt'; locks_ticket (tts s t) tk; locks_ticket (tts s t') tk' \ \ False" apply (cases "tts s t"; simp; cases "tts s t'"; simp) apply (cases "tts s t'"; simp) using A.invar_reachable[OF is_invar1, of s] apply (clarsimp_all simp: invar1_def) unfolding locks_ticket_def has_ticket_def apply metis+ done subsubsection \Fairness: Ordered Lock Acquisition\ text \We first show an auxiliary lemma: Consider a segment of a run from \i\ to \j\. Every thread that waits for a ticket in between the current ticket at \i\ and the current ticket at \j\ will be granted the lock in between \i\ and \j\. \ lemma fair_aux: assumes R: "A.is_run s" assumes A: "i k" "k < cc (s j)" "tl. i\l \ l tts (s l) t = HOLD k" proof - interpret A: run as\<^sub>0 alstep s by unfold_locales fact from A show ?thesis proof (induction "j-i" arbitrary: i) case 0 then show ?case by auto next case (Suc i') hence [simp]: "i'=j - Suc i" by auto note IH=Suc.hyps(1)[OF this] obtain t' where "alstep t' (s i) (s (Suc i))" by (rule A.stepE) then show ?case using Suc.prems proof cases case (1 c n ts c' n' s') note [simp] = "1"(1,2,3) from A.run_invar[OF is_invar1, of i] have "invar1 (c,n,ts)" by auto note I1 = this[unfolded invar1_def, simplified] from "1"(4) show ?thesis proof (cases rule: astep_sng.cases) case enter_wait then show ?thesis using IH Suc.prems apply (auto) by (metis "1"(2) Suc_leD Suc_lessI fst_conv leD thread.distinct(1)) next case (loop_wait k) then show ?thesis using IH Suc.prems apply (auto) by (metis "1"(2) Suc_leD Suc_lessI fst_conv leD) next case exit_wait then show ?thesis apply (cases "t'=t") subgoal using Suc.prems apply clarsimp by (metis "1"(2) Suc_leD Suc_lessI fst_conv fun_upd_same leD less_or_eq_imp_le snd_conv) subgoal using Suc.prems IH apply auto by (metis "1"(2) Suc_leD Suc_lessI fst_conv leD) done next case (start_release k) then show ?thesis using IH Suc.prems apply (auto) by (metis "1"(2) Suc_leD Suc_lessI fst_conv leD thread.distinct(7)) next case (release k) then show ?thesis apply (cases "t'=t") using I1 IH Suc.prems apply (auto) by (metis "1"(2) "1"(3) Suc_leD Suc_leI Suc_lessI fst_conv locks_ticket_simps(4) le_antisym not_less_eq_eq has_ticket_simps(2) has_ticket_simps(4)) qed qed qed qed lemma s_case_expand: "(case s of (c, n, ts) \ P c n ts) = P (cc s) (nn s) (tts s)" by (auto split: prod.splits) text \ A version of the fairness lemma which is very detailed on the actual ticket numbers. We will weaken this later. \ lemma fair_aux2: assumes RUN: "A.is_run s" assumes ACQ: "t0 alstep s by unfold_locales fact from ACQ WAIT have [simp]: "t\t'" "t'\t" by auto from ACQ have [simp]: "nn (s i) = k \ nn (s (Suc i)) = Suc k \ cc (s (Suc i)) = cc (s i) \ tts (s (Suc i)) = (tts (s i))(t:=WAIT k)" apply (rule_tac A.stepE[of i]) apply (erule alstep.cases) apply (erule astep_sng.cases) by (auto simp: nth_list_update split: if_splits) from A.run_invar[OF is_invar1, of i] have "invar1 (s i)" by auto note I1 = this[unfolded invar1_def, unfolded s_case_expand, simplified] from WAIT I1 have "k' < k" by fastforce from ACQ HOLD have "Suc i \ j" by auto with HOLD have "Suc i < j" by auto have X1: "cc (s i) \ k'" using I1 WAIT by fastforce have X2: "k' < cc (s j)" using A.run_invar[OF is_invar1, of j, unfolded invar1_def s_case_expand] using \k' < k\ \t HOLD(2) apply clarsimp by (metis locks_ticket_simps(3) has_ticket_simps(3)) from fair_aux[OF RUN \Suc i < j\, of k' t', simplified] obtain l where "l\Suc i" "l WAIT tk" obtains l where "ij" "tts (s l) t = HOLD tk" proof - interpret A: run as\<^sub>0 alstep s by unfold_locales fact from WAIT(2) NWAIT have "\l. i l\j \ tts (s l) t = HOLD tk" proof (induction "j-i" arbitrary: i) case 0 then show ?case by auto next case (Suc i') hence [simp]: "i'=j - Suc i" by auto note IH=Suc.hyps(1)[OF this] obtain t' where "alstep t' (s i) (s (Suc i))" by (rule A.stepE) then show ?case apply - apply (cases "t=t'";erule alstep.cases; erule astep_sng.cases) apply auto using IH Suc.prems Suc.hyps(2) apply (auto) apply (metis Suc_lessD Suc_lessI fun_upd_same snd_conv) apply (metis Suc_lessD Suc_lessI fun_upd_other snd_conv) apply (metis Suc.prems(1) Suc_lessD Suc_lessI fun_upd_triv) apply (metis Suc_lessD Suc_lessI fun_upd_other snd_conv) apply (metis Suc_lessD Suc_lessI fun_upd_other snd_conv) apply (metis Suc_lessD Suc_lessI fun_upd_other snd_conv) done qed thus ?thesis using that by blast qed text \ Finally we can show fairness, which we state as follows: Whenever a thread \t\ gets a ticket, all other threads \t'\ waiting for the lock will be granted the lock before \t\. \ theorem fair: assumes RUN: "A.is_run s" assumes ACQ: "t \Thread \t\ calls \acquire\ in step \i\\ assumes HOLD: "i \Thread \t\ holds lock in step \j\\ assumes WAIT: "t' \Thread \t'\ waits for lock at step \i\\ obtains l where "i \Then, \t'\ gets lock earlier\ proof - obtain k where Wk: "tts (s (Suc i)) t = WAIT k" using ACQ by (cases "tts (s (Suc i)) t") auto obtain k' where Wk': "tts (s i) t' = WAIT k'" using WAIT by (cases "tts (s i) t'") auto from ACQ HOLD have "Suc i \ j" by auto with HOLD have "Suc i < j" by auto obtain j' where H': "Suc i < j'" "j' \ j" "tts (s j') t = HOLD k" apply (rule find_hold_position[OF RUN \t Wk \Suc i < j\]) using HOLD(2) by auto show ?thesis apply (rule fair_aux2[OF RUN ACQ(1,2) Wk _ H'(3) WAIT(1) Wk']) subgoal using H'(1) by simp subgoal apply (erule that) using H'(2) by auto done qed subsubsection \Liveness\ text \For all tickets in between the current and the next ticket, there is a thread that has this ticket\ definition "invar2 \ \(c,n,ts). \k. c\k \ k (\t0_def) subgoal for s s' apply (clarsimp simp: invar2_def) apply (erule alstep.cases; erule astep_sng.cases; clarsimp) apply (metis less_antisym has_ticket_simps(1)) subgoal by (metis has_ticket_simps(2)) subgoal by (metis has_ticket_simps(2)) subgoal by (metis has_ticket_simps(3)) subgoal apply (frule A.invar_reachable[OF is_invar1]) unfolding invar1_def apply clarsimp by (metis Suc_leD locks_ticket_simps(4) not_less_eq_eq has_ticket_simps(4)) done done text \If a thread t is waiting for a lock, the current lock is also used by a thread\ corollary current_lock_used: assumes R: "A.reachable (c,n,ts)" assumes WAIT: "tUsed tickets are unique (Corollary from invariant 1)\ lemma has_ticket_unique: "\A.reachable (c,n,ts); t \ t'=t" apply (drule A.invar_reachable[OF is_invar1]) by (auto simp: invar1_def) text \We define the thread that holds a specified ticket\ definition "tkt_thread \ \ts k. THE t. t has_ticket (ts t) k" lemma tkt_thread_eq: assumes R: "A.reachable (c,n,ts)" assumes A: "tFor the inductive argument, we will use this measure, that decreases as a single thread progresses through its phases. \ definition "tweight s \ case s of WAIT _ \ 3::nat | HOLD _ \ 2 | REL _ \ 1 | INIT \ 0" text \We show progress: Every thread that waits for the lock will eventually hold the lock.\ theorem progress: assumes FRUN: "A.is_fair_run s" assumes A: "tj>i. is_HOLD (tts (s j) t)" proof - interpret A: fair_run as\<^sub>0 alstep s by unfold_locales fact from A obtain k where Wk: "tts (s i) t = WAIT k" by (cases "tts (s i) t") auto text \We use the following induction scheme: \<^item> Either the current thread increases (if we reach \k\, we are done) \<^item> (lex) the thread using the current ticket makes a step \<^item> (lex) another thread makes a step \ define lrel where "lrel \ inv_image (measure id <*lex*> measure id <*lex*> measure id) (\i. ( k-cc (s i), tweight (tts (s i) (tkt_thread (tts (s i)) (cc (s i)))), A.dist_step (tkt_thread (tts (s i)) (cc (s i))) i ))" have "wf lrel" unfolding lrel_def by auto then show ?thesis using A(1) Wk proof (induction i) case (less i) text \We name the components of this and the next state\ obtain c n ts where [simp]: "s i = (c,n,ts)" by (cases "s i") from A.run_reachable[of i] have R: "A.reachable (c,n,ts)" by simp obtain c' n' ts' where [simp]: "s (Suc i) = (c',n',ts')" by (cases "s (Suc i)") from A.run_reachable[of "Suc i"] have R': "A.reachable (c',n',ts')" by simp from less.prems have WAIT[simp]: "ts t = WAIT k" by simp { text \If thread \t\ left waiting state, we are done\ assume "ts' t \ WAIT k" hence "ts' t = HOLD k" using less.prems apply (rule_tac A.stepE[of i]) apply (auto elim!: alstep.cases astep_sng.cases split: if_splits) done hence ?case by auto } moreover { assume [simp]: "ts' t = WAIT k" text \Otherwise, we obtain the thread \tt\ that holds the current lock\ obtain tt where UTT: "tttt have "A.can_step tt (s i)" by (simp add: never_blocked) hence ?case proof (cases rule: A.rstep_cases) case (other t') \ \Another thread than \tt\ makes a step.\ text \The current ticket and \tt\'s state remain the same\ hence [simp]: "c' = c \ ts' tt = ts tt" using has_ticket_unique[OF R UTT, of t'] unfolding A.rstep_def using holds_only_current[OF R, of t'] by (force elim!: alstep.cases astep_sng.cases) text \Thus, \tt\ is still using the current ticket\ have [simp]: "tkt_thread ts' c = tt" using UTT tkt_thread_eq[OF R', of tt c] by auto text \Only the distance to \tt\'s next step has decreased\ have "(Suc i, i) \ lrel" unfolding lrel_def tweight_def by (simp add: other) text \And we can apply the induction hypothesis\ with less.IH[of "Suc i"] \t show ?thesis apply (auto) using Suc_lessD by blast next case THIS: this \ \The thread \tt\ that uses the current ticket makes a step\ show ?thesis proof (cases "\k'. ts tt = REL k'") case True \ \\tt\ has finished releasing the lock \ then have [simp]: "ts tt = REL c" using UTT by auto text \Thus, current increases\ have [simp]: "c' = Suc c" using THIS apply - unfolding A.rstep_def apply (erule alstep.cases, erule astep_sng.cases) by auto text \But is still less than \k\\ from A.invar_reachable[OF is_invar1 R] have "k>c" apply (auto simp: invar1_def) by (metis UTT WAIT \ts tt = REL c\ le_neq_implies_less less.prems(1) thread.distinct(9) has_ticket_simps(2)) text \And we can apply the induction hypothesis\ hence "(Suc i, i) \ lrel" unfolding lrel_def by auto with less.IH[of "Suc i"] \t show ?thesis apply (auto) using Suc_lessD by blast next case False \ \\tt\ has acquired the lock, or started releasing it\ text \Hence, current remains the same, but the weight of \tt\ decreases\ hence [simp]: " c' = c \ tweight (ts tt) > tweight (ts' tt) \ has_ticket (ts' tt) c" using THIS UTT apply - unfolding A.rstep_def apply (erule alstep.cases, erule astep_sng.cases) by (auto simp: has_ticket_def tweight_def) text \\tt\ still holds the current lock\ have [simp]: "tkt_thread ts' c = tt" using tkt_thread_eq[OF R' \tt, of c] by simp text \And we can apply the IH\ have "(Suc i, i) \ lrel" unfolding lrel_def by auto with less.IH[of "Suc i"] \t show ?thesis apply (auto) using Suc_lessD by blast qed qed } ultimately show ?case by blast qed qed subsection \Refinement 2: Bounding the Counters\ text \We fix the \k\ from the task description, which must be positive\ consts k::nat specification (k) k_not0[simp]: "k\0" by auto lemma k_gt0[simp]: "0System's state: Current ticket, next ticket, thread states\ type_synonym bstate = "nat \ nat \ (nat \ thread)" text \The step relation of a single thread\ inductive bstep_sng where enter_wait: "bstep_sng (c,n,INIT) (c,(n+1) mod (k*N),WAIT (n mod N))" | loop_wait: "c\tk \ bstep_sng (c,n,WAIT tk) (c,n,WAIT tk)" | exit_wait: "bstep_sng (c,n,WAIT c) (c,n,HOLD c)" | start_release: "bstep_sng (c,n,HOLD tk) (c,n,REL tk)" | release: "bstep_sng (c,n,REL tk) ((tk+1) mod N,n,INIT)" text \The step relation of the system, labeled with the thread \t\ that performs the step\ inductive blstep for t where "\ t \ blstep t (c,n,ts) (c',n',ts(t:=s'))" text \Initial state of the system\ definition "bs\<^sub>0 \ (0, 0, \_. INIT)" interpretation B: system bs\<^sub>0 blstep . lemma b_never_blocked: "B.can_step l s \ l0 blstep apply unfold_locales subgoal for s using b_never_blocked[of 0 s] unfolding B.can_step_def by auto done subsubsection \Simulation\ text \We show that the abstract system simulates the concrete one.\ text \A few lemmas to ease the automation further below\ lemma nat_sum_gtZ_iff[simp]: "finite s \ sum f s \ (0::nat) \ (\x\s. f x \ 0)" by simp lemma n_eq_Suc_sub1_conv[simp]: "n = Suc (n - Suc 0) \ n\0" by auto lemma mod_mult_mod_eq[mod_simps]: "x mod (k * N) mod N = x mod N" by (meson dvd_eq_mod_eq_0 mod_mod_cancel mod_mult_self2_is_0) - lemma mod_eq_imp_eq_aux: "b mod N = (a::nat) mod N \ a\b \ b b=a" - by (metis Groups.add_ac add_0_right - le_add_diff_inverse less_diff_conv2 nat_minus_mod - nat_minus_mod_plus_right mod_if) +lemma mod_eq_imp_eq_aux: "b mod N = (a::nat) mod N \ a\b \ b b=a" + using nat_mod_eq_lemma by force lemma mod_eq_imp_eq: "\b \ x; x < b + N; b \ y; y < b + N; x mod N = y mod N \ \ x=y" proof - assume a1: "b \ y" assume a2: "y < b + N" assume a3: "b \ x" assume a4: "x < b + N" assume a5: "x mod N = y mod N" have f6: "x < y + N" using a4 a1 by linarith have "y < x + N" using a3 a2 by linarith then show ?thesis using f6 a5 by (metis (no_types) mod_eq_imp_eq_aux nat_le_linear) qed text \Map the ticket of a thread\ fun map_ticket where "map_ticket f INIT = INIT" | "map_ticket f (WAIT tk) = WAIT (f tk)" | "map_ticket f (HOLD tk) = HOLD (f tk)" | "map_ticket f (REL tk) = REL (f tk)" lemma map_ticket_addsimps[simp]: "map_ticket f t = INIT \ t=INIT" "map_ticket f t = WAIT tk \ (\tk'. tk=f tk' \ t=WAIT tk')" "map_ticket f t = HOLD tk \ (\tk'. tk=f tk' \ t=HOLD tk')" "map_ticket f t = REL tk \ (\tk'. tk=f tk' \ t=REL tk')" by (cases t; auto)+ text \We define the number of threads that use a ticket\ fun ni_weight :: "thread \ nat" where "ni_weight INIT = 0" | "ni_weight _ = 1" lemma ni_weight_le1[simp]: "ni_weight s \ Suc 0" by (cases s) auto definition "num_ni ts \ \ i=0.._. INIT) = 0" by (auto simp: num_ni_def) lemma num_ni_upd: "t num_ni (ts(t:=s)) = num_ni ts - ni_weight (ts t) + ni_weight s" by (auto simp: num_ni_def if_distrib[of ni_weight] sum.If_cases algebra_simps simp: sum_diff1_nat ) lemma num_ni_nz_if[simp]: "\t < N; ts t \ INIT\ \ num_ni ts \ 0" apply (cases "ts t") by (simp_all add: num_ni_def) force+ lemma num_ni_leN: "num_ni ts \ N" apply (clarsimp simp: num_ni_def) apply (rule order_trans) apply (rule sum_bounded_above[where K=1]) apply auto done text \We provide an additional invariant, considering the distance of \c\ and \n\. Although we could probably get this from the previous invariants, it is easy enough to prove directly. \ definition "invar3 \ \(c,n,ts). n = c + num_ni ts" lemma is_invar3: "A.is_invar invar3" apply (rule) subgoal by (auto simp: invar3_def as\<^sub>0_def) subgoal for s s' apply clarify apply (erule alstep.cases, erule astep_sng.cases) apply (auto simp: invar3_def num_ni_upd) using holds_only_current by fastforce done text \We establish a simulation relation: The concrete counters are the abstract ones, wrapped around. \ definition "sim_rel1 \ \(c,n,ts) (ci,ni,tsi). ci = c mod N \ ni = n mod (k*N) \ tsi = (map_ticket (\t. t mod N)) o ts " lemma sraux: "sim_rel1 (c,n,ts) (ci,ni,tsi) \ ci = c mod N \ ni = n mod (k*N)" by (auto simp: sim_rel1_def Let_def) lemma sraux2: "\sim_rel1 (c,n,ts) (ci,ni,tsi); t \ tsi t = map_ticket (\x. x mod N) (ts t)" by (auto simp: sim_rel1_def Let_def) interpretation sim1: simulationI as\<^sub>0 alstep bs\<^sub>0 blstep sim_rel1 proof unfold_locales show "sim_rel1 as\<^sub>0 bs\<^sub>0" by (auto simp: sim_rel1_def as\<^sub>0_def bs\<^sub>0_def) next fix as bs t bs' assume Ra_aux: "A.reachable as" and Rc_aux: "B.reachable bs" and SIM: "sim_rel1 as bs" and CS: "blstep t bs bs'" obtain c n ts where [simp]: "as=(c,n,ts)" by (cases as) obtain ci ni tsi where [simp]: "bs=(ci,ni,tsi)" by (cases bs) obtain ci' ni' tsi' where [simp]: "bs'=(ci',ni',tsi')" by (cases bs') from Ra_aux have Ra: "A.reachable (c,n,ts)" by simp from Rc_aux have Rc: "B.reachable (ci,ni,tsi)" by simp from CS have "ttk" "tkt apply fastforce using \t num_ni_leN[of ts] by fastforce from SIM CS have "\as'. alstep t as as' \ sim_rel1 as' bs'" apply simp apply (erule blstep.cases) apply (erule bstep_sng.cases) apply clarsimp_all subgoal apply (intro exI conjI) apply (rule alstep.intros) apply (simp add: sim_rel1_def Let_def) apply (simp add: sraux sraux2) apply (rule astep_sng.enter_wait) apply (simp add: sim_rel1_def; intro conjI ext) apply (auto simp: sim_rel1_def Let_def mod_simps) done subgoal apply (clarsimp simp: sraux sraux2) apply (intro exI conjI) apply (rule alstep.intros) apply (simp add: sim_rel1_def Let_def) apply clarsimp apply (rule astep_sng.loop_wait) apply (auto simp: sim_rel1_def Let_def mod_simps) done subgoal apply (clarsimp simp: sraux sraux2) subgoal for tk' apply (subgoal_tac "tk'=c") apply (intro exI conjI) apply (rule alstep.intros) apply (simp add: sim_rel1_def Let_def) apply clarsimp apply (rule astep_sng.exit_wait) apply (auto simp: sim_rel1_def Let_def mod_simps) [] apply (clarsimp simp: sim_rel1_def) apply (erule mod_eq_imp_eq_aux) apply (auto simp: AUX1) done done subgoal apply (clarsimp simp: sraux sraux2) apply (intro exI conjI) apply (rule alstep.intros) apply (simp add: sim_rel1_def Let_def) apply clarsimp apply (rule astep_sng.start_release) apply (auto simp: sim_rel1_def Let_def mod_simps) done subgoal apply (clarsimp simp: sraux sraux2) apply (intro exI conjI) apply (rule alstep.intros) apply (simp add: sim_rel1_def Let_def) apply clarsimp apply (rule astep_sng.release) apply (auto simp: sim_rel1_def Let_def mod_simps) done done then show "\as'. sim_rel1 as' bs' \ alstep t as as'" by blast next fix as bs l assume "A.reachable as" "B.reachable bs" "sim_rel1 as bs" "A.can_step l as" then show "B.can_step l bs" using b_never_blocked never_blocked by simp qed subsubsection \Transfer of Properties\ text \We transfer a few properties over the simulation, which we need for the next refinement step. \ lemma xfer_locks_ticket: assumes "locks_ticket (map_ticket (\t. t mod N) (ts t)) tki" obtains tk where "tki=tk mod N" "locks_ticket (ts t) tk" using assms unfolding locks_ticket_def by auto lemma b_holds_only_current: "\B.reachable (c, n, ts); t < N; locks_ticket (ts t) tk\ \ tk = c" apply (rule sim1.xfer_reachable, assumption) apply (clarsimp simp: sim_rel1_def) apply (erule xfer_locks_ticket)+ using holds_only_current by blast lemma b_mutual_exclusion': "\B.reachable s; tt'; locks_ticket (tts s t) tk; locks_ticket (tts s t') tk' \ \ False" apply (rule sim1.xfer_reachable, assumption) apply (clarsimp simp: sim_rel1_def) apply (erule xfer_locks_ticket)+ apply (drule (3) mutual_exclusion'; simp) done lemma xfer_has_ticket: assumes "has_ticket (map_ticket (\t. t mod N) (ts t)) tki" obtains tk where "tki=tk mod N" "has_ticket (ts t) tk" using assms unfolding has_ticket_def by auto lemma has_ticket_in_range: assumes Ra: "A.reachable (c,n,ts)" and "ttk \ tktk \ tkt apply fastforce using \t num_ni_leN[of ts] by fastforce qed lemma b_has_ticket_unique: "\B.reachable (ci,ni,tsi); t \ t'=t" apply (rule sim1.xfer_reachable, assumption) apply (auto simp: sim_rel1_def) subgoal for c n ts apply (erule xfer_has_ticket)+ apply simp subgoal for tk tk' apply (subgoal_tac "tk=tk'") apply simp apply (frule (4) has_ticket_unique, assumption) apply (frule (2) has_ticket_in_range[where tk=tk]) apply (frule (2) has_ticket_in_range[where tk=tk']) apply (auto simp: mod_simps) apply (rule mod_eq_imp_eq; (assumption|simp)) done done done subsection \Refinement 3: Using an Array\text_raw \\label{sec:refine3}\ text \Finally, we use an array instead of a counter, thus obtaining the exact data structures from the challenge assignment. Note that we model the array by a list of Booleans here. \ text \System's state: Current ticket array, next ticket, thread states\ type_synonym cstate = "bool list \ nat \ (nat \ thread)" text \The step relation of a single thread\ inductive cstep_sng where enter_wait: "cstep_sng (p,n,INIT) (p,(n+1) mod (k*N),WAIT (n mod N))" | loop_wait: "\p!tk \ cstep_sng (p,n,WAIT tk) (p,n,WAIT tk)" | exit_wait: "p!tk \ cstep_sng (p,n,WAIT tk) (p,n,HOLD tk)" | start_release: "cstep_sng (p,n,HOLD tk) (p[tk:=False],n,REL tk)" | release: "cstep_sng (p,n,REL tk) (p[(tk+1) mod N := True],n,INIT)" text \The step relation of the system, labeled with the thread \t\ that performs the step\ inductive clstep for t where "\ t \ clstep t (c,n,ts) (c',n',ts(t:=s'))" text \Initial state of the system\ definition "cs\<^sub>0 \ ((replicate N False)[0:=True], 0, \_. INIT)" interpretation C: system cs\<^sub>0 clstep . lemma c_never_blocked: "C.can_step l s \ l0 clstep apply unfold_locales subgoal for s using c_never_blocked[of 0 s] unfolding C.can_step_def by auto done text \We establish another invariant that states that the ticket numbers are bounded.\ definition "invar4 \ \(c,n,ts). c (\ttk. has_ticket (ts t) tk \ tk0_def) subgoal for s s' apply clarify apply (erule blstep.cases, erule bstep_sng.cases) unfolding invar4_def - apply safe - apply (metis N_gt0 fun_upd_other fun_upd_same mod_mod_trivial - nat_mod_lem has_ticket_simps(2)) + apply safe + apply (metis N_gt0 fun_upd_apply has_ticket_simps(2) mod_less_divisor) apply (metis fun_upd_triv) apply (metis fun_upd_other fun_upd_same has_ticket_simps(3)) apply (metis fun_upd_other fun_upd_same has_ticket_def has_ticket_simps(4)) using mod_less_divisor apply blast apply (metis fun_upd_apply thread.distinct(1) thread.distinct(3) thread.distinct(5) has_ticket_def) done done text \We define a predicate that describes that a thread of the system is at the release sequence point --- in this case, the array does not have a set bit, otherwise, the set bit corresponds to the current ticket.\ definition "is_REL_state \ \ts. \ttk. ts t = REL tk" lemma is_REL_state_simps[simp]: "t is_REL_state (ts(t:=REL tk))" "t \is_REL (ts t) \ \is_REL s' \ is_REL_state (ts(t:=s')) \ is_REL_state ts" unfolding is_REL_state_def apply (auto; fail) [] apply auto [] by (metis thread.disc(12)) lemma is_REL_state_aux1: assumes R: "B.reachable (c,n,ts)" assumes REL: "is_REL_state ts" assumes "tc" using REL unfolding is_REL_state_def apply clarify subgoal for t' tk' using b_has_ticket_unique[OF R \t, of tk t'] using b_holds_only_current[OF R, of t' tk'] by (auto) done lemma is_REL_state_aux2: assumes R: "B.reachable (c,n,ts)" assumes A: "tis_REL_state (ts(t:=INIT))" using b_holds_only_current[OF R] A using b_mutual_exclusion'[OF R] apply (clarsimp simp: is_REL_state_def) by fastforce text \Simulation relation that implements current ticket by array\ definition "sim_rel2 \ \(c,n,ts) (ci,ni,tsi). (if is_REL_state ts then ci = replicate N False else ci = (replicate N False)[c:=True] ) \ ni = n \ tsi = ts " interpretation sim2: simulationI bs\<^sub>0 blstep cs\<^sub>0 clstep sim_rel2 proof unfold_locales show "sim_rel2 bs\<^sub>0 cs\<^sub>0" by (auto simp: sim_rel2_def bs\<^sub>0_def cs\<^sub>0_def is_REL_state_def) next fix bs cs t cs' assume Rc_aux: "B.reachable bs" and Rd_aux: "C.reachable cs" and SIM: "sim_rel2 bs cs" and CS: "clstep t cs cs'" obtain c n ts where [simp]: "bs=(c,n,ts)" by (cases bs) obtain ci ni tsi where [simp]: "cs=(ci,ni,tsi)" by (cases cs) obtain ci' ni' tsi' where [simp]: "cs'=(ci',ni',tsi')" by (cases cs') from Rc_aux have Rc: "B.reachable (c,n,ts)" by simp from Rd_aux have Rd: "C.reachable (ci,ni,tsi)" by simp from CS have "tt by (auto simp: invar4_def) have HOLD_AUX: "tk=c" if "ts t = HOLD tk" for tk using b_holds_only_current[OF Rc \t, of tk] that by auto have REL_AUX: "tk=c" if "ts t = REL tk" "tt, of tk] that by auto have [simp]: "c (replicate N False)[c := True]" "(replicate N False)[c := True] \ replicate N False" apply (auto simp: list_eq_iff_nth_eq nth_list_update) using \c < N\ by blast+ have [simp]: "(replicate N False)[c := True] ! d \ d=c" if "dbs'. blstep t bs bs' \ sim_rel2 bs' cs'" apply simp apply (subst (asm) sim_rel2_def) apply (erule clstep.cases) apply (erule cstep_sng.cases) apply clarsimp_all subgoal apply (intro exI conjI) apply (rule blstep.intros) apply (simp) apply clarsimp apply (rule bstep_sng.enter_wait) apply (auto simp: sim_rel2_def split: if_splits) done subgoal for tk' apply (intro exI conjI) apply (rule blstep.intros) apply (simp) apply clarsimp apply (rule bstep_sng.loop_wait) subgoal apply (clarsimp simp: sim_rel2_def split: if_splits) apply (frule (2) is_REL_state_aux1[OF Rc]) by simp subgoal by (auto simp: sim_rel2_def split: if_splits) done subgoal apply (intro exI conjI) apply (rule blstep.intros) apply (simp) apply (clarsimp split: if_splits) apply (rule bstep_sng.exit_wait) apply (auto simp: sim_rel2_def split: if_splits) done subgoal apply (intro exI conjI) apply (rule blstep.intros) apply (simp) apply clarsimp apply (rule bstep_sng.start_release) apply (auto simp: sim_rel2_def dest: HOLD_AUX split: if_splits) done subgoal apply (intro exI conjI) apply (rule blstep.intros) apply (simp) apply clarsimp apply (rule bstep_sng.release) apply (auto simp: sim_rel2_def dest: is_REL_state_aux2[OF Rc] split: if_splits) by (metis fun_upd_triv is_REL_state_simps(1)) done then show "\bs'. sim_rel2 bs' cs' \ blstep t bs bs'" by blast next fix bs cs l assume "B.reachable bs" "C.reachable cs" "sim_rel2 bs cs" "B.can_step l bs" then show "C.can_step l cs" using c_never_blocked b_never_blocked by simp qed subsection \Transfer Setup\ text \We set up the final simulation relation, and the transfer of the concepts used in the correctness statements. \ definition "sim_rel \ sim_rel1 OO sim_rel2" interpretation sim: simulation as\<^sub>0 alstep cs\<^sub>0 clstep sim_rel unfolding sim_rel_def by (rule sim_trans) unfold_locales lemma xfer_holds: assumes "sim_rel s cs" shows "is_HOLD (tts cs t) \ is_HOLD (tts s t)" using assms unfolding sim_rel_def sim_rel1_def sim_rel2_def by (cases "tts cs t") auto lemma xfer_waits: assumes "sim_rel s cs" shows "is_WAIT (tts cs t) \ is_WAIT (tts s t)" using assms unfolding sim_rel_def sim_rel1_def sim_rel2_def by (cases "tts cs t") auto lemma xfer_init: assumes "sim_rel s cs" shows "tts cs t = INIT \ tts s t = INIT" using assms unfolding sim_rel_def sim_rel1_def sim_rel2_def by auto subsection \Main Theorems\ text_raw \\label{sec:main_theorems}\ subsubsection \Trusted Code Base\ text \ Note that the trusted code base for these theorems is only the formalization of the concrete system as defined in Section~\ref{sec:refine3}. The simulation setup and the abstract systems are only auxiliary constructions for the proof. For completeness, we display the relevant definitions of reachability, runs, and fairness here: @{lemma [display, source] "C.step s s' = (\l. clstep l s s')" by simp} @{thm [display] C.reachable_def[no_vars] C.run_definitions[no_vars]} \ subsubsection \Safety\ text \We show that there is no reachable state in which two different threads hold the lock.\ (* ALWAYS (NOT (HOLD ** HOLD)) *) theorem final_mutual_exclusion: "\C.reachable s; tt'; is_HOLD (tts s t); is_HOLD (tts s t') \ \ False" apply (erule sim.xfer_reachable) apply (simp add: xfer_holds) by (erule (5) mutual_exclusion) subsubsection \Fairness\ text \We show that, whenever a thread \t\ draws a ticket, all other threads \t'\ waiting for the lock will be granted the lock before \t\. \ (* ALWAYS (INIT t\<^sub>1 \ WAIT t\<^sub>2 \ X (WAIT t\<^sub>1) \ HOLD t\<^sub>2 BEFORE HOLD t\<^sub>1 ) *) theorem final_fair: assumes RUN: "C.is_run s" assumes ACQ: "t \Thread \t\ draws ticket in step \i\\ assumes HOLD: "i \Thread \t\ holds lock in step \j\\ assumes WAIT: "t' \Thread \t'\ waits for lock at step \i\\ obtains l where "i \Then, \t'\ gets lock earlier\ using RUN proof (rule sim.xfer_run) fix as assume Ra: "A.is_run as" and SIM[rule_format]: "\i. sim_rel (as i) (s i)" note XFER = xfer_init[OF SIM] xfer_holds[OF SIM] xfer_waits[OF SIM] show ?thesis using assms apply (simp add: XFER) apply (erule (6) fair[OF Ra]) apply (erule (1) that) apply (simp add: XFER) done qed subsubsection \Liveness\ text \We show that, for a fair run, every thread that waits for the lock will eventually hold the lock.\ (* ALWAYS (WAIT \ EVENTUALLY HOLD) *) theorem final_progress: assumes FRUN: "C.is_fair_run s" assumes WAIT: "tj>i. is_HOLD (tts (s j) t)" using FRUN proof (rule sim.xfer_fair_run) fix as assume Ra: "A.is_fair_run as" and SIM[rule_format]: "\i. sim_rel (as i) (s i)" note XFER = xfer_init[OF SIM] xfer_holds[OF SIM] xfer_waits[OF SIM] show ?thesis using assms apply (simp add: XFER) apply (erule (1) progress[OF Ra]) done qed end diff --git a/thys/Word_Lib/Bitwise.thy b/thys/Word_Lib/Bitwise.thy --- a/thys/Word_Lib/Bitwise.thy +++ b/thys/Word_Lib/Bitwise.thy @@ -1,525 +1,525 @@ (* Authors: Thomas Sewell, NICTA and Sascha Boehme, TU Muenchen *) theory Bitwise imports "HOL-Word.Word" "HOL-Word.Misc_Arithmetic" begin text \Helper constants used in defining addition\ definition xor3 :: "bool \ bool \ bool \ bool" where "xor3 a b c = (a = (b = c))" definition carry :: "bool \ bool \ bool \ bool" where "carry a b c = ((a \ (b \ c)) \ (b \ c))" lemma carry_simps: "carry True a b = (a \ b)" "carry a True b = (a \ b)" "carry a b True = (a \ b)" "carry False a b = (a \ b)" "carry a False b = (a \ b)" "carry a b False = (a \ b)" by (auto simp add: carry_def) lemma xor3_simps: "xor3 True a b = (a = b)" "xor3 a True b = (a = b)" "xor3 a b True = (a = b)" "xor3 False a b = (a \ b)" "xor3 a False b = (a \ b)" "xor3 a b False = (a \ b)" by (simp_all add: xor3_def) text \Breaking up word equalities into equalities on their bit lists. Equalities are generated and manipulated in the reverse order to \<^const>\to_bl\.\ lemma bl_word_sub: "to_bl (x - y) = to_bl (x + (- y))" by simp lemma rbl_word_1: "rev (to_bl (1 :: 'a::len word)) = takefill False (LENGTH('a)) [True]" apply (rule_tac s="rev (to_bl (word_succ (0 :: 'a word)))" in trans) apply simp apply (simp only: rtb_rbl_ariths(1)[OF refl]) apply simp apply (case_tac "LENGTH('a)") apply simp apply (simp add: takefill_alt) done lemma rbl_word_if: "rev (to_bl (if P then x else y)) = map2 (If P) (rev (to_bl x)) (rev (to_bl y))" by (simp add: split_def) lemma rbl_add_carry_Cons: "(if car then rbl_succ else id) (rbl_add (x # xs) (y # ys)) = xor3 x y car # (if carry x y car then rbl_succ else id) (rbl_add xs ys)" by (simp add: carry_def xor3_def) lemma rbl_add_suc_carry_fold: "length xs = length ys \ \car. (if car then rbl_succ else id) (rbl_add xs ys) = (foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\_. [])) car" apply (erule list_induct2) apply simp apply (simp only: rbl_add_carry_Cons) apply simp done lemma to_bl_plus_carry: "to_bl (x + y) = rev (foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (rev (zip (to_bl x) (to_bl y))) (\_. []) False)" using rbl_add_suc_carry_fold[where xs="rev (to_bl x)" and ys="rev (to_bl y)"] apply (simp add: word_add_rbl[OF refl refl]) apply (drule_tac x=False in spec) apply (simp add: zip_rev) done definition "rbl_plus cin xs ys = foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\_. []) cin" lemma rbl_plus_simps: "rbl_plus cin (x # xs) (y # ys) = xor3 x y cin # rbl_plus (carry x y cin) xs ys" "rbl_plus cin [] ys = []" "rbl_plus cin xs [] = []" by (simp_all add: rbl_plus_def) lemma rbl_word_plus: "rev (to_bl (x + y)) = rbl_plus False (rev (to_bl x)) (rev (to_bl y))" by (simp add: rbl_plus_def to_bl_plus_carry zip_rev) definition "rbl_succ2 b xs = (if b then rbl_succ xs else xs)" lemma rbl_succ2_simps: "rbl_succ2 b [] = []" "rbl_succ2 b (x # xs) = (b \ x) # rbl_succ2 (x \ b) xs" by (simp_all add: rbl_succ2_def) lemma twos_complement: "- x = word_succ (NOT x)" using arg_cong[OF word_add_not[where x=x], where f="\a. a - x + 1"] by (simp add: word_succ_p1 word_sp_01[unfolded word_succ_p1] del: word_add_not) lemma rbl_word_neg: "rev (to_bl (- x)) = rbl_succ2 True (map Not (rev (to_bl x)))" for x :: \'a::len word\ by (simp add: twos_complement word_succ_rbl[OF refl] bl_word_not rev_map rbl_succ2_def) lemma rbl_word_cat: "rev (to_bl (word_cat x y :: 'a::len word)) = takefill False (LENGTH('a)) (rev (to_bl y) @ rev (to_bl x))" by (simp add: word_cat_bl word_rev_tf) lemma rbl_word_slice: "rev (to_bl (slice n w :: 'a::len word)) = takefill False (LENGTH('a)) (drop n (rev (to_bl w)))" apply (simp add: slice_take word_rev_tf rev_take) apply (cases "n < LENGTH('b)", simp_all) done lemma rbl_word_ucast: "rev (to_bl (ucast x :: 'a::len word)) = takefill False (LENGTH('a)) (rev (to_bl x))" apply (simp add: to_bl_ucast takefill_alt) apply (simp add: rev_drop) apply (cases "LENGTH('a) < LENGTH('b)") apply simp_all done lemma rbl_shiftl: "rev (to_bl (w << n)) = takefill False (size w) (replicate n False @ rev (to_bl w))" by (simp add: bl_shiftl takefill_alt word_size rev_drop) lemma rbl_shiftr: "rev (to_bl (w >> n)) = takefill False (size w) (drop n (rev (to_bl w)))" by (simp add: shiftr_slice rbl_word_slice word_size) definition "drop_nonempty v n xs = (if n < length xs then drop n xs else [last (v # xs)])" lemma drop_nonempty_simps: "drop_nonempty v (Suc n) (x # xs) = drop_nonempty x n xs" "drop_nonempty v 0 (x # xs) = (x # xs)" "drop_nonempty v n [] = [v]" by (simp_all add: drop_nonempty_def) definition "takefill_last x n xs = takefill (last (x # xs)) n xs" lemma takefill_last_simps: "takefill_last z (Suc n) (x # xs) = x # takefill_last x n xs" "takefill_last z 0 xs = []" "takefill_last z n [] = replicate n z" by (simp_all add: takefill_last_def) (simp_all add: takefill_alt) lemma rbl_sshiftr: "rev (to_bl (w >>> n)) = takefill_last False (size w) (drop_nonempty False n (rev (to_bl w)))" apply (cases "n < size w") apply (simp add: bl_sshiftr takefill_last_def word_size takefill_alt rev_take last_rev drop_nonempty_def) apply (subgoal_tac "(w >>> n) = of_bl (replicate (size w) (msb w))") apply (simp add: word_size takefill_last_def takefill_alt last_rev word_msb_alt word_rev_tf drop_nonempty_def take_Cons') apply (case_tac "LENGTH('a)", simp_all) apply (rule word_eqI) apply (simp add: nth_sshiftr word_size test_bit_of_bl msb_nth) done lemma nth_word_of_int: "(word_of_int x :: 'a::len word) !! n = (n < LENGTH('a) \ bin_nth x n)" apply (simp add: test_bit_bl word_size to_bl_of_bin) apply (subst conj_cong[OF refl], erule bin_nth_bl) apply auto done lemma nth_scast: "(scast (x :: 'a::len word) :: 'b::len word) !! n = (n < LENGTH('b) \ (if n < LENGTH('a) - 1 then x !! n else x !! (LENGTH('a) - 1)))" by (simp add: scast_def nth_sint) lemma rbl_word_scast: "rev (to_bl (scast x :: 'a::len word)) = takefill_last False (LENGTH('a)) (rev (to_bl x))" apply (rule nth_equalityI) apply (simp add: word_size takefill_last_def) apply (clarsimp simp: nth_scast takefill_last_def - nth_takefill word_size nth_rev to_bl_nth) + nth_takefill word_size rev_nth to_bl_nth) apply (cases "LENGTH('b)") apply simp apply (clarsimp simp: less_Suc_eq_le linorder_not_less last_rev word_msb_alt[symmetric] msb_nth) done definition rbl_mul :: "bool list \ bool list \ bool list" where "rbl_mul xs ys = foldr (\x sm. rbl_plus False (map ((\) x) ys) (False # sm)) xs []" lemma rbl_mul_simps: "rbl_mul (x # xs) ys = rbl_plus False (map ((\) x) ys) (False # rbl_mul xs ys)" "rbl_mul [] ys = []" by (simp_all add: rbl_mul_def) lemma takefill_le2: "length xs \ n \ takefill x m (takefill x n xs) = takefill x m xs" by (simp add: takefill_alt replicate_add[symmetric]) lemma take_rbl_plus: "\n b. take n (rbl_plus b xs ys) = rbl_plus b (take n xs) (take n ys)" apply (simp add: rbl_plus_def take_zip[symmetric]) apply (rule_tac list="zip xs ys" in list.induct) apply simp apply (clarsimp simp: split_def) apply (case_tac n, simp_all) done lemma word_rbl_mul_induct: "length xs \ size y \ rbl_mul xs (rev (to_bl y)) = take (length xs) (rev (to_bl (of_bl (rev xs) * y)))" for y :: "'a::len word" proof (induct xs) case Nil show ?case by (simp add: rbl_mul_simps) next case (Cons z zs) have rbl_word_plus': "to_bl (x + y) = rev (rbl_plus False (rev (to_bl x)) (rev (to_bl y)))" for x y :: "'a word" by (simp add: rbl_word_plus[symmetric]) have mult_bit: "to_bl (of_bl [z] * y) = map ((\) z) (to_bl y)" by (cases z) (simp cong: map_cong, simp add: map_replicate_const cong: map_cong) have shiftl: "of_bl xs * 2 * y = (of_bl xs * y) << 1" for xs by (simp add: shiftl_t2n) have zip_take_triv: "\xs ys n. n = length ys \ zip (take n xs) ys = zip xs ys" by (rule nth_equalityI) simp_all from Cons show ?case apply (simp add: trans [OF of_bl_append add.commute] rbl_mul_simps rbl_word_plus' distrib_right mult_bit shiftl rbl_shiftl) apply (simp add: takefill_alt word_size rev_map take_rbl_plus min_def) apply (simp add: rbl_plus_def zip_take_triv) done qed lemma rbl_word_mul: "rev (to_bl (x * y)) = rbl_mul (rev (to_bl x)) (rev (to_bl y))" for x :: "'a::len word" using word_rbl_mul_induct[where xs="rev (to_bl x)" and y=y] by (simp add: word_size) text \Breaking up inequalities into bitlist properties.\ definition "rev_bl_order F xs ys = (length xs = length ys \ ((xs = ys \ F) \ (\n < length xs. drop (Suc n) xs = drop (Suc n) ys \ \ xs ! n \ ys ! n)))" lemma rev_bl_order_simps: "rev_bl_order F [] [] = F" "rev_bl_order F (x # xs) (y # ys) = rev_bl_order ((y \ \ x) \ ((y \ \ x) \ F)) xs ys" apply (simp_all add: rev_bl_order_def) apply (rule conj_cong[OF refl]) apply (cases "xs = ys") apply (simp add: nth_Cons') apply blast apply (simp add: nth_Cons') apply safe apply (rule_tac x="n - 1" in exI) apply simp apply (rule_tac x="Suc n" in exI) apply simp done lemma rev_bl_order_rev_simp: "length xs = length ys \ rev_bl_order F (xs @ [x]) (ys @ [y]) = ((y \ \ x) \ ((y \ \ x) \ rev_bl_order F xs ys))" by (induct arbitrary: F rule: list_induct2) (auto simp: rev_bl_order_simps) lemma rev_bl_order_bl_to_bin: "length xs = length ys \ rev_bl_order True xs ys = (bl_to_bin (rev xs) \ bl_to_bin (rev ys)) \ rev_bl_order False xs ys = (bl_to_bin (rev xs) < bl_to_bin (rev ys))" apply (induct xs ys rule: list_induct2) apply (simp_all add: rev_bl_order_simps bl_to_bin_app_cat) apply (auto simp add: bl_to_bin_def add1_zle_eq) done lemma word_le_rbl: "x \ y \ rev_bl_order True (rev (to_bl x)) (rev (to_bl y))" for x y :: "'a::len word" by (simp add: rev_bl_order_bl_to_bin word_le_def) lemma word_less_rbl: "x < y \ rev_bl_order False (rev (to_bl x)) (rev (to_bl y))" for x y :: "'a::len word" by (simp add: word_less_alt rev_bl_order_bl_to_bin) lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases "msb x") apply (rule word_sint.Abs_eqD[where 'a='a], simp_all) apply (simp add: word_size wi_hom_syms word_of_int_2p_len) apply (simp add: sints_num word_size) apply (rule conjI) apply (simp add: le_diff_eq') apply (rule order_trans[where y="2 ^ (LENGTH('a) - 1)"]) apply (simp add: power_Suc[symmetric]) apply (simp add: linorder_not_less[symmetric] mask_eq_iff[symmetric]) apply (rule notI, drule word_eqD[where x="size x - 1"]) apply (simp add: msb_nth word_ops_nth_size word_size) apply (simp add: order_less_le_trans[where y=0]) apply (rule word_uint.Abs_eqD[where 'a='a], simp_all) apply (simp add: linorder_not_less uints_num word_msb_sint) apply (rule order_less_le_trans[OF sint_lt]) apply simp done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" apply (simp add: word_sle_def word_sint_msb_eq word_size word_le_def) apply safe apply (rule order_trans[OF _ uint_ge_0]) apply (simp add: order_less_imp_le) apply (erule notE[OF leD]) apply (rule order_less_le_trans[OF _ uint_ge_0]) apply simp done lemma word_sless_msb_less: "x (msb y \ msb x) \ ((msb x \ \ msb y) \ x < y)" by (auto simp add: word_sless_def word_sle_msb_le) definition "map_last f xs = (if xs = [] then [] else butlast xs @ [f (last xs)])" lemma map_last_simps: "map_last f [] = []" "map_last f [x] = [f x]" "map_last f (x # y # zs) = x # map_last f (y # zs)" by (simp_all add: map_last_def) lemma word_sle_rbl: "x <=s y \ rev_bl_order True (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))" using word_msb_alt[where w=x] word_msb_alt[where w=y] apply (simp add: word_sle_msb_le word_le_rbl) apply (subgoal_tac "length (to_bl x) = length (to_bl y)") apply (cases "to_bl x", simp) apply (cases "to_bl y", simp) apply (clarsimp simp: map_last_def rev_bl_order_rev_simp) apply auto done lemma word_sless_rbl: "x rev_bl_order False (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))" using word_msb_alt[where w=x] word_msb_alt[where w=y] apply (simp add: word_sless_msb_less word_less_rbl) apply (subgoal_tac "length (to_bl x) = length (to_bl y)") apply (cases "to_bl x", simp) apply (cases "to_bl y", simp) apply (clarsimp simp: map_last_def rev_bl_order_rev_simp) apply auto done text \Lemmas for unpacking \<^term>\rev (to_bl n)\ for numerals n and also for irreducible values and expressions.\ lemma rev_bin_to_bl_simps: "rev (bin_to_bl 0 x) = []" "rev (bin_to_bl (Suc n) (numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (numeral nm))" "rev (bin_to_bl (Suc n) (numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (numeral nm))" "rev (bin_to_bl (Suc n) (numeral (num.One))) = True # replicate n False" "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (- numeral nm))" "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.One))) = True # replicate n True" "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm + num.One))) = True # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm + num.One))) = False # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.One + num.One))) = False # rev (bin_to_bl n (- numeral num.One))" by (simp_all add: bin_to_bl_aux_append bin_to_bl_zero_aux bin_to_bl_minus1_aux replicate_append_same) lemma to_bl_upt: "to_bl x = rev (map ((!!) x) [0 ..< size x])" apply (rule nth_equalityI) apply (simp add: word_size) - apply (auto simp: to_bl_nth word_size nth_rev) + apply (auto simp: to_bl_nth word_size rev_nth) done lemma rev_to_bl_upt: "rev (to_bl x) = map ((!!) x) [0 ..< size x]" by (simp add: to_bl_upt) lemma upt_eq_list_intros: "j \ i \ [i ..< j] = []" "i = x \ x < j \ [x + 1 ..< j] = xs \ [i ..< j] = (x # xs)" by (simp_all add: upt_eq_Cons_conv) subsection \Tactic definition\ ML \ structure Word_Bitwise_Tac = struct val word_ss = simpset_of \<^theory_context>\Word\; fun mk_nat_clist ns = fold_rev (Thm.mk_binop \<^cterm>\Cons :: nat \ _\) ns \<^cterm>\[] :: nat list\; fun upt_conv ctxt ct = case Thm.term_of ct of (\<^const>\upt\ $ n $ m) => let val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m); val ns = map (Numeral.mk_cnumber \<^ctyp>\nat\) (i upto (j - 1)) |> mk_nat_clist; val prop = Thm.mk_binop \<^cterm>\(=) :: nat list \ _\ ct ns |> Thm.apply \<^cterm>\Trueprop\; in try (fn () => Goal.prove_internal ctxt [] prop (K (REPEAT_DETERM (resolve_tac ctxt @{thms upt_eq_list_intros} 1 ORELSE simp_tac (put_simpset word_ss ctxt) 1))) |> mk_meta_eq) () end | _ => NONE; val expand_upt_simproc = Simplifier.make_simproc \<^context> "expand_upt" {lhss = [\<^term>\upt x y\], proc = K upt_conv}; fun word_len_simproc_fn ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\len_of\, _) $ t => (let val T = fastype_of t |> dest_Type |> snd |> the_single val n = Numeral.mk_cnumber \<^ctyp>\nat\ (Word_Lib.dest_binT T); val prop = Thm.mk_binop \<^cterm>\(=) :: nat \ _\ ct n |> Thm.apply \<^cterm>\Trueprop\; in Goal.prove_internal ctxt [] prop (K (simp_tac (put_simpset word_ss ctxt) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE | TYPE _ => NONE) | _ => NONE); val word_len_simproc = Simplifier.make_simproc \<^context> "word_len" {lhss = [\<^term>\len_of x\], proc = K word_len_simproc_fn}; (* convert 5 or nat 5 to Suc 4 when n_sucs = 1, Suc (Suc 4) when n_sucs = 2, or just 5 (discarding nat) when n_sucs = 0 *) fun nat_get_Suc_simproc_fn n_sucs ctxt ct = let val (f $ arg) = Thm.term_of ct; val n = (case arg of \<^term>\nat\ $ n => n | n => n) |> HOLogic.dest_number |> snd; val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs) else (n, 0); val arg' = funpow i HOLogic.mk_Suc (HOLogic.mk_number \<^typ>\nat\ j); val _ = if arg = arg' then raise TERM ("", []) else (); fun propfn g = HOLogic.mk_eq (g arg, g arg') |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt; val eq1 = Goal.prove_internal ctxt [] (propfn I) (K (simp_tac (put_simpset word_ss ctxt) 1)); in Goal.prove_internal ctxt [] (propfn (curry (op $) f)) (K (simp_tac (put_simpset HOL_ss ctxt addsimps [eq1]) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE; fun nat_get_Suc_simproc n_sucs ts = Simplifier.make_simproc \<^context> "nat_get_Suc" {lhss = map (fn t => t $ \<^term>\n :: nat\) ts, proc = K (nat_get_Suc_simproc_fn n_sucs)}; val no_split_ss = simpset_of (put_simpset HOL_ss \<^context> |> Splitter.del_split @{thm if_split}); val expand_word_eq_sss = (simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms word_eq_rbl_eq word_le_rbl word_less_rbl word_sle_rbl word_sless_rbl}), map simpset_of [ put_simpset no_split_ss \<^context> addsimps @{thms rbl_word_plus rbl_word_and rbl_word_or rbl_word_not rbl_word_neg bl_word_sub rbl_word_xor rbl_word_cat rbl_word_slice rbl_word_scast rbl_word_ucast rbl_shiftl rbl_shiftr rbl_sshiftr rbl_word_if}, put_simpset no_split_ss \<^context> addsimps @{thms to_bl_numeral to_bl_neg_numeral to_bl_0 rbl_word_1}, put_simpset no_split_ss \<^context> addsimps @{thms rev_rev_ident rev_replicate rev_map to_bl_upt word_size} addsimprocs [word_len_simproc], put_simpset no_split_ss \<^context> addsimps @{thms list.simps split_conv replicate.simps list.map zip_Cons_Cons zip_Nil drop_Suc_Cons drop_0 drop_Nil foldr.simps list.map zip.simps(1) zip_Nil zip_Cons_Cons takefill_Suc_Cons takefill_Suc_Nil takefill.Z rbl_succ2_simps rbl_plus_simps rev_bin_to_bl_simps append.simps takefill_last_simps drop_nonempty_simps rev_bl_order_simps} addsimprocs [expand_upt_simproc, nat_get_Suc_simproc 4 [\<^term>\replicate\, \<^term>\takefill x\, \<^term>\drop\, \<^term>\bin_to_bl\, \<^term>\takefill_last x\, \<^term>\drop_nonempty x\]], put_simpset no_split_ss \<^context> addsimps @{thms xor3_simps carry_simps if_bool_simps} ]) fun tac ctxt = let val (ss, sss) = expand_word_eq_sss; in foldr1 (op THEN_ALL_NEW) ((CHANGED o safe_full_simp_tac (put_simpset ss ctxt)) :: map (fn ss => safe_full_simp_tac (put_simpset ss ctxt)) sss) end; end \ method_setup word_bitwise = \Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (Word_Bitwise_Tac.tac ctxt 1))\ "decomposer for word equalities and inequalities into bit propositions" end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,6178 +1,6178 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Word_EqI Word_Enum "HOL-Library.Sublist" begin lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma same_length_is_parallel: assumes len: "\y \ set as. length y = x" shows "\x \ set as. \y \ set as - {x}. x \ y" proof (rule, rule) fix x y assume xi: "x \ set as" and yi: "y \ set as - {x}" from len obtain q where len': "\y \ set as. length y = q" .. show "x \ y" proof (rule not_equal_is_parallel) from xi yi show "x \ y" by auto from xi yi len' show "length x = length y" by (auto dest: bspec) qed qed text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_unat.Rep_eqD) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma set_enum_word8_def: "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < LENGTH('a)" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "x < y \ [x::'a::len word .e. y] = x # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons, unat_arith) apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply (rule map_cong[OF _ refl]) apply (rule arg_cong2[where f = "\x y. [x ..< y]"], unat_arith) apply (rule refl) done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv by simp qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma mask_shift: "(x && ~~ (mask y)) >> y = x >> y" by word_eqI lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_def) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_unat.Rep_eqD) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by (case_tac "size w \ n", clarsimp simp: shiftl_zero_size) (clarsimp simp: not_le ucast_bl bl_shiftl bang_eq test_bit_of_bl rev_nth nth_append) lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_cases: fixes a :: "'a::len word" shows "length [a .e. b] = (if a \ b then Suc (unat b) - unat a else 0)" apply (case_tac "a \ b") apply (clarsimp) apply (clarsimp simp: upto_enum_def) apply unat_arith done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto declare of_nat_power [simp del] (* TODO: move to word *) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt word_unat_power unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma is_aligned_0'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply (simp add: word_unat.Rep_inject [symmetric]) apply simp done lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof cases assume szv: "sz < LENGTH('a::len)" show ?thesis proof (cases "sz = 0") case True then show ?thesis using kv szv by (simp add: unat_of_nat) next case False then have sne: "0 < sz" .. have uk: "unat (of_nat k :: 'a word) = k" apply (subst unat_of_nat) apply (simp add: nat_mod_eq less_trans[OF kv] sne) done show ?thesis using szv apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+ done qed next assume "\ sz < LENGTH('a)" with kv show ?thesis by (simp add: not_less power_overflow) qed lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_def word_mod_by_0) qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF szv]]) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_def power_overflow) then show ?thesis .. qed lemma less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by (simp add: ucast_def test_bit_bin word_ubin.eq_norm nth_bintr word_size) (fast elim!: bin_nth_uint_imp) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by (meson Word.nth_ucast test_bit_conj_lt le_def upper_bits_unset_is_l2p) lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) by word_eqI_solve lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: word_unat_power) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" apply (simp add: word_less_nat_alt) apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply simp+ done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp apply assumption done ultimately show ?thesis by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" apply (erule udvd_minus_le') apply (simp add: udvd_def)+ done lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma prefix_length_less: "strict_prefix xs ys \ length xs < length ys" apply (clarsimp simp: strict_prefix_def) apply (frule prefix_length_le) apply (rule ccontr, simp) apply (clarsimp simp: prefix_def) done lemmas take_less = take_strict_prefix lemma not_prefix_longer: "\ length xs > length ys \ \ \ prefix xs ys" by (clarsimp dest!: prefix_length_le) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by (subst unat_of_nat_eq[where x=n, symmetric], simp+) lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: unat_of_nat) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_def [symmetric]) apply (subst unat_def [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" by (drule unat_of_nat_eq, simp) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_def) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q as bs" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ \ P (x # xs) (y # ys)" shows "P as bs" proof - define as' where "as' == as" define bs' where "bs' == bs" have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp then show ?thesis using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 show ?case by fact next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" by (auto simp: as'_def bs'_def) qed qed qed lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma shiftr_less_t2n': "\ x && mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p[symmetric]) apply word_eqI apply (erule_tac x="na + n" in allE) apply fastforce done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) && ~~ (mask m) = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_def le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p[symmetric]) apply word_eqI apply (erule_tac x="na - n" in allE) apply auto done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" by word_eqI lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" apply (simp add: unat_def ucast_def) apply (subst word_uint.eq_norm) apply (subst nat_mod_distrib) apply simp apply simp apply (subst nat_power_eq) apply simp apply simp done lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (simp add: scast_def) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) - apply (simp add: nth_rev) + apply (simp add: rev_nth) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" by (rule word_plus_and_or_coroll, word_eqI) blast lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma take_prefix: "(take (length xs) ys = xs) = prefix xs ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ (mask n)) = x - (x && (mask n))" by (simp add: field_simps word_plus_and_or_coroll2) lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_def word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m\size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" by (simp add: le_mask_iff shiftl_shiftr3) lemma and_not_mask_twice: "(w && ~~ (mask n)) && ~~ (mask m) = w && ~~ (mask (max m n))" apply (simp add: and_not_mask) apply (case_tac "n x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by word_eqI lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y && z = z \ \ x && y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_def, folded limited_and_def] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by simp lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] compl_of_1 shiftl_shiftr1[unfolded word_size mask_def] shiftl_shiftr2[unfolded word_size mask_def] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" by safe word_eqI_solve lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n, simp) apply (case_tac n; simp) done lemma word_and_1_shiftl: "x && (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" by word_eqI_solve lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x && (mask n << m) = ((x >> m) && mask n) << m" by word_eqI_solve lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma word_FF_is_mask: "0xFF = mask 8" by (simp add: mask_def) lemma word_1FF_is_mask: "0x1FF = mask 9" by (simp add: mask_def) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed lemma word_div_less: "m < n \ m div n = 0" for m :: "'a :: len word" by (simp add: unat_mono word_arith_nat_defs(6)) lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]; simp) done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemmas if_fun_split = if_apply_def2 lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp add: word_less_nat_alt) apply (subst unat_of_nat_eq) apply (erule order_less_trans) apply simp+ done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (induct x) apply clarsimp+ by (metis Suc_eq_plus1 add_lessD1 less_irrefl one_add_one unatSuc word_less_nat_alt) lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w && mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma shiftr_mask_eq: "(x >> n) && mask (size x - n) = x >> n" for x :: "'a :: len word" by word_eqI_solve lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) && mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply (simp add: word_unat.norm_eq_iff [symmetric]) done lemma mask_AND_NOT_mask: "(w && ~~ (mask n)) && mask n = 0" by word_eqI lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ (mask n)) + (w && mask n) = w" by (subst word_plus_and_or_coroll; word_eqI_solve) lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ (mask n) = y && ~~ (mask n)" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size test_bit_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size test_bit_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ (mask n)) !! m)" by (simp add: neg_mask_test_bit test_bit_conj_lt) also have "\ = ((y && ~~ (mask n)) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_test_bit test_bit_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: unat_of_nat) done lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply simp done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply word_eqI done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = 1" by (simp add: mask_def) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" by (simp add: scast_def ucast_def sint_eq_uint) lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" by unat_arith lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: "\ \ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P; \ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P; \ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P \ \ P" apply atomize_elim apply (case_tac "len_of TYPE ('a) = 1") apply clarsimp apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}") apply (metis UNIV_I insert_iff singletonE) apply (subst word_unat.univ) apply (clarsimp simp: unats_def image_def) apply (rule set_eqI, rule iffI) apply clarsimp apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0) apply clarsimp apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc) apply clarsimp apply (metis One_nat_def arith_is_1 le_def len_gt_0) done lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply clarsimp apply (metis Suc_pred int_word_uint len_gt_0 power_Suc uint_eq_0 word_of_int_2p word_pow_0) done lemma sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_def word_sless_alt sbintrunc_If) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) apply (metis word_ubin.eq_norm) done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) apply (metis word_ubin.eq_norm) done (* Signed word arithmetic overflow constraints. *) lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply (simp add: sint_word_ariths scast_def) apply (simp add: wi_hom_mult) apply (subst word_sint.Abs_inject, simp_all) apply (simp add: sints_def range_sbintrunc abs_less_iff) apply clarsimp apply (simp add: sints_def range_sbintrunc word_size) apply (auto elim: order_less_le_trans order_trans[rotated]) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less algebra_split_simps) apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (clarsimp simp: sgn_if) apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans) apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) apply (metis wi_hom_neg word_sint.Rep_inverse') done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: sdiv_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x] sdiv_word_def[where a=0] sdiv_word_def[where a=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y] word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (case_tac "a \ - (2 ^ (LENGTH('a) - 1)) \ b \ -1") apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def minus_word.abs_eq [symmetric] times_word.abs_eq [symmetric]) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x] smod_word_def[where a=0] smod_word_def[where a=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y] word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" apply (clarsimp simp: word_of_int int_word_sint) apply (subst int_mod_eq') apply simp apply (subst (2) power_minus_simp) apply clarsimp apply clarsimp apply clarsimp done lemma of_int_sint [simp]: "of_int (sint a) = a" apply (insert word_sint.Rep [where x=a]) apply (clarsimp simp: word_of_int) done lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_def test_bit_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (clarsimp simp: word_of_int ucast_def) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_of_int) apply (subst word.abs_eq_iff) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_mod_power_int min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_def sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (clarsimp simp: scast_def) apply (metis down_cast_same is_up_down scast_def ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_bl ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) 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 signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply (subst sint_eq_uint) apply (clarsimp simp: msb_nth nth_ucast is_down) apply (metis Suc_leI Suc_pred len_gt_0) apply (clarsimp simp: uint_up_ucast is_up is_down) done lemma word_less_nowrapI': "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = 2 ^ n" by (clarsimp simp: mask_def) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp add: ucast_nat_def [symmetric]) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp add: ucast_nat_def [symmetric]) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) apply (subst (asm) test_bit_conj_lt [symmetric]) apply clarsimp apply arith done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" by (simp add: test_bit_word_eq to_bool_def and_one_eq mod_2_eq_odd) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) && 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) - apply (simp add: nth_ucast nth_shiftr nth_rev field_simps) + apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ (mask n)) + (2 ^ n - 1) = x || mask n" unfolding mask_2pm1[symmetric] by (subst word_plus_and_or_coroll; word_eqI_solve) lemma subtract_mask: "p - (p && mask n) = (p && ~~ (mask n))" "p - (p && ~~ (mask n)) = (p && mask n)" by (simp add: field_simps word_plus_and_or_coroll2)+ lemma and_neg_mask_plus_mask_mono: "(p && ~~ (mask n)) + mask n \ p" apply (rule word_le_minus_cancel[where x = "p && ~~ (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_def word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr && ~~ (mask n)) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply simp done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ (mask n)) = p)" apply (subst(asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll; word_eqI; blast) using word_plus_and_or_coroll2[where x="p + d" and w="mask n"] by simp lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" by word_eqI_solve lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a && mask n) && ~~ (mask m) = (ptr + a && ~~ (mask m) ) && mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w && mask n) && ~~ (mask n) = 0" by (clarsimp simp add: mask_def Parity.bit_eq_iff bit_and_iff bit_not_iff bit_mask_iff) lemma and_and_not[simp]:"(a && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" by (clarsimp simp add: mask_def Parity.bit_eq_iff bit_and_iff bit_not_iff shiftl_word_eq bit_push_bit_iff) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma ex_mask_1[simp]: "(\x. mask x = 1)" apply (rule_tac x=1 in exI) apply (simp add:mask_def) done lemma not_switch:"~~ a = x \ a = ~~ x" by auto (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_def) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma swap_with_xor: "\(x::'a::len word) = a xor b; y = b xor x; z = x xor y\ \ z = b \ y = a" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma scast_nop1: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop[simp] = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ (mask n)) && mask n = x && mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) || y && mask n) && ~~ (mask m) = x && ~~ (mask m)" by (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff mask_eq_mask bit_mask_iff) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w && ~~(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" proof (cases \2 \ LENGTH('a)\) case False then have *: \LENGTH('a) = 1\ by simp then have \x = 0 \ x = 1\ by (metis One_nat_def less_irrefl_nat sint_1_cases) then show ?thesis by (auto simp add: word_arith_nat_defs(6) *) next case True then show ?thesis using shiftr1_unfold [symmetric, of x] uint_2_id [where ?'a = 'a] by (simp add: shiftr1_def word_div_def) qed lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" unfolding word_lsb_def bin_last_def by (metis (no_types, hide_lams) nat_mod_distrib nat_numeral not_mod_2_eq_1_eq_0 numeral_One uint_eq_0 uint_nonnegative unat_0 unat_def zero_le_numeral) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" using word_overflow_unat [of x] apply (simp only: shiftr1_is_div_2 flip: odd_iff_lsb) apply (cases \2 \ LENGTH('a)\) apply (auto simp add: test_bit_def' uint_nat word_arith_nat_div dest: overflow_imp_lsb) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) || (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_def flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) || (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x && mask n || x && ~~ (mask n) = x" apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) && ~~ (mask n) = p && ~~ (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0..UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by (simp add: ucast_def word_of_int_minus) have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis by (simp add: ucast_def word_ubin.eq_norm bintrunc_mod2p) qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_def not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (~~ (2 ^ n :: 'a::len word)) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x && y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) && mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply (simp add: shiftr1_def) apply transfer apply simp done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial; simp?) apply (rule lt2p_lem) apply (simp add: assms) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp add: ucast_nat_def[symmetric]) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst Word.ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) shiftl_shiftr_id shiftl_t2n) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) unfolding is_aligned_def by (metis le_unat_uoi nat_dvd_not_less order_less_imp_le unat_power_lower) lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt2 lt3 unat_mono word_less_minus_mono_left by blast qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: "(max_word :: 'a::len word) = mask LENGTH('a)" unfolding mask_def by simp lemmas mask_len_max = max_word_mask[symmetric] lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_def is_aligned_mask mask_def word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma complement_mask: "complement (2 ^ n - 1) = ~~ (mask n)" unfolding complement_def mask_def by simp lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis complement_mask is_aligned_add_helper p_assoc_help power_2_ge_iff) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst complement_mask) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (metis unat_div unat_less_helper unat_power_lower) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff unat_def zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 && ~~ (mask sz)" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_def[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ (mask n) = y; n \ m \ \ x && ~~ (mask m) = y && ~~ (mask m)" by word_eqI_solve lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q && ~~ (mask n)) = (p + q) && ~~ (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 && ~~ (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma mask_lower_twice: "n \ m \ (x && ~~ (mask n)) && ~~ (mask m) = x && ~~ (mask m)" by word_eqI_solve lemma mask_lower_twice2: "(a && ~~ (mask n)) && ~~ (mask m) = a && ~~ (mask (max n m))" by word_eqI_solve lemma ucast_and_neg_mask: "ucast (x && ~~ (mask n)) = ucast x && ~~ (mask n)" by word_eqI_solve lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" by word_eqI_solve lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" by word_eqI lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q && ~~ (mask sz)) = (p - ((alignUp q sz) && ~~ (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) by (simp add: alignUp_distance and_mask_0_iff_le_mask is_aligned_neg_mask_eq mask_out_add_aligned) lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by (simp add: is_down ucast_id ucast_ucast_a) lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" by (simp add: ucast_def uint_or flip: or_word.abs_eq) lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w && w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2^n = y * (2^n::'a::len word)" shows "x = y" proof (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size] lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w || ~~ (mask (Suc n)) else w && mask (Suc n))" by word_eqI (auto dest: less_antisym) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if') lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply (rule iffI) apply (word_eqI, rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size) apply (erule subst, rule sign_extended_sign_extend) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by word_eqI lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by word_eqI_solve lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr && ~~ (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num) lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: "bin_cat a n b \ uints m" if "a \ uints l" "m \ l + n" using that proof (induction n arbitrary: b m) case 0 then have "uints l \ uints m" by (intro uints_monoI) auto then show ?case using 0 by auto next case (Suc n) then show ?case using Bit_in_uintsI by auto qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" by (auto simp: word_cat_def word_uint.Abs_inject word_of_int_bin_cat_eq_iff that) lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = 2^n + mask n" by (simp add: mask_def) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_def) (erule is_aligned_no_overflow') lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" by - word_eqI thus ?thesis by simp qed lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" apply (induct xs) apply simp apply (simp add: of_bl_Cons mask_Suc) apply (rule conjI; clarsimp) apply (erule word_plus_mono_right) apply (rule is_aligned_no_overflow_mask) apply (rule is_aligned_triv) apply (simp add: word_le_nat_alt) apply (subst unat_add_lem') apply (rule is_aligned_mask_offset_unat) apply (rule is_aligned_triv) apply (simp add: mask_def) apply simp done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_def) lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def word_le_nat_alt) apply (drule le_imp_less_Suc) apply (simp add: Suc_2p_unat_mask) by (metis Groups.mult_ac(2) Suc_leI linorder_not_less mult_le_mono order_refl times_nat.simps(2)) 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) by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a && b = 0" by word_eqI_solve lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a || b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a || b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (clarsimp simp: word_eqI_simps) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" by (clarsimp simp: le_mask_high_bits word_size nth_ucast) lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: "\ p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp done lemma aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma max_word_not_less[simp]: "\ max_word < x" by (simp add: not_less) lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: "m && (2 ^ n - 1) = 0 \ (x + y * m) && ~~(mask n) = (x && ~~(mask n)) + y * m" by (metis Groups.add_ac(2) is_aligned_mask mask_def mask_eqs(5) mask_out_add_aligned mult_zero_right shiftl_1 word_bw_comms(1) word_log_esimps(1)) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y && mask n = 0 \ x + y && ~~(mask n) = (x && ~~(mask n)) + y" by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x && mask n = 0; y && mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_def) done lemma sub_right_shift: "\ x && mask n = 0; y && mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y && mask n = mask n \ (x && y) && mask n = x && mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y && mask n = 0 \ (x && y) && mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) && b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) && b < c" by (metis word_and_le1 xtr7) lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x && mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x && mask n = 0; m \ n \ \ x && mask m = 0" by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) && mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) && 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by simp lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat word_unat.Rep_inverse) lemma shiftr_and_eq_shiftl: "(w >> n) && x = y \ w && (x << n) = (y << n)" for y :: "'a:: len word" by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) lemma neg_mask_combine: "~~(mask a) && ~~(mask b) = ~~(mask (max a b))" by (auto simp: word_ops_nth_size word_size intro!: word_eqI) lemma neg_mask_twice: "x && ~~(mask n) && ~~(mask m) = x && ~~(mask (max n m))" by (metis neg_mask_combine) lemma multiple_mask_trivia: "n \ m \ (x && ~~(mask n)) + (x && mask n && ~~(mask m)) = x && ~~(mask m)" apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p && ~~(mask n) = x" using add_mask_lower_bits by auto lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' && ~~(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis le_word_or2 neg_mask_add_mask and.right_idem) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_def) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (simp only: is_aligned_add_or flip: neg_mask_in_mask_range) by (metis less_mask_eq mask_subsume) lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_def) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_def) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_def) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) apply (erule nonemptyE) apply simp apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_def) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply (intro range_subsetI) apply (rule is_aligned_no_wrap' [OF al xsz]) apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps) lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p && ~~(mask n') \ p'; p' && ~~(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x && mask n >> m" in spec) apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1) done lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" by (clarsimp simp: word_le_def uint_nat of_nat_inverse) lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x && ~~(mask sz)) + 2 ^ sz - 1" by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" apply (case_tac "LENGTH('a) = 1") apply (clarsimp simp: ucast_def) apply (metis (hide_lams, mono_tags) One_nat_def len_signed plus_word.abs_eq uint_word_arith_bintrs(1) word_ubin.Abs_norm) apply (clarsimp simp: ucast_def) apply (metis le_refl len_signed plus_word.abs_eq uint_word_arith_bintrs(1) wi_bintr) done lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) && (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) && mask m) = unat (x && mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using small_powers_of_2 by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2 not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower word_clz_max word_of_nat_less wsst_TYs(3)) lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using small_powers_of_2 uint_nat apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less word_size) by (simp add: uint_nat) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_def power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt by (metis (no_types, lifting)) lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (simp add: ucast_def word_of_int) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n && msk = mask n \ \ (x mod m) && msk = x mod m" by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x && mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m && ~~(mask n)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr && ~~(mask sz')) - (ptr && ~~(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") proof - assume lt: "sz' \ sz" hence "?lhs = ptr && (mask sz && ~~(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma of_bl_length2: "length xs + c < LENGTH('a) \ of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) && ~~(mask sz) = 0" by (simp add: Word_Lemmas.of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr && ~~(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) && ~~(mask sz)) + unat (ptr && mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis (no_types) Word_Lemmas.of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d && mask n) = unat d \ unat (p + d && ~~(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) && ~~(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) && mask high_bits = x >> low_bits" by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x && mask low_bits = 0) = (x = 0)" using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) && mask n = q && mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma from_to_bool_last_bit: "from_bool (to_bool (x && 1)) = x && 1" by (metis from_bool_to_bool_iff word_and_1) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) && mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x && ~~(mask sz)) + (x && mask sz) \ (x && ~~(mask sz)) + ((x && mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using Word_Lemmas.of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "~~(mask n) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * ~~(mask n) = - (x && ~~(mask n))" by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) && ~~(mask n) = p + (x && ~~(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) && ~~(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) && mask (m::nat) = mask (min m n)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n && mask m = (if n < m then 2 ^ n else 0)" by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (max_word :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p && mask a >> b) :: 'a :: len word) = p && mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p && mask a >> b) :: 'a :: len word) = unat (p && mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p && mask a) && ~~(mask b) = 0" by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p && mask a << c) && ~~(mask b) = 0" by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p && ~~(mask a)) && mask b = 0" using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p && ~~(mask a)) + ((p && mask a) && ~~(mask b)) + (p && mask b) = p" by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p && mask a >> b << b) = (p && mask a) && ~~(mask b)" by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p && mask b) \ \ (p && ~~(mask a)) + (p && mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (metis scast_def word_of_int) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (metis cast_simps(23) scast_scast_id(2)) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x && mask n \ of_nat ` set [0 ..< 2 ^ n]" proof - have "x && mask n \ {0 .. 2 ^ n - 1}" by (simp add: mask_def word_and_le1) also have "... = of_nat ` {0 .. 2 ^ n - 1}" apply (rule set_eqI, rule iffI) apply (clarsimp simp: image_iff) apply (rule_tac x="unat x" in bexI; simp) using len apply (simp add: word_le_nat_alt unat_2tp_if unat_minus_one) using len apply (clarsimp simp: word_le_nat_alt unat_2tp_if unat_minus_one) apply (subst unat_of_nat_eq; simp add: nat_le_Suc_less) apply (erule less_le_trans) apply simp done also have "{0::nat .. 2^n - 1} = set [0 ..< 2^n]" by (auto simp: nat_le_Suc_less) finally show ?thesis . qed lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" by (simp add: Word_Lemmas.of_nat_power not_msb_from_less sint_eq_uint) lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower) lemma sint_ctz: "LENGTH('a) > 2 \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") apply (rule conjI) apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) by (rule small_powers_of_2, simp) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w && mask (size w - n)" by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp add: mask_def NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x && mask LENGTH('b) = y && mask LENGTH('b))" by (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" by (metis assms sbintrunc_bintrunc ucast_def word_ubin.eq_norm) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: "scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))" unfolding scast_def sint_uint word_size apply (subst word_eq_iff) apply (rule iffI) apply (rule ballI) apply (drule_tac x=i in spec) apply (subst (asm) test_bit_sbintrunc_ucast; simp) apply (rule allI) apply (case_tac "n < LENGTH('a)") apply (subst test_bit_sbintrunc_ucast) apply simp apply (case_tac "n \ LENGTH('b)") apply (drule_tac x=n in bspec) by auto lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ ~~(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x && mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (~~x) = ~~(ucast x) && mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (~~x) = ~~(UCAST('a \ 'b) x)" by word_eqI lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) && ~~(mask n) = a" by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1)) lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (subst word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule is_aligned_AND_less_0) apply (simp add: is_aligned_mask) apply (rule order_less_le_trans) apply (rule of_bl_length2) apply simp apply (simp add: two_power_increasing) apply simp apply (rule nth_equalityI) apply (simp only: len_bin_to_bl) apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric]) apply (simp add: nth_shiftr nth_shiftl shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric]) - apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl nth_rev) + apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl rev_nth) apply arith done end