src/HOL/Codatatype/Examples/Infinite_Derivation_Trees/Tree.thy
author blanchet
Sat, 08 Sep 2012 21:30:31 +0200
changeset 49222 cbe8c859817c
parent 49220 a6260b4fb410
child 49463 83ac281bcdc2
permissions -rw-r--r--
for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     1
(*  Title:      Gram_Tree.thy
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
Trees with nonterminal internal nodes and terminal leafs.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
header {* Trees with nonterminal internal nodes and terminal leafs *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
theory Tree
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
imports Prelim
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
typedecl N  typedecl T
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
49076
d2ed455fa3d2 compile
blanchet
parents: 48975
diff changeset
    18
codata_raw Tree: 'Tree = "N \<times> (T + 'Tree) fset"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
section {* Sugar notations for Tree *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
subsection{* Setup for map, set, pred *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
(* These should be eventually inferred from compositionality *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    27
lemma pre_Tree_map:
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    28
"pre_Tree_map f (n, as) = (n, map_fset (id \<oplus> f) as)"
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    29
unfolding pre_Tree_map_def id_apply
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    30
sum_map_def by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    32
lemma pre_Tree_map':
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    33
"pre_Tree_map f n_as = (fst n_as, map_fset (id \<oplus> f) (snd n_as))"
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    34
using pre_Tree_map by(cases n_as, simp)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    35
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    37
definition
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    38
"llift2 \<phi> as1 as2 \<longleftrightarrow>
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    39
 (\<forall> n. Inl n \<in> fset as1 \<longleftrightarrow> Inl n \<in> fset as2) \<and>
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    40
 (\<forall> tr1. Inr tr1 \<in> fset as1 \<longrightarrow> (\<exists> tr2. Inr tr2 \<in> fset as2 \<and> \<phi> tr1 tr2)) \<and>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
 (\<forall> tr2. Inr tr2 \<in> fset as2 \<longrightarrow> (\<exists> tr1. Inr tr1 \<in> fset as1 \<and> \<phi> tr1 tr2))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    43
lemma pre_Tree_pred: "pre_Tree_pred \<phi> (n1,as1) (n2,as2) \<longleftrightarrow> n1 = n2 \<and> llift2 \<phi> as1 as2"
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
    44
unfolding llift2_def pre_Tree.pred_unfold
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
apply (auto split: sum.splits)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
apply (metis sumE)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
apply (metis sumE)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
apply (metis sumE)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
apply (metis sumE)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
apply (metis sumE sum.simps(1,2,4))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
apply (metis sumE sum.simps(1,2,4))
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
done
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
subsection{* Constructors *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    57
definition NNode :: "N \<Rightarrow> (T + Tree)fset \<Rightarrow> Tree"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
where "NNode n as \<equiv> Tree_fld (n,as)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    59
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    60
lemmas ctor_defs = NNode_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
subsection {* Pre-selectors *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
(* These are mere auxiliaries *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
definition "asNNode tr \<equiv> SOME n_as. NNode (fst n_as) (snd n_as) = tr"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    68
lemmas pre_sel_defs = asNNode_def
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
subsection {* Selectors *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
(* One for each pair (constructor, constructor argument) *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    74
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
(* For NNode: *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
definition root :: "Tree \<Rightarrow> N" where "root tr = fst (asNNode tr)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
definition ccont :: "Tree \<Rightarrow> (T + Tree)fset" where "ccont tr = snd (asNNode tr)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
lemmas sel_defs = root_def ccont_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
subsection {* Basic properties *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
(* Constructors versus selectors *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
lemma NNode_surj: "\<exists> n as. NNode n as = tr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
unfolding NNode_def
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    87
by (metis Tree.fld_unf pair_collapse)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    89
lemma NNode_asNNode:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
"NNode (fst (asNNode tr)) (snd (asNNode tr)) = tr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
  obtain n as where "NNode n as = tr" using NNode_surj[of tr] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
  hence "NNode (fst (n,as)) (snd (n,as)) = tr" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
  thus ?thesis unfolding asNNode_def by(rule someI)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
    97
theorem NNode_root_ccont[simp]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
"NNode (root tr) (ccont tr) = tr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
using NNode_asNNode unfolding root_def ccont_def .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
(* Constructors *)
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   102
theorem TTree_simps[simp]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
"NNode n as = NNode n' as' \<longleftrightarrow> n = n' \<and> as = as'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
unfolding ctor_defs Tree.fld_inject by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
theorem TTree_cases[elim, case_names NNode Choice]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
assumes NNode: "\<And> n as. tr = NNode n as \<Longrightarrow> phi"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
shows phi
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
proof(cases rule: Tree.fld_exhaust[of tr])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
  fix x assume "tr = Tree_fld x"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
  thus ?thesis
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   112
  apply(cases x)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
    using NNode unfolding ctor_defs apply blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
  done
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
(* Constructors versus selectors *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
theorem TTree_sel_ctor[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
"root (NNode n as) = n"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
"ccont (NNode n as) = as"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
unfolding root_def ccont_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
by (metis (no_types) NNode_asNNode TTree_simps)+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
subsection{* Coinduction *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
theorem TTree_coind_Node[elim, consumes 1, case_names NNode, induct pred: "HOL.eq"]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   128
assumes phi: "\<phi> tr1 tr2" and
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   129
NNode: "\<And> n1 n2 as1 as2.
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   130
          \<lbrakk>\<phi> (NNode n1 as1) (NNode n2 as2)\<rbrakk> \<Longrightarrow>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
          n1 = n2 \<and> llift2 \<phi> as1 as2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
shows "tr1 = tr2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
apply(rule mp[OF Tree.pred_coinduct[of \<phi> tr1 tr2] phi]) proof clarify
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
  fix tr1 tr2  assume \<phi>: "\<phi> tr1 tr2"
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
   135
  show "pre_Tree_pred \<phi> (Tree_unf tr1) (Tree_unf tr2)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
  apply(cases rule: Tree.fld_exhaust[of tr1], cases rule: Tree.fld_exhaust[of tr2])
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   137
  apply (simp add: Tree.unf_fld)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
  apply(case_tac x, case_tac xa, simp)
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
   139
  unfolding pre_Tree_pred apply(rule NNode) using \<phi> unfolding NNode_def by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   142
theorem TTree_coind[elim, consumes 1, case_names LLift]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   143
assumes phi: "\<phi> tr1 tr2" and
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   144
LLift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
                   root tr1 = root tr2 \<and> llift2 \<phi> (ccont tr1) (ccont tr2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
shows "tr1 = tr2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   147
using phi apply(induct rule: TTree_coind_Node)
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   148
using LLift by (metis TTree_sel_ctor)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   149
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   150
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
subsection {* Coiteration *}
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   153
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   154
(* Preliminaries: *)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
declare Tree.unf_fld[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
declare Tree.fld_unf[simp]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
lemma Tree_unf_NNode[simp]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
"Tree_unf (NNode n as) = (n,as)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
unfolding NNode_def Tree.unf_fld ..
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   161
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
lemma Tree_unf_root_ccont:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
"Tree_unf tr = (root tr, ccont tr)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
unfolding root_def ccont_def
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   165
by (metis (lifting) NNode_asNNode Tree_unf_NNode)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
(* Coiteration *)
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   168
definition TTree_coit ::
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + 'b) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   170
where "TTree_coit rt ct \<equiv> Tree_unf_coiter <rt,ct>"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   172
lemma Tree_coit_coit:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   173
"Tree_unf_coiter s = TTree_coit (fst o s) (snd o s)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
apply(rule ext)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
unfolding TTree_coit_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   177
theorem TTree_coit:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   178
"root (TTree_coit rt ct b) = rt b"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
"ccont (TTree_coit rt ct b) = map_fset (id \<oplus> TTree_coit rt ct) (ct b)"
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49220
diff changeset
   180
using Tree.unf_coiters[of "<rt,ct>" b] unfolding Tree_coit_coit fst_convol snd_convol
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
   181
unfolding pre_Tree_map' fst_convol' snd_convol'
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   182
unfolding Tree_unf_root_ccont by simp_all
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   184
(* Corecursion, stronger than coiteration *)
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   185
definition TTree_corec ::
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + (Tree + 'b)) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   187
where "TTree_corec rt ct \<equiv> Tree_unf_corec <rt,ct>"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   189
lemma Tree_unf_corec_corec:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   190
"Tree_unf_corec s = TTree_corec (fst o s) (snd o s)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
apply(rule ext)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
unfolding TTree_corec_def by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   193
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   194
theorem TTree_corec:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   195
"root (TTree_corec rt ct b) = rt b"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
"ccont (TTree_corec rt ct b) = map_fset (id \<oplus> ([[id, TTree_corec rt ct]]) ) (ct b)"
49222
cbe8c859817c for compatiblity with old datatype package: not only "recs" with "s", but also "iters" and their "fld_"/"unf_" variants
blanchet
parents: 49220
diff changeset
   197
using Tree.unf_corecs[of "<rt,ct>" b] unfolding Tree_unf_corec_corec fst_convol snd_convol
49220
a6260b4fb410 imported patch debugging
blanchet
parents: 49128
diff changeset
   198
unfolding pre_Tree_map' fst_convol' snd_convol'
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
unfolding Tree_unf_root_ccont by simp_all
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
subsection{* The characteristic theorems transported from fset to set *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   203
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   204
definition "Node n as \<equiv> NNode n (the_inv fset as)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   205
definition "cont \<equiv> fset o ccont"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   206
definition "coit rt ct \<equiv> TTree_coit rt (the_inv fset o ct)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   207
definition "corec rt ct \<equiv> TTree_corec rt (the_inv fset o ct)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   208
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   209
definition lift ("_ ^#" 200) where
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   210
"lift \<phi> as \<longleftrightarrow> (\<forall> tr. Inr tr \<in> as \<longrightarrow> \<phi> tr)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   211
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   212
definition lift2 ("_ ^#2" 200) where
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   213
"lift2 \<phi> as1 as2 \<longleftrightarrow>
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   214
 (\<forall> n. Inl n \<in> as1 \<longleftrightarrow> Inl n \<in> as2) \<and>
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   215
 (\<forall> tr1. Inr tr1 \<in> as1 \<longrightarrow> (\<exists> tr2. Inr tr2 \<in> as2 \<and> \<phi> tr1 tr2)) \<and>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
 (\<forall> tr2. Inr tr2 \<in> as2 \<longrightarrow> (\<exists> tr1. Inr tr1 \<in> as1 \<and> \<phi> tr1 tr2))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   217
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   218
definition liftS ("_ ^#s" 200) where
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
"liftS trs = {as. Inr -` as \<subseteq> trs}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   221
lemma lift2_llift2:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   222
"\<lbrakk>finite as1; finite as2\<rbrakk> \<Longrightarrow>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   223
 lift2 \<phi> as1 as2 \<longleftrightarrow> llift2 \<phi> (the_inv fset as1) (the_inv fset as2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
unfolding lift2_def llift2_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   225
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   226
lemma llift2_lift2:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
"llift2 \<phi> as1 as2 \<longleftrightarrow> lift2 \<phi> (fset as1) (fset as2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
using lift2_llift2 by (metis finite_fset fset_cong fset_to_fset)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
lemma mono_lift:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   231
assumes "(\<phi>^#) as"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
and "\<And> tr. \<phi> tr \<Longrightarrow> \<phi>' tr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
shows "(\<phi>'^#) as"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
using assms unfolding lift_def[abs_def] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   235
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
lemma mono_liftS:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
assumes "trs1 \<subseteq> trs2 "
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   238
shows "(trs1 ^#s) \<subseteq> (trs2 ^#s)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
using assms unfolding liftS_def[abs_def] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   241
lemma lift_mono:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
assumes "\<phi> \<le> \<phi>'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
shows "(\<phi>^#) \<le> (\<phi>'^#)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
using assms unfolding lift_def[abs_def] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
lemma mono_lift2:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   247
assumes "(\<phi>^#2) as1 as2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   248
and "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow> \<phi>' tr1 tr2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   249
shows "(\<phi>'^#2) as1 as2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   250
using assms unfolding lift2_def[abs_def] by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   251
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   252
lemma lift2_mono:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
assumes "\<phi> \<le> \<phi>'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   254
shows "(\<phi>^#2) \<le> (\<phi>'^#2)"
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   255
using assms unfolding lift2_def[abs_def] by blast
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
lemma finite_cont[simp]: "finite (cont tr)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
unfolding cont_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   260
theorem Node_root_cont[simp]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
"Node (root tr) (cont tr) = tr"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
using NNode_root_ccont unfolding Node_def cont_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
by (metis cont_def finite_cont fset_cong fset_to_fset o_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   265
theorem Tree_simps[simp]:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
assumes "finite as" and "finite as'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
shows "Node n as = Node n' as' \<longleftrightarrow> n = n' \<and> as = as'"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
using assms TTree_simps unfolding Node_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
by (metis fset_to_fset)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
theorem Tree_cases[elim, case_names Node Choice]:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
assumes Node: "\<And> n as. \<lbrakk>finite as; tr = Node n as\<rbrakk> \<Longrightarrow> phi"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
shows phi
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   274
apply(cases rule: TTree_cases[of tr])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   275
using Node unfolding Node_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   276
by (metis Node Node_root_cont finite_cont)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
theorem Tree_sel_ctor[simp]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   279
"root (Node n as) = n"
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   280
"finite as \<Longrightarrow> cont (Node n as) = as"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
unfolding Node_def cont_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
theorems root_Node = Tree_sel_ctor(1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
theorems cont_Node = Tree_sel_ctor(2)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
theorem Tree_coind_Node[elim, consumes 1, case_names Node]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   287
assumes phi: "\<phi> tr1 tr2" and
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   288
Node:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   289
"\<And> n1 n2 as1 as2.
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   290
   \<lbrakk>finite as1; finite as2; \<phi> (Node n1 as1) (Node n2 as2)\<rbrakk>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   291
   \<Longrightarrow> n1 = n2 \<and> (\<phi>^#2) as1 as2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   292
shows "tr1 = tr2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   293
using phi apply(induct rule: TTree_coind_Node)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   294
unfolding llift2_lift2 apply(rule Node)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   295
unfolding Node_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   296
apply (metis finite_fset)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   297
apply (metis finite_fset)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   298
by (metis finite_fset fset_cong fset_to_fset)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   299
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
theorem Tree_coind[elim, consumes 1, case_names Lift, induct pred: "HOL.eq"]:
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   301
assumes phi: "\<phi> tr1 tr2" and
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   302
Lift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   303
                  root tr1 = root tr2 \<and> (\<phi>^#2) (cont tr1) (cont tr2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   304
shows "tr1 = tr2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   305
using phi apply(induct rule: TTree_coind)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   306
unfolding llift2_lift2 apply(rule Lift[unfolded cont_def comp_def]) .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   307
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   308
theorem coit:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   309
"root (coit rt ct b) = rt b"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   310
"finite (ct b) \<Longrightarrow> cont (coit rt ct b) = image (id \<oplus> coit rt ct) (ct b)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   311
using TTree_coit[of rt "the_inv fset \<circ> ct" b] unfolding coit_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   312
apply - apply metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   313
unfolding cont_def comp_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   314
by (metis (no_types) fset_to_fset map_fset_image)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   315
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   316
49128
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   317
theorem corec:
1a86ef0a0210 renamed low-level (co)iterators and (co)recursors with "fld_"/"unf_" prefix
blanchet
parents: 49076
diff changeset
   318
"root (corec rt ct b) = rt b"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   319
"finite (ct b) \<Longrightarrow> cont (corec rt ct b) = image (id \<oplus> ([[id, corec rt ct]])) (ct b)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   320
using TTree_corec[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   321
apply - apply metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   322
unfolding cont_def comp_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   323
by (metis (no_types) fset_to_fset map_fset_image)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   324
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   325
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   326
end