233
+ − 1
theory Closures2
221
+ − 2
imports
+ − 3
Closures
222
+ − 4
Higman
221
+ − 5
(* "~~/src/HOL/Proofs/Extraction/Higman" *)
220
+ − 6
begin
+ − 7
223
+ − 8
section {* Closure under @{text SUBSEQ} and @{text SUPSEQ} *}
+ − 9
+ − 10
(* compatibility with Higman theory by Stefan *)
222
+ − 11
notation
+ − 12
emb ("_ \<preceq> _")
+ − 13
+ − 14
declare emb0 [intro]
+ − 15
declare emb1 [intro]
+ − 16
declare emb2 [intro]
+ − 17
+ − 18
lemma letter_UNIV:
+ − 19
shows "UNIV = {A, B::letter}"
+ − 20
apply(auto)
+ − 21
apply(case_tac x)
+ − 22
apply(auto)
+ − 23
done
+ − 24
+ − 25
instance letter :: finite
+ − 26
apply(default)
+ − 27
apply(simp add: letter_UNIV)
+ − 28
done
+ − 29
+ − 30
hide_const A
+ − 31
hide_const B
233
+ − 32
hide_const R
222
+ − 33
+ − 34
(*
221
+ − 35
inductive
+ − 36
emb :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<preceq> _")
220
+ − 37
where
221
+ − 38
emb0 [intro]: "emb [] y"
+ − 39
| emb1 [intro]: "emb x y \<Longrightarrow> emb x (c # y)"
+ − 40
| emb2 [intro]: "emb x y \<Longrightarrow> emb (c # x) (c # y)"
222
+ − 41
*)
220
+ − 42
221
+ − 43
lemma emb_refl [intro]:
220
+ − 44
shows "x \<preceq> x"
222
+ − 45
by (induct x) (auto)
220
+ − 46
221
+ − 47
lemma emb_right [intro]:
220
+ − 48
assumes a: "x \<preceq> y"
+ − 49
shows "x \<preceq> y @ y'"
222
+ − 50
using a
+ − 51
by (induct arbitrary: y') (auto)
220
+ − 52
221
+ − 53
lemma emb_left [intro]:
220
+ − 54
assumes a: "x \<preceq> y"
+ − 55
shows "x \<preceq> y' @ y"
221
+ − 56
using a by (induct y') (auto)
220
+ − 57
221
+ − 58
lemma emb_appendI [intro]:
220
+ − 59
assumes a: "x \<preceq> x'"
+ − 60
and b: "y \<preceq> y'"
+ − 61
shows "x @ y \<preceq> x' @ y'"
221
+ − 62
using a b by (induct) (auto)
+ − 63
+ − 64
lemma emb_cons_leftD:
+ − 65
assumes "a # x \<preceq> y"
+ − 66
shows "\<exists>y1 y2. y = y1 @ [a] @ y2 \<and> x \<preceq> y2"
+ − 67
using assms
+ − 68
apply(induct x\<equiv>"a # x" y\<equiv>"y" arbitrary: a x y)
+ − 69
apply(auto)
+ − 70
apply(metis append_Cons)
+ − 71
done
+ − 72
+ − 73
lemma emb_append_leftD:
+ − 74
assumes "x1 @ x2 \<preceq> y"
+ − 75
shows "\<exists>y1 y2. y = y1 @ y2 \<and> x1 \<preceq> y1 \<and> x2 \<preceq> y2"
+ − 76
using assms
+ − 77
apply(induct x1 arbitrary: x2 y)
+ − 78
apply(auto)
+ − 79
apply(drule emb_cons_leftD)
+ − 80
apply(auto)
+ − 81
apply(drule_tac x="x2" in meta_spec)
+ − 82
apply(drule_tac x="y2" in meta_spec)
+ − 83
apply(auto)
+ − 84
apply(rule_tac x="y1 @ (a # y1a)" in exI)
+ − 85
apply(rule_tac x="y2a" in exI)
+ − 86
apply(auto)
220
+ − 87
done
+ − 88
221
+ − 89
lemma emb_trans:
+ − 90
assumes a: "x1 \<preceq> x2"
+ − 91
and b: "x2 \<preceq> x3"
+ − 92
shows "x1 \<preceq> x3"
+ − 93
using a b
+ − 94
apply(induct arbitrary: x3)
+ − 95
apply(metis emb0)
+ − 96
apply(metis emb_cons_leftD emb_left)
+ − 97
apply(drule_tac emb_cons_leftD)
+ − 98
apply(auto)
220
+ − 99
done
+ − 100
221
+ − 101
lemma emb_strict_length:
+ − 102
assumes a: "x \<preceq> y" "x \<noteq> y"
+ − 103
shows "length x < length y"
+ − 104
using a
+ − 105
by (induct) (auto simp add: less_Suc_eq)
+ − 106
+ − 107
lemma emb_antisym:
+ − 108
assumes a: "x \<preceq> y" "y \<preceq> x"
+ − 109
shows "x = y"
+ − 110
using a emb_strict_length
+ − 111
by (metis not_less_iff_gr_or_eq)
+ − 112
+ − 113
lemma emb_wf:
+ − 114
shows "wf {(x, y). x \<preceq> y \<and> x \<noteq> y}"
+ − 115
proof -
+ − 116
have "wf (measure length)" by simp
+ − 117
moreover
+ − 118
have "{(x, y). x \<preceq> y \<and> x \<noteq> y} \<subseteq> measure length"
+ − 119
unfolding measure_def by (auto simp add: emb_strict_length)
+ − 120
ultimately
+ − 121
show "wf {(x, y). x \<preceq> y \<and> x \<noteq> y}" by (rule wf_subset)
+ − 122
qed
+ − 123
+ − 124
subsection {* Sub- and Supsequences *}
220
+ − 125
+ − 126
definition
+ − 127
"SUBSEQ A \<equiv> {x. \<exists>y \<in> A. x \<preceq> y}"
+ − 128
+ − 129
definition
221
+ − 130
"SUPSEQ A \<equiv> {x. \<exists>y \<in> A. y \<preceq> x}"
220
+ − 131
221
+ − 132
lemma SUPSEQ_simps [simp]:
+ − 133
shows "SUPSEQ {} = {}"
+ − 134
and "SUPSEQ {[]} = UNIV"
+ − 135
unfolding SUPSEQ_def by auto
220
+ − 136
221
+ − 137
lemma SUPSEQ_atom [simp]:
223
+ − 138
shows "SUPSEQ {[c]} = UNIV \<cdot> {[c]} \<cdot> UNIV"
221
+ − 139
unfolding SUPSEQ_def conc_def
+ − 140
by (auto) (metis append_Cons emb_cons_leftD)
220
+ − 141
221
+ − 142
lemma SUPSEQ_union [simp]:
+ − 143
shows "SUPSEQ (A \<union> B) = SUPSEQ A \<union> SUPSEQ B"
+ − 144
unfolding SUPSEQ_def by auto
220
+ − 145
221
+ − 146
lemma SUPSEQ_conc [simp]:
+ − 147
shows "SUPSEQ (A \<cdot> B) = SUPSEQ A \<cdot> SUPSEQ B"
+ − 148
unfolding SUPSEQ_def conc_def
+ − 149
by (auto) (metis emb_append_leftD)
220
+ − 150
221
+ − 151
lemma SUPSEQ_star [simp]:
+ − 152
shows "SUPSEQ (A\<star>) = UNIV"
220
+ − 153
apply(subst star_unfold_left)
221
+ − 154
apply(simp only: SUPSEQ_union)
220
+ − 155
apply(simp)
+ − 156
done
+ − 157
+ − 158
definition
+ − 159
Allreg :: "'a::finite rexp"
+ − 160
where
+ − 161
"Allreg \<equiv> \<Uplus>(Atom ` UNIV)"
+ − 162
+ − 163
lemma Allreg_lang [simp]:
221
+ − 164
shows "lang Allreg = (\<Union>a. {[a]})"
+ − 165
unfolding Allreg_def by auto
220
+ − 166
+ − 167
lemma [simp]:
221
+ − 168
shows "(\<Union>a. {[a]})\<star> = UNIV"
220
+ − 169
apply(auto)
+ − 170
apply(induct_tac x rule: list.induct)
+ − 171
apply(auto)
+ − 172
apply(subgoal_tac "[a] @ list \<in> (\<Union>a. {[a]})\<star>")
+ − 173
apply(simp)
+ − 174
apply(rule append_in_starI)
+ − 175
apply(auto)
+ − 176
done
+ − 177
+ − 178
lemma Star_Allreg_lang [simp]:
221
+ − 179
shows "lang (Star Allreg) = UNIV"
+ − 180
by simp
220
+ − 181
221
+ − 182
fun
+ − 183
UP :: "'a::finite rexp \<Rightarrow> 'a rexp"
+ − 184
where
+ − 185
"UP (Zero) = Zero"
+ − 186
| "UP (One) = Star Allreg"
+ − 187
| "UP (Atom c) = Times (Star Allreg) (Times (Atom c) (Star Allreg))"
+ − 188
| "UP (Plus r1 r2) = Plus (UP r1) (UP r2)"
+ − 189
| "UP (Times r1 r2) = Times (UP r1) (UP r2)"
+ − 190
| "UP (Star r) = Star Allreg"
220
+ − 191
221
+ − 192
lemma lang_UP:
222
+ − 193
fixes r::"letter rexp"
221
+ − 194
shows "lang (UP r) = SUPSEQ (lang r)"
+ − 195
by (induct r) (simp_all)
+ − 196
+ − 197
lemma regular_SUPSEQ:
222
+ − 198
fixes A::"letter lang"
220
+ − 199
assumes "regular A"
+ − 200
shows "regular (SUPSEQ A)"
+ − 201
proof -
222
+ − 202
from assms obtain r::"letter rexp" where "lang r = A" by auto
221
+ − 203
then have "lang (UP r) = SUPSEQ A" by (simp add: lang_UP)
+ − 204
then show "regular (SUPSEQ A)" by auto
220
+ − 205
qed
221
+ − 206
+ − 207
lemma SUPSEQ_subset:
+ − 208
shows "A \<subseteq> SUPSEQ A"
+ − 209
unfolding SUPSEQ_def by auto
+ − 210
223
+ − 211
lemma SUBSEQ_complement:
+ − 212
shows "- (SUBSEQ A) = SUPSEQ (- (SUBSEQ A))"
+ − 213
proof -
+ − 214
have "- (SUBSEQ A) \<subseteq> SUPSEQ (- (SUBSEQ A))"
+ − 215
by (rule SUPSEQ_subset)
+ − 216
moreover
+ − 217
have "SUPSEQ (- (SUBSEQ A)) \<subseteq> - (SUBSEQ A)"
+ − 218
proof (rule ccontr)
+ − 219
assume "\<not> (SUPSEQ (- (SUBSEQ A)) \<subseteq> - (SUBSEQ A))"
+ − 220
then obtain x where
+ − 221
a: "x \<in> SUPSEQ (- (SUBSEQ A))" and
+ − 222
b: "x \<notin> - (SUBSEQ A)" by auto
+ − 223
+ − 224
from a obtain y where c: "y \<in> - (SUBSEQ A)" and d: "y \<preceq> x"
+ − 225
by (auto simp add: SUPSEQ_def)
221
+ − 226
223
+ − 227
from b have "x \<in> SUBSEQ A" by simp
+ − 228
then obtain x' where f: "x' \<in> A" and e: "x \<preceq> x'"
+ − 229
by (auto simp add: SUBSEQ_def)
+ − 230
+ − 231
from d e have "y \<preceq> x'" by (rule emb_trans)
+ − 232
then have "y \<in> SUBSEQ A" using f
+ − 233
by (auto simp add: SUBSEQ_def)
+ − 234
with c show "False" by simp
+ − 235
qed
+ − 236
ultimately show "- (SUBSEQ A) = SUPSEQ (- (SUBSEQ A))" by simp
+ − 237
qed
+ − 238
+ − 239
lemma Higman_antichains:
+ − 240
assumes a: "\<forall>x \<in> A. \<forall>y \<in> A. x \<noteq> y \<longrightarrow> \<not>(x \<preceq> y) \<and> \<not>(y \<preceq> x)"
+ − 241
shows "finite A"
+ − 242
proof (rule ccontr)
+ − 243
assume "infinite A"
+ − 244
then obtain f::"nat \<Rightarrow> letter list" where b: "inj f" and c: "range f \<subseteq> A"
+ − 245
by (auto simp add: infinite_iff_countable_subset)
+ − 246
from higman_idx obtain i j where d: "i < j" and e: "f i \<preceq> f j" by blast
+ − 247
have "f i \<noteq> f j" using b d by (auto simp add: inj_on_def)
+ − 248
moreover
231
+ − 249
have "f i \<in> A" using c by auto
223
+ − 250
moreover
231
+ − 251
have "f j \<in> A" using c by auto
223
+ − 252
ultimately have "\<not>(f i \<preceq> f j)" using a by simp
+ − 253
with e show "False" by simp
+ − 254
qed
221
+ − 255
+ − 256
definition
223
+ − 257
minimal :: "letter list \<Rightarrow> letter lang \<Rightarrow> bool"
222
+ − 258
where
233
+ − 259
"minimal x A \<equiv> (\<forall>y \<in> A. y \<preceq> x \<longrightarrow> x \<preceq> y)"
220
+ − 260
221
+ − 261
lemma main_lemma:
233
+ − 262
shows "\<exists>M. finite M \<and> SUPSEQ A = SUPSEQ M"
221
+ − 263
proof -
223
+ − 264
def M \<equiv> "{x \<in> A. minimal x A}"
221
+ − 265
have "finite M"
223
+ − 266
unfolding M_def minimal_def
+ − 267
by (rule Higman_antichains) (auto simp add: emb_antisym)
221
+ − 268
moreover
+ − 269
have "SUPSEQ A \<subseteq> SUPSEQ M"
223
+ − 270
proof
+ − 271
fix x
+ − 272
assume "x \<in> SUPSEQ A"
+ − 273
then obtain y where "y \<in> A" and "y \<preceq> x" by (auto simp add: SUPSEQ_def)
+ − 274
then have a: "y \<in> {y' \<in> A. y' \<preceq> x}" by simp
+ − 275
obtain z where b: "z \<in> A" "z \<preceq> x" and c: "\<forall>y. y \<preceq> z \<and> y \<noteq> z \<longrightarrow> y \<notin> {y' \<in> A. y' \<preceq> x}"
+ − 276
using wfE_min[OF emb_wf a] by auto
+ − 277
then have "z \<in> M"
+ − 278
unfolding M_def minimal_def
+ − 279
by (auto intro: emb_trans)
233
+ − 280
with b(2) show "x \<in> SUPSEQ M"
223
+ − 281
by (auto simp add: SUPSEQ_def)
+ − 282
qed
221
+ − 283
moreover
+ − 284
have "SUPSEQ M \<subseteq> SUPSEQ A"
+ − 285
by (auto simp add: SUPSEQ_def M_def)
+ − 286
ultimately
233
+ − 287
show "\<exists>M. finite M \<and> SUPSEQ A = SUPSEQ M" by blast
221
+ − 288
qed
+ − 289
+ − 290
lemma closure_SUPSEQ:
222
+ − 291
fixes A::"letter lang"
221
+ − 292
shows "regular (SUPSEQ A)"
+ − 293
proof -
+ − 294
obtain M where a: "finite M" and b: "SUPSEQ A = SUPSEQ M"
+ − 295
using main_lemma by blast
+ − 296
have "regular M" using a by (rule finite_regular)
+ − 297
then have "regular (SUPSEQ M)" by (rule regular_SUPSEQ)
+ − 298
then show "regular (SUPSEQ A)" using b by simp
+ − 299
qed
+ − 300
+ − 301
lemma closure_SUBSEQ:
222
+ − 302
fixes A::"letter lang"
221
+ − 303
shows "regular (SUBSEQ A)"
+ − 304
proof -
+ − 305
have "regular (SUPSEQ (- SUBSEQ A))" by (rule closure_SUPSEQ)
223
+ − 306
then have "regular (- SUBSEQ A)" by (subst SUBSEQ_complement) (simp)
221
+ − 307
then have "regular (- (- (SUBSEQ A)))" by (rule closure_complement)
+ − 308
then show "regular (SUBSEQ A)" by simp
+ − 309
qed
+ − 310
220
+ − 311
end