src/HOL/BNF/Examples/Derivation_Trees/DTree.thy
author wenzelm
Thu, 14 Feb 2013 16:25:13 +0100
changeset 51120 4b3a062b6538
parent 50530 6266e44b3396
child 51410 f0865a641e76
permissions -rw-r--r--
obsolete;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50530
6266e44b3396 updated some headers;
wenzelm
parents: 49882
diff changeset
     1
(*  Title:      HOL/BNF/Examples/Derivation_Trees/DTree.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
     5
Derivation trees with nonterminal internal nodes and terminal leaves.
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
     8
header {* Trees with Nonterminal Internal Nodes and Terminal Leaves *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    10
theory DTree
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
imports Prelim
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    14
hide_fact (open) Quotient_Product.prod_rel_def
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49222
diff changeset
    15
49631
9ce0c2cbadb8 modernized example
blanchet
parents: 49594
diff changeset
    16
typedecl N
9ce0c2cbadb8 modernized example
blanchet
parents: 49594
diff changeset
    17
typedecl T
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    19
codata dtree = NNode (root: N) (ccont: "(T + dtree) fset")
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
49882
946efb120c42 tuned for document output
traytel
parents: 49879
diff changeset
    21
subsection{* Transporting the Characteristic Lemmas from @{text "fset"} to @{text "set"} *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
definition "Node n as \<equiv> NNode n (the_inv fset as)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
definition "cont \<equiv> fset o ccont"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    25
definition "unfold rt ct \<equiv> dtree_unfold rt (the_inv fset o ct)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    26
definition "corec rt qt ct dt \<equiv> dtree_corec rt qt (the_inv fset o ct) (the_inv fset o dt)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    28
lemma finite_cont[simp]: "finite (cont tr)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
unfolding cont_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    31
lemma Node_root_cont[simp]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
"Node (root tr) (cont tr) = tr"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    33
using dtree.collapse unfolding Node_def cont_def
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
by (metis cont_def finite_cont fset_cong fset_to_fset o_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    36
lemma dtree_simps[simp]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
assumes "finite as" and "finite as'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
shows "Node n as = Node n' as' \<longleftrightarrow> n = n' \<and> as = as'"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    39
using assms dtree.inject unfolding Node_def
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
by (metis fset_to_fset)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    42
lemma dtree_cases[elim, case_names Node Choice]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
assumes Node: "\<And> n as. \<lbrakk>finite as; tr = Node n as\<rbrakk> \<Longrightarrow> phi"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
shows phi
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    45
apply(cases rule: dtree.exhaust[of tr])
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
using Node unfolding Node_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
by (metis Node Node_root_cont finite_cont)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    49
lemma dtree_sel_ctor[simp]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    50
"root (Node n as) = n"
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    51
"finite as \<Longrightarrow> cont (Node n as) = as"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
unfolding Node_def cont_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    54
lemmas root_Node = dtree_sel_ctor(1)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    55
lemmas cont_Node = dtree_sel_ctor(2)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    57
lemma dtree_cong:
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    58
assumes "root tr = root tr'" and "cont tr = cont tr'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    59
shows "tr = tr'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    60
by (metis Node_root_cont assms)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    62
lemma set_rel_cont:
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    63
"set_rel \<chi> (cont tr1) (cont tr2) = fset_rel \<chi> (ccont tr1) (ccont tr2)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    64
unfolding cont_def comp_def fset_rel_fset ..
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    65
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    66
lemma dtree_coinduct[elim, consumes 1, case_names Lift, induct pred: "HOL.eq"]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    67
assumes phi: "\<phi> tr1 tr2" and
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    68
Lift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    69
                  root tr1 = root tr2 \<and> set_rel (sum_rel op = \<phi>) (cont tr1) (cont tr2)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
shows "tr1 = tr2"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    71
using phi apply(elim dtree.coinduct)
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    72
apply(rule Lift[unfolded set_rel_cont]) .
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    73
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    74
lemma unfold:
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    75
"root (unfold rt ct b) = rt b"
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    76
"finite (ct b) \<Longrightarrow> cont (unfold rt ct b) = image (id \<oplus> unfold rt ct) (ct b)"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    77
using dtree.sel_unfold[of rt "the_inv fset \<circ> ct" b] unfolding unfold_def
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
apply - apply metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
unfolding cont_def comp_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
by (metis (no_types) fset_to_fset map_fset_image)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    82
lemma corec:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49631
diff changeset
    83
"root (corec rt qt ct dt b) = rt b"
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49631
diff changeset
    84
"\<lbrakk>finite (ct b); finite (dt b)\<rbrakk> \<Longrightarrow>
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49631
diff changeset
    85
 cont (corec rt qt ct dt b) =
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49631
diff changeset
    86
 (if qt b then ct b else image (id \<oplus> corec rt qt ct dt) (dt b))"
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    87
using dtree.sel_corec[of rt qt "the_inv fset \<circ> ct" "the_inv fset \<circ> dt" b]
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    88
unfolding corec_def
49631
9ce0c2cbadb8 modernized example
blanchet
parents: 49594
diff changeset
    89
apply -
9ce0c2cbadb8 modernized example
blanchet
parents: 49594
diff changeset
    90
apply simp
9ce0c2cbadb8 modernized example
blanchet
parents: 49594
diff changeset
    91
unfolding cont_def comp_def id_def
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
by (metis (no_types) fset_to_fset map_fset_image)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
end