isabelle2021-1: AsmRefine

Signed-off-by: Gerwin Klein <gerwin.klein@proofcraft.systems>
This commit is contained in:
Gerwin Klein 2022-01-17 09:21:13 +11:00 committed by Gerwin Klein
parent 68bb97ef66
commit a6c3ac2901
3 changed files with 11 additions and 14 deletions

View File

@ -25,7 +25,7 @@ lemma fold_all_htd_updates':
ptr_arr_retyps n p = all_htd_updates TYPE('a) 4 (ptr_val p) (of_nat n)"
"\<lbrakk> n < 2 ^ word_bits \<rbrakk> \<Longrightarrow>
ptr_arr_retyps (2 ^ n) p = all_htd_updates TYPE('a) 5 (ptr_val p) (of_nat n)"
by (simp_all add: all_htd_updates_def fun_eq_iff word_bits_conv take_bit_nat_eq_self)
by (simp_all add: all_htd_updates_def fun_eq_iff word_bits_conv take_bit_nat_eq_self unat_of_nat)
lemma upcast_unat_less_2p_length:
"is_up UCAST('a :: len \<rightarrow> 'b :: len) \<Longrightarrow> unat (x :: 'a word) < 2 ^ LENGTH('b)"
@ -89,13 +89,15 @@ proof -
apply (cut_tac x=a in sint_range')
apply (clarsimp simp add: abs_if word_size)
done
note sint_of_int_eq[simp] signed_take_bit_int_eq_self[simp]
show ?thesis using mag len
apply (cases "b = 1")
apply (case_tac "size a", simp_all)[1]
apply (case_tac nat, simp_all add: sint_word_ariths word_size)[1]
apply (simp add: sdiv_int_def sdiv_word_def del: of_int_mult)
apply (simp add: sbintrunc_eq_in_range range_sbintrunc sgn_if)
apply (safe, simp_all add: word_size sint_int_min)
apply (safe, simp_all add: word_size sint_int_min sint_int_max_plus_1;
simp add: sint_word_ariths)
done
qed

View File

@ -1543,7 +1543,7 @@ lemma heap_update_list_If1:
apply (clarsimp simp: nth_append)
apply (rule refl)
apply (simp add: nth_append split del: if_split cong: if_cong)
apply (auto simp: addr_card linorder_not_less less_Suc_eq take_bit_nat_eq_self
apply (auto simp: addr_card linorder_not_less less_Suc_eq take_bit_nat_eq_self unat_of_nat
dest: word_unat.Rep_inverse')
done
@ -1932,7 +1932,7 @@ and mk_simpl_to_graph_thm funs hints cache nm ctxt tm = let
THEN_ALL_NEW (TRY o simpl_to_graph_cache_tac funs hints cache nm ctxt)
THEN_ALL_NEW (TRY o eq_impl_assume_tac ctxt)) 1
|> Seq.hd
|> Drule.generalize ([], ["n", "trS"])
|> Drule.generalize (Names.empty, Names.make_set ["n", "trS"])
|> SOME
end handle TERM (s, _) => (tracing ("mk_simpl_to_graph_thm: " ^ s); NONE)
| Empty => (tracing "mk_simpl_to_graph_thm: raised Empty on:";
@ -1988,7 +1988,7 @@ fun simpl_to_graph_While_tac hints nm ctxt =
val rl_inst = infer_instantiate ctxt [(("G",0), Thm.cterm_of ctxt gd)]
@{thm simpl_to_graph_While_inst}
in
resolve_tac ctxt [Thm.trivial ct |> Drule.generalize ([], ["n", "trS"])] i
resolve_tac ctxt [Thm.trivial ct |> Drule.generalize (Names.empty, Names.make_set ["n", "trS"])] i
THEN resolve_tac ctxt [rl_inst] i
THEN resolve_tac ctxt @{thms refl} i
THEN inst_graph_tac ctxt i
@ -2131,7 +2131,7 @@ fun init_graph_refines_proof funs nm ctxt = let
val thin_While_assums_rule =
@{thm thin_rl[where V="simpl_to_graph SG GG f nn (add_cont (com.While C c) con) n tS P I e e2"]}
|> Drule.generalize ([], ["SG", "GG", "f", "nn", "C", "c", "con", "n", "tS", "P", "I", "e", "e2"])
|> Drule.generalize (Names.empty, Names.make_set ["SG", "GG", "f", "nn", "C", "c", "con", "n", "tS", "P", "I", "e", "e2"])
fun eq_impl_unassume_tac t = let
val hyps = t |> Thm.chyps_of

View File

@ -134,12 +134,7 @@ lemma drop_sign_isomorphism_bitwise:
"drop_sign (scast z) = scast z"
"ucast x = ucast (drop_sign x)"
"scast x = scast (drop_sign x)"
by (rule word_eqI
| simp add: word_size drop_sign_def nth_ucast nth_shiftl
nth_shiftr nth_sshiftr word_ops_nth_size
nth_scast
| safe
| simp add: test_bit_bin)+
by (all \<open>(word_eqI_solve simp: drop_sign_def test_bit_bin)\<close>)
lemma drop_sign_of_nat:
"drop_sign (of_nat n) = of_nat n"
@ -272,7 +267,7 @@ lemma fold_of_nat_eq_Ifs[simplified word_bits_conv]:
\<Longrightarrow> foldr (\<lambda>n v. if x = of_nat n then f n else v) [0 ..< m] (f m)
= f (unat (machine_word_truncate_nat m x))"
apply (rule fold_of_nat_eq_Ifs_proof)
apply (simp_all add: machine_word_truncate_nat_def word_bits_def take_bit_nat_eq_self)
apply (simp_all add: machine_word_truncate_nat_def word_bits_def take_bit_nat_eq_self unat_of_nat)
done
lemma less_is_non_zero_p1':
@ -761,7 +756,7 @@ fun dest_ptr_add_assertion ctxt = SUBGOAL (fn (t, i) =>
fun tactic_check' (ss, t) = (ss, tactic_check (hd ss) t)
fun graph_refine_proof_tacs csenv ctxt = let
val ctxt = ctxt delsimps @{thms shiftl_numeral}
val ctxt = ctxt delsimps @{thms shiftl_numeral_numeral shiftl1_is_mult}
|> Splitter.del_split @{thm if_split}
|> Simplifier.del_cong @{thm if_weak_cong}