Theory CProof
theory CProof
imports
"umm_heap/SepFrame"
"Simpl.Vcg"
"umm_heap/StructSupport"
"umm_heap/ArrayAssertion"
"AutoCorres_Utils"
"Match_Cterm"
"ML_Infer_Instantiate"
begin
ML_file "General.ML"
ML_file "SourcePos.ML"
ML_file "SourceFile.ML"
ML_file "Region.ML"
ML_file "Binaryset.ML"
ML_file "Feedback.ML"
ML_file "basics.ML"
ML_file "MString.ML"
ML_file "TargetNumbers-sig.ML"
ML_file "./umm_heap/ARM/TargetNumbers_ARM.ML"
ML_file "./umm_heap/ARM64/TargetNumbers_ARM64.ML"
ML_file "./umm_heap/ARM_HYP/TargetNumbers_ARM_HYP.ML"
ML_file "./umm_heap/RISCV64/TargetNumbers_RISCV64.ML"
ML_file "./umm_heap/X64/TargetNumbers_X64.ML"
ML_file "./umm_heap/TargetNumbers.ML"
ML_file "RegionExtras.ML"
ML_file "Absyn-CType.ML"
ML_file "Absyn-Expr.ML"
ML_file "Absyn-StmtDecl.ML"
ML_file "Absyn.ML"
ML_file "Absyn-Serial.ML"
ML_file "../lib/ml-helpers/StringExtras.ML"
ML_file "name_generation.ML"
setup ‹
Hoare.add_foldcongsimps [@{thm "update_update"}, @{thm "o_def"}]
›
syntax
"_quote" :: "'b => ('a => 'b)" ("([.[_].])" [0] 1000)
syntax
"_heap" :: "'b ⇒ ('a ⇒ 'b)"
"_heap_state" :: "'a" ("ζ")
"_heap_stateOld" :: "('a ⇒ 'b) ⇒ 'b" ("⇗_⇖ζ" [100] 100)
"_derefCur" :: "('a ⇒ 'b) ⇒ 'b" ("⋆_" [100] 100)
"_derefOld" :: "'a ⇒ ('a ⇒ 'b) ⇒ 'b" ("⇗_⇖⋆_" [100,100] 100)
translations
"{|b|}" => "CONST Collect (_quote (_heap b))"
definition sep_app :: "(heap_state ⇒ bool) ⇒ heap_state ⇒ bool" where
"sep_app P s ≡ P s"
definition hrs_id :: "heap_raw_state ⇒ heap_raw_state" where
"hrs_id ≡ id"
declare hrs_id_def [simp add]
parse_translation ‹
let
fun ac x = Syntax.const "_antiquoteCur" $ Syntax.const x
fun aco x y = Syntax.const y $ (Syntax.const "globals" $ x)
fun hd a = a NameGeneration.global_heap_var
fun d a = Syntax.const "hrs_htd" $ hd a
fun repl (Abs (s,T,t)) = Abs (s,T,repl t)
| repl (Const ("_h_t_valid",_)$x) = Syntax.const "h_t_valid" $ d ac $ Syntax.const "c_guard" $ x
| repl (Const ("_derefCur",_)$x) = Syntax.const "the" $
(Syntax.const "lift_t" $ hd ac $ x)
| repl (Const ("_derefOld",_)$x$y) = Syntax.const "the" $
(Syntax.const "lift_t" $ hd (aco x) $ y)
| repl (Const ("_heap_state",_)) = Syntax.const "hrs_id" $ hd ac
| repl (Const ("_heap_stateOld",_)$x) = Syntax.const "hrs_id" $ hd (aco x)
| repl (x$y) = repl x $ repl y
| repl x = x
fun heap_assert_tr [b] = repl b
| heap_assert_tr ts = raise TERM ("heap_assert_tr", ts);
in [("_heap",K heap_assert_tr)] end
›
parse_translation ‹
let
fun ac x = Syntax.const "_antiquoteCur" $ Syntax.const x
fun aco x y = Syntax.const y $ (Syntax.const "globals" $ x)
fun hd a = Syntax.const "lift_state" $ (a NameGeneration.global_heap_var)
fun st2 (Abs (s,T,t)) n = Abs (s,T,st2 t (n+1))
| st2 (Bound k) n = Bound (if k < n then k else k + 1)
| st2 (x$y) n = st2 x n $ st2 y n
| st2 x _ = x
fun st1 (Abs (s,T,t)) n = Abs (s,T,st1 t (n+1))
| st1 (Bound k) n = Bound (if k < n then k else k + 1)
| st1 (Const ("sep_empty",_)) n = Syntax.const "sep_empty" $ Bound n
| st1 (Const ("sep_map",_)$x$y) n = Syntax.const "sep_map" $
(st2 x n) $ (st2 y n) $ Bound n
| st1 (Const ("sep_map'",_)$x$y) n = Syntax.const "sep_map'" $
(st2 x n) $ (st2 y n) $ Bound n
| st1 (Const ("sep_conj",_)$x$y) n = Syntax.const "sep_conj" $
(nst2 x n) $ (nst2 y n) $ Bound n
| st1 (Const ("sep_impl",_)$x$y) n = Syntax.const "sep_impl" $
(nst2 x n) $ (nst2 y n) $ Bound n
| st1 (x$y) n = st1 x n $ st1 y n
| st1 x _ = x
and new_heap t = Abs ("s",dummyT,st1 t 0)
and nst2 x n = new_heap (st2 x n);
fun sep_tr [t] = Syntax.const "sep_app" $ t $ hd ac
| sep_tr ts = raise TERM ("sep_tr", ts);
in [("_sep_assert",K sep_tr)] end
›
lemma c_null_guard:
"c_null_guard (p::'a::mem_type ptr) ⟹ p ≠ NULL"
by (fastforce simp: c_null_guard_def intro: intvl_self)
lemma (in mem_type) c_guard_no_wrap:
fixes p :: "'a ptr"
assumes cgrd: "c_guard p"
shows "ptr_val p ≤ ptr_val p + of_nat (size_of TYPE('a) - 1)"
using cgrd unfolding c_guard_def c_null_guard_def
apply -
apply (erule conjE)
apply (erule contrapos_np)
apply (simp add: intvl_def)
apply (drule word_wrap_of_natD)
apply (erule exE)
apply (rule exI)
apply (simp add: nat_le_Suc_less size_of_def wf_size_desc_gt(1))
done
lemma word_le_unat_bound:
fixes a::"'a ::len word"
assumes "a ≤ a + b"
shows "unat a + unat b < 2 ^ LENGTH('a)"
using assms no_olen_add_nat by blast
lemma (in mem_type) c_guard_no_wrap':
fixes p :: "'a ptr"
assumes cgrd: "c_guard p"
shows "unat (ptr_val p) + size_of TYPE('a) ≤ addr_card"
proof -
have szb: "size_of TYPE('a) < addr_card"
by (simp add: local.max_size)
have not_null: "0 < size_of TYPE('a)"
by (simp add: size_of_def wf_size_desc_gt(1))
have sz_le: "size_of TYPE('a) - Suc 0 < 2 ^ LENGTH(addr_bitsize)"
using len_of_addr_card less_imp_diff_less szb by simp
with szb
have eq: "unat (word_of_nat (size_of TYPE('a) - 1)::addr_bitsize word) = size_of (TYPE('a)) - 1"
apply (subst unat_of_nat_eq)
apply (simp_all )
done
from word_le_unat_bound [OF c_guard_no_wrap [OF cgrd], simplified eq]
show ?thesis
by (simp add: addr_card)
qed
definition
c_fnptr_guard_def: "c_fnptr_guard (fnptr::unit ptr) ≡ ptr_val fnptr ≠ 0"
lemma c_fnptr_guard_NULL [simp]: "c_fnptr_guard NULL = False"
by (simp add: c_fnptr_guard_def)
lemma c_guardD:
"c_guard (p::'a::mem_type ptr) ⟹ ptr_aligned p ∧ p ≠ NULL"
by (clarsimp simp: c_guard_def c_null_guard)
lemma c_guard_ptr_aligned:
"c_guard p ⟹ ptr_aligned p"
by (simp add: c_guard_def)
lemma c_guard_NULL:
"c_guard (p::'a::mem_type ptr) ⟹ p ≠ NULL"
by (simp add: c_guardD)
lemma c_guard_NULL_simp [simp]:
"¬ c_guard (NULL::'a::mem_type ptr)"
by (force dest: c_guard_NULL)
lemma c_guard_mono:
"guard_mono (c_guard::'a::mem_type ptr_guard) (c_guard::'b::mem_type ptr_guard)"
apply(clarsimp simp: guard_mono_def c_guard_def)
subgoal premises prems for n f p
proof -
have "guard_mono (ptr_aligned::'a::mem_type ptr_guard) (ptr_aligned::'b::mem_type ptr_guard)"
using prems by - (rule ptr_aligned_mono)
then show ?thesis
using prems
apply -
apply(clarsimp simp: guard_mono_def ptr_aligned_def c_null_guard_def typ_uinfo_t_def)
apply(frule field_lookup_export_uinfo_Some_rev)
apply clarsimp
apply(drule field_tag_sub [where p=p])
apply(clarsimp simp: field_lvalue_def field_offset_def field_offset_untyped_def typ_uinfo_t_def)
apply(metis (mono_tags, opaque_lifting) export_size_of subsetD typ_uinfo_t_def)
done
qed
done
lemma c_guard_NULL_fl:
"⟦ c_guard (p::'a::mem_type ptr); field_lookup (typ_info_t TYPE('a)) f 0 = Some (t,n);
export_uinfo t = typ_uinfo_t TYPE('b::mem_type) ⟧
⟹ 0 < &(p→f)"
using c_guard_mono
apply(clarsimp simp: guard_mono_def)
apply((erule allE)+, erule impE)
apply(fastforce dest: field_lookup_export_uinfo_Some simp: typ_uinfo_t_def)
apply(drule field_lookup_export_uinfo_Some)
apply(simp add: field_lvalue_def field_offset_def field_offset_untyped_def typ_uinfo_t_def)
apply(clarsimp simp: c_guard_def)
apply(drule c_null_guard)+
apply(clarsimp simp: word_neq_0_conv)
done
lemma c_guard_ptr_aligned_fl:
"⟦ c_guard (p::'a::mem_type ptr); field_lookup (typ_info_t TYPE('a)) f 0 = Some (t,n);
export_uinfo t = typ_uinfo_t TYPE('b::mem_type) ⟧ ⟹
ptr_aligned ((Ptr &(p→f))::'b ptr)"
using c_guard_mono
apply(clarsimp simp: guard_mono_def)
apply((erule allE)+, erule impE)
apply(fastforce dest: field_lookup_export_uinfo_Some simp: typ_uinfo_t_def)
apply(drule field_lookup_export_uinfo_Some)
apply(simp add: c_guard_def field_lvalue_def field_offset_def field_offset_untyped_def typ_uinfo_t_def)
done
syntax
"_sep_map" :: "'a::c_type ptr ⇒ 'a ⇒ heap_assert" ("_ ↦ _" [56,51] 56)
"_sep_map_any" :: "'a::c_type ptr ⇒ heap_assert" ("_ ↦ -" [56] 56)
"_sep_map'" :: "'a::c_type ptr ⇒ 'a ⇒ heap_assert" ("_ ↪ _" [56,51] 56)
"_sep_map'_any" :: "'a::c_type ptr ⇒ heap_assert" ("_ ↪ -" [56] 56)
"_tagd" :: "'a::c_type ptr ⇒ heap_assert" ("⊢⇩s _" [99] 100)
translations
"p ↦ v" == "p ↦⇧i⇩(CONST c_guard) v"
"p ↦ -" == "p ↦⇧i⇩(CONST c_guard) -"
"p ↪ v" == "p ↪⇧i⇩(CONST c_guard) v"
"p ↪ -" == "p ↪⇧i⇩(CONST c_guard) -"
"⊢⇩s p" == "CONST c_guard ⊢⇩s⇧i p"
term "x ↦ y"
term "(x ↦ y ∧⇧* y ↦ z) s"
term "(x ↦ y ∧⇧* y ↦ z) ∧⇧* x ↦ y"
term "⊢⇩s p"
lemma sep_map_NULL [simp]:
"NULL ↦ (v::'a::mem_type) = sep_false"
by (force simp: sep_map_def sep_map_inv_def c_guard_def
dest: lift_typ_heap_g sep_conjD c_null_guard)
lemma sep_map'_NULL_simp [simp]:
"NULL ↪ (v::'a::mem_type) = sep_false"
apply(simp add: sep_map'_def sep_map'_inv_def sep_conj_ac)
apply(subst sep_conj_com, subst sep_map_inv_def [symmetric])
apply simp
done
lemma sep_map'_ptr_aligned:
"(p ↪ v) s ⟹ ptr_aligned p"
by (rule c_guard_ptr_aligned) (erule sep_map'_g)
lemma sep_map'_NULL:
"(p ↪ (v::'a::mem_type)) s ⟹ p ≠ NULL"
by (rule c_guard_NULL) (erule sep_map'_g)
lemma tagd_sep_false [simp]:
"⊢⇩s (NULL::'a::mem_type ptr) = sep_false"
by (auto simp: tagd_inv_def tagd_def dest!: sep_conjD s_valid_g)
syntax (output)
"_Deref" :: "'b ⇒ 'b" ("*_" [1000] 1000)
"_AssignH" :: "'b => 'b => ('a,'p,'f) com" ("(2*_ :==/ _)" [30, 30] 23)
print_translation ‹
let
fun deref (Const ("_antiquoteCur",_)$Const (h,_)) p =
if h=NameGeneration.global_heap then Syntax.const "_Deref" $ p else
raise Match
| deref h p = raise Match
fun lift_tr [h,p] = deref h p
| lift_tr ts = raise Match
fun deref_update (Const ("heap_update",_)$l$r$(Const ("_antiquoteCur",_)$
Const (h,_))) =
if h=NameGeneration.global_heap then Syntax.const "_AssignH" $ l $ r
else raise Match
| deref_update x = raise Match
fun deref_assign (Const ("_antiquoteCur",_)$Const (h,_)) r =
if h=NameGeneration.global_heap then deref_update r else
raise Match
| deref_assign l r = raise Match
fun assign_tr [l,r] = deref_assign l r
| assign_tr ts = raise Match
in [("CTypesDefs.lift",K lift_tr),("_Assign",K assign_tr)] end
›
print_translation ‹
let
fun sep_app_tr [l,r] = Syntax.const "_sep_assert" $ l
| sep_app_tr ts = raise Match
in [("sep_app",K sep_app_tr)] end
›
syntax "_h_t_valid" :: "'a::c_type ptr ⇒ bool" ("⊨⇩t _" [99] 100)
abbreviation "lift_t_c" :: "heap_mem × heap_typ_desc ⇒ 'a::c_type typ_heap" where
"lift_t_c s == lift_t c_guard s"
syntax "_h_t_valid" :: "heap_typ_desc ⇒ 'a::c_type ptr ⇒ bool" ("_ ⊨⇩t _" [99,99] 100)
translations
"d ⊨⇩t p" == "d,CONST c_guard ⊨⇩t p"
lemma h_t_valid_c_guard:
"d ⊨⇩t p ⟹ c_guard p"
by (clarsimp simp: h_t_valid_def)
lemma h_t_valid_NULL [simp]:
"¬ d ⊨⇩t (NULL::'a::mem_type ptr)"
by (clarsimp simp: h_t_valid_def dest!: c_guard_NULL)
lemma h_t_valid_ptr_aligned:
"d ⊨⇩t p ⟹ ptr_aligned p"
by (auto simp: h_t_valid_def dest: c_guard_ptr_aligned)
lemma lift_t_NULL:
"lift_t_c s (NULL::'a::mem_type ptr) = None"
by (cases s) (auto simp: lift_t_if)
lemma lift_t_ptr_aligned [simp]:
"lift_t_c s p = Some v ⟹ ptr_aligned p"
by (auto dest: lift_t_g c_guard_ptr_aligned)
lemmas c_typ_rewrs = typ_rewrs h_t_valid_ptr_aligned lift_t_NULL