Theory CharacterDataMonad

(***********************************************************************************
 * Copyright (c) 2016-2018 The University of Sheffield, UK
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are met:
 *
 * * Redistributions of source code must retain the above copyright notice, this
 *   list of conditions and the following disclaimer.
 *
 * * Redistributions in binary form must reproduce the above copyright notice,
 *   this list of conditions and the following disclaimer in the documentation
 *   and/or other materials provided with the distribution.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
 * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
 * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
 * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
 * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
 * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
 * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
 * 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‹CharacterData›
text‹In this theory, we introduce the monadic method setup for the CharacterData class.›
theory CharacterDataMonad
  imports
    ElementMonad
    "../classes/CharacterDataClass"
begin

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" 


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  node_ptr_kinds_M|r = |h'  node_ptr_kinds_M|r"
  shows "|h  character_data_ptr_kinds_M|r = |h'  character_data_ptr_kinds_M|r"
  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: 
  "reads (node_ptr. {preserved (get_MObject 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 
      character_data_ptr_kinds_def preserved_def)
  by (smt (verit) node_ptr_kinds_small preserved_def unit_all_impI)

global_interpretation l_dummy defines get_MCharacterData = "l_get_M.a_get_M getCharacterData" .
lemma get_M_is_l_get_M: "l_get_M getCharacterData type_wf character_data_ptr_kinds"
  apply(simp add: getCharacterData_type_wf l_get_M_def)
  by (metis (no_types, opaque_lifting) NodeMonad.get_M_is_l_get_M bind_eq_Some_conv 
      character_data_ptr_kinds_commutes getCharacterData_def l_get_M_def option.distinct(1))
lemmas get_M_defs = get_MCharacterData_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]

adhoc_overloading get_M get_MCharacterData

locale l_get_MCharacterData_lemmas = l_type_wfCharacterData
begin
sublocale l_get_MElement_lemmas by unfold_locales

interpretation l_get_M getCharacterData type_wf character_data_ptr_kinds
  apply(unfold_locales)
   apply (simp add: getCharacterData_type_wf local.type_wfCharacterData)
  by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_MCharacterData_ok = get_M_ok[folded get_MCharacterData_def]
end

global_interpretation l_get_MCharacterData_lemmas type_wf by unfold_locales


global_interpretation l_put_M type_wf character_data_ptr_kinds getCharacterData putCharacterData 
  rewrites "a_get_M = get_MCharacterData" defines put_MCharacterData = a_put_M
   apply (simp add: get_M_is_l_get_M l_put_M_def)
  by (simp add: get_MCharacterData_def)

lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_MCharacterData


locale l_put_MCharacterData_lemmas = l_type_wfCharacterData
begin
sublocale l_put_MElement_lemmas by unfold_locales

interpretation l_put_M type_wf character_data_ptr_kinds getCharacterData putCharacterData
  apply(unfold_locales)
  using getCharacterData_type_wf l_type_wfCharacterData.type_wfCharacterData local.l_type_wfCharacterData_axioms
   apply blast
  by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_MCharacterData_ok = put_M_ok[folded put_MCharacterData_def]
end

global_interpretation l_put_MCharacterData_lemmas type_wf by unfold_locales



lemma CharacterData_simp1 [simp]: 
  "(x. getter (setter (λ_. v) x) = v)  h  put_MCharacterData character_data_ptr setter v h h' 
    h'  get_MCharacterData character_data_ptr getter r v"
  by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma CharacterData_simp2 [simp]: 
  "character_data_ptr  character_data_ptr' 
    h  put_MCharacterData character_data_ptr setter v h h' 
   preserved (get_MCharacterData 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]: "
  (x. getter (setter (λ_. v) x) = getter x) 
   h  put_MCharacterData character_data_ptr setter v h h' 
   preserved (get_MCharacterData character_data_ptr' getter) h h'"
  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  put_MCharacterData character_data_ptr setter v h h' 
   preserved (get_MElement element_ptr getter) h h'"
  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  put_MElement element_ptr setter v h h' 
   preserved (get_MCharacterData character_data_ptr getter) h h'"
  by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def 
      split: option.splits dest: get_heap_E)
lemma CharacterData_simp6 [simp]: 
  "(x. getter (cast (setter (λ_. v) x)) = getter (cast x)) 
    h  put_MCharacterData character_data_ptr setter v h h'
   preserved (get_MObject 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 
      getCharacterData_def getNode_def preserved_def putCharacterData_def putNode_def 
      bind_eq_Some_conv split: option.splits)
lemma CharacterData_simp7 [simp]: 
  "(x. getter (cast (setter (λ_. v) x)) = getter (cast x)) 
   h  put_MCharacterData character_data_ptr setter v h h' 
   preserved (get_MNode 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 
      getCharacterData_def getNode_def preserved_def putCharacterData_def putNode_def 
      bind_eq_Some_conv split: option.splits)

lemma CharacterData_simp8 [simp]: 
  "cast character_data_ptr  node_ptr 
    h  put_MCharacterData character_data_ptr setter v h h' 
    preserved (get_MNode node_ptr getter) h h'"
  by(auto simp add: put_M_defs get_M_defs getCharacterData_def putCharacterData_def NodeMonad.get_M_defs 
      preserved_def split: option.splits dest: get_heap_E)
lemma CharacterData_simp9 [simp]: 
  "h  put_MCharacterData character_data_ptr setter v h h' 
    (x. getter (cast (setter (λ_. v) x)) = getter (cast x)) 
    preserved (get_MNode node_ptr getter) h h'"
  apply(cases "cast character_data_ptr  node_ptr")
  by(auto simp add: put_M_defs get_M_defs getCharacterData_def putCharacterData_def 
      NodeMonad.get_M_defs preserved_def split: option.splits bind_splits 
      dest: get_heap_E)
lemma CharacterData_simp10 [simp]: 
  "cast character_data_ptr  node_ptr 
    h  put_MNode node_ptr setter v h h' 
    preserved (get_MCharacterData character_data_ptr getter) h h'"
  by(auto simp add: NodeMonad.put_M_defs get_M_defs getCharacterData_def NodeMonad.get_M_defs 
      preserved_def split: option.splits dest: get_heap_E)

lemma CharacterData_simp11 [simp]: 
  "cast character_data_ptr  object_ptr 
   h  put_MCharacterData character_data_ptr setter v h h' 
   preserved (get_MObject object_ptr getter) h h'"
  by(auto simp add: put_M_defs get_M_defs getCharacterData_def getNode_def putNode_def putCharacterData_def 
      ObjectMonad.get_M_defs preserved_def 
      split: option.splits dest: get_heap_E)

lemma CharacterData_simp12 [simp]:
  "h  put_MCharacterData character_data_ptr setter v h h' 
    (x. getter (cast (setter (λ_. v) x)) = getter (cast x)) 
    preserved (get_MObject object_ptr getter) h h'"
  apply(cases "cast character_data_ptr  object_ptr")
   apply(auto simp add: put_M_defs get_M_defs getCharacterData_def putCharacterData_def 
      getNode_def putNode_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 getCharacterData_def putCharacterData_def 
      getNode_def putNode_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  object_ptr  h  put_MObject object_ptr setter v h h' 
     preserved (get_MCharacterData character_data_ptr getter) h h'"
  by(auto simp add: ObjectMonad.put_M_defs get_M_defs getCharacterData_def getNode_def putNode_def 
      ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E)

lemma new_element_get_MCharacterData: 
  "h  new_element h h'  preserved (get_MCharacterData ptr getter) h h'"
  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‹Creating CharacterData›

definition new_character_data :: "(_, (_) character_data_ptr) dom_prog"
  where
    "new_character_data = do {
      h  get_heap;
      (new_ptr, h')  return (newCharacterData h);
      return_heap h';
      return new_ptr
    }"

lemma new_character_data_ok [simp]:
  "h  ok new_character_data"
  by(auto simp add: new_character_data_def split: prod.splits)

lemma new_character_data_ptr_in_heap:
  assumes "h  new_character_data h h'"
    and "h  new_character_data r new_character_data_ptr"
  shows "new_character_data_ptr |∈| character_data_ptr_kinds h'"
  using assms
  unfolding new_character_data_def
  by(auto simp add: new_character_data_def newCharacterData_def Let_def putCharacterData_ptr_in_heap 
      is_OK_returns_result_I
      elim!: bind_returns_result_E bind_returns_heap_E)

lemma new_character_data_ptr_not_in_heap:
  assumes "h  new_character_data h h'"
    and "h  new_character_data r new_character_data_ptr"
  shows "new_character_data_ptr |∉| character_data_ptr_kinds h"
  using assms newCharacterData_ptr_not_in_heap
  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:
  assumes "h  new_character_data h h'"
    and "h  new_character_data r new_character_data_ptr"
  shows "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast new_character_data_ptr|}"
  using assms newCharacterData_new_ptr
  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:
  assumes "h  new_character_data r new_character_data_ptr"
  shows "is_character_data_ptr new_character_data_ptr"
  using assms newCharacterData_is_character_data_ptr
  by(auto simp add: new_character_data_def elim!: bind_returns_result_E split: prod.splits)

lemma new_character_data_child_nodes:
  assumes "h  new_character_data h h'"
  assumes "h  new_character_data r new_character_data_ptr"
  shows "h'  get_M new_character_data_ptr val r ''''"
  using assms
  by(auto simp add: get_M_defs new_character_data_def newCharacterData_def Let_def 
      split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)

lemma new_character_data_get_MObject: 
  "h  new_character_data h h'  h  new_character_data r new_character_data_ptr 
     ptr  cast new_character_data_ptr  preserved (get_MObject ptr getter) h h'"
  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_MNode: 
  "h  new_character_data h h'  h  new_character_data r new_character_data_ptr 
     ptr  cast new_character_data_ptr  preserved (get_MNode ptr getter) h h'"
  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_MElement: 
  "h  new_character_data h h'  h  new_character_data r new_character_data_ptr 
     preserved (get_MElement ptr getter) h h'"
  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_MCharacterData: 
  "h  new_character_data h h'  h  new_character_data r new_character_data_ptr 
     ptr  new_character_data_ptr  preserved (get_MCharacterData ptr getter) h h'"
  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‹Modified Heaps›

lemma get_CharacterData_ptr_simp [simp]: 
  "getCharacterData character_data_ptr (putObject ptr obj h) 
      = (if ptr = cast character_data_ptr then cast obj else get character_data_ptr h)"
  by(auto simp add: getCharacterData_def split: option.splits Option.bind_splits)

lemma Character_data_ptr_kinds_simp [simp]: 
  "character_data_ptr_kinds (putObject ptr obj h) = character_data_ptr_kinds h |∪| 
                            (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)

lemma type_wf_put_I:
  assumes "type_wf h"
  assumes "ElementClass.type_wf (putObject ptr obj h)"
  assumes "is_character_data_ptr_kind ptr  is_character_data_kind obj"
  shows "type_wf (putObject ptr obj h)"
  using assms
  by(auto simp add: type_wf_defs split: option.splits)

lemma type_wf_put_ptr_not_in_heap_E:
  assumes "type_wf (putObject ptr obj h)"
  assumes "ptr |∉| 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 
      split: option.splits if_splits)[1]
  using assms(2) node_ptr_kinds_commutes by blast

lemma type_wf_put_ptr_in_heap_E:
  assumes "type_wf (putObject ptr obj h)"
  assumes "ptr |∈| object_ptr_kinds h"
  assumes "ElementClass.type_wf h"
  assumes "is_character_data_ptr_kind ptr  is_character_data_kind (the (get ptr h))"
  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.getObject_type_wf assms(2) bind.bind_lunit
      castNode2CharacterData_inv castObject2Node_inv getCharacterData_def getNode_def option.collapse)

subsection‹Preserving Types›

lemma new_element_type_wf_preserved [simp]:
  assumes "h  new_element h h'"
  shows "type_wf h = type_wf h'"
  using assms
  apply(auto simp add: new_element_def newElement_def Let_def putElement_def putNode_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_wfElement assms new_element_type_wf_preserved apply blast
  using element_ptrs_def apply force
  using CharacterDataClass.type_wfElement 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 
      fimage_eqI is_element_ptr_ref)

lemma new_element_is_l_new_element: "l_new_element type_wf"
  using l_new_element.intro new_element_type_wf_preserved
  by blast

lemma put_MElement_tag_name_type_wf_preserved [simp]:
  "h  put_M element_ptr tag_name_update v h h'  type_wf h = type_wf h'"
  apply(auto simp add: ElementMonad.put_M_defs putElement_def putNode_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 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 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 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 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 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 getElement_def)
  apply metis
  done


lemma put_MElement_child_nodes_type_wf_preserved [simp]: 
  "h  put_M element_ptr child_nodes_update v h h'  type_wf h = type_wf h'"
  apply(auto simp add: ElementMonad.put_M_defs putElement_def putNode_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 
      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 
      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 
      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 
      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 
      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 getElement_def)
  apply metis
  done

lemma put_MElement_attrs_type_wf_preserved [simp]: 
  "h  put_M element_ptr attrs_update v h h'  type_wf h = type_wf h'"
  apply(auto simp add: ElementMonad.put_M_defs putElement_def putNode_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 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 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 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 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 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 getElement_def)
  apply metis
  done

lemma put_MElement_shadow_root_opt_type_wf_preserved [simp]: 
  "h  put_M element_ptr shadow_root_opt_update v h h'  type_wf h = type_wf h'"
  apply(auto simp add: ElementMonad.put_M_defs putElement_def putNode_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 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 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 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 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 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 getElement_def)
  apply metis
  done


lemma new_character_data_type_wf_preserved [simp]: 
  "h  new_character_data h h'  type_wf h = type_wf h'"
  apply(auto simp add: new_character_data_def newCharacterData_def Let_def putCharacterData_def putNode_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 newCharacterData_def newCharacterData_ptr_not_in_heap)

locale l_new_character_data = l_type_wf +
  assumes new_character_data_types_preserved: "h  new_character_data h h'  type_wf h = type_wf h'"

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_MCharacterData_val_type_wf_preserved [simp]: 
  "h  put_M character_data_ptr val_update v h h'  type_wf h = type_wf h'"
  apply(auto simp add: CharacterDataMonad.put_M_defs putCharacterData_def putNode_def 
      CharacterDataClass.type_wfNode CharacterDataClass.type_wfObject
      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 
      split: option.splits)[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.a_type_wf_def
      split: option.splits)[1]
   apply (metis (no_types, lifting) bind_eq_Some_conv getCharacterData_def)
  apply metis
  done

lemma character_data_ptr_kinds_small:
  assumes "object_ptr. preserved (get_MObject 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 
      object_ptr_kinds_preserved_small[OF assms])

lemma character_data_ptr_kinds_preserved:
  assumes "writes SW setter h h'"
  assumes "h  setter h h'"
  assumes "h h'. w  SW. h  w h h' 
           (object_ptr. preserved (get_MObject object_ptr RObject.nothing) h h')"
  shows "character_data_ptr_kinds h = character_data_ptr_kinds h'"
  using writes_small_big[OF assms]
  apply(simp add: reflp_def transp_def preserved_def character_data_ptr_kinds_def)
  by (metis assms node_ptr_kinds_preserved)


lemma type_wf_preserved_small:
  assumes "object_ptr. preserved (get_MObject object_ptr RObject.nothing) h h'"
  assumes "node_ptr. preserved (get_MNode node_ptr RNode.nothing) h h'"
  assumes "element_ptr. preserved (get_MElement element_ptr RElement.nothing) h h'"
  assumes "character_data_ptr. preserved (get_MCharacterData 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)]  
    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)] 
      split: option.splits)[1]
   apply(force)
  by force

lemma type_wf_preserved:
  assumes "writes SW setter h h'"
  assumes "h  setter h h'"
  assumes "h h' w. w  SW  h  w h h' 
           object_ptr. preserved (get_MObject object_ptr RObject.nothing) h h'"
  assumes "h h' w. w  SW  h  w h h' 
           node_ptr. preserved (get_MNode node_ptr RNode.nothing) h h'"
  assumes "h h' w. w  SW  h  w h h' 
           element_ptr. preserved (get_MElement element_ptr RElement.nothing) h h'"
  assumes "h h' w. w  SW  h  w h h' 
           character_data_ptr. preserved (get_MCharacterData character_data_ptr 
                                              RCharacterData.nothing) h h'"
  shows "type_wf h = type_wf h'"
proof -
  have "h h' w. w  SW  h  w h h'  type_wf h = type_wf h'"
    using assms type_wf_preserved_small by fast
  with assms(1) assms(2) show ?thesis
    apply(rule writes_small_big)
    by(auto simp add: reflp_def transp_def)
qed

lemma type_wf_drop: "type_wf h  type_wf (Heap (fmdrop ptr (the_heap h)))"
  apply(auto simp add: type_wf_def ElementMonad.type_wf_drop  
      l_type_wf_defCharacterData.a_type_wf_def)[1]
  using type_wf_drop
  by (metis (no_types, lifting) ElementClass.type_wfObject ObjectClass.getObject_type_wf
      character_data_ptr_kinds_commutes fmlookup_drop getCharacterData_def getNode_def
      getObject_def node_ptr_kinds_commutes object_ptr_kinds_code5)

end