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,1413 +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" for a b :: \'a::len word\ 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" by transfer (simp add: smod_by_div_abs) context includes bit_operations_syntax begin 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" end lift_definition eval_bin_op_impl :: "bin_op \ uint32 \ uint32 \ uint32" is eval_bin_op . lemma [code]: "eval_bin_op_impl bop v1 v2 = eval_bin_op_impl_aux bop v1 v2" apply (cases bop) 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" +| "eval_un_op_impl_aux uo_not v = Bit_Operations.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/CakeML/generated/Lem_word.thy b/thys/CakeML/generated/Lem_word.thy --- a/thys/CakeML/generated/Lem_word.thy +++ b/thys/CakeML/generated/Lem_word.thy @@ -1,1024 +1,1024 @@ chapter \Generated by Lem from \word.lem\.\ theory "Lem_word" imports Main "Lem_bool" "Lem_maybe" "Lem_num" "Lem_basic_classes" "Lem_list" "HOL-Library.Word" begin \ \\open import Bool Maybe Num Basic_classes List\\ \ \\open import {isabelle} `HOL-Library.Word`\\ \ \\open import {hol} `wordsTheory` `wordsLib`\\ \ \\ ========================================================================== \\ \ \\ Define general purpose word, i.e. sequences of bits of arbitrary length \\ \ \\ ========================================================================== \\ datatype bitSequence = BitSeq " nat option " " \ \\ length of the sequence, Nothing means infinite length \\ bool " " bool \ \\ sign of the word, used to fill up after concrete value is exhausted \\ list " \ \\ the initial part of the sequence, least significant bit first \\ \ \\val bitSeqEq : bitSequence -> bitSequence -> bool\\ \ \\val boolListFrombitSeq : nat -> bitSequence -> list bool\\ fun boolListFrombitSeqAux :: " nat \ 'a \ 'a list \ 'a list " where " boolListFrombitSeqAux n s bl = ( if n =( 0 :: nat) then [] else (case bl of [] => List.replicate n s | b # bl' => b # (boolListFrombitSeqAux (n-( 1 :: nat)) s bl') ))" fun boolListFrombitSeq :: " nat \ bitSequence \(bool)list " where " boolListFrombitSeq n (BitSeq _ s bl) = ( boolListFrombitSeqAux n s bl )" \ \\val bitSeqFromBoolList : list bool -> maybe bitSequence\\ definition bitSeqFromBoolList :: "(bool)list \(bitSequence)option " where " bitSeqFromBoolList bl = ( (case dest_init bl of None => None | Some (bl', s) => Some (BitSeq (Some (List.length bl)) s bl') ))" \ \\ cleans up the representation of a bitSequence without changing its semantics \\ \ \\val cleanBitSeq : bitSequence -> bitSequence\\ fun cleanBitSeq :: " bitSequence \ bitSequence " where " cleanBitSeq (BitSeq len s bl) = ( (case len of None => (BitSeq len s (List.rev (dropWhile ((\) s) (List.rev bl)))) | Some n => (BitSeq len s (List.rev (dropWhile ((\) s) (List.rev (List.take (n-( 1 :: nat)) bl))))) ))" \ \\val bitSeqTestBit : bitSequence -> nat -> maybe bool\\ fun bitSeqTestBit :: " bitSequence \ nat \(bool)option " where " bitSeqTestBit (BitSeq None s bl) pos = ( if pos < List.length bl then index bl pos else Some s )" |" bitSeqTestBit (BitSeq(Some l) s bl) pos = ( if (pos \ l) then None else if ((pos = (l -( 1 :: nat))) \ (pos \ List.length bl)) then Some s else index bl pos )" \ \\val bitSeqSetBit : bitSequence -> nat -> bool -> bitSequence\\ fun bitSeqSetBit :: " bitSequence \ nat \ bool \ bitSequence " where " bitSeqSetBit (BitSeq len s bl) pos v = ( (let bl' = (if (pos < List.length bl) then bl else bl @ List.replicate pos s) in (let bl'' = (List.list_update bl' pos v) in (let bs' = (BitSeq len s bl'') in cleanBitSeq bs'))))" \ \\val resizeBitSeq : maybe nat -> bitSequence -> bitSequence\\ definition resizeBitSeq :: "(nat)option \ bitSequence \ bitSequence " where " resizeBitSeq new_len bs = ( (case cleanBitSeq bs of (BitSeq len s bl) => (let shorten_opt = ((case (new_len, len) of (None, _) => None | (Some l1, None) => Some l1 | (Some l1, Some l2) => if (l1 < l2) then Some l1 else None )) in (case shorten_opt of None => BitSeq new_len s bl | Some l1 => ( (let bl' = (List.take l1 (bl @ [s])) in (case dest_init bl' of None => (BitSeq len s bl) \ \\ do nothing if size 0 is requested \\ | Some (bl'', s') => cleanBitSeq (BitSeq new_len s' bl'') ))) )) ) )" \ \\val bitSeqNot : bitSequence -> bitSequence\\ fun bitSeqNot :: " bitSequence \ bitSequence " where " bitSeqNot (BitSeq len s bl) = ( BitSeq len (\ s) (List.map (\ x. \ x) bl))" \ \\val bitSeqBinop : (bool -> bool -> bool) -> bitSequence -> bitSequence -> bitSequence\\ \ \\val bitSeqBinopAux : (bool -> bool -> bool) -> bool -> list bool -> bool -> list bool -> list bool\\ fun bitSeqBinopAux :: "(bool \ bool \ bool)\ bool \(bool)list \ bool \(bool)list \(bool)list " where " bitSeqBinopAux binop s1 ([]) s2 ([]) = ( [])" |" bitSeqBinopAux binop s1 (b1 # bl1') s2 ([]) = ( (binop b1 s2) # bitSeqBinopAux binop s1 bl1' s2 [])" |" bitSeqBinopAux binop s1 ([]) s2 (b2 # bl2') = ( (binop s1 b2) # bitSeqBinopAux binop s1 [] s2 bl2' )" |" bitSeqBinopAux binop s1 (b1 # bl1') s2 (b2 # bl2') = ( (binop b1 b2) # bitSeqBinopAux binop s1 bl1' s2 bl2' )" definition bitSeqBinop :: "(bool \ bool \ bool)\ bitSequence \ bitSequence \ bitSequence " where " bitSeqBinop binop bs1 bs2 = ( ( (case cleanBitSeq bs1 of (BitSeq len1 s1 bl1) => (case cleanBitSeq bs2 of (BitSeq len2 s2 bl2) => (let len = ((case (len1, len2) of (Some l1, Some l2) => Some (max l1 l2) | _ => None )) in (let s = (binop s1 s2) in (let bl = (bitSeqBinopAux binop s1 bl1 s2 bl2) in cleanBitSeq (BitSeq len s bl)))) ) ) ))" definition bitSeqAnd :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqAnd = ( bitSeqBinop (\))" definition bitSeqOr :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqOr = ( bitSeqBinop (\))" definition bitSeqXor :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqXor = ( bitSeqBinop (\ b1 b2. \ (b1 \ b2)))" \ \\val bitSeqShiftLeft : bitSequence -> nat -> bitSequence\\ fun bitSeqShiftLeft :: " bitSequence \ nat \ bitSequence " where " bitSeqShiftLeft (BitSeq len s bl) n = ( cleanBitSeq (BitSeq len s (List.replicate n False @ bl)))" \ \\val bitSeqArithmeticShiftRight : bitSequence -> nat -> bitSequence\\ definition bitSeqArithmeticShiftRight :: " bitSequence \ nat \ bitSequence " where " bitSeqArithmeticShiftRight bs n = ( (case cleanBitSeq bs of (BitSeq len s bl) => cleanBitSeq (BitSeq len s (List.drop n bl)) ) )" \ \\val bitSeqLogicalShiftRight : bitSequence -> nat -> bitSequence\\ definition bitSeqLogicalShiftRight :: " bitSequence \ nat \ bitSequence " where " bitSeqLogicalShiftRight bs n = ( if (n =( 0 :: nat)) then cleanBitSeq bs else (case cleanBitSeq bs of (BitSeq len s bl) => (case len of None => cleanBitSeq (BitSeq len s (List.drop n bl)) | Some l => cleanBitSeq (BitSeq len False ((List.drop n bl) @ List.replicate l s)) ) ) )" \ \\ integerFromBoolList sign bl creates an integer from a list of bits (least significant bit first) and an explicitly given sign bit. It uses two's complement encoding. \\ \ \\val integerFromBoolList : (bool * list bool) -> integer\\ fun integerFromBoolListAux :: " int \(bool)list \ int " where " integerFromBoolListAux (acc1 :: int) (([]) :: bool list) = ( acc1 )" |" integerFromBoolListAux (acc1 :: int) ((True # bl') :: bool list) = ( integerFromBoolListAux ((acc1 *( 2 :: int)) +( 1 :: int)) bl' )" |" integerFromBoolListAux (acc1 :: int) ((False # bl') :: bool list) = ( integerFromBoolListAux (acc1 *( 2 :: int)) bl' )" fun integerFromBoolList :: " bool*(bool)list \ int " where " integerFromBoolList (sign, bl) = ( if sign then - (integerFromBoolListAux(( 0 :: int)) (List.rev (List.map (\ x. \ x) bl)) +( 1 :: int)) else integerFromBoolListAux(( 0 :: int)) (List.rev bl))" \ \\ [boolListFromInteger i] creates a sign bit and a list of booleans from an integer. The len_opt tells it when to stop.\\ \ \\val boolListFromInteger : integer -> bool * list bool\\ fun boolListFromNatural :: "(bool)list \ nat \(bool)list " where " boolListFromNatural acc1 (remainder :: nat) = ( if (remainder >( 0 :: nat)) then (boolListFromNatural (((remainder mod( 2 :: nat)) =( 1 :: nat)) # acc1) (remainder div( 2 :: nat))) else List.rev acc1 )" definition boolListFromInteger :: " int \ bool*(bool)list " where " boolListFromInteger (i :: int) = ( if (i <( 0 :: int)) then (True, List.map (\ x. \ x) (boolListFromNatural [] (nat (abs (- (i +( 1 :: int))))))) else (False, boolListFromNatural [] (nat (abs i))))" \ \\ [bitSeqFromInteger len_opt i] encodes [i] as a bitsequence with [len_opt] bits. If there are not enough bits, truncation happens \\ \ \\val bitSeqFromInteger : maybe nat -> integer -> bitSequence\\ definition bitSeqFromInteger :: "(nat)option \ int \ bitSequence " where " bitSeqFromInteger len_opt i = ( (let (s, bl) = (boolListFromInteger i) in resizeBitSeq len_opt (BitSeq None s bl)))" \ \\val integerFromBitSeq : bitSequence -> integer\\ definition integerFromBitSeq :: " bitSequence \ int " where " integerFromBitSeq bs = ( (case cleanBitSeq bs of (BitSeq len s bl) => integerFromBoolList (s, bl) ) )" \ \\ Now we can via translation to integers map arithmetic operations to bitSequences \\ \ \\val bitSeqArithUnaryOp : (integer -> integer) -> bitSequence -> bitSequence\\ definition bitSeqArithUnaryOp :: "(int \ int)\ bitSequence \ bitSequence " where " bitSeqArithUnaryOp uop bs = ( (case bs of (BitSeq len _ _) => bitSeqFromInteger len (uop (integerFromBitSeq bs)) ) )" \ \\val bitSeqArithBinOp : (integer -> integer -> integer) -> bitSequence -> bitSequence -> bitSequence\\ definition bitSeqArithBinOp :: "(int \ int \ int)\ bitSequence \ bitSequence \ bitSequence " where " bitSeqArithBinOp binop bs1 bs2 = ( (case bs1 of (BitSeq len1 _ _) => (case bs2 of (BitSeq len2 _ _) => (let len = ((case (len1, len2) of (Some l1, Some l2) => Some (max l1 l2) | _ => None )) in bitSeqFromInteger len (binop (integerFromBitSeq bs1) (integerFromBitSeq bs2))) ) ) )" \ \\val bitSeqArithBinTest : forall 'a. (integer -> integer -> 'a) -> bitSequence -> bitSequence -> 'a\\ definition bitSeqArithBinTest :: "(int \ int \ 'a)\ bitSequence \ bitSequence \ 'a " where " bitSeqArithBinTest binop bs1 bs2 = ( binop (integerFromBitSeq bs1) (integerFromBitSeq bs2))" \ \\ now instantiate the number interface for bit-sequences \\ \ \\val bitSeqFromNumeral : numeral -> bitSequence\\ \ \\val bitSeqLess : bitSequence -> bitSequence -> bool\\ definition bitSeqLess :: " bitSequence \ bitSequence \ bool " where " bitSeqLess bs1 bs2 = ( bitSeqArithBinTest (<) bs1 bs2 )" \ \\val bitSeqLessEqual : bitSequence -> bitSequence -> bool\\ definition bitSeqLessEqual :: " bitSequence \ bitSequence \ bool " where " bitSeqLessEqual bs1 bs2 = ( bitSeqArithBinTest (\) bs1 bs2 )" \ \\val bitSeqGreater : bitSequence -> bitSequence -> bool\\ definition bitSeqGreater :: " bitSequence \ bitSequence \ bool " where " bitSeqGreater bs1 bs2 = ( bitSeqArithBinTest (>) bs1 bs2 )" \ \\val bitSeqGreaterEqual : bitSequence -> bitSequence -> bool\\ definition bitSeqGreaterEqual :: " bitSequence \ bitSequence \ bool " where " bitSeqGreaterEqual bs1 bs2 = ( bitSeqArithBinTest (\) bs1 bs2 )" \ \\val bitSeqCompare : bitSequence -> bitSequence -> ordering\\ definition bitSeqCompare :: " bitSequence \ bitSequence \ ordering " where " bitSeqCompare bs1 bs2 = ( bitSeqArithBinTest (genericCompare (<) (=)) bs1 bs2 )" definition instance_Basic_classes_Ord_Word_bitSequence_dict :: "(bitSequence)Ord_class " where " instance_Basic_classes_Ord_Word_bitSequence_dict = ((| compare_method = bitSeqCompare, isLess_method = bitSeqLess, isLessEqual_method = bitSeqLessEqual, isGreater_method = bitSeqGreater, isGreaterEqual_method = bitSeqGreaterEqual |) )" \ \\ arithmetic negation, don't mix up with bitwise negation \\ \ \\val bitSeqNegate : bitSequence -> bitSequence\\ definition bitSeqNegate :: " bitSequence \ bitSequence " where " bitSeqNegate bs = ( bitSeqArithUnaryOp (\ i. - i) bs )" definition instance_Num_NumNegate_Word_bitSequence_dict :: "(bitSequence)NumNegate_class " where " instance_Num_NumNegate_Word_bitSequence_dict = ((| numNegate_method = bitSeqNegate |) )" \ \\val bitSeqAdd : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqAdd :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqAdd bs1 bs2 = ( bitSeqArithBinOp (+) bs1 bs2 )" definition instance_Num_NumAdd_Word_bitSequence_dict :: "(bitSequence)NumAdd_class " where " instance_Num_NumAdd_Word_bitSequence_dict = ((| numAdd_method = bitSeqAdd |) )" \ \\val bitSeqMinus : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMinus :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMinus bs1 bs2 = ( bitSeqArithBinOp (-) bs1 bs2 )" definition instance_Num_NumMinus_Word_bitSequence_dict :: "(bitSequence)NumMinus_class " where " instance_Num_NumMinus_Word_bitSequence_dict = ((| numMinus_method = bitSeqMinus |) )" \ \\val bitSeqSucc : bitSequence -> bitSequence\\ definition bitSeqSucc :: " bitSequence \ bitSequence " where " bitSeqSucc bs = ( bitSeqArithUnaryOp (\ n. n +( 1 :: int)) bs )" definition instance_Num_NumSucc_Word_bitSequence_dict :: "(bitSequence)NumSucc_class " where " instance_Num_NumSucc_Word_bitSequence_dict = ((| succ_method = bitSeqSucc |) )" \ \\val bitSeqPred : bitSequence -> bitSequence\\ definition bitSeqPred :: " bitSequence \ bitSequence " where " bitSeqPred bs = ( bitSeqArithUnaryOp (\ n. n -( 1 :: int)) bs )" definition instance_Num_NumPred_Word_bitSequence_dict :: "(bitSequence)NumPred_class " where " instance_Num_NumPred_Word_bitSequence_dict = ((| pred_method = bitSeqPred |) )" \ \\val bitSeqMult : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMult :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMult bs1 bs2 = ( bitSeqArithBinOp (*) bs1 bs2 )" definition instance_Num_NumMult_Word_bitSequence_dict :: "(bitSequence)NumMult_class " where " instance_Num_NumMult_Word_bitSequence_dict = ((| numMult_method = bitSeqMult |) )" \ \\val bitSeqPow : bitSequence -> nat -> bitSequence\\ definition bitSeqPow :: " bitSequence \ nat \ bitSequence " where " bitSeqPow bs n = ( bitSeqArithUnaryOp (\ i . i ^ n) bs )" definition instance_Num_NumPow_Word_bitSequence_dict :: "(bitSequence)NumPow_class " where " instance_Num_NumPow_Word_bitSequence_dict = ((| numPow_method = bitSeqPow |) )" \ \\val bitSeqDiv : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqDiv :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqDiv bs1 bs2 = ( bitSeqArithBinOp (div) bs1 bs2 )" definition instance_Num_NumIntegerDivision_Word_bitSequence_dict :: "(bitSequence)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Word_bitSequence_dict = ((| div_method = bitSeqDiv |) )" definition instance_Num_NumDivision_Word_bitSequence_dict :: "(bitSequence)NumDivision_class " where " instance_Num_NumDivision_Word_bitSequence_dict = ((| numDivision_method = bitSeqDiv |) )" \ \\val bitSeqMod : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMod :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMod bs1 bs2 = ( bitSeqArithBinOp (mod) bs1 bs2 )" definition instance_Num_NumRemainder_Word_bitSequence_dict :: "(bitSequence)NumRemainder_class " where " instance_Num_NumRemainder_Word_bitSequence_dict = ((| mod_method = bitSeqMod |) )" \ \\val bitSeqMin : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMin :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMin bs1 bs2 = ( bitSeqArithBinOp min bs1 bs2 )" \ \\val bitSeqMax : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMax :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMax bs1 bs2 = ( bitSeqArithBinOp max bs1 bs2 )" definition instance_Basic_classes_OrdMaxMin_Word_bitSequence_dict :: "(bitSequence)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Word_bitSequence_dict = ((| max_method = bitSeqMax, min_method = bitSeqMin |) )" \ \\ ========================================================================== \\ \ \\ Interface for bitoperations \\ \ \\ ========================================================================== \\ record 'a WordNot_class= lnot_method ::" 'a \ 'a " record 'a WordAnd_class= land_method ::" 'a \ 'a \ 'a " record 'a WordOr_class= lor_method ::" 'a \ 'a \ 'a " record 'a WordXor_class= lxor_method ::" 'a \ 'a \ 'a " record 'a WordLsl_class= lsl_method ::" 'a \ nat \ 'a " record 'a WordLsr_class= lsr_method ::" 'a \ nat \ 'a " record 'a WordAsr_class= asr_method ::" 'a \ nat \ 'a " \ \\ ----------------------- \\ \ \\ bitSequence \\ \ \\ ----------------------- \\ definition instance_Word_WordNot_Word_bitSequence_dict :: "(bitSequence)WordNot_class " where " instance_Word_WordNot_Word_bitSequence_dict = ((| lnot_method = bitSeqNot |) )" definition instance_Word_WordAnd_Word_bitSequence_dict :: "(bitSequence)WordAnd_class " where " instance_Word_WordAnd_Word_bitSequence_dict = ((| land_method = bitSeqAnd |) )" definition instance_Word_WordOr_Word_bitSequence_dict :: "(bitSequence)WordOr_class " where " instance_Word_WordOr_Word_bitSequence_dict = ((| lor_method = bitSeqOr |) )" definition instance_Word_WordXor_Word_bitSequence_dict :: "(bitSequence)WordXor_class " where " instance_Word_WordXor_Word_bitSequence_dict = ((| lxor_method = bitSeqXor |) )" definition instance_Word_WordLsl_Word_bitSequence_dict :: "(bitSequence)WordLsl_class " where " instance_Word_WordLsl_Word_bitSequence_dict = ((| lsl_method = bitSeqShiftLeft |) )" definition instance_Word_WordLsr_Word_bitSequence_dict :: "(bitSequence)WordLsr_class " where " instance_Word_WordLsr_Word_bitSequence_dict = ((| lsr_method = bitSeqLogicalShiftRight |) )" definition instance_Word_WordAsr_Word_bitSequence_dict :: "(bitSequence)WordAsr_class " where " instance_Word_WordAsr_Word_bitSequence_dict = ((| asr_method = bitSeqArithmeticShiftRight |) )" \ \\ ----------------------- \\ \ \\ int32 \\ \ \\ ----------------------- \\ \ \\val int32Lnot : int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordNot_Num_int32_dict :: "( 32 word)WordNot_class " where " instance_Word_WordNot_Num_int32_dict = ((| - lnot_method = NOT|) )" + lnot_method = Bit_Operations.not|) )" \ \\val int32Lor : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordOr_Num_int32_dict :: "( 32 word)WordOr_class " where " instance_Word_WordOr_Num_int32_dict = ((| lor_method = Bit_Operations.or|) )" \ \\val int32Lxor : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordXor_Num_int32_dict :: "( 32 word)WordXor_class " where " instance_Word_WordXor_Num_int32_dict = ((| lxor_method = Bit_Operations.xor|) )" \ \\val int32Land : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordAnd_Num_int32_dict :: "( 32 word)WordAnd_class " where " instance_Word_WordAnd_Num_int32_dict = ((| land_method = Bit_Operations.and|) )" \ \\val int32Lsl : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordLsl_Num_int32_dict :: "( 32 word)WordLsl_class " where " instance_Word_WordLsl_Num_int32_dict = ((| lsl_method = (\w n. push_bit n w)|) )" \ \\val int32Lsr : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordLsr_Num_int32_dict :: "( 32 word)WordLsr_class " where " instance_Word_WordLsr_Num_int32_dict = ((| lsr_method = (\w n. drop_bit n w)|) )" \ \\val int32Asr : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordAsr_Num_int32_dict :: "( 32 word)WordAsr_class " where " instance_Word_WordAsr_Num_int32_dict = ((| asr_method = (\w n. signed_drop_bit n w)|) )" \ \\ ----------------------- \\ \ \\ int64 \\ \ \\ ----------------------- \\ \ \\val int64Lnot : int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordNot_Num_int64_dict :: "( 64 word)WordNot_class " where " instance_Word_WordNot_Num_int64_dict = ((| - lnot_method = NOT|) )" + lnot_method = Bit_Operations.not|) )" \ \\val int64Lor : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordOr_Num_int64_dict :: "( 64 word)WordOr_class " where " instance_Word_WordOr_Num_int64_dict = ((| lor_method = Bit_Operations.or|) )" \ \\val int64Lxor : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordXor_Num_int64_dict :: "( 64 word)WordXor_class " where " instance_Word_WordXor_Num_int64_dict = ((| lxor_method = Bit_Operations.xor|) )" \ \\val int64Land : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordAnd_Num_int64_dict :: "( 64 word)WordAnd_class " where " instance_Word_WordAnd_Num_int64_dict = ((| land_method = Bit_Operations.and|) )" \ \\val int64Lsl : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordLsl_Num_int64_dict :: "( 64 word)WordLsl_class " where " instance_Word_WordLsl_Num_int64_dict = ((| lsl_method = (\w n. push_bit n w)|) )" \ \\val int64Lsr : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordLsr_Num_int64_dict :: "( 64 word)WordLsr_class " where " instance_Word_WordLsr_Num_int64_dict = ((| lsr_method = (\w n. drop_bit n w)|) )" \ \\val int64Asr : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordAsr_Num_int64_dict :: "( 64 word)WordAsr_class " where " instance_Word_WordAsr_Num_int64_dict = ((| asr_method = (\w n. signed_drop_bit n w)|) )" \ \\ ----------------------- \\ \ \\ Words via bit sequences \\ \ \\ ----------------------- \\ \ \\val defaultLnot : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a\\ definition defaultLnot :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a " where " defaultLnot fromBitSeq toBitSeq x = ( fromBitSeq (bitSeqNegate (toBitSeq x)))" \ \\val defaultLand : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLand :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLand fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqAnd (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLor : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLor :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLor fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqOr (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLxor : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLxor :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLxor fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqXor (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLsl : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultLsl :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultLsl fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqShiftLeft (toBitSeq x) n))" \ \\val defaultLsr : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultLsr :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultLsr fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqLogicalShiftRight (toBitSeq x) n))" \ \\val defaultAsr : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultAsr :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultAsr fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqArithmeticShiftRight (toBitSeq x) n))" \ \\ ----------------------- \\ \ \\ integer \\ \ \\ ----------------------- \\ \ \\val integerLnot : integer -> integer\\ definition integerLnot :: " int \ int " where " integerLnot i = ( - (i +( 1 :: int)))" definition instance_Word_WordNot_Num_integer_dict :: "(int)WordNot_class " where " instance_Word_WordNot_Num_integer_dict = ((| lnot_method = integerLnot |) )" \ \\val integerLor : integer -> integer -> integer\\ definition integerLor :: " int \ int \ int " where " integerLor i1 i2 = ( defaultLor integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordOr_Num_integer_dict :: "(int)WordOr_class " where " instance_Word_WordOr_Num_integer_dict = ((| lor_method = integerLor |) )" \ \\val integerLxor : integer -> integer -> integer\\ definition integerLxor :: " int \ int \ int " where " integerLxor i1 i2 = ( defaultLxor integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordXor_Num_integer_dict :: "(int)WordXor_class " where " instance_Word_WordXor_Num_integer_dict = ((| lxor_method = integerLxor |) )" \ \\val integerLand : integer -> integer -> integer\\ definition integerLand :: " int \ int \ int " where " integerLand i1 i2 = ( defaultLand integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordAnd_Num_integer_dict :: "(int)WordAnd_class " where " instance_Word_WordAnd_Num_integer_dict = ((| land_method = integerLand |) )" \ \\val integerLsl : integer -> nat -> integer\\ definition integerLsl :: " int \ nat \ int " where " integerLsl i n = ( defaultLsl integerFromBitSeq (bitSeqFromInteger None) i n )" definition instance_Word_WordLsl_Num_integer_dict :: "(int)WordLsl_class " where " instance_Word_WordLsl_Num_integer_dict = ((| lsl_method = integerLsl |) )" \ \\val integerAsr : integer -> nat -> integer\\ definition integerAsr :: " int \ nat \ int " where " integerAsr i n = ( defaultAsr integerFromBitSeq (bitSeqFromInteger None) i n )" definition instance_Word_WordLsr_Num_integer_dict :: "(int)WordLsr_class " where " instance_Word_WordLsr_Num_integer_dict = ((| lsr_method = integerAsr |) )" definition instance_Word_WordAsr_Num_integer_dict :: "(int)WordAsr_class " where " instance_Word_WordAsr_Num_integer_dict = ((| asr_method = integerAsr |) )" \ \\ ----------------------- \\ \ \\ int \\ \ \\ ----------------------- \\ \ \\ sometimes it is convenient to be able to perform bit-operations on ints. However, since int is not well-defined (it has different size on different systems), it should be used very carefully and only for operations that don't depend on the bitwidth of int \\ \ \\val intFromBitSeq : bitSequence -> int\\ definition intFromBitSeq :: " bitSequence \ int " where " intFromBitSeq bs = ( (integerFromBitSeq (resizeBitSeq (Some(( 31 :: nat))) bs)))" \ \\val bitSeqFromInt : int -> bitSequence\\ definition bitSeqFromInt :: " int \ bitSequence " where " bitSeqFromInt i = ( bitSeqFromInteger (Some(( 31 :: nat))) ( i))" \ \\val intLnot : int -> int\\ definition intLnot :: " int \ int " where " intLnot i = ( - (i +( 1 :: int)))" definition instance_Word_WordNot_Num_int_dict :: "(int)WordNot_class " where " instance_Word_WordNot_Num_int_dict = ((| lnot_method = intLnot |) )" \ \\val intLor : int -> int -> int\\ definition intLor :: " int \ int \ int " where " intLor i1 i2 = ( defaultLor intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordOr_Num_int_dict :: "(int)WordOr_class " where " instance_Word_WordOr_Num_int_dict = ((| lor_method = intLor |) )" \ \\val intLxor : int -> int -> int\\ definition intLxor :: " int \ int \ int " where " intLxor i1 i2 = ( defaultLxor intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordXor_Num_int_dict :: "(int)WordXor_class " where " instance_Word_WordXor_Num_int_dict = ((| lxor_method = intLxor |) )" \ \\val intLand : int -> int -> int\\ definition intLand :: " int \ int \ int " where " intLand i1 i2 = ( defaultLand intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordAnd_Num_int_dict :: "(int)WordAnd_class " where " instance_Word_WordAnd_Num_int_dict = ((| land_method = intLand |) )" \ \\val intLsl : int -> nat -> int\\ definition intLsl :: " int \ nat \ int " where " intLsl i n = ( defaultLsl intFromBitSeq bitSeqFromInt i n )" definition instance_Word_WordLsl_Num_int_dict :: "(int)WordLsl_class " where " instance_Word_WordLsl_Num_int_dict = ((| lsl_method = intLsl |) )" \ \\val intAsr : int -> nat -> int\\ definition intAsr :: " int \ nat \ int " where " intAsr i n = ( defaultAsr intFromBitSeq bitSeqFromInt i n )" definition instance_Word_WordAsr_Num_int_dict :: "(int)WordAsr_class " where " instance_Word_WordAsr_Num_int_dict = ((| asr_method = intAsr |) )" \ \\ ----------------------- \\ \ \\ natural \\ \ \\ ----------------------- \\ \ \\ some operations work also on positive numbers \\ \ \\val naturalFromBitSeq : bitSequence -> natural\\ definition naturalFromBitSeq :: " bitSequence \ nat " where " naturalFromBitSeq bs = ( nat (abs (integerFromBitSeq bs)))" \ \\val bitSeqFromNatural : maybe nat -> natural -> bitSequence\\ definition bitSeqFromNatural :: "(nat)option \ nat \ bitSequence " where " bitSeqFromNatural len n = ( bitSeqFromInteger len (int n))" \ \\val naturalLor : natural -> natural -> natural\\ definition naturalLor :: " nat \ nat \ nat " where " naturalLor i1 i2 = ( defaultLor naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordOr_Num_natural_dict :: "(nat)WordOr_class " where " instance_Word_WordOr_Num_natural_dict = ((| lor_method = naturalLor |) )" \ \\val naturalLxor : natural -> natural -> natural\\ definition naturalLxor :: " nat \ nat \ nat " where " naturalLxor i1 i2 = ( defaultLxor naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordXor_Num_natural_dict :: "(nat)WordXor_class " where " instance_Word_WordXor_Num_natural_dict = ((| lxor_method = naturalLxor |) )" \ \\val naturalLand : natural -> natural -> natural\\ definition naturalLand :: " nat \ nat \ nat " where " naturalLand i1 i2 = ( defaultLand naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordAnd_Num_natural_dict :: "(nat)WordAnd_class " where " instance_Word_WordAnd_Num_natural_dict = ((| land_method = naturalLand |) )" \ \\val naturalLsl : natural -> nat -> natural\\ definition naturalLsl :: " nat \ nat \ nat " where " naturalLsl i n = ( defaultLsl naturalFromBitSeq (bitSeqFromNatural None) i n )" definition instance_Word_WordLsl_Num_natural_dict :: "(nat)WordLsl_class " where " instance_Word_WordLsl_Num_natural_dict = ((| lsl_method = naturalLsl |) )" \ \\val naturalAsr : natural -> nat -> natural\\ definition naturalAsr :: " nat \ nat \ nat " where " naturalAsr i n = ( defaultAsr naturalFromBitSeq (bitSeqFromNatural None) i n )" definition instance_Word_WordLsr_Num_natural_dict :: "(nat)WordLsr_class " where " instance_Word_WordLsr_Num_natural_dict = ((| lsr_method = naturalAsr |) )" definition instance_Word_WordAsr_Num_natural_dict :: "(nat)WordAsr_class " where " instance_Word_WordAsr_Num_natural_dict = ((| asr_method = naturalAsr |) )" \ \\ ----------------------- \\ \ \\ nat \\ \ \\ ----------------------- \\ \ \\ sometimes it is convenient to be able to perform bit-operations on nats. However, since nat is not well-defined (it has different size on different systems), it should be used very carefully and only for operations that don't depend on the bitwidth of nat \\ \ \\val natFromBitSeq : bitSequence -> nat\\ definition natFromBitSeq :: " bitSequence \ nat " where " natFromBitSeq bs = ( (naturalFromBitSeq (resizeBitSeq (Some(( 31 :: nat))) bs)))" \ \\val bitSeqFromNat : nat -> bitSequence\\ definition bitSeqFromNat :: " nat \ bitSequence " where " bitSeqFromNat i = ( bitSeqFromNatural (Some(( 31 :: nat))) ( i))" \ \\val natLor : nat -> nat -> nat\\ definition natLor :: " nat \ nat \ nat " where " natLor i1 i2 = ( defaultLor natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordOr_nat_dict :: "(nat)WordOr_class " where " instance_Word_WordOr_nat_dict = ((| lor_method = natLor |) )" \ \\val natLxor : nat -> nat -> nat\\ definition natLxor :: " nat \ nat \ nat " where " natLxor i1 i2 = ( defaultLxor natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordXor_nat_dict :: "(nat)WordXor_class " where " instance_Word_WordXor_nat_dict = ((| lxor_method = natLxor |) )" \ \\val natLand : nat -> nat -> nat\\ definition natLand :: " nat \ nat \ nat " where " natLand i1 i2 = ( defaultLand natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordAnd_nat_dict :: "(nat)WordAnd_class " where " instance_Word_WordAnd_nat_dict = ((| land_method = natLand |) )" \ \\val natLsl : nat -> nat -> nat\\ definition natLsl :: " nat \ nat \ nat " where " natLsl i n = ( defaultLsl natFromBitSeq bitSeqFromNat i n )" definition instance_Word_WordLsl_nat_dict :: "(nat)WordLsl_class " where " instance_Word_WordLsl_nat_dict = ((| lsl_method = natLsl |) )" \ \\val natAsr : nat -> nat -> nat\\ definition natAsr :: " nat \ nat \ nat " where " natAsr i n = ( defaultAsr natFromBitSeq bitSeqFromNat i n )" definition instance_Word_WordAsr_nat_dict :: "(nat)WordAsr_class " where " instance_Word_WordAsr_nat_dict = ((| asr_method = natAsr |) )" end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/Code_Interface.thy b/thys/Iptables_Semantics/Primitive_Matchers/Code_Interface.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/Code_Interface.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/Code_Interface.thy @@ -1,207 +1,208 @@ theory Code_Interface imports Common_Primitive_toString IP_Addresses.IP_Address_Parser "../Call_Return_Unfolding" Transform No_Spoof "../Simple_Firewall/SimpleFw_Compliance" Simple_Firewall.SimpleFw_toString Simple_Firewall.Service_Matrix "../Semantics_Ternary/Optimizing" (*do we use this?*) "../Semantics_Goto" Native_Word.Code_Target_Bits_Int "HOL-Library.Code_Target_Nat" "HOL-Library.Code_Target_Int" begin (*Note: common_primitive_match_expr_ipv4_toString can be really slow*) section\Code Interface\ text\HACK: rewrite quotes such that they are better printable by Isabelle\ definition quote_rewrite :: "string \ string" where "quote_rewrite \ map (\c. if c = char_of_nat 34 then CHR ''~'' else c)" lemma "quote_rewrite (''foo''@[char_of_nat 34]) = ''foo~''" by eval text\The parser returns the @{typ "'i::len common_primitive ruleset"} not as a map but as an association list. This function converts it\ (*this is only to tighten the types*) definition map_of_string_ipv4 :: "(string \ 32 common_primitive rule list) list \ string \ 32 common_primitive rule list" where "map_of_string_ipv4 rs = map_of rs" definition map_of_string_ipv6 :: "(string \ 128 common_primitive rule list) list \ string \ 128 common_primitive rule list" where "map_of_string_ipv6 rs = map_of rs" definition map_of_string :: "(string \ 'i common_primitive rule list) list \ string \ 'i common_primitive rule list" where "map_of_string rs = map_of rs" definition unfold_ruleset_CHAIN_safe :: "string \ action \ 'i::len common_primitive ruleset \ 'i common_primitive rule list option" where "unfold_ruleset_CHAIN_safe = unfold_optimize_ruleset_CHAIN optimize_primitive_univ" lemma "(unfold_ruleset_CHAIN_safe chain a rs = Some rs') \ simple_ruleset rs'" by(simp add: Let_def unfold_ruleset_CHAIN_safe_def unfold_optimize_ruleset_CHAIN_def split: if_split_asm) (*This is just for legacy code compatibility. Use the new _safe function instead*) definition unfold_ruleset_CHAIN :: "string \ action \ 'i::len common_primitive ruleset \ 'i common_primitive rule list" where "unfold_ruleset_CHAIN chain default_action rs = the (unfold_ruleset_CHAIN_safe chain default_action rs)" definition unfold_ruleset_FORWARD :: "action \ 'i::len common_primitive ruleset \ 'i::len common_primitive rule list" where "unfold_ruleset_FORWARD = unfold_ruleset_CHAIN ''FORWARD''" definition unfold_ruleset_INPUT :: "action \ 'i::len common_primitive ruleset \ 'i::len common_primitive rule list" where "unfold_ruleset_INPUT = unfold_ruleset_CHAIN ''INPUT''" definition unfold_ruleset_OUTPUT :: "action \ 'i::len common_primitive ruleset \ 'i::len common_primitive rule list" where "unfold_ruleset_OUTPUT \ unfold_ruleset_CHAIN ''OUTPUT''" lemma "let fw = [''FORWARD'' \ []] in unfold_ruleset_FORWARD action.Drop fw = [Rule (MatchAny :: 32 common_primitive match_expr) action.Drop]" by eval (* only used for ML/Haskell code to convert types *) definition nat_to_8word :: "nat \ 8 word" where "nat_to_8word i \ of_nat i" definition nat_to_16word :: "nat \ 16 word" where "nat_to_16word i \ of_nat i" definition integer_to_16word :: "integer \ 16 word" where "integer_to_16word i \ nat_to_16word (nat_of_integer i)" context begin private definition is_pos_Extra :: "'i::len common_primitive negation_type \ bool" where "is_pos_Extra a \ (case a of Pos (Extra _) \ True | _ \ False)" private definition get_pos_Extra :: "'i::len common_primitive negation_type \ string" where "get_pos_Extra a \ (case a of Pos (Extra e) \ e | _ \ undefined)" fun compress_parsed_extra :: "'i::len common_primitive negation_type list \ 'i common_primitive negation_type list" where "compress_parsed_extra [] = []" | "compress_parsed_extra (a1#a2#as) = (if is_pos_Extra a1 \ is_pos_Extra a2 then compress_parsed_extra (Pos (Extra (get_pos_Extra a1@'' ''@get_pos_Extra a2))#as) else a1#compress_parsed_extra (a2#as) )" | "compress_parsed_extra (a#as) = a#compress_parsed_extra as" lemma "compress_parsed_extra (map Pos [Extra ''-m'', (Extra ''recent'' :: 32 common_primitive), Extra ''--update'', Extra ''--seconds'', Extra ''60'', IIface (Iface ''foobar''), Extra ''--name'', Extra ''DEFAULT'', Extra ''--rsource'']) = map Pos [Extra ''-m recent --update --seconds 60'', IIface (Iface ''foobar''), Extra ''--name DEFAULT --rsource'']" by eval private lemma eval_ternary_And_Unknown_Unkown: "eval_ternary_And TernaryUnknown (eval_ternary_And TernaryUnknown tv) = eval_ternary_And TernaryUnknown tv" by(cases tv) (simp_all) private lemma is_pos_Extra_alist_and: "is_pos_Extra a \ alist_and (a#as) = MatchAnd (Match (Extra (get_pos_Extra a))) (alist_and as)" apply(cases a) apply(simp_all add: get_pos_Extra_def is_pos_Extra_def) apply(rename_tac e) by(case_tac e)(simp_all) private lemma compress_parsed_extra_matchexpr_helper: "ternary_ternary_eval (map_match_tac common_matcher p (alist_and (compress_parsed_extra as))) = ternary_ternary_eval (map_match_tac common_matcher p (alist_and as))" proof(induction as rule: compress_parsed_extra.induct) case 1 thus ?case by(simp) next case (2 a1 a2) thus ?case apply(simp add: is_pos_Extra_alist_and) apply(cases a1) apply(simp_all add: eval_ternary_And_Unknown_Unkown) done next case 3 thus ?case by(simp) qed text\This lemma justifies that it is okay to fold together the parsed unknown tokens\ lemma compress_parsed_extra_matchexpr: "matches (common_matcher, \) (alist_and (compress_parsed_extra as)) = matches (common_matcher, \) (alist_and as)" apply(simp add: fun_eq_iff) apply(intro allI) apply(rule matches_iff_apply_f) apply(simp add: compress_parsed_extra_matchexpr_helper) done end subsection\L4 Ports Parser Helper\ context begin text\Replace all matches on ports with the unspecified @{term 0} protocol with the given @{typ primitive_protocol}.\ private definition fill_l4_protocol_raw :: "primitive_protocol \ 'i::len common_primitive negation_type list \ 'i common_primitive negation_type list" where "fill_l4_protocol_raw protocol \ NegPos_map (\ m. case m of Src_Ports (L4Ports x pts) \ if x \ 0 then undefined else Src_Ports (L4Ports protocol pts) | Dst_Ports (L4Ports x pts) \ if x \ 0 then undefined else Dst_Ports (L4Ports protocol pts) | MultiportPorts (L4Ports x pts) \ if x \ 0 then undefined else MultiportPorts (L4Ports protocol pts) | Prot _ \ undefined \ \there should be no more match on the protocol if it was parsed from an iptables-save line\ | m \ m )" lemma "fill_l4_protocol_raw TCP [Neg (Dst (IpAddrNetmask (ipv4addr_of_dotdecimal (127, 0, 0, 0)) 8)), Pos (Src_Ports (L4Ports 0 [(22,22)]))] = [Neg (Dst (IpAddrNetmask 0x7F000000 8)), Pos (Src_Ports (L4Ports 6 [(0x16, 0x16)]))]" by eval fun fill_l4_protocol :: "'i::len common_primitive negation_type list \ 'i::len common_primitive negation_type list" where "fill_l4_protocol [] = []" | "fill_l4_protocol (Pos (Prot (Proto protocol)) # ms) = Pos (Prot (Proto protocol)) # fill_l4_protocol_raw protocol ms" | "fill_l4_protocol (Pos (Src_Ports _) # _) = undefined" | (*need to find proto first*) "fill_l4_protocol (Pos (Dst_Ports _) # _) = undefined" | "fill_l4_protocol (Pos (MultiportPorts _) # _) = undefined" | "fill_l4_protocol (Neg (Src_Ports _) # _) = undefined" | "fill_l4_protocol (Neg (Dst_Ports _) # _) = undefined" | "fill_l4_protocol (Neg (MultiportPorts _) # _) = undefined" | "fill_l4_protocol (m # ms) = m # fill_l4_protocol ms" lemma "fill_l4_protocol [ Neg (Dst (IpAddrNetmask (ipv4addr_of_dotdecimal (127, 0, 0, 0)) 8)) , Neg (Prot (Proto UDP)) , Pos (Src (IpAddrNetmask (ipv4addr_of_dotdecimal (127, 0, 0, 0)) 8)) , Pos (Prot (Proto TCP)) , Pos (Extra ''foo'') , Pos (Src_Ports (L4Ports 0 [(22,22)])) , Neg (Extra ''Bar'')] = [ Neg (Dst (IpAddrNetmask 0x7F000000 8)) , Neg (Prot (Proto UDP)) , Pos (Src (IpAddrNetmask 0x7F000000 8)) , Pos (Prot (Proto TCP)) , Pos (Extra ''foo'') , Pos (Src_Ports (L4Ports TCP [(0x16, 0x16)])) , Neg (Extra ''Bar'')]" by eval end (*currently unused and unverifed. may be needed for future use*) definition prefix_to_strange_inverse_cisco_mask:: "nat \ (nat \ nat \ nat \ nat)" where - "prefix_to_strange_inverse_cisco_mask n \ dotdecimal_of_ipv4addr ( (NOT (((mask n)::ipv4addr) << (32 - n))) )" + "prefix_to_strange_inverse_cisco_mask n \ dotdecimal_of_ipv4addr (Bit_Operations.not (mask n << 32 - n))" + lemma "prefix_to_strange_inverse_cisco_mask 8 = (0, 255, 255, 255)" by eval lemma "prefix_to_strange_inverse_cisco_mask 16 = (0, 0, 255, 255)" by eval lemma "prefix_to_strange_inverse_cisco_mask 24 = (0, 0, 0, 255)" by eval lemma "prefix_to_strange_inverse_cisco_mask 32 = (0, 0, 0, 0)" by eval end diff --git a/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy b/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy --- a/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy +++ b/thys/LOFT/LinuxRouter_OpenFlow_Translation.thy @@ -1,1057 +1,1057 @@ theory LinuxRouter_OpenFlow_Translation imports IP_Addresses.CIDR_Split Automatic_Refinement.Misc (*TODO@Peter: rename and make available at better place :)*) Simple_Firewall.Generic_SimpleFw Semantics_OpenFlow OpenFlow_Matches OpenFlow_Action Routing.Linux_Router "Pure-ex.Guess" begin hide_const Misc.uncurry hide_fact Misc.uncurry_def (* For reference: iiface :: "iface" --"in-interface" oiface :: "iface" --"out-interface" src :: "(ipv4addr \ nat) " --"source IP address" dst :: "(ipv4addr \ nat) " --"destination" proto :: "protocol" sports :: "(16 word \ 16 word)" --"source-port first:last" dports :: "(16 word \ 16 word)" --"destination-port first:last" p_iiface :: string p_oiface :: string p_src :: ipv4addr p_dst :: ipv4addr p_proto :: primitive_protocol p_sport :: "16 word" p_dport :: "16 word" p_tcp_flags :: "tcp_flag set" p_payload :: string *) definition "route2match r = \iiface = ifaceAny, oiface = ifaceAny, src = (0,0), dst=(pfxm_prefix (routing_match r),pfxm_length (routing_match r)), proto=ProtoAny, sports=(0,- 1), ports=(0,- 1)\" definition toprefixmatch where "toprefixmatch m \ (let pm = PrefixMatch (fst m) (snd m) in if pm = PrefixMatch 0 0 then None else Some pm)" lemma prefix_match_semantics_simple_match: assumes some: "toprefixmatch m = Some pm" assumes vld: "valid_prefix pm" shows "prefix_match_semantics pm = simple_match_ip m" using some by(cases m) (clarsimp simp add: toprefixmatch_def ipset_from_cidr_def pfxm_mask_def fun_eq_iff prefix_match_semantics_ipset_from_netmask[OF vld] NOT_mask_shifted_lenword[symmetric] split: if_splits) definition simple_match_to_of_match_single :: "(32, 'a) simple_match_scheme \ char list option \ protocol \ (16 word \ 16 word) option \ (16 word \ 16 word) option \ of_match_field set" where "simple_match_to_of_match_single m iif prot sport dport \ uncurry L4Src ` option2set sport \ uncurry L4Dst ` option2set dport \ IPv4Proto ` (case prot of ProtoAny \ {} | Proto p \ {p}) \ \protocol is an 8 word option anyway...\ \ IngressPort ` option2set iif \ IPv4Src ` option2set (toprefixmatch (src m)) \ IPv4Dst ` option2set (toprefixmatch (dst m)) \ {EtherType 0x0800}" (* okay, we need to make sure that no packets are output on the interface they were input on. So for rules that don't have an input interface, we'd need to do a product over all interfaces, if we stay naive. The more smart way would be to insert a rule with the same match condition that additionally matches the input interface and drops. However, I'm afraid this is going to be very tricky to verify\ *) definition simple_match_to_of_match :: "32 simple_match \ string list \ of_match_field set list" where "simple_match_to_of_match m ifs \ (let npm = (\p. fst p = 0 \ snd p = - 1); sb = (\p. (if npm p then [None] else if fst p \ snd p - then map (Some \ (\pfx. (pfxm_prefix pfx, NOT (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst p) (snd p))) else [])) + then map (Some \ (\pfx. (pfxm_prefix pfx, Bit_Operations.not (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst p) (snd p))) else [])) in [simple_match_to_of_match_single m iif (proto m) sport dport. iif \ (if iiface m = ifaceAny then [None] else [Some i. i \ ifs, match_iface (iiface m) i]), sport \ sb (sports m), dport \ sb (dports m)] )" (* I wonder\ should I check whether list_all (match_iface (iiface m)) ifs instead of iiface m = ifaceAny? It would be pretty stupid if that wasn't the same, but you know\ *) lemma smtoms_eq_hlp: "simple_match_to_of_match_single r a b c d = simple_match_to_of_match_single r f g h i \ (a = f \ b = g \ c = h \ d = i)" (* In case this proof breaks: there are two alternate proofs in the repo. They are of similar quality, though. Good luck. *) proof(rule iffI,goal_cases) case 1 thus ?case proof(intro conjI) have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (IngressPort x)" by simp show "a = f" using 1 by(cases a; cases f) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Proto x\ \ P (IPv4Proto x)" by simp show "b = g" using 1 by(cases b; cases g) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (uncurry L4Src x)" by simp show "c = h" using 1 by(cases c; cases h) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ next have *: "\P z x. \\x :: of_match_field. P x; z = Some x\ \ P (uncurry L4Dst x)" by simp show "d = i" using 1 by(cases d; cases i) (simp add: option2set_None simple_match_to_of_match_single_def toprefixmatch_def option2set_def; subst(asm) set_eq_iff; drule (1) *; simp split: option.splits uncurry_splits protocol.splits)+ qed qed simp lemma simple_match_to_of_match_generates_prereqs: "simple_match_valid m \ r \ set (simple_match_to_of_match m ifs) \ all_prerequisites r" unfolding simple_match_to_of_match_def Let_def proof(clarsimp, goal_cases) case (1 xiface xsrcp xdstp) note o = this show ?case unfolding simple_match_to_of_match_single_def all_prerequisites_def unfolding ball_Un proof((intro conjI; ((simp;fail)| - )), goal_cases) case 1 have e: "(fst (sports m) = 0 \ snd (sports m) = - 1) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" using o(1) unfolding simple_match_valid_alt Let_def by(clarsimp split: if_splits) show ?case using o(3) e by(elim disjE; simp add: option2set_def split: if_splits prod.splits uncurry_splits) next case 2 have e: "(fst (dports m) = 0 \ snd (dports m) = - 1) \ proto m = Proto TCP \ proto m = Proto UDP \ proto m = Proto L4_Protocol.SCTP" using o(1) unfolding simple_match_valid_alt Let_def by(clarsimp split: if_splits) show ?case using o(4) e by(elim disjE; simp add: option2set_def split: if_splits prod.splits uncurry_splits) qed qed lemma and_assoc: "a \ b \ c \ (a \ b) \ c" by simp lemmas custom_simpset = Let_def set_concat set_map map_map comp_def concat_map_maps set_maps UN_iff fun_app_def Set.image_iff abbreviation "simple_fw_prefix_to_wordinterval \ prefix_to_wordinterval \ uncurry PrefixMatch" lemma simple_match_port_alt: "simple_match_port m p \ p \ wordinterval_to_set (uncurry WordInterval m)" by(simp split: uncurry_splits) lemma simple_match_src_alt: "simple_match_valid r \ simple_match_ip (src r) p \ prefix_match_semantics (PrefixMatch (fst (src r)) (snd (src r))) p" by(cases "(src r)") (simp add: prefix_match_semantics_ipset_from_netmask2 prefix_to_wordset_ipset_from_cidr simple_match_valid_def valid_prefix_fw_def) lemma simple_match_dst_alt: "simple_match_valid r \ simple_match_ip (dst r) p \ prefix_match_semantics (PrefixMatch (fst (dst r)) (snd (dst r))) p" by(cases "(dst r)") (simp add: prefix_match_semantics_ipset_from_netmask2 prefix_to_wordset_ipset_from_cidr simple_match_valid_def valid_prefix_fw_def) lemma "x \ set (wordinterval_CIDR_split_prefixmatch w) \ valid_prefix x" using wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1] . lemma simple_match_to_of_matchI: assumes mv: "simple_match_valid r" assumes mm: "simple_matches r p" assumes ii: "p_iiface p \ set ifs" assumes ippkt: "p_l2type p = 0x800" shows eq: "\gr \ set (simple_match_to_of_match r ifs). OF_match_fields gr p = Some True" proof - let ?npm = "\p. fst p = 0 \ snd p = - 1" let ?sb = "\p r. (if ?npm p then None else Some r)" obtain si where si: "case si of Some ssi \ p_sport p \ prefix_to_wordset ssi | None \ True" "case si of None \ True | Some ssi \ ssi \ set ( wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))" "si = None \ ?npm (sports r)" proof(cases "?npm (sports r)", goal_cases) case 1 (* True *) hence "(case None of None \ True | Some ssi \ p_sport p \ prefix_to_wordset ssi) \ (case None of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r))))" by simp with 1 show ?thesis by blast next case 2 (* False *) from mm have "p_sport p \ wordinterval_to_set (uncurry WordInterval (sports r))" by(simp only: simple_matches.simps simple_match_port_alt) then obtain ssi where ssi: "ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))" "p_sport p \ prefix_to_wordset ssi" using wordinterval_CIDR_split_existential by fast hence "(case Some ssi of None \ True | Some ssi \ p_sport p \ prefix_to_wordset ssi) \ (case Some ssi of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r))))" by simp with 2 show ?thesis by blast qed obtain di where di: "case di of Some ddi \ p_dport p \ prefix_to_wordset ddi | None \ True" "case di of None \ True | Some ddi \ ddi \ set ( wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))" "di = None \ ?npm (dports r)" proof(cases "?npm (dports r)", goal_cases) case 1 hence "(case None of None \ True | Some ssi \ p_dport p \ prefix_to_wordset ssi) \ (case None of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r))))" by simp with 1 show ?thesis by blast next case 2 from mm have "p_dport p \ wordinterval_to_set (uncurry WordInterval (dports r))" by(simp only: simple_matches.simps simple_match_port_alt) then obtain ddi where ddi: "ddi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))" "p_dport p \ prefix_to_wordset ddi" using wordinterval_CIDR_split_existential by fast hence "(case Some ddi of None \ True | Some ssi \ p_dport p \ prefix_to_wordset ssi) \ (case Some ddi of None \ True | Some ssi \ ssi \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r))))" by simp with 2 show ?thesis by blast qed show ?thesis proof let ?mf = "map_option (apsnd (wordNOT \ mask \ (-) 16) \ prefix_match_dtor)" let ?gr = "simple_match_to_of_match_single r (if iiface r = ifaceAny then None else Some (p_iiface p)) (if proto r = ProtoAny then ProtoAny else Proto (p_proto p)) (?mf si) (?mf di)" note mfu = simple_match_port.simps[of "fst (sports r)" "snd (sports r)", unfolded surjective_pairing[of "sports r",symmetric]] simple_match_port.simps[of "fst (dports r)" "snd (dports r)", unfolded surjective_pairing[of "dports r",symmetric]] note u = mm[unfolded simple_matches.simps mfu ord_class.atLeastAtMost_iff simple_packet_unext_def simple_packet.simps] note of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs] from u have ple: "fst (sports r) \ snd (sports r)" "fst (dports r) \ snd (dports r)" by force+ show eg: "?gr \ set (simple_match_to_of_match r ifs)" unfolding simple_match_to_of_match_def unfolding custom_simpset unfolding smtoms_eq_hlp proof(intro bexI, (intro conjI; ((rule refl)?)), goal_cases) case 2 thus ?case using ple(2) di apply(simp add: pfxm_mask_def prefix_match_dtor_def Set.image_iff split: option.splits prod.splits uncurry_splits) apply(erule bexI[rotated]) apply(simp split: prefix_match.splits) done next case 3 thus ?case using ple(1) si apply(simp add: pfxm_mask_def prefix_match_dtor_def Set.image_iff split: option.splits prod.splits uncurry_splits) apply(erule bexI[rotated]) apply(simp split: prefix_match.splits) done next case 4 thus ?case using u ii by(clarsimp simp: set_maps split: if_splits) next case 1 thus ?case using ii u by simp_all (metis match_proto.elims(2)) qed have dpm: "di = Some (PrefixMatch x1 x2) \ p_dport p && ~~ (mask (16 - x2)) = x1" for x1 x2 proof - have *: "di = Some (PrefixMatch x1 x2) \ prefix_match_semantics (the di) (p_dport p) \ p_dport p && ~~ (mask (16 - x2)) = x1" by(clarsimp simp: prefix_match_semantics_def pfxm_mask_def word_bw_comms;fail) have **: "pfx \ set (wordinterval_CIDR_split_prefixmatch ra) \ prefix_match_semantics pfx a = (a \ prefix_to_wordset pfx)" for pfx ra and a :: "16 word" by (fact prefix_match_semantics_wordset[OF wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1]]) have "\di = Some (PrefixMatch x1 x2); p_dport p \ prefix_to_wordset (PrefixMatch x1 x2); PrefixMatch x1 x2 \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (dports r)))\ \ p_dport p && ~~ (mask (16 - x2)) = x1" using di(1,2) using * ** by auto thus "di = Some (PrefixMatch x1 x2) \ p_dport p && ~~ (mask (16 - x2)) = x1" using di(1,2) by auto qed have spm: "si = Some (PrefixMatch x1 x2) \ p_sport p && ~~ (mask (16 - x2)) = x1" for x1 x2 using si proof - have *: "si = Some (PrefixMatch x1 x2) \ prefix_match_semantics (the si) (p_sport p) \ p_sport p && ~~ (mask (16 - x2)) = x1" by(clarsimp simp: prefix_match_semantics_def pfxm_mask_def word_bw_comms;fail) have **: "pfx \ set (wordinterval_CIDR_split_prefixmatch ra) \ prefix_match_semantics pfx a = (a \ prefix_to_wordset pfx)" for pfx ra and a :: "16 word" by (fact prefix_match_semantics_wordset[OF wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1]]) have "\si = Some (PrefixMatch x1 x2); p_sport p \ prefix_to_wordset (PrefixMatch x1 x2); PrefixMatch x1 x2 \ set (wordinterval_CIDR_split_prefixmatch (uncurry WordInterval (sports r)))\ \ p_sport p && ~~ (mask (16 - x2)) = x1" using si(1,2) using * ** by auto thus "si = Some (PrefixMatch x1 x2) \ p_sport p && ~~ (mask (16 - x2)) = x1" using si(1,2) by auto qed show "OF_match_fields ?gr p = Some True" unfolding of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs[OF mv eg]] by(cases si; cases di) (simp_all add: simple_match_to_of_match_single_def OF_match_fields_unsafe_def spm option2set_def u ippkt prefix_match_dtor_def toprefixmatch_def dpm simple_match_dst_alt[OF mv, symmetric] simple_match_src_alt[OF mv, symmetric] split: prefix_match.splits) qed qed lemma prefix_match_00[simp,intro!]: "prefix_match_semantics (PrefixMatch 0 0) p" by (simp add: valid_prefix_def zero_prefix_match_all) lemma simple_match_to_of_matchD: assumes eg: "gr \ set (simple_match_to_of_match r ifs)" assumes mo: "OF_match_fields gr p = Some True" assumes me: "match_iface (oiface r) (p_oiface p)" assumes mv: "simple_match_valid r" shows "simple_matches r p" proof - from mv have validpfx: "valid_prefix (uncurry PrefixMatch (src r))" "valid_prefix (uncurry PrefixMatch (dst r))" "\pm. toprefixmatch (src r) = Some pm \ valid_prefix pm" "\pm. toprefixmatch (dst r) = Some pm \ valid_prefix pm" unfolding simple_match_valid_def valid_prefix_fw_def toprefixmatch_def by(simp_all split: uncurry_splits if_splits) from mo have mo: "OF_match_fields_unsafe gr p" unfolding of_safe_unsafe_match_eq[OF simple_match_to_of_match_generates_prereqs[OF mv eg]] by simp note this[unfolded OF_match_fields_unsafe_def] note eg[unfolded simple_match_to_of_match_def simple_match_to_of_match_single_def custom_simpset option2set_def] then guess x .. moreover from this(2) guess xa .. moreover from this(2) guess xb .. note xx = calculation(1,3) this { fix a b xc xa fix pp :: "16 word" have "\pp && ~~ (pfxm_mask xc) = pfxm_prefix xc\ \ prefix_match_semantics xc (pp)" for xc by(simp add: prefix_match_semantics_def word_bw_comms;fail) moreover have "pp \ wordinterval_to_set (WordInterval a b) \ a \ pp \ pp \ b" by simp moreover have "xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)) \ pp \ prefix_to_wordset xc \ pp \ wordinterval_to_set (WordInterval a b)" by(subst wordinterval_CIDR_split_prefixmatch) blast moreover have "\xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)); xa = Some (pfxm_prefix xc, ~~ (pfxm_mask xc)); prefix_match_semantics xc (pp)\ \ pp \ prefix_to_wordset xc" apply(subst(asm)(1) prefix_match_semantics_wordset) apply(erule wordinterval_CIDR_split_prefixmatch_all_valid_Ball[THEN bspec, THEN conjunct1];fail) apply assumption done ultimately have "\xc \ set (wordinterval_CIDR_split_prefixmatch (WordInterval a b)); xa = Some (pfxm_prefix xc, ~~ (pfxm_mask xc)); pp && ~~ (pfxm_mask xc) = pfxm_prefix xc\ \ a \ pp \ pp \ b" by metis } note l4port_logic = this show ?thesis unfolding simple_matches.simps proof(unfold and_assoc, (rule)+) show "match_iface (iiface r) (p_iiface p)" apply(cases "iiface r = ifaceAny") apply (simp add: match_ifaceAny) using xx(1) mo unfolding xx(4) OF_match_fields_unsafe_def apply(simp only: if_False set_maps UN_iff) apply(clarify) apply(rename_tac a; subgoal_tac "match_iface (iiface r) a") apply(clarsimp simp add: option2set_def;fail) apply(rule ccontr,simp;fail) done next show "match_iface (oiface r) (p_oiface p)" using me . next show "simple_match_ip (src r) (p_src p)" using mo unfolding xx(4) OF_match_fields_unsafe_def toprefixmatch_def by(clarsimp simp add: simple_packet_unext_def option2set_def validpfx simple_match_src_alt[OF mv] toprefixmatch_def split: if_splits) next show "simple_match_ip (dst r) (p_dst p)" using mo unfolding xx(4) OF_match_fields_unsafe_def toprefixmatch_def by(clarsimp simp add: simple_packet_unext_def option2set_def validpfx simple_match_dst_alt[OF mv] toprefixmatch_def split: if_splits) next show "match_proto (proto r) (p_proto p)" using mo unfolding xx(4) OF_match_fields_unsafe_def using xx(1) by(clarsimp simp add: singleton_iff simple_packet_unext_def option2set_def prefix_match_semantics_simple_match ball_Un split: if_splits protocol.splits) next show "simple_match_port (sports r) (p_sport p)" using mo xx(2) unfolding xx(4) OF_match_fields_unsafe_def by(cases "sports r") (clarsimp simp add: l4port_logic simple_packet_unext_def option2set_def prefix_match_semantics_simple_match split: if_splits) next show "simple_match_port (dports r) (p_dport p)" using mo xx(3) unfolding xx(4) OF_match_fields_unsafe_def by(cases "dports r") (clarsimp simp add: l4port_logic simple_packet_unext_def option2set_def prefix_match_semantics_simple_match split: if_splits) qed qed primrec annotate_rlen where "annotate_rlen [] = []" | "annotate_rlen (a#as) = (length as, a) # annotate_rlen as" lemma "annotate_rlen ''asdf'' = [(3, CHR ''a''), (2, CHR ''s''), (1, CHR ''d''), (0, CHR ''f'')]" by simp lemma fst_annotate_rlen_le: "(k, a) \ set (annotate_rlen l) \ k < length l" by(induction l arbitrary: k; simp; force) lemma distinct_fst_annotate_rlen: "distinct (map fst (annotate_rlen l))" using fst_annotate_rlen_le by(induction l) (simp, fastforce) lemma distinct_annotate_rlen: "distinct (annotate_rlen l)" using distinct_fst_annotate_rlen unfolding distinct_map by blast lemma in_annotate_rlen: "(a,x) \ set (annotate_rlen l) \ x \ set l" by(induction l) (simp_all, blast) lemma map_snd_annotate_rlen: "map snd (annotate_rlen l) = l" by(induction l) simp_all lemma "sorted_descending (map fst (annotate_rlen l))" by(induction l; clarsimp) (force dest: fst_annotate_rlen_le) lemma "annotate_rlen l = zip (rev [0.. *) primrec annotate_rlen_code where "annotate_rlen_code [] = (0,[])" | "annotate_rlen_code (a#as) = (case annotate_rlen_code as of (r,aas) \ (Suc r, (r, a) # aas))" lemma annotate_rlen_len: "fst (annotate_rlen_code r) = length r" by(induction r) (clarsimp split: prod.splits)+ lemma annotate_rlen_code[code]: "annotate_rlen s = snd (annotate_rlen_code s)" proof(induction s) case (Cons s ss) thus ?case using annotate_rlen_len[of ss] by(clarsimp split: prod.split) qed simp lemma suc2plus_inj_on: "inj_on (of_nat :: nat \ ('l :: len) word) {0..unat (max_word :: 'l word)}" proof(rule inj_onI) let ?mmw = "(max_word :: 'l word)" let ?mstp = "(of_nat :: nat \ 'l word)" fix x y :: nat assume "x \ {0..unat ?mmw}" "y \ {0..unat ?mmw}" hence se: "x \ unat ?mmw" "y \ unat ?mmw" by simp_all assume eq: "?mstp x = ?mstp y" note f = le_unat_uoi[OF se(1)] le_unat_uoi[OF se(2)] show "x = y" using eq le_unat_uoi se by metis qed lemma distinct_of_nat_list: (* TODO: Move to CaesarWordLemmaBucket *) "distinct l \ \e \ set l. e \ unat (max_word :: ('l::len) word) \ distinct (map (of_nat :: nat \ 'l word) l)" proof(induction l) let ?mmw = "(max_word :: 'l word)" let ?mstp = "(of_nat :: nat \ 'l word)" case (Cons a as) have "distinct as" "\e\set as. e \ unat ?mmw" using Cons.prems by simp_all note mIH = Cons.IH[OF this] moreover have "?mstp a \ ?mstp ` set as" proof have representable_set: "set as \ {0..unat ?mmw}" using \\e\set (a # as). e \ unat max_word\ by fastforce have a_reprbl: "a \ {0..unat ?mmw}" using \\e\set (a # as). e \ unat max_word\ by simp assume "?mstp a \ ?mstp ` set as" with inj_on_image_mem_iff[OF suc2plus_inj_on a_reprbl representable_set] have "a \ set as" by simp with \distinct (a # as)\ show False by simp qed ultimately show ?case by simp qed simp lemma annotate_first_le_hlp: "length l < unat (max_word :: ('l :: len) word) \ \e\set (map fst (annotate_rlen l)). e \ unat (max_word :: 'l word)" by(clarsimp) (meson fst_annotate_rlen_le less_trans nat_less_le) lemmas distinct_of_prio_hlp = distinct_of_nat_list[OF distinct_fst_annotate_rlen annotate_first_le_hlp] (* don't need these right now, but maybe later? *) lemma fst_annotate_rlen: "map fst (annotate_rlen l) = rev [0.. (of_nat :: nat \ ('l :: len) word)" assumes "length l \ unat (max_word :: 'l word)" shows "sorted_descending (map won (rev [0.. unat (max_word :: ('l :: len) word)" shows "sorted_descending (map fst (map (apfst (of_nat :: nat \ 'l word)) (annotate_rlen l)))" proof - let ?won = "(of_nat :: nat \ 'l word)" have "sorted_descending (map ?won (rev [0..l3 device to l2 forwarding\ definition "lr_of_tran_s3 ifs ard = ( [(p, b, case a of simple_action.Accept \ [Forward c] | simple_action.Drop \ []). (p,r,(c,a)) \ ard, b \ simple_match_to_of_match r ifs])" definition "oif_ne_iif_p1 ifs \ [(simple_match_any\oiface := Iface oif, iiface := Iface iif\, simple_action.Accept). oif \ ifs, iif \ ifs, oif \ iif]" definition "oif_ne_iif_p2 ifs = [(simple_match_any\oiface := Iface i, iiface := Iface i\, simple_action.Drop). i \ ifs]" definition "oif_ne_iif ifs = oif_ne_iif_p2 ifs @ oif_ne_iif_p1 ifs" (* order irrelephant *) (*value "oif_ne_iif [''a'', ''b'']"*) (* I first tried something like "oif_ne_iif ifs \ [(simple_match_any\oiface := Iface oi, iiface := Iface ii\, if oi = ii then simple_action.Drop else simple_action.Accept). oi \ ifs, ii \ ifs]", but making the statement I wanted with that was really tricky. Much easier to have the second element constant and do it separately. *) definition "lr_of_tran_s4 ard ifs \ generalized_fw_join ard (oif_ne_iif ifs)" definition "lr_of_tran_s1 rt = [(route2match r, output_iface (routing_action r)). r \ rt]" definition "lr_of_tran_fbs rt fw ifs \ let gfw = map simple_rule_dtor fw; \ \generalized simple fw, hopefully for FORWARD\ frt = lr_of_tran_s1 rt; \ \rt as fw\ prd = generalized_fw_join frt gfw in prd " definition "pack_OF_entries ifs ard \ (map (split3 OFEntry) (lr_of_tran_s3 ifs ard))" definition "no_oif_match \ list_all (\m. oiface (match_sel m) = ifaceAny)" definition "lr_of_tran rt fw ifs \ if \ (no_oif_match fw \ has_default_policy fw \ simple_fw_valid fw \ valid_prefixes rt \ has_default_route rt \ distinct ifs) then Inl ''Error in creating OpenFlow table: prerequisites not satisifed'' else ( let nrd = lr_of_tran_fbs rt fw ifs; ard = map (apfst of_nat) (annotate_rlen nrd) \ \give them a priority\ in if length nrd < unat (- 1 :: 16 word) then Inr (pack_OF_entries ifs ard) else Inl ''Error in creating OpenFlow table: priority number space exhausted'') " definition "is_iface_name i \ i \ [] \ \Iface.iface_name_is_wildcard i" definition "is_iface_list ifs \ distinct ifs \ list_all is_iface_name ifs" lemma max_16_word_max[simp]: "(a :: 16 word) \ 0xffff" proof - have "0xFFFF = (- 1 :: 16 word)" by simp then show ?thesis by (simp only:) simp qed lemma replicate_FT_hlp: "x \ 16 \ y \ 16 \ replicate (16 - x) False @ replicate x True = replicate (16 - y) False @ replicate y True \ x = y" proof - let ?ns = "{0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15,16}" assume "x \ 16 \ y \ 16" hence "x \ ?ns" "y \ ?ns" by(simp; presburger)+ moreover assume "replicate (16 - x) False @ replicate x True = replicate (16 - y) False @ replicate y True" ultimately show "x = y" by simp (elim disjE; simp_all add: numeral_eq_Suc) (* that's only 289 subgoals after the elim *) qed lemma mask_inj_hlp1: "inj_on (mask :: nat \ 16 word) {0..16}" proof(intro inj_onI, goal_cases) case (1 x y) from 1(3) have oe: "of_bl (replicate (16 - x) False @ replicate x True) = (of_bl (replicate (16 - y) False @ replicate y True) :: 16 word)" unfolding mask_bl of_bl_rep_False . have "\z. z \ 16 \ length (replicate (16 - z) False @ replicate z True) = 16" by auto with 1(1,2) have ps: "replicate (16 - x) False @ replicate x True \ {bl. length bl = LENGTH(16)}" " replicate (16 - y) False @ replicate y True \ {bl. length bl = LENGTH(16)}" by simp_all from inj_onD[OF word_bl.Abs_inj_on, OF oe ps] show ?case using 1(1,2) by(fastforce intro: replicate_FT_hlp) qed lemma distinct_simple_match_to_of_match_portlist_hlp: fixes ps :: "(16 word \ 16 word)" shows "distinct ifs \ distinct (if fst ps = 0 \ snd ps = max_word then [None] else if fst ps \ snd ps then map (Some \ (\pfx. (pfxm_prefix pfx, ~~ (pfxm_mask pfx)))) (wordinterval_CIDR_split_prefixmatch (WordInterval (fst ps) (snd ps))) else [])" proof - assume di: "distinct ifs" { define wis where "wis = set (wordinterval_CIDR_split_prefixmatch (WordInterval (fst ps) (snd ps)))" fix x y :: "16 prefix_match" obtain xm xn ym yn where xyd[simp]: "x = PrefixMatch xm xn" "y = PrefixMatch ym yn" by(cases x; cases y) assume iw: "x \ wis" "y \ wis" and et: "(pfxm_prefix x, ~~ (pfxm_mask x)) = (pfxm_prefix y, ~~ (pfxm_mask y))" hence le16: "xn \ 16" "yn \ 16" unfolding wis_def using wordinterval_CIDR_split_prefixmatch_all_valid_Ball[unfolded Ball_def, THEN spec, THEN mp] by force+ with et have "16 - xn = 16 - yn" unfolding pfxm_mask_def by(auto intro: mask_inj_hlp1[THEN inj_onD]) hence "x = y" using et le16 using diff_diff_cancel by simp } note * = this show ?thesis apply(clarsimp simp add: smtoms_eq_hlp distinct_map wordinterval_CIDR_split_distinct) apply(subst comp_inj_on_iff[symmetric]; intro inj_onI) using * by simp_all qed lemma distinct_simple_match_to_of_match: "distinct ifs \ distinct (simple_match_to_of_match m ifs)" apply(unfold simple_match_to_of_match_def Let_def) apply(rule distinct_3lcomprI) subgoal by(induction ifs; clarsimp) subgoal by(fact distinct_simple_match_to_of_match_portlist_hlp) subgoal by(fact distinct_simple_match_to_of_match_portlist_hlp) subgoal by(simp_all add: smtoms_eq_hlp) done lemma inj_inj_on: "inj F \ inj_on F A" using subset_inj_on by auto (* TODO: include Word_Lib *) lemma no_overlaps_lroft_hlp2: "distinct (map fst amr) \ (\r. distinct (fm r)) \ distinct (concat (map (\(p, r, c, a). map (\b. (p, b, fs a c)) (fm r)) amr))" by(induction amr; force intro: injI inj_onI simp add: distinct_map split: prod.splits) lemma distinct_lroft_s3: "\distinct (map fst amr); distinct ifs\ \ distinct (lr_of_tran_s3 ifs amr)" unfolding lr_of_tran_s3_def by(erule no_overlaps_lroft_hlp2, simp add: distinct_simple_match_to_of_match) lemma no_overlaps_lroft_hlp3: "distinct (map fst amr) \ (aa, ab, ac) \ set (lr_of_tran_s3 ifs amr) \ (ba, bb, bc) \ set (lr_of_tran_s3 ifs amr) \ ac \ bc \ aa \ ba" apply(unfold lr_of_tran_s3_def) apply(clarsimp) apply(clarsimp split: simple_action.splits) apply(metis map_of_eq_Some_iff old.prod.inject option.inject) apply(metis map_of_eq_Some_iff old.prod.inject option.inject simple_action.distinct(2))+ done lemma no_overlaps_lroft_s3_hlp_hlp: (* I hlps *) "\distinct (map fst amr); OF_match_fields_unsafe ab p; ab \ ad \ ba \ bb; OF_match_fields_unsafe ad p; (ac, ab, ba) \ set (lr_of_tran_s3 ifs amr); (ac, ad, bb) \ set (lr_of_tran_s3 ifs amr)\ \ False" proof(elim disjE, goal_cases) case 1 have 4: "\distinct (map fst amr); (ac, ab, x1, x2) \ set amr; (ac, bb, x4, x5) \ set amr; ab \ bb\ \ False" for ab x1 x2 bb x4 x5 by (meson distinct_map_fstD old.prod.inject) have conjunctSomeProtoAnyD: "Some ProtoAny = simple_proto_conjunct a (Proto b) \ False" for a b using conjunctProtoD by force have 5: "\OF_match_fields_unsafe am p; OF_match_fields_unsafe bm p; am \ bm; am \ set (simple_match_to_of_match ab ifs); bm \ set (simple_match_to_of_match bb ifs); \ ab \ bb\ \ False" for ab bb am bm by(clarify | unfold simple_match_to_of_match_def smtoms_eq_hlp Let_def set_concat set_map de_Morgan_conj not_False_eq_True)+ (auto dest: conjunctSomeProtoAnyD cidrsplit_no_overlaps simp add: OF_match_fields_unsafe_def simple_match_to_of_match_single_def option2set_def comp_def split: if_splits cong: smtoms_eq_hlp) (*1min*) from 1 show ?case using 4 5 by(clarsimp simp add: lr_of_tran_s3_def) blast qed(metis no_overlaps_lroft_hlp3) lemma no_overlaps_lroft_s3_hlp: "distinct (map fst amr) \ distinct ifs \ no_overlaps OF_match_fields_unsafe (map (split3 OFEntry) (lr_of_tran_s3 ifs amr))" apply(rule no_overlapsI[rotated]) apply(subst distinct_map, rule conjI) subgoal by(erule (1) distinct_lroft_s3) subgoal apply(rule inj_inj_on) apply(rule injI) apply(rename_tac x y, case_tac x, case_tac y) apply(simp add: split3_def;fail) done subgoal apply(unfold check_no_overlap_def) apply(clarify) apply(unfold set_map) apply(clarify) apply(unfold split3_def prod.simps flow_entry_match.simps flow_entry_match.sel de_Morgan_conj) apply(clarsimp simp only:) apply(erule (1) no_overlaps_lroft_s3_hlp_hlp) apply simp apply assumption apply assumption apply simp done done lemma lr_of_tran_no_overlaps: assumes "distinct ifs" shows "Inr t = (lr_of_tran rt fw ifs) \ no_overlaps OF_match_fields_unsafe t" apply(unfold lr_of_tran_def Let_def pack_OF_entries_def) apply(simp split: if_splits) apply(thin_tac "t = _") apply(drule distinct_of_prio_hlp) apply(rule no_overlaps_lroft_s3_hlp[rotated]) subgoal by(simp add: assms) subgoal by(simp add: o_assoc) done lemma sorted_lr_of_tran_s3_hlp: "\x\set f. fst x \ a \ b \ set (lr_of_tran_s3 s f) \ fst b \ a" by(auto simp add: lr_of_tran_s3_def) lemma lr_of_tran_s3_Cons: "lr_of_tran_s3 ifs (a#ard) = ( [(p, b, case a of simple_action.Accept \ [Forward c] | simple_action.Drop \ []). (p,r,(c,a)) \ [a], b \ simple_match_to_of_match r ifs]) @ lr_of_tran_s3 ifs ard" by(clarsimp simp: lr_of_tran_s3_def) lemma sorted_lr_of_tran_s3: "sorted_descending (map fst f) \ sorted_descending (map fst (lr_of_tran_s3 s f))" apply(induction f) subgoal by(simp add: lr_of_tran_s3_def) apply(clarsimp simp: lr_of_tran_s3_Cons map_concat comp_def) apply(unfold sorted_descending_append) apply(simp add: sorted_descending_alt rev_map sorted_lr_of_tran_s3_hlp sorted_const) done lemma sorted_lr_of_tran_hlp: "(ofe_prio \ split3 OFEntry) = fst" by(simp add: fun_eq_iff comp_def split3_def) lemma lr_of_tran_sorted_descending: "Inr r = lr_of_tran rt fw ifs \ sorted_descending (map ofe_prio r)" apply(unfold lr_of_tran_def Let_def) apply(simp split: if_splits) apply(thin_tac "r = _") apply(unfold sorted_lr_of_tran_hlp pack_OF_entries_def split3_def[abs_def] fun_app_def map_map comp_def prod.case_distrib) apply(simp add: fst_def[symmetric]) apply(rule sorted_lr_of_tran_s3) apply(drule sorted_annotated[OF less_or_eq_imp_le, OF disjI1]) apply(simp add: o_assoc) done lemma lr_of_tran_s1_split: "lr_of_tran_s1 (a # rt) = (route2match a, output_iface (routing_action a)) # lr_of_tran_s1 rt" by(unfold lr_of_tran_s1_def list.map, rule) lemma route2match_correct: "valid_prefix (routing_match a) \ prefix_match_semantics (routing_match a) (p_dst p) \ simple_matches (route2match a) (p)" by(simp add: route2match_def simple_matches.simps match_ifaceAny match_iface_refl ipset_from_cidr_0 prefix_match_semantics_ipset_from_netmask2) lemma s1_correct: "valid_prefixes rt \ has_default_route (rt::('i::len) prefix_routing) \ \rm ra. generalized_sfw (lr_of_tran_s1 rt) p = Some (rm,ra) \ ra = output_iface (routing_table_semantics rt (p_dst p))" apply(induction rt) apply(simp;fail) apply(drule valid_prefixes_split) apply(clarsimp) apply(erule disjE) subgoal for a rt apply(case_tac a) apply(rename_tac routing_m metric routing_action) apply(case_tac routing_m) apply(simp add: valid_prefix_def pfxm_mask_def prefix_match_semantics_def generalized_sfw_def lr_of_tran_s1_def route2match_def simple_matches.simps match_ifaceAny match_iface_refl ipset_from_cidr_0 max_word_mask[where 'a = 'i, symmetric, simplified]) done subgoal apply(rule conjI) apply(simp add: generalized_sfw_def lr_of_tran_s1_def route2match_correct;fail) apply(simp add: route2match_def simple_matches.simps prefix_match_semantics_ipset_from_netmask2 lr_of_tran_s1_split generalized_sfw_simps) done done definition "to_OF_action a \ (case a of (p,d) \ (case d of simple_action.Accept \ [Forward p] | simple_action.Drop \ []))" definition "from_OF_action a = (case a of [] \ ('''',simple_action.Drop) | [Forward p] \ (p, simple_action.Accept))" lemma OF_match_linear_not_noD: "OF_match_linear \ oms p \ NoAction \ \ome. ome \ set oms \ \ (ofe_fields ome) p" apply(induction oms) apply(simp) apply(simp split: if_splits) apply blast+ done lemma s3_noaction_hlp: "\simple_match_valid ac; \simple_matches ac p; match_iface (oiface ac) (p_oiface p)\ \ OF_match_linear OF_match_fields_safe (map (\x. split3 OFEntry (x1, x, case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ [])) (simple_match_to_of_match ac ifs)) p = NoAction" apply(rule ccontr) apply(drule OF_match_linear_not_noD) apply(clarsimp) apply(rename_tac x) apply(subgoal_tac "all_prerequisites x") apply(drule simple_match_to_of_matchD) apply(simp add: split3_def) apply(subst(asm) of_match_fields_safe_eq2) apply(simp;fail)+ using simple_match_to_of_match_generates_prereqs by blast lemma aux: \v = Some x \ the v = x\ by simp lemma s3_correct: assumes vsfwm: "list_all simple_match_valid (map (fst \ snd) ard)" assumes ippkt: "p_l2type p = 0x800" assumes iiifs: "p_iiface p \ set ifs" assumes oiifs: "list_all (\m. oiface (fst (snd m)) = ifaceAny) ard" shows "OF_match_linear OF_match_fields_safe (pack_OF_entries ifs ard) p = Action ao \ (\r af. generalized_sfw (map snd ard) p = (Some (r,af)) \ (if snd af = simple_action.Drop then ao = [] else ao = [Forward (fst af)]))" unfolding pack_OF_entries_def lr_of_tran_s3_def fun_app_def using vsfwm oiifs apply(induction ard) subgoal by(simp add: generalized_sfw_simps) apply simp apply(clarsimp simp add: generalized_sfw_simps split: prod.splits) apply(intro conjI) (* make two subgoals from one *) subgoal for ard x1 ac ad ba apply(clarsimp simp add: OF_match_linear_append split: prod.splits) apply(drule simple_match_to_of_matchI[rotated]) apply(rule iiifs) apply(rule ippkt) apply blast apply(clarsimp simp add: comp_def) apply(drule OF_match_linear_match_allsameaction[where \=OF_match_fields_safe and pri = x1 and oms = "simple_match_to_of_match ac ifs" and act = "case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ []"]) apply(unfold OF_match_fields_safe_def comp_def) apply(erule aux) apply(clarsimp) apply(intro iffI) subgoal apply(rule exI[where x = ac]) apply(rule exI[where x = ad]) apply(rule exI[where x = ba]) apply(clarsimp simp: split3_def split: simple_action.splits flowtable_behavior.splits if_splits) done subgoal apply(clarsimp) apply(rename_tac b) apply(case_tac b) apply(simp_all) done done subgoal for ard x1 ac ad ba apply(simp add: OF_match_linear_append OF_match_fields_safe_def comp_def) apply(clarify) apply(subgoal_tac "OF_match_linear OF_match_fields_safe (map (\x. split3 OFEntry (x1, x, case ba of simple_action.Accept \ [Forward ad] | simple_action.Drop \ [])) (simple_match_to_of_match ac ifs)) p = NoAction") apply(simp;fail) apply(erule (1) s3_noaction_hlp) apply(simp add: match_ifaceAny;fail) done done context notes valid_prefix_00[simp, intro!] begin lemma lr_of_tran_s1_valid: "valid_prefixes rt \ gsfw_valid (lr_of_tran_s1 rt)" unfolding lr_of_tran_s1_def route2match_def gsfw_valid_def list_all_iff apply(clarsimp simp: simple_match_valid_def valid_prefix_fw_def) apply(intro conjI) apply force apply(simp add: valid_prefixes_alt_def) done end lemma simple_match_valid_fbs_rlen: "\valid_prefixes rt; simple_fw_valid fw; (a, aa, ab, b) \ set (annotate_rlen (lr_of_tran_fbs rt fw ifs))\ \ simple_match_valid aa" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def] have "gsfw_valid (map simple_rule_dtor fw)" using gsfw_validI 1 by blast moreover have "gsfw_valid (lr_of_tran_s1 rt)" using 1 lr_of_tran_s1_valid by blast ultimately have "gsfw_valid (generalized_fw_join (lr_of_tran_s1 rt) (map simple_rule_dtor fw))" using gsfw_join_valid by blast moreover have "(aa, ab, b) \ set (lr_of_tran_fbs rt fw ifs)" using 1 using in_annotate_rlen by fast ultimately show ?thesis unfolding lr_of_tran_fbs_def Let_def gsfw_valid_def list_all_iff by fastforce qed lemma simple_match_valid_fbs: "\valid_prefixes rt; simple_fw_valid fw\ \ list_all simple_match_valid (map fst (lr_of_tran_fbs rt fw ifs))" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def] have "gsfw_valid (map simple_rule_dtor fw)" using gsfw_validI 1 by blast moreover have "gsfw_valid (lr_of_tran_s1 rt)" using 1 lr_of_tran_s1_valid by blast ultimately have "gsfw_valid (generalized_fw_join (lr_of_tran_s1 rt) (map simple_rule_dtor fw))" using gsfw_join_valid by blast thus ?thesis unfolding lr_of_tran_fbs_def Let_def gsfw_valid_def list_all_iff by fastforce qed lemma lr_of_tran_prereqs: "valid_prefixes rt \ simple_fw_valid fw \ lr_of_tran rt fw ifs = Inr oft \ list_all (all_prerequisites \ ofe_fields) oft" unfolding lr_of_tran_def pack_OF_entries_def lr_of_tran_s3_def Let_def apply(simp add: map_concat comp_def prod.case_distrib split3_def split: if_splits) apply(simp add: list_all_iff) apply(clarsimp) apply(drule simple_match_valid_fbs_rlen[rotated]) apply(simp add: list_all_iff;fail) apply(simp add: list_all_iff;fail) apply(rule simple_match_to_of_match_generates_prereqs; assumption) done (* TODO: move. where? *) lemma OF_unsafe_safe_match3_eq: " list_all (all_prerequisites \ ofe_fields) oft \ OF_priority_match OF_match_fields_unsafe oft = OF_priority_match OF_match_fields_safe oft" unfolding OF_priority_match_def[abs_def] proof(goal_cases) case 1 from 1 have "\packet. [f\oft . OF_match_fields_unsafe (ofe_fields f) packet] = [f\oft . OF_match_fields_safe (ofe_fields f) packet]" apply(clarsimp simp add: list_all_iff of_match_fields_safe_eq) using of_match_fields_safe_eq by(metis (mono_tags, lifting) filter_cong) thus ?case by metis qed lemma OF_unsafe_safe_match_linear_eq: " list_all (all_prerequisites \ ofe_fields) oft \ OF_match_linear OF_match_fields_unsafe oft = OF_match_linear OF_match_fields_safe oft" unfolding fun_eq_iff by(induction oft) (clarsimp simp add: list_all_iff of_match_fields_safe_eq)+ lemma simple_action_ne[simp]: "b \ simple_action.Accept \ b = simple_action.Drop" "b \ simple_action.Drop \ b = simple_action.Accept" using simple_action.exhaust by blast+ lemma map_snd_apfst: "map snd (map (apfst x) l) = map snd l" unfolding map_map comp_def snd_apfst .. lemma match_ifaceAny_eq: "oiface m = ifaceAny \ simple_matches m p = simple_matches m (p\p_oiface := any\)" by(cases m) (simp add: simple_matches.simps match_ifaceAny) lemma no_oif_matchD: "no_oif_match fw \ simple_fw fw p = simple_fw fw (p\p_oiface := any\)" by(induction fw) (auto simp add: no_oif_match_def simple_fw_alt dest: match_ifaceAny_eq) lemma lr_of_tran_fbs_acceptD: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" shows "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Accept) \ simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\)" proof(goal_cases) case 1 note 1[unfolded lr_of_tran_fbs_def Let_def, THEN generalized_fw_joinD] then guess r1 .. then guess r2 .. note r12 = this note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this from r12 rmra have oifra: "oif = ra" by simp from r12 have sfw: "simple_fw fw p = Decision FinalAllow" using simple_fw_iff_generalized_fw_accept by blast note ifupdateirrel = no_oif_matchD[OF s2, where any = " output_iface (routing_table_semantics rt (p_dst p))" and p = p, symmetric] show ?case unfolding simple_linux_router_nol12_def by(simp add: Let_def ifupdateirrel sfw oifra rmra split: Option.bind_splits option.splits) qed lemma lr_of_tran_fbs_acceptI: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" "has_default_policy fw" shows "simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ \r. generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Accept)" proof(goal_cases) from s2 have nud: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy state.distinct(1)) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1 have "simple_fw fw p = Decision FinalAllow" by(simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) then obtain r where r: "generalized_sfw (map simple_rule_dtor fw) p = Some (r, simple_action.Accept)" using simple_fw_iff_generalized_fw_accept by blast have oif_def: "oif = output_iface (routing_table_semantics rt (p_dst p))" using 1 by(cases p) (simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this show ?case unfolding lr_of_tran_fbs_def Let_def apply(rule exI) apply(rule generalized_fw_joinI) unfolding oif_def using rmra apply simp apply(rule r) done qed lemma lr_of_tran_fbs_dropD: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" shows "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop) \ simple_linux_router_nol12 rt fw p = None" proof(goal_cases) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1[unfolded lr_of_tran_fbs_def Let_def, THEN generalized_fw_joinD] obtain rr fr where "generalized_sfw (lr_of_tran_s1 rt) p = Some (rr, oif) \ generalized_sfw (map simple_rule_dtor fw) p = Some (fr, simple_action.Drop) \ Some r = simple_match_and rr fr" by presburger hence fd: "\u. simple_fw fw (p\p_oiface := u\) = Decision FinalDeny" unfolding ifupdateirrel using simple_fw_iff_generalized_fw_drop by blast show ?thesis by(clarsimp simp: simple_linux_router_nol12_def Let_def fd split: Option.bind_splits) qed lemma lr_of_tran_fbs_dropI: assumes s1: "valid_prefixes rt" "has_default_route rt" assumes s2: "no_oif_match fw" "has_default_policy fw" shows "simple_linux_router_nol12 rt fw p = None \ \r oif. generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" proof(goal_cases) from s2 have nud: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy state.distinct(1)) note ifupdateirrel = no_oif_matchD[OF s2(1), symmetric] case 1 from 1 have "simple_fw fw p = Decision FinalDeny" by(simp add: simple_linux_router_nol12_def Let_def nud ifupdateirrel split: Option.bind_splits state.splits final_decision.splits) then obtain r where r: "generalized_sfw (map simple_rule_dtor fw) p = Some (r, simple_action.Drop)" using simple_fw_iff_generalized_fw_drop by blast note s1_correct[OF s1, of p] then guess rm .. then guess ra .. note rmra = this show ?case unfolding lr_of_tran_fbs_def Let_def apply(rule exI) apply(rule exI[where x = ra]) apply(rule generalized_fw_joinI) using rmra apply simp apply(rule r) done qed lemma no_oif_match_fbs: "no_oif_match fw \ list_all (\m. oiface (fst (snd m)) = ifaceAny) (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))" proof(goal_cases) case 1 have c: "\mr ar mf af f a. \(mr, ar) \ set (lr_of_tran_s1 rt); (mf, af) \ simple_rule_dtor ` set fw; simple_match_and mr mf = Some a\ \ oiface a = ifaceAny" proof(goal_cases) case (1 mr ar mf af f a) have "oiface mr = ifaceAny" using 1(1) unfolding lr_of_tran_s1_def route2match_def by(clarsimp simp add: Set.image_iff) moreover have "oiface mf = ifaceAny" using 1(2) \no_oif_match fw\ unfolding no_oif_match_def simple_rule_dtor_def[abs_def] by(clarsimp simp: list_all_iff split: simple_rule.splits) fastforce ultimately show ?case using 1(3) by(cases a; cases mr; cases mf) (simp add: iface_conjunct_ifaceAny split: option.splits) qed have la: "list_all (\m. oiface (fst m) = ifaceAny) (lr_of_tran_fbs rt fw ifs)" unfolding lr_of_tran_fbs_def Let_def list_all_iff apply(clarify) apply(subst(asm) generalized_sfw_join_set) apply(clarsimp) using c by blast thus ?case proof(goal_cases) case 1 have *: "(\m. oiface (fst (snd m)) = ifaceAny) = (\m. oiface (fst m) = ifaceAny) \ snd" unfolding comp_def .. show ?case unfolding * list_all_map[symmetric] map_snd_apfst map_snd_annotate_rlen using la . qed qed lemma lr_of_tran_correct: fixes p :: "(32, 'a) simple_packet_ext_scheme" assumes nerr: "lr_of_tran rt fw ifs = Inr oft" and ippkt: "p_l2type p = 0x800" and ifvld: "p_iiface p \ set ifs" shows "OF_priority_match OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = (Some (p\p_oiface := oif\))" "OF_priority_match OF_match_fields_safe oft p = Action [] \ simple_linux_router_nol12 rt fw p = None" (* fun stuff: *) "OF_priority_match OF_match_fields_safe oft p \ NoAction" "OF_priority_match OF_match_fields_safe oft p \ Undefined" "OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" proof - have s1: "valid_prefixes rt" "has_default_route rt" and s2: "has_default_policy fw" "simple_fw_valid fw" "no_oif_match fw" and difs: "distinct ifs" using nerr unfolding lr_of_tran_def by(simp_all split: if_splits) have "no_oif_match fw" using nerr unfolding lr_of_tran_def by(simp split: if_splits) note s2 = s2 this have unsafe_safe_eq: "OF_priority_match OF_match_fields_unsafe oft = OF_priority_match OF_match_fields_safe oft" "OF_match_linear OF_match_fields_unsafe oft = OF_match_linear OF_match_fields_safe oft" apply(subst OF_unsafe_safe_match3_eq; (rule lr_of_tran_prereqs s1 s2 nerr refl)+) apply(subst OF_unsafe_safe_match_linear_eq; (rule lr_of_tran_prereqs s1 s2 nerr refl)+) done have lin: "OF_priority_match OF_match_fields_safe oft = OF_match_linear OF_match_fields_safe oft" using OF_eq[OF lr_of_tran_no_overlaps lr_of_tran_sorted_descending, OF difs nerr[symmetric] nerr[symmetric]] unfolding fun_eq_iff unsafe_safe_eq by metis let ?ard = "map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs))" have oft_def: "oft = pack_OF_entries ifs ?ard" using nerr unfolding lr_of_tran_def Let_def by(simp split: if_splits) have vld: "list_all simple_match_valid (map (fst \ snd) ?ard)" unfolding fun_app_def map_map[symmetric] snd_apfst map_snd_apfst map_snd_annotate_rlen using simple_match_valid_fbs[OF s1(1) s2(2)] . have *: "list_all (\m. oiface (fst (snd m)) = ifaceAny) ?ard" using no_oif_match_fbs[OF s2(3)] . have not_undec: "\p. simple_fw fw p \ Undecided" by (metis has_default_policy s2(1) state.simps(3)) have w1_1: "\oif. OF_match_linear OF_match_fields_safe oft p = Action [Forward oif] \ simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ oif = output_iface (routing_table_semantics rt (p_dst p))" proof(intro conjI, goal_cases) case (1 oif) note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD1, unfolded oft_def[symmetric], OF 1] hence "\r. generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, (oif, simple_action.Accept))" by(clarsimp split: if_splits) then obtain r where "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, (oif, simple_action.Accept))" unfolding map_map comp_def snd_apfst map_snd_annotate_rlen by blast thus ?case using lr_of_tran_fbs_acceptD[OF s1 s2(3)] by metis thus "oif = output_iface (routing_table_semantics rt (p_dst p))" by(cases p) (clarsimp simp: simple_linux_router_nol12_def Let_def not_undec split: Option.bind_splits state.splits final_decision.splits) qed have w1_2: "\oif. simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\) \ OF_match_linear OF_match_fields_safe oft p = Action [Forward oif]" proof(goal_cases) case (1 oif) note lr_of_tran_fbs_acceptI[OF s1 s2(3) s2(1) this, of ifs] then guess r .. note r = this hence "generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, (oif, simple_action.Accept))" unfolding map_snd_apfst map_snd_annotate_rlen . moreover note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD2, unfolded oft_def[symmetric], of "[Forward oif]"] ultimately show ?case by simp qed show w1: "\oif. (OF_priority_match OF_match_fields_safe oft p = Action [Forward oif]) = (simple_linux_router_nol12 rt fw p = Some (p\p_oiface := oif\))" unfolding lin using w1_1 w1_2 by blast show w2: "(OF_priority_match OF_match_fields_safe oft p = Action []) = (simple_linux_router_nol12 rt fw p = None)" unfolding lin proof(rule iffI, goal_cases) case 1 note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD1, unfolded oft_def[symmetric], OF 1] then obtain r oif where roif: "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" unfolding map_snd_apfst map_snd_annotate_rlen by(clarsimp split: if_splits) note lr_of_tran_fbs_dropD[OF s1 s2(3) this] thus ?case . next case 2 note lr_of_tran_fbs_dropI[OF s1 s2(3) s2(1) this, of ifs] then obtain r oif where "generalized_sfw (lr_of_tran_fbs rt fw ifs) p = Some (r, oif, simple_action.Drop)" by blast hence "generalized_sfw (map snd (map (apfst of_nat) (annotate_rlen (lr_of_tran_fbs rt fw ifs)))) p = Some (r, oif, simple_action.Drop)" unfolding map_snd_apfst map_snd_annotate_rlen . moreover note s3_correct[OF vld ippkt ifvld(1) *, THEN iffD2, unfolded oft_def[symmetric], of "[]"] ultimately show ?case by force qed have lr_determ: "\a. simple_linux_router_nol12 rt fw p = Some a \ a = p\p_oiface := output_iface (routing_table_semantics rt (p_dst p))\" by(clarsimp simp: simple_linux_router_nol12_def Let_def not_undec split: Option.bind_splits state.splits final_decision.splits) show notno: "OF_priority_match OF_match_fields_safe oft p \ NoAction" apply(cases "simple_linux_router_nol12 rt fw p") using w2 apply(simp) using w1[of "output_iface (routing_table_semantics rt (p_dst p))"] apply(simp) apply(drule lr_determ) apply(simp) done show notub: "OF_priority_match OF_match_fields_safe oft p \ Undefined" unfolding lin using OF_match_linear_ne_Undefined . show notmult: "\ls. OF_priority_match OF_match_fields_safe oft p = Action ls \ length ls \ 1" apply(cases "simple_linux_router_nol12 rt fw p") using w2 apply(simp) using w1[of "output_iface (routing_table_semantics rt (p_dst p))"] apply(simp) apply(drule lr_determ) apply(clarsimp) done show "\ls. length ls \ 1 \ OF_priority_match OF_match_fields_safe oft p = Action ls" apply(cases "OF_priority_match OF_match_fields_safe oft p") using notmult apply blast using notno apply blast using notub apply blast done qed end diff --git a/thys/Native_Word/Bits_Integer.thy b/thys/Native_Word/Bits_Integer.thy --- a/thys/Native_Word/Bits_Integer.thy +++ b/thys/Native_Word/Bits_Integer.thy @@ -1,680 +1,680 @@ (* Title: Bits_Integer.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Bit operations for target language integers\ theory Bits_Integer imports "Word_Lib.Bit_Comprehension" Code_Int_Integer_Conversion Code_Symbolic_Bits_Int begin lemmas [transfer_rule] = identity_quotient fun_quotient Quotient_integer[folded integer.pcr_cr_eq] lemma undefined_transfer: assumes "Quotient R Abs Rep T" shows "T (Rep undefined) undefined" using assms unfolding Quotient_alt_def by blast bundle undefined_transfer = undefined_transfer[transfer_rule] section \More lemmas about @{typ integer}s\ context includes integer.lifting begin lemma bitval_integer_transfer [transfer_rule]: "(rel_fun (=) pcr_integer) of_bool of_bool" by(auto simp add: of_bool_def integer.pcr_cr_eq cr_integer_def) lemma integer_of_nat_less_0_conv [simp]: "\ integer_of_nat n < 0" by(transfer) simp lemma int_of_integer_pow: "int_of_integer (x ^ n) = int_of_integer x ^ n" by(induct n) simp_all lemma pow_integer_transfer [transfer_rule]: "(rel_fun pcr_integer (rel_fun (=) pcr_integer)) (^) (^)" by(auto 4 3 simp add: integer.pcr_cr_eq cr_integer_def int_of_integer_pow) lemma sub1_lt_0_iff [simp]: "Code_Numeral.sub n num.One < 0 \ False" by(cases n)(simp_all add: Code_Numeral.sub_code) lemma nat_of_integer_numeral [simp]: "nat_of_integer (numeral n) = numeral n" by transfer simp lemma nat_of_integer_sub1_conv_pred_numeral [simp]: "nat_of_integer (Code_Numeral.sub n num.One) = pred_numeral n" by(cases n)(simp_all add: Code_Numeral.sub_code) lemma nat_of_integer_1 [simp]: "nat_of_integer 1 = 1" by transfer simp lemma dup_1 [simp]: "Code_Numeral.dup 1 = 2" by transfer simp section \Bit operations on @{typ integer}\ text \Bit operations on @{typ integer} are the same as on @{typ int}\ lift_definition bin_rest_integer :: "integer \ integer" is \\k . k div 2\ . lift_definition bin_last_integer :: "integer \ bool" is odd . lift_definition Bit_integer :: "integer \ bool \ integer" is \\k b. of_bool b + 2 * k\ . end instantiation integer :: lsb begin context includes integer.lifting begin lift_definition lsb_integer :: "integer \ bool" is lsb . instance by (standard; transfer) (fact lsb_odd) end end instantiation integer :: msb begin context includes integer.lifting begin lift_definition msb_integer :: "integer \ bool" is msb . instance .. end end instantiation integer :: set_bit begin context includes integer.lifting begin lift_definition set_bit_integer :: "integer \ nat \ bool \ integer" is set_bit . instance apply standard apply transfer apply (simp add: bit_simps) done end end abbreviation (input) wf_set_bits_integer where "wf_set_bits_integer \ wf_set_bits_int" section \Target language implementations\ text \ Unfortunately, this is not straightforward, because these API functions have different signatures and preconditions on the parameters: \begin{description} \item[Standard ML] Shifts in IntInf are given as word, but not IntInf. \item[Haskell] In the Data.Bits.Bits type class, shifts and bit indices are given as Int rather than Integer. \end{description} Additional constants take only parameters of type @{typ integer} rather than @{typ nat} and check the preconditions as far as possible (e.g., being non-negative) in a portable way. Manual implementations inside code\_printing perform the remaining range checks and convert these @{typ integer}s into the right type. For normalisation by evaluation, we derive custom code equations, because NBE does not know these code\_printing serialisations and would otherwise loop. \ code_identifier code_module Bits_Integer \ (SML) Bits_Int and (OCaml) Bits_Int and (Haskell) Bits_Int and (Scala) Bits_Int code_printing code_module Bits_Integer \ (SML) \structure Bits_Integer : sig val set_bit : IntInf.int -> IntInf.int -> bool -> IntInf.int val shiftl : IntInf.int -> IntInf.int -> IntInf.int val shiftr : IntInf.int -> IntInf.int -> IntInf.int val test_bit : IntInf.int -> IntInf.int -> bool end = struct val maxWord = IntInf.pow (2, Word.wordSize); fun set_bit x n b = if n < maxWord then if b then IntInf.orb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) else IntInf.andb (x, IntInf.notb (IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n)))) else raise (Fail ("Bit index too large: " ^ IntInf.toString n)); fun shiftl x n = if n < maxWord then IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)) else raise (Fail ("Shift operand too large: " ^ IntInf.toString n)); fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else raise (Fail ("Shift operand too large: " ^ IntInf.toString n)); fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else raise (Fail ("Bit index too large: " ^ IntInf.toString n)); end; (*struct Bits_Integer*)\ code_reserved SML Bits_Integer code_printing code_module Bits_Integer \ (OCaml) \module Bits_Integer : sig val shiftl : Z.t -> Z.t -> Z.t val shiftr : Z.t -> Z.t -> Z.t val test_bit : Z.t -> Z.t -> bool end = struct (* We do not need an explicit range checks here, because Big_int.int_of_big_int raises Failure if the argument does not fit into an int. *) let shiftl x n = Z.shift_left x (Z.to_int n);; let shiftr x n = Z.shift_right x (Z.to_int n);; let test_bit x n = Z.testbit x (Z.to_int n);; end;; (*struct Bits_Integer*)\ code_reserved OCaml Bits_Integer code_printing code_module Data_Bits \ (Haskell) \ module Data_Bits where { import qualified Data.Bits; {- The ...Bounded functions assume that the Integer argument for the shift or bit index fits into an Int, is non-negative and (for types of fixed bit width) less than bitSize -} infixl 7 .&.; infixl 6 `xor`; infixl 5 .|.; (.&.) :: Data.Bits.Bits a => a -> a -> a; (.&.) = (Data.Bits..&.); xor :: Data.Bits.Bits a => a -> a -> a; xor = Data.Bits.xor; (.|.) :: Data.Bits.Bits a => a -> a -> a; (.|.) = (Data.Bits..|.); complement :: Data.Bits.Bits a => a -> a; complement = Data.Bits.complement; testBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool; testBitUnbounded x b | b <= toInteger (Prelude.maxBound :: Int) = Data.Bits.testBit x (fromInteger b) | otherwise = error ("Bit index too large: " ++ show b) ; testBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool; testBitBounded x b = Data.Bits.testBit x (fromInteger b); setBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool -> a; setBitUnbounded x n b | n <= toInteger (Prelude.maxBound :: Int) = if b then Data.Bits.setBit x (fromInteger n) else Data.Bits.clearBit x (fromInteger n) | otherwise = error ("Bit index too large: " ++ show n) ; setBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool -> a; setBitBounded x n True = Data.Bits.setBit x (fromInteger n); setBitBounded x n False = Data.Bits.clearBit x (fromInteger n); shiftlUnbounded :: Data.Bits.Bits a => a -> Integer -> a; shiftlUnbounded x n | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftL x (fromInteger n) | otherwise = error ("Shift operand too large: " ++ show n) ; shiftlBounded :: Data.Bits.Bits a => a -> Integer -> a; shiftlBounded x n = Data.Bits.shiftL x (fromInteger n); shiftrUnbounded :: Data.Bits.Bits a => a -> Integer -> a; shiftrUnbounded x n | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftR x (fromInteger n) | otherwise = error ("Shift operand too large: " ++ show n) ; shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Integer -> a; shiftrBounded x n = Data.Bits.shiftR x (fromInteger n); }\ and \ \@{theory HOL.Quickcheck_Narrowing} maps @{typ integer} to Haskell's Prelude.Int type instead of Integer. For compatibility with the Haskell target, we nevertheless provide bounded and unbounded functions.\ (Haskell_Quickcheck) \ module Data_Bits where { import qualified Data.Bits; {- The functions assume that the Int argument for the shift or bit index is non-negative and (for types of fixed bit width) less than bitSize -} infixl 7 .&.; infixl 6 `xor`; infixl 5 .|.; (.&.) :: Data.Bits.Bits a => a -> a -> a; (.&.) = (Data.Bits..&.); xor :: Data.Bits.Bits a => a -> a -> a; xor = Data.Bits.xor; (.|.) :: Data.Bits.Bits a => a -> a -> a; (.|.) = (Data.Bits..|.); complement :: Data.Bits.Bits a => a -> a; complement = Data.Bits.complement; testBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool; testBitUnbounded = Data.Bits.testBit; testBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool; testBitBounded = Data.Bits.testBit; setBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool -> a; setBitUnbounded x n True = Data.Bits.setBit x n; setBitUnbounded x n False = Data.Bits.clearBit x n; setBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool -> a; setBitBounded x n True = Data.Bits.setBit x n; setBitBounded x n False = Data.Bits.clearBit x n; shiftlUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftlUnbounded = Data.Bits.shiftL; shiftlBounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftlBounded = Data.Bits.shiftL; shiftrUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftrUnbounded = Data.Bits.shiftR; shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Prelude.Int -> a; shiftrBounded = Data.Bits.shiftR; }\ code_reserved Haskell Data_Bits code_printing code_module Bits_Integer \ (Scala) \object Bits_Integer { def setBit(x: BigInt, n: BigInt, b: Boolean) : BigInt = if (n.isValidInt) if (b) x.setBit(n.toInt) else x.clearBit(n.toInt) else sys.error("Bit index too large: " + n.toString) def shiftl(x: BigInt, n: BigInt) : BigInt = if (n.isValidInt) x << n.toInt else sys.error("Shift index too large: " + n.toString) def shiftr(x: BigInt, n: BigInt) : BigInt = if (n.isValidInt) x << n.toInt else sys.error("Shift index too large: " + n.toString) def testBit(x: BigInt, n: BigInt) : Boolean = if (n.isValidInt) x.testBit(n.toInt) else sys.error("Bit index too large: " + n.toString) } /* object Bits_Integer */\ code_printing constant "Bit_Operations.and :: integer \ integer \ integer" \ (SML) "IntInf.andb ((_),/ (_))" and (OCaml) "Z.logand" and (Haskell) "((Data'_Bits..&.) :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "((Data'_Bits..&.) :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 3 "&" | constant "Bit_Operations.or :: integer \ integer \ integer" \ (SML) "IntInf.orb ((_),/ (_))" and (OCaml) "Z.logor" and (Haskell) "((Data'_Bits..|.) :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "((Data'_Bits..|.) :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 1 "|" | constant "Bit_Operations.xor :: integer \ integer \ integer" \ (SML) "IntInf.xorb ((_),/ (_))" and (OCaml) "Z.logxor" and (Haskell) "(Data'_Bits.xor :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.xor :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 2 "^" -| constant "NOT :: integer \ integer" \ +| constant "Bit_Operations.not :: integer \ integer" \ (SML) "IntInf.notb" and (OCaml) "Z.lognot" and (Haskell) "(Data'_Bits.complement :: Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.complement :: Prelude.Int -> Prelude.Int)" and (Scala) "_.unary'_~" code_printing constant bin_rest_integer \ (SML) "IntInf.div ((_), 2)" and (OCaml) "Z.shift'_right/ _/ 1" and (Haskell) "(Data'_Bits.shiftrUnbounded _ 1 :: Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftrUnbounded _ 1 :: Prelude.Int)" and (Scala) "_ >> 1" context includes integer.lifting bit_operations_syntax begin lemma bitNOT_integer_code [code]: fixes i :: integer shows "NOT i = - i - 1" by transfer (simp add: not_int_def) lemma bin_rest_integer_code [code nbe]: "bin_rest_integer i = i div 2" by transfer rule lemma bin_last_integer_code [code]: "bin_last_integer i \ i AND 1 \ 0" by transfer (simp add: and_one_eq odd_iff_mod_2_eq_one) lemma bin_last_integer_nbe [code nbe]: "bin_last_integer i \ i mod 2 \ 0" by transfer (simp add: odd_iff_mod_2_eq_one) lemma bitval_bin_last_integer [code_unfold]: "of_bool (bin_last_integer i) = i AND 1" by transfer (simp add: and_one_eq mod_2_eq_odd) end definition integer_test_bit :: "integer \ integer \ bool" where "integer_test_bit x n = (if n < 0 then undefined x n else bit x (nat_of_integer n))" declare [[code drop: \bit :: integer \ nat \ bool\]] lemma bit_integer_code [code]: "bit x n \ integer_test_bit x (integer_of_nat n)" by (simp add: integer_test_bit_def) lemma integer_test_bit_code [code]: "integer_test_bit x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_test_bit 0 0 = False" "integer_test_bit 0 (Code_Numeral.Pos n) = False" "integer_test_bit (Code_Numeral.Pos num.One) 0 = True" "integer_test_bit (Code_Numeral.Pos (num.Bit0 n)) 0 = False" "integer_test_bit (Code_Numeral.Pos (num.Bit1 n)) 0 = True" "integer_test_bit (Code_Numeral.Pos num.One) (Code_Numeral.Pos n') = False" "integer_test_bit (Code_Numeral.Pos (num.Bit0 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Pos n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Pos (num.Bit1 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Pos n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Neg num.One) 0 = True" "integer_test_bit (Code_Numeral.Neg (num.Bit0 n)) 0 = False" "integer_test_bit (Code_Numeral.Neg (num.Bit1 n)) 0 = True" "integer_test_bit (Code_Numeral.Neg num.One) (Code_Numeral.Pos n') = True" "integer_test_bit (Code_Numeral.Neg (num.Bit0 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Neg n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Neg (num.Bit1 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Neg (n + num.One)) (Code_Numeral.sub n' num.One)" by (simp_all add: integer_test_bit_def bit_integer_def flip: bit_not_int_iff') code_printing constant integer_test_bit \ (SML) "Bits'_Integer.test'_bit" and (OCaml) "Bits'_Integer.test'_bit" and (Haskell) "(Data'_Bits.testBitUnbounded :: Integer -> Integer -> Bool)" and (Haskell_Quickcheck) "(Data'_Bits.testBitUnbounded :: Prelude.Int -> Prelude.Int -> Bool)" and (Scala) "Bits'_Integer.testBit" context includes integer.lifting bit_operations_syntax begin lemma lsb_integer_code [code]: fixes x :: integer shows "lsb x = bit x 0" by transfer(simp add: lsb_int_def) definition integer_set_bit :: "integer \ integer \ bool \ integer" where [code del]: "integer_set_bit x n b = (if n < 0 then undefined x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_integer_code [code]: "set_bit x i b = integer_set_bit x (integer_of_nat i) b" by(simp add: integer_set_bit_def) lemma set_bit_integer_conv_masks: fixes x :: integer shows "set_bit x i b = (if b then x OR (push_bit i 1) else x AND NOT (push_bit i 1))" by transfer (simp add: int_set_bit_False_conv_NAND int_set_bit_True_conv_OR) end code_printing constant integer_set_bit \ (SML) "Bits'_Integer.set'_bit" and (Haskell) "(Data'_Bits.setBitUnbounded :: Integer -> Integer -> Bool -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.setBitUnbounded :: Prelude.Int -> Prelude.Int -> Bool -> Prelude.Int)" and (Scala) "Bits'_Integer.setBit" text \ OCaml.Big\_int does not have a method for changing an individual bit, so we emulate that with masks. We prefer an Isabelle implementation, because this then takes care of the signs for AND and OR. \ context includes bit_operations_syntax begin lemma integer_set_bit_code [code]: "integer_set_bit x n b = (if n < 0 then undefined x n b else if b then x OR (push_bit (nat_of_integer n) 1) else x AND NOT (push_bit (nat_of_integer n) 1))" by (auto simp add: integer_set_bit_def not_less set_bit_eq set_bit_def unset_bit_def) end definition integer_shiftl :: "integer \ integer \ integer" where [code del]: "integer_shiftl x n = (if n < 0 then undefined x n else push_bit (nat_of_integer n) x)" declare [[code drop: \push_bit :: nat \ integer \ integer\]] lemma shiftl_integer_code [code]: fixes x :: integer shows "push_bit n x = integer_shiftl x (integer_of_nat n)" by(auto simp add: integer_shiftl_def) context includes integer.lifting begin lemma shiftl_integer_conv_mult_pow2: fixes x :: integer shows "push_bit n x = x * 2 ^ n" by (fact push_bit_eq_mult) lemma integer_shiftl_code [code]: "integer_shiftl x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_shiftl x 0 = x" "integer_shiftl x (Code_Numeral.Pos n) = integer_shiftl (Code_Numeral.dup x) (Code_Numeral.sub n num.One)" "integer_shiftl 0 (Code_Numeral.Pos n) = 0" apply (simp_all add: integer_shiftl_def numeral_eq_Suc) apply transfer apply (simp add: ac_simps) done end code_printing constant integer_shiftl \ (SML) "Bits'_Integer.shiftl" and (OCaml) "Bits'_Integer.shiftl" and (Haskell) "(Data'_Bits.shiftlUnbounded :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftlUnbounded :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) "Bits'_Integer.shiftl" definition integer_shiftr :: "integer \ integer \ integer" where [code del]: "integer_shiftr x n = (if n < 0 then undefined x n else drop_bit (nat_of_integer n) x)" declare [[code drop: \drop_bit :: nat \ integer \ integer\]] lemma shiftr_integer_conv_div_pow2: includes integer.lifting fixes x :: integer shows "drop_bit n x = x div 2 ^ n" by (fact drop_bit_eq_div) lemma shiftr_integer_code [code]: fixes x :: integer shows "drop_bit n x = integer_shiftr x (integer_of_nat n)" by(auto simp add: integer_shiftr_def) code_printing constant integer_shiftr \ (SML) "Bits'_Integer.shiftr" and (OCaml) "Bits'_Integer.shiftr" and (Haskell) "(Data'_Bits.shiftrUnbounded :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftrUnbounded :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) "Bits'_Integer.shiftr" lemma integer_shiftr_code [code]: includes integer.lifting shows "integer_shiftr x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_shiftr x 0 = x" "integer_shiftr 0 (Code_Numeral.Pos n) = 0" "integer_shiftr (Code_Numeral.Pos num.One) (Code_Numeral.Pos n) = 0" "integer_shiftr (Code_Numeral.Pos (num.Bit0 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Pos n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Pos (num.Bit1 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Pos n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Neg num.One) (Code_Numeral.Pos n) = -1" "integer_shiftr (Code_Numeral.Neg (num.Bit0 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Neg n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Neg (num.Bit1 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Neg (Num.inc n')) (Code_Numeral.sub n num.One)" apply (simp_all add: integer_shiftr_def numeral_eq_Suc drop_bit_Suc) apply transfer apply simp apply transfer apply simp apply transfer apply (simp add: add_One) done context includes integer.lifting begin lemma Bit_integer_code [code]: "Bit_integer i False = push_bit 1 i" "Bit_integer i True = push_bit 1 i + 1" by (transfer; simp)+ lemma msb_integer_code [code]: "msb (x :: integer) \ x < 0" by transfer (simp add: msb_int_def) end context includes integer.lifting natural.lifting bit_operations_syntax begin lemma bitAND_integer_unfold [code]: "x AND y = (if x = 0 then 0 else if x = - 1 then y else Bit_integer (bin_rest_integer x AND bin_rest_integer y) (bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps and_int_rec [of _ \_ * 2\] and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] elim!: evenE oddE) lemma bitOR_integer_unfold [code]: "x OR y = (if x = 0 then y else if x = - 1 then - 1 else Bit_integer (bin_rest_integer x OR bin_rest_integer y) (bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps or_int_rec [of _ \_ * 2\] or_int_rec [of _ \1 + _ * 2\] or_int_rec [of \1 + _ * 2\] elim!: evenE oddE) lemma bitXOR_integer_unfold [code]: "x XOR y = (if x = 0 then y else if x = - 1 then NOT y else Bit_integer (bin_rest_integer x XOR bin_rest_integer y) (\ bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps xor_int_rec [of _ \_ * 2\] xor_int_rec [of \_ * 2\] xor_int_rec [of \1 + _ * 2\] elim!: evenE oddE) end section \Test code generator setup\ context includes bit_operations_syntax begin definition bit_integer_test :: "bool" where "bit_integer_test = (([ -1 AND 3, 1 AND -3, 3 AND 5, -3 AND (- 5) , -3 OR 1, 1 OR -3, 3 OR 5, -3 OR (- 5) , NOT 1, NOT (- 3) , -1 XOR 3, 1 XOR (- 3), 3 XOR 5, -5 XOR (- 3) , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , push_bit 2 1, push_bit 3 (- 1) , drop_bit 3 100, drop_bit 3 (- 100)] :: integer list) = [ 3, 1, 1, -7 , -3, -3, 7, -1 , -2, 2 , -4, -4, 6, 6 , 21, -1, 4, -7 , 4, -8 , 12, -13] \ [ bit (5 :: integer) 4, bit (5 :: integer) 2, bit (-5 :: integer) 4, bit (-5 :: integer) 2 , lsb (5 :: integer), lsb (4 :: integer), lsb (-1 :: integer), lsb (-2 :: integer), msb (5 :: integer), msb (0 :: integer), msb (-1 :: integer), msb (-2 :: integer)] = [ False, True, True, False, True, False, True, False, False, False, True, True])" export_code bit_integer_test checking SML Haskell? Haskell_Quickcheck? OCaml? Scala notepad begin have bit_integer_test by eval have bit_integer_test by normalization have bit_integer_test by code_simp end ML_val \val true = @{code bit_integer_test}\ lemma "x AND y = x OR (y :: integer)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: integer) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] oops lemma "(f :: integer \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) end hide_const bit_integer_test hide_fact bit_integer_test_def end diff --git a/thys/Native_Word/Uint.thy b/thys/Native_Word/Uint.thy --- a/thys/Native_Word/Uint.thy +++ b/thys/Native_Word/Uint.thy @@ -1,824 +1,824 @@ (* Title: Uint.thy Author: Peter Lammich, TU Munich Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of default size\ theory Uint imports Code_Target_Word_Base Word_Type_Copies begin text \ This theory provides access to words in the target languages of the code generator whose bit width is the default of the target language. To that end, the type \uint\ models words of width \dflt_size\, but \dflt_size\ is known only to be positive. Usage restrictions: Default-size words (type \uint\) cannot be used for evaluation, because the results depend on the particular choice of word size in the target language and implementation. Symbolic evaluation has not yet been set up for \uint\. \ text \The default size type\ typedecl dflt_size instantiation dflt_size :: typerep begin definition "typerep_class.typerep \ \_ :: dflt_size itself. Typerep.Typerep (STR ''Uint.dflt_size'') []" instance .. end consts dflt_size_aux :: "nat" specification (dflt_size_aux) dflt_size_aux_g0: "dflt_size_aux > 0" by auto hide_fact dflt_size_aux_def instantiation dflt_size :: len begin definition "len_of_dflt_size (_ :: dflt_size itself) \ dflt_size_aux" instance by(intro_classes)(simp add: len_of_dflt_size_def dflt_size_aux_g0) end abbreviation "dflt_size \ len_of (TYPE (dflt_size))" context includes integer.lifting begin lift_definition dflt_size_integer :: integer is "int dflt_size" . declare dflt_size_integer_def[code del] \ \The code generator will substitute a machine-dependent value for this constant\ lemma dflt_size_by_int[code]: "dflt_size = nat_of_integer dflt_size_integer" by transfer simp lemma dflt_size[simp]: "dflt_size > 0" "dflt_size \ Suc 0" "\ dflt_size < Suc 0" using len_gt_0[where 'a=dflt_size] by (simp_all del: len_gt_0) end section \Type definition and primitive operations\ typedef uint = \UNIV :: dflt_size word set\ .. global_interpretation uint: word_type_copy Abs_uint Rep_uint using type_definition_uint by (rule word_type_copy.intro) setup_lifting type_definition_uint declare uint.of_word_of [code abstype] declare Quotient_uint [transfer_rule] instantiation uint :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint :: uint is 0 . lift_definition one_uint :: uint is 1 . lift_definition plus_uint :: \uint \ uint \ uint\ is \(+)\ . lift_definition uminus_uint :: \uint \ uint\ is uminus . lift_definition minus_uint :: \uint \ uint \ uint\ is \(-)\ . lift_definition times_uint :: \uint \ uint \ uint\ is \(*)\ . lift_definition divide_uint :: \uint \ uint \ uint\ is \(div)\ . lift_definition modulo_uint :: \uint \ uint \ uint\ is \(mod)\ . lift_definition equal_uint :: \uint \ uint \ bool\ is \HOL.equal\ . lift_definition less_eq_uint :: \uint \ uint \ bool\ is \(\)\ . lift_definition less_uint :: \uint \ uint \ bool\ is \(<)\ . global_interpretation uint: word_type_copy_ring Abs_uint Rep_uint by standard (fact zero_uint.rep_eq one_uint.rep_eq plus_uint.rep_eq uminus_uint.rep_eq minus_uint.rep_eq times_uint.rep_eq divide_uint.rep_eq modulo_uint.rep_eq equal_uint.rep_eq less_eq_uint.rep_eq less_uint.rep_eq)+ instance proof - show \OFCLASS(uint, comm_ring_1_class)\ by (rule uint.of_class_comm_ring_1) show \OFCLASS(uint, semiring_modulo_class)\ by (fact uint.of_class_semiring_modulo) show \OFCLASS(uint, equal_class)\ by (fact uint.of_class_equal) show \OFCLASS(uint, linorder_class)\ by (fact uint.of_class_linorder) qed end instantiation uint :: ring_bit_operations begin lift_definition bit_uint :: \uint \ nat \ bool\ is bit . -lift_definition not_uint :: \uint \ uint\ is NOT . +lift_definition not_uint :: \uint \ uint\ is \Bit_Operations.not\ . lift_definition and_uint :: \uint \ uint \ uint\ is \Bit_Operations.and\ . lift_definition or_uint :: \uint \ uint \ uint\ is \Bit_Operations.or\ . lift_definition xor_uint :: \uint \ uint \ uint\ is \Bit_Operations.xor\ . lift_definition mask_uint :: \nat \ uint\ is mask . lift_definition push_bit_uint :: \nat \ uint \ uint\ is push_bit . lift_definition drop_bit_uint :: \nat \ uint \ uint\ is drop_bit . lift_definition signed_drop_bit_uint :: \nat \ uint \ uint\ is signed_drop_bit . lift_definition take_bit_uint :: \nat \ uint \ uint\ is take_bit . lift_definition set_bit_uint :: \nat \ uint \ uint\ is Bit_Operations.set_bit . lift_definition unset_bit_uint :: \nat \ uint \ uint\ is unset_bit . lift_definition flip_bit_uint :: \nat \ uint \ uint\ is flip_bit . global_interpretation uint: word_type_copy_bits Abs_uint Rep_uint signed_drop_bit_uint by standard (fact bit_uint.rep_eq not_uint.rep_eq and_uint.rep_eq or_uint.rep_eq xor_uint.rep_eq mask_uint.rep_eq push_bit_uint.rep_eq drop_bit_uint.rep_eq signed_drop_bit_uint.rep_eq take_bit_uint.rep_eq set_bit_uint.rep_eq unset_bit_uint.rep_eq flip_bit_uint.rep_eq)+ instance by (fact uint.of_class_ring_bit_operations) end lift_definition uint_of_nat :: \nat \ uint\ is word_of_nat . lift_definition nat_of_uint :: \uint \ nat\ is unat . lift_definition uint_of_int :: \int \ uint\ is word_of_int . lift_definition int_of_uint :: \uint \ int\ is uint . context includes integer.lifting begin lift_definition Uint :: \integer \ uint\ is word_of_int . lift_definition integer_of_uint :: \uint \ integer\ is uint . end global_interpretation uint: word_type_copy_more Abs_uint Rep_uint signed_drop_bit_uint uint_of_nat nat_of_uint uint_of_int int_of_uint Uint integer_of_uint apply standard apply (simp_all add: uint_of_nat.rep_eq nat_of_uint.rep_eq uint_of_int.rep_eq int_of_uint.rep_eq Uint.rep_eq integer_of_uint.rep_eq integer_eq_iff) done instantiation uint :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint :: \uint \ nat\ is size . lift_definition msb_uint :: \uint \ bool\ is msb . lift_definition lsb_uint :: \uint \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint :: \uint \ nat \ bool \ uint\ where set_bit_uint_eq: \set_bit_uint a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint_transfer [transfer_rule]: \(cr_uint ===> (=) ===> (\) ===> cr_uint) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint_eq [abs_def]) transfer_prover end lift_definition set_bits_uint :: \(nat \ bool) \ uint\ is set_bits . lift_definition set_bits_aux_uint :: \(nat \ bool) \ nat \ uint \ uint\ is set_bits_aux . global_interpretation uint: word_type_copy_misc Abs_uint Rep_uint signed_drop_bit_uint uint_of_nat nat_of_uint uint_of_int int_of_uint Uint integer_of_uint dflt_size set_bits_aux_uint by (standard; transfer) simp_all instance using uint.of_class_bit_comprehension uint.of_class_set_bit uint.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint \ (SML) \ structure Uint : sig val set_bit : Word.word -> IntInf.int -> bool -> Word.word val shiftl : Word.word -> IntInf.int -> Word.word val shiftr : Word.word -> IntInf.int -> Word.word val shiftr_signed : Word.word -> IntInf.int -> Word.word val test_bit : Word.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word.orb (x, mask) else Word.andb (x, Word.notb mask) end fun shiftl x n = Word.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word.andb (x, Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word.fromInt 0 end; (* struct Uint *)\ code_reserved SML Uint code_printing code_module Uint \ (Haskell) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Integer dflt_size = Prelude.toInteger (bitSize_aux (0::Word)) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize\ and (Haskell_Quickcheck) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Int dflt_size = bitSize_aux (0::Word) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize \ code_reserved Haskell Uint dflt_size text \ OCaml and Scala provide only signed bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint" \ (OCaml) \module Uint : sig type t = int val dflt_size : Z.t val less : t -> t -> bool val less_eq : t -> t -> bool val set_bit : t -> Z.t -> bool -> t val shiftl : t -> Z.t -> t val shiftr : t -> Z.t -> t val shiftr_signed : t -> Z.t -> t val test_bit : t -> Z.t -> bool val int_mask : int val int32_mask : int32 val int64_mask : int64 end = struct type t = int let dflt_size = Z.of_int Sys.int_size;; (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if x<0 then y<0 && x 0;; let int_mask = if Sys.int_size < 32 then lnot 0 else 0xFFFFFFFF;; let int32_mask = if Sys.int_size < 32 then Int32.pred (Int32.shift_left Int32.one Sys.int_size) else Int32.of_string "0xFFFFFFFF";; let int64_mask = if Sys.int_size < 64 then Int64.pred (Int64.shift_left Int64.one Sys.int_size) else Int64.of_string "0xFFFFFFFFFFFFFFFF";; end;; (*struct Uint*)\ code_reserved OCaml Uint code_printing code_module Uint \ (Scala) \object Uint { def dflt_size : BigInt = BigInt(32) def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint */\ code_reserved Scala Uint text \ OCaml's conversion from Big\_int to int demands that the value fits into a signed integer. The following justifies the implementation. \ context includes integer.lifting bit_operations_syntax begin definition wivs_mask :: int where "wivs_mask = 2^ dflt_size - 1" lift_definition wivs_mask_integer :: integer is wivs_mask . lemma [code]: "wivs_mask_integer = 2 ^ dflt_size - 1" by transfer (simp add: wivs_mask_def) definition wivs_shift :: int where "wivs_shift = 2 ^ dflt_size" lift_definition wivs_shift_integer :: integer is wivs_shift . lemma [code]: "wivs_shift_integer = 2 ^ dflt_size" by transfer (simp add: wivs_shift_def) definition wivs_index :: nat where "wivs_index == dflt_size - 1" lift_definition wivs_index_integer :: integer is "int wivs_index". lemma wivs_index_integer_code[code]: "wivs_index_integer = dflt_size_integer - 1" by transfer (simp add: wivs_index_def of_nat_diff) definition wivs_overflow :: int where "wivs_overflow == 2^ (dflt_size - 1)" lift_definition wivs_overflow_integer :: integer is wivs_overflow . lemma [code]: "wivs_overflow_integer = 2 ^ (dflt_size - 1)" by transfer (simp add: wivs_overflow_def) definition wivs_least :: int where "wivs_least == - wivs_overflow" lift_definition wivs_least_integer :: integer is wivs_least . lemma [code]: "wivs_least_integer = - (2 ^ (dflt_size - 1))" by transfer (simp add: wivs_overflow_def wivs_least_def) definition Uint_signed :: "integer \ uint" where "Uint_signed i = (if i < wivs_least_integer \ wivs_overflow_integer \ i then undefined Uint i else Uint i)" lemma Uint_code [code]: "Uint i = (let i' = i AND wivs_mask_integer in if bit i' wivs_index then Uint_signed (i' - wivs_shift_integer) else Uint_signed i')" including undefined_transfer unfolding Uint_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed wivs_mask_def wivs_index_def wivs_overflow_def wivs_least_def wivs_shift_def) done lemma Uint_signed_code [code]: "Rep_uint (Uint_signed i) = (if i < wivs_least_integer \ i \ wivs_overflow_integer then Rep_uint (undefined Uint i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint_signed_def Uint_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint_inverse) end text \ Avoid @{term Abs_uint} in generated code, use @{term Rep_uint'} instead. The symbolic implementations for code\_simp use @{term Rep_uint}. The new destructor @{term Rep_uint'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint} ([code abstract] equations for @{typ uint} may use @{term Rep_uint} because these instances will be folded away.) \ definition Rep_uint' where [simp]: "Rep_uint' = Rep_uint" lemma Rep_uint'_code [code]: "Rep_uint' x = (BITS n. bit x n)" unfolding Rep_uint'_def by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint' :: "dflt_size word \ uint" is "\x :: dflt_size word. x" . lemma Abs_uint'_code [code]: "Abs_uint' x = Uint (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint \ _"]] lemma term_of_uint_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint.uint.Abs_uint'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR (STR ''Uint.dflt_size'') []], TR (STR ''Uint.uint'') []])) (term_of_class.term_of (Rep_uint' x))" by(simp add: term_of_anything) text \Important: We must prevent the reflection oracle (eval-tac) to use our machine-dependent type. \ code_printing type_constructor uint \ (SML) "Word.word" and (Haskell) "Uint.Word" and (OCaml) "Uint.t" and (Scala) "Int" and (Eval) "*** \"Error: Machine dependent type\" ***" and (Quickcheck) "Word.word" | constant dflt_size_integer \ (SML) "(IntInf.fromLarge (Int.toLarge Word.wordSize))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.wordSize" and (Haskell) "Uint.dflt'_size" and (OCaml) "Uint.dflt'_size" and (Scala) "Uint.dflt'_size" | constant Uint \ (SML) "Word.fromLargeInt (IntInf.toLarge _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint.Word)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint.Word)" and (Scala) "_.intValue" | constant Uint_signed \ (OCaml) "Z.to'_int" | constant "0 :: uint" \ (SML) "(Word.fromInt 0)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 0)" and (Haskell) "(0 :: Uint.Word)" and (OCaml) "0" and (Scala) "0" | constant "1 :: uint" \ (SML) "(Word.fromInt 1)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 1)" and (Haskell) "(1 :: Uint.Word)" and (OCaml) "1" and (Scala) "1" | constant "plus :: uint \ _ " \ (SML) "Word.+ ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Pervasives.(+)" and (Scala) infixl 7 "+" | constant "uminus :: uint \ _" \ (SML) "Word.~" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.~" and (Haskell) "negate" and (OCaml) "Pervasives.(~-)" and (Scala) "!(- _)" | constant "minus :: uint \ _" \ (SML) "Word.- ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Pervasives.(-)" and (Scala) infixl 7 "-" | constant "times :: uint \ _ \ _" \ (SML) "Word.* ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Pervasives.( * )" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint \ _ \ bool" \ (SML) "!((_ : Word.word) = _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "!((_ : Word.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Pervasives.(=):Uint.t -> Uint.t -> bool)" and (Scala) infixl 5 "==" | class_instance uint :: equal \ (Haskell) - | constant "less_eq :: uint \ _ \ bool" \ (SML) "Word.<= ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint.less'_eq" and (Scala) "Uint.less'_eq" | constant "less :: uint \ _ \ bool" \ (SML) "Word.< ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint.less" and (Scala) "Uint.less" -| constant "NOT :: uint \ _" \ +| constant "Bit_Operations.not :: uint \ _" \ (SML) "Word.notb" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Pervasives.lnot" and (Scala) "_.unary'_~" | constant "Bit_Operations.and :: uint \ _" \ (SML) "Word.andb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Pervasives.(land)" and (Scala) infixl 3 "&" | constant "Bit_Operations.or :: uint \ _" \ (SML) "Word.orb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Pervasives.(lor)" and (Scala) infixl 1 "|" | constant "Bit_Operations.xor :: uint \ _" \ (SML) "Word.xorb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Pervasives.(lxor)" and (Scala) infixl 2 "^" definition uint_divmod :: "uint \ uint \ uint \ uint" where "uint_divmod x y = (if y = 0 then (undefined ((div) :: uint \ _) x (0 :: uint), undefined ((mod) :: uint \ _) x (0 :: uint)) else (x div y, x mod y))" definition uint_div :: "uint \ uint \ uint" where "uint_div x y = fst (uint_divmod x y)" definition uint_mod :: "uint \ uint \ uint" where "uint_mod x y = snd (uint_divmod x y)" lemma div_uint_code [code]: "x div y = (if y = 0 then 0 else uint_div x y)" including undefined_transfer unfolding uint_divmod_def uint_div_def by transfer(simp add: word_div_def) lemma mod_uint_code [code]: "x mod y = (if y = 0 then x else uint_mod x y)" including undefined_transfer unfolding uint_mod_def uint_divmod_def by transfer(simp add: word_mod_def) definition uint_sdiv :: "uint \ uint \ uint" where [code del]: "uint_sdiv x y = (if y = 0 then undefined ((div) :: uint \ _) x (0 :: uint) else Abs_uint (Rep_uint x sdiv Rep_uint y))" definition div0_uint :: "uint \ uint" where [code del]: "div0_uint x = undefined ((div) :: uint \ _) x (0 :: uint)" declare [[code abort: div0_uint]] definition mod0_uint :: "uint \ uint" where [code del]: "mod0_uint x = undefined ((mod) :: uint \ _) x (0 :: uint)" declare [[code abort: mod0_uint]] definition wivs_overflow_uint :: uint where "wivs_overflow_uint \ push_bit (dflt_size - 1) 1" lemma Rep_uint_wivs_overflow_uint_eq: \Rep_uint wivs_overflow_uint = 2 ^ (dflt_size - Suc 0)\ by (simp add: wivs_overflow_uint_def one_uint.rep_eq push_bit_uint.rep_eq uint.word_of_power push_bit_eq_mult) lemma wivs_overflow_uint_greater_eq_0: \wivs_overflow_uint > 0\ apply (simp add: less_uint.rep_eq zero_uint.rep_eq Rep_uint_wivs_overflow_uint_eq) apply transfer apply (simp add: take_bit_push_bit push_bit_eq_mult) done lemma uint_divmod_code [code]: "uint_divmod x y = (if wivs_overflow_uint \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint x, mod0_uint x) else let q = push_bit 1 (uint_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof (cases \y = 0\) case True moreover have \x \ 0\ by transfer simp moreover note wivs_overflow_uint_greater_eq_0 ultimately show ?thesis by (auto simp add: uint_divmod_def div0_uint_def mod0_uint_def not_less) next case False then show ?thesis including undefined_transfer unfolding uint_divmod_def uint_sdiv_def div0_uint_def mod0_uint_def wivs_overflow_uint_def apply transfer apply (simp add: divmod_via_sdivmod push_bit_of_1) done qed lemma uint_sdiv_code [code]: "Rep_uint (uint_sdiv x y) = (if y = 0 then Rep_uint (undefined ((div) :: uint \ _) x (0 :: uint)) else Rep_uint x sdiv Rep_uint y)" unfolding uint_sdiv_def by(simp add: Abs_uint_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint_divmod_code} computes both with division only. \ code_printing constant uint_div \ (SML) "Word.div ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint_mod \ (SML) "Word.mod ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint_divmod \ (Haskell) "divmod" | constant uint_sdiv \ (OCaml) "Pervasives.('/)" and (Scala) "_ '/ _" definition uint_test_bit :: "uint \ integer \ bool" where [code del]: "uint_test_bit x n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint_code [code]: "bit x n \ n < dflt_size \ uint_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint_test_bit_def by (transfer, simp, transfer, simp) lemma uint_test_bit_code [code]: "uint_test_bit w n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) w n else bit (Rep_uint w) (nat_of_integer n))" unfolding uint_test_bit_def by(simp add: bit_uint.rep_eq) code_printing constant uint_test_bit \ (SML) "Uint.test'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint.test'_bit" and (Scala) "Uint.test'_bit" definition uint_set_bit :: "uint \ integer \ bool \ uint" where [code del]: "uint_set_bit x n b = (if n < 0 \ dflt_size_integer \ n then undefined (set_bit :: uint \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint_code [code]: "set_bit x n b = (if n < dflt_size then uint_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by (transfer) (auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint_set_bit_code [code]: "Rep_uint (uint_set_bit w n b) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (set_bit :: uint \ _) w n b) else set_bit (Rep_uint w) (nat_of_integer n) b)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by transfer simp code_printing constant uint_set_bit \ (SML) "Uint.set'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint.set'_bit" and (Scala) "Uint.set'_bit" definition uint_shiftl :: "uint \ integer \ uint" where [code del]: "uint_shiftl x n = (if n < 0 \ dflt_size_integer \ n then undefined (push_bit :: nat \ uint \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint_code [code]: "push_bit n x = (if n < dflt_size then uint_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftl_def by (transfer fixing: n) simp lemma uint_shiftl_code [code]: "Rep_uint (uint_shiftl w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (push_bit :: nat \ uint \ _) w n) else push_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_shiftl_def by transfer simp code_printing constant uint_shiftl \ (SML) "Uint.shiftl" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint.shiftl" and (Scala) "Uint.shiftl" definition uint_shiftr :: "uint \ integer \ uint" where [code del]: "uint_shiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined (drop_bit :: nat \ uint \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint_code [code]: "drop_bit n x = (if n < dflt_size then uint_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftr_def by (transfer fixing: n) simp lemma uint_shiftr_code [code]: "Rep_uint (uint_shiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (drop_bit :: nat \ uint \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_shiftr_def by transfer simp code_printing constant uint_shiftr \ (SML) "Uint.shiftr" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint.shiftr" and (Scala) "Uint.shiftr" definition uint_sshiftr :: "uint \ integer \ uint" where [code del]: "uint_sshiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined signed_drop_bit_uint n x else signed_drop_bit_uint (nat_of_integer n) x)" lemma sshiftr_uint_code [code]: "signed_drop_bit_uint n x = (if n < dflt_size then uint_sshiftr x (integer_of_nat n) else if bit x wivs_index then -1 else 0)" including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer(simp add: not_less signed_drop_bit_beyond word_size wivs_index_def) lemma uint_sshiftr_code [code]: "Rep_uint (uint_sshiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined signed_drop_bit_uint n w) else signed_drop_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer simp code_printing constant uint_sshiftr \ (SML) "Uint.shiftr'_signed" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint.Int) _)) :: Uint.Word)" and (OCaml) "Uint.shiftr'_signed" and (Scala) "Uint.shiftr'_signed" lemma uint_msb_test_bit: "msb x \ bit (x :: uint) wivs_index" by transfer (simp add: msb_word_iff_bit wivs_index_def) lemma msb_uint_code [code]: "msb x \ uint_test_bit x wivs_index_integer" apply(simp add: uint_test_bit_def uint_msb_test_bit wivs_index_integer_code dflt_size_integer_def wivs_index_def) by (metis (full_types) One_nat_def dflt_size(2) less_iff_diff_less_0 nat_of_integer_of_nat of_nat_1 of_nat_diff of_nat_less_0_iff wivs_index_def) lemma uint_of_int_code [code]: "uint_of_int i = (BITS n. bit i n)" by transfer (simp add: word_of_int_conv_set_bits) section \Quickcheck setup\ definition uint_of_natural :: "natural \ uint" where "uint_of_natural x \ Uint (integer_of_natural x)" instantiation uint :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint \ qc_random_cnv uint_of_natural" definition "exhaustive_uint \ qc_exhaustive_cnv uint_of_natural" definition "full_exhaustive_uint \ qc_full_exhaustive_cnv uint_of_natural" instance .. end instantiation uint :: narrowing begin interpretation quickcheck_narrowing_samples "\i. (Uint i, Uint (- i))" "0" "Typerep.Typerep (STR ''Uint.uint'') []" . definition "narrowing_uint d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint itself \ _"]] lemmas partial_term_of_uint [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint16.thy b/thys/Native_Word/Uint16.thy --- a/thys/Native_Word/Uint16.thy +++ b/thys/Native_Word/Uint16.thy @@ -1,536 +1,536 @@ (* Title: Uint16.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 16 bits\ theory Uint16 imports Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for ML code generation: This theory assumes that the ML system provides a Word16 implementation (mlton does, but PolyML 5.5 does not). Therefore, the code setup lives in the target \SML_word\ rather than \SML\. This ensures that code generation still works as long as \uint16\ is not involved. For the target \SML\ itself, no special code generation for this type is set up. Nevertheless, it should work by emulation via \<^typ>\16 word\ if the theory \Code_Target_Bits_Int\ is imported. Restriction for OCaml code generation: OCaml does not provide an int16 type, so no special code generation for this type is set up. \ section \Type definition and primitive operations\ typedef uint16 = \UNIV :: 16 word set\ .. global_interpretation uint16: word_type_copy Abs_uint16 Rep_uint16 using type_definition_uint16 by (rule word_type_copy.intro) setup_lifting type_definition_uint16 declare uint16.of_word_of [code abstype] declare Quotient_uint16 [transfer_rule] instantiation uint16 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint16 :: uint16 is 0 . lift_definition one_uint16 :: uint16 is 1 . lift_definition plus_uint16 :: \uint16 \ uint16 \ uint16\ is \(+)\ . lift_definition uminus_uint16 :: \uint16 \ uint16\ is uminus . lift_definition minus_uint16 :: \uint16 \ uint16 \ uint16\ is \(-)\ . lift_definition times_uint16 :: \uint16 \ uint16 \ uint16\ is \(*)\ . lift_definition divide_uint16 :: \uint16 \ uint16 \ uint16\ is \(div)\ . lift_definition modulo_uint16 :: \uint16 \ uint16 \ uint16\ is \(mod)\ . lift_definition equal_uint16 :: \uint16 \ uint16 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint16 :: \uint16 \ uint16 \ bool\ is \(\)\ . lift_definition less_uint16 :: \uint16 \ uint16 \ bool\ is \(<)\ . global_interpretation uint16: word_type_copy_ring Abs_uint16 Rep_uint16 by standard (fact zero_uint16.rep_eq one_uint16.rep_eq plus_uint16.rep_eq uminus_uint16.rep_eq minus_uint16.rep_eq times_uint16.rep_eq divide_uint16.rep_eq modulo_uint16.rep_eq equal_uint16.rep_eq less_eq_uint16.rep_eq less_uint16.rep_eq)+ instance proof - show \OFCLASS(uint16, comm_ring_1_class)\ by (rule uint16.of_class_comm_ring_1) show \OFCLASS(uint16, semiring_modulo_class)\ by (fact uint16.of_class_semiring_modulo) show \OFCLASS(uint16, equal_class)\ by (fact uint16.of_class_equal) show \OFCLASS(uint16, linorder_class)\ by (fact uint16.of_class_linorder) qed end instantiation uint16 :: ring_bit_operations begin lift_definition bit_uint16 :: \uint16 \ nat \ bool\ is bit . -lift_definition not_uint16 :: \uint16 \ uint16\ is NOT . +lift_definition not_uint16 :: \uint16 \ uint16\ is \Bit_Operations.not\ . lift_definition and_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.and\ . lift_definition or_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.or\ . lift_definition xor_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.xor\ . lift_definition mask_uint16 :: \nat \ uint16\ is mask . lift_definition push_bit_uint16 :: \nat \ uint16 \ uint16\ is push_bit . lift_definition drop_bit_uint16 :: \nat \ uint16 \ uint16\ is drop_bit . lift_definition signed_drop_bit_uint16 :: \nat \ uint16 \ uint16\ is signed_drop_bit . lift_definition take_bit_uint16 :: \nat \ uint16 \ uint16\ is take_bit . lift_definition set_bit_uint16 :: \nat \ uint16 \ uint16\ is Bit_Operations.set_bit . lift_definition unset_bit_uint16 :: \nat \ uint16 \ uint16\ is unset_bit . lift_definition flip_bit_uint16 :: \nat \ uint16 \ uint16\ is flip_bit . global_interpretation uint16: word_type_copy_bits Abs_uint16 Rep_uint16 signed_drop_bit_uint16 by standard (fact bit_uint16.rep_eq not_uint16.rep_eq and_uint16.rep_eq or_uint16.rep_eq xor_uint16.rep_eq mask_uint16.rep_eq push_bit_uint16.rep_eq drop_bit_uint16.rep_eq signed_drop_bit_uint16.rep_eq take_bit_uint16.rep_eq set_bit_uint16.rep_eq unset_bit_uint16.rep_eq flip_bit_uint16.rep_eq)+ instance by (fact uint16.of_class_ring_bit_operations) end lift_definition uint16_of_nat :: \nat \ uint16\ is word_of_nat . lift_definition nat_of_uint16 :: \uint16 \ nat\ is unat . lift_definition uint16_of_int :: \int \ uint16\ is word_of_int . lift_definition int_of_uint16 :: \uint16 \ int\ is uint . context includes integer.lifting begin lift_definition Uint16 :: \integer \ uint16\ is word_of_int . lift_definition integer_of_uint16 :: \uint16 \ integer\ is uint . end global_interpretation uint16: word_type_copy_more Abs_uint16 Rep_uint16 signed_drop_bit_uint16 uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 apply standard apply (simp_all add: uint16_of_nat.rep_eq nat_of_uint16.rep_eq uint16_of_int.rep_eq int_of_uint16.rep_eq Uint16.rep_eq integer_of_uint16.rep_eq integer_eq_iff) done instantiation uint16 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint16 :: \uint16 \ nat\ is size . lift_definition msb_uint16 :: \uint16 \ bool\ is msb . lift_definition lsb_uint16 :: \uint16 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint16 :: \uint16 \ nat \ bool \ uint16\ where set_bit_uint16_eq: \set_bit_uint16 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint16_transfer [transfer_rule]: \(cr_uint16 ===> (=) ===> (\) ===> cr_uint16) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint16_eq [abs_def]) transfer_prover end lift_definition set_bits_uint16 :: \(nat \ bool) \ uint16\ is set_bits . lift_definition set_bits_aux_uint16 :: \(nat \ bool) \ nat \ uint16 \ uint16\ is set_bits_aux . global_interpretation uint16: word_type_copy_misc Abs_uint16 Rep_uint16 signed_drop_bit_uint16 uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 16 set_bits_aux_uint16 by (standard; transfer) simp_all instance using uint16.of_class_bit_comprehension uint16.of_class_set_bit uint16.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint16 \ (SML_word) \(* Test that words can handle numbers between 0 and 15 *) val _ = if 4 <= Word.wordSize then () else raise (Fail ("wordSize less than 4")); structure Uint16 : sig val set_bit : Word16.word -> IntInf.int -> bool -> Word16.word val shiftl : Word16.word -> IntInf.int -> Word16.word val shiftr : Word16.word -> IntInf.int -> Word16.word val shiftr_signed : Word16.word -> IntInf.int -> Word16.word val test_bit : Word16.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word16.orb (x, mask) else Word16.andb (x, Word16.notb mask) end fun shiftl x n = Word16.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word16.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word16.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word16.andb (x, Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word16.fromInt 0 end; (* struct Uint16 *)\ code_reserved SML_word Uint16 code_printing code_module Uint16 \ (Haskell) \module Uint16(Int16, Word16) where import Data.Int(Int16) import Data.Word(Word16)\ code_reserved Haskell Uint16 text \Scala provides unsigned 16-bit numbers as Char.\ code_printing code_module Uint16 \ (Scala) \object Uint16 { def set_bit(x: scala.Char, n: BigInt, b: Boolean) : scala.Char = if (b) (x | (1.toChar << n.intValue)).toChar else (x & (1.toChar << n.intValue).unary_~).toChar def shiftl(x: scala.Char, n: BigInt) : scala.Char = (x << n.intValue).toChar def shiftr(x: scala.Char, n: BigInt) : scala.Char = (x >>> n.intValue).toChar def shiftr_signed(x: scala.Char, n: BigInt) : scala.Char = (x.toShort >> n.intValue).toChar def test_bit(x: scala.Char, n: BigInt) : Boolean = (x & (1.toChar << n.intValue)) != 0 } /* object Uint16 */\ code_reserved Scala Uint16 text \ Avoid @{term Abs_uint16} in generated code, use @{term Rep_uint16'} instead. The symbolic implementations for code\_simp use @{term Rep_uint16}. The new destructor @{term Rep_uint16'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint16} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint16} ([code abstract] equations for @{typ uint16} may use @{term Rep_uint16} because these instances will be folded away.) To convert @{typ "16 word"} values into @{typ uint16}, use @{term "Abs_uint16'"}. \ definition Rep_uint16' where [simp]: "Rep_uint16' = Rep_uint16" lemma Rep_uint16'_transfer [transfer_rule]: "rel_fun cr_uint16 (=) (\x. x) Rep_uint16'" unfolding Rep_uint16'_def by(rule uint16.rep_transfer) lemma Rep_uint16'_code [code]: "Rep_uint16' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint16' :: "16 word \ uint16" is "\x :: 16 word. x" . lemma Abs_uint16'_code [code]: "Abs_uint16' x = Uint16 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint16 \ _"]] lemma term_of_uint16_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint16.uint16.Abs_uint16'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]], TR (STR ''Uint16.uint16'') []])) (term_of_class.term_of (Rep_uint16' x))" by(simp add: term_of_anything) lemma Uint16_code [code]: "Rep_uint16 (Uint16 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint16_def int_of_integer_symbolic_def by(simp add: Abs_uint16_inverse) code_printing type_constructor uint16 \ (SML_word) "Word16.word" and (Haskell) "Uint16.Word16" and (Scala) "scala.Char" | constant Uint16 \ (SML_word) "Word16.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint16.Word16)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Word16)" and (Scala) "_.charValue" | constant "0 :: uint16" \ (SML_word) "(Word16.fromInt 0)" and (Haskell) "(0 :: Uint16.Word16)" and (Scala) "0" | constant "1 :: uint16" \ (SML_word) "(Word16.fromInt 1)" and (Haskell) "(1 :: Uint16.Word16)" and (Scala) "1" | constant "plus :: uint16 \ _ \ _" \ (SML_word) "Word16.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toChar" | constant "uminus :: uint16 \ _" \ (SML_word) "Word16.~" and (Haskell) "negate" and (Scala) "(- _).toChar" | constant "minus :: uint16 \ _" \ (SML_word) "Word16.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toChar" | constant "times :: uint16 \ _ \ _" \ (SML_word) "Word16.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toChar" | constant "HOL.equal :: uint16 \ _ \ bool" \ (SML_word) "!((_ : Word16.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint16 :: equal \ (Haskell) - | constant "less_eq :: uint16 \ _ \ bool" \ (SML_word) "Word16.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) infixl 4 "<=" | constant "less :: uint16 \ _ \ bool" \ (SML_word) "Word16.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) infixl 4 "<" -| constant "NOT :: uint16 \ _" \ +| constant "Bit_Operations.not :: uint16 \ _" \ (SML_word) "Word16.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toChar" | constant "Bit_Operations.and :: uint16 \ _" \ (SML_word) "Word16.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toChar" | constant "Bit_Operations.or :: uint16 \ _" \ (SML_word) "Word16.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toChar" | constant "Bit_Operations.xor :: uint16 \ _" \ (SML_word) "Word16.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (Scala) "(_ ^ _).toChar" definition uint16_div :: "uint16 \ uint16 \ uint16" where "uint16_div x y = (if y = 0 then undefined ((div) :: uint16 \ _) x (0 :: uint16) else x div y)" definition uint16_mod :: "uint16 \ uint16 \ uint16" where "uint16_mod x y = (if y = 0 then undefined ((mod) :: uint16 \ _) x (0 :: uint16) else x mod y)" context includes undefined_transfer begin lemma div_uint16_code [code]: "x div y = (if y = 0 then 0 else uint16_div x y)" unfolding uint16_div_def by transfer (simp add: word_div_def) lemma mod_uint16_code [code]: "x mod y = (if y = 0 then x else uint16_mod x y)" unfolding uint16_mod_def by transfer (simp add: word_mod_def) lemma uint16_div_code [code]: "Rep_uint16 (uint16_div x y) = (if y = 0 then Rep_uint16 (undefined ((div) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x div Rep_uint16 y)" unfolding uint16_div_def by transfer simp lemma uint16_mod_code [code]: "Rep_uint16 (uint16_mod x y) = (if y = 0 then Rep_uint16 (undefined ((mod) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x mod Rep_uint16 y)" unfolding uint16_mod_def by transfer simp end code_printing constant uint16_div \ (SML_word) "Word16.div ((_), (_))" and (Haskell) "Prelude.div" and (Scala) "(_ '/ _).toChar" | constant uint16_mod \ (SML_word) "Word16.mod ((_), (_))" and (Haskell) "Prelude.mod" and (Scala) "(_ % _).toChar" definition uint16_test_bit :: "uint16 \ integer \ bool" where [code del]: "uint16_test_bit x n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint16_code [code]: "bit x n \ n < 16 \ uint16_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint16_test_bit_def by (transfer, simp, transfer, simp) lemma uint16_test_bit_code [code]: "uint16_test_bit w n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) w n else bit (Rep_uint16 w) (nat_of_integer n))" unfolding uint16_test_bit_def by (simp add: bit_uint16.rep_eq) code_printing constant uint16_test_bit \ (SML_word) "Uint16.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint16.test'_bit" definition uint16_set_bit :: "uint16 \ integer \ bool \ uint16" where [code del]: "uint16_set_bit x n b = (if n < 0 \ 15 < n then undefined (set_bit :: uint16 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint16_code [code]: "set_bit x n b = (if n < 16 then uint16_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint16_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint16_set_bit_code [code]: "Rep_uint16 (uint16_set_bit w n b) = (if n < 0 \ 15 < n then Rep_uint16 (undefined (set_bit :: uint16 \ _) w n b) else set_bit (Rep_uint16 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint16_set_bit_def by transfer simp code_printing constant uint16_set_bit \ (SML_word) "Uint16.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint16.set'_bit" definition uint16_shiftl :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftl x n = (if n < 0 \ 16 \ n then undefined (push_bit :: nat \ uint16 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint16_code [code]: "push_bit n x = (if n < 16 then uint16_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftl_def by transfer simp lemma uint16_shiftl_code [code]: "Rep_uint16 (uint16_shiftl w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (push_bit :: nat \ uint16 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftl_def by transfer simp code_printing constant uint16_shiftl \ (SML_word) "Uint16.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint16.shiftl" definition uint16_shiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftr x n = (if n < 0 \ 16 \ n then undefined (drop_bit :: nat \ uint16 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint16_code [code]: "drop_bit n x = (if n < 16 then uint16_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftr_def by transfer simp lemma uint16_shiftr_code [code]: "Rep_uint16 (uint16_shiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (drop_bit :: nat \ uint16 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftr_def by transfer simp code_printing constant uint16_shiftr \ (SML_word) "Uint16.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint16.shiftr" definition uint16_sshiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_sshiftr x n = (if n < 0 \ 16 \ n then undefined signed_drop_bit_uint16 n x else signed_drop_bit_uint16 (nat_of_integer n) x)" lemma sshiftr_uint16_code [code]: "signed_drop_bit_uint16 n x = (if n < 16 then uint16_sshiftr x (integer_of_nat n) else if bit x 15 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint16_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) lemma uint16_sshiftr_code [code]: "Rep_uint16 (uint16_sshiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined signed_drop_bit_uint16 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_sshiftr_def by transfer simp code_printing constant uint16_sshiftr \ (SML_word) "Uint16.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Int16) _)) :: Uint16.Word16)" and (Scala) "Uint16.shiftr'_signed" lemma uint16_msb_test_bit: "msb x \ bit (x :: uint16) 15" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint16_test_bit x 15" by (simp add: uint16_test_bit_def uint16_msb_test_bit) lemma uint16_of_int_code [code]: "uint16_of_int i = Uint16 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint16_code [code]: "int_of_uint16 x = int_of_integer (integer_of_uint16 x)" by (simp add: int_of_uint16.rep_eq integer_of_uint16_def) lemma uint16_of_nat_code [code]: "uint16_of_nat = uint16_of_int \ int" by transfer (simp add: fun_eq_iff) lemma nat_of_uint16_code [code]: "nat_of_uint16 x = nat_of_integer (integer_of_uint16 x)" unfolding integer_of_uint16_def including integer.lifting by transfer simp lemma integer_of_uint16_code [code]: "integer_of_uint16 n = integer_of_int (uint (Rep_uint16' n))" unfolding integer_of_uint16_def by transfer auto code_printing constant "integer_of_uint16" \ (SML_word) "Word16.toInt _ : IntInf.int" and (Haskell) "Prelude.toInteger" and (Scala) "BigInt" section \Quickcheck setup\ definition uint16_of_natural :: "natural \ uint16" where "uint16_of_natural x \ Uint16 (integer_of_natural x)" instantiation uint16 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint16 \ qc_random_cnv uint16_of_natural" definition "exhaustive_uint16 \ qc_exhaustive_cnv uint16_of_natural" definition "full_exhaustive_uint16 \ qc_full_exhaustive_cnv uint16_of_natural" instance .. end instantiation uint16 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint16 i in (x, 0xFFFF - x)" "0" "Typerep.Typerep (STR ''Uint16.uint16'') []" . definition "narrowing_uint16 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint16 itself \ _"]] lemmas partial_term_of_uint16 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint32.thy b/thys/Native_Word/Uint32.thy --- a/thys/Native_Word/Uint32.thy +++ b/thys/Native_Word/Uint32.thy @@ -1,691 +1,691 @@ (* Title: Uint32.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 32 bits\ theory Uint32 imports Code_Target_Word_Base Word_Type_Copies begin section \Type definition and primitive operations\ typedef uint32 = \UNIV :: 32 word set\ .. global_interpretation uint32: word_type_copy Abs_uint32 Rep_uint32 using type_definition_uint32 by (rule word_type_copy.intro) setup_lifting type_definition_uint32 declare uint32.of_word_of [code abstype] declare Quotient_uint32 [transfer_rule] instantiation uint32 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint32 :: uint32 is 0 . lift_definition one_uint32 :: uint32 is 1 . lift_definition plus_uint32 :: \uint32 \ uint32 \ uint32\ is \(+)\ . lift_definition uminus_uint32 :: \uint32 \ uint32\ is uminus . lift_definition minus_uint32 :: \uint32 \ uint32 \ uint32\ is \(-)\ . lift_definition times_uint32 :: \uint32 \ uint32 \ uint32\ is \(*)\ . lift_definition divide_uint32 :: \uint32 \ uint32 \ uint32\ is \(div)\ . lift_definition modulo_uint32 :: \uint32 \ uint32 \ uint32\ is \(mod)\ . lift_definition equal_uint32 :: \uint32 \ uint32 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint32 :: \uint32 \ uint32 \ bool\ is \(\)\ . lift_definition less_uint32 :: \uint32 \ uint32 \ bool\ is \(<)\ . global_interpretation uint32: word_type_copy_ring Abs_uint32 Rep_uint32 by standard (fact zero_uint32.rep_eq one_uint32.rep_eq plus_uint32.rep_eq uminus_uint32.rep_eq minus_uint32.rep_eq times_uint32.rep_eq divide_uint32.rep_eq modulo_uint32.rep_eq equal_uint32.rep_eq less_eq_uint32.rep_eq less_uint32.rep_eq)+ instance proof - show \OFCLASS(uint32, comm_ring_1_class)\ by (rule uint32.of_class_comm_ring_1) show \OFCLASS(uint32, semiring_modulo_class)\ by (fact uint32.of_class_semiring_modulo) show \OFCLASS(uint32, equal_class)\ by (fact uint32.of_class_equal) show \OFCLASS(uint32, linorder_class)\ by (fact uint32.of_class_linorder) qed end instantiation uint32 :: ring_bit_operations begin lift_definition bit_uint32 :: \uint32 \ nat \ bool\ is bit . -lift_definition not_uint32 :: \uint32 \ uint32\ is NOT . +lift_definition not_uint32 :: \uint32 \ uint32\ is \Bit_Operations.not\ . lift_definition and_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.and\ . lift_definition or_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.or\ . lift_definition xor_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.xor\ . lift_definition mask_uint32 :: \nat \ uint32\ is mask . lift_definition push_bit_uint32 :: \nat \ uint32 \ uint32\ is push_bit . lift_definition drop_bit_uint32 :: \nat \ uint32 \ uint32\ is drop_bit . lift_definition signed_drop_bit_uint32 :: \nat \ uint32 \ uint32\ is signed_drop_bit . lift_definition take_bit_uint32 :: \nat \ uint32 \ uint32\ is take_bit . lift_definition set_bit_uint32 :: \nat \ uint32 \ uint32\ is Bit_Operations.set_bit . lift_definition unset_bit_uint32 :: \nat \ uint32 \ uint32\ is unset_bit . lift_definition flip_bit_uint32 :: \nat \ uint32 \ uint32\ is flip_bit . global_interpretation uint32: word_type_copy_bits Abs_uint32 Rep_uint32 signed_drop_bit_uint32 by standard (fact bit_uint32.rep_eq not_uint32.rep_eq and_uint32.rep_eq or_uint32.rep_eq xor_uint32.rep_eq mask_uint32.rep_eq push_bit_uint32.rep_eq drop_bit_uint32.rep_eq signed_drop_bit_uint32.rep_eq take_bit_uint32.rep_eq set_bit_uint32.rep_eq unset_bit_uint32.rep_eq flip_bit_uint32.rep_eq)+ instance by (fact uint32.of_class_ring_bit_operations) end lift_definition uint32_of_nat :: \nat \ uint32\ is word_of_nat . lift_definition nat_of_uint32 :: \uint32 \ nat\ is unat . lift_definition uint32_of_int :: \int \ uint32\ is word_of_int . lift_definition int_of_uint32 :: \uint32 \ int\ is uint . context includes integer.lifting begin lift_definition Uint32 :: \integer \ uint32\ is word_of_int . lift_definition integer_of_uint32 :: \uint32 \ integer\ is uint . end global_interpretation uint32: word_type_copy_more Abs_uint32 Rep_uint32 signed_drop_bit_uint32 uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 apply standard apply (simp_all add: uint32_of_nat.rep_eq nat_of_uint32.rep_eq uint32_of_int.rep_eq int_of_uint32.rep_eq Uint32.rep_eq integer_of_uint32.rep_eq integer_eq_iff) done instantiation uint32 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint32 :: \uint32 \ nat\ is size . lift_definition msb_uint32 :: \uint32 \ bool\ is msb . lift_definition lsb_uint32 :: \uint32 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint32 :: \uint32 \ nat \ bool \ uint32\ where set_bit_uint32_eq: \set_bit_uint32 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint32_transfer [transfer_rule]: \(cr_uint32 ===> (=) ===> (\) ===> cr_uint32) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint32_eq [abs_def]) transfer_prover end lift_definition set_bits_uint32 :: \(nat \ bool) \ uint32\ is set_bits . lift_definition set_bits_aux_uint32 :: \(nat \ bool) \ nat \ uint32 \ uint32\ is set_bits_aux . global_interpretation uint32: word_type_copy_misc Abs_uint32 Rep_uint32 signed_drop_bit_uint32 uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 32 set_bits_aux_uint32 by (standard; transfer) simp_all instance using uint32.of_class_bit_comprehension uint32.of_class_set_bit uint32.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint32 \ (SML) \(* Test that words can handle numbers between 0 and 31 *) val _ = if 5 <= Word.wordSize then () else raise (Fail ("wordSize less than 5")); structure Uint32 : sig val set_bit : Word32.word -> IntInf.int -> bool -> Word32.word val shiftl : Word32.word -> IntInf.int -> Word32.word val shiftr : Word32.word -> IntInf.int -> Word32.word val shiftr_signed : Word32.word -> IntInf.int -> Word32.word val test_bit : Word32.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word32.orb (x, mask) else Word32.andb (x, Word32.notb mask) end fun shiftl x n = Word32.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word32.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word32.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word32.andb (x, Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word32.fromInt 0 end; (* struct Uint32 *)\ code_reserved SML Uint32 code_printing code_module Uint32 \ (Haskell) \module Uint32(Int32, Word32) where import Data.Int(Int32) import Data.Word(Word32)\ code_reserved Haskell Uint32 text \ OCaml and Scala provide only signed 32bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint32" \ (OCaml) \module Uint32 : sig val less : int32 -> int32 -> bool val less_eq : int32 -> int32 -> bool val set_bit : int32 -> Z.t -> bool -> int32 val shiftl : int32 -> Z.t -> int32 val shiftr : int32 -> Z.t -> int32 val shiftr_signed : int32 -> Z.t -> int32 val test_bit : int32 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y < 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y < 0;; let less_eq x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y <= 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y <= 0;; let set_bit x n b = let mask = Int32.shift_left Int32.one (Z.to_int n) in if b then Int32.logor x mask else Int32.logand x (Int32.lognot mask);; let shiftl x n = Int32.shift_left x (Z.to_int n);; let shiftr x n = Int32.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int32.shift_right x (Z.to_int n);; let test_bit x n = Int32.compare (Int32.logand x (Int32.shift_left Int32.one (Z.to_int n))) Int32.zero <> 0;; end;; (*struct Uint32*)\ code_reserved OCaml Uint32 code_printing code_module Uint32 \ (Scala) \object Uint32 { def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint32 */\ code_reserved Scala Uint32 text \ OCaml's conversion from Big\_int to int32 demands that the value fits int a signed 32-bit integer. The following justifies the implementation. \ context includes bit_operations_syntax begin definition Uint32_signed :: "integer \ uint32" where "Uint32_signed i = (if i < -(0x80000000) \ i \ 0x80000000 then undefined Uint32 i else Uint32 i)" lemma Uint32_code [code]: "Uint32 i = (let i' = i AND 0xFFFFFFFF in if bit i' 31 then Uint32_signed (i' - 0x100000000) else Uint32_signed i')" including undefined_transfer integer.lifting unfolding Uint32_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done lemma Uint32_signed_code [code]: "Rep_uint32 (Uint32_signed i) = (if i < -(0x80000000) \ i \ 0x80000000 then Rep_uint32 (undefined Uint32 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint32_signed_def Uint32_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint32_inverse) end text \ Avoid @{term Abs_uint32} in generated code, use @{term Rep_uint32'} instead. The symbolic implementations for code\_simp use @{term Rep_uint32}. The new destructor @{term Rep_uint32'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint32} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint32} ([code abstract] equations for @{typ uint32} may use @{term Rep_uint32} because these instances will be folded away.) To convert @{typ "32 word"} values into @{typ uint32}, use @{term "Abs_uint32'"}. \ definition Rep_uint32' where [simp]: "Rep_uint32' = Rep_uint32" lemma Rep_uint32'_transfer [transfer_rule]: "rel_fun cr_uint32 (=) (\x. x) Rep_uint32'" unfolding Rep_uint32'_def by(rule uint32.rep_transfer) lemma Rep_uint32'_code [code]: "Rep_uint32' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint32' :: "32 word \ uint32" is "\x :: 32 word. x" . lemma Abs_uint32'_code [code]: "Abs_uint32' x = Uint32 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint32 \ _"]] lemma term_of_uint32_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint32.uint32.Abs_uint32'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]], TR (STR ''Uint32.uint32'') []])) (term_of_class.term_of (Rep_uint32' x))" by(simp add: term_of_anything) code_printing type_constructor uint32 \ (SML) "Word32.word" and (Haskell) "Uint32.Word32" and (OCaml) "int32" and (Scala) "Int" and (Eval) "Word32.word" | constant Uint32 \ (SML) "Word32.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint32.Word32)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Word32)" and (Scala) "_.intValue" | constant Uint32_signed \ (OCaml) "Z.to'_int32" | constant "0 :: uint32" \ (SML) "(Word32.fromInt 0)" and (Haskell) "(0 :: Uint32.Word32)" and (OCaml) "Int32.zero" and (Scala) "0" | constant "1 :: uint32" \ (SML) "(Word32.fromInt 1)" and (Haskell) "(1 :: Uint32.Word32)" and (OCaml) "Int32.one" and (Scala) "1" | constant "plus :: uint32 \ _ " \ (SML) "Word32.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Int32.add" and (Scala) infixl 7 "+" | constant "uminus :: uint32 \ _" \ (SML) "Word32.~" and (Haskell) "negate" and (OCaml) "Int32.neg" and (Scala) "!(- _)" | constant "minus :: uint32 \ _" \ (SML) "Word32.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Int32.sub" and (Scala) infixl 7 "-" | constant "times :: uint32 \ _ \ _" \ (SML) "Word32.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Int32.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint32 \ _ \ bool" \ (SML) "!((_ : Word32.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int32.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint32 :: equal \ (Haskell) - | constant "less_eq :: uint32 \ _ \ bool" \ (SML) "Word32.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint32.less'_eq" and (Scala) "Uint32.less'_eq" | constant "less :: uint32 \ _ \ bool" \ (SML) "Word32.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint32.less" and (Scala) "Uint32.less" -| constant "NOT :: uint32 \ _" \ +| constant "Bit_Operations.not :: uint32 \ _" \ (SML) "Word32.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int32.lognot" and (Scala) "_.unary'_~" | constant "Bit_Operations.and :: uint32 \ _" \ (SML) "Word32.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int32.logand" and (Scala) infixl 3 "&" | constant "Bit_Operations.or :: uint32 \ _" \ (SML) "Word32.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int32.logor" and (Scala) infixl 1 "|" | constant "Bit_Operations.xor :: uint32 \ _" \ (SML) "Word32.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Int32.logxor" and (Scala) infixl 2 "^" definition uint32_divmod :: "uint32 \ uint32 \ uint32 \ uint32" where "uint32_divmod x y = (if y = 0 then (undefined ((div) :: uint32 \ _) x (0 :: uint32), undefined ((mod) :: uint32 \ _) x (0 :: uint32)) else (x div y, x mod y))" definition uint32_div :: "uint32 \ uint32 \ uint32" where "uint32_div x y = fst (uint32_divmod x y)" definition uint32_mod :: "uint32 \ uint32 \ uint32" where "uint32_mod x y = snd (uint32_divmod x y)" lemma div_uint32_code [code]: "x div y = (if y = 0 then 0 else uint32_div x y)" including undefined_transfer unfolding uint32_divmod_def uint32_div_def by transfer (simp add: word_div_def) lemma mod_uint32_code [code]: "x mod y = (if y = 0 then x else uint32_mod x y)" including undefined_transfer unfolding uint32_mod_def uint32_divmod_def by transfer (simp add: word_mod_def) definition uint32_sdiv :: "uint32 \ uint32 \ uint32" where [code del]: "uint32_sdiv x y = (if y = 0 then undefined ((div) :: uint32 \ _) x (0 :: uint32) else Abs_uint32 (Rep_uint32 x sdiv Rep_uint32 y))" definition div0_uint32 :: "uint32 \ uint32" where [code del]: "div0_uint32 x = undefined ((div) :: uint32 \ _) x (0 :: uint32)" declare [[code abort: div0_uint32]] definition mod0_uint32 :: "uint32 \ uint32" where [code del]: "mod0_uint32 x = undefined ((mod) :: uint32 \ _) x (0 :: uint32)" declare [[code abort: mod0_uint32]] lemma uint32_divmod_code [code]: "uint32_divmod x y = (if 0x80000000 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint32 x, mod0_uint32 x) else let q = push_bit 1 (uint32_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint32_divmod_def uint32_sdiv_def div0_uint32_def mod0_uint32_def less_eq_uint32.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint32_sdiv_code [code]: "Rep_uint32 (uint32_sdiv x y) = (if y = 0 then Rep_uint32 (undefined ((div) :: uint32 \ _) x (0 :: uint32)) else Rep_uint32 x sdiv Rep_uint32 y)" unfolding uint32_sdiv_def by(simp add: Abs_uint32_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint32_divmod_code} computes both with division only. \ code_printing constant uint32_div \ (SML) "Word32.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint32_mod \ (SML) "Word32.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint32_divmod \ (Haskell) "divmod" | constant uint32_sdiv \ (OCaml) "Int32.div" and (Scala) "_ '/ _" definition uint32_test_bit :: "uint32 \ integer \ bool" where [code del]: "uint32_test_bit x n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint32_code [code]: "bit x n \ n < 32 \ uint32_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint32_test_bit_def by (transfer, simp, transfer, simp) lemma uint32_test_bit_code [code]: "uint32_test_bit w n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) w n else bit (Rep_uint32 w) (nat_of_integer n))" unfolding uint32_test_bit_def by(simp add: bit_uint32.rep_eq) code_printing constant uint32_test_bit \ (SML) "Uint32.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint32.test'_bit" and (Scala) "Uint32.test'_bit" and (Eval) "(fn w => fn n => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_test'_bit out of bounds\") else Uint32.test'_bit w n)" definition uint32_set_bit :: "uint32 \ integer \ bool \ uint32" where [code del]: "uint32_set_bit x n b = (if n < 0 \ 31 < n then undefined (set_bit :: uint32 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint32_code [code]: "set_bit x n b = (if n < 32 then uint32_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint32_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint32_set_bit_code [code]: "Rep_uint32 (uint32_set_bit w n b) = (if n < 0 \ 31 < n then Rep_uint32 (undefined (set_bit :: uint32 \ _) w n b) else set_bit (Rep_uint32 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint32_set_bit_def by transfer simp code_printing constant uint32_set_bit \ (SML) "Uint32.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint32.set'_bit" and (Scala) "Uint32.set'_bit" and (Eval) "(fn w => fn n => fn b => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_set'_bit out of bounds\") else Uint32.set'_bit x n b)" definition uint32_shiftl :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftl x n = (if n < 0 \ 32 \ n then undefined (push_bit :: nat \ uint32 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint32_code [code]: "push_bit n x = (if n < 32 then uint32_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftl_def by transfer simp lemma uint32_shiftl_code [code]: "Rep_uint32 (uint32_shiftl w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (push_bit :: nat \ uint32 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftl_def by transfer simp code_printing constant uint32_shiftl \ (SML) "Uint32.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint32.shiftl" and (Scala) "Uint32.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftl out of bounds\" else Uint32.shiftl x i)" definition uint32_shiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftr x n = (if n < 0 \ 32 \ n then undefined (drop_bit :: nat \ uint32 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint32_code [code]: "drop_bit n x = (if n < 32 then uint32_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftr_def by transfer simp lemma uint32_shiftr_code [code]: "Rep_uint32 (uint32_shiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (drop_bit :: nat \ uint32 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftr_def by transfer simp code_printing constant uint32_shiftr \ (SML) "Uint32.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint32.shiftr" and (Scala) "Uint32.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr out of bounds\" else Uint32.shiftr x i)" definition uint32_sshiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_sshiftr x n = (if n < 0 \ 32 \ n then undefined signed_drop_bit_uint32 n x else signed_drop_bit_uint32 (nat_of_integer n) x)" lemma sshiftr_uint32_code [code]: "signed_drop_bit_uint32 n x = (if n < 32 then uint32_sshiftr x (integer_of_nat n) else if bit x 31 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint32_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) lemma uint32_sshiftr_code [code]: "Rep_uint32 (uint32_sshiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined signed_drop_bit_uint32 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_sshiftr_def by transfer simp code_printing constant uint32_sshiftr \ (SML) "Uint32.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Int32) _)) :: Uint32.Word32)" and (OCaml) "Uint32.shiftr'_signed" and (Scala) "Uint32.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr'_signed out of bounds\" else Uint32.shiftr'_signed x i)" context includes bit_operations_syntax begin lemma uint32_msb_test_bit: "msb x \ bit (x :: uint32) 31" by transfer (simp add: msb_word_iff_bit) lemma msb_uint32_code [code]: "msb x \ uint32_test_bit x 31" by (simp add: uint32_test_bit_def uint32_msb_test_bit) lemma uint32_of_int_code [code]: "uint32_of_int i = Uint32 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint32_code [code]: "int_of_uint32 x = int_of_integer (integer_of_uint32 x)" including integer.lifting by transfer simp lemma uint32_of_nat_code [code]: "uint32_of_nat = uint32_of_int \ int" by transfer (simp add: fun_eq_iff) lemma nat_of_uint32_code [code]: "nat_of_uint32 x = nat_of_integer (integer_of_uint32 x)" unfolding integer_of_uint32_def including integer.lifting by transfer simp definition integer_of_uint32_signed :: "uint32 \ integer" where "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_uint32 n)" lemma integer_of_uint32_signed_code [code]: "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_int (uint (Rep_uint32' n)))" by (simp add: integer_of_uint32_signed_def integer_of_uint32_def) lemma integer_of_uint32_code [code]: "integer_of_uint32 n = (if bit n 31 then integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 else integer_of_uint32_signed n)" proof - have \integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 = Bit_Operations.set_bit 31 (integer_of_uint32_signed (take_bit 31 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint32 n = Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))\ if \bit n 31\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint32 n) m = bit (Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint32_signed_def bit_simps) qed end code_printing constant "integer_of_uint32" \ (SML) "IntInf.fromLarge (Word32.toLargeInt _) : IntInf.int" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint32_signed" \ (OCaml) "Z.of'_int32" and (Scala) "BigInt" section \Quickcheck setup\ definition uint32_of_natural :: "natural \ uint32" where "uint32_of_natural x \ Uint32 (integer_of_natural x)" instantiation uint32 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint32 \ qc_random_cnv uint32_of_natural" definition "exhaustive_uint32 \ qc_exhaustive_cnv uint32_of_natural" definition "full_exhaustive_uint32 \ qc_full_exhaustive_cnv uint32_of_natural" instance .. end instantiation uint32 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint32 i in (x, 0xFFFFFFFF - x)" "0" "Typerep.Typerep (STR ''Uint32.uint32'') []" . definition "narrowing_uint32 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint32 itself \ _"]] lemmas partial_term_of_uint32 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint64.thy b/thys/Native_Word/Uint64.thy --- a/thys/Native_Word/Uint64.thy +++ b/thys/Native_Word/Uint64.thy @@ -1,888 +1,888 @@ (* Title: Uint64.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 64 bits\ theory Uint64 imports Code_Target_Word_Base Word_Type_Copies begin text \ PolyML (in version 5.7) provides a Word64 structure only when run in 64-bit mode. Therefore, we by default provide an implementation of 64-bit words using \verb$IntInf.int$ and masking. The code target \texttt{SML\_word} replaces this implementation and maps the operations directly to the \verb$Word64$ structure provided by the Standard ML implementations. The \verb$Eval$ target used by @{command value} and @{method eval} dynamically tests at runtime for the version of PolyML and uses PolyML's Word64 structure if it detects a 64-bit version which does not suffer from a division bug found in PolyML 5.6. \ section \Type definition and primitive operations\ typedef uint64 = \UNIV :: 64 word set\ .. global_interpretation uint64: word_type_copy Abs_uint64 Rep_uint64 using type_definition_uint64 by (rule word_type_copy.intro) setup_lifting type_definition_uint64 declare uint64.of_word_of [code abstype] declare Quotient_uint64 [transfer_rule] instantiation uint64 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint64 :: uint64 is 0 . lift_definition one_uint64 :: uint64 is 1 . lift_definition plus_uint64 :: \uint64 \ uint64 \ uint64\ is \(+)\ . lift_definition uminus_uint64 :: \uint64 \ uint64\ is uminus . lift_definition minus_uint64 :: \uint64 \ uint64 \ uint64\ is \(-)\ . lift_definition times_uint64 :: \uint64 \ uint64 \ uint64\ is \(*)\ . lift_definition divide_uint64 :: \uint64 \ uint64 \ uint64\ is \(div)\ . lift_definition modulo_uint64 :: \uint64 \ uint64 \ uint64\ is \(mod)\ . lift_definition equal_uint64 :: \uint64 \ uint64 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint64 :: \uint64 \ uint64 \ bool\ is \(\)\ . lift_definition less_uint64 :: \uint64 \ uint64 \ bool\ is \(<)\ . global_interpretation uint64: word_type_copy_ring Abs_uint64 Rep_uint64 by standard (fact zero_uint64.rep_eq one_uint64.rep_eq plus_uint64.rep_eq uminus_uint64.rep_eq minus_uint64.rep_eq times_uint64.rep_eq divide_uint64.rep_eq modulo_uint64.rep_eq equal_uint64.rep_eq less_eq_uint64.rep_eq less_uint64.rep_eq)+ instance proof - show \OFCLASS(uint64, comm_ring_1_class)\ by (rule uint64.of_class_comm_ring_1) show \OFCLASS(uint64, semiring_modulo_class)\ by (fact uint64.of_class_semiring_modulo) show \OFCLASS(uint64, equal_class)\ by (fact uint64.of_class_equal) show \OFCLASS(uint64, linorder_class)\ by (fact uint64.of_class_linorder) qed end instantiation uint64 :: ring_bit_operations begin lift_definition bit_uint64 :: \uint64 \ nat \ bool\ is bit . -lift_definition not_uint64 :: \uint64 \ uint64\ is NOT . +lift_definition not_uint64 :: \uint64 \ uint64\ is \Bit_Operations.not\ . lift_definition and_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.and\ . lift_definition or_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.or\ . lift_definition xor_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.xor\ . lift_definition mask_uint64 :: \nat \ uint64\ is mask . lift_definition push_bit_uint64 :: \nat \ uint64 \ uint64\ is push_bit . lift_definition drop_bit_uint64 :: \nat \ uint64 \ uint64\ is drop_bit . lift_definition signed_drop_bit_uint64 :: \nat \ uint64 \ uint64\ is signed_drop_bit . lift_definition take_bit_uint64 :: \nat \ uint64 \ uint64\ is take_bit . lift_definition set_bit_uint64 :: \nat \ uint64 \ uint64\ is Bit_Operations.set_bit . lift_definition unset_bit_uint64 :: \nat \ uint64 \ uint64\ is unset_bit . lift_definition flip_bit_uint64 :: \nat \ uint64 \ uint64\ is flip_bit . global_interpretation uint64: word_type_copy_bits Abs_uint64 Rep_uint64 signed_drop_bit_uint64 by standard (fact bit_uint64.rep_eq not_uint64.rep_eq and_uint64.rep_eq or_uint64.rep_eq xor_uint64.rep_eq mask_uint64.rep_eq push_bit_uint64.rep_eq drop_bit_uint64.rep_eq signed_drop_bit_uint64.rep_eq take_bit_uint64.rep_eq set_bit_uint64.rep_eq unset_bit_uint64.rep_eq flip_bit_uint64.rep_eq)+ instance by (fact uint64.of_class_ring_bit_operations) end lift_definition uint64_of_nat :: \nat \ uint64\ is word_of_nat . lift_definition nat_of_uint64 :: \uint64 \ nat\ is unat . lift_definition uint64_of_int :: \int \ uint64\ is word_of_int . lift_definition int_of_uint64 :: \uint64 \ int\ is uint . context includes integer.lifting begin lift_definition Uint64 :: \integer \ uint64\ is word_of_int . lift_definition integer_of_uint64 :: \uint64 \ integer\ is uint . end global_interpretation uint64: word_type_copy_more Abs_uint64 Rep_uint64 signed_drop_bit_uint64 uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 apply standard apply (simp_all add: uint64_of_nat.rep_eq nat_of_uint64.rep_eq uint64_of_int.rep_eq int_of_uint64.rep_eq Uint64.rep_eq integer_of_uint64.rep_eq integer_eq_iff) done instantiation uint64 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint64 :: \uint64 \ nat\ is size . lift_definition msb_uint64 :: \uint64 \ bool\ is msb . lift_definition lsb_uint64 :: \uint64 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint64 :: \uint64 \ nat \ bool \ uint64\ where set_bit_uint64_eq: \set_bit_uint64 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint64_transfer [transfer_rule]: \(cr_uint64 ===> (=) ===> (\) ===> cr_uint64) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint64_eq [abs_def]) transfer_prover end lift_definition set_bits_uint64 :: \(nat \ bool) \ uint64\ is set_bits . lift_definition set_bits_aux_uint64 :: \(nat \ bool) \ nat \ uint64 \ uint64\ is set_bits_aux . global_interpretation uint64: word_type_copy_misc Abs_uint64 Rep_uint64 signed_drop_bit_uint64 uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 64 set_bits_aux_uint64 by (standard; transfer) simp_all instance using uint64.of_class_bit_comprehension uint64.of_class_set_bit uint64.of_class_lsb by simp_all standard end section \Code setup\ text \ For SML, we generate an implementation of unsigned 64-bit words using \verb$IntInf.int$. If @{ML "LargeWord.wordSize > 63"} of the Isabelle/ML runtime environment holds, then we assume that there is also a \Word64\ structure available and accordingly replace the implementation for the target \verb$Eval$. \ code_printing code_module "Uint64" \ (SML) \(* Test that words can handle numbers between 0 and 63 *) val _ = if 6 <= Word.wordSize then () else raise (Fail ("wordSize less than 6")); structure Uint64 : sig eqtype uint64; val zero : uint64; val one : uint64; val fromInt : IntInf.int -> uint64; val toInt : uint64 -> IntInf.int; val toLarge : uint64 -> LargeWord.word; val fromLarge : LargeWord.word -> uint64 val plus : uint64 -> uint64 -> uint64; val minus : uint64 -> uint64 -> uint64; val times : uint64 -> uint64 -> uint64; val divide : uint64 -> uint64 -> uint64; val modulus : uint64 -> uint64 -> uint64; val negate : uint64 -> uint64; val less_eq : uint64 -> uint64 -> bool; val less : uint64 -> uint64 -> bool; val notb : uint64 -> uint64; val andb : uint64 -> uint64 -> uint64; val orb : uint64 -> uint64 -> uint64; val xorb : uint64 -> uint64 -> uint64; val shiftl : uint64 -> IntInf.int -> uint64; val shiftr : uint64 -> IntInf.int -> uint64; val shiftr_signed : uint64 -> IntInf.int -> uint64; val set_bit : uint64 -> IntInf.int -> bool -> uint64; val test_bit : uint64 -> IntInf.int -> bool; end = struct type uint64 = IntInf.int; val mask = 0xFFFFFFFFFFFFFFFF : IntInf.int; val zero = 0 : IntInf.int; val one = 1 : IntInf.int; fun fromInt x = IntInf.andb(x, mask); fun toInt x = x fun toLarge x = LargeWord.fromLargeInt (IntInf.toLarge x); fun fromLarge x = IntInf.fromLarge (LargeWord.toLargeInt x); fun plus x y = IntInf.andb(IntInf.+(x, y), mask); fun minus x y = IntInf.andb(IntInf.-(x, y), mask); fun negate x = IntInf.andb(IntInf.~(x), mask); fun times x y = IntInf.andb(IntInf.*(x, y), mask); fun divide x y = IntInf.div(x, y); fun modulus x y = IntInf.mod(x, y); fun less_eq x y = IntInf.<=(x, y); fun less x y = IntInf.<(x, y); fun notb x = IntInf.andb(IntInf.notb(x), mask); fun orb x y = IntInf.orb(x, y); fun andb x y = IntInf.andb(x, y); fun xorb x y = IntInf.xorb(x, y); val maxWord = IntInf.pow (2, Word.wordSize); fun shiftl x n = if n < maxWord then IntInf.andb(IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)), mask) else 0; fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else 0; val msb_mask = 0x8000000000000000 : IntInf.int; fun shiftr_signed x i = if IntInf.andb(x, msb_mask) = 0 then shiftr x i else if i >= 64 then 0xFFFFFFFFFFFFFFFF else let val x' = shiftr x i val m' = IntInf.andb(IntInf.<<(mask, Word.max(0w64 - Word.fromLargeInt (IntInf.toLarge i), 0w0)), mask) in IntInf.orb(x', m') end; fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else false; fun set_bit x n b = if n < 64 then if b then IntInf.orb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) else IntInf.andb (x, IntInf.notb (IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n)))) else x; end \ code_reserved SML Uint64 setup \ let val polyml64 = LargeWord.wordSize > 63; (* PolyML 5.6 has bugs in its Word64 implementation. We test for one such bug and refrain from using Word64 in that case. Testing is done with dynamic code evaluation such that the compiler does not choke on the Word64 structure, which need not be present in a 32bit environment. *) val error_msg = "Buggy Word64 structure"; val test_code = "val _ = if Word64.div (0w18446744073709551611 : Word64.word, 0w3) = 0w6148914691236517203 then ()\n" ^ "else raise (Fail \"" ^ error_msg ^ "\");"; val f = Exn.interruptible_capture (fn () => ML_Compiler.eval ML_Compiler.flags Position.none (ML_Lex.tokenize test_code)) val use_Word64 = polyml64 andalso (case f () of Exn.Res _ => true | Exn.Exn (e as ERROR m) => if String.isSuffix error_msg m then false else Exn.reraise e | Exn.Exn e => Exn.reraise e) ; val newline = "\n"; val content = "structure Uint64 : sig" ^ newline ^ " eqtype uint64;" ^ newline ^ " val zero : uint64;" ^ newline ^ " val one : uint64;" ^ newline ^ " val fromInt : IntInf.int -> uint64;" ^ newline ^ " val toInt : uint64 -> IntInf.int;" ^ newline ^ " val toLarge : uint64 -> LargeWord.word;" ^ newline ^ " val fromLarge : LargeWord.word -> uint64" ^ newline ^ " val plus : uint64 -> uint64 -> uint64;" ^ newline ^ " val minus : uint64 -> uint64 -> uint64;" ^ newline ^ " val times : uint64 -> uint64 -> uint64;" ^ newline ^ " val divide : uint64 -> uint64 -> uint64;" ^ newline ^ " val modulus : uint64 -> uint64 -> uint64;" ^ newline ^ " val negate : uint64 -> uint64;" ^ newline ^ " val less_eq : uint64 -> uint64 -> bool;" ^ newline ^ " val less : uint64 -> uint64 -> bool;" ^ newline ^ " val notb : uint64 -> uint64;" ^ newline ^ " val andb : uint64 -> uint64 -> uint64;" ^ newline ^ " val orb : uint64 -> uint64 -> uint64;" ^ newline ^ " val xorb : uint64 -> uint64 -> uint64;" ^ newline ^ " val shiftl : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr_signed : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val set_bit : uint64 -> IntInf.int -> bool -> uint64;" ^ newline ^ " val test_bit : uint64 -> IntInf.int -> bool;" ^ newline ^ "end = struct" ^ newline ^ "" ^ newline ^ "type uint64 = Word64.word;" ^ newline ^ "" ^ newline ^ "val zero = (0wx0 : uint64);" ^ newline ^ "" ^ newline ^ "val one = (0wx1 : uint64);" ^ newline ^ "" ^ newline ^ "fun fromInt x = Word64.fromLargeInt (IntInf.toLarge x);" ^ newline ^ "" ^ newline ^ "fun toInt x = IntInf.fromLarge (Word64.toLargeInt x);" ^ newline ^ "" ^ newline ^ "fun fromLarge x = Word64.fromLarge x;" ^ newline ^ "" ^ newline ^ "fun toLarge x = Word64.toLarge x;" ^ newline ^ "" ^ newline ^ "fun plus x y = Word64.+(x, y);" ^ newline ^ "" ^ newline ^ "fun minus x y = Word64.-(x, y);" ^ newline ^ "" ^ newline ^ "fun negate x = Word64.~(x);" ^ newline ^ "" ^ newline ^ "fun times x y = Word64.*(x, y);" ^ newline ^ "" ^ newline ^ "fun divide x y = Word64.div(x, y);" ^ newline ^ "" ^ newline ^ "fun modulus x y = Word64.mod(x, y);" ^ newline ^ "" ^ newline ^ "fun less_eq x y = Word64.<=(x, y);" ^ newline ^ "" ^ newline ^ "fun less x y = Word64.<(x, y);" ^ newline ^ "" ^ newline ^ "fun set_bit x n b =" ^ newline ^ " let val mask = Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ " in if b then Word64.orb (x, mask)" ^ newline ^ " else Word64.andb (x, Word64.notb mask)" ^ newline ^ " end" ^ newline ^ "" ^ newline ^ "fun shiftl x n =" ^ newline ^ " Word64.<< (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr x n =" ^ newline ^ " Word64.>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr_signed x n =" ^ newline ^ " Word64.~>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun test_bit x n =" ^ newline ^ " Word64.andb (x, Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word64.fromInt 0" ^ newline ^ "" ^ newline ^ "val notb = Word64.notb" ^ newline ^ "" ^ newline ^ "fun andb x y = Word64.andb(x, y);" ^ newline ^ "" ^ newline ^ "fun orb x y = Word64.orb(x, y);" ^ newline ^ "" ^ newline ^ "fun xorb x y = Word64.xorb(x, y);" ^ newline ^ "" ^ newline ^ "end (*struct Uint64*)" val target_SML64 = "SML_word"; in (if use_Word64 then Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(Code_Runtime.target, SOME (content, []))])) else I) #> Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(target_SML64, SOME (content, []))])) end \ code_printing code_module Uint64 \ (Haskell) \module Uint64(Int64, Word64) where import Data.Int(Int64) import Data.Word(Word64)\ code_reserved Haskell Uint64 text \ OCaml and Scala provide only signed 64bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint64" \ (OCaml) \module Uint64 : sig val less : int64 -> int64 -> bool val less_eq : int64 -> int64 -> bool val set_bit : int64 -> Z.t -> bool -> int64 val shiftl : int64 -> Z.t -> int64 val shiftr : int64 -> Z.t -> int64 val shiftr_signed : int64 -> Z.t -> int64 val test_bit : int64 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y < 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y < 0;; let less_eq x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y <= 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y <= 0;; let set_bit x n b = let mask = Int64.shift_left Int64.one (Z.to_int n) in if b then Int64.logor x mask else Int64.logand x (Int64.lognot mask);; let shiftl x n = Int64.shift_left x (Z.to_int n);; let shiftr x n = Int64.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int64.shift_right x (Z.to_int n);; let test_bit x n = Int64.compare (Int64.logand x (Int64.shift_left Int64.one (Z.to_int n))) Int64.zero <> 0;; end;; (*struct Uint64*)\ code_reserved OCaml Uint64 code_printing code_module Uint64 \ (Scala) \object Uint64 { def less(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Long, n: BigInt, b: Boolean) : Long = if (b) x | (1L << n.intValue) else x & (1L << n.intValue).unary_~ def shiftl(x: Long, n: BigInt) : Long = x << n.intValue def shiftr(x: Long, n: BigInt) : Long = x >>> n.intValue def shiftr_signed(x: Long, n: BigInt) : Long = x >> n.intValue def test_bit(x: Long, n: BigInt) : Boolean = (x & (1L << n.intValue)) != 0 } /* object Uint64 */\ code_reserved Scala Uint64 text \ OCaml's conversion from Big\_int to int64 demands that the value fits int a signed 64-bit integer. The following justifies the implementation. \ context includes bit_operations_syntax begin definition Uint64_signed :: "integer \ uint64" where "Uint64_signed i = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then undefined Uint64 i else Uint64 i)" lemma Uint64_code [code]: "Uint64 i = (let i' = i AND 0xFFFFFFFFFFFFFFFF in if bit i' 63 then Uint64_signed (i' - 0x10000000000000000) else Uint64_signed i')" including undefined_transfer integer.lifting unfolding Uint64_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done lemma Uint64_signed_code [code]: "Rep_uint64 (Uint64_signed i) = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then Rep_uint64 (undefined Uint64 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint64_signed_def Uint64_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint64_inverse) end text \ Avoid @{term Abs_uint64} in generated code, use @{term Rep_uint64'} instead. The symbolic implementations for code\_simp use @{term Rep_uint64}. The new destructor @{term Rep_uint64'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint64} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint64} ([code abstract] equations for @{typ uint64} may use @{term Rep_uint64} because these instances will be folded away.) To convert @{typ "64 word"} values into @{typ uint64}, use @{term "Abs_uint64'"}. \ definition Rep_uint64' where [simp]: "Rep_uint64' = Rep_uint64" lemma Rep_uint64'_transfer [transfer_rule]: "rel_fun cr_uint64 (=) (\x. x) Rep_uint64'" unfolding Rep_uint64'_def by(rule uint64.rep_transfer) lemma Rep_uint64'_code [code]: "Rep_uint64' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint64' :: "64 word \ uint64" is "\x :: 64 word. x" . lemma Abs_uint64'_code [code]: "Abs_uint64' x = Uint64 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint64 \ _"]] lemma term_of_uint64_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint64.uint64.Abs_uint64'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]]], TR (STR ''Uint64.uint64'') []])) (term_of_class.term_of (Rep_uint64' x))" by(simp add: term_of_anything) code_printing type_constructor uint64 \ (SML) "Uint64.uint64" and (Haskell) "Uint64.Word64" and (OCaml) "int64" and (Scala) "Long" | constant Uint64 \ (SML) "Uint64.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint64.Word64)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Word64)" and (Scala) "_.longValue" | constant Uint64_signed \ (OCaml) "Z.to'_int64" | constant "0 :: uint64" \ (SML) "Uint64.zero" and (Haskell) "(0 :: Uint64.Word64)" and (OCaml) "Int64.zero" and (Scala) "0" | constant "1 :: uint64" \ (SML) "Uint64.one" and (Haskell) "(1 :: Uint64.Word64)" and (OCaml) "Int64.one" and (Scala) "1" | constant "plus :: uint64 \ _ " \ (SML) "Uint64.plus" and (Haskell) infixl 6 "+" and (OCaml) "Int64.add" and (Scala) infixl 7 "+" | constant "uminus :: uint64 \ _" \ (SML) "Uint64.negate" and (Haskell) "negate" and (OCaml) "Int64.neg" and (Scala) "!(- _)" | constant "minus :: uint64 \ _" \ (SML) "Uint64.minus" and (Haskell) infixl 6 "-" and (OCaml) "Int64.sub" and (Scala) infixl 7 "-" | constant "times :: uint64 \ _ \ _" \ (SML) "Uint64.times" and (Haskell) infixl 7 "*" and (OCaml) "Int64.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint64 \ _ \ bool" \ (SML) "!((_ : Uint64.uint64) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int64.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint64 :: equal \ (Haskell) - | constant "less_eq :: uint64 \ _ \ bool" \ (SML) "Uint64.less'_eq" and (Haskell) infix 4 "<=" and (OCaml) "Uint64.less'_eq" and (Scala) "Uint64.less'_eq" | constant "less :: uint64 \ _ \ bool" \ (SML) "Uint64.less" and (Haskell) infix 4 "<" and (OCaml) "Uint64.less" and (Scala) "Uint64.less" -| constant "NOT :: uint64 \ _" \ +| constant "Bit_Operations.not :: uint64 \ _" \ (SML) "Uint64.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int64.lognot" and (Scala) "_.unary'_~" | constant "Bit_Operations.and :: uint64 \ _" \ (SML) "Uint64.andb" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int64.logand" and (Scala) infixl 3 "&" | constant "Bit_Operations.or :: uint64 \ _" \ (SML) "Uint64.orb" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int64.logor" and (Scala) infixl 1 "|" | constant "Bit_Operations.xor :: uint64 \ _" \ (SML) "Uint64.xorb" and (Haskell) "Data'_Bits.xor" and (OCaml) "Int64.logxor" and (Scala) infixl 2 "^" definition uint64_divmod :: "uint64 \ uint64 \ uint64 \ uint64" where "uint64_divmod x y = (if y = 0 then (undefined ((div) :: uint64 \ _) x (0 :: uint64), undefined ((mod) :: uint64 \ _) x (0 :: uint64)) else (x div y, x mod y))" definition uint64_div :: "uint64 \ uint64 \ uint64" where "uint64_div x y = fst (uint64_divmod x y)" definition uint64_mod :: "uint64 \ uint64 \ uint64" where "uint64_mod x y = snd (uint64_divmod x y)" lemma div_uint64_code [code]: "x div y = (if y = 0 then 0 else uint64_div x y)" including undefined_transfer unfolding uint64_divmod_def uint64_div_def by transfer (simp add: word_div_def) lemma mod_uint64_code [code]: "x mod y = (if y = 0 then x else uint64_mod x y)" including undefined_transfer unfolding uint64_mod_def uint64_divmod_def by transfer (simp add: word_mod_def) definition uint64_sdiv :: "uint64 \ uint64 \ uint64" where [code del]: "uint64_sdiv x y = (if y = 0 then undefined ((div) :: uint64 \ _) x (0 :: uint64) else Abs_uint64 (Rep_uint64 x sdiv Rep_uint64 y))" definition div0_uint64 :: "uint64 \ uint64" where [code del]: "div0_uint64 x = undefined ((div) :: uint64 \ _) x (0 :: uint64)" declare [[code abort: div0_uint64]] definition mod0_uint64 :: "uint64 \ uint64" where [code del]: "mod0_uint64 x = undefined ((mod) :: uint64 \ _) x (0 :: uint64)" declare [[code abort: mod0_uint64]] lemma uint64_divmod_code [code]: "uint64_divmod x y = (if 0x8000000000000000 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint64 x, mod0_uint64 x) else let q = push_bit 1 (uint64_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint64_divmod_def uint64_sdiv_def div0_uint64_def mod0_uint64_def less_eq_uint64.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint64_sdiv_code [code]: "Rep_uint64 (uint64_sdiv x y) = (if y = 0 then Rep_uint64 (undefined ((div) :: uint64 \ _) x (0 :: uint64)) else Rep_uint64 x sdiv Rep_uint64 y)" unfolding uint64_sdiv_def by(simp add: Abs_uint64_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint64_divmod_code} computes both with division only. \ code_printing constant uint64_div \ (SML) "Uint64.divide" and (Haskell) "Prelude.div" | constant uint64_mod \ (SML) "Uint64.modulus" and (Haskell) "Prelude.mod" | constant uint64_divmod \ (Haskell) "divmod" | constant uint64_sdiv \ (OCaml) "Int64.div" and (Scala) "_ '/ _" definition uint64_test_bit :: "uint64 \ integer \ bool" where [code del]: "uint64_test_bit x n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint64_code [code]: "bit x n \ n < 64 \ uint64_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint64_test_bit_def by transfer (auto dest: bit_imp_le_length) lemma uint64_test_bit_code [code]: "uint64_test_bit w n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) w n else bit (Rep_uint64 w) (nat_of_integer n))" unfolding uint64_test_bit_def by(simp add: bit_uint64.rep_eq) code_printing constant uint64_test_bit \ (SML) "Uint64.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint64.test'_bit" and (Scala) "Uint64.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_test'_bit out of bounds\") else Uint64.test'_bit x i)" definition uint64_set_bit :: "uint64 \ integer \ bool \ uint64" where [code del]: "uint64_set_bit x n b = (if n < 0 \ 63 < n then undefined (set_bit :: uint64 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint64_code [code]: "set_bit x n b = (if n < 64 then uint64_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint64_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint64_set_bit_code [code]: "Rep_uint64 (uint64_set_bit w n b) = (if n < 0 \ 63 < n then Rep_uint64 (undefined (set_bit :: uint64 \ _) w n b) else set_bit (Rep_uint64 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint64_set_bit_def by transfer simp code_printing constant uint64_set_bit \ (SML) "Uint64.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint64.set'_bit" and (Scala) "Uint64.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_set'_bit out of bounds\") else Uint64.set'_bit x i b)" definition uint64_shiftl :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftl x n = (if n < 0 \ 64 \ n then undefined (push_bit :: nat \ uint64 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint64_code [code]: "push_bit n x = (if n < 64 then uint64_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftl_def by transfer simp lemma uint64_shiftl_code [code]: "Rep_uint64 (uint64_shiftl w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (push_bit :: nat \ uint64 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftl_def by transfer simp code_printing constant uint64_shiftl \ (SML) "Uint64.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint64.shiftl" and (Scala) "Uint64.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftl out of bounds\") else Uint64.shiftl x i)" definition uint64_shiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftr x n = (if n < 0 \ 64 \ n then undefined (drop_bit :: nat \ uint64 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint64_code [code]: "drop_bit n x = (if n < 64 then uint64_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftr_def by transfer simp lemma uint64_shiftr_code [code]: "Rep_uint64 (uint64_shiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (drop_bit :: nat \ uint64 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftr_def by transfer simp code_printing constant uint64_shiftr \ (SML) "Uint64.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint64.shiftr" and (Scala) "Uint64.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr out of bounds\") else Uint64.shiftr x i)" definition uint64_sshiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_sshiftr x n = (if n < 0 \ 64 \ n then undefined signed_drop_bit_uint64 n x else signed_drop_bit_uint64 (nat_of_integer n) x)" lemma sshiftr_uint64_code [code]: "signed_drop_bit_uint64 n x = (if n < 64 then uint64_sshiftr x (integer_of_nat n) else if bit x 63 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint64_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) lemma uint64_sshiftr_code [code]: "Rep_uint64 (uint64_sshiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined signed_drop_bit_uint64 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_sshiftr_def by transfer simp code_printing constant uint64_sshiftr \ (SML) "Uint64.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Int64) _)) :: Uint64.Word64)" and (OCaml) "Uint64.shiftr'_signed" and (Scala) "Uint64.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr'_signed out of bounds\") else Uint64.shiftr'_signed x i)" context includes bit_operations_syntax begin lemma uint64_msb_test_bit: "msb x \ bit (x :: uint64) 63" by transfer (simp add: msb_word_iff_bit) lemma msb_uint64_code [code]: "msb x \ uint64_test_bit x 63" by (simp add: uint64_test_bit_def uint64_msb_test_bit) lemma uint64_of_int_code [code]: "uint64_of_int i = Uint64 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint64_code [code]: "int_of_uint64 x = int_of_integer (integer_of_uint64 x)" including integer.lifting by transfer simp lemma uint64_of_nat_code [code]: "uint64_of_nat = uint64_of_int \ int" by transfer (simp add: fun_eq_iff) lemma nat_of_uint64_code [code]: "nat_of_uint64 x = nat_of_integer (integer_of_uint64 x)" unfolding integer_of_uint64_def including integer.lifting by transfer simp definition integer_of_uint64_signed :: "uint64 \ integer" where "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_uint64 n)" lemma integer_of_uint64_signed_code [code]: "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_int (uint (Rep_uint64' n)))" by (simp add: integer_of_uint64_signed_def integer_of_uint64_def) lemma integer_of_uint64_code [code]: "integer_of_uint64 n = (if bit n 63 then integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 else integer_of_uint64_signed n)" proof - have \integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 = Bit_Operations.set_bit 63 (integer_of_uint64_signed (take_bit 63 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint64 n = Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))\ if \bit n 63\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint64 n) m = bit (Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint64_signed_def bit_simps) qed end code_printing constant "integer_of_uint64" \ (SML) "Uint64.toInt" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint64_signed" \ (OCaml) "Z.of'_int64" and (Scala) "BigInt" section \Quickcheck setup\ definition uint64_of_natural :: "natural \ uint64" where "uint64_of_natural x \ Uint64 (integer_of_natural x)" instantiation uint64 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint64 \ qc_random_cnv uint64_of_natural" definition "exhaustive_uint64 \ qc_exhaustive_cnv uint64_of_natural" definition "full_exhaustive_uint64 \ qc_full_exhaustive_cnv uint64_of_natural" instance .. end instantiation uint64 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint64 i in (x, 0xFFFFFFFFFFFFFFFF - x)" "0" "Typerep.Typerep (STR ''Uint64.uint64'') []" . definition "narrowing_uint64 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint64 itself \ _"]] lemmas partial_term_of_uint64 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint8.thy b/thys/Native_Word/Uint8.thy --- a/thys/Native_Word/Uint8.thy +++ b/thys/Native_Word/Uint8.thy @@ -1,610 +1,610 @@ (* Title: Uint8.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 8 bits\ theory Uint8 imports Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for OCaml code generation: OCaml does not provide an int8 type, so no special code generation for this type is set up. If the theory \Code_Target_Bits_Int\ is imported, the type \uint8\ is emulated via \<^typ>\8 word\. \ section \Type definition and primitive operations\ typedef uint8 = \UNIV :: 8 word set\ .. global_interpretation uint8: word_type_copy Abs_uint8 Rep_uint8 using type_definition_uint8 by (rule word_type_copy.intro) setup_lifting type_definition_uint8 declare uint8.of_word_of [code abstype] declare Quotient_uint8 [transfer_rule] instantiation uint8 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint8 :: uint8 is 0 . lift_definition one_uint8 :: uint8 is 1 . lift_definition plus_uint8 :: \uint8 \ uint8 \ uint8\ is \(+)\ . lift_definition uminus_uint8 :: \uint8 \ uint8\ is uminus . lift_definition minus_uint8 :: \uint8 \ uint8 \ uint8\ is \(-)\ . lift_definition times_uint8 :: \uint8 \ uint8 \ uint8\ is \(*)\ . lift_definition divide_uint8 :: \uint8 \ uint8 \ uint8\ is \(div)\ . lift_definition modulo_uint8 :: \uint8 \ uint8 \ uint8\ is \(mod)\ . lift_definition equal_uint8 :: \uint8 \ uint8 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint8 :: \uint8 \ uint8 \ bool\ is \(\)\ . lift_definition less_uint8 :: \uint8 \ uint8 \ bool\ is \(<)\ . global_interpretation uint8: word_type_copy_ring Abs_uint8 Rep_uint8 by standard (fact zero_uint8.rep_eq one_uint8.rep_eq plus_uint8.rep_eq uminus_uint8.rep_eq minus_uint8.rep_eq times_uint8.rep_eq divide_uint8.rep_eq modulo_uint8.rep_eq equal_uint8.rep_eq less_eq_uint8.rep_eq less_uint8.rep_eq)+ instance proof - show \OFCLASS(uint8, comm_ring_1_class)\ by (rule uint8.of_class_comm_ring_1) show \OFCLASS(uint8, semiring_modulo_class)\ by (fact uint8.of_class_semiring_modulo) show \OFCLASS(uint8, equal_class)\ by (fact uint8.of_class_equal) show \OFCLASS(uint8, linorder_class)\ by (fact uint8.of_class_linorder) qed end instantiation uint8 :: ring_bit_operations begin lift_definition bit_uint8 :: \uint8 \ nat \ bool\ is bit . -lift_definition not_uint8 :: \uint8 \ uint8\ is NOT . +lift_definition not_uint8 :: \uint8 \ uint8\ is \Bit_Operations.not\ . lift_definition and_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.and\ . lift_definition or_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.or\ . lift_definition xor_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.xor\ . lift_definition mask_uint8 :: \nat \ uint8\ is mask . lift_definition push_bit_uint8 :: \nat \ uint8 \ uint8\ is push_bit . lift_definition drop_bit_uint8 :: \nat \ uint8 \ uint8\ is drop_bit . lift_definition signed_drop_bit_uint8 :: \nat \ uint8 \ uint8\ is signed_drop_bit . lift_definition take_bit_uint8 :: \nat \ uint8 \ uint8\ is take_bit . lift_definition set_bit_uint8 :: \nat \ uint8 \ uint8\ is Bit_Operations.set_bit . lift_definition unset_bit_uint8 :: \nat \ uint8 \ uint8\ is unset_bit . lift_definition flip_bit_uint8 :: \nat \ uint8 \ uint8\ is flip_bit . global_interpretation uint8: word_type_copy_bits Abs_uint8 Rep_uint8 signed_drop_bit_uint8 by standard (fact bit_uint8.rep_eq not_uint8.rep_eq and_uint8.rep_eq or_uint8.rep_eq xor_uint8.rep_eq mask_uint8.rep_eq push_bit_uint8.rep_eq drop_bit_uint8.rep_eq signed_drop_bit_uint8.rep_eq take_bit_uint8.rep_eq set_bit_uint8.rep_eq unset_bit_uint8.rep_eq flip_bit_uint8.rep_eq)+ instance by (fact uint8.of_class_ring_bit_operations) end lift_definition uint8_of_nat :: \nat \ uint8\ is word_of_nat . lift_definition nat_of_uint8 :: \uint8 \ nat\ is unat . lift_definition uint8_of_int :: \int \ uint8\ is word_of_int . lift_definition int_of_uint8 :: \uint8 \ int\ is uint . context includes integer.lifting begin lift_definition Uint8 :: \integer \ uint8\ is word_of_int . lift_definition integer_of_uint8 :: \uint8 \ integer\ is uint . end global_interpretation uint8: word_type_copy_more Abs_uint8 Rep_uint8 signed_drop_bit_uint8 uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 apply standard apply (simp_all add: uint8_of_nat.rep_eq nat_of_uint8.rep_eq uint8_of_int.rep_eq int_of_uint8.rep_eq Uint8.rep_eq integer_of_uint8.rep_eq integer_eq_iff) done instantiation uint8 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint8 :: \uint8 \ nat\ is size . lift_definition msb_uint8 :: \uint8 \ bool\ is msb . lift_definition lsb_uint8 :: \uint8 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint8 :: \uint8 \ nat \ bool \ uint8\ where set_bit_uint8_eq: \set_bit_uint8 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint8_transfer [transfer_rule]: \(cr_uint8 ===> (=) ===> (\) ===> cr_uint8) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint8_eq [abs_def]) transfer_prover end lift_definition set_bits_uint8 :: \(nat \ bool) \ uint8\ is set_bits . lift_definition set_bits_aux_uint8 :: \(nat \ bool) \ nat \ uint8 \ uint8\ is set_bits_aux . global_interpretation uint8: word_type_copy_misc Abs_uint8 Rep_uint8 signed_drop_bit_uint8 uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 8 set_bits_aux_uint8 by (standard; transfer) simp_all instance using uint8.of_class_bit_comprehension uint8.of_class_set_bit uint8.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint8 \ (SML) \(* Test that words can handle numbers between 0 and 3 *) val _ = if 3 <= Word.wordSize then () else raise (Fail ("wordSize less than 3")); structure Uint8 : sig val set_bit : Word8.word -> IntInf.int -> bool -> Word8.word val shiftl : Word8.word -> IntInf.int -> Word8.word val shiftr : Word8.word -> IntInf.int -> Word8.word val shiftr_signed : Word8.word -> IntInf.int -> Word8.word val test_bit : Word8.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word8.orb (x, mask) else Word8.andb (x, Word8.notb mask) end fun shiftl x n = Word8.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word8.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word8.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word8.andb (x, Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word8.fromInt 0 end; (* struct Uint8 *)\ code_reserved SML Uint8 code_printing code_module Uint8 \ (Haskell) \module Uint8(Int8, Word8) where import Data.Int(Int8) import Data.Word(Word8)\ code_reserved Haskell Uint8 text \ Scala provides only signed 8bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module Uint8 \ (Scala) \object Uint8 { def less(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Byte, n: BigInt, b: Boolean) : Byte = if (b) (x | (1 << n.intValue)).toByte else (x & (1 << n.intValue).unary_~).toByte def shiftl(x: Byte, n: BigInt) : Byte = (x << n.intValue).toByte def shiftr(x: Byte, n: BigInt) : Byte = ((x & 255) >>> n.intValue).toByte def shiftr_signed(x: Byte, n: BigInt) : Byte = (x >> n.intValue).toByte def test_bit(x: Byte, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint8 */\ code_reserved Scala Uint8 text \ Avoid @{term Abs_uint8} in generated code, use @{term Rep_uint8'} instead. The symbolic implementations for code\_simp use @{term Rep_uint8}. The new destructor @{term Rep_uint8'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint8} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint8} ([code abstract] equations for @{typ uint8} may use @{term Rep_uint8} because these instances will be folded away.) To convert @{typ "8 word"} values into @{typ uint8}, use @{term "Abs_uint8'"}. \ definition Rep_uint8' where [simp]: "Rep_uint8' = Rep_uint8" lemma Rep_uint8'_transfer [transfer_rule]: "rel_fun cr_uint8 (=) (\x. x) Rep_uint8'" unfolding Rep_uint8'_def by(rule uint8.rep_transfer) lemma Rep_uint8'_code [code]: "Rep_uint8' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint8' :: "8 word \ uint8" is "\x :: 8 word. x" . lemma Abs_uint8'_code [code]: "Abs_uint8' x = Uint8 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint8 \ _"]] lemma term_of_uint8_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint8.uint8.Abs_uint8'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]], TR (STR ''Uint8.uint8'') []])) (term_of_class.term_of (Rep_uint8' x))" by(simp add: term_of_anything) lemma Uin8_code [code]: "Rep_uint8 (Uint8 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint8_def int_of_integer_symbolic_def by(simp add: Abs_uint8_inverse) code_printing type_constructor uint8 \ (SML) "Word8.word" and (Haskell) "Uint8.Word8" and (Scala) "Byte" | constant Uint8 \ (SML) "Word8.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint8.Word8)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Word8)" and (Scala) "_.byteValue" | constant "0 :: uint8" \ (SML) "(Word8.fromInt 0)" and (Haskell) "(0 :: Uint8.Word8)" and (Scala) "0.toByte" | constant "1 :: uint8" \ (SML) "(Word8.fromInt 1)" and (Haskell) "(1 :: Uint8.Word8)" and (Scala) "1.toByte" | constant "plus :: uint8 \ _ \ _" \ (SML) "Word8.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toByte" | constant "uminus :: uint8 \ _" \ (SML) "Word8.~" and (Haskell) "negate" and (Scala) "(- _).toByte" | constant "minus :: uint8 \ _" \ (SML) "Word8.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toByte" | constant "times :: uint8 \ _ \ _" \ (SML) "Word8.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toByte" | constant "HOL.equal :: uint8 \ _ \ bool" \ (SML) "!((_ : Word8.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint8 :: equal \ (Haskell) - | constant "less_eq :: uint8 \ _ \ bool" \ (SML) "Word8.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) "Uint8.less'_eq" | constant "less :: uint8 \ _ \ bool" \ (SML) "Word8.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) "Uint8.less" -| constant "NOT :: uint8 \ _" \ +| constant "Bit_Operations.not :: uint8 \ _" \ (SML) "Word8.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toByte" | constant "Bit_Operations.and :: uint8 \ _" \ (SML) "Word8.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toByte" | constant "Bit_Operations.or :: uint8 \ _" \ (SML) "Word8.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toByte" | constant "Bit_Operations.xor :: uint8 \ _" \ (SML) "Word8.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (Scala) "(_ ^ _).toByte" definition uint8_divmod :: "uint8 \ uint8 \ uint8 \ uint8" where "uint8_divmod x y = (if y = 0 then (undefined ((div) :: uint8 \ _) x (0 :: uint8), undefined ((mod) :: uint8 \ _) x (0 :: uint8)) else (x div y, x mod y))" definition uint8_div :: "uint8 \ uint8 \ uint8" where "uint8_div x y = fst (uint8_divmod x y)" definition uint8_mod :: "uint8 \ uint8 \ uint8" where "uint8_mod x y = snd (uint8_divmod x y)" lemma div_uint8_code [code]: "x div y = (if y = 0 then 0 else uint8_div x y)" including undefined_transfer unfolding uint8_divmod_def uint8_div_def by transfer (simp add: word_div_def) lemma mod_uint8_code [code]: "x mod y = (if y = 0 then x else uint8_mod x y)" including undefined_transfer unfolding uint8_mod_def uint8_divmod_def by transfer (simp add: word_mod_def) definition uint8_sdiv :: "uint8 \ uint8 \ uint8" where "uint8_sdiv x y = (if y = 0 then undefined ((div) :: uint8 \ _) x (0 :: uint8) else Abs_uint8 (Rep_uint8 x sdiv Rep_uint8 y))" definition div0_uint8 :: "uint8 \ uint8" where [code del]: "div0_uint8 x = undefined ((div) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: div0_uint8]] definition mod0_uint8 :: "uint8 \ uint8" where [code del]: "mod0_uint8 x = undefined ((mod) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: mod0_uint8]] lemma uint8_divmod_code [code]: "uint8_divmod x y = (if 0x80 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint8 x, mod0_uint8 x) else let q = push_bit 1 (uint8_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint8_divmod_def uint8_sdiv_def div0_uint8_def mod0_uint8_def less_eq_uint8.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint8_sdiv_code [code]: "Rep_uint8 (uint8_sdiv x y) = (if y = 0 then Rep_uint8 (undefined ((div) :: uint8 \ _) x (0 :: uint8)) else Rep_uint8 x sdiv Rep_uint8 y)" unfolding uint8_sdiv_def by(simp add: Abs_uint8_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint8_divmod_code} computes both with division only. \ code_printing constant uint8_div \ (SML) "Word8.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint8_mod \ (SML) "Word8.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint8_divmod \ (Haskell) "divmod" | constant uint8_sdiv \ (Scala) "(_ '/ _).toByte" definition uint8_test_bit :: "uint8 \ integer \ bool" where [code del]: "uint8_test_bit x n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint8_code [code]: "bit x n \ n < 8 \ uint8_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint8_test_bit_def by (transfer, simp, transfer, simp) lemma uint8_test_bit_code [code]: "uint8_test_bit w n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) w n else bit (Rep_uint8 w) (nat_of_integer n))" unfolding uint8_test_bit_def by (simp add: bit_uint8.rep_eq) code_printing constant uint8_test_bit \ (SML) "Uint8.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint8.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_test'_bit out of bounds\") else Uint8.test'_bit x i)" definition uint8_set_bit :: "uint8 \ integer \ bool \ uint8" where [code del]: "uint8_set_bit x n b = (if n < 0 \ 7 < n then undefined (set_bit :: uint8 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint8_code [code]: "set_bit x n b = (if n < 8 then uint8_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint8_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint8_set_bit_code [code]: "Rep_uint8 (uint8_set_bit w n b) = (if n < 0 \ 7 < n then Rep_uint8 (undefined (set_bit :: uint8 \ _) w n b) else set_bit (Rep_uint8 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint8_set_bit_def by transfer simp code_printing constant uint8_set_bit \ (SML) "Uint8.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint8.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_set'_bit out of bounds\") else Uint8.set'_bit x i b)" definition uint8_shiftl :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftl x n = (if n < 0 \ 8 \ n then undefined (push_bit :: nat \ uint8 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint8_code [code]: "push_bit n x = (if n < 8 then uint8_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftl_def by transfer simp lemma uint8_shiftl_code [code]: "Rep_uint8 (uint8_shiftl w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (push_bit :: nat \ uint8 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftl_def by transfer simp code_printing constant uint8_shiftl \ (SML) "Uint8.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint8.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftl out of bounds\") else Uint8.shiftl x i)" definition uint8_shiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftr x n = (if n < 0 \ 8 \ n then undefined (drop_bit :: _ \ _ \ uint8) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint8_code [code]: "drop_bit n x = (if n < 8 then uint8_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftr_def by transfer simp lemma uint8_shiftr_code [code]: "Rep_uint8 (uint8_shiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (drop_bit :: _ \ _ \ uint8) w n) else drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftr_def by transfer simp code_printing constant uint8_shiftr \ (SML) "Uint8.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint8.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftr out of bounds\") else Uint8.shiftr x i)" definition uint8_sshiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_sshiftr x n = (if n < 0 \ 8 \ n then undefined signed_drop_bit_uint8 n x else signed_drop_bit_uint8 (nat_of_integer n) x)" lemma sshiftr_uint8_code [code]: "signed_drop_bit_uint8 n x = (if n < 8 then uint8_sshiftr x (integer_of_nat n) else if bit x 7 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint8_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) lemma uint8_sshiftr_code [code]: "Rep_uint8 (uint8_sshiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined signed_drop_bit_uint8 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_sshiftr_def by transfer simp code_printing constant uint8_sshiftr \ (SML) "Uint8.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Int8) _)) :: Uint8.Word8)" and (Scala) "Uint8.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_sshiftr out of bounds\") else Uint8.shiftr'_signed x i)" context includes bit_operations_syntax begin lemma uint8_msb_test_bit: "msb x \ bit (x :: uint8) 7" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint8_test_bit x 7" by (simp add: uint8_test_bit_def uint8_msb_test_bit) lemma uint8_of_int_code [code]: "uint8_of_int i = Uint8 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint8_code [code]: "int_of_uint8 x = int_of_integer (integer_of_uint8 x)" by (simp add: int_of_uint8.rep_eq integer_of_uint8_def) lemma uint8_of_nat_code [code]: "uint8_of_nat = uint8_of_int \ int" by transfer (simp add: fun_eq_iff) lemma nat_of_uint8_code [code]: "nat_of_uint8 x = nat_of_integer (integer_of_uint8 x)" unfolding integer_of_uint8_def including integer.lifting by transfer simp definition integer_of_uint8_signed :: "uint8 \ integer" where "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_uint8 n)" lemma integer_of_uint8_signed_code [code]: "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_int (uint (Rep_uint8' n)))" by (simp add: integer_of_uint8_signed_def integer_of_uint8_def) lemma integer_of_uint8_code [code]: "integer_of_uint8 n = (if bit n 7 then integer_of_uint8_signed (n AND 0x7F) OR 0x80 else integer_of_uint8_signed n)" proof - have \integer_of_uint8_signed (n AND 0x7F) OR 0x80 = Bit_Operations.set_bit 7 (integer_of_uint8_signed (take_bit 7 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint8 n = Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))\ if \bit n 7\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint8 n) m = bit (Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint8_signed_def bit_simps) qed end code_printing constant "integer_of_uint8" \ (SML) "IntInf.fromLarge (Word8.toLargeInt _)" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint8_signed" \ (Scala) "BigInt" section \Quickcheck setup\ definition uint8_of_natural :: "natural \ uint8" where "uint8_of_natural x \ Uint8 (integer_of_natural x)" instantiation uint8 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint8 \ qc_random_cnv uint8_of_natural" definition "exhaustive_uint8 \ qc_exhaustive_cnv uint8_of_natural" definition "full_exhaustive_uint8 \ qc_full_exhaustive_cnv uint8_of_natural" instance .. end instantiation uint8 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint8 i in (x, 0xFF - x)" "0" "Typerep.Typerep (STR ''Uint8.uint8'') []" . definition "narrowing_uint8 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint8 itself \ _"]] lemmas partial_term_of_uint8 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Word_Type_Copies.thy b/thys/Native_Word/Word_Type_Copies.thy --- a/thys/Native_Word/Word_Type_Copies.thy +++ b/thys/Native_Word/Word_Type_Copies.thy @@ -1,341 +1,341 @@ (* Title: Word_Type_Copies.thy Author: Florian Haftmann, TU Muenchen *) chapter \Systematic approach towards type copies of word type\ theory Word_Type_Copies imports "HOL-Library.Word" "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" Code_Target_Word_Base begin lemma int_of_integer_unsigned_eq [simp]: \int_of_integer (unsigned w) = uint w\ by transfer simp lemma int_of_integer_signed_eq [simp]: \int_of_integer (signed w) = sint w\ by transfer simp lemma word_of_int_of_integer_eq [simp]: \word_of_int (int_of_integer k) = word_of_integer k\ by (simp add: word_of_integer.rep_eq) text \The lifting machinery is not localized, hence the abstract proofs are carried out using morphisms.\ locale word_type_copy = fixes of_word :: \'b::len word \ 'a\ and word_of :: \'a \ 'b word\ assumes type_definition: \type_definition word_of of_word UNIV\ begin lemma word_of_word: \word_of (of_word w) = w\ using type_definition by (simp add: type_definition_def) lemma of_word_of [code abstype]: \of_word (word_of p) = p\ \ \Use an abstract type for code generation to disable pattern matching on \<^term>\of_word\.\ using type_definition by (simp add: type_definition_def) lemma word_of_eqI: \p = q\ if \word_of p = word_of q\ proof - from that have \of_word (word_of p) = of_word (word_of q)\ by simp then show ?thesis by (simp add: of_word_of) qed lemma eq_iff_word_of: \p = q \ word_of p = word_of q\ by (auto intro: word_of_eqI) end bundle constraintless begin declaration \ let val cs = map (rpair NONE o fst o dest_Const) [\<^term>\0\, \<^term>\(+)\, \<^term>\uminus\, \<^term>\(-)\, \<^term>\1\, \<^term>\(*)\, \<^term>\(div)\, \<^term>\(mod)\, \<^term>\HOL.equal\, \<^term>\(\)\, \<^term>\(<)\, \<^term>\(dvd)\, \<^term>\of_bool\, \<^term>\numeral\, \<^term>\of_nat\, \<^term>\bit\, - \<^term>\NOT\, \<^term>\Bit_Operations.and\, \<^term>\Bit_Operations.or\, \<^term>\Bit_Operations.xor\, \<^term>\mask\, + \<^term>\Bit_Operations.not\, \<^term>\Bit_Operations.and\, \<^term>\Bit_Operations.or\, \<^term>\Bit_Operations.xor\, \<^term>\mask\, \<^term>\push_bit\, \<^term>\drop_bit\, \<^term>\take_bit\, \<^term>\Bit_Operations.set_bit\, \<^term>\unset_bit\, \<^term>\flip_bit\, \<^term>\msb\, \<^term>\lsb\, \<^term>\size\, \<^term>\Generic_set_bit.set_bit\, \<^term>\set_bits\] in K (Context.mapping I (fold Proof_Context.add_const_constraint cs)) end \ end locale word_type_copy_ring = word_type_copy opening constraintless + constrains word_of :: \'a \ 'b::len word\ assumes word_of_0 [code]: \word_of 0 = 0\ and word_of_1 [code]: \word_of 1 = 1\ and word_of_add [code]: \word_of (p + q) = word_of p + word_of q\ and word_of_minus [code]: \word_of (- p) = - (word_of p)\ and word_of_diff [code]: \word_of (p - q) = word_of p - word_of q\ and word_of_mult [code]: \word_of (p * q) = word_of p * word_of q\ and word_of_div [code]: \word_of (p div q) = word_of p div word_of q\ and word_of_mod [code]: \word_of (p mod q) = word_of p mod word_of q\ and equal_iff_word_of [code]: \HOL.equal p q \ HOL.equal (word_of p) (word_of q)\ and less_eq_iff_word_of [code]: \p \ q \ word_of p \ word_of q\ and less_iff_word_of [code]: \p < q \ word_of p < word_of q\ begin lemma of_class_comm_ring_1: \OFCLASS('a, comm_ring_1_class)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_add word_of_minus word_of_diff word_of_mult algebra_simps) lemma of_class_semiring_modulo: \OFCLASS('a, semiring_modulo_class)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_add word_of_minus word_of_diff word_of_mult word_of_mod word_of_div algebra_simps mod_mult_div_eq) lemma of_class_equal: \OFCLASS('a, equal_class)\ by standard (simp add: eq_iff_word_of equal_iff_word_of equal) lemma of_class_linorder: \OFCLASS('a, linorder_class)\ by standard (auto simp add: eq_iff_word_of less_eq_iff_word_of less_iff_word_of) end locale word_type_copy_bits = word_type_copy_ring opening constraintless bit_operations_syntax + constrains word_of :: \'a::{comm_ring_1, semiring_modulo, equal, linorder} \ 'b::len word\ fixes signed_drop_bit :: \nat \ 'a \ 'a\ assumes bit_eq_word_of [code]: \bit p = bit (word_of p)\ and word_of_not [code]: \word_of (NOT p) = NOT (word_of p)\ and word_of_and [code]: \word_of (p AND q) = word_of p AND word_of q\ and word_of_or [code]: \word_of (p OR q) = word_of p OR word_of q\ and word_of_xor [code]: \word_of (p XOR q) = word_of p XOR word_of q\ and word_of_mask [code]: \word_of (mask n) = mask n\ and word_of_push_bit [code]: \word_of (push_bit n p) = push_bit n (word_of p)\ and word_of_drop_bit [code]: \word_of (drop_bit n p) = drop_bit n (word_of p)\ and word_of_signed_drop_bit [code]: \word_of (signed_drop_bit n p) = Word.signed_drop_bit n (word_of p)\ and word_of_take_bit [code]: \word_of (take_bit n p) = take_bit n (word_of p)\ and word_of_set_bit [code]: \word_of (Bit_Operations.set_bit n p) = Bit_Operations.set_bit n (word_of p)\ and word_of_unset_bit [code]: \word_of (unset_bit n p) = unset_bit n (word_of p)\ and word_of_flip_bit [code]: \word_of (flip_bit n p) = flip_bit n (word_of p)\ begin lemma word_of_bool: \word_of (of_bool n) = of_bool n\ by (simp add: word_of_0 word_of_1) lemma word_of_nat: \word_of (of_nat n) = of_nat n\ by (induction n) (simp_all add: word_of_0 word_of_1 word_of_add) lemma word_of_numeral [simp]: \word_of (numeral n) = numeral n\ proof - have \word_of (of_nat (numeral n)) = of_nat (numeral n)\ by (simp only: word_of_nat) then show ?thesis by simp qed lemma word_of_power: \word_of (p ^ n) = word_of p ^ n\ by (induction n) (simp_all add: word_of_1 word_of_mult word_of_numeral) lemma even_iff_word_of: \2 dvd p \ 2 dvd word_of p\ (is \?P \ ?Q\) proof assume ?P then obtain q where \p = 2 * q\ .. then show ?Q by (simp add: word_of_mult word_of_numeral) next assume ?Q then obtain w where \word_of p = 2 * w\ .. then have \of_word (word_of p) = of_word (2 * w)\ by simp then have \p = 2 * of_word w\ by (simp add: eq_iff_word_of word_of_word word_of_mult word_of_numeral) then show ?P by simp qed lemma of_class_ring_bit_operations: \OFCLASS('a, ring_bit_operations_class)\ proof - have induct: \P p\ if stable: \(\p. p div 2 = p \ P p)\ and rec: \(\p b. P p \ (of_bool b + 2 * p) div 2 = p \ P (of_bool b + 2 * p))\ for p :: 'a and P proof - from stable have stable': \(\p. word_of p div 2 = word_of p \ P p)\ by (simp add: eq_iff_word_of word_of_div word_of_numeral) from rec have rec': \(\p b. P p \ (of_bool b + 2 * word_of p) div 2 = word_of p \ P (of_bool b + 2 * p))\ by (simp add: eq_iff_word_of word_of_add word_of_bool word_of_mult word_of_div word_of_numeral) define w where \w = word_of p\ then have \p = of_word w\ by (simp add: of_word_of) also have \P (of_word w)\ proof (induction w rule: bits_induct) case (stable w) show ?case by (rule stable') (simp add: word_of_word stable) next case (rec w b) have \P (of_bool b + 2 * of_word w)\ by (rule rec') (simp_all add: word_of_word rec) also have \of_bool b + 2 * of_word w = of_word (of_bool b + 2 * w)\ by (simp add: eq_iff_word_of word_of_word word_of_add word_of_1 word_of_mult word_of_numeral word_of_0) finally show ?case . qed finally show \P p\ . qed have \class.semiring_parity_axioms (+) (0::'a) (*) 1 (mod)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_numeral even_iff_word_of word_of_mod even_iff_mod_2_eq_zero) with of_class_semiring_modulo have \OFCLASS('a, semiring_parity_class)\ by (rule semiring_parity_class.intro) moreover have \class.semiring_bits_axioms (+) (-) (0::'a) (*) 1 (div) (mod) bit\ apply (standard, fact induct) apply (simp_all only: eq_iff_word_of word_of_0 word_of_1 word_of_bool word_of_numeral word_of_add word_of_diff word_of_mult word_of_div word_of_mod word_of_power even_iff_word_of bit_eq_word_of push_bit_take_bit drop_bit_take_bit even_drop_bit_iff_not_bit flip: push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod mask_eq_exp_minus_1) apply (auto simp add: ac_simps bit_simps drop_bit_exp_eq push_bit_of_1) done ultimately have \OFCLASS('a, semiring_bits_class)\ by (rule semiring_bits_class.intro) moreover have \class.semiring_bit_operations_axioms (+) (-) (*) (1::'a) (div) (mod) bit (AND) (OR) (XOR) mask Bit_Operations.set_bit unset_bit flip_bit push_bit drop_bit take_bit\ by standard (simp_all add: eq_iff_word_of word_of_push_bit word_of_power bit_eq_word_of word_of_and word_of_or word_of_xor word_of_mask word_of_diff word_of_1 bit_simps word_of_set_bit set_bit_eq_or word_of_unset_bit word_of_flip_bit flip_bit_eq_xor word_of_mult word_of_drop_bit word_of_div word_of_take_bit word_of_mod flip: mask_eq_exp_minus_1 push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod) ultimately have \OFCLASS('a, semiring_bit_operations_class)\ by (rule semiring_bit_operations_class.intro) moreover have \OFCLASS('a, ring_parity_class)\ using \OFCLASS('a, semiring_parity_class)\ by (rule ring_parity_class.intro) standard moreover have \class.ring_bit_operations_axioms (+) (-) (0::'a) (*) 1 bit uminus NOT\ by standard (simp_all add: eq_iff_word_of word_of_power word_of_numeral bit_eq_word_of word_of_diff word_of_1 bit_simps linorder_not_le word_of_not word_of_0 word_of_minus minus_eq_not_minus_1) ultimately show \OFCLASS('a, ring_bit_operations_class)\ by (rule ring_bit_operations_class.intro) qed lemma [code]: \take_bit n a = a AND mask n\ for a :: 'a by (simp add: eq_iff_word_of word_of_take_bit word_of_and word_of_mask take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: 'a) OR mask n\ \mask 0 = (0 :: 'a)\ by (simp_all add: eq_iff_word_of word_of_mask word_of_or word_of_push_bit word_of_0 word_of_1 mask_Suc_exp push_bit_of_1) lemma [code]: \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: 'a by (simp add: eq_iff_word_of word_of_set_bit word_of_or word_of_push_bit word_of_1 set_bit_eq_or) lemma [code]: \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: 'a by (simp add: eq_iff_word_of word_of_unset_bit word_of_and word_of_not word_of_push_bit word_of_1 unset_bit_eq_and_not) lemma [code]: \flip_bit n w = w XOR push_bit n 1\ for w :: 'a by (simp add: eq_iff_word_of word_of_flip_bit word_of_xor word_of_push_bit word_of_1 flip_bit_eq_xor) end locale word_type_copy_more = word_type_copy_bits + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ fixes of_nat :: \nat \ 'a\ and nat_of :: \'a \ nat\ and of_int :: \int \ 'a\ and int_of :: \'a \ int\ and of_integer :: \integer \ 'a\ and integer_of :: \'a \ integer\ assumes word_of_nat_eq: \word_of (of_nat n) = word_of_nat n\ and nat_of_eq_word_of: \nat_of p = unat (word_of p)\ and word_of_int_eq: \word_of (of_int k) = word_of_int k\ and int_of_eq_word_of: \int_of p = uint (word_of p)\ and word_of_integer_eq: \word_of (of_integer l) = word_of_integer l\ and integer_of_eq_word_of: \integer_of p = unsigned (word_of p)\ begin lemma of_word_numeral [code_post]: \of_word (numeral n) = numeral n\ by (simp add: eq_iff_word_of word_of_word word_of_numeral) lemma of_word_0 [code_post]: \of_word 0 = 0\ by (simp add: eq_iff_word_of word_of_0 word_of_word) lemma of_word_1 [code_post]: \of_word 1 = 1\ by (simp add: eq_iff_word_of word_of_1 word_of_word) text \Use pretty numerals from integer for pretty printing\ lemma numeral_eq_integer [code_unfold]: \numeral n = of_integer (numeral n)\ by (simp add: eq_iff_word_of word_of_integer_eq word_of_numeral word_of_integer.rep_eq) lemma neg_numeral_eq_integer [code_unfold]: \- numeral n = of_integer (- numeral n)\ by (simp add: eq_iff_word_of word_of_integer_eq word_of_minus word_of_numeral word_of_integer.rep_eq) end locale word_type_copy_misc = word_type_copy_more opening constraintless bit_operations_syntax + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ fixes size :: nat and set_bits_aux :: \(nat \ bool) \ nat \ 'a \ 'a\ assumes size_eq_length: \size = LENGTH('b::len)\ and msb_iff_word_of [code]: \msb p \ msb (word_of p)\ and lsb_iff_word_of [code]: \lsb p \ lsb (word_of p)\ and size_eq_word_of: \Nat.size (p :: 'a) = Nat.size (word_of p)\ and word_of_generic_set_bit [code]: \word_of (Generic_set_bit.set_bit p n b) = Generic_set_bit.set_bit (word_of p) n b\ and word_of_set_bits: \word_of (set_bits P) = set_bits P\ and word_of_set_bits_aux: \word_of (set_bits_aux P n p) = Code_Target_Word_Base.set_bits_aux P n (word_of p)\ begin lemma size_eq [code]: \Nat.size p = size\ for p :: 'a by (simp add: size_eq_length size_eq_word_of word_size) lemma set_bits_aux_code [code]: \set_bits_aux f n w = (if n = 0 then w else let n' = n - 1 in set_bits_aux f n' (push_bit 1 w OR (if f n' then 1 else 0)))\ by (simp add: eq_iff_word_of word_of_set_bits_aux Let_def word_of_mult word_of_or word_of_0 word_of_1 set_bits_aux_rec [of f n]) lemma set_bits_code [code]: \set_bits P = set_bits_aux P size 0\ by (simp add: fun_eq_iff eq_iff_word_of word_of_set_bits word_of_set_bits_aux word_of_0 size_eq_length set_bits_conv_set_bits_aux) lemma of_class_lsb: \OFCLASS('a, lsb_class)\ by standard (simp add: fun_eq_iff lsb_iff_word_of even_iff_word_of lsb_odd) lemma of_class_set_bit: \OFCLASS('a, set_bit_class)\ by standard (simp add: eq_iff_word_of word_of_generic_set_bit bit_eq_word_of word_of_power word_of_0 bit_simps linorder_not_le) lemma of_class_bit_comprehension: \OFCLASS('a, bit_comprehension_class)\ by standard (simp add: eq_iff_word_of word_of_set_bits bit_eq_word_of set_bits_bit_eq) end end diff --git a/thys/Word_Lib/Ancient_Numeral.thy b/thys/Word_Lib/Ancient_Numeral.thy --- a/thys/Word_Lib/Ancient_Numeral.thy +++ b/thys/Word_Lib/Ancient_Numeral.thy @@ -1,237 +1,237 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Ancient_Numeral imports Main Reversed_Bit_Lists Legacy_Aliases begin definition Bit :: "int \ bool \ int" (infixl "BIT" 90) where "k BIT b = (if b then 1 else 0) + k + k" lemma Bit_B0: "k BIT False = k + k" by (simp add: Bit_def) lemma Bit_B1: "k BIT True = k + k + 1" by (simp add: Bit_def) lemma Bit_B0_2t: "k BIT False = 2 * k" by (rule trans, rule Bit_B0) simp lemma Bit_B1_2t: "k BIT True = 2 * k + 1" by (rule trans, rule Bit_B1) simp lemma uminus_Bit_eq: "- k BIT b = (- k - of_bool b) BIT b" by (cases b) (simp_all add: Bit_def) lemma power_BIT: "2 ^ Suc n - 1 = (2 ^ n - 1) BIT True" by (simp add: Bit_B1) lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" by (simp add: Bit_def) lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" by (simp add: Bit_def) lemma even_BIT [simp]: "even (x BIT b) \ \ b" by (simp add: Bit_def) lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" by simp lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \ u = v \ b = c" by (auto simp: Bit_def) arith+ lemma BIT_bin_simps [simp]: "numeral k BIT False = numeral (Num.Bit0 k)" "numeral k BIT True = numeral (Num.Bit1 k)" "(- numeral k) BIT False = - numeral (Num.Bit0 k)" "(- numeral k) BIT True = - numeral (Num.BitM k)" by (simp_all only: Bit_B0 Bit_B1 numeral.simps numeral_BitM) lemma BIT_special_simps [simp]: shows "0 BIT False = 0" and "0 BIT True = 1" and "1 BIT False = 2" and "1 BIT True = 3" and "(- 1) BIT False = - 2" and "(- 1) BIT True = - 1" by (simp_all add: Bit_def) lemma Bit_eq_0_iff: "w BIT b = 0 \ w = 0 \ \ b" by (auto simp: Bit_def) arith lemma Bit_eq_m1_iff: "w BIT b = -1 \ w = -1 \ b" by (auto simp: Bit_def) arith lemma expand_BIT: "numeral (Num.Bit0 w) = numeral w BIT False" "numeral (Num.Bit1 w) = numeral w BIT True" "- numeral (Num.Bit0 w) = (- numeral w) BIT False" "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" by (simp_all add: BitM_inc_eq add_One) lemma less_Bits: "v BIT b < w BIT c \ v < w \ v \ w \ \ b \ c" by (auto simp: Bit_def) lemma le_Bits: "v BIT b \ w BIT c \ v < w \ v \ w \ (\ b \ c)" by (auto simp: Bit_def) lemma pred_BIT_simps [simp]: "x BIT False - 1 = (x - 1) BIT True" "x BIT True - 1 = x BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma succ_BIT_simps [simp]: "x BIT False + 1 = x BIT True" "x BIT True + 1 = (x + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma add_BIT_simps [simp]: "x BIT False + y BIT False = (x + y) BIT False" "x BIT False + y BIT True = (x + y) BIT True" "x BIT True + y BIT False = (x + y) BIT True" "x BIT True + y BIT True = (x + y + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma mult_BIT_simps [simp]: "x BIT False * y = (x * y) BIT False" "x * y BIT False = (x * y) BIT False" "x BIT True * y = (x * y) BIT False + y" by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) lemma B_mod_2': "X = 2 \ (w BIT True) mod X = 1 \ (w BIT False) mod X = 0" by (simp add: Bit_B0 Bit_B1) lemma bin_ex_rl: "\w b. w BIT b = bin" by (metis bin_rl_simp) lemma bin_exhaust: "(\x b. bin = x BIT b \ Q) \ Q" by (metis bin_ex_rl) lemma bin_abs_lem: "bin = (w BIT b) \ bin \ -1 \ bin \ 0 \ nat \w\ < nat \bin\" apply clarsimp apply (unfold Bit_def) apply (cases b) apply (clarsimp, arith) apply (clarsimp, arith) done lemma bin_induct: assumes PPls: "P 0" and PMin: "P (- 1)" and PBit: "\bin bit. P bin \ P (bin BIT bit)" shows "P bin" apply (rule_tac P=P and a=bin and f1="nat \ abs" in wf_measure [THEN wf_induct]) apply (simp add: measure_def inv_image_def) apply (case_tac x rule: bin_exhaust) apply (frule bin_abs_lem) apply (auto simp add : PPls PMin PBit) done lemma Bit_div2: "(w BIT b) div 2 = w" by (fact bin_rest_BIT) lemma twice_conv_BIT: "2 * x = x BIT False" by (simp add: Bit_def) lemma BIT_lt0 [simp]: "x BIT b < 0 \ x < 0" by(cases b)(auto simp add: Bit_def) lemma BIT_ge0 [simp]: "x BIT b \ 0 \ x \ 0" by(cases b)(auto simp add: Bit_def) lemma bin_to_bl_aux_Bit_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (w BIT b) bl = bin_to_bl_aux (n - 1) w (b # bl)" by (cases n) auto lemma bl_to_bin_BIT: "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" by (simp add: bl_to_bin_append Bit_def) lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \ b" by simp lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" by (simp add: bit_Suc) lemma bin_nth_minus [simp]: "0 < n \ bin_nth (w BIT b) n = bin_nth w (n - 1)" by (cases n) (simp_all add: bit_Suc) lemma bin_sign_simps [simp]: "bin_sign (w BIT b) = bin_sign w" by (simp add: bin_sign_def Bit_def) lemma bin_nth_Bit: "bin_nth (w BIT b) n \ n = 0 \ b \ (\m. n = Suc m \ bin_nth w m)" by (cases n) auto lemmas sbintrunc_Suc_BIT [simp] = signed_take_bit_Suc [where a="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b lemmas sbintrunc_0_BIT_B0 [simp] = signed_take_bit_0 [where a="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] for w lemmas sbintrunc_0_BIT_B1 [simp] = signed_take_bit_0 [where a="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] for w lemma sbintrunc_Suc_minus_Is: \0 < n \ sbintrunc (n - 1) w = y \ sbintrunc n (w BIT b) = y BIT b\ by (cases n) (simp_all add: Bit_def signed_take_bit_Suc) lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" by (auto simp add: Bit_def concat_bit_Suc) -lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" - by (simp add: not_int_def Bit_def) - context includes bit_operations_syntax begin +lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" + by (simp add: not_int_def Bit_def) + lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" using and_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" using or_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" using xor_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) end lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" for bit proof - have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" by (simp add: mod_mult_mult1) also have "\ = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" by (simp add: ac_simps pos_zmod_mult_2) also have "\ = (2 * bin + 1) mod 2 ^ Suc n" by (simp only: mod_simps) finally show ?thesis by (auto simp add: Bit_def) qed lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" by(simp add: Bit_def) lemma int_lsb_BIT [simp]: fixes x :: int shows "lsb (x BIT b) \ b" by(simp add: lsb_int_def) lemma int_shiftr_BIT [simp]: fixes x :: int shows int_shiftr0: "drop_bit 0 x = x" and int_shiftr_Suc: "drop_bit (Suc n) (x BIT b) = drop_bit n x" by (simp_all add: drop_bit_Suc) lemma msb_BIT [simp]: "msb (x BIT b) = msb x" by(simp add: msb_int_def) end \ No newline at end of file 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,507 +1,507 @@ (* * Copyright Thomas Sewell, NICTA and Sascha Boehme, TU Muenchen * * SPDX-License-Identifier: BSD-2-Clause *) theory Bitwise imports "HOL-Library.Word" More_Arithmetic Reversed_Bit_Lists Bit_Shifts_Infix_Syntax 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)" +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: bit_simps word_size test_bit_of_bl msb_nth) done lemma nth_word_of_int: "bit (word_of_int x :: 'a::len word) n = (n < LENGTH('a) \ bit 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: "bit (scast (x :: 'a::len word) :: 'b::len word) n = (n < LENGTH('b) \ (if n < LENGTH('a) - 1 then bit x n else bit x (LENGTH('a) - 1)))" apply transfer apply (auto simp add: bit_signed_take_bit_iff min_def) done 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 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: push_bit_eq_mult) 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) apply (simp add: zip_take_triv) apply (simp only: mult.commute [of _ 2] mult.assoc bl_shiftl1) apply (simp flip: butlast_rev add: take_butlast) 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 concat_bit_Suc) 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) 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 (bit x) [0 ..< size x])" by (simp add: to_bl_eq_rev word_size rev_map) lemma rev_to_bl_upt: "rev (to_bl x) = map (bit 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\ lemma if_bool_simps: "If p True y = (p \ y) \ If p False y = (\ p \ y) \ If p y True = (p \ y) \ If p y False = (p \ y)" by auto 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) = dest_comb (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/Legacy_Aliases.thy b/thys/Word_Lib/Legacy_Aliases.thy --- a/thys/Word_Lib/Legacy_Aliases.thy +++ b/thys/Word_Lib/Legacy_Aliases.thy @@ -1,271 +1,271 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Legacy aliases\ theory Legacy_Aliases imports "HOL-Library.Word" begin context abstract_boolean_algebra begin lemma conj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" by (fact conj.assoc) lemma conj_commute: "x \<^bold>\ y = y \<^bold>\ x" by (fact conj.commute) lemmas conj_left_commute = conj.left_commute lemmas conj_ac = conj.assoc conj.commute conj.left_commute lemma conj_one_left: "\<^bold>1 \<^bold>\ x = x" by (fact conj.left_neutral) lemma conj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" by (fact conj.left_idem) lemma conj_absorb: "x \<^bold>\ x = x" by (fact conj.idem) lemma disj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" by (fact disj.assoc) lemma disj_commute: "x \<^bold>\ y = y \<^bold>\ x" by (fact disj.commute) lemmas disj_left_commute = disj.left_commute lemmas disj_ac = disj.assoc disj.commute disj.left_commute lemma disj_zero_left: "\<^bold>0 \<^bold>\ x = x" by (fact disj.left_neutral) lemma disj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" by (fact disj.left_idem) lemma disj_absorb: "x \<^bold>\ x = x" by (fact disj.idem) end context abstract_boolean_algebra_sym_diff begin lemmas xor_assoc = xor.assoc lemmas xor_commute = xor.commute lemmas xor_left_commute = xor.left_commute lemmas xor_ac = xor.assoc xor.commute xor.left_commute lemma xor_zero_right: "x \<^bold>\ \<^bold>0 = x" by (fact xor.comm_neutral) lemma xor_zero_left: "\<^bold>0 \<^bold>\ x = x" by (fact xor.left_neutral) end abbreviation (input) test_bit :: \'a::semiring_bits \ nat \ bool\ where \test_bit \ bit\ abbreviation (input) bin_nth :: \int \ nat \ bool\ where \bin_nth \ bit\ abbreviation (input) bin_last :: \int \ bool\ where \bin_last \ odd\ abbreviation (input) bin_rest :: \int \ int\ where \bin_rest w \ w div 2\ abbreviation (input) bintrunc :: \nat \ int \ int\ where \bintrunc \ take_bit\ abbreviation (input) sbintrunc :: \nat \ int \ int\ where \sbintrunc \ signed_take_bit\ abbreviation (input) bin_cat :: \int \ nat \ int \ int\ where \bin_cat k n l \ concat_bit n l k\ abbreviation (input) norm_sint :: \nat \ int \ int\ where \norm_sint n \ signed_take_bit (n - 1)\ abbreviation (input) max_word :: \'a::len word\ where \max_word \ - 1\ abbreviation (input) complement :: \'a::len word \ 'a word\ where \complement \ not\ lemma complement_mask: - "complement (2 ^ n - 1) = NOT (mask n)" + "complement (2 ^ n - 1) = not (mask n)" unfolding mask_eq_decr_exp by simp abbreviation (input) shiftl1 :: \'a::len word \ 'a word\ where \shiftl1 \ (*) 2\ abbreviation (input) shiftr1 :: \'a::len word \ 'a word\ where \shiftr1 w \ w div 2\ abbreviation (input) sshiftr1 :: \'a::len word \ 'a word\ where \sshiftr1 \ signed_drop_bit (Suc 0)\ context includes bit_operations_syntax begin abbreviation (input) bshiftr1 :: \bool \ 'a::len word \ 'a word\ where \bshiftr1 b w \ w div 2 OR push_bit (LENGTH('a) - Suc 0) (of_bool b) \ end lemma shiftr1_1: "shiftr1 (1::'a::len word) = 0" by (fact bits_1_div_2) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by (rule bit_word_eqI) (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) lemma shiftl1_eq: \shiftl1 w = word_of_int (2 * uint w)\ by (rule bit_word_eqI) (simp add: bit_simps) lemma bit_shiftl1_iff: \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_shiftr1_iff: \bit (shiftr1 w) n \ bit w (Suc n)\ by (simp add: bit_Suc) lemma shiftr1_eq: \shiftr1 w = word_of_int (uint w div 2)\ by (rule bit_word_eqI) (simp add: bit_simps flip: bit_Suc) lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" by (rule bit_word_eqI) (auto simp add: bit_simps Suc_diff_Suc simp flip: bit_Suc) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (fact mult_2) lemma shiftr1_bintr: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (take_bit LENGTH('a) (numeral w) div 2)" by (rule bit_word_eqI) (simp add: bit_simps bit_numeral_iff [where ?'a = int] flip: bit_Suc) lemma sshiftr1_sbintr: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" apply (cases \LENGTH('a)\) apply simp_all apply (rule bit_word_eqI) apply (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) done lemma shiftl1_wi: "shiftl1 (word_of_int w) = word_of_int (2 * w)" by transfer simp lemma shiftl1_numeral: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0: "shiftl1 0 = 0" by (fact mult_zero_right) lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" by (fact shiftl1_eq) lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" by simp lemma shiftr1_0: "shiftr1 0 = 0" by (fact bits_div_0) lemma sshiftr1_0: "sshiftr1 0 = 0" by (fact signed_drop_bit_of_0) lemma sshiftr1_n1: "sshiftr1 (- 1) = - 1" by (fact signed_drop_bit_of_minus_1) lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" by (rule bit_eqI) (simp add: bit_simps flip: bit_Suc) lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" by (rule bit_eqI) (auto simp add: bit_simps ac_simps min_def simp flip: bit_Suc elim: le_SucE) lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ by (auto simp add: bit_simps simp flip: bit_Suc) lemma nth_shiftl1: "bit (shiftl1 w) n \ n < size w \ n > 0 \ bit w (n - 1)" by (auto simp add: word_size bit_simps) lemma nth_shiftr1: "bit (shiftr1 w) n = bit w (Suc n)" by (simp add: bit_Suc) lemma nth_sshiftr1: "bit (sshiftr1 w) n = (if n = size w - 1 then bit w n else bit w (Suc n))" by (auto simp add: word_size bit_simps) lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" by (induction x) simp_all lemma le_shiftr1: \shiftr1 u \ shiftr1 v\ if \u \ v\ using that by (simp add: word_le_nat_alt unat_div div_le_mono) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" by (meson dual_order.antisym le_cases le_shiftr1) abbreviation (input) setBit :: \'a::len word \ nat \ 'a word\ where \setBit w n \ set_bit n w\ abbreviation (input) clearBit :: \'a::len word \ nat \ 'a word\ where \clearBit w n \ unset_bit n w\ lemma bit_setBit_iff: \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_clearBit_iff: \bit (clearBit w m) n \ m \ n \ bit w n\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] end diff --git a/thys/Word_Lib/Most_significant_bit.thy b/thys/Word_Lib/Most_significant_bit.thy --- a/thys/Word_Lib/Most_significant_bit.thy +++ b/thys/Word_Lib/Most_significant_bit.thy @@ -1,193 +1,193 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Dedicated operation for the most significant bit\ theory Most_significant_bit imports "HOL-Library.Word" More_Word More_Arithmetic begin class msb = fixes msb :: \'a \ bool\ instantiation int :: msb begin definition \msb x \ x < 0\ for x :: int instance .. end lemma msb_bin_rest [simp]: "msb (x div 2) = msb x" for x :: int by (simp add: msb_int_def) context includes bit_operations_syntax begin lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" by(simp add: msb_int_def) -end - lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" by(simp add: msb_int_def not_less) +end + lemma msb_shiftl [simp]: "msb (push_bit n (x :: int)) \ msb x" by(simp add: msb_int_def) lemma msb_shiftr [simp]: "msb (drop_bit r (x :: int)) \ msb x" by(simp add: msb_int_def) lemma msb_0 [simp]: "msb (0 :: int) = False" by(simp add: msb_int_def) lemma msb_1 [simp]: "msb (1 :: int) = False" by(simp add: msb_int_def) lemma msb_numeral [simp]: "msb (numeral n :: int) = False" "msb (- numeral n :: int) = True" by(simp_all add: msb_int_def) instantiation word :: (len) msb begin definition msb_word :: \'a word \ bool\ where msb_word_iff_bit: \msb w \ bit w (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ instance .. end lemma msb_word_eq: \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ by (simp add: msb_word_iff_bit) lemma word_msb_sint: "msb w \ sint w < 0" by (simp add: msb_word_eq bit_last_iff) lemma msb_word_iff_sless_0: \msb w \ w by (simp add: word_msb_sint word_sless_alt) lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bit x (LENGTH('a) - 1)" by (simp add: msb_word_iff_bit bit_simps) lemma word_msb_numeral [simp]: "msb (numeral w::'a::len word) = bit (numeral w :: int) (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule msb_word_of_int) lemma word_msb_neg_numeral [simp]: "msb (- numeral w::'a::len word) = bit (- numeral w :: int) (LENGTH('a) - 1)" unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" by (simp add: msb_word_iff_bit) lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" by (simp add: msb_word_iff_bit le_Suc_eq) lemma word_msb_nth: "msb w = bit (uint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: msb_word_iff_bit bit_simps) lemma msb_nth: "msb w = bit w (LENGTH('a) - 1)" for w :: "'a::len word" by (fact msb_word_eq) lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" by (simp add: msb_word_eq not_le) lemma msb_shift: "msb w \ drop_bit (LENGTH('a) - 1) w \ 0" for w :: "'a::len word" by (simp add: msb_word_eq bit_iff_odd_drop_bit drop_bit_eq_zero_iff_not_bit_last) lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_iff_bit word_size) apply transfer apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" apply (simp add: word_sle_eq 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_eq word_sle_msb_le) 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 add: bit_simps) done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_iff_bit) apply transfer apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" apply (cases \LENGTH('a)\) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) apply (metis le_less_Suc_eq test_bit_bin) done lemma msb_ucast_eq: "LENGTH('a) = LENGTH('b) \ msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" by (simp add: msb_word_eq bit_simps) lemma msb_big: \msb a \ 2 ^ (LENGTH('a) - Suc 0) \ a\ for a :: \'a::len word\ using bang_is_le [of a \LENGTH('a) - Suc 0\] apply (auto simp add: msb_nth word_le_not_less) apply (rule ccontr) apply (erule notE) apply (rule ccontr) apply (clarsimp simp: not_less) apply (subgoal_tac "a = take_bit (LENGTH('a) - Suc 0) a") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply auto apply (simp flip: take_bit_eq_mask) apply (rule sym) apply (simp add: take_bit_eq_self_iff_drop_bit_eq_0 drop_bit_eq_zero_iff_not_bit_last) done end diff --git a/thys/Word_Lib/Word_Lib_Sumo.thy b/thys/Word_Lib/Word_Lib_Sumo.thy --- a/thys/Word_Lib/Word_Lib_Sumo.thy +++ b/thys/Word_Lib/Word_Lib_Sumo.thy @@ -1,135 +1,135 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) section \Ancient comprehensive Word Library\ theory Word_Lib_Sumo imports "HOL-Library.Word" Aligned Bit_Comprehension Bit_Shifts_Infix_Syntax Bits_Int Bitwise_Signed Bitwise Enumeration_Word Generic_set_bit Hex_Words Least_significant_bit More_Arithmetic More_Divides More_Sublist Even_More_List More_Misc Strict_part_mono Legacy_Aliases Most_significant_bit Next_and_Prev Norm_Words Reversed_Bit_Lists Rsplit Signed_Words Syntax_Bundles Typedef_Morphisms Type_Syntax Word_EqI Word_Lemmas Word_8 Word_16 Word_32 Word_Syntax Signed_Division_Word More_Word_Operations Many_More begin unbundle bit_operations_syntax unbundle bit_projection_infix_syntax declare word_induct2[induct type] declare word_nat_cases[cases type] declare signed_take_bit_Suc [simp] (* these generate take_bit terms, which we often don't want for concrete lengths *) lemmas of_int_and_nat = unsigned_of_nat unsigned_of_int signed_of_int signed_of_nat bundle no_take_bit begin - declare of_int_and_nat[simp del] +declare of_int_and_nat[simp del] end lemmas bshiftr1_def = bshiftr1_eq lemmas is_down_def = is_down_eq lemmas is_up_def = is_up_eq lemmas mask_def = mask_eq lemmas scast_def = scast_eq lemmas shiftl1_def = shiftl1_eq lemmas shiftr1_def = shiftr1_eq lemmas sshiftr1_def = sshiftr1_eq lemmas sshiftr_def = sshiftr_eq_funpow_sshiftr1 lemmas to_bl_def = to_bl_eq lemmas ucast_def = ucast_eq lemmas unat_def = unat_eq_nat_uint lemmas word_cat_def = word_cat_eq lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl lemmas word_rotl_def = word_rotl_eq lemmas word_rotr_def = word_rotr_eq lemmas word_sle_def = word_sle_eq lemmas word_sless_def = word_sless_eq lemmas uint_0 = uint_nonnegative lemmas uint_lt = uint_bounded lemmas uint_mod_same = uint_idem lemmas of_nth_def = word_set_bits_def lemmas of_nat_word_eq_iff = word_of_nat_eq_iff lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff lemmas of_int_word_eq_iff = word_of_int_eq_iff lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff lemmas word_next_def = word_next_unfold lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp word32_and_max_simp lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemmas and_bang = word_and_nth lemmas sdiv_int_def = signed_divide_int_def lemmas smod_int_def = signed_modulo_int_def (* shortcut for some specific lengths *) lemma word_fixed_sint_1[simp]: "sint (1::8 word) = 1" "sint (1::16 word) = 1" "sint (1::32 word) = 1" "sint (1::64 word) = 1" by (auto simp: sint_word_ariths) declare of_nat_diff [simp] (* Haskellish names/syntax *) notation (input) bit ("testBit") lemmas cast_simps = cast_simps ucast_down_bl (* 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 (auto simp add: bit_simps not_le dest: bit_imp_le_length) end