src/HOL/BNF/Examples/Derivation_Trees/Parallel.thy
author wenzelm
Fri, 09 Aug 2013 15:14:59 +0200
changeset 52940 6fce81e92e7c
parent 49882 946efb120c42
child 54538 ba7392b52a7c
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
     1
(*  Title:      HOL/BNF/Examples/Derivation_Trees/Parallel.thy
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     2
    Author:     Andrei Popescu, TU Muenchen
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     3
    Copyright   2012
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     4
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     5
Parallel composition.
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     6
*)
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     7
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     8
header {* Parallel Composition *}
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 48975
diff changeset
     9
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    10
theory Parallel
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    11
imports DTree
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
49839
9cbec40e88ea disambiguated grammar
traytel
parents: 49514
diff changeset
    14
no_notation plus_class.plus (infixl "+" 65)
9cbec40e88ea disambiguated grammar
traytel
parents: 49514
diff changeset
    15
no_notation Sublist.parallel (infixl "\<parallel>" 50)
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    16
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl "+" 60)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    19
axiomatization where
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
    Nplus_comm: "(a::N) + b = b + (a::N)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
and Nplus_assoc: "((a::N) + b) + c = a + (b + c)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
49882
946efb120c42 tuned for document output
traytel
parents: 49880
diff changeset
    23
subsection{* Corecursive Definition of Parallel Composition *}
48975
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
fun par_r where "par_r (tr1,tr2) = root tr1 + root tr2"
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    26
fun par_c where
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    27
"par_c (tr1,tr2) =
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    28
 Inl ` (Inl -` (cont tr1 \<union> cont tr2)) \<union>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    29
 Inr ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
declare par_r.simps[simp del]  declare par_c.simps[simp del]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    33
definition par :: "dtree \<times> dtree \<Rightarrow> dtree" where
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    34
"par \<equiv> unfold par_r par_c"
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
abbreviation par_abbr (infixr "\<parallel>" 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
lemma finite_par_c: "finite (par_c (tr1, tr2))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
unfolding par_c.simps apply(rule finite_UnI)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
  apply (metis finite_Un finite_cont finite_imageI finite_vimageI inj_Inl)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  apply(intro finite_imageI finite_cartesian_product finite_vimageI)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    42
  using finite_cont by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
lemma root_par: "root (tr1 \<parallel> tr2) = root tr1 + root tr2"
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    45
using unfold(1)[of par_r par_c "(tr1,tr2)"] unfolding par_def par_r.simps by simp
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    47
lemma cont_par:
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
"cont (tr1 \<parallel> tr2) = (id \<oplus> par) ` par_c (tr1,tr2)"
49508
1e205327f059 adapted examples to renamings
blanchet
parents: 49463
diff changeset
    49
using unfold(2)[of par_c "(tr1,tr2)" par_r, OF finite_par_c]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
unfolding par_def ..
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    51
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
lemma Inl_cont_par[simp]:
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    53
"Inl -` (cont (tr1 \<parallel> tr2)) = Inl -` (cont tr1 \<union> cont tr2)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    54
unfolding cont_par par_c.simps by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    56
lemma Inr_cont_par[simp]:
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
    57
"Inr -` (cont (tr1 \<parallel> tr2)) = par ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    58
unfolding cont_par par_c.simps by auto
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
lemma Inl_in_cont_par:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
"Inl t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (Inl t \<in> cont tr1 \<or> Inl t \<in> cont tr2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
using Inl_cont_par[of tr1 tr2] unfolding vimage_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
lemma Inr_in_cont_par:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
"Inr t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (t \<in> par ` (Inr -` cont tr1 \<times> Inr -` cont tr2))"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
using Inr_cont_par[of tr1 tr2] unfolding vimage_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
49882
946efb120c42 tuned for document output
traytel
parents: 49880
diff changeset
    69
subsection{* Structural Coinduction Proofs *}
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    70
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    71
lemma set_rel_sum_rel_eq[simp]:
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    72
"set_rel (sum_rel (op =) \<phi>) A1 A2 \<longleftrightarrow>
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    73
 Inl -` A1 = Inl -` A2 \<and> set_rel \<phi> (Inr -` A1) (Inr -` A2)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    74
unfolding set_rel_sum_rel set_rel_eq ..
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    76
(* Detailed proofs of commutativity and associativity: *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
theorem par_com: "tr1 \<parallel> tr2 = tr2 \<parallel> tr1"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
proof-
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    79
  let ?\<theta> = "\<lambda> trA trB. \<exists> tr1 tr2. trA = tr1 \<parallel> tr2 \<and> trB = tr2 \<parallel> tr1"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
  {fix trA trB
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    81
   assume "?\<theta> trA trB" hence "trA = trB"
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
    82
   apply (induct rule: dtree_coinduct)
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    83
   unfolding set_rel_sum_rel set_rel_eq unfolding set_rel_def proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    84
     fix tr1 tr2  show "root (tr1 \<parallel> tr2) = root (tr2 \<parallel> tr1)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
     unfolding root_par by (rule Nplus_comm)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
   next
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    87
     fix n tr1 tr2 assume "Inl n \<in> cont (tr1 \<parallel> tr2)" thus "n \<in> Inl -` (cont (tr2 \<parallel> tr1))"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    88
     unfolding Inl_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    89
   next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    90
     fix n tr1 tr2 assume "Inl n \<in> cont (tr2 \<parallel> tr1)" thus "n \<in> Inl -` (cont (tr1 \<parallel> tr2))"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    91
     unfolding Inl_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    92
   next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    93
     fix tr1 tr2 trA' assume "Inr trA' \<in> cont (tr1 \<parallel> tr2)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    94
     then obtain tr1' tr2' where "trA' = tr1' \<parallel> tr2'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    95
     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    96
     unfolding Inr_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    97
     thus "\<exists> trB' \<in> Inr -` (cont (tr2 \<parallel> tr1)). ?\<theta> trA' trB'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    98
     apply(intro bexI[of _ "tr2' \<parallel> tr1'"]) unfolding Inr_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
    99
   next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   100
     fix tr1 tr2 trB' assume "Inr trB' \<in> cont (tr2 \<parallel> tr1)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   101
     then obtain tr1' tr2' where "trB' = tr2' \<parallel> tr1'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   102
     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   103
     unfolding Inr_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   104
     thus "\<exists> trA' \<in> Inr -` (cont (tr1 \<parallel> tr2)). ?\<theta> trA' trB'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   105
     apply(intro bexI[of _ "tr1' \<parallel> tr2'"]) unfolding Inr_in_cont_par by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
   qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
  thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   110
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   111
lemma par_assoc: "(tr1 \<parallel> tr2) \<parallel> tr3 = tr1 \<parallel> (tr2 \<parallel> tr3)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
proof-
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   113
  let ?\<theta> =
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   114
  "\<lambda> trA trB. \<exists> tr1 tr2 tr3. trA = (tr1 \<parallel> tr2) \<parallel> tr3 \<and> trB = tr1 \<parallel> (tr2 \<parallel> tr3)"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
  {fix trA trB
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   116
   assume "?\<theta> trA trB" hence "trA = trB"
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
   117
   apply (induct rule: dtree_coinduct)
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   118
   unfolding set_rel_sum_rel set_rel_eq unfolding set_rel_def proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   119
     fix tr1 tr2 tr3  show "root ((tr1 \<parallel> tr2) \<parallel> tr3) = root (tr1 \<parallel> (tr2 \<parallel> tr3))"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
     unfolding root_par by (rule Nplus_assoc)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
   next
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
   122
     fix n tr1 tr2 tr3 assume "Inl n \<in> (cont ((tr1 \<parallel> tr2) \<parallel> tr3))"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   123
     thus "n \<in> Inl -` (cont (tr1 \<parallel> tr2 \<parallel> tr3))" unfolding Inl_in_cont_par by simp
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   124
   next
49879
5e323695f26e tuned whitespace
traytel
parents: 49877
diff changeset
   125
     fix n tr1 tr2 tr3 assume "Inl n \<in> (cont (tr1 \<parallel> tr2 \<parallel> tr3))"
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   126
     thus "n \<in> Inl -` (cont ((tr1 \<parallel> tr2) \<parallel> tr3))" unfolding Inl_in_cont_par by simp
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   127
   next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   128
     fix trA' tr1 tr2 tr3 assume "Inr trA' \<in> cont ((tr1 \<parallel> tr2) \<parallel> tr3)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   129
     then obtain tr1' tr2' tr3' where "trA' = (tr1' \<parallel> tr2') \<parallel> tr3'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   130
     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   131
     and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   132
     thus "\<exists> trB' \<in> Inr -` (cont (tr1 \<parallel> tr2 \<parallel> tr3)). ?\<theta> trA' trB'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   133
     apply(intro bexI[of _ "tr1' \<parallel> tr2' \<parallel> tr3'"])
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   134
     unfolding Inr_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   135
   next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   136
     fix trB' tr1 tr2 tr3 assume "Inr trB' \<in> cont (tr1 \<parallel> tr2 \<parallel> tr3)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   137
     then obtain tr1' tr2' tr3' where "trB' = tr1' \<parallel> (tr2' \<parallel> tr3')"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   138
     and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   139
     and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   140
     thus "\<exists> trA' \<in> Inr -` cont ((tr1 \<parallel> tr2) \<parallel> tr3). ?\<theta> trA' trB'"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   141
     apply(intro bexI[of _ "(tr1' \<parallel> tr2') \<parallel> tr3'"])
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49871
diff changeset
   142
     unfolding Inr_in_cont_par by auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
   qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   144
  }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
  thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   146
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   147
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   148
end