Regular_Set.thy
author urbanc
Mon, 13 Feb 2012 21:34:19 +0000
changeset 316 0423e4d7c77b
parent 203 5d724fe0e096
child 372 2c56b20032a7
permissions -rw-r--r--
more conclusion
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     1
(*  Author: Tobias Nipkow, Alex Krauss  *)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     2
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     3
header "Regular sets"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     4
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     5
theory Regular_Set
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     6
imports Main
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     7
begin
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     8
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
     9
type_synonym 'a lang = "'a list set"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    10
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    11
definition conc :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a lang" (infixr "@@" 75) where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    12
"A @@ B = {xs@ys | xs ys. xs:A & ys:B}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    13
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    14
overloading lang_pow == "compow :: nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    15
begin
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    16
  primrec lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    17
  "lang_pow 0 A = {[]}" |
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    18
  "lang_pow (Suc n) A = A @@ (lang_pow n A)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    19
end
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    20
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    21
definition star :: "'a lang \<Rightarrow> 'a lang" where
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    22
"star A = (\<Union>n. A ^^ n)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    23
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    24
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    25
subsection{* @{term "op @@"} *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    26
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    27
lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A @@ B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    28
by (auto simp add: conc_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    29
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    30
lemma concE[elim]: 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    31
assumes "w \<in> A @@ B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    32
obtains u v where "u \<in> A" "v \<in> B" "w = u@v"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    33
using assms by (auto simp: conc_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    34
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    35
lemma conc_mono: "A \<subseteq> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> A @@ B \<subseteq> C @@ D"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    36
by (auto simp: conc_def) 
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    37
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    38
lemma conc_empty[simp]: shows "{} @@ A = {}" and "A @@ {} = {}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    39
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    40
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    41
lemma conc_epsilon[simp]: shows "{[]} @@ A = A" and "A @@ {[]} = A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    42
by (simp_all add:conc_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    43
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    44
lemma conc_assoc: "(A @@ B) @@ C = A @@ (B @@ C)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    45
by (auto elim!: concE) (simp only: append_assoc[symmetric] concI)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    46
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    47
lemma conc_Un_distrib:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    48
shows "A @@ (B \<union> C) = A @@ B \<union> A @@ C"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    49
and   "(A \<union> B) @@ C = A @@ C \<union> B @@ C"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    50
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    51
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    52
lemma conc_UNION_distrib:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    53
shows "A @@ UNION I M = UNION I (%i. A @@ M i)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    54
and   "UNION I M @@ A = UNION I (%i. M i @@ A)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    55
by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    56
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    57
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    58
subsection{* @{term "A ^^ n"} *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    59
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    60
lemma lang_pow_add: "A ^^ (n + m) = A ^^ n @@ A ^^ m"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    61
by (induct n) (auto simp: conc_assoc)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    62
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    63
lemma lang_pow_empty: "{} ^^ n = (if n = 0 then {[]} else {})"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    64
by (induct n) auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    65
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    66
lemma lang_pow_empty_Suc[simp]: "({}::'a lang) ^^ Suc n = {}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    67
by (simp add: lang_pow_empty)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    68
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    69
lemma conc_pow_comm:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    70
  shows "A @@ (A ^^ n) = (A ^^ n) @@ A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
    71
by (induct n) (simp_all add: conc_assoc[symmetric])
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    72
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    73
lemma length_lang_pow_ub:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    74
  "ALL w : A. length w \<le> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<le> k*n"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    75
by(induct n arbitrary: w) (fastsimp simp: conc_def)+
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    76
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    77
lemma length_lang_pow_lb:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    78
  "ALL w : A. length w \<ge> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<ge> k*n"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    79
by(induct n arbitrary: w) (fastsimp simp: conc_def)+
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    80
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    81
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    82
subsection{* @{const star} *}
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    83
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    84
lemma star_if_lang_pow[simp]: "w : A ^^ n \<Longrightarrow> w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    85
by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    86
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    87
lemma Nil_in_star[iff]: "[] : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    88
proof (rule star_if_lang_pow)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    89
  show "[] : A ^^ 0" by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    90
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    91
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    92
lemma star_if_lang[simp]: assumes "w : A" shows "w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    93
proof (rule star_if_lang_pow)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    94
  show "w : A ^^ 1" using `w : A` by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    95
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    96
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    97
lemma append_in_starI[simp]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    98
assumes "u : star A" and "v : star A" shows "u@v : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
    99
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   100
  from `u : star A` obtain m where "u : A ^^ m" by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   101
  moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   102
  from `v : star A` obtain n where "v : A ^^ n" by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   103
  ultimately have "u@v : A ^^ (m+n)" by (simp add: lang_pow_add)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   104
  thus ?thesis by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   105
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   106
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   107
lemma conc_star_star: "star A @@ star A = star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   108
by (auto simp: conc_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   109
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   110
lemma conc_star_comm:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   111
  shows "A @@ star A = star A @@ A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   112
unfolding star_def conc_pow_comm conc_UNION_distrib
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   113
by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   114
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   115
lemma star_induct[consumes 1, case_names Nil append, induct set: star]:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   116
assumes "w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   117
  and "P []"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   118
  and step: "!!u v. u : A \<Longrightarrow> v : star A \<Longrightarrow> P v \<Longrightarrow> P (u@v)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   119
shows "P w"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   120
proof -
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   121
  { fix n have "w : A ^^ n \<Longrightarrow> P w"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   122
    by (induct n arbitrary: w) (auto intro: `P []` step star_if_lang_pow) }
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   123
  with `w : star A` show "P w" by (auto simp: star_def)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   124
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   125
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   126
lemma star_empty[simp]: "star {} = {[]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   127
by (auto elim: star_induct)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   128
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   129
lemma star_epsilon[simp]: "star {[]} = {[]}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   130
by (auto elim: star_induct)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   131
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   132
lemma star_idemp[simp]: "star (star A) = star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   133
by (auto elim: star_induct)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   134
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   135
lemma star_unfold_left: "star A = A @@ star A \<union> {[]}" (is "?L = ?R")
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   136
proof
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   137
  show "?L \<subseteq> ?R" by (rule, erule star_induct) auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   138
qed auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   139
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   140
lemma concat_in_star: "set ws \<subseteq> A \<Longrightarrow> concat ws : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   141
by (induct ws) simp_all
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   142
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   143
lemma in_star_iff_concat:
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   144
  "w : star A = (EX ws. set ws \<subseteq> A & w = concat ws)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   145
  (is "_ = (EX ws. ?R w ws)")
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   146
proof
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   147
  assume "w : star A" thus "EX ws. ?R w ws"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   148
  proof induct
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   149
    case Nil have "?R [] []" by simp
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   150
    thus ?case ..
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   151
  next
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   152
    case (append u v)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   153
    moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   154
    then obtain ws where "set ws \<subseteq> A \<and> v = concat ws" by blast
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   155
    ultimately have "?R (u@v) (u#ws)" by auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   156
    thus ?case ..
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   157
  qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   158
next
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   159
  assume "EX us. ?R w us" thus "w : star A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   160
  by (auto simp: concat_in_star)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   161
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   162
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   163
lemma star_conv_concat: "star A = {concat ws|ws. set ws \<subseteq> A}"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   164
by (fastsimp simp: in_star_iff_concat)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   165
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   166
lemma star_insert_eps[simp]: "star (insert [] A) = star(A)"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   167
proof-
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   168
  { fix us
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   169
    have "set us \<subseteq> insert [] A \<Longrightarrow> EX vs. concat us = concat vs \<and> set vs \<subseteq> A"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   170
      (is "?P \<Longrightarrow> EX vs. ?Q vs")
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   171
    proof
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   172
      let ?vs = "filter (%u. u \<noteq> []) us"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   173
      show "?P \<Longrightarrow> ?Q ?vs" by (induct us) auto
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   174
    qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   175
  } thus ?thesis by (auto simp: star_conv_concat)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   176
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   177
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   178
lemma star_decom: 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   179
  assumes a: "x \<in> star A" "x \<noteq> []"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   180
  shows "\<exists>a b. x = a @ b \<and> a \<noteq> [] \<and> a \<in> A \<and> b \<in> star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   181
using a by (induct rule: star_induct) (blast)+
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   182
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   183
subsection {* Arden's Lemma *}
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   184
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   185
lemma arden_helper:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   186
  assumes eq: "X = A @@ X \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   187
  shows "X = (A ^^ Suc n) @@ X \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   188
proof (induct n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   189
  case 0 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   190
  show "X = (A ^^ Suc 0) @@ X \<union> (\<Union>m\<le>0. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   191
    using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   192
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   193
  case (Suc n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   194
  have ih: "X = (A ^^ Suc n) @@ X \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   195
  also have "\<dots> = (A ^^ Suc n) @@ (A @@ X \<union> B) \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)" using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   196
  also have "\<dots> = (A ^^ Suc (Suc n)) @@ X \<union> ((A ^^ Suc n) @@ B) \<union> (\<Union>m\<le>n. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   197
    by (simp add: conc_Un_distrib conc_assoc[symmetric] conc_pow_comm)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   198
  also have "\<dots> = (A ^^ Suc (Suc n)) @@ X \<union> (\<Union>m\<le>Suc n. (A ^^ m) @@ B)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   199
    by (auto simp add: le_Suc_eq)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   200
  finally show "X = (A ^^ Suc (Suc n)) @@ X \<union> (\<Union>m\<le>Suc n. (A ^^ m) @@ B)" .
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   201
qed
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   202
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   203
lemma Arden:
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   204
  assumes "[] \<notin> A" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   205
  shows "X = A @@ X \<union> B \<longleftrightarrow> X = star A @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   206
proof
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   207
  assume eq: "X = A @@ X \<union> B"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   208
  { fix w assume "w : X"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   209
    let ?n = "size w"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   210
    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   211
      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   212
    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   213
      by (metis length_lang_pow_lb nat_mult_1)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   214
    hence "ALL u : A^^(?n+1)@@X. length u \<ge> ?n+1"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   215
      by(auto simp only: conc_def length_append)
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   216
    hence "w \<notin> A^^(?n+1)@@X" by auto
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   217
    hence "w : star A @@ B" using `w : X` using arden_helper[OF eq, where n="?n"]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   218
      by (auto simp add: star_def conc_UNION_distrib)
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   219
  } moreover
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   220
  { fix w assume "w : star A @@ B"
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   221
    hence "EX n. w : A^^n @@ B" by(auto simp: conc_def star_def)
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   222
    hence "w : X" using arden_helper[OF eq] by blast
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   223
  } ultimately show "X = star A @@ B" by blast 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   224
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   225
  assume eq: "X = star A @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   226
  have "star A = A @@ star A \<union> {[]}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   227
    by (rule star_unfold_left)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   228
  then have "star A @@ B = (A @@ star A \<union> {[]}) @@ B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   229
    by metis
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   230
  also have "\<dots> = (A @@ star A) @@ B \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   231
    unfolding conc_Un_distrib by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   232
  also have "\<dots> = A @@ (star A @@ B) \<union> B" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   233
    by (simp only: conc_assoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   234
  finally show "X = A @@ X \<union> B" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   235
    using eq by blast 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   236
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   237
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   238
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   239
lemma reversed_arden_helper:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   240
  assumes eq: "X = X @@ A \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   241
  shows "X = X @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   242
proof (induct n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   243
  case 0 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   244
  show "X = X @@ (A ^^ Suc 0) \<union> (\<Union>m\<le>0. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   245
    using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   246
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   247
  case (Suc n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   248
  have ih: "X = X @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   249
  also have "\<dots> = (X @@ A \<union> B) @@ (A ^^ Suc n) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))" using eq by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   250
  also have "\<dots> = X @@ (A ^^ Suc (Suc n)) \<union> (B @@ (A ^^ Suc n)) \<union> (\<Union>m\<le>n. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   251
    by (simp add: conc_Un_distrib conc_assoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   252
  also have "\<dots> = X @@ (A ^^ Suc (Suc n)) \<union> (\<Union>m\<le>Suc n. B @@ (A ^^ m))"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   253
    by (auto simp add: le_Suc_eq)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   254
  finally show "X = X @@ (A ^^ Suc (Suc n)) \<union> (\<Union>m\<le>Suc n. B @@ (A ^^ m))" .
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   255
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   256
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   257
theorem reversed_Arden:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   258
  assumes nemp: "[] \<notin> A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   259
  shows "X = X @@ A \<union> B \<longleftrightarrow> X = B @@ star A"
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   260
proof
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   261
 assume eq: "X = X @@ A \<union> B"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   262
  { fix w assume "w : X"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   263
    let ?n = "size w"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   264
    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   265
      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   266
    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   267
      by (metis length_lang_pow_lb nat_mult_1)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   268
    hence "ALL u : X @@ A^^(?n+1). length u \<ge> ?n+1"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   269
      by(auto simp only: conc_def length_append)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   270
    hence "w \<notin> X @@ A^^(?n+1)" by auto
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   271
    hence "w : B @@ star A" using `w : X` using reversed_arden_helper[OF eq, where n="?n"]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   272
      by (auto simp add: star_def conc_UNION_distrib)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   273
  } moreover
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   274
  { fix w assume "w : B @@ star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   275
    hence "EX n. w : B @@ A^^n" by (auto simp: conc_def star_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   276
    hence "w : X" using reversed_arden_helper[OF eq] by blast
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   277
  } ultimately show "X = B @@ star A" by blast 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   278
next 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   279
  assume eq: "X = B @@ star A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   280
  have "star A = {[]} \<union> star A @@ A" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   281
    unfolding conc_star_comm[symmetric]
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   282
    by(metis Un_commute star_unfold_left)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   283
  then have "B @@ star A = B @@ ({[]} \<union> star A @@ A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   284
    by metis
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   285
  also have "\<dots> = B \<union> B @@ (star A @@ A)"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   286
    unfolding conc_Un_distrib by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   287
  also have "\<dots> = B \<union> (B @@ star A) @@ A" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   288
    by (simp only: conc_assoc)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   289
  finally show "X = X @@ A \<union> B" 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 170
diff changeset
   290
    using eq by blast 
170
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   291
qed
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   292
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   293
b1258b7d2789 made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
urbanc
parents:
diff changeset
   294
end