renamed 'BNF_Examples' to 'Datatype_Examples' (cf. 'datatypes.pdf')
authorblanchet
Thu, 11 Sep 2014 19:26:59 +0200
changeset 58309 a09ec6daaa19
parent 58308 0ccba1b6d00b
child 58310 91ea607a34d8
renamed 'BNF_Examples' to 'Datatype_Examples' (cf. 'datatypes.pdf')
src/Doc/Datatypes/Datatypes.thy
src/HOL/BNF_Examples/Brackin.thy
src/HOL/BNF_Examples/Compat.thy
src/HOL/BNF_Examples/Derivation_Trees/DTree.thy
src/HOL/BNF_Examples/Derivation_Trees/Gram_Lang.thy
src/HOL/BNF_Examples/Derivation_Trees/Parallel.thy
src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy
src/HOL/BNF_Examples/Instructions.thy
src/HOL/BNF_Examples/IsaFoR_Datatypes.thy
src/HOL/BNF_Examples/Koenig.thy
src/HOL/BNF_Examples/Lambda_Term.thy
src/HOL/BNF_Examples/Misc_Codatatype.thy
src/HOL/BNF_Examples/Misc_Datatype.thy
src/HOL/BNF_Examples/Misc_Primcorec.thy
src/HOL/BNF_Examples/Misc_Primrec.thy
src/HOL/BNF_Examples/Process.thy
src/HOL/BNF_Examples/SML.thy
src/HOL/BNF_Examples/Stream.thy
src/HOL/BNF_Examples/Stream_Processor.thy
src/HOL/BNF_Examples/TreeFI.thy
src/HOL/BNF_Examples/TreeFsetI.thy
src/HOL/BNF_Examples/Verilog.thy
src/HOL/Datatype_Examples/Brackin.thy
src/HOL/Datatype_Examples/Compat.thy
src/HOL/Datatype_Examples/Derivation_Trees/DTree.thy
src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy
src/HOL/Datatype_Examples/Derivation_Trees/Parallel.thy
src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy
src/HOL/Datatype_Examples/Instructions.thy
src/HOL/Datatype_Examples/IsaFoR_Datatypes.thy
src/HOL/Datatype_Examples/Koenig.thy
src/HOL/Datatype_Examples/Lambda_Term.thy
src/HOL/Datatype_Examples/Misc_Codatatype.thy
src/HOL/Datatype_Examples/Misc_Datatype.thy
src/HOL/Datatype_Examples/Misc_Primcorec.thy
src/HOL/Datatype_Examples/Misc_Primrec.thy
src/HOL/Datatype_Examples/Process.thy
src/HOL/Datatype_Examples/SML.thy
src/HOL/Datatype_Examples/Stream.thy
src/HOL/Datatype_Examples/Stream_Processor.thy
src/HOL/Datatype_Examples/TreeFI.thy
src/HOL/Datatype_Examples/TreeFsetI.thy
src/HOL/Datatype_Examples/Verilog.thy
src/HOL/ROOT
--- a/src/Doc/Datatypes/Datatypes.thy	Thu Sep 11 19:20:23 2014 +0200
+++ b/src/Doc/Datatypes/Datatypes.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -1112,7 +1112,7 @@
 text {*
 Primitive recursion is illustrated through concrete examples based on the
 datatypes defined in Section~\ref{ssec:datatype-introductory-examples}. More
-examples can be found in the directory \verb|~~/src/HOL/BNF_Examples|.
+examples can be found in the directory \verb|~~/src/HOL/Datatype_Examples|.
 *}
 
 
@@ -1847,7 +1847,7 @@
 \keyw{prim\-corec\-ursive} commands, which support primitive corecursion, or
 using the more general \keyw{partial_function} command. Here, the focus is on
 the first two. More examples can be found in the directory
-\verb|~~/src/HOL/BNF_Examples|.
+\verb|~~/src/HOL/Datatype_Examples|.
 
 Whereas recursive functions consume datatypes one constructor at a time,
 corecursive functions construct codatatypes one constructor at a time.
@@ -1891,8 +1891,8 @@
 text {*
 Primitive corecursion is illustrated through concrete examples based on the
 codatatypes defined in Section~\ref{ssec:codatatype-introductory-examples}. More
-examples can be found in the directory \verb|~~/src/HOL/BNF_Examples|. The code
-view is favored in the examples below. Sections
+examples can be found in the directory \verb|~~/src/HOL/Datatype_Examples|. The
+code view is favored in the examples below. Sections
 \ref{ssec:primrec-constructor-view} and \ref{ssec:primrec-destructor-view}
 present the same examples expressed using the constructor and destructor views.
 *}
--- a/src/HOL/BNF_Examples/Brackin.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-(*  Title:      HOL/Datatype_Benchmark/Brackin.thy
-
-A couple of datatypes from Steve Brackin's work.
-*)
-
-theory Brackin imports Main begin
-
-datatype T =
-    X1 | X2 | X3 | X4 | X5 | X6 | X7 | X8 | X9 | X10 | X11 |
-    X12 | X13 | X14 | X15 | X16 | X17 | X18 | X19 | X20 | X21 |
-    X22 | X23 | X24 | X25 | X26 | X27 | X28 | X29 | X30 | X31 |
-    X32 | X33 | X34
-
-datatype ('a, 'b, 'c) TY1 =
-      NoF
-    | Fk 'a "('a, 'b, 'c) TY2"
-and ('a, 'b, 'c) TY2 =
-      Ta bool
-    | Td bool
-    | Tf "('a, 'b, 'c) TY1"
-    | Tk bool
-    | Tp bool
-    | App 'a "('a, 'b, 'c) TY1" "('a, 'b, 'c) TY2" "('a, 'b, 'c) TY3"
-    | Pair "('a, 'b, 'c) TY2" "('a, 'b, 'c) TY2"
-and ('a, 'b, 'c) TY3 =
-      NoS
-    | Fresh "('a, 'b, 'c) TY2"
-    | Trustworthy 'a
-    | PrivateKey 'a 'b 'c
-    | PublicKey 'a 'b 'c
-    | Conveyed 'a "('a, 'b, 'c) TY2"
-    | Possesses 'a "('a, 'b, 'c) TY2"
-    | Received 'a "('a, 'b, 'c) TY2"
-    | Recognizes 'a "('a, 'b, 'c) TY2"
-    | NeverMalFromSelf 'a 'b "('a, 'b, 'c) TY2"
-    | Sends 'a "('a, 'b, 'c) TY2" 'b
-    | SharedSecret 'a "('a, 'b, 'c) TY2" 'b
-    | Believes 'a "('a, 'b, 'c) TY3"
-    | And "('a, 'b, 'c) TY3" "('a, 'b, 'c) TY3"
-
-end
--- a/src/HOL/BNF_Examples/Compat.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,236 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Compat.thy
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2014
-
-Tests for compatibility with the old datatype package.
-*)
-
-header \<open> Tests for Compatibility with the Old Datatype Package \<close>
-
-theory Compat
-imports Main
-begin
-
-subsection \<open> Viewing and Registering New-Style Datatypes as Old-Style Ones \<close>
-
-ML \<open>
-fun check_len n xs label =
-  length xs = n orelse error ("Expected length " ^ string_of_int (length xs) ^ " for " ^ label);
-
-fun check_lens (n1, n2, n3) (xs1, xs2, xs3) =
-  check_len n1 xs1 "old" andalso check_len n2 xs2 "unfold" andalso check_len n3 xs3 "keep";
-
-fun get_descrs thy lens T_name =
-  (these (Option.map #descr (Old_Datatype_Data.get_info thy T_name)),
-   these (Option.map #descr (BNF_LFP_Compat.get_info thy BNF_LFP_Compat.Unfold_Nesting T_name)),
-   these (Option.map #descr (BNF_LFP_Compat.get_info thy BNF_LFP_Compat.Keep_Nesting T_name)))
-  |> tap (check_lens lens);
-\<close>
-
-old_datatype 'a old_lst = Old_Nl | Old_Cns 'a "'a old_lst"
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name old_lst}; \<close>
-
-datatype_new 'a lst = Nl | Cns 'a "'a lst"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name lst}; \<close>
-
-datatype_compat lst
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name lst}; \<close>
-
-datatype_new 'b w = W | W' "'b w \<times> 'b list"
-
-(* no support for sums of products:
-datatype_compat w
-*)
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name w}; \<close>
-
-datatype_new ('c, 'b) s = L 'c | R 'b
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name s}; \<close>
-
-datatype_new 'd x = X | X' "('d x lst, 'd list) s"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name x}; \<close>
-
-datatype_compat s
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name s}; \<close>
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name x}; \<close>
-
-datatype_compat x
-
-ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name x}; \<close>
-
-thm x.induct x.rec
-thm compat_x.induct compat_x.rec
-
-datatype_new 'a tttre = TTTre 'a "'a tttre lst lst lst"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name tttre}; \<close>
-
-datatype_compat tttre
-
-ML \<open> get_descrs @{theory} (4, 4, 1) @{type_name tttre}; \<close>
-
-thm tttre.induct tttre.rec
-thm compat_tttre.induct compat_tttre.rec
-
-datatype_new 'a ftre = FEmp | FTre "'a \<Rightarrow> 'a ftre lst"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name ftre}; \<close>
-
-datatype_compat ftre
-
-ML \<open> get_descrs @{theory} (2, 2, 1) @{type_name ftre}; \<close>
-
-thm ftre.induct ftre.rec
-thm compat_ftre.induct compat_ftre.rec
-
-datatype_new 'a btre = BTre 'a "'a btre lst" "'a btre lst"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name btre}; \<close>
-
-datatype_compat btre
-
-ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name btre}; \<close>
-
-thm btre.induct btre.rec
-thm compat_btre.induct compat_btre.rec
-
-datatype_new 'a foo = Foo | Foo' 'a "'a bar" and 'a bar = Bar | Bar' 'a "'a foo"
-
-ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name foo}; \<close>
-ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name bar}; \<close>
-
-datatype_compat foo bar
-
-ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name foo}; \<close>
-ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name bar}; \<close>
-
-datatype_new 'a tre = Tre 'a "'a tre lst"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name tre}; \<close>
-
-datatype_compat tre
-
-ML \<open> get_descrs @{theory} (2, 2, 1) @{type_name tre}; \<close>
-
-thm tre.induct tre.rec
-thm compat_tre.induct compat_tre.rec
-
-datatype_new 'a f = F 'a and 'a g = G 'a
-
-ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name f}; \<close>
-ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name g}; \<close>
-
-datatype_new h = H "h f" | H'
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name h}; \<close>
-
-datatype_compat f g
-
-ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name f}; \<close>
-ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name g}; \<close>
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name h}; \<close>
-
-datatype_compat h
-
-ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name h}; \<close>
-
-thm h.induct h.rec
-thm compat_h.induct compat_h.rec
-
-datatype_new myunit = MyUnity
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name myunit}; \<close>
-
-datatype_compat myunit
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name myunit}; \<close>
-
-datatype_new mylist = MyNil | MyCons nat mylist
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name mylist}; \<close>
-
-datatype_compat mylist
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name mylist}; \<close>
-
-datatype_new foo' = FooNil | FooCons bar' foo' and bar' = Bar
-
-ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name foo'}; \<close>
-ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name bar'}; \<close>
-
-datatype_compat bar' foo'
-
-ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name foo'}; \<close>
-ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name bar'}; \<close>
-
-old_datatype funky = Funky "funky tre" | Funky'
-
-ML \<open> get_descrs @{theory} (3, 3, 3) @{type_name funky}; \<close>
-
-old_datatype fnky = Fnky "nat tre"
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name fnky}; \<close>
-
-datatype_new tree = Tree "tree foo"
-
-ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name tree}; \<close>
-
-datatype_compat tree
-
-ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name tree}; \<close>
-
-thm tree.induct tree.rec
-thm compat_tree.induct compat_tree.rec
-
-
-subsection \<open> Creating New-Style Datatypes Using Old-Style Interfaces \<close>
-
-ML \<open>
-val l_specs =
-  [((@{binding l}, [("'a", @{sort type})], NoSyn),
-   [(@{binding N}, [], NoSyn),
-    (@{binding C}, [@{typ 'a}, Type (Sign.full_name @{theory} @{binding l}, [@{typ 'a}])], NoSyn)])];
-\<close>
-
-setup \<open> snd o BNF_LFP_Compat.add_datatype BNF_LFP_Compat.Unfold_Nesting l_specs; \<close>
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name l}; \<close>
-
-thm l.exhaust l.map l.induct l.rec l.size
-
-ML \<open>
-val t_specs =
-  [((@{binding t}, [("'b", @{sort type})], NoSyn),
-   [(@{binding T}, [@{typ 'b}, Type (@{type_name l},
-       [Type (Sign.full_name @{theory} @{binding t}, [@{typ 'b}])])], NoSyn)])];
-\<close>
-
-setup \<open> snd o BNF_LFP_Compat.add_datatype BNF_LFP_Compat.Unfold_Nesting t_specs; \<close>
-
-ML \<open> get_descrs @{theory} (2, 2, 1) @{type_name t}; \<close>
-
-thm t.exhaust t.map t.induct t.rec t.size
-thm compat_t.induct compat_t.rec
-
-ML \<open>
-val ft_specs =
-  [((@{binding ft}, [("'a", @{sort type})], NoSyn),
-   [(@{binding FT0}, [], NoSyn),
-    (@{binding FT}, [@{typ 'a} --> Type (Sign.full_name @{theory} @{binding ft}, [@{typ 'a}])],
-     NoSyn)])];
-\<close>
-
-setup \<open> snd o BNF_LFP_Compat.add_datatype BNF_LFP_Compat.Unfold_Nesting ft_specs; \<close>
-
-ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name ft}; \<close>
-
-thm ft.exhaust ft.induct ft.rec ft.size
-thm compat_ft.induct compat_ft.rec
-
-end
--- a/src/HOL/BNF_Examples/Derivation_Trees/DTree.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,90 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Derivation_Trees/DTree.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Derivation trees with nonterminal internal nodes and terminal leaves.
-*)
-
-header {* Trees with Nonterminal Internal Nodes and Terminal Leaves *}
-
-theory DTree
-imports Prelim
-begin
-
-typedecl N
-typedecl T
-
-codatatype dtree = NNode (root: N) (ccont: "(T + dtree) fset")
-
-subsection{* Transporting the Characteristic Lemmas from @{text "fset"} to @{text "set"} *}
-
-definition "Node n as \<equiv> NNode n (the_inv fset as)"
-definition "cont \<equiv> fset o ccont"
-definition "unfold rt ct \<equiv> corec_dtree rt (the_inv fset o image (map_sum id Inr) o ct)"
-definition "corec rt ct \<equiv> corec_dtree rt (the_inv fset o ct)"
-
-lemma finite_cont[simp]: "finite (cont tr)"
-  unfolding cont_def comp_apply by (cases tr, clarsimp)
-
-lemma Node_root_cont[simp]:
-  "Node (root tr) (cont tr) = tr"
-  unfolding Node_def cont_def comp_apply
-  apply (rule trans[OF _ dtree.collapse])
-  apply (rule arg_cong2[OF refl the_inv_into_f_f[unfolded inj_on_def]])
-  apply (simp_all add: fset_inject)
-  done
-
-lemma dtree_simps[simp]:
-assumes "finite as" and "finite as'"
-shows "Node n as = Node n' as' \<longleftrightarrow> n = n' \<and> as = as'"
-using assms dtree.inject unfolding Node_def
-by (metis fset_to_fset)
-
-lemma dtree_cases[elim, case_names Node Choice]:
-assumes Node: "\<And> n as. \<lbrakk>finite as; tr = Node n as\<rbrakk> \<Longrightarrow> phi"
-shows phi
-apply(cases rule: dtree.exhaust[of tr])
-using Node unfolding Node_def
-by (metis Node Node_root_cont finite_cont)
-
-lemma dtree_sel_ctr[simp]:
-"root (Node n as) = n"
-"finite as \<Longrightarrow> cont (Node n as) = as"
-unfolding Node_def cont_def by auto
-
-lemmas root_Node = dtree_sel_ctr(1)
-lemmas cont_Node = dtree_sel_ctr(2)
-
-lemma dtree_cong:
-assumes "root tr = root tr'" and "cont tr = cont tr'"
-shows "tr = tr'"
-by (metis Node_root_cont assms)
-
-lemma rel_set_cont:
-"rel_set \<chi> (cont tr1) (cont tr2) = rel_fset \<chi> (ccont tr1) (ccont tr2)"
-unfolding cont_def comp_def rel_fset_fset ..
-
-lemma dtree_coinduct[elim, consumes 1, case_names Lift, induct pred: "HOL.eq"]:
-assumes phi: "\<phi> tr1 tr2" and
-Lift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
-                  root tr1 = root tr2 \<and> rel_set (rel_sum op = \<phi>) (cont tr1) (cont tr2)"
-shows "tr1 = tr2"
-using phi apply(elim dtree.coinduct)
-apply(rule Lift[unfolded rel_set_cont]) .
-
-lemma unfold:
-"root (unfold rt ct b) = rt b"
-"finite (ct b) \<Longrightarrow> cont (unfold rt ct b) = image (id \<oplus> unfold rt ct) (ct b)"
-using dtree.corec_sel[of rt "the_inv fset o image (map_sum id Inr) o ct" b] unfolding unfold_def
-apply blast
-unfolding cont_def comp_def
-by (simp add: case_sum_o_inj map_sum.compositionality image_image)
-
-lemma corec:
-"root (corec rt ct b) = rt b"
-"finite (ct b) \<Longrightarrow> cont (corec rt ct b) = image (id \<oplus> ([[id, corec rt ct]])) (ct b)"
-using dtree.corec_sel[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
-unfolding cont_def comp_def id_def
-by simp_all
-
-end
--- a/src/HOL/BNF_Examples/Derivation_Trees/Gram_Lang.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1357 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Derivation_Trees/Gram_Lang.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Language of a grammar.
-*)
-
-header {* Language of a Grammar *}
-
-theory Gram_Lang
-imports DTree "~~/src/HOL/Library/Infinite_Set"
-begin
-
-
-(* We assume that the sets of terminals, and the left-hand sides of
-productions are finite and that the grammar has no unused nonterminals. *)
-consts P :: "(N \<times> (T + N) set) set"
-axiomatization where
-    finite_N: "finite (UNIV::N set)"
-and finite_in_P: "\<And> n tns. (n,tns) \<in> P \<longrightarrow> finite tns"
-and used: "\<And> n. \<exists> tns. (n,tns) \<in> P"
-
-
-subsection{* Tree Basics: frontier, interior, etc. *}
-
-
-(* Frontier *)
-
-inductive inFr :: "N set \<Rightarrow> dtree \<Rightarrow> T \<Rightarrow> bool" where
-Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr ns tr t"
-|
-Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inFr ns tr1 t\<rbrakk> \<Longrightarrow> inFr ns tr t"
-
-definition "Fr ns tr \<equiv> {t. inFr ns tr t}"
-
-lemma inFr_root_in: "inFr ns tr t \<Longrightarrow> root tr \<in> ns"
-by (metis inFr.simps)
-
-lemma inFr_mono:
-assumes "inFr ns tr t" and "ns \<subseteq> ns'"
-shows "inFr ns' tr t"
-using assms apply(induct arbitrary: ns' rule: inFr.induct)
-using Base Ind by (metis inFr.simps set_mp)+
-
-lemma inFr_Ind_minus:
-assumes "inFr ns1 tr1 t" and "Inr tr1 \<in> cont tr"
-shows "inFr (insert (root tr) ns1) tr t"
-using assms apply(induct rule: inFr.induct)
-  apply (metis inFr.simps insert_iff)
-  by (metis inFr.simps inFr_mono insertI1 subset_insertI)
-
-(* alternative definition *)
-inductive inFr2 :: "N set \<Rightarrow> dtree \<Rightarrow> T \<Rightarrow> bool" where
-Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr2 ns tr t"
-|
-Ind: "\<lbrakk>Inr tr1 \<in> cont tr; inFr2 ns1 tr1 t\<rbrakk>
-      \<Longrightarrow> inFr2 (insert (root tr) ns1) tr t"
-
-lemma inFr2_root_in: "inFr2 ns tr t \<Longrightarrow> root tr \<in> ns"
-apply(induct rule: inFr2.induct) by auto
-
-lemma inFr2_mono:
-assumes "inFr2 ns tr t" and "ns \<subseteq> ns'"
-shows "inFr2 ns' tr t"
-using assms apply(induct arbitrary: ns' rule: inFr2.induct)
-using Base Ind
-apply (metis subsetD) by (metis inFr2.simps insert_absorb insert_subset)
-
-lemma inFr2_Ind:
-assumes "inFr2 ns tr1 t" and "root tr \<in> ns" and "Inr tr1 \<in> cont tr"
-shows "inFr2 ns tr t"
-using assms apply(induct rule: inFr2.induct)
-  apply (metis inFr2.simps insert_absorb)
-  by (metis inFr2.simps insert_absorb)
-
-lemma inFr_inFr2:
-"inFr = inFr2"
-apply (rule ext)+  apply(safe)
-  apply(erule inFr.induct)
-    apply (metis (lifting) inFr2.Base)
-    apply (metis (lifting) inFr2_Ind)
-  apply(erule inFr2.induct)
-    apply (metis (lifting) inFr.Base)
-    apply (metis (lifting) inFr_Ind_minus)
-done
-
-lemma not_root_inFr:
-assumes "root tr \<notin> ns"
-shows "\<not> inFr ns tr t"
-by (metis assms inFr_root_in)
-
-lemma not_root_Fr:
-assumes "root tr \<notin> ns"
-shows "Fr ns tr = {}"
-using not_root_inFr[OF assms] unfolding Fr_def by auto
-
-
-(* Interior *)
-
-inductive inItr :: "N set \<Rightarrow> dtree \<Rightarrow> N \<Rightarrow> bool" where
-Base: "root tr \<in> ns \<Longrightarrow> inItr ns tr (root tr)"
-|
-Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inItr ns tr1 n\<rbrakk> \<Longrightarrow> inItr ns tr n"
-
-definition "Itr ns tr \<equiv> {n. inItr ns tr n}"
-
-lemma inItr_root_in: "inItr ns tr n \<Longrightarrow> root tr \<in> ns"
-by (metis inItr.simps)
-
-lemma inItr_mono:
-assumes "inItr ns tr n" and "ns \<subseteq> ns'"
-shows "inItr ns' tr n"
-using assms apply(induct arbitrary: ns' rule: inItr.induct)
-using Base Ind by (metis inItr.simps set_mp)+
-
-
-(* The subtree relation *)
-
-inductive subtr where
-Refl: "root tr \<in> ns \<Longrightarrow> subtr ns tr tr"
-|
-Step: "\<lbrakk>root tr3 \<in> ns; subtr ns tr1 tr2; Inr tr2 \<in> cont tr3\<rbrakk> \<Longrightarrow> subtr ns tr1 tr3"
-
-lemma subtr_rootL_in:
-assumes "subtr ns tr1 tr2"
-shows "root tr1 \<in> ns"
-using assms apply(induct rule: subtr.induct) by auto
-
-lemma subtr_rootR_in:
-assumes "subtr ns tr1 tr2"
-shows "root tr2 \<in> ns"
-using assms apply(induct rule: subtr.induct) by auto
-
-lemmas subtr_roots_in = subtr_rootL_in subtr_rootR_in
-
-lemma subtr_mono:
-assumes "subtr ns tr1 tr2" and "ns \<subseteq> ns'"
-shows "subtr ns' tr1 tr2"
-using assms apply(induct arbitrary: ns' rule: subtr.induct)
-using Refl Step by (metis subtr.simps set_mp)+
-
-lemma subtr_trans_Un:
-assumes "subtr ns12 tr1 tr2" and "subtr ns23 tr2 tr3"
-shows "subtr (ns12 \<union> ns23) tr1 tr3"
-proof-
-  have "subtr ns23 tr2 tr3  \<Longrightarrow>
-        (\<forall> ns12 tr1. subtr ns12 tr1 tr2 \<longrightarrow> subtr (ns12 \<union> ns23) tr1 tr3)"
-  apply(induct  rule: subtr.induct, safe)
-    apply (metis subtr_mono sup_commute sup_ge2)
-    by (metis (lifting) Step UnI2)
-  thus ?thesis using assms by auto
-qed
-
-lemma subtr_trans:
-assumes "subtr ns tr1 tr2" and "subtr ns tr2 tr3"
-shows "subtr ns tr1 tr3"
-using subtr_trans_Un[OF assms] by simp
-
-lemma subtr_StepL:
-assumes r: "root tr1 \<in> ns" and tr12: "Inr tr1 \<in> cont tr2" and s: "subtr ns tr2 tr3"
-shows "subtr ns tr1 tr3"
-apply(rule subtr_trans[OF _ s])
-apply(rule Step[of tr2 ns tr1 tr1])
-apply(rule subtr_rootL_in[OF s])
-apply(rule Refl[OF r])
-apply(rule tr12)
-done
-
-(* alternative definition: *)
-inductive subtr2 where
-Refl: "root tr \<in> ns \<Longrightarrow> subtr2 ns tr tr"
-|
-Step: "\<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr2 ns tr2 tr3\<rbrakk> \<Longrightarrow> subtr2 ns tr1 tr3"
-
-lemma subtr2_rootL_in:
-assumes "subtr2 ns tr1 tr2"
-shows "root tr1 \<in> ns"
-using assms apply(induct rule: subtr2.induct) by auto
-
-lemma subtr2_rootR_in:
-assumes "subtr2 ns tr1 tr2"
-shows "root tr2 \<in> ns"
-using assms apply(induct rule: subtr2.induct) by auto
-
-lemmas subtr2_roots_in = subtr2_rootL_in subtr2_rootR_in
-
-lemma subtr2_mono:
-assumes "subtr2 ns tr1 tr2" and "ns \<subseteq> ns'"
-shows "subtr2 ns' tr1 tr2"
-using assms apply(induct arbitrary: ns' rule: subtr2.induct)
-using Refl Step by (metis subtr2.simps set_mp)+
-
-lemma subtr2_trans_Un:
-assumes "subtr2 ns12 tr1 tr2" and "subtr2 ns23 tr2 tr3"
-shows "subtr2 (ns12 \<union> ns23) tr1 tr3"
-proof-
-  have "subtr2 ns12 tr1 tr2  \<Longrightarrow>
-        (\<forall> ns23 tr3. subtr2 ns23 tr2 tr3 \<longrightarrow> subtr2 (ns12 \<union> ns23) tr1 tr3)"
-  apply(induct  rule: subtr2.induct, safe)
-    apply (metis subtr2_mono sup_commute sup_ge2)
-    by (metis Un_iff subtr2.simps)
-  thus ?thesis using assms by auto
-qed
-
-lemma subtr2_trans:
-assumes "subtr2 ns tr1 tr2" and "subtr2 ns tr2 tr3"
-shows "subtr2 ns tr1 tr3"
-using subtr2_trans_Un[OF assms] by simp
-
-lemma subtr2_StepR:
-assumes r: "root tr3 \<in> ns" and tr23: "Inr tr2 \<in> cont tr3" and s: "subtr2 ns tr1 tr2"
-shows "subtr2 ns tr1 tr3"
-apply(rule subtr2_trans[OF s])
-apply(rule Step[of _ _ tr3])
-apply(rule subtr2_rootR_in[OF s])
-apply(rule tr23)
-apply(rule Refl[OF r])
-done
-
-lemma subtr_subtr2:
-"subtr = subtr2"
-apply (rule ext)+  apply(safe)
-  apply(erule subtr.induct)
-    apply (metis (lifting) subtr2.Refl)
-    apply (metis (lifting) subtr2_StepR)
-  apply(erule subtr2.induct)
-    apply (metis (lifting) subtr.Refl)
-    apply (metis (lifting) subtr_StepL)
-done
-
-lemma subtr_inductL[consumes 1, case_names Refl Step]:
-assumes s: "subtr ns tr1 tr2" and Refl: "\<And>ns tr. \<phi> ns tr tr"
-and Step:
-"\<And>ns tr1 tr2 tr3.
-   \<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr ns tr2 tr3; \<phi> ns tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> ns tr1 tr3"
-shows "\<phi> ns tr1 tr2"
-using s unfolding subtr_subtr2 apply(rule subtr2.induct)
-using Refl Step unfolding subtr_subtr2 by auto
-
-lemma subtr_UNIV_inductL[consumes 1, case_names Refl Step]:
-assumes s: "subtr UNIV tr1 tr2" and Refl: "\<And>tr. \<phi> tr tr"
-and Step:
-"\<And>tr1 tr2 tr3.
-   \<lbrakk>Inr tr1 \<in> cont tr2; subtr UNIV tr2 tr3; \<phi> tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> tr1 tr3"
-shows "\<phi> tr1 tr2"
-using s apply(induct rule: subtr_inductL)
-apply(rule Refl) using Step subtr_mono by (metis subset_UNIV)
-
-(* Subtree versus frontier: *)
-lemma subtr_inFr:
-assumes "inFr ns tr t" and "subtr ns tr tr1"
-shows "inFr ns tr1 t"
-proof-
-  have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inFr ns tr t \<longrightarrow> inFr ns tr1 t)"
-  apply(induct rule: subtr.induct, safe) by (metis inFr.Ind)
-  thus ?thesis using assms by auto
-qed
-
-corollary Fr_subtr:
-"Fr ns tr = \<Union> {Fr ns tr' | tr'. subtr ns tr' tr}"
-unfolding Fr_def proof safe
-  fix t assume t: "inFr ns tr t"  hence "root tr \<in> ns" by (rule inFr_root_in)
-  thus "t \<in> \<Union>{{t. inFr ns tr' t} |tr'. subtr ns tr' tr}"
-  apply(intro UnionI[of "{t. inFr ns tr t}" _ t]) using t subtr.Refl by auto
-qed(metis subtr_inFr)
-
-lemma inFr_subtr:
-assumes "inFr ns tr t"
-shows "\<exists> tr'. subtr ns tr' tr \<and> Inl t \<in> cont tr'"
-using assms apply(induct rule: inFr.induct) apply safe
-  apply (metis subtr.Refl)
-  by (metis (lifting) subtr.Step)
-
-corollary Fr_subtr_cont:
-"Fr ns tr = \<Union> {Inl -` cont tr' | tr'. subtr ns tr' tr}"
-unfolding Fr_def
-apply safe
-apply (frule inFr_subtr)
-apply auto
-by (metis inFr.Base subtr_inFr subtr_rootL_in)
-
-(* Subtree versus interior: *)
-lemma subtr_inItr:
-assumes "inItr ns tr n" and "subtr ns tr tr1"
-shows "inItr ns tr1 n"
-proof-
-  have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inItr ns tr n \<longrightarrow> inItr ns tr1 n)"
-  apply(induct rule: subtr.induct, safe) by (metis inItr.Ind)
-  thus ?thesis using assms by auto
-qed
-
-corollary Itr_subtr:
-"Itr ns tr = \<Union> {Itr ns tr' | tr'. subtr ns tr' tr}"
-unfolding Itr_def apply safe
-apply (metis (lifting, mono_tags) UnionI inItr_root_in mem_Collect_eq subtr.Refl)
-by (metis subtr_inItr)
-
-lemma inItr_subtr:
-assumes "inItr ns tr n"
-shows "\<exists> tr'. subtr ns tr' tr \<and> root tr' = n"
-using assms apply(induct rule: inItr.induct) apply safe
-  apply (metis subtr.Refl)
-  by (metis (lifting) subtr.Step)
-
-corollary Itr_subtr_cont:
-"Itr ns tr = {root tr' | tr'. subtr ns tr' tr}"
-unfolding Itr_def apply safe
-  apply (metis (lifting, mono_tags) inItr_subtr)
-  by (metis inItr.Base subtr_inItr subtr_rootL_in)
-
-
-subsection{* The Immediate Subtree Function *}
-
-(* production of: *)
-abbreviation "prodOf tr \<equiv> (id \<oplus> root) ` (cont tr)"
-(* subtree of: *)
-definition "subtrOf tr n \<equiv> SOME tr'. Inr tr' \<in> cont tr \<and> root tr' = n"
-
-lemma subtrOf:
-assumes n: "Inr n \<in> prodOf tr"
-shows "Inr (subtrOf tr n) \<in> cont tr \<and> root (subtrOf tr n) = n"
-proof-
-  obtain tr' where "Inr tr' \<in> cont tr \<and> root tr' = n"
-  using n unfolding image_def by (metis (lifting) Inr_oplus_elim assms)
-  thus ?thesis unfolding subtrOf_def by(rule someI)
-qed
-
-lemmas Inr_subtrOf = subtrOf[THEN conjunct1]
-lemmas root_subtrOf[simp] = subtrOf[THEN conjunct2]
-
-lemma Inl_prodOf: "Inl -` (prodOf tr) = Inl -` (cont tr)"
-proof safe
-  fix t ttr assume "Inl t = (id \<oplus> root) ttr" and "ttr \<in> cont tr"
-  thus "t \<in> Inl -` cont tr" by(cases ttr, auto)
-next
-  fix t assume "Inl t \<in> cont tr" thus "t \<in> Inl -` prodOf tr"
-  by (metis (lifting) id_def image_iff map_sum.simps(1) vimageI2)
-qed
-
-lemma root_prodOf:
-assumes "Inr tr' \<in> cont tr"
-shows "Inr (root tr') \<in> prodOf tr"
-by (metis (lifting) assms image_iff map_sum.simps(2))
-
-
-subsection{* Well-Formed Derivation Trees *}
-
-hide_const wf
-
-coinductive wf where
-dtree: "\<lbrakk>(root tr, (id \<oplus> root) ` (cont tr)) \<in> P; inj_on root (Inr -` cont tr);
-        \<And> tr'. tr' \<in> Inr -` (cont tr) \<Longrightarrow> wf tr'\<rbrakk> \<Longrightarrow> wf tr"
-
-(* destruction rules: *)
-lemma wf_P:
-assumes "wf tr"
-shows "(root tr, (id \<oplus> root) ` (cont tr)) \<in> P"
-using assms wf.simps[of tr] by auto
-
-lemma wf_inj_on:
-assumes "wf tr"
-shows "inj_on root (Inr -` cont tr)"
-using assms wf.simps[of tr] by auto
-
-lemma wf_inj[simp]:
-assumes "wf tr" and "Inr tr1 \<in> cont tr" and "Inr tr2 \<in> cont tr"
-shows "root tr1 = root tr2 \<longleftrightarrow> tr1 = tr2"
-using assms wf_inj_on unfolding inj_on_def by auto
-
-lemma wf_cont:
-assumes "wf tr" and "Inr tr' \<in> cont tr"
-shows "wf tr'"
-using assms wf.simps[of tr] by auto
-
-
-(* coinduction:*)
-lemma wf_coind[elim, consumes 1, case_names Hyp]:
-assumes phi: "\<phi> tr"
-and Hyp:
-"\<And> tr. \<phi> tr \<Longrightarrow>
-       (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
-       inj_on root (Inr -` cont tr) \<and>
-       (\<forall> tr' \<in> Inr -` (cont tr). \<phi> tr' \<or> wf tr')"
-shows "wf tr"
-apply(rule wf.coinduct[of \<phi> tr, OF phi])
-using Hyp by blast
-
-lemma wf_raw_coind[elim, consumes 1, case_names Hyp]:
-assumes phi: "\<phi> tr"
-and Hyp:
-"\<And> tr. \<phi> tr \<Longrightarrow>
-       (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
-       inj_on root (Inr -` cont tr) \<and>
-       (\<forall> tr' \<in> Inr -` (cont tr). \<phi> tr')"
-shows "wf tr"
-using phi apply(induct rule: wf_coind)
-using Hyp by (metis (mono_tags))
-
-lemma wf_subtr_inj_on:
-assumes d: "wf tr1" and s: "subtr ns tr tr1"
-shows "inj_on root (Inr -` cont tr)"
-using s d apply(induct rule: subtr.induct)
-apply (metis (lifting) wf_inj_on) by (metis wf_cont)
-
-lemma wf_subtr_P:
-assumes d: "wf tr1" and s: "subtr ns tr tr1"
-shows "(root tr, (id \<oplus> root) ` cont tr) \<in> P"
-using s d apply(induct rule: subtr.induct)
-apply (metis (lifting) wf_P) by (metis wf_cont)
-
-lemma subtrOf_root[simp]:
-assumes tr: "wf tr" and cont: "Inr tr' \<in> cont tr"
-shows "subtrOf tr (root tr') = tr'"
-proof-
-  have 0: "Inr (subtrOf tr (root tr')) \<in> cont tr" using Inr_subtrOf
-  by (metis (lifting) cont root_prodOf)
-  have "root (subtrOf tr (root tr')) = root tr'"
-  using root_subtrOf by (metis (lifting) cont root_prodOf)
-  thus ?thesis unfolding wf_inj[OF tr 0 cont] .
-qed
-
-lemma surj_subtrOf:
-assumes "wf tr" and 0: "Inr tr' \<in> cont tr"
-shows "\<exists> n. Inr n \<in> prodOf tr \<and> subtrOf tr n = tr'"
-apply(rule exI[of _ "root tr'"])
-using root_prodOf[OF 0] subtrOf_root[OF assms] by simp
-
-lemma wf_subtr:
-assumes "wf tr1" and "subtr ns tr tr1"
-shows "wf tr"
-proof-
-  have "(\<exists> ns tr1. wf tr1 \<and> subtr ns tr tr1) \<Longrightarrow> wf tr"
-  proof (induct rule: wf_raw_coind)
-    case (Hyp tr)
-    then obtain ns tr1 where tr1: "wf tr1" and tr_tr1: "subtr ns tr tr1" by auto
-    show ?case proof safe
-      show "(root tr, (id \<oplus> root) ` cont tr) \<in> P" using wf_subtr_P[OF tr1 tr_tr1] .
-    next
-      show "inj_on root (Inr -` cont tr)" using wf_subtr_inj_on[OF tr1 tr_tr1] .
-    next
-      fix tr' assume tr': "Inr tr' \<in> cont tr"
-      have tr_tr1: "subtr (ns \<union> {root tr'}) tr tr1" using subtr_mono[OF tr_tr1] by auto
-      have "subtr (ns \<union> {root tr'}) tr' tr1" using subtr_StepL[OF _ tr' tr_tr1] by auto
-      thus "\<exists>ns' tr1. wf tr1 \<and> subtr ns' tr' tr1" using tr1 by blast
-    qed
-  qed
-  thus ?thesis using assms by auto
-qed
-
-
-subsection{* Default Trees *}
-
-(* Pick a left-hand side of a production for each nonterminal *)
-definition S where "S n \<equiv> SOME tns. (n,tns) \<in> P"
-
-lemma S_P: "(n, S n) \<in> P"
-using used unfolding S_def by(rule someI_ex)
-
-lemma finite_S: "finite (S n)"
-using S_P finite_in_P by auto
-
-
-(* The default tree of a nonterminal *)
-definition deftr :: "N \<Rightarrow> dtree" where
-"deftr \<equiv> unfold id S"
-
-lemma deftr_simps[simp]:
-"root (deftr n) = n"
-"cont (deftr n) = image (id \<oplus> deftr) (S n)"
-using unfold(1)[of id S n] unfold(2)[of S n id, OF finite_S]
-unfolding deftr_def by simp_all
-
-lemmas root_deftr = deftr_simps(1)
-lemmas cont_deftr = deftr_simps(2)
-
-lemma root_o_deftr[simp]: "root o deftr = id"
-by (rule ext, auto)
-
-lemma wf_deftr: "wf (deftr n)"
-proof-
-  {fix tr assume "\<exists> n. tr = deftr n" hence "wf tr"
-   apply(induct rule: wf_raw_coind) apply safe
-   unfolding deftr_simps image_comp map_sum.comp id_comp
-   root_o_deftr map_sum.id image_id id_apply apply(rule S_P)
-   unfolding inj_on_def by auto
-  }
-  thus ?thesis by auto
-qed
-
-
-subsection{* Hereditary Substitution *}
-
-(* Auxiliary concept: The root-ommiting frontier: *)
-definition "inFrr ns tr t \<equiv> \<exists> tr'. Inr tr' \<in> cont tr \<and> inFr ns tr' t"
-definition "Frr ns tr \<equiv> {t. \<exists> tr'. Inr tr' \<in> cont tr \<and> t \<in> Fr ns tr'}"
-
-context
-fixes tr0 :: dtree
-begin
-
-definition "hsubst_r tr \<equiv> root tr"
-definition "hsubst_c tr \<equiv> if root tr = root tr0 then cont tr0 else cont tr"
-
-(* Hereditary substitution: *)
-definition hsubst :: "dtree \<Rightarrow> dtree" where
-"hsubst \<equiv> unfold hsubst_r hsubst_c"
-
-lemma finite_hsubst_c: "finite (hsubst_c n)"
-unfolding hsubst_c_def by (metis (full_types) finite_cont)
-
-lemma root_hsubst[simp]: "root (hsubst tr) = root tr"
-using unfold(1)[of hsubst_r hsubst_c tr] unfolding hsubst_def hsubst_r_def by simp
-
-lemma root_o_subst[simp]: "root o hsubst = root"
-unfolding comp_def root_hsubst ..
-
-lemma cont_hsubst_eq[simp]:
-assumes "root tr = root tr0"
-shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr0)"
-apply(subst id_comp[symmetric, of id]) unfolding id_comp
-using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
-unfolding hsubst_def hsubst_c_def using assms by simp
-
-lemma hsubst_eq:
-assumes "root tr = root tr0"
-shows "hsubst tr = hsubst tr0"
-apply(rule dtree_cong) using assms cont_hsubst_eq by auto
-
-lemma cont_hsubst_neq[simp]:
-assumes "root tr \<noteq> root tr0"
-shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr)"
-apply(subst id_comp[symmetric, of id]) unfolding id_comp
-using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
-unfolding hsubst_def hsubst_c_def using assms by simp
-
-lemma Inl_cont_hsubst_eq[simp]:
-assumes "root tr = root tr0"
-shows "Inl -` cont (hsubst tr) = Inl -` (cont tr0)"
-unfolding cont_hsubst_eq[OF assms] by simp
-
-lemma Inr_cont_hsubst_eq[simp]:
-assumes "root tr = root tr0"
-shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr0"
-unfolding cont_hsubst_eq[OF assms] by simp
-
-lemma Inl_cont_hsubst_neq[simp]:
-assumes "root tr \<noteq> root tr0"
-shows "Inl -` cont (hsubst tr) = Inl -` (cont tr)"
-unfolding cont_hsubst_neq[OF assms] by simp
-
-lemma Inr_cont_hsubst_neq[simp]:
-assumes "root tr \<noteq> root tr0"
-shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr"
-unfolding cont_hsubst_neq[OF assms] by simp
-
-lemma wf_hsubst:
-assumes tr0: "wf tr0" and tr: "wf tr"
-shows "wf (hsubst tr)"
-proof-
-  {fix tr1 have "(\<exists> tr. wf tr \<and> tr1 = hsubst tr) \<Longrightarrow> wf tr1"
-   proof (induct rule: wf_raw_coind)
-     case (Hyp tr1) then obtain tr
-     where dtr: "wf tr" and tr1: "tr1 = hsubst tr" by auto
-     show ?case unfolding tr1 proof safe
-       show "(root (hsubst tr), prodOf (hsubst tr)) \<in> P"
-       unfolding tr1 apply(cases "root tr = root tr0")
-       using  wf_P[OF dtr] wf_P[OF tr0]
-       by (auto simp add: image_comp map_sum.comp)
-       show "inj_on root (Inr -` cont (hsubst tr))"
-       apply(cases "root tr = root tr0") using wf_inj_on[OF dtr] wf_inj_on[OF tr0]
-       unfolding inj_on_def by (auto, blast)
-       fix tr' assume "Inr tr' \<in> cont (hsubst tr)"
-       thus "\<exists>tra. wf tra \<and> tr' = hsubst tra"
-       apply(cases "root tr = root tr0", simp_all)
-         apply (metis wf_cont tr0)
-         by (metis dtr wf_cont)
-     qed
-   qed
-  }
-  thus ?thesis using assms by blast
-qed
-
-lemma Frr: "Frr ns tr = {t. inFrr ns tr t}"
-unfolding inFrr_def Frr_def Fr_def by auto
-
-lemma inFr_hsubst_imp:
-assumes "inFr ns (hsubst tr) t"
-shows "t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
-       inFr (ns - {root tr0}) tr t"
-proof-
-  {fix tr1
-   have "inFr ns tr1 t \<Longrightarrow>
-   (\<And> tr. tr1 = hsubst tr \<Longrightarrow> (t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
-                              inFr (ns - {root tr0}) tr t))"
-   proof(induct rule: inFr.induct)
-     case (Base tr1 ns t tr)
-     hence rtr: "root tr1 \<in> ns" and t_tr1: "Inl t \<in> cont tr1" and tr1: "tr1 = hsubst tr"
-     by auto
-     show ?case
-     proof(cases "root tr1 = root tr0")
-       case True
-       hence "t \<in> Inl -` (cont tr0)" using t_tr1 unfolding tr1 by auto
-       thus ?thesis by simp
-     next
-       case False
-       hence "inFr (ns - {root tr0}) tr t" using t_tr1 unfolding tr1 apply simp
-       by (metis Base.prems Diff_iff root_hsubst inFr.Base rtr singletonE)
-       thus ?thesis by simp
-     qed
-   next
-     case (Ind tr1 ns tr1' t) note IH = Ind(4)
-     have rtr1: "root tr1 \<in> ns" and tr1'_tr1: "Inr tr1' \<in> cont tr1"
-     and t_tr1': "inFr ns tr1' t" and tr1: "tr1 = hsubst tr" using Ind by auto
-     have rtr1: "root tr1 = root tr" unfolding tr1 by simp
-     show ?case
-     proof(cases "root tr1 = root tr0")
-       case True
-       then obtain tr' where tr'_tr0: "Inr tr' \<in> cont tr0" and tr1': "tr1' = hsubst tr'"
-       using tr1'_tr1 unfolding tr1 by auto
-       show ?thesis using IH[OF tr1'] proof (elim disjE)
-         assume "inFr (ns - {root tr0}) tr' t"
-         thus ?thesis using tr'_tr0 unfolding inFrr_def by auto
-       qed auto
-     next
-       case False
-       then obtain tr' where tr'_tr: "Inr tr' \<in> cont tr" and tr1': "tr1' = hsubst tr'"
-       using tr1'_tr1 unfolding tr1 by auto
-       show ?thesis using IH[OF tr1'] proof (elim disjE)
-         assume "inFr (ns - {root tr0}) tr' t"
-         thus ?thesis using tr'_tr unfolding inFrr_def
-         by (metis Diff_iff False Ind(1) empty_iff inFr2_Ind inFr_inFr2 insert_iff rtr1)
-       qed auto
-     qed
-   qed
-  }
-  thus ?thesis using assms by auto
-qed
-
-lemma inFr_hsubst_notin:
-assumes "inFr ns tr t" and "root tr0 \<notin> ns"
-shows "inFr ns (hsubst tr) t"
-using assms apply(induct rule: inFr.induct)
-apply (metis Inl_cont_hsubst_neq inFr2.Base inFr_inFr2 root_hsubst vimageD vimageI2)
-by (metis (lifting) Inr_cont_hsubst_neq inFr.Ind rev_image_eqI root_hsubst vimageD vimageI2)
-
-lemma inFr_hsubst_minus:
-assumes "inFr (ns - {root tr0}) tr t"
-shows "inFr ns (hsubst tr) t"
-proof-
-  have 1: "inFr (ns - {root tr0}) (hsubst tr) t"
-  using inFr_hsubst_notin[OF assms] by simp
-  show ?thesis using inFr_mono[OF 1] by auto
-qed
-
-lemma inFr_self_hsubst:
-assumes "root tr0 \<in> ns"
-shows
-"inFr ns (hsubst tr0) t \<longleftrightarrow>
- t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t"
-(is "?A \<longleftrightarrow> ?B \<or> ?C")
-apply(intro iffI)
-apply (metis inFr_hsubst_imp Diff_iff inFr_root_in insertI1) proof(elim disjE)
-  assume ?B thus ?A apply(intro inFr.Base) using assms by auto
-next
-  assume ?C then obtain tr where
-  tr_tr0: "Inr tr \<in> cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t"
-  unfolding inFrr_def by auto
-  def tr1 \<equiv> "hsubst tr"
-  have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto
-  have "Inr tr1 \<in> cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto
-  thus ?A using 1 inFr.Ind assms by (metis root_hsubst)
-qed
-
-lemma Fr_self_hsubst:
-assumes "root tr0 \<in> ns"
-shows "Fr ns (hsubst tr0) = Inl -` (cont tr0) \<union> Frr (ns - {root tr0}) tr0"
-using inFr_self_hsubst[OF assms] unfolding Frr Fr_def by auto
-
-end (* context *)
-
-
-subsection{* Regular Trees *}
-
-definition "reg f tr \<equiv> \<forall> tr'. subtr UNIV tr' tr \<longrightarrow> tr' = f (root tr')"
-definition "regular tr \<equiv> \<exists> f. reg f tr"
-
-lemma reg_def2: "reg f tr \<longleftrightarrow> (\<forall> ns tr'. subtr ns tr' tr \<longrightarrow> tr' = f (root tr'))"
-unfolding reg_def using subtr_mono by (metis subset_UNIV)
-
-lemma regular_def2: "regular tr \<longleftrightarrow> (\<exists> f. reg f tr \<and> (\<forall> n. root (f n) = n))"
-unfolding regular_def proof safe
-  fix f assume f: "reg f tr"
-  def g \<equiv> "\<lambda> n. if inItr UNIV tr n then f n else deftr n"
-  show "\<exists>g. reg g tr \<and> (\<forall>n. root (g n) = n)"
-  apply(rule exI[of _ g])
-  using f deftr_simps(1) unfolding g_def reg_def apply safe
-    apply (metis (lifting) inItr.Base subtr_inItr subtr_rootL_in)
-    by (metis (full_types) inItr_subtr)
-qed auto
-
-lemma reg_root:
-assumes "reg f tr"
-shows "f (root tr) = tr"
-using assms unfolding reg_def
-by (metis (lifting) iso_tuple_UNIV_I subtr.Refl)
-
-
-lemma reg_Inr_cont:
-assumes "reg f tr" and "Inr tr' \<in> cont tr"
-shows "reg f tr'"
-by (metis (lifting) assms iso_tuple_UNIV_I reg_def subtr.Step)
-
-lemma reg_subtr:
-assumes "reg f tr" and "subtr ns tr' tr"
-shows "reg f tr'"
-using assms unfolding reg_def using subtr_trans[of UNIV tr] UNIV_I
-by (metis UNIV_eq_I UnCI Un_upper1 iso_tuple_UNIV_I subtr_mono subtr_trans)
-
-lemma regular_subtr:
-assumes r: "regular tr" and s: "subtr ns tr' tr"
-shows "regular tr'"
-using r reg_subtr[OF _ s] unfolding regular_def by auto
-
-lemma subtr_deftr:
-assumes "subtr ns tr' (deftr n)"
-shows "tr' = deftr (root tr')"
-proof-
-  {fix tr have "subtr ns tr' tr \<Longrightarrow> (\<forall> n. tr = deftr n \<longrightarrow> tr' = deftr (root tr'))"
-   apply (induct rule: subtr.induct)
-   proof(metis (lifting) deftr_simps(1), safe)
-     fix tr3 ns tr1 tr2 n
-     assume 1: "root (deftr n) \<in> ns" and 2: "subtr ns tr1 tr2"
-     and IH: "\<forall>n. tr2 = deftr n \<longrightarrow> tr1 = deftr (root tr1)"
-     and 3: "Inr tr2 \<in> cont (deftr n)"
-     have "tr2 \<in> deftr ` UNIV"
-     using 3 unfolding deftr_simps image_def
-     by (metis (lifting, full_types) 3 CollectI Inr_oplus_iff cont_deftr
-         iso_tuple_UNIV_I)
-     then obtain n where "tr2 = deftr n" by auto
-     thus "tr1 = deftr (root tr1)" using IH by auto
-   qed
-  }
-  thus ?thesis using assms by auto
-qed
-
-lemma reg_deftr: "reg deftr (deftr n)"
-unfolding reg_def using subtr_deftr by auto
-
-lemma wf_subtrOf_Union:
-assumes "wf tr"
-shows "\<Union>{K tr' |tr'. Inr tr' \<in> cont tr} =
-       \<Union>{K (subtrOf tr n) |n. Inr n \<in> prodOf tr}"
-unfolding Union_eq Bex_def mem_Collect_eq proof safe
-  fix x xa tr'
-  assume x: "x \<in> K tr'" and tr'_tr: "Inr tr' \<in> cont tr"
-  show "\<exists>X. (\<exists>n. X = K (subtrOf tr n) \<and> Inr n \<in> prodOf tr) \<and> x \<in> X"
-  apply(rule exI[of _ "K (subtrOf tr (root tr'))"]) apply(intro conjI)
-    apply(rule exI[of _ "root tr'"]) apply (metis (lifting) root_prodOf tr'_tr)
-    by (metis (lifting) assms subtrOf_root tr'_tr x)
-next
-  fix x X n ttr
-  assume x: "x \<in> K (subtrOf tr n)" and n: "Inr n = (id \<oplus> root) ttr" and ttr: "ttr \<in> cont tr"
-  show "\<exists>X. (\<exists>tr'. X = K tr' \<and> Inr tr' \<in> cont tr) \<and> x \<in> X"
-  apply(rule exI[of _ "K (subtrOf tr n)"]) apply(intro conjI)
-    apply(rule exI[of _ "subtrOf tr n"]) apply (metis imageI n subtrOf ttr)
-    using x .
-qed
-
-
-
-
-subsection {* Paths in a Regular Tree *}
-
-inductive path :: "(N \<Rightarrow> dtree) \<Rightarrow> N list \<Rightarrow> bool" for f where
-Base: "path f [n]"
-|
-Ind: "\<lbrakk>path f (n1 # nl); Inr (f n1) \<in> cont (f n)\<rbrakk>
-      \<Longrightarrow> path f (n # n1 # nl)"
-
-lemma path_NE:
-assumes "path f nl"
-shows "nl \<noteq> Nil"
-using assms apply(induct rule: path.induct) by auto
-
-lemma path_post:
-assumes f: "path f (n # nl)" and nl: "nl \<noteq> []"
-shows "path f nl"
-proof-
-  obtain n1 nl1 where nl: "nl = n1 # nl1" using nl by (cases nl, auto)
-  show ?thesis using assms unfolding nl using path.simps by (metis (lifting) list.inject)
-qed
-
-lemma path_post_concat:
-assumes "path f (nl1 @ nl2)" and "nl2 \<noteq> Nil"
-shows "path f nl2"
-using assms apply (induct nl1)
-apply (metis append_Nil) by (metis Nil_is_append_conv append_Cons path_post)
-
-lemma path_concat:
-assumes "path f nl1" and "path f ((last nl1) # nl2)"
-shows "path f (nl1 @ nl2)"
-using assms apply(induct rule: path.induct) apply simp
-by (metis append_Cons last.simps list.simps(3) path.Ind)
-
-lemma path_distinct:
-assumes "path f nl"
-shows "\<exists> nl'. path f nl' \<and> hd nl' = hd nl \<and> last nl' = last nl \<and>
-              set nl' \<subseteq> set nl \<and> distinct nl'"
-using assms proof(induct rule: length_induct)
-  case (1 nl)  hence p_nl: "path f nl" by simp
-  then obtain n nl1 where nl: "nl = n # nl1" by (metis list.exhaust path_NE)
-  show ?case
-  proof(cases nl1)
-    case Nil
-    show ?thesis apply(rule exI[of _ nl]) using path.Base unfolding nl Nil by simp
-  next
-    case (Cons n1 nl2)
-    hence p1: "path f nl1" by (metis list.simps(3) nl p_nl path_post)
-    show ?thesis
-    proof(cases "n \<in> set nl1")
-      case False
-      obtain nl1' where p1': "path f nl1'" and hd_nl1': "hd nl1' = hd nl1" and
-      l_nl1': "last nl1' = last nl1" and d_nl1': "distinct nl1'"
-      and s_nl1': "set nl1' \<subseteq> set nl1"
-      using 1(1)[THEN allE[of _ nl1]] p1 unfolding nl by auto
-      obtain nl2' where nl1': "nl1' = n1 # nl2'" using path_NE[OF p1'] hd_nl1'
-      unfolding Cons by(cases nl1', auto)
-      show ?thesis apply(intro exI[of _ "n # nl1'"]) unfolding nl proof safe
-        show "path f (n # nl1')" unfolding nl1'
-        apply(rule path.Ind, metis nl1' p1')
-        by (metis (lifting) Cons list.inject nl p1 p_nl path.simps path_NE)
-      qed(insert l_nl1' Cons nl1' s_nl1' d_nl1' False, auto)
-    next
-      case True
-      then obtain nl11 nl12 where nl1: "nl1 = nl11 @ n # nl12"
-      by (metis split_list)
-      have p12: "path f (n # nl12)"
-      apply(rule path_post_concat[of _ "n # nl11"]) using p_nl[unfolded nl nl1] by auto
-      obtain nl12' where p1': "path f nl12'" and hd_nl12': "hd nl12' = n" and
-      l_nl12': "last nl12' = last (n # nl12)" and d_nl12': "distinct nl12'"
-      and s_nl12': "set nl12' \<subseteq> {n} \<union> set nl12"
-      using 1(1)[THEN allE[of _ "n # nl12"]] p12 unfolding nl nl1 by auto
-      thus ?thesis apply(intro exI[of _ nl12']) unfolding nl nl1 by auto
-    qed
-  qed
-qed
-
-lemma path_subtr:
-assumes f: "\<And> n. root (f n) = n"
-and p: "path f nl"
-shows "subtr (set nl) (f (last nl)) (f (hd nl))"
-using p proof (induct rule: path.induct)
-  case (Ind n1 nl n)  let ?ns1 = "insert n1 (set nl)"
-  have "path f (n1 # nl)"
-  and "subtr ?ns1 (f (last (n1 # nl))) (f n1)"
-  and fn1: "Inr (f n1) \<in> cont (f n)" using Ind by simp_all
-  hence fn1_flast:  "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n1)"
-  by (metis subset_insertI subtr_mono)
-  have 1: "last (n # n1 # nl) = last (n1 # nl)" by auto
-  have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)"
-  using f subtr.Step[OF _ fn1_flast fn1] by auto
-  thus ?case unfolding 1 by simp
-qed (metis f list.sel(1) last_ConsL last_in_set not_Cons_self2 subtr.Refl)
-
-lemma reg_subtr_path_aux:
-assumes f: "reg f tr" and n: "subtr ns tr1 tr"
-shows "\<exists> nl. path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
-using n f proof(induct rule: subtr.induct)
-  case (Refl tr ns)
-  thus ?case
-  apply(intro exI[of _ "[root tr]"]) apply simp by (metis (lifting) path.Base reg_root)
-next
-  case (Step tr ns tr2 tr1)
-  hence rtr: "root tr \<in> ns" and tr1_tr: "Inr tr1 \<in> cont tr"
-  and tr2_tr1: "subtr ns tr2 tr1" and tr: "reg f tr" by auto
-  have tr1: "reg f tr1" using reg_subtr[OF tr] rtr tr1_tr
-  by (metis (lifting) Step.prems iso_tuple_UNIV_I reg_def subtr.Step)
-  obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1"
-  and last_nl: "f (last nl) = tr2" and set: "set nl \<subseteq> ns" using Step(3)[OF tr1] by auto
-  have 0: "path f (root tr # nl)" apply (subst path.simps)
-  using f_nl nl reg_root tr tr1_tr by (metis list.sel(1) neq_Nil_conv)
-  show ?case apply(rule exI[of _ "(root tr) # nl"])
-  using 0 reg_root tr last_nl nl path_NE rtr set by auto
-qed
-
-lemma reg_subtr_path:
-assumes f: "reg f tr" and n: "subtr ns tr1 tr"
-shows "\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
-using reg_subtr_path_aux[OF assms] path_distinct[of f]
-by (metis (lifting) order_trans)
-
-lemma subtr_iff_path:
-assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
-shows "subtr ns tr1 tr \<longleftrightarrow>
-       (\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns)"
-proof safe
-  fix nl assume p: "path f nl" and nl: "set nl \<subseteq> ns"
-  have "subtr (set nl) (f (last nl)) (f (hd nl))"
-  apply(rule path_subtr) using p f by simp_all
-  thus "subtr ns (f (last nl)) (f (hd nl))"
-  using subtr_mono nl by auto
-qed(insert reg_subtr_path[OF r], auto)
-
-lemma inFr_iff_path:
-assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
-shows
-"inFr ns tr t \<longleftrightarrow>
- (\<exists> nl tr1. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and>
-            set nl \<subseteq> ns \<and> Inl t \<in> cont tr1)"
-apply safe
-apply (metis (no_types) inFr_subtr r reg_subtr_path)
-by (metis f inFr.Base path_subtr subtr_inFr subtr_mono subtr_rootL_in)
-
-
-
-subsection{* The Regular Cut of a Tree *}
-
-context fixes tr0 :: dtree
-begin
-
-(* Picking a subtree of a certain root: *)
-definition "pick n \<equiv> SOME tr. subtr UNIV tr tr0 \<and> root tr = n"
-
-lemma pick:
-assumes "inItr UNIV tr0 n"
-shows "subtr UNIV (pick n) tr0 \<and> root (pick n) = n"
-proof-
-  have "\<exists> tr. subtr UNIV tr tr0 \<and> root tr = n"
-  using assms by (metis (lifting) inItr_subtr)
-  thus ?thesis unfolding pick_def by(rule someI_ex)
-qed
-
-lemmas subtr_pick = pick[THEN conjunct1]
-lemmas root_pick = pick[THEN conjunct2]
-
-lemma wf_pick:
-assumes tr0: "wf tr0" and n: "inItr UNIV tr0 n"
-shows "wf (pick n)"
-using wf_subtr[OF tr0 subtr_pick[OF n]] .
-
-definition "H_r n \<equiv> root (pick n)"
-definition "H_c n \<equiv> (id \<oplus> root) ` cont (pick n)"
-
-(* The regular tree of a function: *)
-definition H :: "N \<Rightarrow> dtree" where
-"H \<equiv> unfold H_r H_c"
-
-lemma finite_H_c: "finite (H_c n)"
-unfolding H_c_def by (metis finite_cont finite_imageI)
-
-lemma root_H_pick: "root (H n) = root (pick n)"
-using unfold(1)[of H_r H_c n] unfolding H_def H_r_def by simp
-
-lemma root_H[simp]:
-assumes "inItr UNIV tr0 n"
-shows "root (H n) = n"
-unfolding root_H_pick root_pick[OF assms] ..
-
-lemma cont_H[simp]:
-"cont (H n) = (id \<oplus> (H o root)) ` cont (pick n)"
-apply(subst id_comp[symmetric, of id]) unfolding map_sum.comp[symmetric]
-unfolding image_comp [symmetric] H_c_def [symmetric]
-using unfold(2) [of H_c n H_r, OF finite_H_c]
-unfolding H_def ..
-
-lemma Inl_cont_H[simp]:
-"Inl -` (cont (H n)) = Inl -` (cont (pick n))"
-unfolding cont_H by simp
-
-lemma Inr_cont_H:
-"Inr -` (cont (H n)) = (H \<circ> root) ` (Inr -` cont (pick n))"
-unfolding cont_H by simp
-
-lemma subtr_H:
-assumes n: "inItr UNIV tr0 n" and "subtr UNIV tr1 (H n)"
-shows "\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = H n1"
-proof-
-  {fix tr ns assume "subtr UNIV tr1 tr"
-   hence "tr = H n \<longrightarrow> (\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = H n1)"
-   proof (induct rule: subtr_UNIV_inductL)
-     case (Step tr2 tr1 tr)
-     show ?case proof
-       assume "tr = H n"
-       then obtain n1 where tr2: "Inr tr2 \<in> cont tr1"
-       and tr1_tr: "subtr UNIV tr1 tr" and n1: "inItr UNIV tr0 n1" and tr1: "tr1 = H n1"
-       using Step by auto
-       obtain tr2' where tr2: "tr2 = H (root tr2')"
-       and tr2': "Inr tr2' \<in> cont (pick n1)"
-       using tr2 Inr_cont_H[of n1]
-       unfolding tr1 image_def comp_def using vimage_eq by auto
-       have "inItr UNIV tr0 (root tr2')"
-       using inItr.Base inItr.Ind n1 pick subtr_inItr tr2' by (metis iso_tuple_UNIV_I)
-       thus "\<exists>n2. inItr UNIV tr0 n2 \<and> tr2 = H n2" using tr2 by blast
-     qed
-   qed(insert n, auto)
-  }
-  thus ?thesis using assms by auto
-qed
-
-lemma root_H_root:
-assumes n: "inItr UNIV tr0 n" and t_tr: "t_tr \<in> cont (pick n)"
-shows "(id \<oplus> (root \<circ> H \<circ> root)) t_tr = (id \<oplus> root) t_tr"
-using assms apply(cases t_tr)
-  apply (metis (lifting) map_sum.simps(1))
-  using pick H_def H_r_def unfold(1)
-      inItr.Base comp_apply subtr_StepL subtr_inItr map_sum.simps(2)
-  by (metis UNIV_I)
-
-lemma H_P:
-assumes tr0: "wf tr0" and n: "inItr UNIV tr0 n"
-shows "(n, (id \<oplus> root) ` cont (H n)) \<in> P" (is "?L \<in> P")
-proof-
-  have "?L = (n, (id \<oplus> root) ` cont (pick n))"
-  unfolding cont_H image_comp map_sum.comp id_comp comp_assoc[symmetric]
-  unfolding Pair_eq apply(rule conjI[OF refl]) apply(rule image_cong[OF refl])
-  by (rule root_H_root[OF n])
-  moreover have "... \<in> P" by (metis (lifting) wf_pick root_pick wf_P n tr0)
-  ultimately show ?thesis by simp
-qed
-
-lemma wf_H:
-assumes tr0: "wf tr0" and "inItr UNIV tr0 n"
-shows "wf (H n)"
-proof-
-  {fix tr have "\<exists> n. inItr UNIV tr0 n \<and> tr = H n \<Longrightarrow> wf tr"
-   proof (induct rule: wf_raw_coind)
-     case (Hyp tr)
-     then obtain n where n: "inItr UNIV tr0 n" and tr: "tr = H n" by auto
-     show ?case apply safe
-     apply (metis (lifting) H_P root_H n tr tr0)
-     unfolding tr Inr_cont_H unfolding inj_on_def apply clarsimp using root_H
-     apply (metis UNIV_I inItr.Base n pick subtr2.simps subtr_inItr subtr_subtr2)
-     by (metis n subtr.Refl subtr_StepL subtr_H tr UNIV_I)
-   qed
-  }
-  thus ?thesis using assms by blast
-qed
-
-(* The regular cut of a tree: *)
-definition "rcut \<equiv> H (root tr0)"
-
-lemma reg_rcut: "reg H rcut"
-unfolding reg_def rcut_def
-by (metis inItr.Base root_H subtr_H UNIV_I)
-
-lemma rcut_reg:
-assumes "reg H tr0"
-shows "rcut = tr0"
-using assms unfolding rcut_def reg_def by (metis subtr.Refl UNIV_I)
-
-lemma rcut_eq: "rcut = tr0 \<longleftrightarrow> reg H tr0"
-using reg_rcut rcut_reg by metis
-
-lemma regular_rcut: "regular rcut"
-using reg_rcut unfolding regular_def by blast
-
-lemma Fr_rcut: "Fr UNIV rcut \<subseteq> Fr UNIV tr0"
-proof safe
-  fix t assume "t \<in> Fr UNIV rcut"
-  then obtain tr where t: "Inl t \<in> cont tr" and tr: "subtr UNIV tr (H (root tr0))"
-  using Fr_subtr[of UNIV "H (root tr0)"] unfolding rcut_def
-  by (metis (full_types) Fr_def inFr_subtr mem_Collect_eq)
-  obtain n where n: "inItr UNIV tr0 n" and tr: "tr = H n" using tr
-  by (metis (lifting) inItr.Base subtr_H UNIV_I)
-  have "Inl t \<in> cont (pick n)" using t using Inl_cont_H[of n] unfolding tr
-  by (metis (lifting) vimageD vimageI2)
-  moreover have "subtr UNIV (pick n) tr0" using subtr_pick[OF n] ..
-  ultimately show "t \<in> Fr UNIV tr0" unfolding Fr_subtr_cont by auto
-qed
-
-lemma wf_rcut:
-assumes "wf tr0"
-shows "wf rcut"
-unfolding rcut_def using wf_H[OF assms inItr.Base] by simp
-
-lemma root_rcut[simp]: "root rcut = root tr0"
-unfolding rcut_def
-by (metis (lifting) root_H inItr.Base reg_def reg_root subtr_rootR_in)
-
-end (* context *)
-
-
-subsection{* Recursive Description of the Regular Tree Frontiers *}
-
-lemma regular_inFr:
-assumes r: "regular tr" and In: "root tr \<in> ns"
-and t: "inFr ns tr t"
-shows "t \<in> Inl -` (cont tr) \<or>
-       (\<exists> tr'. Inr tr' \<in> cont tr \<and> inFr (ns - {root tr}) tr' t)"
-(is "?L \<or> ?R")
-proof-
-  obtain f where r: "reg f tr" and f: "\<And>n. root (f n) = n"
-  using r unfolding regular_def2 by auto
-  obtain nl tr1 where d_nl: "distinct nl" and p: "path f nl" and hd_nl: "f (hd nl) = tr"
-  and l_nl: "f (last nl) = tr1" and s_nl: "set nl \<subseteq> ns" and t_tr1: "Inl t \<in> cont tr1"
-  using t unfolding inFr_iff_path[OF r f] by auto
-  obtain n nl1 where nl: "nl = n # nl1" by (metis (lifting) p path.simps)
-  hence f_n: "f n = tr" using hd_nl by simp
-  have n_nl1: "n \<notin> set nl1" using d_nl unfolding nl by auto
-  show ?thesis
-  proof(cases nl1)
-    case Nil hence "tr = tr1" using f_n l_nl unfolding nl by simp
-    hence ?L using t_tr1 by simp thus ?thesis by simp
-  next
-    case (Cons n1 nl2) note nl1 = Cons
-    have 1: "last nl1 = last nl" "hd nl1 = n1" unfolding nl nl1 by simp_all
-    have p1: "path f nl1" and n1_tr: "Inr (f n1) \<in> cont tr"
-    using path.simps[of f nl] p f_n unfolding nl nl1 by auto
-    have r1: "reg f (f n1)" using reg_Inr_cont[OF r n1_tr] .
-    have 0: "inFr (set nl1) (f n1) t" unfolding inFr_iff_path[OF r1 f]
-    apply(intro exI[of _ nl1], intro exI[of _ tr1])
-    using d_nl unfolding 1 l_nl unfolding nl using p1 t_tr1 by auto
-    have root_tr: "root tr = n" by (metis f f_n)
-    have "inFr (ns - {root tr}) (f n1) t" apply(rule inFr_mono[OF 0])
-    using s_nl unfolding root_tr unfolding nl using n_nl1 by auto
-    thus ?thesis using n1_tr by auto
-  qed
-qed
-
-lemma regular_Fr:
-assumes r: "regular tr" and In: "root tr \<in> ns"
-shows "Fr ns tr =
-       Inl -` (cont tr) \<union>
-       \<Union> {Fr (ns - {root tr}) tr' | tr'. Inr tr' \<in> cont tr}"
-unfolding Fr_def
-using In inFr.Base regular_inFr[OF assms] apply safe
-apply (simp, metis (full_types) mem_Collect_eq)
-apply simp
-by (simp, metis (lifting) inFr_Ind_minus insert_Diff)
-
-
-subsection{* The Generated Languages *}
-
-(* The (possibly inifinite tree) generated language *)
-definition "L ns n \<equiv> {Fr ns tr | tr. wf tr \<and> root tr = n}"
-
-(* The regular-tree generated language *)
-definition "Lr ns n \<equiv> {Fr ns tr | tr. wf tr \<and> root tr = n \<and> regular tr}"
-
-lemma L_rec_notin:
-assumes "n \<notin> ns"
-shows "L ns n = {{}}"
-using assms unfolding L_def apply safe
-  using not_root_Fr apply force
-  apply(rule exI[of _ "deftr n"])
-  by (metis (no_types) wf_deftr not_root_Fr root_deftr)
-
-lemma Lr_rec_notin:
-assumes "n \<notin> ns"
-shows "Lr ns n = {{}}"
-using assms unfolding Lr_def apply safe
-  using not_root_Fr apply force
-  apply(rule exI[of _ "deftr n"])
-  by (metis (no_types) regular_def wf_deftr not_root_Fr reg_deftr root_deftr)
-
-lemma wf_subtrOf:
-assumes "wf tr" and "Inr n \<in> prodOf tr"
-shows "wf (subtrOf tr n)"
-by (metis assms wf_cont subtrOf)
-
-lemma Lr_rec_in:
-assumes n: "n \<in> ns"
-shows "Lr ns n \<subseteq>
-{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
-    (n,tns) \<in> P \<and>
-    (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n')}"
-(is "Lr ns n \<subseteq> {?F tns K | tns K. (n,tns) \<in> P \<and> ?\<phi> tns K}")
-proof safe
-  fix ts assume "ts \<in> Lr ns n"
-  then obtain tr where dtr: "wf tr" and r: "root tr = n" and tr: "regular tr"
-  and ts: "ts = Fr ns tr" unfolding Lr_def by auto
-  def tns \<equiv> "(id \<oplus> root) ` (cont tr)"
-  def K \<equiv> "\<lambda> n'. Fr (ns - {n}) (subtrOf tr n')"
-  show "\<exists>tns K. ts = ?F tns K \<and> (n, tns) \<in> P \<and> ?\<phi> tns K"
-  apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI)
-    show "ts = Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns}"
-    unfolding ts regular_Fr[OF tr n[unfolded r[symmetric]]]
-    unfolding tns_def K_def r[symmetric]
-    unfolding Inl_prodOf wf_subtrOf_Union[OF dtr] ..
-    show "(n, tns) \<in> P" unfolding tns_def r[symmetric] using wf_P[OF dtr] .
-    fix n' assume "Inr n' \<in> tns" thus "K n' \<in> Lr (ns - {n}) n'"
-    unfolding K_def Lr_def mem_Collect_eq apply(intro exI[of _ "subtrOf tr n'"])
-    using dtr tr apply(intro conjI refl)  unfolding tns_def
-      apply(erule wf_subtrOf[OF dtr])
-      apply (metis subtrOf)
-      by (metis Inr_subtrOf UNIV_I regular_subtr subtr.simps)
-  qed
-qed
-
-lemma hsubst_aux:
-fixes n ftr tns
-assumes n: "n \<in> ns" and tns: "finite tns" and
-1: "\<And> n'. Inr n' \<in> tns \<Longrightarrow> wf (ftr n')"
-defines "tr \<equiv> Node n ((id \<oplus> ftr) ` tns)"  defines "tr' \<equiv> hsubst tr tr"
-shows "Fr ns tr' = Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
-(is "_ = ?B") proof-
-  have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
-  unfolding tr_def using tns by auto
-  have Frr: "Frr (ns - {n}) tr = \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
-  unfolding Frr_def ctr by auto
-  have "Fr ns tr' = Inl -` (cont tr) \<union> Frr (ns - {n}) tr"
-  using Fr_self_hsubst[OF n[unfolded rtr[symmetric]]] unfolding tr'_def rtr ..
-  also have "... = ?B" unfolding ctr Frr by simp
-  finally show ?thesis .
-qed
-
-lemma L_rec_in:
-assumes n: "n \<in> ns"
-shows "
-{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
-    (n,tns) \<in> P \<and>
-    (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n')}
- \<subseteq> L ns n"
-proof safe
-  fix tns K
-  assume P: "(n, tns) \<in> P" and 0: "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n'"
-  {fix n' assume "Inr n' \<in> tns"
-   hence "K n' \<in> L (ns - {n}) n'" using 0 by auto
-   hence "\<exists> tr'. K n' = Fr (ns - {n}) tr' \<and> wf tr' \<and> root tr' = n'"
-   unfolding L_def mem_Collect_eq by auto
-  }
-  then obtain ftr where 0: "\<And> n'. Inr n' \<in> tns \<Longrightarrow>
-  K n' = Fr (ns - {n}) (ftr n') \<and> wf (ftr n') \<and> root (ftr n') = n'"
-  by metis
-  def tr \<equiv> "Node n ((id \<oplus> ftr) ` tns)"  def tr' \<equiv> "hsubst tr tr"
-  have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
-  unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P)
-  have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong)
-  unfolding ctr apply simp apply simp apply safe
-  using 0 unfolding image_def apply force apply simp by (metis 0 vimageI2)
-  have 1: "{K n' |n'. Inr n' \<in> tns} = {Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
-  using 0 by auto
-  have dtr: "wf tr" apply(rule wf.dtree)
-    apply (metis (lifting) P prtr rtr)
-    unfolding inj_on_def ctr using 0 by auto
-  hence dtr': "wf tr'" unfolding tr'_def by (metis wf_hsubst)
-  have tns: "finite tns" using finite_in_P P by simp
-  have "Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns} \<in> L ns n"
-  unfolding L_def mem_Collect_eq apply(intro exI[of _ tr'] conjI)
-  using dtr' 0 hsubst_aux[OF assms tns, of ftr] unfolding tr_def tr'_def by auto
-  thus "Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} \<in> L ns n" unfolding 1 .
-qed
-
-lemma card_N: "(n::N) \<in> ns \<Longrightarrow> card (ns - {n}) < card ns"
-by (metis finite_N Diff_UNIV Diff_infinite_finite card_Diff1_less finite.emptyI)
-
-function LL where
-"LL ns n =
- (if n \<notin> ns then {{}} else
- {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
-    (n,tns) \<in> P \<and>
-    (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n')})"
-by(pat_completeness, auto)
-termination apply(relation "inv_image (measure card) fst")
-using card_N by auto
-
-declare LL.simps[code]
-declare LL.simps[simp del]
-
-lemma Lr_LL: "Lr ns n \<subseteq> LL ns n"
-proof (induct ns arbitrary: n rule: measure_induct[of card])
-  case (1 ns n) show ?case proof(cases "n \<in> ns")
-    case False thus ?thesis unfolding Lr_rec_notin[OF False] by (simp add: LL.simps)
-  next
-    case True show ?thesis apply(rule subset_trans)
-    using Lr_rec_in[OF True] apply assumption
-    unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
-      fix tns K
-      assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
-      assume "(n, tns) \<in> P"
-      and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n'"
-      thus "\<exists>tnsa Ka.
-             Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
-             Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
-             (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> LL (ns - {n}) n')"
-      apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
-    qed
-  qed
-qed
-
-lemma LL_L: "LL ns n \<subseteq> L ns n"
-proof (induct ns arbitrary: n rule: measure_induct[of card])
-  case (1 ns n) show ?case proof(cases "n \<in> ns")
-    case False thus ?thesis unfolding L_rec_notin[OF False] by (simp add: LL.simps)
-  next
-    case True show ?thesis apply(rule subset_trans)
-    prefer 2 using L_rec_in[OF True] apply assumption
-    unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
-      fix tns K
-      assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
-      assume "(n, tns) \<in> P"
-      and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n'"
-      thus "\<exists>tnsa Ka.
-             Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
-             Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
-             (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> L (ns - {n}) n')"
-      apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
-    qed
-  qed
-qed
-
-(* The subsumpsion relation between languages *)
-definition "subs L1 L2 \<equiv> \<forall> ts2 \<in> L2. \<exists> ts1 \<in> L1. ts1 \<subseteq> ts2"
-
-lemma incl_subs[simp]: "L2 \<subseteq> L1 \<Longrightarrow> subs L1 L2"
-unfolding subs_def by auto
-
-lemma subs_refl[simp]: "subs L1 L1" unfolding subs_def by auto
-
-lemma subs_trans: "\<lbrakk>subs L1 L2; subs L2 L3\<rbrakk> \<Longrightarrow> subs L1 L3"
-unfolding subs_def by (metis subset_trans)
-
-(* Language equivalence *)
-definition "leqv L1 L2 \<equiv> subs L1 L2 \<and> subs L2 L1"
-
-lemma subs_leqv[simp]: "leqv L1 L2 \<Longrightarrow> subs L1 L2"
-unfolding leqv_def by auto
-
-lemma subs_leqv_sym[simp]: "leqv L1 L2 \<Longrightarrow> subs L2 L1"
-unfolding leqv_def by auto
-
-lemma leqv_refl[simp]: "leqv L1 L1" unfolding leqv_def by auto
-
-lemma leqv_trans:
-assumes 12: "leqv L1 L2" and 23: "leqv L2 L3"
-shows "leqv L1 L3"
-using assms unfolding leqv_def by (metis (lifting) subs_trans)
-
-lemma leqv_sym: "leqv L1 L2 \<Longrightarrow> leqv L2 L1"
-unfolding leqv_def by auto
-
-lemma leqv_Sym: "leqv L1 L2 \<longleftrightarrow> leqv L2 L1"
-unfolding leqv_def by auto
-
-lemma Lr_incl_L: "Lr ns ts \<subseteq> L ns ts"
-unfolding Lr_def L_def by auto
-
-lemma Lr_subs_L: "subs (Lr UNIV ts) (L UNIV ts)"
-unfolding subs_def proof safe
-  fix ts2 assume "ts2 \<in> L UNIV ts"
-  then obtain tr where ts2: "ts2 = Fr UNIV tr" and dtr: "wf tr" and rtr: "root tr = ts"
-  unfolding L_def by auto
-  thus "\<exists>ts1\<in>Lr UNIV ts. ts1 \<subseteq> ts2"
-  apply(intro bexI[of _ "Fr UNIV (rcut tr)"])
-  unfolding Lr_def L_def using Fr_rcut wf_rcut root_rcut regular_rcut by auto
-qed
-
-lemma Lr_leqv_L: "leqv (Lr UNIV ts) (L UNIV ts)"
-using Lr_subs_L unfolding leqv_def by (metis (lifting) Lr_incl_L incl_subs)
-
-lemma LL_leqv_L: "leqv (LL UNIV ts) (L UNIV ts)"
-by (metis (lifting) LL_L Lr_LL Lr_subs_L incl_subs leqv_def subs_trans)
-
-lemma LL_leqv_Lr: "leqv (LL UNIV ts) (Lr UNIV ts)"
-using Lr_leqv_L LL_leqv_L by (metis leqv_Sym leqv_trans)
-
-end
--- a/src/HOL/BNF_Examples/Derivation_Trees/Parallel.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,147 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Derivation_Trees/Parallel.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Parallel composition.
-*)
-
-header {* Parallel Composition *}
-
-theory Parallel
-imports DTree
-begin
-
-no_notation plus_class.plus (infixl "+" 65)
-
-consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl "+" 60)
-
-axiomatization where
-    Nplus_comm: "(a::N) + b = b + (a::N)"
-and Nplus_assoc: "((a::N) + b) + c = a + (b + c)"
-
-subsection{* Corecursive Definition of Parallel Composition *}
-
-fun par_r where "par_r (tr1,tr2) = root tr1 + root tr2"
-fun par_c where
-"par_c (tr1,tr2) =
- Inl ` (Inl -` (cont tr1 \<union> cont tr2)) \<union>
- Inr ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
-
-declare par_r.simps[simp del]  declare par_c.simps[simp del]
-
-definition par :: "dtree \<times> dtree \<Rightarrow> dtree" where
-"par \<equiv> unfold par_r par_c"
-
-abbreviation par_abbr (infixr "\<parallel>" 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
-
-lemma finite_par_c: "finite (par_c (tr1, tr2))"
-unfolding par_c.simps apply(rule finite_UnI)
-  apply (metis finite_Un finite_cont finite_imageI finite_vimageI inj_Inl)
-  apply(intro finite_imageI finite_cartesian_product finite_vimageI)
-  using finite_cont by auto
-
-lemma root_par: "root (tr1 \<parallel> tr2) = root tr1 + root tr2"
-using unfold(1)[of par_r par_c "(tr1,tr2)"] unfolding par_def par_r.simps by simp
-
-lemma cont_par:
-"cont (tr1 \<parallel> tr2) = (id \<oplus> par) ` par_c (tr1,tr2)"
-using unfold(2)[of par_c "(tr1,tr2)" par_r, OF finite_par_c]
-unfolding par_def ..
-
-lemma Inl_cont_par[simp]:
-"Inl -` (cont (tr1 \<parallel> tr2)) = Inl -` (cont tr1 \<union> cont tr2)"
-unfolding cont_par par_c.simps by auto
-
-lemma Inr_cont_par[simp]:
-"Inr -` (cont (tr1 \<parallel> tr2)) = par ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
-unfolding cont_par par_c.simps by auto
-
-lemma Inl_in_cont_par:
-"Inl t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (Inl t \<in> cont tr1 \<or> Inl t \<in> cont tr2)"
-using Inl_cont_par[of tr1 tr2] unfolding vimage_def by auto
-
-lemma Inr_in_cont_par:
-"Inr t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (t \<in> par ` (Inr -` cont tr1 \<times> Inr -` cont tr2))"
-using Inr_cont_par[of tr1 tr2] unfolding vimage_def by auto
-
-
-subsection{* Structural Coinduction Proofs *}
-
-lemma rel_set_rel_sum_eq[simp]:
-"rel_set (rel_sum (op =) \<phi>) A1 A2 \<longleftrightarrow>
- Inl -` A1 = Inl -` A2 \<and> rel_set \<phi> (Inr -` A1) (Inr -` A2)"
-unfolding rel_set_rel_sum rel_set_eq ..
-
-(* Detailed proofs of commutativity and associativity: *)
-theorem par_com: "tr1 \<parallel> tr2 = tr2 \<parallel> tr1"
-proof-
-  let ?\<theta> = "\<lambda> trA trB. \<exists> tr1 tr2. trA = tr1 \<parallel> tr2 \<and> trB = tr2 \<parallel> tr1"
-  {fix trA trB
-   assume "?\<theta> trA trB" hence "trA = trB"
-   apply (induct rule: dtree_coinduct)
-   unfolding rel_set_rel_sum rel_set_eq unfolding rel_set_def proof safe
-     fix tr1 tr2  show "root (tr1 \<parallel> tr2) = root (tr2 \<parallel> tr1)"
-     unfolding root_par by (rule Nplus_comm)
-   next
-     fix n tr1 tr2 assume "Inl n \<in> cont (tr1 \<parallel> tr2)" thus "n \<in> Inl -` (cont (tr2 \<parallel> tr1))"
-     unfolding Inl_in_cont_par by auto
-   next
-     fix n tr1 tr2 assume "Inl n \<in> cont (tr2 \<parallel> tr1)" thus "n \<in> Inl -` (cont (tr1 \<parallel> tr2))"
-     unfolding Inl_in_cont_par by auto
-   next
-     fix tr1 tr2 trA' assume "Inr trA' \<in> cont (tr1 \<parallel> tr2)"
-     then obtain tr1' tr2' where "trA' = tr1' \<parallel> tr2'"
-     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
-     unfolding Inr_in_cont_par by auto
-     thus "\<exists> trB' \<in> Inr -` (cont (tr2 \<parallel> tr1)). ?\<theta> trA' trB'"
-     apply(intro bexI[of _ "tr2' \<parallel> tr1'"]) unfolding Inr_in_cont_par by auto
-   next
-     fix tr1 tr2 trB' assume "Inr trB' \<in> cont (tr2 \<parallel> tr1)"
-     then obtain tr1' tr2' where "trB' = tr2' \<parallel> tr1'"
-     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
-     unfolding Inr_in_cont_par by auto
-     thus "\<exists> trA' \<in> Inr -` (cont (tr1 \<parallel> tr2)). ?\<theta> trA' trB'"
-     apply(intro bexI[of _ "tr1' \<parallel> tr2'"]) unfolding Inr_in_cont_par by auto
-   qed
-  }
-  thus ?thesis by blast
-qed
-
-lemma par_assoc: "(tr1 \<parallel> tr2) \<parallel> tr3 = tr1 \<parallel> (tr2 \<parallel> tr3)"
-proof-
-  let ?\<theta> =
-  "\<lambda> trA trB. \<exists> tr1 tr2 tr3. trA = (tr1 \<parallel> tr2) \<parallel> tr3 \<and> trB = tr1 \<parallel> (tr2 \<parallel> tr3)"
-  {fix trA trB
-   assume "?\<theta> trA trB" hence "trA = trB"
-   apply (induct rule: dtree_coinduct)
-   unfolding rel_set_rel_sum rel_set_eq unfolding rel_set_def proof safe
-     fix tr1 tr2 tr3  show "root ((tr1 \<parallel> tr2) \<parallel> tr3) = root (tr1 \<parallel> (tr2 \<parallel> tr3))"
-     unfolding root_par by (rule Nplus_assoc)
-   next
-     fix n tr1 tr2 tr3 assume "Inl n \<in> (cont ((tr1 \<parallel> tr2) \<parallel> tr3))"
-     thus "n \<in> Inl -` (cont (tr1 \<parallel> tr2 \<parallel> tr3))" unfolding Inl_in_cont_par by simp
-   next
-     fix n tr1 tr2 tr3 assume "Inl n \<in> (cont (tr1 \<parallel> tr2 \<parallel> tr3))"
-     thus "n \<in> Inl -` (cont ((tr1 \<parallel> tr2) \<parallel> tr3))" unfolding Inl_in_cont_par by simp
-   next
-     fix trA' tr1 tr2 tr3 assume "Inr trA' \<in> cont ((tr1 \<parallel> tr2) \<parallel> tr3)"
-     then obtain tr1' tr2' tr3' where "trA' = (tr1' \<parallel> tr2') \<parallel> tr3'"
-     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
-     and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
-     thus "\<exists> trB' \<in> Inr -` (cont (tr1 \<parallel> tr2 \<parallel> tr3)). ?\<theta> trA' trB'"
-     apply(intro bexI[of _ "tr1' \<parallel> tr2' \<parallel> tr3'"])
-     unfolding Inr_in_cont_par by auto
-   next
-     fix trB' tr1 tr2 tr3 assume "Inr trB' \<in> cont (tr1 \<parallel> tr2 \<parallel> tr3)"
-     then obtain tr1' tr2' tr3' where "trB' = tr1' \<parallel> (tr2' \<parallel> tr3')"
-     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
-     and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
-     thus "\<exists> trA' \<in> Inr -` cont ((tr1 \<parallel> tr2) \<parallel> tr3). ?\<theta> trA' trB'"
-     apply(intro bexI[of _ "(tr1' \<parallel> tr2') \<parallel> tr3'"])
-     unfolding Inr_in_cont_par by auto
-   qed
-  }
-  thus ?thesis by blast
-qed
-
-end
--- a/src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,64 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Derivation_Trees/Prelim.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Preliminaries.
-*)
-
-header {* Preliminaries *}
-
-theory Prelim
-imports "~~/src/HOL/Library/FSet"
-begin
-
-notation BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
-
-declare fset_to_fset[simp]
-
-lemma fst_snd_convol_o[simp]: "\<langle>fst o s, snd o s\<rangle> = s"
-apply(rule ext) by (simp add: convol_def)
-
-abbreviation sm_abbrev (infix "\<oplus>" 60)
-where "f \<oplus> g \<equiv> Sum_Type.map_sum f g"
-
-lemma map_sum_InlD: "(f \<oplus> g) z = Inl x \<Longrightarrow> \<exists>y. z = Inl y \<and> f y = x"
-by (cases z) auto
-
-lemma map_sum_InrD: "(f \<oplus> g) z = Inr x \<Longrightarrow> \<exists>y. z = Inr y \<and> g y = x"
-by (cases z) auto
-
-abbreviation case_sum_abbrev ("[[_,_]]" 800)
-where "[[f,g]] \<equiv> Sum_Type.case_sum f g"
-
-lemma Inl_oplus_elim:
-assumes "Inl tr \<in> (id \<oplus> f) ` tns"
-shows "Inl tr \<in> tns"
-using assms apply clarify by (case_tac x, auto)
-
-lemma Inl_oplus_iff[simp]: "Inl tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> Inl tr \<in> tns"
-using Inl_oplus_elim
-by (metis id_def image_iff map_sum.simps(1))
-
-lemma Inl_m_oplus[simp]: "Inl -` (id \<oplus> f) ` tns = Inl -` tns"
-using Inl_oplus_iff unfolding vimage_def by auto
-
-lemma Inr_oplus_elim:
-assumes "Inr tr \<in> (id \<oplus> f) ` tns"
-shows "\<exists> n. Inr n \<in> tns \<and> f n = tr"
-using assms apply clarify by (case_tac x, auto)
-
-lemma Inr_oplus_iff[simp]:
-"Inr tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> (\<exists> n. Inr n \<in> tns \<and> f n = tr)"
-apply (rule iffI)
- apply (metis Inr_oplus_elim)
-by (metis image_iff map_sum.simps(2))
-
-lemma Inr_m_oplus[simp]: "Inr -` (id \<oplus> f) ` tns = f ` (Inr -` tns)"
-using Inr_oplus_iff unfolding vimage_def by auto
-
-lemma Inl_Inr_image_cong:
-assumes "Inl -` A = Inl -` B" and "Inr -` A = Inr -` B"
-shows "A = B"
-apply safe using assms apply(case_tac x, auto) by(case_tac x, auto)
-
-end
\ No newline at end of file
--- a/src/HOL/BNF_Examples/Instructions.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,162 +0,0 @@
-(*  Title:      HOL/Datatype_Benchmark/Instructions.thy
-
-Example from Konrad: 68000 instruction set.
-*)
-
-theory Instructions imports Main begin
-
-datatype Size = Byte | Word | Long
-
-datatype DataRegister =
-  RegD0
-| RegD1
-| RegD2
-| RegD3
-| RegD4
-| RegD5
-| RegD6
-| RegD7
-
-datatype AddressRegister =
-  RegA0
-| RegA1
-| RegA2
-| RegA3
-| RegA4
-| RegA5
-| RegA6
-| RegA7
-
-datatype DataOrAddressRegister =
-  data DataRegister
-| address AddressRegister
-
-datatype Condition =
-  Hi
-| Ls
-| Cc
-| Cs
-| Ne
-| Eq
-| Vc
-| Vs
-| Pl
-| Mi
-| Ge
-| Lt
-| Gt
-| Le
-
-datatype AddressingMode =
-  immediate nat
-| direct DataOrAddressRegister
-| indirect AddressRegister
-| postinc AddressRegister
-| predec AddressRegister
-| indirectdisp nat AddressRegister
-| indirectindex nat AddressRegister DataOrAddressRegister Size
-| absolute nat
-| pcdisp nat
-| pcindex nat DataOrAddressRegister Size
-
-datatype M68kInstruction =
-  ABCD AddressingMode AddressingMode
-| ADD Size AddressingMode AddressingMode
-| ADDA Size AddressingMode AddressRegister
-| ADDI Size nat AddressingMode
-| ADDQ Size nat AddressingMode
-| ADDX Size AddressingMode AddressingMode
-| AND Size AddressingMode AddressingMode
-| ANDI Size nat AddressingMode
-| ANDItoCCR nat
-| ANDItoSR nat
-| ASL Size AddressingMode DataRegister
-| ASLW AddressingMode
-| ASR Size AddressingMode DataRegister
-| ASRW AddressingMode
-| Bcc Condition Size nat
-| BTST Size AddressingMode AddressingMode
-| BCHG Size AddressingMode AddressingMode
-| BCLR Size AddressingMode AddressingMode
-| BSET Size AddressingMode AddressingMode
-| BRA Size nat
-| BSR Size nat
-| CHK AddressingMode DataRegister
-| CLR Size AddressingMode
-| CMP Size AddressingMode DataRegister
-| CMPA Size AddressingMode AddressRegister
-| CMPI Size nat AddressingMode
-| CMPM Size AddressRegister AddressRegister
-| DBT DataRegister nat
-| DBF DataRegister nat
-| DBcc Condition DataRegister nat
-| DIVS AddressingMode DataRegister
-| DIVU AddressingMode DataRegister
-| EOR Size DataRegister AddressingMode
-| EORI Size nat AddressingMode
-| EORItoCCR nat
-| EORItoSR nat
-| EXG DataOrAddressRegister DataOrAddressRegister
-| EXT Size DataRegister
-| ILLEGAL
-| JMP AddressingMode
-| JSR AddressingMode
-| LEA AddressingMode AddressRegister
-| LINK AddressRegister nat
-| LSL Size AddressingMode DataRegister
-| LSLW AddressingMode
-| LSR Size AddressingMode DataRegister
-| LSRW AddressingMode
-| MOVE Size AddressingMode AddressingMode
-| MOVEtoCCR AddressingMode
-| MOVEtoSR AddressingMode
-| MOVEfromSR AddressingMode
-| MOVEtoUSP AddressingMode
-| MOVEfromUSP AddressingMode
-| MOVEA Size AddressingMode AddressRegister
-| MOVEMto Size AddressingMode "DataOrAddressRegister list"
-| MOVEMfrom Size "DataOrAddressRegister list" AddressingMode
-| MOVEP Size AddressingMode AddressingMode
-| MOVEQ nat DataRegister
-| MULS AddressingMode DataRegister
-| MULU AddressingMode DataRegister
-| NBCD AddressingMode
-| NEG Size AddressingMode
-| NEGX Size AddressingMode
-| NOP
-| NOT Size AddressingMode
-| OR Size AddressingMode AddressingMode
-| ORI Size nat AddressingMode
-| ORItoCCR nat
-| ORItoSR nat
-| PEA AddressingMode
-| RESET
-| ROL Size AddressingMode DataRegister
-| ROLW AddressingMode
-| ROR Size AddressingMode DataRegister
-| RORW AddressingMode
-| ROXL Size AddressingMode DataRegister
-| ROXLW AddressingMode
-| ROXR Size AddressingMode DataRegister
-| ROXRW AddressingMode
-| RTE
-| RTR
-| RTS
-| SBCD AddressingMode AddressingMode
-| ST AddressingMode
-| SF AddressingMode
-| Scc Condition AddressingMode
-| STOP nat
-| SUB Size AddressingMode AddressingMode
-| SUBA Size AddressingMode AddressingMode
-| SUBI Size nat AddressingMode
-| SUBQ Size nat AddressingMode
-| SUBX Size AddressingMode AddressingMode
-| SWAP DataRegister
-| TAS AddressingMode
-| TRAP nat
-| TRAPV
-| TST Size AddressingMode
-| UNLK AddressRegister
-
-end
--- a/src/HOL/BNF_Examples/IsaFoR_Datatypes.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,380 +0,0 @@
-(*  Title:      HOL/BNF_Examples/IsaFoR_Datatypes.thy
-    Author:     Rene Thiemann, UIBK
-    Copyright   2014
-
-Benchmark consisting of datatypes defined in IsaFoR.
-*)
-
-header {* Benchmark Consisting of Datatypes Defined in IsaFoR *}
-
-theory IsaFoR_Datatypes
-imports Real
-begin
-
-datatype_new (discs_sels) ('f, 'l) lab =
-    Lab "('f, 'l) lab" 'l
-  | FunLab "('f, 'l) lab" "('f, 'l) lab list"
-  | UnLab 'f
-  | Sharp "('f, 'l) lab"
-
-datatype_new (discs_sels) 'f projL = Projection "(('f \<times> nat) \<times> nat) list"
-
-datatype_new (discs_sels) ('f, 'v) "term" = Var 'v | Fun 'f "('f, 'v) term list"
-datatype_new (discs_sels) ('f, 'v) ctxt =
-    Hole ("\<box>")
-  | More 'f "('f, 'v) term list" "('f, 'v) ctxt" "('f, 'v) term list"
-
-type_synonym ('f, 'v) rule = "('f, 'v) term \<times> ('f, 'v) term"
-type_synonym ('f, 'v) trs  = "('f, 'v) rule set"
-
-type_synonym ('f, 'v) rules = "('f, 'v) rule list"
-type_synonym ('f, 'l, 'v) ruleLL  = "(('f, 'l) lab, 'v) rule"
-type_synonym ('f, 'l, 'v) trsLL   = "(('f, 'l) lab, 'v) rules"
-type_synonym ('f, 'l, 'v) termsLL = "(('f, 'l) lab, 'v) term list"
-
-datatype_new (discs_sels) pos = Empty ("\<epsilon>") | PCons "nat" "pos" (infixr "<#" 70)
-
-type_synonym  ('f, 'v) prseq = "(pos \<times> ('f, 'v) rule \<times> bool \<times> ('f, 'v) term) list"
-type_synonym  ('f, 'v) rseq = "(pos \<times> ('f, 'v) rule \<times> ('f, 'v) term) list"
-
-type_synonym ('f, 'l, 'v) rseqL   = "((('f, 'l) lab, 'v) rule \<times> (('f, 'l) lab, 'v) rseq) list"
-type_synonym ('f, 'l, 'v) dppLL   =
-  "bool \<times> bool \<times> ('f, 'l, 'v) trsLL \<times> ('f, 'l, 'v) trsLL \<times>
-  ('f, 'l, 'v) termsLL \<times>
-  ('f, 'l, 'v) trsLL \<times> ('f, 'l, 'v) trsLL"
-
-type_synonym ('f, 'l, 'v) qreltrsLL =
-  "bool \<times> ('f, 'l, 'v) termsLL \<times> ('f, 'l, 'v) trsLL \<times> ('f, 'l, 'v) trsLL"
-
-type_synonym ('f, 'l, 'v) qtrsLL =
-  "bool \<times> ('f, 'l, 'v) termsLL \<times> ('f, 'l, 'v) trsLL"
-
-datatype_new (discs_sels) location = H | A | B | R
-
-type_synonym ('f, 'v) forb_pattern = "('f, 'v) ctxt \<times> ('f, 'v) term \<times> location"
-type_synonym ('f, 'v) forb_patterns = "('f, 'v) forb_pattern set"
-
-type_synonym ('f, 'l, 'v) fptrsLL =
-  "(('f, 'l) lab, 'v) forb_pattern list \<times> ('f, 'l, 'v) trsLL"
-
-type_synonym ('f, 'l, 'v) prob = "('f, 'l, 'v) qreltrsLL + ('f, 'l, 'v) dppLL"
-
-type_synonym ('f, 'a) lpoly_inter = "'f \<times> nat \<Rightarrow> ('a \<times> 'a list)"
-type_synonym ('f, 'a) lpoly_interL = "(('f \<times> nat) \<times> ('a \<times> 'a list)) list"
-
-type_synonym 'v monom = "('v \<times> nat) list"
-type_synonym ('v, 'a) poly = "('v monom \<times> 'a) list"
-type_synonym ('f, 'a) poly_inter_list = "(('f \<times> nat) \<times> (nat, 'a) poly) list"
-type_synonym 'a vec = "'a list"
-type_synonym 'a mat = "'a vec list"
-
-datatype_new (discs_sels) arctic = MinInfty | Num_arc int
-datatype_new (discs_sels) 'a arctic_delta = MinInfty_delta | Num_arc_delta 'a
-datatype_new (discs_sels) order_tag = Lex | Mul
-
-type_synonym 'f status_prec_repr = "(('f \<times> nat) \<times> (nat \<times> order_tag)) list"
-
-datatype_new (discs_sels) af_entry =
-    Collapse nat
-  | AFList "nat list"
-
-type_synonym 'f afs_list = "(('f \<times> nat) \<times> af_entry) list"
-type_synonym 'f prec_weight_repr = "(('f \<times> nat) \<times> (nat \<times> nat \<times> (nat list option))) list \<times> nat"
-
-datatype_new (discs_sels) 'f redtriple_impl =
-    Int_carrier "('f, int) lpoly_interL"
-  | Int_nl_carrier "('f, int) poly_inter_list"
-  | Rat_carrier "('f, rat) lpoly_interL"
-  | Rat_nl_carrier rat "('f, rat) poly_inter_list"
-  | Real_carrier "('f, real) lpoly_interL"
-  | Real_nl_carrier real "('f, real) poly_inter_list"
-  | Arctic_carrier "('f, arctic) lpoly_interL"
-  | Arctic_rat_carrier "('f, rat arctic_delta) lpoly_interL"
-  | Int_mat_carrier nat nat "('f, int mat) lpoly_interL"
-  | Rat_mat_carrier nat nat "('f, rat mat) lpoly_interL"
-  | Real_mat_carrier nat nat "('f, real mat) lpoly_interL"
-  | Arctic_mat_carrier nat "('f, arctic mat) lpoly_interL"
-  | Arctic_rat_mat_carrier nat "('f, rat arctic_delta mat) lpoly_interL"
-  | RPO "'f status_prec_repr" "'f afs_list"
-  | KBO "'f prec_weight_repr" "'f afs_list"
-
-datatype_new (discs_sels) list_order_type = MS_Ext | Max_Ext | Min_Ext  | Dms_Ext
-type_synonym 'f scnp_af = "(('f \<times> nat) \<times> (nat \<times> nat) list) list"
-
-datatype_new (discs_sels) 'f root_redtriple_impl = SCNP list_order_type "'f scnp_af" "'f redtriple_impl"
-
-type_synonym 'f sig_map_list = "(('f \<times> nat) \<times> 'f list) list"
-type_synonym ('f, 'v) uncurry_info = "'f \<times> 'f sig_map_list \<times> ('f, 'v) rules \<times> ('f, 'v) rules"
-
-datatype_new (discs_sels) arithFun =
-    Arg nat
-  | Const nat
-  | Sum "arithFun list"
-  | Max "arithFun list"
-  | Min "arithFun list"
-  | Prod "arithFun list"
-  | IfEqual arithFun arithFun arithFun arithFun
-
-datatype_new (discs_sels) 'f sl_inter = SL_Inter nat "(('f \<times> nat) \<times> arithFun) list"
-datatype_new (discs_sels) ('f, 'v) sl_variant =
-    Rootlab "('f \<times> nat) option"
-  | Finitelab "'f sl_inter"
-  | QuasiFinitelab "'f sl_inter" 'v
-
-type_synonym ('f, 'v) crit_pair_joins = "(('f, 'v) term \<times> ('f, 'v) rseq \<times> ('f, 'v) term \<times> ('f, 'v) rseq) list"
-
-datatype_new (discs_sels) 'f join_info = Guided "('f, string) crit_pair_joins" | Join_NF | Join_BFS nat
-
-type_synonym unknown_info = string
-
-type_synonym dummy_prf = unit
-
-datatype_new (discs_sels) ('f, 'v) complex_constant_removal_prf = Complex_Constant_Removal_Proof
-  "('f, 'v) term"
-  "(('f, 'v) rule \<times> ('f, 'v) rule) list"
-
-datatype_new (discs_sels) ('f, 'v) cond_constraint =
-    CC_cond bool "('f, 'v) rule"
-  | CC_rewr "('f, 'v) term" "('f, 'v) term"
-  | CC_impl "('f, 'v) cond_constraint list" "('f, 'v) cond_constraint"
-  | CC_all 'v "('f, 'v) cond_constraint"
-
-type_synonym ('f, 'v, 'w) gsubstL = "('v \<times> ('f, 'w) term) list"
-type_synonym ('f, 'v) substL = "('f, 'v, 'v) gsubstL"
-
-datatype_new (discs_sels) ('f, 'v) cond_constraint_prf =
-    Final
-  | Delete_Condition "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
-  | Different_Constructor "('f, 'v) cond_constraint"
-  | Same_Constructor "('f, 'v) cond_constraint" "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
-  | Variable_Equation 'v "('f, 'v) term" "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
-  | Funarg_Into_Var "('f, 'v) cond_constraint" nat 'v "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
-  | Simplify_Condition "('f, 'v) cond_constraint" "('f, 'v) substL" "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
-  | Induction "('f, 'v) cond_constraint" "('f, 'v) cond_constraint list" "(('f, 'v) rule \<times> (('f, 'v) term \<times> 'v list) list \<times> ('f, 'v) cond_constraint \<times> ('f, 'v) cond_constraint_prf) list"
-
-datatype_new (discs_sels) ('f, 'v) cond_red_pair_prf =
-  Cond_Red_Pair_Prf
-    'f "(('f, 'v) cond_constraint \<times> ('f, 'v) rules \<times> ('f, 'v) cond_constraint_prf) list" nat nat
-
-datatype_new (discs_sels) ('q, 'f) ta_rule = TA_rule 'f "'q list" 'q ("_ _ \<rightarrow> _")
-datatype_new (discs_sels) ('q, 'f) tree_automaton = Tree_Automaton "'q list" "('q, 'f) ta_rule list" "('q \<times> 'q) list"
-datatype_new (discs_sels) 'q ta_relation =
-    Decision_Proc
-  | Id_Relation
-  | Some_Relation "('q \<times> 'q) list"
-
-datatype_new (discs_sels) boundstype = Roof | Match
-datatype_new (discs_sels) ('f, 'q) bounds_info = Bounds_Info boundstype nat "'q list" "('q, 'f \<times> nat) tree_automaton" "'q ta_relation"
-
-datatype_new (discs_sels) ('f, 'v) pat_eqv_prf =
-    Pat_Dom_Renaming "('f, 'v) substL"
-  | Pat_Irrelevant "('f, 'v) substL" "('f, 'v) substL"
-  | Pat_Simplify "('f, 'v) substL" "('f, 'v) substL"
-
-datatype_new (discs_sels) pat_rule_pos = Pat_Base | Pat_Pump | Pat_Close
-
-datatype_new (discs_sels) ('f, 'v) pat_rule_prf =
-    Pat_OrigRule "('f, 'v) rule" bool
-  | Pat_InitPump "('f, 'v) pat_rule_prf" "('f, 'v) substL" "('f, 'v) substL"
-  | Pat_InitPumpCtxt "('f, 'v) pat_rule_prf" "('f, 'v) substL" pos 'v
-  | Pat_Equiv "('f, 'v) pat_rule_prf" bool "('f, 'v) pat_eqv_prf"
-  | Pat_Narrow "('f, 'v) pat_rule_prf" "('f, 'v) pat_rule_prf" pos
-  | Pat_Inst "('f, 'v) pat_rule_prf" "('f, 'v) substL" pat_rule_pos
-  | Pat_Rewr "('f, 'v) pat_rule_prf" "('f, 'v) term \<times> ('f, 'v) rseq" pat_rule_pos 'v
-  | Pat_Exp_Sigma "('f, 'v) pat_rule_prf" nat
-
-datatype_new (discs_sels) ('f, 'v) non_loop_prf =
-    Non_Loop_Prf "('f, 'v) pat_rule_prf" "('f, 'v) substL" "('f, 'v) substL" nat nat pos
-
-datatype_new (discs_sels) ('f, 'l, 'v) problem =
-    SN_TRS "('f, 'l, 'v) qreltrsLL"
-  | SN_FP_TRS "('f, 'l, 'v) fptrsLL"
-  | Finite_DPP "('f, 'l, 'v) dppLL"
-  | Unknown_Problem unknown_info
-  | Not_SN_TRS "('f, 'l, 'v) qtrsLL"
-  | Not_RelSN_TRS "('f, 'l, 'v) qreltrsLL"
-  | Infinite_DPP "('f, 'l, 'v) dppLL"
-  | Not_SN_FP_TRS "('f, 'l, 'v) fptrsLL"
-
-declare [[bnf_timing]]
-
-datatype_new (discs_sels) ('f, 'l, 'v, 'a, 'b, 'c, 'd, 'e) generic_assm_proof =
-    SN_assm_proof "('f, 'l, 'v) qreltrsLL" 'a
-  | Finite_assm_proof "('f, 'l, 'v) dppLL" 'b
-  | SN_FP_assm_proof "('f, 'l, 'v) fptrsLL" 'c
-  | Not_SN_assm_proof "('f, 'l, 'v) qtrsLL" 'a
-  | Infinite_assm_proof "('f, 'l, 'v) dppLL" 'b
-  | Not_RelSN_assm_proof "('f, 'l, 'v) qreltrsLL" 'c
-  | Not_SN_FP_assm_proof "('f, 'l, 'v) fptrsLL" 'd
-  | Unknown_assm_proof unknown_info 'e
-
-type_synonym ('f, 'l, 'v, 'a, 'b, 'c, 'd) assm_proof = "('f, 'l, 'v, 'a, 'b, 'c, dummy_prf, 'd) generic_assm_proof"
-
-datatype_new (discs_sels) ('f, 'l, 'v) assm =
-    SN_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) qreltrsLL"
-  | SN_FP_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) fptrsLL"
-  | Finite_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) dppLL"
-  | Unknown_assm "('f, 'l, 'v) problem list" unknown_info
-  | Not_SN_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) qtrsLL"
-  | Not_RelSN_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) qreltrsLL"
-  | Not_SN_FP_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) fptrsLL"
-  | Infinite_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) dppLL"
-
-fun satisfied :: "('f, 'l, 'v) problem \<Rightarrow> bool" where
-  "satisfied (SN_TRS t) = (t = t)"
-| "satisfied (SN_FP_TRS t) = (t \<noteq> t)"
-| "satisfied (Finite_DPP d) = (d \<noteq> d)"
-| "satisfied (Unknown_Problem s) = (s = ''foo'')"
-| "satisfied (Not_SN_TRS (nfs, q, r)) = (length q = length r)"
-| "satisfied (Not_RelSN_TRS (nfs, q, r, rw)) = (r = rw)"
-| "satisfied (Infinite_DPP d) = (d = d)"
-| "satisfied (Not_SN_FP_TRS t) = (t = t)"
-
-fun collect_assms :: "('tp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('dpp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('fptp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('unk \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('f, 'l, 'v, 'tp, 'dpp, 'fptp, 'unk) assm_proof \<Rightarrow> ('f, 'l, 'v) assm list" where
-  "collect_assms tp dpp fptp unk (SN_assm_proof t prf) = tp prf"
-| "collect_assms tp dpp fptp unk (SN_FP_assm_proof t prf) = fptp prf"
-| "collect_assms tp dpp fptp unk (Finite_assm_proof d prf) = dpp prf"
-| "collect_assms tp dpp fptp unk (Unknown_assm_proof p prf) = unk prf"
-| "collect_assms _ _ _ _ _ = []"
-
-fun collect_neg_assms :: "('tp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('dpp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('rtp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('fptp \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('unk \<Rightarrow> ('f, 'l, 'v) assm list)
-  \<Rightarrow> ('f, 'l, 'v, 'tp, 'dpp, 'rtp, 'fptp, 'unk) generic_assm_proof \<Rightarrow> ('f, 'l, 'v) assm list" where
-  "collect_neg_assms tp dpp rtp fptp unk (Not_SN_assm_proof t prf) = tp prf"
-| "collect_neg_assms tp dpp rtp fptp unk (Infinite_assm_proof d prf) = dpp prf"
-| "collect_neg_assms tp dpp rtp fptp unk (Not_RelSN_assm_proof t prf) = rtp prf"
-| "collect_neg_assms tp dpp rtp fptp unk (Not_SN_FP_assm_proof t prf) = fptp prf"
-| "collect_neg_assms tp dpp rtp fptp unk (Unknown_assm_proof p prf) = unk prf"
-| "collect_neg_assms tp dpp rtp fptp unk _ = []"
-
-datatype_new (discs_sels) ('f, 'l, 'v) dp_nontermination_proof =
-    DP_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) prseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
-  | DP_Nonloop "(('f, 'l) lab, 'v) non_loop_prf"
-  | DP_Rule_Removal "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Q_Increase "('f, 'l, 'v) termsLL" "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Q_Reduction "('f, 'l, 'v) termsLL" "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Termination_Switch "('f, 'l) lab join_info" "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Instantiation "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Rewriting "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL" "(('f, 'l) lab, 'v) rule" pos "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Narrowing "('f, 'l, 'v) ruleLL" pos "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_nontermination_proof"
-  | DP_Assume_Infinite  "('f, 'l, 'v) dppLL"
-      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
-       ('f, 'l, 'v) dp_nontermination_proof,
-       ('f, 'l, 'v) reltrs_nontermination_proof,
-       ('f, 'l, 'v) fp_nontermination_proof,
-       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
-and ('f, 'l, 'v) "trs_nontermination_proof" =
-    TRS_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) rseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
-  | TRS_Not_Well_Formed
-  | TRS_Rule_Removal "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_nontermination_proof"
-  | TRS_String_Reversal "('f, 'l, 'v) trs_nontermination_proof"
-  | TRS_DP_Trans "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_nontermination_proof"
-  | TRS_Nonloop "(('f, 'l) lab, 'v) non_loop_prf"
-  | TRS_Q_Increase "('f, 'l, 'v) termsLL" "('f, 'l, 'v) trs_nontermination_proof"
-  | TRS_Assume_Not_SN  "('f, 'l, 'v) qtrsLL"
-      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
-       ('f, 'l, 'v) dp_nontermination_proof,
-       ('f, 'l, 'v) reltrs_nontermination_proof,
-       ('f, 'l, 'v) fp_nontermination_proof,
-       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
-and ('f, 'l, 'v)"reltrs_nontermination_proof" =
-    Rel_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) prseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
-  | Rel_Not_Well_Formed
-  | Rel_Rule_Removal "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) reltrs_nontermination_proof"
-  | Rel_R_Not_SN "('f, 'l, 'v) trs_nontermination_proof"
-  | Rel_TRS_Assume_Not_SN  "('f, 'l, 'v) qreltrsLL"
-      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
-       ('f, 'l, 'v) dp_nontermination_proof,
-       ('f, 'l, 'v) reltrs_nontermination_proof,
-       ('f, 'l, 'v) fp_nontermination_proof,
-       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
-and ('f, 'l, 'v) "fp_nontermination_proof" =
-    FPTRS_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) rseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
-  | FPTRS_Rule_Removal "('f, 'l, 'v) trsLL" "('f, 'l, 'v) fp_nontermination_proof"
-  | FPTRS_Assume_Not_SN  "('f, 'l, 'v) fptrsLL"
-      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
-       ('f, 'l, 'v) dp_nontermination_proof,
-       ('f, 'l, 'v) reltrs_nontermination_proof,
-       ('f, 'l, 'v) fp_nontermination_proof,
-       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
-and ('f, 'l, 'v) neg_unknown_proof =
-    Assume_NT_Unknown unknown_info
-      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
-       ('f, 'l, 'v) dp_nontermination_proof,
-       ('f, 'l, 'v) reltrs_nontermination_proof,
-       ('f, 'l, 'v) fp_nontermination_proof,
-       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
-
-datatype_new (discs_sels) ('f, 'l, 'v) dp_termination_proof =
-    P_is_Empty
-  | Subterm_Criterion_Proc "('f, 'l) lab projL" "('f, 'l, 'v) rseqL"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Redpair_Proc "('f, 'l) lab root_redtriple_impl + ('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL"  "('f, 'l, 'v) dp_termination_proof"
-  | Redpair_UR_Proc "('f, 'l) lab root_redtriple_impl + ('f, 'l) lab redtriple_impl"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Usable_Rules_Proc "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Dep_Graph_Proc "(('f, 'l, 'v) dp_termination_proof option \<times> ('f, 'l, 'v) trsLL) list"
-  | Mono_Redpair_Proc "('f, 'l) lab redtriple_impl"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Mono_Redpair_UR_Proc "('f, 'l) lab redtriple_impl"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Size_Change_Subterm_Proc "((('f, 'l) lab, 'v) rule \<times> ((nat \<times> nat) list \<times> (nat \<times> nat) list)) list"
-  | Size_Change_Redpair_Proc "('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL option"
-      "((('f, 'l) lab, 'v) rule \<times> ((nat \<times> nat) list \<times> (nat \<times> nat) list)) list"
-  | Uncurry_Proc "nat option" "(('f, 'l) lab, 'v) uncurry_info"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Fcc_Proc "('f, 'l) lab" "(('f, 'l) lab, 'v) ctxt list"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Split_Proc
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL"
-      "('f, 'l, 'v) dp_termination_proof" "('f, 'l, 'v) dp_termination_proof"
-  | Semlab_Proc
-      "(('f, 'l) lab, 'v) sl_variant" "('f, 'l, 'v) trsLL"
-      "(('f, 'l) lab, 'v) term list" "('f, 'l, 'v) trsLL"
-      "('f, 'l, 'v) dp_termination_proof"
-  | Switch_Innermost_Proc "('f, 'l) lab join_info" "('f, 'l, 'v) dp_termination_proof"
-  | Rewriting_Proc
-      "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL"
-      "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL" pos "('f, 'l, 'v) dp_termination_proof"
-  | Instantiation_Proc "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Forward_Instantiation_Proc
-      "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) dp_termination_proof"
-  | Narrowing_Proc "('f, 'l, 'v) ruleLL" pos "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Assume_Finite
-      "('f, 'l, 'v) dppLL" "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
-  | Unlab_Proc "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Q_Reduction_Proc "('f, 'l, 'v) termsLL" "('f, 'l, 'v) dp_termination_proof"
-  | Complex_Constant_Removal_Proc "(('f, 'l) lab, 'v) complex_constant_removal_prf" "('f, 'l, 'v) dp_termination_proof"
-  | General_Redpair_Proc
-      "('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL"
-      "(('f, 'l) lab, 'v) cond_red_pair_prf" "('f, 'l, 'v) dp_termination_proof list"
-  | To_Trs_Proc "('f, 'l, 'v) trs_termination_proof"
-and ('f, 'l, 'v) trs_termination_proof =
-    DP_Trans bool bool "(('f, 'l) lab, 'v) rules" "('f, 'l, 'v) dp_termination_proof"
-  | Rule_Removal "('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
-  | String_Reversal "('f, 'l, 'v) trs_termination_proof"
-  | Bounds "(('f, 'l) lab, 'v) bounds_info"
-  | Uncurry "(('f, 'l) lab, 'v) uncurry_info" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
-  | Semlab
-      "(('f, 'l) lab, 'v) sl_variant" "(('f, 'l) lab, 'v) term list"
-      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
-  | R_is_Empty
-  | Fcc "(('f, 'l) lab, 'v) ctxt list" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
-  | Split "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof" "('f, 'l, 'v) trs_termination_proof"
-  | Switch_Innermost "('f, 'l) lab join_info" "('f, 'l, 'v) trs_termination_proof"
-  | Drop_Equality "('f, 'l, 'v) trs_termination_proof"
-  | Remove_Nonapplicable_Rules "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
-  | Assume_SN "('f, 'l, 'v) qreltrsLL" "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
-and ('f, 'l, 'v) unknown_proof =
-    Assume_Unknown unknown_info "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
-and ('f, 'l, 'v) fptrs_termination_proof =
-    Assume_FP_SN "('f, 'l, 'v) fptrsLL" "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
-
-end
--- a/src/HOL/BNF_Examples/Koenig.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,122 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Koenig.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Koenig's lemma.
-*)
-
-header {* Koenig's Lemma *}
-
-theory Koenig
-imports TreeFI Stream
-begin
-
-(* infinite trees: *)
-coinductive infiniteTr where
-"\<lbrakk>tr' \<in> set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
-
-lemma infiniteTr_strong_coind[consumes 1, case_names sub]:
-assumes *: "phi tr" and
-**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> set (sub tr). phi tr' \<or> infiniteTr tr'"
-shows "infiniteTr tr"
-using assms by (elim infiniteTr.coinduct) blast
-
-lemma infiniteTr_coind[consumes 1, case_names sub, induct pred: infiniteTr]:
-assumes *: "phi tr" and
-**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> set (sub tr). phi tr'"
-shows "infiniteTr tr"
-using assms by (elim infiniteTr.coinduct) blast
-
-lemma infiniteTr_sub[simp]:
-"infiniteTr tr \<Longrightarrow> (\<exists> tr' \<in> set (sub tr). infiniteTr tr')"
-by (erule infiniteTr.cases) blast
-
-primcorec konigPath where
-  "shd (konigPath t) = lab t"
-| "stl (konigPath t) = konigPath (SOME tr. tr \<in> set (sub t) \<and> infiniteTr tr)"
-
-(* proper paths in trees: *)
-coinductive properPath where
-"\<lbrakk>shd as = lab tr; tr' \<in> set (sub tr); properPath (stl as) tr'\<rbrakk> \<Longrightarrow>
- properPath as tr"
-
-lemma properPath_strong_coind[consumes 1, case_names shd_lab sub]:
-assumes *: "phi as tr" and
-**: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and
-***: "\<And> as tr.
-         phi as tr \<Longrightarrow>
-         \<exists> tr' \<in> set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
-shows "properPath as tr"
-using assms by (elim properPath.coinduct) blast
-
-lemma properPath_coind[consumes 1, case_names shd_lab sub, induct pred: properPath]:
-assumes *: "phi as tr" and
-**: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and
-***: "\<And> as tr.
-         phi as tr \<Longrightarrow>
-         \<exists> tr' \<in> set (sub tr). phi (stl as) tr'"
-shows "properPath as tr"
-using properPath_strong_coind[of phi, OF * **] *** by blast
-
-lemma properPath_shd_lab:
-"properPath as tr \<Longrightarrow> shd as = lab tr"
-by (erule properPath.cases) blast
-
-lemma properPath_sub:
-"properPath as tr \<Longrightarrow>
- \<exists> tr' \<in> set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
-by (erule properPath.cases) blast
-
-(* prove the following by coinduction *)
-theorem Konig:
-  assumes "infiniteTr tr"
-  shows "properPath (konigPath tr) tr"
-proof-
-  {fix as
-   assume "infiniteTr tr \<and> as = konigPath tr" hence "properPath as tr"
-   proof (coinduction arbitrary: tr as rule: properPath_coind)
-     case (sub tr as)
-     let ?t = "SOME t'. t' \<in> set (sub tr) \<and> infiniteTr t'"
-     from sub have "\<exists>t' \<in> set (sub tr). infiniteTr t'" by simp
-     then have "\<exists>t'. t' \<in> set (sub tr) \<and> infiniteTr t'" by blast
-     then have "?t \<in> set (sub tr) \<and> infiniteTr ?t" by (rule someI_ex)
-     moreover have "stl (konigPath tr) = konigPath ?t" by simp
-     ultimately show ?case using sub by blast
-   qed simp
-  }
-  thus ?thesis using assms by blast
-qed
-
-(* some more stream theorems *)
-
-primcorec plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<oplus>" 66) where
-  "shd (plus xs ys) = shd xs + shd ys"
-| "stl (plus xs ys) = plus (stl xs) (stl ys)"
-
-definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<cdot>" 68) where
-  [simp]: "scalar n = smap (\<lambda>x. n * x)"
-
-primcorec ones :: "nat stream" where "ones = 1 ## ones"
-primcorec twos :: "nat stream" where "twos = 2 ## twos"
-definition ns :: "nat \<Rightarrow> nat stream" where [simp]: "ns n = scalar n ones"
-
-lemma "ones \<oplus> ones = twos"
-  by coinduction simp
-
-lemma "n \<cdot> twos = ns (2 * n)"
-  by coinduction simp
-
-lemma prod_scalar: "(n * m) \<cdot> xs = n \<cdot> m \<cdot> xs"
-  by (coinduction arbitrary: xs) auto
-
-lemma scalar_plus: "n \<cdot> (xs \<oplus> ys) = n \<cdot> xs \<oplus> n \<cdot> ys"
-  by (coinduction arbitrary: xs ys) (auto simp: add_mult_distrib2)
-
-lemma plus_comm: "xs \<oplus> ys = ys \<oplus> xs"
-  by (coinduction arbitrary: xs ys) auto
-
-lemma plus_assoc: "(xs \<oplus> ys) \<oplus> zs = xs \<oplus> ys \<oplus> zs"
-  by (coinduction arbitrary: xs ys zs) auto
-
-end
--- a/src/HOL/BNF_Examples/Lambda_Term.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Lambda_Term.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Lambda-terms.
-*)
-
-header {* Lambda-Terms *}
-
-theory Lambda_Term
-imports "~~/src/HOL/Library/FSet"
-begin
-
-section {* Datatype definition *}
-
-datatype_new 'a trm =
-  Var 'a |
-  App "'a trm" "'a trm" |
-  Lam 'a "'a trm" |
-  Lt "('a \<times> 'a trm) fset" "'a trm"
-
-
-subsection {* Example: The set of all variables varsOf and free variables fvarsOf of a term *}
-
-primrec varsOf :: "'a trm \<Rightarrow> 'a set" where
-  "varsOf (Var a) = {a}"
-| "varsOf (App f x) = varsOf f \<union> varsOf x"
-| "varsOf (Lam x b) = {x} \<union> varsOf b"
-| "varsOf (Lt F t) = varsOf t \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| fimage (map_prod id varsOf) F})"
-
-primrec fvarsOf :: "'a trm \<Rightarrow> 'a set" where
-  "fvarsOf (Var x) = {x}"
-| "fvarsOf (App t1 t2) = fvarsOf t1 \<union> fvarsOf t2"
-| "fvarsOf (Lam x t) = fvarsOf t - {x}"
-| "fvarsOf (Lt xts t) = fvarsOf t - {x | x X. (x,X) |\<in>| fimage (map_prod id varsOf) xts} \<union>
-    (\<Union> {X | x X. (x,X) |\<in>| fimage (map_prod id varsOf) xts})"
-
-lemma diff_Un_incl_triv: "\<lbrakk>A \<subseteq> D; C \<subseteq> E\<rbrakk> \<Longrightarrow> A - B \<union> C \<subseteq> D \<union> E" by blast
-
-lemma in_fimage_map_prod_fset_iff[simp]:
-  "(x, y) |\<in>| fimage (map_prod f g) xts \<longleftrightarrow> (\<exists> t1 t2. (t1, t2) |\<in>| xts \<and> x = f t1 \<and> y = g t2)"
-  by force
-
-lemma fvarsOf_varsOf: "fvarsOf t \<subseteq> varsOf t"
-proof induct
-  case (Lt xts t) thus ?case unfolding fvarsOf.simps varsOf.simps by (elim diff_Un_incl_triv) auto
-qed auto
-
-end
--- a/src/HOL/BNF_Examples/Misc_Codatatype.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,122 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Misc_Codatatype.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Miscellaneous codatatype definitions.
-*)
-
-header {* Miscellaneous Codatatype Definitions *}
-
-theory Misc_Codatatype
-imports "~~/src/HOL/Library/FSet"
-begin
-
-codatatype simple = X1 | X2 | X3 | X4
-
-codatatype simple' = X1' unit | X2' unit | X3' unit | X4' unit
-
-codatatype simple'' = X1'' nat int | X2''
-
-codatatype 'a stream = Stream (shd: 'a) (stl: "'a stream")
-
-codatatype 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
-
-codatatype ('b, 'c :: ord, 'd, 'e) some_passive =
-  SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
-
-codatatype lambda =
-  Var string |
-  App lambda lambda |
-  Abs string lambda |
-  Let "(string \<times> lambda) fset" lambda
-
-codatatype 'a par_lambda =
-  PVar 'a |
-  PApp "'a par_lambda" "'a par_lambda" |
-  PAbs 'a "'a par_lambda" |
-  PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
-
-(*
-  ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
-  ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
-*)
-
-codatatype 'a p = P "'a + 'a p"
-
-codatatype 'a J1 = J11 'a "'a J1" | J12 'a "'a J2"
-and 'a J2 = J21 | J22 "'a J1" "'a J2"
-
-codatatype 'a tree = TEmpty | TNode 'a "'a forest"
-and 'a forest = FNil | FCons "'a tree" "'a forest"
-
-codatatype 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
-and 'a branch = Branch 'a "'a tree'"
-
-codatatype 'a bin_rose_tree = BRTree 'a "'a bin_rose_tree mylist" "'a bin_rose_tree mylist"
-
-codatatype ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
-and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
-and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
-
-codatatype ('a, 'b, 'c) some_killing =
-  SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
-and ('a, 'b, 'c) in_here =
-  IH1 'b 'a | IH2 'c
-
-codatatype ('a, 'b, 'c) some_killing' =
-  SK' "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing' + ('a, 'b, 'c) in_here'"
-and ('a, 'b, 'c) in_here' =
-  IH1' 'b | IH2' 'c
-
-codatatype ('a, 'b, 'c) some_killing'' =
-  SK'' "'a \<Rightarrow> ('a, 'b, 'c) in_here''"
-and ('a, 'b, 'c) in_here'' =
-  IH1'' 'b 'a | IH2'' 'c
-
-codatatype ('b, 'c) less_killing = LK "'b \<Rightarrow> 'c"
-
-codatatype 'b poly_unit = U "'b \<Rightarrow> 'b poly_unit"
-codatatype 'b cps = CPS1 'b | CPS2 "'b \<Rightarrow> 'b cps"
-
-codatatype ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs =
-  FR "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow>
-      ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs"
-
-codatatype ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
-        'b18, 'b19, 'b20) fun_rhs' =
-  FR' "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> 'b10 \<Rightarrow> 'b11 \<Rightarrow> 'b12 \<Rightarrow> 'b13 \<Rightarrow> 'b14 \<Rightarrow>
-       'b15 \<Rightarrow> 'b16 \<Rightarrow> 'b17 \<Rightarrow> 'b18 \<Rightarrow> 'b19 \<Rightarrow> 'b20 \<Rightarrow>
-       ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
-        'b18, 'b19, 'b20) fun_rhs'"
-
-codatatype ('a, 'b, 'c) wit3_F1 = W1 'a "('a, 'b, 'c) wit3_F1" "('a, 'b, 'c) wit3_F2"
-and ('a, 'b, 'c) wit3_F2 = W2 'b "('a, 'b, 'c) wit3_F2"
-and ('a, 'b, 'c) wit3_F3 = W31 'a 'b "('a, 'b, 'c) wit3_F1" | W32 'c 'a 'b "('a, 'b, 'c) wit3_F1"
-
-codatatype ('c, 'e, 'g) coind_wit1 =
-       CW1 'c "('c, 'e, 'g) coind_wit1" "('c, 'e, 'g) ind_wit" "('c, 'e, 'g) coind_wit2"
-and ('c, 'e, 'g) coind_wit2 =
-       CW21 "('c, 'e, 'g) coind_wit2" 'e | CW22 'c 'g
-and ('c, 'e, 'g) ind_wit =
-       IW1 | IW2 'c
-
-codatatype ('b, 'a) bar = BAR "'a \<Rightarrow> 'b"
-codatatype ('a, 'b, 'c, 'd) foo = FOO "'d + 'b \<Rightarrow> 'c + 'a"
-
-codatatype 'a dead_foo = A
-codatatype ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
-
-(* SLOW, MEMORY-HUNGRY
-codatatype ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
-and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
-and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
-and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
-and ('a, 'c) D5 = A5 "('a, 'c) D6"
-and ('a, 'c) D6 = A6 "('a, 'c) D7"
-and ('a, 'c) D7 = A7 "('a, 'c) D8"
-and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
-*)
-
-end
--- a/src/HOL/BNF_Examples/Misc_Datatype.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,338 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Misc_Datatype.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Miscellaneous datatype definitions.
-*)
-
-header {* Miscellaneous Datatype Definitions *}
-
-theory Misc_Datatype
-imports "~~/src/HOL/Library/Countable" "~~/src/HOL/Library/FSet"
-begin
-
-datatype_new (discs_sels) simple = X1 | X2 | X3 | X4
-
-datatype_new (discs_sels) simple' = X1' unit | X2' unit | X3' unit | X4' unit
-
-datatype_new (discs_sels) simple'' = X1'' nat int | X2''
-
-datatype_new (discs_sels) 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
-
-datatype_new (discs_sels) ('b, 'c :: ord, 'd, 'e) some_passive =
-  SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
-
-datatype_new (discs_sels) hfset = HFset "hfset fset"
-
-datatype_new (discs_sels) lambda =
-  Var string |
-  App lambda lambda |
-  Abs string lambda |
-  Let "(string \<times> lambda) fset" lambda
-
-datatype_new (discs_sels) 'a par_lambda =
-  PVar 'a |
-  PApp "'a par_lambda" "'a par_lambda" |
-  PAbs 'a "'a par_lambda" |
-  PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
-
-(*
-  ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
-  ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
-*)
-
-datatype_new (discs_sels) 'a I1 = I11 'a "'a I1" | I12 'a "'a I2"
-and 'a I2 = I21 | I22 "'a I1" "'a I2"
-
-datatype_new (discs_sels) 'a tree = TEmpty | TNode 'a "'a forest"
-and 'a forest = FNil | FCons "'a tree" "'a forest"
-
-datatype_new (discs_sels) 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
-and 'a branch = Branch 'a "'a tree'"
-
-datatype_new (discs_sels) 'a bin_rose_tree = BRTree 'a "'a bin_rose_tree mylist" "'a bin_rose_tree mylist"
-
-datatype_new (discs_sels) ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
-and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
-and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
-
-datatype_new (discs_sels) 'a ftree = FTLeaf 'a | FTNode "'a \<Rightarrow> 'a ftree"
-
-datatype_new (discs_sels) ('a, 'b, 'c) some_killing =
-  SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
-and ('a, 'b, 'c) in_here =
-  IH1 'b 'a | IH2 'c
-
-datatype_new (discs_sels) 'b nofail1 = NF11 "'b nofail1" 'b | NF12 'b
-datatype_new (discs_sels) 'b nofail2 = NF2 "('b nofail2 \<times> 'b \<times> 'b nofail2 \<times> 'b) list"
-datatype_new (discs_sels) 'b nofail3 = NF3 'b "('b nofail3 \<times> 'b \<times> 'b nofail3 \<times> 'b) fset"
-datatype_new (discs_sels) 'b nofail4 = NF4 "('b nofail4 \<times> ('b nofail4 \<times> 'b \<times> 'b nofail4 \<times> 'b) fset) list"
-
-(*
-datatype_new (discs_sels) 'b fail = F "'b fail" 'b "'b fail" "'b list"
-datatype_new (discs_sels) 'b fail = F "'b fail" 'b "'b fail" 'b
-datatype_new (discs_sels) 'b fail = F1 "'b fail" 'b | F2 "'b fail"
-datatype_new (discs_sels) 'b fail = F "'b fail" 'b
-*)
-
-datatype_new (discs_sels) l1 = L1 "l2 list"
-and l2 = L21 "l1 fset" | L22 l2
-
-datatype_new (discs_sels) kk1 = KK1 kk2
-and kk2 = KK2 kk3
-and kk3 = KK3 "kk1 list"
-
-datatype_new (discs_sels) t1 = T11 t3 | T12 t2
-and t2 = T2 t1
-and t3 = T3
-
-datatype_new (discs_sels) t1' = T11' t2' | T12' t3'
-and t2' = T2' t1'
-and t3' = T3'
-
-(*
-datatype_new (discs_sels) fail1 = F1 fail2
-and fail2 = F2 fail3
-and fail3 = F3 fail1
-
-datatype_new (discs_sels) fail1 = F1 "fail2 list" fail2
-and fail2 = F2 "fail2 fset" fail3
-and fail3 = F3 fail1
-
-datatype_new (discs_sels) fail1 = F1 "fail2 list" fail2
-and fail2 = F2 "fail1 fset" fail1
-*)
-
-(* SLOW
-datatype_new (discs_sels) ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
-and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
-and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
-and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
-and ('a, 'c) D5 = A5 "('a, 'c) D6"
-and ('a, 'c) D6 = A6 "('a, 'c) D7"
-and ('a, 'c) D7 = A7 "('a, 'c) D8"
-and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
-*)
-
-(* fail:
-datatype_new (discs_sels) tt1 = TT11 tt2 tt3 | TT12 tt2 tt4
-and tt2 = TT2
-and tt3 = TT3 tt4
-and tt4 = TT4 tt1
-*)
-
-datatype_new (discs_sels) k1 = K11 k2 k3 | K12 k2 k4
-and k2 = K2
-and k3 = K3 k4
-and k4 = K4
-
-datatype_new (discs_sels) tt1 = TT11 tt3 tt2 | TT12 tt2 tt4
-and tt2 = TT2
-and tt3 = TT3 tt1
-and tt4 = TT4
-
-(* SLOW
-datatype_new (discs_sels) s1 = S11 s2 s3 s4 | S12 s3 | S13 s2 s6 | S14 s4 s2 | S15 s2 s2
-and s2 = S21 s7 s5 | S22 s5 s4 s6
-and s3 = S31 s1 s7 s2 | S32 s3 s3 | S33 s4 s5
-and s4 = S4 s5
-and s5 = S5
-and s6 = S61 s6 | S62 s1 s2 | S63 s6
-and s7 = S71 s8 | S72 s5
-and s8 = S8 nat
-*)
-
-datatype_new (discs_sels) 'a deadbar = DeadBar "'a \<Rightarrow> 'a"
-datatype_new (discs_sels) 'a deadbar_option = DeadBarOption "'a option \<Rightarrow> 'a option"
-datatype_new (discs_sels) ('a, 'b) bar = Bar "'b \<Rightarrow> 'a"
-datatype_new (discs_sels) ('a, 'b, 'c, 'd) foo = Foo "'d + 'b \<Rightarrow> 'c + 'a"
-datatype_new (discs_sels) 'a deadfoo = DeadFoo "'a \<Rightarrow> 'a + 'a"
-
-datatype_new (discs_sels) 'a dead_foo = A
-datatype_new (discs_sels) ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
-
-datatype_new (discs_sels) d1 = D
-datatype_new (discs_sels) d1' = is_D: D
-
-datatype_new (discs_sels) d2 = D nat
-datatype_new (discs_sels) d2' = is_D: D nat
-
-datatype_new (discs_sels) d3 = D | E
-datatype_new (discs_sels) d3' = D | is_E: E
-datatype_new (discs_sels) d3'' = is_D: D | E
-datatype_new (discs_sels) d3''' = is_D: D | is_E: E
-
-datatype_new (discs_sels) d4 = D nat | E
-datatype_new (discs_sels) d4' = D nat | is_E: E
-datatype_new (discs_sels) d4'' = is_D: D nat | E
-datatype_new (discs_sels) d4''' = is_D: D nat | is_E: E
-
-datatype_new (discs_sels) d5 = D nat | E int
-datatype_new (discs_sels) d5' = D nat | is_E: E int
-datatype_new (discs_sels) d5'' = is_D: D nat | E int
-datatype_new (discs_sels) d5''' = is_D: D nat | is_E: E int
-
-instance simple :: countable
-  by countable_datatype
-
-instance simple'' :: countable
-  by countable_datatype
-
-instance mylist :: (countable) countable
-  by countable_datatype
-
-instance some_passive :: (countable, "{countable,ord}", countable, countable) countable
-  by countable_datatype
-
-(* TODO: Enable once "fset" is registered as countable:
-
-instance hfset :: countable
-  by countable_datatype
-
-instance lambda :: countable
-  by countable_datatype
-
-instance par_lambda :: (countable) countable
-  by countable_datatype
-*)
-
-instance I1 and I2 :: (countable) countable
-  by countable_datatype
-
-instance tree and forest :: (countable) countable
-  by countable_datatype
-
-instance tree' and branch :: (countable) countable
-  by countable_datatype
-
-instance bin_rose_tree :: (countable) countable
-  by countable_datatype
-
-instance exp and trm and factor :: (countable, countable) countable
-  by countable_datatype
-
-instance nofail1 :: (countable) countable
-  by countable_datatype
-
-instance nofail2 :: (countable) countable
-  by countable_datatype
-
-(* TODO: Enable once "fset" is registered as countable:
-
-instance nofail3 :: (countable) countable
-  by countable_datatype
-
-instance nofail4 :: (countable) countable
-  by countable_datatype
-
-instance l1 and l2 :: countable
-  by countable_datatype
-*)
-
-instance kk1 and kk2 :: countable
-  by countable_datatype
-
-instance t1 and t2 and t3 :: countable
-  by countable_datatype
-
-instance t1' and t2' and t3' :: countable
-  by countable_datatype
-
-instance k1 and k2 and k3 and k4 :: countable
-  by countable_datatype
-
-instance tt1 and tt2 and tt3 and tt4 :: countable
-  by countable_datatype
-
-instance d1 :: countable
-  by countable_datatype
-
-instance d1' :: countable
-  by countable_datatype
-
-instance d2 :: countable
-  by countable_datatype
-
-instance d2' :: countable
-  by countable_datatype
-
-instance d3 :: countable
-  by countable_datatype
-
-instance d3' :: countable
-  by countable_datatype
-
-instance d3'' :: countable
-  by countable_datatype
-
-instance d3''' :: countable
-  by countable_datatype
-
-instance d4 :: countable
-  by countable_datatype
-
-instance d4' :: countable
-  by countable_datatype
-
-instance d4'' :: countable
-  by countable_datatype
-
-instance d4''' :: countable
-  by countable_datatype
-
-instance d5 :: countable
-  by countable_datatype
-
-instance d5' :: countable
-  by countable_datatype
-
-instance d5'' :: countable
-  by countable_datatype
-
-instance d5''' :: countable
-  by countable_datatype
-
-datatype_compat simple
-datatype_compat simple'
-datatype_compat simple''
-datatype_compat mylist
-datatype_compat some_passive
-datatype_compat I1 I2
-datatype_compat tree forest
-datatype_compat tree' branch
-datatype_compat bin_rose_tree
-datatype_compat exp trm factor
-datatype_compat ftree
-datatype_compat nofail1
-datatype_compat kk1 kk2 kk3
-datatype_compat t1 t2 t3
-datatype_compat t1' t2' t3'
-datatype_compat k1 k2 k3 k4
-datatype_compat tt1 tt2 tt3 tt4
-datatype_compat deadbar
-datatype_compat deadbar_option
-datatype_compat bar
-datatype_compat foo
-datatype_compat deadfoo
-datatype_compat dead_foo
-datatype_compat use_dead_foo
-datatype_compat d1
-datatype_compat d1'
-datatype_compat d2
-datatype_compat d2'
-datatype_compat d3
-datatype_compat d3'
-datatype_compat d3''
-datatype_compat d3'''
-datatype_compat d4
-datatype_compat d4'
-datatype_compat d4''
-datatype_compat d4'''
-datatype_compat d5
-datatype_compat d5'
-datatype_compat d5''
-datatype_compat d5'''
-
-end
--- a/src/HOL/BNF_Examples/Misc_Primcorec.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,138 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Misc_Primcorec.thy
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Miscellaneous primitive corecursive function definitions.
-*)
-
-header {* Miscellaneous Primitive Corecursive Function Definitions *}
-
-theory Misc_Primcorec
-imports Misc_Codatatype
-begin
-
-primcorec simple_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple" where
-  "simple_of_bools b b' = (if b then if b' then X1 else X2 else if b' then X3 else X4)"
-
-primcorec simple'_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple'" where
-  "simple'_of_bools b b' =
-     (if b then if b' then X1' () else X2' () else if b' then X3' () else X4' ())"
-
-primcorec inc_simple'' :: "nat \<Rightarrow> simple'' \<Rightarrow> simple''" where
-  "inc_simple'' k s = (case s of X1'' n i \<Rightarrow> X1'' (n + k) (i + int k) | X2'' \<Rightarrow> X2'')"
-
-primcorec sinterleave :: "'a stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
-  "sinterleave s s' = Stream (shd s) (sinterleave s' (stl s))"
-
-primcorec myapp :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" where
-  "myapp xs ys =
-     (if xs = MyNil then ys
-      else if ys = MyNil then xs
-      else MyCons (myhd xs) (myapp (mytl xs) ys))"
-
-primcorec shuffle_sp :: "('a \<Colon> ord, 'b \<Colon> ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
-  "shuffle_sp sp =
-     (case sp of
-       SP1 sp' \<Rightarrow> SP1 (shuffle_sp sp')
-     | SP2 a \<Rightarrow> SP3 a
-     | SP3 b \<Rightarrow> SP4 b
-     | SP4 c \<Rightarrow> SP5 c
-     | SP5 d \<Rightarrow> SP2 d)"
-
-primcorec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
-  "rename_lam f l =
-     (case l of
-       Var s \<Rightarrow> Var (f s)
-     | App l l' \<Rightarrow> App (rename_lam f l) (rename_lam f l')
-     | Abs s l \<Rightarrow> Abs (f s) (rename_lam f l)
-     | Let SL l \<Rightarrow> Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l))"
-
-primcorec
-  j1_sum :: "('a\<Colon>{zero,one,plus}) \<Rightarrow> 'a J1" and
-  j2_sum :: "'a \<Rightarrow> 'a J2"
-where
-  "n = 0 \<Longrightarrow> is_J11 (j1_sum n)" |
-  "un_J111 (j1_sum _) = 0" |
-  "un_J112 (j1_sum _) = j1_sum 0" |
-  "un_J121 (j1_sum n) = n + 1" |
-  "un_J122 (j1_sum n) = j2_sum (n + 1)" |
-  "n = 0 \<Longrightarrow> j2_sum n = J21" |
-  "un_J221 (j2_sum n) = j1_sum (n + 1)" |
-  "un_J222 (j2_sum n) = j2_sum (n + 1)"
-
-primcorec forest_of_mylist :: "'a tree mylist \<Rightarrow> 'a forest" where
-  "forest_of_mylist ts =
-     (case ts of
-       MyNil \<Rightarrow> FNil
-     | MyCons t ts \<Rightarrow> FCons t (forest_of_mylist ts))"
-
-primcorec mylist_of_forest :: "'a forest \<Rightarrow> 'a tree mylist" where
-  "mylist_of_forest f =
-     (case f of
-       FNil \<Rightarrow> MyNil
-     | FCons t ts \<Rightarrow> MyCons t (mylist_of_forest ts))"
-
-primcorec semi_stream :: "'a stream \<Rightarrow> 'a stream" where
-  "semi_stream s = Stream (shd s) (semi_stream (stl (stl s)))"
-
-primcorec
-  tree'_of_stream :: "'a stream \<Rightarrow> 'a tree'" and
-  branch_of_stream :: "'a stream \<Rightarrow> 'a branch"
-where
-  "tree'_of_stream s =
-     TNode' (branch_of_stream (semi_stream s)) (branch_of_stream (semi_stream (stl s)))" |
-  "branch_of_stream s = (case s of Stream h t \<Rightarrow> Branch h (tree'_of_stream t))"
-
-primcorec
-  id_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
-  id_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
-  id_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
-where
-  "id_tree t = (case t of BRTree a ts ts' \<Rightarrow> BRTree a (id_trees1 ts) (id_trees2 ts'))" |
-  "id_trees1 ts = (case ts of
-       MyNil \<Rightarrow> MyNil
-     | MyCons t ts \<Rightarrow> MyCons (id_tree t) (id_trees1 ts))" |
-  "id_trees2 ts = (case ts of
-       MyNil \<Rightarrow> MyNil
-     | MyCons t ts \<Rightarrow> MyCons (id_tree t) (id_trees2 ts))"
-
-primcorec
-  trunc_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
-  trunc_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
-  trunc_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
-where
-  "trunc_tree t = (case t of BRTree a ts ts' \<Rightarrow> BRTree a (trunc_trees1 ts) (trunc_trees2 ts'))" |
-  "trunc_trees1 ts = (case ts of
-       MyNil \<Rightarrow> MyNil
-     | MyCons t _ \<Rightarrow> MyCons (trunc_tree t) MyNil)" |
-  "trunc_trees2 ts = (case ts of
-       MyNil \<Rightarrow> MyNil
-     | MyCons t ts \<Rightarrow> MyCons (trunc_tree t) MyNil)"
-
-primcorec
-  freeze_exp :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) exp \<Rightarrow> ('a, 'b) exp" and
-  freeze_trm :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) trm \<Rightarrow> ('a, 'b) trm" and
-  freeze_factor :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) factor \<Rightarrow> ('a, 'b) factor"
-where
-  "freeze_exp g e =
-     (case e of
-       Term t \<Rightarrow> Term (freeze_trm g t)
-     | Sum t e \<Rightarrow> Sum (freeze_trm g t) (freeze_exp g e))" |
-  "freeze_trm g t =
-     (case t of
-       Factor f \<Rightarrow> Factor (freeze_factor g f)
-     | Prod f t \<Rightarrow> Prod (freeze_factor g f) (freeze_trm g t))" |
-  "freeze_factor g f =
-     (case f of
-       C a \<Rightarrow> C a
-     | V b \<Rightarrow> C (g b)
-     | Paren e \<Rightarrow> Paren (freeze_exp g e))"
-
-primcorec poly_unity :: "'a poly_unit" where
-  "poly_unity = U (\<lambda>_. poly_unity)"
-
-primcorec build_cps :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> bool stream) \<Rightarrow> 'a \<Rightarrow> bool stream \<Rightarrow> 'a cps" where
-  "shd b \<Longrightarrow> build_cps f g a b = CPS1 a" |
-  "_ \<Longrightarrow> build_cps f g a b = CPS2 (\<lambda>a. build_cps f g (f a) (g a))"
-
-end
--- a/src/HOL/BNF_Examples/Misc_Primrec.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,136 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Misc_Primrec.thy
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Miscellaneous primitive recursive function definitions.
-*)
-
-header {* Miscellaneous Primitive Recursive Function Definitions *}
-
-theory Misc_Primrec
-imports Misc_Datatype
-begin
-
-primrec nat_of_simple :: "simple \<Rightarrow> nat" where
-  "nat_of_simple X1 = 1" |
-  "nat_of_simple X2 = 2" |
-  "nat_of_simple X3 = 3" |
-  "nat_of_simple X4 = 4"
-
-primrec simple_of_simple' :: "simple' \<Rightarrow> simple" where
-  "simple_of_simple' (X1' _) = X1" |
-  "simple_of_simple' (X2' _) = X2" |
-  "simple_of_simple' (X3' _) = X3" |
-  "simple_of_simple' (X4' _) = X4"
-
-primrec inc_simple'' :: "nat \<Rightarrow> simple'' \<Rightarrow> simple''" where
-  "inc_simple'' k (X1'' n i) = X1'' (n + k) (i + int k)" |
-  "inc_simple'' _ X2'' = X2''"
-
-primrec myapp :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" where
-  "myapp MyNil ys = ys" |
-  "myapp (MyCons x xs) ys = MyCons x (myapp xs ys)"
-
-primrec myrev :: "'a mylist \<Rightarrow> 'a mylist" where
-  "myrev MyNil = MyNil" |
-  "myrev (MyCons x xs) = myapp (myrev xs) (MyCons x MyNil)"
-
-primrec shuffle_sp :: "('a \<Colon> ord, 'b \<Colon> ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
-  "shuffle_sp (SP1 sp) = SP1 (shuffle_sp sp)" |
-  "shuffle_sp (SP2 a) = SP3 a" |
-  "shuffle_sp (SP3 b) = SP4 b" |
-  "shuffle_sp (SP4 c) = SP5 c" |
-  "shuffle_sp (SP5 d) = SP2 d"
-
-primrec
-  hf_size :: "hfset \<Rightarrow> nat"
-where
-  "hf_size (HFset X) = 1 + setsum id (fset (fimage hf_size X))"
-
-primrec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
-  "rename_lam f (Var s) = Var (f s)" |
-  "rename_lam f (App l l') = App (rename_lam f l) (rename_lam f l')" |
-  "rename_lam f (Abs s l) = Abs (f s) (rename_lam f l)" |
-  "rename_lam f (Let SL l) = Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l)"
-
-primrec
-  sum_i1 :: "('a\<Colon>{zero,plus}) I1 \<Rightarrow> 'a" and
-  sum_i2 :: "'a I2 \<Rightarrow> 'a"
-where
-  "sum_i1 (I11 n i) = n + sum_i1 i" |
-  "sum_i1 (I12 n i) = n + sum_i2 i" |
-  "sum_i2 I21 = 0" |
-  "sum_i2 (I22 i j) = sum_i1 i + sum_i2 j"
-
-primrec forest_of_mylist :: "'a tree mylist \<Rightarrow> 'a forest" where
-  "forest_of_mylist MyNil = FNil" |
-  "forest_of_mylist (MyCons t ts) = FCons t (forest_of_mylist ts)"
-
-primrec mylist_of_forest :: "'a forest \<Rightarrow> 'a tree mylist" where
-  "mylist_of_forest FNil = MyNil" |
-  "mylist_of_forest (FCons t ts) = MyCons t (mylist_of_forest ts)"
-
-definition frev :: "'a forest \<Rightarrow> 'a forest" where
-  "frev = forest_of_mylist \<circ> myrev \<circ> mylist_of_forest"
-
-primrec
-  mirror_tree :: "'a tree \<Rightarrow> 'a tree" and
-  mirror_forest :: "'a forest \<Rightarrow> 'a forest"
-where
-  "mirror_tree TEmpty = TEmpty" |
-  "mirror_tree (TNode x ts) = TNode x (mirror_forest ts)" |
-  "mirror_forest FNil = FNil" |
-  "mirror_forest (FCons t ts) = frev (FCons (mirror_tree t) (mirror_forest ts))"
-
-primrec
-  mylist_of_tree' :: "'a tree' \<Rightarrow> 'a mylist" and
-  mylist_of_branch :: "'a branch \<Rightarrow> 'a mylist"
-where
-  "mylist_of_tree' TEmpty' = MyNil" |
-  "mylist_of_tree' (TNode' b b') = myapp (mylist_of_branch b) (mylist_of_branch b')" |
-  "mylist_of_branch (Branch x t) = MyCons x (mylist_of_tree' t)"
-
-primrec
-  id_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
-  id_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
-  id_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
-where
-  "id_tree (BRTree a ts ts') = BRTree a (id_trees1 ts) (id_trees2 ts')" |
-  "id_trees1 MyNil = MyNil" |
-  "id_trees1 (MyCons t ts) = MyCons (id_tree t) (id_trees1 ts)" |
-  "id_trees2 MyNil = MyNil" |
-  "id_trees2 (MyCons t ts) = MyCons (id_tree t) (id_trees2 ts)"
-
-primrec
-  trunc_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
-  trunc_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
-  trunc_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
-where
-  "trunc_tree (BRTree a ts ts') = BRTree a (trunc_trees1 ts) (trunc_trees2 ts')" |
-  "trunc_trees1 MyNil = MyNil" |
-  "trunc_trees1 (MyCons t ts) = MyCons (id_tree t) MyNil" |
-  "trunc_trees2 MyNil = MyNil" |
-  "trunc_trees2 (MyCons t ts) = MyCons (id_tree t) MyNil"
-
-primrec
-  is_ground_exp :: "('a, 'b) exp \<Rightarrow> bool" and
-  is_ground_trm :: "('a, 'b) trm \<Rightarrow> bool" and
-  is_ground_factor :: "('a, 'b) factor \<Rightarrow> bool"
-where
-  "is_ground_exp (Term t) \<longleftrightarrow> is_ground_trm t" |
-  "is_ground_exp (Sum t e) \<longleftrightarrow> is_ground_trm t \<and> is_ground_exp e" |
-  "is_ground_trm (Factor f) \<longleftrightarrow> is_ground_factor f" |
-  "is_ground_trm (Prod f t) \<longleftrightarrow> is_ground_factor f \<and> is_ground_trm t" |
-  "is_ground_factor (C _) \<longleftrightarrow> True" |
-  "is_ground_factor (V _) \<longleftrightarrow> False" |
-  "is_ground_factor (Paren e) \<longleftrightarrow> is_ground_exp e"
-
-primrec map_ftreeA :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
-  "map_ftreeA f (FTLeaf x) = FTLeaf (f x)" |
-  "map_ftreeA f (FTNode g) = FTNode (map_ftreeA f \<circ> g)"
-
-primrec map_ftreeB :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a ftree \<Rightarrow> 'b ftree" where
-  "map_ftreeB f (FTLeaf x) = FTLeaf (f x)" |
-  "map_ftreeB f (FTNode g) = FTNode (map_ftreeB f \<circ> g \<circ> the_inv f)"
-
-end
--- a/src/HOL/BNF_Examples/Process.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,278 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Process.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Processes.
-*)
-
-header {* Processes *}
-
-theory Process
-imports Stream 
-begin
-
-codatatype 'a process =
-  isAction: Action (prefOf: 'a) (contOf: "'a process") |
-  isChoice: Choice (ch1Of: "'a process") (ch2Of: "'a process")
-
-(* Read: prefix of, continuation of, choice 1 of, choice 2 of *)
-
-section {* Customization *}
-
-subsection {* Basic properties *}
-
-declare
-  rel_pre_process_def[simp]
-  rel_sum_def[simp]
-  rel_prod_def[simp]
-
-(* Constructors versus discriminators *)
-theorem isAction_isChoice:
-"isAction p \<or> isChoice p"
-by (rule process.exhaust_disc) auto
-
-theorem not_isAction_isChoice: "\<not> (isAction p \<and> isChoice p)"
-by (cases rule: process.exhaust[of p]) auto
-
-
-subsection{* Coinduction *}
-
-theorem process_coind[elim, consumes 1, case_names iss Action Choice, induct pred: "HOL.eq"]:
-  assumes phi: "\<phi> p p'" and
-  iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
-  Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> \<phi> p p'" and
-  Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> \<phi> p p' \<and> \<phi> q q'"
-  shows "p = p'"
-  using assms
-  by (coinduct rule: process.coinduct) (metis process.collapse(1,2) process.disc(3))
-
-(* Stronger coinduction, up to equality: *)
-theorem process_strong_coind[elim, consumes 1, case_names iss Action Choice]:
-  assumes phi: "\<phi> p p'" and
-  iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
-  Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> (\<phi> p p' \<or> p = p')" and
-  Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> (\<phi> p p' \<or> p = p') \<and> (\<phi> q q' \<or> q = q')"
-  shows "p = p'"
-  using assms
-  by (coinduct rule: process.coinduct_strong) (metis process.collapse(1,2) process.disc(3))
-
-
-subsection {* Coiteration (unfold) *}
-
-
-section{* Coinductive definition of the notion of trace *}
-coinductive trace where
-"trace p as \<Longrightarrow> trace (Action a p) (a ## as)"
-|
-"trace p as \<or> trace q as \<Longrightarrow> trace (Choice p q) as"
-
-
-section{* Examples of corecursive definitions: *}
-
-subsection{* Single-guard fixpoint definition *}
-
-primcorec BX where
-  "isAction BX"
-| "prefOf BX = ''a''"
-| "contOf BX = BX"
-
-
-subsection{* Multi-guard fixpoint definitions, simulated with auxiliary arguments *}
-
-datatype_new x_y_ax = x | y | ax
-
-primcorec F :: "x_y_ax \<Rightarrow> char list process" where
-  "xyax = x \<Longrightarrow> isChoice (F xyax)"
-| "ch1Of (F xyax) = F ax"
-| "ch2Of (F xyax) = F y"
-| "prefOf (F xyax) = (if xyax = y then ''b'' else ''a'')"
-| "contOf (F xyax) = F x"
-
-definition "X = F x"  definition "Y = F y"  definition "AX = F ax"
-
-lemma X_Y_AX: "X = Choice AX Y"  "Y = Action ''b'' X"  "AX = Action ''a'' X"
-unfolding X_def Y_def AX_def by (subst F.code, simp)+
-
-(* end product: *)
-lemma X_AX:
-"X = Choice AX (Action ''b'' X)"
-"AX = Action ''a'' X"
-using X_Y_AX by simp_all
-
-
-
-section{* Case study: Multi-guard fixpoint definitions, without auxiliary arguments *}
-
-hide_const x y ax X Y AX
-
-(* Process terms *)
-datatype_new ('a,'pvar) process_term =
- VAR 'pvar |
- PROC "'a process" |
- ACT 'a "('a,'pvar) process_term" | CH "('a,'pvar) process_term" "('a,'pvar) process_term"
-
-(* below, sys represents a system of equations *)
-fun isACT where
-"isACT sys (VAR X) =
- (case sys X of ACT a T \<Rightarrow> True |PROC p \<Rightarrow> isAction p |_ \<Rightarrow> False)"
-|
-"isACT sys (PROC p) = isAction p"
-|
-"isACT sys (ACT a T) = True"
-|
-"isACT sys (CH T1 T2) = False"
-
-fun PREF where
-"PREF sys (VAR X) =
- (case sys X of ACT a T \<Rightarrow> a | PROC p \<Rightarrow> prefOf p)"
-|
-"PREF sys (PROC p) = prefOf p"
-|
-"PREF sys (ACT a T) = a"
-
-fun CONT where
-"CONT sys (VAR X) =
- (case sys X of ACT a T \<Rightarrow> T | PROC p \<Rightarrow> PROC (contOf p))"
-|
-"CONT sys (PROC p) = PROC (contOf p)"
-|
-"CONT sys (ACT a T) = T"
-
-fun CH1 where
-"CH1 sys (VAR X) =
- (case sys X of CH T1 T2 \<Rightarrow> T1 |PROC p \<Rightarrow> PROC (ch1Of p))"
-|
-"CH1 sys (PROC p) = PROC (ch1Of p)"
-|
-"CH1 sys (CH T1 T2) = T1"
-
-fun CH2 where
-"CH2 sys (VAR X) =
- (case sys X of CH T1 T2 \<Rightarrow> T2 |PROC p \<Rightarrow> PROC (ch2Of p))"
-|
-"CH2 sys (PROC p) = PROC (ch2Of p)"
-|
-"CH2 sys (CH T1 T2) = T2"
-
-definition "guarded sys \<equiv> \<forall> X Y. sys X \<noteq> VAR Y"
-
-primcorec solution where
-  "isACT sys T \<Longrightarrow> solution sys T = Action (PREF sys T) (solution sys (CONT sys T))"
-| "_ \<Longrightarrow> solution sys T = Choice (solution sys (CH1 sys T)) (solution sys (CH2 sys T))"
-
-lemma isACT_VAR:
-assumes g: "guarded sys"
-shows "isACT sys (VAR X) \<longleftrightarrow> isACT sys (sys X)"
-using g unfolding guarded_def by (cases "sys X") auto
-
-lemma solution_VAR:
-assumes g: "guarded sys"
-shows "solution sys (VAR X) = solution sys (sys X)"
-proof(cases "isACT sys (VAR X)")
-  case True
-  hence T: "isACT sys (sys X)" unfolding isACT_VAR[OF g] .
-  show ?thesis
-  unfolding solution.ctr(1)[OF T] using solution.ctr(1)[of sys "VAR X"] True g
-  unfolding guarded_def by (cases "sys X", auto)
-next
-  case False note FFalse = False
-  hence TT: "\<not> isACT sys (sys X)" unfolding isACT_VAR[OF g] .
-  show ?thesis
-  unfolding solution.ctr(2)[OF TT] using solution.ctr(2)[of sys "VAR X"] FFalse g
-  unfolding guarded_def by (cases "sys X", auto)
-qed
-
-lemma solution_PROC[simp]:
-"solution sys (PROC p) = p"
-proof-
-  {fix q assume "q = solution sys (PROC p)"
-   hence "p = q"
-   proof (coinduct rule: process_coind)
-     case (iss p p')
-     from isAction_isChoice[of p] show ?case
-     proof
-       assume p: "isAction p"
-       hence 0: "isACT sys (PROC p)" by simp
-       thus ?thesis using iss not_isAction_isChoice by auto
-     next
-       assume "isChoice p"
-       hence 0: "\<not> isACT sys (PROC p)"
-       using not_isAction_isChoice by auto
-       thus ?thesis using iss isAction_isChoice by auto
-     qed
-   next
-     case (Action a a' p p')
-     hence 0: "isACT sys (PROC (Action a p))" by simp
-     show ?case using Action unfolding solution.ctr(1)[OF 0] by simp
-   next
-     case (Choice p q p' q')
-     hence 0: "\<not> isACT sys (PROC (Choice p q))" using not_isAction_isChoice by auto
-     show ?case using Choice unfolding solution.ctr(2)[OF 0] by simp
-   qed
-  }
-  thus ?thesis by metis
-qed
-
-lemma solution_ACT[simp]:
-"solution sys (ACT a T) = Action a (solution sys T)"
-by (metis CONT.simps(3) PREF.simps(3) isACT.simps(3) solution.ctr(1))
-
-lemma solution_CH[simp]:
-"solution sys (CH T1 T2) = Choice (solution sys T1) (solution sys T2)"
-by (metis CH1.simps(3) CH2.simps(3) isACT.simps(4) solution.ctr(2))
-
-
-(* Example: *)
-
-fun sys where
-"sys 0 = CH (VAR (Suc 0)) (ACT ''b'' (VAR 0))"
-|
-"sys (Suc 0) = ACT ''a'' (VAR 0)"
-| (* dummy guarded term for variables outside the system: *)
-"sys X = ACT ''a'' (VAR 0)"
-
-lemma guarded_sys:
-"guarded sys"
-unfolding guarded_def proof (intro allI)
-  fix X Y show "sys X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
-qed
-
-(* the actual processes: *)
-definition "x \<equiv> solution sys (VAR 0)"
-definition "ax \<equiv> solution sys (VAR (Suc 0))"
-
-(* end product: *)
-lemma x_ax:
-"x = Choice ax (Action ''b'' x)"
-"ax = Action ''a'' x"
-unfolding x_def ax_def by (subst solution_VAR[OF guarded_sys], simp)+
-
-
-(* Thanks to the inclusion of processes as process terms, one can
-also consider parametrized systems of equations---here, x is a (semantic)
-process parameter: *)
-
-fun sys' where
-"sys' 0 = CH (PROC x) (ACT ''b'' (VAR 0))"
-|
-"sys' (Suc 0) = CH (ACT ''a'' (VAR 0)) (PROC x)"
-| (* dummy guarded term : *)
-"sys' X = ACT ''a'' (VAR 0)"
-
-lemma guarded_sys':
-"guarded sys'"
-unfolding guarded_def proof (intro allI)
-  fix X Y show "sys' X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
-qed
-
-(* the actual processes: *)
-definition "y \<equiv> solution sys' (VAR 0)"
-definition "ay \<equiv> solution sys' (VAR (Suc 0))"
-
-(* end product: *)
-lemma y_ay:
-"y = Choice x (Action ''b'' y)"
-"ay = Choice (Action ''a'' y) x"
-unfolding y_def ay_def by (subst solution_VAR[OF guarded_sys'], simp)+
-
-end
--- a/src/HOL/BNF_Examples/SML.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,91 +0,0 @@
-(*  Title:      HOL/Datatype_Benchmark/SML.thy
-
-Example from Myra: part of the syntax of SML.
-*)
-
-theory SML imports Main begin
-
-datatype
-  string = EMPTY_STRING | CONS_STRING nat string
-
-datatype
-   strid = STRID string
-
-datatype
-   var = VAR string
-
-datatype
-   con = CON string
-
-datatype
-   scon = SCINT nat | SCSTR string
-
-datatype
-   excon = EXCON string
-
-datatype
-   label = LABEL string
-
-datatype
-  'a nonemptylist = Head_and_tail 'a "'a list"
-
-datatype
-  'a long = BASE 'a | QUALIFIED strid "'a long"
-
-datatype
-   atpat_e = WILDCARDatpat_e
-           | SCONatpat_e scon
-           | VARatpat_e var
-           | CONatpat_e "con long"
-           | EXCONatpat_e "excon long"
-           | RECORDatpat_e "patrow_e option"
-           | PARatpat_e pat_e
-and
-   patrow_e = DOTDOTDOT_e
-            | PATROW_e label pat_e "patrow_e option"
-and
-   pat_e = ATPATpat_e atpat_e
-         | CONpat_e "con long" atpat_e
-         | EXCONpat_e "excon long" atpat_e
-         | LAYEREDpat_e var pat_e
-and
-   conbind_e = CONBIND_e con "conbind_e option"
-and
-   datbind_e = DATBIND_e conbind_e "datbind_e option"
-and
-   exbind_e = EXBIND1_e excon "exbind_e option"
-            | EXBIND2_e excon "excon long" "exbind_e option"
-and
-   atexp_e = SCONatexp_e scon
-           | VARatexp_e "var long"
-           | CONatexp_e "con long"
-           | EXCONatexp_e "excon long"
-           | RECORDatexp_e "exprow_e option"
-           | LETatexp_e dec_e exp_e
-           | PARatexp_e exp_e
-and
-   exprow_e = EXPROW_e label exp_e "exprow_e option"
-and
-   exp_e = ATEXPexp_e atexp_e
-         | APPexp_e exp_e atexp_e
-         | HANDLEexp_e exp_e match_e
-         | RAISEexp_e exp_e
-         | FNexp_e match_e
-and
-   match_e = MATCH_e mrule_e "match_e option"
-and
-   mrule_e = MRULE_e pat_e exp_e
-and
-   dec_e = VALdec_e valbind_e
-         | DATATYPEdec_e datbind_e
-         | ABSTYPEdec_e datbind_e dec_e
-         | EXCEPTdec_e exbind_e
-         | LOCALdec_e dec_e dec_e
-         | OPENdec_e "strid long nonemptylist"
-         | EMPTYdec_e
-         | SEQdec_e dec_e dec_e
-and
-   valbind_e = PLAINvalbind_e pat_e exp_e "valbind_e option"
-             | RECvalbind_e valbind_e
-
-end
--- a/src/HOL/BNF_Examples/Stream.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,582 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Stream.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012, 2013
-
-Infinite streams.
-*)
-
-header {* Infinite Streams *}
-
-theory Stream
-imports "~~/src/HOL/Library/Nat_Bijection"
-begin
-
-codatatype (sset: 'a) stream =
-  SCons (shd: 'a) (stl: "'a stream") (infixr "##" 65)
-for
-  map: smap
-  rel: stream_all2
-
-(*for code generation only*)
-definition smember :: "'a \<Rightarrow> 'a stream \<Rightarrow> bool" where
-  [code_abbrev]: "smember x s \<longleftrightarrow> x \<in> sset s"
-
-lemma smember_code[code, simp]: "smember x (y ## s) = (if x = y then True else smember x s)"
-  unfolding smember_def by auto
-
-hide_const (open) smember
-
-lemmas smap_simps[simp] = stream.map_sel
-lemmas shd_sset = stream.set_sel(1)
-lemmas stl_sset = stream.set_sel(2)
-
-theorem sset_induct[consumes 1, case_names shd stl, induct set: sset]:
-  assumes "y \<in> sset s" and "\<And>s. P (shd s) s" and "\<And>s y. \<lbrakk>y \<in> sset (stl s); P y (stl s)\<rbrakk> \<Longrightarrow> P y s"
-  shows "P y s"
-using assms by induct (metis stream.sel(1), auto)
-
-
-subsection {* prepend list to stream *}
-
-primrec shift :: "'a list \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr "@-" 65) where
-  "shift [] s = s"
-| "shift (x # xs) s = x ## shift xs s"
-
-lemma smap_shift[simp]: "smap f (xs @- s) = map f xs @- smap f s"
-  by (induct xs) auto
-
-lemma shift_append[simp]: "(xs @ ys) @- s = xs @- ys @- s"
-  by (induct xs) auto
-
-lemma shift_simps[simp]:
-   "shd (xs @- s) = (if xs = [] then shd s else hd xs)"
-   "stl (xs @- s) = (if xs = [] then stl s else tl xs @- s)"
-  by (induct xs) auto
-
-lemma sset_shift[simp]: "sset (xs @- s) = set xs \<union> sset s"
-  by (induct xs) auto
-
-lemma shift_left_inj[simp]: "xs @- s1 = xs @- s2 \<longleftrightarrow> s1 = s2"
-  by (induct xs) auto
-
-
-subsection {* set of streams with elements in some fixed set *}
-
-coinductive_set
-  streams :: "'a set \<Rightarrow> 'a stream set"
-  for A :: "'a set"
-where
-  Stream[intro!, simp, no_atp]: "\<lbrakk>a \<in> A; s \<in> streams A\<rbrakk> \<Longrightarrow> a ## s \<in> streams A"
-
-lemma shift_streams: "\<lbrakk>w \<in> lists A; s \<in> streams A\<rbrakk> \<Longrightarrow> w @- s \<in> streams A"
-  by (induct w) auto
-
-lemma streams_Stream: "x ## s \<in> streams A \<longleftrightarrow> x \<in> A \<and> s \<in> streams A"
-  by (auto elim: streams.cases)
-
-lemma streams_stl: "s \<in> streams A \<Longrightarrow> stl s \<in> streams A"
-  by (cases s) (auto simp: streams_Stream)
-
-lemma streams_shd: "s \<in> streams A \<Longrightarrow> shd s \<in> A"
-  by (cases s) (auto simp: streams_Stream)
-
-lemma sset_streams:
-  assumes "sset s \<subseteq> A"
-  shows "s \<in> streams A"
-using assms proof (coinduction arbitrary: s)
-  case streams then show ?case by (cases s) simp
-qed
-
-lemma streams_sset:
-  assumes "s \<in> streams A"
-  shows "sset s \<subseteq> A"
-proof
-  fix x assume "x \<in> sset s" from this `s \<in> streams A` show "x \<in> A"
-    by (induct s) (auto intro: streams_shd streams_stl)
-qed
-
-lemma streams_iff_sset: "s \<in> streams A \<longleftrightarrow> sset s \<subseteq> A"
-  by (metis sset_streams streams_sset)
-
-lemma streams_mono:  "s \<in> streams A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> s \<in> streams B"
-  unfolding streams_iff_sset by auto
-
-lemma smap_streams: "s \<in> streams A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> smap f s \<in> streams B"
-  unfolding streams_iff_sset stream.set_map by auto
-
-lemma streams_empty: "streams {} = {}"
-  by (auto elim: streams.cases)
-
-lemma streams_UNIV[simp]: "streams UNIV = UNIV"
-  by (auto simp: streams_iff_sset)
-
-subsection {* nth, take, drop for streams *}
-
-primrec snth :: "'a stream \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where
-  "s !! 0 = shd s"
-| "s !! Suc n = stl s !! n"
-
-lemma snth_smap[simp]: "smap f s !! n = f (s !! n)"
-  by (induct n arbitrary: s) auto
-
-lemma shift_snth_less[simp]: "p < length xs \<Longrightarrow> (xs @- s) !! p = xs ! p"
-  by (induct p arbitrary: xs) (auto simp: hd_conv_nth nth_tl)
-
-lemma shift_snth_ge[simp]: "p \<ge> length xs \<Longrightarrow> (xs @- s) !! p = s !! (p - length xs)"
-  by (induct p arbitrary: xs) (auto simp: Suc_diff_eq_diff_pred)
-
-lemma shift_snth: "(xs @- s) !! n = (if n < length xs then xs ! n else s !! (n - length xs))"
-  by auto
-
-lemma snth_sset[simp]: "s !! n \<in> sset s"
-  by (induct n arbitrary: s) (auto intro: shd_sset stl_sset)
-
-lemma sset_range: "sset s = range (snth s)"
-proof (intro equalityI subsetI)
-  fix x assume "x \<in> sset s"
-  thus "x \<in> range (snth s)"
-  proof (induct s)
-    case (stl s x)
-    then obtain n where "x = stl s !! n" by auto
-    thus ?case by (auto intro: range_eqI[of _ _ "Suc n"])
-  qed (auto intro: range_eqI[of _ _ 0])
-qed auto
-
-primrec stake :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a list" where
-  "stake 0 s = []"
-| "stake (Suc n) s = shd s # stake n (stl s)"
-
-lemma length_stake[simp]: "length (stake n s) = n"
-  by (induct n arbitrary: s) auto
-
-lemma stake_smap[simp]: "stake n (smap f s) = map f (stake n s)"
-  by (induct n arbitrary: s) auto
-
-lemma take_stake: "take n (stake m s) = stake (min n m) s"
-proof (induct m arbitrary: s n)
-  case (Suc m) thus ?case by (cases n) auto
-qed simp
-
-primrec sdrop :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
-  "sdrop 0 s = s"
-| "sdrop (Suc n) s = sdrop n (stl s)"
-
-lemma sdrop_simps[simp]:
-  "shd (sdrop n s) = s !! n" "stl (sdrop n s) = sdrop (Suc n) s"
-  by (induct n arbitrary: s)  auto
-
-lemma sdrop_smap[simp]: "sdrop n (smap f s) = smap f (sdrop n s)"
-  by (induct n arbitrary: s) auto
-
-lemma sdrop_stl: "sdrop n (stl s) = stl (sdrop n s)"
-  by (induct n) auto
-
-lemma drop_stake: "drop n (stake m s) = stake (m - n) (sdrop n s)"
-proof (induct m arbitrary: s n)
-  case (Suc m) thus ?case by (cases n) auto
-qed simp
-
-lemma stake_sdrop: "stake n s @- sdrop n s = s"
-  by (induct n arbitrary: s) auto
-
-lemma id_stake_snth_sdrop:
-  "s = stake i s @- s !! i ## sdrop (Suc i) s"
-  by (subst stake_sdrop[symmetric, of _ i]) (metis sdrop_simps stream.collapse)
-
-lemma smap_alt: "smap f s = s' \<longleftrightarrow> (\<forall>n. f (s !! n) = s' !! n)" (is "?L = ?R")
-proof
-  assume ?R
-  then have "\<And>n. smap f (sdrop n s) = sdrop n s'"
-    by coinduction (auto intro: exI[of _ 0] simp del: sdrop.simps(2))
-  then show ?L using sdrop.simps(1) by metis
-qed auto
-
-lemma stake_invert_Nil[iff]: "stake n s = [] \<longleftrightarrow> n = 0"
-  by (induct n) auto
-
-lemma sdrop_shift: "sdrop i (w @- s) = drop i w @- sdrop (i - length w) s"
-  by (induct i arbitrary: w s) (auto simp: drop_tl drop_Suc neq_Nil_conv)
-
-lemma stake_shift: "stake i (w @- s) = take i w @ stake (i - length w) s"
-  by (induct i arbitrary: w s) (auto simp: neq_Nil_conv)
-
-lemma stake_add[simp]: "stake m s @ stake n (sdrop m s) = stake (m + n) s"
-  by (induct m arbitrary: s) auto
-
-lemma sdrop_add[simp]: "sdrop n (sdrop m s) = sdrop (m + n) s"
-  by (induct m arbitrary: s) auto
-
-lemma sdrop_snth: "sdrop n s !! m = s !! (n + m)"
-  by (induct n arbitrary: m s) auto
-
-partial_function (tailrec) sdrop_while :: "('a \<Rightarrow> bool) \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where 
-  "sdrop_while P s = (if P (shd s) then sdrop_while P (stl s) else s)"
-
-lemma sdrop_while_SCons[code]:
-  "sdrop_while P (a ## s) = (if P a then sdrop_while P s else a ## s)"
-  by (subst sdrop_while.simps) simp
-
-lemma sdrop_while_sdrop_LEAST:
-  assumes "\<exists>n. P (s !! n)"
-  shows "sdrop_while (Not o P) s = sdrop (LEAST n. P (s !! n)) s"
-proof -
-  from assms obtain m where "P (s !! m)" "\<And>n. P (s !! n) \<Longrightarrow> m \<le> n"
-    and *: "(LEAST n. P (s !! n)) = m" by atomize_elim (auto intro: LeastI Least_le)
-  thus ?thesis unfolding *
-  proof (induct m arbitrary: s)
-    case (Suc m)
-    hence "sdrop_while (Not \<circ> P) (stl s) = sdrop m (stl s)"
-      by (metis (full_types) not_less_eq_eq snth.simps(2))
-    moreover from Suc(3) have "\<not> (P (s !! 0))" by blast
-    ultimately show ?case by (subst sdrop_while.simps) simp
-  qed (metis comp_apply sdrop.simps(1) sdrop_while.simps snth.simps(1))
-qed
-
-primcorec sfilter where
-  "shd (sfilter P s) = shd (sdrop_while (Not o P) s)"
-| "stl (sfilter P s) = sfilter P (stl (sdrop_while (Not o P) s))"
-
-lemma sfilter_Stream: "sfilter P (x ## s) = (if P x then x ## sfilter P s else sfilter P s)"
-proof (cases "P x")
-  case True thus ?thesis by (subst sfilter.ctr) (simp add: sdrop_while_SCons)
-next
-  case False thus ?thesis by (subst (1 2) sfilter.ctr) (simp add: sdrop_while_SCons)
-qed
-
-
-subsection {* unary predicates lifted to streams *}
-
-definition "stream_all P s = (\<forall>p. P (s !! p))"
-
-lemma stream_all_iff[iff]: "stream_all P s \<longleftrightarrow> Ball (sset s) P"
-  unfolding stream_all_def sset_range by auto
-
-lemma stream_all_shift[simp]: "stream_all P (xs @- s) = (list_all P xs \<and> stream_all P s)"
-  unfolding stream_all_iff list_all_iff by auto
-
-lemma stream_all_Stream: "stream_all P (x ## X) \<longleftrightarrow> P x \<and> stream_all P X"
-  by simp
-
-
-subsection {* recurring stream out of a list *}
-
-primcorec cycle :: "'a list \<Rightarrow> 'a stream" where
-  "shd (cycle xs) = hd xs"
-| "stl (cycle xs) = cycle (tl xs @ [hd xs])"
-
-lemma cycle_decomp: "u \<noteq> [] \<Longrightarrow> cycle u = u @- cycle u"
-proof (coinduction arbitrary: u)
-  case Eq_stream then show ?case using stream.collapse[of "cycle u"]
-    by (auto intro!: exI[of _ "tl u @ [hd u]"])
-qed
-
-lemma cycle_Cons[code]: "cycle (x # xs) = x ## cycle (xs @ [x])"
-  by (subst cycle.ctr) simp
-
-lemma cycle_rotated: "\<lbrakk>v \<noteq> []; cycle u = v @- s\<rbrakk> \<Longrightarrow> cycle (tl u @ [hd u]) = tl v @- s"
-  by (auto dest: arg_cong[of _ _ stl])
-
-lemma stake_append: "stake n (u @- s) = take (min (length u) n) u @ stake (n - length u) s"
-proof (induct n arbitrary: u)
-  case (Suc n) thus ?case by (cases u) auto
-qed auto
-
-lemma stake_cycle_le[simp]:
-  assumes "u \<noteq> []" "n < length u"
-  shows "stake n (cycle u) = take n u"
-using min_absorb2[OF less_imp_le_nat[OF assms(2)]]
-  by (subst cycle_decomp[OF assms(1)], subst stake_append) auto
-
-lemma stake_cycle_eq[simp]: "u \<noteq> [] \<Longrightarrow> stake (length u) (cycle u) = u"
-  by (subst cycle_decomp) (auto simp: stake_shift)
-
-lemma sdrop_cycle_eq[simp]: "u \<noteq> [] \<Longrightarrow> sdrop (length u) (cycle u) = cycle u"
-  by (subst cycle_decomp) (auto simp: sdrop_shift)
-
-lemma stake_cycle_eq_mod_0[simp]: "\<lbrakk>u \<noteq> []; n mod length u = 0\<rbrakk> \<Longrightarrow>
-   stake n (cycle u) = concat (replicate (n div length u) u)"
-  by (induct "n div length u" arbitrary: n u) (auto simp: stake_add[symmetric])
-
-lemma sdrop_cycle_eq_mod_0[simp]: "\<lbrakk>u \<noteq> []; n mod length u = 0\<rbrakk> \<Longrightarrow>
-   sdrop n (cycle u) = cycle u"
-  by (induct "n div length u" arbitrary: n u) (auto simp: sdrop_add[symmetric])
-
-lemma stake_cycle: "u \<noteq> [] \<Longrightarrow>
-   stake n (cycle u) = concat (replicate (n div length u) u) @ take (n mod length u) u"
-  by (subst mod_div_equality[of n "length u", symmetric], unfold stake_add[symmetric]) auto
-
-lemma sdrop_cycle: "u \<noteq> [] \<Longrightarrow> sdrop n (cycle u) = cycle (rotate (n mod length u) u)"
-  by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric])
-
-
-subsection {* iterated application of a function *}
-
-primcorec siterate where
-  "shd (siterate f x) = x"
-| "stl (siterate f x) = siterate f (f x)"
-
-lemma stake_Suc: "stake (Suc n) s = stake n s @ [s !! n]"
-  by (induct n arbitrary: s) auto
-
-lemma snth_siterate[simp]: "siterate f x !! n = (f^^n) x"
-  by (induct n arbitrary: x) (auto simp: funpow_swap1)
-
-lemma sdrop_siterate[simp]: "sdrop n (siterate f x) = siterate f ((f^^n) x)"
-  by (induct n arbitrary: x) (auto simp: funpow_swap1)
-
-lemma stake_siterate[simp]: "stake n (siterate f x) = map (\<lambda>n. (f^^n) x) [0 ..< n]"
-  by (induct n arbitrary: x) (auto simp del: stake.simps(2) simp: stake_Suc)
-
-lemma sset_siterate: "sset (siterate f x) = {(f^^n) x | n. True}"
-  by (auto simp: sset_range)
-
-lemma smap_siterate: "smap f (siterate f x) = siterate f (f x)"
-  by (coinduction arbitrary: x) auto
-
-
-subsection {* stream repeating a single element *}
-
-abbreviation "sconst \<equiv> siterate id"
-
-lemma shift_replicate_sconst[simp]: "replicate n x @- sconst x = sconst x"
-  by (subst (3) stake_sdrop[symmetric]) (simp add: map_replicate_trivial)
-
-lemma sset_sconst[simp]: "sset (sconst x) = {x}"
-  by (simp add: sset_siterate)
-
-lemma sconst_alt: "s = sconst x \<longleftrightarrow> sset s = {x}"
-proof
-  assume "sset s = {x}"
-  then show "s = sconst x"
-  proof (coinduction arbitrary: s)
-    case Eq_stream
-    then have "shd s = x" "sset (stl s) \<subseteq> {x}" by (case_tac [!] s) auto
-    then have "sset (stl s) = {x}" by (cases "stl s") auto
-    with `shd s = x` show ?case by auto
-  qed
-qed simp
-
-lemma same_cycle: "sconst x = cycle [x]"
-  by coinduction auto
-
-lemma smap_sconst: "smap f (sconst x) = sconst (f x)"
-  by coinduction auto
-
-lemma sconst_streams: "x \<in> A \<Longrightarrow> sconst x \<in> streams A"
-  by (simp add: streams_iff_sset)
-
-
-subsection {* stream of natural numbers *}
-
-abbreviation "fromN \<equiv> siterate Suc"
-
-abbreviation "nats \<equiv> fromN 0"
-
-lemma sset_fromN[simp]: "sset (fromN n) = {n ..}"
-  by (auto simp add: sset_siterate le_iff_add)
-
-lemma stream_smap_fromN: "s = smap (\<lambda>j. let i = j - n in s !! i) (fromN n)"
-  by (coinduction arbitrary: s n)
-    (force simp: neq_Nil_conv Let_def snth.simps(2)[symmetric] Suc_diff_Suc
-      intro: stream.map_cong split: if_splits simp del: snth.simps(2))
-
-lemma stream_smap_nats: "s = smap (snth s) nats"
-  using stream_smap_fromN[where n = 0] by simp
-
-
-subsection {* flatten a stream of lists *}
-
-primcorec flat where
-  "shd (flat ws) = hd (shd ws)"
-| "stl (flat ws) = flat (if tl (shd ws) = [] then stl ws else tl (shd ws) ## stl ws)"
-
-lemma flat_Cons[simp, code]: "flat ((x # xs) ## ws) = x ## flat (if xs = [] then ws else xs ## ws)"
-  by (subst flat.ctr) simp
-
-lemma flat_Stream[simp]: "xs \<noteq> [] \<Longrightarrow> flat (xs ## ws) = xs @- flat ws"
-  by (induct xs) auto
-
-lemma flat_unfold: "shd ws \<noteq> [] \<Longrightarrow> flat ws = shd ws @- flat (stl ws)"
-  by (cases ws) auto
-
-lemma flat_snth: "\<forall>xs \<in> sset s. xs \<noteq> [] \<Longrightarrow> flat s !! n = (if n < length (shd s) then 
-  shd s ! n else flat (stl s) !! (n - length (shd s)))"
-  by (metis flat_unfold not_less shd_sset shift_snth_ge shift_snth_less)
-
-lemma sset_flat[simp]: "\<forall>xs \<in> sset s. xs \<noteq> [] \<Longrightarrow> 
-  sset (flat s) = (\<Union>xs \<in> sset s. set xs)" (is "?P \<Longrightarrow> ?L = ?R")
-proof safe
-  fix x assume ?P "x : ?L"
-  then obtain m where "x = flat s !! m" by (metis image_iff sset_range)
-  with `?P` obtain n m' where "x = s !! n ! m'" "m' < length (s !! n)"
-  proof (atomize_elim, induct m arbitrary: s rule: less_induct)
-    case (less y)
-    thus ?case
-    proof (cases "y < length (shd s)")
-      case True thus ?thesis by (metis flat_snth less(2,3) snth.simps(1))
-    next
-      case False
-      hence "x = flat (stl s) !! (y - length (shd s))" by (metis less(2,3) flat_snth)
-      moreover
-      { from less(2) have *: "length (shd s) > 0" by (cases s) simp_all
-        with False have "y > 0" by (cases y) simp_all
-        with * have "y - length (shd s) < y" by simp
-      }
-      moreover have "\<forall>xs \<in> sset (stl s). xs \<noteq> []" using less(2) by (cases s) auto
-      ultimately have "\<exists>n m'. x = stl s !! n ! m' \<and> m' < length (stl s !! n)" by (intro less(1)) auto
-      thus ?thesis by (metis snth.simps(2))
-    qed
-  qed
-  thus "x \<in> ?R" by (auto simp: sset_range dest!: nth_mem)
-next
-  fix x xs assume "xs \<in> sset s" ?P "x \<in> set xs" thus "x \<in> ?L"
-    by (induct rule: sset_induct)
-      (metis UnI1 flat_unfold shift.simps(1) sset_shift,
-       metis UnI2 flat_unfold shd_sset stl_sset sset_shift)
-qed
-
-
-subsection {* merge a stream of streams *}
-
-definition smerge :: "'a stream stream \<Rightarrow> 'a stream" where
-  "smerge ss = flat (smap (\<lambda>n. map (\<lambda>s. s !! n) (stake (Suc n) ss) @ stake n (ss !! n)) nats)"
-
-lemma stake_nth[simp]: "m < n \<Longrightarrow> stake n s ! m = s !! m"
-  by (induct n arbitrary: s m) (auto simp: nth_Cons', metis Suc_pred snth.simps(2))
-
-lemma snth_sset_smerge: "ss !! n !! m \<in> sset (smerge ss)"
-proof (cases "n \<le> m")
-  case False thus ?thesis unfolding smerge_def
-    by (subst sset_flat)
-      (auto simp: stream.set_map in_set_conv_nth simp del: stake.simps
-        intro!: exI[of _ n, OF disjI2] exI[of _ m, OF mp])
-next
-  case True thus ?thesis unfolding smerge_def
-    by (subst sset_flat)
-      (auto simp: stream.set_map in_set_conv_nth image_iff simp del: stake.simps snth.simps
-        intro!: exI[of _ m, OF disjI1] bexI[of _ "ss !! n"] exI[of _ n, OF mp])
-qed
-
-lemma sset_smerge: "sset (smerge ss) = UNION (sset ss) sset"
-proof safe
-  fix x assume "x \<in> sset (smerge ss)"
-  thus "x \<in> UNION (sset ss) sset"
-    unfolding smerge_def by (subst (asm) sset_flat)
-      (auto simp: stream.set_map in_set_conv_nth sset_range simp del: stake.simps, fast+)
-next
-  fix s x assume "s \<in> sset ss" "x \<in> sset s"
-  thus "x \<in> sset (smerge ss)" using snth_sset_smerge by (auto simp: sset_range)
-qed
-
-
-subsection {* product of two streams *}
-
-definition sproduct :: "'a stream \<Rightarrow> 'b stream \<Rightarrow> ('a \<times> 'b) stream" where
-  "sproduct s1 s2 = smerge (smap (\<lambda>x. smap (Pair x) s2) s1)"
-
-lemma sset_sproduct: "sset (sproduct s1 s2) = sset s1 \<times> sset s2"
-  unfolding sproduct_def sset_smerge by (auto simp: stream.set_map)
-
-
-subsection {* interleave two streams *}
-
-primcorec sinterleave where
-  "shd (sinterleave s1 s2) = shd s1"
-| "stl (sinterleave s1 s2) = sinterleave s2 (stl s1)"
-
-lemma sinterleave_code[code]:
-  "sinterleave (x ## s1) s2 = x ## sinterleave s2 s1"
-  by (subst sinterleave.ctr) simp
-
-lemma sinterleave_snth[simp]:
-  "even n \<Longrightarrow> sinterleave s1 s2 !! n = s1 !! (n div 2)"
-   "odd n \<Longrightarrow> sinterleave s1 s2 !! n = s2 !! (n div 2)"
-  by (induct n arbitrary: s1 s2)
-    (auto dest: even_nat_Suc_div_2 odd_nat_plus_one_div_two[folded nat_2])
-
-lemma sset_sinterleave: "sset (sinterleave s1 s2) = sset s1 \<union> sset s2"
-proof (intro equalityI subsetI)
-  fix x assume "x \<in> sset (sinterleave s1 s2)"
-  then obtain n where "x = sinterleave s1 s2 !! n" unfolding sset_range by blast
-  thus "x \<in> sset s1 \<union> sset s2" by (cases "even n") auto
-next
-  fix x assume "x \<in> sset s1 \<union> sset s2"
-  thus "x \<in> sset (sinterleave s1 s2)"
-  proof
-    assume "x \<in> sset s1"
-    then obtain n where "x = s1 !! n" unfolding sset_range by blast
-    hence "sinterleave s1 s2 !! (2 * n) = x" by simp
-    thus ?thesis unfolding sset_range by blast
-  next
-    assume "x \<in> sset s2"
-    then obtain n where "x = s2 !! n" unfolding sset_range by blast
-    hence "sinterleave s1 s2 !! (2 * n + 1) = x" by simp
-    thus ?thesis unfolding sset_range by blast
-  qed
-qed
-
-
-subsection {* zip *}
-
-primcorec szip where
-  "shd (szip s1 s2) = (shd s1, shd s2)"
-| "stl (szip s1 s2) = szip (stl s1) (stl s2)"
-
-lemma szip_unfold[code]: "szip (a ## s1) (b ## s2) = (a, b) ## (szip s1 s2)"
-  by (subst szip.ctr) simp
-
-lemma snth_szip[simp]: "szip s1 s2 !! n = (s1 !! n, s2 !! n)"
-  by (induct n arbitrary: s1 s2) auto
-
-lemma stake_szip[simp]:
-  "stake n (szip s1 s2) = zip (stake n s1) (stake n s2)"
-  by (induct n arbitrary: s1 s2) auto
-
-lemma sdrop_szip[simp]: "sdrop n (szip s1 s2) = szip (sdrop n s1) (sdrop n s2)"
-  by (induct n arbitrary: s1 s2) auto
-
-lemma smap_szip_fst:
-  "smap (\<lambda>x. f (fst x)) (szip s1 s2) = smap f s1"
-  by (coinduction arbitrary: s1 s2) auto
-
-lemma smap_szip_snd:
-  "smap (\<lambda>x. g (snd x)) (szip s1 s2) = smap g s2"
-  by (coinduction arbitrary: s1 s2) auto
-
-
-subsection {* zip via function *}
-
-primcorec smap2 where
-  "shd (smap2 f s1 s2) = f (shd s1) (shd s2)"
-| "stl (smap2 f s1 s2) = smap2 f (stl s1) (stl s2)"
-
-lemma smap2_unfold[code]:
-  "smap2 f (a ## s1) (b ## s2) = f a b ## (smap2 f s1 s2)"
-  by (subst smap2.ctr) simp
-
-lemma smap2_szip:
-  "smap2 f s1 s2 = smap (split f) (szip s1 s2)"
-  by (coinduction arbitrary: s1 s2) auto
-
-lemma smap_smap2[simp]:
-  "smap f (smap2 g s1 s2) = smap2 (\<lambda>x y. f (g x y)) s1 s2"
-  unfolding smap2_szip stream.map_comp o_def split_def ..
-
-lemma smap2_alt:
-  "(smap2 f s1 s2 = s) = (\<forall>n. f (s1 !! n) (s2 !! n) = s !! n)"
-  unfolding smap2_szip smap_alt by auto
-
-lemma snth_smap2[simp]:
-  "smap2 f s1 s2 !! n = f (s1 !! n) (s2 !! n)"
-  by (induct n arbitrary: s1 s2) auto
-
-lemma stake_smap2[simp]:
-  "stake n (smap2 f s1 s2) = map (split f) (zip (stake n s1) (stake n s2))"
-  by (induct n arbitrary: s1 s2) auto
-
-lemma sdrop_smap2[simp]:
-  "sdrop n (smap2 f s1 s2) = smap2 f (sdrop n s1) (sdrop n s2)"
-  by (induct n arbitrary: s1 s2) auto
-
-end
--- a/src/HOL/BNF_Examples/Stream_Processor.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,189 +0,0 @@
-(*  Title:      HOL/BNF_Examples/Stream_Processor.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2014
-
-Stream processors---a syntactic representation of continuous functions on streams.
-*)
-
-header {* Stream Processors---A Syntactic Representation of Continuous Functions on Streams *}
-
-theory Stream_Processor
-imports Stream "~~/src/HOL/Library/BNF_Axiomatization"
-begin
-
-section {* Continuous Functions on Streams *}
-
-datatype_new ('a, 'b, 'c) sp\<^sub>\<mu> = Get "'a \<Rightarrow> ('a, 'b, 'c) sp\<^sub>\<mu>" | Put "'b" "'c"
-codatatype ('a, 'b) sp\<^sub>\<nu> = In (out: "('a, 'b, ('a, 'b) sp\<^sub>\<nu>) sp\<^sub>\<mu>")
-
-primrec run\<^sub>\<mu> :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> 'a stream \<Rightarrow> ('b \<times> 'c) \<times> 'a stream" where
-  "run\<^sub>\<mu> (Get f) s = run\<^sub>\<mu> (f (shd s)) (stl s)"
-| "run\<^sub>\<mu> (Put b sp) s = ((b, sp), s)"
-
-primcorec run\<^sub>\<nu> :: "('a, 'b) sp\<^sub>\<nu> \<Rightarrow> 'a stream \<Rightarrow> 'b stream" where
-  "run\<^sub>\<nu> sp s = (let ((h, sp), s) = run\<^sub>\<mu> (out sp) s in h ## run\<^sub>\<nu> sp s)"
-
-primcorec copy :: "('a, 'a) sp\<^sub>\<nu>" where
-  "copy = In (Get (\<lambda>a. Put a copy))"
-
-lemma run\<^sub>\<nu>_copy: "run\<^sub>\<nu> copy s = s"
-  by (coinduction arbitrary: s) simp
-
-text {*
-To use the function package for the definition of composition the
-wellfoundedness of the subtree relation needs to be proved first.
-*}
-
-definition "sub \<equiv> {(f a, Get f) | a f. True}"
-
-lemma subI[intro]: "(f a, Get f) \<in> sub"
-  unfolding sub_def by blast
-
-lemma wf_sub[simp, intro]: "wf sub"
-proof (rule wfUNIVI)
-  fix P  :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> bool" and x
-  assume "\<forall>x. (\<forall>y. (y, x) \<in> sub \<longrightarrow> P y) \<longrightarrow> P x"
-  hence I: "\<And>x. (\<forall>y. (\<exists>a f. y = f a \<and> x = Get f) \<longrightarrow> P y) \<Longrightarrow> P x" unfolding sub_def by blast
-  show "P x" by (induct x) (auto intro: I)
-qed
-
-function
-  sp\<^sub>\<mu>_comp :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> ('d, 'a, ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu> \<Rightarrow> ('d, 'b, 'c \<times> ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu>"
-  (infixl "o\<^sub>\<mu>" 65)
-where
-  "Put b sp o\<^sub>\<mu> fsp = Put b (sp, In fsp)"
-| "Get f o\<^sub>\<mu> Put b sp = f b o\<^sub>\<mu> out sp"
-| "Get f o\<^sub>\<mu> Get g = Get (\<lambda>a. Get f o\<^sub>\<mu> g a)"
-by pat_completeness auto
-termination by (relation "lex_prod sub sub") auto
-
-primcorec sp\<^sub>\<nu>_comp (infixl "o\<^sub>\<nu>" 65) where
-  "out (sp o\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sub>\<nu> sp') (out sp o\<^sub>\<mu> out sp')"
-
-lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp: "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
-proof (rule ext, unfold comp_apply)
-  fix s
-  show "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') s = run\<^sub>\<nu> sp (run\<^sub>\<nu> sp' s)"
-  proof (coinduction arbitrary: sp sp' s)
-    case Eq_stream
-    show ?case
-    proof (induct "out sp" "out sp'" arbitrary: sp sp' s rule: sp\<^sub>\<mu>_comp.induct)
-      case (1 b sp'')
-      show ?case by (auto simp add: 1[symmetric])
-    next
-      case (2 f b sp'')
-      from 2(1)[of "In (f b)" sp''] show ?case by (simp add: 2(2,3)[symmetric])
-    next
-      case (3 f h)
-      from 3(1)[of _ "shd s" "In (h (shd s))", OF 3(2)] show ?case by (simp add: 3(2,3)[symmetric])
-    qed
-  qed
-qed
-
-text {* Alternative definition of composition using primrec instead of function *}
-
-primrec sp\<^sub>\<mu>_comp2R  where
-  "sp\<^sub>\<mu>_comp2R f (Put b sp) = f b (out sp)"
-| "sp\<^sub>\<mu>_comp2R f (Get h) = Get (sp\<^sub>\<mu>_comp2R f o h)"
-
-primrec sp\<^sub>\<mu>_comp2 (infixl "o\<^sup>*\<^sub>\<mu>" 65) where
-  "Put b sp o\<^sup>*\<^sub>\<mu> fsp = Put b (sp, In fsp)"
-| "Get f o\<^sup>*\<^sub>\<mu> fsp = sp\<^sub>\<mu>_comp2R (op o\<^sup>*\<^sub>\<mu> o f) fsp"
-
-primcorec sp\<^sub>\<nu>_comp2 (infixl "o\<^sup>*\<^sub>\<nu>" 65) where
-  "out (sp o\<^sup>*\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sup>*\<^sub>\<nu> sp') (out sp o\<^sup>*\<^sub>\<mu> out sp')"
-
-lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp2: "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
-proof (rule ext, unfold comp_apply)
-  fix s
-  show "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') s = run\<^sub>\<nu> sp (run\<^sub>\<nu> sp' s)"
-  proof (coinduction arbitrary: sp sp' s)
-    case Eq_stream
-    show ?case
-    proof (induct "out sp" arbitrary: sp sp' s)
-      case (Put b sp'')
-      show ?case by (auto simp add: Put[symmetric])
-    next
-      case (Get f)
-      then show ?case
-      proof (induct "out sp'" arbitrary: sp sp' s)
-        case (Put b sp'')
-        from Put(2)[of "In (f b)" sp''] show ?case by (simp add: Put(1,3)[symmetric])
-      next
-        case (Get h)
-        from Get(1)[OF _ Get(3,4), of "In (h (shd s))"] show ?case by (simp add: Get(2,4)[symmetric])
-      qed
-    qed
-  qed
-qed
-
-text {* The two definitions are equivalent *}
-
-lemma sp\<^sub>\<mu>_comp_sp\<^sub>\<mu>_comp2[simp]: "sp o\<^sub>\<mu> sp' = sp o\<^sup>*\<^sub>\<mu> sp'"
-  by (induct sp sp' rule: sp\<^sub>\<mu>_comp.induct) auto
-
-(*will be provided by the package*)
-lemma sp\<^sub>\<mu>_rel_map_map[unfolded vimage2p_def, simp]:
-  "rel_sp\<^sub>\<mu> R1 R2 (map_sp\<^sub>\<mu> f1 f2 sp) (map_sp\<^sub>\<mu> g1 g2 sp') =
-  rel_sp\<^sub>\<mu> (BNF_Def.vimage2p f1 g1 R1) (BNF_Def.vimage2p f2 g2 R2) sp sp'"
-by (tactic {*
-  let val ks = 1 upto 2;
-  in
-    BNF_Tactics.unfold_thms_tac @{context}
-      @{thms sp\<^sub>\<mu>.rel_compp sp\<^sub>\<mu>.rel_conversep sp\<^sub>\<mu>.rel_Grp vimage2p_Grp} THEN
-    HEADGOAL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
-      BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac,
-      rtac @{thm conversepI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
-      BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks,
-      REPEAT_DETERM o eresolve_tac @{thms relcomppE conversepE GrpE},
-      hyp_subst_tac @{context}, atac])
-  end
-*})
-
-lemma sp\<^sub>\<mu>_rel_self: "\<lbrakk>op = \<le> R1; op = \<le> R2\<rbrakk> \<Longrightarrow> rel_sp\<^sub>\<mu> R1 R2 x x"
-  by (erule (1) predicate2D[OF sp\<^sub>\<mu>.rel_mono]) (simp only: sp\<^sub>\<mu>.rel_eq)
-
-lemma sp\<^sub>\<nu>_comp_sp\<^sub>\<nu>_comp2: "sp o\<^sub>\<nu> sp' = sp o\<^sup>*\<^sub>\<nu> sp'"
-  by (coinduction arbitrary: sp sp') (auto intro!: sp\<^sub>\<mu>_rel_self)
-
-
-section {* Generalization to an Arbitrary BNF as Codomain *}
-
-bnf_axiomatization ('a, 'b) F for map: F
-
-notation BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
-
-definition \<theta> :: "('p,'a) F * 'b \<Rightarrow> ('p,'a * 'b) F" where
-  "\<theta> xb = F id \<langle>id, \<lambda> a. (snd xb)\<rangle> (fst xb)"
-
-(* The strength laws for \<theta>: *)
-lemma \<theta>_natural: "F id (map_prod f g) o \<theta> = \<theta> o map_prod (F id f) g"
-  unfolding \<theta>_def F.map_comp comp_def id_apply convol_def map_prod_def split_beta fst_conv snd_conv ..
-
-definition assl :: "'a * ('b * 'c) \<Rightarrow> ('a * 'b) * 'c" where
-  "assl abc = ((fst abc, fst (snd abc)), snd (snd abc))"
-
-lemma \<theta>_rid: "F id fst o \<theta> = fst"
-  unfolding \<theta>_def F.map_comp F.map_id comp_def id_apply convol_def fst_conv sym[OF id_def] ..
-
-lemma \<theta>_assl: "F id assl o \<theta> = \<theta> o map_prod \<theta> id o assl"
-  unfolding assl_def \<theta>_def F.map_comp comp_def id_apply convol_def map_prod_def split fst_conv snd_conv ..
-
-datatype_new ('a, 'b, 'c) spF\<^sub>\<mu> = GetF "'a \<Rightarrow> ('a, 'b, 'c) spF\<^sub>\<mu>" | PutF "('b,'c) F"
-codatatype ('a, 'b) spF\<^sub>\<nu> = InF (outF: "('a, 'b, ('a, 'b) spF\<^sub>\<nu>) spF\<^sub>\<mu>")
-
-codatatype 'b JF = Ctor (dtor: "('b, 'b JF) F")
-
-(* Definition of run for an arbitrary final coalgebra as codomain: *)
-
-primrec
-  runF\<^sub>\<mu> :: "('a, 'b, ('a, 'b) spF\<^sub>\<nu>) spF\<^sub>\<mu> \<Rightarrow> 'a stream \<Rightarrow> (('b, ('a, 'b) spF\<^sub>\<nu>) F) \<times> 'a stream" 
-where
-  "runF\<^sub>\<mu> (GetF f) s = (runF\<^sub>\<mu> o f) (shd s) (stl s)"
-| "runF\<^sub>\<mu> (PutF x) s = (x, s)"
-
-primcorec runF\<^sub>\<nu> :: "('a, 'b) spF\<^sub>\<nu> \<Rightarrow> 'a stream \<Rightarrow> 'b JF" where
-  "runF\<^sub>\<nu> sp s = (let (x, s) = runF\<^sub>\<mu> (outF sp) s in Ctor (F id (\<lambda> sp. runF\<^sub>\<nu> sp s) x))"
-
-end
--- a/src/HOL/BNF_Examples/TreeFI.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,46 +0,0 @@
-(*  Title:      HOL/BNF_Examples/TreeFI.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Finitely branching possibly infinite trees.
-*)
-
-header {* Finitely Branching Possibly Infinite Trees *}
-
-theory TreeFI
-imports Main
-begin
-
-codatatype 'a treeFI = Tree (lab: 'a) (sub: "'a treeFI list")
-
-(* Tree reverse:*)
-primcorec trev where
-  "lab (trev t) = lab t"
-| "sub (trev t) = map trev (rev (sub t))"
-
-lemma treeFI_coinduct:
-  assumes *: "phi x y"
-  and step: "\<And>a b. phi a b \<Longrightarrow>
-     lab a = lab b \<and>
-     length (sub a) = length (sub b) \<and>
-     (\<forall>i < length (sub a). phi (nth (sub a) i) (nth (sub b) i))"
-  shows "x = y"
-using * proof (coinduction arbitrary: x y)
-  case (Eq_treeFI t1 t2)
-  from conjunct1[OF conjunct2[OF step[OF Eq_treeFI]]] conjunct2[OF conjunct2[OF step[OF Eq_treeFI]]]
-  have "list_all2 phi (sub t1) (sub t2)"
-  proof (induction "sub t1" "sub t2" arbitrary: t1 t2 rule: list_induct2)
-    case (Cons x xs y ys)
-    note sub = Cons(3,4)[symmetric] and phi = mp[OF spec[OF Cons(5)], unfolded sub]
-      and IH = Cons(2)[of "Tree (lab t1) (tl (sub t1))" "Tree (lab t2) (tl (sub t2))",
-        unfolded sub, simplified]
-    from phi[of 0] show ?case unfolding sub by (auto intro!: IH dest: phi[simplified, OF Suc_mono])
-  qed simp
-  with conjunct1[OF step[OF Eq_treeFI]] show ?case by simp
-qed
-
-lemma trev_trev: "trev (trev tr) = tr"
-  by (coinduction arbitrary: tr rule: treeFI_coinduct) (auto simp add: rev_map)
-
-end
--- a/src/HOL/BNF_Examples/TreeFsetI.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-(*  Title:      HOL/BNF_Examples/TreeFsetI.thy
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Finitely branching possibly infinite trees, with sets of children.
-*)
-
-header {* Finitely Branching Possibly Infinite Trees, with Sets of Children *}
-
-theory TreeFsetI
-imports "~~/src/HOL/Library/FSet"
-begin
-
-codatatype 'a treeFsetI = Tree (lab: 'a) (sub: "'a treeFsetI fset")
-
-(* tree map (contrived example): *)
-primcorec tmap where
-"lab (tmap f t) = f (lab t)" |
-"sub (tmap f t) = fimage (tmap f) (sub t)"
-
-lemma "tmap (f o g) x = tmap f (tmap g x)"
-  by (coinduction arbitrary: x) (auto simp: rel_fset_alt)
-
-end
--- a/src/HOL/BNF_Examples/Verilog.thy	Thu Sep 11 19:20:23 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,138 +0,0 @@
-(*  Title:      HOL/Datatype_Benchmark/Verilog.thy
-
-Example from Daryl: a Verilog grammar.
-*)
-
-theory Verilog imports Main begin
-
-datatype
-  Source_text
-     = module string "string list" "Module_item list"
-     | Source_textMeta string
-and
-  Module_item
-     = "declaration" Declaration
-     | initial Statement
-     | always Statement
-     | assign Lvalue Expression
-     | Module_itemMeta string
-and
-  Declaration
-     = reg_declaration "Range option" "string list"
-     | net_declaration "Range option" "string list"
-     | input_declaration "Range option" "string list"
-     | output_declaration "Range option" "string list"
-     | DeclarationMeta string
-and
-  Range = range Expression Expression | RangeMeta string
-and
-  Statement
-     = clock_statement Clock Statement_or_null
-     | blocking_assignment Lvalue Expression
-     | non_blocking_assignment Lvalue Expression
-     | conditional_statement
-          Expression Statement_or_null "Statement_or_null option"
-     | case_statement Expression "Case_item list"
-     | while_loop Expression Statement
-     | repeat_loop Expression Statement
-     | for_loop
-          Lvalue Expression Expression Lvalue Expression Statement
-     | forever_loop Statement
-     | disable string
-     | seq_block "string option" "Statement list"
-     | StatementMeta string
-and
-  Statement_or_null
-     = statement Statement | null_statement | Statement_or_nullMeta string
-and
-  Clock
-     = posedge string
-     | negedge string
-     | clock string
-     | ClockMeta string
-and
-  Case_item
-     = case_item "Expression list" Statement_or_null
-     | default_case_item Statement_or_null
-     | Case_itemMeta string
-and
-  Expression
-     = plus Expression Expression
-     | minus Expression Expression
-     | lshift Expression Expression
-     | rshift Expression Expression
-     | lt Expression Expression
-     | leq Expression Expression
-     | gt Expression Expression
-     | geq Expression Expression
-     | logeq Expression Expression
-     | logneq Expression Expression
-     | caseeq Expression Expression
-     | caseneq Expression Expression
-     | bitand Expression Expression
-     | bitxor Expression Expression
-     | bitor Expression Expression
-     | logand Expression Expression
-     | logor Expression Expression
-     | conditional Expression Expression Expression
-     | positive Primary
-     | negative Primary
-     | lognot Primary
-     | bitnot Primary
-     | reducand Primary
-     | reducxor Primary
-     | reducor Primary
-     | reducnand Primary
-     | reducxnor Primary
-     | reducnor Primary
-     | primary Primary
-     | ExpressionMeta string
-and
-  Primary
-     = primary_number Number
-     | primary_IDENTIFIER string
-     | primary_bit_select string Expression
-     | primary_part_select string Expression Expression
-     | primary_gen_bit_select Expression Expression
-     | primary_gen_part_select Expression Expression Expression
-     | primary_concatenation Concatenation
-     | primary_multiple_concatenation Multiple_concatenation
-     | brackets Expression
-     | PrimaryMeta string
-and
-  Lvalue
-     = lvalue string
-     | lvalue_bit_select string Expression
-     | lvalue_part_select string Expression Expression
-     | lvalue_concatenation Concatenation
-     | LvalueMeta string
-and
-  Number
-     = decimal string
-     | based "string option" string
-     | NumberMeta string
-and
-  Concatenation
-     = concatenation "Expression list" | ConcatenationMeta string
-and
-  Multiple_concatenation
-     = multiple_concatenation Expression "Expression list"
-     | Multiple_concatenationMeta string
-and
-  meta
-     = Meta_Source_text Source_text
-     | Meta_Module_item Module_item
-     | Meta_Declaration Declaration
-     | Meta_Range Range
-     | Meta_Statement Statement
-     | Meta_Statement_or_null Statement_or_null
-     | Meta_Clock Clock
-     | Meta_Case_item Case_item
-     | Meta_Expression Expression
-     | Meta_Primary Primary
-     | Meta_Lvalue Lvalue
-     | Meta_Number Number
-     | Meta_Concatenation Concatenation
-     | Meta_Multiple_concatenation Multiple_concatenation
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Brackin.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,41 @@
+(*  Title:      HOL/Datatype_Benchmark/Brackin.thy
+
+A couple of datatypes from Steve Brackin's work.
+*)
+
+theory Brackin imports Main begin
+
+datatype T =
+    X1 | X2 | X3 | X4 | X5 | X6 | X7 | X8 | X9 | X10 | X11 |
+    X12 | X13 | X14 | X15 | X16 | X17 | X18 | X19 | X20 | X21 |
+    X22 | X23 | X24 | X25 | X26 | X27 | X28 | X29 | X30 | X31 |
+    X32 | X33 | X34
+
+datatype ('a, 'b, 'c) TY1 =
+      NoF
+    | Fk 'a "('a, 'b, 'c) TY2"
+and ('a, 'b, 'c) TY2 =
+      Ta bool
+    | Td bool
+    | Tf "('a, 'b, 'c) TY1"
+    | Tk bool
+    | Tp bool
+    | App 'a "('a, 'b, 'c) TY1" "('a, 'b, 'c) TY2" "('a, 'b, 'c) TY3"
+    | Pair "('a, 'b, 'c) TY2" "('a, 'b, 'c) TY2"
+and ('a, 'b, 'c) TY3 =
+      NoS
+    | Fresh "('a, 'b, 'c) TY2"
+    | Trustworthy 'a
+    | PrivateKey 'a 'b 'c
+    | PublicKey 'a 'b 'c
+    | Conveyed 'a "('a, 'b, 'c) TY2"
+    | Possesses 'a "('a, 'b, 'c) TY2"
+    | Received 'a "('a, 'b, 'c) TY2"
+    | Recognizes 'a "('a, 'b, 'c) TY2"
+    | NeverMalFromSelf 'a 'b "('a, 'b, 'c) TY2"
+    | Sends 'a "('a, 'b, 'c) TY2" 'b
+    | SharedSecret 'a "('a, 'b, 'c) TY2" 'b
+    | Believes 'a "('a, 'b, 'c) TY3"
+    | And "('a, 'b, 'c) TY3" "('a, 'b, 'c) TY3"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Compat.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,236 @@
+(*  Title:      HOL/Datatype_Examples/Compat.thy
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2014
+
+Tests for compatibility with the old datatype package.
+*)
+
+header \<open> Tests for Compatibility with the Old Datatype Package \<close>
+
+theory Compat
+imports Main
+begin
+
+subsection \<open> Viewing and Registering New-Style Datatypes as Old-Style Ones \<close>
+
+ML \<open>
+fun check_len n xs label =
+  length xs = n orelse error ("Expected length " ^ string_of_int (length xs) ^ " for " ^ label);
+
+fun check_lens (n1, n2, n3) (xs1, xs2, xs3) =
+  check_len n1 xs1 "old" andalso check_len n2 xs2 "unfold" andalso check_len n3 xs3 "keep";
+
+fun get_descrs thy lens T_name =
+  (these (Option.map #descr (Old_Datatype_Data.get_info thy T_name)),
+   these (Option.map #descr (BNF_LFP_Compat.get_info thy BNF_LFP_Compat.Unfold_Nesting T_name)),
+   these (Option.map #descr (BNF_LFP_Compat.get_info thy BNF_LFP_Compat.Keep_Nesting T_name)))
+  |> tap (check_lens lens);
+\<close>
+
+old_datatype 'a old_lst = Old_Nl | Old_Cns 'a "'a old_lst"
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name old_lst}; \<close>
+
+datatype_new 'a lst = Nl | Cns 'a "'a lst"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name lst}; \<close>
+
+datatype_compat lst
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name lst}; \<close>
+
+datatype_new 'b w = W | W' "'b w \<times> 'b list"
+
+(* no support for sums of products:
+datatype_compat w
+*)
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name w}; \<close>
+
+datatype_new ('c, 'b) s = L 'c | R 'b
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name s}; \<close>
+
+datatype_new 'd x = X | X' "('d x lst, 'd list) s"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name x}; \<close>
+
+datatype_compat s
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name s}; \<close>
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name x}; \<close>
+
+datatype_compat x
+
+ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name x}; \<close>
+
+thm x.induct x.rec
+thm compat_x.induct compat_x.rec
+
+datatype_new 'a tttre = TTTre 'a "'a tttre lst lst lst"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name tttre}; \<close>
+
+datatype_compat tttre
+
+ML \<open> get_descrs @{theory} (4, 4, 1) @{type_name tttre}; \<close>
+
+thm tttre.induct tttre.rec
+thm compat_tttre.induct compat_tttre.rec
+
+datatype_new 'a ftre = FEmp | FTre "'a \<Rightarrow> 'a ftre lst"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name ftre}; \<close>
+
+datatype_compat ftre
+
+ML \<open> get_descrs @{theory} (2, 2, 1) @{type_name ftre}; \<close>
+
+thm ftre.induct ftre.rec
+thm compat_ftre.induct compat_ftre.rec
+
+datatype_new 'a btre = BTre 'a "'a btre lst" "'a btre lst"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name btre}; \<close>
+
+datatype_compat btre
+
+ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name btre}; \<close>
+
+thm btre.induct btre.rec
+thm compat_btre.induct compat_btre.rec
+
+datatype_new 'a foo = Foo | Foo' 'a "'a bar" and 'a bar = Bar | Bar' 'a "'a foo"
+
+ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name foo}; \<close>
+ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name bar}; \<close>
+
+datatype_compat foo bar
+
+ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name foo}; \<close>
+ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name bar}; \<close>
+
+datatype_new 'a tre = Tre 'a "'a tre lst"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name tre}; \<close>
+
+datatype_compat tre
+
+ML \<open> get_descrs @{theory} (2, 2, 1) @{type_name tre}; \<close>
+
+thm tre.induct tre.rec
+thm compat_tre.induct compat_tre.rec
+
+datatype_new 'a f = F 'a and 'a g = G 'a
+
+ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name f}; \<close>
+ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name g}; \<close>
+
+datatype_new h = H "h f" | H'
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name h}; \<close>
+
+datatype_compat f g
+
+ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name f}; \<close>
+ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name g}; \<close>
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name h}; \<close>
+
+datatype_compat h
+
+ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name h}; \<close>
+
+thm h.induct h.rec
+thm compat_h.induct compat_h.rec
+
+datatype_new myunit = MyUnity
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name myunit}; \<close>
+
+datatype_compat myunit
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name myunit}; \<close>
+
+datatype_new mylist = MyNil | MyCons nat mylist
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name mylist}; \<close>
+
+datatype_compat mylist
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name mylist}; \<close>
+
+datatype_new foo' = FooNil | FooCons bar' foo' and bar' = Bar
+
+ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name foo'}; \<close>
+ML \<open> get_descrs @{theory} (0, 2, 2) @{type_name bar'}; \<close>
+
+datatype_compat bar' foo'
+
+ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name foo'}; \<close>
+ML \<open> get_descrs @{theory} (2, 2, 2) @{type_name bar'}; \<close>
+
+old_datatype funky = Funky "funky tre" | Funky'
+
+ML \<open> get_descrs @{theory} (3, 3, 3) @{type_name funky}; \<close>
+
+old_datatype fnky = Fnky "nat tre"
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name fnky}; \<close>
+
+datatype_new tree = Tree "tree foo"
+
+ML \<open> get_descrs @{theory} (0, 1, 1) @{type_name tree}; \<close>
+
+datatype_compat tree
+
+ML \<open> get_descrs @{theory} (3, 3, 1) @{type_name tree}; \<close>
+
+thm tree.induct tree.rec
+thm compat_tree.induct compat_tree.rec
+
+
+subsection \<open> Creating New-Style Datatypes Using Old-Style Interfaces \<close>
+
+ML \<open>
+val l_specs =
+  [((@{binding l}, [("'a", @{sort type})], NoSyn),
+   [(@{binding N}, [], NoSyn),
+    (@{binding C}, [@{typ 'a}, Type (Sign.full_name @{theory} @{binding l}, [@{typ 'a}])], NoSyn)])];
+\<close>
+
+setup \<open> snd o BNF_LFP_Compat.add_datatype BNF_LFP_Compat.Unfold_Nesting l_specs; \<close>
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name l}; \<close>
+
+thm l.exhaust l.map l.induct l.rec l.size
+
+ML \<open>
+val t_specs =
+  [((@{binding t}, [("'b", @{sort type})], NoSyn),
+   [(@{binding T}, [@{typ 'b}, Type (@{type_name l},
+       [Type (Sign.full_name @{theory} @{binding t}, [@{typ 'b}])])], NoSyn)])];
+\<close>
+
+setup \<open> snd o BNF_LFP_Compat.add_datatype BNF_LFP_Compat.Unfold_Nesting t_specs; \<close>
+
+ML \<open> get_descrs @{theory} (2, 2, 1) @{type_name t}; \<close>
+
+thm t.exhaust t.map t.induct t.rec t.size
+thm compat_t.induct compat_t.rec
+
+ML \<open>
+val ft_specs =
+  [((@{binding ft}, [("'a", @{sort type})], NoSyn),
+   [(@{binding FT0}, [], NoSyn),
+    (@{binding FT}, [@{typ 'a} --> Type (Sign.full_name @{theory} @{binding ft}, [@{typ 'a}])],
+     NoSyn)])];
+\<close>
+
+setup \<open> snd o BNF_LFP_Compat.add_datatype BNF_LFP_Compat.Unfold_Nesting ft_specs; \<close>
+
+ML \<open> get_descrs @{theory} (1, 1, 1) @{type_name ft}; \<close>
+
+thm ft.exhaust ft.induct ft.rec ft.size
+thm compat_ft.induct compat_ft.rec
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/DTree.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,90 @@
+(*  Title:      HOL/Datatype_Examples/Derivation_Trees/DTree.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Derivation trees with nonterminal internal nodes and terminal leaves.
+*)
+
+header {* Trees with Nonterminal Internal Nodes and Terminal Leaves *}
+
+theory DTree
+imports Prelim
+begin
+
+typedecl N
+typedecl T
+
+codatatype dtree = NNode (root: N) (ccont: "(T + dtree) fset")
+
+subsection{* Transporting the Characteristic Lemmas from @{text "fset"} to @{text "set"} *}
+
+definition "Node n as \<equiv> NNode n (the_inv fset as)"
+definition "cont \<equiv> fset o ccont"
+definition "unfold rt ct \<equiv> corec_dtree rt (the_inv fset o image (map_sum id Inr) o ct)"
+definition "corec rt ct \<equiv> corec_dtree rt (the_inv fset o ct)"
+
+lemma finite_cont[simp]: "finite (cont tr)"
+  unfolding cont_def comp_apply by (cases tr, clarsimp)
+
+lemma Node_root_cont[simp]:
+  "Node (root tr) (cont tr) = tr"
+  unfolding Node_def cont_def comp_apply
+  apply (rule trans[OF _ dtree.collapse])
+  apply (rule arg_cong2[OF refl the_inv_into_f_f[unfolded inj_on_def]])
+  apply (simp_all add: fset_inject)
+  done
+
+lemma dtree_simps[simp]:
+assumes "finite as" and "finite as'"
+shows "Node n as = Node n' as' \<longleftrightarrow> n = n' \<and> as = as'"
+using assms dtree.inject unfolding Node_def
+by (metis fset_to_fset)
+
+lemma dtree_cases[elim, case_names Node Choice]:
+assumes Node: "\<And> n as. \<lbrakk>finite as; tr = Node n as\<rbrakk> \<Longrightarrow> phi"
+shows phi
+apply(cases rule: dtree.exhaust[of tr])
+using Node unfolding Node_def
+by (metis Node Node_root_cont finite_cont)
+
+lemma dtree_sel_ctr[simp]:
+"root (Node n as) = n"
+"finite as \<Longrightarrow> cont (Node n as) = as"
+unfolding Node_def cont_def by auto
+
+lemmas root_Node = dtree_sel_ctr(1)
+lemmas cont_Node = dtree_sel_ctr(2)
+
+lemma dtree_cong:
+assumes "root tr = root tr'" and "cont tr = cont tr'"
+shows "tr = tr'"
+by (metis Node_root_cont assms)
+
+lemma rel_set_cont:
+"rel_set \<chi> (cont tr1) (cont tr2) = rel_fset \<chi> (ccont tr1) (ccont tr2)"
+unfolding cont_def comp_def rel_fset_fset ..
+
+lemma dtree_coinduct[elim, consumes 1, case_names Lift, induct pred: "HOL.eq"]:
+assumes phi: "\<phi> tr1 tr2" and
+Lift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
+                  root tr1 = root tr2 \<and> rel_set (rel_sum op = \<phi>) (cont tr1) (cont tr2)"
+shows "tr1 = tr2"
+using phi apply(elim dtree.coinduct)
+apply(rule Lift[unfolded rel_set_cont]) .
+
+lemma unfold:
+"root (unfold rt ct b) = rt b"
+"finite (ct b) \<Longrightarrow> cont (unfold rt ct b) = image (id \<oplus> unfold rt ct) (ct b)"
+using dtree.corec_sel[of rt "the_inv fset o image (map_sum id Inr) o ct" b] unfolding unfold_def
+apply blast
+unfolding cont_def comp_def
+by (simp add: case_sum_o_inj map_sum.compositionality image_image)
+
+lemma corec:
+"root (corec rt ct b) = rt b"
+"finite (ct b) \<Longrightarrow> cont (corec rt ct b) = image (id \<oplus> ([[id, corec rt ct]])) (ct b)"
+using dtree.corec_sel[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
+unfolding cont_def comp_def id_def
+by simp_all
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,1357 @@
+(*  Title:      HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Language of a grammar.
+*)
+
+header {* Language of a Grammar *}
+
+theory Gram_Lang
+imports DTree "~~/src/HOL/Library/Infinite_Set"
+begin
+
+
+(* We assume that the sets of terminals, and the left-hand sides of
+productions are finite and that the grammar has no unused nonterminals. *)
+consts P :: "(N \<times> (T + N) set) set"
+axiomatization where
+    finite_N: "finite (UNIV::N set)"
+and finite_in_P: "\<And> n tns. (n,tns) \<in> P \<longrightarrow> finite tns"
+and used: "\<And> n. \<exists> tns. (n,tns) \<in> P"
+
+
+subsection{* Tree Basics: frontier, interior, etc. *}
+
+
+(* Frontier *)
+
+inductive inFr :: "N set \<Rightarrow> dtree \<Rightarrow> T \<Rightarrow> bool" where
+Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr ns tr t"
+|
+Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inFr ns tr1 t\<rbrakk> \<Longrightarrow> inFr ns tr t"
+
+definition "Fr ns tr \<equiv> {t. inFr ns tr t}"
+
+lemma inFr_root_in: "inFr ns tr t \<Longrightarrow> root tr \<in> ns"
+by (metis inFr.simps)
+
+lemma inFr_mono:
+assumes "inFr ns tr t" and "ns \<subseteq> ns'"
+shows "inFr ns' tr t"
+using assms apply(induct arbitrary: ns' rule: inFr.induct)
+using Base Ind by (metis inFr.simps set_mp)+
+
+lemma inFr_Ind_minus:
+assumes "inFr ns1 tr1 t" and "Inr tr1 \<in> cont tr"
+shows "inFr (insert (root tr) ns1) tr t"
+using assms apply(induct rule: inFr.induct)
+  apply (metis inFr.simps insert_iff)
+  by (metis inFr.simps inFr_mono insertI1 subset_insertI)
+
+(* alternative definition *)
+inductive inFr2 :: "N set \<Rightarrow> dtree \<Rightarrow> T \<Rightarrow> bool" where
+Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr2 ns tr t"
+|
+Ind: "\<lbrakk>Inr tr1 \<in> cont tr; inFr2 ns1 tr1 t\<rbrakk>
+      \<Longrightarrow> inFr2 (insert (root tr) ns1) tr t"
+
+lemma inFr2_root_in: "inFr2 ns tr t \<Longrightarrow> root tr \<in> ns"
+apply(induct rule: inFr2.induct) by auto
+
+lemma inFr2_mono:
+assumes "inFr2 ns tr t" and "ns \<subseteq> ns'"
+shows "inFr2 ns' tr t"
+using assms apply(induct arbitrary: ns' rule: inFr2.induct)
+using Base Ind
+apply (metis subsetD) by (metis inFr2.simps insert_absorb insert_subset)
+
+lemma inFr2_Ind:
+assumes "inFr2 ns tr1 t" and "root tr \<in> ns" and "Inr tr1 \<in> cont tr"
+shows "inFr2 ns tr t"
+using assms apply(induct rule: inFr2.induct)
+  apply (metis inFr2.simps insert_absorb)
+  by (metis inFr2.simps insert_absorb)
+
+lemma inFr_inFr2:
+"inFr = inFr2"
+apply (rule ext)+  apply(safe)
+  apply(erule inFr.induct)
+    apply (metis (lifting) inFr2.Base)
+    apply (metis (lifting) inFr2_Ind)
+  apply(erule inFr2.induct)
+    apply (metis (lifting) inFr.Base)
+    apply (metis (lifting) inFr_Ind_minus)
+done
+
+lemma not_root_inFr:
+assumes "root tr \<notin> ns"
+shows "\<not> inFr ns tr t"
+by (metis assms inFr_root_in)
+
+lemma not_root_Fr:
+assumes "root tr \<notin> ns"
+shows "Fr ns tr = {}"
+using not_root_inFr[OF assms] unfolding Fr_def by auto
+
+
+(* Interior *)
+
+inductive inItr :: "N set \<Rightarrow> dtree \<Rightarrow> N \<Rightarrow> bool" where
+Base: "root tr \<in> ns \<Longrightarrow> inItr ns tr (root tr)"
+|
+Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inItr ns tr1 n\<rbrakk> \<Longrightarrow> inItr ns tr n"
+
+definition "Itr ns tr \<equiv> {n. inItr ns tr n}"
+
+lemma inItr_root_in: "inItr ns tr n \<Longrightarrow> root tr \<in> ns"
+by (metis inItr.simps)
+
+lemma inItr_mono:
+assumes "inItr ns tr n" and "ns \<subseteq> ns'"
+shows "inItr ns' tr n"
+using assms apply(induct arbitrary: ns' rule: inItr.induct)
+using Base Ind by (metis inItr.simps set_mp)+
+
+
+(* The subtree relation *)
+
+inductive subtr where
+Refl: "root tr \<in> ns \<Longrightarrow> subtr ns tr tr"
+|
+Step: "\<lbrakk>root tr3 \<in> ns; subtr ns tr1 tr2; Inr tr2 \<in> cont tr3\<rbrakk> \<Longrightarrow> subtr ns tr1 tr3"
+
+lemma subtr_rootL_in:
+assumes "subtr ns tr1 tr2"
+shows "root tr1 \<in> ns"
+using assms apply(induct rule: subtr.induct) by auto
+
+lemma subtr_rootR_in:
+assumes "subtr ns tr1 tr2"
+shows "root tr2 \<in> ns"
+using assms apply(induct rule: subtr.induct) by auto
+
+lemmas subtr_roots_in = subtr_rootL_in subtr_rootR_in
+
+lemma subtr_mono:
+assumes "subtr ns tr1 tr2" and "ns \<subseteq> ns'"
+shows "subtr ns' tr1 tr2"
+using assms apply(induct arbitrary: ns' rule: subtr.induct)
+using Refl Step by (metis subtr.simps set_mp)+
+
+lemma subtr_trans_Un:
+assumes "subtr ns12 tr1 tr2" and "subtr ns23 tr2 tr3"
+shows "subtr (ns12 \<union> ns23) tr1 tr3"
+proof-
+  have "subtr ns23 tr2 tr3  \<Longrightarrow>
+        (\<forall> ns12 tr1. subtr ns12 tr1 tr2 \<longrightarrow> subtr (ns12 \<union> ns23) tr1 tr3)"
+  apply(induct  rule: subtr.induct, safe)
+    apply (metis subtr_mono sup_commute sup_ge2)
+    by (metis (lifting) Step UnI2)
+  thus ?thesis using assms by auto
+qed
+
+lemma subtr_trans:
+assumes "subtr ns tr1 tr2" and "subtr ns tr2 tr3"
+shows "subtr ns tr1 tr3"
+using subtr_trans_Un[OF assms] by simp
+
+lemma subtr_StepL:
+assumes r: "root tr1 \<in> ns" and tr12: "Inr tr1 \<in> cont tr2" and s: "subtr ns tr2 tr3"
+shows "subtr ns tr1 tr3"
+apply(rule subtr_trans[OF _ s])
+apply(rule Step[of tr2 ns tr1 tr1])
+apply(rule subtr_rootL_in[OF s])
+apply(rule Refl[OF r])
+apply(rule tr12)
+done
+
+(* alternative definition: *)
+inductive subtr2 where
+Refl: "root tr \<in> ns \<Longrightarrow> subtr2 ns tr tr"
+|
+Step: "\<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr2 ns tr2 tr3\<rbrakk> \<Longrightarrow> subtr2 ns tr1 tr3"
+
+lemma subtr2_rootL_in:
+assumes "subtr2 ns tr1 tr2"
+shows "root tr1 \<in> ns"
+using assms apply(induct rule: subtr2.induct) by auto
+
+lemma subtr2_rootR_in:
+assumes "subtr2 ns tr1 tr2"
+shows "root tr2 \<in> ns"
+using assms apply(induct rule: subtr2.induct) by auto
+
+lemmas subtr2_roots_in = subtr2_rootL_in subtr2_rootR_in
+
+lemma subtr2_mono:
+assumes "subtr2 ns tr1 tr2" and "ns \<subseteq> ns'"
+shows "subtr2 ns' tr1 tr2"
+using assms apply(induct arbitrary: ns' rule: subtr2.induct)
+using Refl Step by (metis subtr2.simps set_mp)+
+
+lemma subtr2_trans_Un:
+assumes "subtr2 ns12 tr1 tr2" and "subtr2 ns23 tr2 tr3"
+shows "subtr2 (ns12 \<union> ns23) tr1 tr3"
+proof-
+  have "subtr2 ns12 tr1 tr2  \<Longrightarrow>
+        (\<forall> ns23 tr3. subtr2 ns23 tr2 tr3 \<longrightarrow> subtr2 (ns12 \<union> ns23) tr1 tr3)"
+  apply(induct  rule: subtr2.induct, safe)
+    apply (metis subtr2_mono sup_commute sup_ge2)
+    by (metis Un_iff subtr2.simps)
+  thus ?thesis using assms by auto
+qed
+
+lemma subtr2_trans:
+assumes "subtr2 ns tr1 tr2" and "subtr2 ns tr2 tr3"
+shows "subtr2 ns tr1 tr3"
+using subtr2_trans_Un[OF assms] by simp
+
+lemma subtr2_StepR:
+assumes r: "root tr3 \<in> ns" and tr23: "Inr tr2 \<in> cont tr3" and s: "subtr2 ns tr1 tr2"
+shows "subtr2 ns tr1 tr3"
+apply(rule subtr2_trans[OF s])
+apply(rule Step[of _ _ tr3])
+apply(rule subtr2_rootR_in[OF s])
+apply(rule tr23)
+apply(rule Refl[OF r])
+done
+
+lemma subtr_subtr2:
+"subtr = subtr2"
+apply (rule ext)+  apply(safe)
+  apply(erule subtr.induct)
+    apply (metis (lifting) subtr2.Refl)
+    apply (metis (lifting) subtr2_StepR)
+  apply(erule subtr2.induct)
+    apply (metis (lifting) subtr.Refl)
+    apply (metis (lifting) subtr_StepL)
+done
+
+lemma subtr_inductL[consumes 1, case_names Refl Step]:
+assumes s: "subtr ns tr1 tr2" and Refl: "\<And>ns tr. \<phi> ns tr tr"
+and Step:
+"\<And>ns tr1 tr2 tr3.
+   \<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr ns tr2 tr3; \<phi> ns tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> ns tr1 tr3"
+shows "\<phi> ns tr1 tr2"
+using s unfolding subtr_subtr2 apply(rule subtr2.induct)
+using Refl Step unfolding subtr_subtr2 by auto
+
+lemma subtr_UNIV_inductL[consumes 1, case_names Refl Step]:
+assumes s: "subtr UNIV tr1 tr2" and Refl: "\<And>tr. \<phi> tr tr"
+and Step:
+"\<And>tr1 tr2 tr3.
+   \<lbrakk>Inr tr1 \<in> cont tr2; subtr UNIV tr2 tr3; \<phi> tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> tr1 tr3"
+shows "\<phi> tr1 tr2"
+using s apply(induct rule: subtr_inductL)
+apply(rule Refl) using Step subtr_mono by (metis subset_UNIV)
+
+(* Subtree versus frontier: *)
+lemma subtr_inFr:
+assumes "inFr ns tr t" and "subtr ns tr tr1"
+shows "inFr ns tr1 t"
+proof-
+  have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inFr ns tr t \<longrightarrow> inFr ns tr1 t)"
+  apply(induct rule: subtr.induct, safe) by (metis inFr.Ind)
+  thus ?thesis using assms by auto
+qed
+
+corollary Fr_subtr:
+"Fr ns tr = \<Union> {Fr ns tr' | tr'. subtr ns tr' tr}"
+unfolding Fr_def proof safe
+  fix t assume t: "inFr ns tr t"  hence "root tr \<in> ns" by (rule inFr_root_in)
+  thus "t \<in> \<Union>{{t. inFr ns tr' t} |tr'. subtr ns tr' tr}"
+  apply(intro UnionI[of "{t. inFr ns tr t}" _ t]) using t subtr.Refl by auto
+qed(metis subtr_inFr)
+
+lemma inFr_subtr:
+assumes "inFr ns tr t"
+shows "\<exists> tr'. subtr ns tr' tr \<and> Inl t \<in> cont tr'"
+using assms apply(induct rule: inFr.induct) apply safe
+  apply (metis subtr.Refl)
+  by (metis (lifting) subtr.Step)
+
+corollary Fr_subtr_cont:
+"Fr ns tr = \<Union> {Inl -` cont tr' | tr'. subtr ns tr' tr}"
+unfolding Fr_def
+apply safe
+apply (frule inFr_subtr)
+apply auto
+by (metis inFr.Base subtr_inFr subtr_rootL_in)
+
+(* Subtree versus interior: *)
+lemma subtr_inItr:
+assumes "inItr ns tr n" and "subtr ns tr tr1"
+shows "inItr ns tr1 n"
+proof-
+  have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inItr ns tr n \<longrightarrow> inItr ns tr1 n)"
+  apply(induct rule: subtr.induct, safe) by (metis inItr.Ind)
+  thus ?thesis using assms by auto
+qed
+
+corollary Itr_subtr:
+"Itr ns tr = \<Union> {Itr ns tr' | tr'. subtr ns tr' tr}"
+unfolding Itr_def apply safe
+apply (metis (lifting, mono_tags) UnionI inItr_root_in mem_Collect_eq subtr.Refl)
+by (metis subtr_inItr)
+
+lemma inItr_subtr:
+assumes "inItr ns tr n"
+shows "\<exists> tr'. subtr ns tr' tr \<and> root tr' = n"
+using assms apply(induct rule: inItr.induct) apply safe
+  apply (metis subtr.Refl)
+  by (metis (lifting) subtr.Step)
+
+corollary Itr_subtr_cont:
+"Itr ns tr = {root tr' | tr'. subtr ns tr' tr}"
+unfolding Itr_def apply safe
+  apply (metis (lifting, mono_tags) inItr_subtr)
+  by (metis inItr.Base subtr_inItr subtr_rootL_in)
+
+
+subsection{* The Immediate Subtree Function *}
+
+(* production of: *)
+abbreviation "prodOf tr \<equiv> (id \<oplus> root) ` (cont tr)"
+(* subtree of: *)
+definition "subtrOf tr n \<equiv> SOME tr'. Inr tr' \<in> cont tr \<and> root tr' = n"
+
+lemma subtrOf:
+assumes n: "Inr n \<in> prodOf tr"
+shows "Inr (subtrOf tr n) \<in> cont tr \<and> root (subtrOf tr n) = n"
+proof-
+  obtain tr' where "Inr tr' \<in> cont tr \<and> root tr' = n"
+  using n unfolding image_def by (metis (lifting) Inr_oplus_elim assms)
+  thus ?thesis unfolding subtrOf_def by(rule someI)
+qed
+
+lemmas Inr_subtrOf = subtrOf[THEN conjunct1]
+lemmas root_subtrOf[simp] = subtrOf[THEN conjunct2]
+
+lemma Inl_prodOf: "Inl -` (prodOf tr) = Inl -` (cont tr)"
+proof safe
+  fix t ttr assume "Inl t = (id \<oplus> root) ttr" and "ttr \<in> cont tr"
+  thus "t \<in> Inl -` cont tr" by(cases ttr, auto)
+next
+  fix t assume "Inl t \<in> cont tr" thus "t \<in> Inl -` prodOf tr"
+  by (metis (lifting) id_def image_iff map_sum.simps(1) vimageI2)
+qed
+
+lemma root_prodOf:
+assumes "Inr tr' \<in> cont tr"
+shows "Inr (root tr') \<in> prodOf tr"
+by (metis (lifting) assms image_iff map_sum.simps(2))
+
+
+subsection{* Well-Formed Derivation Trees *}
+
+hide_const wf
+
+coinductive wf where
+dtree: "\<lbrakk>(root tr, (id \<oplus> root) ` (cont tr)) \<in> P; inj_on root (Inr -` cont tr);
+        \<And> tr'. tr' \<in> Inr -` (cont tr) \<Longrightarrow> wf tr'\<rbrakk> \<Longrightarrow> wf tr"
+
+(* destruction rules: *)
+lemma wf_P:
+assumes "wf tr"
+shows "(root tr, (id \<oplus> root) ` (cont tr)) \<in> P"
+using assms wf.simps[of tr] by auto
+
+lemma wf_inj_on:
+assumes "wf tr"
+shows "inj_on root (Inr -` cont tr)"
+using assms wf.simps[of tr] by auto
+
+lemma wf_inj[simp]:
+assumes "wf tr" and "Inr tr1 \<in> cont tr" and "Inr tr2 \<in> cont tr"
+shows "root tr1 = root tr2 \<longleftrightarrow> tr1 = tr2"
+using assms wf_inj_on unfolding inj_on_def by auto
+
+lemma wf_cont:
+assumes "wf tr" and "Inr tr' \<in> cont tr"
+shows "wf tr'"
+using assms wf.simps[of tr] by auto
+
+
+(* coinduction:*)
+lemma wf_coind[elim, consumes 1, case_names Hyp]:
+assumes phi: "\<phi> tr"
+and Hyp:
+"\<And> tr. \<phi> tr \<Longrightarrow>
+       (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
+       inj_on root (Inr -` cont tr) \<and>
+       (\<forall> tr' \<in> Inr -` (cont tr). \<phi> tr' \<or> wf tr')"
+shows "wf tr"
+apply(rule wf.coinduct[of \<phi> tr, OF phi])
+using Hyp by blast
+
+lemma wf_raw_coind[elim, consumes 1, case_names Hyp]:
+assumes phi: "\<phi> tr"
+and Hyp:
+"\<And> tr. \<phi> tr \<Longrightarrow>
+       (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
+       inj_on root (Inr -` cont tr) \<and>
+       (\<forall> tr' \<in> Inr -` (cont tr). \<phi> tr')"
+shows "wf tr"
+using phi apply(induct rule: wf_coind)
+using Hyp by (metis (mono_tags))
+
+lemma wf_subtr_inj_on:
+assumes d: "wf tr1" and s: "subtr ns tr tr1"
+shows "inj_on root (Inr -` cont tr)"
+using s d apply(induct rule: subtr.induct)
+apply (metis (lifting) wf_inj_on) by (metis wf_cont)
+
+lemma wf_subtr_P:
+assumes d: "wf tr1" and s: "subtr ns tr tr1"
+shows "(root tr, (id \<oplus> root) ` cont tr) \<in> P"
+using s d apply(induct rule: subtr.induct)
+apply (metis (lifting) wf_P) by (metis wf_cont)
+
+lemma subtrOf_root[simp]:
+assumes tr: "wf tr" and cont: "Inr tr' \<in> cont tr"
+shows "subtrOf tr (root tr') = tr'"
+proof-
+  have 0: "Inr (subtrOf tr (root tr')) \<in> cont tr" using Inr_subtrOf
+  by (metis (lifting) cont root_prodOf)
+  have "root (subtrOf tr (root tr')) = root tr'"
+  using root_subtrOf by (metis (lifting) cont root_prodOf)
+  thus ?thesis unfolding wf_inj[OF tr 0 cont] .
+qed
+
+lemma surj_subtrOf:
+assumes "wf tr" and 0: "Inr tr' \<in> cont tr"
+shows "\<exists> n. Inr n \<in> prodOf tr \<and> subtrOf tr n = tr'"
+apply(rule exI[of _ "root tr'"])
+using root_prodOf[OF 0] subtrOf_root[OF assms] by simp
+
+lemma wf_subtr:
+assumes "wf tr1" and "subtr ns tr tr1"
+shows "wf tr"
+proof-
+  have "(\<exists> ns tr1. wf tr1 \<and> subtr ns tr tr1) \<Longrightarrow> wf tr"
+  proof (induct rule: wf_raw_coind)
+    case (Hyp tr)
+    then obtain ns tr1 where tr1: "wf tr1" and tr_tr1: "subtr ns tr tr1" by auto
+    show ?case proof safe
+      show "(root tr, (id \<oplus> root) ` cont tr) \<in> P" using wf_subtr_P[OF tr1 tr_tr1] .
+    next
+      show "inj_on root (Inr -` cont tr)" using wf_subtr_inj_on[OF tr1 tr_tr1] .
+    next
+      fix tr' assume tr': "Inr tr' \<in> cont tr"
+      have tr_tr1: "subtr (ns \<union> {root tr'}) tr tr1" using subtr_mono[OF tr_tr1] by auto
+      have "subtr (ns \<union> {root tr'}) tr' tr1" using subtr_StepL[OF _ tr' tr_tr1] by auto
+      thus "\<exists>ns' tr1. wf tr1 \<and> subtr ns' tr' tr1" using tr1 by blast
+    qed
+  qed
+  thus ?thesis using assms by auto
+qed
+
+
+subsection{* Default Trees *}
+
+(* Pick a left-hand side of a production for each nonterminal *)
+definition S where "S n \<equiv> SOME tns. (n,tns) \<in> P"
+
+lemma S_P: "(n, S n) \<in> P"
+using used unfolding S_def by(rule someI_ex)
+
+lemma finite_S: "finite (S n)"
+using S_P finite_in_P by auto
+
+
+(* The default tree of a nonterminal *)
+definition deftr :: "N \<Rightarrow> dtree" where
+"deftr \<equiv> unfold id S"
+
+lemma deftr_simps[simp]:
+"root (deftr n) = n"
+"cont (deftr n) = image (id \<oplus> deftr) (S n)"
+using unfold(1)[of id S n] unfold(2)[of S n id, OF finite_S]
+unfolding deftr_def by simp_all
+
+lemmas root_deftr = deftr_simps(1)
+lemmas cont_deftr = deftr_simps(2)
+
+lemma root_o_deftr[simp]: "root o deftr = id"
+by (rule ext, auto)
+
+lemma wf_deftr: "wf (deftr n)"
+proof-
+  {fix tr assume "\<exists> n. tr = deftr n" hence "wf tr"
+   apply(induct rule: wf_raw_coind) apply safe
+   unfolding deftr_simps image_comp map_sum.comp id_comp
+   root_o_deftr map_sum.id image_id id_apply apply(rule S_P)
+   unfolding inj_on_def by auto
+  }
+  thus ?thesis by auto
+qed
+
+
+subsection{* Hereditary Substitution *}
+
+(* Auxiliary concept: The root-ommiting frontier: *)
+definition "inFrr ns tr t \<equiv> \<exists> tr'. Inr tr' \<in> cont tr \<and> inFr ns tr' t"
+definition "Frr ns tr \<equiv> {t. \<exists> tr'. Inr tr' \<in> cont tr \<and> t \<in> Fr ns tr'}"
+
+context
+fixes tr0 :: dtree
+begin
+
+definition "hsubst_r tr \<equiv> root tr"
+definition "hsubst_c tr \<equiv> if root tr = root tr0 then cont tr0 else cont tr"
+
+(* Hereditary substitution: *)
+definition hsubst :: "dtree \<Rightarrow> dtree" where
+"hsubst \<equiv> unfold hsubst_r hsubst_c"
+
+lemma finite_hsubst_c: "finite (hsubst_c n)"
+unfolding hsubst_c_def by (metis (full_types) finite_cont)
+
+lemma root_hsubst[simp]: "root (hsubst tr) = root tr"
+using unfold(1)[of hsubst_r hsubst_c tr] unfolding hsubst_def hsubst_r_def by simp
+
+lemma root_o_subst[simp]: "root o hsubst = root"
+unfolding comp_def root_hsubst ..
+
+lemma cont_hsubst_eq[simp]:
+assumes "root tr = root tr0"
+shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr0)"
+apply(subst id_comp[symmetric, of id]) unfolding id_comp
+using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
+unfolding hsubst_def hsubst_c_def using assms by simp
+
+lemma hsubst_eq:
+assumes "root tr = root tr0"
+shows "hsubst tr = hsubst tr0"
+apply(rule dtree_cong) using assms cont_hsubst_eq by auto
+
+lemma cont_hsubst_neq[simp]:
+assumes "root tr \<noteq> root tr0"
+shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr)"
+apply(subst id_comp[symmetric, of id]) unfolding id_comp
+using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
+unfolding hsubst_def hsubst_c_def using assms by simp
+
+lemma Inl_cont_hsubst_eq[simp]:
+assumes "root tr = root tr0"
+shows "Inl -` cont (hsubst tr) = Inl -` (cont tr0)"
+unfolding cont_hsubst_eq[OF assms] by simp
+
+lemma Inr_cont_hsubst_eq[simp]:
+assumes "root tr = root tr0"
+shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr0"
+unfolding cont_hsubst_eq[OF assms] by simp
+
+lemma Inl_cont_hsubst_neq[simp]:
+assumes "root tr \<noteq> root tr0"
+shows "Inl -` cont (hsubst tr) = Inl -` (cont tr)"
+unfolding cont_hsubst_neq[OF assms] by simp
+
+lemma Inr_cont_hsubst_neq[simp]:
+assumes "root tr \<noteq> root tr0"
+shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr"
+unfolding cont_hsubst_neq[OF assms] by simp
+
+lemma wf_hsubst:
+assumes tr0: "wf tr0" and tr: "wf tr"
+shows "wf (hsubst tr)"
+proof-
+  {fix tr1 have "(\<exists> tr. wf tr \<and> tr1 = hsubst tr) \<Longrightarrow> wf tr1"
+   proof (induct rule: wf_raw_coind)
+     case (Hyp tr1) then obtain tr
+     where dtr: "wf tr" and tr1: "tr1 = hsubst tr" by auto
+     show ?case unfolding tr1 proof safe
+       show "(root (hsubst tr), prodOf (hsubst tr)) \<in> P"
+       unfolding tr1 apply(cases "root tr = root tr0")
+       using  wf_P[OF dtr] wf_P[OF tr0]
+       by (auto simp add: image_comp map_sum.comp)
+       show "inj_on root (Inr -` cont (hsubst tr))"
+       apply(cases "root tr = root tr0") using wf_inj_on[OF dtr] wf_inj_on[OF tr0]
+       unfolding inj_on_def by (auto, blast)
+       fix tr' assume "Inr tr' \<in> cont (hsubst tr)"
+       thus "\<exists>tra. wf tra \<and> tr' = hsubst tra"
+       apply(cases "root tr = root tr0", simp_all)
+         apply (metis wf_cont tr0)
+         by (metis dtr wf_cont)
+     qed
+   qed
+  }
+  thus ?thesis using assms by blast
+qed
+
+lemma Frr: "Frr ns tr = {t. inFrr ns tr t}"
+unfolding inFrr_def Frr_def Fr_def by auto
+
+lemma inFr_hsubst_imp:
+assumes "inFr ns (hsubst tr) t"
+shows "t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
+       inFr (ns - {root tr0}) tr t"
+proof-
+  {fix tr1
+   have "inFr ns tr1 t \<Longrightarrow>
+   (\<And> tr. tr1 = hsubst tr \<Longrightarrow> (t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
+                              inFr (ns - {root tr0}) tr t))"
+   proof(induct rule: inFr.induct)
+     case (Base tr1 ns t tr)
+     hence rtr: "root tr1 \<in> ns" and t_tr1: "Inl t \<in> cont tr1" and tr1: "tr1 = hsubst tr"
+     by auto
+     show ?case
+     proof(cases "root tr1 = root tr0")
+       case True
+       hence "t \<in> Inl -` (cont tr0)" using t_tr1 unfolding tr1 by auto
+       thus ?thesis by simp
+     next
+       case False
+       hence "inFr (ns - {root tr0}) tr t" using t_tr1 unfolding tr1 apply simp
+       by (metis Base.prems Diff_iff root_hsubst inFr.Base rtr singletonE)
+       thus ?thesis by simp
+     qed
+   next
+     case (Ind tr1 ns tr1' t) note IH = Ind(4)
+     have rtr1: "root tr1 \<in> ns" and tr1'_tr1: "Inr tr1' \<in> cont tr1"
+     and t_tr1': "inFr ns tr1' t" and tr1: "tr1 = hsubst tr" using Ind by auto
+     have rtr1: "root tr1 = root tr" unfolding tr1 by simp
+     show ?case
+     proof(cases "root tr1 = root tr0")
+       case True
+       then obtain tr' where tr'_tr0: "Inr tr' \<in> cont tr0" and tr1': "tr1' = hsubst tr'"
+       using tr1'_tr1 unfolding tr1 by auto
+       show ?thesis using IH[OF tr1'] proof (elim disjE)
+         assume "inFr (ns - {root tr0}) tr' t"
+         thus ?thesis using tr'_tr0 unfolding inFrr_def by auto
+       qed auto
+     next
+       case False
+       then obtain tr' where tr'_tr: "Inr tr' \<in> cont tr" and tr1': "tr1' = hsubst tr'"
+       using tr1'_tr1 unfolding tr1 by auto
+       show ?thesis using IH[OF tr1'] proof (elim disjE)
+         assume "inFr (ns - {root tr0}) tr' t"
+         thus ?thesis using tr'_tr unfolding inFrr_def
+         by (metis Diff_iff False Ind(1) empty_iff inFr2_Ind inFr_inFr2 insert_iff rtr1)
+       qed auto
+     qed
+   qed
+  }
+  thus ?thesis using assms by auto
+qed
+
+lemma inFr_hsubst_notin:
+assumes "inFr ns tr t" and "root tr0 \<notin> ns"
+shows "inFr ns (hsubst tr) t"
+using assms apply(induct rule: inFr.induct)
+apply (metis Inl_cont_hsubst_neq inFr2.Base inFr_inFr2 root_hsubst vimageD vimageI2)
+by (metis (lifting) Inr_cont_hsubst_neq inFr.Ind rev_image_eqI root_hsubst vimageD vimageI2)
+
+lemma inFr_hsubst_minus:
+assumes "inFr (ns - {root tr0}) tr t"
+shows "inFr ns (hsubst tr) t"
+proof-
+  have 1: "inFr (ns - {root tr0}) (hsubst tr) t"
+  using inFr_hsubst_notin[OF assms] by simp
+  show ?thesis using inFr_mono[OF 1] by auto
+qed
+
+lemma inFr_self_hsubst:
+assumes "root tr0 \<in> ns"
+shows
+"inFr ns (hsubst tr0) t \<longleftrightarrow>
+ t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t"
+(is "?A \<longleftrightarrow> ?B \<or> ?C")
+apply(intro iffI)
+apply (metis inFr_hsubst_imp Diff_iff inFr_root_in insertI1) proof(elim disjE)
+  assume ?B thus ?A apply(intro inFr.Base) using assms by auto
+next
+  assume ?C then obtain tr where
+  tr_tr0: "Inr tr \<in> cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t"
+  unfolding inFrr_def by auto
+  def tr1 \<equiv> "hsubst tr"
+  have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto
+  have "Inr tr1 \<in> cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto
+  thus ?A using 1 inFr.Ind assms by (metis root_hsubst)
+qed
+
+lemma Fr_self_hsubst:
+assumes "root tr0 \<in> ns"
+shows "Fr ns (hsubst tr0) = Inl -` (cont tr0) \<union> Frr (ns - {root tr0}) tr0"
+using inFr_self_hsubst[OF assms] unfolding Frr Fr_def by auto
+
+end (* context *)
+
+
+subsection{* Regular Trees *}
+
+definition "reg f tr \<equiv> \<forall> tr'. subtr UNIV tr' tr \<longrightarrow> tr' = f (root tr')"
+definition "regular tr \<equiv> \<exists> f. reg f tr"
+
+lemma reg_def2: "reg f tr \<longleftrightarrow> (\<forall> ns tr'. subtr ns tr' tr \<longrightarrow> tr' = f (root tr'))"
+unfolding reg_def using subtr_mono by (metis subset_UNIV)
+
+lemma regular_def2: "regular tr \<longleftrightarrow> (\<exists> f. reg f tr \<and> (\<forall> n. root (f n) = n))"
+unfolding regular_def proof safe
+  fix f assume f: "reg f tr"
+  def g \<equiv> "\<lambda> n. if inItr UNIV tr n then f n else deftr n"
+  show "\<exists>g. reg g tr \<and> (\<forall>n. root (g n) = n)"
+  apply(rule exI[of _ g])
+  using f deftr_simps(1) unfolding g_def reg_def apply safe
+    apply (metis (lifting) inItr.Base subtr_inItr subtr_rootL_in)
+    by (metis (full_types) inItr_subtr)
+qed auto
+
+lemma reg_root:
+assumes "reg f tr"
+shows "f (root tr) = tr"
+using assms unfolding reg_def
+by (metis (lifting) iso_tuple_UNIV_I subtr.Refl)
+
+
+lemma reg_Inr_cont:
+assumes "reg f tr" and "Inr tr' \<in> cont tr"
+shows "reg f tr'"
+by (metis (lifting) assms iso_tuple_UNIV_I reg_def subtr.Step)
+
+lemma reg_subtr:
+assumes "reg f tr" and "subtr ns tr' tr"
+shows "reg f tr'"
+using assms unfolding reg_def using subtr_trans[of UNIV tr] UNIV_I
+by (metis UNIV_eq_I UnCI Un_upper1 iso_tuple_UNIV_I subtr_mono subtr_trans)
+
+lemma regular_subtr:
+assumes r: "regular tr" and s: "subtr ns tr' tr"
+shows "regular tr'"
+using r reg_subtr[OF _ s] unfolding regular_def by auto
+
+lemma subtr_deftr:
+assumes "subtr ns tr' (deftr n)"
+shows "tr' = deftr (root tr')"
+proof-
+  {fix tr have "subtr ns tr' tr \<Longrightarrow> (\<forall> n. tr = deftr n \<longrightarrow> tr' = deftr (root tr'))"
+   apply (induct rule: subtr.induct)
+   proof(metis (lifting) deftr_simps(1), safe)
+     fix tr3 ns tr1 tr2 n
+     assume 1: "root (deftr n) \<in> ns" and 2: "subtr ns tr1 tr2"
+     and IH: "\<forall>n. tr2 = deftr n \<longrightarrow> tr1 = deftr (root tr1)"
+     and 3: "Inr tr2 \<in> cont (deftr n)"
+     have "tr2 \<in> deftr ` UNIV"
+     using 3 unfolding deftr_simps image_def
+     by (metis (lifting, full_types) 3 CollectI Inr_oplus_iff cont_deftr
+         iso_tuple_UNIV_I)
+     then obtain n where "tr2 = deftr n" by auto
+     thus "tr1 = deftr (root tr1)" using IH by auto
+   qed
+  }
+  thus ?thesis using assms by auto
+qed
+
+lemma reg_deftr: "reg deftr (deftr n)"
+unfolding reg_def using subtr_deftr by auto
+
+lemma wf_subtrOf_Union:
+assumes "wf tr"
+shows "\<Union>{K tr' |tr'. Inr tr' \<in> cont tr} =
+       \<Union>{K (subtrOf tr n) |n. Inr n \<in> prodOf tr}"
+unfolding Union_eq Bex_def mem_Collect_eq proof safe
+  fix x xa tr'
+  assume x: "x \<in> K tr'" and tr'_tr: "Inr tr' \<in> cont tr"
+  show "\<exists>X. (\<exists>n. X = K (subtrOf tr n) \<and> Inr n \<in> prodOf tr) \<and> x \<in> X"
+  apply(rule exI[of _ "K (subtrOf tr (root tr'))"]) apply(intro conjI)
+    apply(rule exI[of _ "root tr'"]) apply (metis (lifting) root_prodOf tr'_tr)
+    by (metis (lifting) assms subtrOf_root tr'_tr x)
+next
+  fix x X n ttr
+  assume x: "x \<in> K (subtrOf tr n)" and n: "Inr n = (id \<oplus> root) ttr" and ttr: "ttr \<in> cont tr"
+  show "\<exists>X. (\<exists>tr'. X = K tr' \<and> Inr tr' \<in> cont tr) \<and> x \<in> X"
+  apply(rule exI[of _ "K (subtrOf tr n)"]) apply(intro conjI)
+    apply(rule exI[of _ "subtrOf tr n"]) apply (metis imageI n subtrOf ttr)
+    using x .
+qed
+
+
+
+
+subsection {* Paths in a Regular Tree *}
+
+inductive path :: "(N \<Rightarrow> dtree) \<Rightarrow> N list \<Rightarrow> bool" for f where
+Base: "path f [n]"
+|
+Ind: "\<lbrakk>path f (n1 # nl); Inr (f n1) \<in> cont (f n)\<rbrakk>
+      \<Longrightarrow> path f (n # n1 # nl)"
+
+lemma path_NE:
+assumes "path f nl"
+shows "nl \<noteq> Nil"
+using assms apply(induct rule: path.induct) by auto
+
+lemma path_post:
+assumes f: "path f (n # nl)" and nl: "nl \<noteq> []"
+shows "path f nl"
+proof-
+  obtain n1 nl1 where nl: "nl = n1 # nl1" using nl by (cases nl, auto)
+  show ?thesis using assms unfolding nl using path.simps by (metis (lifting) list.inject)
+qed
+
+lemma path_post_concat:
+assumes "path f (nl1 @ nl2)" and "nl2 \<noteq> Nil"
+shows "path f nl2"
+using assms apply (induct nl1)
+apply (metis append_Nil) by (metis Nil_is_append_conv append_Cons path_post)
+
+lemma path_concat:
+assumes "path f nl1" and "path f ((last nl1) # nl2)"
+shows "path f (nl1 @ nl2)"
+using assms apply(induct rule: path.induct) apply simp
+by (metis append_Cons last.simps list.simps(3) path.Ind)
+
+lemma path_distinct:
+assumes "path f nl"
+shows "\<exists> nl'. path f nl' \<and> hd nl' = hd nl \<and> last nl' = last nl \<and>
+              set nl' \<subseteq> set nl \<and> distinct nl'"
+using assms proof(induct rule: length_induct)
+  case (1 nl)  hence p_nl: "path f nl" by simp
+  then obtain n nl1 where nl: "nl = n # nl1" by (metis list.exhaust path_NE)
+  show ?case
+  proof(cases nl1)
+    case Nil
+    show ?thesis apply(rule exI[of _ nl]) using path.Base unfolding nl Nil by simp
+  next
+    case (Cons n1 nl2)
+    hence p1: "path f nl1" by (metis list.simps(3) nl p_nl path_post)
+    show ?thesis
+    proof(cases "n \<in> set nl1")
+      case False
+      obtain nl1' where p1': "path f nl1'" and hd_nl1': "hd nl1' = hd nl1" and
+      l_nl1': "last nl1' = last nl1" and d_nl1': "distinct nl1'"
+      and s_nl1': "set nl1' \<subseteq> set nl1"
+      using 1(1)[THEN allE[of _ nl1]] p1 unfolding nl by auto
+      obtain nl2' where nl1': "nl1' = n1 # nl2'" using path_NE[OF p1'] hd_nl1'
+      unfolding Cons by(cases nl1', auto)
+      show ?thesis apply(intro exI[of _ "n # nl1'"]) unfolding nl proof safe
+        show "path f (n # nl1')" unfolding nl1'
+        apply(rule path.Ind, metis nl1' p1')
+        by (metis (lifting) Cons list.inject nl p1 p_nl path.simps path_NE)
+      qed(insert l_nl1' Cons nl1' s_nl1' d_nl1' False, auto)
+    next
+      case True
+      then obtain nl11 nl12 where nl1: "nl1 = nl11 @ n # nl12"
+      by (metis split_list)
+      have p12: "path f (n # nl12)"
+      apply(rule path_post_concat[of _ "n # nl11"]) using p_nl[unfolded nl nl1] by auto
+      obtain nl12' where p1': "path f nl12'" and hd_nl12': "hd nl12' = n" and
+      l_nl12': "last nl12' = last (n # nl12)" and d_nl12': "distinct nl12'"
+      and s_nl12': "set nl12' \<subseteq> {n} \<union> set nl12"
+      using 1(1)[THEN allE[of _ "n # nl12"]] p12 unfolding nl nl1 by auto
+      thus ?thesis apply(intro exI[of _ nl12']) unfolding nl nl1 by auto
+    qed
+  qed
+qed
+
+lemma path_subtr:
+assumes f: "\<And> n. root (f n) = n"
+and p: "path f nl"
+shows "subtr (set nl) (f (last nl)) (f (hd nl))"
+using p proof (induct rule: path.induct)
+  case (Ind n1 nl n)  let ?ns1 = "insert n1 (set nl)"
+  have "path f (n1 # nl)"
+  and "subtr ?ns1 (f (last (n1 # nl))) (f n1)"
+  and fn1: "Inr (f n1) \<in> cont (f n)" using Ind by simp_all
+  hence fn1_flast:  "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n1)"
+  by (metis subset_insertI subtr_mono)
+  have 1: "last (n # n1 # nl) = last (n1 # nl)" by auto
+  have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)"
+  using f subtr.Step[OF _ fn1_flast fn1] by auto
+  thus ?case unfolding 1 by simp
+qed (metis f list.sel(1) last_ConsL last_in_set not_Cons_self2 subtr.Refl)
+
+lemma reg_subtr_path_aux:
+assumes f: "reg f tr" and n: "subtr ns tr1 tr"
+shows "\<exists> nl. path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
+using n f proof(induct rule: subtr.induct)
+  case (Refl tr ns)
+  thus ?case
+  apply(intro exI[of _ "[root tr]"]) apply simp by (metis (lifting) path.Base reg_root)
+next
+  case (Step tr ns tr2 tr1)
+  hence rtr: "root tr \<in> ns" and tr1_tr: "Inr tr1 \<in> cont tr"
+  and tr2_tr1: "subtr ns tr2 tr1" and tr: "reg f tr" by auto
+  have tr1: "reg f tr1" using reg_subtr[OF tr] rtr tr1_tr
+  by (metis (lifting) Step.prems iso_tuple_UNIV_I reg_def subtr.Step)
+  obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1"
+  and last_nl: "f (last nl) = tr2" and set: "set nl \<subseteq> ns" using Step(3)[OF tr1] by auto
+  have 0: "path f (root tr # nl)" apply (subst path.simps)
+  using f_nl nl reg_root tr tr1_tr by (metis list.sel(1) neq_Nil_conv)
+  show ?case apply(rule exI[of _ "(root tr) # nl"])
+  using 0 reg_root tr last_nl nl path_NE rtr set by auto
+qed
+
+lemma reg_subtr_path:
+assumes f: "reg f tr" and n: "subtr ns tr1 tr"
+shows "\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
+using reg_subtr_path_aux[OF assms] path_distinct[of f]
+by (metis (lifting) order_trans)
+
+lemma subtr_iff_path:
+assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
+shows "subtr ns tr1 tr \<longleftrightarrow>
+       (\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns)"
+proof safe
+  fix nl assume p: "path f nl" and nl: "set nl \<subseteq> ns"
+  have "subtr (set nl) (f (last nl)) (f (hd nl))"
+  apply(rule path_subtr) using p f by simp_all
+  thus "subtr ns (f (last nl)) (f (hd nl))"
+  using subtr_mono nl by auto
+qed(insert reg_subtr_path[OF r], auto)
+
+lemma inFr_iff_path:
+assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
+shows
+"inFr ns tr t \<longleftrightarrow>
+ (\<exists> nl tr1. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and>
+            set nl \<subseteq> ns \<and> Inl t \<in> cont tr1)"
+apply safe
+apply (metis (no_types) inFr_subtr r reg_subtr_path)
+by (metis f inFr.Base path_subtr subtr_inFr subtr_mono subtr_rootL_in)
+
+
+
+subsection{* The Regular Cut of a Tree *}
+
+context fixes tr0 :: dtree
+begin
+
+(* Picking a subtree of a certain root: *)
+definition "pick n \<equiv> SOME tr. subtr UNIV tr tr0 \<and> root tr = n"
+
+lemma pick:
+assumes "inItr UNIV tr0 n"
+shows "subtr UNIV (pick n) tr0 \<and> root (pick n) = n"
+proof-
+  have "\<exists> tr. subtr UNIV tr tr0 \<and> root tr = n"
+  using assms by (metis (lifting) inItr_subtr)
+  thus ?thesis unfolding pick_def by(rule someI_ex)
+qed
+
+lemmas subtr_pick = pick[THEN conjunct1]
+lemmas root_pick = pick[THEN conjunct2]
+
+lemma wf_pick:
+assumes tr0: "wf tr0" and n: "inItr UNIV tr0 n"
+shows "wf (pick n)"
+using wf_subtr[OF tr0 subtr_pick[OF n]] .
+
+definition "H_r n \<equiv> root (pick n)"
+definition "H_c n \<equiv> (id \<oplus> root) ` cont (pick n)"
+
+(* The regular tree of a function: *)
+definition H :: "N \<Rightarrow> dtree" where
+"H \<equiv> unfold H_r H_c"
+
+lemma finite_H_c: "finite (H_c n)"
+unfolding H_c_def by (metis finite_cont finite_imageI)
+
+lemma root_H_pick: "root (H n) = root (pick n)"
+using unfold(1)[of H_r H_c n] unfolding H_def H_r_def by simp
+
+lemma root_H[simp]:
+assumes "inItr UNIV tr0 n"
+shows "root (H n) = n"
+unfolding root_H_pick root_pick[OF assms] ..
+
+lemma cont_H[simp]:
+"cont (H n) = (id \<oplus> (H o root)) ` cont (pick n)"
+apply(subst id_comp[symmetric, of id]) unfolding map_sum.comp[symmetric]
+unfolding image_comp [symmetric] H_c_def [symmetric]
+using unfold(2) [of H_c n H_r, OF finite_H_c]
+unfolding H_def ..
+
+lemma Inl_cont_H[simp]:
+"Inl -` (cont (H n)) = Inl -` (cont (pick n))"
+unfolding cont_H by simp
+
+lemma Inr_cont_H:
+"Inr -` (cont (H n)) = (H \<circ> root) ` (Inr -` cont (pick n))"
+unfolding cont_H by simp
+
+lemma subtr_H:
+assumes n: "inItr UNIV tr0 n" and "subtr UNIV tr1 (H n)"
+shows "\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = H n1"
+proof-
+  {fix tr ns assume "subtr UNIV tr1 tr"
+   hence "tr = H n \<longrightarrow> (\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = H n1)"
+   proof (induct rule: subtr_UNIV_inductL)
+     case (Step tr2 tr1 tr)
+     show ?case proof
+       assume "tr = H n"
+       then obtain n1 where tr2: "Inr tr2 \<in> cont tr1"
+       and tr1_tr: "subtr UNIV tr1 tr" and n1: "inItr UNIV tr0 n1" and tr1: "tr1 = H n1"
+       using Step by auto
+       obtain tr2' where tr2: "tr2 = H (root tr2')"
+       and tr2': "Inr tr2' \<in> cont (pick n1)"
+       using tr2 Inr_cont_H[of n1]
+       unfolding tr1 image_def comp_def using vimage_eq by auto
+       have "inItr UNIV tr0 (root tr2')"
+       using inItr.Base inItr.Ind n1 pick subtr_inItr tr2' by (metis iso_tuple_UNIV_I)
+       thus "\<exists>n2. inItr UNIV tr0 n2 \<and> tr2 = H n2" using tr2 by blast
+     qed
+   qed(insert n, auto)
+  }
+  thus ?thesis using assms by auto
+qed
+
+lemma root_H_root:
+assumes n: "inItr UNIV tr0 n" and t_tr: "t_tr \<in> cont (pick n)"
+shows "(id \<oplus> (root \<circ> H \<circ> root)) t_tr = (id \<oplus> root) t_tr"
+using assms apply(cases t_tr)
+  apply (metis (lifting) map_sum.simps(1))
+  using pick H_def H_r_def unfold(1)
+      inItr.Base comp_apply subtr_StepL subtr_inItr map_sum.simps(2)
+  by (metis UNIV_I)
+
+lemma H_P:
+assumes tr0: "wf tr0" and n: "inItr UNIV tr0 n"
+shows "(n, (id \<oplus> root) ` cont (H n)) \<in> P" (is "?L \<in> P")
+proof-
+  have "?L = (n, (id \<oplus> root) ` cont (pick n))"
+  unfolding cont_H image_comp map_sum.comp id_comp comp_assoc[symmetric]
+  unfolding Pair_eq apply(rule conjI[OF refl]) apply(rule image_cong[OF refl])
+  by (rule root_H_root[OF n])
+  moreover have "... \<in> P" by (metis (lifting) wf_pick root_pick wf_P n tr0)
+  ultimately show ?thesis by simp
+qed
+
+lemma wf_H:
+assumes tr0: "wf tr0" and "inItr UNIV tr0 n"
+shows "wf (H n)"
+proof-
+  {fix tr have "\<exists> n. inItr UNIV tr0 n \<and> tr = H n \<Longrightarrow> wf tr"
+   proof (induct rule: wf_raw_coind)
+     case (Hyp tr)
+     then obtain n where n: "inItr UNIV tr0 n" and tr: "tr = H n" by auto
+     show ?case apply safe
+     apply (metis (lifting) H_P root_H n tr tr0)
+     unfolding tr Inr_cont_H unfolding inj_on_def apply clarsimp using root_H
+     apply (metis UNIV_I inItr.Base n pick subtr2.simps subtr_inItr subtr_subtr2)
+     by (metis n subtr.Refl subtr_StepL subtr_H tr UNIV_I)
+   qed
+  }
+  thus ?thesis using assms by blast
+qed
+
+(* The regular cut of a tree: *)
+definition "rcut \<equiv> H (root tr0)"
+
+lemma reg_rcut: "reg H rcut"
+unfolding reg_def rcut_def
+by (metis inItr.Base root_H subtr_H UNIV_I)
+
+lemma rcut_reg:
+assumes "reg H tr0"
+shows "rcut = tr0"
+using assms unfolding rcut_def reg_def by (metis subtr.Refl UNIV_I)
+
+lemma rcut_eq: "rcut = tr0 \<longleftrightarrow> reg H tr0"
+using reg_rcut rcut_reg by metis
+
+lemma regular_rcut: "regular rcut"
+using reg_rcut unfolding regular_def by blast
+
+lemma Fr_rcut: "Fr UNIV rcut \<subseteq> Fr UNIV tr0"
+proof safe
+  fix t assume "t \<in> Fr UNIV rcut"
+  then obtain tr where t: "Inl t \<in> cont tr" and tr: "subtr UNIV tr (H (root tr0))"
+  using Fr_subtr[of UNIV "H (root tr0)"] unfolding rcut_def
+  by (metis (full_types) Fr_def inFr_subtr mem_Collect_eq)
+  obtain n where n: "inItr UNIV tr0 n" and tr: "tr = H n" using tr
+  by (metis (lifting) inItr.Base subtr_H UNIV_I)
+  have "Inl t \<in> cont (pick n)" using t using Inl_cont_H[of n] unfolding tr
+  by (metis (lifting) vimageD vimageI2)
+  moreover have "subtr UNIV (pick n) tr0" using subtr_pick[OF n] ..
+  ultimately show "t \<in> Fr UNIV tr0" unfolding Fr_subtr_cont by auto
+qed
+
+lemma wf_rcut:
+assumes "wf tr0"
+shows "wf rcut"
+unfolding rcut_def using wf_H[OF assms inItr.Base] by simp
+
+lemma root_rcut[simp]: "root rcut = root tr0"
+unfolding rcut_def
+by (metis (lifting) root_H inItr.Base reg_def reg_root subtr_rootR_in)
+
+end (* context *)
+
+
+subsection{* Recursive Description of the Regular Tree Frontiers *}
+
+lemma regular_inFr:
+assumes r: "regular tr" and In: "root tr \<in> ns"
+and t: "inFr ns tr t"
+shows "t \<in> Inl -` (cont tr) \<or>
+       (\<exists> tr'. Inr tr' \<in> cont tr \<and> inFr (ns - {root tr}) tr' t)"
+(is "?L \<or> ?R")
+proof-
+  obtain f where r: "reg f tr" and f: "\<And>n. root (f n) = n"
+  using r unfolding regular_def2 by auto
+  obtain nl tr1 where d_nl: "distinct nl" and p: "path f nl" and hd_nl: "f (hd nl) = tr"
+  and l_nl: "f (last nl) = tr1" and s_nl: "set nl \<subseteq> ns" and t_tr1: "Inl t \<in> cont tr1"
+  using t unfolding inFr_iff_path[OF r f] by auto
+  obtain n nl1 where nl: "nl = n # nl1" by (metis (lifting) p path.simps)
+  hence f_n: "f n = tr" using hd_nl by simp
+  have n_nl1: "n \<notin> set nl1" using d_nl unfolding nl by auto
+  show ?thesis
+  proof(cases nl1)
+    case Nil hence "tr = tr1" using f_n l_nl unfolding nl by simp
+    hence ?L using t_tr1 by simp thus ?thesis by simp
+  next
+    case (Cons n1 nl2) note nl1 = Cons
+    have 1: "last nl1 = last nl" "hd nl1 = n1" unfolding nl nl1 by simp_all
+    have p1: "path f nl1" and n1_tr: "Inr (f n1) \<in> cont tr"
+    using path.simps[of f nl] p f_n unfolding nl nl1 by auto
+    have r1: "reg f (f n1)" using reg_Inr_cont[OF r n1_tr] .
+    have 0: "inFr (set nl1) (f n1) t" unfolding inFr_iff_path[OF r1 f]
+    apply(intro exI[of _ nl1], intro exI[of _ tr1])
+    using d_nl unfolding 1 l_nl unfolding nl using p1 t_tr1 by auto
+    have root_tr: "root tr = n" by (metis f f_n)
+    have "inFr (ns - {root tr}) (f n1) t" apply(rule inFr_mono[OF 0])
+    using s_nl unfolding root_tr unfolding nl using n_nl1 by auto
+    thus ?thesis using n1_tr by auto
+  qed
+qed
+
+lemma regular_Fr:
+assumes r: "regular tr" and In: "root tr \<in> ns"
+shows "Fr ns tr =
+       Inl -` (cont tr) \<union>
+       \<Union> {Fr (ns - {root tr}) tr' | tr'. Inr tr' \<in> cont tr}"
+unfolding Fr_def
+using In inFr.Base regular_inFr[OF assms] apply safe
+apply (simp, metis (full_types) mem_Collect_eq)
+apply simp
+by (simp, metis (lifting) inFr_Ind_minus insert_Diff)
+
+
+subsection{* The Generated Languages *}
+
+(* The (possibly inifinite tree) generated language *)
+definition "L ns n \<equiv> {Fr ns tr | tr. wf tr \<and> root tr = n}"
+
+(* The regular-tree generated language *)
+definition "Lr ns n \<equiv> {Fr ns tr | tr. wf tr \<and> root tr = n \<and> regular tr}"
+
+lemma L_rec_notin:
+assumes "n \<notin> ns"
+shows "L ns n = {{}}"
+using assms unfolding L_def apply safe
+  using not_root_Fr apply force
+  apply(rule exI[of _ "deftr n"])
+  by (metis (no_types) wf_deftr not_root_Fr root_deftr)
+
+lemma Lr_rec_notin:
+assumes "n \<notin> ns"
+shows "Lr ns n = {{}}"
+using assms unfolding Lr_def apply safe
+  using not_root_Fr apply force
+  apply(rule exI[of _ "deftr n"])
+  by (metis (no_types) regular_def wf_deftr not_root_Fr reg_deftr root_deftr)
+
+lemma wf_subtrOf:
+assumes "wf tr" and "Inr n \<in> prodOf tr"
+shows "wf (subtrOf tr n)"
+by (metis assms wf_cont subtrOf)
+
+lemma Lr_rec_in:
+assumes n: "n \<in> ns"
+shows "Lr ns n \<subseteq>
+{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
+    (n,tns) \<in> P \<and>
+    (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n')}"
+(is "Lr ns n \<subseteq> {?F tns K | tns K. (n,tns) \<in> P \<and> ?\<phi> tns K}")
+proof safe
+  fix ts assume "ts \<in> Lr ns n"
+  then obtain tr where dtr: "wf tr" and r: "root tr = n" and tr: "regular tr"
+  and ts: "ts = Fr ns tr" unfolding Lr_def by auto
+  def tns \<equiv> "(id \<oplus> root) ` (cont tr)"
+  def K \<equiv> "\<lambda> n'. Fr (ns - {n}) (subtrOf tr n')"
+  show "\<exists>tns K. ts = ?F tns K \<and> (n, tns) \<in> P \<and> ?\<phi> tns K"
+  apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI)
+    show "ts = Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns}"
+    unfolding ts regular_Fr[OF tr n[unfolded r[symmetric]]]
+    unfolding tns_def K_def r[symmetric]
+    unfolding Inl_prodOf wf_subtrOf_Union[OF dtr] ..
+    show "(n, tns) \<in> P" unfolding tns_def r[symmetric] using wf_P[OF dtr] .
+    fix n' assume "Inr n' \<in> tns" thus "K n' \<in> Lr (ns - {n}) n'"
+    unfolding K_def Lr_def mem_Collect_eq apply(intro exI[of _ "subtrOf tr n'"])
+    using dtr tr apply(intro conjI refl)  unfolding tns_def
+      apply(erule wf_subtrOf[OF dtr])
+      apply (metis subtrOf)
+      by (metis Inr_subtrOf UNIV_I regular_subtr subtr.simps)
+  qed
+qed
+
+lemma hsubst_aux:
+fixes n ftr tns
+assumes n: "n \<in> ns" and tns: "finite tns" and
+1: "\<And> n'. Inr n' \<in> tns \<Longrightarrow> wf (ftr n')"
+defines "tr \<equiv> Node n ((id \<oplus> ftr) ` tns)"  defines "tr' \<equiv> hsubst tr tr"
+shows "Fr ns tr' = Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
+(is "_ = ?B") proof-
+  have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
+  unfolding tr_def using tns by auto
+  have Frr: "Frr (ns - {n}) tr = \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
+  unfolding Frr_def ctr by auto
+  have "Fr ns tr' = Inl -` (cont tr) \<union> Frr (ns - {n}) tr"
+  using Fr_self_hsubst[OF n[unfolded rtr[symmetric]]] unfolding tr'_def rtr ..
+  also have "... = ?B" unfolding ctr Frr by simp
+  finally show ?thesis .
+qed
+
+lemma L_rec_in:
+assumes n: "n \<in> ns"
+shows "
+{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
+    (n,tns) \<in> P \<and>
+    (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n')}
+ \<subseteq> L ns n"
+proof safe
+  fix tns K
+  assume P: "(n, tns) \<in> P" and 0: "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n'"
+  {fix n' assume "Inr n' \<in> tns"
+   hence "K n' \<in> L (ns - {n}) n'" using 0 by auto
+   hence "\<exists> tr'. K n' = Fr (ns - {n}) tr' \<and> wf tr' \<and> root tr' = n'"
+   unfolding L_def mem_Collect_eq by auto
+  }
+  then obtain ftr where 0: "\<And> n'. Inr n' \<in> tns \<Longrightarrow>
+  K n' = Fr (ns - {n}) (ftr n') \<and> wf (ftr n') \<and> root (ftr n') = n'"
+  by metis
+  def tr \<equiv> "Node n ((id \<oplus> ftr) ` tns)"  def tr' \<equiv> "hsubst tr tr"
+  have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
+  unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P)
+  have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong)
+  unfolding ctr apply simp apply simp apply safe
+  using 0 unfolding image_def apply force apply simp by (metis 0 vimageI2)
+  have 1: "{K n' |n'. Inr n' \<in> tns} = {Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
+  using 0 by auto
+  have dtr: "wf tr" apply(rule wf.dtree)
+    apply (metis (lifting) P prtr rtr)
+    unfolding inj_on_def ctr using 0 by auto
+  hence dtr': "wf tr'" unfolding tr'_def by (metis wf_hsubst)
+  have tns: "finite tns" using finite_in_P P by simp
+  have "Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns} \<in> L ns n"
+  unfolding L_def mem_Collect_eq apply(intro exI[of _ tr'] conjI)
+  using dtr' 0 hsubst_aux[OF assms tns, of ftr] unfolding tr_def tr'_def by auto
+  thus "Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} \<in> L ns n" unfolding 1 .
+qed
+
+lemma card_N: "(n::N) \<in> ns \<Longrightarrow> card (ns - {n}) < card ns"
+by (metis finite_N Diff_UNIV Diff_infinite_finite card_Diff1_less finite.emptyI)
+
+function LL where
+"LL ns n =
+ (if n \<notin> ns then {{}} else
+ {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
+    (n,tns) \<in> P \<and>
+    (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n')})"
+by(pat_completeness, auto)
+termination apply(relation "inv_image (measure card) fst")
+using card_N by auto
+
+declare LL.simps[code]
+declare LL.simps[simp del]
+
+lemma Lr_LL: "Lr ns n \<subseteq> LL ns n"
+proof (induct ns arbitrary: n rule: measure_induct[of card])
+  case (1 ns n) show ?case proof(cases "n \<in> ns")
+    case False thus ?thesis unfolding Lr_rec_notin[OF False] by (simp add: LL.simps)
+  next
+    case True show ?thesis apply(rule subset_trans)
+    using Lr_rec_in[OF True] apply assumption
+    unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
+      fix tns K
+      assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
+      assume "(n, tns) \<in> P"
+      and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n'"
+      thus "\<exists>tnsa Ka.
+             Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
+             Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
+             (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> LL (ns - {n}) n')"
+      apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
+    qed
+  qed
+qed
+
+lemma LL_L: "LL ns n \<subseteq> L ns n"
+proof (induct ns arbitrary: n rule: measure_induct[of card])
+  case (1 ns n) show ?case proof(cases "n \<in> ns")
+    case False thus ?thesis unfolding L_rec_notin[OF False] by (simp add: LL.simps)
+  next
+    case True show ?thesis apply(rule subset_trans)
+    prefer 2 using L_rec_in[OF True] apply assumption
+    unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
+      fix tns K
+      assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
+      assume "(n, tns) \<in> P"
+      and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n'"
+      thus "\<exists>tnsa Ka.
+             Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
+             Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
+             (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> L (ns - {n}) n')"
+      apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
+    qed
+  qed
+qed
+
+(* The subsumpsion relation between languages *)
+definition "subs L1 L2 \<equiv> \<forall> ts2 \<in> L2. \<exists> ts1 \<in> L1. ts1 \<subseteq> ts2"
+
+lemma incl_subs[simp]: "L2 \<subseteq> L1 \<Longrightarrow> subs L1 L2"
+unfolding subs_def by auto
+
+lemma subs_refl[simp]: "subs L1 L1" unfolding subs_def by auto
+
+lemma subs_trans: "\<lbrakk>subs L1 L2; subs L2 L3\<rbrakk> \<Longrightarrow> subs L1 L3"
+unfolding subs_def by (metis subset_trans)
+
+(* Language equivalence *)
+definition "leqv L1 L2 \<equiv> subs L1 L2 \<and> subs L2 L1"
+
+lemma subs_leqv[simp]: "leqv L1 L2 \<Longrightarrow> subs L1 L2"
+unfolding leqv_def by auto
+
+lemma subs_leqv_sym[simp]: "leqv L1 L2 \<Longrightarrow> subs L2 L1"
+unfolding leqv_def by auto
+
+lemma leqv_refl[simp]: "leqv L1 L1" unfolding leqv_def by auto
+
+lemma leqv_trans:
+assumes 12: "leqv L1 L2" and 23: "leqv L2 L3"
+shows "leqv L1 L3"
+using assms unfolding leqv_def by (metis (lifting) subs_trans)
+
+lemma leqv_sym: "leqv L1 L2 \<Longrightarrow> leqv L2 L1"
+unfolding leqv_def by auto
+
+lemma leqv_Sym: "leqv L1 L2 \<longleftrightarrow> leqv L2 L1"
+unfolding leqv_def by auto
+
+lemma Lr_incl_L: "Lr ns ts \<subseteq> L ns ts"
+unfolding Lr_def L_def by auto
+
+lemma Lr_subs_L: "subs (Lr UNIV ts) (L UNIV ts)"
+unfolding subs_def proof safe
+  fix ts2 assume "ts2 \<in> L UNIV ts"
+  then obtain tr where ts2: "ts2 = Fr UNIV tr" and dtr: "wf tr" and rtr: "root tr = ts"
+  unfolding L_def by auto
+  thus "\<exists>ts1\<in>Lr UNIV ts. ts1 \<subseteq> ts2"
+  apply(intro bexI[of _ "Fr UNIV (rcut tr)"])
+  unfolding Lr_def L_def using Fr_rcut wf_rcut root_rcut regular_rcut by auto
+qed
+
+lemma Lr_leqv_L: "leqv (Lr UNIV ts) (L UNIV ts)"
+using Lr_subs_L unfolding leqv_def by (metis (lifting) Lr_incl_L incl_subs)
+
+lemma LL_leqv_L: "leqv (LL UNIV ts) (L UNIV ts)"
+by (metis (lifting) LL_L Lr_LL Lr_subs_L incl_subs leqv_def subs_trans)
+
+lemma LL_leqv_Lr: "leqv (LL UNIV ts) (Lr UNIV ts)"
+using Lr_leqv_L LL_leqv_L by (metis leqv_Sym leqv_trans)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Parallel.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,147 @@
+(*  Title:      HOL/Datatype_Examples/Derivation_Trees/Parallel.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Parallel composition.
+*)
+
+header {* Parallel Composition *}
+
+theory Parallel
+imports DTree
+begin
+
+no_notation plus_class.plus (infixl "+" 65)
+
+consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl "+" 60)
+
+axiomatization where
+    Nplus_comm: "(a::N) + b = b + (a::N)"
+and Nplus_assoc: "((a::N) + b) + c = a + (b + c)"
+
+subsection{* Corecursive Definition of Parallel Composition *}
+
+fun par_r where "par_r (tr1,tr2) = root tr1 + root tr2"
+fun par_c where
+"par_c (tr1,tr2) =
+ Inl ` (Inl -` (cont tr1 \<union> cont tr2)) \<union>
+ Inr ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
+
+declare par_r.simps[simp del]  declare par_c.simps[simp del]
+
+definition par :: "dtree \<times> dtree \<Rightarrow> dtree" where
+"par \<equiv> unfold par_r par_c"
+
+abbreviation par_abbr (infixr "\<parallel>" 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
+
+lemma finite_par_c: "finite (par_c (tr1, tr2))"
+unfolding par_c.simps apply(rule finite_UnI)
+  apply (metis finite_Un finite_cont finite_imageI finite_vimageI inj_Inl)
+  apply(intro finite_imageI finite_cartesian_product finite_vimageI)
+  using finite_cont by auto
+
+lemma root_par: "root (tr1 \<parallel> tr2) = root tr1 + root tr2"
+using unfold(1)[of par_r par_c "(tr1,tr2)"] unfolding par_def par_r.simps by simp
+
+lemma cont_par:
+"cont (tr1 \<parallel> tr2) = (id \<oplus> par) ` par_c (tr1,tr2)"
+using unfold(2)[of par_c "(tr1,tr2)" par_r, OF finite_par_c]
+unfolding par_def ..
+
+lemma Inl_cont_par[simp]:
+"Inl -` (cont (tr1 \<parallel> tr2)) = Inl -` (cont tr1 \<union> cont tr2)"
+unfolding cont_par par_c.simps by auto
+
+lemma Inr_cont_par[simp]:
+"Inr -` (cont (tr1 \<parallel> tr2)) = par ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
+unfolding cont_par par_c.simps by auto
+
+lemma Inl_in_cont_par:
+"Inl t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (Inl t \<in> cont tr1 \<or> Inl t \<in> cont tr2)"
+using Inl_cont_par[of tr1 tr2] unfolding vimage_def by auto
+
+lemma Inr_in_cont_par:
+"Inr t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (t \<in> par ` (Inr -` cont tr1 \<times> Inr -` cont tr2))"
+using Inr_cont_par[of tr1 tr2] unfolding vimage_def by auto
+
+
+subsection{* Structural Coinduction Proofs *}
+
+lemma rel_set_rel_sum_eq[simp]:
+"rel_set (rel_sum (op =) \<phi>) A1 A2 \<longleftrightarrow>
+ Inl -` A1 = Inl -` A2 \<and> rel_set \<phi> (Inr -` A1) (Inr -` A2)"
+unfolding rel_set_rel_sum rel_set_eq ..
+
+(* Detailed proofs of commutativity and associativity: *)
+theorem par_com: "tr1 \<parallel> tr2 = tr2 \<parallel> tr1"
+proof-
+  let ?\<theta> = "\<lambda> trA trB. \<exists> tr1 tr2. trA = tr1 \<parallel> tr2 \<and> trB = tr2 \<parallel> tr1"
+  {fix trA trB
+   assume "?\<theta> trA trB" hence "trA = trB"
+   apply (induct rule: dtree_coinduct)
+   unfolding rel_set_rel_sum rel_set_eq unfolding rel_set_def proof safe
+     fix tr1 tr2  show "root (tr1 \<parallel> tr2) = root (tr2 \<parallel> tr1)"
+     unfolding root_par by (rule Nplus_comm)
+   next
+     fix n tr1 tr2 assume "Inl n \<in> cont (tr1 \<parallel> tr2)" thus "n \<in> Inl -` (cont (tr2 \<parallel> tr1))"
+     unfolding Inl_in_cont_par by auto
+   next
+     fix n tr1 tr2 assume "Inl n \<in> cont (tr2 \<parallel> tr1)" thus "n \<in> Inl -` (cont (tr1 \<parallel> tr2))"
+     unfolding Inl_in_cont_par by auto
+   next
+     fix tr1 tr2 trA' assume "Inr trA' \<in> cont (tr1 \<parallel> tr2)"
+     then obtain tr1' tr2' where "trA' = tr1' \<parallel> tr2'"
+     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+     unfolding Inr_in_cont_par by auto
+     thus "\<exists> trB' \<in> Inr -` (cont (tr2 \<parallel> tr1)). ?\<theta> trA' trB'"
+     apply(intro bexI[of _ "tr2' \<parallel> tr1'"]) unfolding Inr_in_cont_par by auto
+   next
+     fix tr1 tr2 trB' assume "Inr trB' \<in> cont (tr2 \<parallel> tr1)"
+     then obtain tr1' tr2' where "trB' = tr2' \<parallel> tr1'"
+     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+     unfolding Inr_in_cont_par by auto
+     thus "\<exists> trA' \<in> Inr -` (cont (tr1 \<parallel> tr2)). ?\<theta> trA' trB'"
+     apply(intro bexI[of _ "tr1' \<parallel> tr2'"]) unfolding Inr_in_cont_par by auto
+   qed
+  }
+  thus ?thesis by blast
+qed
+
+lemma par_assoc: "(tr1 \<parallel> tr2) \<parallel> tr3 = tr1 \<parallel> (tr2 \<parallel> tr3)"
+proof-
+  let ?\<theta> =
+  "\<lambda> trA trB. \<exists> tr1 tr2 tr3. trA = (tr1 \<parallel> tr2) \<parallel> tr3 \<and> trB = tr1 \<parallel> (tr2 \<parallel> tr3)"
+  {fix trA trB
+   assume "?\<theta> trA trB" hence "trA = trB"
+   apply (induct rule: dtree_coinduct)
+   unfolding rel_set_rel_sum rel_set_eq unfolding rel_set_def proof safe
+     fix tr1 tr2 tr3  show "root ((tr1 \<parallel> tr2) \<parallel> tr3) = root (tr1 \<parallel> (tr2 \<parallel> tr3))"
+     unfolding root_par by (rule Nplus_assoc)
+   next
+     fix n tr1 tr2 tr3 assume "Inl n \<in> (cont ((tr1 \<parallel> tr2) \<parallel> tr3))"
+     thus "n \<in> Inl -` (cont (tr1 \<parallel> tr2 \<parallel> tr3))" unfolding Inl_in_cont_par by simp
+   next
+     fix n tr1 tr2 tr3 assume "Inl n \<in> (cont (tr1 \<parallel> tr2 \<parallel> tr3))"
+     thus "n \<in> Inl -` (cont ((tr1 \<parallel> tr2) \<parallel> tr3))" unfolding Inl_in_cont_par by simp
+   next
+     fix trA' tr1 tr2 tr3 assume "Inr trA' \<in> cont ((tr1 \<parallel> tr2) \<parallel> tr3)"
+     then obtain tr1' tr2' tr3' where "trA' = (tr1' \<parallel> tr2') \<parallel> tr3'"
+     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+     and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
+     thus "\<exists> trB' \<in> Inr -` (cont (tr1 \<parallel> tr2 \<parallel> tr3)). ?\<theta> trA' trB'"
+     apply(intro bexI[of _ "tr1' \<parallel> tr2' \<parallel> tr3'"])
+     unfolding Inr_in_cont_par by auto
+   next
+     fix trB' tr1 tr2 tr3 assume "Inr trB' \<in> cont (tr1 \<parallel> tr2 \<parallel> tr3)"
+     then obtain tr1' tr2' tr3' where "trB' = tr1' \<parallel> (tr2' \<parallel> tr3')"
+     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+     and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
+     thus "\<exists> trA' \<in> Inr -` cont ((tr1 \<parallel> tr2) \<parallel> tr3). ?\<theta> trA' trB'"
+     apply(intro bexI[of _ "(tr1' \<parallel> tr2') \<parallel> tr3'"])
+     unfolding Inr_in_cont_par by auto
+   qed
+  }
+  thus ?thesis by blast
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,64 @@
+(*  Title:      HOL/Datatype_Examples/Derivation_Trees/Prelim.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Preliminaries.
+*)
+
+header {* Preliminaries *}
+
+theory Prelim
+imports "~~/src/HOL/Library/FSet"
+begin
+
+notation BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
+
+declare fset_to_fset[simp]
+
+lemma fst_snd_convol_o[simp]: "\<langle>fst o s, snd o s\<rangle> = s"
+apply(rule ext) by (simp add: convol_def)
+
+abbreviation sm_abbrev (infix "\<oplus>" 60)
+where "f \<oplus> g \<equiv> Sum_Type.map_sum f g"
+
+lemma map_sum_InlD: "(f \<oplus> g) z = Inl x \<Longrightarrow> \<exists>y. z = Inl y \<and> f y = x"
+by (cases z) auto
+
+lemma map_sum_InrD: "(f \<oplus> g) z = Inr x \<Longrightarrow> \<exists>y. z = Inr y \<and> g y = x"
+by (cases z) auto
+
+abbreviation case_sum_abbrev ("[[_,_]]" 800)
+where "[[f,g]] \<equiv> Sum_Type.case_sum f g"
+
+lemma Inl_oplus_elim:
+assumes "Inl tr \<in> (id \<oplus> f) ` tns"
+shows "Inl tr \<in> tns"
+using assms apply clarify by (case_tac x, auto)
+
+lemma Inl_oplus_iff[simp]: "Inl tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> Inl tr \<in> tns"
+using Inl_oplus_elim
+by (metis id_def image_iff map_sum.simps(1))
+
+lemma Inl_m_oplus[simp]: "Inl -` (id \<oplus> f) ` tns = Inl -` tns"
+using Inl_oplus_iff unfolding vimage_def by auto
+
+lemma Inr_oplus_elim:
+assumes "Inr tr \<in> (id \<oplus> f) ` tns"
+shows "\<exists> n. Inr n \<in> tns \<and> f n = tr"
+using assms apply clarify by (case_tac x, auto)
+
+lemma Inr_oplus_iff[simp]:
+"Inr tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> (\<exists> n. Inr n \<in> tns \<and> f n = tr)"
+apply (rule iffI)
+ apply (metis Inr_oplus_elim)
+by (metis image_iff map_sum.simps(2))
+
+lemma Inr_m_oplus[simp]: "Inr -` (id \<oplus> f) ` tns = f ` (Inr -` tns)"
+using Inr_oplus_iff unfolding vimage_def by auto
+
+lemma Inl_Inr_image_cong:
+assumes "Inl -` A = Inl -` B" and "Inr -` A = Inr -` B"
+shows "A = B"
+apply safe using assms apply(case_tac x, auto) by(case_tac x, auto)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Instructions.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,162 @@
+(*  Title:      HOL/Datatype_Benchmark/Instructions.thy
+
+Example from Konrad: 68000 instruction set.
+*)
+
+theory Instructions imports Main begin
+
+datatype Size = Byte | Word | Long
+
+datatype DataRegister =
+  RegD0
+| RegD1
+| RegD2
+| RegD3
+| RegD4
+| RegD5
+| RegD6
+| RegD7
+
+datatype AddressRegister =
+  RegA0
+| RegA1
+| RegA2
+| RegA3
+| RegA4
+| RegA5
+| RegA6
+| RegA7
+
+datatype DataOrAddressRegister =
+  data DataRegister
+| address AddressRegister
+
+datatype Condition =
+  Hi
+| Ls
+| Cc
+| Cs
+| Ne
+| Eq
+| Vc
+| Vs
+| Pl
+| Mi
+| Ge
+| Lt
+| Gt
+| Le
+
+datatype AddressingMode =
+  immediate nat
+| direct DataOrAddressRegister
+| indirect AddressRegister
+| postinc AddressRegister
+| predec AddressRegister
+| indirectdisp nat AddressRegister
+| indirectindex nat AddressRegister DataOrAddressRegister Size
+| absolute nat
+| pcdisp nat
+| pcindex nat DataOrAddressRegister Size
+
+datatype M68kInstruction =
+  ABCD AddressingMode AddressingMode
+| ADD Size AddressingMode AddressingMode
+| ADDA Size AddressingMode AddressRegister
+| ADDI Size nat AddressingMode
+| ADDQ Size nat AddressingMode
+| ADDX Size AddressingMode AddressingMode
+| AND Size AddressingMode AddressingMode
+| ANDI Size nat AddressingMode
+| ANDItoCCR nat
+| ANDItoSR nat
+| ASL Size AddressingMode DataRegister
+| ASLW AddressingMode
+| ASR Size AddressingMode DataRegister
+| ASRW AddressingMode
+| Bcc Condition Size nat
+| BTST Size AddressingMode AddressingMode
+| BCHG Size AddressingMode AddressingMode
+| BCLR Size AddressingMode AddressingMode
+| BSET Size AddressingMode AddressingMode
+| BRA Size nat
+| BSR Size nat
+| CHK AddressingMode DataRegister
+| CLR Size AddressingMode
+| CMP Size AddressingMode DataRegister
+| CMPA Size AddressingMode AddressRegister
+| CMPI Size nat AddressingMode
+| CMPM Size AddressRegister AddressRegister
+| DBT DataRegister nat
+| DBF DataRegister nat
+| DBcc Condition DataRegister nat
+| DIVS AddressingMode DataRegister
+| DIVU AddressingMode DataRegister
+| EOR Size DataRegister AddressingMode
+| EORI Size nat AddressingMode
+| EORItoCCR nat
+| EORItoSR nat
+| EXG DataOrAddressRegister DataOrAddressRegister
+| EXT Size DataRegister
+| ILLEGAL
+| JMP AddressingMode
+| JSR AddressingMode
+| LEA AddressingMode AddressRegister
+| LINK AddressRegister nat
+| LSL Size AddressingMode DataRegister
+| LSLW AddressingMode
+| LSR Size AddressingMode DataRegister
+| LSRW AddressingMode
+| MOVE Size AddressingMode AddressingMode
+| MOVEtoCCR AddressingMode
+| MOVEtoSR AddressingMode
+| MOVEfromSR AddressingMode
+| MOVEtoUSP AddressingMode
+| MOVEfromUSP AddressingMode
+| MOVEA Size AddressingMode AddressRegister
+| MOVEMto Size AddressingMode "DataOrAddressRegister list"
+| MOVEMfrom Size "DataOrAddressRegister list" AddressingMode
+| MOVEP Size AddressingMode AddressingMode
+| MOVEQ nat DataRegister
+| MULS AddressingMode DataRegister
+| MULU AddressingMode DataRegister
+| NBCD AddressingMode
+| NEG Size AddressingMode
+| NEGX Size AddressingMode
+| NOP
+| NOT Size AddressingMode
+| OR Size AddressingMode AddressingMode
+| ORI Size nat AddressingMode
+| ORItoCCR nat
+| ORItoSR nat
+| PEA AddressingMode
+| RESET
+| ROL Size AddressingMode DataRegister
+| ROLW AddressingMode
+| ROR Size AddressingMode DataRegister
+| RORW AddressingMode
+| ROXL Size AddressingMode DataRegister
+| ROXLW AddressingMode
+| ROXR Size AddressingMode DataRegister
+| ROXRW AddressingMode
+| RTE
+| RTR
+| RTS
+| SBCD AddressingMode AddressingMode
+| ST AddressingMode
+| SF AddressingMode
+| Scc Condition AddressingMode
+| STOP nat
+| SUB Size AddressingMode AddressingMode
+| SUBA Size AddressingMode AddressingMode
+| SUBI Size nat AddressingMode
+| SUBQ Size nat AddressingMode
+| SUBX Size AddressingMode AddressingMode
+| SWAP DataRegister
+| TAS AddressingMode
+| TRAP nat
+| TRAPV
+| TST Size AddressingMode
+| UNLK AddressRegister
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/IsaFoR_Datatypes.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,380 @@
+(*  Title:      HOL/Datatype_Examples/IsaFoR_Datatypes.thy
+    Author:     Rene Thiemann, UIBK
+    Copyright   2014
+
+Benchmark consisting of datatypes defined in IsaFoR.
+*)
+
+header {* Benchmark Consisting of Datatypes Defined in IsaFoR *}
+
+theory IsaFoR_Datatypes
+imports Real
+begin
+
+datatype_new (discs_sels) ('f, 'l) lab =
+    Lab "('f, 'l) lab" 'l
+  | FunLab "('f, 'l) lab" "('f, 'l) lab list"
+  | UnLab 'f
+  | Sharp "('f, 'l) lab"
+
+datatype_new (discs_sels) 'f projL = Projection "(('f \<times> nat) \<times> nat) list"
+
+datatype_new (discs_sels) ('f, 'v) "term" = Var 'v | Fun 'f "('f, 'v) term list"
+datatype_new (discs_sels) ('f, 'v) ctxt =
+    Hole ("\<box>")
+  | More 'f "('f, 'v) term list" "('f, 'v) ctxt" "('f, 'v) term list"
+
+type_synonym ('f, 'v) rule = "('f, 'v) term \<times> ('f, 'v) term"
+type_synonym ('f, 'v) trs  = "('f, 'v) rule set"
+
+type_synonym ('f, 'v) rules = "('f, 'v) rule list"
+type_synonym ('f, 'l, 'v) ruleLL  = "(('f, 'l) lab, 'v) rule"
+type_synonym ('f, 'l, 'v) trsLL   = "(('f, 'l) lab, 'v) rules"
+type_synonym ('f, 'l, 'v) termsLL = "(('f, 'l) lab, 'v) term list"
+
+datatype_new (discs_sels) pos = Empty ("\<epsilon>") | PCons "nat" "pos" (infixr "<#" 70)
+
+type_synonym  ('f, 'v) prseq = "(pos \<times> ('f, 'v) rule \<times> bool \<times> ('f, 'v) term) list"
+type_synonym  ('f, 'v) rseq = "(pos \<times> ('f, 'v) rule \<times> ('f, 'v) term) list"
+
+type_synonym ('f, 'l, 'v) rseqL   = "((('f, 'l) lab, 'v) rule \<times> (('f, 'l) lab, 'v) rseq) list"
+type_synonym ('f, 'l, 'v) dppLL   =
+  "bool \<times> bool \<times> ('f, 'l, 'v) trsLL \<times> ('f, 'l, 'v) trsLL \<times>
+  ('f, 'l, 'v) termsLL \<times>
+  ('f, 'l, 'v) trsLL \<times> ('f, 'l, 'v) trsLL"
+
+type_synonym ('f, 'l, 'v) qreltrsLL =
+  "bool \<times> ('f, 'l, 'v) termsLL \<times> ('f, 'l, 'v) trsLL \<times> ('f, 'l, 'v) trsLL"
+
+type_synonym ('f, 'l, 'v) qtrsLL =
+  "bool \<times> ('f, 'l, 'v) termsLL \<times> ('f, 'l, 'v) trsLL"
+
+datatype_new (discs_sels) location = H | A | B | R
+
+type_synonym ('f, 'v) forb_pattern = "('f, 'v) ctxt \<times> ('f, 'v) term \<times> location"
+type_synonym ('f, 'v) forb_patterns = "('f, 'v) forb_pattern set"
+
+type_synonym ('f, 'l, 'v) fptrsLL =
+  "(('f, 'l) lab, 'v) forb_pattern list \<times> ('f, 'l, 'v) trsLL"
+
+type_synonym ('f, 'l, 'v) prob = "('f, 'l, 'v) qreltrsLL + ('f, 'l, 'v) dppLL"
+
+type_synonym ('f, 'a) lpoly_inter = "'f \<times> nat \<Rightarrow> ('a \<times> 'a list)"
+type_synonym ('f, 'a) lpoly_interL = "(('f \<times> nat) \<times> ('a \<times> 'a list)) list"
+
+type_synonym 'v monom = "('v \<times> nat) list"
+type_synonym ('v, 'a) poly = "('v monom \<times> 'a) list"
+type_synonym ('f, 'a) poly_inter_list = "(('f \<times> nat) \<times> (nat, 'a) poly) list"
+type_synonym 'a vec = "'a list"
+type_synonym 'a mat = "'a vec list"
+
+datatype_new (discs_sels) arctic = MinInfty | Num_arc int
+datatype_new (discs_sels) 'a arctic_delta = MinInfty_delta | Num_arc_delta 'a
+datatype_new (discs_sels) order_tag = Lex | Mul
+
+type_synonym 'f status_prec_repr = "(('f \<times> nat) \<times> (nat \<times> order_tag)) list"
+
+datatype_new (discs_sels) af_entry =
+    Collapse nat
+  | AFList "nat list"
+
+type_synonym 'f afs_list = "(('f \<times> nat) \<times> af_entry) list"
+type_synonym 'f prec_weight_repr = "(('f \<times> nat) \<times> (nat \<times> nat \<times> (nat list option))) list \<times> nat"
+
+datatype_new (discs_sels) 'f redtriple_impl =
+    Int_carrier "('f, int) lpoly_interL"
+  | Int_nl_carrier "('f, int) poly_inter_list"
+  | Rat_carrier "('f, rat) lpoly_interL"
+  | Rat_nl_carrier rat "('f, rat) poly_inter_list"
+  | Real_carrier "('f, real) lpoly_interL"
+  | Real_nl_carrier real "('f, real) poly_inter_list"
+  | Arctic_carrier "('f, arctic) lpoly_interL"
+  | Arctic_rat_carrier "('f, rat arctic_delta) lpoly_interL"
+  | Int_mat_carrier nat nat "('f, int mat) lpoly_interL"
+  | Rat_mat_carrier nat nat "('f, rat mat) lpoly_interL"
+  | Real_mat_carrier nat nat "('f, real mat) lpoly_interL"
+  | Arctic_mat_carrier nat "('f, arctic mat) lpoly_interL"
+  | Arctic_rat_mat_carrier nat "('f, rat arctic_delta mat) lpoly_interL"
+  | RPO "'f status_prec_repr" "'f afs_list"
+  | KBO "'f prec_weight_repr" "'f afs_list"
+
+datatype_new (discs_sels) list_order_type = MS_Ext | Max_Ext | Min_Ext  | Dms_Ext
+type_synonym 'f scnp_af = "(('f \<times> nat) \<times> (nat \<times> nat) list) list"
+
+datatype_new (discs_sels) 'f root_redtriple_impl = SCNP list_order_type "'f scnp_af" "'f redtriple_impl"
+
+type_synonym 'f sig_map_list = "(('f \<times> nat) \<times> 'f list) list"
+type_synonym ('f, 'v) uncurry_info = "'f \<times> 'f sig_map_list \<times> ('f, 'v) rules \<times> ('f, 'v) rules"
+
+datatype_new (discs_sels) arithFun =
+    Arg nat
+  | Const nat
+  | Sum "arithFun list"
+  | Max "arithFun list"
+  | Min "arithFun list"
+  | Prod "arithFun list"
+  | IfEqual arithFun arithFun arithFun arithFun
+
+datatype_new (discs_sels) 'f sl_inter = SL_Inter nat "(('f \<times> nat) \<times> arithFun) list"
+datatype_new (discs_sels) ('f, 'v) sl_variant =
+    Rootlab "('f \<times> nat) option"
+  | Finitelab "'f sl_inter"
+  | QuasiFinitelab "'f sl_inter" 'v
+
+type_synonym ('f, 'v) crit_pair_joins = "(('f, 'v) term \<times> ('f, 'v) rseq \<times> ('f, 'v) term \<times> ('f, 'v) rseq) list"
+
+datatype_new (discs_sels) 'f join_info = Guided "('f, string) crit_pair_joins" | Join_NF | Join_BFS nat
+
+type_synonym unknown_info = string
+
+type_synonym dummy_prf = unit
+
+datatype_new (discs_sels) ('f, 'v) complex_constant_removal_prf = Complex_Constant_Removal_Proof
+  "('f, 'v) term"
+  "(('f, 'v) rule \<times> ('f, 'v) rule) list"
+
+datatype_new (discs_sels) ('f, 'v) cond_constraint =
+    CC_cond bool "('f, 'v) rule"
+  | CC_rewr "('f, 'v) term" "('f, 'v) term"
+  | CC_impl "('f, 'v) cond_constraint list" "('f, 'v) cond_constraint"
+  | CC_all 'v "('f, 'v) cond_constraint"
+
+type_synonym ('f, 'v, 'w) gsubstL = "('v \<times> ('f, 'w) term) list"
+type_synonym ('f, 'v) substL = "('f, 'v, 'v) gsubstL"
+
+datatype_new (discs_sels) ('f, 'v) cond_constraint_prf =
+    Final
+  | Delete_Condition "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
+  | Different_Constructor "('f, 'v) cond_constraint"
+  | Same_Constructor "('f, 'v) cond_constraint" "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
+  | Variable_Equation 'v "('f, 'v) term" "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
+  | Funarg_Into_Var "('f, 'v) cond_constraint" nat 'v "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
+  | Simplify_Condition "('f, 'v) cond_constraint" "('f, 'v) substL" "('f, 'v) cond_constraint" "('f, 'v) cond_constraint_prf"
+  | Induction "('f, 'v) cond_constraint" "('f, 'v) cond_constraint list" "(('f, 'v) rule \<times> (('f, 'v) term \<times> 'v list) list \<times> ('f, 'v) cond_constraint \<times> ('f, 'v) cond_constraint_prf) list"
+
+datatype_new (discs_sels) ('f, 'v) cond_red_pair_prf =
+  Cond_Red_Pair_Prf
+    'f "(('f, 'v) cond_constraint \<times> ('f, 'v) rules \<times> ('f, 'v) cond_constraint_prf) list" nat nat
+
+datatype_new (discs_sels) ('q, 'f) ta_rule = TA_rule 'f "'q list" 'q ("_ _ \<rightarrow> _")
+datatype_new (discs_sels) ('q, 'f) tree_automaton = Tree_Automaton "'q list" "('q, 'f) ta_rule list" "('q \<times> 'q) list"
+datatype_new (discs_sels) 'q ta_relation =
+    Decision_Proc
+  | Id_Relation
+  | Some_Relation "('q \<times> 'q) list"
+
+datatype_new (discs_sels) boundstype = Roof | Match
+datatype_new (discs_sels) ('f, 'q) bounds_info = Bounds_Info boundstype nat "'q list" "('q, 'f \<times> nat) tree_automaton" "'q ta_relation"
+
+datatype_new (discs_sels) ('f, 'v) pat_eqv_prf =
+    Pat_Dom_Renaming "('f, 'v) substL"
+  | Pat_Irrelevant "('f, 'v) substL" "('f, 'v) substL"
+  | Pat_Simplify "('f, 'v) substL" "('f, 'v) substL"
+
+datatype_new (discs_sels) pat_rule_pos = Pat_Base | Pat_Pump | Pat_Close
+
+datatype_new (discs_sels) ('f, 'v) pat_rule_prf =
+    Pat_OrigRule "('f, 'v) rule" bool
+  | Pat_InitPump "('f, 'v) pat_rule_prf" "('f, 'v) substL" "('f, 'v) substL"
+  | Pat_InitPumpCtxt "('f, 'v) pat_rule_prf" "('f, 'v) substL" pos 'v
+  | Pat_Equiv "('f, 'v) pat_rule_prf" bool "('f, 'v) pat_eqv_prf"
+  | Pat_Narrow "('f, 'v) pat_rule_prf" "('f, 'v) pat_rule_prf" pos
+  | Pat_Inst "('f, 'v) pat_rule_prf" "('f, 'v) substL" pat_rule_pos
+  | Pat_Rewr "('f, 'v) pat_rule_prf" "('f, 'v) term \<times> ('f, 'v) rseq" pat_rule_pos 'v
+  | Pat_Exp_Sigma "('f, 'v) pat_rule_prf" nat
+
+datatype_new (discs_sels) ('f, 'v) non_loop_prf =
+    Non_Loop_Prf "('f, 'v) pat_rule_prf" "('f, 'v) substL" "('f, 'v) substL" nat nat pos
+
+datatype_new (discs_sels) ('f, 'l, 'v) problem =
+    SN_TRS "('f, 'l, 'v) qreltrsLL"
+  | SN_FP_TRS "('f, 'l, 'v) fptrsLL"
+  | Finite_DPP "('f, 'l, 'v) dppLL"
+  | Unknown_Problem unknown_info
+  | Not_SN_TRS "('f, 'l, 'v) qtrsLL"
+  | Not_RelSN_TRS "('f, 'l, 'v) qreltrsLL"
+  | Infinite_DPP "('f, 'l, 'v) dppLL"
+  | Not_SN_FP_TRS "('f, 'l, 'v) fptrsLL"
+
+declare [[bnf_timing]]
+
+datatype_new (discs_sels) ('f, 'l, 'v, 'a, 'b, 'c, 'd, 'e) generic_assm_proof =
+    SN_assm_proof "('f, 'l, 'v) qreltrsLL" 'a
+  | Finite_assm_proof "('f, 'l, 'v) dppLL" 'b
+  | SN_FP_assm_proof "('f, 'l, 'v) fptrsLL" 'c
+  | Not_SN_assm_proof "('f, 'l, 'v) qtrsLL" 'a
+  | Infinite_assm_proof "('f, 'l, 'v) dppLL" 'b
+  | Not_RelSN_assm_proof "('f, 'l, 'v) qreltrsLL" 'c
+  | Not_SN_FP_assm_proof "('f, 'l, 'v) fptrsLL" 'd
+  | Unknown_assm_proof unknown_info 'e
+
+type_synonym ('f, 'l, 'v, 'a, 'b, 'c, 'd) assm_proof = "('f, 'l, 'v, 'a, 'b, 'c, dummy_prf, 'd) generic_assm_proof"
+
+datatype_new (discs_sels) ('f, 'l, 'v) assm =
+    SN_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) qreltrsLL"
+  | SN_FP_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) fptrsLL"
+  | Finite_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) dppLL"
+  | Unknown_assm "('f, 'l, 'v) problem list" unknown_info
+  | Not_SN_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) qtrsLL"
+  | Not_RelSN_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) qreltrsLL"
+  | Not_SN_FP_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) fptrsLL"
+  | Infinite_assm "('f, 'l, 'v) problem list" "('f, 'l, 'v) dppLL"
+
+fun satisfied :: "('f, 'l, 'v) problem \<Rightarrow> bool" where
+  "satisfied (SN_TRS t) = (t = t)"
+| "satisfied (SN_FP_TRS t) = (t \<noteq> t)"
+| "satisfied (Finite_DPP d) = (d \<noteq> d)"
+| "satisfied (Unknown_Problem s) = (s = ''foo'')"
+| "satisfied (Not_SN_TRS (nfs, q, r)) = (length q = length r)"
+| "satisfied (Not_RelSN_TRS (nfs, q, r, rw)) = (r = rw)"
+| "satisfied (Infinite_DPP d) = (d = d)"
+| "satisfied (Not_SN_FP_TRS t) = (t = t)"
+
+fun collect_assms :: "('tp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('dpp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('fptp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('unk \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('f, 'l, 'v, 'tp, 'dpp, 'fptp, 'unk) assm_proof \<Rightarrow> ('f, 'l, 'v) assm list" where
+  "collect_assms tp dpp fptp unk (SN_assm_proof t prf) = tp prf"
+| "collect_assms tp dpp fptp unk (SN_FP_assm_proof t prf) = fptp prf"
+| "collect_assms tp dpp fptp unk (Finite_assm_proof d prf) = dpp prf"
+| "collect_assms tp dpp fptp unk (Unknown_assm_proof p prf) = unk prf"
+| "collect_assms _ _ _ _ _ = []"
+
+fun collect_neg_assms :: "('tp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('dpp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('rtp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('fptp \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('unk \<Rightarrow> ('f, 'l, 'v) assm list)
+  \<Rightarrow> ('f, 'l, 'v, 'tp, 'dpp, 'rtp, 'fptp, 'unk) generic_assm_proof \<Rightarrow> ('f, 'l, 'v) assm list" where
+  "collect_neg_assms tp dpp rtp fptp unk (Not_SN_assm_proof t prf) = tp prf"
+| "collect_neg_assms tp dpp rtp fptp unk (Infinite_assm_proof d prf) = dpp prf"
+| "collect_neg_assms tp dpp rtp fptp unk (Not_RelSN_assm_proof t prf) = rtp prf"
+| "collect_neg_assms tp dpp rtp fptp unk (Not_SN_FP_assm_proof t prf) = fptp prf"
+| "collect_neg_assms tp dpp rtp fptp unk (Unknown_assm_proof p prf) = unk prf"
+| "collect_neg_assms tp dpp rtp fptp unk _ = []"
+
+datatype_new (discs_sels) ('f, 'l, 'v) dp_nontermination_proof =
+    DP_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) prseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
+  | DP_Nonloop "(('f, 'l) lab, 'v) non_loop_prf"
+  | DP_Rule_Removal "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Q_Increase "('f, 'l, 'v) termsLL" "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Q_Reduction "('f, 'l, 'v) termsLL" "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Termination_Switch "('f, 'l) lab join_info" "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Instantiation "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Rewriting "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL" "(('f, 'l) lab, 'v) rule" pos "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Narrowing "('f, 'l, 'v) ruleLL" pos "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_nontermination_proof"
+  | DP_Assume_Infinite  "('f, 'l, 'v) dppLL"
+      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
+       ('f, 'l, 'v) dp_nontermination_proof,
+       ('f, 'l, 'v) reltrs_nontermination_proof,
+       ('f, 'l, 'v) fp_nontermination_proof,
+       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
+and ('f, 'l, 'v) "trs_nontermination_proof" =
+    TRS_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) rseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
+  | TRS_Not_Well_Formed
+  | TRS_Rule_Removal "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_nontermination_proof"
+  | TRS_String_Reversal "('f, 'l, 'v) trs_nontermination_proof"
+  | TRS_DP_Trans "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_nontermination_proof"
+  | TRS_Nonloop "(('f, 'l) lab, 'v) non_loop_prf"
+  | TRS_Q_Increase "('f, 'l, 'v) termsLL" "('f, 'l, 'v) trs_nontermination_proof"
+  | TRS_Assume_Not_SN  "('f, 'l, 'v) qtrsLL"
+      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
+       ('f, 'l, 'v) dp_nontermination_proof,
+       ('f, 'l, 'v) reltrs_nontermination_proof,
+       ('f, 'l, 'v) fp_nontermination_proof,
+       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
+and ('f, 'l, 'v)"reltrs_nontermination_proof" =
+    Rel_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) prseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
+  | Rel_Not_Well_Formed
+  | Rel_Rule_Removal "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) reltrs_nontermination_proof"
+  | Rel_R_Not_SN "('f, 'l, 'v) trs_nontermination_proof"
+  | Rel_TRS_Assume_Not_SN  "('f, 'l, 'v) qreltrsLL"
+      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
+       ('f, 'l, 'v) dp_nontermination_proof,
+       ('f, 'l, 'v) reltrs_nontermination_proof,
+       ('f, 'l, 'v) fp_nontermination_proof,
+       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
+and ('f, 'l, 'v) "fp_nontermination_proof" =
+    FPTRS_Loop "(('f, 'l) lab, 'v) term" "(('f, 'l) lab, 'v) rseq" "(('f, 'l) lab, 'v) substL" "(('f, 'l) lab, 'v) ctxt"
+  | FPTRS_Rule_Removal "('f, 'l, 'v) trsLL" "('f, 'l, 'v) fp_nontermination_proof"
+  | FPTRS_Assume_Not_SN  "('f, 'l, 'v) fptrsLL"
+      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
+       ('f, 'l, 'v) dp_nontermination_proof,
+       ('f, 'l, 'v) reltrs_nontermination_proof,
+       ('f, 'l, 'v) fp_nontermination_proof,
+       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
+and ('f, 'l, 'v) neg_unknown_proof =
+    Assume_NT_Unknown unknown_info
+      "('f, 'l, 'v, ('f, 'l, 'v) trs_nontermination_proof,
+       ('f, 'l, 'v) dp_nontermination_proof,
+       ('f, 'l, 'v) reltrs_nontermination_proof,
+       ('f, 'l, 'v) fp_nontermination_proof,
+       ('f, 'l, 'v) neg_unknown_proof) generic_assm_proof list"
+
+datatype_new (discs_sels) ('f, 'l, 'v) dp_termination_proof =
+    P_is_Empty
+  | Subterm_Criterion_Proc "('f, 'l) lab projL" "('f, 'l, 'v) rseqL"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Redpair_Proc "('f, 'l) lab root_redtriple_impl + ('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL"  "('f, 'l, 'v) dp_termination_proof"
+  | Redpair_UR_Proc "('f, 'l) lab root_redtriple_impl + ('f, 'l) lab redtriple_impl"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Usable_Rules_Proc "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Dep_Graph_Proc "(('f, 'l, 'v) dp_termination_proof option \<times> ('f, 'l, 'v) trsLL) list"
+  | Mono_Redpair_Proc "('f, 'l) lab redtriple_impl"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Mono_Redpair_UR_Proc "('f, 'l) lab redtriple_impl"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Size_Change_Subterm_Proc "((('f, 'l) lab, 'v) rule \<times> ((nat \<times> nat) list \<times> (nat \<times> nat) list)) list"
+  | Size_Change_Redpair_Proc "('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL option"
+      "((('f, 'l) lab, 'v) rule \<times> ((nat \<times> nat) list \<times> (nat \<times> nat) list)) list"
+  | Uncurry_Proc "nat option" "(('f, 'l) lab, 'v) uncurry_info"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Fcc_Proc "('f, 'l) lab" "(('f, 'l) lab, 'v) ctxt list"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Split_Proc
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL"
+      "('f, 'l, 'v) dp_termination_proof" "('f, 'l, 'v) dp_termination_proof"
+  | Semlab_Proc
+      "(('f, 'l) lab, 'v) sl_variant" "('f, 'l, 'v) trsLL"
+      "(('f, 'l) lab, 'v) term list" "('f, 'l, 'v) trsLL"
+      "('f, 'l, 'v) dp_termination_proof"
+  | Switch_Innermost_Proc "('f, 'l) lab join_info" "('f, 'l, 'v) dp_termination_proof"
+  | Rewriting_Proc
+      "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL"
+      "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) ruleLL" pos "('f, 'l, 'v) dp_termination_proof"
+  | Instantiation_Proc "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Forward_Instantiation_Proc
+      "('f, 'l, 'v) ruleLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL option" "('f, 'l, 'v) dp_termination_proof"
+  | Narrowing_Proc "('f, 'l, 'v) ruleLL" pos "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Assume_Finite
+      "('f, 'l, 'v) dppLL" "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
+  | Unlab_Proc "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Q_Reduction_Proc "('f, 'l, 'v) termsLL" "('f, 'l, 'v) dp_termination_proof"
+  | Complex_Constant_Removal_Proc "(('f, 'l) lab, 'v) complex_constant_removal_prf" "('f, 'l, 'v) dp_termination_proof"
+  | General_Redpair_Proc
+      "('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trsLL"
+      "(('f, 'l) lab, 'v) cond_red_pair_prf" "('f, 'l, 'v) dp_termination_proof list"
+  | To_Trs_Proc "('f, 'l, 'v) trs_termination_proof"
+and ('f, 'l, 'v) trs_termination_proof =
+    DP_Trans bool bool "(('f, 'l) lab, 'v) rules" "('f, 'l, 'v) dp_termination_proof"
+  | Rule_Removal "('f, 'l) lab redtriple_impl" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
+  | String_Reversal "('f, 'l, 'v) trs_termination_proof"
+  | Bounds "(('f, 'l) lab, 'v) bounds_info"
+  | Uncurry "(('f, 'l) lab, 'v) uncurry_info" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
+  | Semlab
+      "(('f, 'l) lab, 'v) sl_variant" "(('f, 'l) lab, 'v) term list"
+      "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
+  | R_is_Empty
+  | Fcc "(('f, 'l) lab, 'v) ctxt list" "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
+  | Split "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof" "('f, 'l, 'v) trs_termination_proof"
+  | Switch_Innermost "('f, 'l) lab join_info" "('f, 'l, 'v) trs_termination_proof"
+  | Drop_Equality "('f, 'l, 'v) trs_termination_proof"
+  | Remove_Nonapplicable_Rules "('f, 'l, 'v) trsLL" "('f, 'l, 'v) trs_termination_proof"
+  | Assume_SN "('f, 'l, 'v) qreltrsLL" "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
+and ('f, 'l, 'v) unknown_proof =
+    Assume_Unknown unknown_info "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
+and ('f, 'l, 'v) fptrs_termination_proof =
+    Assume_FP_SN "('f, 'l, 'v) fptrsLL" "('f, 'l, 'v, ('f, 'l, 'v) trs_termination_proof, ('f, 'l, 'v) dp_termination_proof, ('f, 'l, 'v) fptrs_termination_proof, ('f, 'l, 'v) unknown_proof) assm_proof list"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Koenig.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,122 @@
+(*  Title:      HOL/Datatype_Examples/Koenig.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Koenig's lemma.
+*)
+
+header {* Koenig's Lemma *}
+
+theory Koenig
+imports TreeFI Stream
+begin
+
+(* infinite trees: *)
+coinductive infiniteTr where
+"\<lbrakk>tr' \<in> set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
+
+lemma infiniteTr_strong_coind[consumes 1, case_names sub]:
+assumes *: "phi tr" and
+**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> set (sub tr). phi tr' \<or> infiniteTr tr'"
+shows "infiniteTr tr"
+using assms by (elim infiniteTr.coinduct) blast
+
+lemma infiniteTr_coind[consumes 1, case_names sub, induct pred: infiniteTr]:
+assumes *: "phi tr" and
+**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> set (sub tr). phi tr'"
+shows "infiniteTr tr"
+using assms by (elim infiniteTr.coinduct) blast
+
+lemma infiniteTr_sub[simp]:
+"infiniteTr tr \<Longrightarrow> (\<exists> tr' \<in> set (sub tr). infiniteTr tr')"
+by (erule infiniteTr.cases) blast
+
+primcorec konigPath where
+  "shd (konigPath t) = lab t"
+| "stl (konigPath t) = konigPath (SOME tr. tr \<in> set (sub t) \<and> infiniteTr tr)"
+
+(* proper paths in trees: *)
+coinductive properPath where
+"\<lbrakk>shd as = lab tr; tr' \<in> set (sub tr); properPath (stl as) tr'\<rbrakk> \<Longrightarrow>
+ properPath as tr"
+
+lemma properPath_strong_coind[consumes 1, case_names shd_lab sub]:
+assumes *: "phi as tr" and
+**: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and
+***: "\<And> as tr.
+         phi as tr \<Longrightarrow>
+         \<exists> tr' \<in> set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
+shows "properPath as tr"
+using assms by (elim properPath.coinduct) blast
+
+lemma properPath_coind[consumes 1, case_names shd_lab sub, induct pred: properPath]:
+assumes *: "phi as tr" and
+**: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and
+***: "\<And> as tr.
+         phi as tr \<Longrightarrow>
+         \<exists> tr' \<in> set (sub tr). phi (stl as) tr'"
+shows "properPath as tr"
+using properPath_strong_coind[of phi, OF * **] *** by blast
+
+lemma properPath_shd_lab:
+"properPath as tr \<Longrightarrow> shd as = lab tr"
+by (erule properPath.cases) blast
+
+lemma properPath_sub:
+"properPath as tr \<Longrightarrow>
+ \<exists> tr' \<in> set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
+by (erule properPath.cases) blast
+
+(* prove the following by coinduction *)
+theorem Konig:
+  assumes "infiniteTr tr"
+  shows "properPath (konigPath tr) tr"
+proof-
+  {fix as
+   assume "infiniteTr tr \<and> as = konigPath tr" hence "properPath as tr"
+   proof (coinduction arbitrary: tr as rule: properPath_coind)
+     case (sub tr as)
+     let ?t = "SOME t'. t' \<in> set (sub tr) \<and> infiniteTr t'"
+     from sub have "\<exists>t' \<in> set (sub tr). infiniteTr t'" by simp
+     then have "\<exists>t'. t' \<in> set (sub tr) \<and> infiniteTr t'" by blast
+     then have "?t \<in> set (sub tr) \<and> infiniteTr ?t" by (rule someI_ex)
+     moreover have "stl (konigPath tr) = konigPath ?t" by simp
+     ultimately show ?case using sub by blast
+   qed simp
+  }
+  thus ?thesis using assms by blast
+qed
+
+(* some more stream theorems *)
+
+primcorec plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<oplus>" 66) where
+  "shd (plus xs ys) = shd xs + shd ys"
+| "stl (plus xs ys) = plus (stl xs) (stl ys)"
+
+definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<cdot>" 68) where
+  [simp]: "scalar n = smap (\<lambda>x. n * x)"
+
+primcorec ones :: "nat stream" where "ones = 1 ## ones"
+primcorec twos :: "nat stream" where "twos = 2 ## twos"
+definition ns :: "nat \<Rightarrow> nat stream" where [simp]: "ns n = scalar n ones"
+
+lemma "ones \<oplus> ones = twos"
+  by coinduction simp
+
+lemma "n \<cdot> twos = ns (2 * n)"
+  by coinduction simp
+
+lemma prod_scalar: "(n * m) \<cdot> xs = n \<cdot> m \<cdot> xs"
+  by (coinduction arbitrary: xs) auto
+
+lemma scalar_plus: "n \<cdot> (xs \<oplus> ys) = n \<cdot> xs \<oplus> n \<cdot> ys"
+  by (coinduction arbitrary: xs ys) (auto simp: add_mult_distrib2)
+
+lemma plus_comm: "xs \<oplus> ys = ys \<oplus> xs"
+  by (coinduction arbitrary: xs ys) auto
+
+lemma plus_assoc: "(xs \<oplus> ys) \<oplus> zs = xs \<oplus> ys \<oplus> zs"
+  by (coinduction arbitrary: xs ys zs) auto
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Lambda_Term.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,50 @@
+(*  Title:      HOL/Datatype_Examples/Lambda_Term.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Lambda-terms.
+*)
+
+header {* Lambda-Terms *}
+
+theory Lambda_Term
+imports "~~/src/HOL/Library/FSet"
+begin
+
+section {* Datatype definition *}
+
+datatype_new 'a trm =
+  Var 'a |
+  App "'a trm" "'a trm" |
+  Lam 'a "'a trm" |
+  Lt "('a \<times> 'a trm) fset" "'a trm"
+
+
+subsection {* Example: The set of all variables varsOf and free variables fvarsOf of a term *}
+
+primrec varsOf :: "'a trm \<Rightarrow> 'a set" where
+  "varsOf (Var a) = {a}"
+| "varsOf (App f x) = varsOf f \<union> varsOf x"
+| "varsOf (Lam x b) = {x} \<union> varsOf b"
+| "varsOf (Lt F t) = varsOf t \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| fimage (map_prod id varsOf) F})"
+
+primrec fvarsOf :: "'a trm \<Rightarrow> 'a set" where
+  "fvarsOf (Var x) = {x}"
+| "fvarsOf (App t1 t2) = fvarsOf t1 \<union> fvarsOf t2"
+| "fvarsOf (Lam x t) = fvarsOf t - {x}"
+| "fvarsOf (Lt xts t) = fvarsOf t - {x | x X. (x,X) |\<in>| fimage (map_prod id varsOf) xts} \<union>
+    (\<Union> {X | x X. (x,X) |\<in>| fimage (map_prod id varsOf) xts})"
+
+lemma diff_Un_incl_triv: "\<lbrakk>A \<subseteq> D; C \<subseteq> E\<rbrakk> \<Longrightarrow> A - B \<union> C \<subseteq> D \<union> E" by blast
+
+lemma in_fimage_map_prod_fset_iff[simp]:
+  "(x, y) |\<in>| fimage (map_prod f g) xts \<longleftrightarrow> (\<exists> t1 t2. (t1, t2) |\<in>| xts \<and> x = f t1 \<and> y = g t2)"
+  by force
+
+lemma fvarsOf_varsOf: "fvarsOf t \<subseteq> varsOf t"
+proof induct
+  case (Lt xts t) thus ?case unfolding fvarsOf.simps varsOf.simps by (elim diff_Un_incl_triv) auto
+qed auto
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Misc_Codatatype.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,122 @@
+(*  Title:      HOL/Datatype_Examples/Misc_Codatatype.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Miscellaneous codatatype definitions.
+*)
+
+header {* Miscellaneous Codatatype Definitions *}
+
+theory Misc_Codatatype
+imports "~~/src/HOL/Library/FSet"
+begin
+
+codatatype simple = X1 | X2 | X3 | X4
+
+codatatype simple' = X1' unit | X2' unit | X3' unit | X4' unit
+
+codatatype simple'' = X1'' nat int | X2''
+
+codatatype 'a stream = Stream (shd: 'a) (stl: "'a stream")
+
+codatatype 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
+
+codatatype ('b, 'c :: ord, 'd, 'e) some_passive =
+  SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
+
+codatatype lambda =
+  Var string |
+  App lambda lambda |
+  Abs string lambda |
+  Let "(string \<times> lambda) fset" lambda
+
+codatatype 'a par_lambda =
+  PVar 'a |
+  PApp "'a par_lambda" "'a par_lambda" |
+  PAbs 'a "'a par_lambda" |
+  PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
+
+(*
+  ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
+  ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
+*)
+
+codatatype 'a p = P "'a + 'a p"
+
+codatatype 'a J1 = J11 'a "'a J1" | J12 'a "'a J2"
+and 'a J2 = J21 | J22 "'a J1" "'a J2"
+
+codatatype 'a tree = TEmpty | TNode 'a "'a forest"
+and 'a forest = FNil | FCons "'a tree" "'a forest"
+
+codatatype 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
+and 'a branch = Branch 'a "'a tree'"
+
+codatatype 'a bin_rose_tree = BRTree 'a "'a bin_rose_tree mylist" "'a bin_rose_tree mylist"
+
+codatatype ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
+and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
+and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
+
+codatatype ('a, 'b, 'c) some_killing =
+  SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
+and ('a, 'b, 'c) in_here =
+  IH1 'b 'a | IH2 'c
+
+codatatype ('a, 'b, 'c) some_killing' =
+  SK' "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing' + ('a, 'b, 'c) in_here'"
+and ('a, 'b, 'c) in_here' =
+  IH1' 'b | IH2' 'c
+
+codatatype ('a, 'b, 'c) some_killing'' =
+  SK'' "'a \<Rightarrow> ('a, 'b, 'c) in_here''"
+and ('a, 'b, 'c) in_here'' =
+  IH1'' 'b 'a | IH2'' 'c
+
+codatatype ('b, 'c) less_killing = LK "'b \<Rightarrow> 'c"
+
+codatatype 'b poly_unit = U "'b \<Rightarrow> 'b poly_unit"
+codatatype 'b cps = CPS1 'b | CPS2 "'b \<Rightarrow> 'b cps"
+
+codatatype ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs =
+  FR "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow>
+      ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs"
+
+codatatype ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
+        'b18, 'b19, 'b20) fun_rhs' =
+  FR' "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> 'b10 \<Rightarrow> 'b11 \<Rightarrow> 'b12 \<Rightarrow> 'b13 \<Rightarrow> 'b14 \<Rightarrow>
+       'b15 \<Rightarrow> 'b16 \<Rightarrow> 'b17 \<Rightarrow> 'b18 \<Rightarrow> 'b19 \<Rightarrow> 'b20 \<Rightarrow>
+       ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
+        'b18, 'b19, 'b20) fun_rhs'"
+
+codatatype ('a, 'b, 'c) wit3_F1 = W1 'a "('a, 'b, 'c) wit3_F1" "('a, 'b, 'c) wit3_F2"
+and ('a, 'b, 'c) wit3_F2 = W2 'b "('a, 'b, 'c) wit3_F2"
+and ('a, 'b, 'c) wit3_F3 = W31 'a 'b "('a, 'b, 'c) wit3_F1" | W32 'c 'a 'b "('a, 'b, 'c) wit3_F1"
+
+codatatype ('c, 'e, 'g) coind_wit1 =
+       CW1 'c "('c, 'e, 'g) coind_wit1" "('c, 'e, 'g) ind_wit" "('c, 'e, 'g) coind_wit2"
+and ('c, 'e, 'g) coind_wit2 =
+       CW21 "('c, 'e, 'g) coind_wit2" 'e | CW22 'c 'g
+and ('c, 'e, 'g) ind_wit =
+       IW1 | IW2 'c
+
+codatatype ('b, 'a) bar = BAR "'a \<Rightarrow> 'b"
+codatatype ('a, 'b, 'c, 'd) foo = FOO "'d + 'b \<Rightarrow> 'c + 'a"
+
+codatatype 'a dead_foo = A
+codatatype ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
+
+(* SLOW, MEMORY-HUNGRY
+codatatype ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
+and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
+and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
+and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
+and ('a, 'c) D5 = A5 "('a, 'c) D6"
+and ('a, 'c) D6 = A6 "('a, 'c) D7"
+and ('a, 'c) D7 = A7 "('a, 'c) D8"
+and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Misc_Datatype.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,338 @@
+(*  Title:      HOL/Datatype_Examples/Misc_Datatype.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Miscellaneous datatype definitions.
+*)
+
+header {* Miscellaneous Datatype Definitions *}
+
+theory Misc_Datatype
+imports "~~/src/HOL/Library/Countable" "~~/src/HOL/Library/FSet"
+begin
+
+datatype_new (discs_sels) simple = X1 | X2 | X3 | X4
+
+datatype_new (discs_sels) simple' = X1' unit | X2' unit | X3' unit | X4' unit
+
+datatype_new (discs_sels) simple'' = X1'' nat int | X2''
+
+datatype_new (discs_sels) 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
+
+datatype_new (discs_sels) ('b, 'c :: ord, 'd, 'e) some_passive =
+  SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
+
+datatype_new (discs_sels) hfset = HFset "hfset fset"
+
+datatype_new (discs_sels) lambda =
+  Var string |
+  App lambda lambda |
+  Abs string lambda |
+  Let "(string \<times> lambda) fset" lambda
+
+datatype_new (discs_sels) 'a par_lambda =
+  PVar 'a |
+  PApp "'a par_lambda" "'a par_lambda" |
+  PAbs 'a "'a par_lambda" |
+  PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
+
+(*
+  ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
+  ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
+*)
+
+datatype_new (discs_sels) 'a I1 = I11 'a "'a I1" | I12 'a "'a I2"
+and 'a I2 = I21 | I22 "'a I1" "'a I2"
+
+datatype_new (discs_sels) 'a tree = TEmpty | TNode 'a "'a forest"
+and 'a forest = FNil | FCons "'a tree" "'a forest"
+
+datatype_new (discs_sels) 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
+and 'a branch = Branch 'a "'a tree'"
+
+datatype_new (discs_sels) 'a bin_rose_tree = BRTree 'a "'a bin_rose_tree mylist" "'a bin_rose_tree mylist"
+
+datatype_new (discs_sels) ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
+and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
+and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
+
+datatype_new (discs_sels) 'a ftree = FTLeaf 'a | FTNode "'a \<Rightarrow> 'a ftree"
+
+datatype_new (discs_sels) ('a, 'b, 'c) some_killing =
+  SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
+and ('a, 'b, 'c) in_here =
+  IH1 'b 'a | IH2 'c
+
+datatype_new (discs_sels) 'b nofail1 = NF11 "'b nofail1" 'b | NF12 'b
+datatype_new (discs_sels) 'b nofail2 = NF2 "('b nofail2 \<times> 'b \<times> 'b nofail2 \<times> 'b) list"
+datatype_new (discs_sels) 'b nofail3 = NF3 'b "('b nofail3 \<times> 'b \<times> 'b nofail3 \<times> 'b) fset"
+datatype_new (discs_sels) 'b nofail4 = NF4 "('b nofail4 \<times> ('b nofail4 \<times> 'b \<times> 'b nofail4 \<times> 'b) fset) list"
+
+(*
+datatype_new (discs_sels) 'b fail = F "'b fail" 'b "'b fail" "'b list"
+datatype_new (discs_sels) 'b fail = F "'b fail" 'b "'b fail" 'b
+datatype_new (discs_sels) 'b fail = F1 "'b fail" 'b | F2 "'b fail"
+datatype_new (discs_sels) 'b fail = F "'b fail" 'b
+*)
+
+datatype_new (discs_sels) l1 = L1 "l2 list"
+and l2 = L21 "l1 fset" | L22 l2
+
+datatype_new (discs_sels) kk1 = KK1 kk2
+and kk2 = KK2 kk3
+and kk3 = KK3 "kk1 list"
+
+datatype_new (discs_sels) t1 = T11 t3 | T12 t2
+and t2 = T2 t1
+and t3 = T3
+
+datatype_new (discs_sels) t1' = T11' t2' | T12' t3'
+and t2' = T2' t1'
+and t3' = T3'
+
+(*
+datatype_new (discs_sels) fail1 = F1 fail2
+and fail2 = F2 fail3
+and fail3 = F3 fail1
+
+datatype_new (discs_sels) fail1 = F1 "fail2 list" fail2
+and fail2 = F2 "fail2 fset" fail3
+and fail3 = F3 fail1
+
+datatype_new (discs_sels) fail1 = F1 "fail2 list" fail2
+and fail2 = F2 "fail1 fset" fail1
+*)
+
+(* SLOW
+datatype_new (discs_sels) ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
+and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
+and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
+and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
+and ('a, 'c) D5 = A5 "('a, 'c) D6"
+and ('a, 'c) D6 = A6 "('a, 'c) D7"
+and ('a, 'c) D7 = A7 "('a, 'c) D8"
+and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
+*)
+
+(* fail:
+datatype_new (discs_sels) tt1 = TT11 tt2 tt3 | TT12 tt2 tt4
+and tt2 = TT2
+and tt3 = TT3 tt4
+and tt4 = TT4 tt1
+*)
+
+datatype_new (discs_sels) k1 = K11 k2 k3 | K12 k2 k4
+and k2 = K2
+and k3 = K3 k4
+and k4 = K4
+
+datatype_new (discs_sels) tt1 = TT11 tt3 tt2 | TT12 tt2 tt4
+and tt2 = TT2
+and tt3 = TT3 tt1
+and tt4 = TT4
+
+(* SLOW
+datatype_new (discs_sels) s1 = S11 s2 s3 s4 | S12 s3 | S13 s2 s6 | S14 s4 s2 | S15 s2 s2
+and s2 = S21 s7 s5 | S22 s5 s4 s6
+and s3 = S31 s1 s7 s2 | S32 s3 s3 | S33 s4 s5
+and s4 = S4 s5
+and s5 = S5
+and s6 = S61 s6 | S62 s1 s2 | S63 s6
+and s7 = S71 s8 | S72 s5
+and s8 = S8 nat
+*)
+
+datatype_new (discs_sels) 'a deadbar = DeadBar "'a \<Rightarrow> 'a"
+datatype_new (discs_sels) 'a deadbar_option = DeadBarOption "'a option \<Rightarrow> 'a option"
+datatype_new (discs_sels) ('a, 'b) bar = Bar "'b \<Rightarrow> 'a"
+datatype_new (discs_sels) ('a, 'b, 'c, 'd) foo = Foo "'d + 'b \<Rightarrow> 'c + 'a"
+datatype_new (discs_sels) 'a deadfoo = DeadFoo "'a \<Rightarrow> 'a + 'a"
+
+datatype_new (discs_sels) 'a dead_foo = A
+datatype_new (discs_sels) ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
+
+datatype_new (discs_sels) d1 = D
+datatype_new (discs_sels) d1' = is_D: D
+
+datatype_new (discs_sels) d2 = D nat
+datatype_new (discs_sels) d2' = is_D: D nat
+
+datatype_new (discs_sels) d3 = D | E
+datatype_new (discs_sels) d3' = D | is_E: E
+datatype_new (discs_sels) d3'' = is_D: D | E
+datatype_new (discs_sels) d3''' = is_D: D | is_E: E
+
+datatype_new (discs_sels) d4 = D nat | E
+datatype_new (discs_sels) d4' = D nat | is_E: E
+datatype_new (discs_sels) d4'' = is_D: D nat | E
+datatype_new (discs_sels) d4''' = is_D: D nat | is_E: E
+
+datatype_new (discs_sels) d5 = D nat | E int
+datatype_new (discs_sels) d5' = D nat | is_E: E int
+datatype_new (discs_sels) d5'' = is_D: D nat | E int
+datatype_new (discs_sels) d5''' = is_D: D nat | is_E: E int
+
+instance simple :: countable
+  by countable_datatype
+
+instance simple'' :: countable
+  by countable_datatype
+
+instance mylist :: (countable) countable
+  by countable_datatype
+
+instance some_passive :: (countable, "{countable,ord}", countable, countable) countable
+  by countable_datatype
+
+(* TODO: Enable once "fset" is registered as countable:
+
+instance hfset :: countable
+  by countable_datatype
+
+instance lambda :: countable
+  by countable_datatype
+
+instance par_lambda :: (countable) countable
+  by countable_datatype
+*)
+
+instance I1 and I2 :: (countable) countable
+  by countable_datatype
+
+instance tree and forest :: (countable) countable
+  by countable_datatype
+
+instance tree' and branch :: (countable) countable
+  by countable_datatype
+
+instance bin_rose_tree :: (countable) countable
+  by countable_datatype
+
+instance exp and trm and factor :: (countable, countable) countable
+  by countable_datatype
+
+instance nofail1 :: (countable) countable
+  by countable_datatype
+
+instance nofail2 :: (countable) countable
+  by countable_datatype
+
+(* TODO: Enable once "fset" is registered as countable:
+
+instance nofail3 :: (countable) countable
+  by countable_datatype
+
+instance nofail4 :: (countable) countable
+  by countable_datatype
+
+instance l1 and l2 :: countable
+  by countable_datatype
+*)
+
+instance kk1 and kk2 :: countable
+  by countable_datatype
+
+instance t1 and t2 and t3 :: countable
+  by countable_datatype
+
+instance t1' and t2' and t3' :: countable
+  by countable_datatype
+
+instance k1 and k2 and k3 and k4 :: countable
+  by countable_datatype
+
+instance tt1 and tt2 and tt3 and tt4 :: countable
+  by countable_datatype
+
+instance d1 :: countable
+  by countable_datatype
+
+instance d1' :: countable
+  by countable_datatype
+
+instance d2 :: countable
+  by countable_datatype
+
+instance d2' :: countable
+  by countable_datatype
+
+instance d3 :: countable
+  by countable_datatype
+
+instance d3' :: countable
+  by countable_datatype
+
+instance d3'' :: countable
+  by countable_datatype
+
+instance d3''' :: countable
+  by countable_datatype
+
+instance d4 :: countable
+  by countable_datatype
+
+instance d4' :: countable
+  by countable_datatype
+
+instance d4'' :: countable
+  by countable_datatype
+
+instance d4''' :: countable
+  by countable_datatype
+
+instance d5 :: countable
+  by countable_datatype
+
+instance d5' :: countable
+  by countable_datatype
+
+instance d5'' :: countable
+  by countable_datatype
+
+instance d5''' :: countable
+  by countable_datatype
+
+datatype_compat simple
+datatype_compat simple'
+datatype_compat simple''
+datatype_compat mylist
+datatype_compat some_passive
+datatype_compat I1 I2
+datatype_compat tree forest
+datatype_compat tree' branch
+datatype_compat bin_rose_tree
+datatype_compat exp trm factor
+datatype_compat ftree
+datatype_compat nofail1
+datatype_compat kk1 kk2 kk3
+datatype_compat t1 t2 t3
+datatype_compat t1' t2' t3'
+datatype_compat k1 k2 k3 k4
+datatype_compat tt1 tt2 tt3 tt4
+datatype_compat deadbar
+datatype_compat deadbar_option
+datatype_compat bar
+datatype_compat foo
+datatype_compat deadfoo
+datatype_compat dead_foo
+datatype_compat use_dead_foo
+datatype_compat d1
+datatype_compat d1'
+datatype_compat d2
+datatype_compat d2'
+datatype_compat d3
+datatype_compat d3'
+datatype_compat d3''
+datatype_compat d3'''
+datatype_compat d4
+datatype_compat d4'
+datatype_compat d4''
+datatype_compat d4'''
+datatype_compat d5
+datatype_compat d5'
+datatype_compat d5''
+datatype_compat d5'''
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Misc_Primcorec.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,138 @@
+(*  Title:      HOL/Datatype_Examples/Misc_Primcorec.thy
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Miscellaneous primitive corecursive function definitions.
+*)
+
+header {* Miscellaneous Primitive Corecursive Function Definitions *}
+
+theory Misc_Primcorec
+imports Misc_Codatatype
+begin
+
+primcorec simple_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple" where
+  "simple_of_bools b b' = (if b then if b' then X1 else X2 else if b' then X3 else X4)"
+
+primcorec simple'_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple'" where
+  "simple'_of_bools b b' =
+     (if b then if b' then X1' () else X2' () else if b' then X3' () else X4' ())"
+
+primcorec inc_simple'' :: "nat \<Rightarrow> simple'' \<Rightarrow> simple''" where
+  "inc_simple'' k s = (case s of X1'' n i \<Rightarrow> X1'' (n + k) (i + int k) | X2'' \<Rightarrow> X2'')"
+
+primcorec sinterleave :: "'a stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
+  "sinterleave s s' = Stream (shd s) (sinterleave s' (stl s))"
+
+primcorec myapp :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" where
+  "myapp xs ys =
+     (if xs = MyNil then ys
+      else if ys = MyNil then xs
+      else MyCons (myhd xs) (myapp (mytl xs) ys))"
+
+primcorec shuffle_sp :: "('a \<Colon> ord, 'b \<Colon> ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
+  "shuffle_sp sp =
+     (case sp of
+       SP1 sp' \<Rightarrow> SP1 (shuffle_sp sp')
+     | SP2 a \<Rightarrow> SP3 a
+     | SP3 b \<Rightarrow> SP4 b
+     | SP4 c \<Rightarrow> SP5 c
+     | SP5 d \<Rightarrow> SP2 d)"
+
+primcorec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
+  "rename_lam f l =
+     (case l of
+       Var s \<Rightarrow> Var (f s)
+     | App l l' \<Rightarrow> App (rename_lam f l) (rename_lam f l')
+     | Abs s l \<Rightarrow> Abs (f s) (rename_lam f l)
+     | Let SL l \<Rightarrow> Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l))"
+
+primcorec
+  j1_sum :: "('a\<Colon>{zero,one,plus}) \<Rightarrow> 'a J1" and
+  j2_sum :: "'a \<Rightarrow> 'a J2"
+where
+  "n = 0 \<Longrightarrow> is_J11 (j1_sum n)" |
+  "un_J111 (j1_sum _) = 0" |
+  "un_J112 (j1_sum _) = j1_sum 0" |
+  "un_J121 (j1_sum n) = n + 1" |
+  "un_J122 (j1_sum n) = j2_sum (n + 1)" |
+  "n = 0 \<Longrightarrow> j2_sum n = J21" |
+  "un_J221 (j2_sum n) = j1_sum (n + 1)" |
+  "un_J222 (j2_sum n) = j2_sum (n + 1)"
+
+primcorec forest_of_mylist :: "'a tree mylist \<Rightarrow> 'a forest" where
+  "forest_of_mylist ts =
+     (case ts of
+       MyNil \<Rightarrow> FNil
+     | MyCons t ts \<Rightarrow> FCons t (forest_of_mylist ts))"
+
+primcorec mylist_of_forest :: "'a forest \<Rightarrow> 'a tree mylist" where
+  "mylist_of_forest f =
+     (case f of
+       FNil \<Rightarrow> MyNil
+     | FCons t ts \<Rightarrow> MyCons t (mylist_of_forest ts))"
+
+primcorec semi_stream :: "'a stream \<Rightarrow> 'a stream" where
+  "semi_stream s = Stream (shd s) (semi_stream (stl (stl s)))"
+
+primcorec
+  tree'_of_stream :: "'a stream \<Rightarrow> 'a tree'" and
+  branch_of_stream :: "'a stream \<Rightarrow> 'a branch"
+where
+  "tree'_of_stream s =
+     TNode' (branch_of_stream (semi_stream s)) (branch_of_stream (semi_stream (stl s)))" |
+  "branch_of_stream s = (case s of Stream h t \<Rightarrow> Branch h (tree'_of_stream t))"
+
+primcorec
+  id_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
+  id_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
+  id_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
+where
+  "id_tree t = (case t of BRTree a ts ts' \<Rightarrow> BRTree a (id_trees1 ts) (id_trees2 ts'))" |
+  "id_trees1 ts = (case ts of
+       MyNil \<Rightarrow> MyNil
+     | MyCons t ts \<Rightarrow> MyCons (id_tree t) (id_trees1 ts))" |
+  "id_trees2 ts = (case ts of
+       MyNil \<Rightarrow> MyNil
+     | MyCons t ts \<Rightarrow> MyCons (id_tree t) (id_trees2 ts))"
+
+primcorec
+  trunc_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
+  trunc_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
+  trunc_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
+where
+  "trunc_tree t = (case t of BRTree a ts ts' \<Rightarrow> BRTree a (trunc_trees1 ts) (trunc_trees2 ts'))" |
+  "trunc_trees1 ts = (case ts of
+       MyNil \<Rightarrow> MyNil
+     | MyCons t _ \<Rightarrow> MyCons (trunc_tree t) MyNil)" |
+  "trunc_trees2 ts = (case ts of
+       MyNil \<Rightarrow> MyNil
+     | MyCons t ts \<Rightarrow> MyCons (trunc_tree t) MyNil)"
+
+primcorec
+  freeze_exp :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) exp \<Rightarrow> ('a, 'b) exp" and
+  freeze_trm :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) trm \<Rightarrow> ('a, 'b) trm" and
+  freeze_factor :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) factor \<Rightarrow> ('a, 'b) factor"
+where
+  "freeze_exp g e =
+     (case e of
+       Term t \<Rightarrow> Term (freeze_trm g t)
+     | Sum t e \<Rightarrow> Sum (freeze_trm g t) (freeze_exp g e))" |
+  "freeze_trm g t =
+     (case t of
+       Factor f \<Rightarrow> Factor (freeze_factor g f)
+     | Prod f t \<Rightarrow> Prod (freeze_factor g f) (freeze_trm g t))" |
+  "freeze_factor g f =
+     (case f of
+       C a \<Rightarrow> C a
+     | V b \<Rightarrow> C (g b)
+     | Paren e \<Rightarrow> Paren (freeze_exp g e))"
+
+primcorec poly_unity :: "'a poly_unit" where
+  "poly_unity = U (\<lambda>_. poly_unity)"
+
+primcorec build_cps :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> bool stream) \<Rightarrow> 'a \<Rightarrow> bool stream \<Rightarrow> 'a cps" where
+  "shd b \<Longrightarrow> build_cps f g a b = CPS1 a" |
+  "_ \<Longrightarrow> build_cps f g a b = CPS2 (\<lambda>a. build_cps f g (f a) (g a))"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Misc_Primrec.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,136 @@
+(*  Title:      HOL/Datatype_Examples/Misc_Primrec.thy
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Miscellaneous primitive recursive function definitions.
+*)
+
+header {* Miscellaneous Primitive Recursive Function Definitions *}
+
+theory Misc_Primrec
+imports Misc_Datatype
+begin
+
+primrec nat_of_simple :: "simple \<Rightarrow> nat" where
+  "nat_of_simple X1 = 1" |
+  "nat_of_simple X2 = 2" |
+  "nat_of_simple X3 = 3" |
+  "nat_of_simple X4 = 4"
+
+primrec simple_of_simple' :: "simple' \<Rightarrow> simple" where
+  "simple_of_simple' (X1' _) = X1" |
+  "simple_of_simple' (X2' _) = X2" |
+  "simple_of_simple' (X3' _) = X3" |
+  "simple_of_simple' (X4' _) = X4"
+
+primrec inc_simple'' :: "nat \<Rightarrow> simple'' \<Rightarrow> simple''" where
+  "inc_simple'' k (X1'' n i) = X1'' (n + k) (i + int k)" |
+  "inc_simple'' _ X2'' = X2''"
+
+primrec myapp :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" where
+  "myapp MyNil ys = ys" |
+  "myapp (MyCons x xs) ys = MyCons x (myapp xs ys)"
+
+primrec myrev :: "'a mylist \<Rightarrow> 'a mylist" where
+  "myrev MyNil = MyNil" |
+  "myrev (MyCons x xs) = myapp (myrev xs) (MyCons x MyNil)"
+
+primrec shuffle_sp :: "('a \<Colon> ord, 'b \<Colon> ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
+  "shuffle_sp (SP1 sp) = SP1 (shuffle_sp sp)" |
+  "shuffle_sp (SP2 a) = SP3 a" |
+  "shuffle_sp (SP3 b) = SP4 b" |
+  "shuffle_sp (SP4 c) = SP5 c" |
+  "shuffle_sp (SP5 d) = SP2 d"
+
+primrec
+  hf_size :: "hfset \<Rightarrow> nat"
+where
+  "hf_size (HFset X) = 1 + setsum id (fset (fimage hf_size X))"
+
+primrec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
+  "rename_lam f (Var s) = Var (f s)" |
+  "rename_lam f (App l l') = App (rename_lam f l) (rename_lam f l')" |
+  "rename_lam f (Abs s l) = Abs (f s) (rename_lam f l)" |
+  "rename_lam f (Let SL l) = Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l)"
+
+primrec
+  sum_i1 :: "('a\<Colon>{zero,plus}) I1 \<Rightarrow> 'a" and
+  sum_i2 :: "'a I2 \<Rightarrow> 'a"
+where
+  "sum_i1 (I11 n i) = n + sum_i1 i" |
+  "sum_i1 (I12 n i) = n + sum_i2 i" |
+  "sum_i2 I21 = 0" |
+  "sum_i2 (I22 i j) = sum_i1 i + sum_i2 j"
+
+primrec forest_of_mylist :: "'a tree mylist \<Rightarrow> 'a forest" where
+  "forest_of_mylist MyNil = FNil" |
+  "forest_of_mylist (MyCons t ts) = FCons t (forest_of_mylist ts)"
+
+primrec mylist_of_forest :: "'a forest \<Rightarrow> 'a tree mylist" where
+  "mylist_of_forest FNil = MyNil" |
+  "mylist_of_forest (FCons t ts) = MyCons t (mylist_of_forest ts)"
+
+definition frev :: "'a forest \<Rightarrow> 'a forest" where
+  "frev = forest_of_mylist \<circ> myrev \<circ> mylist_of_forest"
+
+primrec
+  mirror_tree :: "'a tree \<Rightarrow> 'a tree" and
+  mirror_forest :: "'a forest \<Rightarrow> 'a forest"
+where
+  "mirror_tree TEmpty = TEmpty" |
+  "mirror_tree (TNode x ts) = TNode x (mirror_forest ts)" |
+  "mirror_forest FNil = FNil" |
+  "mirror_forest (FCons t ts) = frev (FCons (mirror_tree t) (mirror_forest ts))"
+
+primrec
+  mylist_of_tree' :: "'a tree' \<Rightarrow> 'a mylist" and
+  mylist_of_branch :: "'a branch \<Rightarrow> 'a mylist"
+where
+  "mylist_of_tree' TEmpty' = MyNil" |
+  "mylist_of_tree' (TNode' b b') = myapp (mylist_of_branch b) (mylist_of_branch b')" |
+  "mylist_of_branch (Branch x t) = MyCons x (mylist_of_tree' t)"
+
+primrec
+  id_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
+  id_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
+  id_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
+where
+  "id_tree (BRTree a ts ts') = BRTree a (id_trees1 ts) (id_trees2 ts')" |
+  "id_trees1 MyNil = MyNil" |
+  "id_trees1 (MyCons t ts) = MyCons (id_tree t) (id_trees1 ts)" |
+  "id_trees2 MyNil = MyNil" |
+  "id_trees2 (MyCons t ts) = MyCons (id_tree t) (id_trees2 ts)"
+
+primrec
+  trunc_tree :: "'a bin_rose_tree \<Rightarrow> 'a bin_rose_tree" and
+  trunc_trees1 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist" and
+  trunc_trees2 :: "'a bin_rose_tree mylist \<Rightarrow> 'a bin_rose_tree mylist"
+where
+  "trunc_tree (BRTree a ts ts') = BRTree a (trunc_trees1 ts) (trunc_trees2 ts')" |
+  "trunc_trees1 MyNil = MyNil" |
+  "trunc_trees1 (MyCons t ts) = MyCons (id_tree t) MyNil" |
+  "trunc_trees2 MyNil = MyNil" |
+  "trunc_trees2 (MyCons t ts) = MyCons (id_tree t) MyNil"
+
+primrec
+  is_ground_exp :: "('a, 'b) exp \<Rightarrow> bool" and
+  is_ground_trm :: "('a, 'b) trm \<Rightarrow> bool" and
+  is_ground_factor :: "('a, 'b) factor \<Rightarrow> bool"
+where
+  "is_ground_exp (Term t) \<longleftrightarrow> is_ground_trm t" |
+  "is_ground_exp (Sum t e) \<longleftrightarrow> is_ground_trm t \<and> is_ground_exp e" |
+  "is_ground_trm (Factor f) \<longleftrightarrow> is_ground_factor f" |
+  "is_ground_trm (Prod f t) \<longleftrightarrow> is_ground_factor f \<and> is_ground_trm t" |
+  "is_ground_factor (C _) \<longleftrightarrow> True" |
+  "is_ground_factor (V _) \<longleftrightarrow> False" |
+  "is_ground_factor (Paren e) \<longleftrightarrow> is_ground_exp e"
+
+primrec map_ftreeA :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
+  "map_ftreeA f (FTLeaf x) = FTLeaf (f x)" |
+  "map_ftreeA f (FTNode g) = FTNode (map_ftreeA f \<circ> g)"
+
+primrec map_ftreeB :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a ftree \<Rightarrow> 'b ftree" where
+  "map_ftreeB f (FTLeaf x) = FTLeaf (f x)" |
+  "map_ftreeB f (FTNode g) = FTNode (map_ftreeB f \<circ> g \<circ> the_inv f)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Process.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,278 @@
+(*  Title:      HOL/Datatype_Examples/Process.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Processes.
+*)
+
+header {* Processes *}
+
+theory Process
+imports Stream 
+begin
+
+codatatype 'a process =
+  isAction: Action (prefOf: 'a) (contOf: "'a process") |
+  isChoice: Choice (ch1Of: "'a process") (ch2Of: "'a process")
+
+(* Read: prefix of, continuation of, choice 1 of, choice 2 of *)
+
+section {* Customization *}
+
+subsection {* Basic properties *}
+
+declare
+  rel_pre_process_def[simp]
+  rel_sum_def[simp]
+  rel_prod_def[simp]
+
+(* Constructors versus discriminators *)
+theorem isAction_isChoice:
+"isAction p \<or> isChoice p"
+by (rule process.exhaust_disc) auto
+
+theorem not_isAction_isChoice: "\<not> (isAction p \<and> isChoice p)"
+by (cases rule: process.exhaust[of p]) auto
+
+
+subsection{* Coinduction *}
+
+theorem process_coind[elim, consumes 1, case_names iss Action Choice, induct pred: "HOL.eq"]:
+  assumes phi: "\<phi> p p'" and
+  iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
+  Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> \<phi> p p'" and
+  Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> \<phi> p p' \<and> \<phi> q q'"
+  shows "p = p'"
+  using assms
+  by (coinduct rule: process.coinduct) (metis process.collapse(1,2) process.disc(3))
+
+(* Stronger coinduction, up to equality: *)
+theorem process_strong_coind[elim, consumes 1, case_names iss Action Choice]:
+  assumes phi: "\<phi> p p'" and
+  iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
+  Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> (\<phi> p p' \<or> p = p')" and
+  Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> (\<phi> p p' \<or> p = p') \<and> (\<phi> q q' \<or> q = q')"
+  shows "p = p'"
+  using assms
+  by (coinduct rule: process.coinduct_strong) (metis process.collapse(1,2) process.disc(3))
+
+
+subsection {* Coiteration (unfold) *}
+
+
+section{* Coinductive definition of the notion of trace *}
+coinductive trace where
+"trace p as \<Longrightarrow> trace (Action a p) (a ## as)"
+|
+"trace p as \<or> trace q as \<Longrightarrow> trace (Choice p q) as"
+
+
+section{* Examples of corecursive definitions: *}
+
+subsection{* Single-guard fixpoint definition *}
+
+primcorec BX where
+  "isAction BX"
+| "prefOf BX = ''a''"
+| "contOf BX = BX"
+
+
+subsection{* Multi-guard fixpoint definitions, simulated with auxiliary arguments *}
+
+datatype_new x_y_ax = x | y | ax
+
+primcorec F :: "x_y_ax \<Rightarrow> char list process" where
+  "xyax = x \<Longrightarrow> isChoice (F xyax)"
+| "ch1Of (F xyax) = F ax"
+| "ch2Of (F xyax) = F y"
+| "prefOf (F xyax) = (if xyax = y then ''b'' else ''a'')"
+| "contOf (F xyax) = F x"
+
+definition "X = F x"  definition "Y = F y"  definition "AX = F ax"
+
+lemma X_Y_AX: "X = Choice AX Y"  "Y = Action ''b'' X"  "AX = Action ''a'' X"
+unfolding X_def Y_def AX_def by (subst F.code, simp)+
+
+(* end product: *)
+lemma X_AX:
+"X = Choice AX (Action ''b'' X)"
+"AX = Action ''a'' X"
+using X_Y_AX by simp_all
+
+
+
+section{* Case study: Multi-guard fixpoint definitions, without auxiliary arguments *}
+
+hide_const x y ax X Y AX
+
+(* Process terms *)
+datatype_new ('a,'pvar) process_term =
+ VAR 'pvar |
+ PROC "'a process" |
+ ACT 'a "('a,'pvar) process_term" | CH "('a,'pvar) process_term" "('a,'pvar) process_term"
+
+(* below, sys represents a system of equations *)
+fun isACT where
+"isACT sys (VAR X) =
+ (case sys X of ACT a T \<Rightarrow> True |PROC p \<Rightarrow> isAction p |_ \<Rightarrow> False)"
+|
+"isACT sys (PROC p) = isAction p"
+|
+"isACT sys (ACT a T) = True"
+|
+"isACT sys (CH T1 T2) = False"
+
+fun PREF where
+"PREF sys (VAR X) =
+ (case sys X of ACT a T \<Rightarrow> a | PROC p \<Rightarrow> prefOf p)"
+|
+"PREF sys (PROC p) = prefOf p"
+|
+"PREF sys (ACT a T) = a"
+
+fun CONT where
+"CONT sys (VAR X) =
+ (case sys X of ACT a T \<Rightarrow> T | PROC p \<Rightarrow> PROC (contOf p))"
+|
+"CONT sys (PROC p) = PROC (contOf p)"
+|
+"CONT sys (ACT a T) = T"
+
+fun CH1 where
+"CH1 sys (VAR X) =
+ (case sys X of CH T1 T2 \<Rightarrow> T1 |PROC p \<Rightarrow> PROC (ch1Of p))"
+|
+"CH1 sys (PROC p) = PROC (ch1Of p)"
+|
+"CH1 sys (CH T1 T2) = T1"
+
+fun CH2 where
+"CH2 sys (VAR X) =
+ (case sys X of CH T1 T2 \<Rightarrow> T2 |PROC p \<Rightarrow> PROC (ch2Of p))"
+|
+"CH2 sys (PROC p) = PROC (ch2Of p)"
+|
+"CH2 sys (CH T1 T2) = T2"
+
+definition "guarded sys \<equiv> \<forall> X Y. sys X \<noteq> VAR Y"
+
+primcorec solution where
+  "isACT sys T \<Longrightarrow> solution sys T = Action (PREF sys T) (solution sys (CONT sys T))"
+| "_ \<Longrightarrow> solution sys T = Choice (solution sys (CH1 sys T)) (solution sys (CH2 sys T))"
+
+lemma isACT_VAR:
+assumes g: "guarded sys"
+shows "isACT sys (VAR X) \<longleftrightarrow> isACT sys (sys X)"
+using g unfolding guarded_def by (cases "sys X") auto
+
+lemma solution_VAR:
+assumes g: "guarded sys"
+shows "solution sys (VAR X) = solution sys (sys X)"
+proof(cases "isACT sys (VAR X)")
+  case True
+  hence T: "isACT sys (sys X)" unfolding isACT_VAR[OF g] .
+  show ?thesis
+  unfolding solution.ctr(1)[OF T] using solution.ctr(1)[of sys "VAR X"] True g
+  unfolding guarded_def by (cases "sys X", auto)
+next
+  case False note FFalse = False
+  hence TT: "\<not> isACT sys (sys X)" unfolding isACT_VAR[OF g] .
+  show ?thesis
+  unfolding solution.ctr(2)[OF TT] using solution.ctr(2)[of sys "VAR X"] FFalse g
+  unfolding guarded_def by (cases "sys X", auto)
+qed
+
+lemma solution_PROC[simp]:
+"solution sys (PROC p) = p"
+proof-
+  {fix q assume "q = solution sys (PROC p)"
+   hence "p = q"
+   proof (coinduct rule: process_coind)
+     case (iss p p')
+     from isAction_isChoice[of p] show ?case
+     proof
+       assume p: "isAction p"
+       hence 0: "isACT sys (PROC p)" by simp
+       thus ?thesis using iss not_isAction_isChoice by auto
+     next
+       assume "isChoice p"
+       hence 0: "\<not> isACT sys (PROC p)"
+       using not_isAction_isChoice by auto
+       thus ?thesis using iss isAction_isChoice by auto
+     qed
+   next
+     case (Action a a' p p')
+     hence 0: "isACT sys (PROC (Action a p))" by simp
+     show ?case using Action unfolding solution.ctr(1)[OF 0] by simp
+   next
+     case (Choice p q p' q')
+     hence 0: "\<not> isACT sys (PROC (Choice p q))" using not_isAction_isChoice by auto
+     show ?case using Choice unfolding solution.ctr(2)[OF 0] by simp
+   qed
+  }
+  thus ?thesis by metis
+qed
+
+lemma solution_ACT[simp]:
+"solution sys (ACT a T) = Action a (solution sys T)"
+by (metis CONT.simps(3) PREF.simps(3) isACT.simps(3) solution.ctr(1))
+
+lemma solution_CH[simp]:
+"solution sys (CH T1 T2) = Choice (solution sys T1) (solution sys T2)"
+by (metis CH1.simps(3) CH2.simps(3) isACT.simps(4) solution.ctr(2))
+
+
+(* Example: *)
+
+fun sys where
+"sys 0 = CH (VAR (Suc 0)) (ACT ''b'' (VAR 0))"
+|
+"sys (Suc 0) = ACT ''a'' (VAR 0)"
+| (* dummy guarded term for variables outside the system: *)
+"sys X = ACT ''a'' (VAR 0)"
+
+lemma guarded_sys:
+"guarded sys"
+unfolding guarded_def proof (intro allI)
+  fix X Y show "sys X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
+qed
+
+(* the actual processes: *)
+definition "x \<equiv> solution sys (VAR 0)"
+definition "ax \<equiv> solution sys (VAR (Suc 0))"
+
+(* end product: *)
+lemma x_ax:
+"x = Choice ax (Action ''b'' x)"
+"ax = Action ''a'' x"
+unfolding x_def ax_def by (subst solution_VAR[OF guarded_sys], simp)+
+
+
+(* Thanks to the inclusion of processes as process terms, one can
+also consider parametrized systems of equations---here, x is a (semantic)
+process parameter: *)
+
+fun sys' where
+"sys' 0 = CH (PROC x) (ACT ''b'' (VAR 0))"
+|
+"sys' (Suc 0) = CH (ACT ''a'' (VAR 0)) (PROC x)"
+| (* dummy guarded term : *)
+"sys' X = ACT ''a'' (VAR 0)"
+
+lemma guarded_sys':
+"guarded sys'"
+unfolding guarded_def proof (intro allI)
+  fix X Y show "sys' X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
+qed
+
+(* the actual processes: *)
+definition "y \<equiv> solution sys' (VAR 0)"
+definition "ay \<equiv> solution sys' (VAR (Suc 0))"
+
+(* end product: *)
+lemma y_ay:
+"y = Choice x (Action ''b'' y)"
+"ay = Choice (Action ''a'' y) x"
+unfolding y_def ay_def by (subst solution_VAR[OF guarded_sys'], simp)+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/SML.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,91 @@
+(*  Title:      HOL/Datatype_Benchmark/SML.thy
+
+Example from Myra: part of the syntax of SML.
+*)
+
+theory SML imports Main begin
+
+datatype
+  string = EMPTY_STRING | CONS_STRING nat string
+
+datatype
+   strid = STRID string
+
+datatype
+   var = VAR string
+
+datatype
+   con = CON string
+
+datatype
+   scon = SCINT nat | SCSTR string
+
+datatype
+   excon = EXCON string
+
+datatype
+   label = LABEL string
+
+datatype
+  'a nonemptylist = Head_and_tail 'a "'a list"
+
+datatype
+  'a long = BASE 'a | QUALIFIED strid "'a long"
+
+datatype
+   atpat_e = WILDCARDatpat_e
+           | SCONatpat_e scon
+           | VARatpat_e var
+           | CONatpat_e "con long"
+           | EXCONatpat_e "excon long"
+           | RECORDatpat_e "patrow_e option"
+           | PARatpat_e pat_e
+and
+   patrow_e = DOTDOTDOT_e
+            | PATROW_e label pat_e "patrow_e option"
+and
+   pat_e = ATPATpat_e atpat_e
+         | CONpat_e "con long" atpat_e
+         | EXCONpat_e "excon long" atpat_e
+         | LAYEREDpat_e var pat_e
+and
+   conbind_e = CONBIND_e con "conbind_e option"
+and
+   datbind_e = DATBIND_e conbind_e "datbind_e option"
+and
+   exbind_e = EXBIND1_e excon "exbind_e option"
+            | EXBIND2_e excon "excon long" "exbind_e option"
+and
+   atexp_e = SCONatexp_e scon
+           | VARatexp_e "var long"
+           | CONatexp_e "con long"
+           | EXCONatexp_e "excon long"
+           | RECORDatexp_e "exprow_e option"
+           | LETatexp_e dec_e exp_e
+           | PARatexp_e exp_e
+and
+   exprow_e = EXPROW_e label exp_e "exprow_e option"
+and
+   exp_e = ATEXPexp_e atexp_e
+         | APPexp_e exp_e atexp_e
+         | HANDLEexp_e exp_e match_e
+         | RAISEexp_e exp_e
+         | FNexp_e match_e
+and
+   match_e = MATCH_e mrule_e "match_e option"
+and
+   mrule_e = MRULE_e pat_e exp_e
+and
+   dec_e = VALdec_e valbind_e
+         | DATATYPEdec_e datbind_e
+         | ABSTYPEdec_e datbind_e dec_e
+         | EXCEPTdec_e exbind_e
+         | LOCALdec_e dec_e dec_e
+         | OPENdec_e "strid long nonemptylist"
+         | EMPTYdec_e
+         | SEQdec_e dec_e dec_e
+and
+   valbind_e = PLAINvalbind_e pat_e exp_e "valbind_e option"
+             | RECvalbind_e valbind_e
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Stream.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,582 @@
+(*  Title:      HOL/Datatype_Examples/Stream.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012, 2013
+
+Infinite streams.
+*)
+
+header {* Infinite Streams *}
+
+theory Stream
+imports "~~/src/HOL/Library/Nat_Bijection"
+begin
+
+codatatype (sset: 'a) stream =
+  SCons (shd: 'a) (stl: "'a stream") (infixr "##" 65)
+for
+  map: smap
+  rel: stream_all2
+
+(*for code generation only*)
+definition smember :: "'a \<Rightarrow> 'a stream \<Rightarrow> bool" where
+  [code_abbrev]: "smember x s \<longleftrightarrow> x \<in> sset s"
+
+lemma smember_code[code, simp]: "smember x (y ## s) = (if x = y then True else smember x s)"
+  unfolding smember_def by auto
+
+hide_const (open) smember
+
+lemmas smap_simps[simp] = stream.map_sel
+lemmas shd_sset = stream.set_sel(1)
+lemmas stl_sset = stream.set_sel(2)
+
+theorem sset_induct[consumes 1, case_names shd stl, induct set: sset]:
+  assumes "y \<in> sset s" and "\<And>s. P (shd s) s" and "\<And>s y. \<lbrakk>y \<in> sset (stl s); P y (stl s)\<rbrakk> \<Longrightarrow> P y s"
+  shows "P y s"
+using assms by induct (metis stream.sel(1), auto)
+
+
+subsection {* prepend list to stream *}
+
+primrec shift :: "'a list \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr "@-" 65) where
+  "shift [] s = s"
+| "shift (x # xs) s = x ## shift xs s"
+
+lemma smap_shift[simp]: "smap f (xs @- s) = map f xs @- smap f s"
+  by (induct xs) auto
+
+lemma shift_append[simp]: "(xs @ ys) @- s = xs @- ys @- s"
+  by (induct xs) auto
+
+lemma shift_simps[simp]:
+   "shd (xs @- s) = (if xs = [] then shd s else hd xs)"
+   "stl (xs @- s) = (if xs = [] then stl s else tl xs @- s)"
+  by (induct xs) auto
+
+lemma sset_shift[simp]: "sset (xs @- s) = set xs \<union> sset s"
+  by (induct xs) auto
+
+lemma shift_left_inj[simp]: "xs @- s1 = xs @- s2 \<longleftrightarrow> s1 = s2"
+  by (induct xs) auto
+
+
+subsection {* set of streams with elements in some fixed set *}
+
+coinductive_set
+  streams :: "'a set \<Rightarrow> 'a stream set"
+  for A :: "'a set"
+where
+  Stream[intro!, simp, no_atp]: "\<lbrakk>a \<in> A; s \<in> streams A\<rbrakk> \<Longrightarrow> a ## s \<in> streams A"
+
+lemma shift_streams: "\<lbrakk>w \<in> lists A; s \<in> streams A\<rbrakk> \<Longrightarrow> w @- s \<in> streams A"
+  by (induct w) auto
+
+lemma streams_Stream: "x ## s \<in> streams A \<longleftrightarrow> x \<in> A \<and> s \<in> streams A"
+  by (auto elim: streams.cases)
+
+lemma streams_stl: "s \<in> streams A \<Longrightarrow> stl s \<in> streams A"
+  by (cases s) (auto simp: streams_Stream)
+
+lemma streams_shd: "s \<in> streams A \<Longrightarrow> shd s \<in> A"
+  by (cases s) (auto simp: streams_Stream)
+
+lemma sset_streams:
+  assumes "sset s \<subseteq> A"
+  shows "s \<in> streams A"
+using assms proof (coinduction arbitrary: s)
+  case streams then show ?case by (cases s) simp
+qed
+
+lemma streams_sset:
+  assumes "s \<in> streams A"
+  shows "sset s \<subseteq> A"
+proof
+  fix x assume "x \<in> sset s" from this `s \<in> streams A` show "x \<in> A"
+    by (induct s) (auto intro: streams_shd streams_stl)
+qed
+
+lemma streams_iff_sset: "s \<in> streams A \<longleftrightarrow> sset s \<subseteq> A"
+  by (metis sset_streams streams_sset)
+
+lemma streams_mono:  "s \<in> streams A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> s \<in> streams B"
+  unfolding streams_iff_sset by auto
+
+lemma smap_streams: "s \<in> streams A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> smap f s \<in> streams B"
+  unfolding streams_iff_sset stream.set_map by auto
+
+lemma streams_empty: "streams {} = {}"
+  by (auto elim: streams.cases)
+
+lemma streams_UNIV[simp]: "streams UNIV = UNIV"
+  by (auto simp: streams_iff_sset)
+
+subsection {* nth, take, drop for streams *}
+
+primrec snth :: "'a stream \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where
+  "s !! 0 = shd s"
+| "s !! Suc n = stl s !! n"
+
+lemma snth_smap[simp]: "smap f s !! n = f (s !! n)"
+  by (induct n arbitrary: s) auto
+
+lemma shift_snth_less[simp]: "p < length xs \<Longrightarrow> (xs @- s) !! p = xs ! p"
+  by (induct p arbitrary: xs) (auto simp: hd_conv_nth nth_tl)
+
+lemma shift_snth_ge[simp]: "p \<ge> length xs \<Longrightarrow> (xs @- s) !! p = s !! (p - length xs)"
+  by (induct p arbitrary: xs) (auto simp: Suc_diff_eq_diff_pred)
+
+lemma shift_snth: "(xs @- s) !! n = (if n < length xs then xs ! n else s !! (n - length xs))"
+  by auto
+
+lemma snth_sset[simp]: "s !! n \<in> sset s"
+  by (induct n arbitrary: s) (auto intro: shd_sset stl_sset)
+
+lemma sset_range: "sset s = range (snth s)"
+proof (intro equalityI subsetI)
+  fix x assume "x \<in> sset s"
+  thus "x \<in> range (snth s)"
+  proof (induct s)
+    case (stl s x)
+    then obtain n where "x = stl s !! n" by auto
+    thus ?case by (auto intro: range_eqI[of _ _ "Suc n"])
+  qed (auto intro: range_eqI[of _ _ 0])
+qed auto
+
+primrec stake :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a list" where
+  "stake 0 s = []"
+| "stake (Suc n) s = shd s # stake n (stl s)"
+
+lemma length_stake[simp]: "length (stake n s) = n"
+  by (induct n arbitrary: s) auto
+
+lemma stake_smap[simp]: "stake n (smap f s) = map f (stake n s)"
+  by (induct n arbitrary: s) auto
+
+lemma take_stake: "take n (stake m s) = stake (min n m) s"
+proof (induct m arbitrary: s n)
+  case (Suc m) thus ?case by (cases n) auto
+qed simp
+
+primrec sdrop :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
+  "sdrop 0 s = s"
+| "sdrop (Suc n) s = sdrop n (stl s)"
+
+lemma sdrop_simps[simp]:
+  "shd (sdrop n s) = s !! n" "stl (sdrop n s) = sdrop (Suc n) s"
+  by (induct n arbitrary: s)  auto
+
+lemma sdrop_smap[simp]: "sdrop n (smap f s) = smap f (sdrop n s)"
+  by (induct n arbitrary: s) auto
+
+lemma sdrop_stl: "sdrop n (stl s) = stl (sdrop n s)"
+  by (induct n) auto
+
+lemma drop_stake: "drop n (stake m s) = stake (m - n) (sdrop n s)"
+proof (induct m arbitrary: s n)
+  case (Suc m) thus ?case by (cases n) auto
+qed simp
+
+lemma stake_sdrop: "stake n s @- sdrop n s = s"
+  by (induct n arbitrary: s) auto
+
+lemma id_stake_snth_sdrop:
+  "s = stake i s @- s !! i ## sdrop (Suc i) s"
+  by (subst stake_sdrop[symmetric, of _ i]) (metis sdrop_simps stream.collapse)
+
+lemma smap_alt: "smap f s = s' \<longleftrightarrow> (\<forall>n. f (s !! n) = s' !! n)" (is "?L = ?R")
+proof
+  assume ?R
+  then have "\<And>n. smap f (sdrop n s) = sdrop n s'"
+    by coinduction (auto intro: exI[of _ 0] simp del: sdrop.simps(2))
+  then show ?L using sdrop.simps(1) by metis
+qed auto
+
+lemma stake_invert_Nil[iff]: "stake n s = [] \<longleftrightarrow> n = 0"
+  by (induct n) auto
+
+lemma sdrop_shift: "sdrop i (w @- s) = drop i w @- sdrop (i - length w) s"
+  by (induct i arbitrary: w s) (auto simp: drop_tl drop_Suc neq_Nil_conv)
+
+lemma stake_shift: "stake i (w @- s) = take i w @ stake (i - length w) s"
+  by (induct i arbitrary: w s) (auto simp: neq_Nil_conv)
+
+lemma stake_add[simp]: "stake m s @ stake n (sdrop m s) = stake (m + n) s"
+  by (induct m arbitrary: s) auto
+
+lemma sdrop_add[simp]: "sdrop n (sdrop m s) = sdrop (m + n) s"
+  by (induct m arbitrary: s) auto
+
+lemma sdrop_snth: "sdrop n s !! m = s !! (n + m)"
+  by (induct n arbitrary: m s) auto
+
+partial_function (tailrec) sdrop_while :: "('a \<Rightarrow> bool) \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where 
+  "sdrop_while P s = (if P (shd s) then sdrop_while P (stl s) else s)"
+
+lemma sdrop_while_SCons[code]:
+  "sdrop_while P (a ## s) = (if P a then sdrop_while P s else a ## s)"
+  by (subst sdrop_while.simps) simp
+
+lemma sdrop_while_sdrop_LEAST:
+  assumes "\<exists>n. P (s !! n)"
+  shows "sdrop_while (Not o P) s = sdrop (LEAST n. P (s !! n)) s"
+proof -
+  from assms obtain m where "P (s !! m)" "\<And>n. P (s !! n) \<Longrightarrow> m \<le> n"
+    and *: "(LEAST n. P (s !! n)) = m" by atomize_elim (auto intro: LeastI Least_le)
+  thus ?thesis unfolding *
+  proof (induct m arbitrary: s)
+    case (Suc m)
+    hence "sdrop_while (Not \<circ> P) (stl s) = sdrop m (stl s)"
+      by (metis (full_types) not_less_eq_eq snth.simps(2))
+    moreover from Suc(3) have "\<not> (P (s !! 0))" by blast
+    ultimately show ?case by (subst sdrop_while.simps) simp
+  qed (metis comp_apply sdrop.simps(1) sdrop_while.simps snth.simps(1))
+qed
+
+primcorec sfilter where
+  "shd (sfilter P s) = shd (sdrop_while (Not o P) s)"
+| "stl (sfilter P s) = sfilter P (stl (sdrop_while (Not o P) s))"
+
+lemma sfilter_Stream: "sfilter P (x ## s) = (if P x then x ## sfilter P s else sfilter P s)"
+proof (cases "P x")
+  case True thus ?thesis by (subst sfilter.ctr) (simp add: sdrop_while_SCons)
+next
+  case False thus ?thesis by (subst (1 2) sfilter.ctr) (simp add: sdrop_while_SCons)
+qed
+
+
+subsection {* unary predicates lifted to streams *}
+
+definition "stream_all P s = (\<forall>p. P (s !! p))"
+
+lemma stream_all_iff[iff]: "stream_all P s \<longleftrightarrow> Ball (sset s) P"
+  unfolding stream_all_def sset_range by auto
+
+lemma stream_all_shift[simp]: "stream_all P (xs @- s) = (list_all P xs \<and> stream_all P s)"
+  unfolding stream_all_iff list_all_iff by auto
+
+lemma stream_all_Stream: "stream_all P (x ## X) \<longleftrightarrow> P x \<and> stream_all P X"
+  by simp
+
+
+subsection {* recurring stream out of a list *}
+
+primcorec cycle :: "'a list \<Rightarrow> 'a stream" where
+  "shd (cycle xs) = hd xs"
+| "stl (cycle xs) = cycle (tl xs @ [hd xs])"
+
+lemma cycle_decomp: "u \<noteq> [] \<Longrightarrow> cycle u = u @- cycle u"
+proof (coinduction arbitrary: u)
+  case Eq_stream then show ?case using stream.collapse[of "cycle u"]
+    by (auto intro!: exI[of _ "tl u @ [hd u]"])
+qed
+
+lemma cycle_Cons[code]: "cycle (x # xs) = x ## cycle (xs @ [x])"
+  by (subst cycle.ctr) simp
+
+lemma cycle_rotated: "\<lbrakk>v \<noteq> []; cycle u = v @- s\<rbrakk> \<Longrightarrow> cycle (tl u @ [hd u]) = tl v @- s"
+  by (auto dest: arg_cong[of _ _ stl])
+
+lemma stake_append: "stake n (u @- s) = take (min (length u) n) u @ stake (n - length u) s"
+proof (induct n arbitrary: u)
+  case (Suc n) thus ?case by (cases u) auto
+qed auto
+
+lemma stake_cycle_le[simp]:
+  assumes "u \<noteq> []" "n < length u"
+  shows "stake n (cycle u) = take n u"
+using min_absorb2[OF less_imp_le_nat[OF assms(2)]]
+  by (subst cycle_decomp[OF assms(1)], subst stake_append) auto
+
+lemma stake_cycle_eq[simp]: "u \<noteq> [] \<Longrightarrow> stake (length u) (cycle u) = u"
+  by (subst cycle_decomp) (auto simp: stake_shift)
+
+lemma sdrop_cycle_eq[simp]: "u \<noteq> [] \<Longrightarrow> sdrop (length u) (cycle u) = cycle u"
+  by (subst cycle_decomp) (auto simp: sdrop_shift)
+
+lemma stake_cycle_eq_mod_0[simp]: "\<lbrakk>u \<noteq> []; n mod length u = 0\<rbrakk> \<Longrightarrow>
+   stake n (cycle u) = concat (replicate (n div length u) u)"
+  by (induct "n div length u" arbitrary: n u) (auto simp: stake_add[symmetric])
+
+lemma sdrop_cycle_eq_mod_0[simp]: "\<lbrakk>u \<noteq> []; n mod length u = 0\<rbrakk> \<Longrightarrow>
+   sdrop n (cycle u) = cycle u"
+  by (induct "n div length u" arbitrary: n u) (auto simp: sdrop_add[symmetric])
+
+lemma stake_cycle: "u \<noteq> [] \<Longrightarrow>
+   stake n (cycle u) = concat (replicate (n div length u) u) @ take (n mod length u) u"
+  by (subst mod_div_equality[of n "length u", symmetric], unfold stake_add[symmetric]) auto
+
+lemma sdrop_cycle: "u \<noteq> [] \<Longrightarrow> sdrop n (cycle u) = cycle (rotate (n mod length u) u)"
+  by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric])
+
+
+subsection {* iterated application of a function *}
+
+primcorec siterate where
+  "shd (siterate f x) = x"
+| "stl (siterate f x) = siterate f (f x)"
+
+lemma stake_Suc: "stake (Suc n) s = stake n s @ [s !! n]"
+  by (induct n arbitrary: s) auto
+
+lemma snth_siterate[simp]: "siterate f x !! n = (f^^n) x"
+  by (induct n arbitrary: x) (auto simp: funpow_swap1)
+
+lemma sdrop_siterate[simp]: "sdrop n (siterate f x) = siterate f ((f^^n) x)"
+  by (induct n arbitrary: x) (auto simp: funpow_swap1)
+
+lemma stake_siterate[simp]: "stake n (siterate f x) = map (\<lambda>n. (f^^n) x) [0 ..< n]"
+  by (induct n arbitrary: x) (auto simp del: stake.simps(2) simp: stake_Suc)
+
+lemma sset_siterate: "sset (siterate f x) = {(f^^n) x | n. True}"
+  by (auto simp: sset_range)
+
+lemma smap_siterate: "smap f (siterate f x) = siterate f (f x)"
+  by (coinduction arbitrary: x) auto
+
+
+subsection {* stream repeating a single element *}
+
+abbreviation "sconst \<equiv> siterate id"
+
+lemma shift_replicate_sconst[simp]: "replicate n x @- sconst x = sconst x"
+  by (subst (3) stake_sdrop[symmetric]) (simp add: map_replicate_trivial)
+
+lemma sset_sconst[simp]: "sset (sconst x) = {x}"
+  by (simp add: sset_siterate)
+
+lemma sconst_alt: "s = sconst x \<longleftrightarrow> sset s = {x}"
+proof
+  assume "sset s = {x}"
+  then show "s = sconst x"
+  proof (coinduction arbitrary: s)
+    case Eq_stream
+    then have "shd s = x" "sset (stl s) \<subseteq> {x}" by (case_tac [!] s) auto
+    then have "sset (stl s) = {x}" by (cases "stl s") auto
+    with `shd s = x` show ?case by auto
+  qed
+qed simp
+
+lemma same_cycle: "sconst x = cycle [x]"
+  by coinduction auto
+
+lemma smap_sconst: "smap f (sconst x) = sconst (f x)"
+  by coinduction auto
+
+lemma sconst_streams: "x \<in> A \<Longrightarrow> sconst x \<in> streams A"
+  by (simp add: streams_iff_sset)
+
+
+subsection {* stream of natural numbers *}
+
+abbreviation "fromN \<equiv> siterate Suc"
+
+abbreviation "nats \<equiv> fromN 0"
+
+lemma sset_fromN[simp]: "sset (fromN n) = {n ..}"
+  by (auto simp add: sset_siterate le_iff_add)
+
+lemma stream_smap_fromN: "s = smap (\<lambda>j. let i = j - n in s !! i) (fromN n)"
+  by (coinduction arbitrary: s n)
+    (force simp: neq_Nil_conv Let_def snth.simps(2)[symmetric] Suc_diff_Suc
+      intro: stream.map_cong split: if_splits simp del: snth.simps(2))
+
+lemma stream_smap_nats: "s = smap (snth s) nats"
+  using stream_smap_fromN[where n = 0] by simp
+
+
+subsection {* flatten a stream of lists *}
+
+primcorec flat where
+  "shd (flat ws) = hd (shd ws)"
+| "stl (flat ws) = flat (if tl (shd ws) = [] then stl ws else tl (shd ws) ## stl ws)"
+
+lemma flat_Cons[simp, code]: "flat ((x # xs) ## ws) = x ## flat (if xs = [] then ws else xs ## ws)"
+  by (subst flat.ctr) simp
+
+lemma flat_Stream[simp]: "xs \<noteq> [] \<Longrightarrow> flat (xs ## ws) = xs @- flat ws"
+  by (induct xs) auto
+
+lemma flat_unfold: "shd ws \<noteq> [] \<Longrightarrow> flat ws = shd ws @- flat (stl ws)"
+  by (cases ws) auto
+
+lemma flat_snth: "\<forall>xs \<in> sset s. xs \<noteq> [] \<Longrightarrow> flat s !! n = (if n < length (shd s) then 
+  shd s ! n else flat (stl s) !! (n - length (shd s)))"
+  by (metis flat_unfold not_less shd_sset shift_snth_ge shift_snth_less)
+
+lemma sset_flat[simp]: "\<forall>xs \<in> sset s. xs \<noteq> [] \<Longrightarrow> 
+  sset (flat s) = (\<Union>xs \<in> sset s. set xs)" (is "?P \<Longrightarrow> ?L = ?R")
+proof safe
+  fix x assume ?P "x : ?L"
+  then obtain m where "x = flat s !! m" by (metis image_iff sset_range)
+  with `?P` obtain n m' where "x = s !! n ! m'" "m' < length (s !! n)"
+  proof (atomize_elim, induct m arbitrary: s rule: less_induct)
+    case (less y)
+    thus ?case
+    proof (cases "y < length (shd s)")
+      case True thus ?thesis by (metis flat_snth less(2,3) snth.simps(1))
+    next
+      case False
+      hence "x = flat (stl s) !! (y - length (shd s))" by (metis less(2,3) flat_snth)
+      moreover
+      { from less(2) have *: "length (shd s) > 0" by (cases s) simp_all
+        with False have "y > 0" by (cases y) simp_all
+        with * have "y - length (shd s) < y" by simp
+      }
+      moreover have "\<forall>xs \<in> sset (stl s). xs \<noteq> []" using less(2) by (cases s) auto
+      ultimately have "\<exists>n m'. x = stl s !! n ! m' \<and> m' < length (stl s !! n)" by (intro less(1)) auto
+      thus ?thesis by (metis snth.simps(2))
+    qed
+  qed
+  thus "x \<in> ?R" by (auto simp: sset_range dest!: nth_mem)
+next
+  fix x xs assume "xs \<in> sset s" ?P "x \<in> set xs" thus "x \<in> ?L"
+    by (induct rule: sset_induct)
+      (metis UnI1 flat_unfold shift.simps(1) sset_shift,
+       metis UnI2 flat_unfold shd_sset stl_sset sset_shift)
+qed
+
+
+subsection {* merge a stream of streams *}
+
+definition smerge :: "'a stream stream \<Rightarrow> 'a stream" where
+  "smerge ss = flat (smap (\<lambda>n. map (\<lambda>s. s !! n) (stake (Suc n) ss) @ stake n (ss !! n)) nats)"
+
+lemma stake_nth[simp]: "m < n \<Longrightarrow> stake n s ! m = s !! m"
+  by (induct n arbitrary: s m) (auto simp: nth_Cons', metis Suc_pred snth.simps(2))
+
+lemma snth_sset_smerge: "ss !! n !! m \<in> sset (smerge ss)"
+proof (cases "n \<le> m")
+  case False thus ?thesis unfolding smerge_def
+    by (subst sset_flat)
+      (auto simp: stream.set_map in_set_conv_nth simp del: stake.simps
+        intro!: exI[of _ n, OF disjI2] exI[of _ m, OF mp])
+next
+  case True thus ?thesis unfolding smerge_def
+    by (subst sset_flat)
+      (auto simp: stream.set_map in_set_conv_nth image_iff simp del: stake.simps snth.simps
+        intro!: exI[of _ m, OF disjI1] bexI[of _ "ss !! n"] exI[of _ n, OF mp])
+qed
+
+lemma sset_smerge: "sset (smerge ss) = UNION (sset ss) sset"
+proof safe
+  fix x assume "x \<in> sset (smerge ss)"
+  thus "x \<in> UNION (sset ss) sset"
+    unfolding smerge_def by (subst (asm) sset_flat)
+      (auto simp: stream.set_map in_set_conv_nth sset_range simp del: stake.simps, fast+)
+next
+  fix s x assume "s \<in> sset ss" "x \<in> sset s"
+  thus "x \<in> sset (smerge ss)" using snth_sset_smerge by (auto simp: sset_range)
+qed
+
+
+subsection {* product of two streams *}
+
+definition sproduct :: "'a stream \<Rightarrow> 'b stream \<Rightarrow> ('a \<times> 'b) stream" where
+  "sproduct s1 s2 = smerge (smap (\<lambda>x. smap (Pair x) s2) s1)"
+
+lemma sset_sproduct: "sset (sproduct s1 s2) = sset s1 \<times> sset s2"
+  unfolding sproduct_def sset_smerge by (auto simp: stream.set_map)
+
+
+subsection {* interleave two streams *}
+
+primcorec sinterleave where
+  "shd (sinterleave s1 s2) = shd s1"
+| "stl (sinterleave s1 s2) = sinterleave s2 (stl s1)"
+
+lemma sinterleave_code[code]:
+  "sinterleave (x ## s1) s2 = x ## sinterleave s2 s1"
+  by (subst sinterleave.ctr) simp
+
+lemma sinterleave_snth[simp]:
+  "even n \<Longrightarrow> sinterleave s1 s2 !! n = s1 !! (n div 2)"
+   "odd n \<Longrightarrow> sinterleave s1 s2 !! n = s2 !! (n div 2)"
+  by (induct n arbitrary: s1 s2)
+    (auto dest: even_nat_Suc_div_2 odd_nat_plus_one_div_two[folded nat_2])
+
+lemma sset_sinterleave: "sset (sinterleave s1 s2) = sset s1 \<union> sset s2"
+proof (intro equalityI subsetI)
+  fix x assume "x \<in> sset (sinterleave s1 s2)"
+  then obtain n where "x = sinterleave s1 s2 !! n" unfolding sset_range by blast
+  thus "x \<in> sset s1 \<union> sset s2" by (cases "even n") auto
+next
+  fix x assume "x \<in> sset s1 \<union> sset s2"
+  thus "x \<in> sset (sinterleave s1 s2)"
+  proof
+    assume "x \<in> sset s1"
+    then obtain n where "x = s1 !! n" unfolding sset_range by blast
+    hence "sinterleave s1 s2 !! (2 * n) = x" by simp
+    thus ?thesis unfolding sset_range by blast
+  next
+    assume "x \<in> sset s2"
+    then obtain n where "x = s2 !! n" unfolding sset_range by blast
+    hence "sinterleave s1 s2 !! (2 * n + 1) = x" by simp
+    thus ?thesis unfolding sset_range by blast
+  qed
+qed
+
+
+subsection {* zip *}
+
+primcorec szip where
+  "shd (szip s1 s2) = (shd s1, shd s2)"
+| "stl (szip s1 s2) = szip (stl s1) (stl s2)"
+
+lemma szip_unfold[code]: "szip (a ## s1) (b ## s2) = (a, b) ## (szip s1 s2)"
+  by (subst szip.ctr) simp
+
+lemma snth_szip[simp]: "szip s1 s2 !! n = (s1 !! n, s2 !! n)"
+  by (induct n arbitrary: s1 s2) auto
+
+lemma stake_szip[simp]:
+  "stake n (szip s1 s2) = zip (stake n s1) (stake n s2)"
+  by (induct n arbitrary: s1 s2) auto
+
+lemma sdrop_szip[simp]: "sdrop n (szip s1 s2) = szip (sdrop n s1) (sdrop n s2)"
+  by (induct n arbitrary: s1 s2) auto
+
+lemma smap_szip_fst:
+  "smap (\<lambda>x. f (fst x)) (szip s1 s2) = smap f s1"
+  by (coinduction arbitrary: s1 s2) auto
+
+lemma smap_szip_snd:
+  "smap (\<lambda>x. g (snd x)) (szip s1 s2) = smap g s2"
+  by (coinduction arbitrary: s1 s2) auto
+
+
+subsection {* zip via function *}
+
+primcorec smap2 where
+  "shd (smap2 f s1 s2) = f (shd s1) (shd s2)"
+| "stl (smap2 f s1 s2) = smap2 f (stl s1) (stl s2)"
+
+lemma smap2_unfold[code]:
+  "smap2 f (a ## s1) (b ## s2) = f a b ## (smap2 f s1 s2)"
+  by (subst smap2.ctr) simp
+
+lemma smap2_szip:
+  "smap2 f s1 s2 = smap (split f) (szip s1 s2)"
+  by (coinduction arbitrary: s1 s2) auto
+
+lemma smap_smap2[simp]:
+  "smap f (smap2 g s1 s2) = smap2 (\<lambda>x y. f (g x y)) s1 s2"
+  unfolding smap2_szip stream.map_comp o_def split_def ..
+
+lemma smap2_alt:
+  "(smap2 f s1 s2 = s) = (\<forall>n. f (s1 !! n) (s2 !! n) = s !! n)"
+  unfolding smap2_szip smap_alt by auto
+
+lemma snth_smap2[simp]:
+  "smap2 f s1 s2 !! n = f (s1 !! n) (s2 !! n)"
+  by (induct n arbitrary: s1 s2) auto
+
+lemma stake_smap2[simp]:
+  "stake n (smap2 f s1 s2) = map (split f) (zip (stake n s1) (stake n s2))"
+  by (induct n arbitrary: s1 s2) auto
+
+lemma sdrop_smap2[simp]:
+  "sdrop n (smap2 f s1 s2) = smap2 f (sdrop n s1) (sdrop n s2)"
+  by (induct n arbitrary: s1 s2) auto
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Stream_Processor.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,189 @@
+(*  Title:      HOL/Datatype_Examples/Stream_Processor.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2014
+
+Stream processors---a syntactic representation of continuous functions on streams.
+*)
+
+header {* Stream Processors---A Syntactic Representation of Continuous Functions on Streams *}
+
+theory Stream_Processor
+imports Stream "~~/src/HOL/Library/BNF_Axiomatization"
+begin
+
+section {* Continuous Functions on Streams *}
+
+datatype_new ('a, 'b, 'c) sp\<^sub>\<mu> = Get "'a \<Rightarrow> ('a, 'b, 'c) sp\<^sub>\<mu>" | Put "'b" "'c"
+codatatype ('a, 'b) sp\<^sub>\<nu> = In (out: "('a, 'b, ('a, 'b) sp\<^sub>\<nu>) sp\<^sub>\<mu>")
+
+primrec run\<^sub>\<mu> :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> 'a stream \<Rightarrow> ('b \<times> 'c) \<times> 'a stream" where
+  "run\<^sub>\<mu> (Get f) s = run\<^sub>\<mu> (f (shd s)) (stl s)"
+| "run\<^sub>\<mu> (Put b sp) s = ((b, sp), s)"
+
+primcorec run\<^sub>\<nu> :: "('a, 'b) sp\<^sub>\<nu> \<Rightarrow> 'a stream \<Rightarrow> 'b stream" where
+  "run\<^sub>\<nu> sp s = (let ((h, sp), s) = run\<^sub>\<mu> (out sp) s in h ## run\<^sub>\<nu> sp s)"
+
+primcorec copy :: "('a, 'a) sp\<^sub>\<nu>" where
+  "copy = In (Get (\<lambda>a. Put a copy))"
+
+lemma run\<^sub>\<nu>_copy: "run\<^sub>\<nu> copy s = s"
+  by (coinduction arbitrary: s) simp
+
+text {*
+To use the function package for the definition of composition the
+wellfoundedness of the subtree relation needs to be proved first.
+*}
+
+definition "sub \<equiv> {(f a, Get f) | a f. True}"
+
+lemma subI[intro]: "(f a, Get f) \<in> sub"
+  unfolding sub_def by blast
+
+lemma wf_sub[simp, intro]: "wf sub"
+proof (rule wfUNIVI)
+  fix P  :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> bool" and x
+  assume "\<forall>x. (\<forall>y. (y, x) \<in> sub \<longrightarrow> P y) \<longrightarrow> P x"
+  hence I: "\<And>x. (\<forall>y. (\<exists>a f. y = f a \<and> x = Get f) \<longrightarrow> P y) \<Longrightarrow> P x" unfolding sub_def by blast
+  show "P x" by (induct x) (auto intro: I)
+qed
+
+function
+  sp\<^sub>\<mu>_comp :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> ('d, 'a, ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu> \<Rightarrow> ('d, 'b, 'c \<times> ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu>"
+  (infixl "o\<^sub>\<mu>" 65)
+where
+  "Put b sp o\<^sub>\<mu> fsp = Put b (sp, In fsp)"
+| "Get f o\<^sub>\<mu> Put b sp = f b o\<^sub>\<mu> out sp"
+| "Get f o\<^sub>\<mu> Get g = Get (\<lambda>a. Get f o\<^sub>\<mu> g a)"
+by pat_completeness auto
+termination by (relation "lex_prod sub sub") auto
+
+primcorec sp\<^sub>\<nu>_comp (infixl "o\<^sub>\<nu>" 65) where
+  "out (sp o\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sub>\<nu> sp') (out sp o\<^sub>\<mu> out sp')"
+
+lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp: "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
+proof (rule ext, unfold comp_apply)
+  fix s
+  show "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') s = run\<^sub>\<nu> sp (run\<^sub>\<nu> sp' s)"
+  proof (coinduction arbitrary: sp sp' s)
+    case Eq_stream
+    show ?case
+    proof (induct "out sp" "out sp'" arbitrary: sp sp' s rule: sp\<^sub>\<mu>_comp.induct)
+      case (1 b sp'')
+      show ?case by (auto simp add: 1[symmetric])
+    next
+      case (2 f b sp'')
+      from 2(1)[of "In (f b)" sp''] show ?case by (simp add: 2(2,3)[symmetric])
+    next
+      case (3 f h)
+      from 3(1)[of _ "shd s" "In (h (shd s))", OF 3(2)] show ?case by (simp add: 3(2,3)[symmetric])
+    qed
+  qed
+qed
+
+text {* Alternative definition of composition using primrec instead of function *}
+
+primrec sp\<^sub>\<mu>_comp2R  where
+  "sp\<^sub>\<mu>_comp2R f (Put b sp) = f b (out sp)"
+| "sp\<^sub>\<mu>_comp2R f (Get h) = Get (sp\<^sub>\<mu>_comp2R f o h)"
+
+primrec sp\<^sub>\<mu>_comp2 (infixl "o\<^sup>*\<^sub>\<mu>" 65) where
+  "Put b sp o\<^sup>*\<^sub>\<mu> fsp = Put b (sp, In fsp)"
+| "Get f o\<^sup>*\<^sub>\<mu> fsp = sp\<^sub>\<mu>_comp2R (op o\<^sup>*\<^sub>\<mu> o f) fsp"
+
+primcorec sp\<^sub>\<nu>_comp2 (infixl "o\<^sup>*\<^sub>\<nu>" 65) where
+  "out (sp o\<^sup>*\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sup>*\<^sub>\<nu> sp') (out sp o\<^sup>*\<^sub>\<mu> out sp')"
+
+lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp2: "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
+proof (rule ext, unfold comp_apply)
+  fix s
+  show "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') s = run\<^sub>\<nu> sp (run\<^sub>\<nu> sp' s)"
+  proof (coinduction arbitrary: sp sp' s)
+    case Eq_stream
+    show ?case
+    proof (induct "out sp" arbitrary: sp sp' s)
+      case (Put b sp'')
+      show ?case by (auto simp add: Put[symmetric])
+    next
+      case (Get f)
+      then show ?case
+      proof (induct "out sp'" arbitrary: sp sp' s)
+        case (Put b sp'')
+        from Put(2)[of "In (f b)" sp''] show ?case by (simp add: Put(1,3)[symmetric])
+      next
+        case (Get h)
+        from Get(1)[OF _ Get(3,4), of "In (h (shd s))"] show ?case by (simp add: Get(2,4)[symmetric])
+      qed
+    qed
+  qed
+qed
+
+text {* The two definitions are equivalent *}
+
+lemma sp\<^sub>\<mu>_comp_sp\<^sub>\<mu>_comp2[simp]: "sp o\<^sub>\<mu> sp' = sp o\<^sup>*\<^sub>\<mu> sp'"
+  by (induct sp sp' rule: sp\<^sub>\<mu>_comp.induct) auto
+
+(*will be provided by the package*)
+lemma sp\<^sub>\<mu>_rel_map_map[unfolded vimage2p_def, simp]:
+  "rel_sp\<^sub>\<mu> R1 R2 (map_sp\<^sub>\<mu> f1 f2 sp) (map_sp\<^sub>\<mu> g1 g2 sp') =
+  rel_sp\<^sub>\<mu> (BNF_Def.vimage2p f1 g1 R1) (BNF_Def.vimage2p f2 g2 R2) sp sp'"
+by (tactic {*
+  let val ks = 1 upto 2;
+  in
+    BNF_Tactics.unfold_thms_tac @{context}
+      @{thms sp\<^sub>\<mu>.rel_compp sp\<^sub>\<mu>.rel_conversep sp\<^sub>\<mu>.rel_Grp vimage2p_Grp} THEN
+    HEADGOAL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
+      BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac,
+      rtac @{thm conversepI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
+      BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks,
+      REPEAT_DETERM o eresolve_tac @{thms relcomppE conversepE GrpE},
+      hyp_subst_tac @{context}, atac])
+  end
+*})
+
+lemma sp\<^sub>\<mu>_rel_self: "\<lbrakk>op = \<le> R1; op = \<le> R2\<rbrakk> \<Longrightarrow> rel_sp\<^sub>\<mu> R1 R2 x x"
+  by (erule (1) predicate2D[OF sp\<^sub>\<mu>.rel_mono]) (simp only: sp\<^sub>\<mu>.rel_eq)
+
+lemma sp\<^sub>\<nu>_comp_sp\<^sub>\<nu>_comp2: "sp o\<^sub>\<nu> sp' = sp o\<^sup>*\<^sub>\<nu> sp'"
+  by (coinduction arbitrary: sp sp') (auto intro!: sp\<^sub>\<mu>_rel_self)
+
+
+section {* Generalization to an Arbitrary BNF as Codomain *}
+
+bnf_axiomatization ('a, 'b) F for map: F
+
+notation BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
+
+definition \<theta> :: "('p,'a) F * 'b \<Rightarrow> ('p,'a * 'b) F" where
+  "\<theta> xb = F id \<langle>id, \<lambda> a. (snd xb)\<rangle> (fst xb)"
+
+(* The strength laws for \<theta>: *)
+lemma \<theta>_natural: "F id (map_prod f g) o \<theta> = \<theta> o map_prod (F id f) g"
+  unfolding \<theta>_def F.map_comp comp_def id_apply convol_def map_prod_def split_beta fst_conv snd_conv ..
+
+definition assl :: "'a * ('b * 'c) \<Rightarrow> ('a * 'b) * 'c" where
+  "assl abc = ((fst abc, fst (snd abc)), snd (snd abc))"
+
+lemma \<theta>_rid: "F id fst o \<theta> = fst"
+  unfolding \<theta>_def F.map_comp F.map_id comp_def id_apply convol_def fst_conv sym[OF id_def] ..
+
+lemma \<theta>_assl: "F id assl o \<theta> = \<theta> o map_prod \<theta> id o assl"
+  unfolding assl_def \<theta>_def F.map_comp comp_def id_apply convol_def map_prod_def split fst_conv snd_conv ..
+
+datatype_new ('a, 'b, 'c) spF\<^sub>\<mu> = GetF "'a \<Rightarrow> ('a, 'b, 'c) spF\<^sub>\<mu>" | PutF "('b,'c) F"
+codatatype ('a, 'b) spF\<^sub>\<nu> = InF (outF: "('a, 'b, ('a, 'b) spF\<^sub>\<nu>) spF\<^sub>\<mu>")
+
+codatatype 'b JF = Ctor (dtor: "('b, 'b JF) F")
+
+(* Definition of run for an arbitrary final coalgebra as codomain: *)
+
+primrec
+  runF\<^sub>\<mu> :: "('a, 'b, ('a, 'b) spF\<^sub>\<nu>) spF\<^sub>\<mu> \<Rightarrow> 'a stream \<Rightarrow> (('b, ('a, 'b) spF\<^sub>\<nu>) F) \<times> 'a stream" 
+where
+  "runF\<^sub>\<mu> (GetF f) s = (runF\<^sub>\<mu> o f) (shd s) (stl s)"
+| "runF\<^sub>\<mu> (PutF x) s = (x, s)"
+
+primcorec runF\<^sub>\<nu> :: "('a, 'b) spF\<^sub>\<nu> \<Rightarrow> 'a stream \<Rightarrow> 'b JF" where
+  "runF\<^sub>\<nu> sp s = (let (x, s) = runF\<^sub>\<mu> (outF sp) s in Ctor (F id (\<lambda> sp. runF\<^sub>\<nu> sp s) x))"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/TreeFI.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,46 @@
+(*  Title:      HOL/Datatype_Examples/TreeFI.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Finitely branching possibly infinite trees.
+*)
+
+header {* Finitely Branching Possibly Infinite Trees *}
+
+theory TreeFI
+imports Main
+begin
+
+codatatype 'a treeFI = Tree (lab: 'a) (sub: "'a treeFI list")
+
+(* Tree reverse:*)
+primcorec trev where
+  "lab (trev t) = lab t"
+| "sub (trev t) = map trev (rev (sub t))"
+
+lemma treeFI_coinduct:
+  assumes *: "phi x y"
+  and step: "\<And>a b. phi a b \<Longrightarrow>
+     lab a = lab b \<and>
+     length (sub a) = length (sub b) \<and>
+     (\<forall>i < length (sub a). phi (nth (sub a) i) (nth (sub b) i))"
+  shows "x = y"
+using * proof (coinduction arbitrary: x y)
+  case (Eq_treeFI t1 t2)
+  from conjunct1[OF conjunct2[OF step[OF Eq_treeFI]]] conjunct2[OF conjunct2[OF step[OF Eq_treeFI]]]
+  have "list_all2 phi (sub t1) (sub t2)"
+  proof (induction "sub t1" "sub t2" arbitrary: t1 t2 rule: list_induct2)
+    case (Cons x xs y ys)
+    note sub = Cons(3,4)[symmetric] and phi = mp[OF spec[OF Cons(5)], unfolded sub]
+      and IH = Cons(2)[of "Tree (lab t1) (tl (sub t1))" "Tree (lab t2) (tl (sub t2))",
+        unfolded sub, simplified]
+    from phi[of 0] show ?case unfolding sub by (auto intro!: IH dest: phi[simplified, OF Suc_mono])
+  qed simp
+  with conjunct1[OF step[OF Eq_treeFI]] show ?case by simp
+qed
+
+lemma trev_trev: "trev (trev tr) = tr"
+  by (coinduction arbitrary: tr rule: treeFI_coinduct) (auto simp add: rev_map)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/TreeFsetI.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,25 @@
+(*  Title:      HOL/Datatype_Examples/TreeFsetI.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Finitely branching possibly infinite trees, with sets of children.
+*)
+
+header {* Finitely Branching Possibly Infinite Trees, with Sets of Children *}
+
+theory TreeFsetI
+imports "~~/src/HOL/Library/FSet"
+begin
+
+codatatype 'a treeFsetI = Tree (lab: 'a) (sub: "'a treeFsetI fset")
+
+(* tree map (contrived example): *)
+primcorec tmap where
+"lab (tmap f t) = f (lab t)" |
+"sub (tmap f t) = fimage (tmap f) (sub t)"
+
+lemma "tmap (f o g) x = tmap f (tmap g x)"
+  by (coinduction arbitrary: x) (auto simp: rel_fset_alt)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Datatype_Examples/Verilog.thy	Thu Sep 11 19:26:59 2014 +0200
@@ -0,0 +1,138 @@
+(*  Title:      HOL/Datatype_Benchmark/Verilog.thy
+
+Example from Daryl: a Verilog grammar.
+*)
+
+theory Verilog imports Main begin
+
+datatype
+  Source_text
+     = module string "string list" "Module_item list"
+     | Source_textMeta string
+and
+  Module_item
+     = "declaration" Declaration
+     | initial Statement
+     | always Statement
+     | assign Lvalue Expression
+     | Module_itemMeta string
+and
+  Declaration
+     = reg_declaration "Range option" "string list"
+     | net_declaration "Range option" "string list"
+     | input_declaration "Range option" "string list"
+     | output_declaration "Range option" "string list"
+     | DeclarationMeta string
+and
+  Range = range Expression Expression | RangeMeta string
+and
+  Statement
+     = clock_statement Clock Statement_or_null
+     | blocking_assignment Lvalue Expression
+     | non_blocking_assignment Lvalue Expression
+     | conditional_statement
+          Expression Statement_or_null "Statement_or_null option"
+     | case_statement Expression "Case_item list"
+     | while_loop Expression Statement
+     | repeat_loop Expression Statement
+     | for_loop
+          Lvalue Expression Expression Lvalue Expression Statement
+     | forever_loop Statement
+     | disable string
+     | seq_block "string option" "Statement list"
+     | StatementMeta string
+and
+  Statement_or_null
+     = statement Statement | null_statement | Statement_or_nullMeta string
+and
+  Clock
+     = posedge string
+     | negedge string
+     | clock string
+     | ClockMeta string
+and
+  Case_item
+     = case_item "Expression list" Statement_or_null
+     | default_case_item Statement_or_null
+     | Case_itemMeta string
+and
+  Expression
+     = plus Expression Expression
+     | minus Expression Expression
+     | lshift Expression Expression
+     | rshift Expression Expression
+     | lt Expression Expression
+     | leq Expression Expression
+     | gt Expression Expression
+     | geq Expression Expression
+     | logeq Expression Expression
+     | logneq Expression Expression
+     | caseeq Expression Expression
+     | caseneq Expression Expression
+     | bitand Expression Expression
+     | bitxor Expression Expression
+     | bitor Expression Expression
+     | logand Expression Expression
+     | logor Expression Expression
+     | conditional Expression Expression Expression
+     | positive Primary
+     | negative Primary
+     | lognot Primary
+     | bitnot Primary
+     | reducand Primary
+     | reducxor Primary
+     | reducor Primary
+     | reducnand Primary
+     | reducxnor Primary
+     | reducnor Primary
+     | primary Primary
+     | ExpressionMeta string
+and
+  Primary
+     = primary_number Number
+     | primary_IDENTIFIER string
+     | primary_bit_select string Expression
+     | primary_part_select string Expression Expression
+     | primary_gen_bit_select Expression Expression
+     | primary_gen_part_select Expression Expression Expression
+     | primary_concatenation Concatenation
+     | primary_multiple_concatenation Multiple_concatenation
+     | brackets Expression
+     | PrimaryMeta string
+and
+  Lvalue
+     = lvalue string
+     | lvalue_bit_select string Expression
+     | lvalue_part_select string Expression Expression
+     | lvalue_concatenation Concatenation
+     | LvalueMeta string
+and
+  Number
+     = decimal string
+     | based "string option" string
+     | NumberMeta string
+and
+  Concatenation
+     = concatenation "Expression list" | ConcatenationMeta string
+and
+  Multiple_concatenation
+     = multiple_concatenation Expression "Expression list"
+     | Multiple_concatenationMeta string
+and
+  meta
+     = Meta_Source_text Source_text
+     | Meta_Module_item Module_item
+     | Meta_Declaration Declaration
+     | Meta_Range Range
+     | Meta_Statement Statement
+     | Meta_Statement_or_null Statement_or_null
+     | Meta_Clock Clock
+     | Meta_Case_item Case_item
+     | Meta_Expression Expression
+     | Meta_Primary Primary
+     | Meta_Lvalue Lvalue
+     | Meta_Number Number
+     | Meta_Concatenation Concatenation
+     | Meta_Multiple_concatenation Multiple_concatenation
+
+end
--- a/src/HOL/ROOT	Thu Sep 11 19:20:23 2014 +0200
+++ b/src/HOL/ROOT	Thu Sep 11 19:26:59 2014 +0200
@@ -730,9 +730,9 @@
     "root.tex"
     "root.bib"
 
-session "HOL-BNF_Examples" in BNF_Examples = HOL +
+session "HOL-Datatype_Examples" in Datatype_Examples = HOL +
   description {*
-    Examples for Bounded Natural Functors.
+    (Co)datatype Examples.
   *}
   options [document = false]
   theories