Theory Core_DOM_BaseTest

(***********************************************************************************
 * 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‹Common Test Setup›
text‹This theory provides the common test setup that is used by all formalized test cases.›

theory Core_DOM_BaseTest
  imports
  (*<*) 
  "../preliminaries/Testing_Utils"
  (*>*) 
  "../Core_DOM"
begin

definition "assert_throws e p = do {
  h  get_heap;
  (if (h  p e e) then return () else error AssertException)
}"
notation assert_throws ("assert'_throws'(_, _')")

definition "test p h  h  ok p"


definition field_access :: "(string  (_, (_) object_ptr option) dom_prog)  string 
                                     (_, (_) object_ptr option) dom_prog"  (infix "." 80)
  where
    "field_access m field = m field"

definition assert_equals :: "'a  'a  (_, unit) dom_prog"
  where
    "assert_equals l r = (if l = r then return () else error AssertException)"
definition assert_equals_with_message :: "'a  'a  'b  (_, unit) dom_prog"
  where
    "assert_equals_with_message l r _ = (if l = r then return () else error AssertException)"
notation assert_equals ("assert'_equals'(_, _')")
notation assert_equals_with_message ("assert'_equals'(_, _, _')")
notation assert_equals ("assert'_array'_equals'(_, _')")
notation assert_equals_with_message ("assert'_array'_equals'(_, _, _')")

definition assert_not_equals :: "'a  'a  (_, unit) dom_prog"
  where
    "assert_not_equals l r = (if l  r then return () else error AssertException)"
definition assert_not_equals_with_message :: "'a  'a  'b  (_, unit) dom_prog"
  where
    "assert_not_equals_with_message l r _ = (if l  r then return () else error AssertException)"
notation assert_not_equals ("assert'_not'_equals'(_, _')")
notation assert_not_equals_with_message ("assert'_not'_equals'(_, _, _')")
notation assert_not_equals ("assert'_array'_not'_equals'(_, _')")
notation assert_not_equals_with_message ("assert'_array'_not'_equals'(_, _, _')")

definition removeWhiteSpaceOnlyTextNodes :: "((_) object_ptr option)  (_, unit) dom_prog"
  where
    "removeWhiteSpaceOnlyTextNodes _ = return ()"


subsection ‹Making the functions under test compatible with untyped languages such as JavaScript›

fun set_attribute_with_null :: "((_) object_ptr option)  attr_key  attr_value  (_, unit) dom_prog"
  where
    "set_attribute_with_null (Some ptr) k v = (case cast ptr of
      Some element_ptr  set_attribute element_ptr k (Some v))"
fun set_attribute_with_null2 :: "((_) object_ptr option)  attr_key  attr_value option  (_, unit) dom_prog"
  where
    "set_attribute_with_null2 (Some ptr) k v = (case cast ptr of
      Some element_ptr  set_attribute element_ptr k v)"
notation set_attribute_with_null ("_ . setAttribute'(_, _')")
notation set_attribute_with_null2 ("_ . setAttribute'(_, _')")

fun get_child_nodesCore_DOM_with_null :: "((_) object_ptr option)  (_, (_) object_ptr option list) dom_prog"
  where
    "get_child_nodesCore_DOM_with_null (Some ptr) = do {
      children  get_child_nodes ptr;
      return (map (Some  cast) children)
    }"
notation get_child_nodesCore_DOM_with_null ("_ . childNodes")

fun create_element_with_null :: "((_) object_ptr option)  string  (_, ((_) object_ptr option)) dom_prog"
  where
    "create_element_with_null (Some owner_document_obj) tag = (case cast owner_document_obj of
      Some owner_document  do {
        element_ptr  create_element owner_document tag;
        return (Some (cast element_ptr))})"
notation create_element_with_null ("_ . createElement'(_')")

fun create_character_data_with_null :: "((_) object_ptr option)  string  (_, ((_) object_ptr option)) dom_prog"
  where
    "create_character_data_with_null (Some owner_document_obj) tag = (case cast owner_document_obj of
      Some owner_document  do {
        character_data_ptr  create_character_data owner_document tag;
        return (Some (cast character_data_ptr))})"
notation create_character_data_with_null ("_ . createTextNode'(_')")

definition create_document_with_null :: "string  (_, ((_::linorder) object_ptr option)) dom_prog"
  where
    "create_document_with_null title = do {
      new_document_ptr  create_document;
      html  create_element new_document_ptr ''html'';
      append_child (cast new_document_ptr) (cast html);
      heap  create_element new_document_ptr ''heap'';
      append_child (cast html) (cast heap);
      body  create_element new_document_ptr ''body'';
      append_child (cast html) (cast body);
      return (Some (cast new_document_ptr))
    }"
abbreviation "create_document_with_null2 _ _ _  create_document_with_null ''''"
notation create_document_with_null ("createDocument'(_')")
notation create_document_with_null2 ("createDocument'(_, _, _')")

fun get_element_by_id_with_null :: "((_::linorder) object_ptr option)  string  (_, ((_) object_ptr option)) dom_prog"
  where                                                                     
    "get_element_by_id_with_null (Some ptr) id' = do {
        element_ptr_opt  get_element_by_id ptr id';
        (case element_ptr_opt of
          Some element_ptr  return (Some (castelement_ptr2object_ptr element_ptr))
        | None  return None)}"
  | "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)  string  (_, ((_) 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  map_M (return  Some  castelement_ptr2object_ptr)"
notation get_elements_by_class_name_with_null ("_ . getElementsByClassName'(_')")

fun get_elements_by_tag_name_with_null ::
"((_::linorder) object_ptr option)  string  (_, ((_) object_ptr option) list) dom_prog"
  where                                                                     
    "get_elements_by_tag_name_with_null (Some ptr) tag =
      get_elements_by_tag_name ptr tag  map_M (return  Some  castelement_ptr2object_ptr)"
notation get_elements_by_tag_name_with_null ("_ . getElementsByTagName'(_')")

fun insert_before_with_null ::
"((_::linorder) object_ptr option)  ((_) object_ptr option)  ((_) object_ptr option) 
(_, ((_) 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  do {
        (case ref_child_obj_opt of
          Some ref_child_obj  insert_before ptr child (cast ref_child_obj)
        | None  insert_before ptr child None);
        return (Some child_obj)}
    | None  error HierarchyRequestError)"
notation insert_before_with_null ("_ . insertBefore'(_, _')")

fun append_child_with_null :: "((_::linorder) object_ptr option)  ((_) object_ptr option) 
(_, unit) dom_prog"
  where
    "append_child_with_null (Some ptr) (Some child_obj) = (case cast child_obj of
      Some child  append_child ptr child
    | None  error SegmentationFault)"
notation append_child_with_null ("_ . appendChild'(_')")

fun get_body :: "((_::linorder) object_ptr option)  (_, ((_) object_ptr option)) dom_prog"
  where
    "get_body ptr = do {
       ptrs  ptr . getElementsByTagName(''body'');
       return (hd ptrs)
    }"
notation get_body ("_ . body")

fun get_document_element_with_null :: "((_::linorder) object_ptr option) 
(_, ((_) object_ptr option)) dom_prog"
  where
    "get_document_element_with_null (Some ptr) = (case castobject_ptr2document_ptr ptr of
    Some document_ptr  do {
      element_ptr_opt  get_M document_ptr document_element;
      return (case element_ptr_opt of
        Some element_ptr  Some (castelement_ptr2object_ptr element_ptr)
      | None  None)})"
notation get_document_element_with_null ("_ . documentElement")

fun get_owner_document_with_null :: "((_::linorder) object_ptr option) 
(_, ((_) object_ptr option)) dom_prog"
  where
    "get_owner_document_with_null (Some ptr) = (do {
        document_ptr  get_owner_document ptr;
        return (Some (castdocument_ptr2object_ptr document_ptr))})"
notation get_owner_document_with_null ("_ . ownerDocument")

fun remove_with_null :: "((_::linorder) object_ptr option)  ((_) object_ptr option) 
(_, ((_) object_ptr option)) dom_prog"
  where
    "remove_with_null (Some ptr) (Some child) = (case cast child of
      Some child_node  do {
        remove child_node;
        return (Some child)}
    | None  error NotFoundError)"
  | "remove_with_null None _ = error TypeError"
  | "remove_with_null _ None = error TypeError"
notation remove_with_null ("_ . remove'(')")

fun remove_child_with_null :: "((_::linorder) object_ptr option)  ((_) object_ptr option) 
(_, ((_) object_ptr option)) dom_prog"
  where
    "remove_child_with_null (Some ptr) (Some child) = (case cast child of
      Some child_node  do {
        remove_child ptr child_node;
        return (Some child)}
    | None  error NotFoundError)"
  | "remove_child_with_null None _ = error TypeError"
  | "remove_child_with_null _ None = error TypeError"
notation remove_child_with_null ("_ . removeChild")

fun get_tag_name_with_null :: "((_) object_ptr option)  (_, attr_value) dom_prog"
  where
    "get_tag_name_with_null (Some ptr) = (case cast ptr of
      Some element_ptr  get_M element_ptr tag_name)"
notation get_tag_name_with_null ("_ . tagName")

abbreviation "remove_attribute_with_null ptr k  set_attribute_with_null2 ptr k None"
notation remove_attribute_with_null ("_ . removeAttribute'(_')")

fun get_attribute_with_null :: "((_) object_ptr option)  attr_key  (_, attr_value option) dom_prog"
  where
    "get_attribute_with_null (Some ptr) k = (case cast ptr of 
      Some element_ptr  get_attribute element_ptr k)"
fun get_attribute_with_null2 :: "((_) object_ptr option)  attr_key  (_, attr_value) dom_prog"
  where
    "get_attribute_with_null2 (Some ptr) k = (case cast ptr of 
      Some element_ptr  do {
        a  get_attribute element_ptr k;
        return (the a)})"
notation get_attribute_with_null ("_ . getAttribute'(_')")
notation get_attribute_with_null2 ("_ . getAttribute'(_')")

fun get_parent_with_null :: "((_::linorder) object_ptr option)  (_, (_) object_ptr option) dom_prog"
  where
    "get_parent_with_null (Some ptr) = (case cast ptr of
      Some node_ptr  get_parent node_ptr)"
notation get_parent_with_null ("_ . parentNode")

fun first_child_with_null :: "((_) object_ptr option)  (_, ((_) object_ptr option)) dom_prog"
  where
    "first_child_with_null (Some ptr) = do {
      child_opt  first_child ptr;
      return (case child_opt of
        Some child  Some (cast child)
      | None  None)}"
notation first_child_with_null ("_ . firstChild")

fun adopt_node_with_null ::
"((_::linorder) object_ptr option)  ((_) object_ptr option) (_, ((_) object_ptr option)) dom_prog"
  where
    "adopt_node_with_null (Some ptr) (Some child) = (case cast ptr of
      Some document_ptr  (case cast child of
        Some child_node  do {
          adopt_node document_ptr child_node;
          return (Some child)}))"
notation adopt_node_with_null ("_ . adoptNode'(_')")
      

definition createTestTree ::
"((_::linorder) object_ptr option)  (_, (string   (_, ((_) object_ptr option)) dom_prog)) dom_prog"
  where
   "createTestTree ref = return (λid. get_element_by_id_with_null ref id)"

end