Fixed long lines and simp lemmas without names.

This commit is contained in:
Michael Herzberg 2020-07-22 22:11:21 +01:00
parent 99a6566ed0
commit f955f2fa56
35 changed files with 6466 additions and 5933 deletions

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
*******************************************************************************\***)
@ -31,7 +31,7 @@ section\<open>Basic Data Types\<close>
text\<open>
\label{sec:Core_DOM_Basic_Datatypes}
This theory formalizes the primitive data types used by the DOM standard~\cite{dom-specification}.
\<close>
\<close>
theory Core_DOM_Basic_Datatypes
imports
Main
@ -39,16 +39,16 @@ begin
type_synonym USVString = string
text\<open>
In the official standard, the type @{type "USVString"} corresponds to the set of all possible
In the official standard, the type @{type "USVString"} corresponds to the set of all possible
sequences of Unicode scalar values. As we are not interested in analyzing the specifics of Unicode
strings, we just model @{type "USVString"} using the standard type @{type "string"} of Isabelle/HOL.
\<close>
\<close>
type_synonym DOMString = string
text\<open>
In the official standard, the type @{type "DOMString"} corresponds to the set of all possible
sequences of code units, commonly interpreted as UTF-16 encoded strings. Again, as we are not
interested in analyzing the specifics of Unicode strings, we just model @{type "DOMString"} using
In the official standard, the type @{type "DOMString"} corresponds to the set of all possible
sequences of code units, commonly interpreted as UTF-16 encoded strings. Again, as we are not
interested in analyzing the specifics of Unicode strings, we just model @{type "DOMString"} using
the standard type @{type "string"} of Isabelle/HOL.
\<close>

File diff suppressed because it is too large Load Diff

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)

View File

@ -23,18 +23,18 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>The Class Infrastructure\<close>
text\<open>In this theory, we introduce the basic infrastructure for our encoding
text\<open>In this theory, we introduce the basic infrastructure for our encoding
of classes.\<close>
theory BaseClass
imports
"HOL-Library.Finite_Map"
"../pointers/Ref"
"../Core_DOM_Basic_Datatypes"
"../Core_DOM_Basic_Datatypes"
begin
named_theorems instances
@ -43,26 +43,26 @@ consts get :: 'a
consts put :: 'a
consts delete :: 'a
text \<open>Overall, the definition of the class types follows closely the one of the pointer
types. Instead of datatypes, we use records for our classes. This allows us to, first,
text \<open>Overall, the definition of the class types follows closely the one of the pointer
types. Instead of datatypes, we use records for our classes. This allows us to, first,
make use of record inheritance, which is, in addition to the type synonyms of
previous class types, the second place where the inheritance relationship of
previous class types, the second place where the inheritance relationship of
our types manifest. Second, we get a convenient notation to define classes, in
addition to automatically generated getter and setter functions.\<close>
text \<open>Along with our class types, we also develop our heap type, which is a finite
map at its core. It is important to note that while the map stores a mapping
from @{term "object_ptr"} to @{term "Object"}, we restrict the type variables
of the record extension slot of @{term "Object"} in such a way that allows
down-casting, but requires a bit of taking-apart and re-assembling of our records
text \<open>Along with our class types, we also develop our heap type, which is a finite
map at its core. It is important to note that while the map stores a mapping
from @{term "object_ptr"} to @{term "Object"}, we restrict the type variables
of the record extension slot of @{term "Object"} in such a way that allows
down-casting, but requires a bit of taking-apart and re-assembling of our records
before they are stored in the heap.\<close>
text \<open>Throughout the theory files, we will use underscore case to reference pointer
text \<open>Throughout the theory files, we will use underscore case to reference pointer
types, and camel case for class types.\<close>
text \<open>Every class type contains at least one attribute; nothing. This is used for
two purposes: first, the record package does not allow records without any
attributes. Second, we will use the getter of nothing later to check whether a
text \<open>Every class type contains at least one attribute; nothing. This is used for
two purposes: first, the record package does not allow records without any
attributes. Second, we will use the getter of nothing later to check whether a
class of the correct type could be retrieved, for which we will be able to use
our infrastructure regarding the behaviour of getters across different heaps.\<close>

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -36,45 +36,45 @@ begin
subsubsection\<open>CharacterData\<close>
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, defined
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, defined
\autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
record RCharacterData = RNode +
nothing :: unit
val :: DOMString
register_default_tvars "'CharacterData RCharacterData_ext"
register_default_tvars "'CharacterData RCharacterData_ext"
type_synonym 'CharacterData CharacterData = "'CharacterData option RCharacterData_scheme"
register_default_tvars "'CharacterData CharacterData"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
register_default_tvars "'CharacterData CharacterData"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'CharacterData option RCharacterData_ext + 'Node, 'Element) Node"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node,
'Element, 'CharacterData) Node"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'CharacterData option RCharacterData_ext + 'Node,
'Element) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object,
'Node, 'Element, 'CharacterData) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'CharacterData option RCharacterData_ext + 'Node, 'Element) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition character_data_ptr_kinds :: "(_) heap \<Rightarrow> (_) character_data_ptr fset"
where
"character_data_ptr_kinds heap = the |`| (cast |`| (ffilter is_character_data_ptr_kind
where
"character_data_ptr_kinds heap = the |`| (cast |`| (ffilter is_character_data_ptr_kind
(node_ptr_kinds heap)))"
lemma character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (Heap (fmupd (cast character_data_ptr) character_data (the_heap h)))
"character_data_ptr_kinds (Heap (fmupd (cast character_data_ptr) character_data (the_heap h)))
= {|character_data_ptr|} |\<union>| character_data_ptr_kinds h"
apply(auto simp add: character_data_ptr_kinds_def)[1]
by force
@ -94,7 +94,7 @@ adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^su
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^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 :: "(_) Object \<Rightarrow> (_) CharacterData option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^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 obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj of Some node \<Rightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 node
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^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 obj \<equiv> (case cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj of Some node \<Rightarrow> cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 node
| None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^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
@ -123,15 +123,15 @@ abbreviation is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^su
adhoc_overloading is_character_data_kind is_character_data_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma character_data_ptr_kinds_commutes [simp]:
"cast character_data_ptr |\<in>| node_ptr_kinds h
"cast character_data_ptr |\<in>| node_ptr_kinds h
\<longleftrightarrow> character_data_ptr |\<in>| character_data_ptr_kinds h"
apply(auto simp add: character_data_ptr_kinds_def)[1]
by (metis character_data_ptr_casts_commute2 comp_eq_dest_lhs ffmember_filter fimage_eqI
by (metis character_data_ptr_casts_commute2 comp_eq_dest_lhs ffmember_filter fimage_eqI
is_character_data_ptr_kind_none
option.distinct(1) option.sel)
definition 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 :: "(_) character_data_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) CharacterData option"
where
where
"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 character_data_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr) h) cast"
adhoc_overloading get 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
@ -160,11 +160,12 @@ sublocale l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas b
lemma 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_type_wf:
assumes "type_wf h"
shows "character_data_ptr |\<in>| character_data_ptr_kinds h
shows "character_data_ptr |\<in>| character_data_ptr_kinds h
\<longleftrightarrow> 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 character_data_ptr h \<noteq> None"
using l_type_wf\<^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_axioms assms
apply(simp add: type_wf_defs 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 l_type_wf\<^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)
by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember.rep_eq local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3))
by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember.rep_eq
local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3))
end
global_interpretation l_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_lemmas type_wf
@ -172,7 +173,7 @@ global_interpretation l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^su
definition put\<^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 :: "(_) character_data_ptr \<Rightarrow> (_) CharacterData \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
"put\<^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 character_data_ptr character_data = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr)
"put\<^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 character_data_ptr character_data = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast character_data_ptr)
(cast character_data)"
adhoc_overloading put put\<^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
@ -196,16 +197,16 @@ lemma cast\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 node = None \<longleftrightarrow> \<not> (\<exists>character_data. cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_some [simp]:
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 node = Some character_data \<longleftrightarrow> cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_inv [simp]:
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_inv [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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 (cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data) = Some character_data"
by simp
@ -214,19 +215,19 @@ lemma cast_element_not_character_data [simp]:
"(cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e character_data \<noteq> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element)"
by(auto simp add: cast\<^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\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RNode.extend_def)
lemma get_CharacterData_simp1 [simp]:
"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 character_data_ptr (put\<^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 character_data_ptr character_data h)
lemma get_CharacterData_simp1 [simp]:
"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 character_data_ptr (put\<^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 character_data_ptr character_data h)
= Some character_data"
by(auto simp add: 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 put\<^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)
lemma get_CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr' \<Longrightarrow> 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 character_data_ptr
lemma get_CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr' \<Longrightarrow> 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 character_data_ptr
(put\<^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 character_data_ptr' character_data h) = 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 character_data_ptr h"
by(auto simp add: 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 put\<^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)
lemma get_CharacterData_simp3 [simp]:
lemma get_CharacterData_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^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 character_data_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^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)
lemma get_CharacterData_simp4 [simp]:
lemma get_CharacterData_simp4 [simp]:
"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 element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t character_data_ptr f h) = 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 element_ptr h"
by(auto simp add: 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 put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -244,7 +245,7 @@ abbreviation "create_character_data_obj val_arg
definition new\<^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 :: "(_) heap \<Rightarrow> ((_) character_data_ptr \<times> (_) heap)"
where
"new\<^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 h =
(let new_character_data_ptr = character_data_ptr.Ref (Suc (fMax (character_data_ptr.the_ref
(let new_character_data_ptr = character_data_ptr.Ref (Suc (fMax (character_data_ptr.the_ref
|`| (character_data_ptrs h)))) in
(new_character_data_ptr, put new_character_data_ptr (create_character_data_obj '''') h))"
@ -255,17 +256,19 @@ lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>
unfolding new\<^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 Let_def
using put\<^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_ptr_in_heap by blast
lemma new_character_data_ptr_new:
"character_data_ptr.Ref (Suc (fMax (finsert 0 (character_data_ptr.the_ref |`| character_data_ptrs h))))
lemma new_character_data_ptr_new:
"character_data_ptr.Ref (Suc (fMax (finsert 0 (character_data_ptr.the_ref |`| character_data_ptrs h))))
|\<notin>| character_data_ptrs h"
by (metis Suc_n_not_le_n character_data_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
by (metis Suc_n_not_le_n character_data_ptr.sel(1) fMax_ge fimage_finsert finsertI1
finsertI2 set_finsert)
lemma new\<^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_ptr_not_in_heap:
assumes "new\<^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 h = (new_character_data_ptr, h')"
shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms
unfolding new\<^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
by (metis Pair_inject character_data_ptrs_def fMax_finsert fempty_iff ffmember_filter fimage_is_fempty is_character_data_ptr_ref max_0L new_character_data_ptr_new)
by (metis Pair_inject character_data_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_character_data_ptr_ref max_0L new_character_data_ptr_new)
lemma new\<^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_new_ptr:
assumes "new\<^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 h = (new_character_data_ptr, h')"
@ -313,7 +316,7 @@ definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = (known_ptr ptr \<or> is_character_data_ptr ptr)"
lemma known_ptr_not_character_data_ptr:
lemma known_ptr_not_character_data_ptr:
"\<not>is_character_data_ptr ptr \<Longrightarrow> a_known_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
@ -331,13 +334,15 @@ lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object
apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce
lemma known_ptrs_preserved:
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^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 known_ptr defines known_ptrs = a_known_ptrs .

View File

@ -23,18 +23,18 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Document\<close>
text\<open>In this theory, we introduce the types for the Document class.\<close>
theory DocumentClass
imports
imports
CharacterDataClass
begin
begin
text\<open>The type @{type "doctype"} is a type synonym for @{type "string"}, defined
text\<open>The type @{type "doctype"} is a type synonym for @{type "string"}, defined
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
record ('node_ptr, 'element_ptr, 'character_data_ptr) RDocument = RObject +
@ -42,35 +42,35 @@ record ('node_ptr, 'element_ptr, 'character_data_ptr) RDocument = RObject +
doctype :: doctype
document_element :: "(_) element_ptr option"
disconnected_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document) Document
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_scheme"
register_default_tvars
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'Document) Document"
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node,
'Element, 'CharacterData, 'Document) Object
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option)
RDocument_ext + 'Object, 'Node, 'Element, 'CharacterData) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'Document) Object"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'Document) Object"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_ext + 'Object, 'Node,
= "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr,
('node_ptr, 'element_ptr, 'character_data_ptr, 'Document option) RDocument_ext + 'Object, 'Node,
'Element, 'CharacterData) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition document_ptr_kinds :: "(_) heap \<Rightarrow> (_) document_ptr fset"
where
"document_ptr_kinds heap = the |`| (cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`|
"document_ptr_kinds heap = the |`| (cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`|
(ffilter is_document_ptr_kind (object_ptr_kinds heap)))"
definition document_ptrs :: "(_) heap \<Rightarrow> (_) document_ptr fset"
@ -86,7 +86,7 @@ adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^su
definition cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:: "(_) Document \<Rightarrow> (_) Object"
where
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = (RObject.extend (RObject.truncate document)
"cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = (RObject.extend (RObject.truncate document)
(Inr (Inl (RObject.more document))))"
adhoc_overloading cast cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
@ -94,20 +94,20 @@ definition is_document_kind :: "(_) Object \<Rightarrow> bool"
where
"is_document_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
lemma document_ptr_kinds_simp [simp]:
"document_ptr_kinds (Heap (fmupd (cast document_ptr) document (the_heap h)))
lemma document_ptr_kinds_simp [simp]:
"document_ptr_kinds (Heap (fmupd (cast document_ptr) document (the_heap h)))
= {|document_ptr|} |\<union>| document_ptr_kinds h"
apply(auto simp add: document_ptr_kinds_def)[1]
by force
lemma document_ptr_kinds_commutes [simp]:
lemma document_ptr_kinds_commutes [simp]:
"cast document_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> document_ptr |\<in>| document_ptr_kinds h"
apply(auto simp add: object_ptr_kinds_def document_ptr_kinds_def)[1]
by (metis (no_types, lifting) document_ptr_casts_commute2 document_ptr_document_ptr_cast
by (metis (no_types, lifting) document_ptr_casts_commute2 document_ptr_document_ptr_cast
ffmember_filter fimage_eqI fset.map_comp option.sel)
definition get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) document_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Document option"
where
where
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h = Option.bind (get (cast document_ptr) h) cast"
adhoc_overloading get get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
@ -115,7 +115,7 @@ locale l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^s
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (CharacterDataClass.type_wf h \<and>
"a_type_wf h = (CharacterDataClass.type_wf h \<and>
(\<forall>document_ptr \<in> fset (document_ptr_kinds h). get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
@ -136,7 +136,8 @@ lemma get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_w
shows "document_ptr |\<in>| document_ptr_kinds h \<longleftrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None"
using l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by (metis document_ptr_kinds_commutes fmember.rep_eq is_none_bind is_none_simps(1) is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
by (metis document_ptr_kinds_commutes fmember.rep_eq is_none_bind is_none_simps(1)
is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
end
global_interpretation l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
@ -164,15 +165,15 @@ lemma cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub
apply(simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = None \<longleftrightarrow> \<not> (\<exists>document. cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t obj = Some document \<longleftrightarrow> cast document = obj"
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
by(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def
split: sum.splits)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document) = Some document"
@ -183,24 +184,26 @@ lemma cast_document_not_node [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node \<noteq> cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t document"
by(auto simp add: cast\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
lemma get_document_ptr_simp1 [simp]:
lemma get_document_ptr_simp1 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr document h) = Some document"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp2 [simp]:
"document_ptr \<noteq> document_ptr'
lemma get_document_ptr_simp2 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' document h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp3 [simp]:
lemma get_document_ptr_simp3 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp4 [simp]: "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
lemma get_document_ptr_simp4 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_document_ptr_simp5 [simp]:
lemma get_document_ptr_simp5 [simp]:
"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 character_data_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = 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 character_data_ptr h"
by(auto simp add: 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 put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_document_ptr_simp6 [simp]: "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^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 character_data_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
lemma get_document_ptr_simp6 [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^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 character_data_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]:
@ -217,18 +220,18 @@ lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>
abbreviation
abbreviation
create_document_obj :: "char list \<Rightarrow> (_) element_ptr option \<Rightarrow> (_) node_ptr list \<Rightarrow> (_) Document"
where
"create_document_obj doctype_arg document_element_arg disconnected_nodes_arg
\<equiv> \<lparr> RObject.nothing = (), RDocument.nothing = (), doctype = doctype_arg,
\<equiv> \<lparr> RObject.nothing = (), RDocument.nothing = (), doctype = doctype_arg,
document_element = document_element_arg,
disconnected_nodes = disconnected_nodes_arg, \<dots> = None \<rparr>"
definition new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_)heap \<Rightarrow> ((_) document_ptr \<times> (_) heap)"
where
"new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_document_ptr = document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| (document_ptrs h)))))
"new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_document_ptr = document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| (document_ptrs h)))))
in
(new_document_ptr, put new_document_ptr (create_document_obj '''' None []) h))"
@ -239,8 +242,8 @@ lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in
unfolding new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_document_ptr_new:
"document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| document_ptrs h))))
lemma new_document_ptr_new:
"document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| document_ptrs h))))
|\<notin>| document_ptrs h"
by (metis Suc_n_not_le_n document_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
@ -249,7 +252,7 @@ lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_no
shows "new_document_ptr |\<notin>| document_ptr_kinds h"
using assms
unfolding new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis Pair_inject document_ptrs_def fMax_finsert fempty_iff ffmember_filter
by (metis Pair_inject document_ptrs_def fMax_finsert fempty_iff ffmember_filter
fimage_is_fempty is_document_ptr_ref max_0L new_document_ptr_new)
lemma new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr:
@ -321,13 +324,15 @@ lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object
apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce
lemma known_ptrs_preserved:
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset:
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -41,33 +41,33 @@ subsubsection\<open>Node\<close>
record RNode = RObject
+ nothing :: unit
register_default_tvars "'Node RNode_ext"
register_default_tvars "'Node RNode_ext"
type_synonym 'Node Node = "'Node RNode_scheme"
register_default_tvars "'Node Node"
register_default_tvars "'Node Node"
type_synonym ('Object, 'Node) Object = "('Node RNode_ext + 'Object) Object"
register_default_tvars "('Object, 'Node) Object"
register_default_tvars "('Object, 'Node) Object"
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node) heap
= "('node_ptr node_ptr + 'object_ptr, 'Node RNode_ext + 'Object) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'Object, 'Node) heap"
"('object_ptr, 'node_ptr, 'Object, 'Node) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit) heap"
definition node_ptr_kinds :: "(_) heap \<Rightarrow> (_) node_ptr fset"
where
"node_ptr_kinds heap =
"node_ptr_kinds heap =
(the |`| (cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_node_ptr_kind (object_ptr_kinds heap))))"
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (Heap (fmupd (cast node_ptr) node (the_heap h)))
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (Heap (fmupd (cast node_ptr) node (the_heap h)))
= {|node_ptr|} |\<union>| node_ptr_kinds h"
apply(auto simp add: node_ptr_kinds_def)[1]
by force
definition cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) Object \<Rightarrow> (_) Node option"
where
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = (case RObject.more obj of Inl node
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = (case RObject.more obj of Inl node
\<Rightarrow> Some (RObject.extend (RObject.truncate obj) node) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e
@ -81,7 +81,7 @@ definition is_node_kind :: "(_) Object \<Rightarrow> bool"
"is_node_kind ptr \<longleftrightarrow> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr \<noteq> None"
definition get\<^sub>N\<^sub>o\<^sub>d\<^sub>e :: "(_) node_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Node option"
where
where
"get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h = Option.bind (get (cast node_ptr) h) cast"
adhoc_overloading get get\<^sub>N\<^sub>o\<^sub>d\<^sub>e
@ -89,7 +89,7 @@ locale l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ObjectClass.type_wf h
"a_type_wf h = (ObjectClass.type_wf h
\<and> (\<forall>node_ptr \<in> fset( node_ptr_kinds h). get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e defines type_wf = a_type_wf .
@ -110,8 +110,8 @@ lemma get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf:
shows "node_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None"
using l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_axioms assms
apply(simp add: type_wf_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
by (metis bind_eq_None_conv ffmember_filter fimage_eqI fmember.rep_eq is_node_ptr_kind_cast
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf node_ptr_casts_commute2 node_ptr_kinds_def option.sel option.simps(3))
by (metis bind_eq_None_conv ffmember_filter fimage_eqI fmember.rep_eq is_node_ptr_kind_cast
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf node_ptr_casts_commute2 node_ptr_kinds_def option.sel option.simps(3))
end
global_interpretation l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf
@ -127,7 +127,7 @@ lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ptr_in_heap:
shows "node_ptr |\<in>| node_ptr_kinds h'"
using assms
unfolding put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def node_ptr_kinds_def
by (metis ffmember_filter fimage_eqI is_node_ptr_kind_cast node_ptr_casts_commute2
by (metis ffmember_filter fimage_eqI is_node_ptr_kind_cast node_ptr_casts_commute2
option.sel put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap)
lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs:
@ -136,14 +136,14 @@ lemma put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_put_ptrs:
using assms
by (simp add: put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs)
lemma node_ptr_kinds_commutes [simp]:
lemma node_ptr_kinds_commutes [simp]:
"cast node_ptr |\<in>| object_ptr_kinds h \<longleftrightarrow> node_ptr |\<in>| node_ptr_kinds h"
apply(auto simp add: node_ptr_kinds_def split: option.splits)[1]
by (metis (no_types, lifting) ffmember_filter fimage_eqI fset.map_comp
by (metis (no_types, lifting) ffmember_filter fimage_eqI fset.map_comp
is_node_ptr_kind_none node_ptr_casts_commute2
option.distinct(1) option.sel)
lemma node_empty [simp]:
lemma node_empty [simp]:
"\<lparr>RObject.nothing = (), RNode.nothing = (), \<dots> = RNode.more node\<rparr> = node"
by simp
@ -151,7 +151,7 @@ lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub
apply(simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def)
by (metis (full_types) RObject.surjective old.unit.exhaust)
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none [simp]:
lemma cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none [simp]:
"cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e obj = None \<longleftrightarrow> \<not> (\<exists>node. cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node = obj)"
apply(auto simp add: cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def RObject.extend_def split: sum.splits)[1]
by (metis (full_types) RObject.select_convs(2) RObject.surjective old.unit.exhaust)
@ -181,23 +181,28 @@ definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset
known_ptrs_new_ptr
by blast
lemma get_node_ptr_simp1 [simp]: "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h) = Some node"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma get_node_ptr_simp2 [simp]:
lemma get_node_ptr_simp2 [simp]:
"node_ptr \<noteq> node_ptr' \<Longrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr' node h) = get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)

View File

@ -23,12 +23,12 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Object\<close>
text\<open>In this theory, we introduce the definition of the class Object. This class is the
text\<open>In this theory, we introduce the definition of the class Object. This class is the
common superclass of our class model.\<close>
theory ObjectClass
@ -39,27 +39,27 @@ begin
record RObject =
nothing :: unit
register_default_tvars "'Object RObject_ext"
register_default_tvars "'Object RObject_ext"
type_synonym 'Object Object = "'Object RObject_scheme"
register_default_tvars "'Object Object"
register_default_tvars "'Object Object"
datatype ('object_ptr, 'Object) heap = Heap (the_heap: "((_) object_ptr, (_) Object) fmap")
register_default_tvars "('object_ptr, 'Object) heap"
register_default_tvars "('object_ptr, 'Object) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit) heap"
definition object_ptr_kinds :: "(_) heap \<Rightarrow> (_) object_ptr fset"
where
"object_ptr_kinds = fmdom \<circ> the_heap"
lemma object_ptr_kinds_simp [simp]:
"object_ptr_kinds (Heap (fmupd object_ptr object (the_heap h)))
lemma object_ptr_kinds_simp [simp]:
"object_ptr_kinds (Heap (fmupd object_ptr object (the_heap h)))
= {|object_ptr|} |\<union>| object_ptr_kinds h"
by(auto simp add: object_ptr_kinds_def)
definition get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Object option"
where
"get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = fmlookup (the_heap h) ptr"
adhoc_overloading get get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
adhoc_overloading get get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
locale l_type_wf_def\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
begin
@ -102,7 +102,7 @@ lemma put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_put_ptrs:
assumes "put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h = h'"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|object_ptr|}"
using assms
by (metis comp_apply fmdom_fmupd funion_finsert_right heap.sel object_ptr_kinds_def
by (metis comp_apply fmdom_fmupd funion_finsert_right heap.sel object_ptr_kinds_def
sup_bot.right_neutral put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemma object_more_extend_id [simp]: "more (extend x y) = y"
@ -117,7 +117,7 @@ definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
where
"a_known_ptr ptr = False"
lemma known_ptr_not_object_ptr:
lemma known_ptr_not_object_ptr:
"a_known_ptr ptr \<Longrightarrow> \<not>is_object_ptr ptr \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptr_def)
end
@ -127,9 +127,13 @@ lemmas known_ptr_defs = a_known_ptr_def
locale l_known_ptrs = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool" +
fixes known_ptrs :: "(_) heap \<Rightarrow> bool"
assumes known_ptrs_known_ptr: "known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
assumes known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h = known_ptrs h'"
assumes known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'"
assumes known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'"
assumes known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h = known_ptrs h'"
assumes known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'"
assumes known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
known_ptrs h \<Longrightarrow> known_ptrs h'"
locale l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool"
begin
@ -137,16 +141,20 @@ definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
lemma known_ptrs_known_ptr:
"a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
@ -159,8 +167,8 @@ lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
lemma get_object_ptr_simp1 [simp]: "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h) = Some object"
by(simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
lemma get_object_ptr_simp2 [simp]:
"object_ptr \<noteq> object_ptr'
lemma get_object_ptr_simp2 [simp]:
"object_ptr \<noteq> object_ptr'
\<Longrightarrow> get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr' object h) = get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr h"
by(simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
@ -169,11 +177,11 @@ subsection\<open>Limited Heap Modifications\<close>
definition heap_unchanged_except :: "(_) object_ptr set \<Rightarrow> (_) heap \<Rightarrow> (_) heap \<Rightarrow> bool"
where
"heap_unchanged_except S h h' = (\<forall>ptr \<in> (fset (object_ptr_kinds h)
"heap_unchanged_except S h h' = (\<forall>ptr \<in> (fset (object_ptr_kinds h)
\<union> (fset (object_ptr_kinds h'))) - S. get ptr h = get ptr h')"
definition delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "(_) object_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) heap option" where
"delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = (if ptr |\<in>| object_ptr_kinds h then Some (Heap (fmdrop ptr (the_heap h)))
"delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr h = (if ptr |\<in>| object_ptr_kinds h then Some (Heap (fmdrop ptr (the_heap h)))
else None)"
lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_pointer_removed:
@ -201,15 +209,15 @@ definition "create_heap xs = Heap (fmap_of_list xs)"
code_datatype ObjectClass.heap.Heap create_heap
lemma object_ptr_kinds_code3 [code]:
lemma object_ptr_kinds_code3 [code]:
"fmlookup (the_heap (create_heap xs)) x = map_of xs x"
by(auto simp add: create_heap_def fmlookup_of_list)
lemma object_ptr_kinds_code4 [code]:
lemma object_ptr_kinds_code4 [code]:
"the_heap (create_heap xs) = fmap_of_list xs"
by(simp add: create_heap_def)
lemma object_ptr_kinds_code5 [code]:
lemma object_ptr_kinds_code5 [code]:
"the_heap (Heap x) = x"
by simp

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -46,7 +46,7 @@ consts put_M :: 'a
consts get_M :: 'a
consts delete_M :: 'a
lemma sorted_list_of_set_eq [dest]:
lemma sorted_list_of_set_eq [dest]:
"sorted_list_of_set (fset x) = sorted_list_of_set (fset y) \<Longrightarrow> x = y"
by (metis finite_fset fset_inject sorted_list_of_set(1))
@ -70,18 +70,18 @@ lemma ptr_kinds_M_pure [simp]: "pure a_ptr_kinds_M h"
lemma ptr_kinds_ptr_kinds_M [simp]: "ptr \<in> set |h \<turnstile> a_ptr_kinds_M|\<^sub>r \<longleftrightarrow> ptr |\<in>| ptr_kinds h"
by(simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds [simp]:
lemma ptr_kinds_M_ptr_kinds [simp]:
"h \<turnstile> a_ptr_kinds_M \<rightarrow>\<^sub>r xa \<longleftrightarrow> xa = sorted_list_of_set (fset (ptr_kinds h))"
by(auto simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds_returns_result [simp]:
lemma ptr_kinds_M_ptr_kinds_returns_result [simp]:
"h \<turnstile> a_ptr_kinds_M \<bind> f \<rightarrow>\<^sub>r x \<longleftrightarrow> h \<turnstile> f (sorted_list_of_set (fset (ptr_kinds h))) \<rightarrow>\<^sub>r x"
by(auto simp add: a_ptr_kinds_M_def)
lemma ptr_kinds_M_ptr_kinds_returns_heap [simp]:
lemma ptr_kinds_M_ptr_kinds_returns_heap [simp]:
"h \<turnstile> a_ptr_kinds_M \<bind> f \<rightarrow>\<^sub>h h' \<longleftrightarrow> h \<turnstile> f (sorted_list_of_set (fset (ptr_kinds h))) \<rightarrow>\<^sub>h h'"
by(auto simp add: a_ptr_kinds_M_def)
end
locale l_get_M =
locale l_get_M =
fixes get :: "'ptr \<Rightarrow> 'heap \<Rightarrow> 'obj option"
fixes type_wf :: "'heap \<Rightarrow> bool"
fixes ptr_kinds :: "'heap \<Rightarrow> 'ptr fset"
@ -129,14 +129,14 @@ lemma put_M_ok:
lemma put_M_ptr_in_heap:
"h \<turnstile> ok (a_put_M ptr setter v) \<Longrightarrow> ptr |\<in>| ptr_kinds h"
by(auto simp add: a_put_M_def intro!: bind_is_OK_I2 elim: get_M_ptr_in_heap
by(auto simp add: a_put_M_def intro!: bind_is_OK_I2 elim: get_M_ptr_in_heap
dest: is_OK_returns_result_I elim!: bind_is_OK_E)
end
subsection \<open>Setup for Defining Partial Functions\<close>
lemma execute_admissible:
lemma execute_admissible:
"ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e)))
((\<lambda>a. \<forall>(h::'heap) h2 (r::'result). h \<turnstile> a = Inr (r, h2) \<longrightarrow> P h h2 r) \<circ> Prog)"
proof (unfold comp_def, rule ccpo.admissibleI, clarify)
@ -153,16 +153,16 @@ proof (unfold comp_def, rule ccpo.admissibleI, clarify)
by force
qed
lemma execute_admissible2:
lemma execute_admissible2:
"ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e)))
((\<lambda>a. \<forall>(h::'heap) h' h2 h2' (r::'result) r'.
((\<lambda>a. \<forall>(h::'heap) h' h2 h2' (r::'result) r'.
h \<turnstile> a = Inr (r, h2) \<longrightarrow> h' \<turnstile> a = Inr (r', h2') \<longrightarrow> P h h' h2 h2' r r') \<circ> Prog)"
proof (unfold comp_def, rule ccpo.admissibleI, clarify)
fix A :: "('heap \<Rightarrow> 'e + 'result \<times> 'heap) set"
let ?lub = "Prog (fun_lub (flat_lub (Inl e)) A)"
fix h h' h2 h2' r r'
assume 1: "Complete_Partial_Order.chain (fun_ord (flat_ord (Inl e))) A"
and 2 [rule_format]: "\<forall>xa\<in>A. \<forall>h h' h2 h2' r r'. h \<turnstile> Prog xa = Inr (r, h2)
and 2 [rule_format]: "\<forall>xa\<in>A. \<forall>h h' h2 h2' r r'. h \<turnstile> Prog xa = Inr (r, h2)
\<longrightarrow> h' \<turnstile> Prog xa = Inr (r', h2') \<longrightarrow> P h h' h2 h2' r r'"
and 4: "h \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r, h2)"
and 5: "h' \<turnstile> Prog (fun_lub (flat_lub (Inl e)) A) = Inr (r', h2')"
@ -180,18 +180,18 @@ proof (unfold comp_def, rule ccpo.admissibleI, clarify)
"f \<in> A" and
"h \<turnstile> Prog f = Inr (r, h2)" and
"h' \<turnstile> Prog f = Inr (r', h2')"
using 1 4 5
using 1 4 5
apply(auto simp add: chain_def fun_ord_def flat_ord_def execute_def)[1]
by (metis Inl_Inr_False)
then show "P h h' h2 h2' r r'"
by(fact 2)
qed
definition dom_prog_ord ::
definition dom_prog_ord ::
"('heap, exception, 'result) prog \<Rightarrow> ('heap, exception, 'result) prog \<Rightarrow> bool" where
"dom_prog_ord = img_ord (\<lambda>a b. execute b a) (fun_ord (flat_ord (Inl NonTerminationException)))"
definition dom_prog_lub ::
definition dom_prog_lub ::
"('heap, exception, 'result) prog set \<Rightarrow> ('heap, exception, 'result) prog" where
"dom_prog_lub = img_lub (\<lambda>a b. execute b a) Prog (fun_lub (flat_lub (Inl NonTerminationException)))"
@ -200,7 +200,7 @@ lemma dom_prog_lub_empty: "dom_prog_lub {} = error NonTerminationException"
lemma dom_prog_interpretation: "partial_function_definitions dom_prog_ord dom_prog_lub"
proof -
have "partial_function_definitions (fun_ord (flat_ord (Inl NonTerminationException)))
have "partial_function_definitions (fun_ord (flat_ord (Inl NonTerminationException)))
(fun_lub (flat_lub (Inl NonTerminationException)))"
by (rule partial_function_lift) (rule flat_interpretation)
then show ?thesis
@ -212,15 +212,15 @@ interpretation dom_prog: partial_function_definitions dom_prog_ord dom_prog_lub
rewrites "dom_prog_lub {} \<equiv> error NonTerminationException"
by (fact dom_prog_interpretation)(simp add: dom_prog_lub_empty)
lemma admissible_dom_prog:
lemma admissible_dom_prog:
"dom_prog.admissible (\<lambda>f. \<forall>x h h' r. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h' \<longrightarrow> P x h h' r)"
proof (rule admissible_fun[OF dom_prog_interpretation])
fix x
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h'
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h'
\<longrightarrow> P x h h' r)"
unfolding dom_prog_ord_def dom_prog_lub_def
proof (intro admissible_image partial_function_lift flat_interpretation)
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
(fun_ord (flat_ord (Inl NonTerminationException)))
((\<lambda>a. \<forall>h h' r. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> P x h h' r) \<circ> Prog)"
by(auto simp add: execute_admissible returns_result_def returns_heap_def split: sum.splits)
@ -234,20 +234,20 @@ proof (rule admissible_fun[OF dom_prog_interpretation])
qed
lemma admissible_dom_prog2:
"dom_prog.admissible (\<lambda>f. \<forall>x h h2 h' h2' r r2. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h'
"dom_prog.admissible (\<lambda>f. \<forall>x h h2 h' h2' r r2. h \<turnstile> f x \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> f x \<rightarrow>\<^sub>h h'
\<longrightarrow> h2 \<turnstile> f x \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> f x \<rightarrow>\<^sub>h h2' \<longrightarrow> P x h h2 h' h2' r r2)"
proof (rule admissible_fun[OF dom_prog_interpretation])
fix x
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r
show "ccpo.admissible dom_prog_lub dom_prog_ord (\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r
\<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2' \<longrightarrow> P x h h2 h' h2' r r2)"
unfolding dom_prog_ord_def dom_prog_lub_def
proof (intro admissible_image partial_function_lift flat_interpretation)
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
show "ccpo.admissible (fun_lub (flat_lub (Inl NonTerminationException)))
(fun_ord (flat_ord (Inl NonTerminationException)))
((\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2'
((\<lambda>a. \<forall>h h2 h' h2' r r2. h \<turnstile> a \<rightarrow>\<^sub>r r \<longrightarrow> h \<turnstile> a \<rightarrow>\<^sub>h h' \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>r r2 \<longrightarrow> h2 \<turnstile> a \<rightarrow>\<^sub>h h2'
\<longrightarrow> P x h h2 h' h2' r r2) \<circ> Prog)"
by(auto simp add: returns_result_def returns_heap_def intro!: ccpo.admissibleI
dest!: ccpo.admissibleD[OF execute_admissible2[where P="P x"]]
by(auto simp add: returns_result_def returns_heap_def intro!: ccpo.admissibleI
dest!: ccpo.admissibleD[OF execute_admissible2[where P="P x"]]
split: sum.splits)
next
show "\<And>x y. (\<lambda>b. b \<turnstile> x) = (\<lambda>b. b \<turnstile> y) \<Longrightarrow> x = y"
@ -266,7 +266,7 @@ lemma fixp_induct_dom_prog:
assumes mono: "\<And>x. monotone (fun_ord dom_prog_ord) dom_prog_ord (\<lambda>f. U (F (C f)) x)"
assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub dom_prog_lub) (fun_ord dom_prog_ord) (\<lambda>f. U (F (C f))))"
assumes inverse2: "\<And>f. U (C f) = f"
assumes step: "\<And>f x h h' r. (\<And>x h h' r. h \<turnstile> (U f x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U f x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r)
assumes step: "\<And>f x h h' r. (\<And>x h h' r. h \<turnstile> (U f x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U f x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r)
\<Longrightarrow> h \<turnstile> (U (F f) x) \<rightarrow>\<^sub>r r \<Longrightarrow> h \<turnstile> (U (F f) x) \<rightarrow>\<^sub>h h' \<Longrightarrow> P x h h' r"
assumes defined: "h \<turnstile> (U f x) \<rightarrow>\<^sub>r r" and "h \<turnstile> (U f x) \<rightarrow>\<^sub>h h'"
shows "P x h h' r"
@ -315,7 +315,7 @@ proof (rule monotoneI)
proof (rule dom_prog_ordI)
fix h
from 1 show "h \<turnstile> ?L \<rightarrow>\<^sub>e NonTerminationException \<or> h \<turnstile> ?L = h \<turnstile> ?R"
apply(rule dom_prog_ordE)
apply(rule dom_prog_ordE)
apply(auto)[1]
using bind_cong by fastforce
qed
@ -358,7 +358,7 @@ lemma mono_dom_prog1 [partial_function_mono]:
assumes "\<And>x. (mono_dom_prog (\<lambda>f. g f x))"
shows "mono_dom_prog (\<lambda>f. map_M (g f) xs)"
using assms
apply (induct xs)
apply (induct xs)
by(auto simp add: call_mono dom_prog.const_mono intro!: bind_mono)
lemma mono_dom_prog2 [partial_function_mono]:
@ -366,10 +366,10 @@ lemma mono_dom_prog2 [partial_function_mono]:
assumes "\<And>x. (mono_dom_prog (\<lambda>f. g f x))"
shows "mono_dom_prog (\<lambda>f. forall_M (g f) xs)"
using assms
apply (induct xs)
apply (induct xs)
by(auto simp add: call_mono dom_prog.const_mono intro!: bind_mono)
lemma sorted_list_set_cong [simp]:
lemma sorted_list_set_cong [simp]:
"sorted_list_of_set (fset FS) = sorted_list_of_set (fset FS') \<longleftrightarrow> FS = FS'"
by auto

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -35,36 +35,36 @@ theory CharacterDataMonad
"../classes/CharacterDataClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog"
global_interpretation l_ptr_kinds_M character_data_ptr_kinds
global_interpretation l_ptr_kinds_M character_data_ptr_kinds
defines character_data_ptr_kinds_M = a_ptr_kinds_M .
lemmas character_data_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma character_data_ptr_kinds_M_eq:
assumes "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> character_data_ptr_kinds_M|\<^sub>r = |h' \<turnstile> character_data_ptr_kinds_M|\<^sub>r"
using assms
by(auto simp add: character_data_ptr_kinds_M_defs node_ptr_kinds_M_defs
using assms
by(auto simp add: character_data_ptr_kinds_M_defs node_ptr_kinds_M_defs
character_data_ptr_kinds_def)
lemma character_data_ptr_kinds_M_reads:
lemma character_data_ptr_kinds_M_reads:
"reads (\<Union>node_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node_ptr RObject.nothing)}) character_data_ptr_kinds_M h h'"
using node_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs character_data_ptr_kinds_M_defs
apply (simp add: reads_def node_ptr_kinds_M_defs character_data_ptr_kinds_M_defs
character_data_ptr_kinds_def preserved_def)
by (smt node_ptr_kinds_small preserved_def unit_all_impI)
global_interpretation l_dummy defines get_M\<^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 = "l_get_M.a_get_M 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" .
lemma get_M_is_l_get_M: "l_get_M 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 type_wf character_data_ptr_kinds"
apply(simp add: 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_type_wf l_get_M_def)
by (metis (no_types, hide_lams) NodeMonad.get_M_is_l_get_M bind_eq_Some_conv
by (metis (no_types, hide_lams) NodeMonad.get_M_is_l_get_M bind_eq_Some_conv
character_data_ptr_kinds_commutes 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 l_get_M_def option.distinct(1))
lemmas get_M_defs = get_M\<^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[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
@ -84,7 +84,7 @@ end
global_interpretation l_get_M\<^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_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf character_data_ptr_kinds 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 put\<^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
global_interpretation l_put_M type_wf character_data_ptr_kinds 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 put\<^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
rewrites "a_get_M = get_M\<^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" defines put_M\<^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 = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^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)
@ -109,98 +109,98 @@ global_interpretation l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^
lemma CharacterData_simp1 [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = v) \<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp1 [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = v) \<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> h' \<turnstile> get_M\<^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 character_data_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp2 [simp]:
"character_data_ptr \<noteq> character_data_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^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 character_data_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp3 [simp]: "
(\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
(\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^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 character_data_ptr' getter) h h'"
apply(cases "character_data_ptr = character_data_ptr'")
apply(cases "character_data_ptr = character_data_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp4 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp4 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp5 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp5 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^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 character_data_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp6 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma CharacterData_simp6 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply (cases "cast character_data_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
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 preserved_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply (cases "cast character_data_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
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 preserved_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp7 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp7 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast character_data_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
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 preserved_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs
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 preserved_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp8 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp8 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs 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 put\<^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 NodeMonad.get_M_defs
by(auto simp add: put_M_defs get_M_defs 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 put\<^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 NodeMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp9 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma CharacterData_simp9 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast character_data_ptr \<noteq> node_ptr")
by(auto simp add: put_M_defs get_M_defs 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 put\<^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
NodeMonad.get_M_defs preserved_def split: option.splits bind_splits
by(auto simp add: put_M_defs get_M_defs 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 put\<^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
NodeMonad.get_M_defs preserved_def split: option.splits bind_splits
dest: get_heap_E)
lemma CharacterData_simp10 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp10 [simp]:
"cast character_data_ptr \<noteq> node_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^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 character_data_ptr getter) h h'"
by(auto simp add: NodeMonad.put_M_defs get_M_defs 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 NodeMonad.get_M_defs
by(auto simp add: NodeMonad.put_M_defs get_M_defs 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 NodeMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp11 [simp]:
"cast character_data_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp11 [simp]:
"cast character_data_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs 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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^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
ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs 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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^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
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma CharacterData_simp12 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast character_data_ptr \<noteq> object_ptr")
apply(auto simp add: put_M_defs get_M_defs 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 put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
apply(auto simp add: put_M_defs get_M_defs 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 put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)[1]
by(auto simp add: put_M_defs get_M_defs 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 put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs 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 put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)[1]
lemma CharacterData_simp13 [simp]:
"cast character_data_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
lemma CharacterData_simp13 [simp]:
"cast character_data_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^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 character_data_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs 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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
by(auto simp add: ObjectMonad.put_M_defs get_M_defs 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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma new_element_get_M\<^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:
lemma new_element_get_M\<^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:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^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 ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def split: prod.splits option.splits
by(auto simp add: new_element_def get_M_defs preserved_def split: prod.splits option.splits
elim!: bind_returns_result_E bind_returns_heap_E)
@ -225,7 +225,7 @@ lemma new_character_data_ptr_in_heap:
shows "new_character_data_ptr |\<in>| character_data_ptr_kinds h'"
using assms
unfolding new_character_data_def
by(auto simp add: new_character_data_def new\<^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 Let_def put\<^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_ptr_in_heap
by(auto simp add: new_character_data_def new\<^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 Let_def put\<^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_ptr_in_heap
is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
@ -234,7 +234,7 @@ lemma new_character_data_ptr_not_in_heap:
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms new\<^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_ptr_not_in_heap
by(auto simp add: new_character_data_def split: prod.splits
by(auto simp add: new_character_data_def split: prod.splits
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_new_ptr:
@ -242,7 +242,7 @@ lemma new_character_data_new_ptr:
and "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using assms new\<^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_new_ptr
by(auto simp add: new_character_data_def split: prod.splits
by(auto simp add: new_character_data_def split: prod.splits
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_is_character_data_ptr:
@ -256,41 +256,41 @@ lemma new_character_data_child_nodes:
assumes "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr"
shows "h' \<turnstile> get_M new_character_data_ptr val \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_character_data_def new\<^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 Let_def
by(auto simp add: get_M_defs new_character_data_def new\<^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 Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
lemma new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> cast new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: new_character_data_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
lemma new_character_data_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> cast new_character_data_ptr \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_character_data_def NodeMonad.get_M_defs preserved_def
by(auto simp add: new_character_data_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
lemma new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def ElementMonad.get_M_defs preserved_def
by(auto simp add: new_character_data_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^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:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
lemma new_character_data_get_M\<^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:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr
\<Longrightarrow> ptr \<noteq> new_character_data_ptr \<Longrightarrow> preserved (get_M\<^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 ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_CharacterData_ptr_simp [simp]:
"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 character_data_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
lemma get_CharacterData_ptr_simp [simp]:
"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 character_data_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast character_data_ptr then cast obj else get character_data_ptr h)"
by(auto simp add: 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 split: option.splits Option.bind_splits)
lemma Character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = character_data_ptr_kinds h |\<union>|
lemma Character_data_ptr_kinds_simp [simp]:
"character_data_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = character_data_ptr_kinds h |\<union>|
(if is_character_data_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: character_data_ptr_kinds_def is_node_ptr_kind_def split: option.splits)
@ -307,7 +307,7 @@ lemma type_wf_put_ptr_not_in_heap_E:
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E
apply(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
@ -319,7 +319,8 @@ lemma type_wf_put_ptr_in_heap_E:
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
by (metis (no_types, lifting) ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf assms(2) bind.bind_lunit cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv 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 notin_fset option.collapse)
by (metis (no_types, lifting) ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf assms(2) bind.bind_lunit
cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^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_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv 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 notin_fset option.collapse)
subsection\<open>Preserving Types\<close>
@ -327,13 +328,13 @@ lemma new_element_type_wf_preserved [simp]:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>h h'"
shows "type_wf h = type_wf h'"
using assms
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I split: if_splits)[1]
using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast
using element_ptrs_def apply fastforce
using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_element_ptr_ref)
lemma new_element_is_l_new_element: "l_new_element type_wf"
@ -342,20 +343,20 @@ lemma new_element_is_l_new_element: "l_new_element type_wf"
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
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]
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]
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]
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]
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]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -363,70 +364,70 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_typ
done
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]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_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]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
apply (metis finite_set_in)
done
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'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
apply (metis finite_set_in)
done
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'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.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]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -434,11 +435,11 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_
done
lemma new_character_data_type_wf_preserved [simp]:
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_character_data_def new\<^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 Let_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
apply(auto simp add: new_character_data_def new\<^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 Let_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(simp_all add: type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs is_node_kind_def)
by (meson new\<^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 new\<^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_ptr_not_in_heap)
@ -450,36 +451,36 @@ lemma new_character_data_is_l_new_character_data: "l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma put_M\<^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_val_type_wf_preserved [simp]:
lemma put_M\<^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_val_type_wf_preserved [simp]:
"h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
CharacterDataClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e CharacterDataClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
split: option.splits)[1]
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 CharacterDataMonad.get_M_defs
ObjectClass.a_type_wf_def
split: option.splits)[1]
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in 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)
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in 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)
apply (metis finite_set_in)
done
lemma character_data_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
by(simp add: character_data_ptr_kinds_def node_ptr_kinds_def preserved_def
by(simp add: character_data_ptr_kinds_def node_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma character_data_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
using writes_small_big[OF assms]
@ -491,27 +492,27 @@ lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved (get_M\<^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 character_data_ptr
assumes "\<And>character_data_ptr. preserved (get_M\<^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 character_data_ptr
RCharacterData.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3)]
using type_wf_preserved_small[OF assms(1) assms(2) assms(3)]
allI[OF assms(4), of id, simplified] character_data_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs character_data_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs character_data_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
apply(force)
by force
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved (get_M\<^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 character_data_ptr
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved (get_M\<^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 character_data_ptr
RCharacterData.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
@ -523,9 +524,11 @@ proof -
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_def ElementMonad.type_wf_drop
apply(auto simp add: type_wf_def ElementMonad.type_wf_drop
l_type_wf_def\<^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.a_type_wf_def)[1]
using type_wf_drop
by (metis (no_types, lifting) ElementClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf character_data_ptr_kinds_commutes finite_set_in 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 node_ptr_kinds_commutes object_ptr_kinds_code5)
by (metis (no_types, lifting) ElementClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
character_data_ptr_kinds_commutes finite_set_in 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 node_ptr_kinds_commutes object_ptr_kinds_code5)
end

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -36,11 +36,11 @@ theory DocumentMonad
"../classes/DocumentClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'result) dom_prog"
global_interpretation l_ptr_kinds_M document_ptr_kinds defines document_ptr_kinds_M = a_ptr_kinds_M .
@ -49,21 +49,21 @@ lemmas document_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma document_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> document_ptr_kinds_M|\<^sub>r = |h' \<turnstile> document_ptr_kinds_M|\<^sub>r"
using assms
using assms
by(auto simp add: document_ptr_kinds_M_defs object_ptr_kinds_M_defs document_ptr_kinds_def)
lemma document_ptr_kinds_M_reads:
lemma document_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)}) document_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def object_ptr_kinds_M_defs document_ptr_kinds_M_defs
document_ptr_kinds_def preserved_def cong del: image_cong_simp)
apply (simp add: reads_def object_ptr_kinds_M_defs document_ptr_kinds_M_defs
document_ptr_kinds_def preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, hide_lams) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)
done
global_interpretation l_dummy defines get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = "l_get_M.a_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf document_ptr_kinds"
apply(simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf l_get_M_def)
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs bind_eq_None_conv
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs bind_eq_None_conv
document_ptr_kinds_commutes get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.simps(3))
lemmas get_M_defs = get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
@ -83,7 +83,7 @@ end
global_interpretation l_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf document_ptr_kinds get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
global_interpretation l_put_M type_wf document_ptr_kinds get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t" defines put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -106,84 +106,84 @@ end
global_interpretation l_put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
lemma document_put_get [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
lemma document_put_get [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Mdocument_preserved1 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Mdocument_preserved1 [simp]:
"document_ptr \<noteq> document_ptr'
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma document_put_get_preserved [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
lemma document_put_get_preserved [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = getter x)
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr' getter) h h'"
apply(cases "document_ptr = document_ptr'")
apply(cases "document_ptr = document_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved2 [simp]:
lemma get_M_Mdocument_preserved2 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved3 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Mdocument_preserved3 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved4 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Mdocument_preserved4 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast document_ptr \<noteq> object_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mdocument_preserved5 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Mdocument_preserved5 [simp]:
"cast document_ptr \<noteq> object_ptr
\<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved6 [simp]:
lemma get_M_Mdocument_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved7 [simp]:
lemma get_M_Mdocument_preserved7 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved8 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Mdocument_preserved8 [simp]:
"h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^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 character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved9 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Mdocument_preserved9 [simp]:
"h \<turnstile> put_M\<^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 character_data_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mdocument_preserved10 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Mdocument_preserved10 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast document_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
apply(cases "cast document_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
lemma new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
lemma new_character_data_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
@ -236,7 +236,7 @@ lemma new_document_doctype:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr doctype \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_document_element:
@ -244,7 +244,7 @@ lemma new_document_document_element:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr document_element \<rightarrow>\<^sub>r None"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_disconnected_nodes:
@ -252,33 +252,33 @@ lemma new_document_disconnected_nodes:
assumes "h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr"
shows "h' \<turnstile> get_M new_document_ptr disconnected_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
lemma new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> ptr \<noteq> cast new_document_ptr \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: new_document_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
lemma new_document_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_document_def NodeMonad.get_M_defs preserved_def
by(auto simp add: new_document_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
lemma new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def ElementMonad.get_M_defs preserved_def
by(auto simp add: new_document_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^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:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
lemma new_document_get_M\<^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:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr
\<Longrightarrow> preserved (get_M\<^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 ptr getter) h h'"
by(auto simp add: new_document_def CharacterDataMonad.get_M_defs preserved_def
by(auto simp add: new_document_def CharacterDataMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> ptr \<noteq> new_document_ptr
lemma new_document_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_document \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> new_document \<rightarrow>\<^sub>r new_document_ptr \<Longrightarrow> ptr \<noteq> new_document_ptr
\<Longrightarrow> preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_document_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
@ -287,13 +287,13 @@ lemma new_document_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n
subsection \<open>Modified Heaps\<close>
lemma get_document_ptr_simp [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
lemma get_document_ptr_simp [simp]:
"get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast document_ptr then cast obj else get document_ptr h)"
by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: option.splits Option.bind_splits)
lemma document_ptr_kidns_simp [simp]:
"document_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
lemma document_ptr_kidns_simp [simp]:
"document_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= document_ptr_kinds h |\<union>| (if is_document_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: document_ptr_kinds_def split: option.splits)
@ -310,7 +310,7 @@ lemma type_wf_put_ptr_not_in_heap_E:
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_not_in_heap_E
by(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
@ -320,145 +320,155 @@ lemma type_wf_put_ptr_in_heap_E:
assumes "is_document_ptr_kind ptr \<Longrightarrow> is_document_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E
apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def is_document_kind_def notin_fset option.exhaust_sel)
by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
is_document_kind_def notin_fset option.exhaust_sel)
subsection \<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]:
lemma new_element_type_wf_preserved [simp]:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def element_ptrs_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply fastforce
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_element_ptr_ref)
lemma new_element_is_l_new_element [instances]:
lemma new_element_is_l_new_element [instances]:
"l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
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]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
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'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
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'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def
bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse
option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
lemma new_character_data_type_wf_preserved [simp]:
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_kind_def
new_character_data_def new\<^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 Let_def put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
dest!: get_heap_E
elim!: bind_returns_heap_E2 bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
by (meson new\<^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 new\<^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_ptr_not_in_heap)
lemma new_character_data_is_l_new_character_data [instances]:
lemma new_character_data_is_l_new_character_data [instances]:
"l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma put_M\<^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_val_type_wf_preserved [simp]:
lemma put_M\<^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_val_type_wf_preserved [simp]:
"h \<turnstile> put_M character_data_ptr val_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
apply(auto simp add: CharacterDataMonad.put_M_defs put\<^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 put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_kind_def
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
dest!: get_heap_E elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis bind.bind_lzero finite_set_in 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 option.distinct(1) option.exhaust_sel)
apply (metis bind.bind_lzero finite_set_in 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 option.distinct(1) option.exhaust_sel)
by (metis finite_set_in)
lemma new_document_type_wf_preserved [simp]: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
apply(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I ElementMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
split: option.splits)[1]
using document_ptrs_def apply fastforce
apply (simp add: is_document_kind_def)
apply (metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter fimage_eqI is_document_ptr_ref)
apply (metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_document_ptr_ref)
done
locale l_new_document = l_type_wf +
@ -468,59 +478,59 @@ lemma new_document_is_l_new_document [instances]: "l_new_document type_wf"
using l_new_document.intro new_document_type_wf_preserved
by blast
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doctype_type_wf_preserved [simp]:
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doctype_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr doctype_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: get_M_defs)[1]
by (metis (mono_tags) error_returns_result finite_set_in option.exhaust_sel option.simps(4))
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr document_element_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^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
DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e
DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_ptr_kind_none
cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none is_document_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: get_M_defs is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs
apply(auto simp add: get_M_defs is_document_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs
split: option.splits)[1]
by (metis finite_set_in)
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr disconnected_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: put_M_defs put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
DocumentClass.type_wf\<^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
@ -529,13 +539,13 @@ lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disc
DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
is_node_ptr_kind_none
cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none is_document_kind_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def get_M_defs type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_document_kind_def get_M_defs type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
by (metis finite_set_in)
@ -547,7 +557,7 @@ lemma document_ptr_kinds_small:
lemma document_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "document_ptr_kinds h = document_ptr_kinds h'"
using writes_small_big[OF assms]
@ -558,33 +568,33 @@ lemma type_wf_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>character_data_ptr. preserved
assumes "\<And>character_data_ptr. preserved
(get_M\<^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 character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
shows "DocumentClass.type_wf h = DocumentClass.type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4)]
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4)]
allI[OF assms(5), of id, simplified] document_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply force
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs document_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
by force
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>character_data_ptr. preserved
(get_M\<^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 character_data_ptr RCharacterData.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>document_ptr. preserved (get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr RDocument.nothing) h h'"
shows "DocumentClass.type_wf h = DocumentClass.type_wf h'"
proof -
@ -599,5 +609,6 @@ lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel)
by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop
document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel)
end

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -35,11 +35,11 @@ theory ElementMonad
"ElementClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element,'result) dom_prog"
global_interpretation l_ptr_kinds_M element_ptr_kinds defines element_ptr_kinds_M = a_ptr_kinds_M .
@ -49,10 +49,10 @@ lemmas element_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma element_ptr_kinds_M_eq:
assumes "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> element_ptr_kinds_M|\<^sub>r = |h' \<turnstile> element_ptr_kinds_M|\<^sub>r"
using assms
using assms
by(auto simp add: element_ptr_kinds_M_defs node_ptr_kinds_M_defs element_ptr_kinds_def)
lemma element_ptr_kinds_M_reads:
lemma element_ptr_kinds_M_reads:
"reads (\<Union>element_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t element_ptr RObject.nothing)}) element_ptr_kinds_M h h'"
apply (simp add: reads_def node_ptr_kinds_M_defs element_ptr_kinds_M_defs element_ptr_kinds_def
node_ptr_kinds_M_reads preserved_def cong del: image_cong_simp)
@ -62,8 +62,8 @@ lemma element_ptr_kinds_M_reads:
global_interpretation l_dummy defines get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = "l_get_M.a_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t type_wf element_ptr_kinds"
apply(simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf l_get_M_def)
by (metis (no_types, lifting) ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs
bind_eq_Some_conv bind_eq_Some_conv element_ptr_kinds_commutes get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
by (metis (no_types, lifting) ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf ObjectClass.type_wf_defs
bind_eq_Some_conv bind_eq_Some_conv element_ptr_kinds_commutes get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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 node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
@ -84,8 +84,8 @@ end
global_interpretation l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf element_ptr_kinds get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t"
global_interpretation l_put_M type_wf element_ptr_kinds get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
rewrites "a_get_M = get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t"
defines put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -109,74 +109,74 @@ end
global_interpretation l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
lemma element_put_get [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
lemma element_put_get [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (setter (\<lambda>_. v) x) = v)
\<Longrightarrow> h' \<turnstile> get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter \<rightarrow>\<^sub>r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Element_preserved1 [simp]:
"element_ptr \<noteq> element_ptr' \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Element_preserved1 [simp]:
"element_ptr \<noteq> element_ptr' \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma element_put_get_preserved [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = getter x) \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
lemma element_put_get_preserved [simp]:
"(\<And>x. getter (setter (\<lambda>_. v) x) = getter x) \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' getter) h h'"
apply(cases "element_ptr = element_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def
apply(cases "element_ptr = element_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved3 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Element_preserved3 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast element_ptr = object_ptr")
by (auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
apply(cases "cast element_ptr = object_ptr")
by (auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma get_M_Element_preserved4 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Element_preserved4 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast element_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
apply(cases "cast element_ptr = node_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv
split: option.splits)
lemma get_M_Element_preserved5 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Element_preserved5 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Element_preserved6 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
apply(cases "cast element_ptr \<noteq> node_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Element_preserved7 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Element_preserved7 [simp]:
"cast element_ptr \<noteq> node_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def NodeMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved8 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Element_preserved8 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Element_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Element_preserved9 [simp]:
"h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast element_ptr \<noteq> object_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Element_preserved10 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Element_preserved10 [simp]:
"cast element_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection\<open>Creating Elements\<close>
@ -208,7 +208,7 @@ lemma new_element_ptr_not_in_heap:
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "new_element_ptr |\<notin>| element_ptr_kinds h"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
bind_returns_heap_E)
lemma new_element_new_ptr:
@ -216,7 +216,7 @@ lemma new_element_new_ptr:
and "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using assms new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_new_ptr
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
by(auto simp add: new_element_def split: prod.splits elim!: bind_returns_result_E
bind_returns_heap_E)
lemma new_element_is_element_ptr:
@ -230,7 +230,7 @@ lemma new_element_child_nodes:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr child_nodes \<rightarrow>\<^sub>r []"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_tag_name:
@ -238,7 +238,7 @@ lemma new_element_tag_name:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr tag_name \<rightarrow>\<^sub>r ''''"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_attrs:
@ -246,7 +246,7 @@ lemma new_element_attrs:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr attrs \<rightarrow>\<^sub>r fmempty"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_shadow_root_opt:
@ -254,35 +254,35 @@ lemma new_element_shadow_root_opt:
assumes "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr"
shows "h' \<turnstile> get_M new_element_ptr shadow_root_opt \<rightarrow>\<^sub>r None"
using assms
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
lemma new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: new_element_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
lemma new_element_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> cast new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'"
by(auto simp add: new_element_def NodeMonad.get_M_defs preserved_def
by(auto simp add: new_element_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> new_element_ptr
lemma new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t:
"h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr \<Longrightarrow> ptr \<noteq> new_element_ptr
\<Longrightarrow> preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_Element_ptr_simp [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
lemma get_Element_ptr_simp [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= (if ptr = cast element_ptr then cast obj else get element_ptr h)"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def split: option.splits Option.bind_splits)
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= element_ptr_kinds h |\<union>| (if is_element_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: element_ptr_kinds_def is_node_ptr_kind_def split: option.splits)
@ -299,7 +299,7 @@ lemma type_wf_put_ptr_not_in_heap_E:
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: NodeMonad.type_wf_put_ptr_not_in_heap_E
apply(auto simp add: type_wf_defs elim!: NodeMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
@ -318,8 +318,8 @@ lemma type_wf_put_ptr_in_heap_E:
subsection\<open>Preserving Types\<close>
lemma new_element_type_wf_preserved [simp]: "h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
new_element_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
new_element_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
split: prod.splits if_splits elim!: bind_returns_heap_E)[1]
apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter finite_set_in
@ -340,39 +340,39 @@ lemma new_element_is_l_new_element: "l_new_element type_wf"
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr tag_name_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
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]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
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'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs Let_def
put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs Let_def
put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
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'"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
@ -381,15 +381,15 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_
lemma put_M_pointers_preserved:
assumes "h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
using assms
apply(auto simp add: put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
using assms
apply(auto simp add: put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
elim!: bind_returns_heap_E2 dest!: get_heap_E)[1]
by (meson get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap is_OK_returns_result_I)
lemma element_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "element_ptr_kinds h = element_ptr_kinds h'"
using writes_small_big[OF assms]
@ -400,7 +400,7 @@ lemma element_ptr_kinds_preserved:
lemma element_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "element_ptr_kinds h = element_ptr_kinds h'"
by(simp add: element_ptr_kinds_def node_ptr_kinds_def preserved_def
by(simp add: element_ptr_kinds_def node_ptr_kinds_def preserved_def
object_ptr_kinds_preserved_small[OF assms])
lemma type_wf_preserved_small:
@ -410,19 +410,19 @@ lemma type_wf_preserved_small:
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2)] allI[OF assms(3), of id, simplified]
apply(auto simp add: type_wf_defs )[1]
apply(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
apply(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
split: option.splits,force)[1]
by(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
by(auto simp add: preserved_def get_M_defs element_ptr_kinds_small[OF assms(1)]
split: option.splits,force)
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
@ -434,8 +434,8 @@ proof -
qed
lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
node_ptr_kinds_def object_ptr_kinds_def is_node_ptr_kind_def
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
node_ptr_kinds_def object_ptr_kinds_def is_node_ptr_kind_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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)[1]
apply (metis (no_types, lifting) element_ptr_kinds_commutes finite_set_in fmdom_notD fmdom_notI
fmlookup_drop heap.sel node_ptr_kinds_commutes o_apply object_ptr_kinds_def)

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -37,7 +37,7 @@ begin
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog"
register_default_tvars "('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog"
global_interpretation l_ptr_kinds_M node_ptr_kinds defines node_ptr_kinds_M = a_ptr_kinds_M .
@ -46,14 +46,14 @@ lemmas node_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma node_ptr_kinds_M_eq:
assumes "|h \<turnstile> object_ptr_kinds_M|\<^sub>r = |h' \<turnstile> object_ptr_kinds_M|\<^sub>r"
shows "|h \<turnstile> node_ptr_kinds_M|\<^sub>r = |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using assms
using assms
by(auto simp add: node_ptr_kinds_M_defs object_ptr_kinds_M_defs node_ptr_kinds_def)
global_interpretation l_dummy defines get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = "l_get_M.a_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds"
apply(simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf l_get_M_def)
by (metis ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind_eq_None_conv get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
by (metis ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind_eq_None_conv get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
@ -72,15 +72,15 @@ end
global_interpretation l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales
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'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def
object_ptr_kinds_M_reads preserved_def)
by (smt object_ptr_kinds_preserved_small preserved_def unit_all_impI)
global_interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
rewrites "a_get_M = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e"
global_interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
rewrites "a_get_M = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e"
defines put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
@ -102,40 +102,40 @@ end
global_interpretation l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales
lemma get_M_Object_preserved1 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x)) \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Object_preserved1 [simp]:
"(\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x)) \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast node_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv
apply(cases "cast node_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def
bind_eq_Some_conv
split: option.splits)
lemma get_M_Object_preserved2 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Object_preserved2 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Object_preserved3 [simp]:
"h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
lemma get_M_Object_preserved3 [simp]:
"h \<turnstile> put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \<rightarrow>\<^sub>h h' \<Longrightarrow> (\<And>x. getter (cast (setter (\<lambda>_. v) x)) = getter (cast x))
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'"
apply(cases "cast node_ptr \<noteq> object_ptr")
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Object_preserved4 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
lemma get_M_Object_preserved4 [simp]:
"cast node_ptr \<noteq> object_ptr \<Longrightarrow> h \<turnstile> put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection\<open>Modified Heaps\<close>
lemma get_node_ptr_simp [simp]:
lemma get_node_ptr_simp [simp]:
"get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = cast node_ptr then cast obj else get node_ptr h)"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)
= node_ptr_kinds h |\<union>| (if is_node_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: node_ptr_kinds_def)
@ -155,7 +155,7 @@ lemma type_wf_put_ptr_not_in_heap_E:
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: ObjectMonad.type_wf_put_ptr_not_in_heap_E
by(auto simp add: type_wf_defs elim!: ObjectMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
@ -179,7 +179,7 @@ lemma node_ptr_kinds_small:
lemma node_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h'. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h'
\<longrightarrow> (\<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h')"
shows "node_ptr_kinds h = node_ptr_kinds h'"
using writes_small_big[OF assms]
@ -202,9 +202,9 @@ lemma type_wf_preserved_small:
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> \<forall>node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -37,7 +37,7 @@ begin
type_synonym ('object_ptr, 'Object, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'Object, 'result) dom_prog"
register_default_tvars "('object_ptr, 'Object, 'result) dom_prog"
global_interpretation l_ptr_kinds_M object_ptr_kinds defines object_ptr_kinds_M = a_ptr_kinds_M .
lemmas object_ptr_kinds_M_defs = a_ptr_kinds_M_def
@ -63,16 +63,16 @@ end
global_interpretation l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas type_wf
by (simp add: l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_def l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms)
lemma object_ptr_kinds_M_reads:
lemma object_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)}) object_ptr_kinds_M h h'"
apply(auto simp add: object_ptr_kinds_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf type_wf_defs reads_def
preserved_def get_M_defs
apply(auto simp add: object_ptr_kinds_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf type_wf_defs reads_def
preserved_def get_M_defs
split: option.splits)[1]
using a_type_wf_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf by blast+
global_interpretation l_put_M type_wf object_ptr_kinds get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
rewrites "a_get_M = get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"
global_interpretation l_put_M type_wf object_ptr_kinds get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
rewrites "a_get_M = get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"
defines put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)
@ -108,35 +108,35 @@ lemma check_in_heap_ptr_in_heap: "ptr |\<in>| object_ptr_kinds h \<longleftright
by(auto simp add: check_in_heap_def)
lemma check_in_heap_pure [simp]: "pure (check_in_heap ptr) h"
by(auto simp add: check_in_heap_def intro!: bind_pure_I)
lemma check_in_heap_is_OK [simp]:
lemma check_in_heap_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (check_in_heap ptr \<bind> f) = h \<turnstile> ok (f ())"
by(simp add: check_in_heap_def)
lemma check_in_heap_returns_result [simp]:
lemma check_in_heap_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> (check_in_heap ptr \<bind> f) \<rightarrow>\<^sub>r x = h \<turnstile> f () \<rightarrow>\<^sub>r x"
by(simp add: check_in_heap_def)
lemma check_in_heap_returns_heap [simp]:
lemma check_in_heap_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> h \<turnstile> (check_in_heap ptr \<bind> f) \<rightarrow>\<^sub>h h' = h \<turnstile> f () \<rightarrow>\<^sub>h h'"
by(simp add: check_in_heap_def)
lemma check_in_heap_reads:
lemma check_in_heap_reads:
"reads {preserved (get_M object_ptr nothing)} (check_in_heap object_ptr) h h'"
apply(simp add: check_in_heap_def reads_def preserved_def)
by (metis a_type_wf_def get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap is_OK_returns_result_E
by (metis a_type_wf_def get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ptr_in_heap is_OK_returns_result_E
is_OK_returns_result_I unit_all_impI)
subsection\<open>Invoke\<close>
fun invoke_rec :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list \<Rightarrow> (_) object_ptr \<Rightarrow> 'args
fun invoke_rec :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list \<Rightarrow> (_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog"
where
"invoke_rec ((P, f)#xs) ptr args = (if P ptr then f ptr args else invoke_rec xs ptr args)"
| "invoke_rec [] ptr args = error InvokeError"
definition invoke :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list
definition invoke :: "(((_) object_ptr \<Rightarrow> bool) \<times> ((_) object_ptr \<Rightarrow> 'args
\<Rightarrow> (_, 'result) dom_prog)) list
\<Rightarrow> (_) object_ptr \<Rightarrow> 'args \<Rightarrow> (_, 'result) dom_prog"
where
where
"invoke xs ptr args = do { check_in_heap ptr; invoke_rec xs ptr args}"
lemma invoke_split: "P (invoke ((Pred, f) # xs) ptr args) =
@ -156,16 +156,16 @@ lemma invoke_ptr_in_heap: "h \<turnstile> ok (invoke xs ptr args) \<Longrightarr
lemma invoke_pure [simp]: "pure (invoke [] ptr args) h"
by(auto simp add: invoke_def intro!: bind_pure_I)
lemma invoke_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
lemma invoke_is_OK [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> ok (invoke ((Pred, f) # xs) ptr args) = h \<turnstile> ok (f ptr args)"
by(simp add: invoke_def)
lemma invoke_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
lemma invoke_returns_result [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> (invoke ((Pred, f) # xs) ptr args) \<rightarrow>\<^sub>r x = h \<turnstile> f ptr args \<rightarrow>\<^sub>r x"
by(simp add: invoke_def)
lemma invoke_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
lemma invoke_returns_heap [simp]:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> Pred ptr
\<Longrightarrow> h \<turnstile> (invoke ((Pred, f) # xs) ptr args) \<rightarrow>\<^sub>h h' = h \<turnstile> f ptr args \<rightarrow>\<^sub>h h'"
by(simp add: invoke_def)
@ -182,7 +182,7 @@ lemma invoke_empty_reads [simp]: "\<forall>P \<in> S. reflp P \<and> transp P \<
subsection\<open>Modified Heaps\<close>
lemma get_object_ptr_simp [simp]:
lemma get_object_ptr_simp [simp]:
"get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = object_ptr then Some obj else get object_ptr h)"
by(auto simp add: get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: option.splits Option.bind_splits)
@ -220,17 +220,17 @@ lemma object_ptr_kinds_preserved_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
using assms
apply(auto simp add: object_ptr_kinds_def preserved_def get_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
apply(auto simp add: object_ptr_kinds_def preserved_def get_M_defs get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: option.splits)[1]
apply (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember.rep_eq
apply (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember.rep_eq
old.unit.exhaust option.case_eq_if return_returns_result)
by (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember.rep_eq
by (metis (mono_tags, lifting) domIff error_returns_result fmdom.rep_eq fmember.rep_eq
old.unit.exhaust option.case_eq_if return_returns_result)
lemma object_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' w object_ptr. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
assumes "\<And>h h' w object_ptr. w \<in> SW \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h'
\<Longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
shows "object_ptr_kinds h = object_ptr_kinds h'"
proof -

View File

@ -23,25 +23,25 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>CharacterData\<close>
text\<open>In this theory, we introduce the typed pointers for the class CharacterData.\<close>
text\<open>In this theory, we introduce the typed pointers for the class CharacterData.\<close>
theory CharacterDataPointer
imports
ElementPointer
begin
datatype 'character_data_ptr character_data_ptr = Ref (the_ref: ref) | Ext 'character_data_ptr
register_default_tvars "'character_data_ptr character_data_ptr"
register_default_tvars "'character_data_ptr character_data_ptr"
type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr
= "('character_data_ptr character_data_ptr + 'node_ptr, 'element_ptr) node_ptr"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr"
register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr) object_ptr
= "('object_ptr, 'character_data_ptr character_data_ptr + 'node_ptr, 'element_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr) object_ptr"
definition cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) character_data_ptr \<Rightarrow> (_) node_ptr"
where
@ -53,7 +53,7 @@ abbreviation cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>
definition cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) node_ptr \<Rightarrow> (_) character_data_ptr option"
where
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = (case node_ptr of
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = (case node_ptr of
node_ptr.Ext (Inr (Inl character_data_ptr)) \<Rightarrow> Some character_data_ptr
| _ \<Rightarrow> None)"
@ -63,29 +63,29 @@ abbreviation cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>
Some node_ptr \<Rightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
| None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
adhoc_overloading cast cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r
consts is_character_data_ptr_kind :: 'a
definition is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) node_ptr \<Rightarrow> bool"
where
"is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr
"is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr
of Some _ \<Rightarrow> True | _ \<Rightarrow> False)"
abbreviation is_character_data_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
where
"is_character_data_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
"is_character_data_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some node_ptr \<Rightarrow> is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
| None \<Rightarrow> False)"
adhoc_overloading is_character_data_ptr_kind is_character_data_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
adhoc_overloading is_character_data_ptr_kind is_character_data_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r
lemmas is_character_data_ptr_kind_def = is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
consts is_character_data_ptr :: 'a
definition is_character_data_ptr\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) character_data_ptr \<Rightarrow> bool"
where
"is_character_data_ptr\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr
"is_character_data_ptr\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr
of character_data_ptr.Ref _ \<Rightarrow> True | _ \<Rightarrow> False)"
abbreviation is_character_data_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) node_ptr \<Rightarrow> bool"
@ -105,7 +105,7 @@ adhoc_overloading is_character_data_ptr
lemmas is_character_data_ptr_def = is_character_data_ptr\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
consts is_character_data_ptr_ext :: 'a
abbreviation
abbreviation
"is_character_data_ptr_ext\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> \<not> is_character_data_ptr\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
abbreviation "is_character_data_ptr_ext\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
@ -121,17 +121,17 @@ adhoc_overloading is_character_data_ptr_ext
instantiation character_data_ptr :: (linorder) linorder
begin
definition
definition
less_eq_character_data_ptr :: "(_::linorder) character_data_ptr \<Rightarrow> (_) character_data_ptr \<Rightarrow> bool"
where
where
"less_eq_character_data_ptr x y \<equiv> (case x of Ext i \<Rightarrow> (case y of Ext j \<Rightarrow> i \<le> j | Ref _ \<Rightarrow> False)
| Ref i \<Rightarrow> (case y of Ext _ \<Rightarrow> True | Ref j \<Rightarrow> i \<le> j))"
definition
definition
less_character_data_ptr :: "(_::linorder) character_data_ptr \<Rightarrow> (_) character_data_ptr \<Rightarrow> bool"
where "less_character_data_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance
apply(standard)
by(auto simp add: less_eq_character_data_ptr_def less_character_data_ptr_def
instance
apply(standard)
by(auto simp add: less_eq_character_data_ptr_def less_character_data_ptr_def
split: character_data_ptr.splits)
end
@ -141,20 +141,20 @@ lemma is_character_data_ptr_ref [simp]: "is_character_data_ptr (character_data_p
lemma cast_element_ptr_not_character_data_ptr [simp]:
"(cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr \<noteq> cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r character_data_ptr)"
"(cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r character_data_ptr \<noteq> cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr)"
unfolding cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
unfolding cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto)
lemma is_character_data_ptr_kind_not_element_ptr [simp]:
lemma is_character_data_ptr_kind_not_element_ptr [simp]:
"\<not> is_character_data_ptr_kind (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr)"
unfolding is_character_data_ptr_kind_def cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by auto
lemma is_element_ptr_kind_not_character_data_ptr [simp]:
lemma is_element_ptr_kind_not_character_data_ptr [simp]:
"\<not> is_element_ptr_kind (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r character_data_ptr)"
using is_element_ptr_kind_obtains by fastforce
lemma is_character_data_ptr_kind\<^sub>_cast [simp]:
lemma is_character_data_ptr_kind\<^sub>_cast [simp]:
"is_character_data_ptr_kind (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r character_data_ptr)"
by (simp add: cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by (simp add: cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma character_data_ptr_casts_commute [simp]:
@ -171,14 +171,14 @@ lemma character_data_ptr_casts_commute3 [simp]:
assumes "is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
shows "cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (the (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) = node_ptr"
using assms
by(auto simp add: is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto simp add: is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: node_ptr.splits sum.splits)
lemma is_character_data_ptr_kind_obtains:
assumes "is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
obtains character_data_ptr where "cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r character_data_ptr = node_ptr"
by (metis assms is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def case_optionE
by (metis assms is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def case_optionE
character_data_ptr_casts_commute)
lemma is_character_data_ptr_kind_none:
@ -188,11 +188,11 @@ lemma is_character_data_ptr_kind_none:
unfolding is_character_data_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: node_ptr.splits sum.splits)
lemma cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
lemma cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
"cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r x = cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r y \<longleftrightarrow> x = y"
by(simp add: cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
lemma cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r (node_ptr.Ext (Inr (Inr node_ext_ptr))) = None"
by(simp add: cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)

View File

@ -23,22 +23,22 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Document\<close>
text\<open>In this theory, we introduce the typed pointers for the class Document.\<close>
text\<open>In this theory, we introduce the typed pointers for the class Document.\<close>
theory DocumentPointer
imports
CharacterDataPointer
begin
datatype 'document_ptr document_ptr = Ref (the_ref: ref) | Ext 'document_ptr
register_default_tvars "'document_ptr document_ptr"
register_default_tvars "'document_ptr document_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr) object_ptr
= "('document_ptr document_ptr + 'object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr) object_ptr"
definition cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_)document_ptr \<Rightarrow> (_) object_ptr"
where
@ -46,8 +46,8 @@ definition cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\
definition cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> (_) document_ptr option"
where
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of
object_ptr.Ext (Inr (Inl document_ptr)) \<Rightarrow> Some document_ptr
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of
object_ptr.Ext (Inr (Inl document_ptr)) \<Rightarrow> Some document_ptr
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
@ -55,7 +55,7 @@ adhoc_overloading cast cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^su
definition is_document_ptr_kind :: "(_) object_ptr \<Rightarrow> bool"
where
"is_document_ptr_kind ptr = (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
"is_document_ptr_kind ptr = (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
Some _ \<Rightarrow> True | None \<Rightarrow> False)"
consts is_document_ptr :: 'a
@ -86,8 +86,8 @@ definition less_eq_document_ptr :: "(_::linorder) document_ptr \<Rightarrow> (_)
| Ref i \<Rightarrow> (case y of Ext _ \<Rightarrow> True | Ref j \<Rightarrow> i \<le> j))"
definition less_document_ptr :: "(_::linorder) document_ptr \<Rightarrow> (_) document_ptr \<Rightarrow> bool"
where "less_document_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance
apply(standard)
instance
apply(standard)
by(auto simp add: less_eq_document_ptr_def less_document_ptr_def split: document_ptr.splits)
end
@ -97,17 +97,17 @@ lemma is_document_ptr_ref [simp]: "is_document_ptr (document_ptr.Ref n)"
lemma cast_document_ptr_not_node_ptr [simp]:
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr \<noteq> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \<noteq> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"
unfolding cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
unfolding cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by auto
lemma document_ptr_no_node_ptr_cast [simp]:
lemma document_ptr_no_node_ptr_cast [simp]:
"\<not> is_document_ptr_kind (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)"
by(simp add: cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def)
lemma node_ptr_no_document_ptr_cast [simp]:
lemma node_ptr_no_document_ptr_cast [simp]:
"\<not> is_node_ptr_kind (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)"
using is_node_ptr_kind_obtains by fastforce
lemma document_ptr_document_ptr_cast [simp]:
lemma document_ptr_document_ptr_cast [simp]:
"is_document_ptr_kind (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)"
by (simp add: cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_document_ptr_kind_def)
@ -116,7 +116,7 @@ lemma document_ptr_casts_commute [simp]:
unfolding cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: object_ptr.splits sum.splits)
lemma document_ptr_casts_commute2 [simp]:
lemma document_ptr_casts_commute2 [simp]:
"(cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) = Some document_ptr)"
by simp
@ -140,11 +140,11 @@ lemma is_document_ptr_kind_none:
unfolding is_document_ptr_kind_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by (auto split: object_ptr.splits sum.splits)
lemma cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
lemma cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r x = cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r y \<longleftrightarrow> x = y"
by(simp add: cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
lemma cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (object_ptr.Ext (Inr (Inr (Inr object_ext_ptr)))) = None"
by(simp add: cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)

View File

@ -23,26 +23,26 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Element\<close>
text\<open>In this theory, we introduce the typed pointers for the class Element.\<close>
text\<open>In this theory, we introduce the typed pointers for the class Element.\<close>
theory ElementPointer
imports
NodePointer
begin
datatype 'element_ptr element_ptr = Ref (the_ref: ref) | Ext 'element_ptr
register_default_tvars "'element_ptr element_ptr"
register_default_tvars "'element_ptr element_ptr"
type_synonym ('node_ptr, 'element_ptr) node_ptr
= "('element_ptr element_ptr + 'node_ptr) node_ptr"
register_default_tvars "('node_ptr, 'element_ptr) node_ptr"
register_default_tvars "('node_ptr, 'element_ptr) node_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr) object_ptr
= "('object_ptr, 'element_ptr element_ptr + 'node_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr) object_ptr"
definition cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) element_ptr \<Rightarrow> (_) element_ptr"
@ -59,16 +59,16 @@ abbreviation cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>
definition cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) node_ptr \<Rightarrow> (_) element_ptr option"
where
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = (case node_ptr of node_ptr.Ext (Inl element_ptr)
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = (case node_ptr of node_ptr.Ext (Inl element_ptr)
\<Rightarrow> Some element_ptr | _ \<Rightarrow> None)"
abbreviation cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> (_) element_ptr option"
where
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
Some node_ptr \<Rightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
Some node_ptr \<Rightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
| None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
adhoc_overloading cast cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
consts is_element_ptr_kind :: 'a
@ -78,8 +78,8 @@ definition is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<
abbreviation is_element_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
where
"is_element_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some node_ptr \<Rightarrow> is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
"is_element_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some node_ptr \<Rightarrow> is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
| None \<Rightarrow> False)"
adhoc_overloading is_element_ptr_kind is_element_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r
@ -92,14 +92,14 @@ definition is_element_ptr\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>
abbreviation is_element_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) node_ptr \<Rightarrow> bool"
where
"is_element_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some element_ptr \<Rightarrow> is_element_ptr\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr
"is_element_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some element_ptr \<Rightarrow> is_element_ptr\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr
| _ \<Rightarrow> False)"
abbreviation is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
where
"is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some node_ptr \<Rightarrow> is_element_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
"is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some node_ptr \<Rightarrow> is_element_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr
| None \<Rightarrow> False)"
adhoc_overloading is_element_ptr is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_element_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_element_ptr\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
@ -116,16 +116,16 @@ adhoc_overloading is_element_ptr_ext is_element_ptr_ext\<^sub>o\<^sub>b\<^sub>j\
instantiation element_ptr :: (linorder) linorder
begin
definition
definition
less_eq_element_ptr :: "(_::linorder) element_ptr \<Rightarrow> (_)element_ptr \<Rightarrow> bool"
where
where
"less_eq_element_ptr x y \<equiv> (case x of Ext i \<Rightarrow> (case y of Ext j \<Rightarrow> i \<le> j | Ref _ \<Rightarrow> False)
| Ref i \<Rightarrow> (case y of Ext _ \<Rightarrow> True | Ref j \<Rightarrow> i \<le> j))"
definition
definition
less_element_ptr :: "(_::linorder) element_ptr \<Rightarrow> (_) element_ptr \<Rightarrow> bool"
where "less_element_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance
apply(standard)
instance
apply(standard)
by(auto simp add: less_eq_element_ptr_def less_element_ptr_def split: element_ptr.splits)
end
@ -137,7 +137,7 @@ lemma element_ptr_casts_commute [simp]:
unfolding cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: node_ptr.splits sum.splits)
lemma element_ptr_casts_commute2 [simp]:
lemma element_ptr_casts_commute2 [simp]:
"(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr) = Some element_ptr)"
by simp
@ -145,7 +145,7 @@ lemma element_ptr_casts_commute3 [simp]:
assumes "is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr"
shows "cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (the (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) = node_ptr"
using assms
by(auto simp add: is_element_ptr_kind_def cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto simp add: is_element_ptr_kind_def cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: node_ptr.splits sum.splits)
lemma is_element_ptr_kind_obtains:
@ -160,15 +160,15 @@ lemma is_element_ptr_kind_none:
unfolding is_element_ptr_kind_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: node_ptr.splits sum.splits)
lemma is_element_ptr_kind_cast [simp]:
lemma is_element_ptr_kind_cast [simp]:
"is_element_ptr_kind (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr)"
by (metis element_ptr_casts_commute is_element_ptr_kind_none option.distinct(1))
lemma cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
lemma cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
"cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r x = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r y \<longleftrightarrow> x = y"
by(simp add: cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
lemma cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (node_ptr.Ext (Inr (Inr node_ext_ptr))) = None"
by(simp add: cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)

View File

@ -23,22 +23,22 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Node\<close>
text\<open>In this theory, we introduce the typed pointers for the class Node.\<close>
text\<open>In this theory, we introduce the typed pointers for the class Node.\<close>
theory NodePointer
imports
ObjectPointer
begin
datatype 'node_ptr node_ptr = Ext 'node_ptr
register_default_tvars "'node_ptr node_ptr"
register_default_tvars "'node_ptr node_ptr"
type_synonym ('object_ptr, 'node_ptr) object_ptr = "('node_ptr node_ptr + 'object_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr) object_ptr"
definition cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) node_ptr \<Rightarrow> (_) object_ptr"
where
@ -46,7 +46,7 @@ definition cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\
definition cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> (_) node_ptr option"
where
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r object_ptr = (case object_ptr of object_ptr.Ext (Inl node_ptr)
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r object_ptr = (case object_ptr of object_ptr.Ext (Inl node_ptr)
\<Rightarrow> Some node_ptr | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r
@ -61,17 +61,17 @@ definition less_eq_node_ptr :: "(_::linorder) node_ptr \<Rightarrow> (_) node_pt
where "less_eq_node_ptr x y \<equiv> (case x of Ext i \<Rightarrow> (case y of Ext j \<Rightarrow> i \<le> j))"
definition less_node_ptr :: "(_::linorder) node_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> bool"
where "less_node_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance
instance
apply(standard)
by(auto simp add: less_eq_node_ptr_def less_node_ptr_def split: node_ptr.splits)
end
lemma node_ptr_casts_commute [simp]:
lemma node_ptr_casts_commute [simp]:
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = Some node_ptr \<longleftrightarrow> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr = ptr"
unfolding cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: object_ptr.splits sum.splits)
lemma node_ptr_casts_commute2 [simp]:
lemma node_ptr_casts_commute2 [simp]:
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr) = Some node_ptr"
by simp
@ -79,7 +79,7 @@ lemma node_ptr_casts_commute3 [simp]:
assumes "is_node_ptr_kind ptr"
shows "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (the (cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)) = ptr"
using assms
by(auto simp add: is_node_ptr_kind_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto simp add: is_node_ptr_kind_def cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: object_ptr.splits sum.splits)
lemma is_node_ptr_kind_obtains:
@ -97,15 +97,15 @@ lemma is_node_ptr_kind_none:
lemma is_node_ptr_kind_cast [simp]: "is_node_ptr_kind (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)"
unfolding is_node_ptr_kind_def by simp
lemma cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
lemma cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r x = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r y \<longleftrightarrow> x = y"
by(simp add: cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
lemma cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (object_ptr.Ext (Inr (Inr (Inr object_ext_ptr)))) = None"
by(simp add: cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma node_ptr_inclusion [simp]:
lemma node_ptr_inclusion [simp]:
"cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` node_ptrs \<longleftrightarrow> node_ptr \<in> node_ptrs"
by auto
end

View File

@ -23,12 +23,12 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Object\<close>
text\<open>In this theory, we introduce the typed pointer for the class Object. This class is the
text\<open>In this theory, we introduce the typed pointer for the class Object. This class is the
common superclass of our class model.\<close>
theory ObjectPointer
imports
@ -36,7 +36,7 @@ theory ObjectPointer
begin
datatype 'object_ptr object_ptr = Ext 'object_ptr
register_default_tvars "'object_ptr object_ptr"
register_default_tvars "'object_ptr object_ptr"
instantiation object_ptr :: (linorder) linorder
begin
@ -44,7 +44,7 @@ definition less_eq_object_ptr :: "'object_ptr::linorder object_ptr \<Rightarrow>
where "less_eq_object_ptr x y \<equiv> (case x of Ext i \<Rightarrow> (case y of Ext j \<Rightarrow> i \<le> j))"
definition less_object_ptr :: "'object_ptr::linorder object_ptr \<Rightarrow> 'object_ptr object_ptr \<Rightarrow> bool"
where "less_object_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance by(standard, auto simp add: less_eq_object_ptr_def less_object_ptr_def
instance by(standard, auto simp add: less_eq_object_ptr_def less_object_ptr_def
split: object_ptr.splits)
end

View File

@ -23,16 +23,16 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>References\<close>
text\<open>
This theory, we introduce a generic reference. All our typed pointers include such
a reference, which allows us to distinguish pointers of the same type, but also to
This theory, we introduce a generic reference. All our typed pointers include such
a reference, which allows us to distinguish pointers of the same type, but also to
iterate over all pointers in a set.\<close>
theory
theory
Ref
imports
"HOL-Library.Adhoc_Overloading"

View File

@ -23,15 +23,15 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>The Heap Error Monad\<close>
text \<open>In this theory, we define a heap and error monad for modeling exceptions.
This allows us to define composite methods similar to stateful programming in Haskell,
text \<open>In this theory, we define a heap and error monad for modeling exceptions.
This allows us to define composite methods similar to stateful programming in Haskell,
but also to stay close to the official DOM specification.\<close>
theory
theory
Heap_Error_Monad
imports
Hiding_Type_Variables
@ -45,22 +45,22 @@ register_default_tvars "('heap, 'e, 'result) prog" (print, parse)
subsection \<open>Basic Functions\<close>
definition
definition
bind :: "(_, 'result) prog \<Rightarrow> ('result \<Rightarrow> (_, 'result2) prog) \<Rightarrow> (_, 'result2) prog"
where
"bind f g = Prog (\<lambda>h. (case (the_prog f) h of Inr (x, h') \<Rightarrow> (the_prog (g x)) h'
"bind f g = Prog (\<lambda>h. (case (the_prog f) h of Inr (x, h') \<Rightarrow> (the_prog (g x)) h'
| Inl exception \<Rightarrow> Inl exception))"
adhoc_overloading Monad_Syntax.bind bind
definition
execute :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> ('e + 'result \<times> 'heap)"
definition
execute :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> ('e + 'result \<times> 'heap)"
("((_)/ \<turnstile> (_))" [51, 52] 55)
where
"execute h p = (the_prog p) h"
definition
returns_result :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'result \<Rightarrow> bool"
definition
returns_result :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'result \<Rightarrow> bool"
("((_)/ \<turnstile> (_)/ \<rightarrow>\<^sub>r (_))" [60, 35, 61] 65)
where
"returns_result h p r \<longleftrightarrow> (case h \<turnstile> p of Inr (r', _) \<Rightarrow> r = r' | Inl _ \<Rightarrow> False)"
@ -73,8 +73,8 @@ fun select_result ("|(_)|\<^sub>r")
lemma returns_result_eq [elim]: "h \<turnstile> f \<rightarrow>\<^sub>r y \<Longrightarrow> h \<turnstile> f \<rightarrow>\<^sub>r y' \<Longrightarrow> y = y'"
by(auto simp add: returns_result_def split: sum.splits)
definition
returns_heap :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'heap \<Rightarrow> bool"
definition
returns_heap :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'heap \<Rightarrow> bool"
("((_)/ \<turnstile> (_)/ \<rightarrow>\<^sub>h (_))" [60, 35, 61] 65)
where
"returns_heap h p h' \<longleftrightarrow> (case h \<turnstile> p of Inr (_ , h'') \<Rightarrow> h' = h'' | Inl _ \<Rightarrow> False)"
@ -87,13 +87,14 @@ fun select_heap ("|(_)|\<^sub>h")
lemma returns_heap_eq [elim]: "h \<turnstile> f \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> f \<rightarrow>\<^sub>h h'' \<Longrightarrow> h' = h''"
by(auto simp add: returns_heap_def split: sum.splits)
definition
returns_result_heap :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'result \<Rightarrow> 'heap \<Rightarrow> bool"
definition
returns_result_heap :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'result \<Rightarrow> 'heap \<Rightarrow> bool"
("((_)/ \<turnstile> (_)/ \<rightarrow>\<^sub>r (_) \<rightarrow>\<^sub>h (_))" [60, 35, 61, 62] 65)
where
"returns_result_heap h p r h' \<longleftrightarrow> h \<turnstile> p \<rightarrow>\<^sub>r r \<and> h \<turnstile> p \<rightarrow>\<^sub>h h'"
lemma return_result_heap_code [code]: "returns_result_heap h p r h' \<longleftrightarrow> (case h \<turnstile> p of Inr (r', h'') \<Rightarrow> r = r' \<and> h' = h'' | Inl _ \<Rightarrow> False)"
lemma return_result_heap_code [code]:
"returns_result_heap h p r h' \<longleftrightarrow> (case h \<turnstile> p of Inr (r', h'') \<Rightarrow> r = r' \<and> h' = h'' | Inl _ \<Rightarrow> False)"
by(auto simp add: returns_result_heap_def returns_result_def returns_heap_def split: sum.splits)
fun select_result_heap ("|(_)|\<^sub>r\<^sub>h")
@ -101,8 +102,8 @@ fun select_result_heap ("|(_)|\<^sub>r\<^sub>h")
"select_result_heap (Inr (r, h)) = (r, h)"
| "select_result_heap (Inl _) = undefined"
definition
returns_error :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'e \<Rightarrow> bool"
definition
returns_error :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'e \<Rightarrow> bool"
("((_)/ \<turnstile> (_)/ \<rightarrow>\<^sub>e (_))" [60, 35, 61] 65)
where
"returns_error h p e = (case h \<turnstile> p of Inr _ \<Rightarrow> False | Inl e' \<Rightarrow> e = e')"
@ -147,13 +148,13 @@ lemma returns_result_select_result [simp]:
by (simp add: select_result_I)
lemma select_result_E:
assumes "P |h \<turnstile> f|\<^sub>r" and "h \<turnstile> ok f"
assumes "P |h \<turnstile> f|\<^sub>r" and "h \<turnstile> ok f"
obtains x where "h \<turnstile> f \<rightarrow>\<^sub>r x" and "P x"
using assms
by(auto simp add: is_OK_def returns_result_def split: sum.splits)
lemma select_result_eq: "(\<And>x .h \<turnstile> f \<rightarrow>\<^sub>r x = h' \<turnstile> f \<rightarrow>\<^sub>r x) \<Longrightarrow> |h \<turnstile> f|\<^sub>r = |h' \<turnstile> f|\<^sub>r"
by (metis (no_types, lifting) is_OK_def old.sum.simps(6) select_result.elims
by (metis (no_types, lifting) is_OK_def old.sum.simps(6) select_result.elims
select_result_I select_result_I2)
definition error :: "'e \<Rightarrow> ('heap, 'e, 'result) prog"
@ -252,7 +253,7 @@ lemma pure_returns_heap_eq:
"h \<turnstile> f \<rightarrow>\<^sub>h h' \<Longrightarrow> pure f h \<Longrightarrow> h = h'"
by (meson pure_def is_OK_returns_heap_I returns_heap_eq)
lemma pure_eq_iff:
lemma pure_eq_iff:
"(\<forall>h' x. h \<turnstile> f \<rightarrow>\<^sub>r x \<longrightarrow> h \<turnstile> f \<rightarrow>\<^sub>h h' \<longrightarrow> h = h') \<longleftrightarrow> pure f h"
by(auto simp add: pure_def)
@ -265,7 +266,7 @@ lemma bind_assoc [simp]:
lemma bind_returns_result_E:
assumes "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>r y"
obtains x h' where "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> g x \<rightarrow>\<^sub>r y"
using assms by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
using assms by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
split: sum.splits)
lemma bind_returns_result_E2:
@ -279,14 +280,14 @@ lemma bind_returns_result_E3:
using assms returns_result_eq bind_returns_result_E2 by metis
lemma bind_returns_result_E4:
assumes "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>r y" and "h \<turnstile> f \<rightarrow>\<^sub>r x"
assumes "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>r y" and "h \<turnstile> f \<rightarrow>\<^sub>r x"
obtains h' where "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> g x \<rightarrow>\<^sub>r y"
using assms returns_result_eq bind_returns_result_E by metis
lemma bind_returns_heap_E:
assumes "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>h h''"
obtains x h' where "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> g x \<rightarrow>\<^sub>h h''"
using assms by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
using assms by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
split: sum.splits)
lemma bind_returns_heap_E2 [elim]:
@ -295,7 +296,7 @@ lemma bind_returns_heap_E2 [elim]:
using assms pure_returns_heap_eq by (fastforce elim: bind_returns_heap_E)
lemma bind_returns_heap_E3 [elim]:
assumes "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>h h'" and "h \<turnstile> f \<rightarrow>\<^sub>r x" and "pure f h"
assumes "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>h h'" and "h \<turnstile> f \<rightarrow>\<^sub>r x" and "pure f h"
shows "h \<turnstile> g x \<rightarrow>\<^sub>h h'"
using assms pure_returns_heap_eq returns_result_eq by (fastforce elim: bind_returns_heap_E)
@ -315,7 +316,7 @@ lemma bind_returns_error_I3:
assumes "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> g x \<rightarrow>\<^sub>e e"
shows "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>e e"
using assms
by(auto simp add: returns_error_def bind_def execute_def returns_heap_def returns_result_def
by(auto simp add: returns_error_def bind_def execute_def returns_heap_def returns_result_def
split: sum.splits)
lemma bind_returns_error_I2 [intro]:
@ -327,22 +328,22 @@ lemma bind_returns_error_I2 [intro]:
lemma bind_is_OK_E [elim]:
assumes "h \<turnstile> ok (f \<bind> g)"
obtains x h' where "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> ok (g x)"
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def is_OK_def execute_def
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def is_OK_def execute_def
split: sum.splits)
lemma bind_is_OK_E2:
assumes "h \<turnstile> ok (f \<bind> g)" and "h \<turnstile> f \<rightarrow>\<^sub>r x"
obtains h' where "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> ok (g x)"
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def is_OK_def execute_def
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def is_OK_def execute_def
split: sum.splits)
lemma bind_returns_result_I [intro]:
assumes "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> g x \<rightarrow>\<^sub>r y"
shows "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>r y"
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
split: sum.splits)
lemma bind_pure_returns_result_I [intro]:
@ -359,8 +360,8 @@ lemma bind_pure_returns_result_I2 [intro]:
lemma bind_returns_heap_I [intro]:
assumes "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> g x \<rightarrow>\<^sub>h h''"
shows "h \<turnstile> f \<bind> g \<rightarrow>\<^sub>h h''"
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
using assms
by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
split: sum.splits)
lemma bind_returns_heap_I2 [intro]:
@ -372,13 +373,13 @@ lemma bind_returns_heap_I2 [intro]:
lemma bind_is_OK_I [intro]:
assumes "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'" and "h' \<turnstile> ok (g x)"
shows "h \<turnstile> ok (f \<bind> g)"
by (meson assms(1) assms(2) assms(3) bind_returns_heap_I is_OK_returns_heap_E
by (meson assms(1) assms(2) assms(3) bind_returns_heap_I is_OK_returns_heap_E
is_OK_returns_heap_I)
lemma bind_is_OK_I2 [intro]:
assumes "h \<turnstile> ok f" and "\<And>x h'. h \<turnstile> f \<rightarrow>\<^sub>r x \<Longrightarrow> h \<turnstile> f \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> ok (g x)"
shows "h \<turnstile> ok (f \<bind> g)"
using assms by blast
using assms by blast
lemma bind_is_OK_pure_I [intro]:
assumes "pure f h" and "h \<turnstile> ok f" and "\<And>x. h \<turnstile> f \<rightarrow>\<^sub>r x \<Longrightarrow> h \<turnstile> ok (g x)"
@ -394,15 +395,15 @@ lemma bind_pure_I:
lemma pure_pure:
assumes "h \<turnstile> ok f" and "pure f h"
shows "h \<turnstile> f \<rightarrow>\<^sub>h h"
using assms returns_heap_eq
using assms returns_heap_eq
unfolding pure_def
by auto
lemma bind_returns_error_eq:
lemma bind_returns_error_eq:
assumes "h \<turnstile> f \<rightarrow>\<^sub>e e"
and "h \<turnstile> g \<rightarrow>\<^sub>e e"
shows "h \<turnstile> f = h \<turnstile> g"
using assms
using assms
by(auto simp add: returns_error_def split: sum.splits)
subsection \<open>Map\<close>
@ -416,7 +417,7 @@ fun map_M :: "('x \<Rightarrow> ('heap, 'e, 'result) prog) \<Rightarrow> 'x list
return (y # ys)
}"
lemma map_M_ok_I [intro]:
lemma map_M_ok_I [intro]:
"(\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (f x)) \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (f x) h) \<Longrightarrow> h \<turnstile> ok (map_M f xs)"
apply(induct xs)
by (simp_all add: bind_is_OK_I2 bind_is_OK_pure_I)
@ -452,38 +453,16 @@ fun forall_M :: "('y \<Rightarrow> ('heap, 'e, 'result) prog) \<Rightarrow> 'y l
P x;
forall_M P xs
}"
(*
lemma forall_M_elim:
assumes "h \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True" and "\<And>x h. x \<in> set xs \<Longrightarrow> pure (P x) h"
shows "\<forall>x \<in> set xs. h \<turnstile> P x \<rightarrow>\<^sub>r True"
apply(insert assms, induct xs)
apply(simp)
apply(auto elim!: bind_returns_result_E)[1]
by (metis (full_types) pure_returns_heap_eq) *)
lemma pure_forall_M_I: "(\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> pure (forall_M P xs) h"
apply(induct xs)
by(auto intro!: bind_pure_I)
(*
lemma forall_M_pure_I:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> P x \<rightarrow>\<^sub>r True" and "\<And>x h. x \<in> set xs \<Longrightarrow> pure (P x)h"
shows "h \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True"
apply(insert assms, induct xs)
apply(simp)
by(fastforce)
lemma forall_M_pure_eq:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> P x \<rightarrow>\<^sub>r True \<longleftrightarrow> h' \<turnstile> P x \<rightarrow>\<^sub>r True"
and "\<And>x h. x \<in> set xs \<Longrightarrow> pure (P x) h"
shows "(h \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True) \<longleftrightarrow> h' \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True"
using assms
by(auto intro!: forall_M_pure_I dest!: forall_M_elim) *)
subsection \<open>Fold\<close>
fun fold_M :: "('result \<Rightarrow> 'y \<Rightarrow> ('heap, 'e, 'result) prog) \<Rightarrow> 'result \<Rightarrow> 'y list
\<Rightarrow> ('heap, 'e, 'result) prog"
where
where
"fold_M f d [] = return d" |
"fold_M f d (x # xs) = do { y \<leftarrow> f d x; fold_M f y xs }"
@ -503,10 +482,11 @@ fun filter_M :: "('x \<Rightarrow> ('heap, 'e, bool) prog) \<Rightarrow> 'x list
}"
lemma filter_M_pure_I [intro]: "(\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> pure (filter_M P xs)h"
apply(induct xs)
apply(induct xs)
by(auto intro!: bind_pure_I)
lemma filter_M_is_OK_I [intro]: "(\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (P x)) \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> h \<turnstile> ok (filter_M P xs)"
lemma filter_M_is_OK_I [intro]:
"(\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (P x)) \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> h \<turnstile> ok (filter_M P xs)"
apply(induct xs)
apply(simp)
by(auto intro!: bind_is_OK_pure_I)
@ -518,7 +498,8 @@ lemma filter_M_not_more_elements:
by(auto elim!: bind_returns_result_E2 split: if_splits intro!: set_ConsD)
lemma filter_M_in_result_if_ok:
assumes "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys" and "\<And>h x. x \<in> set xs \<Longrightarrow> pure (P x) h" and "x \<in> set xs" and "h \<turnstile> P x \<rightarrow>\<^sub>r True"
assumes "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys" and "\<And>h x. x \<in> set xs \<Longrightarrow> pure (P x) h" and "x \<in> set xs" and
"h \<turnstile> P x \<rightarrow>\<^sub>r True"
shows "x \<in> set ys"
apply(insert assms, induct xs arbitrary: ys)
apply(simp)
@ -539,13 +520,13 @@ lemma filter_M_empty_I:
apply(induct xs)
by(auto intro!: bind_pure_returns_result_I)
lemma filter_M_subset_2: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys \<Longrightarrow> h' \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys'
\<Longrightarrow> (\<And>x. pure (P x) h) \<Longrightarrow> (\<And>x. pure (P x) h')
\<Longrightarrow> (\<forall>b. \<forall>x \<in> set xs. h \<turnstile> P x \<rightarrow>\<^sub>r True \<longrightarrow> h' \<turnstile> P x \<rightarrow>\<^sub>r b \<longrightarrow> b)
lemma filter_M_subset_2: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys \<Longrightarrow> h' \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys'
\<Longrightarrow> (\<And>x. pure (P x) h) \<Longrightarrow> (\<And>x. pure (P x) h')
\<Longrightarrow> (\<forall>b. \<forall>x \<in> set xs. h \<turnstile> P x \<rightarrow>\<^sub>r True \<longrightarrow> h' \<turnstile> P x \<rightarrow>\<^sub>r b \<longrightarrow> b)
\<Longrightarrow> set ys \<subseteq> set ys'"
proof -
assume 1: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys" and 2: "h' \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys'"
and 3: "(\<And>x. pure (P x) h)" and "(\<And>x. pure (P x) h')"
assume 1: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys" and 2: "h' \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys'"
and 3: "(\<And>x. pure (P x) h)" and "(\<And>x. pure (P x) h')"
and 4: "\<forall>b. \<forall>x\<in>set xs. h \<turnstile> P x \<rightarrow>\<^sub>r True \<longrightarrow> h' \<turnstile> P x \<rightarrow>\<^sub>r b \<longrightarrow> b"
have h1: "\<forall>x \<in> set xs. h' \<turnstile> ok (P x)"
using 2 3 \<open>(\<And>x. pure (P x) h')\<close>
@ -583,17 +564,17 @@ lemma filter_M_distinct: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys
apply(auto elim!: bind_returns_result_E)[1]
by fastforce
lemma filter_M_filter: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h)
lemma filter_M_filter: "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h)
\<Longrightarrow> (\<forall>x \<in> set xs. h \<turnstile> ok P x) \<and> ys = filter (\<lambda>x. |h \<turnstile> P x|\<^sub>r) xs"
apply(induct xs arbitrary: ys)
by(auto elim!: bind_returns_result_E2)
lemma filter_M_filter2: "(\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h \<and> h \<turnstile> ok P x)
lemma filter_M_filter2: "(\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h \<and> h \<turnstile> ok P x)
\<Longrightarrow> filter (\<lambda>x. |h \<turnstile> P x|\<^sub>r) xs = ys \<Longrightarrow> h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys"
apply(induct xs arbitrary: ys)
by(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)
lemma filter_ex1: "\<exists>!x \<in> set xs. P x \<Longrightarrow> P x \<Longrightarrow> x \<in> set xs \<Longrightarrow> distinct xs
lemma filter_ex1: "\<exists>!x \<in> set xs. P x \<Longrightarrow> P x \<Longrightarrow> x \<in> set xs \<Longrightarrow> distinct xs
\<Longrightarrow> filter P xs = [x]"
apply(auto)[1]
apply(induct xs)
@ -612,16 +593,16 @@ lemma filter_M_ex1:
proof -
have *: "\<exists>!x \<in> set xs. |h \<turnstile> P x|\<^sub>r"
apply(insert assms(1) assms(3) assms(4))
apply(drule filter_M_filter)
apply(drule filter_M_filter)
apply(simp)
apply(auto simp add: select_result_I2)[1]
by (metis (full_types) is_OK_returns_result_E select_result_I2)
then show ?thesis
apply(insert assms(1) assms(4))
apply(drule filter_M_filter)
apply(auto)[1]
by (metis * assms(2) assms(5) assms(6) distinct_filter
distinct_length_2_or_more filter_empty_conv filter_set list.exhaust
apply(auto)[1]
by (metis * assms(2) assms(5) assms(6) distinct_filter
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)
qed
@ -631,7 +612,7 @@ lemma filter_M_eq:
shows "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys \<longleftrightarrow> h' \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys"
using assms
apply (induct xs arbitrary: ys)
by(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I
by(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I
dest: returns_result_eq)
@ -696,8 +677,8 @@ subsection\<open>Miscellaneous Rules\<close>
lemma execute_bind_simp:
assumes "h \<turnstile> f \<rightarrow>\<^sub>r x" and "h \<turnstile> f \<rightarrow>\<^sub>h h'"
shows "h \<turnstile> f \<bind> g = h' \<turnstile> g x"
using assms
by(auto simp add: returns_result_def returns_heap_def bind_def execute_def
using assms
by(auto simp add: returns_result_def returns_heap_def bind_def execute_def
split: sum.splits)
lemma bind_cong [fundef_cong]:
@ -706,8 +687,8 @@ lemma bind_cong [fundef_cong]:
assumes "h \<turnstile> f1 = h \<turnstile> f2"
and "\<And>y h'. h \<turnstile> f1 \<rightarrow>\<^sub>r y \<Longrightarrow> h \<turnstile> f1 \<rightarrow>\<^sub>h h' \<Longrightarrow> h' \<turnstile> g1 y = h' \<turnstile> g2 y"
shows "h \<turnstile> (f1 \<bind> g1) = h \<turnstile> (f2 \<bind> g2)"
apply(insert assms, cases "h \<turnstile> f1")
by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
apply(insert assms, cases "h \<turnstile> f1")
by(auto simp add: bind_def returns_result_def returns_heap_def execute_def
split: sum.splits)
lemma bind_cong_2:
@ -730,7 +711,8 @@ definition preserved :: "('heap, 'e, 'result) prog \<Rightarrow> 'heap \<Rightar
where
"preserved f h h' \<longleftrightarrow> (\<forall>x. h \<turnstile> f \<rightarrow>\<^sub>r x \<longleftrightarrow> h' \<turnstile> f \<rightarrow>\<^sub>r x)"
lemma preserved_code [code]: "preserved f h h' = (((h \<turnstile> ok f) \<and> (h' \<turnstile> ok f) \<and> |h \<turnstile> f|\<^sub>r = |h' \<turnstile> f|\<^sub>r) \<or> ((\<not>h \<turnstile> ok f) \<and> (\<not>h' \<turnstile> ok f)))"
lemma preserved_code [code]:
"preserved f h h' = (((h \<turnstile> ok f) \<and> (h' \<turnstile> ok f) \<and> |h \<turnstile> f|\<^sub>r = |h' \<turnstile> f|\<^sub>r) \<or> ((\<not>h \<turnstile> ok f) \<and> (\<not>h' \<turnstile> ok f)))"
apply(auto simp add: preserved_def)[1]
apply (meson is_OK_returns_result_E is_OK_returns_result_I)+
done
@ -741,17 +723,17 @@ lemma transp_preserved_f [simp]: "transp (preserved f)"
by(auto simp add: preserved_def transp_def)
definition
definition
all_args :: "('a \<Rightarrow> ('heap, 'e, 'result) prog) \<Rightarrow> ('heap, 'e, 'result) prog set"
where
"all_args f = (\<Union>arg. {f arg})"
definition
reads :: "('heap \<Rightarrow> 'heap \<Rightarrow> bool) set \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'heap
definition
reads :: "('heap \<Rightarrow> 'heap \<Rightarrow> bool) set \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'heap
\<Rightarrow> 'heap \<Rightarrow> bool"
where
"reads S getter h h' \<longleftrightarrow> (\<forall>P \<in> S. reflp P \<and> transp P) \<and> ((\<forall>P \<in> S. P h h')
"reads S getter h h' \<longleftrightarrow> (\<forall>P \<in> S. reflp P \<and> transp P) \<and> ((\<forall>P \<in> S. P h h')
\<longrightarrow> preserved getter h h')"
lemma reads_singleton [simp]: "reads {preserved f} f h h'"
@ -763,18 +745,21 @@ lemma reads_bind_pure:
and "\<And>x. h \<turnstile> f \<rightarrow>\<^sub>r x \<Longrightarrow> reads S (g x) h h'"
shows "reads S (f \<bind> g) h h'"
using assms
by(auto simp add: reads_def pure_pure preserved_def
intro!: bind_pure_returns_result_I is_OK_returns_result_I
dest: pure_returns_heap_eq
by(auto simp add: reads_def pure_pure preserved_def
intro!: bind_pure_returns_result_I is_OK_returns_result_I
dest: pure_returns_heap_eq
elim!: bind_returns_result_E)
lemma reads_insert_writes_set_left: "\<forall>P \<in> S. reflp P \<and> transp P \<Longrightarrow> reads {getter} f h h' \<Longrightarrow> reads (insert getter S) f h h'"
lemma reads_insert_writes_set_left:
"\<forall>P \<in> S. reflp P \<and> transp P \<Longrightarrow> reads {getter} f h h' \<Longrightarrow> reads (insert getter S) f h h'"
unfolding reads_def by simp
lemma reads_insert_writes_set_right: "reflp getter \<Longrightarrow> transp getter \<Longrightarrow> reads S f h h' \<Longrightarrow> reads (insert getter S) f h h'"
lemma reads_insert_writes_set_right:
"reflp getter \<Longrightarrow> transp getter \<Longrightarrow> reads S f h h' \<Longrightarrow> reads (insert getter S) f h h'"
unfolding reads_def by blast
lemma reads_subset: "reads S f h h' \<Longrightarrow> \<forall>P \<in> S' - S. reflp P \<and> transp P \<Longrightarrow> S \<subseteq> S' \<Longrightarrow> reads S' f h h'"
lemma reads_subset:
"reads S f h h' \<Longrightarrow> \<forall>P \<in> S' - S. reflp P \<and> transp P \<Longrightarrow> S \<subseteq> S' \<Longrightarrow> reads S' f h h'"
by(auto simp add: reads_def)
lemma return_reads [simp]: "reads {} (return x) h h'"
@ -795,10 +780,10 @@ lemma filter_M_reads:
apply(induct xs)
by(auto intro: reads_subset[OF return_reads] intro!: reads_bind_pure)
definition writes ::
definition writes ::
"('heap, 'e, 'result) prog set \<Rightarrow> ('heap, 'e, 'result2) prog \<Rightarrow> 'heap \<Rightarrow> 'heap \<Rightarrow> bool"
where
"writes S setter h h'
where
"writes S setter h h'
\<longleftrightarrow> (h \<turnstile> setter \<rightarrow>\<^sub>h h' \<longrightarrow> (\<exists>progs. set progs \<subseteq> S \<and> h \<turnstile> iterate_M progs \<rightarrow>\<^sub>h h'))"
lemma writes_singleton [simp]: "writes (all_args f) (f a) h h'"
@ -847,7 +832,7 @@ lemma writes_pure [simp]:
by (metis bot.extremum iterate_M.simps(1) list.set(1) pure_returns_heap_eq return_returns_heap)
lemma writes_bind:
assumes "\<And>h2. writes S f h h2"
assumes "\<And>h2. writes S f h h2"
assumes "\<And>x h2. h \<turnstile> f \<rightarrow>\<^sub>r x \<Longrightarrow> h \<turnstile> f \<rightarrow>\<^sub>h h2 \<Longrightarrow> writes S (g x) h2 h'"
shows "writes S (f \<bind> g) h h'"
using assms

View File

@ -26,26 +26,26 @@
*
* SPDX-License-Identifier: BSD-2-Clause
* Repository: https://git.logicalhacking.com/adbrucker/isabelle-hacks/
* Dependencies: None (assert.thy is used for testing the theory but it is
* Dependencies: None (assert.thy is used for testing the theory but it is
* not required for providing the functionality of this hack)
***********************************************************************************)
(*
This file is based on commit 8a5e95421521c36ab71ab2711435a9bc0fa2c5cc from upstream
(https://git.logicalhacking.com/adbrucker/isabelle-hacks/). Merely the dependency to
Assert.thy has been removed by disabling the example section (which include assert
(*
This file is based on commit 8a5e95421521c36ab71ab2711435a9bc0fa2c5cc from upstream
(https://git.logicalhacking.com/adbrucker/isabelle-hacks/). Merely the dependency to
Assert.thy has been removed by disabling the example section (which include assert
checks).
*)
section\<open>Hiding Type Variables\<close>
text\<open> This theory\footnote{This theory can be used ``stand-alone,'' i.e., this theory is
text\<open> This theory\footnote{This theory can be used ``stand-alone,'' i.e., this theory is
not specific to the DOM formalization. The latest version is part of the ``Isabelle Hacks''
repository: \url{https://git.logicalhacking.com/adbrucker/isabelle-hacks/}.} implements
a mechanism for declaring default type variables for data types. This comes handy for complex
repository: \url{https://git.logicalhacking.com/adbrucker/isabelle-hacks/}.} implements
a mechanism for declaring default type variables for data types. This comes handy for complex
data types with many type variables.\<close>
theory
"Hiding_Type_Variables"
imports
imports
Main
keywords
"register_default_tvars"
@ -58,40 +58,40 @@ ML\<open>
signature HIDE_TVAR = sig
datatype print_mode = print_all | print | noprint
datatype tvar_subst = right | left
datatype parse_mode = parse | noparse
datatype parse_mode = parse | noparse
type hide_varT = {
name: string,
tvars: typ list,
typ_syn_tab : (string * typ list*string) Symtab.table,
typ_syn_tab : (string * typ list*string) Symtab.table,
print_mode: print_mode,
parse_mode: parse_mode
}
}
val parse_print_mode : string -> print_mode
val parse_parse_mode : string -> parse_mode
val register : string -> print_mode option -> parse_mode option ->
val register : string -> print_mode option -> parse_mode option ->
theory -> theory
val update_mode : string -> print_mode option -> parse_mode option ->
theory -> theory
val lookup : theory -> string -> hide_varT option
val hide_tvar_tr' : string -> Proof.context -> term list -> term
val hide_tvar_tr' : string -> Proof.context -> term list -> term
val hide_tvar_ast_tr : Proof.context -> Ast.ast list -> Ast.ast
val hide_tvar_subst_ast_tr : tvar_subst -> Proof.context -> Ast.ast list
val hide_tvar_subst_ast_tr : tvar_subst -> Proof.context -> Ast.ast list
-> Ast.ast
val hide_tvar_subst_return_ast_tr : tvar_subst -> Proof.context
val hide_tvar_subst_return_ast_tr : tvar_subst -> Proof.context
-> Ast.ast list -> Ast.ast
end
structure Hide_Tvar : HIDE_TVAR = struct
datatype print_mode = print_all | print | noprint
datatype tvar_subst = right | left
datatype parse_mode = parse | noparse
datatype parse_mode = parse | noparse
type hide_varT = {
name: string,
tvars: typ list,
typ_syn_tab : (string * typ list*string) Symtab.table,
typ_syn_tab : (string * typ list*string) Symtab.table,
print_mode: print_mode,
parse_mode: parse_mode
}
}
type hide_tvar_tab = (hide_varT) Symtab.table
fun hide_tvar_eq (a, a') = (#name a) = (#name a')
fun merge_tvar_tab (tab,tab') = Symtab.merge hide_tvar_eq (tab,tab')
@ -109,27 +109,27 @@ structure Hide_Tvar : HIDE_TVAR = struct
| parse_print_mode "print" = print
| parse_print_mode "noprint" = noprint
| parse_print_mode s = error("Print mode not supported: "^s)
fun parse_parse_mode "parse" = parse
| parse_parse_mode "noparse" = noparse
| parse_parse_mode s = error("Parse mode not supported: "^s)
fun update_mode typ_str print_mode parse_mode thy =
let
let
val ctx = Toplevel.context_of(Toplevel.theory_toplevel thy)
val typ = Syntax.parse_typ ctx typ_str (* no type checking *)
val name = case typ of
val name = case typ of
Type(name,_) => name
| _ => error("Complex type not (yet) supported.")
fun update tab =
let
val old_entry = (case Symtab.lookup tab name of
SOME t => t
let
val old_entry = (case Symtab.lookup tab name of
SOME t => t
| NONE => error ("Type shorthand not registered: "^name))
val print_m = case print_mode of
SOME m => m
| NONE => #print_mode old_entry
val parse_m = case parse_mode of
val parse_m = case parse_mode of
SOME m => m
| NONE => #parse_mode old_entry
val entry = {
@ -139,48 +139,48 @@ structure Hide_Tvar : HIDE_TVAR = struct
print_mode = print_m,
parse_mode = parse_m
}
in
in
Symtab.update (name,entry) tab
end
in
in
Context.theory_of ( (Data.map update) (Context.Theory thy))
end
fun lookup thy name =
let
val tab = (Data.get o Context.Theory) thy
in
in
Symtab.lookup tab name
end
fun obtain_normalized_vname lookup_table vname =
fun obtain_normalized_vname lookup_table vname =
case List.find (fn e => fst e = vname) lookup_table of
SOME (_,idx) => (lookup_table, Int.toString idx)
| NONE => let
fun max_idx [] = 0
| NONE => let
fun max_idx [] = 0
| max_idx ((_,idx)::lt) = Int.max(idx,max_idx lt)
val idx = (max_idx lookup_table ) + 1
in
((vname,idx)::lookup_table, Int.toString idx) end
fun normalize_typvar_type lt (Type (a, Ts)) =
let
let
fun switch (a,b) = (b,a)
val (Ts', lt') = fold_map (fn t => fn lt => switch (normalize_typvar_type lt t)) Ts lt
in
in
(lt', Type (a, Ts'))
end
| normalize_typvar_type lt (TFree (vname, S)) =
let
| normalize_typvar_type lt (TFree (vname, S)) =
let
val (lt, vname) = obtain_normalized_vname lt (vname)
in
in
(lt, TFree( vname, S))
end
| normalize_typvar_type lt (TVar (xi, S)) =
let
| normalize_typvar_type lt (TVar (xi, S)) =
let
val (lt, vname) = obtain_normalized_vname lt (Term.string_of_vname xi)
in
in
(lt, TFree( vname, S))
end
@ -195,26 +195,26 @@ structure Hide_Tvar : HIDE_TVAR = struct
fun normalize_typvar_term lt (Const (a, t)) = (lt, Const(a, t))
| normalize_typvar_term lt (Free (a, t)) = let
| normalize_typvar_term lt (Free (a, t)) = let
val (lt, vname) = obtain_normalized_vname lt a
in
(lt, Free(vname,t))
end
| normalize_typvar_term lt (Var (xi, t)) =
let
let
val (lt, vname) = obtain_normalized_vname lt (Term.string_of_vname xi)
in
(lt, Free(vname,t))
end
| normalize_typvar_term lt (Bound (i)) = (lt, Bound(i))
| normalize_typvar_term lt (Abs(s,ty,tr)) =
let
| normalize_typvar_term lt (Bound (i)) = (lt, Bound(i))
| normalize_typvar_term lt (Abs(s,ty,tr)) =
let
val (lt,tr) = normalize_typvar_term lt tr
in
(lt, Abs(s,ty,tr))
end
| normalize_typvar_term lt (t1$t2) =
let
let
val (lt,t1) = normalize_typvar_term lt t1
val (lt,t2) = normalize_typvar_term lt t2
in
@ -222,78 +222,78 @@ structure Hide_Tvar : HIDE_TVAR = struct
end
fun normalize_typvar_term' t = snd(normalize_typvar_term [] t)
fun normalize_typvar_term' t = snd(normalize_typvar_term [] t)
fun key_of_term (Const(s,_)) = if String.isPrefix "\<^type>" s
then Lexicon.unmark_type s
else ""
| key_of_term (Free(s,_)) = s
| key_of_term (Var(xi,_)) = Term.string_of_vname xi
| key_of_term (Var(xi,_)) = Term.string_of_vname xi
| key_of_term (Bound(_)) = error("Bound() not supported in key_of_term")
| key_of_term (Abs(_,_,_)) = error("Abs() not supported in key_of_term")
| key_of_term (t1$t2) = (key_of_term t1)^(key_of_term t2)
val key_of_term' = key_of_term o normalize_typvar_term'
fun hide_tvar_tr' tname ctx terms =
let
val mtyp = Syntax.parse_typ ctx tname (* no type checking *)
val mtyp = Syntax.parse_typ ctx tname (* no type checking *)
val (fq_name, _) = case mtyp of
Type(s,ts) => (s,ts)
| _ => error("Complex type not (yet) supported.")
| _ => error("Complex type not (yet) supported.")
val local_name_of = hd o rev o String.fields (fn c => c = #".")
fun hide_type tname = Syntax.const("(_) "^tname)
fun hide_type tname = Syntax.const("(_) "^tname)
val reg_type_as_term = Term.list_comb(Const(Lexicon.mark_type tname,dummyT),terms)
val key = key_of_term' reg_type_as_term
val actual_tvars_key = key_of_term reg_type_as_term
in
case lookup (Proof_Context.theory_of ctx) fq_name of
case lookup (Proof_Context.theory_of ctx) fq_name of
NONE => raise Match
| SOME e => let
val (tname,default_tvars_key) =
val (tname,default_tvars_key) =
case Symtab.lookup (#typ_syn_tab e) key of
NONE => (local_name_of tname, "")
| SOME (s,_,tv) => (local_name_of s,tv)
in
in
case (#print_mode e) of
print_all => hide_type tname
| print => if default_tvars_key=actual_tvars_key
then hide_type tname
else raise Match
| noprint => raise Match
end
end
end
fun hide_tvar_ast_tr ctx ast=
let
fun hide_tvar_ast_tr ctx ast=
let
val thy = Proof_Context.theory_of ctx
fun parse_ast ((Ast.Constant const)::[]) = (const,NONE)
| parse_ast ((Ast.Constant sort)::(Ast.Constant const)::[])
= (const,SOME sort)
| parse_ast _ = error("AST type not supported.")
| parse_ast ((Ast.Constant sort)::(Ast.Constant const)::[])
= (const,SOME sort)
| parse_ast _ = error("AST type not supported.")
val (decorated_name, decorated_sort) = parse_ast ast
val (decorated_name, decorated_sort) = parse_ast ast
val name = Lexicon.unmark_type decorated_name
val default_info = case lookup thy name of
val default_info = case lookup thy name of
NONE => error("No default type vars registered: "^name)
| SOME e => e
val _ = if #parse_mode default_info = noparse
val _ = if #parse_mode default_info = noparse
then error("Default type vars disabled (option noparse): "^name)
else ()
fun name_of_tvar tvar = case tvar of (TFree(n,_)) => n
fun name_of_tvar tvar = case tvar of (TFree(n,_)) => n
| _ => error("Unsupported type structure.")
val type_vars_ast =
let fun mk_tvar n =
case decorated_sort of
val type_vars_ast =
let fun mk_tvar n =
case decorated_sort of
NONE => Ast.Variable(name_of_tvar n)
| SOME sort => Ast.Appl([Ast.Constant("_ofsort"),
Ast.Variable(name_of_tvar n),
@ -303,15 +303,15 @@ structure Hide_Tvar : HIDE_TVAR = struct
end
in
Ast.Appl ((Ast.Constant decorated_name)::type_vars_ast)
end
end
fun register typ_str print_mode parse_mode thy =
let
let
val ctx = Toplevel.context_of(Toplevel.theory_toplevel thy)
val typ = Syntax.parse_typ ctx typ_str
val (name,tvars) = case typ of Type(name,tvars) => (name,tvars)
| _ => error("Unsupported type structure.")
val base_typ = Syntax.read_typ ctx typ_str
val (base_name,base_tvars) = case base_typ of Type(name,tvars) => (name,tvars)
| _ => error("Unsupported type structure.")
@ -319,10 +319,10 @@ structure Hide_Tvar : HIDE_TVAR = struct
val base_key = key_of_type' base_typ
val base_tvar_key = key_of_type base_typ
val print_m = case print_mode of
val print_m = case print_mode of
SOME m => m
| NONE => print_all
val parse_m = case parse_mode of
val parse_m = case parse_mode of
SOME m => m
| NONE => parse
val entry = {
@ -333,8 +333,8 @@ structure Hide_Tvar : HIDE_TVAR = struct
parse_mode = parse_m
}
val base_entry = if name = base_name
then
val base_entry = if name = base_name
then
{
name = "",
tvars = [],
@ -342,7 +342,7 @@ structure Hide_Tvar : HIDE_TVAR = struct
print_mode = noprint,
parse_mode = noparse
}
else case lookup thy base_name of
else case lookup thy base_name of
SOME e => e
| NONE => error ("No entry found for "^base_name^
" (via "^name^")")
@ -351,15 +351,15 @@ structure Hide_Tvar : HIDE_TVAR = struct
name = #name base_entry,
tvars = #tvars base_entry,
typ_syn_tab = Symtab.update (base_key, (name, base_tvars, base_tvar_key))
(#typ_syn_tab (base_entry)),
(#typ_syn_tab (base_entry)),
print_mode = #print_mode base_entry,
parse_mode = #parse_mode base_entry
}
fun reg tab = let
fun reg tab = let
val tab = Symtab.update_new(name, entry) tab
val tab = if name = base_name
then tab
val tab = if name = base_name
then tab
else Symtab.update(base_name, base_entry) tab
in
tab
@ -368,13 +368,13 @@ structure Hide_Tvar : HIDE_TVAR = struct
val thy = Sign.print_translation
[(Lexicon.mark_type name, hide_tvar_tr' name)] thy
in
in
Context.theory_of ( (Data.map reg) (Context.Theory thy))
handle Symtab.DUP _ => error("Type shorthand already registered: "^name)
end
fun hide_tvar_subst_ast_tr hole ctx (ast::[]) =
let
let
val thy = Proof_Context.theory_of ctx
val (decorated_name, args) = case ast
@ -385,23 +385,23 @@ structure Hide_Tvar : HIDE_TVAR = struct
val default_info = case lookup thy name of
NONE => error("No default type vars registered: "^name)
| SOME e => e
val _ = if #parse_mode default_info = noparse
val _ = if #parse_mode default_info = noparse
then error("Default type vars disabled (option noparse): "^name)
else ()
fun name_of_tvar tvar = case tvar of (TFree(n,_)) => n
fun name_of_tvar tvar = case tvar of (TFree(n,_)) => n
| _ => error("Unsupported type structure.")
val type_vars_ast = map (fn n => Ast.Variable(name_of_tvar n)) (#tvars default_info)
val type_vars_ast = case hole of
val type_vars_ast = case hole of
right => (List.rev(List.drop(List.rev type_vars_ast, List.length args)))@args
| left => args@List.drop(type_vars_ast, List.length args)
in
Ast.Appl ((Ast.Constant decorated_name)::type_vars_ast)
end
end
| hide_tvar_subst_ast_tr _ _ _ = error("hide_tvar_subst_ast_tr: empty AST.")
fun hide_tvar_subst_return_ast_tr hole ctx (retval::constructor::[]) =
fun hide_tvar_subst_return_ast_tr hole ctx (retval::constructor::[]) =
hide_tvar_subst_ast_tr hole ctx [Ast.Appl (constructor::retval::[])]
| hide_tvar_subst_return_ast_tr _ _ _ =
| hide_tvar_subst_return_ast_tr _ _ _ =
error("hide_tvar_subst_return_ast_tr: error in parsing AST")
@ -411,7 +411,7 @@ end
subsection\<open>Register Parse Translations\<close>
syntax "_tvars_wildcard" :: "type \<Rightarrow> type" ("'('_') _")
syntax "_tvars_wildcard" :: "type \<Rightarrow> type" ("'('_') _")
syntax "_tvars_wildcard_retval" :: "type \<Rightarrow> type \<Rightarrow> type" ("'('_, _') _")
syntax "_tvars_wildcard_sort" :: "sort \<Rightarrow> type \<Rightarrow> type" ("'('_::_') _")
syntax "_tvars_wildcard_right" :: "type \<Rightarrow> type" ("_ '_..")
@ -431,42 +431,42 @@ subsection\<open>Register Top-Level Isar Commands\<close>
ML\<open>
val modeP = (Parse.$$$ "("
|-- (Parse.name --| Parse.$$$ ","
-- Parse.name --|
-- Parse.name --|
Parse.$$$ ")"))
val typ_modeP = Parse.typ -- (Scan.optional modeP ("print_all","parse"))
val _ = Outer_Syntax.command @{command_keyword "register_default_tvars"}
"Register default variables (and hiding mechanims) for a type."
(typ_modeP >> (fn (typ,(print_m,parse_m)) =>
(Toplevel.theory
(Hide_Tvar.register typ
(SOME (Hide_Tvar.parse_print_mode print_m))
(SOME (Hide_Tvar.parse_parse_mode parse_m))))));
(typ_modeP >> (fn (typ,(print_m,parse_m)) =>
(Toplevel.theory
(Hide_Tvar.register typ
(SOME (Hide_Tvar.parse_print_mode print_m))
(SOME (Hide_Tvar.parse_parse_mode parse_m))))));
val _ = Outer_Syntax.command @{command_keyword "update_default_tvars_mode"}
"Update print and/or parse mode or the default type variables for a certain type."
(typ_modeP >> (fn (typ,(print_m,parse_m)) =>
(Toplevel.theory
(Hide_Tvar.update_mode typ
(SOME (Hide_Tvar.parse_print_mode print_m))
(SOME (Hide_Tvar.parse_parse_mode parse_m))))));
(typ_modeP >> (fn (typ,(print_m,parse_m)) =>
(Toplevel.theory
(Hide_Tvar.update_mode typ
(SOME (Hide_Tvar.parse_print_mode print_m))
(SOME (Hide_Tvar.parse_parse_mode parse_m))))));
\<close>
(*
section\<open>Examples\<close>
subsection\<open>Print Translation\<close>
datatype ('a, 'b) hide_tvar_foobar = hide_tvar_foo 'a | hide_tvar_bar 'b
datatype ('a, 'b) hide_tvar_foobar = hide_tvar_foo 'a | hide_tvar_bar 'b
type_synonym ('a, 'b, 'c, 'd) hide_tvar_baz = "('a+'b, 'a \<times> 'b) hide_tvar_foobar"
definition hide_tvar_f::"('a, 'b) hide_tvar_foobar \<Rightarrow> ('a, 'b) hide_tvar_foobar \<Rightarrow> ('a, 'b) hide_tvar_foobar"
definition hide_tvar_f::"('a, 'b) hide_tvar_foobar \<Rightarrow> ('a, 'b) hide_tvar_foobar \<Rightarrow> ('a, 'b) hide_tvar_foobar"
where "hide_tvar_f a b = a"
definition hide_tvar_g::"('a, 'b, 'c, 'd) hide_tvar_baz \<Rightarrow> ('a, 'b, 'c, 'd) hide_tvar_baz \<Rightarrow> ('a, 'b, 'c, 'd) hide_tvar_baz"
definition hide_tvar_g::"('a, 'b, 'c, 'd) hide_tvar_baz \<Rightarrow> ('a, 'b, 'c, 'd) hide_tvar_baz \<Rightarrow> ('a, 'b, 'c, 'd) hide_tvar_baz"
where "hide_tvar_g a b = a"
assert[string_of_thm_equal,
thm_def="hide_tvar_f_def",
thm_def="hide_tvar_f_def",
str="hide_tvar_f (a::('a, 'b) hide_tvar_foobar) (b::('a, 'b) hide_tvar_foobar) = a"]
assert[string_of_thm_equal,
thm_def="hide_tvar_g_def",
thm_def="hide_tvar_g_def",
str="hide_tvar_g (a::('a + 'b, 'a \<times> 'b) hide_tvar_foobar) (b::('a + 'b, 'a \<times> 'b) hide_tvar_foobar) = a"]
register_default_tvars "('alpha, 'beta) hide_tvar_foobar" (print_all,parse)
@ -477,7 +477,7 @@ assert[string_of_thm_equal,
thm_def="hide_tvar_f_def",
str="hide_tvar_f (a::('a, 'b) hide_tvar_foobar) (b::('a, 'b) hide_tvar_foobar) = a"]
assert[string_of_thm_equal,
thm_def="hide_tvar_g_def",
thm_def="hide_tvar_g_def",
str="hide_tvar_g (a::('a + 'b, 'a \<times> 'b) hide_tvar_foobar) (b::('a + 'b, 'a \<times> 'b) hide_tvar_foobar) = a"]
update_default_tvars_mode "_ hide_tvar_foobar" (print_all,noparse)
@ -501,29 +501,29 @@ definition hide_tvar_A' :: "'x \<Rightarrow> (('x,'b) hide_tvar_foobar) .._"
assert[string_of_thm_equal,
thm_def="hide_tvar_A'_def", str="hide_tvar_A' (x::'x) = hide_tvar_foo x"]
definition hide_tvar_B' :: "(_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar"
definition hide_tvar_B' :: "(_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar"
where "hide_tvar_B' x y = x"
assert[string_of_thm_equal,
thm_def="hide_tvar_A'_def", str="hide_tvar_A' (x::'x) = hide_tvar_foo x"]
definition hide_tvar_B :: "(_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar"
definition hide_tvar_B :: "(_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar"
where "hide_tvar_B x y = x"
assert[string_of_thm_equal,
thm_def="hide_tvar_B_def", str="hide_tvar_B (x::(_) hide_tvar_foobar) (y::(_) hide_tvar_foobar) = x"]
definition hide_tvar_C :: "(_) hide_tvar_baz \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_baz"
definition hide_tvar_C :: "(_) hide_tvar_baz \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_baz"
where "hide_tvar_C x y = x"
assert[string_of_thm_equal,
thm_def="hide_tvar_C_def", str="hide_tvar_C (x::(_) hide_tvar_baz) (y::(_) hide_tvar_foobar) = x"]
definition hide_tvar_E :: "(_::linorder) hide_tvar_baz \<Rightarrow> (_::linorder) hide_tvar_foobar \<Rightarrow> (_::linorder) hide_tvar_baz"
definition hide_tvar_E :: "(_::linorder) hide_tvar_baz \<Rightarrow> (_::linorder) hide_tvar_foobar \<Rightarrow> (_::linorder) hide_tvar_baz"
where "hide_tvar_E x y = x"
assert[string_of_thm_equal,
thm_def="hide_tvar_C_def", str="hide_tvar_C (x::(_) hide_tvar_baz) (y::(_) hide_tvar_foobar) = x"]
definition hide_tvar_X :: "(_, 'retval::linorder) hide_tvar_baz
\<Rightarrow> (_,'retval) hide_tvar_foobar
definition hide_tvar_X :: "(_, 'retval::linorder) hide_tvar_baz
\<Rightarrow> (_,'retval) hide_tvar_foobar
\<Rightarrow> (_,'retval) hide_tvar_baz"
where "hide_tvar_X x y = x"
*)
@ -531,52 +531,52 @@ definition hide_tvar_X :: "(_, 'retval::linorder) hide_tvar_baz
subsection\<open>Introduction\<close>
text\<open>
When modelling object-oriented data models in HOL with the goal of preserving \<^emph>\<open>extensibility\<close>
(e.g., as described in~\cite{brucker.ea:extensible:2008-b,brucker:interactive:2007}) one needs
When modelling object-oriented data models in HOL with the goal of preserving \<^emph>\<open>extensibility\<close>
(e.g., as described in~\cite{brucker.ea:extensible:2008-b,brucker:interactive:2007}) one needs
to define type constructors with a large number of type variables. This can reduce the readability
of the overall formalization. Thus, we use a short-hand notation in cases were the names of
the type variables are known from the context. In more detail, this theory sets up both
configurable print and parse translations that allows for replacing @{emph \<open>all\<close>} type variables
by \<open>(_)\<close>, e.g., a five-ary constructor \<open>('a, 'b, 'c, 'd, 'e) hide_tvar_foo\<close> can
be shorted to \<open>(_) hide_tvar_foo\<close>. The use of this shorthand in output (printing) and
input (parsing) is, on a per-type basis, user-configurable using the top-level commands
\<open>register_default_tvars\<close> (for registering the names of the default type variables and
the print/parse mode) and \<open>update_default_tvars_mode\<close> (for changing the print/parse mode
dynamically).
of the overall formalization. Thus, we use a short-hand notation in cases were the names of
the type variables are known from the context. In more detail, this theory sets up both
configurable print and parse translations that allows for replacing @{emph \<open>all\<close>} type variables
by \<open>(_)\<close>, e.g., a five-ary constructor \<open>('a, 'b, 'c, 'd, 'e) hide_tvar_foo\<close> can
be shorted to \<open>(_) hide_tvar_foo\<close>. The use of this shorthand in output (printing) and
input (parsing) is, on a per-type basis, user-configurable using the top-level commands
\<open>register_default_tvars\<close> (for registering the names of the default type variables and
the print/parse mode) and \<open>update_default_tvars_mode\<close> (for changing the print/parse mode
dynamically).
The input also supports short-hands for declaring default sorts (e.g., \<open>(_::linorder)\<close>
specifies that all default variables need to be instances of the sort (type class)
@{class \<open>linorder\<close>} and short-hands of overriding a suffice (or prefix) of the default type
variables. For example, \<open>('state) hide_tvar_foo _.\<close> is a short-hand for
\<open>('a, 'b, 'c, 'd, 'state) hide_tvar_foo\<close>. In this document, we omit the implementation
details (we refer the interested reader to theory file) and continue directly with a few
examples.
The input also supports short-hands for declaring default sorts (e.g., \<open>(_::linorder)\<close>
specifies that all default variables need to be instances of the sort (type class)
@{class \<open>linorder\<close>} and short-hands of overriding a suffice (or prefix) of the default type
variables. For example, \<open>('state) hide_tvar_foo _.\<close> is a short-hand for
\<open>('a, 'b, 'c, 'd, 'state) hide_tvar_foo\<close>. In this document, we omit the implementation
details (we refer the interested reader to theory file) and continue directly with a few
examples.
\<close>
subsection\<open>Example\<close>
text\<open>Given the following type definition:\<close>
datatype ('a, 'b) hide_tvar_foobar = hide_tvar_foo 'a | hide_tvar_bar 'b
datatype ('a, 'b) hide_tvar_foobar = hide_tvar_foo 'a | hide_tvar_bar 'b
type_synonym ('a, 'b, 'c, 'd) hide_tvar_baz = "('a+'b, 'a \<times> 'b) hide_tvar_foobar"
text\<open>We can register default values for the type variables for the abstract
data type as well as the type synonym:\<close>
data type as well as the type synonym:\<close>
register_default_tvars "('alpha, 'beta) hide_tvar_foobar" (print_all,parse)
register_default_tvars "('alpha, 'beta, 'gamma, 'delta) hide_tvar_baz" (print_all,parse)
text\<open>This allows us to write\<close>
definition hide_tvar_f::"(_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar"
definition hide_tvar_f::"(_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar \<Rightarrow> (_) hide_tvar_foobar"
where "hide_tvar_f a b = a"
definition hide_tvar_g::"(_) hide_tvar_baz \<Rightarrow> (_) hide_tvar_baz \<Rightarrow> (_) hide_tvar_baz"
definition hide_tvar_g::"(_) hide_tvar_baz \<Rightarrow> (_) hide_tvar_baz \<Rightarrow> (_) hide_tvar_baz"
where "hide_tvar_g a b = a"
text\<open>Instead of specifying the type variables explicitely. This makes, in particular
for type constructors with a large number of type variables, definitions much
more concise. This syntax is also used in the output of antiquotations, e.g.,
@{term[show_types] "x = hide_tvar_g"}. Both the print translation and the parse
for type constructors with a large number of type variables, definitions much
more concise. This syntax is also used in the output of antiquotations, e.g.,
@{term[show_types] "x = hide_tvar_g"}. Both the print translation and the parse
translation can be disabled for each type individually:\<close>
update_default_tvars_mode "_ hide_tvar_foobar" (noprint,noparse)
update_default_tvars_mode "_ hide_tvar_foobar" (noprint,noparse)
text\<open> Now, Isabelle's interactive output and the antiquotations will show
text\<open> Now, Isabelle's interactive output and the antiquotations will show
all type variables, e.g., @{term[show_types] "x = hide_tvar_g"}.\<close>

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -31,10 +31,10 @@ theory Testing_Utils
imports Main
begin
ML \<open>
val _ = Theory.setup
val _ = Theory.setup
(Method.setup @{binding timed_code_simp}
(Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo (fn a => fn b => fn tac =>
let
let
val start = Time.now ();
val result = Code_Simp.dynamic_tac a b tac;
val t = Time.now() - start;
@ -75,10 +75,12 @@ val _ = Theory.setup
handle Timeout.TIMEOUT _ => NONE;
val t2 = Time.now() - start2;
in
if length (Seq.list_of result) > 0 then (Output.information ("eval took " ^ (Time.toString t)); File.append (Path.explode "/tmp/isabellebench") (Time.toString t ^ ",")) else ();
if length (Seq.list_of result) > 0 then (Output.information ("eval took " ^ (Time.toString t));
File.append (Path.explode "/tmp/isabellebench") (Time.toString t ^ ",")) else ();
(case result2_opt of
SOME result2 =>
(if length (Seq.list_of result2) > 0 then (Output.information ("code_simp took " ^ (Time.toString t2)); File.append (Path.explode "/tmp/isabellebench") (Time.toString t2 ^ "\n")) else ())
SOME result2 =>
(if length (Seq.list_of result2) > 0 then (Output.information ("code_simp took " ^ (Time.toString t2));
File.append (Path.explode "/tmp/isabellebench") (Time.toString t2 ^ "\n")) else ())
| NONE => (Output.information "code_simp timed out after 600s"; File.append (Path.explode "/tmp/isabellebench") (">600.000\n")));
result
end)))

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -32,9 +32,9 @@ text\<open>This theory provides the common test setup that is used by all formal
theory Core_DOM_BaseTest
imports
(*<*)
(*<*)
"../preliminaries/Testing_Utils"
(*>*)
(*>*)
"../Core_DOM"
begin
@ -47,7 +47,7 @@ notation assert_throws ("assert'_throws'(_, _')")
definition "test p h \<longleftrightarrow> h \<turnstile> ok p"
definition field_access :: "(string \<Rightarrow> (_, (_) object_ptr option) dom_prog) \<Rightarrow> string
definition field_access :: "(string \<Rightarrow> (_, (_) object_ptr option) dom_prog) \<Rightarrow> string
\<Rightarrow> (_, (_) object_ptr option) dom_prog" (infix "." 80)
where
"field_access m field = m field"
@ -133,7 +133,7 @@ notation create_document_with_null ("createDocument'(_')")
notation create_document_with_null2 ("createDocument'(_, _, _')")
fun get_element_by_id_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
where
where
"get_element_by_id_with_null (Some ptr) id' = do {
element_ptr_opt \<leftarrow> get_element_by_id ptr id';
(case element_ptr_opt of
@ -142,19 +142,23 @@ fun get_element_by_id_with_null :: "((_::linorder) object_ptr option) \<Rightarr
| "get_element_by_id_with_null _ _ = error SegmentationFault"
notation get_element_by_id_with_null ("_ . getElementById'(_')")
fun get_elements_by_class_name_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog"
where
fun get_elements_by_class_name_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog"
where
"get_elements_by_class_name_with_null (Some ptr) class_name =
get_elements_by_class_name ptr class_name \<bind> map_M (return \<circ> Some \<circ> cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r)"
notation get_elements_by_class_name_with_null ("_ . getElementsByClassName'(_')")
fun get_elements_by_tag_name_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog"
where
fun get_elements_by_tag_name_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog"
where
"get_elements_by_tag_name_with_null (Some ptr) tag =
get_elements_by_tag_name ptr tag \<bind> map_M (return \<circ> Some \<circ> cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r)"
notation get_elements_by_tag_name_with_null ("_ . getElementsByTagName'(_')")
fun insert_before_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
fun insert_before_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where
"insert_before_with_null (Some ptr) (Some child_obj) ref_child_obj_opt = (case cast child_obj of
Some child \<Rightarrow> do {
@ -165,7 +169,8 @@ fun insert_before_with_null :: "((_::linorder) object_ptr option) \<Rightarrow>
| None \<Rightarrow> error HierarchyRequestError)"
notation insert_before_with_null ("_ . insertBefore'(_, _')")
fun append_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, unit) dom_prog"
fun append_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, unit) dom_prog"
where
"append_child_with_null (Some ptr) (Some child_obj) = (case cast child_obj of
Some child \<Rightarrow> append_child ptr child
@ -180,7 +185,8 @@ fun get_body :: "((_::linorder) object_ptr option) \<Rightarrow> (_, ((_) object
}"
notation get_body ("_ . body")
fun get_document_element_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
fun get_document_element_with_null :: "((_::linorder) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where
"get_document_element_with_null (Some ptr) = (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
Some document_ptr \<Rightarrow> do {
@ -190,14 +196,16 @@ fun get_document_element_with_null :: "((_::linorder) object_ptr option) \<Right
| None \<Rightarrow> None)})"
notation get_document_element_with_null ("_ . documentElement")
fun get_owner_document_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
fun get_owner_document_with_null :: "((_::linorder) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where
"get_owner_document_with_null (Some ptr) = (do {
document_ptr \<leftarrow> get_owner_document ptr;
return (Some (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr))})"
notation get_owner_document_with_null ("_ . ownerDocument")
fun remove_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
fun remove_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where
"remove_with_null (Some ptr) (Some child) = (case cast child of
Some child_node \<Rightarrow> do {
@ -208,7 +216,8 @@ fun remove_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) ob
| "remove_with_null _ None = error TypeError"
notation remove_with_null ("_ . remove'(')")
fun remove_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
fun remove_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where
"remove_child_with_null (Some ptr) (Some child) = (case cast child of
Some child_node \<Rightarrow> do {
@ -230,11 +239,11 @@ notation remove_attribute_with_null ("_ . removeAttribute'(_')")
fun get_attribute_with_null :: "((_) object_ptr option) \<Rightarrow> attr_key \<Rightarrow> (_, attr_value option) dom_prog"
where
"get_attribute_with_null (Some ptr) k = (case cast ptr of
"get_attribute_with_null (Some ptr) k = (case cast ptr of
Some element_ptr \<Rightarrow> get_attribute element_ptr k)"
fun get_attribute_with_null2 :: "((_) object_ptr option) \<Rightarrow> attr_key \<Rightarrow> (_, attr_value) dom_prog"
where
"get_attribute_with_null2 (Some ptr) k = (case cast ptr of
"get_attribute_with_null2 (Some ptr) k = (case cast ptr of
Some element_ptr \<Rightarrow> do {
a \<leftarrow> get_attribute element_ptr k;
return (the a)})"
@ -256,7 +265,8 @@ fun first_child_with_null :: "((_) object_ptr option) \<Rightarrow> (_, ((_) obj
| None \<Rightarrow> None)}"
notation first_child_with_null ("_ . firstChild")
fun adopt_node_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog"
fun adopt_node_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>(_, ((_) object_ptr option)) dom_prog"
where
"adopt_node_with_null (Some ptr) (Some child) = (case cast ptr of
Some document_ptr \<Rightarrow> (case cast child of
@ -264,9 +274,10 @@ fun adopt_node_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_
adopt_node document_ptr child_node;
return (Some child)}))"
notation adopt_node_with_null ("_ . adoptNode'(_')")
definition createTestTree :: "((_::linorder) object_ptr option) \<Rightarrow> (_, (string \<Rightarrow> (_, ((_) object_ptr option)) dom_prog)) dom_prog"
definition createTestTree ::
"((_::linorder) object_ptr option) \<Rightarrow> (_, (string \<Rightarrow> (_, ((_) object_ptr option)) dom_prog)) dom_prog"
where
"createTestTree ref = return (\<lambda>id. get_element_by_id_with_null ref id)"

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)

File diff suppressed because it is too large Load Diff

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -34,9 +34,9 @@ theory ElementClass
"NodeClass"
"ShadowRootPointer"
begin
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
type_synonym attr_key = DOMString
type_synonym attr_key = DOMString
type_synonym attr_value = DOMString
type_synonym attrs = "(attr_key, attr_value) fmap"
type_synonym tag_name = DOMString
@ -46,36 +46,45 @@ record ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr) RElement
child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
attrs :: attrs
shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_scheme"
register_default_tvars
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_scheme"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node
= "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Node"
register_default_tvars
= "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext
+ 'Node) Node"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object
= "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Object"
register_default_tvars
= "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_ext + 'Node) Object"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object"
type_synonym
('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap
= "('document_ptr document_ptr + 'shadow_root_ptr shadow_root_ptr + 'object_ptr, 'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object,
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) heap"
('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap
= "('document_ptr document_ptr + 'shadow_root_ptr shadow_root_ptr + 'object_ptr,
'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object,
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext +
'Node) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap"
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition element_ptr_kinds :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where
"element_ptr_kinds heap = the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_element_ptr_kind (node_ptr_kinds heap)))"
"element_ptr_kinds heap =
the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_element_ptr_kind (node_ptr_kinds heap)))"
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) = {|element_ptr|} |\<union>| element_ptr_kinds h"
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) =
{|element_ptr|} |\<union>| element_ptr_kinds h"
apply(auto simp add: element_ptr_kinds_def)[1]
by force
@ -85,7 +94,8 @@ definition element_ptrs :: "(_) heap \<Rightarrow> (_) element_ptr fset"
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Node \<Rightarrow> (_) Element option"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = (case RNode.more node of Inl element \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> None)"
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node =
(case RNode.more node of Inl element \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \<Rightarrow> (_) Element option"
@ -116,15 +126,15 @@ abbreviation is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ::
"is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
adhoc_overloading is_element_kind is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma element_ptr_kinds_commutes [simp]:
lemma element_ptr_kinds_commutes [simp]:
"cast element_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> element_ptr |\<in>| element_ptr_kinds h"
apply(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)[1]
by (metis (no_types, lifting) element_ptr_casts_commute2 ffmember_filter fimage_eqI
fset.map_comp is_element_ptr_kind_none node_ptr_casts_commute3
by (metis (no_types, lifting) element_ptr_casts_commute2 ffmember_filter fimage_eqI
fset.map_comp is_element_ptr_kind_none node_ptr_casts_commute3
node_ptr_kinds_commutes node_ptr_kinds_def option.sel option.simps(3))
definition get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Element option"
where
where
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) h) cast"
adhoc_overloading get get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
@ -163,9 +173,9 @@ global_interpretation l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^su
by unfold_locales
definition put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) Element \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
where
"put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) (cast element)"
adhoc_overloading put put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
adhoc_overloading put put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h = h'"
@ -182,30 +192,30 @@ lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs:
lemma cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]:
lemma cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]:
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e x = cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def)
by (metis (full_types) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = None \<longleftrightarrow> \<not> (\<exists>element. cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = Some element \<longleftrightarrow> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element) = Some element"
by simp
lemma get_elment_ptr_simp1 [simp]:
lemma get_elment_ptr_simp1 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h) = Some element"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_elment_ptr_simp2 [simp]:
"element_ptr \<noteq> element_ptr'
lemma get_elment_ptr_simp2 [simp]:
"element_ptr \<noteq> element_ptr'
\<Longrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' element h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -217,9 +227,9 @@ abbreviation "create_element_obj tag_name_arg child_nodes_arg attrs_arg shadow_r
definition new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) heap \<Rightarrow> ((_) element_ptr \<times> (_) heap)"
where
"new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_element_ptr = element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref
|`| (element_ptrs h)))))
"new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_element_ptr = element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref
|`| (element_ptrs h)))))
in
(new_element_ptr, put new_element_ptr (create_element_obj '''' [] fmempty None) h))"
@ -230,7 +240,7 @@ lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
unfolding new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_element_ptr_new:
lemma new_element_ptr_new:
"element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref |`| element_ptrs h)))) |\<notin>| element_ptrs h"
by (metis Suc_n_not_le_n element_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
@ -293,16 +303,20 @@ definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
lemma known_ptrs_known_ptr:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> known_ptr ptr"
apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .

View File

@ -23,28 +23,28 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>ShadowRoot\<close>
text\<open>In this theory, we introduce the typed pointers for the class ShadowRoot. Note that, in
this document, we will not make use of ShadowRoots nor will we discuss their particular properties.
text\<open>In this theory, we introduce the typed pointers for the class ShadowRoot. Note that, in
this document, we will not make use of ShadowRoots nor will we discuss their particular properties.
We only include them here, as they are required for future work and they cannot be added alter
following the object-oriented extensibility of our data model.\<close>
following the object-oriented extensibility of our data model.\<close>
theory ShadowRootPointer
imports
"DocumentPointer"
begin
datatype 'shadow_root_ptr shadow_root_ptr = Ref (the_ref: ref) | Ext 'shadow_root_ptr
register_default_tvars "'shadow_root_ptr shadow_root_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
register_default_tvars "'shadow_root_ptr shadow_root_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr
= "('shadow_root_ptr shadow_root_ptr + 'object_ptr, 'node_ptr, 'element_ptr,
= "('shadow_root_ptr shadow_root_ptr + 'object_ptr, 'node_ptr, 'element_ptr,
'character_data_ptr, 'document_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr"
definition cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr"
where
@ -57,7 +57,7 @@ definition cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\
definition cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> (_) shadow_root_ptr option"
where
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of
object_ptr.Ext (Inr (Inr (Inl shadow_root_ptr))) \<Rightarrow> Some shadow_root_ptr
object_ptr.Ext (Inr (Inr (Inl shadow_root_ptr))) \<Rightarrow> Some shadow_root_ptr
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
@ -65,13 +65,13 @@ adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^su
definition is_shadow_root_ptr_kind :: "(_) object_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr_kind ptr = (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some _ \<Rightarrow> True
"is_shadow_root_ptr_kind ptr = (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some _ \<Rightarrow> True
| None \<Rightarrow> False)"
consts is_shadow_root_ptr :: 'a
definition is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of shadow_root_ptr.Ref _ \<Rightarrow> True
"is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of shadow_root_ptr.Ref _ \<Rightarrow> True
| _ \<Rightarrow> False)"
abbreviation is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
@ -92,16 +92,16 @@ adhoc_overloading is_shadow_root_ptr_ext is_shadow_root_ptr_ext\<^sub>o\<^sub>b\
instantiation shadow_root_ptr :: (linorder) linorder
begin
definition
definition
less_eq_shadow_root_ptr :: "(_::linorder) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> bool"
where
where
"less_eq_shadow_root_ptr x y \<equiv> (case x of Ext i \<Rightarrow> (case y of Ext j \<Rightarrow> i \<le> j | Ref _ \<Rightarrow> False)
| Ref i \<Rightarrow> (case y of Ext _ \<Rightarrow> True | Ref j \<Rightarrow> i \<le> j))"
definition less_shadow_root_ptr :: "(_::linorder) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> bool"
where "less_shadow_root_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance
instance
apply(standard)
by(auto simp add: less_eq_shadow_root_ptr_def less_shadow_root_ptr_def
by(auto simp add: less_eq_shadow_root_ptr_def less_shadow_root_ptr_def
split: shadow_root_ptr.splits)
end
@ -122,21 +122,21 @@ lemma cast_shadow_root_ptr_not_document_ptr [simp]:
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr \<noteq> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr"
unfolding cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def by auto
lemma shadow_root_ptr_no_node_ptr_cast [simp]:
lemma shadow_root_ptr_no_node_ptr_cast [simp]:
"\<not> is_shadow_root_ptr_kind (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)"
by(simp add: cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_shadow_root_ptr_kind_def)
lemma node_ptr_no_shadow_root_ptr_cast [simp]:
lemma node_ptr_no_shadow_root_ptr_cast [simp]:
"\<not> is_node_ptr_kind (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr)"
using is_node_ptr_kind_obtains by fastforce
lemma shadow_root_ptr_no_document_ptr_cast [simp]:
lemma shadow_root_ptr_no_document_ptr_cast [simp]:
"\<not> is_shadow_root_ptr_kind (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)"
by(simp add: cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_shadow_root_ptr_kind_def)
lemma document_ptr_no_shadow_root_ptr_cast [simp]:
lemma document_ptr_no_shadow_root_ptr_cast [simp]:
"\<not> is_document_ptr_kind (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr)"
using is_document_ptr_kind_obtains by fastforce
lemma shadow_root_ptr_shadow_root_ptr_cast [simp]:
lemma shadow_root_ptr_shadow_root_ptr_cast [simp]:
"is_shadow_root_ptr_kind (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr)"
by (simp add: cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def is_shadow_root_ptr_kind_def)
@ -145,7 +145,7 @@ lemma shadow_root_ptr_casts_commute [simp]:
unfolding cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: object_ptr.splits sum.splits)
lemma shadow_root_ptr_casts_commute2 [simp]:
lemma shadow_root_ptr_casts_commute2 [simp]:
"(cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) = Some shadow_root_ptr)"
by simp
@ -169,11 +169,11 @@ lemma is_shadow_root_ptr_kind_none:
unfolding is_shadow_root_ptr_kind_def cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by (auto split: object_ptr.splits sum.splits)
lemma cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
lemma cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
"cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r x = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r y \<longleftrightarrow> x = y"
by(simp add: cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
lemma cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (object_ptr.Ext (Inr (Inr (Inr object_ext_ptr)))) = None"
by(simp add: cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)

File diff suppressed because it is too large Load Diff

View File

@ -23,7 +23,7 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
@ -34,9 +34,9 @@ theory ElementClass
"NodeClass"
"ShadowRootPointer"
begin
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
type_synonym attr_key = DOMString
type_synonym attr_key = DOMString
type_synonym attr_value = DOMString
type_synonym attrs = "(attr_key, attr_value) fmap"
type_synonym tag_name = DOMString
@ -46,36 +46,43 @@ record ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr) RElement
child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list"
attrs :: attrs
shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element
= "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_scheme"
register_default_tvars
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node
= "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Node"
register_default_tvars
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node"
type_synonym
type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object
= "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Object"
register_default_tvars
= "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_ext + 'Node) Object"
register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object"
type_synonym
('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap
= "(('document_ptr, 'shadow_root_ptr) document_ptr + 'object_ptr, 'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object,
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) heap"
('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap
= "(('document_ptr, 'shadow_root_ptr) document_ptr + 'object_ptr, 'element_ptr element_ptr +
'character_data_ptr character_data_ptr + 'node_ptr, 'Object,
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option)
RElement_ext + 'Node) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap"
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr,
'Object, 'Node, 'Element) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition element_ptr_kinds :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where
"element_ptr_kinds heap = the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_element_ptr_kind (node_ptr_kinds heap)))"
"element_ptr_kinds heap = the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`|
(ffilter is_element_ptr_kind (node_ptr_kinds heap)))"
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) = {|element_ptr|} |\<union>| element_ptr_kinds h"
lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) =
{|element_ptr|} |\<union>| element_ptr_kinds h"
apply(auto simp add: element_ptr_kinds_def)[1]
by force
@ -85,7 +92,8 @@ definition element_ptrs :: "(_) heap \<Rightarrow> (_) element_ptr fset"
definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Node \<Rightarrow> (_) Element option"
where
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = (case RNode.more node of Inl element \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> None)"
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = (case RNode.more node of Inl element \<Rightarrow>
Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \<Rightarrow> (_) Element option"
@ -116,15 +124,15 @@ abbreviation is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ::
"is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr \<equiv> cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr \<noteq> None"
adhoc_overloading is_element_kind is_element_kind\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t
lemma element_ptr_kinds_commutes [simp]:
lemma element_ptr_kinds_commutes [simp]:
"cast element_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> element_ptr |\<in>| element_ptr_kinds h"
apply(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)[1]
by (metis (no_types, lifting) element_ptr_casts_commute2 ffmember_filter fimage_eqI
fset.map_comp is_element_ptr_kind_none node_ptr_casts_commute3
by (metis (no_types, lifting) element_ptr_casts_commute2 ffmember_filter fimage_eqI
fset.map_comp is_element_ptr_kind_none node_ptr_casts_commute3
node_ptr_kinds_commutes node_ptr_kinds_def option.sel option.simps(3))
definition get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) heap \<Rightarrow> (_) Element option"
where
where
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h = Option.bind (get\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) h) cast"
adhoc_overloading get get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
@ -156,16 +164,16 @@ lemma get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf:
using l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv element_ptr_kinds_commutes notin_fset
option.distinct(1))
option.distinct(1))
end
global_interpretation l_get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf
by unfold_locales
definition put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) element_ptr \<Rightarrow> (_) Element \<Rightarrow> (_) heap \<Rightarrow> (_) heap"
where
where
"put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element = put\<^sub>N\<^sub>o\<^sub>d\<^sub>e (cast element_ptr) (cast element)"
adhoc_overloading put put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
adhoc_overloading put put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
assumes "put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h = h'"
@ -182,30 +190,30 @@ lemma put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_put_ptrs:
lemma cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]:
lemma cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inject [simp]:
"cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e x = cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e y \<longleftrightarrow> x = y"
apply(simp add: cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def)
by (metis (full_types) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_none [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = None \<longleftrightarrow> \<not> (\<exists>element. cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node)"
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
apply(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)[1]
by (metis (full_types) RNode.select_convs(2) RNode.surjective old.unit.exhaust)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_some [simp]:
"cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = Some element \<longleftrightarrow> cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element = node"
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
by(auto simp add: cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def RObject.extend_def RNode.extend_def
split: sum.splits)
lemma cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_inv [simp]: "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t (cast\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e element) = Some element"
by simp
lemma get_elment_ptr_simp1 [simp]:
lemma get_elment_ptr_simp1 [simp]:
"get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr element h) = Some element"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
lemma get_elment_ptr_simp2 [simp]:
"element_ptr \<noteq> element_ptr'
lemma get_elment_ptr_simp2 [simp]:
"element_ptr \<noteq> element_ptr'
\<Longrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr' element h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h"
by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
@ -217,9 +225,9 @@ abbreviation "create_element_obj tag_name_arg child_nodes_arg attrs_arg shadow_r
definition new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) heap \<Rightarrow> ((_) element_ptr \<times> (_) heap)"
where
"new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_element_ptr = element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref
|`| (element_ptrs h)))))
"new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_element_ptr = element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref
|`| (element_ptrs h)))))
in
(new_element_ptr, put new_element_ptr (create_element_obj '''' [] fmempty None) h))"
@ -230,7 +238,7 @@ lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap:
unfolding new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def
using put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap by blast
lemma new_element_ptr_new:
lemma new_element_ptr_new:
"element_ptr.Ref (Suc (fMax (finsert 0 (element_ptr.the_ref |`| element_ptrs h)))) |\<notin>| element_ptrs h"
by (metis Suc_n_not_le_n element_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert)
@ -293,22 +301,27 @@ definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
lemma known_ptrs_known_ptr:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> known_ptr ptr"
apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD)
lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def by blast
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def
by blast
end

View File

@ -23,31 +23,31 @@
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>ShadowRoot\<close>
text\<open>In this theory, we introduce the typed pointers for the class ShadowRoot. Note that, in
this document, we will not make use of ShadowRoots nor will we discuss their particular properties.
text\<open>In this theory, we introduce the typed pointers for the class ShadowRoot. Note that, in
this document, we will not make use of ShadowRoots nor will we discuss their particular properties.
We only include them here, as they are required for future work and they cannot be added alter
following the object-oriented extensibility of our data model.\<close>
following the object-oriented extensibility of our data model.\<close>
theory ShadowRootPointer
imports
"DocumentPointer"
begin
datatype 'shadow_root_ptr shadow_root_ptr = Ref (the_ref: ref) | Ext 'shadow_root_ptr
register_default_tvars "'shadow_root_ptr shadow_root_ptr"
register_default_tvars "'shadow_root_ptr shadow_root_ptr"
type_synonym ('document_ptr, 'shadow_root_ptr) document_ptr
= "('shadow_root_ptr shadow_root_ptr + 'document_ptr) document_ptr"
register_default_tvars "('document_ptr, 'shadow_root_ptr) document_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
register_default_tvars "('document_ptr, 'shadow_root_ptr) document_ptr"
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr
= "('object_ptr, 'node_ptr, 'element_ptr,
= "('object_ptr, 'node_ptr, 'element_ptr,
'character_data_ptr, 'shadow_root_ptr shadow_root_ptr + 'document_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr"
definition cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr"
@ -64,28 +64,29 @@ abbreviation cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>
definition cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) document_ptr \<Rightarrow> (_) shadow_root_ptr option"
where
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = (case document_ptr of document_ptr.Ext (Inl shadow_root_ptr)
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = (case document_ptr of document_ptr.Ext (Inl shadow_root_ptr)
\<Rightarrow> Some shadow_root_ptr | _ \<Rightarrow> None)"
abbreviation cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> (_) shadow_root_ptr option"
where
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
Some document_ptr \<Rightarrow> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr
"cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of
Some document_ptr \<Rightarrow> cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr
| None \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
consts is_shadow_root_ptr_kind :: 'a
definition is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) document_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some _ \<Rightarrow> True | _ \<Rightarrow> False)"
"is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr =
(case cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some _ \<Rightarrow> True | _ \<Rightarrow> False)"
abbreviation is_shadow_root_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some document_ptr \<Rightarrow> is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr
| None \<Rightarrow> False)"
"is_shadow_root_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some document_ptr \<Rightarrow> is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr
| None \<Rightarrow> False)"
adhoc_overloading is_shadow_root_ptr_kind is_shadow_root_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
lemmas is_shadow_root_ptr_kind_def = is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
@ -93,44 +94,47 @@ lemmas is_shadow_root_ptr_kind_def = is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^s
consts is_shadow_root_ptr :: 'a
definition is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of shadow_root_ptr.Ref _ \<Rightarrow> True
"is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case ptr of shadow_root_ptr.Ref _ \<Rightarrow> True
| _ \<Rightarrow> False)"
abbreviation is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) document_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some shadow_root_ptr \<Rightarrow> is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr
| _ \<Rightarrow> False)"
"is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some shadow_root_ptr \<Rightarrow> is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr
| _ \<Rightarrow> False)"
abbreviation is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
where
"is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some document_ptr \<Rightarrow> is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr
"is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> (case cast ptr of
Some document_ptr \<Rightarrow> is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr
| None \<Rightarrow> False)"
adhoc_overloading is_shadow_root_ptr is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
adhoc_overloading is_shadow_root_ptr is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
lemmas is_shadow_root_ptr_def = is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
consts is_shadow_root_ptr_ext :: 'a
abbreviation "is_shadow_root_ptr_ext\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> \<not> is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
abbreviation "is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)"
abbreviation "is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv>
is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)"
abbreviation "is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)"
abbreviation "is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv>
is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)"
adhoc_overloading is_shadow_root_ptr_ext is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
instantiation shadow_root_ptr :: (linorder) linorder
begin
definition
definition
less_eq_shadow_root_ptr :: "(_::linorder) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> bool"
where
where
"less_eq_shadow_root_ptr x y \<equiv> (case x of Ext i \<Rightarrow> (case y of Ext j \<Rightarrow> i \<le> j | Ref _ \<Rightarrow> False)
| Ref i \<Rightarrow> (case y of Ext _ \<Rightarrow> True | Ref j \<Rightarrow> i \<le> j))"
definition less_shadow_root_ptr :: "(_::linorder) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr \<Rightarrow> bool"
where "less_shadow_root_ptr x y \<equiv> x \<le> y \<and> \<not> y \<le> x"
instance
instance
apply(standard)
by(auto simp add: less_eq_shadow_root_ptr_def less_shadow_root_ptr_def
by(auto simp add: less_eq_shadow_root_ptr_def less_shadow_root_ptr_def
split: shadow_root_ptr.splits)
end
@ -139,11 +143,12 @@ lemma is_shadow_root_ptr_ref [simp]: "is_shadow_root_ptr (shadow_root_ptr.Ref n)
by(simp add: is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma shadow_root_ptr_casts_commute [simp]:
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = Some shadow_root_ptr \<longleftrightarrow> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = document_ptr"
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr =
Some shadow_root_ptr \<longleftrightarrow> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = document_ptr"
unfolding cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: document_ptr.splits sum.splits)
lemma shadow_root_ptr_casts_commute2 [simp]:
lemma shadow_root_ptr_casts_commute2 [simp]:
"(cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr) = Some shadow_root_ptr)"
by simp
@ -151,7 +156,7 @@ lemma shadow_root_ptr_casts_commute3 [simp]:
assumes "is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"
shows "cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (the (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)) = document_ptr"
using assms
by(auto simp add: is_shadow_root_ptr_kind_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto simp add: is_shadow_root_ptr_kind_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
split: document_ptr.splits sum.splits)
lemma is_shadow_root_ptr_kind_obtains:
@ -166,19 +171,20 @@ lemma is_shadow_root_ptr_kind_none:
unfolding is_shadow_root_ptr_kind_def cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: document_ptr.splits sum.splits)
lemma is_shadow_root_ptr_kind_cast [simp]:
lemma is_shadow_root_ptr_kind_cast [simp]:
"is_shadow_root_ptr_kind (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr)"
by (metis shadow_root_ptr_casts_commute is_shadow_root_ptr_kind_none option.distinct(1))
lemma cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
lemma cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject [simp]:
"cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r x = cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r y \<longleftrightarrow> x = y"
by(simp add: cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
lemma cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_ext_none [simp]:
"cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (document_ptr.Ext (Inr (Inr document_ext_ptr))) = None"
by(simp add: cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma is_shadow_root_ptr_implies_kind [dest]: "is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<Longrightarrow> is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
lemma is_shadow_root_ptr_implies_kind [dest]:
"is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<Longrightarrow> is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
by(auto split: option.splits)
lemma is_shadow_root_ptr_kind_not_document_ptr [simp]: "\<not>is_shadow_root_ptr_kind (document_ptr.Ref x)"