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