Theory Document_adoptNode
section‹Testing Document\_adoptNode›
text‹This theory contains the test cases for Document\_adoptNode.›
theory Document_adoptNode
imports
"Core_DOM_BaseTest"
begin
definition Document_adoptNode_heap :: heap⇩f⇩i⇩n⇩a⇩l where
"Document_adoptNode_heap = create_heap [(cast (document_ptr.Ref 1), cast (create_document_obj html (Some (cast (element_ptr.Ref 1))) [])),
(cast (element_ptr.Ref 1), cast (create_element_obj ''html'' [cast (element_ptr.Ref 2), cast (element_ptr.Ref 8)] fmempty None)),
(cast (element_ptr.Ref 2), cast (create_element_obj ''head'' [cast (element_ptr.Ref 3), cast (element_ptr.Ref 4), cast (element_ptr.Ref 5), cast (element_ptr.Ref 6), cast (element_ptr.Ref 7)] fmempty None)),
(cast (element_ptr.Ref 3), cast (create_element_obj ''meta'' [] (fmap_of_list [(''charset'', ''utf-8'')]) None)),
(cast (element_ptr.Ref 4), cast (create_element_obj ''title'' [cast (character_data_ptr.Ref 1)] fmempty None)),
(cast (character_data_ptr.Ref 1), cast (create_character_data_obj ''Document.adoptNode'')),
(cast (element_ptr.Ref 5), cast (create_element_obj ''link'' [] (fmap_of_list [(''rel'', ''help''), (''href'', ''https://dom.spec.whatwg.org/#dom-document-adoptnode'')]) None)),
(cast (element_ptr.Ref 6), cast (create_element_obj ''script'' [] (fmap_of_list [(''src'', ''/resources/testharness.js'')]) None)),
(cast (element_ptr.Ref 7), cast (create_element_obj ''script'' [] (fmap_of_list [(''src'', ''/resources/testharnessreport.js'')]) None)),
(cast (element_ptr.Ref 8), cast (create_element_obj ''body'' [cast (element_ptr.Ref 9), cast (element_ptr.Ref 10), cast (element_ptr.Ref 11)] fmempty None)),
(cast (element_ptr.Ref 9), cast (create_element_obj ''div'' [] (fmap_of_list [(''id'', ''log'')]) None)),
(cast (element_ptr.Ref 10), cast (create_element_obj ''x<'' [cast (character_data_ptr.Ref 2)] fmempty None)),
(cast (character_data_ptr.Ref 2), cast (create_character_data_obj ''x'')),
(cast (element_ptr.Ref 11), cast (create_element_obj ''script'' [cast (character_data_ptr.Ref 3)] fmempty None)),
(cast (character_data_ptr.Ref 3), cast (create_character_data_obj ''%3C%3Cscript%3E%3E''))]"
definition Document_adoptNode_document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "Document_adoptNode_document = Some (cast (document_ptr.Ref 1))"
text ‹"Adopting an Element called 'x<' should work."›
lemma "test (do {
tmp0 ← Document_adoptNode_document . getElementsByTagName(''x<'');
y ← return (tmp0 ! 0);
child ← y . firstChild;
tmp1 ← y . parentNode;
tmp2 ← Document_adoptNode_document . body;
assert_equals(tmp1, tmp2);
tmp3 ← y . ownerDocument;
assert_equals(tmp3, Document_adoptNode_document);
tmp4 ← Document_adoptNode_document . adoptNode(y);
assert_equals(tmp4, y);
tmp5 ← y . parentNode;
assert_equals(tmp5, None);
tmp6 ← y . firstChild;
assert_equals(tmp6, child);
tmp7 ← y . ownerDocument;
assert_equals(tmp7, Document_adoptNode_document);
tmp8 ← child . ownerDocument;
assert_equals(tmp8, Document_adoptNode_document);
doc ← createDocument(None, None, None);
tmp9 ← doc . adoptNode(y);
assert_equals(tmp9, y);
tmp10 ← y . parentNode;
assert_equals(tmp10, None);
tmp11 ← y . firstChild;
assert_equals(tmp11, child);
tmp12 ← y . ownerDocument;
assert_equals(tmp12, doc);
tmp13 ← child . ownerDocument;
assert_equals(tmp13, doc)
}) Document_adoptNode_heap"
by eval
text ‹"Adopting an Element called ':good:times:' should work."›
lemma "test (do {
x ← Document_adoptNode_document . createElement('':good:times:'');
tmp0 ← Document_adoptNode_document . adoptNode(x);
assert_equals(tmp0, x);
doc ← createDocument(None, None, None);
tmp1 ← doc . adoptNode(x);
assert_equals(tmp1, x);
tmp2 ← x . parentNode;
assert_equals(tmp2, None);
tmp3 ← x . ownerDocument;
assert_equals(tmp3, doc)
}) Document_adoptNode_heap"
by eval
end