Integrated latest changes of official AFP entry.

This commit is contained in:
Achim D. Brucker 2019-01-07 08:12:07 +00:00
parent 6cec568e45
commit 34ac13f630
6 changed files with 14 additions and 20 deletions

View File

@ -5882,7 +5882,6 @@ proof -
moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r = {}" moreover have "set |h3 \<turnstile> get_disconnected_nodes x|\<^sub>r \<inter> set |h3 \<turnstile> get_disconnected_nodes y|\<^sub>r = {}"
using calculation by(auto dest: distinct_concat_map_E(1)) using calculation by(auto dest: distinct_concat_map_E(1))
ultimately show "False" ultimately show "False"
apply(-)
apply(cases "x = document_ptr") apply(cases "x = document_ptr")
apply (metis (no_types) NodeMonad.ptr_kinds_ptr_kinds_M \<open>a_all_ptrs_in_heap h\<close> apply (metis (no_types) NodeMonad.ptr_kinds_ptr_kinds_M \<open>a_all_ptrs_in_heap h\<close>
\<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>cast new_character_data_ptr \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
@ -5912,7 +5911,6 @@ proof -
and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r" and 6: "x \<in> set |h' \<turnstile> get_disconnected_nodes xb|\<^sub>r"
show "False" show "False"
using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
apply -
apply(cases "xb = document_ptr") apply(cases "xb = document_ptr")
apply (metis (no_types, hide_lams) "3" "4" "6" apply (metis (no_types, hide_lams) "3" "4" "6"
\<open>\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close> \<open>\<And>p. p |\<in>| object_ptr_kinds h3 \<Longrightarrow> cast new_character_data_ptr \<notin> set |h3 \<turnstile> get_child_nodes p|\<^sub>r\<close>
@ -5927,11 +5925,11 @@ proof -
then have "a_owner_document_valid h'" then have "a_owner_document_valid h'"
using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close> using disc_nodes_h3 \<open>document_ptr |\<in>| document_ptr_kinds h\<close>
apply(simp add: a_owner_document_valid_def) apply(simp add: a_owner_document_valid_def)
apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 ) apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )
apply(simp add: object_ptr_kinds_eq_h2) apply(simp add: object_ptr_kinds_eq_h2)
apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 ) apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )
apply(simp add: document_ptr_kinds_eq_h2) apply(simp add: document_ptr_kinds_eq_h2)
apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 ) apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )
apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h ) apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )
apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h
disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1]

View File

@ -1,7 +1,7 @@
chapter AFP chapter AFP
session "Core_DOM-devel" (AFP) = "HOL-Library" + session "Core_DOM-devel" (AFP) = "HOL-Library" +
options [timeout = 4800, document = pdf, document_variants="document:outline=/proof,/ML"] options [timeout = 600]
theories theories
Core_DOM Core_DOM
Core_DOM_Tests Core_DOM_Tests

View File

@ -32,7 +32,7 @@ text\<open>In this theory, we introduce the basic infrastructure for our encodin
of classes.\<close> of classes.\<close>
theory BaseClass theory BaseClass
imports imports
"~~/src/HOL/Library/Finite_Map" "HOL-Library.Finite_Map"
"../pointers/Ref" "../pointers/Ref"
"../Core_DOM_Basic_Datatypes" "../Core_DOM_Basic_Datatypes"
begin begin

View File

@ -364,10 +364,8 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_type_typ
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result apply (metis (mono_tags, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.exhaust_sel)
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4)) by (metis (no_types, lifting) option.exhaust_sel )
by (metis (no_types, lifting) ElementMonad.a_get_M_def error_returns_result
get_heap_returns_result option.exhaust_sel option.simps(4))
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]: lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
@ -394,8 +392,7 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4)) get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4))
by (metis (no_types, lifting) ElementMonad.a_get_M_def error_returns_result get_heap_returns_result by (metis (no_types, lifting) option.exhaust_sel)
option.exhaust_sel option.simps(4))
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]: lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'" "h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -416,7 +413,7 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_w
ElementMonad.get_M_defs split: option.splits)[1] ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4)) apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4))
by (metis (no_types, lifting) ElementMonad.a_get_M_def error_returns_result get_heap_returns_result option.exhaust_sel option.simps(4)) by (metis (no_types, lifting) option.exhaust_sel)
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]: lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'" "h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -437,8 +434,7 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_
ElementMonad.get_M_defs split: option.splits)[1] ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) apply (metis (mono_tags, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by (metis (no_types, lifting) ElementMonad.a_get_M_def error_returns_result get_heap_returns_result by (metis (no_types, lifting) option.exhaust_sel)
option.exhaust_sel option.simps(4))
lemma new_character_data_type_wf_preserved [simp]: lemma new_character_data_type_wf_preserved [simp]:
@ -536,4 +532,4 @@ lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_
fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel
node_ptr_kinds_commutes) node_ptr_kinds_commutes)
end end

View File

@ -75,7 +75,7 @@ global_interpretation l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by
lemma node_ptr_kinds_M_reads: lemma node_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'" "reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads using object_ptr_kinds_M_reads
apply(simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_M_defs node_ptr_kinds_def apply(simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def
object_ptr_kinds_M_reads preserved_def) object_ptr_kinds_M_reads preserved_def)
by (metis (mono_tags, hide_lams) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def) by (metis (mono_tags, hide_lams) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)

View File

@ -591,7 +591,7 @@ lemma filter_M_ex1:
and "h \<turnstile> P x \<rightarrow>\<^sub>r True" and "h \<turnstile> P x \<rightarrow>\<^sub>r True"
shows "ys = [x]" shows "ys = [x]"
proof - proof -
have "\<exists>!x \<in> set xs. |h \<turnstile> P x|\<^sub>r" have *: "\<exists>!x \<in> set xs. |h \<turnstile> P x|\<^sub>r"
apply(insert assms(1) assms(3) assms(4)) apply(insert assms(1) assms(3) assms(4))
apply(drule filter_M_filter) apply(drule filter_M_filter)
apply(simp) apply(simp)
@ -601,7 +601,7 @@ proof -
apply(insert assms(1) assms(4)) apply(insert assms(1) assms(4))
apply(drule filter_M_filter) apply(drule filter_M_filter)
apply(auto)[1] apply(auto)[1]
by (metis \<open>\<exists>!x. x \<in> set xs \<and> |h \<turnstile> P x|\<^sub>r\<close> assms(2) assms(5) assms(6) distinct_filter by (metis * assms(2) assms(5) assms(6) distinct_filter
distinct_length_2_or_more filter_empty_conv filter_set list.exhaust distinct_length_2_or_more filter_empty_conv filter_set list.exhaust
list.set_intros(1) list.set_intros(2) member_filter select_result_I2) list.set_intros(1) list.set_intros(2) member_filter select_result_I2)
qed qed